From 02128fe4e46d10e3d15574bc6aa38baab14384be Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 23 Aug 2022 14:57:14 +0000 Subject: [PATCH 001/828] feat(topology/algebra/module/character_space): add facts about `character_space.union_zero` (#16209) This adds that the `character_space` along with the zero map is a closed subspace of the `weak_dual`. The point of this is eventually to show that in the non-unital case, the `character_space` is a locally compact Hausdorff space, under appropriate assumptions. --- .../algebra/module/character_space.lean | 28 ++++++++++++++----- 1 file changed, 21 insertions(+), 7 deletions(-) diff --git a/src/topology/algebra/module/character_space.lean b/src/topology/algebra/module/character_space.lean index ff31ed7f83132..825ecb535bea4 100644 --- a/src/topology/algebra/module/character_space.lean +++ b/src/topology/algebra/module/character_space.lean @@ -84,6 +84,23 @@ def to_non_unital_alg_hom (φ : character_space 𝕜 A) : A →ₙₐ[𝕜] 𝕜 @[simp] lemma coe_to_non_unital_alg_hom (φ : character_space 𝕜 A) : ⇑(to_non_unital_alg_hom φ) = φ := rfl +variables (𝕜 A) + +lemma union_zero : + character_space 𝕜 A ∪ {0} = {φ : weak_dual 𝕜 A | ∀ (x y : A), φ (x * y) = (φ x) * (φ y)} := +le_antisymm + (by { rintros φ (hφ | h₀), { exact hφ.2 }, { exact λ x y, by simp [set.eq_of_mem_singleton h₀] }}) + (λ φ hφ, or.elim (em $ φ = 0) (λ h₀, or.inr h₀) (λ h₀, or.inl ⟨h₀, hφ⟩)) + +/-- The `character_space 𝕜 A` along with `0` is always a closed set in `weak_dual 𝕜 A`. -/ +lemma union_zero_is_closed [t2_space 𝕜] [has_continuous_mul 𝕜] : + is_closed (character_space 𝕜 A ∪ {0}) := +begin + simp only [union_zero, set.set_of_forall], + exact is_closed_Inter (λ x, is_closed_Inter $ λ y, is_closed_eq (eval_continuous _) $ + (eval_continuous _).mul (eval_continuous _)) +end + end non_unital_non_assoc_semiring section unital @@ -125,17 +142,14 @@ begin simpa using h.1, end +/-- under suitable mild assumptions on `𝕜`, the character space is a closed set in +`weak_dual 𝕜 A`. -/ lemma is_closed [nontrivial 𝕜] [t2_space 𝕜] [has_continuous_mul 𝕜] : is_closed (character_space 𝕜 A) := begin - rw [eq_set_map_one_map_mul], + rw [eq_set_map_one_map_mul, set.set_of_and], refine is_closed.inter (is_closed_eq (eval_continuous _) continuous_const) _, - change is_closed {φ : weak_dual 𝕜 A | ∀ x y : A, φ (x * y) = φ x * φ y}, - rw [set.set_of_forall], - refine is_closed_Inter (λ a, _), - rw [set.set_of_forall], - exact is_closed_Inter (λ _, is_closed_eq (eval_continuous _) - ((eval_continuous _).mul (eval_continuous _))) + simpa only [(union_zero 𝕜 A).symm] using union_zero_is_closed _ _, end end unital From e919421b24b74d339627c9dc68b579b36ebf6f0d Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Tue, 23 Aug 2022 15:50:52 +0000 Subject: [PATCH 002/828] feat(category_theory): functors preserving epis and kernels are exact (#16129) --- src/category_theory/abelian/exact.lean | 41 +++++++++++++++++++++++--- 1 file changed, 37 insertions(+), 4 deletions(-) diff --git a/src/category_theory/abelian/exact.lean b/src/category_theory/abelian/exact.lean index 2d3e9b1c1d974..79419368d87ef 100644 --- a/src/category_theory/abelian/exact.lean +++ b/src/category_theory/abelian/exact.lean @@ -125,7 +125,7 @@ begin end /-- The dual result is true even in non-abelian categories, see - `category_theory.exact_epi_comp_iff`. -/ + `category_theory.exact_comp_mono_iff`. -/ lemma exact_epi_comp_iff {W : C} (h : W ⟶ X) [epi h] : exact (h ≫ f) g ↔ exact f g := begin refine ⟨λ hfg, _, λ h, exact_epi_comp h⟩, @@ -135,7 +135,7 @@ begin exact zero_of_epi_comp h (by rw [← hfg.1, category.assoc]) end -/-- If `(f, g)` is exact, then `images.image.ι f` is a kernel of `g`. -/ +/-- If `(f, g)` is exact, then `abelian.image.ι f` is a kernel of `g`. -/ def is_limit_image (h : exact f g) : is_limit (kernel_fork.of_ι (abelian.image.ι f) (image_ι_comp_eq_zero h.1) : kernel_fork g) := @@ -227,6 +227,12 @@ begin rw [← this, category.assoc, cokernel.condition, comp_zero] end +lemma exact_iff_exact_image_ι : exact f g ↔ exact (abelian.image.ι f) g := +by conv_lhs { rw ← abelian.image.fac f }; apply exact_epi_comp_iff + +lemma exact_iff_exact_coimage_π : exact f g ↔ exact f (coimage.π g) := +by conv_lhs { rw ← abelian.coimage.fac g}; apply exact_comp_mono_iff + section variables (Z) @@ -419,7 +425,7 @@ def preserves_cokernels_of_map_exact (X Y : A) (f : X ⟶ Y) : /-- A functor which preserves exactness is left exact, i.e. preserves finite limits. This is part of the inverse implication to `functor.map_exact`. -/ -def preserves_finite_limits_of_map_exact : limits.preserves_finite_limits L := +def preserves_finite_limits_of_map_exact : preserves_finite_limits L := begin letI := preserves_zero_morphisms_of_map_exact L h, letI := preserves_kernels_of_map_exact L h, @@ -428,7 +434,7 @@ end /-- A functor which preserves exactness is right exact, i.e. preserves finite colimits. This is part of the inverse implication to `functor.map_exact`. -/ -def preserves_finite_colimits_of_map_exact : limits.preserves_finite_colimits L := +def preserves_finite_colimits_of_map_exact : preserves_finite_colimits L := begin letI := preserves_zero_morphisms_of_map_exact L h, letI := preserves_cokernels_of_map_exact L h, @@ -437,6 +443,33 @@ end end +section + +/-- A functor preserving zero morphisms, monos, and cokernels preserves finite limits. -/ +def preserves_finite_limits_of_preserves_monos_and_cokernels + [preserves_zero_morphisms L] [preserves_monomorphisms L] + [∀ {X Y} (f : X ⟶ Y), preserves_colimit (parallel_pair f 0) L] : preserves_finite_limits L := +begin + apply preserves_finite_limits_of_map_exact, + intros X Y Z f g h, + rw [← abelian.coimage.fac g, L.map_comp, exact_comp_mono_iff], + exact exact_of_is_cokernel _ _ _ + (is_colimit_cofork_map_of_is_colimit' L _ (is_colimit_coimage f g h)) +end + +/-- A functor preserving zero morphisms, epis, and kernels preserves finite colimits. -/ +def preserves_finite_colimits_of_preserves_epis_and_kernels + [preserves_zero_morphisms L] [preserves_epimorphisms L] + [∀ {X Y} (f : X ⟶ Y), preserves_limit (parallel_pair f 0) L] : preserves_finite_colimits L := +begin + apply preserves_finite_colimits_of_map_exact, + intros X Y Z f g h, + rw [← abelian.image.fac f, L.map_comp, exact_epi_comp_iff], + exact exact_of_is_kernel _ _ _ (is_limit_fork_map_of_is_limit' L _ (is_limit_image f g h)) +end + +end + end functor end category_theory From cf9eb049e32484dc7d35f8f17da6021fd0f47135 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Tue, 23 Aug 2022 16:58:32 +0000 Subject: [PATCH 003/828] feat(topology/metric_space): Add first countability instances (#15942) This PR proves the following two facts: - pseudo-metric/metrizable spaces are first countable - uniform groups that have a countably generated neighborhood at the origin have a countably generated uniformity The instance for uniform groups allows for an easy way to prove that a topological group is metrizable, since the neighborhood filter is a more common object than the uniformity. In total this PR gives the proof that a topological group is metrizable iff it the neighborhood filter at the origin is countably generated. Co-authored-by: Moritz Doll --- src/topology/algebra/uniform_group.lean | 7 +++++++ src/topology/metric_space/closeds.lean | 2 +- src/topology/metric_space/metrizable.lean | 11 +++++++++++ 3 files changed, 19 insertions(+), 1 deletion(-) diff --git a/src/topology/algebra/uniform_group.lean b/src/topology/algebra/uniform_group.lean index 98090af313e52..53c720bdbcd1a 100644 --- a/src/topology/algebra/uniform_group.lean +++ b/src/topology/algebra/uniform_group.lean @@ -205,6 +205,13 @@ end 𝓤 α = comap (λx:α×α, x.1 / x.2) (𝓝 (1:α)) := by { rw [← comap_swap_uniformity, uniformity_eq_comap_nhds_one, comap_comap, (∘)], refl } +variables {α} + +@[to_additive] theorem uniform_group.uniformity_countably_generated + [(𝓝 (1 : α)).is_countably_generated] : + (𝓤 α).is_countably_generated := +by { rw uniformity_eq_comap_nhds_one, exact filter.comap.is_countably_generated _ _ } + open mul_opposite @[to_additive] diff --git a/src/topology/metric_space/closeds.lean b/src/topology/metric_space/closeds.lean index bad5ff62afd01..2655f95e0c8a8 100644 --- a/src/topology/metric_space/closeds.lean +++ b/src/topology/metric_space/closeds.lean @@ -377,7 +377,7 @@ begin -- we have proved that `d` is a good approximation of `t` as requested exact ⟨d, ‹d ∈ v›, Dtc⟩ }, end, - apply uniform_space.second_countable_of_separable, + exact uniform_space.second_countable_of_separable (nonempty_compacts α), end end --section diff --git a/src/topology/metric_space/metrizable.lean b/src/topology/metric_space/metrizable.lean index ea81fc5b2eb16..d158d8e7691aa 100644 --- a/src/topology/metric_space/metrizable.lean +++ b/src/topology/metric_space/metrizable.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import topology.urysohns_lemma import topology.continuous_function.bounded +import topology.uniform_space.cauchy /-! # Metrizability of a T₃ topological space with second countable topology @@ -62,6 +63,16 @@ begin exact ⟨⟨hf.comap_pseudo_metric_space, rfl⟩⟩ end +/-- Every pseudo-metrizable space is first countable. -/ +@[priority 100] +instance pseudo_metrizable_space.first_countable_topology [h : pseudo_metrizable_space X] : + topological_space.first_countable_topology X := +begin + unfreezingI { rcases h with ⟨_, hm⟩, rw ←hm }, + exact @uniform_space.first_countable_topology X pseudo_metric_space.to_uniform_space + emetric.uniformity.filter.is_countably_generated, +end + instance pseudo_metrizable_space.subtype [pseudo_metrizable_space X] (s : set X) : pseudo_metrizable_space s := inducing_coe.pseudo_metrizable_space From 6f1ad8258079beb8437c2e402d4218f4e6b7f9f6 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 23 Aug 2022 16:58:33 +0000 Subject: [PATCH 004/828] chore(algebra/algebra/spectrum): generalize some spectrum results from `alg_hom` to `alg_hom_class` (#16177) - [x] depends on: #16178 --- src/algebra/algebra/spectrum.lean | 25 ++++++++++--------- src/analysis/normed_space/spectrum.lean | 21 ++++++++++------ .../algebra/module/character_space.lean | 2 +- 3 files changed, 28 insertions(+), 20 deletions(-) diff --git a/src/algebra/algebra/spectrum.lean b/src/algebra/algebra/spectrum.lean index cb2b97534386c..2a005b72ac92f 100644 --- a/src/algebra/algebra/spectrum.lean +++ b/src/algebra/algebra/spectrum.lean @@ -437,32 +437,33 @@ namespace alg_hom section comm_semiring -variables {R : Type*} {A B : Type*} [comm_ring R] [ring A] [algebra R A] [ring B] [algebra R B] +variables {F R A B : Type*} [comm_ring R] [ring A] [algebra R A] [ring B] [algebra R B] +variables [alg_hom_class F R A B] local notation `σ` := spectrum R local notation `↑ₐ` := algebra_map R A -lemma mem_resolvent_set_apply (φ : A →ₐ[R] B) {a : A} {r : R} (h : r ∈ resolvent_set R a) : - r ∈ resolvent_set R (φ a) := -by simpa only [map_sub, commutes] using h.map φ +lemma mem_resolvent_set_apply (φ : F) {a : A} {r : R} (h : r ∈ resolvent_set R a) : + r ∈ resolvent_set R ((φ : A → B) a) := +by simpa only [map_sub, alg_hom_class.commutes] using h.map φ -lemma spectrum_apply_subset (φ : A →ₐ[R] B) (a : A) : σ (φ a) ⊆ σ a := +lemma spectrum_apply_subset (φ : F) (a : A) : σ ((φ : A → B) a) ⊆ σ a := λ _, mt (mem_resolvent_set_apply φ) end comm_semiring section comm_ring -variables {R : Type*} {A B : Type*} [comm_ring R] [ring A] [algebra R A] [ring B] [algebra R B] +variables {F R A B : Type*} [comm_ring R] [ring A] [algebra R A] [ring B] [algebra R B] +variables [alg_hom_class F R A R] local notation `σ` := spectrum R local notation `↑ₐ` := algebra_map R A -lemma apply_mem_spectrum [nontrivial R] (φ : A →ₐ[R] R) (a : A) : φ a ∈ σ a := +lemma apply_mem_spectrum [nontrivial R] (φ : F) (a : A) : φ a ∈ σ a := begin - have h : ↑ₐ(φ a) - a ∈ φ.to_ring_hom.ker, - { simp only [ring_hom.mem_ker, coe_to_ring_hom, commutes, algebra.id.map_eq_id, - to_ring_hom_eq_coe, ring_hom.id_apply, sub_self, map_sub] }, - simp only [spectrum.mem_iff, ←mem_nonunits_iff, - coe_subset_nonunits (φ.to_ring_hom.ker_ne_top) h], + have h : ↑ₐ(φ a) - a ∈ (φ : A →+* R).ker, + { simp only [ring_hom.mem_ker, map_sub, ring_hom.coe_coe, alg_hom_class.commutes, + algebra.id.map_eq_id, ring_hom.id_apply, sub_self], }, + simp only [spectrum.mem_iff, ←mem_nonunits_iff, coe_subset_nonunits ((φ : A →+* R).ker_ne_top) h], end end comm_ring diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 18f9ec5c54eaf..55fb086b524ed 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -413,14 +413,21 @@ section normed_field variables [normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] [complete_space A] local notation `↑ₐ` := algebra_map 𝕜 A + +/-- An algebra homomorphism into the base field, as a continuous linear map (since it is +automatically bounded). -/ +instance [norm_one_class A] : continuous_linear_map_class (A →ₐ[𝕜] 𝕜) 𝕜 A 𝕜 := +{ map_continuous := λ φ, add_monoid_hom_class.continuous_of_bound φ 1 $ + λ a, (one_mul ∥a∥).symm ▸ spectrum.norm_le_norm_of_mem (apply_mem_spectrum φ _), + .. alg_hom_class.linear_map_class } + /-- An algebra homomorphism into the base field, as a continuous linear map (since it is automatically bounded). -/ -@[simps] def to_continuous_linear_map [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : A →L[𝕜] 𝕜 := -φ.to_linear_map.mk_continuous_of_exists_bound $ - ⟨1, λ a, (one_mul ∥a∥).symm ▸ spectrum.norm_le_norm_of_mem (φ.apply_mem_spectrum _)⟩ +def to_continuous_linear_map [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : A →L[𝕜] 𝕜 := +{ cont := map_continuous φ, .. φ.to_linear_map } -lemma continuous [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : continuous φ := -φ.to_continuous_linear_map.continuous +@[simp] lemma coe_to_continuous_linear_map [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : + ⇑φ.to_continuous_linear_map = φ := rfl end normed_field @@ -431,8 +438,8 @@ local notation `↑ₐ` := algebra_map 𝕜 A @[simp] lemma to_continuous_linear_map_norm [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : ∥φ.to_continuous_linear_map∥ = 1 := continuous_linear_map.op_norm_eq_of_bounds zero_le_one - (λ a, (one_mul ∥a∥).symm ▸ spectrum.norm_le_norm_of_mem (φ.apply_mem_spectrum _)) - (λ _ _ h, by simpa only [to_continuous_linear_map_apply, mul_one, map_one, norm_one] using h 1) + (λ a, (one_mul ∥a∥).symm ▸ spectrum.norm_le_norm_of_mem (apply_mem_spectrum φ _)) + (λ _ _ h, by simpa only [coe_to_continuous_linear_map, map_one, norm_one, mul_one] using h 1) end nontrivially_normed_field diff --git a/src/topology/algebra/module/character_space.lean b/src/topology/algebra/module/character_space.lean index 825ecb535bea4..dfce8973e4cbc 100644 --- a/src/topology/algebra/module/character_space.lean +++ b/src/topology/algebra/module/character_space.lean @@ -160,7 +160,7 @@ variables [comm_ring 𝕜] [no_zero_divisors 𝕜] [topological_space 𝕜] [has [has_continuous_const_smul 𝕜 𝕜] [topological_space A] [ring A] [algebra 𝕜 A] lemma apply_mem_spectrum [nontrivial 𝕜] (φ : character_space 𝕜 A) (a : A) : φ a ∈ spectrum 𝕜 a := -(to_alg_hom φ).apply_mem_spectrum a +alg_hom.apply_mem_spectrum φ a end ring From 04763f325311ce56b5f69094961b9109200e73ea Mon Sep 17 00:00:00 2001 From: tb65536 Date: Wed, 24 Aug 2022 03:40:54 +0000 Subject: [PATCH 005/828] feat(group_theory/index): `card_dvd_of_surjective` (#16133) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Mathlib already has `card_dvd_of_injective : fintype.card G ∣ fintype.card H` This PR adds: `nat_card_dvd_of_injective : nat.card G ∣ nat.card H` `nat_card_dvd_of_surjective : nat.card H ∣ nat.card G` `card_dvd_of_surjective : fintype.card H ∣ fintype.card G` --- src/group_theory/index.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/src/group_theory/index.lean b/src/group_theory/index.lean index 6d1281e746d6c..52bcb63b48b14 100644 --- a/src/group_theory/index.lean +++ b/src/group_theory/index.lean @@ -166,6 +166,24 @@ by rw [relindex, subgroup_of_self, index_top] lemma card_mul_index : nat.card H * H.index = nat.card G := by { rw [←relindex_bot_left, ←index_bot], exact relindex_mul_index bot_le } +@[to_additive] lemma nat_card_dvd_of_injective {G H : Type*} [group G] [group H] (f : G →* H) + (hf : function.injective f) : nat.card G ∣ nat.card H := +begin + rw nat.card_congr (monoid_hom.of_injective hf).to_equiv, + exact dvd.intro f.range.index f.range.card_mul_index, +end + +@[to_additive] lemma nat_card_dvd_of_surjective {G H : Type*} [group G] [group H] (f : G →* H) + (hf : function.surjective f) : nat.card H ∣ nat.card G := +begin + rw ← nat.card_congr (quotient_group.quotient_ker_equiv_of_surjective f hf).to_equiv, + exact dvd.intro_left (nat.card f.ker) f.ker.card_mul_index, +end + +@[to_additive] lemma card_dvd_of_surjective {G H : Type*} [group G] [group H] [fintype G] + [fintype H] (f : G →* H) (hf : function.surjective f) : fintype.card H ∣ fintype.card G := +by simp only [←nat.card_eq_fintype_card, nat_card_dvd_of_surjective f hf] + @[to_additive] lemma index_map {G' : Type*} [group G'] (f : G →* G') : (H.map f).index = (H ⊔ f.ker).index * f.range.index := by rw [←comap_map_eq, index_comap, relindex_mul_index (H.map_le_range f)] From 757e73c5f0411c82c8f175d7e1c7d5388a4e67ce Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 24 Aug 2022 05:21:04 +0000 Subject: [PATCH 006/828] feat(data/seq/seq): prove `seq.ext` earlier (#15830) We heavily golf `seq.ext` and move it to almost the beginning of the file. Doing this breaks the proof of `seq.of_list_cons`, which we also change and golf by adding a few trivial `simp` lemmas (not all of them are needed, but might as well add them). --- src/data/seq/seq.lean | 64 ++++++++++++++++++------------------------- 1 file changed, 26 insertions(+), 38 deletions(-) diff --git a/src/data/seq/seq.lean b/src/data/seq/seq.lean index 17c086e5bcc02..36f0179570d69 100644 --- a/src/data/seq/seq.lean +++ b/src/data/seq/seq.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import data.list.basic import data.lazy_list import data.nat.basic import data.stream.init @@ -49,6 +50,15 @@ end⟩ /-- Get the nth element of a sequence (if it exists) -/ def nth : seq α → ℕ → option α := subtype.val +@[simp] theorem nth_mk (f hf) : @nth α ⟨f, hf⟩ = f := rfl + +@[simp] theorem nth_nil (n : ℕ) : (@nil α).nth n = none := rfl +@[simp] theorem nth_cons_zero (a : α) (s : seq α) : (cons a s).nth 0 = some a := rfl +@[simp] theorem nth_cons_succ (a : α) (s : seq α) (n : ℕ) : (cons a s).nth (n + 1) = s.nth n := rfl + +@[ext] protected lemma ext {s t : seq α} (h : ∀ n : ℕ, s.nth n = t.nth n) : s = t := +subtype.eq $ funext h + /-- A sequence has terminated at position `n` if the value at position `n` equals `none`. -/ def terminated_at (s : seq α) (n : ℕ) : Prop := s.nth n = none @@ -161,6 +171,8 @@ by rw [head_eq_destruct, destruct_cons]; refl @[simp] theorem tail_cons (a : α) (s) : tail (cons a s) = s := by cases s with f al; apply subtype.eq; dsimp [tail, cons]; rw [stream.tail_cons] +@[simp] theorem nth_tail (s : seq α) (n) : nth (tail s) n = nth s (n + 1) := rfl + def cases_on {C : seq α → Sort v} (s : seq α) (h1 : C nil) (h2 : ∀ x s, C (cons x s)) : C s := begin induction H : destruct s with v v, @@ -222,17 +234,6 @@ begin dsimp [corec.F], rw h, refl end -/-- Embed a list as a sequence -/ -def of_list (l : list α) : seq α := -⟨list.nth l, λ n h, begin - induction l with a l IH generalizing n, refl, - dsimp [list.nth], cases n with n; dsimp [list.nth] at h, - { contradiction }, - { apply IH _ h } -end⟩ - -instance coe_list : has_coe (list α) (seq α) := ⟨of_list⟩ - section bisim variable (R : seq α → seq α → Prop) @@ -291,6 +292,20 @@ begin rw [h1, h2], apply H end +/-- Embed a list as a sequence -/ +def of_list (l : list α) : seq α := +⟨list.nth l, λ n h, begin + rw list.nth_eq_none_iff at h ⊢, + exact h.trans (nat.le_succ n) +end⟩ + +instance coe_list : has_coe (list α) (seq α) := ⟨of_list⟩ + +@[simp] theorem of_list_nil : of_list [] = (nil : seq α) := rfl +@[simp] theorem of_list_nth (l : list α) (n : ℕ) : (of_list l).nth n = l.nth n := rfl +@[simp] theorem of_list_cons (a : α) (l : list α) : of_list (a :: l) = cons a (of_list l) := +by ext1 (_|n); refl + /-- Embed an infinite stream as a sequence -/ def of_stream (s : stream α) : seq α := ⟨s.map some, λ n h, by contradiction⟩ @@ -589,12 +604,6 @@ begin { refine ⟨nil, S, T, _, _⟩; simp } end -@[simp] theorem of_list_nil : of_list [] = (nil : seq α) := rfl - -@[simp] theorem of_list_cons (a : α) (l) : - of_list (a :: l) = cons a (of_list l) := -by ext (_|n) : 2; simp [of_list, cons, stream.nth, stream.cons] - @[simp] theorem of_stream_cons (a : α) (s) : of_stream (a :: s) = cons a (of_stream s) := by apply subtype.eq; simp [of_stream, cons]; rw stream.map_cons @@ -624,27 +633,6 @@ theorem dropn_add (s : seq α) (m) : ∀ n, drop s (m + n) = drop (drop s m) n theorem dropn_tail (s : seq α) (n) : drop (tail s) n = drop s (n + 1) := by rw add_comm; symmetry; apply dropn_add -theorem nth_tail : ∀ (s : seq α) n, nth (tail s) n = nth s (n + 1) -| ⟨f, al⟩ n := rfl - -@[ext] -protected lemma ext (s s': seq α) (hyp : ∀ (n : ℕ), s.nth n = s'.nth n) : s = s' := -begin - let ext := (λ (s s' : seq α), ∀ n, s.nth n = s'.nth n), - apply seq.eq_of_bisim ext _ hyp, - -- we have to show that ext is a bisimulation - clear hyp s s', - assume s s' (hyp : ext s s'), - unfold seq.destruct, - rw (hyp 0), - cases (s'.nth 0), - { simp [seq.bisim_o] }, -- option.none - { -- option.some - suffices : ext s.tail s'.tail, by simpa, - assume n, - simp only [seq.nth_tail _ n, (hyp $ n + 1)] } -end - @[simp] theorem head_dropn (s : seq α) (n) : head (drop s n) = nth s n := begin induction n with n IH generalizing s, { refl }, From 6cc2de236bf6af3b7ef549325db52a0d88451bde Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 24 Aug 2022 05:21:06 +0000 Subject: [PATCH 007/828] feat(analysis/complex/arg): `same_ray_iff_arg_div_eq_zero` (#16218) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add the lemma that `same_ray ℝ x y ↔ arg (x / y) = 0`. --- src/analysis/complex/arg.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/analysis/complex/arg.lean b/src/analysis/complex/arg.lean index 919daf70a0dcb..34e540b1e5dd0 100644 --- a/src/analysis/complex/arg.lean +++ b/src/analysis/complex/arg.lean @@ -37,6 +37,14 @@ begin rw [mul_comm, eq_comm] end +lemma same_ray_iff_arg_div_eq_zero : same_ray ℝ x y ↔ arg (x / y) = 0 := +begin + rw [←real.angle.to_real_zero, ←arg_coe_angle_eq_iff_eq_to_real, same_ray_iff], + by_cases hx : x = 0, { simp [hx] }, + by_cases hy : y = 0, { simp [hy] }, + simp [hx, hy, arg_div_coe_angle, sub_eq_zero] +end + lemma abs_add_eq_iff : (x + y).abs = x.abs + y.abs ↔ x = 0 ∨ y = 0 ∨ x.arg = y.arg := same_ray_iff_norm_add.symm.trans same_ray_iff From 44de64f183393284a16016dfb2a48ac97382f2bd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 24 Aug 2022 07:00:05 +0000 Subject: [PATCH 008/828] chore(algebra/order/with_zero): Move unrelated lemmas (#15676) Move a bunch of lemmas that were not about `whatever_with_zero` from `algebra.order.with_zero` to `algebra.group_power.order`. Delete `nat.le_zero_iff` in favor of `le_zero_iff`. --- archive/imo/imo1988_q6.lean | 4 +- archive/imo/imo2013_q5.lean | 2 +- src/algebra/big_operators/basic.lean | 2 +- src/algebra/group_power/order.lean | 103 +++++++++++++- src/algebra/lie/solvable.lean | 2 +- src/algebra/order/with_zero.lean | 128 ------------------ .../inner_product_space/projection.lean | 2 +- src/combinatorics/configuration.lean | 2 +- src/data/list/big_operators.lean | 4 +- src/data/list/zip.lean | 4 +- src/data/nat/basic.lean | 4 +- src/data/nat/choose/factorization.lean | 2 +- src/data/nat/factorial/basic.lean | 2 +- src/data/ordmap/ordset.lean | 16 +-- src/data/polynomial/reverse.lean | 2 +- src/data/polynomial/unit_trinomial.lean | 2 +- src/data/real/hyperreal.lean | 2 +- src/data/zmod/basic.lean | 2 +- src/field_theory/polynomial_galois_group.lean | 2 +- src/group_theory/p_group.lean | 2 +- src/group_theory/perm/cycle/type.lean | 2 +- src/group_theory/perm/support.lean | 2 +- src/information_theory/hamming.lean | 2 +- src/linear_algebra/basis.lean | 2 +- src/number_theory/padics/padic_val.lean | 2 +- src/ring_theory/multiplicity.lean | 2 +- src/topology/path_connected.lean | 2 +- 27 files changed, 134 insertions(+), 169 deletions(-) diff --git a/archive/imo/imo1988_q6.lean b/archive/imo/imo1988_q6.lean index c06f98146b3ce..ffceec71d3b10 100644 --- a/archive/imo/imo1988_q6.lean +++ b/archive/imo/imo1988_q6.lean @@ -205,7 +205,7 @@ begin { -- Show that the claim is true if a = b. intros x hx, suffices : k ≤ 1, - { rw [nat.le_add_one_iff, nat.le_zero_iff] at this, + { rw [nat.le_add_one_iff, le_zero_iff] at this, rcases this with rfl|rfl, { use 0, simp }, { use 1, simp } }, @@ -286,7 +286,7 @@ begin end, } }, { -- Show the base case. intros x y h h_base, - obtain rfl|rfl : x = 0 ∨ x = 1 := by rwa [nat.le_add_one_iff, nat.le_zero_iff] at h_base, + obtain rfl|rfl : x = 0 ∨ x = 1 := by rwa [nat.le_add_one_iff, le_zero_iff] at h_base, { simpa using h, }, { simp only [mul_one, one_mul, add_comm, zero_add] at h, have y_dvd : y ∣ y * k := dvd_mul_right y k, diff --git a/archive/imo/imo2013_q5.lean b/archive/imo/imo2013_q5.lean index 925a5b83d9596..261f9526df2d7 100644 --- a/archive/imo/imo2013_q5.lean +++ b/archive/imo/imo2013_q5.lean @@ -184,7 +184,7 @@ begin have hxp : 0 < x := by positivity, have hNp : 0 < N, - { by_contra' H, rw [nat.le_zero_iff.mp H] at hN, linarith }, + { by_contra' H, rw [le_zero_iff.mp H] at hN, linarith }, have h2 := calc f x + f (a^N - x) ≤ f (x + (a^N - x)) : H2 x (a^N - x) hxp (zero_lt_one.trans h_big_enough) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 84b21f722d20d..ffc7e8ceaa609 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -949,7 +949,7 @@ begin induction m with m hm, { simp }, erw [prod_range_succ, hm], - simp [hu] + simp [hu, @zero_le' ℕ], end @[to_additive] diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index 38428a814aa03..0580656f20723 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -13,15 +13,22 @@ Note that some lemmas are in `algebra/group_power/lemmas.lean` as they import fi depend on this file. -/ +open function + variables {A G M R : Type*} +section monoid +variable [monoid M] + section preorder +variable [preorder M] -variables [monoid M] [preorder M] [covariant_class M M (*) (≤)] +section left +variables [covariant_class M M (*) (≤)] {x : M} @[to_additive nsmul_le_nsmul_of_le_right, mono] -lemma pow_le_pow_of_le_left' [covariant_class M M (function.swap (*)) (≤)] - {a b : M} (hab : a ≤ b) : ∀ i : ℕ, a ^ i ≤ b ^ i +lemma pow_le_pow_of_le_left' [covariant_class M M (swap (*)) (≤)] {a b : M} (hab : a ≤ b) : + ∀ i : ℕ, a ^ i ≤ b ^ i | 0 := by simp | (k+1) := by { rw [pow_succ, pow_succ], exact mul_le_mul' hab (pow_le_pow_of_le_left' k) } @@ -75,11 +82,53 @@ lemma pow_strict_mono_left [covariant_class M M (*) (<)] {a : M} (ha : 1 < a) : strict_mono ((^) a : ℕ → M) := λ m n, pow_lt_pow' ha +@[to_additive left.pow_nonneg] +lemma left.one_le_pow_of_le (hx : 1 ≤ x) : ∀ {n : ℕ}, 1 ≤ x^n +| 0 := (pow_zero x).ge +| (n + 1) := by { rw pow_succ, exact left.one_le_mul hx left.one_le_pow_of_le } + +@[to_additive left.pow_nonpos] +lemma left.pow_le_one_of_le (hx : x ≤ 1) : ∀ {n : ℕ}, x^n ≤ 1 +| 0 := (pow_zero _).le +| (n + 1) := by { rw pow_succ, exact left.mul_le_one hx left.pow_le_one_of_le } + +end left + +section right +variables [covariant_class M M (swap (*)) (≤)] {x : M} + +@[to_additive right.pow_nonneg] +lemma right.one_le_pow_of_le (hx : 1 ≤ x) : ∀ {n : ℕ}, 1 ≤ x^n +| 0 := (pow_zero _).ge +| (n + 1) := by { rw pow_succ, exact right.one_le_mul hx right.one_le_pow_of_le } + +@[to_additive right.pow_nonpos] +lemma right.pow_le_one_of_le (hx : x ≤ 1) : ∀ {n : ℕ}, x^n ≤ 1 +| 0 := (pow_zero _).le +| (n + 1) := by { rw pow_succ, exact right.mul_le_one hx right.pow_le_one_of_le } + +end right + +@[to_additive left.pow_neg] +lemma left.pow_lt_one_of_lt [covariant_class M M (*) (<)] {n : ℕ} {x : M} (hn : 0 < n) (h : x < 1) : + x^n < 1 := +nat.le_induction ((pow_one _).trans_lt h) (λ n _ ih, by { rw pow_succ, exact mul_lt_one h ih }) _ + (nat.succ_le_iff.2 hn) + +@[to_additive right.pow_neg] +lemma right.pow_lt_one_of_lt [covariant_class M M (swap (*)) (<)] {n : ℕ} {x : M} + (hn : 0 < n) (h : x < 1) : + x^n < 1 := +nat.le_induction ((pow_one _).trans_lt h) + (λ n _ ih, by { rw pow_succ, exact right.mul_lt_one h ih }) _ (nat.succ_le_iff.2 hn) + end preorder section linear_order +variables [linear_order M] -variables [monoid M] [linear_order M] [covariant_class M M (*) (≤)] +section covariant_le +variables [covariant_class M M (*) (≤)] @[to_additive nsmul_nonneg_iff] lemma one_le_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 ≤ x ^ n ↔ 1 ≤ x := @@ -109,7 +158,21 @@ lemma pow_le_pow_iff' (ha : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := (pow_strict_ @[to_additive nsmul_lt_nsmul_iff] lemma pow_lt_pow_iff' (ha : 1 < a) : a ^ m < a ^ n ↔ m < n := (pow_strict_mono_left ha).lt_iff_lt +end covariant_le + +@[to_additive left.nsmul_neg_iff] +lemma left.pow_lt_one_iff [covariant_class M M (*) (<)] {n : ℕ} {x : M} (hn : 0 < n) : + x^n < 1 ↔ x < 1 := +by { haveI := has_mul.to_covariant_class_left M, exact pow_lt_one_iff hn.ne' } + +@[to_additive right.nsmul_neg_iff] +lemma right.pow_lt_one_iff [covariant_class M M (swap (*)) (<)] {n : ℕ} {x : M} (hn : 0 < n) : + x^n < 1 ↔ x < 1 := +⟨λ H, not_le.mp $ λ k, H.not_le $ by { haveI := has_mul.to_covariant_class_right M, + exact right.one_le_pow_of_le k }, right.pow_lt_one_of_lt hn⟩ + end linear_order +end monoid section div_inv_monoid @@ -404,3 +467,35 @@ sub_nonneg.mp ((sub_add_eq_add_sub _ _ _).subst ((sub_sq a b).subst (sq_nonneg _ alias two_mul_le_add_sq ← two_mul_le_add_pow_two end linear_ordered_comm_ring + +section linear_ordered_comm_monoid_with_zero +variables [linear_ordered_comm_monoid_with_zero M] [no_zero_divisors M] {a : M} {n : ℕ} + +lemma pow_pos_iff (hn : 0 < n) : 0 < a ^ n ↔ 0 < a := by simp_rw [zero_lt_iff, pow_ne_zero_iff hn] + +end linear_ordered_comm_monoid_with_zero + +section linear_ordered_comm_group_with_zero +variables [linear_ordered_comm_group_with_zero M] {a : M} {m n : ℕ} + +lemma pow_lt_pow_succ (ha : 1 < a) : a ^ n < a ^ n.succ := +by { rw [←one_mul (a ^ n), pow_succ], + exact mul_lt_right₀ _ ha (pow_ne_zero _ (zero_lt_one₀.trans ha).ne') } + +lemma pow_lt_pow₀ (ha : 1 < a) (hmn : m < n) : a ^ m < a ^ n := +by { induction hmn with n hmn ih, exacts [pow_lt_pow_succ ha, lt_trans ih (pow_lt_pow_succ ha)] } + +end linear_ordered_comm_group_with_zero + +namespace monoid_hom +variables [ring R] [monoid M] [linear_order M] [covariant_class M M (*) (≤)] (f : R →* M) + +lemma map_neg_one : f (-1) = 1 := +(pow_eq_one_iff (nat.succ_ne_zero 1)).1 $ by rw [←map_pow, neg_one_sq, map_one] + +@[simp] lemma map_neg (x : R) : f (-x) = f x := +by rw [←neg_one_mul, map_mul, map_neg_one, one_mul] + +lemma map_sub_swap (x y : R) : f (x - y) = f (y - x) := by rw [←map_neg, neg_sub] + +end monoid_hom diff --git a/src/algebra/lie/solvable.lean b/src/algebra/lie/solvable.lean index 4dfacb432f908..ec058f075296a 100644 --- a/src/algebra/lie/solvable.lean +++ b/src/algebra/lie/solvable.lean @@ -78,7 +78,7 @@ end D k I ≤ D l J := begin revert l, induction k with k ih; intros l h₂, - { rw nat.le_zero_iff at h₂, rw [h₂, derived_series_of_ideal_zero], exact h₁, }, + { rw le_zero_iff at h₂, rw [h₂, derived_series_of_ideal_zero], exact h₁, }, { have h : l = k.succ ∨ l ≤ k, by rwa [le_iff_eq_or_lt, nat.lt_succ_iff] at h₂, cases h, { rw [h, derived_series_of_ideal_succ, derived_series_of_ideal_succ], diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index 2329a5f97ff01..e7b9859484e26 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -3,9 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Johan Commelin, Patrick Massot -/ - import algebra.order.group -import tactic.abel /-! # Linearly ordered commutative groups and monoids with a zero element adjoined @@ -50,98 +48,6 @@ instance [linear_ordered_add_comm_group_with_top α] : ..multiplicative.linear_ordered_comm_monoid_with_zero, ..multiplicative.nontrivial } -section monoid -variable [monoid α] - -section preorder -variable [preorder α] - -section left -variable [covariant_class α α (*) (≤)] - -lemma left.one_le_pow_of_le : ∀ {n : ℕ} {x : α}, 1 ≤ x → 1 ≤ x^n -| 0 x _ := (pow_zero x).symm.le -| (n + 1) x H := calc 1 ≤ x : H - ... = x * 1 : (mul_one x).symm - ... ≤ x * x ^ n : mul_le_mul_left' (left.one_le_pow_of_le H) x - ... = x ^ n.succ : (pow_succ x n).symm - -end left - -section right -variable [covariant_class α α (function.swap (*)) (≤)] - -lemma right.one_le_pow_of_le {x : α} (H : 1 ≤ x) : - ∀ {n : ℕ}, 1 ≤ x^n -| 0 := (pow_zero _).symm.le -| (n + 1) := calc 1 ≤ x : H - ... = 1 * x : (one_mul x).symm - ... ≤ x ^ n * x : mul_le_mul_right' right.one_le_pow_of_le x - ... = x ^ n.succ : (pow_succ' x n).symm - -lemma right.pow_le_one_of_le {x : α} (H : x ≤ 1) : - ∀ {n : ℕ}, x^n ≤ 1 -| 0 := (pow_zero _).le -| (n + 1) := calc x ^ n.succ = x ^ n * x : pow_succ' x n - ... ≤ 1 * x : mul_le_mul_right' right.pow_le_one_of_le x - ... = x : one_mul x - ... ≤ 1 : H - -end right - -lemma pow_le_pow_of_le [covariant_class α α (*) (≤)] [covariant_class α α (function.swap (*)) (≤)] - {x y : α} (H : x ≤ y) : - ∀ {n : ℕ} , x^n ≤ y^n -| 0 := (pow_zero _).le.trans (pow_zero _).symm.le -| (n + 1) := calc x ^ n.succ = x * x ^ n : pow_succ x n - ... ≤ y * x ^ n : mul_le_mul_right' H (x ^ n) - ... ≤ y * y ^ n : mul_le_mul_left' pow_le_pow_of_le y - ... = y ^ n.succ : (pow_succ y n).symm - -lemma left.pow_lt_one_of_lt [covariant_class α α (*) (<)] {n : ℕ} {x : α} (n0 : 0 < n) (H : x < 1) : - x^n < 1 := -begin - refine nat.le_induction ((pow_one _).le.trans_lt H) (λ n n1 hn, _) _ (nat.succ_le_iff.mpr n0), - calc x ^ (n + 1) = x * x ^ n : pow_succ x n - ... < x * 1 : mul_lt_mul_left' hn x - ... = x : mul_one x - ... < 1 : H -end - -lemma left.pow_lt_one_iff {α: Type*} [monoid α] [linear_order α] - [covariant_class α α (*) (<)] {n : ℕ} {x : α} (n0 : 0 < n) : - x^n < 1 ↔ x < 1 := -⟨λ H, not_le.mp (λ k, not_le.mpr H (by - { haveI := has_mul.to_covariant_class_left α, - exact left.one_le_pow_of_le k})), left.pow_lt_one_of_lt n0⟩ - -lemma right.pow_lt_one_of_lt [covariant_class α α (function.swap (*)) (<)] {n : ℕ} {x : α} - (n0 : 0 < n) (H : x < 1) : - x^n < 1 := -begin - refine nat.le_induction ((pow_one _).le.trans_lt H) (λ n n1 hn, _) _ (nat.succ_le_iff.mpr n0), - calc x ^ (n + 1) = x ^ n * x : pow_succ' x n - ... < 1 * x : mul_lt_mul_right' hn x - ... = x : one_mul x - ... < 1 : H -end - -lemma right.pow_lt_one_iff {α: Type*} [monoid α] [linear_order α] - [covariant_class α α (function.swap (*)) (<)] {n : ℕ} {x : α} (n0 : 0 < n) : - x^n < 1 ↔ x < 1 := -⟨λ H, not_le.mp (λ k, not_le.mpr H (by - { haveI := has_mul.to_covariant_class_right α, - exact right.one_le_pow_of_le k})), right.pow_lt_one_of_lt n0⟩ - -end preorder - -section left_right -variables [linear_order α] - [covariant_class α α (*) (≤)] [covariant_class α α (function.swap (*)) (≤)] - -end left_right - -end monoid instance [linear_ordered_comm_monoid α] : linear_ordered_comm_monoid_with_zero (with_zero α) := { mul_le_mul_left := λ x y, mul_le_mul_left', @@ -191,9 +97,6 @@ lemma zero_lt_iff : 0 < a ↔ a ≠ 0 := lemma ne_zero_of_lt (h : b < a) : a ≠ 0 := λ h1, not_lt_zero' $ show b < 0, from h1 ▸ h -lemma pow_pos_iff [no_zero_divisors α] {n : ℕ} (hn : 0 < n) : 0 < a ^ n ↔ 0 < a := -by simp_rw [zero_lt_iff, pow_ne_zero_iff hn] - instance : linear_ordered_add_comm_monoid_with_top (additive αᵒᵈ) := { top := (0 : α), top_add' := λ a, (zero_mul a : (0 : α) * a = 0), @@ -265,13 +168,6 @@ by { rw mul_comm at *, exact mul_inv_lt_of_lt_mul₀ h } lemma mul_lt_right₀ (c : α) (h : a < b) (hc : c ≠ 0) : a * c < b * c := by { contrapose! h, exact le_of_le_mul_right hc h } -lemma pow_lt_pow_succ {x : α} {n : ℕ} (hx : 1 < x) : x ^ n < x ^ n.succ := -by { rw [← one_mul (x ^ n), pow_succ], -exact mul_lt_right₀ _ hx (pow_ne_zero _ $ ne_of_gt (lt_trans zero_lt_one₀ hx)) } - -lemma pow_lt_pow₀ {x : α} {m n : ℕ} (hx : 1 < x) (hmn : m < n) : x ^ m < x ^ n := -by { induction hmn with n hmn ih, exacts [pow_lt_pow_succ hx, lt_trans ih (pow_lt_pow_succ hx)] } - lemma inv_lt_inv₀ (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ < b⁻¹ ↔ b < a := show (units.mk0 a ha)⁻¹ < (units.mk0 b hb)⁻¹ ↔ (units.mk0 b hb) < (units.mk0 a ha), from inv_lt_inv_iff @@ -350,27 +246,3 @@ instance : linear_ordered_add_comm_group_with_top (additive αᵒᵈ) := ..additive.sub_neg_monoid, ..additive.linear_ordered_add_comm_monoid_with_top, ..additive.nontrivial } - -namespace monoid_hom - -variables {R : Type*} [ring R] (f : R →* α) - -theorem map_neg_one : f (-1) = 1 := -(pow_eq_one_iff (nat.succ_ne_zero 1)).1 $ - calc f (-1) ^ 2 = f (-1) * f(-1) : sq _ - ... = f ((-1) * - 1) : (f.map_mul _ _).symm - ... = f ( - - 1) : congr_arg _ (neg_one_mul _) - ... = f 1 : congr_arg _ (neg_neg _) - ... = 1 : map_one f - -@[simp] lemma map_neg (x : R) : f (-x) = f x := -calc f (-x) = f (-1 * x) : congr_arg _ (neg_one_mul _).symm - ... = f (-1) * f x : map_mul _ _ _ - ... = 1 * f x : _root_.congr_arg (λ g, g * (f x)) (map_neg_one f) - ... = f x : one_mul _ - -lemma map_sub_swap (x y : R) : f (x - y) = f (y - x) := -calc f (x - y) = f (-(y - x)) : congr_arg _ (neg_sub _ _).symm - ... = _ : map_neg _ _ - -end monoid_hom diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 05762771767d5..73d8df8dd8ca3 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -1019,7 +1019,7 @@ begin { -- Base case: `n = 0`, the fixed subspace is the whole space, so `φ = id` refine ⟨[], rfl.le, show φ = 1, from _⟩, have : (continuous_linear_map.id ℝ F - φ.to_continuous_linear_equiv).ker = ⊤, - { rwa [nat.le_zero_iff, finrank_eq_zero, submodule.orthogonal_eq_bot_iff] at hn }, + { rwa [le_zero_iff, finrank_eq_zero, submodule.orthogonal_eq_bot_iff] at hn }, symmetry, ext x, have := linear_map.congr_fun (linear_map.ker_eq_top.mp this) x, diff --git a/src/combinatorics/configuration.lean b/src/combinatorics/configuration.lean index d87b2a2ebb7c3..353471a704eb0 100644 --- a/src/combinatorics/configuration.lean +++ b/src/combinatorics/configuration.lean @@ -130,7 +130,7 @@ begin finset.one_lt_card_iff.mp (nat.one_lt_iff_ne_zero_and_ne_one.mpr ⟨hs₀, hs₁⟩), exact (eq_or_eq (hp₁ l₁ hl₁) (hp₂ l₁ hl₁) (hp₁ l₂ hl₂) (hp₂ l₂ hl₂)).resolve_right hl₃ }, by_cases hs₃ : sᶜ.card = 0, - { rw [hs₃, nat.le_zero_iff], + { rw [hs₃, le_zero_iff], rw [finset.card_compl, tsub_eq_zero_iff_le, has_le.le.le_iff_eq (finset.card_le_univ _), eq_comm, finset.card_eq_iff_eq_univ] at hs₃ ⊢, rw hs₃, diff --git a/src/data/list/big_operators.lean b/src/data/list/big_operators.lean index 3de22fdefeacc..345c581d5a53a 100644 --- a/src/data/list/big_operators.lean +++ b/src/data/list/big_operators.lean @@ -117,7 +117,7 @@ end @[simp, to_additive] lemma prod_take_mul_prod_drop : ∀ (L : list M) (i : ℕ), (L.take i).prod * (L.drop i).prod = L.prod -| [] i := by simp +| [] i := by simp [@zero_le' ℕ] | L 0 := by simp | (h :: t) (n+1) := by { dsimp, rw [prod_cons, prod_cons, mul_assoc, prod_take_mul_prod_drop] } @@ -151,7 +151,7 @@ lemma prod_update_nth : ∀ (L : list M) (n : ℕ) (a : M), (L.take n).prod * (if n < L.length then a else 1) * (L.drop (n + 1)).prod | (x :: xs) 0 a := by simp [update_nth] | (x :: xs) (i+1) a := by simp [update_nth, prod_update_nth xs i a, mul_assoc] -| [] _ _ := by simp [update_nth, (nat.zero_le _).not_lt] +| [] _ _ := by simp [update_nth, (nat.zero_le _).not_lt, @zero_le' ℕ] open mul_opposite diff --git a/src/data/list/zip.lean b/src/data/list/zip.lean index 137526647df63..19ce67b061c94 100644 --- a/src/data/list/zip.lean +++ b/src/data/list/zip.lean @@ -395,8 +395,8 @@ variables [comm_monoid α] @[to_additive] lemma prod_mul_prod_eq_prod_zip_with_mul_prod_drop : ∀ (L L' : list α), L.prod * L'.prod = (zip_with (*) L L').prod * (L.drop L'.length).prod * (L'.drop L.length).prod -| [] ys := by simp -| xs [] := by simp +| [] ys := by simp [@zero_le' ℕ] +| xs [] := by simp [@zero_le' ℕ] | (x :: xs) (y :: ys) := begin simp only [drop, length, zip_with_cons_cons, prod_cons], rw [mul_assoc x, mul_comm xs.prod, mul_assoc y, mul_comm ys.prod, diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index bb19d1ada78af..e4b00182ed6ed 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import algebra.order.ring +import algebra.order.with_zero /-! # Basic operations on the natural numbers @@ -231,9 +232,6 @@ add_right_eq_self.mp $ le_antisymm ((two_mul a).symm.trans_le h) le_add_self lemma eq_zero_of_mul_le {a b : ℕ} (hb : 2 ≤ b) (h : b * a ≤ a) : a = 0 := eq_zero_of_double_le $ le_trans (nat.mul_le_mul_right _ hb) h -theorem le_zero_iff {i : ℕ} : i ≤ 0 ↔ i = 0 := -⟨nat.eq_zero_of_le_zero, λ h, h ▸ le_refl i⟩ - lemma zero_max {m : ℕ} : max 0 m = m := max_eq_right (zero_le _) diff --git a/src/data/nat/choose/factorization.lean b/src/data/nat/choose/factorization.lean index 0dc2dfedbecde..f856c8de6eb62 100644 --- a/src/data/nat/choose/factorization.lean +++ b/src/data/nat/choose/factorization.lean @@ -51,7 +51,7 @@ A `pow` form of `nat.factorization_choose_le` lemma pow_factorization_choose_le (hn : 0 < n) : p ^ (choose n k).factorization p ≤ n := begin cases le_or_lt p 1, - { exact (pow_le_pow_of_le h).trans ((le_of_eq (one_pow _)).trans hn) }, + { exact (pow_le_pow_of_le_left' h _).trans ((le_of_eq (one_pow _)).trans hn) }, { exact (pow_le_iff_le_log h hn).mpr factorization_choose_le_log }, end diff --git a/src/data/nat/factorial/basic.lean b/src/data/nat/factorial/basic.lean index 6860e7000836f..7c5f330db94f8 100644 --- a/src/data/nat/factorial/basic.lean +++ b/src/data/nat/factorial/basic.lean @@ -145,7 +145,7 @@ begin { exact (n.add_factorial_succ_lt_factorial_add_succ i2).le }, { rw [←add_assoc, factorial_succ (1 + n), add_mul, one_mul, add_comm 1 n], exact (add_le_add_iff_right _).mpr (one_le_mul (nat.le_add_left 1 n) (n + 1).factorial_pos) }, - rw [nat.le_zero_iff.mp (nat.succ_le_succ_iff.mp i0), zero_add, zero_add] + rw [le_zero_iff.mp (nat.succ_le_succ_iff.mp i0), zero_add, zero_add] end lemma add_factorial_le_factorial_add (i : ℕ) {n : ℕ} (n1 : 1 ≤ n) : diff --git a/src/data/ordmap/ordset.lean b/src/data/ordmap/ordset.lean index 17375b3c9e6c2..f3a2d52782cb1 100644 --- a/src/data/ordmap/ordset.lean +++ b/src/data/ordmap/ordset.lean @@ -566,7 +566,7 @@ begin { have : size rrl = 0 ∧ size rrr = 0, { have := balanced_sz_zero.1 hr.1.symm, rwa [size, sr.2.2.1, nat.succ_le_succ_iff, - nat.le_zero_iff, add_eq_zero_iff] at this }, + le_zero_iff, add_eq_zero_iff] at this }, cases sr.2.2.2.1.size_eq_zero.1 this.1, cases sr.2.2.2.2.size_eq_zero.1 this.2, obtain rfl : rrs = 1 := sr.2.2.1, @@ -575,7 +575,7 @@ begin { have : size rll = 0 ∧ size rlr = 0, { have := balanced_sz_zero.1 hr.1, rwa [size, sr.2.1.1, nat.succ_le_succ_iff, - nat.le_zero_iff, add_eq_zero_iff] at this }, + le_zero_iff, add_eq_zero_iff] at this }, cases sr.2.1.2.1.size_eq_zero.1 this.1, cases sr.2.1.2.2.size_eq_zero.1 this.2, obtain rfl : rls = 1 := sr.2.1.1, @@ -596,7 +596,7 @@ begin { have : size lrl = 0 ∧ size lrr = 0, { have := balanced_sz_zero.1 hl.1.symm, rwa [size, sl.2.2.1, nat.succ_le_succ_iff, - nat.le_zero_iff, add_eq_zero_iff] at this }, + le_zero_iff, add_eq_zero_iff] at this }, cases sl.2.2.2.1.size_eq_zero.1 this.1, cases sl.2.2.2.2.size_eq_zero.1 this.2, obtain rfl : lrs = 1 := sl.2.2.1, @@ -605,7 +605,7 @@ begin { have : size lll = 0 ∧ size llr = 0, { have := balanced_sz_zero.1 hl.1, rwa [size, sl.2.1.1, nat.succ_le_succ_iff, - nat.le_zero_iff, add_eq_zero_iff] at this }, + le_zero_iff, add_eq_zero_iff] at this }, cases sl.2.1.2.1.size_eq_zero.1 this.1, cases sl.2.1.2.2.size_eq_zero.1 this.2, obtain rfl : lls = 1 := sl.2.1.1, @@ -660,7 +660,7 @@ begin { have : size rl = 0 ∧ size rr = 0, { have := H1 rfl, rwa [size, sr.1, nat.succ_le_succ_iff, - nat.le_zero_iff, add_eq_zero_iff] at this }, + le_zero_iff, add_eq_zero_iff] at this }, cases sr.2.1.size_eq_zero.1 this.1, cases sr.2.2.size_eq_zero.1 this.2, rw sr.eq_node', refl }, @@ -1039,7 +1039,7 @@ begin { rw nat.succ_add at mm, rcases mm with _|⟨_,⟨⟩⟩ } }, rcases hm.3.1.resolve_left mm with ⟨mm₁, mm₂⟩, cases nat.eq_zero_or_pos (size ml) with ml0 ml0, - { rw [ml0, mul_zero, nat.le_zero_iff] at mm₂, + { rw [ml0, mul_zero, le_zero_iff] at mm₂, rw [ml0, mm₂] at mm, cases mm dec_trivial }, have : 2 * size l ≤ size ml + size mr + 1, { have := nat.mul_le_mul_left _ lr₁, @@ -1116,9 +1116,9 @@ begin { exact le_trans hb₂ (nat.mul_le_mul_left _ $ le_trans (nat.le_add_left _ _) (nat.le_add_right _ _)) } }, { cases nat.eq_zero_or_pos (size rl) with rl0 rl0, - { rw [rl0, not_lt, nat.le_zero_iff, nat.mul_eq_zero] at h, + { rw [rl0, not_lt, le_zero_iff, nat.mul_eq_zero] at h, replace h := h.resolve_left dec_trivial, - rw [rl0, h, nat.le_zero_iff, nat.mul_eq_zero] at H2, + rw [rl0, h, le_zero_iff, nat.mul_eq_zero] at H2, rw [hr.2.size_eq, rl0, h, H2.resolve_left dec_trivial] at H1, cases H1 dec_trivial }, refine hl.node4_l hr.left hr.right rl0 _, diff --git a/src/data/polynomial/reverse.lean b/src/data/polynomial/reverse.lean index c74102003ecff..eb328346d277f 100644 --- a/src/data/polynomial/reverse.lean +++ b/src/data/polynomial/reverse.lean @@ -265,7 +265,7 @@ lemma reverse_nat_trailing_degree (f : R[X]) : f.reverse.nat_trailing_degree = begin by_cases hf : f = 0, { rw [hf, reverse_zero, nat_trailing_degree_zero] }, - { rw ← nat.le_zero_iff, + { rw ← le_zero_iff, apply nat_trailing_degree_le_of_ne_zero, rw [coeff_zero_reverse], exact mt leading_coeff_eq_zero.mp hf }, diff --git a/src/data/polynomial/unit_trinomial.lean b/src/data/polynomial/unit_trinomial.lean index 48f117061837c..f2148fd2305e4 100644 --- a/src/data/polynomial/unit_trinomial.lean +++ b/src/data/polynomial/unit_trinomial.lean @@ -341,7 +341,7 @@ begin replace hp := hp.leading_coeff_is_unit, rw leading_coeff_mul at hp, replace hp := is_unit_of_mul_is_unit_left hp, - rw [not_lt, nat.le_zero_iff] at this, + rw [not_lt, le_zero_iff] at this, rwa [eq_C_of_nat_degree_eq_zero this, is_unit_C, ←this] }, intro hq'', rw nat_degree_pos_iff_degree_pos at hq'', diff --git a/src/data/real/hyperreal.lean b/src/data/real/hyperreal.lean index 019101769669a..056a5dc833097 100644 --- a/src/data/real/hyperreal.lean +++ b/src/data/real/hyperreal.lean @@ -72,7 +72,7 @@ lemma inv_epsilon_eq_omega : ε⁻¹ = ω := @inv_inv _ _ ω lemma epsilon_pos : 0 < ε := suffices ∀ᶠ i in hyperfilter ℕ, (0 : ℝ) < (i : ℕ)⁻¹, by rwa lt_def, have h0' : {n : ℕ | ¬ 0 < n} = {0} := -by simp only [not_lt, (set.set_of_eq_eq_singleton).symm]; ext; exact nat.le_zero_iff, +by simp only [not_lt, (set.set_of_eq_eq_singleton).symm]; ext; exact le_bot_iff, begin simp only [inv_pos, nat.cast_pos], exact mem_hyperfilter_of_finite_compl (by convert set.finite_singleton _), diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index 00fcd1af39fd7..76ef1f9022cf9 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -717,7 +717,7 @@ begin apply nat.mod_eq_of_lt, apply nat.sub_lt (fact.out (0 < n)), contrapose! h, - rwa [nat.le_zero_iff, val_eq_zero] at h, + rwa [le_zero_iff, val_eq_zero] at h, end /-- `val_min_abs x` returns the integer in the same equivalence class as `x` that is closest to `0`, diff --git a/src/field_theory/polynomial_galois_group.lean b/src/field_theory/polynomial_galois_group.lean index cb29cbc08bc77..7a2697257221c 100644 --- a/src/field_theory/polynomial_galois_group.lean +++ b/src/field_theory/polynomial_galois_group.lean @@ -363,7 +363,7 @@ lemma card_complex_roots_eq_card_real_add_card_not_gal_inv (p : ℚ[X]) : begin by_cases hp : p = 0, { simp_rw [hp, root_set_zero, set.to_finset_eq_empty_iff.mpr rfl, finset.card_empty, zero_add], - refine eq.symm (nat.le_zero_iff.mp ((finset.card_le_univ _).trans (le_of_eq _))), + refine eq.symm (le_zero_iff.mp ((finset.card_le_univ _).trans (le_of_eq _))), simp_rw [hp, root_set_zero, fintype.card_eq_zero_iff], apply_instance }, have inj : function.injective (is_scalar_tower.to_alg_hom ℚ ℝ ℂ) := (algebra_map ℝ ℂ).injective, diff --git a/src/group_theory/p_group.lean b/src/group_theory/p_group.lean index c2e9f15055037..e7114d7c01225 100644 --- a/src/group_theory/p_group.lean +++ b/src/group_theory/p_group.lean @@ -143,7 +143,7 @@ begin (λ b, quotient.induction_on' b (λ b _ hb, _)) (λ a ha _, by { rw [key, mem_fixed_points_iff_card_orbit_eq_one.mp a.2] })), obtain ⟨k, hk⟩ := hG.card_orbit b, - have : k = 0 := nat.le_zero_iff.1 (nat.le_of_lt_succ (lt_of_not_ge (mt (pow_dvd_pow p) + have : k = 0 := le_zero_iff.1 (nat.le_of_lt_succ (lt_of_not_ge (mt (pow_dvd_pow p) (by rwa [pow_one, ←hk, ←nat.modeq_zero_iff_dvd, ←zmod.eq_iff_modeq_nat, ←key, nat.cast_zero])))), exact ⟨⟨b, mem_fixed_points_iff_card_orbit_eq_one.2 $ by rw [hk, this, pow_zero]⟩, diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index 771575489f91a..b60349105f90d 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -225,7 +225,7 @@ lemma is_cycle_of_prime_order {σ : perm α} (h1 : (order_of σ).prime) begin obtain ⟨n, hn⟩ := cycle_type_prime_order h1, rw [←σ.sum_cycle_type, hn, multiset.sum_repeat, nsmul_eq_mul, nat.cast_id, mul_lt_mul_right - (order_of_pos σ), nat.succ_lt_succ_iff, nat.lt_succ_iff, nat.le_zero_iff] at h2, + (order_of_pos σ), nat.succ_lt_succ_iff, nat.lt_succ_iff, le_zero_iff] at h2, rw [←card_cycle_type_eq_one, hn, card_repeat, h2], end diff --git a/src/group_theory/perm/support.lean b/src/group_theory/perm/support.lean index e1944f2fcaffc..19830aa6d1a6a 100644 --- a/src/group_theory/perm/support.lean +++ b/src/group_theory/perm/support.lean @@ -539,7 +539,7 @@ begin end @[simp] lemma card_support_le_one {f : perm α} : f.support.card ≤ 1 ↔ f = 1 := -by rw [le_iff_lt_or_eq, nat.lt_succ_iff, nat.le_zero_iff, card_support_eq_zero, +by rw [le_iff_lt_or_eq, nat.lt_succ_iff, le_zero_iff, card_support_eq_zero, or_iff_not_imp_right, imp_iff_right f.card_support_ne_one] lemma two_le_card_support_of_ne_one {f : perm α} (h : f ≠ 1) : diff --git a/src/information_theory/hamming.lean b/src/information_theory/hamming.lean index cc11ad8a3649f..541c080cd68dc 100644 --- a/src/information_theory/hamming.lean +++ b/src/information_theory/hamming.lean @@ -88,7 +88,7 @@ hamming_dist_eq_zero.not /-- Corresponds to `dist_pos`. -/ @[simp] lemma hamming_dist_pos {x y : Π i, β i} : 0 < hamming_dist x y ↔ x ≠ y := -by rw [←hamming_dist_ne_zero, iff_not_comm, not_lt, nat.le_zero_iff] +by rw [←hamming_dist_ne_zero, iff_not_comm, not_lt, le_zero_iff] @[simp] lemma hamming_dist_lt_one {x y : Π i, β i} : hamming_dist x y < 1 ↔ x = y := by rw [nat.lt_one_iff, hamming_dist_eq_zero] diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index 8c2455db0bd0a..6086acf42aaa6 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -1116,7 +1116,7 @@ begin induction n with n rank_ih generalizing N, { suffices : N = ⊥, { rwa this }, - apply eq_bot_of_rank_eq_zero b _ (λ m v hv, nat.le_zero_iff.mp (rank_le v hv)) }, + apply eq_bot_of_rank_eq_zero b _ (λ m v hv, le_zero_iff.mp (rank_le v hv)) }, apply ih, intros N' N'_le x x_mem x_ortho, apply rank_ih, diff --git a/src/number_theory/padics/padic_val.lean b/src/number_theory/padics/padic_val.lean index e297e0f631f3d..e89f21ed2e3c2 100644 --- a/src/number_theory/padics/padic_val.lean +++ b/src/number_theory/padics/padic_val.lean @@ -484,7 +484,7 @@ begin split_ifs, { rw part_enat.coe_le_iff, exact λ hn, or.inr (hn _) }, - { simp only [true_and, not_lt, ne.def, not_false_iff, nat.le_zero_iff, hp.out.ne_one] at h, + { simp only [true_and, not_lt, ne.def, not_false_iff, le_zero_iff, hp.out.ne_one] at h, exact λ hn, or.inl h } }, { rintro (rfl|h), { exact dvd_zero (p ^ n) }, diff --git a/src/ring_theory/multiplicity.lean b/src/ring_theory/multiplicity.lean index d138e40faf236..19eeea7a630ef 100644 --- a/src/ring_theory/multiplicity.lean +++ b/src/ring_theory/multiplicity.lean @@ -227,7 +227,7 @@ lemma dvd_iff_multiplicity_pos {a b : α} : (0 : part_enat) < multiplicity a b lemma finite_nat_iff {a b : ℕ} : finite a b ↔ (a ≠ 1 ∧ 0 < b) := begin rw [← not_iff_not, not_finite_iff_forall, not_and_distrib, ne.def, - not_not, not_lt, nat.le_zero_iff], + not_not, not_lt, le_zero_iff], exact ⟨λ h, or_iff_not_imp_right.2 (λ hb, have ha : a ≠ 0, from λ ha, by simpa [ha] using h 1, by_contradiction (λ ha1 : a ≠ 1, diff --git a/src/topology/path_connected.lean b/src/topology/path_connected.lean index 4397b7b123f04..b751dc4ea436d 100644 --- a/src/topology/path_connected.lean +++ b/src/topology/path_connected.lean @@ -799,7 +799,7 @@ begin induction n with n hn, { use path.refl (p' 0), { split, - { rintros i hi, rw nat.le_zero_iff.mp hi, exact ⟨0, rfl⟩ }, + { rintros i hi, rw le_zero_iff.mp hi, exact ⟨0, rfl⟩ }, { rw range_subset_iff, rintros x, exact hp' 0 le_rfl } } }, { rcases hn (λ i hi, hp' i $ nat.le_succ_of_le hi) with ⟨γ₀, hγ₀⟩, rcases h.joined_in (p' n) (hp' n n.le_succ) (p' $ n+1) (hp' (n+1) $ le_rfl) with ⟨γ₁, hγ₁⟩, From 94f13c6596fa647be1d6778806598255f78a49ad Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 24 Aug 2022 09:35:12 +0000 Subject: [PATCH 009/828] feat(ring_theory/localization/away): finitely presented as R[X]/(rX-1) (#12455) + Show `localization.away r` is isomorphic to `R[X]/(rX-1)` (implemented as `adjoin_root (C r * X - 1)`) as `R`-algebras: `localization.away_equiv_adjoin`. Lemmas introduced for this purpose are: `alg_hom_ext`, `root_is_inv` and `alg_hom_subsingleton` under namespace `adjoin_root`, `ideal.quotient.alg_hom_ext`, `is_localization.away.mul_inv_self`, and `is_localization.alg_hom_subsingleton` (which says any two R-alg_hom from from a localization of R to another R-algebra are equal). + Deduce that the R-algebra S is finitely presented if S is a localization of R away from some `r : R`: `is_localization.away.finite_presentation`. Lemmas introduced for this purpose are `algebra.finite_presentation.polynomial` and `adjoin_root.finite_presentation`, and the `finite_type` versions are also added. + Golf `algebra.finite_finite_type/presentation.mv_polynomial` and fix typo in `mem_adjoint_support`. --- src/ring_theory/adjoin_root.lean | 26 +++++++++++++++++ src/ring_theory/finiteness.lean | 37 +++++++++---------------- src/ring_theory/ideal/operations.lean | 4 +++ src/ring_theory/localization/away.lean | 35 ++++++++++++++++++++++- src/ring_theory/localization/basic.lean | 6 ++++ 5 files changed, 83 insertions(+), 25 deletions(-) diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index 46b947b05040f..3156cdeb9947a 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -102,6 +102,12 @@ lemma algebra_map_eq' [comm_semiring S] [algebra S R] : variables {S} +lemma finite_type : algebra.finite_type R (adjoin_root f) := +(algebra.finite_type.polynomial R).of_surjective _ (ideal.quotient.mkₐ_surjective R _) + +lemma finite_presentation : algebra.finite_presentation R (adjoin_root f) := +(algebra.finite_presentation.polynomial R).quotient (submodule.fg_span_singleton f) + /-- The adjoined root. -/ def root : adjoin_root f := mk f X @@ -109,6 +115,12 @@ variables {f} instance has_coe_t : has_coe_t R (adjoin_root f) := ⟨of f⟩ +/-- Two `R`-`alg_hom` from `adjoin_root f` to the same `R`-algebra are the same iff + they agree on `root f`. -/ +@[ext] lemma alg_hom_ext [semiring S] [algebra R S] {g₁ g₂ : adjoin_root f →ₐ[R] S} + (h : g₁ (root f) = g₂ (root f)) : g₁ = g₂ := +ideal.quotient.alg_hom_ext R $ polynomial.alg_hom_ext h + @[simp] lemma mk_eq_mk {g h : R[X]} : mk f g = mk f h ↔ f ∣ g - h := ideal.quotient.eq.trans ideal.mem_span_singleton @@ -200,6 +212,20 @@ lift_root hfx @[simp] lemma lift_hom_of {x : R} : lift_hom f a hfx (of f x) = algebra_map _ _ x := lift_of hfx +section adjoin_inv + +@[simp] lemma root_is_inv (r : R) : of _ r * root (C r * X - 1) = 1 := +by convert sub_eq_zero.1 ((eval₂_sub _).symm.trans $ eval₂_root $ C r * X - 1); + simp only [eval₂_mul, eval₂_C, eval₂_X, eval₂_one] + +lemma alg_hom_subsingleton {S : Type*} [comm_ring S] [algebra R S] {r : R} : + subsingleton (adjoin_root (C r * X - 1) →ₐ[R] S) := +⟨λ f g, alg_hom_ext (@inv_unique _ _ (algebra_map R S r) _ _ + (by rw [← f.commutes, ← f.map_mul, algebra_map_eq, root_is_inv, map_one]) + (by rw [← g.commutes, ← g.map_mul, algebra_map_eq, root_is_inv, map_one]))⟩ + +end adjoin_inv + end comm_ring section irreducible diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index e92ccb25f6b93..eddc0aa982511 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -177,20 +177,12 @@ namespace finite_type lemma self : finite_type R R := ⟨⟨{1}, subsingleton.elim _ _⟩⟩ -section -open_locale classical +protected lemma polynomial : finite_type R R[X] := +⟨⟨{polynomial.X}, by { rw finset.coe_singleton, exact polynomial.adjoin_X }⟩⟩ protected lemma mv_polynomial (ι : Type*) [fintype ι] : finite_type R (mv_polynomial ι R) := -⟨⟨finset.univ.image mv_polynomial.X, begin - rw eq_top_iff, refine λ p, mv_polynomial.induction_on' p - (λ u x, finsupp.induction u (subalgebra.algebra_map_mem _ x) - (λ i n f hif hn ih, _)) - (λ p q ihp ihq, subalgebra.add_mem _ ihp ihq), - rw [add_comm, mv_polynomial.monomial_add_single], - exact subalgebra.mul_mem _ ih - (subalgebra.pow_mem _ (subset_adjoin $ finset.mem_image_of_mem _ $ finset.mem_univ _) _) -end⟩⟩ -end +⟨⟨finset.univ.image mv_polynomial.X, + by {rw [finset.coe_image, finset.coe_univ, set.image_univ], exact mv_polynomial.adjoin_range_X}⟩⟩ lemma of_restrict_scalars_finite_type [algebra A B] [is_scalar_tower R A B] [hB : finite_type R B] : finite_type A B := @@ -335,21 +327,18 @@ variable (R) /-- The ring of polynomials in finitely many variables is finitely presented. -/ protected lemma mv_polynomial (ι : Type u_2) [fintype ι] : finite_presentation R (mv_polynomial ι R) := -begin - have equiv := mv_polynomial.rename_equiv R (fintype.equiv_fin ι), - refine ⟨_, alg_equiv.to_alg_hom equiv.symm, _⟩, - split, - { exact (alg_equiv.symm equiv).surjective }, - suffices hinj : function.injective equiv.symm.to_alg_hom.to_ring_hom, - { rw [(ring_hom.injective_iff_ker_eq_bot _).1 hinj], - exact submodule.fg_bot }, - exact (alg_equiv.symm equiv).injective -end +let eqv := (mv_polynomial.rename_equiv R $ fintype.equiv_fin ι).symm in +⟨fintype.card ι, eqv, eqv.surjective, + ((ring_hom.injective_iff_ker_eq_bot _).1 eqv.injective).symm ▸ submodule.fg_bot⟩ /-- `R` is finitely presented as `R`-algebra. -/ lemma self : finite_presentation R R := equiv (finite_presentation.mv_polynomial R pempty) (mv_polynomial.is_empty_alg_equiv R pempty) +/-- `R[X]` is finitely presented as `R`-algebra. -/ +lemma polynomial : finite_presentation R R[X] := +equiv (finite_presentation.mv_polynomial R punit) (mv_polynomial.punit_alg_equiv R) + variable {R} /-- The quotient of a finitely presented algebra by a finitely generated ideal is finitely @@ -839,7 +828,7 @@ section semiring variables [comm_semiring R] [monoid M] /-- An element of `monoid_algebra R M` is in the subalgebra generated by its support. -/ -lemma mem_adjoint_support (f : monoid_algebra R M) : f ∈ adjoin R (of R M '' f.support) := +lemma mem_adjoin_support (f : monoid_algebra R M) : f ∈ adjoin R (of R M '' f.support) := begin suffices : span R (of R M '' f.support) ≤ (adjoin R (of R M '' f.support)).to_submodule, { exact this (mem_span_support f) }, @@ -859,7 +848,7 @@ begin ⋃ (g : monoid_algebra R M) (H : g ∈ S), of R M '' g.support, { intros s hs, exact set.mem_Union₂.2 ⟨f, ⟨hf, hs⟩⟩ }, - exact adjoin_mono hincl (mem_adjoint_support f) + exact adjoin_mono hincl (mem_adjoin_support f) end /-- If a set `S` generates, as algebra, `monoid_algebra R M`, then the image of the union of the diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 88786d5395d29..2feafe45fcf7b 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -1763,6 +1763,10 @@ by apply_instance def quotient.mkₐ (I : ideal A) : A →ₐ[R₁] A ⧸ I := ⟨λ a, submodule.quotient.mk a, rfl, λ _ _, rfl, rfl, λ _ _, rfl, λ _, rfl⟩ +lemma quotient.alg_hom_ext {I : ideal A} {S} [semiring S] [algebra R₁ S] ⦃f g : A ⧸ I →ₐ[R₁] S⦄ + (h : f.comp (quotient.mkₐ R₁ I) = g.comp (quotient.mkₐ R₁ I)) : f = g := +alg_hom.ext $ λ x, quotient.induction_on' x $ alg_hom.congr_fun h + lemma quotient.alg_map_eq (I : ideal A) : algebra_map R₁ (A ⧸ I) = (algebra_map A (A ⧸ I)).comp (algebra_map R₁ A) := rfl diff --git a/src/ring_theory/localization/away.lean b/src/ring_theory/localization/away.lean index 94e5a33c55037..897861524ceb7 100644 --- a/src/ring_theory/localization/away.lean +++ b/src/ring_theory/localization/away.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baanen -/ import ring_theory.localization.basic +import ring_theory.adjoin_root /-! # Localizations away from an element @@ -21,10 +22,12 @@ See `src/ring_theory/localization/basic.lean` for a design overview. localization, ring localization, commutative ring localization, characteristic predicate, commutative ring, field of fractions -/ + +section comm_semiring + variables {R : Type*} [comm_semiring R] (M : submonoid R) {S : Type*} [comm_semiring S] variables [algebra R S] {P : Type*} [comm_semiring P] - namespace is_localization section away @@ -44,6 +47,9 @@ variables [is_localization.away x S] noncomputable def inv_self : S := mk' S (1 : R) ⟨x, submonoid.mem_powers _⟩ +@[simp] lemma mul_inv_self : algebra_map R S x * inv_self x = 1 := +by { convert is_localization.mk'_mul_mk'_eq_one _ 1, symmetry, apply is_localization.mk'_one } + variables {g : R →+* P} /-- Given `x : R`, a localization map `F : R →+* S` away from `x`, and a map of `comm_semiring`s @@ -161,3 +167,30 @@ abbreviation away_map (f : R →+* P) (r : R) : is_localization.away.map _ _ f r end localization + +end comm_semiring + +open polynomial adjoin_root localization + +variables {R : Type*} [comm_ring R] + +local attribute [instance] is_localization.alg_hom_subsingleton adjoin_root.alg_hom_subsingleton + +/-- The `R`-`alg_equiv` between the localization of `R` away from `r` and + `R` with an inverse of `r` adjoined. -/ +noncomputable def localization.away_equiv_adjoin (r : R) : away r ≃ₐ[R] adjoin_root (C r * X - 1) := +alg_equiv.of_alg_hom + { commutes' := is_localization.away.away_map.lift_eq r + (is_unit_of_mul_eq_one _ _ $ root_is_inv r), .. away_lift _ r _ } + (lift_hom _ (is_localization.away.inv_self r) $ by simp only + [map_sub, map_mul, aeval_C, aeval_X, is_localization.away.mul_inv_self, aeval_one, sub_self]) + (subsingleton.elim _ _) + (subsingleton.elim _ _) + +lemma is_localization.adjoin_inv (r : R) : is_localization.away r (adjoin_root $ C r * X - 1) := +is_localization.is_localization_of_alg_equiv _ (localization.away_equiv_adjoin r) + +lemma is_localization.away.finite_presentation (r : R) {S} [comm_ring S] [algebra R S] + [is_localization.away r S] : algebra.finite_presentation R S := +(adjoin_root.finite_presentation _).equiv $ (localization.away_equiv_adjoin r).symm.trans $ + is_localization.alg_equiv (submonoid.powers r) _ _ diff --git a/src/ring_theory/localization/basic.lean b/src/ring_theory/localization/basic.lean index bc8650d08e7d4..67da6bb0ee506 100644 --- a/src/ring_theory/localization/basic.lean +++ b/src/ring_theory/localization/basic.lean @@ -469,6 +469,12 @@ lemma ring_hom_ext ⦃j k : S →+* P⦄ (h : j.comp (algebra_map R S) = k.comp (algebra_map R S)) : j = k := ring_hom.coe_monoid_hom_injective $ monoid_hom_ext M $ monoid_hom.ext $ ring_hom.congr_fun h +/- This is not an instance because the submonoid `M` would become a metavariable + in typeclass search. -/ +lemma alg_hom_subsingleton [algebra R P] : subsingleton (S →ₐ[R] P) := +⟨λ f g, alg_hom.coe_ring_hom_injective $ is_localization.ring_hom_ext M $ + by rw [f.comp_algebra_map, g.comp_algebra_map]⟩ + /-- To show `j` and `k` agree on the whole localization, it suffices to show they agree on the image of the base ring, if they preserve `1` and `*`. -/ protected lemma ext (j k : S → P) (hj1 : j 1 = 1) (hk1 : k 1 = 1) From a9402e0a11843994283cd0a918e61fa51360e26a Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 24 Aug 2022 09:35:13 +0000 Subject: [PATCH 010/828] refactor(algebraic_geometry/*): Make `LocallyRingedSpace.hom` a custom structure. (#15973) We also define `algebraic_geometry.Scheme.hom` as an type alias for morphisms between schemes to enable dot notation. Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/algebraic_geometry/AffineScheme.lean | 8 ++-- .../Gamma_Spec_adjunction.lean | 6 +-- src/algebraic_geometry/Scheme.lean | 29 +++++-------- src/algebraic_geometry/Spec.lean | 10 +++-- src/algebraic_geometry/gluing.lean | 3 +- .../locally_ringed_space.lean | 25 +++++------ .../locally_ringed_space/has_colimits.lean | 13 +++--- src/algebraic_geometry/morphisms/basic.lean | 4 +- .../morphisms/ring_hom_properties.lean | 2 +- src/algebraic_geometry/open_immersion.lean | 42 ++++++++++++------- 10 files changed, 73 insertions(+), 69 deletions(-) diff --git a/src/algebraic_geometry/AffineScheme.lean b/src/algebraic_geometry/AffineScheme.lean index 86f196c957aad..c80a5ebebafd5 100644 --- a/src/algebraic_geometry/AffineScheme.lean +++ b/src/algebraic_geometry/AffineScheme.lean @@ -146,7 +146,7 @@ def Scheme.affine_opens (X : Scheme) : set (opens X.carrier) := { U : opens X.carrier | is_affine_open U } lemma range_is_affine_open_of_open_immersion {X Y : Scheme} [is_affine X] (f : X ⟶ Y) - [H : is_open_immersion f] : is_affine_open ⟨set.range f.1.base, H.base_open.open_range⟩ := + [H : is_open_immersion f] : is_affine_open f.opens_range := begin refine is_affine_of_iso (is_open_immersion.iso_of_range_eq f (Y.of_restrict _) _).inv, exact subtype.range_coe.symm, @@ -222,7 +222,7 @@ end lemma is_affine_open.image_is_open_immersion {X Y : Scheme} {U : opens X.carrier} (hU : is_affine_open U) - (f : X ⟶ Y) [H : is_open_immersion f] : is_affine_open (H.open_functor.obj U) := + (f : X ⟶ Y) [H : is_open_immersion f] : is_affine_open (f.opens_functor.obj U) := begin haveI : is_affine _ := hU, convert range_is_affine_open_of_open_immersion (X.of_restrict U.open_embedding ≫ f), @@ -317,7 +317,7 @@ begin (X.basic_open f : set X.carrier), { rw [set.image_preimage_eq_inter_range, set.inter_eq_left_iff_subset, hU.from_Spec_range], exact Scheme.basic_open_subset _ _ }, - rw [subtype.coe_mk, Scheme.comp_val_base, ← this, coe_comp, set.range_comp], + rw [Scheme.hom.opens_range_coe, Scheme.comp_val_base, ← this, coe_comp, set.range_comp], congr' 1, refine (congr_arg coe $ Scheme.preimage_basic_open hU.from_Spec f).trans _, refine eq.trans _ (prime_spectrum.localization_away_comap_range (localization.away f) f).symm, @@ -396,7 +396,7 @@ begin swap, exact ⟨x, h⟩, have : U.open_embedding.is_open_map.functor.obj ((X.restrict U.open_embedding).basic_open r) = X.basic_open (X.presheaf.map (eq_to_hom U.open_embedding_obj_top.symm).op r), - { refine (is_open_immersion.image_basic_open (X.of_restrict U.open_embedding) r).trans _, + { refine (Scheme.image_basic_open (X.of_restrict U.open_embedding) r).trans _, erw ← Scheme.basic_open_res_eq _ _ (eq_to_hom U.open_embedding_obj_top).op, rw [← comp_apply, ← category_theory.functor.map_comp, ← op_comp, eq_to_hom_trans, eq_to_hom_refl, op_id, category_theory.functor.map_id], diff --git a/src/algebraic_geometry/Gamma_Spec_adjunction.lean b/src/algebraic_geometry/Gamma_Spec_adjunction.lean index 1c6a650e38820..ebf8d621b59e4 100644 --- a/src/algebraic_geometry/Gamma_Spec_adjunction.lean +++ b/src/algebraic_geometry/Gamma_Spec_adjunction.lean @@ -183,10 +183,10 @@ begin end /-- The canonical morphism from `X` to the spectrum of its global sections. -/ -@[simps coe_base] +@[simps val_base] def to_Γ_Spec : X ⟶ Spec.LocallyRingedSpace_obj (Γ.obj (op X)) := { val := X.to_Γ_Spec_SheafedSpace, - property := + prop := begin intro x, let p : prime_spectrum (Γ.obj (op X)) := X.to_Γ_Spec_fun x, @@ -246,7 +246,7 @@ def identity_to_Γ_Spec : 𝟭 LocallyRingedSpace.{u} ⟶ Γ.right_op ⋙ Spec.t apply LocallyRingedSpace.comp_ring_hom_ext, { ext1 x, dsimp [Spec.Top_map, LocallyRingedSpace.to_Γ_Spec_fun], - rw [← subtype.val_eq_coe, ← local_ring.comap_closed_point (PresheafedSpace.stalk_map _ x), + rw [← local_ring.comap_closed_point (PresheafedSpace.stalk_map _ x), ← prime_spectrum.comap_comp_apply, ← prime_spectrum.comap_comp_apply], congr' 2, exact (PresheafedSpace.stalk_map_germ f.1 ⊤ ⟨x,trivial⟩).symm, diff --git a/src/algebraic_geometry/Scheme.lean b/src/algebraic_geometry/Scheme.lean index d071a492cc6c7..b6dde5bd50afa 100644 --- a/src/algebraic_geometry/Scheme.lean +++ b/src/algebraic_geometry/Scheme.lean @@ -39,11 +39,16 @@ structure Scheme extends to_LocallyRingedSpace : LocallyRingedSpace := namespace Scheme +/-- A morphism between schemes is a morphism between the underlying locally ringed spaces. -/ +@[nolint has_nonempty_instance] -- There isn't nessecarily a morphism between two schemes. +def hom (X Y : Scheme) : Type* := +X.to_LocallyRingedSpace ⟶ Y.to_LocallyRingedSpace + /-- Schemes are a full subcategory of locally ringed spaces. -/ instance : category Scheme := -induced_category.category Scheme.to_LocallyRingedSpace +{ hom := hom, ..(induced_category.category Scheme.to_LocallyRingedSpace) } /-- The structure sheaf of a Scheme. -/ protected abbreviation sheaf (X : Scheme) := X.to_SheafedSpace.sheaf @@ -61,16 +66,11 @@ def forget_to_LocallyRingedSpace : Scheme ⥤ LocallyRingedSpace := def forget_to_Top : Scheme ⥤ Top := Scheme.forget_to_LocallyRingedSpace ⋙ LocallyRingedSpace.forget_to_Top -instance {X Y : Scheme} : has_lift_t (X ⟶ Y) - (X.to_SheafedSpace ⟶ Y.to_SheafedSpace) := (@@coe_to_lift $ @@coe_base coe_subtype) - -lemma id_val_base (X : Scheme) : (subtype.val (𝟙 X)).base = 𝟙 _ := rfl - -@[simp] lemma id_coe_base (X : Scheme) : - (↑(𝟙 X) : X.to_SheafedSpace ⟶ X.to_SheafedSpace).base = 𝟙 _ := rfl +@[simp] +lemma id_val_base (X : Scheme) : (𝟙 X : _).1.base = 𝟙 _ := rfl @[simp] lemma id_app {X : Scheme} (U : (opens X.carrier)ᵒᵖ) : - (subtype.val (𝟙 X)).c.app U = X.presheaf.map + (𝟙 X : _).val.c.app U = X.presheaf.map (eq_to_hom (by { induction U using opposite.rec, cases U, refl })) := PresheafedSpace.id_c_app X.to_PresheafedSpace U @@ -80,7 +80,7 @@ lemma comp_val {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : @[reassoc, simp] lemma comp_coe_base {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : - (↑(f ≫ g) : X.to_SheafedSpace ⟶ Z.to_SheafedSpace).base = f.val.base ≫ g.val.base := rfl + (f ≫ g).val.base = f.val.base ≫ g.val.base := rfl @[reassoc, elementwise] lemma comp_val_base {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : @@ -211,19 +211,13 @@ RingedSpace.basic_open_res_eq _ i f lemma basic_open_subset : X.basic_open f ⊆ U := RingedSpace.basic_open_subset _ _ +@[simp] lemma preimage_basic_open {X Y : Scheme} (f : X ⟶ Y) {U : opens Y.carrier} (r : Y.presheaf.obj $ op U) : (opens.map f.1.base).obj (Y.basic_open r) = @Scheme.basic_open X ((opens.map f.1.base).obj U) (f.1.c.app _ r) := LocallyRingedSpace.preimage_basic_open f r -@[simp] -lemma preimage_basic_open' {X Y : Scheme} (f : X ⟶ Y) {U : opens Y.carrier} - (r : Y.presheaf.obj $ op U) : - (opens.map (↑f : X.to_SheafedSpace ⟶ Y.to_SheafedSpace).base).obj (Y.basic_open r) = - @Scheme.basic_open X ((opens.map f.1.base).obj U) (f.1.c.app _ r) := -LocallyRingedSpace.preimage_basic_open f r - @[simp] lemma basic_open_zero (U : opens X.carrier) : X.basic_open (0 : X.presheaf.obj $ op U) = ∅ := LocallyRingedSpace.basic_open_zero _ U @@ -232,7 +226,6 @@ LocallyRingedSpace.basic_open_zero _ U lemma basic_open_mul : X.basic_open (f * g) = X.basic_open f ⊓ X.basic_open g := RingedSpace.basic_open_mul _ _ _ -@[simp] lemma basic_open_of_is_unit {f : X.presheaf.obj (op U)} (hf : is_unit f) : X.basic_open f = U := RingedSpace.basic_open_of_is_unit _ hf diff --git a/src/algebraic_geometry/Spec.lean b/src/algebraic_geometry/Spec.lean index ef8082757c8d8..05e1db2032b95 100644 --- a/src/algebraic_geometry/Spec.lean +++ b/src/algebraic_geometry/Spec.lean @@ -188,7 +188,7 @@ The induced map of a ring homomorphism on the prime spectra, as a morphism of lo -/ @[simps] def Spec.LocallyRingedSpace_map {R S : CommRing} (f : R ⟶ S) : Spec.LocallyRingedSpace_obj S ⟶ Spec.LocallyRingedSpace_obj R := -subtype.mk (Spec.SheafedSpace_map f) $ λ p, is_local_ring_hom.mk $ λ a ha, +LocallyRingedSpace.hom.mk (Spec.SheafedSpace_map f) $ λ p, is_local_ring_hom.mk $ λ a ha, begin -- Here, we are showing that the map on prime spectra induced by `f` is really a morphism of -- *locally* ringed spaces, i.e. that the induced map on the stalks is a local ring homomorphism. @@ -197,17 +197,19 @@ begin rw iso.inv_hom_id_apply at ha, replace ha := is_local_ring_hom.map_nonunit _ ha, convert ring_hom.is_unit_map (stalk_iso R (prime_spectrum.comap f p)).inv ha, - rw iso.hom_inv_id_apply, + rw iso.hom_inv_id_apply end @[simp] lemma Spec.LocallyRingedSpace_map_id (R : CommRing) : Spec.LocallyRingedSpace_map (𝟙 R) = 𝟙 (Spec.LocallyRingedSpace_obj R) := -subtype.ext $ by { rw [Spec.LocallyRingedSpace_map_coe, Spec.SheafedSpace_map_id], refl } +LocallyRingedSpace.hom.ext _ _ $ + by { rw [Spec.LocallyRingedSpace_map_val, Spec.SheafedSpace_map_id], refl } lemma Spec.LocallyRingedSpace_map_comp {R S T : CommRing} (f : R ⟶ S) (g : S ⟶ T) : Spec.LocallyRingedSpace_map (f ≫ g) = Spec.LocallyRingedSpace_map g ≫ Spec.LocallyRingedSpace_map f := -subtype.ext $ by { rw [Spec.LocallyRingedSpace_map_coe, Spec.SheafedSpace_map_comp], refl } +LocallyRingedSpace.hom.ext _ _ $ + by { rw [Spec.LocallyRingedSpace_map_val, Spec.SheafedSpace_map_comp], refl } /-- Spec, as a contravariant functor from commutative rings to locally ringed spaces. diff --git a/src/algebraic_geometry/gluing.lean b/src/algebraic_geometry/gluing.lean index 9fc37f88a00ca..18c770363b79a 100644 --- a/src/algebraic_geometry/gluing.lean +++ b/src/algebraic_geometry/gluing.lean @@ -351,7 +351,8 @@ instance from_glued_stalk_iso (x : 𝒰.glued_cover.glued.carrier) : is_iso (PresheafedSpace.stalk_map 𝒰.from_glued.val x) := begin obtain ⟨i, x, rfl⟩ := 𝒰.glued_cover.ι_jointly_surjective x, - have := PresheafedSpace.stalk_map.congr_hom _ _ (congr_arg subtype.val $ 𝒰.ι_from_glued i) x, + have := PresheafedSpace.stalk_map.congr_hom _ _ + (congr_arg LocallyRingedSpace.hom.val $ 𝒰.ι_from_glued i) x, erw PresheafedSpace.stalk_map.comp at this, rw ← is_iso.eq_comp_inv at this, rw this, diff --git a/src/algebraic_geometry/locally_ringed_space.lean b/src/algebraic_geometry/locally_ringed_space.lean index ab03f1d82675c..74ee81b462f4b 100644 --- a/src/algebraic_geometry/locally_ringed_space.lean +++ b/src/algebraic_geometry/locally_ringed_space.lean @@ -63,15 +63,13 @@ def 𝒪 : sheaf CommRing X.to_Top := X.to_SheafedSpace.sheaf /-- A morphism of locally ringed spaces is a morphism of ringed spaces such that the morphims induced on stalks are local ring homomorphisms. -/ -def hom (X Y : LocallyRingedSpace) : Type* := -{ f : X.to_SheafedSpace ⟶ Y.to_SheafedSpace // - ∀ x, is_local_ring_hom (PresheafedSpace.stalk_map f x) } +@[ext] +structure hom (X Y : LocallyRingedSpace.{u}) : Type u := +(val : X.to_SheafedSpace ⟶ Y.to_SheafedSpace) +(prop : ∀ x, is_local_ring_hom (PresheafedSpace.stalk_map val x)) instance : quiver LocallyRingedSpace := ⟨hom⟩ -@[ext] lemma hom_ext {X Y : LocallyRingedSpace} (f g : hom X Y) (w : f.1 = g.1) : f = g := -subtype.eq w - /-- The stalk of a locally ringed space, just as a `CommRing`. -/ @@ -103,7 +101,6 @@ def id (X : LocallyRingedSpace) : hom X X := instance (X : LocallyRingedSpace) : inhabited (hom X X) := ⟨id X⟩ /-- Composition of morphisms of locally ringed spaces. -/ -@[simps] def comp {X Y Z : LocallyRingedSpace} (f : hom X Y) (g : hom Y Z) : hom X Z := ⟨f.val ≫ g.val, λ x, begin @@ -116,9 +113,9 @@ instance : category LocallyRingedSpace := { hom := hom, id := id, comp := λ X Y Z f g, comp f g, - comp_id' := by { intros, ext1, simp, }, - id_comp' := by { intros, ext1, simp, }, - assoc' := by { intros, ext1, simp, }, }. + comp_id' := by { intros, ext1, simp [comp], }, + id_comp' := by { intros, ext1, simp [comp], }, + assoc' := by { intros, ext1, simp [comp], }, }. /-- The forgetful functor from `LocallyRingedSpace` to `SheafedSpace CommRing`. -/ @[simps] def forget_to_SheafedSpace : LocallyRingedSpace ⥤ SheafedSpace CommRing := @@ -151,7 +148,7 @@ See also `iso_of_SheafedSpace_iso`. @[simps] def hom_of_SheafedSpace_hom_of_is_iso {X Y : LocallyRingedSpace} (f : X.to_SheafedSpace ⟶ Y.to_SheafedSpace) [is_iso f] : X ⟶ Y := -subtype.mk f $ λ x, +hom.mk f $ λ x, -- Here we need to see that the stalk maps are really local ring homomorphisms. -- This can be solved by type class inference, because stalk maps of isomorphisms are isomorphisms -- and isomorphisms are local ring homomorphisms. @@ -171,14 +168,14 @@ def iso_of_SheafedSpace_iso {X Y : LocallyRingedSpace} (f : X.to_SheafedSpace ≅ Y.to_SheafedSpace) : X ≅ Y := { hom := hom_of_SheafedSpace_hom_of_is_iso f.hom, inv := hom_of_SheafedSpace_hom_of_is_iso f.inv, - hom_inv_id' := hom_ext _ _ f.hom_inv_id, - inv_hom_id' := hom_ext _ _ f.inv_hom_id } + hom_inv_id' := hom.ext _ _ f.hom_inv_id, + inv_hom_id' := hom.ext _ _ f.inv_hom_id } instance : reflects_isomorphisms forget_to_SheafedSpace := { reflects := λ X Y f i, { out := by exactI ⟨hom_of_SheafedSpace_hom_of_is_iso (category_theory.inv (forget_to_SheafedSpace.map f)), - hom_ext _ _ (is_iso.hom_inv_id _), hom_ext _ _ (is_iso.inv_hom_id _)⟩ } } + hom.ext _ _ (is_iso.hom_inv_id _), hom.ext _ _ (is_iso.inv_hom_id _)⟩ } } instance is_SheafedSpace_iso {X Y : LocallyRingedSpace} (f : X ⟶ Y) [is_iso f] : is_iso f.1 := diff --git a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean index cdb16170c412f..29e06353efd20 100644 --- a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean +++ b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean @@ -90,9 +90,10 @@ def coproduct_cofan_is_colimit : is_colimit (coproduct_cofan F) := ((forget_to_SheafedSpace.map_cocone s).ι.app i) y) := (s.ι.app i).2 y, apply_instance end⟩, - fac' := λ s j, subtype.eq (colimit.ι_desc _ _), - uniq' := λ s f h, subtype.eq (is_colimit.uniq _ (forget_to_SheafedSpace.map_cocone s) f.1 - (λ j, congr_arg subtype.val (h j))) } + fac' := λ s j, LocallyRingedSpace.hom.ext _ _ (colimit.ι_desc _ _), + uniq' := λ s f h, LocallyRingedSpace.hom.ext _ _ + (is_colimit.uniq _ (forget_to_SheafedSpace.map_cocone s) f.1 + (λ j, congr_arg LocallyRingedSpace.hom.val (h j))) } instance : has_coproducts.{u} LocallyRingedSpace.{u} := λ ι, ⟨λ F, ⟨⟨⟨_, coproduct_cofan_is_colimit F⟩⟩⟩⟩ @@ -235,7 +236,7 @@ def coequalizer : LocallyRingedSpace := noncomputable def coequalizer_cofork : cofork f g := @cofork.of_π _ _ _ _ f g (coequalizer f g) ⟨coequalizer.π f.1 g.1, infer_instance⟩ - (subtype.eq (coequalizer.condition f.1 g.1)) + (LocallyRingedSpace.hom.ext _ _ (coequalizer.condition f.1 g.1)) lemma is_local_ring_hom_stalk_map_congr {X Y : RingedSpace} (f g : X ⟶ Y) (H : f = g) (x) (h : is_local_ring_hom (PresheafedSpace.stalk_map f x)) : @@ -259,10 +260,10 @@ begin apply is_local_ring_hom_stalk_map_congr _ _ (coequalizer.π_desc s.π.1 e).symm y, apply_instance }, split, - exact subtype.eq (coequalizer.π_desc _ _), + { exact LocallyRingedSpace.hom.ext _ _ (coequalizer.π_desc _ _) }, intros m h, replace h : (coequalizer_cofork f g).π.1 ≫ m.1 = s.π.1 := by { rw ← h, refl }, - apply subtype.eq, + apply LocallyRingedSpace.hom.ext, apply (colimit.is_colimit (parallel_pair f.1 g.1)).uniq (cofork.of_π s.π.1 e) m.1, rintro ⟨⟩, { rw [← (colimit.cocone (parallel_pair f.val g.val)).w walking_parallel_pair_hom.left, diff --git a/src/algebraic_geometry/morphisms/basic.lean b/src/algebraic_geometry/morphisms/basic.lean index e0a48423e7776..a1d4c72ad5966 100644 --- a/src/algebraic_geometry/morphisms/basic.lean +++ b/src/algebraic_geometry/morphisms/basic.lean @@ -232,7 +232,7 @@ begin all_goals { ext1, exact subtype.range_coe } }, tfae_have : 1 → 5, { intro H, - refine ⟨Y.carrier, λ x, is_open_immersion.opens_range (Y.affine_cover.map x), _, + refine ⟨Y.carrier, λ x, (Y.affine_cover.map x).opens_range, _, λ i, range_is_affine_open_of_open_immersion _, _⟩, { rw eq_top_iff, intros x _, erw opens.mem_supr, exact⟨x, Y.affine_cover.covers x⟩ }, { intro i, exact H ⟨_, range_is_affine_open_of_open_immersion _⟩ } }, @@ -357,7 +357,7 @@ begin tfae_have : 4 → 3, { intros H 𝒰 i, rw ← hP.1.arrow_mk_iso_iff (morphism_restrict_opens_range f _), - exact H (is_open_immersion.opens_range $ 𝒰.map i) }, + exact H (𝒰.map i).opens_range }, tfae_have : 3 → 2, { exact λ H, ⟨Y.affine_cover, H Y.affine_cover⟩ }, tfae_have : 4 → 5, diff --git a/src/algebraic_geometry/morphisms/ring_hom_properties.lean b/src/algebraic_geometry/morphisms/ring_hom_properties.lean index e240172a464f3..89deac727eda6 100644 --- a/src/algebraic_geometry/morphisms/ring_hom_properties.lean +++ b/src/algebraic_geometry/morphisms/ring_hom_properties.lean @@ -286,7 +286,7 @@ begin { dsimp [functor.op], conv_lhs { rw opens.open_embedding_obj_top }, conv_rhs { rw opens.open_embedding_obj_top }, - erw is_open_immersion.image_basic_open (X.of_restrict U.1.open_embedding), + erw Scheme.image_basic_open (X.of_restrict U.1.open_embedding), erw PresheafedSpace.is_open_immersion.of_restrict_inv_app_apply, rw Scheme.basic_open_res_eq }, { apply_instance } diff --git a/src/algebraic_geometry/open_immersion.lean b/src/algebraic_geometry/open_immersion.lean index ac30093b33e51..44448d5e81040 100644 --- a/src/algebraic_geometry/open_immersion.lean +++ b/src/algebraic_geometry/open_immersion.lean @@ -610,8 +610,7 @@ omit H @[simp] lemma LocallyRingedSpace_to_LocallyRingedSpace {X Y : LocallyRingedSpace} (f : X ⟶ Y) [LocallyRingedSpace.is_open_immersion f] : - @to_LocallyRingedSpace X.to_PresheafedSpace Y (@@coe (@@coe_to_lift (@@coe_base coe_subtype)) f) - (show is_open_immersion f.val, by apply_instance) = X := + to_LocallyRingedSpace Y f.1 = X := by unfreezingI { cases X, delta to_LocallyRingedSpace, simp } end to_LocallyRingedSpace @@ -901,7 +900,8 @@ begin rw ← is_iso.eq_inv_comp at this, rw this, apply_instance }, - { exact subtype.eq (PresheafedSpace.is_open_immersion.pullback_cone_of_left_condition _ _) }, + { exact LocallyRingedSpace.hom.ext _ _ + (PresheafedSpace.is_open_immersion.pullback_cone_of_left_condition _ _) }, end instance : LocallyRingedSpace.is_open_immersion (pullback_cone_of_left f g).snd := @@ -912,24 +912,26 @@ def pullback_cone_of_left_is_limit : is_limit (pullback_cone_of_left f g) := pullback_cone.is_limit_aux' _ $ λ s, begin use PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift f.1 g.1 - (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg subtype.val s.condition)), + (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.hom.val s.condition)), { intro x, have := PresheafedSpace.stalk_map.congr_hom _ _ (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 - (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg subtype.val s.condition))) x, + (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.hom.val s.condition))) x, change _ = _ ≫ PresheafedSpace.stalk_map s.snd.1 x at this, rw [PresheafedSpace.stalk_map.comp, ← is_iso.eq_inv_comp] at this, rw this, apply_instance }, split, - exact subtype.eq (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_fst f.1 g.1 _), + { exact LocallyRingedSpace.hom.ext _ _ + (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_fst f.1 g.1 _) }, split, - exact subtype.eq (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 _), + { exact LocallyRingedSpace.hom.ext _ _ + (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 _) }, intros m h₁ h₂, rw ← cancel_mono (pullback_cone_of_left f g).snd, - exact (h₂.trans (subtype.eq + exact (h₂.trans (LocallyRingedSpace.hom.ext _ _ (PresheafedSpace.is_open_immersion.pullback_cone_of_left_lift_snd f.1 g.1 - (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg subtype.val s.condition))).symm)) + (pullback_cone.mk s.fst.1 s.snd.1 (congr_arg LocallyRingedSpace.hom.val s.condition))).symm)) end instance has_pullback_of_left : @@ -995,7 +997,7 @@ begin apply_with limits.comp_preserves_limit { instances := ff }, apply_instance, apply preserves_limit_of_iso_diagram _ (diagram_iso_cospan.{u} _).symm, - dsimp [SheafedSpace.forget_to_PresheafedSpace, -subtype.val_eq_coe], + dsimp [SheafedSpace.forget_to_PresheafedSpace], apply_instance, end @@ -1599,6 +1601,14 @@ LocallyRingedSpace.is_open_immersion.lift_uniq f g H' l hl hom_inv_id' := by { rw ← cancel_mono f, simp }, inv_hom_id' := by { rw ← cancel_mono g, simp } } +end is_open_immersion + +namespace Scheme + +/-- The functor `opens X ⥤ opens Y` associated with an open immersion `f : X ⟶ Y`. -/ +abbreviation hom.opens_functor {X Y : Scheme} (f : X ⟶ Y) [H : is_open_immersion f] : + opens X.carrier ⥤ opens Y.carrier := H.open_functor + lemma image_basic_open {X Y : Scheme} (f : X ⟶ Y) [H : is_open_immersion f] {U : opens X.carrier} (r : X.presheaf.obj (op U)) : H.base_open.is_open_map.functor.obj (X.basic_open r) = Y.basic_open (H.inv_app U r) := @@ -1618,10 +1628,10 @@ end /-- The image of an open immersion as an open set. -/ @[simps] -def opens_range (f : X ⟶ Y) [H : is_open_immersion f] : opens Y.carrier := +def hom.opens_range {X Y : Scheme} (f : X ⟶ Y) [H : is_open_immersion f] : opens Y.carrier := ⟨_, H.base_open.open_range⟩ -end is_open_immersion +end Scheme /-- The functor taking open subsets of `X` to open subschemes of `X`. -/ @[simps obj_left obj_hom map_left] @@ -1792,7 +1802,7 @@ by { delta morphism_restrict, apply_instance } lemma morphism_restrict_base_coe {X Y : Scheme} (f : X ⟶ Y) (U : opens Y.carrier) (x) : @coe U Y.carrier _ ((f ∣_ U).1.base x) = f.1.base x.1 := -congr_arg (λ f, PresheafedSpace.hom.base (subtype.val f) x) (morphism_restrict_ι f U) +congr_arg (λ f, PresheafedSpace.hom.base (LocallyRingedSpace.hom.val f) x) (morphism_restrict_ι f U) lemma image_morphism_restrict_preimage {X Y : Scheme} (f : X ⟶ Y) (U : opens Y.carrier) (V : opens U) : @@ -1848,9 +1858,9 @@ end along the immersion. -/ def morphism_restrict_opens_range {X Y U : Scheme} (f : X ⟶ Y) (g : U ⟶ Y) [hg : is_open_immersion g] : - arrow.mk (f ∣_ is_open_immersion.opens_range g) ≅ arrow.mk (pullback.snd : pullback f g ⟶ _) := + arrow.mk (f ∣_ g.opens_range) ≅ arrow.mk (pullback.snd : pullback f g ⟶ _) := begin - let V : opens Y.carrier := is_open_immersion.opens_range g, + let V : opens Y.carrier := g.opens_range, let e := is_open_immersion.iso_of_range_eq g (Y.of_restrict V.open_embedding) (by exact subtype.range_coe.symm), let t : pullback f g ⟶ pullback f (Y.of_restrict V.open_embedding) := @@ -1899,7 +1909,7 @@ def morphism_restrict_restrict_basic_open {X Y : Scheme} (f : X ⟶ Y) (U : open begin refine morphism_restrict_restrict _ _ _ ≪≫ morphism_restrict_eq _ _, have e := Scheme.preimage_basic_open (Y.of_restrict U.open_embedding) r, - erw [Scheme.of_restrict_coe_c_app, opens.adjunction_counit_app_self, eq_to_hom_op] at e, + erw [Scheme.of_restrict_val_c_app, opens.adjunction_counit_app_self, eq_to_hom_op] at e, rw [← (Y.restrict U.open_embedding).basic_open_res_eq _ (eq_to_hom U.inclusion_map_eq_top).op, ← comp_apply], erw ← Y.presheaf.map_comp, From fa0bbc7973fbbf5d738f070d628e1d02b3682fb8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 24 Aug 2022 09:35:14 +0000 Subject: [PATCH 011/828] feat(algebra/order/floor): `floor`/`ceil` are preserved under order ring homs (#16025) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `⌊f a⌋ = ⌊a⌋`, `⌈f a⌉ = ⌈a⌉` and similar for `f` a strictly monotone ring homomorphism. --- src/algebra/order/floor.lean | 60 +++++++++++++++++++++++++++++-- src/algebra/order/hom/monoid.lean | 10 ++++++ 2 files changed, 67 insertions(+), 3 deletions(-) diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index 87ec96e477564..aaa5276505d00 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -37,8 +37,6 @@ for `nnnorm`. ## TODO -Some `nat.floor` and `nat.ceil` lemmas require `linear_ordered_ring α`. Is `has_ordered_sub` enough? - `linear_ordered_ring`/`linear_ordered_semiring` can be relaxed to `order_ring`/`order_semiring` in many lemmas. @@ -48,7 +46,7 @@ rounding, floor, ceil -/ open set -variables {α : Type*} +variables {F α β : Type*} /-! ### Floor semiring -/ @@ -737,6 +735,62 @@ end end round +namespace nat +variables [linear_ordered_semiring α] [linear_ordered_semiring β] [floor_semiring α] + [floor_semiring β] [ring_hom_class F α β] {a : α} {b : β} +include β + +lemma floor_congr (h : ∀ n : ℕ, (n : α) ≤ a ↔ (n : β) ≤ b) : ⌊a⌋₊ = ⌊b⌋₊ := +begin + have h₀ : 0 ≤ a ↔ 0 ≤ b := by simpa only [cast_zero] using h 0, + obtain ha | ha := lt_or_le a 0, + { rw [floor_of_nonpos ha.le, floor_of_nonpos (le_of_not_le $ h₀.not.mp ha.not_le)] }, + exact (le_floor $ (h _).1 $ floor_le ha).antisymm (le_floor $ (h _).2 $ floor_le $ h₀.1 ha), +end + +lemma ceil_congr (h : ∀ n : ℕ, a ≤ n ↔ b ≤ n) : ⌈a⌉₊ = ⌈b⌉₊ := +(ceil_le.2 $ (h _).2 $ le_ceil _).antisymm $ ceil_le.2 $ (h _).1 $ le_ceil _ + +lemma map_floor (f : F) (hf : strict_mono f) (a : α) : ⌊f a⌋₊ = ⌊a⌋₊ := +floor_congr $ λ n, by rw [←map_nat_cast f, hf.le_iff_le] + +lemma map_ceil (f : F) (hf : strict_mono f) (a : α) : ⌈f a⌉₊ = ⌈a⌉₊ := +ceil_congr $ λ n, by rw [←map_nat_cast f, hf.le_iff_le] + +end nat + +namespace int +variables [linear_ordered_ring α] [linear_ordered_ring β] [floor_ring α] [floor_ring β] + [ring_hom_class F α β] {a : α} {b : β} +include β + +lemma floor_congr (h : ∀ n : ℤ, (n : α) ≤ a ↔ (n : β) ≤ b) : ⌊a⌋ = ⌊b⌋ := +(le_floor.2 $ (h _).1 $ floor_le _).antisymm $ le_floor.2 $ (h _).2 $ floor_le _ + +lemma ceil_congr (h : ∀ n : ℤ, a ≤ n ↔ b ≤ n) : ⌈a⌉ = ⌈b⌉ := +(ceil_le.2 $ (h _).2 $ le_ceil _).antisymm $ ceil_le.2 $ (h _).1 $ le_ceil _ + +lemma map_floor (f : F) (hf : strict_mono f) (a : α) : ⌊f a⌋ = ⌊a⌋ := +floor_congr $ λ n, by rw [←map_int_cast f, hf.le_iff_le] + +lemma map_ceil (f : F) (hf : strict_mono f) (a : α) : ⌈f a⌉ = ⌈a⌉ := +ceil_congr $ λ n, by rw [←map_int_cast f, hf.le_iff_le] + +lemma map_fract (f : F) (hf : strict_mono f) (a : α) : fract (f a) = f (fract a) := +by simp_rw [fract, map_sub, map_int_cast, map_floor _ hf] + +end int + +namespace int +variables [linear_ordered_field α] [linear_ordered_field β] [floor_ring α] [floor_ring β] + [ring_hom_class F α β] {a : α} {b : β} +include β + +lemma map_round (f : F) (hf : strict_mono f) (a : α) : round (f a) = round a := +by simp_rw [round, ←map_floor _ hf, map_add, one_div, map_inv₀, map_bit0, map_one] + +end int + variables {α} [linear_ordered_ring α] [floor_ring α] /-! #### A floor ring as a floor semiring -/ diff --git a/src/algebra/order/hom/monoid.lean b/src/algebra/order/hom/monoid.lean index c4ad367c017c9..5d9ab16c0c6de 100644 --- a/src/algebra/order/hom/monoid.lean +++ b/src/algebra/order/hom/monoid.lean @@ -161,6 +161,16 @@ instance [order_monoid_with_zero_hom_class F α β] : has_coe_t F (α →*₀o end monoid_with_zero +section ordered_add_comm_monoid +variables [ordered_add_comm_monoid α] [ordered_add_comm_monoid β] [order_add_monoid_hom_class F α β] + (f : F) {a : α} +include β + +lemma map_nonneg (ha : 0 ≤ a) : 0 ≤ f a := by { rw ←map_zero f, exact order_hom_class.mono _ ha } +lemma map_nonpos (ha : a ≤ 0) : f a ≤ 0 := by { rw ←map_zero f, exact order_hom_class.mono _ ha } + +end ordered_add_comm_monoid + namespace order_monoid_hom section preorder variables [preorder α] [preorder β] [preorder γ] [preorder δ] [mul_one_class α] From 7f06271e7b45dced7cd32cecf14c809c177a4797 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 24 Aug 2022 09:35:15 +0000 Subject: [PATCH 012/828] feat(algebraic_geometry/limits): Empty scheme is initial. (#16086) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/algebraic_geometry/Scheme.lean | 12 +- src/algebraic_geometry/limits.lean | 133 +++++++++++++++++++++ src/algebraic_geometry/open_immersion.lean | 32 +++++ src/category_theory/sites/sheaf.lean | 4 + 4 files changed, 178 insertions(+), 3 deletions(-) create mode 100644 src/algebraic_geometry/limits.lean diff --git a/src/algebraic_geometry/Scheme.lean b/src/algebraic_geometry/Scheme.lean index b6dde5bd50afa..d3ba561ccede5 100644 --- a/src/algebraic_geometry/Scheme.lean +++ b/src/algebraic_geometry/Scheme.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import algebraic_geometry.Spec +import algebra.category.Ring.constructions /-! # The category of schemes @@ -154,10 +155,15 @@ The spectrum, as a contravariant functor from commutative rings to schemes. map_comp' := λ R S T f g, by rw [unop_comp, Spec_map_comp] } /-- -The empty scheme, as `Spec 0`. +The empty scheme. -/ -def empty : Scheme := -Spec_obj (CommRing.of punit) +@[simps] +def {u} empty : Scheme.{u} := +{ carrier := Top.of pempty, + presheaf := (category_theory.functor.const _).obj (CommRing.of punit), + is_sheaf := presheaf.is_sheaf_of_is_terminal _ CommRing.punit_is_terminal, + local_ring := λ x, pempty.elim x, + local_affine := λ x, pempty.elim x } instance : has_emptyc Scheme := ⟨empty⟩ diff --git a/src/algebraic_geometry/limits.lean b/src/algebraic_geometry/limits.lean new file mode 100644 index 0000000000000..21a2b67b86130 --- /dev/null +++ b/src/algebraic_geometry/limits.lean @@ -0,0 +1,133 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import algebraic_geometry.pullbacks +import algebraic_geometry.AffineScheme +import category_theory.limits.constructions.finite_products_of_binary_products +import category_theory.limits.constructions.equalizers + +/-! +# (Co)Limits of Schemes + +We construct various limits and colimits in the category of schemes. + +* The existence of fibred products was shown in `algebraic_geometry/pullbacks.lean`. +* `Spec ℤ` is the terminal object. +* The preceding two results imply that `Scheme` has all finite limits. +* The empty scheme is the (strict) initial object. + +## Todo + +* Coproducts exists (and the forgetful functors preserve them). + +-/ + +universe u + +open category_theory category_theory.limits opposite topological_space + +namespace algebraic_geometry + +/-- `Spec ℤ` is the terminal object in the category of schemes. -/ +noncomputable +def Spec_Z_is_terminal : is_terminal (Scheme.Spec.obj (op $ CommRing.of ℤ)) := +@@is_terminal.is_terminal_obj _ _ Scheme.Spec _ infer_instance + (terminal_op_of_initial CommRing.Z_is_initial) + +instance : has_terminal Scheme := has_terminal_of_has_terminal_of_preserves_limit Scheme.Spec + +instance : is_affine (⊤_ Scheme.{u}) := +is_affine_of_iso (preserves_terminal.iso Scheme.Spec).inv + +instance : has_finite_limits Scheme := +has_finite_limits_of_has_terminal_and_pullbacks + +section initial + +/-- The map from the empty scheme. -/ +@[simps] +def Scheme.empty_to (X : Scheme.{u}) : ∅ ⟶ X := +⟨{ base := ⟨λ x, pempty.elim x, by continuity⟩, + c := { app := λ U, CommRing.punit_is_terminal.from _ } }, λ x, pempty.elim x⟩ + +@[ext] +lemma Scheme.empty_ext {X : Scheme.{u}} (f g : ∅ ⟶ X) : f = g := +by { ext a, exact pempty.elim a } + +lemma Scheme.eq_empty_to {X : Scheme.{u}} (f : ∅ ⟶ X) : f = Scheme.empty_to X := +Scheme.empty_ext f (Scheme.empty_to X) + +instance (X : Scheme.{u}) : unique (∅ ⟶ X) := +⟨⟨Scheme.empty_to _⟩, λ _, Scheme.empty_ext _ _⟩ + +/-- The empty scheme is the initial object in the category of schemes. -/ +def empty_is_initial : is_initial (∅ : Scheme.{u}) := +is_initial.of_unique _ + +@[simp] +lemma empty_is_initial_to : empty_is_initial.to = Scheme.empty_to := rfl + +instance : is_empty Scheme.empty.carrier := +show is_empty pempty, by apply_instance + +instance Spec_punit_is_empty : is_empty (Scheme.Spec.obj (op $ CommRing.of punit)).carrier := +⟨prime_spectrum.punit⟩ + +@[priority 100] +instance is_open_immersion_of_is_empty {X Y : Scheme} (f : X ⟶ Y) [is_empty X.carrier] : + is_open_immersion f := +begin + apply_with is_open_immersion.of_stalk_iso { instances := ff }, + { apply open_embedding_of_continuous_injective_open, + { continuity }, + { rintro (i : X.carrier), exact is_empty_elim i }, + { intros U hU, convert is_open_empty, ext, apply (iff_false _).mpr, + exact λ x, is_empty_elim (show X.carrier, from x.some) } }, + { rintro (i : X.carrier), exact is_empty_elim i } +end + +@[priority 100] +instance is_iso_of_is_empty {X Y : Scheme} (f : X ⟶ Y) [is_empty Y.carrier] : is_iso f := +begin + haveI : is_empty X.carrier := ⟨λ x, is_empty_elim (show Y.carrier, from f.1.base x)⟩, + haveI : epi f.1.base, + { rw Top.epi_iff_surjective, rintro (x : Y.carrier), exact is_empty_elim x }, + apply is_open_immersion.to_iso +end + +/-- A scheme is initial if its underlying space is empty . -/ +noncomputable +def is_initial_of_is_empty {X : Scheme} [is_empty X.carrier] : is_initial X := +empty_is_initial.of_iso (as_iso $ empty_is_initial.to _) + +/-- `Spec 0` is the initial object in the category of schemes. -/ +noncomputable +def Spec_punit_is_initial : is_initial (Scheme.Spec.obj (op $ CommRing.of punit)) := +empty_is_initial.of_iso (as_iso $ empty_is_initial.to _) + +@[priority 100] +instance is_affine_of_is_empty {X : Scheme} [is_empty X.carrier] : is_affine X := +is_affine_of_iso (inv (empty_is_initial.to X) ≫ + empty_is_initial.to (Scheme.Spec.obj (op $ CommRing.of punit))) + +instance : has_initial Scheme := +has_initial_of_unique Scheme.empty + +instance initial_is_empty : is_empty (⊥_ Scheme).carrier := +⟨λ x, ((initial.to Scheme.empty : _).1.base x).elim⟩ + +lemma bot_is_affine_open (X : Scheme) : is_affine_open (⊥ : opens X.carrier) := +begin + convert range_is_affine_open_of_open_immersion (initial.to X), + ext, + exact (false_iff _).mpr (λ x, is_empty_elim (show (⊥_ Scheme).carrier, from x.some)), +end + +instance : has_strict_initial_objects Scheme := +has_strict_initial_objects_of_initial_is_strict (λ A f, by apply_instance) + +end initial + +end algebraic_geometry diff --git a/src/algebraic_geometry/open_immersion.lean b/src/algebraic_geometry/open_immersion.lean index 44448d5e81040..dc1263efad0f6 100644 --- a/src/algebraic_geometry/open_immersion.lean +++ b/src/algebraic_geometry/open_immersion.lean @@ -1430,6 +1430,38 @@ instance of_is_iso [is_iso g] : is_open_immersion g := @@LocallyRingedSpace.is_open_immersion.of_is_iso _ (show is_iso ((induced_functor _).map g), by apply_instance) +lemma to_iso {X Y : Scheme} (f : X ⟶ Y) [h : is_open_immersion f] + [epi f.1.base] : is_iso f := +@@is_iso_of_reflects_iso _ _ f (Scheme.forget_to_LocallyRingedSpace ⋙ + LocallyRingedSpace.forget_to_SheafedSpace ⋙ SheafedSpace.forget_to_PresheafedSpace) + (@@PresheafedSpace.is_open_immersion.to_iso _ f.1 h _) _ + +lemma of_stalk_iso {X Y : Scheme} (f : X ⟶ Y) (hf : open_embedding f.1.base) + [∀ x, is_iso (PresheafedSpace.stalk_map f.1 x)] : is_open_immersion f := +SheafedSpace.is_open_immersion.of_stalk_iso f.1 hf + +lemma iff_stalk_iso {X Y : Scheme} (f : X ⟶ Y) : + is_open_immersion f ↔ open_embedding f.1.base ∧ ∀ x, is_iso (PresheafedSpace.stalk_map f.1 x) := +⟨λ H, ⟨H.1, by exactI infer_instance⟩, λ ⟨h₁, h₂⟩, @@is_open_immersion.of_stalk_iso f h₁ h₂⟩ + +lemma _root_.algebraic_geometry.is_iso_iff_is_open_immersion {X Y : Scheme} (f : X ⟶ Y) : + is_iso f ↔ is_open_immersion f ∧ epi f.1.base := +⟨λ H, by exactI ⟨infer_instance, infer_instance⟩, λ ⟨h₁, h₂⟩, @@is_open_immersion.to_iso f h₁ h₂⟩ + +lemma _root_.algebraic_geometry.is_iso_iff_stalk_iso {X Y : Scheme} (f : X ⟶ Y) : + is_iso f ↔ is_iso f.1.base ∧ ∀ x, is_iso (PresheafedSpace.stalk_map f.1 x) := +begin + rw [is_iso_iff_is_open_immersion, is_open_immersion.iff_stalk_iso, and_comm, ← and_assoc], + refine and_congr ⟨_, _⟩ iff.rfl, + { rintro ⟨h₁, h₂⟩, + convert_to is_iso (Top.iso_of_homeo (homeomorph.homeomorph_of_continuous_open + (equiv.of_bijective _ ⟨h₂.inj, (Top.epi_iff_surjective _).mp h₁⟩) + h₂.continuous h₂.is_open_map)).hom, + { ext, refl }, + { apply_instance } }, + { intro H, exactI ⟨infer_instance, (Top.homeo_of_iso (as_iso f.1.base)).open_embedding⟩ } +end + /-- A open immersion induces an isomorphism from the domain onto the image -/ def iso_restrict : X ≅ (Z.restrict H.base_open : _) := ⟨H.iso_restrict.hom, H.iso_restrict.inv, H.iso_restrict.hom_inv_id, H.iso_restrict.inv_hom_id⟩ diff --git a/src/category_theory/sites/sheaf.lean b/src/category_theory/sites/sheaf.lean index 6e8c10b4a65a2..0da437d8201d8 100644 --- a/src/category_theory/sites/sheaf.lean +++ b/src/category_theory/sites/sheaf.lean @@ -212,6 +212,10 @@ forall_congr $ λ a, ⟨presieve.is_sheaf_iso J (iso_whisker_right e _), variable (J) +lemma is_sheaf_of_is_terminal {X : A} (hX : is_terminal X) : + presheaf.is_sheaf J ((category_theory.functor.const _).obj X) := +λ _ _ _ _ _ _, ⟨hX.from _, λ _ _ _, hX.hom_ext _ _, λ _ _, hX.hom_ext _ _⟩ + end presheaf variables {C : Type u₁} [category.{v₁} C] From 214f197bce8a451a35010134598a8da2a4a2ad6a Mon Sep 17 00:00:00 2001 From: tb65536 Date: Wed, 24 Aug 2022 09:35:16 +0000 Subject: [PATCH 013/828] refactor(field_theory/*): Move and golf `alg_hom.fintype` (#16114) This PR moves `alg_hom.fintype` to `minpoly.lean` so that it can be used later (e.g., in `normal.lean`). --- src/field_theory/fixed.lean | 50 ----------------------------------- src/field_theory/minpoly.lean | 37 ++++++++++++++++++++++++++ 2 files changed, 37 insertions(+), 50 deletions(-) diff --git a/src/field_theory/fixed.lean b/src/field_theory/fixed.lean index 278167c25f643..599e38b164127 100644 --- a/src/field_theory/fixed.lean +++ b/src/field_theory/fixed.lean @@ -290,56 +290,6 @@ lemma cardinal_mk_alg_hom (K : Type u) (V : Type v) (W : Type w) cardinal.mk (V →ₐ[K] W) ≤ finrank W (V →ₗ[K] W) := cardinal_mk_le_finrank_of_linear_independent $ linear_independent_to_linear_map K V W -section alg_hom_fintype - -/-- A technical finiteness result. -/ -noncomputable def fintype.subtype_prod {E : Type*} {X : set E} (hX : X.finite) {L : Type*} - (F : E → multiset L) : fintype (Π x : X, {l : L // l ∈ F x}) := -by { classical, letI : fintype X := set.finite.fintype hX, exact pi.fintype} - -variables (E K : Type*) [field E] [field K] [algebra F E] [algebra F K] [finite_dimensional F E] - -/-- Function from Hom_K(E,L) to pi type Π (x : basis), roots of min poly of x -/ --- Marked as `noncomputable!` since this definition takes multiple seconds to compile, --- and isn't very computable in practice (since neither `finrank` nor `fin_basis` are). -noncomputable! def roots_of_min_poly_pi_type (φ : E →ₐ[F] K) - (x : set.range (finite_dimensional.fin_basis F E : _ → E)) : - {l : K // l ∈ (((minpoly F x.1).map (algebra_map F K)).roots : multiset K)} := -⟨φ x, begin - rw [polynomial.mem_roots_map (minpoly.ne_zero_of_finite_field_extension F x.val), - ← polynomial.alg_hom_eval₂_algebra_map, ← φ.map_zero], - exact congr_arg φ (minpoly.aeval F (x : E)), -end⟩ - -lemma aux_inj_roots_of_min_poly : function.injective (roots_of_min_poly_pi_type F E K) := -begin - intros f g h, - suffices : (f : E →ₗ[F] K) = g, - { rw linear_map.ext_iff at this, - ext x, exact this x }, - rw function.funext_iff at h, - apply linear_map.ext_on (finite_dimensional.fin_basis F E).span_eq, - rintro e he, - have := (h ⟨e, he⟩), - apply_fun subtype.val at this, - exact this, -end - -/-- Given field extensions `E/F` and `K/F`, with `E/F` finite, there are finitely many `F`-algebra - homomorphisms `E →ₐ[K] K`. -/ -noncomputable instance alg_hom.fintype : fintype (E →ₐ[F] K) := -let n := finite_dimensional.finrank F E in -begin - let B : basis (fin n) F E := finite_dimensional.fin_basis F E, - let X := set.range (B : fin n → E), - have hX : X.finite := set.finite_range ⇑B, - refine @fintype.of_injective _ _ - (fintype.subtype_prod hX (λ e, ((minpoly F e).map (algebra_map F K)).roots)) _ - (aux_inj_roots_of_min_poly F E K), -end - -end alg_hom_fintype - noncomputable instance alg_equiv.fintype (K : Type u) (V : Type v) [field K] [field V] [algebra K V] [finite_dimensional K V] : fintype (V ≃ₐ[K] V) := diff --git a/src/field_theory/minpoly.lean b/src/field_theory/minpoly.lean index 6b5b7d3a9a744..2cfa3279a3de6 100644 --- a/src/field_theory/minpoly.lean +++ b/src/field_theory/minpoly.lean @@ -387,6 +387,43 @@ lemma sub_algebra_map {B : Type*} [comm_ring B] [algebra A B] {x : B} minpoly A (x - (algebra_map A B a)) = (minpoly A x).comp (X + C a) := by simpa [sub_eq_add_neg] using add_algebra_map hx (-a) +section alg_hom_fintype + +/-- A technical finiteness result. -/ +noncomputable def fintype.subtype_prod {E : Type*} {X : set E} (hX : X.finite) {L : Type*} + (F : E → multiset L) : fintype (Π x : X, {l : L // l ∈ F x}) := +let hX := finite.fintype hX in by exactI pi.fintype + +variables (F E K : Type*) [field F] [ring E] [comm_ring K] [is_domain K] + [algebra F E] [algebra F K] [finite_dimensional F E] + +/-- Function from Hom_K(E,L) to pi type Π (x : basis), roots of min poly of x -/ +-- Marked as `noncomputable!` since this definition takes multiple seconds to compile, +-- and isn't very computable in practice (since neither `finrank` nor `fin_basis` are). +noncomputable! def roots_of_min_poly_pi_type (φ : E →ₐ[F] K) + (x : range (finite_dimensional.fin_basis F E : _ → E)) : + {l : K // l ∈ (((minpoly F x.1).map (algebra_map F K)).roots : multiset K)} := +⟨φ x, by rw [mem_roots_map (minpoly.ne_zero_of_finite_field_extension F x.val), + subtype.val_eq_coe, ←aeval_def, aeval_alg_hom_apply, minpoly.aeval, map_zero]⟩ + +lemma aux_inj_roots_of_min_poly : injective (roots_of_min_poly_pi_type F E K) := +begin + intros f g h, + suffices : (f : E →ₗ[F] K) = g, + { rwa fun_like.ext'_iff at this ⊢ }, + rw funext_iff at h, + exact linear_map.ext_on (finite_dimensional.fin_basis F E).span_eq + (λ e he, subtype.ext_iff.mp (h ⟨e, he⟩)), +end + +/-- Given field extensions `E/F` and `K/F`, with `E/F` finite, there are finitely many `F`-algebra + homomorphisms `E →ₐ[K] K`. -/ +noncomputable instance alg_hom.fintype : fintype (E →ₐ[F] K) := +@fintype.of_injective _ _ (fintype.subtype_prod (finite_range (finite_dimensional.fin_basis F E)) + (λ e, ((minpoly F e).map (algebra_map F K)).roots)) _ (aux_inj_roots_of_min_poly F E K) + +end alg_hom_fintype + section gcd_domain variables {R S : Type*} (K L : Type*) [comm_ring R] [is_domain R] [normalized_gcd_monoid R] From 53b79573863973b99d320eb6f01a90643c4bc1c7 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Wed, 24 Aug 2022 09:35:18 +0000 Subject: [PATCH 014/828] =?UTF-8?q?feat(analysis/normed=5Fspace/basic):=20?= =?UTF-8?q?if=20`E`=20is=20a=20`normed=5Fspace`=20over=20`=E2=84=9A`=20the?= =?UTF-8?q?n=20`=E2=84=A4=20=E2=88=99=20e`=20is=20discrete=20for=20any=20`?= =?UTF-8?q?e`=20in=20`E`=20(#16135)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/algebra/module/basic.lean | 5 +++++ src/analysis/normed_space/basic.lean | 20 ++++++++++++++++++++ src/group_theory/subgroup/basic.lean | 4 ++++ 3 files changed, 29 insertions(+) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index 4f8f3d0f354db..abe2872e1df65 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -621,6 +621,11 @@ instance division_ring.to_no_zero_smul_divisors : no_zero_smul_divisors R M := end division_ring +@[priority 100] -- see note [lower instance priority] +instance rat_module.no_zero_smul_divisors [add_comm_group M] [module ℚ M] : + no_zero_smul_divisors ℤ M := +⟨λ k x h, by simpa [zsmul_eq_smul_cast ℚ k x] using h⟩ + end no_zero_smul_divisors @[simp] lemma nat.smul_one_eq_coe {R : Type*} [semiring R] (m : ℕ) : diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index c8086b2625ec0..28027c4731564 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -65,6 +65,10 @@ begin by rw [norm_inv, ← mul_assoc, mul_inv_cancel (mt norm_eq_zero.1 h), one_mul] } end +lemma norm_zsmul (α) [normed_field α] [normed_space α β] (n : ℤ) (x : β) : + ∥n • x∥ = ∥(n : α)∥ * ∥x∥ := +by rw [← norm_smul, ← int.smul_one_eq_coe, smul_assoc, one_smul] + @[simp] lemma abs_norm_eq_norm (z : β) : |∥z∥| = ∥z∥ := (abs_eq (norm_nonneg z)).mpr (or.inl rfl) @@ -158,6 +162,22 @@ theorem frontier_closed_ball [normed_space ℝ E] (x : E) {r : ℝ} (hr : r ≠ by rw [frontier, closure_closed_ball, interior_closed_ball x hr, closed_ball_diff_ball] +instance {E : Type*} [normed_add_comm_group E] [normed_space ℚ E] (e : E) : + discrete_topology $ add_subgroup.zmultiples e := +begin + rcases eq_or_ne e 0 with rfl | he, + { rw [add_subgroup.zmultiples_zero_eq_bot], apply_instance, }, + { rw [discrete_topology_iff_open_singleton_zero, is_open_induced_iff], + refine ⟨metric.ball 0 (∥e∥), metric.is_open_ball, _⟩, + ext ⟨x, hx⟩, + obtain ⟨k, rfl⟩ := add_subgroup.mem_zmultiples_iff.mp hx, + rw [mem_preimage, mem_ball_zero_iff, add_subgroup.coe_mk, mem_singleton_iff, + subtype.ext_iff, add_subgroup.coe_mk, add_subgroup.coe_zero, norm_zsmul ℚ k e, + int.norm_cast_rat, int.norm_eq_abs, ← int.cast_abs, zero_lt.mul_lt_iff_lt_one_left + (norm_pos_iff.mpr he), ← @int.cast_one ℝ _, int.cast_lt, int.abs_lt_one_iff, smul_eq_zero, + or_iff_left he], }, +end + /-- A (semi) normed real vector space is homeomorphic to the unit ball in the same space. This homeomorphism sends `x : E` to `(1 + ∥x∥²)^(- ½) • x`. diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index fdcbca0529a78..7e18b8d0da9f7 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2695,6 +2695,10 @@ by rw [zpowers_eq_closure, closure_le, set.singleton_subset_iff, set_like.mem_co @[simp, to_additive zmultiples_eq_bot] lemma zpowers_eq_bot {g : G} : zpowers g = ⊥ ↔ g = 1 := by rw [eq_bot_iff, zpowers_le, mem_bot] +@[simp, to_additive zmultiples_zero_eq_bot] lemma zpowers_one_eq_bot : + subgroup.zpowers (1 : G) = ⊥ := +subgroup.zpowers_eq_bot.mpr rfl + end subgroup namespace monoid_hom From 912a310aa9dcb3d3e8f247c5cf21b1b57598c541 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 24 Aug 2022 09:35:19 +0000 Subject: [PATCH 015/828] refactor(category_theory/morphism_property): stable_under_base_change (#16196) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR redefines `morphism_property.stable_under_base_change P`. This condition now states that for all pullback squares (`is_pullback`), the left map satisfies the property if the right map does. Then, it is automatic that the property `P` respects isos. A constructor is provided to take as an input the assumption that `P` respects isos and the previous condition that if a map `g` satisfies `P`, then `pullback.fst : pullback f g → _` satisfies it `P`. --- src/algebraic_geometry/morphisms/basic.lean | 5 +- .../limits/shapes/comm_sq.lean | 10 +++ src/category_theory/morphism_property.lean | 78 +++++++++++++------ 3 files changed, 66 insertions(+), 27 deletions(-) diff --git a/src/algebraic_geometry/morphisms/basic.lean b/src/algebraic_geometry/morphisms/basic.lean index a1d4c72ad5966..31d5de9befe66 100644 --- a/src/algebraic_geometry/morphisms/basic.lean +++ b/src/algebraic_geometry/morphisms/basic.lean @@ -413,9 +413,10 @@ end lemma is_local.stable_under_base_change {P : affine_target_morphism_property} (hP : P.is_local) (hP' : P.stable_under_base_change) : - (target_affine_locally P).stable_under_base_change := + (target_affine_locally P).stable_under_base_change := +morphism_property.stable_under_base_change.mk (target_affine_locally_respects_iso hP.respects_iso) begin - introv X H, + intros X Y S f g H, rw (hP.target_affine_locally_is_local.open_cover_tfae (pullback.fst : pullback f g ⟶ X)).out 0 1, use S.affine_cover.pullback_cover f, intro i, diff --git a/src/category_theory/limits/shapes/comm_sq.lean b/src/category_theory/limits/shapes/comm_sq.lean index f3ff7bf35e394..a58826108945c 100644 --- a/src/category_theory/limits/shapes/comm_sq.lean +++ b/src/category_theory/limits/shapes/comm_sq.lean @@ -213,6 +213,13 @@ lemma of_iso_pullback (h : comm_sq fst snd f g) [has_pullback f g] (i : P ≅ pu of_is_limit' h (limits.is_limit.of_iso_limit (limit.is_limit _) (@pullback_cone.ext _ _ _ _ _ _ _ (pullback_cone.mk _ _ _) _ i w₁.symm w₂.symm).symm) +lemma of_horiz_is_iso [is_iso fst] [is_iso g] (sq : comm_sq fst snd f g) : + is_pullback fst snd f g := of_is_limit' sq +begin + refine pullback_cone.is_limit.mk _ (λ s, s.fst ≫ inv fst) (by tidy) (λ s, _) (by tidy), + simp only [← cancel_mono g, category.assoc, ← sq.w, is_iso.inv_hom_id_assoc, s.condition], +end + end is_pullback namespace is_pushout @@ -375,6 +382,9 @@ is_pushout.of_is_colimit (is_colimit.of_iso_colimit (limits.pullback_cone.is_limit_equiv_is_colimit_unop h.flip.cone h.flip.is_limit) h.to_comm_sq.flip.cone_unop) +lemma of_vert_is_iso [is_iso snd] [is_iso f] (sq : comm_sq fst snd f g) : + is_pullback fst snd f g := is_pullback.flip (of_horiz_is_iso sq.flip) + end is_pullback namespace is_pushout diff --git a/src/category_theory/morphism_property.lean b/src/category_theory/morphism_property.lean index e05ab6edd756b..31947a267b336 100644 --- a/src/category_theory/morphism_property.lean +++ b/src/category_theory/morphism_property.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang -/ import category_theory.limits.shapes.pullbacks import category_theory.arrow +import category_theory.limits.shapes.comm_sq /-! # Properties of morphisms @@ -15,7 +16,8 @@ The following meta-properties are defined * `respects_iso`: `P` respects isomorphisms if `P f → P (e ≫ f)` and `P f → P (f ≫ e)`, where `e` is an isomorphism. * `stable_under_composition`: `P` is stable under composition if `P f → P g → P (f ≫ g)`. -* `stable_under_base_change`: `P` is stable under base change if `P (Y ⟶ S) → P (X ×[S] Y ⟶ X)`. +* `stable_under_base_change`: `P` is stable under base change if in all pullback + squares, the left map satisfies `P` if the right map satisfies it. -/ @@ -56,8 +58,9 @@ def stable_under_inverse (P : morphism_property C) : Prop := /-- A morphism property is `stable_under_base_change` if the base change of such a morphism still falls in the class. -/ -def stable_under_base_change [has_pullbacks C] (P : morphism_property C) : Prop := -∀ ⦃X Y S : C⦄ (f : X ⟶ S) (g : Y ⟶ S), P g → P (pullback.fst : pullback f g ⟶ X) +def stable_under_base_change (P : morphism_property C) : Prop := +∀ ⦃X Y Y' S : C⦄ ⦃f : X ⟶ S⦄ ⦃g : Y ⟶ S⦄ ⦃f' : Y' ⟶ Y⦄ ⦃g' : Y' ⟶ X⦄ + (sq : is_pullback f' g' g f) (hg : P g), P g' lemma stable_under_composition.respects_iso {P : morphism_property C} (hP : stable_under_composition P) (hP' : ∀ {X Y} (e : X ≅ Y), P e.hom) : respects_iso P := @@ -82,42 +85,67 @@ lemma respects_iso.arrow_mk_iso_iff {P : morphism_property C} P f ↔ P g := hP.arrow_iso_iff e --- This is here to mirror `stable_under_base_change.snd`. -@[nolint unused_arguments] -lemma stable_under_base_change.fst [has_pullbacks C] {P : morphism_property C} - (hP : stable_under_base_change P) (hP' : respects_iso P) {X Y S : C} (f : X ⟶ S) - (g : Y ⟶ S) (H : P g) : P (pullback.fst : pullback f g ⟶ X) := -hP f g H +lemma respects_iso.of_respects_arrow_iso (P : morphism_property C) + (hP : ∀ (f g : arrow C) (e : f ≅ g) (hf : P f.hom), P g.hom) : respects_iso P := +begin + split, + { intros X Y Z e f hf, + refine hP (arrow.mk f) (arrow.mk (e.hom ≫ f)) (arrow.iso_mk e.symm (iso.refl _) _) hf, + dsimp, + simp only [iso.inv_hom_id_assoc, category.comp_id], }, + { intros X Y Z e f hf, + refine hP (arrow.mk f) (arrow.mk (f ≫ e.hom)) (arrow.iso_mk (iso.refl _) e _) hf, + dsimp, + simp only [category.id_comp], }, +end -lemma stable_under_base_change.snd [has_pullbacks C] {P : morphism_property C} - (hP : stable_under_base_change P) (hP' : respects_iso P) {X Y S : C} (f : X ⟶ S) - (g : Y ⟶ S) (H : P f) : P (pullback.snd : pullback f g ⟶ Y) := +lemma stable_under_base_change.mk {P : morphism_property C} [has_pullbacks C] + (hP₁ : respects_iso P) + (hP₂ : ∀ (X Y S : C) (f : X ⟶ S) (g : Y ⟶ S) (hg : P g), P (pullback.fst : pullback f g ⟶ X)) : + stable_under_base_change P := λ X Y Y' S f g f' g' sq hg, begin - rw [← pullback_symmetry_hom_comp_fst, hP'.cancel_left_is_iso], - exact hP g f H + let e := sq.flip.iso_pullback, + rw [← hP₁.cancel_left_is_iso e.inv, sq.flip.iso_pullback_inv_fst], + exact hP₂ _ _ _ f g hg, end +lemma stable_under_base_change.respects_iso {P : morphism_property C} + (hP : stable_under_base_change P) : respects_iso P := +begin + apply respects_iso.of_respects_arrow_iso, + intros f g e, + exact hP (is_pullback.of_horiz_is_iso (comm_sq.mk e.inv.w)), +end + +lemma stable_under_base_change.fst {P : morphism_property C} + (hP : stable_under_base_change P) {X Y S : C} (f : X ⟶ S) (g : Y ⟶ S) [has_pullback f g] + (H : P g) : P (pullback.fst : pullback f g ⟶ X) := +hP (is_pullback.of_has_pullback f g).flip H + +lemma stable_under_base_change.snd {P : morphism_property C} + (hP : stable_under_base_change P) {X Y S : C} (f : X ⟶ S) (g : Y ⟶ S) [has_pullback f g] + (H : P f) : P (pullback.snd : pullback f g ⟶ Y) := +hP (is_pullback.of_has_pullback f g) H + lemma stable_under_base_change.base_change_obj [has_pullbacks C] {P : morphism_property C} - (hP : stable_under_base_change P) (hP' : respects_iso P) {S S' : C} (f : S' ⟶ S) + (hP : stable_under_base_change P) {S S' : C} (f : S' ⟶ S) (X : over S) (H : P X.hom) : P ((base_change f).obj X).hom := -hP.snd hP' X.hom f H +hP.snd X.hom f H lemma stable_under_base_change.base_change_map [has_pullbacks C] {P : morphism_property C} - (hP : stable_under_base_change P) (hP' : respects_iso P) {S S' : C} (f : S' ⟶ S) + (hP : stable_under_base_change P) {S S' : C} (f : S' ⟶ S) {X Y : over S} (g : X ⟶ Y) (H : P g.left) : P ((base_change f).map g).left := begin let e := pullback_right_pullback_fst_iso Y.hom f g.left ≪≫ pullback.congr_hom (g.w.trans (category.comp_id _)) rfl, have : e.inv ≫ pullback.snd = ((base_change f).map g).left, { apply pullback.hom_ext; dsimp; simp }, - rw [← this, hP'.cancel_left_is_iso], - apply hP.snd hP', - exact H + rw [← this, hP.respects_iso.cancel_left_is_iso], + exact hP.snd _ _ H, end lemma stable_under_base_change.pullback_map [has_pullbacks C] {P : morphism_property C} - (hP : stable_under_base_change P) (hP' : respects_iso P) - (hP'' : stable_under_composition P) {S X X' Y Y' : C} + (hP : stable_under_base_change P) (hP' : stable_under_composition P) {S X X' Y Y' : C} {f : X ⟶ S} {g : Y ⟶ S} {f' : X' ⟶ S} {g' : Y' ⟶ S} {i₁ : X ⟶ X'} {i₂ : Y ⟶ Y'} (h₁ : P i₁) (h₂ : P i₂) (e₁ : f = i₁ ≫ f') (e₂ : g = i₂ ≫ g') : P (pullback.map f g f' g' i₁ i₂ (𝟙 _) @@ -131,9 +159,9 @@ begin ((base_change g').map (over.hom_mk _ e₁.symm : over.mk f ⟶ over.mk f')).left, { apply pullback.hom_ext; dsimp; simp }, rw this, - apply hP''; rw hP'.cancel_left_is_iso, - exacts [hP.base_change_map hP' _ (over.hom_mk _ e₂.symm : over.mk g ⟶ over.mk g') h₂, - hP.base_change_map hP' _ (over.hom_mk _ e₁.symm : over.mk f ⟶ over.mk f') h₁], + apply hP'; rw hP.respects_iso.cancel_left_is_iso, + exacts [hP.base_change_map _ (over.hom_mk _ e₂.symm : over.mk g ⟶ over.mk g') h₂, + hP.base_change_map _ (over.hom_mk _ e₁.symm : over.mk f ⟶ over.mk f') h₁], end /-- If `P : morphism_property C` and `F : C ⥤ D`, then From 7362d50a80c3e7f7a1fb5adb1c0b6ba593db103d Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 24 Aug 2022 09:35:21 +0000 Subject: [PATCH 016/828] feat(topology/sheaves/sheaf_condition/opens_le_cover): generalize universe (#16214) + Generalize universe levels in the sheaf condition `opens_le_cover` on topological spaces and its equivalence with the sheaf condition on sites, allowing three different universe parameters as in `Top.{w}`, `C : Type u` and `category.{v} C`. To be used in #15934. + Generalize universes for the sheaf condition `pairwise_intersection` on topological spaces. This sheaf condition also doesn't require any assumption on the category `C`, and its equivalence with `opens_le_cover` could also have universe levels at maximal generality; however, the proof `is_sheaf_opens_le_cover_iff_is_sheaf_pairwise_intersections` breaks due to the `small_category` restriction on [category_theory.functor.initial.is_limit_whisker_equiv](https://leanprover-community.github.io/mathlib_docs/category_theory/limits/final.html#category_theory.functor.initial.is_limit_whisker_equiv), which will take more work to fix, so we do not generalize the equivalence at this time. + Generalize universes for `Top.presheaf.pushforward`; pullback is not generalized because it would need `has_colimits_of_size`. `Top.sheaf.pushforward` is not yet generalized. + Fixate the universe level of [Top.presheaf](https://leanprover-community.github.io/mathlib_docs/topology/sheaves/presheaf.html#Top.presheaf) to be the same as [Top.sheaf](https://leanprover-community.github.io/mathlib_docs/topology/sheaves/sheaf.html#Top.sheaf). --- src/algebraic_geometry/locally_ringed_space.lean | 2 +- src/topology/sheaves/presheaf.lean | 8 ++++---- .../sheaves/sheaf_condition/opens_le_cover.lean | 10 +++++----- .../sheaf_condition/pairwise_intersections.lean | 15 +++++++++------ 4 files changed, 19 insertions(+), 16 deletions(-) diff --git a/src/algebraic_geometry/locally_ringed_space.lean b/src/algebraic_geometry/locally_ringed_space.lean index 74ee81b462f4b..036fc1011fc45 100644 --- a/src/algebraic_geometry/locally_ringed_space.lean +++ b/src/algebraic_geometry/locally_ringed_space.lean @@ -132,7 +132,7 @@ forget_to_SheafedSpace ⋙ SheafedSpace.forget _ @[simp] lemma comp_val {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).val = f.val ≫ g.val := rfl -@[simp] lemma comp_val_c {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) : +@[simp] lemma comp_val_c {X Y Z : LocallyRingedSpace.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).val.c = g.val.c ≫ (presheaf.pushforward _ g.val.base).map f.val.c := rfl lemma comp_val_c_app {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) (U : (opens Z)ᵒᵖ) : diff --git a/src/topology/sheaves/presheaf.lean b/src/topology/sheaves/presheaf.lean index 06147fc2c2434..a46874dd0f145 100644 --- a/src/topology/sheaves/presheaf.lean +++ b/src/topology/sheaves/presheaf.lean @@ -38,7 +38,7 @@ namespace Top /-- The category of `C`-valued presheaves on a (bundled) topological space `X`. -/ @[derive category, nolint has_nonempty_instance] -def presheaf (X : Top.{w}) := (opens X)ᵒᵖ ⥤ C +def presheaf (X : Top.{w}) : Type (max u v w) := (opens X)ᵒᵖ ⥤ C variables {C} @@ -222,16 +222,16 @@ variable (C) /-- The pushforward functor. -/ -def pushforward {X Y : Top.{v}} (f : X ⟶ Y) : X.presheaf C ⥤ Y.presheaf C := +def pushforward {X Y : Top.{w}} (f : X ⟶ Y) : X.presheaf C ⥤ Y.presheaf C := { obj := pushforward_obj f, map := @pushforward_map _ _ X Y f } @[simp] -lemma pushforward_map_app' {X Y : Top.{v}} (f : X ⟶ Y) +lemma pushforward_map_app' {X Y : Top.{w}} (f : X ⟶ Y) {ℱ 𝒢 : X.presheaf C} (α : ℱ ⟶ 𝒢) {U : (opens Y)ᵒᵖ} : ((pushforward C f).map α).app U = α.app (op $ (opens.map f).obj U.unop) := rfl -lemma id_pushforward {X : Top.{v}} : pushforward C (𝟙 X) = 𝟭 (X.presheaf C) := +lemma id_pushforward {X : Top.{w}} : pushforward C (𝟙 X) = 𝟭 (X.presheaf C) := begin apply category_theory.functor.ext, { intros, diff --git a/src/topology/sheaves/sheaf_condition/opens_le_cover.lean b/src/topology/sheaves/sheaf_condition/opens_le_cover.lean index 081f85a57a0bd..14d8c30fb71e0 100644 --- a/src/topology/sheaves/sheaf_condition/opens_le_cover.lean +++ b/src/topology/sheaves/sheaf_condition/opens_le_cover.lean @@ -25,7 +25,7 @@ or equivalently whether we're looking at the first or second object in an equali * This is the definition Lurie uses in [Spectral Algebraic Geometry][LurieSAG]. -/ -universes v u +universes w v u noncomputable theory @@ -38,7 +38,7 @@ open topological_space.opens namespace Top variables {C : Type u} [category.{v} C] -variables {X : Top.{v}} (F : presheaf C X) {ι : Type v} (U : ι → opens X) +variables {X : Top.{w}} (F : presheaf C X) {ι : Type w} (U : ι → opens X) namespace presheaf @@ -48,7 +48,7 @@ namespace sheaf_condition The category of open sets contained in some element of the cover. -/ @[derive category] -def opens_le_cover : Type v := full_subcategory (λ (V : opens X), ∃ i, V ≤ U i) +def opens_le_cover : Type w := full_subcategory (λ (V : opens X), ∃ i, V ≤ U i) instance [inhabited ι] : inhabited (opens_le_cover U) := ⟨⟨⊥, default, bot_le⟩⟩ @@ -93,7 +93,7 @@ A presheaf is a sheaf if `F` sends the cone `(opens_le_cover_cocone U).op` to a mapping down to any `V` which is contained in some `U i`.) -/ def is_sheaf_opens_le_cover : Prop := -∀ ⦃ι : Type v⦄ (U : ι → opens X), nonempty (is_limit (F.map_cone (opens_le_cover_cocone U).op)) +∀ ⦃ι : Type w⦄ (U : ι → opens X), nonempty (is_limit (F.map_cone (opens_le_cover_cocone U).op)) namespace sheaf_condition @@ -215,7 +215,7 @@ in terms of a limit diagram over all `{ V : opens X // ∃ i, V ≤ U i }` is equivalent to the reformulation in terms of a limit diagram over `U i` and `U i ⊓ U j`. -/ -lemma is_sheaf_opens_le_cover_iff_is_sheaf_pairwise_intersections (F : presheaf C X) : +lemma is_sheaf_opens_le_cover_iff_is_sheaf_pairwise_intersections {X : Top.{v}} (F : presheaf C X) : F.is_sheaf_opens_le_cover ↔ F.is_sheaf_pairwise_intersections := forall₂_congr $ λ ι U, equiv.nonempty_congr $ calc is_limit (F.map_cone (opens_le_cover_cocone U).op) diff --git a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean index 098a78851bc74..c9f8c083ec7e5 100644 --- a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean +++ b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean @@ -33,7 +33,7 @@ We express this in two equivalent ways, as noncomputable theory -universes v u +universes w v u open topological_space open Top @@ -43,10 +43,11 @@ open category_theory.limits namespace Top.presheaf -variables {X : Top.{v}} - variables {C : Type u} [category.{v} C] +section +variables {X : Top.{w}} + /-- An alternative formulation of the sheaf condition (which we prove equivalent to the usual one below as @@ -56,7 +57,7 @@ A presheaf is a sheaf if `F` sends the cone `(pairwise.cocone U).op` to a limit (Recall `pairwise.cocone U` has cone point `supr U`, mapping down to the `U i` and the `U i ⊓ U j`.) -/ def is_sheaf_pairwise_intersections (F : presheaf C X) : Prop := -∀ ⦃ι : Type v⦄ (U : ι → opens X), nonempty (is_limit (F.map_cone (pairwise.cocone U).op)) +∀ ⦃ι : Type w⦄ (U : ι → opens X), nonempty (is_limit (F.map_cone (pairwise.cocone U).op)) /-- An alternative formulation of the sheaf condition @@ -68,14 +69,16 @@ A presheaf is a sheaf if `F` preserves the limit of `pairwise.diagram U`. `U i ⊓ U j` mapping into the open sets `U i`. This diagram has limit `supr U`.) -/ def is_sheaf_preserves_limit_pairwise_intersections (F : presheaf C X) : Prop := -∀ ⦃ι : Type v⦄ (U : ι → opens X), nonempty (preserves_limit (pairwise.diagram U).op F) +∀ ⦃ι : Type w⦄ (U : ι → opens X), nonempty (preserves_limit (pairwise.diagram U).op F) + +end /-! The remainder of this file shows that these conditions are equivalent to the usual sheaf condition. -/ -variables [has_products.{v} C] +variables {X : Top.{v}} [has_products.{v} C] namespace sheaf_condition_pairwise_intersections From 9b75911144d8cbcad32543f296805a1ca18da7b0 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 24 Aug 2022 12:36:07 +0000 Subject: [PATCH 017/828] feat(algebra/order): lemmas about `cmp` and arithmetic operations (#16207) --- src/algebra/order/group.lean | 8 +++++++- src/algebra/order/monoid_lemmas.lean | 10 ++++++++++ src/algebra/order/ring.lean | 21 +++++++++++++++++++++ 3 files changed, 38 insertions(+), 1 deletion(-) diff --git a/src/algebra/order/group.lean b/src/algebra/order/group.lean index 9cfc82e3bd4f0..e41cf98dcf8ab 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group.lean @@ -768,7 +768,13 @@ end preorder end comm_group section linear_order -variables [group α] [linear_order α] [covariant_class α α (*) (≤)] +variables [group α] [linear_order α] + +@[simp, to_additive cmp_sub_zero] +lemma cmp_div_one' [covariant_class α α (swap (*)) (≤)] (a b : α) : cmp (a / b) 1 = cmp a b := +by rw [← cmp_mul_right' _ _ b, one_mul, div_mul_cancel'] + +variables [covariant_class α α (*) (≤)] section variable_names variables {a b c : α} diff --git a/src/algebra/order/monoid_lemmas.lean b/src/algebra/order/monoid_lemmas.lean index 629325700a6f4..441534f08ae2f 100644 --- a/src/algebra/order/monoid_lemmas.lean +++ b/src/algebra/order/monoid_lemmas.lean @@ -1047,6 +1047,16 @@ lemma strict_anti_on.mul_antitone' (hf : strict_anti_on f s) (hg : antitone_on g strict_anti_on (λ x, f x * g x) s := λ x hx y hy h, mul_lt_mul_of_lt_of_le (hf hx hy h) (hg hx hy h.le) +@[simp, to_additive cmp_add_left] +lemma cmp_mul_left' {α : Type*} [has_mul α] [linear_order α] [covariant_class α α (*) (<)] + (a b c : α) : cmp (a * b) (a * c) = cmp b c := +(strict_mono_id.const_mul' a).cmp_map_eq b c + +@[simp, to_additive cmp_add_right] +lemma cmp_mul_right' {α : Type*} [has_mul α] [linear_order α] [covariant_class α α (swap (*)) (<)] + (a b c : α) : cmp (a * c) (b * c) = cmp a b := +(strict_mono_id.mul_const' c).cmp_map_eq a b + end mono /-- diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 75f548d42e012..25a42974aacfd 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -909,6 +909,21 @@ ordered_semiring.to_char_zero end linear_ordered_semiring +section mono +variables [linear_ordered_semiring α] {a : α} + +local attribute [instance] linear_ordered_semiring.decidable_lt + +lemma cmp_mul_pos_left (ha : 0 < a) (b c : α) : + cmp (a * b) (a * c) = cmp b c := +(strict_mono_mul_left_of_pos ha).cmp_map_eq b c + +lemma cmp_mul_pos_right (ha : 0 < a) (b c : α) : + cmp (b * a) (c * a) = cmp b c := +(strict_mono_mul_right_of_pos ha).cmp_map_eq b c + +end mono + section linear_ordered_semiring variables [linear_ordered_semiring α] {a b c : α} @@ -1286,6 +1301,12 @@ lemma neg_one_lt_zero : -1 < (0:α) := neg_lt_zero.2 zero_lt_one @[simp] lemma mul_lt_mul_right_of_neg {a b c : α} (h : c < 0) : a * c < b * c ↔ b < a := (strict_anti_mul_right h).lt_iff_lt +lemma cmp_mul_neg_left {a : α} (ha : a < 0) (b c : α) : cmp (a * b) (a * c) = cmp c b := +(strict_anti_mul_left ha).cmp_map_eq b c + +lemma cmp_mul_neg_right {a : α} (ha : a < 0) (b c : α) : cmp (b * a) (c * a) = cmp c b := +(strict_anti_mul_right ha).cmp_map_eq b c + lemma sub_one_lt (a : α) : a - 1 < a := sub_lt_iff_lt_add.2 (lt_add_one a) From 474f96ac0fb72d493b7e79832250c9140b06d17f Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 24 Aug 2022 12:36:08 +0000 Subject: [PATCH 018/828] feat(geometry/euclidean/oriented_angle): more on relation to unoriented angles (#16215) Add more lemmas about the relation of oriented and unoriented angles; in particular, some involving `real.angle.to_real` and `real.angle.sign`. --- src/geometry/euclidean/oriented_angle.lean | 146 +++++++++++++++++++++ 1 file changed, 146 insertions(+) diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 495b2159f92fa..504bb7cc1b7aa 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -834,6 +834,113 @@ lemma oangle_eq_angle_or_eq_neg_angle {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : hb.oangle x y = -inner_product_geometry.angle x y := real.angle.cos_eq_real_cos_iff_eq_or_eq_neg.1 $ hb.cos_oangle_eq_cos_angle hx hy +/-- The unoriented angle between two nonzero vectors is the absolute value of the oriented angle, +converted to a real. -/ +lemma angle_eq_abs_oangle_to_real {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + inner_product_geometry.angle x y = |(hb.oangle x y).to_real| := +begin + have h0 := inner_product_geometry.angle_nonneg x y, + have hpi := inner_product_geometry.angle_le_pi x y, + rcases hb.oangle_eq_angle_or_eq_neg_angle hx hy with (h|h), + { rw [h, eq_comm, real.angle.abs_to_real_coe_eq_self_iff], + exact ⟨h0, hpi⟩ }, + { rw [h, eq_comm, real.angle.abs_to_real_neg_coe_eq_self_iff], + exact ⟨h0, hpi⟩ } +end + +/-- If the sign of the oriented angle between two vectors is zero, either one of the vectors is +zero or the unoriented angle is 0 or π. -/ +lemma eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero {x y : V} + (h : (hb.oangle x y).sign = 0) : + x = 0 ∨ y = 0 ∨ inner_product_geometry.angle x y = 0 ∨ inner_product_geometry.angle x y = π := +begin + by_cases hx : x = 0, { simp [hx] }, + by_cases hy : y = 0, { simp [hy] }, + rw hb.angle_eq_abs_oangle_to_real hx hy, + rw real.angle.sign_eq_zero_iff at h, + rcases h with h|h; + simp [h, real.pi_pos.le] +end + +/-- If two unoriented angles are equal, and the signs of the corresponding oriented angles are +equal, then the oriented angles are equal (even in degenerate cases). -/ +lemma oangle_eq_of_angle_eq_of_sign_eq {w x y z : V} + (h : inner_product_geometry.angle w x = inner_product_geometry.angle y z) + (hs : (hb.oangle w x).sign = (hb.oangle y z).sign) : hb.oangle w x = hb.oangle y z := +begin + by_cases h0 : (w = 0 ∨ x = 0) ∨ (y = 0 ∨ z = 0), + { have hs' : (hb.oangle w x).sign = 0 ∧ (hb.oangle y z).sign = 0, + { rcases h0 with (rfl|rfl)|rfl|rfl, + { simpa using hs.symm }, + { simpa using hs.symm }, + { simpa using hs }, + { simpa using hs } }, + rcases hs' with ⟨hswx, hsyz⟩, + have h' : inner_product_geometry.angle w x = π / 2 ∧ inner_product_geometry.angle y z = π / 2, + { rcases h0 with (rfl|rfl)|rfl|rfl, + { simpa using h.symm }, + { simpa using h.symm }, + { simpa using h }, + { simpa using h } }, + rcases h' with ⟨hwx, hyz⟩, + have hpi : π / 2 ≠ π, + { intro hpi, + rw [div_eq_iff, eq_comm, ←sub_eq_zero, mul_two, add_sub_cancel] at hpi, + { exact real.pi_pos.ne.symm hpi }, + { exact two_ne_zero } }, + have h0wx : (w = 0 ∨ x = 0), + { have h0' := hb.eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero hswx, + simpa [hwx, real.pi_pos.ne.symm, hpi] using h0' }, + have h0yz : (y = 0 ∨ z = 0), + { have h0' := hb.eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero hsyz, + simpa [hyz, real.pi_pos.ne.symm, hpi] using h0' }, + rcases h0wx with h0wx|h0wx; rcases h0yz with h0yz|h0yz; + simp [h0wx, h0yz] }, + { push_neg at h0, + rw real.angle.eq_iff_abs_to_real_eq_of_sign_eq hs, + rwa [hb.angle_eq_abs_oangle_to_real h0.1.1 h0.1.2, + hb.angle_eq_abs_oangle_to_real h0.2.1 h0.2.2] at h } +end + +/-- If the signs of two oriented angles between nonzero vectors are equal, the oriented angles are +equal if and only if the unoriented angles are equal. -/ +lemma oangle_eq_iff_angle_eq_of_sign_eq {w x y z : V} (hw : w ≠ 0) (hx : x ≠ 0) (hy : y ≠ 0) + (hz : z ≠ 0) (hs : (hb.oangle w x).sign = (hb.oangle y z).sign) : + inner_product_geometry.angle w x = inner_product_geometry.angle y z ↔ + hb.oangle w x = hb.oangle y z := +begin + refine ⟨λ h, hb.oangle_eq_of_angle_eq_of_sign_eq h hs, λ h, _⟩, + rw [hb.angle_eq_abs_oangle_to_real hw hx, hb.angle_eq_abs_oangle_to_real hy hz, h] +end + +/-- The oriented angle between two nonzero vectors is zero if and only if the unoriented angle +is zero. -/ +lemma oangle_eq_zero_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + hb.oangle x y = 0 ↔ inner_product_geometry.angle x y = 0 := +begin + refine ⟨λ h, _, λ h, _⟩, + { simpa [hb.angle_eq_abs_oangle_to_real hx hy] }, + { have ha := hb.oangle_eq_angle_or_eq_neg_angle hx hy, + rw h at ha, + simpa using ha } +end + +/-- The oriented angle between two vectors is `π` if and only if the unoriented angle is `π`. -/ +lemma oangle_eq_pi_iff_angle_eq_pi {x y : V} : + hb.oangle x y = π ↔ inner_product_geometry.angle x y = π := +begin + by_cases hx : x = 0, { simp [hx, real.angle.pi_ne_zero.symm, div_eq_mul_inv, mul_right_eq_self₀, + not_or_distrib, real.pi_ne_zero], norm_num }, + by_cases hy : y = 0, { simp [hy, real.angle.pi_ne_zero.symm, div_eq_mul_inv, mul_right_eq_self₀, + not_or_distrib, real.pi_ne_zero], norm_num }, + refine ⟨λ h, _, λ h, _⟩, + { rw [hb.angle_eq_abs_oangle_to_real hx hy, h], + simp [real.pi_pos.le] }, + { have ha := hb.oangle_eq_angle_or_eq_neg_angle hx hy, + rw h at ha, + simpa using ha } +end + end orthonormal namespace orientation @@ -1285,4 +1392,43 @@ lemma oangle_eq_angle_or_eq_neg_angle {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : o.oangle x y = -inner_product_geometry.angle x y := (ob).oangle_eq_angle_or_eq_neg_angle hx hy +/-- The unoriented angle between two nonzero vectors is the absolute value of the oriented angle, +converted to a real. -/ +lemma angle_eq_abs_oangle_to_real {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + inner_product_geometry.angle x y = |(o.oangle x y).to_real| := +(ob).angle_eq_abs_oangle_to_real hx hy + +/-- If the sign of the oriented angle between two vectors is zero, either one of the vectors is +zero or the unoriented angle is 0 or π. -/ +lemma eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero {x y : V} + (h : (o.oangle x y).sign = 0) : + x = 0 ∨ y = 0 ∨ inner_product_geometry.angle x y = 0 ∨ inner_product_geometry.angle x y = π := +(ob).eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero h + +/-- If two unoriented angles are equal, and the signs of the corresponding oriented angles are +equal, then the oriented angles are equal (even in degenerate cases). -/ +lemma oangle_eq_of_angle_eq_of_sign_eq {w x y z : V} + (h : inner_product_geometry.angle w x = inner_product_geometry.angle y z) + (hs : (o.oangle w x).sign = (o.oangle y z).sign) : o.oangle w x = o.oangle y z := +(ob).oangle_eq_of_angle_eq_of_sign_eq h hs + +/-- If the signs of two oriented angles between nonzero vectors are equal, the oriented angles are +equal if and only if the unoriented angles are equal. -/ +lemma oangle_eq_iff_angle_eq_of_sign_eq {w x y z : V} (hw : w ≠ 0) (hx : x ≠ 0) (hy : y ≠ 0) + (hz : z ≠ 0) (hs : (o.oangle w x).sign = (o.oangle y z).sign) : + inner_product_geometry.angle w x = inner_product_geometry.angle y z ↔ + o.oangle w x = o.oangle y z := +(ob).oangle_eq_iff_angle_eq_of_sign_eq hw hx hy hz hs + +/-- The oriented angle between two nonzero vectors is zero if and only if the unoriented angle +is zero. -/ +lemma oangle_eq_zero_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + o.oangle x y = 0 ↔ inner_product_geometry.angle x y = 0 := +(ob).oangle_eq_zero_iff_angle_eq_zero hx hy + +/-- The oriented angle between two vectors is `π` if and only if the unoriented angle is `π`. -/ +lemma oangle_eq_pi_iff_angle_eq_pi {x y : V} : + o.oangle x y = π ↔ inner_product_geometry.angle x y = π := +(ob).oangle_eq_pi_iff_angle_eq_pi + end orientation From 839b92fedff9981cf3fe1c1f623e04b0d127f57c Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 24 Aug 2022 12:36:10 +0000 Subject: [PATCH 019/828] feat(data/nat/enat): new file (#16217) Define `enat := with_top nat`, use notation. --- src/algebra/group/to_additive.lean | 2 +- src/analysis/calculus/affine_map.lean | 2 +- src/analysis/calculus/cont_diff.lean | 113 +++++++++--------- src/analysis/calculus/fderiv_analytic.lean | 2 +- src/analysis/calculus/inverse.lean | 14 +-- src/analysis/calculus/iterated_deriv.lean | 42 +++---- src/analysis/calculus/specific_functions.lean | 4 +- src/analysis/complex/real_deriv.lean | 4 +- src/analysis/convolution.lean | 8 +- .../inner_product_space/calculus.lean | 12 +- .../inner_product_space/euclidean_dist.lean | 2 +- src/analysis/special_functions/arsinh.lean | 4 +- .../special_functions/complex/log_deriv.lean | 2 +- src/analysis/special_functions/log/deriv.lean | 4 +- src/analysis/special_functions/pow_deriv.lean | 6 +- src/analysis/special_functions/sqrt.lean | 4 +- .../trigonometric/arctan_deriv.lean | 4 +- .../trigonometric/complex_deriv.lean | 2 +- .../trigonometric/inverse_deriv.lean | 14 +-- src/data/nat/cast.lean | 26 ---- src/data/nat/enat.lean | 69 +++++++++++ src/data/nat/part_enat.lean | 34 +++--- .../polynomial/degree/trailing_degree.lean | 36 +++--- src/geometry/manifold/algebra/monoid.lean | 4 +- src/geometry/manifold/cont_mdiff.lean | 22 ++-- src/geometry/manifold/cont_mdiff_map.lean | 2 +- src/geometry/manifold/derivation_bundle.lean | 2 +- src/geometry/manifold/diffeomorph.lean | 2 +- src/geometry/manifold/instances/sphere.lean | 2 +- src/geometry/manifold/partition_of_unity.lean | 4 +- .../smooth_manifold_with_corners.lean | 6 +- src/probability/hitting_time.lean | 2 +- src/ring_theory/int/basic.lean | 2 +- .../unique_factorization_domain.lean | 2 +- 34 files changed, 253 insertions(+), 207 deletions(-) create mode 100644 src/data/nat/enat.lean diff --git a/src/algebra/group/to_additive.lean b/src/algebra/group/to_additive.lean index 6acfad5044ff4..310d1a7ce4d8d 100644 --- a/src/algebra/group/to_additive.lean +++ b/src/algebra/group/to_additive.lean @@ -411,7 +411,7 @@ There are some exceptions to this heuristic: * If an identifier has attribute `@[to_additive_ignore_args n1 n2 ...]` then all the arguments in positions `n1`, `n2`, ... will not be checked for unapplied identifiers (start counting from 1). For example, `cont_mdiff_map` has attribute `@[to_additive_ignore_args 21]`, which means - that its 21st argument `(n : with_top ℕ)` can contain `ℕ` + that its 21st argument `(n : ℕ∞)` can contain `ℕ` (usually in the form `has_top.top ℕ ...`) and still be additivized. So `@has_mul.mul (C^∞⟮I, N; I', G⟯) _ f g` will be additivized. diff --git a/src/analysis/calculus/affine_map.lean b/src/analysis/calculus/affine_map.lean index 10a08f23164a3..97f2e2a756d3a 100644 --- a/src/analysis/calculus/affine_map.lean +++ b/src/analysis/calculus/affine_map.lean @@ -24,7 +24,7 @@ variables [normed_add_comm_group V] [normed_space 𝕜 V] variables [normed_add_comm_group W] [normed_space 𝕜 W] /-- A continuous affine map between normed vector spaces is smooth. -/ -lemma cont_diff {n : with_top ℕ} (f : V →A[𝕜] W) : +lemma cont_diff {n : ℕ∞} (f : V →A[𝕜] W) : cont_diff 𝕜 n f := begin rw f.decomp, diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index ec83922e515dc..98f8bed33ad57 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel import analysis.calculus.mean_value import analysis.normed_space.multilinear import analysis.calculus.formal_multilinear_series +import data.nat.enat import tactic.congrm /-! @@ -150,7 +151,7 @@ continuous linear equivs. We use the notation `E [×n]→L[𝕜] F` for the space of continuous multilinear maps on `E^n` with values in `F`. This is the space in which the `n`-th derivative of a function from `E` to `F` lives. -In this file, we denote `⊤ : with_top ℕ` with `∞`. +In this file, we denote `⊤ : ℕ∞` with `∞`. ## Tags @@ -160,7 +161,7 @@ derivative, differentiability, higher derivative, `C^n`, multilinear, Taylor ser noncomputable theory open_locale classical big_operators nnreal -local notation `∞` := (⊤ : with_top ℕ) +local notation `∞` := (⊤ : ℕ∞) universes u v w @@ -176,7 +177,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] {X : Type*} [normed_add_comm_group X] [normed_space 𝕜 X] {s s₁ t u : set E} {f f₁ : E → F} {g : F → G} {x : E} {c : F} -{b : E × F → G} {m n : with_top ℕ} +{b : E × F → G} {m n : ℕ∞} /-! ### Functions with a Taylor series on a domain -/ @@ -185,12 +186,12 @@ variable {p : E → formal_multilinear_series 𝕜 E F} /-- `has_ftaylor_series_up_to_on n f p s` registers the fact that `p 0 = f` and `p (m+1)` is a derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a predicate analogous to `has_fderiv_within_at` but for higher order derivatives. -/ -structure has_ftaylor_series_up_to_on (n : with_top ℕ) +structure has_ftaylor_series_up_to_on (n : ℕ∞) (f : E → F) (p : E → formal_multilinear_series 𝕜 E F) (s : set E) : Prop := (zero_eq : ∀ x ∈ s, (p x 0).uncurry0 = f x) -(fderiv_within : ∀ (m : ℕ) (hm : (m : with_top ℕ) < n), ∀ x ∈ s, +(fderiv_within : ∀ (m : ℕ) (hm : (m : ℕ∞) < n), ∀ x ∈ s, has_fderiv_within_at (λ y, p y m) (p x m.succ).curry_left s x) -(cont : ∀ (m : ℕ) (hm : (m : with_top ℕ) ≤ n), continuous_on (λ x, p x m) s) +(cont : ∀ (m : ℕ) (hm : (m : ℕ∞) ≤ n), continuous_on (λ x, p x m) s) lemma has_ftaylor_series_up_to_on.zero_eq' (h : has_ftaylor_series_up_to_on n f p s) {x : E} (hx : x ∈ s) : @@ -269,7 +270,7 @@ begin (continuous_multilinear_curry_fin1 𝕜 E F (p x 1)) s x, by exact H.congr A (A x hx), rw linear_isometry_equiv.comp_has_fderiv_within_at_iff', - have : ((0 : ℕ) : with_top ℕ) < n := + have : ((0 : ℕ) : ℕ∞) < n := lt_of_lt_of_le (with_top.coe_lt_coe.2 nat.zero_lt_one) hn, convert h.fderiv_within _ this x hx, ext y v, @@ -349,8 +350,8 @@ begin refine ⟨H.zero_eq, H.fderiv_within 0 (with_top.coe_lt_coe.2 (nat.succ_pos n)), _⟩, split, { assume x hx, refl }, - { assume m (hm : (m : with_top ℕ) < n) x (hx : x ∈ s), - have A : (m.succ : with_top ℕ) < n.succ, + { assume m (hm : (m : ℕ∞) < n) x (hx : x ∈ s), + have A : (m.succ : ℕ∞) < n.succ, by { rw with_top.coe_lt_coe at ⊢ hm, exact nat.lt_succ_iff.mpr hm }, change has_fderiv_within_at ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm @@ -362,8 +363,8 @@ begin change (p x m.succ.succ) (snoc (cons y (init v)) (v (last _))) = (p x (nat.succ (nat.succ m))) (cons y v), rw [← cons_snoc_eq_snoc_cons, snoc_init_self] }, - { assume m (hm : (m : with_top ℕ) ≤ n), - have A : (m.succ : with_top ℕ) ≤ n.succ, + { assume m (hm : (m : ℕ∞) ≤ n), + have A : (m.succ : ℕ∞) ≤ n.succ, by { rw with_top.coe_le_coe at ⊢ hm, exact nat.pred_le_iff.mp hm }, change continuous_on ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm ∘ (λ (y : E), p y m.succ)) s, @@ -372,10 +373,10 @@ begin { rintros ⟨Hzero_eq, Hfderiv_zero, Htaylor⟩, split, { exact Hzero_eq }, - { assume m (hm : (m : with_top ℕ) < n.succ) x (hx : x ∈ s), + { assume m (hm : (m : ℕ∞) < n.succ) x (hx : x ∈ s), cases m, { exact Hfderiv_zero x hx }, - { have A : (m : with_top ℕ) < n, + { have A : (m : ℕ∞) < n, by { rw with_top.coe_lt_coe at hm ⊢, exact nat.lt_of_succ_lt_succ hm }, have : has_fderiv_within_at ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm ∘ (λ (y : E), p y m.succ)) ((p x).shift m.succ).curry_left s x := @@ -386,12 +387,12 @@ begin change (p x (nat.succ (nat.succ m))) (cons y v) = (p x m.succ.succ) (snoc (cons y (init v)) (v (last _))), rw [← cons_snoc_eq_snoc_cons, snoc_init_self] } }, - { assume m (hm : (m : with_top ℕ) ≤ n.succ), + { assume m (hm : (m : ℕ∞) ≤ n.succ), cases m, { have : differentiable_on 𝕜 (λ x, p x 0) s := λ x hx, (Hfderiv_zero x hx).differentiable_within_at, exact this.continuous_on }, - { have A : (m : with_top ℕ) ≤ n, + { have A : (m : ℕ∞) ≤ n, by { rw with_top.coe_le_coe at hm ⊢, exact nat.lt_succ_iff.mp hm }, have : continuous_on ((continuous_multilinear_curry_right_equiv' 𝕜 m E F).symm ∘ (λ (y : E), p y m.succ)) s := @@ -411,8 +412,8 @@ depend on the finite order we consider). For instance, a real function which is `C^m` on `(-1/m, 1/m)` for each natural `m`, but not better, is `C^∞` at `0` within `univ`. -/ -def cont_diff_within_at (n : with_top ℕ) (f : E → F) (s : set E) (x : E) := -∀ (m : ℕ), (m : with_top ℕ) ≤ n → +def cont_diff_within_at (n : ℕ∞) (f : E → F) (s : set E) (x : E) := +∀ (m : ℕ), (m : ℕ∞) ≤ n → ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → formal_multilinear_series 𝕜 E F, has_ftaylor_series_up_to_on m f p u @@ -617,7 +618,7 @@ admits continuous derivatives up to order `n` on a neighborhood of `x` in `s`. For `n = ∞`, we only require that this holds up to any finite order (where the neighborhood may depend on the finite order we consider). -/ -definition cont_diff_on (n : with_top ℕ) (f : E → F) (s : set E) := +definition cont_diff_on (n : ℕ∞) (f : E → F) (s : set E) := ∀ x ∈ s, cont_diff_within_at 𝕜 n f s x variable {𝕜} @@ -627,7 +628,7 @@ lemma cont_diff_on.cont_diff_within_at (h : cont_diff_on 𝕜 n f s) (hx : x ∈ h x hx lemma cont_diff_within_at.cont_diff_on {m : ℕ} - (hm : (m : with_top ℕ) ≤ n) (h : cont_diff_within_at 𝕜 n f s x) : + (hm : (m : ℕ∞) ≤ n) (h : cont_diff_within_at 𝕜 n f s x) : ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ cont_diff_on 𝕜 m f u := begin rcases h m hm with ⟨u, u_nhd, p, hp⟩, @@ -888,7 +889,7 @@ iterated_fderiv_within_inter' (mem_nhds_within_of_mem_nhds hu) hs xs begin refine ⟨λ H, H.continuous_on, λ H, _⟩, assume x hx m hm, - have : (m : with_top ℕ) = 0 := le_antisymm hm bot_le, + have : (m : ℕ∞) = 0 := le_antisymm hm bot_le, rw this, refine ⟨insert x s, self_mem_nhds_within, ftaylor_series_within 𝕜 f s, _⟩, rw has_ftaylor_series_up_to_on_zero_iff, @@ -915,12 +916,12 @@ end with the one we have chosen in `iterated_fderiv_within 𝕜 m f s`. -/ theorem has_ftaylor_series_up_to_on.eq_ftaylor_series_of_unique_diff_on (h : has_ftaylor_series_up_to_on n f p s) - {m : ℕ} (hmn : (m : with_top ℕ) ≤ n) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : + {m : ℕ} (hmn : (m : ℕ∞) ≤ n) (hs : unique_diff_on 𝕜 s) (hx : x ∈ s) : p x m = iterated_fderiv_within 𝕜 m f s x := begin induction m with m IH generalizing x, { rw [h.zero_eq' hx, iterated_fderiv_within_zero_eq_comp] }, - { have A : (m : with_top ℕ) < n := lt_of_lt_of_le (with_top.coe_lt_coe.2 (lt_add_one m)) hmn, + { have A : (m : ℕ∞) < n := lt_of_lt_of_le (with_top.coe_lt_coe.2 (lt_add_one m)) hmn, have : has_fderiv_within_at (λ (y : E), iterated_fderiv_within 𝕜 m f s y) (continuous_multilinear_map.curry_left (p x (nat.succ m))) s x := (h.fderiv_within m A x hx).congr (λ y hy, (IH (le_of_lt A) hy).symm) (IH (le_of_lt A) hx).symm, @@ -940,7 +941,7 @@ begin simp only [ftaylor_series_within, continuous_multilinear_map.uncurry0_apply, iterated_fderiv_within_zero_apply] }, { assume m hm x hx, - rcases (h x hx) m.succ (with_top.add_one_le_of_lt hm) with ⟨u, hu, p, Hp⟩, + rcases (h x hx) m.succ (enat.add_one_le_of_lt hm) with ⟨u, hu, p, Hp⟩, rw insert_eq_of_mem hx at hu, rcases mem_nhds_within.1 hu with ⟨o, o_open, xo, ho⟩, rw inter_comm at ho, @@ -976,9 +977,9 @@ begin end lemma cont_diff_on_of_continuous_on_differentiable_on - (Hcont : ∀ (m : ℕ), (m : with_top ℕ) ≤ n → + (Hcont : ∀ (m : ℕ), (m : ℕ∞) ≤ n → continuous_on (λ x, iterated_fderiv_within 𝕜 m f s x) s) - (Hdiff : ∀ (m : ℕ), (m : with_top ℕ) < n → + (Hdiff : ∀ (m : ℕ), (m : ℕ∞) < n → differentiable_on 𝕜 (λ x, iterated_fderiv_within 𝕜 m f s x) s) : cont_diff_on 𝕜 n f s := begin @@ -999,27 +1000,27 @@ begin end lemma cont_diff_on_of_differentiable_on - (h : ∀(m : ℕ), (m : with_top ℕ) ≤ n → differentiable_on 𝕜 (iterated_fderiv_within 𝕜 m f s) s) : + (h : ∀(m : ℕ), (m : ℕ∞) ≤ n → differentiable_on 𝕜 (iterated_fderiv_within 𝕜 m f s) s) : cont_diff_on 𝕜 n f s := cont_diff_on_of_continuous_on_differentiable_on (λ m hm, (h m hm).continuous_on) (λ m hm, (h m (le_of_lt hm))) lemma cont_diff_on.continuous_on_iterated_fderiv_within {m : ℕ} - (h : cont_diff_on 𝕜 n f s) (hmn : (m : with_top ℕ) ≤ n) (hs : unique_diff_on 𝕜 s) : + (h : cont_diff_on 𝕜 n f s) (hmn : (m : ℕ∞) ≤ n) (hs : unique_diff_on 𝕜 s) : continuous_on (iterated_fderiv_within 𝕜 m f s) s := (h.ftaylor_series_within hs).cont m hmn lemma cont_diff_on.differentiable_on_iterated_fderiv_within {m : ℕ} - (h : cont_diff_on 𝕜 n f s) (hmn : (m : with_top ℕ) < n) (hs : unique_diff_on 𝕜 s) : + (h : cont_diff_on 𝕜 n f s) (hmn : (m : ℕ∞) < n) (hs : unique_diff_on 𝕜 s) : differentiable_on 𝕜 (iterated_fderiv_within 𝕜 m f s) s := λ x hx, ((h.ftaylor_series_within hs).fderiv_within m hmn x hx).differentiable_within_at lemma cont_diff_on_iff_continuous_on_differentiable_on (hs : unique_diff_on 𝕜 s) : cont_diff_on 𝕜 n f s ↔ - (∀ (m : ℕ), (m : with_top ℕ) ≤ n → + (∀ (m : ℕ), (m : ℕ∞) ≤ n → continuous_on (λ x, iterated_fderiv_within 𝕜 m f s x) s) - ∧ (∀ (m : ℕ), (m : with_top ℕ) < n → + ∧ (∀ (m : ℕ), (m : ℕ∞) < n → differentiable_on 𝕜 (λ x, iterated_fderiv_within 𝕜 m f s x) s) := begin split, @@ -1091,7 +1092,7 @@ begin exact h.of_le le_top }, { assume h, refine cont_diff_on_top.2 (λ n, _), - have A : (n : with_top ℕ) ≤ ∞ := le_top, + have A : (n : ℕ∞) ≤ ∞ := le_top, apply ((cont_diff_on_succ_iff_fderiv_within hs).2 ⟨h.1, h.2.of_le A⟩).of_le, exact with_top.coe_le_coe.2 (nat.le_succ n) } end @@ -1118,7 +1119,7 @@ begin have : n = ∞, by simpa using hmn, rw this at hf, exact ((cont_diff_on_top_iff_fderiv_within hs).1 hf).2 }, - { change (m.succ : with_top ℕ) ≤ n at hmn, + { change (m.succ : ℕ∞) ≤ n at hmn, exact ((cont_diff_on_succ_iff_fderiv_within hs).1 (hf.of_le hmn)).2 } end @@ -1142,7 +1143,7 @@ lemma cont_diff_within_at.fderiv_within' (hmn : m + 1 ≤ n) : cont_diff_within_at 𝕜 m (fderiv_within 𝕜 f s) s x := begin - have : ∀ k : ℕ, (k + 1 : with_top ℕ) ≤ n → cont_diff_within_at 𝕜 k (fderiv_within 𝕜 f s) s x, + have : ∀ k : ℕ, (k + 1 : ℕ∞) ≤ n → cont_diff_within_at 𝕜 k (fderiv_within 𝕜 f s) s x, { intros k hkn, obtain ⟨v, hv, -, f', hvf', hf'⟩ := (hf.of_le hkn).has_fderiv_within_at_nhds, apply hf'.congr_of_eventually_eq_insert, @@ -1157,7 +1158,7 @@ end lemma cont_diff_within_at.fderiv_within (hf : cont_diff_within_at 𝕜 n f s x) (hs : unique_diff_on 𝕜 s) - (hmn : (m + 1 : with_top ℕ) ≤ n) (hxs : x ∈ s) : + (hmn : (m + 1 : ℕ∞) ≤ n) (hxs : x ∈ s) : cont_diff_within_at 𝕜 m (fderiv_within 𝕜 f s) s x := hf.fderiv_within' (by { rw [insert_eq_of_mem hxs], exact eventually_of_mem self_mem_nhds_within hs}) hmn @@ -1181,12 +1182,12 @@ end /-- `has_ftaylor_series_up_to n f p` registers the fact that `p 0 = f` and `p (m+1)` is a derivative of `p m` for `m < n`, and is continuous for `m ≤ n`. This is a predicate analogous to `has_fderiv_at` but for higher order derivatives. -/ -structure has_ftaylor_series_up_to (n : with_top ℕ) +structure has_ftaylor_series_up_to (n : ℕ∞) (f : E → F) (p : E → formal_multilinear_series 𝕜 E F) : Prop := (zero_eq : ∀ x, (p x 0).uncurry0 = f x) -(fderiv : ∀ (m : ℕ) (hm : (m : with_top ℕ) < n), ∀ x, +(fderiv : ∀ (m : ℕ) (hm : (m : ℕ∞) < n), ∀ x, has_fderiv_at (λ y, p y m) (p x m.succ).curry_left x) -(cont : ∀ (m : ℕ) (hm : (m : with_top ℕ) ≤ n), continuous (λ x, p x m)) +(cont : ∀ (m : ℕ) (hm : (m : ℕ∞) ≤ n), continuous (λ x, p x m)) lemma has_ftaylor_series_up_to.zero_eq' (h : has_ftaylor_series_up_to n f p) (x : E) : @@ -1272,7 +1273,7 @@ variable (𝕜) /-- A function is continuously differentiable up to `n` at a point `x` if, for any integer `k ≤ n`, there is a neighborhood of `x` where `f` admits derivatives up to order `n`, which are continuous. -/ -def cont_diff_at (n : with_top ℕ) (f : E → F) (x : E) := +def cont_diff_at (n : ℕ∞) (f : E → F) (x : E) := cont_diff_within_at 𝕜 n f univ x variable {𝕜} @@ -1347,7 +1348,7 @@ variable (𝕜) order `n`, which are continuous. Contrary to the case of definitions in domains (where derivatives might not be unique) we do not need to localize the definition in space or time. -/ -definition cont_diff (n : with_top ℕ) (f : E → F) := +definition cont_diff (n : ℕ∞) (f : E → F) := ∃ p : E → formal_multilinear_series 𝕜 E F, has_ftaylor_series_up_to n f p variable {𝕜} @@ -1392,8 +1393,8 @@ by { rw ← cont_diff_within_at_univ, simp [cont_diff_within_at_zero, nhds_withi theorem cont_diff_at_one_iff : cont_diff_at 𝕜 1 f x ↔ ∃ f' : E → (E →L[𝕜] F), ∃ u ∈ 𝓝 x, continuous_on f' u ∧ ∀ x ∈ u, has_fderiv_at f (f' x) x := -by simp_rw [show (1 : with_top ℕ) = (0 + 1 : ℕ), from (zero_add 1).symm, - cont_diff_at_succ_iff_has_fderiv_at, show ((0 : ℕ) : with_top ℕ) = 0, from rfl, +by simp_rw [show (1 : ℕ∞) = (0 + 1 : ℕ), from (zero_add 1).symm, + cont_diff_at_succ_iff_has_fderiv_at, show ((0 : ℕ) : ℕ∞) = 0, from rfl, cont_diff_at_zero, exists_mem_and_iff antitone_bforall antitone_continuous_on, and_comm] lemma cont_diff.of_le (h : cont_diff 𝕜 n f) (hmn : m ≤ n) : cont_diff 𝕜 m f := @@ -1519,24 +1520,24 @@ end lemma cont_diff_iff_continuous_differentiable : cont_diff 𝕜 n f ↔ - (∀ (m : ℕ), (m : with_top ℕ) ≤ n → continuous (λ x, iterated_fderiv 𝕜 m f x)) - ∧ (∀ (m : ℕ), (m : with_top ℕ) < n → differentiable 𝕜 (λ x, iterated_fderiv 𝕜 m f x)) := + (∀ (m : ℕ), (m : ℕ∞) ≤ n → continuous (λ x, iterated_fderiv 𝕜 m f x)) + ∧ (∀ (m : ℕ), (m : ℕ∞) < n → differentiable 𝕜 (λ x, iterated_fderiv 𝕜 m f x)) := by simp [cont_diff_on_univ.symm, continuous_iff_continuous_on_univ, differentiable_on_univ.symm, iterated_fderiv_within_univ, cont_diff_on_iff_continuous_on_differentiable_on unique_diff_on_univ] /-- If `f` is `C^n` then its `m`-times iterated derivative is continuous for `m ≤ n`. -/ -lemma cont_diff.continuous_iterated_fderiv {m : ℕ} (hm : (m : with_top ℕ) ≤ n) +lemma cont_diff.continuous_iterated_fderiv {m : ℕ} (hm : (m : ℕ∞) ≤ n) (hf : cont_diff 𝕜 n f) : continuous (λ x, iterated_fderiv 𝕜 m f x) := (cont_diff_iff_continuous_differentiable.mp hf).1 m hm /-- If `f` is `C^n` then its `m`-times iterated derivative is differentiable for `m < n`. -/ -lemma cont_diff.differentiable_iterated_fderiv {m : ℕ} (hm : (m : with_top ℕ) < n) +lemma cont_diff.differentiable_iterated_fderiv {m : ℕ} (hm : (m : ℕ∞) < n) (hf : cont_diff 𝕜 n f) : differentiable 𝕜 (λ x, iterated_fderiv 𝕜 m f x) := (cont_diff_iff_continuous_differentiable.mp hf).2 m hm lemma cont_diff_of_differentiable_iterated_fderiv - (h : ∀(m : ℕ), (m : with_top ℕ) ≤ n → differentiable 𝕜 (iterated_fderiv 𝕜 m f)) : + (h : ∀(m : ℕ), (m : ℕ∞) ≤ n → differentiable 𝕜 (iterated_fderiv 𝕜 m f)) : cont_diff 𝕜 n f := cont_diff_iff_continuous_differentiable.2 ⟨λ m hm, (h m hm).continuous, λ m hm, (h m (le_of_lt hm))⟩ @@ -1959,7 +1960,7 @@ private lemma cont_diff_on.comp_same_univ (hg : cont_diff_on 𝕜 n g t) (hf : cont_diff_on 𝕜 n f s) (st : s ⊆ f ⁻¹' t) : cont_diff_on 𝕜 n (g ∘ f) s := begin - unfreezingI { induction n using with_top.nat_induction with n IH Itop generalizing Eu Fu Gu }, + unfreezingI { induction n using enat.nat_induction with n IH Itop generalizing Eu Fu Gu }, { rw cont_diff_on_zero at hf hg ⊢, exact continuous_on.comp hg hf st }, { rw cont_diff_on_succ_iff_has_fderiv_within_at at hg ⊢, @@ -2414,7 +2415,7 @@ begin induction i with i hi generalizing x, { ext h, simp }, { ext h, - have hi' : (i : with_top ℕ) < i+1 := + have hi' : (i : ℕ∞) < i+1 := with_top.coe_lt_coe.mpr (nat.lt_succ_self _), have hdf : differentiable_on 𝕜 (iterated_fderiv_within 𝕜 i f s) s := hf.differentiable_on_iterated_fderiv_within hi' hu, @@ -2483,7 +2484,7 @@ begin induction i with i hi generalizing x, { ext h, simp }, { ext h, - have hi' : (i : with_top ℕ) < i+1 := + have hi' : (i : ℕ∞) < i+1 := with_top.coe_lt_coe.mpr (nat.lt_succ_self _), calc iterated_fderiv_within 𝕜 (i+1) (-f) s x h = fderiv_within 𝕜 (iterated_fderiv_within 𝕜 i (-f) s) s x (h 0) (fin.tail h) : rfl @@ -2745,7 +2746,7 @@ begin induction i with i hi generalizing x, { ext, simp }, { ext h, - have hi' : (i : with_top ℕ) < i+1 := + have hi' : (i : ℕ∞) < i+1 := with_top.coe_lt_coe.mpr (nat.lt_succ_self _), have hdf : differentiable_on 𝕜 (iterated_fderiv_within 𝕜 i f s) s := hf.differentiable_on_iterated_fderiv_within hi' hu, @@ -2861,7 +2862,7 @@ derivative of inversion as a bilinear map of inversion itself. -/ lemma cont_diff_at_ring_inverse [complete_space R] (x : Rˣ) : cont_diff_at 𝕜 n ring.inverse (x : R) := begin - induction n using with_top.nat_induction with n IH Itop, + induction n using enat.nat_induction with n IH Itop, { intros m hm, refine ⟨{y : R | is_unit y}, _, _⟩, { simp [nhds_within_univ], @@ -2990,7 +2991,7 @@ theorem local_homeomorph.cont_diff_at_symm [complete_space E] cont_diff_at 𝕜 n f.symm a := begin -- We prove this by induction on `n` - induction n using with_top.nat_induction with n IH Itop, + induction n using enat.nat_induction with n IH Itop, { rw cont_diff_at_zero, exact ⟨f.target, is_open.mem_nhds f.open_target ha, f.continuous_inv_fun⟩ }, { obtain ⟨f', ⟨u, hu, hff'⟩, hf'⟩ := cont_diff_at_succ_iff_has_fderiv_at.mp hf, @@ -3079,7 +3080,7 @@ open function finite_dimensional variables [complete_space 𝕜] /-- A family of continuous linear maps is `C^n` on `s` if all its applications are. -/ -lemma cont_diff_on_clm_apply {n : with_top ℕ} {f : E → F →L[𝕜] G} +lemma cont_diff_on_clm_apply {n : ℕ∞} {f : E → F →L[𝕜] G} {s : set E} [finite_dimensional 𝕜 F] : cont_diff_on 𝕜 n f s ↔ ∀ y, cont_diff_on 𝕜 n (λ x, f x y) s := begin @@ -3092,7 +3093,7 @@ begin exact e₂.symm.cont_diff.comp_cont_diff_on (cont_diff_on_pi.mpr (λ i, h _)) end -lemma cont_diff_clm_apply {n : with_top ℕ} {f : E → F →L[𝕜] G} [finite_dimensional 𝕜 F] : +lemma cont_diff_clm_apply {n : ℕ∞} {f : E → F →L[𝕜] G} [finite_dimensional 𝕜 F] : cont_diff 𝕜 n f ↔ ∀ y, cont_diff 𝕜 n (λ x, f x y) := by simp_rw [← cont_diff_on_univ, cont_diff_on_clm_apply] @@ -3316,7 +3317,7 @@ begin exact h.of_le le_top }, { assume h, refine cont_diff_on_top.2 (λ n, _), - have A : (n : with_top ℕ) ≤ ∞ := le_top, + have A : (n : ℕ∞) ≤ ∞ := le_top, apply ((cont_diff_on_succ_iff_deriv_within hs).2 ⟨h.1, h.2.of_le A⟩).of_le, exact with_top.coe_le_coe.2 (nat.le_succ n) } end @@ -3343,7 +3344,7 @@ begin have : n = ∞, by simpa using hmn, rw this at hf, exact ((cont_diff_on_top_iff_deriv_within hs).1 hf).2 }, - { change (m.succ : with_top ℕ) ≤ n at hmn, + { change (m.succ : ℕ∞) ≤ n at hmn, exact ((cont_diff_on_succ_iff_deriv_within hs).1 (hf.of_le hmn)).2 } end diff --git a/src/analysis/calculus/fderiv_analytic.lean b/src/analysis/calculus/fderiv_analytic.lean index ce82fb6b3c0ca..42afc7b3f2353 100644 --- a/src/analysis/calculus/fderiv_analytic.lean +++ b/src/analysis/calculus/fderiv_analytic.lean @@ -121,7 +121,7 @@ begin end /-- An analytic function is infinitely differentiable. -/ -lemma analytic_on.cont_diff_on [complete_space F] (h : analytic_on 𝕜 f s) {n : with_top ℕ} : +lemma analytic_on.cont_diff_on [complete_space F] (h : analytic_on 𝕜 f s) {n : ℕ∞} : cont_diff_on 𝕜 n f s := begin let t := {x | analytic_at 𝕜 f x}, diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index 5006b15c09371..0d145396471ed 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -732,7 +732,7 @@ variables [complete_space E'] (f : E' → F') {f' : E' ≃L[𝕂] F'} {a : E'} /-- Given a `cont_diff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative at `a`, returns a `local_homeomorph` with `to_fun = f` and `a ∈ source`. -/ def to_local_homeomorph - {n : with_top ℕ} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) + {n : ℕ∞} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) (hn : 1 ≤ n) : local_homeomorph E' F' := (hf.has_strict_fderiv_at' hf' hn).to_local_homeomorph f @@ -740,18 +740,18 @@ def to_local_homeomorph variable {f} @[simp] lemma to_local_homeomorph_coe - {n : with_top ℕ} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) + {n : ℕ∞} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) (hn : 1 ≤ n) : (hf.to_local_homeomorph f hf' hn : E' → F') = f := rfl lemma mem_to_local_homeomorph_source - {n : with_top ℕ} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) + {n : ℕ∞} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) (hn : 1 ≤ n) : a ∈ (hf.to_local_homeomorph f hf' hn).source := (hf.has_strict_fderiv_at' hf' hn).mem_to_local_homeomorph_source lemma image_mem_to_local_homeomorph_target - {n : with_top ℕ} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) + {n : ℕ∞} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) (hn : 1 ≤ n) : f a ∈ (hf.to_local_homeomorph f hf' hn).target := (hf.has_strict_fderiv_at' hf' hn).image_mem_to_local_homeomorph_target @@ -759,13 +759,13 @@ lemma image_mem_to_local_homeomorph_target /-- Given a `cont_diff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative at `a`, returns a function that is locally inverse to `f`. -/ def local_inverse - {n : with_top ℕ} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) + {n : ℕ∞} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) (hn : 1 ≤ n) : F' → E' := (hf.has_strict_fderiv_at' hf' hn).local_inverse f f' a lemma local_inverse_apply_image - {n : with_top ℕ} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) + {n : ℕ∞} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) (hn : 1 ≤ n) : hf.local_inverse hf' hn (f a) = a := (hf.has_strict_fderiv_at' hf' hn).local_inverse_apply_image @@ -774,7 +774,7 @@ lemma local_inverse_apply_image at `a`, the inverse function (produced by `cont_diff.to_local_homeomorph`) is also `cont_diff`. -/ lemma to_local_inverse - {n : with_top ℕ} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) + {n : ℕ∞} (hf : cont_diff_at 𝕂 n f a) (hf' : has_fderiv_at f (f' : E' →L[𝕂] F') a) (hn : 1 ≤ n) : cont_diff_at 𝕂 n (hf.local_inverse hf' hn) (f a) := begin diff --git a/src/analysis/calculus/iterated_deriv.lean b/src/analysis/calculus/iterated_deriv.lean index 10795867616d8..57637f6b3ffd9 100644 --- a/src/analysis/calculus/iterated_deriv.lean +++ b/src/analysis/calculus/iterated_deriv.lean @@ -108,10 +108,10 @@ by { simp [iterated_deriv_within, iterated_fderiv_within_one_apply hs hx], refl derivatives are differentiable, then the function is `C^n`. This is not an equivalence in general, but this is an equivalence when the set has unique derivatives, see `cont_diff_on_iff_continuous_on_differentiable_on_deriv`. -/ -lemma cont_diff_on_of_continuous_on_differentiable_on_deriv {n : with_top ℕ} - (Hcont : ∀ (m : ℕ), (m : with_top ℕ) ≤ n → +lemma cont_diff_on_of_continuous_on_differentiable_on_deriv {n : ℕ∞} + (Hcont : ∀ (m : ℕ), (m : ℕ∞) ≤ n → continuous_on (λ x, iterated_deriv_within m f s x) s) - (Hdiff : ∀ (m : ℕ), (m : with_top ℕ) < n → + (Hdiff : ∀ (m : ℕ), (m : ℕ∞) < n → differentiable_on 𝕜 (λ x, iterated_deriv_within m f s x) s) : cont_diff_on 𝕜 n f s := begin @@ -125,8 +125,8 @@ first `n` derivatives are differentiable. This is slightly too strong as the con require on the `n`-th derivative is differentiability instead of continuity, but it has the advantage of avoiding the discussion of continuity in the proof (and for `n = ∞` this is optimal). -/ -lemma cont_diff_on_of_differentiable_on_deriv {n : with_top ℕ} - (h : ∀(m : ℕ), (m : with_top ℕ) ≤ n → differentiable_on 𝕜 (iterated_deriv_within m f s) s) : +lemma cont_diff_on_of_differentiable_on_deriv {n : ℕ∞} + (h : ∀(m : ℕ), (m : ℕ∞) ≤ n → differentiable_on 𝕜 (iterated_deriv_within m f s) s) : cont_diff_on 𝕜 n f s := begin apply cont_diff_on_of_differentiable_on, @@ -136,16 +136,16 @@ end /-- On a set with unique derivatives, a `C^n` function has derivatives up to `n` which are continuous. -/ -lemma cont_diff_on.continuous_on_iterated_deriv_within {n : with_top ℕ} {m : ℕ} - (h : cont_diff_on 𝕜 n f s) (hmn : (m : with_top ℕ) ≤ n) (hs : unique_diff_on 𝕜 s) : +lemma cont_diff_on.continuous_on_iterated_deriv_within {n : ℕ∞} {m : ℕ} + (h : cont_diff_on 𝕜 n f s) (hmn : (m : ℕ∞) ≤ n) (hs : unique_diff_on 𝕜 s) : continuous_on (iterated_deriv_within m f s) s := by simpa only [iterated_deriv_within_eq_equiv_comp, linear_isometry_equiv.comp_continuous_on_iff] using h.continuous_on_iterated_fderiv_within hmn hs /-- On a set with unique derivatives, a `C^n` function has derivatives less than `n` which are differentiable. -/ -lemma cont_diff_on.differentiable_on_iterated_deriv_within {n : with_top ℕ} {m : ℕ} - (h : cont_diff_on 𝕜 n f s) (hmn : (m : with_top ℕ) < n) (hs : unique_diff_on 𝕜 s) : +lemma cont_diff_on.differentiable_on_iterated_deriv_within {n : ℕ∞} {m : ℕ} + (h : cont_diff_on 𝕜 n f s) (hmn : (m : ℕ∞) < n) (hs : unique_diff_on 𝕜 s) : differentiable_on 𝕜 (iterated_deriv_within m f s) s := by simpa only [iterated_deriv_within_eq_equiv_comp, linear_isometry_equiv.comp_differentiable_on_iff] @@ -153,11 +153,11 @@ by simpa only [iterated_deriv_within_eq_equiv_comp, /-- The property of being `C^n`, initially defined in terms of the Fréchet derivative, can be reformulated in terms of the one-dimensional derivative on sets with unique derivatives. -/ -lemma cont_diff_on_iff_continuous_on_differentiable_on_deriv {n : with_top ℕ} +lemma cont_diff_on_iff_continuous_on_differentiable_on_deriv {n : ℕ∞} (hs : unique_diff_on 𝕜 s) : cont_diff_on 𝕜 n f s ↔ - (∀m:ℕ, (m : with_top ℕ) ≤ n → continuous_on (iterated_deriv_within m f s) s) - ∧ (∀m:ℕ, (m : with_top ℕ) < n → differentiable_on 𝕜 (iterated_deriv_within m f s) s) := + (∀m:ℕ, (m : ℕ∞) ≤ n → continuous_on (iterated_deriv_within m f s) s) + ∧ (∀m:ℕ, (m : ℕ∞) < n → differentiable_on 𝕜 (iterated_deriv_within m f s) s) := by simp only [cont_diff_on_iff_continuous_on_differentiable_on hs, iterated_fderiv_within_eq_equiv_comp, linear_isometry_equiv.comp_continuous_on_iff, linear_isometry_equiv.comp_differentiable_on_iff] @@ -232,10 +232,10 @@ by { ext x, simp [iterated_deriv], refl } /-- The property of being `C^n`, initially defined in terms of the Fréchet derivative, can be reformulated in terms of the one-dimensional derivative. -/ -lemma cont_diff_iff_iterated_deriv {n : with_top ℕ} : +lemma cont_diff_iff_iterated_deriv {n : ℕ∞} : cont_diff 𝕜 n f ↔ -(∀m:ℕ, (m : with_top ℕ) ≤ n → continuous (iterated_deriv m f)) -∧ (∀m:ℕ, (m : with_top ℕ) < n → differentiable 𝕜 (iterated_deriv m f)) := +(∀m:ℕ, (m : ℕ∞) ≤ n → continuous (iterated_deriv m f)) +∧ (∀m:ℕ, (m : ℕ∞) < n → differentiable 𝕜 (iterated_deriv m f)) := by simp only [cont_diff_iff_continuous_differentiable, iterated_fderiv_eq_equiv_comp, linear_isometry_equiv.comp_continuous_iff, linear_isometry_equiv.comp_differentiable_iff] @@ -244,19 +244,19 @@ first `n` derivatives are differentiable. This is slightly too strong as the con require on the `n`-th derivative is differentiability instead of continuity, but it has the advantage of avoiding the discussion of continuity in the proof (and for `n = ∞` this is optimal). -/ -lemma cont_diff_of_differentiable_iterated_deriv {n : with_top ℕ} - (h : ∀(m : ℕ), (m : with_top ℕ) ≤ n → differentiable 𝕜 (iterated_deriv m f)) : +lemma cont_diff_of_differentiable_iterated_deriv {n : ℕ∞} + (h : ∀(m : ℕ), (m : ℕ∞) ≤ n → differentiable 𝕜 (iterated_deriv m f)) : cont_diff 𝕜 n f := cont_diff_iff_iterated_deriv.2 ⟨λ m hm, (h m hm).continuous, λ m hm, (h m (le_of_lt hm))⟩ -lemma cont_diff.continuous_iterated_deriv {n : with_top ℕ} (m : ℕ) - (h : cont_diff 𝕜 n f) (hmn : (m : with_top ℕ) ≤ n) : +lemma cont_diff.continuous_iterated_deriv {n : ℕ∞} (m : ℕ) + (h : cont_diff 𝕜 n f) (hmn : (m : ℕ∞) ≤ n) : continuous (iterated_deriv m f) := (cont_diff_iff_iterated_deriv.1 h).1 m hmn -lemma cont_diff.differentiable_iterated_deriv {n : with_top ℕ} (m : ℕ) - (h : cont_diff 𝕜 n f) (hmn : (m : with_top ℕ) < n) : +lemma cont_diff.differentiable_iterated_deriv {n : ℕ∞} (m : ℕ) + (h : cont_diff 𝕜 n f) (hmn : (m : ℕ∞) < n) : differentiable 𝕜 (iterated_deriv m f) := (cont_diff_iff_iterated_deriv.1 h).2 m hmn diff --git a/src/analysis/calculus/specific_functions.lean b/src/analysis/calculus/specific_functions.lean index d02ba83483161..d55ef2521e7f6 100644 --- a/src/analysis/calculus/specific_functions.lean +++ b/src/analysis/calculus/specific_functions.lean @@ -293,7 +293,7 @@ lemma R_pos {c : E} (f : cont_diff_bump_of_inner c) : 0 < f.R := f.r_pos.trans f instance (c : E) : inhabited (cont_diff_bump_of_inner c) := ⟨⟨1, 2, zero_lt_one, one_lt_two⟩⟩ variables [inner_product_space ℝ E] [normed_add_comm_group X] [normed_space ℝ X] -variables {c : E} (f : cont_diff_bump_of_inner c) {x : E} {n : with_top ℕ} +variables {c : E} (f : cont_diff_bump_of_inner c) {x : E} {n : ℕ∞} /-- The function defined by `f : cont_diff_bump_of_inner c`. Use automatic coercion to function instead. -/ @@ -406,7 +406,7 @@ rfl lemma nonneg_normed (x : E) : 0 ≤ f.normed μ x := div_nonneg f.nonneg $ integral_nonneg f.nonneg' -lemma cont_diff_normed {n : with_top ℕ} : cont_diff ℝ n (f.normed μ) := +lemma cont_diff_normed {n : ℕ∞} : cont_diff ℝ n (f.normed μ) := f.cont_diff.div_const lemma continuous_normed : continuous (f.normed μ) := diff --git a/src/analysis/complex/real_deriv.lean b/src/analysis/complex/real_deriv.lean index 9f40a3dcf68f6..c06c96100c8cc 100644 --- a/src/analysis/complex/real_deriv.lean +++ b/src/analysis/complex/real_deriv.lean @@ -64,7 +64,7 @@ begin simpa using (C.comp z (B.comp z A)).has_deriv_at end -theorem cont_diff_at.real_of_complex {n : with_top ℕ} (h : cont_diff_at ℂ n e z) : +theorem cont_diff_at.real_of_complex {n : ℕ∞} (h : cont_diff_at ℂ n e z) : cont_diff_at ℝ n (λ x : ℝ, (e x).re) z := begin have A : cont_diff_at ℝ n (coe : ℝ → ℂ) z, @@ -74,7 +74,7 @@ begin exact C.comp z (B.comp z A) end -theorem cont_diff.real_of_complex {n : with_top ℕ} (h : cont_diff ℂ n e) : +theorem cont_diff.real_of_complex {n : ℕ∞} (h : cont_diff ℂ n e) : cont_diff ℝ n (λ x : ℝ, (e x).re) := cont_diff_iff_cont_diff_at.2 $ λ x, h.cont_diff_at.real_of_complex diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 83e485df2d100..67f117ad07fca 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -684,7 +684,7 @@ end normed_add_comm_group namespace cont_diff_bump_of_inner -variables {n : with_top ℕ} +variables {n : ℕ∞} variables [normed_space ℝ E'] variables [inner_product_space ℝ G] variables [complete_space E'] @@ -758,7 +758,7 @@ variables [normed_space 𝕜 E] variables [normed_space 𝕜 E'] variables [normed_space 𝕜 E''] variables [normed_space ℝ F] [normed_space 𝕜 F] -variables {n : with_top ℕ} +variables {n : ℕ∞} variables [complete_space F] variables [measurable_space G] {μ : measure G} variables (L : E →L[𝕜] E' →L[𝕜] F) @@ -854,7 +854,7 @@ lemma has_compact_support.cont_diff_convolution_right [finite_dimensional 𝕜 G (hcg : has_compact_support g) (hf : locally_integrable f μ) (hg : cont_diff 𝕜 n g) : cont_diff 𝕜 n (f ⋆[L, μ] g) := begin - induction n using with_top.nat_induction with n ih ih generalizing g, + induction n using enat.nat_induction with n ih ih generalizing g, { rw [cont_diff_zero] at hg ⊢, exact hcg.continuous_convolution_right L hf hg }, { have h : ∀ x, has_fderiv_at (f ⋆[L, μ] g) ((f ⋆[L.precompR G, μ] fderiv 𝕜 g) x) x := @@ -888,7 +888,7 @@ variables [normed_space 𝕜 E] variables [normed_space 𝕜 E'] variables [normed_space ℝ F] [normed_space 𝕜 F] variables {f₀ : 𝕜 → E} {g₀ : 𝕜 → E'} -variables {n : with_top ℕ} +variables {n : ℕ∞} variables (L : E →L[𝕜] E' →L[𝕜] F) variables [complete_space F] variables {μ : measure 𝕜} diff --git a/src/analysis/inner_product_space/calculus.lean b/src/analysis/inner_product_space/calculus.lean index da65a4fa98d35..9f95e4b6a1992 100644 --- a/src/analysis/inner_product_space/calculus.lean +++ b/src/analysis/inner_product_space/calculus.lean @@ -53,7 +53,7 @@ lemma differentiable_inner : differentiable ℝ (λ p : E × E, ⟪p.1, p.2⟫) is_bounded_bilinear_map_inner.differentiable_at variables {G : Type*} [normed_add_comm_group G] [normed_space ℝ G] - {f g : G → E} {f' g' : G →L[ℝ] E} {s : set G} {x : G} {n : with_top ℕ} + {f g : G → E} {f' g' : G →L[ℝ] E} {s : set G} {x : G} {n : ℕ∞} include 𝕜 @@ -304,28 +304,28 @@ begin refl end -lemma cont_diff_within_at_euclidean {n : with_top ℕ} : +lemma cont_diff_within_at_euclidean {n : ℕ∞} : cont_diff_within_at 𝕜 n f t y ↔ ∀ i, cont_diff_within_at 𝕜 n (λ x, f x i) t y := begin rw [← (euclidean_space.equiv ι 𝕜).comp_cont_diff_within_at_iff, cont_diff_within_at_pi], refl end -lemma cont_diff_at_euclidean {n : with_top ℕ} : +lemma cont_diff_at_euclidean {n : ℕ∞} : cont_diff_at 𝕜 n f y ↔ ∀ i, cont_diff_at 𝕜 n (λ x, f x i) y := begin rw [← (euclidean_space.equiv ι 𝕜).comp_cont_diff_at_iff, cont_diff_at_pi], refl end -lemma cont_diff_on_euclidean {n : with_top ℕ} : +lemma cont_diff_on_euclidean {n : ℕ∞} : cont_diff_on 𝕜 n f t ↔ ∀ i, cont_diff_on 𝕜 n (λ x, f x i) t := begin rw [← (euclidean_space.equiv ι 𝕜).comp_cont_diff_on_iff, cont_diff_on_pi], refl end -lemma cont_diff_euclidean {n : with_top ℕ} : +lemma cont_diff_euclidean {n : ℕ∞} : cont_diff 𝕜 n f ↔ ∀ i, cont_diff 𝕜 n (λ x, f x i) := begin rw [← (euclidean_space.equiv ι 𝕜).comp_cont_diff_iff, cont_diff_pi], @@ -338,7 +338,7 @@ section diffeomorph_unit_ball open metric (hiding mem_nhds_iff) -variables {n : with_top ℕ} {E : Type*} [inner_product_space ℝ E] +variables {n : ℕ∞} {E : Type*} [inner_product_space ℝ E] lemma cont_diff_homeomorph_unit_ball : cont_diff ℝ n $ λ (x : E), (homeomorph_unit_ball x : E) := diff --git a/src/analysis/inner_product_space/euclidean_dist.lean b/src/analysis/inner_product_space/euclidean_dist.lean index bfdd60a7cc2eb..55b118bbe1293 100644 --- a/src/analysis/inner_product_space/euclidean_dist.lean +++ b/src/analysis/inner_product_space/euclidean_dist.lean @@ -106,7 +106,7 @@ nhds_basis_ball.mem_of_mem hr end euclidean -variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] {f g : F → E} {n : with_top ℕ} +variables {F : Type*} [normed_add_comm_group F] [normed_space ℝ F] {f g : F → E} {n : ℕ∞} lemma cont_diff.euclidean_dist (hf : cont_diff ℝ n f) (hg : cont_diff ℝ n g) (h : ∀ x, f x ≠ g x) : diff --git a/src/analysis/special_functions/arsinh.lean b/src/analysis/special_functions/arsinh.lean index 49d41ae88a67b..545a87e210932 100644 --- a/src/analysis/special_functions/arsinh.lean +++ b/src/analysis/special_functions/arsinh.lean @@ -134,7 +134,7 @@ lemma has_deriv_at_arsinh (x : ℝ) : has_deriv_at arsinh (sqrt (1 + x ^ 2))⁻ lemma differentiable_arsinh : differentiable ℝ arsinh := λ x, (has_deriv_at_arsinh x).differentiable_at -lemma cont_diff_arsinh {n : with_top ℕ} : cont_diff ℝ n arsinh := +lemma cont_diff_arsinh {n : ℕ∞} : cont_diff ℝ n arsinh := sinh_homeomorph.cont_diff_symm_deriv (λ x, (cosh_pos x).ne') has_deriv_at_sinh cont_diff_sinh @[continuity] lemma continuous_arsinh : continuous arsinh := sinh_homeomorph.symm.continuous @@ -168,7 +168,7 @@ end continuous section fderiv variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {s : set E} {a : E} - {f' : E →L[ℝ] ℝ} {n : with_top ℕ} + {f' : E →L[ℝ] ℝ} {n : ℕ∞} lemma has_strict_fderiv_at.arsinh (hf : has_strict_fderiv_at f f' a) : has_strict_fderiv_at (λ x, arsinh (f x)) ((sqrt (1 + (f a) ^ 2))⁻¹ • f') a := diff --git a/src/analysis/special_functions/complex/log_deriv.lean b/src/analysis/special_functions/complex/log_deriv.lean index 176a8c6e0ea00..99ed7f3c415ae 100644 --- a/src/analysis/special_functions/complex/log_deriv.lean +++ b/src/analysis/special_functions/complex/log_deriv.lean @@ -56,7 +56,7 @@ lemma has_strict_fderiv_at_log_real {x : ℂ} (h : 0 < x.re ∨ x.im ≠ 0) : has_strict_fderiv_at log (x⁻¹ • (1 : ℂ →L[ℝ] ℂ)) x := (has_strict_deriv_at_log h).complex_to_real_fderiv -lemma cont_diff_at_log {x : ℂ} (h : 0 < x.re ∨ x.im ≠ 0) {n : with_top ℕ} : +lemma cont_diff_at_log {x : ℂ} (h : 0 < x.re ∨ x.im ≠ 0) {n : ℕ∞} : cont_diff_at ℂ n log x := exp_local_homeomorph.cont_diff_at_symm_deriv (exp_ne_zero $ log x) h (has_deriv_at_exp _) cont_diff_exp.cont_diff_at diff --git a/src/analysis/special_functions/log/deriv.lean b/src/analysis/special_functions/log/deriv.lean index dc732754a436d..aa7477d789953 100644 --- a/src/analysis/special_functions/log/deriv.lean +++ b/src/analysis/special_functions/log/deriv.lean @@ -60,14 +60,14 @@ else (has_deriv_at_log hx).deriv @[simp] lemma deriv_log' : deriv log = has_inv.inv := funext deriv_log -lemma cont_diff_on_log {n : with_top ℕ} : cont_diff_on ℝ n log {0}ᶜ := +lemma cont_diff_on_log {n : ℕ∞} : cont_diff_on ℝ n log {0}ᶜ := begin suffices : cont_diff_on ℝ ⊤ log {0}ᶜ, from this.of_le le_top, refine (cont_diff_on_top_iff_deriv_of_open is_open_compl_singleton).2 _, simp [differentiable_on_log, cont_diff_on_inv] end -lemma cont_diff_at_log {n : with_top ℕ} : cont_diff_at ℝ n log x ↔ x ≠ 0 := +lemma cont_diff_at_log {n : ℕ∞} : cont_diff_at ℝ n log x ↔ x ≠ 0 := ⟨λ h, continuous_at_log_iff.1 h.continuous_at, λ hx, (cont_diff_on_log x hx).cont_diff_at $ is_open.mem_nhds is_open_compl_singleton hx⟩ diff --git a/src/analysis/special_functions/pow_deriv.lean b/src/analysis/special_functions/pow_deriv.lean index d43a491aa916d..8b39292431f8d 100644 --- a/src/analysis/special_functions/pow_deriv.lean +++ b/src/analysis/special_functions/pow_deriv.lean @@ -225,7 +225,7 @@ begin end /-- The function `λ (x, y), x ^ y` is infinitely smooth at `(x, y)` unless `x = 0`. -/ -lemma cont_diff_at_rpow_of_ne (p : ℝ × ℝ) (hp : p.1 ≠ 0) {n : with_top ℕ} : +lemma cont_diff_at_rpow_of_ne (p : ℝ × ℝ) (hp : p.1 ≠ 0) {n : ℕ∞} : cont_diff_at ℝ n (λ p : ℝ × ℝ, p.1 ^ p.2) p := begin cases hp.lt_or_lt with hneg hpos, @@ -302,7 +302,7 @@ lemma deriv_rpow_const' {p : ℝ} (h : 1 ≤ p) : deriv (λ x : ℝ, x ^ p) = λ x, p * x ^ (p - 1) := funext $ λ x, deriv_rpow_const (or.inr h) -lemma cont_diff_at_rpow_const_of_ne {x p : ℝ} {n : with_top ℕ} (h : x ≠ 0) : +lemma cont_diff_at_rpow_const_of_ne {x p : ℝ} {n : ℕ∞} (h : x ≠ 0) : cont_diff_at ℝ n (λ x, x ^ p) x := (cont_diff_at_rpow_of_ne (x, p) h).comp x (cont_diff_at_id.prod cont_diff_at_const) @@ -340,7 +340,7 @@ open real section fderiv variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f g : E → ℝ} {f' g' : E →L[ℝ] ℝ} - {x : E} {s : set E} {c p : ℝ} {n : with_top ℕ} + {x : E} {s : set E} {c p : ℝ} {n : ℕ∞} lemma has_fderiv_within_at.rpow (hf : has_fderiv_within_at f f' s x) (hg : has_fderiv_within_at g g' s x) (h : 0 < f x) : diff --git a/src/analysis/special_functions/sqrt.lean b/src/analysis/special_functions/sqrt.lean index 6a28132dbe8d3..6304b60ea46f0 100644 --- a/src/analysis/special_functions/sqrt.lean +++ b/src/analysis/special_functions/sqrt.lean @@ -57,7 +57,7 @@ lemma has_strict_deriv_at_sqrt {x : ℝ} (hx : x ≠ 0) : has_strict_deriv_at sqrt (1 / (2 * sqrt x)) x := (deriv_sqrt_aux hx).1 -lemma cont_diff_at_sqrt {x : ℝ} {n : with_top ℕ} (hx : x ≠ 0) : +lemma cont_diff_at_sqrt {x : ℝ} {n : ℕ∞} (hx : x ≠ 0) : cont_diff_at ℝ n sqrt x := (deriv_sqrt_aux hx).2 n @@ -98,7 +98,7 @@ end deriv section fderiv -variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {n : with_top ℕ} +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {n : ℕ∞} {s : set E} {x : E} {f' : E →L[ℝ] ℝ} lemma has_fderiv_at.sqrt (hf : has_fderiv_at f f' x) (hx : f x ≠ 0) : diff --git a/src/analysis/special_functions/trigonometric/arctan_deriv.lean b/src/analysis/special_functions/trigonometric/arctan_deriv.lean index 80962df871ee9..7f46be4b3f99d 100644 --- a/src/analysis/special_functions/trigonometric/arctan_deriv.lean +++ b/src/analysis/special_functions/trigonometric/arctan_deriv.lean @@ -85,7 +85,7 @@ lemma differentiable_arctan : differentiable ℝ arctan := differentiable_at_arc @[simp] lemma deriv_arctan : deriv arctan = (λ x, 1 / (1 + x^2)) := funext $ λ x, (has_deriv_at_arctan x).deriv -lemma cont_diff_arctan {n : with_top ℕ} : cont_diff ℝ n arctan := +lemma cont_diff_arctan {n : ℕ∞} : cont_diff ℝ n arctan := cont_diff_iff_cont_diff_at.2 $ λ x, have cos (arctan x) ≠ 0 := (cos_arctan_pos x).ne', tan_local_homeomorph.cont_diff_at_symm_deriv (by simpa) trivial (has_deriv_at_tan this) @@ -132,7 +132,7 @@ end deriv section fderiv variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] {f : E → ℝ} {f' : E →L[ℝ] ℝ} - {x : E} {s : set E} {n : with_top ℕ} + {x : E} {s : set E} {n : ℕ∞} lemma has_strict_fderiv_at.arctan (hf : has_strict_fderiv_at f f' x) : has_strict_fderiv_at (λ x, arctan (f x)) ((1 / (1 + (f x)^2)) • f') x := diff --git a/src/analysis/special_functions/trigonometric/complex_deriv.lean b/src/analysis/special_functions/trigonometric/complex_deriv.lean index 65bd9c6f5e77b..a8a78b04f7acb 100644 --- a/src/analysis/special_functions/trigonometric/complex_deriv.lean +++ b/src/analysis/special_functions/trigonometric/complex_deriv.lean @@ -64,7 +64,7 @@ if h : cos x = 0 then by simp [deriv_zero_of_not_differentiable_at this, h, sq] else (has_deriv_at_tan h).deriv -@[simp] lemma cont_diff_at_tan {x : ℂ} {n : with_top ℕ} : +@[simp] lemma cont_diff_at_tan {x : ℂ} {n : ℕ∞} : cont_diff_at ℂ n tan x ↔ cos x ≠ 0 := ⟨λ h, continuous_at_tan.1 h.continuous_at, cont_diff_sin.cont_diff_at.div cont_diff_cos.cont_diff_at⟩ diff --git a/src/analysis/special_functions/trigonometric/inverse_deriv.lean b/src/analysis/special_functions/trigonometric/inverse_deriv.lean index 897138e6344f9..c449c042d7303 100644 --- a/src/analysis/special_functions/trigonometric/inverse_deriv.lean +++ b/src/analysis/special_functions/trigonometric/inverse_deriv.lean @@ -54,7 +54,7 @@ lemma has_deriv_at_arcsin {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) : has_deriv_at arcsin (1 / sqrt (1 - x ^ 2)) x := (has_strict_deriv_at_arcsin h₁ h₂).has_deriv_at -lemma cont_diff_at_arcsin {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : with_top ℕ} : +lemma cont_diff_at_arcsin {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : ℕ∞} : cont_diff_at ℝ n arcsin x := (deriv_arcsin_aux h₁ h₂).2.of_le le_top @@ -117,14 +117,14 @@ lemma differentiable_on_arcsin : differentiable_on ℝ arcsin {-1, 1}ᶜ := λ x hx, (differentiable_at_arcsin.2 ⟨λ h, hx (or.inl h), λ h, hx (or.inr h)⟩).differentiable_within_at -lemma cont_diff_on_arcsin {n : with_top ℕ} : +lemma cont_diff_on_arcsin {n : ℕ∞} : cont_diff_on ℝ n arcsin {-1, 1}ᶜ := λ x hx, (cont_diff_at_arcsin (mt or.inl hx) (mt or.inr hx)).cont_diff_within_at -lemma cont_diff_at_arcsin_iff {x : ℝ} {n : with_top ℕ} : +lemma cont_diff_at_arcsin_iff {x : ℝ} {n : ℕ∞} : cont_diff_at ℝ n arcsin x ↔ n = 0 ∨ (x ≠ -1 ∧ x ≠ 1) := ⟨λ h, or_iff_not_imp_left.2 $ λ hn, differentiable_at_arcsin.1 $ h.differentiable_at $ - with_top.one_le_iff_pos.2 (pos_iff_ne_zero.2 hn), + enat.one_le_iff_ne_zero.2 hn, λ h, h.elim (λ hn, hn.symm ▸ (cont_diff_zero.2 continuous_arcsin).cont_diff_at) $ λ hx, cont_diff_at_arcsin hx.1 hx.2⟩ @@ -140,7 +140,7 @@ lemma has_deriv_at_arccos {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) : has_deriv_at arccos (-(1 / sqrt (1 - x ^ 2))) x := (has_deriv_at_arcsin h₁ h₂).const_sub (π / 2) -lemma cont_diff_at_arccos {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : with_top ℕ} : +lemma cont_diff_at_arccos {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : ℕ∞} : cont_diff_at ℝ n arccos x := cont_diff_at_const.sub (cont_diff_at_arcsin h₁ h₂) @@ -170,11 +170,11 @@ funext $ λ x, (deriv_const_sub _).trans $ by simp only [deriv_arcsin] lemma differentiable_on_arccos : differentiable_on ℝ arccos {-1, 1}ᶜ := differentiable_on_arcsin.const_sub _ -lemma cont_diff_on_arccos {n : with_top ℕ} : +lemma cont_diff_on_arccos {n : ℕ∞} : cont_diff_on ℝ n arccos {-1, 1}ᶜ := cont_diff_on_const.sub cont_diff_on_arcsin -lemma cont_diff_at_arccos_iff {x : ℝ} {n : with_top ℕ} : +lemma cont_diff_at_arccos_iff {x : ℝ} {n : ℕ∞} : cont_diff_at ℝ n arccos x ↔ n = 0 ∨ (x ≠ -1 ∧ x ≠ 1) := by refine iff.trans ⟨λ h, _, λ h, _⟩ cont_diff_at_arcsin_iff; simpa [arccos] using (@cont_diff_at_const _ _ _ _ _ _ _ _ _ _ (π / 2)).sub h diff --git a/src/data/nat/cast.lean b/src/data/nat/cast.lean index fc2e93c0c411a..8283d0dab77b1 100644 --- a/src/data/nat/cast.lean +++ b/src/data/nat/cast.lean @@ -224,32 +224,6 @@ variables [add_monoid_with_one α] end mul_opposite -namespace with_top -variables [add_monoid_with_one α] - -lemma add_one_le_of_lt {i n : with_top ℕ} (h : i < n) : i + 1 ≤ n := -begin - cases n, { exact le_top }, - cases i, { exact (not_le_of_lt h le_top).elim }, - exact with_top.coe_le_coe.2 (with_top.coe_lt_coe.1 h) -end - -lemma one_le_iff_pos {n : with_top ℕ} : 1 ≤ n ↔ 0 < n := -⟨lt_of_lt_of_le (coe_lt_coe.mpr zero_lt_one), - λ h, by simpa only [zero_add] using add_one_le_of_lt h⟩ - -@[elab_as_eliminator] -lemma nat_induction {P : with_top ℕ → Prop} (a : with_top ℕ) - (h0 : P 0) (hsuc : ∀n:ℕ, P n → P n.succ) (htop : (∀n : ℕ, P n) → P ⊤) : P a := -begin - have A : ∀n:ℕ, P n := λ n, nat.rec_on n h0 hsuc, - cases a, - { exact htop A }, - { exact A a } -end - -end with_top - namespace pi variables {π : α → Type*} [Π a, has_nat_cast (π a)] diff --git a/src/data/nat/enat.lean b/src/data/nat/enat.lean new file mode 100644 index 0000000000000..8f8e7b4060d50 --- /dev/null +++ b/src/data/nat/enat.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import data.nat.lattice +import data.nat.succ_pred + +/-! +# Definition and basic properties of extended natural numbers + +In this file we define `enat` (notation: `ℕ∞`) to be `with_top ℕ` and prove some basic lemmas +about this type. +-/ + +/-- Extended natural numbers `ℕ∞ = with_top ℕ`. -/ +@[derive [has_zero, add_comm_monoid_with_one, canonically_ordered_comm_semiring, nontrivial, + linear_order, order_bot, order_top, has_bot, has_top, canonically_linear_ordered_add_monoid, + has_sub, has_ordered_sub, complete_linear_order, linear_ordered_add_comm_monoid_with_top, + succ_order]] +def enat : Type := with_top ℕ + +notation `ℕ∞` := enat + +namespace enat + +instance : inhabited ℕ∞ := ⟨0⟩ + +variables {m n : ℕ∞} + +@[simp, norm_cast] lemma coe_zero : ((0 : ℕ) : ℕ∞) = 0 := rfl +@[simp, norm_cast] lemma coe_one : ((1 : ℕ) : ℕ∞) = 1 := rfl +@[simp, norm_cast] lemma coe_add (m n : ℕ) : ↑(m + n) = (m + n : ℕ∞) := rfl +@[simp, norm_cast] lemma coe_sub (m n : ℕ) : ↑(m - n) = (m - n : ℕ∞) := rfl +@[simp, norm_cast] lemma coe_mul (m n : ℕ) : ↑(m * n) = (m * n : ℕ∞) := with_top.coe_mul + +instance : can_lift ℕ∞ ℕ := with_top.can_lift + +/-- Conversion of `ℕ∞` to `ℕ` sending `∞` to `0`. -/ +def to_nat : monoid_with_zero_hom ℕ∞ ℕ := +{ to_fun := with_top.untop' 0, + map_one' := rfl, + map_zero' := rfl, + map_mul' := with_top.untop'_zero_mul } + +@[simp] lemma succ_def (m : ℕ∞) : order.succ m = m + 1 := by cases m; refl + +lemma add_one_le_of_lt (h : m < n) : m + 1 ≤ n := +m.succ_def ▸ order.succ_le_of_lt h + +lemma add_one_le_iff (hm : m ≠ ⊤) : m + 1 ≤ n ↔ m < n := +m.succ_def ▸ (order.succ_le_iff_of_not_is_max $ by rwa [is_max_iff_eq_top]) + +lemma one_le_iff_pos : 1 ≤ n ↔ 0 < n := add_one_le_iff with_top.zero_ne_top + +lemma one_le_iff_ne_zero : 1 ≤ n ↔ n ≠ 0 := one_le_iff_pos.trans pos_iff_ne_zero + +lemma le_of_lt_add_one (h : m < n + 1) : m ≤ n := order.le_of_lt_succ $ n.succ_def.symm ▸ h + +@[elab_as_eliminator] +lemma nat_induction {P : ℕ∞ → Prop} (a : ℕ∞) (h0 : P 0) (hsuc : ∀ n : ℕ, P n → P n.succ) + (htop : (∀ n : ℕ, P n) → P ⊤) : P a := +begin + have A : ∀ n : ℕ, P n := λ n, nat.rec_on n h0 hsuc, + cases a, + exacts [htop A, A a] +end + +end enat diff --git a/src/data/nat/part_enat.lean b/src/data/nat/part_enat.lean index c22ddee0b6d57..2c118d87571cf 100644 --- a/src/data/nat/part_enat.lean +++ b/src/data/nat/part_enat.lean @@ -5,14 +5,14 @@ Authors: Chris Hughes -/ import algebra.hom.equiv import data.part -import data.nat.lattice +import data.nat.enat import tactic.norm_num /-! # Natural numbers with infinity The natural numbers and an extra `top` element `⊤`. This implementation uses `part ℕ` as an -implementation. Use `with_top ℕ` instead unless you care about computability. +implementation. Use `ℕ∞` instead unless you care about computability. ## Main definitions @@ -24,11 +24,11 @@ The following instances are defined: There is no additive analogue of `monoid_with_zero`; if there were then `part_enat` could be an `add_monoid_with_top`. -* `to_with_top` : the map from `part_enat` to `with_top ℕ`, with theorems that it plays well +* `to_with_top` : the map from `part_enat` to `ℕ∞`, with theorems that it plays well with `+` and `≤`. -* `with_top_add_equiv : part_enat ≃+ with_top ℕ` -* `with_top_order_iso : part_enat ≃o with_top ℕ` +* `with_top_add_equiv : part_enat ≃+ ℕ∞` +* `with_top_order_iso : part_enat ≃o ℕ∞` ## Implementation details @@ -45,7 +45,7 @@ followed by `@[simp] lemma to_with_top_zero'` whose proof uses `convert`. ## Tags -part_enat, with_top ℕ +part_enat, ℕ∞ -/ open part (hiding some) @@ -392,8 +392,8 @@ by rw [add_comm a, add_comm a, part_enat.add_right_cancel_iff ha] section with_top -/-- Computably converts an `part_enat` to a `with_top ℕ`. -/ -def to_with_top (x : part_enat) [decidable x.dom] : with_top ℕ := x.to_option +/-- Computably converts an `part_enat` to a `ℕ∞`. -/ +def to_with_top (x : part_enat) [decidable x.dom] : ℕ∞ := x.to_option lemma to_with_top_top : to_with_top ⊤ = ⊤ := rfl @@ -430,11 +430,11 @@ open_locale classical @[simp] lemma to_with_top_add {x y : part_enat} : to_with_top (x + y) = to_with_top x + to_with_top y := -by apply part_enat.cases_on y; apply part_enat.cases_on x; simp [← nat.cast_add, ← with_top.coe_add] +by apply part_enat.cases_on y; apply part_enat.cases_on x; simp [← nat.cast_add, ← enat.coe_add] -/-- `equiv` between `part_enat` and `with_top ℕ` (for the order isomorphism see +/-- `equiv` between `part_enat` and `ℕ∞` (for the order isomorphism see `with_top_order_iso`). -/ -noncomputable def with_top_equiv : part_enat ≃ with_top ℕ := +noncomputable def with_top_equiv : part_enat ≃ ℕ∞ := { to_fun := λ x, to_with_top x, inv_fun := λ x, match x with (option.some n) := coe n | none := ⊤ end, left_inv := λ x, by apply part_enat.cases_on x; intros; simp; refl, @@ -455,8 +455,8 @@ to_with_top_le @[simp] lemma with_top_equiv_lt {x y : part_enat} : with_top_equiv x < with_top_equiv y ↔ x < y := to_with_top_lt -/-- `to_with_top` induces an order isomorphism between `part_enat` and `with_top ℕ`. -/ -noncomputable def with_top_order_iso : part_enat ≃o with_top ℕ := +/-- `to_with_top` induces an order isomorphism between `part_enat` and `ℕ∞`. -/ +noncomputable def with_top_order_iso : part_enat ≃o ℕ∞ := { map_rel_iff' := λ _ _, with_top_equiv_le, .. with_top_equiv} @@ -469,16 +469,16 @@ rfl @[simp] lemma with_top_equiv_symm_zero : with_top_equiv.symm 0 = 0 := rfl -@[simp] lemma with_top_equiv_symm_le {x y : with_top ℕ} : +@[simp] lemma with_top_equiv_symm_le {x y : ℕ∞} : with_top_equiv.symm x ≤ with_top_equiv.symm y ↔ x ≤ y := by rw ← with_top_equiv_le; simp -@[simp] lemma with_top_equiv_symm_lt {x y : with_top ℕ} : +@[simp] lemma with_top_equiv_symm_lt {x y : ℕ∞} : with_top_equiv.symm x < with_top_equiv.symm y ↔ x < y := by rw ← with_top_equiv_lt; simp -/-- `to_with_top` induces an additive monoid isomorphism between `part_enat` and `with_top ℕ`. -/ -noncomputable def with_top_add_equiv : part_enat ≃+ with_top ℕ := +/-- `to_with_top` induces an additive monoid isomorphism between `part_enat` and `ℕ∞`. -/ +noncomputable def with_top_add_equiv : part_enat ≃+ ℕ∞ := { map_add' := λ x y, by simp only [with_top_equiv]; convert to_with_top_add, ..with_top_equiv} diff --git a/src/data/polynomial/degree/trailing_degree.lean b/src/data/polynomial/degree/trailing_degree.lean index 2b30c3abde035..8499ebabe851e 100644 --- a/src/data/polynomial/degree/trailing_degree.lean +++ b/src/data/polynomial/degree/trailing_degree.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import data.nat.enat import data.polynomial.degree.definitions /-! @@ -35,7 +36,7 @@ variables [semiring R] {p q r : R[X]} `trailing_degree p = some n` when `p ≠ 0` and `n` is the smallest power of `X` that appears in `p`, otherwise `trailing_degree 0 = ⊤`. -/ -def trailing_degree (p : R[X]) : with_top ℕ := p.support.inf some +def trailing_degree (p : R[X]) : ℕ∞ := p.support.inf some lemma trailing_degree_lt_wf : well_founded (λp q : R[X], trailing_degree p < trailing_degree q) := @@ -70,7 +71,7 @@ lemma trailing_degree_eq_top : trailing_degree p = ⊤ ↔ p = 0 := λ h, by simp [h]⟩ lemma trailing_degree_eq_nat_trailing_degree (hp : p ≠ 0) : - trailing_degree p = (nat_trailing_degree p : with_top ℕ) := + trailing_degree p = (nat_trailing_degree p : ℕ∞) := let ⟨n, hn⟩ := not_forall.1 (mt option.eq_none_iff_forall_not_mem.2 (mt trailing_degree_eq_top.1 hp)) in have hn : trailing_degree p = some n := not_not.1 hn, @@ -93,7 +94,7 @@ end lemma nat_trailing_degree_eq_of_trailing_degree_eq_some {p : R[X]} {n : ℕ} (h : trailing_degree p = n) : nat_trailing_degree p = n := have hp0 : p ≠ 0, from λ hp0, by rw hp0 at h; exact option.no_confusion h, -option.some_inj.1 $ show (nat_trailing_degree p : with_top ℕ) = n, +option.some_inj.1 $ show (nat_trailing_degree p : ℕ∞) = n, by rwa [← trailing_degree_eq_nat_trailing_degree hp0] @[simp] lemma nat_trailing_degree_le_trailing_degree : @@ -110,7 +111,7 @@ nat_trailing_degree p = nat_trailing_degree q := by unfold nat_trailing_degree; rw h lemma le_trailing_degree_of_ne_zero (h : coeff p n ≠ 0) : trailing_degree p ≤ n := -show @has_le.le (with_top ℕ) _ (p.support.inf some : with_top ℕ) (some n : with_top ℕ), +show @has_le.le (ℕ∞) _ (p.support.inf some : ℕ∞) (some n : ℕ∞), from finset.inf_le (mem_support_iff.2 h) lemma nat_trailing_degree_le_of_ne_zero (h : coeff p n ≠ 0) : nat_trailing_degree p ≤ n := @@ -133,7 +134,7 @@ lemma trailing_degree_ne_of_nat_trailing_degree_ne {n : ℕ} : mt $ λ h, by rw [nat_trailing_degree, h, option.get_or_else_coe] theorem nat_trailing_degree_le_of_trailing_degree_le {n : ℕ} {hp : p ≠ 0} - (H : (n : with_top ℕ) ≤ trailing_degree p) : n ≤ nat_trailing_degree p := + (H : (n : ℕ∞) ≤ trailing_degree p) : n ≤ nat_trailing_degree p := begin rw trailing_degree_eq_nat_trailing_degree hp at H, exact with_top.coe_le_coe.mp H, @@ -160,13 +161,13 @@ if ha : a = 0 then by simp [ha] else (nat_trailing_degree_monomial ha).le lemma le_trailing_degree_monomial : ↑n ≤ trailing_degree (monomial n a) := if ha : a = 0 then by simp [ha] else (trailing_degree_monomial ha).ge -@[simp] lemma trailing_degree_C (ha : a ≠ 0) : trailing_degree (C a) = (0 : with_top ℕ) := +@[simp] lemma trailing_degree_C (ha : a ≠ 0) : trailing_degree (C a) = (0 : ℕ∞) := trailing_degree_monomial ha -lemma le_trailing_degree_C : (0 : with_top ℕ) ≤ trailing_degree (C a) := +lemma le_trailing_degree_C : (0 : ℕ∞) ≤ trailing_degree (C a) := le_trailing_degree_monomial -lemma trailing_degree_one_le : (0 : with_top ℕ) ≤ trailing_degree (1 : R[X]) := +lemma trailing_degree_one_le : (0 : ℕ∞) ≤ trailing_degree (1 : R[X]) := by rw [← C_1]; exact le_trailing_degree_C @[simp] lemma nat_trailing_degree_C (a : R) : nat_trailing_degree (C a) = 0 := @@ -183,10 +184,10 @@ by simp only [←C_eq_nat_cast, nat_trailing_degree_C] by rw [C_mul_X_pow_eq_monomial, trailing_degree_monomial ha] lemma le_trailing_degree_C_mul_X_pow (n : ℕ) (a : R) : - (n : with_top ℕ) ≤ trailing_degree (C a * X ^ n) := + (n : ℕ∞) ≤ trailing_degree (C a * X ^ n) := by { rw C_mul_X_pow_eq_monomial, exact le_trailing_degree_monomial } -lemma coeff_eq_zero_of_trailing_degree_lt (h : (n : with_top ℕ) < trailing_degree p) : +lemma coeff_eq_zero_of_trailing_degree_lt (h : (n : ℕ∞) < trailing_degree p) : coeff p n = 0 := not_not.1 (mt le_trailing_degree_of_ne_zero (not_le_of_gt h)) @@ -201,15 +202,15 @@ begin end @[simp] lemma coeff_nat_trailing_degree_pred_eq_zero {p : R[X]} - {hp : (0 : with_top ℕ) < nat_trailing_degree p} : p.coeff (p.nat_trailing_degree - 1) = 0 := + {hp : (0 : ℕ∞) < nat_trailing_degree p} : p.coeff (p.nat_trailing_degree - 1) = 0 := coeff_eq_zero_of_lt_nat_trailing_degree $ nat.sub_lt ((with_top.zero_lt_coe (nat_trailing_degree p)).mp hp) nat.one_pos theorem le_trailing_degree_X_pow (n : ℕ) : - (n : with_top ℕ) ≤ trailing_degree (X^n : R[X]) := + (n : ℕ∞) ≤ trailing_degree (X^n : R[X]) := by simpa only [C_1, one_mul] using le_trailing_degree_C_mul_X_pow n (1:R) -theorem le_trailing_degree_X : (1 : with_top ℕ) ≤ trailing_degree (X : R[X]) := +theorem le_trailing_degree_X : (1 : ℕ∞) ≤ trailing_degree (X : R[X]) := le_trailing_degree_monomial lemma nat_trailing_degree_X_le : (X : R[X]).nat_trailing_degree ≤ 1 := @@ -272,7 +273,7 @@ begin refine (add_le_add (inf_le (mem_support_iff.mpr (left_ne_zero_of_mul hpq))) (inf_le (mem_support_iff.mpr (right_ne_zero_of_mul hpq)))).trans (le_of_eq _), rwa [with_top.some_eq_coe, with_top.some_eq_coe, with_top.some_eq_coe, - ←with_top.coe_add, with_top.coe_eq_coe, ←nat.mem_antidiagonal], + ← with_top.coe_add, with_top.coe_eq_coe, ←nat.mem_antidiagonal], end lemma le_nat_trailing_degree_mul (h : p * q ≠ 0) : @@ -309,7 +310,8 @@ begin have hp : p ≠ 0 := λ hp, h (by rw [hp, trailing_coeff_zero, zero_mul]), have hq : q ≠ 0 := λ hq, h (by rw [hq, trailing_coeff_zero, mul_zero]), refine le_antisymm _ le_trailing_degree_mul, - rw [trailing_degree_eq_nat_trailing_degree hp, trailing_degree_eq_nat_trailing_degree hq], + rw [trailing_degree_eq_nat_trailing_degree hp, trailing_degree_eq_nat_trailing_degree hq, + ← enat.coe_add], apply le_trailing_degree_of_ne_zero, rwa coeff_mul_nat_trailing_degree_add_nat_trailing_degree, end @@ -334,7 +336,7 @@ end semiring section nonzero_semiring variables [semiring R] [nontrivial R] {p q : R[X]} -@[simp] lemma trailing_degree_one : trailing_degree (1 : R[X]) = (0 : with_top ℕ) := +@[simp] lemma trailing_degree_one : trailing_degree (1 : R[X]) = (0 : ℕ∞) := trailing_degree_C one_ne_zero @[simp] lemma trailing_degree_X : trailing_degree (X : R[X]) = 1 := @@ -385,7 +387,7 @@ lemma coeff_nat_trailing_degree_eq_zero_of_trailing_degree_lt coeff q (nat_trailing_degree p) = 0 := coeff_eq_zero_of_trailing_degree_lt $ nat_trailing_degree_le_trailing_degree.trans_lt h -lemma ne_zero_of_trailing_degree_lt {n : with_top ℕ} (h : trailing_degree p < n) : p ≠ 0 := +lemma ne_zero_of_trailing_degree_lt {n : ℕ∞} (h : trailing_degree p < n) : p ≠ 0 := λ h₀, h.not_le (by simp [h₀]) end semiring diff --git a/src/geometry/manifold/algebra/monoid.lean b/src/geometry/manifold/algebra/monoid.lean index ce893551579e5..1daf60fb0b289 100644 --- a/src/geometry/manifold/algebra/monoid.lean +++ b/src/geometry/manifold/algebra/monoid.lean @@ -91,7 +91,7 @@ end section -variables {f g : M → G} {s : set M} {x : M} {n : with_top ℕ} +variables {f g : M → G} {s : set M} {x : M} {n : ℕ∞} @[to_additive] lemma cont_mdiff_within_at.mul (hf : cont_mdiff_within_at I' I n f s x) @@ -255,7 +255,7 @@ variables {ι 𝕜 : Type*} [nontrivially_normed_field 𝕜] {E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] {H' : Type*} [topological_space H'] {I' : model_with_corners 𝕜 E' H'} {M : Type*} [topological_space M] [charted_space H' M] {s : set M} {x : M} -{t : finset ι} {f : ι → M → G} {n : with_top ℕ} {p : ι → Prop} +{t : finset ι} {f : ι → M → G} {n : ℕ∞} {p : ι → Prop} @[to_additive] lemma cont_mdiff_within_at_finset_prod' (h : ∀ i ∈ t, cont_mdiff_within_at I' I n (f i) s x) : diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index e1d310f87da4d..4b7650eb57e3d 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -70,12 +70,12 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {G' : Type*} [topological_space G'] {J' : model_with_corners 𝕜 F' G'} {N' : Type*} [topological_space N'] [charted_space G' N'] [J's : smooth_manifold_with_corners J' N'] -- declare functions, sets, points and smoothness indices -{f f₁ : M → M'} {s s₁ t : set M} {x : M} {m n : with_top ℕ} +{f f₁ : M → M'} {s s₁ t : set M} {x : M} {m n : ℕ∞} /-- Property in the model space of a model with corners of being `C^n` within at set at a point, when read in the model vector space. This property will be lifted to manifolds to define smooth functions between manifolds. -/ -def cont_diff_within_at_prop (n : with_top ℕ) (f : H → H') (s : set H) (x : H) : Prop := +def cont_diff_within_at_prop (n : ℕ∞) (f : H → H') (s : set H) (x : H) : Prop := cont_diff_within_at 𝕜 n (I' ∘ f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) (I x) lemma cont_diff_within_at_prop_self_source {f : E → H'} {s : set E} {x : E} : @@ -96,7 +96,7 @@ iff.rfl /-- Being `Cⁿ` in the model space is a local property, invariant under smooth maps. Therefore, it will lift nicely to manifolds. -/ -lemma cont_diff_within_at_local_invariant_prop (n : with_top ℕ) : +lemma cont_diff_within_at_local_invariant_prop (n : ℕ∞) : (cont_diff_groupoid ∞ I).local_invariant_prop (cont_diff_groupoid ∞ I') (cont_diff_within_at_prop I I' n) := { is_local := @@ -146,7 +146,7 @@ lemma cont_diff_within_at_local_invariant_prop (n : with_top ℕ) : { assume y hy, simp only with mfld_simps at hy, simpa only [hy] with mfld_simps using hs hy.1 } end } -lemma cont_diff_within_at_prop_mono (n : with_top ℕ) +lemma cont_diff_within_at_prop_mono (n : ℕ∞) ⦃s x t⦄ ⦃f : H → H'⦄ (hts : t ⊆ s) (h : cont_diff_within_at_prop I I' n f s x) : cont_diff_within_at_prop I I' n f t x := begin @@ -169,7 +169,7 @@ end /-- A function is `n` times continuously differentiable within a set at a point in a manifold if it is continuous and it is `n` times continuously differentiable in this set around this point, when read in the preferred chart at this point. -/ -def cont_mdiff_within_at (n : with_top ℕ) (f : M → M') (s : set M) (x : M) := +def cont_mdiff_within_at (n : ℕ∞) (f : M → M') (s : set M) (x : M) := lift_prop_within_at (cont_diff_within_at_prop I I' n) f s x /-- Abbreviation for `cont_mdiff_within_at I I' ⊤ f s x`. See also documentation for `smooth`. @@ -180,10 +180,10 @@ cont_mdiff_within_at I I' ⊤ f s x /-- A function is `n` times continuously differentiable at a point in a manifold if it is continuous and it is `n` times continuously differentiable around this point, when read in the preferred chart at this point. -/ -def cont_mdiff_at (n : with_top ℕ) (f : M → M') (x : M) := +def cont_mdiff_at (n : ℕ∞) (f : M → M') (x : M) := cont_mdiff_within_at I I' n f univ x -lemma cont_mdiff_at_iff {n : with_top ℕ} {f : M → M'} {x : M} : +lemma cont_mdiff_at_iff {n : ℕ∞} {f : M → M'} {x : M} : cont_mdiff_at I I' n f x ↔ continuous_at f x ∧ cont_diff_within_at 𝕜 n (ext_chart_at I' (f x) ∘ f ∘ (ext_chart_at I x).symm) (range I) (ext_chart_at I x x) := lift_prop_at_iff.trans $ by { rw [cont_diff_within_at_prop, preimage_univ, univ_inter], refl } @@ -194,7 +194,7 @@ lift_prop_at_iff.trans $ by { rw [cont_diff_within_at_prop, preimage_univ, univ_ /-- A function is `n` times continuously differentiable in a set of a manifold if it is continuous and, for any pair of points, it is `n` times continuously differentiable on this set in the charts around these points. -/ -def cont_mdiff_on (n : with_top ℕ) (f : M → M') (s : set M) := +def cont_mdiff_on (n : ℕ∞) (f : M → M') (s : set M) := ∀ x ∈ s, cont_mdiff_within_at I I' n f s x /-- Abbreviation for `cont_mdiff_on I I' ⊤ f s`. See also documentation for `smooth`. -/ @@ -203,7 +203,7 @@ def cont_mdiff_on (n : with_top ℕ) (f : M → M') (s : set M) := /-- A function is `n` times continuously differentiable in a manifold if it is continuous and, for any pair of points, it is `n` times continuously differentiable in the charts around these points. -/ -def cont_mdiff (n : with_top ℕ) (f : M → M') := +def cont_mdiff (n : ℕ∞) (f : M → M') := ∀ x, cont_mdiff_at I I' n f x /-- Abbreviation for `cont_mdiff I I' ⊤ f`. @@ -699,7 +699,7 @@ lemma cont_mdiff_top : lemma cont_mdiff_within_at_iff_nat : cont_mdiff_within_at I I' n f s x ↔ - (∀m:ℕ, (m : with_top ℕ) ≤ n → cont_mdiff_within_at I I' m f s x) := + (∀m:ℕ, (m : ℕ∞) ≤ n → cont_mdiff_within_at I I' m f s x) := begin refine ⟨λ h m hm, h.of_le hm, λ h, _⟩, cases n, @@ -1552,7 +1552,7 @@ variables (Z : basic_smooth_vector_bundle_core I M E') /-- A version of `cont_mdiff_at_iff_target` when the codomain is the total space of a `basic_smooth_vector_bundle_core`. The continuity condition in the RHS is weaker. -/ lemma cont_mdiff_at_iff_target {f : N → Z.to_topological_vector_bundle_core.total_space} - {x : N} {n : with_top ℕ} : + {x : N} {n : ℕ∞} : cont_mdiff_at J (I.prod 𝓘(𝕜, E')) n f x ↔ continuous_at (bundle.total_space.proj ∘ f) x ∧ cont_mdiff_at J 𝓘(𝕜, E × E') n (ext_chart_at (I.prod 𝓘(𝕜, E')) (f x) ∘ f) x := begin diff --git a/src/geometry/manifold/cont_mdiff_map.lean b/src/geometry/manifold/cont_mdiff_map.lean index eb0ec6c483c5a..69f5cf0955d30 100644 --- a/src/geometry/manifold/cont_mdiff_map.lean +++ b/src/geometry/manifold/cont_mdiff_map.lean @@ -26,7 +26,7 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {H'' : Type*} [topological_space H''] {I'' : model_with_corners 𝕜 E'' H''} {M'' : Type*} [topological_space M''] [charted_space H'' M''] -(n : with_top ℕ) +(n : ℕ∞) /-- Bundled `n` times continuously differentiable maps. -/ @[protect_proj] diff --git a/src/geometry/manifold/derivation_bundle.lean b/src/geometry/manifold/derivation_bundle.lean index 68308fc1b2656..e110fe02e4ee6 100644 --- a/src/geometry/manifold/derivation_bundle.lean +++ b/src/geometry/manifold/derivation_bundle.lean @@ -23,7 +23,7 @@ of the Lie algebra for a Lie group. variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) -(M : Type*) [topological_space M] [charted_space H M] (n : with_top ℕ) +(M : Type*) [topological_space M] [charted_space H M] (n : ℕ∞) open_locale manifold diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index a1b654d4ce246..e452c81399888 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -61,7 +61,7 @@ variables {M : Type*} [topological_space M] [charted_space H M] {M' : Type*} [topological_space M'] [charted_space H' M'] {N : Type*} [topological_space N] [charted_space G N] {N' : Type*} [topological_space N'] [charted_space G' N'] -{n : with_top ℕ} +{n : ℕ∞} section defs diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index f63a5beffe4c5..7424b08e368f3 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -381,7 +381,7 @@ variables {M : Type*} [topological_space M] [charted_space H M] [smooth_manifold /-- If a `cont_mdiff` function `f : M → E`, where `M` is some manifold, takes values in the sphere, then it restricts to a `cont_mdiff` function from `M` to the sphere. -/ lemma cont_mdiff.cod_restrict_sphere {n : ℕ} [fact (finrank ℝ E = n + 1)] - {m : with_top ℕ} {f : M → E} (hf : cont_mdiff I 𝓘(ℝ, E) m f) + {m : ℕ∞} {f : M → E} (hf : cont_mdiff I 𝓘(ℝ, E) m f) (hf' : ∀ x, f x ∈ sphere (0:E) 1) : cont_mdiff I (𝓡 n) m (set.cod_restrict _ _ hf' : M → (sphere (0:E) 1)) := begin diff --git a/src/geometry/manifold/partition_of_unity.lean b/src/geometry/manifold/partition_of_unity.lean index d5c795f33cc34..8d129b488da36 100644 --- a/src/geometry/manifold/partition_of_unity.lean +++ b/src/geometry/manifold/partition_of_unity.lean @@ -122,7 +122,7 @@ variables {ι I M} namespace smooth_partition_of_unity -variables {s : set M} (f : smooth_partition_of_unity ι I M s) {n : with_top ℕ} +variables {s : set M} (f : smooth_partition_of_unity ι I M s) {n : ℕ∞} instance {s : set M} : has_coe_to_fun (smooth_partition_of_unity ι I M s) (λ _, ι → C^∞⟮I, M; 𝓘(ℝ), ℝ⟯) := @@ -458,7 +458,7 @@ end end smooth_partition_of_unity -variables [sigma_compact_space M] [t2_space M] {t : M → set F} {n : with_top ℕ} +variables [sigma_compact_space M] [t2_space M] {t : M → set F} {n : ℕ∞} /-- Let `M` be a σ-compact Hausdorff finite dimensional topological manifold. Let `t : M → set F` be a family of convex sets. Suppose that for each point `x : M` there exists a neighborhood diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index 7125c5da404aa..2c8fff7140c7d 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -14,7 +14,7 @@ half-space (to get manifolds with boundaries) for which the changes of coordinat We define a model with corners as a map `I : H → E` embedding nicely the topological space `H` in the vector space `E` (or more precisely as a structure containing all the relevant properties). Given such a model with corners `I` on `(E, H)`, we define the groupoid of local -homeomorphisms of `H` which are smooth when read in `E` (for any regularity `n : with_top ℕ`). +homeomorphisms of `H` which are smooth when read in `E` (for any regularity `n : ℕ∞`). With this groupoid at hand and the general machinery of charted spaces, we thus get the notion of `C^n` manifold with respect to any model with corners `I` on `(E, H)`. We also introduce a specific type class for `C^∞` manifolds as these are the most commonly used. @@ -116,7 +116,7 @@ universes u v w u' v' w' open set filter function open_locale manifold filter topological_space -localized "notation `∞` := (⊤ : with_top ℕ)" in manifold +localized "notation `∞` := (⊤ : ℕ∞)" in manifold /-! ### Models with corners. -/ @@ -398,7 +398,7 @@ end boundaryless section cont_diff_groupoid /-! ### Smooth functions on models with corners -/ -variables {m n : with_top ℕ} {𝕜 : Type*} [nontrivially_normed_field 𝕜] +variables {m n : ℕ∞} {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) diff --git a/src/probability/hitting_time.lean b/src/probability/hitting_time.lean index 6e148b94e46b7..89b76ec997164 100644 --- a/src/probability/hitting_time.lean +++ b/src/probability/hitting_time.lean @@ -28,7 +28,7 @@ In the definition of the hitting time, we bound the hitting time by an upper and This is to ensure that our result is meaningful in the case we are taking the infimum of an empty set or the infimum of a set which is unbounded from below. With this, we can talk about hitting times indexed by the natural numbers or the reals. By taking the bounds to be -`⊤` and `⊥`, we obtain the standard definition in the case that the index is `with_top ℕ` or `ℝ≥0∞`. +`⊤` and `⊥`, we obtain the standard definition in the case that the index is `ℕ∞` or `ℝ≥0∞`. -/ diff --git a/src/ring_theory/int/basic.lean b/src/ring_theory/int/basic.lean index c138a37d10f56..f17203787592f 100644 --- a/src/ring_theory/int/basic.lean +++ b/src/ring_theory/int/basic.lean @@ -31,7 +31,7 @@ namespace nat instance : wf_dvd_monoid ℕ := ⟨begin refine rel_hom_class.well_founded - (⟨λ (x : ℕ), if x = 0 then (⊤ : with_top ℕ) else x, _⟩ : dvd_not_unit →r (<)) + (⟨λ (x : ℕ), if x = 0 then (⊤ : ℕ∞) else x, _⟩ : dvd_not_unit →r (<)) (with_top.well_founded_lt nat.lt_wf), intros a b h, cases a, diff --git a/src/ring_theory/unique_factorization_domain.lean b/src/ring_theory/unique_factorization_domain.lean index 95d23ef4a22df..d8f2fe739ec38 100644 --- a/src/ring_theory/unique_factorization_domain.lean +++ b/src/ring_theory/unique_factorization_domain.lean @@ -250,7 +250,7 @@ lemma wf_dvd_monoid.of_exists_prime_factors : wf_dvd_monoid α := ⟨begin classical, refine rel_hom_class.well_founded - (rel_hom.mk _ _ : (dvd_not_unit : α → α → Prop) →r ((<) : with_top ℕ → with_top ℕ → Prop)) + (rel_hom.mk _ _ : (dvd_not_unit : α → α → Prop) →r ((<) : ℕ∞ → ℕ∞ → Prop)) (with_top.well_founded_lt nat.lt_wf), { intro a, by_cases h : a = 0, { exact ⊤ }, From 3339976e2bcae9f1c81e620836d1eb736e3c4700 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 24 Aug 2022 12:36:12 +0000 Subject: [PATCH 020/828] refactor(topology/constructions): rename `continuous_subtype_mk` (#16223) * rename `continuous_subtype_mk` to `continuous.subtype_mk` to allow dot notation; * add `continuous.subtype_map`. --- .../topological_simplex.lean | 4 +-- src/analysis/convex/contractible.lean | 7 ++--- src/analysis/normed/field/unit_ball.lean | 3 +- src/analysis/normed/group/ball_sphere.lean | 15 ++++------ src/analysis/normed/group/basic.lean | 29 +++++++++---------- src/analysis/normed_space/ball_action.lean | 15 ++++------ src/analysis/normed_space/basic.lean | 4 +-- src/analysis/special_functions/exp.lean | 2 +- src/analysis/special_functions/log/basic.lean | 2 +- src/dynamics/flow.lean | 3 +- src/geometry/manifold/instances/real.lean | 16 +++++----- src/topology/algebra/module/basic.lean | 10 +++---- src/topology/algebra/order/proj_Icc.lean | 2 +- src/topology/algebra/semigroup.lean | 4 +-- src/topology/bases.lean | 3 +- src/topology/category/CompHaus/default.lean | 2 +- src/topology/constructions.lean | 12 +++++--- src/topology/continuous_function/bounded.lean | 2 +- src/topology/homeomorph.lean | 19 ++++++------ src/topology/instances/ennreal.lean | 2 +- src/topology/instances/ereal.lean | 2 +- src/topology/instances/nnreal.lean | 20 ++++++------- src/topology/instances/real.lean | 2 +- src/topology/local_homeomorph.lean | 8 ++--- src/topology/metric_space/basic.lean | 2 +- src/topology/metric_space/pi_nat.lean | 2 +- src/topology/tietze_extension.lean | 2 +- src/topology/uniform_space/basic.lean | 2 +- src/topology/uniform_space/equiv.lean | 16 +++++----- 29 files changed, 99 insertions(+), 113 deletions(-) diff --git a/src/algebraic_topology/topological_simplex.lean b/src/algebraic_topology/topological_simplex.lean index abf904e11b970..920922530903d 100644 --- a/src/algebraic_topology/topological_simplex.lean +++ b/src/algebraic_topology/topological_simplex.lean @@ -62,8 +62,8 @@ lemma coe_to_Top_map {x y : simplex_category} (f : x ⟶ y) (g : x.to_Top_obj) ( @[continuity] lemma continuous_to_Top_map {x y : simplex_category} (f : x ⟶ y) : continuous (to_Top_map f) := -continuous_subtype_mk _ $ continuous_pi $ λ i, continuous_finset_sum _ $ - λ j hj, continuous.comp (continuous_apply _) continuous_subtype_val +continuous.subtype_mk (continuous_pi $ λ i, continuous_finset_sum _ $ + λ j hj, (continuous_apply _).comp continuous_subtype_val) _ /-- The functor associating the topological `n`-simplex to `[n] : simplex_category`. -/ @[simps] diff --git a/src/analysis/convex/contractible.lean b/src/analysis/convex/contractible.lean index cd5d483023228..6d06aa1c54b9e 100644 --- a/src/analysis/convex/contractible.lean +++ b/src/analysis/convex/contractible.lean @@ -23,10 +23,9 @@ begin refine (contractible_iff_id_nullhomotopic _).2 ⟨⟨x, h.mem hne⟩, ⟨⟨⟨λ p, ⟨p.1.1 • x + (1 - p.1.1) • p.2, _⟩, _⟩, λ x, _, λ x, _⟩⟩⟩, { exact h p.2.2 p.1.2.1 (sub_nonneg.2 p.1.2.2) (add_sub_cancel'_right _ _) }, - { exact continuous_subtype_mk _ - (((continuous_subtype_val.comp continuous_fst).smul continuous_const).add - ((continuous_const.sub $ continuous_subtype_val.comp continuous_fst).smul - (continuous_subtype_val.comp continuous_snd))) }, + { exact ((continuous_subtype_val.fst'.smul continuous_const).add + ((continuous_const.sub continuous_subtype_val.fst').smul + continuous_subtype_val.snd')).subtype_mk _ }, { ext1, simp }, { ext1, simp } end diff --git a/src/analysis/normed/field/unit_ball.lean b/src/analysis/normed/field/unit_ball.lean index 3b52ad13ab74a..9de3087c3a72d 100644 --- a/src/analysis/normed/field/unit_ball.lean +++ b/src/analysis/normed/field/unit_ball.lean @@ -156,8 +156,7 @@ subtype.coe_injective.has_distrib_neg (coe : sphere (0 : 𝕜) 1 → 𝕜) (λ _ instance [normed_division_ring 𝕜] : topological_group (sphere (0 : 𝕜) 1) := { to_has_continuous_mul := (submonoid.unit_sphere 𝕜).has_continuous_mul, - continuous_inv := continuous_subtype_mk _ $ - continuous_subtype_coe.inv₀ ne_zero_of_mem_unit_sphere } + continuous_inv := (continuous_subtype_coe.inv₀ ne_zero_of_mem_unit_sphere).subtype_mk _ } instance [normed_field 𝕜] : comm_group (sphere (0 : 𝕜) 1) := { .. metric.sphere.group, diff --git a/src/analysis/normed/group/ball_sphere.lean b/src/analysis/normed/group/ball_sphere.lean index 6ff647c77a400..64d63bc2ecf10 100644 --- a/src/analysis/normed/group/ball_sphere.lean +++ b/src/analysis/normed/group/ball_sphere.lean @@ -19,38 +19,35 @@ variables {E : Type*} [seminormed_add_comm_group E] {r : ℝ} /-- We equip the sphere, in a seminormed group, with a formal operation of negation, namely the antipodal map. -/ instance : has_involutive_neg (sphere (0 : E) r) := -{ neg := λ w, ⟨-↑w, by simp⟩, +{ neg := subtype.map has_neg.neg $ λ w, by simp, neg_neg := λ x, subtype.ext $ neg_neg x } @[simp] lemma coe_neg_sphere {r : ℝ} (v : sphere (0 : E) r) : ↑(-v) = (-v : E) := rfl -instance : has_continuous_neg (sphere (0 : E) r) := -⟨continuous_subtype_mk _ continuous_subtype_coe.neg⟩ +instance : has_continuous_neg (sphere (0 : E) r) := ⟨continuous_neg.subtype_map _⟩ /-- We equip the ball, in a seminormed group, with a formal operation of negation, namely the antipodal map. -/ instance {r : ℝ} : has_involutive_neg (ball (0 : E) r) := -{ neg := λ w, ⟨-↑w, by simpa using w.coe_prop⟩, +{ neg := subtype.map has_neg.neg $ λ w, by simp, neg_neg := λ x, subtype.ext $ neg_neg x } @[simp] lemma coe_neg_ball {r : ℝ} (v : ball (0 : E) r) : ↑(-v) = (-v : E) := rfl -instance : has_continuous_neg (ball (0 : E) r) := -⟨continuous_subtype_mk _ continuous_subtype_coe.neg⟩ +instance : has_continuous_neg (ball (0 : E) r) := ⟨continuous_neg.subtype_map _⟩ /-- We equip the closed ball, in a seminormed group, with a formal operation of negation, namely the antipodal map. -/ instance {r : ℝ} : has_involutive_neg (closed_ball (0 : E) r) := -{ neg := λ w, ⟨-↑w, by simpa using w.coe_prop⟩, +{ neg := subtype.map has_neg.neg $ λ w, by simp, neg_neg := λ x, subtype.ext $ neg_neg x } @[simp] lemma coe_neg_closed_ball {r : ℝ} (v : closed_ball (0 : E) r) : ↑(-v) = (-v : E) := rfl -instance : has_continuous_neg (closed_ball (0 : E) r) := -⟨continuous_subtype_mk _ continuous_subtype_coe.neg⟩ +instance : has_continuous_neg (closed_ball (0 : E) r) := ⟨continuous_neg.subtype_map _⟩ diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index f893d68663ecf..5458632c8ed25 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -913,20 +913,6 @@ squeeze_zero_norm' (eventually_of_forall h) h' lemma tendsto_norm_sub_self (x : E) : tendsto (λ g : E, ∥g - x∥) (𝓝 x) (𝓝 0) := by simpa [dist_eq_norm] using tendsto_id.dist (tendsto_const_nhds : tendsto (λ g, (x:E)) (𝓝 x) _) -lemma tendsto_norm {x : E} : tendsto (λg : E, ∥g∥) (𝓝 x) (𝓝 ∥x∥) := -by simpa using tendsto_id.dist (tendsto_const_nhds : tendsto (λ g, (0:E)) _ _) - -lemma tendsto_norm_zero : tendsto (λg : E, ∥g∥) (𝓝 0) (𝓝 0) := -by simpa using tendsto_norm_sub_self (0:E) - -@[continuity] -lemma continuous_norm : continuous (λg:E, ∥g∥) := -by simpa using continuous_id.dist (continuous_const : continuous (λ g, (0:E))) - -@[continuity] -lemma continuous_nnnorm : continuous (λ (a : E), ∥a∥₊) := -continuous_subtype_mk _ continuous_norm - lemma lipschitz_with_one_norm : lipschitz_with 1 (norm : E → ℝ) := by simpa only [dist_zero_left] using lipschitz_with.dist_right (0 : E) @@ -937,7 +923,20 @@ lemma uniform_continuous_norm : uniform_continuous (norm : E → ℝ) := lipschitz_with_one_norm.uniform_continuous lemma uniform_continuous_nnnorm : uniform_continuous (λ (a : E), ∥a∥₊) := -uniform_continuous_subtype_mk uniform_continuous_norm _ +uniform_continuous_norm.subtype_mk _ + +@[continuity] lemma continuous_norm : continuous (λg:E, ∥g∥) := +uniform_continuous_norm.continuous + +@[continuity] +lemma continuous_nnnorm : continuous (λ (a : E), ∥a∥₊) := +uniform_continuous_nnnorm.continuous + +lemma tendsto_norm {x : E} : tendsto (λg : E, ∥g∥) (𝓝 x) (𝓝 ∥x∥) := +continuous_norm.continuous_at + +lemma tendsto_norm_zero : tendsto (λg : E, ∥g∥) (𝓝 0) (𝓝 0) := +continuous_norm.tendsto' 0 0 norm_zero /-- A helper lemma used to prove that the (scalar or usual) product of a function that tends to zero and a bounded function tends to zero. This lemma is formulated for any binary operation diff --git a/src/analysis/normed_space/ball_action.lean b/src/analysis/normed_space/ball_action.lean index be948c62c24df..472cc334c00a8 100644 --- a/src/analysis/normed_space/ball_action.lean +++ b/src/analysis/normed_space/ball_action.lean @@ -31,8 +31,7 @@ instance mul_action_closed_ball_ball : mul_action (closed_ball (0 : 𝕜) 1) (ba instance has_continuous_smul_closed_ball_ball : has_continuous_smul (closed_ball (0 : 𝕜) 1) (ball (0 : E) r) := -⟨continuous_subtype_mk _ $ (continuous_subtype_val.comp continuous_fst).smul - (continuous_subtype_val.comp continuous_snd)⟩ +⟨(continuous_subtype_val.fst'.smul continuous_subtype_val.snd').subtype_mk _⟩ instance mul_action_closed_ball_closed_ball : mul_action (closed_ball (0 : 𝕜) 1) (closed_ball (0 : E) r) := @@ -45,8 +44,7 @@ instance mul_action_closed_ball_closed_ball : instance has_continuous_smul_closed_ball_closed_ball : has_continuous_smul (closed_ball (0 : 𝕜) 1) (closed_ball (0 : E) r) := -⟨continuous_subtype_mk _ $ (continuous_subtype_val.comp continuous_fst).smul - (continuous_subtype_val.comp continuous_snd)⟩ +⟨(continuous_subtype_val.fst'.smul continuous_subtype_val.snd').subtype_mk _⟩ end closed_ball @@ -59,8 +57,7 @@ instance mul_action_sphere_ball : mul_action (sphere (0 : 𝕜) 1) (ball (0 : E) instance has_continuous_smul_sphere_ball : has_continuous_smul (sphere (0 : 𝕜) 1) (ball (0 : E) r) := -⟨continuous_subtype_mk _ $ (continuous_subtype_val.comp continuous_fst).smul - (continuous_subtype_val.comp continuous_snd)⟩ +⟨(continuous_subtype_val.fst'.smul continuous_subtype_val.snd').subtype_mk _⟩ instance mul_action_sphere_closed_ball : mul_action (sphere (0 : 𝕜) 1) (closed_ball (0 : E) r) := { smul := λ c x, inclusion sphere_subset_closed_ball c • x, @@ -69,8 +66,7 @@ instance mul_action_sphere_closed_ball : mul_action (sphere (0 : 𝕜) 1) (close instance has_continuous_smul_sphere_closed_ball : has_continuous_smul (sphere (0 : 𝕜) 1) (closed_ball (0 : E) r) := -⟨continuous_subtype_mk _ $ (continuous_subtype_val.comp continuous_fst).smul - (continuous_subtype_val.comp continuous_snd)⟩ +⟨(continuous_subtype_val.fst'.smul continuous_subtype_val.snd').subtype_mk _⟩ instance mul_action_sphere_sphere : mul_action (sphere (0 : 𝕜) 1) (sphere (0 : E) r) := { smul := λ c x, ⟨(c : 𝕜) • x, mem_sphere_zero_iff_norm.2 $ @@ -81,8 +77,7 @@ instance mul_action_sphere_sphere : mul_action (sphere (0 : 𝕜) 1) (sphere (0 instance has_continuous_smul_sphere_sphere : has_continuous_smul (sphere (0 : 𝕜) 1) (sphere (0 : E) r) := -⟨continuous_subtype_mk _ $ (continuous_subtype_val.comp continuous_fst).smul - (continuous_subtype_val.comp continuous_snd)⟩ +⟨(continuous_subtype_val.fst'.smul continuous_subtype_val.snd').subtype_mk _⟩ end sphere diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 28027c4731564..5f0dc2f80afe5 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -208,9 +208,9 @@ def homeomorph_unit_ball [normed_space ℝ E] : by nlinarith [norm_nonneg (y : E), (mem_ball_zero_iff.1 y.2 : ∥(y : E)∥ < 1)], field_simp [norm_smul, smul_smul, this.ne', real.sq_sqrt this.le, ← real.sqrt_div this.le], end, - continuous_to_fun := continuous_subtype_mk _ $ + continuous_to_fun := begin - suffices : continuous (λ x, (1 + ∥x∥^2).sqrt⁻¹), { exact this.smul continuous_id, }, + suffices : continuous (λ x, (1 + ∥x∥^2).sqrt⁻¹), from (this.smul continuous_id).subtype_mk _, refine continuous.inv₀ _ (λ x, real.sqrt_ne_zero'.mpr (by positivity)), continuity, end, diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index 16febaeed988d..11de0b639b955 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -229,7 +229,7 @@ end /-- `real.exp` as an order isomorphism between `ℝ` and `(0, +∞)`. -/ def exp_order_iso : ℝ ≃o Ioi (0 : ℝ) := strict_mono.order_iso_of_surjective _ (exp_strict_mono.cod_restrict exp_pos) $ - (continuous_subtype_mk _ continuous_exp).surjective + (continuous_exp.subtype_mk _).surjective (by simp only [tendsto_Ioi_at_top, subtype.coe_mk, tendsto_exp_at_top]) (by simp [tendsto_exp_at_bot_nhds_within]) diff --git a/src/analysis/special_functions/log/basic.lean b/src/analysis/special_functions/log/basic.lean index 8cd71b90ef6fa..a4ffa13124f6f 100644 --- a/src/analysis/special_functions/log/basic.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -240,7 +240,7 @@ lemma continuous_on_log : continuous_on log {0}ᶜ := begin rw [continuous_on_iff_continuous_restrict, restrict], conv in (log _) { rw [log_of_ne_zero (show (x : ℝ) ≠ 0, from x.2)] }, - exact exp_order_iso.symm.continuous.comp (continuous_subtype_mk _ continuous_subtype_coe.norm) + exact exp_order_iso.symm.continuous.comp (continuous_subtype_coe.norm.subtype_mk _) end @[continuity] lemma continuous_log : continuous (λ x : {x : ℝ // x ≠ 0}, log x) := diff --git a/src/dynamics/flow.lean b/src/dynamics/flow.lean index ca3dcf4650506..7cb4766756e7b 100644 --- a/src/dynamics/flow.lean +++ b/src/dynamics/flow.lean @@ -129,8 +129,7 @@ def from_iter {g : α → α} (h : continuous g) : flow ℕ α := /-- Restriction of a flow onto an invariant set. -/ def restrict {s : set α} (h : is_invariant ϕ s) : flow τ ↥s := { to_fun := λ t, (h t).restrict _ _ _, - cont' := continuous_subtype_mk _ (ϕ.continuous continuous_fst - (continuous_subtype_coe.comp continuous_snd)), + cont' := (ϕ.continuous continuous_fst continuous_subtype_coe.snd').subtype_mk _, map_add' := λ _ _ _, subtype.ext (map_add _ _ _ _), map_zero' := λ _, subtype.ext (map_zero_apply _ _)} diff --git a/src/geometry/manifold/instances/real.lean b/src/geometry/manifold/instances/real.lean index 1e6df2db0d821..eb90335fedeeb 100644 --- a/src/geometry/manifold/instances/real.lean +++ b/src/geometry/manifold/instances/real.lean @@ -98,8 +98,8 @@ def model_with_corners_euclidean_half_space (n : ℕ) [has_zero (fin n)] : unique_diff_on.pi (fin n) (λ _, ℝ) _ _ (λ i ∈ ({0} : set (fin n)), unique_diff_on_Ici 0), by simpa only [singleton_pi] using this, continuous_to_fun := continuous_subtype_val, - continuous_inv_fun := continuous_subtype_mk _ $ continuous_id.update 0 $ - (continuous_apply 0).max continuous_const } + continuous_inv_fun := (continuous_id.update 0 $ + (continuous_apply 0).max continuous_const).subtype_mk _ } /-- Definition of the model with corners `(euclidean_space ℝ (fin n), euclidean_quadrant n)`, used as a @@ -120,8 +120,8 @@ def model_with_corners_euclidean_quadrant (n : ℕ) : unique_diff_on.univ_pi (fin n) (λ _, ℝ) _ (λ i, unique_diff_on_Ici 0), by simpa only [pi_univ_Ici] using this, continuous_to_fun := continuous_subtype_val, - continuous_inv_fun := continuous_subtype_mk _ $ continuous_pi $ λ i, - (continuous_id.max continuous_const).comp (continuous_apply i) } + continuous_inv_fun := continuous.subtype_mk (continuous_pi $ λ i, + (continuous_id.max continuous_const).comp (continuous_apply i)) _ } localized "notation `𝓡 `n := (model_with_corners_self ℝ (euclidean_space ℝ (fin n)) : @@ -170,14 +170,14 @@ def Icc_left_chart (x y : ℝ) [fact (x < y)] : end, continuous_to_fun := begin apply continuous.continuous_on, - apply continuous_subtype_mk, + apply continuous.subtype_mk, have : continuous (λ (z : ℝ) (i : fin 1), z - x) := continuous.sub (continuous_pi $ λi, continuous_id) continuous_const, exact this.comp continuous_subtype_val, end, continuous_inv_fun := begin apply continuous.continuous_on, - apply continuous_subtype_mk, + apply continuous.subtype_mk, have A : continuous (λ z : ℝ, min (z + x) y) := (continuous_id.add continuous_const).min continuous_const, have B : continuous (λz : euclidean_space ℝ (fin 1), z 0) := continuous_apply 0, @@ -225,14 +225,14 @@ def Icc_right_chart (x y : ℝ) [fact (x < y)] : end, continuous_to_fun := begin apply continuous.continuous_on, - apply continuous_subtype_mk, + apply continuous.subtype_mk, have : continuous (λ (z : ℝ) (i : fin 1), y - z) := continuous_const.sub (continuous_pi (λi, continuous_id)), exact this.comp continuous_subtype_val, end, continuous_inv_fun := begin apply continuous.continuous_on, - apply continuous_subtype_mk, + apply continuous.subtype_mk, have A : continuous (λ z : ℝ, max (y - z) x) := (continuous_const.sub continuous_id).max continuous_const, have B : continuous (λz : euclidean_space ℝ (fin 1), z 0) := continuous_apply 0, diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index dd51e5e7a73db..527bff98a2251 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -833,7 +833,7 @@ lemma range_prod_le [module R₁ M₂] [module R₁ M₃] (f : M₁ →L[R₁] M /-- Restrict codomain of a continuous linear map. -/ def cod_restrict (f : M₁ →SL[σ₁₂] M₂) (p : submodule R₂ M₂) (h : ∀ x, f x ∈ p) : M₁ →SL[σ₁₂] p := -{ cont := continuous_subtype_mk h f.continuous, +{ cont := f.continuous.subtype_mk _, to_linear_map := (f : M₁ →ₛₗ[σ₁₂] M₂).cod_restrict p h} @[norm_cast] lemma coe_cod_restrict (f : M₁ →SL[σ₁₂] M₂) (p : submodule R₂ M₂) (h : ∀ x, f x ∈ p) : @@ -1022,15 +1022,15 @@ of `φ` is linearly equivalent to the product over `I`. -/ def infi_ker_proj_equiv {I J : set ι} [decidable_pred (λi, i ∈ I)] (hd : disjoint I J) (hu : set.univ ⊆ I ∪ J) : (⨅i ∈ J, ker (proj i) : submodule R (Πi, φ i)) ≃L[R] (Πi:I, φ i) := -⟨ linear_map.infi_ker_proj_equiv R φ hd hu, - continuous_pi (λ i, begin +{ to_linear_equiv := linear_map.infi_ker_proj_equiv R φ hd hu, + continuous_to_fun := continuous_pi (λ i, begin have := @continuous_subtype_coe _ _ (λ x, x ∈ (⨅i ∈ J, ker (proj i) : submodule R (Πi, φ i))), have := continuous.comp (by exact continuous_apply i) this, exact this end), - continuous_subtype_mk _ (continuous_pi (λ i, begin + continuous_inv_fun := continuous.subtype_mk (continuous_pi (λ i, begin dsimp, split_ifs; [apply continuous_apply, exact continuous_zero] - end)) ⟩ + end)) _ } end pi diff --git a/src/topology/algebra/order/proj_Icc.lean b/src/topology/algebra/order/proj_Icc.lean index 85437ec1d3ed1..a032314bb4a6d 100644 --- a/src/topology/algebra/order/proj_Icc.lean +++ b/src/topology/algebra/order/proj_Icc.lean @@ -28,7 +28,7 @@ variables [topological_space α] [order_topology α] [topological_space β] @[continuity] lemma continuous_proj_Icc : continuous (proj_Icc a b h) := -continuous_subtype_mk _ $ continuous_const.max $ continuous_const.min continuous_id +(continuous_const.max $ continuous_const.min continuous_id).subtype_mk _ lemma quotient_map_proj_Icc : quotient_map (proj_Icc a b h) := quotient_map_iff.2 ⟨proj_Icc_surjective h, λ s, diff --git a/src/topology/algebra/semigroup.lean b/src/topology/algebra/semigroup.lean index 52d583b07b163..8aa53080d36bc 100644 --- a/src/topology/algebra/semigroup.lean +++ b/src/topology/algebra/semigroup.lean @@ -80,8 +80,8 @@ begin mul_assoc := λ p q r, subtype.eq (mul_assoc _ _ _) }, haveI : compact_space M' := is_compact_iff_compact_space.mp s_compact, haveI : nonempty M' := nonempty_subtype.mpr snemp, - have : ∀ p : M', continuous (* p) := λ p, continuous_subtype_mk _ - ((continuous_mul_left p.1).comp continuous_subtype_val), + have : ∀ p : M', continuous (* p) := λ p, + ((continuous_mul_left p.1).comp continuous_subtype_val).subtype_mk _, obtain ⟨⟨m, hm⟩, idem⟩ := exists_idempotent_of_compact_t2_of_continuous_mul_left this, exact ⟨m, hm, subtype.ext_iff.mp idem⟩ end diff --git a/src/topology/bases.lean b/src/topology/bases.lean index 44c214414a832..093b3ee0a425b 100644 --- a/src/topology/bases.lean +++ b/src/topology/bases.lean @@ -491,8 +491,7 @@ end instance separable_space_univ {α : Type*} [topological_space α] [separable_space α] : separable_space (univ : set α) := -(equiv.set.univ α).symm.surjective.dense_range.separable_space - (continuous_subtype_mk _ continuous_id) +(equiv.set.univ α).symm.surjective.dense_range.separable_space (continuous_id.subtype_mk _) /-- If `α` is a separable topological space with a partial order, then there exists a countable dense set `s : set α` that contains those of both bottom and top elements of `α` that actually diff --git a/src/topology/category/CompHaus/default.lean b/src/topology/category/CompHaus/default.lean index 8c58caf89086c..33fb3285809a2 100644 --- a/src/topology/category/CompHaus/default.lean +++ b/src/topology/category/CompHaus/default.lean @@ -216,7 +216,7 @@ begin haveI : t2_space (ulift.{u} $ set.Icc (0:ℝ) 1) := homeomorph.ulift.symm.t2_space, let Z := of (ulift.{u} $ set.Icc (0:ℝ) 1), let g : Y ⟶ Z := ⟨λ y', ⟨⟨φ y', hφ01 y'⟩⟩, - continuous_ulift_up.comp (continuous_subtype_mk (λ y', hφ01 y') φ.continuous)⟩, + continuous_ulift_up.comp (φ.continuous.subtype_mk (λ y', hφ01 y'))⟩, let h : Y ⟶ Z := ⟨λ _, ⟨⟨0, set.left_mem_Icc.mpr zero_le_one⟩⟩, continuous_const⟩, have H : h = g, { rw ← cancel_epi f, diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index bd3f4a4f17f77..5f71b93167519 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -830,12 +830,16 @@ lemma is_closed.closed_embedding_subtype_coe {s : set α} (hs : is_closed s) : inj := subtype.coe_injective, closed_range := (subtype.range_coe : range coe = s).symm ▸ hs } -@[continuity] lemma continuous_subtype_mk {f : β → α} - (hp : ∀x, p (f x)) (h : continuous f) : continuous (λx, (⟨f x, hp x⟩ : subtype p)) := +@[continuity] lemma continuous.subtype_mk {f : β → α} (h : continuous f) + (hp : ∀x, p (f x)) : continuous (λx, (⟨f x, hp x⟩ : subtype p)) := continuous_induced_rng.2 h +lemma continuous.subtype_map {f : α → β} (h : continuous f) {q : β → Prop} + (hpq : ∀ x, p x → q (f x)) : continuous (subtype.map f hpq) := +(h.comp continuous_subtype_coe).subtype_mk _ + lemma continuous_inclusion {s t : set α} (h : s ⊆ t) : continuous (inclusion h) := -continuous_subtype_mk _ continuous_subtype_coe +continuous_id.subtype_map h lemma continuous_at_subtype_coe {p : α → Prop} {a : subtype p} : continuous_at (coe : subtype p → α) a := @@ -912,7 +916,7 @@ lemma continuous_at.restrict_preimage {f : α → β} {s : set β} {x : f ⁻¹' h.restrict _ @[continuity] lemma continuous.cod_restrict {f : α → β} {s : set β} (hf : continuous f) - (hs : ∀ a, f a ∈ s) : continuous (s.cod_restrict f hs) := continuous_subtype_mk hs hf + (hs : ∀ a, f a ∈ s) : continuous (s.cod_restrict f hs) := hf.subtype_mk hs lemma inducing.cod_restrict {e : α → β} (he : inducing e) {s : set β} (hs : ∀ x, e x ∈ s) : inducing (cod_restrict e s hs) := diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index f817af85adeeb..ca045fbf9df68 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -336,7 +336,7 @@ lemma continuous_comp {G : β → γ} {C : ℝ≥0} (H : lipschitz_with C G) : /-- Restriction (in the target) of a bounded continuous function taking values in a subset -/ def cod_restrict (s : set β) (f : α →ᵇ β) (H : ∀x, f x ∈ s) : α →ᵇ s := -⟨⟨s.cod_restrict f H, continuous_subtype_mk _ f.continuous⟩, f.bounded⟩ +⟨⟨s.cod_restrict f H, f.continuous.subtype_mk _⟩, f.bounded⟩ section extend diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 20f2d0e48b701..8531d7282888d 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -166,9 +166,9 @@ protected lemma embedding (h : α ≃ₜ β) : embedding h := /-- Homeomorphism given an embedding. -/ noncomputable def of_embedding (f : α → β) (hf : embedding f) : α ≃ₜ (set.range f) := -{ continuous_to_fun := continuous_subtype_mk _ hf.continuous, +{ continuous_to_fun := hf.continuous.subtype_mk _, continuous_inv_fun := by simp [hf.continuous_iff, continuous_subtype_coe], - .. equiv.of_injective f hf.inj } + to_equiv := equiv.of_injective f hf.inj } protected lemma second_countable_topology [topological_space.second_countable_topology β] (h : α ≃ₜ β) : @@ -322,8 +322,8 @@ end /-- If two sets are equal, then they are homeomorphic. -/ def set_congr {s t : set α} (h : s = t) : s ≃ₜ t := -{ continuous_to_fun := continuous_subtype_mk _ continuous_subtype_val, - continuous_inv_fun := continuous_subtype_mk _ continuous_subtype_val, +{ continuous_to_fun := continuous_inclusion h.subset, + continuous_inv_fun := continuous_inclusion h.symm.subset, to_equiv := equiv.set_congr h } /-- Sum of two homeomorphisms. -/ @@ -446,16 +446,15 @@ A subset of a topological space is homeomorphic to its image under a homeomorphi def set.univ (α : Type*) [topological_space α] : (univ : set α) ≃ₜ α := { to_equiv := equiv.set.univ α, continuous_to_fun := continuous_subtype_coe, - continuous_inv_fun := continuous_subtype_mk _ continuous_id } + continuous_inv_fun := continuous_id.subtype_mk _ } /-- `s ×ˢ t` is homeomorphic to `s × t`. -/ @[simps] def set.prod (s : set α) (t : set β) : ↥(s ×ˢ t) ≃ₜ s × t := { to_equiv := equiv.set.prod s t, - continuous_to_fun := continuous.prod_mk - (continuous_subtype_mk _ (continuous_fst.comp continuous_induced_dom)) - (continuous_subtype_mk _ (continuous_snd.comp continuous_induced_dom)), - continuous_inv_fun := continuous_subtype_mk _ (continuous.prod_mk - (continuous_induced_dom.comp continuous_fst) (continuous_induced_dom.comp continuous_snd)) } + continuous_to_fun := (continuous_subtype_coe.fst.subtype_mk _).prod_mk + (continuous_subtype_coe.snd.subtype_mk _), + continuous_inv_fun := (continuous_subtype_coe.fst'.prod_mk + continuous_subtype_coe.snd').subtype_mk _ } end homeomorph diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index ce2e438aba169..1d6b77812a7a7 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -127,7 +127,7 @@ nnreal.tendsto_coe.2 $ tendsto_to_nnreal ha /-- The set of finite `ℝ≥0∞` numbers is homeomorphic to `ℝ≥0`. -/ def ne_top_homeomorph_nnreal : {a | a ≠ ∞} ≃ₜ ℝ≥0 := { continuous_to_fun := continuous_on_iff_continuous_restrict.1 continuous_on_to_nnreal, - continuous_inv_fun := continuous_subtype_mk _ continuous_coe, + continuous_inv_fun := continuous_coe.subtype_mk _, .. ne_top_equiv_nnreal } /-- The set of finite `ℝ≥0∞` numbers is homeomorphic to `ℝ≥0`. -/ diff --git a/src/topology/instances/ereal.lean b/src/topology/instances/ereal.lean index d470b25d8eb13..ba32823a51d9b 100644 --- a/src/topology/instances/ereal.lean +++ b/src/topology/instances/ereal.lean @@ -129,7 +129,7 @@ lemma continuous_on_to_real : continuous_on ereal.to_real ({⊥, ⊤}ᶜ : set e /-- The set of finite `ereal` numbers is homeomorphic to `ℝ`. -/ def ne_bot_top_homeomorph_real : ({⊥, ⊤}ᶜ : set ereal) ≃ₜ ℝ := { continuous_to_fun := continuous_on_iff_continuous_restrict.1 continuous_on_to_real, - continuous_inv_fun := continuous_subtype_mk _ continuous_coe_real_ereal, + continuous_inv_fun := continuous_coe_real_ereal.subtype_mk _, .. ne_top_bot_equiv_real } diff --git a/src/topology/instances/nnreal.lean b/src/topology/instances/nnreal.lean index ff0cd9279ce0e..be9f755cf5c69 100644 --- a/src/topology/instances/nnreal.lean +++ b/src/topology/instances/nnreal.lean @@ -53,10 +53,10 @@ open_locale nnreal big_operators filter instance : topological_space ℝ≥0 := infer_instance -- short-circuit type class inference instance : topological_semiring ℝ≥0 := -{ continuous_mul := continuous_subtype_mk _ $ - (continuous_subtype_val.comp continuous_fst).mul (continuous_subtype_val.comp continuous_snd), - continuous_add := continuous_subtype_mk _ $ - (continuous_subtype_val.comp continuous_fst).add (continuous_subtype_val.comp continuous_snd) } +{ continuous_mul := + (continuous_subtype_val.fst'.mul continuous_subtype_val.snd').subtype_mk _, + continuous_add := + (continuous_subtype_val.fst'.add continuous_subtype_val.snd').subtype_mk _ } instance : second_countable_topology ℝ≥0 := topological_space.subtype.second_countable_topology _ _ @@ -68,7 +68,7 @@ variable {α : Type*} open filter finset lemma _root_.continuous_real_to_nnreal : continuous real.to_nnreal := -continuous_subtype_mk _ $ continuous_id.max continuous_const +(continuous_id.max continuous_const).subtype_mk _ lemma continuous_coe : continuous (coe : ℝ≥0 → ℝ) := continuous_subtype_val @@ -80,7 +80,7 @@ continuous_subtype_val instance {X : Type*} [topological_space X] : can_lift C(X, ℝ) C(X, ℝ≥0) := { coe := continuous_map.coe_nnreal_real.comp, cond := λ f, ∀ x, 0 ≤ f x, - prf := λ f hf, ⟨⟨λ x, ⟨f x, hf x⟩, continuous_subtype_mk _ f.2⟩, fun_like.ext' rfl⟩ } + prf := λ f hf, ⟨⟨λ x, ⟨f x, hf x⟩, f.2.subtype_mk _⟩, fun_like.ext' rfl⟩ } @[simp, norm_cast] lemma tendsto_coe {f : filter α} {m : α → ℝ≥0} {x : ℝ≥0} : tendsto (λa, (m a : ℝ)) f (𝓝 (x : ℝ)) ↔ tendsto m f (𝓝 x) := @@ -111,17 +111,15 @@ lemma nhds_zero_basis : (𝓝 (0 : ℝ≥0)).has_basis (λ a : ℝ≥0, 0 < a) ( nhds_bot_basis instance : has_continuous_sub ℝ≥0 := -⟨continuous_subtype_mk _ $ - ((continuous_coe.comp continuous_fst).sub - (continuous_coe.comp continuous_snd)).max continuous_const⟩ +⟨((continuous_coe.fst'.sub continuous_coe.snd').max continuous_const).subtype_mk _⟩ instance : has_continuous_inv₀ ℝ≥0 := ⟨λ x hx, tendsto_coe.1 $ (real.tendsto_inv $ nnreal.coe_ne_zero.2 hx).comp continuous_coe.continuous_at⟩ instance : has_continuous_smul ℝ≥0 ℝ := -{ continuous_smul := continuous.comp real.continuous_mul $ continuous.prod_mk - (continuous.comp continuous_subtype_val continuous_fst) continuous_snd } +{ continuous_smul := real.continuous_mul.comp $ + (continuous_subtype_val.comp continuous_fst).prod_mk continuous_snd } @[norm_cast] lemma has_sum_coe {f : α → ℝ≥0} {r : ℝ≥0} : has_sum (λa, (f a : ℝ)) (r : ℝ) ↔ has_sum f r := diff --git a/src/topology/instances/real.lean b/src/topology/instances/real.lean index 1bef704d733ae..226dbffb45a0f 100644 --- a/src/topology/instances/real.lean +++ b/src/topology/instances/real.lean @@ -100,7 +100,7 @@ continuous_iff_continuous_at.mpr $ assume ⟨r, hr⟩, lemma real.continuous.inv [topological_space α] {f : α → ℝ} (h : ∀a, f a ≠ 0) (hf : continuous f) : continuous (λa, (f a)⁻¹) := show continuous ((has_inv.inv ∘ @subtype.val ℝ (λr, r ≠ 0)) ∘ λa, ⟨f a, h a⟩), - from real.continuous_inv.comp (continuous_subtype_mk _ hf) + from real.continuous_inv.comp (hf.subtype_mk _) lemma real.uniform_continuous_const_mul {x : ℝ} : uniform_continuous ((*) x) := uniform_continuous_const_smul x diff --git a/src/topology/local_homeomorph.lean b/src/topology/local_homeomorph.lean index 1ac0bf05a418f..d54e294e7b397 100644 --- a/src/topology/local_homeomorph.lean +++ b/src/topology/local_homeomorph.lean @@ -991,10 +991,10 @@ def to_homeomorph_source_target : e.source ≃ₜ e.target := inv_fun := e.symm_maps_to.restrict _ _ _, left_inv := λ x, subtype.eq $ e.left_inv x.2, right_inv := λ x, subtype.eq $ e.right_inv x.2, - continuous_to_fun := continuous_subtype_mk _ $ - continuous_on_iff_continuous_restrict.1 e.continuous_on, - continuous_inv_fun := continuous_subtype_mk _ $ - continuous_on_iff_continuous_restrict.1 e.symm.continuous_on } + continuous_to_fun := + (continuous_on_iff_continuous_restrict.1 e.continuous_on).subtype_mk _, + continuous_inv_fun := + (continuous_on_iff_continuous_restrict.1 e.symm.continuous_on).subtype_mk _ } lemma second_countable_topology_source [second_countable_topology β] (e : local_homeomorph α β) : diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 12eaf5b40b8ba..6a270b2213e17 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1598,7 +1598,7 @@ lemma tendsto_iff_dist_tendsto_zero {f : β → α} {x : filter β} {a : α} : by rw [← nhds_comap_dist a, tendsto_comap_iff] lemma uniform_continuous_nndist : uniform_continuous (λp:α×α, nndist p.1 p.2) := -uniform_continuous_subtype_mk uniform_continuous_dist _ +uniform_continuous_dist.subtype_mk _ lemma uniform_continuous.nndist [uniform_space β] {f g : β → α} (hf : uniform_continuous f) (hg : uniform_continuous g) : diff --git a/src/topology/metric_space/pi_nat.lean b/src/topology/metric_space/pi_nat.lean index 126e2cd4a3241..31070b61bb700 100644 --- a/src/topology/metric_space/pi_nat.lean +++ b/src/topology/metric_space/pi_nat.lean @@ -669,7 +669,7 @@ begin { assume x, apply subtype.coe_injective.eq_iff.1, simpa only using fs x.val x.property }, - exact ⟨cod_restrict f s A, B, λ x, ⟨x, B x⟩, continuous_subtype_mk _ f_cont⟩, + exact ⟨cod_restrict f s A, B, λ x, ⟨x, B x⟩, f_cont.subtype_mk _⟩, end end pi_nat diff --git a/src/topology/tietze_extension.lean b/src/topology/tietze_extension.lean index cdf65d66ec568..8a5e96357448b 100644 --- a/src/topology/tietze_extension.lean +++ b/src/topology/tietze_extension.lean @@ -348,7 +348,7 @@ begin rcases F.exists_extension_forall_mem_of_closed_embedding hFt (hne.image _) he with ⟨G, hG, hGF⟩, set g : C(Y, ℝ) := ⟨h.symm ∘ cod_restrict G _ (λ y, ht_sub (hG y)), h.symm.continuous.comp $ - continuous_subtype_mk _ G.continuous⟩, + G.continuous.subtype_mk _⟩, have hgG : ∀ {y a}, g y = a ↔ G y = h a, from λ y a, h.to_equiv.symm_apply_eq.trans subtype.ext_iff, refine ⟨g, λ y, _, _⟩, diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index 5aad3035aa56c..c01f3f0b6213d 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -1281,7 +1281,7 @@ lemma uniform_continuous_subtype_coe {p : α → Prop} [uniform_space α] : uniform_continuous (coe : {a : α // p a} → α) := uniform_continuous_subtype_val -lemma uniform_continuous_subtype_mk {p : α → Prop} [uniform_space α] [uniform_space β] +lemma uniform_continuous.subtype_mk {p : α → Prop} [uniform_space α] [uniform_space β] {f : β → α} (hf : uniform_continuous f) (h : ∀x, p (f x)) : uniform_continuous (λx, ⟨f x, h x⟩ : β → subtype p) := uniform_continuous_comap' hf diff --git a/src/topology/uniform_space/equiv.lean b/src/topology/uniform_space/equiv.lean index 037462f508c2a..073dad3b71306 100644 --- a/src/topology/uniform_space/equiv.lean +++ b/src/topology/uniform_space/equiv.lean @@ -165,16 +165,15 @@ protected lemma uniform_embedding (h : α ≃ᵤ β) : uniform_embedding h := /-- Uniform equiv given a uniform embedding. -/ noncomputable def of_uniform_embedding (f : α → β) (hf : uniform_embedding f) : α ≃ᵤ (set.range f) := -{ uniform_continuous_to_fun := uniform_continuous_subtype_mk - hf.to_uniform_inducing.uniform_continuous _, +{ uniform_continuous_to_fun := hf.to_uniform_inducing.uniform_continuous.subtype_mk _, uniform_continuous_inv_fun := by simp [hf.to_uniform_inducing.uniform_continuous_iff, uniform_continuous_subtype_coe], to_equiv := equiv.of_injective f hf.inj } /-- If two sets are equal, then they are uniformly equivalent. -/ def set_congr {s t : set α} (h : s = t) : s ≃ᵤ t := -{ uniform_continuous_to_fun := uniform_continuous_subtype_mk uniform_continuous_subtype_val _, - uniform_continuous_inv_fun := uniform_continuous_subtype_mk uniform_continuous_subtype_val _, +{ uniform_continuous_to_fun := uniform_continuous_subtype_val.subtype_mk _, + uniform_continuous_inv_fun := uniform_continuous_subtype_val.subtype_mk _, to_equiv := equiv.set_congr h } /-- Product of two uniform isomorphisms. -/ @@ -251,11 +250,10 @@ def {u} pi_fin_two (α : fin 2 → Type u) [Π i, uniform_space (α i)] : (Π i, A subset of a uniform space is uniformly isomorphic to its image under a uniform isomorphism. -/ def image (e : α ≃ᵤ β) (s : set α) : s ≃ᵤ e '' s := -{ uniform_continuous_to_fun := uniform_continuous_subtype_mk - (e.uniform_continuous.comp uniform_continuous_subtype_val) (λ x, mem_image_of_mem _ x.2), - uniform_continuous_inv_fun := uniform_continuous_subtype_mk - (e.symm.uniform_continuous.comp uniform_continuous_subtype_val) - (λ x, by simpa using mem_image_of_mem e.symm x.2), +{ uniform_continuous_to_fun := + (e.uniform_continuous.comp uniform_continuous_subtype_val).subtype_mk _, + uniform_continuous_inv_fun := + (e.symm.uniform_continuous.comp uniform_continuous_subtype_val).subtype_mk _, to_equiv := e.to_equiv.image s } end uniform_equiv From 63f2525e803321aa8f653e22268516a376cea6c1 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 24 Aug 2022 12:36:14 +0000 Subject: [PATCH 021/828] chore(data/finset/pi_induction): fix lint (#16224) --- src/data/finset/pi_induction.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/data/finset/pi_induction.lean b/src/data/finset/pi_induction.lean index 23fd8ef8cb3bb..be3fa94749c91 100644 --- a/src/data/finset/pi_induction.lean +++ b/src/data/finset/pi_induction.lean @@ -11,7 +11,7 @@ import data.fintype.basic In this file we prove a few induction principles for functions `Π i : ι, finset (α i)` defined on a finite type. -* `finset.induction_on_pi` is a generic lemma that requires only `[fintype ι]`, `[decidable_eq ι]`, +* `finset.induction_on_pi` is a generic lemma that requires only `[finite ι]`, `[decidable_eq ι]`, and `[Π i, decidable_eq (α i)]`; this version can be seen as a direct generalization of `finset.induction_on`. @@ -25,7 +25,7 @@ finite set, finite type, induction, function open function -variables {ι : Type*} {α : ι → Type*} [fintype ι] [decidable_eq ι] [Π i, decidable_eq (α i)] +variables {ι : Type*} {α : ι → Type*} [finite ι] [decidable_eq ι] [Π i, decidable_eq (α i)] namespace finset @@ -37,6 +37,7 @@ lemma induction_on_pi_of_choice (r : Π i, α i → finset (α i) → Prop) r i x (g i) → p g → p (update g i (insert x (g i)))) : p f := begin + casesI nonempty_fintype ι, induction hs : univ.sigma f using finset.strong_induction_on with s ihs generalizing f, subst s, cases eq_empty_or_nonempty (univ.sigma f) with he hne, { convert h0, simpa [funext_iff] using he }, From 730cde3423dfba8c0e5f06d066a49939daa6bf89 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 24 Aug 2022 15:46:18 +0000 Subject: [PATCH 022/828] =?UTF-8?q?feat(algebra/order/monoid):=20a=20canon?= =?UTF-8?q?ically=20ordered=20add=20monoid=20has=20`0=20=E2=89=A4=201`=20(?= =?UTF-8?q?#16199)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/algebra/order/monoid.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/algebra/order/monoid.lean b/src/algebra/order/monoid.lean index 200f5d691d11b..f80005deee113 100644 --- a/src/algebra/order/monoid.lean +++ b/src/algebra/order/monoid.lean @@ -483,6 +483,10 @@ end canonically_ordered_monoid lemma pos_of_gt {M : Type*} [canonically_ordered_add_monoid M] {n m : M} (h : n < m) : 0 < m := lt_of_le_of_lt (zero_le _) h +@[priority 100] instance canonically_ordered_add_monoid.zero_le_one_class {M : Type*} + [canonically_ordered_add_monoid M] [has_one M] : zero_le_one_class M := +⟨zero_le 1⟩ + /-- A canonically linear-ordered additive monoid is a canonically ordered additive monoid whose ordering is a linear order. -/ @[protect_proj, ancestor canonically_ordered_add_monoid linear_order] From 40427f7889b7751e88e5a8aaa469d9566d070e79 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 24 Aug 2022 15:46:19 +0000 Subject: [PATCH 023/828] feat(analysis/special_functions/trigonometric/angle): continuity and signs (#16204) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add the lemmas that `real.angle.sign` is continuous away from 0 and π, and thus that any function to angles that is continuous on a connected set and does not take the value 0 or π on that set produces angles with constant sign on that set (this is a general principle for use in proving results about when two oriented angles in a geometrical configuration must have the same sign). --- .../trigonometric/angle.lean | 22 ++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/src/analysis/special_functions/trigonometric/angle.lean b/src/analysis/special_functions/trigonometric/angle.lean index e994c599b1955..67809c5979110 100644 --- a/src/analysis/special_functions/trigonometric/angle.lean +++ b/src/analysis/special_functions/trigonometric/angle.lean @@ -6,7 +6,7 @@ Authors: Calle Sönne import analysis.special_functions.trigonometric.basic import algebra.char_zero.quotient import algebra.order.to_interval_mod -import data.sign +import topology.instances.sign /-! # The type of angles @@ -527,6 +527,26 @@ lemma eq_iff_abs_to_real_eq_of_sign_eq {θ ψ : angle} (h : θ.sign = ψ.sign) : θ = ψ ↔ |θ.to_real| = |ψ.to_real| := by simpa [h] using @eq_iff_sign_eq_and_abs_to_real_eq θ ψ +lemma continuous_at_sign {θ : angle} (h0 : θ ≠ 0) (hpi : θ ≠ π) : continuous_at sign θ := +(continuous_at_sign_of_ne_zero (sin_ne_zero_iff.2 ⟨h0, hpi⟩)).comp continuous_sin.continuous_at + +lemma _root_.continuous_on.angle_sign_comp {α : Type*} [topological_space α] {f : α → angle} + {s : set α} (hf : continuous_on f s) (hs : ∀ z ∈ s, f z ≠ 0 ∧ f z ≠ π) : + continuous_on (sign ∘ f) s := +begin + refine (continuous_at.continuous_on (λ θ hθ, _)).comp hf (set.maps_to_image f s), + obtain ⟨z, hz, rfl⟩ := hθ, + exact continuous_at_sign (hs _ hz).1 (hs _ hz).2 +end + +/-- Suppose a function to angles is continuous on a connected set and never takes the values `0` +or `π` on that set. Then the values of the function on that set all have the same sign. -/ +lemma sign_eq_of_continuous_on {α : Type*} [topological_space α] {f : α → angle} {s : set α} + {x y : α} (hc : is_connected s) (hf : continuous_on f s) (hs : ∀ z ∈ s, f z ≠ 0 ∧ f z ≠ π) + (hx : x ∈ s) (hy : y ∈ s) : (f y).sign = (f x).sign := +(hc.image _ (hf.angle_sign_comp hs)).is_preconnected.subsingleton + (set.mem_image_of_mem _ hy) (set.mem_image_of_mem _ hx) + end angle end real From 4e64a3fae82a793473236183a590255cc8a03448 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 24 Aug 2022 18:34:10 +0000 Subject: [PATCH 024/828] feat(algebra/category): forgetful functors from modules/abelian groups preserve epis/monos (#15108) The corresponding `reflects` statements already follow from faithfulness. --- src/algebra/category/Group/epi_mono.lean | 21 +++++++++++++++++-- src/algebra/category/Module/epi_mono.lean | 10 ++++++++- .../concrete_category/basic.lean | 20 ++++++++++++++++-- src/category_theory/functor/epi_mono.lean | 16 ++++++++++++++ 4 files changed, 62 insertions(+), 5 deletions(-) diff --git a/src/algebra/category/Group/epi_mono.lean b/src/algebra/category/Group/epi_mono.lean index ca6fc836c22f5..91929aef51738 100644 --- a/src/algebra/category/Group/epi_mono.lean +++ b/src/algebra/category/Group/epi_mono.lean @@ -321,6 +321,19 @@ iff.trans (epi_iff_surjective _) (add_subgroup.eq_top_iff' f.range).symm end AddGroup +namespace Group +variables {A B : Group.{u}} (f : A ⟶ B) + +@[to_additive] +instance forget_Group_preserves_mono : (forget Group).preserves_monomorphisms := +{ preserves := λ X Y f e, by rwa [mono_iff_injective, ←category_theory.mono_iff_injective] at e } + +@[to_additive] +instance forget_Group_preserves_epi : (forget Group).preserves_epimorphisms := +{ preserves := λ X Y f e, by rwa [epi_iff_surjective, ←category_theory.epi_iff_surjective] at e } + +end Group + namespace CommGroup variables {A B : CommGroup.{u}} (f : A ⟶ B) @@ -353,8 +366,12 @@ lemma epi_iff_surjective : epi f ↔ function.surjective f := by rw [epi_iff_range_eq_top, monoid_hom.range_top_iff_surjective] @[to_additive] -instance : functor.preserves_epimorphisms (forget₂ CommGroup Group) := -{ preserves := λ X Y f e, by rwa [epi_iff_surjective, ←@Group.epi_iff_surjective ⟨X⟩ ⟨Y⟩ f] at e } +instance forget_CommGroup_preserves_mono : (forget CommGroup).preserves_monomorphisms := +{ preserves := λ X Y f e, by rwa [mono_iff_injective, ←category_theory.mono_iff_injective] at e } + +@[to_additive] +instance forget_CommGroup_preserves_epi : (forget CommGroup).preserves_epimorphisms := +{ preserves := λ X Y f e, by rwa [epi_iff_surjective, ←category_theory.epi_iff_surjective] at e } end CommGroup diff --git a/src/algebra/category/Module/epi_mono.lean b/src/algebra/category/Module/epi_mono.lean index d9642ac984022..5664600bd5530 100644 --- a/src/algebra/category/Module/epi_mono.lean +++ b/src/algebra/category/Module/epi_mono.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import linear_algebra.quotient -import category_theory.epi_mono +import algebra.category.Group.epi_mono import algebra.category.Module.basic /-! @@ -55,4 +55,12 @@ instance mono_as_hom'_subtype (U : submodule R X) : mono ↾U.subtype := instance epi_as_hom''_mkq (U : submodule R X) : epi ↿U.mkq := (epi_iff_range_eq_top _).mpr $ submodule.range_mkq _ +instance forget_preserves_epimorphisms : (forget (Module.{v} R)).preserves_epimorphisms := +{ preserves := λ X Y f hf, by rwa [forget_map_eq_coe, category_theory.epi_iff_surjective, + ← epi_iff_surjective] } + +instance forget_preserves_monomorphisms : (forget (Module.{v} R)).preserves_monomorphisms := +{ preserves := λ X Y f hf, by rwa [forget_map_eq_coe, category_theory.mono_iff_injective, + ← mono_iff_injective] } + end Module diff --git a/src/category_theory/concrete_category/basic.lean b/src/category_theory/concrete_category/basic.lean index 29e9eb2079c0b..d9fc670e07014 100644 --- a/src/category_theory/concrete_category/basic.lean +++ b/src/category_theory/concrete_category/basic.lean @@ -155,10 +155,24 @@ class has_forget₂ (C : Type v) (D : Type v') [category C] [concrete_category.{ [concrete_category D] [has_forget₂ C D] : C ⥤ D := has_forget₂.forget₂ -instance forget_faithful (C : Type v) (D : Type v') [category C] [concrete_category C] [category D] +instance forget₂_faithful (C : Type v) (D : Type v') [category C] [concrete_category C] [category D] [concrete_category D] [has_forget₂ C D] : faithful (forget₂ C D) := has_forget₂.forget_comp.faithful_of_comp +instance forget₂_preserves_monomorphisms (C : Type v) (D : Type v') [category C] + [concrete_category C] [category D] [concrete_category D] [has_forget₂ C D] + [(forget C).preserves_monomorphisms] : (forget₂ C D).preserves_monomorphisms := +have (forget₂ C D ⋙ forget D).preserves_monomorphisms, + by { simp only [has_forget₂.forget_comp], apply_instance }, +by exactI functor.preserves_monomorphisms_of_preserves_of_reflects _ (forget D) + +instance forget₂_preserves_epimorphisms (C : Type v) (D : Type v') [category C] + [concrete_category C] [category D] [concrete_category D] [has_forget₂ C D] + [(forget C).preserves_epimorphisms] : (forget₂ C D).preserves_epimorphisms := +have (forget₂ C D ⋙ forget D).preserves_epimorphisms, + by { simp only [has_forget₂.forget_comp], apply_instance }, +by exactI functor.preserves_epimorphisms_of_preserves_of_reflects _ (forget D) + instance induced_category.concrete_category {C : Type v} {D : Type v'} [category D] [concrete_category D] (f : C → D) : concrete_category (induced_category D f) := @@ -191,7 +205,9 @@ has_forget₂ C D := { forget₂ := faithful.div _ _ _ @h_obj _ @h_map, forget_comp := by apply faithful.div_comp } -instance has_forget_to_Type (C : Type v) [category C] [concrete_category C] : +/-- Every forgetful functor factors through the identity functor. This is not a global instance as + it is prone to creating type class resolution loops. -/ +def has_forget_to_Type (C : Type v) [category C] [concrete_category C] : has_forget₂ C (Type u) := { forget₂ := forget C, forget_comp := functor.comp_id _ } diff --git a/src/category_theory/functor/epi_mono.lean b/src/category_theory/functor/epi_mono.lean index cb3bff689fb2a..54039ba411dde 100644 --- a/src/category_theory/functor/epi_mono.lean +++ b/src/category_theory/functor/epi_mono.lean @@ -70,6 +70,22 @@ instance reflects_epimorphisms_comp (F : C ⥤ D) (G : D ⥤ E) [reflects_epimor [reflects_epimorphisms G] : reflects_epimorphisms (F ⋙ G) := { reflects := λ X Y f h, (F.epi_of_epi_map (G.epi_of_epi_map h)) } +lemma preserves_epimorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) + [preserves_epimorphisms (F ⋙ G)] [reflects_epimorphisms G] : preserves_epimorphisms F := +⟨λ X Y f hf, G.epi_of_epi_map $ show epi ((F ⋙ G).map f), by exactI infer_instance⟩ + +lemma preserves_monomorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) + [preserves_monomorphisms (F ⋙ G)] [reflects_monomorphisms G] : preserves_monomorphisms F := +⟨λ X Y f hf, G.mono_of_mono_map $ show mono ((F ⋙ G).map f), by exactI infer_instance⟩ + +lemma reflects_epimorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) + [preserves_epimorphisms G] [reflects_epimorphisms (F ⋙ G)] : reflects_epimorphisms F := +⟨λ X Y f hf, (F ⋙ G).epi_of_epi_map $ show epi (G.map (F.map f)), by exactI infer_instance⟩ + +lemma reflects_monomorphisms_of_preserves_of_reflects (F : C ⥤ D) (G : D ⥤ E) + [preserves_monomorphisms G] [reflects_monomorphisms (F ⋙ G)] : reflects_monomorphisms F := +⟨λ X Y f hf, (F ⋙ G).mono_of_mono_map $ show mono (G.map (F.map f)), by exactI infer_instance⟩ + lemma preserves_monomorphisms.of_iso {F G : C ⥤ D} [preserves_monomorphisms F] (α : F ≅ G) : preserves_monomorphisms G := { preserves := λ X Y f h, From 96ed4226b1ba5a589af198b93cb352f034fb9cc8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 24 Aug 2022 18:34:11 +0000 Subject: [PATCH 025/828] feat(order/heyting_algebra): Heyting algebras (#15305) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define (generalized) Heyting, co-Heyting and bi-Heyting algebras. No lemma has been renamed, even though this might look like so from the diff. `himp` is now a field of `boolean_algebra`, with default value `λ x y, y ⊔ xᶜ`. ## Duplicated lemmas Those lemmas have been duplicated because they are true for Heyting algebras but are also used to prove that Boolean algebras are Heyting algebras: * `sdiff_le'` * `inf_compl_eq_bot'` ## Changes to argument implicitness * `sdiff_inf_self_left` * `sdiff_inf_self_right` --- .../box_integral/partition/filter.lean | 2 +- src/combinatorics/set_family/kleitman.lean | 2 +- src/data/finset/basic.lean | 7 +- src/data/set/basic.lean | 5 +- src/data/set/function.lean | 2 +- src/measure_theory/measure/ae_disjoint.lean | 6 +- src/order/boolean_algebra.lean | 287 +++---- src/order/heyting/basic.lean | 725 ++++++++++++++++++ 8 files changed, 826 insertions(+), 210 deletions(-) create mode 100644 src/order/heyting/basic.lean diff --git a/src/analysis/box_integral/partition/filter.lean b/src/analysis/box_integral/partition/filter.lean index b73ff899457a2..edd2fbd89e18d 100644 --- a/src/analysis/box_integral/partition/filter.lean +++ b/src/analysis/box_integral/partition/filter.lean @@ -349,7 +349,7 @@ begin rcases hπ.4 hD with ⟨π₁, hπ₁U, hc⟩, set π₂ := π.filter (λ J, ¬p J), have : disjoint π₁.Union π₂.Union, - by simpa [π₂, hπ₁U] using (disjoint_diff.mono_left sdiff_le).symm, + by simpa [π₂, hπ₁U] using disjoint_sdiff_self_left.mono_right sdiff_le, refine ⟨π₁.disj_union π₂.to_prepartition this, _, _⟩, { suffices : ↑I \ π.Union ∪ π.Union \ (π.filter p).Union = ↑I \ (π.filter p).Union, by simpa *, have : (π.filter p).Union ⊆ π.Union, from bUnion_subset_bUnion_left (finset.filter_subset _ _), diff --git a/src/combinatorics/set_family/kleitman.lean b/src/combinatorics/set_family/kleitman.lean index daefabc28c79d..6949daf8fde3e 100644 --- a/src/combinatorics/set_family/kleitman.lean +++ b/src/combinatorics/set_family/kleitman.lean @@ -55,7 +55,7 @@ begin refine (card_le_of_subset $ bUnion_mono $ λ j hj, (hf₁ _ hj).1).trans _, nth_rewrite 0 cons_eq_insert i, rw bUnion_insert, - refine (card_mono $ @le_sup_sdiff _ (f' i) _ _).trans ((card_union_le _ _).trans _), + refine (card_mono $ @le_sup_sdiff _ _ _ $ f' i).trans ((card_union_le _ _).trans _), rw [union_sdiff_left, sdiff_eq_inter_compl], refine le_of_mul_le_mul_left _ (pow_pos zero_lt_two $ card α + 1), rw [pow_succ', mul_add, mul_assoc, mul_comm _ 2, mul_assoc], diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index ae47f6cdf0f1e..a28db8d8bfc78 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -1233,8 +1233,11 @@ lemma inter_sdiff (s t u : finset α) : s ∩ (t \ u) = s ∩ t \ u := by { ext lemma sdiff_inter_distrib_right (s t u : finset α) : s \ (t ∩ u) = (s \ t) ∪ (s \ u) := sdiff_inf -@[simp] lemma sdiff_inter_self_left (s t : finset α) : s \ (s ∩ t) = s \ t := sdiff_inf_self_left -@[simp] lemma sdiff_inter_self_right (s t : finset α) : s \ (t ∩ s) = s \ t := sdiff_inf_self_right +@[simp] lemma sdiff_inter_self_left (s t : finset α) : s \ (s ∩ t) = s \ t := +sdiff_inf_self_left _ _ + +@[simp] lemma sdiff_inter_self_right (s t : finset α) : s \ (t ∩ s) = s \ t := +sdiff_inf_self_right _ _ @[simp] lemma sdiff_empty : s \ ∅ = s := sdiff_bot diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 68a16d6150f63..d9a60cb8af577 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1150,10 +1150,9 @@ sup_sdiff_self_left inf_sdiff_self_left @[simp] theorem diff_inter_self_eq_diff {s t : set α} : s \ (t ∩ s) = s \ t := -sdiff_inf_self_right +sdiff_inf_self_right _ _ -@[simp] theorem diff_self_inter {s t : set α} : s \ (s ∩ t) = s \ t := -sdiff_inf_self_left +@[simp] theorem diff_self_inter {s t : set α} : s \ (s ∩ t) = s \ t := sdiff_inf_self_left _ _ @[simp] theorem diff_eq_self {s t : set α} : s \ t = s ↔ t ∩ s ⊆ ∅ := show s \ t = s ↔ t ⊓ s ≤ ⊥, from sdiff_eq_self_iff_disjoint diff --git a/src/data/set/function.lean b/src/data/set/function.lean index 3751674f66ac4..b88e3aaada131 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -1036,7 +1036,7 @@ end lemma injective_piecewise_iff {f g : α → β} : injective (s.piecewise f g) ↔ inj_on f s ∧ inj_on g sᶜ ∧ (∀ (x ∈ s) (y ∉ s), f x ≠ g y) := begin - rw [injective_iff_inj_on_univ, ← union_compl_self s, inj_on_union (@disjoint_compl_right _ s _), + rw [injective_iff_inj_on_univ, ← union_compl_self s, inj_on_union (@disjoint_compl_right _ _ s), (piecewise_eq_on s f g).inj_on_iff, (piecewise_eq_on_compl s f g).inj_on_iff], refine and_congr iff.rfl (and_congr iff.rfl $ forall₄_congr $ λ x hx y hy, _), rw [piecewise_eq_of_mem s f g hx, piecewise_eq_of_not_mem s f g hy] diff --git a/src/measure_theory/measure/ae_disjoint.lean b/src/measure_theory/measure/ae_disjoint.lean index 53efebd9be758..8a806614f013b 100644 --- a/src/measure_theory/measure/ae_disjoint.lean +++ b/src/measure_theory/measure/ae_disjoint.lean @@ -28,7 +28,7 @@ variables {μ} {s t u v : set α} family of measurable null sets `t i` such that `s i \ t i` are pairwise disjoint. -/ lemma exists_null_pairwise_disjoint_diff [encodable ι] {s : ι → set α} (hd : pairwise (ae_disjoint μ on s)) : - ∃ t : ι → set α, (∀ i, measurable_set (t i)) ∧ (∀ i, μ (t i) = 0) ∧ + ∃ t : ι → set α, (∀ i, measurable_set (t i)) ∧ (∀ i, μ (t i) = 0) ∧ pairwise (disjoint on (λ i, s i \ t i)) := begin refine ⟨λ i, to_measurable μ (s i ∩ ⋃ j ∈ ({i}ᶜ : set ι), s j), @@ -115,7 +115,7 @@ lemma of_null_left (h : μ s = 0) : ae_disjoint μ s t := (of_null_right h).symm end ae_disjoint -lemma ae_disjoint_compl_left : ae_disjoint μ sᶜ s := (@disjoint_compl_left _ s _).ae_disjoint -lemma ae_disjoint_compl_right : ae_disjoint μ s sᶜ := (@disjoint_compl_right _ s _).ae_disjoint +lemma ae_disjoint_compl_left : ae_disjoint μ sᶜ s := (@disjoint_compl_left _ _ s).ae_disjoint +lemma ae_disjoint_compl_right : ae_disjoint μ s sᶜ := (@disjoint_compl_right _ _ s).ae_disjoint end measure_theory diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index 587fe912136f4..265df62ad5559 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -3,7 +3,8 @@ 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, Bryan Gin-ge Chen -/ -import order.bounded_order +import order.heyting.basic + /-! # (Generalized) Boolean algebras @@ -69,8 +70,6 @@ Some of the lemmas in this section are from: -/ -export has_sdiff (sdiff) - /-- A generalized Boolean algebra is a distributive lattice with `⊥` and a relative complement operation `\` (called `sdiff`, after "set difference") satisfying `(a ⊓ b) ⊔ (a \ b) = a` and `(a ⊓ b) ⊓ (a \ b) = ⊥`, i.e. `a \ b` is the complement of `b` in `a`. @@ -114,19 +113,13 @@ begin exact (eq_of_inf_eq_sup_eq i s).symm, end -lemma sdiff_le : x \ y ≤ x := +lemma sdiff_le' : x \ y ≤ x := calc x \ y ≤ (x ⊓ y) ⊔ (x \ y) : le_sup_right ... = x : sup_inf_sdiff x y -@[simp] lemma bot_sdiff : ⊥ \ x = ⊥ := le_bot_iff.1 sdiff_le - -lemma inf_sdiff_right : x ⊓ (x \ y) = x \ y := by rw [inf_of_le_right (@sdiff_le _ x y _)] +lemma inf_sdiff_right : x ⊓ (x \ y) = x \ y := inf_of_le_right (@sdiff_le' _ x y _) lemma inf_sdiff_left : (x \ y) ⊓ x = x \ y := by rw [inf_comm, inf_sdiff_right] --- cf. `is_compl_top_bot` -@[simp] lemma sdiff_self : x \ x = ⊥ := -by rw [←inf_inf_sdiff, inf_idem, inf_of_le_right (@sdiff_le _ x x _)] - @[simp] theorem sup_sdiff_self_right : x ⊔ (y \ x) = x ⊔ y := calc x ⊔ (y \ x) = (x ⊔ (x ⊓ y)) ⊔ (y \ x) : by rw sup_inf_self ... = x ⊔ ((y ⊓ x) ⊔ (y \ x)) : by ac_refl @@ -149,8 +142,8 @@ lemma sup_le_of_le_sdiff_left (h : y ≤ z \ x) (hxz : x ≤ z) : x ⊔ y ≤ z lemma sup_le_of_le_sdiff_right (h : x ≤ z \ y) (hyz : y ≤ z) : x ⊔ y ≤ z := (sup_le_sup_right h y).trans (sdiff_sup_cancel hyz).le -@[simp] lemma sup_sdiff_left : x ⊔ (x \ y) = x := by { rw sup_eq_left, exact sdiff_le } -lemma sup_sdiff_right : (x \ y) ⊔ x = x := by rw [sup_comm, sup_sdiff_left] +@[simp] lemma sup_sdiff_left : x ⊔ (x \ y) = x := sup_eq_left.2 sdiff_le' +lemma sup_sdiff_right : (x \ y) ⊔ x = x := sup_eq_right.2 sdiff_le' @[simp] lemma sdiff_inf_sdiff : x \ y ⊓ (y \ x) = ⊥ := eq.symm $ @@ -166,24 +159,37 @@ eq.symm $ lemma disjoint_sdiff_sdiff : disjoint (x \ y) (y \ x) := sdiff_inf_sdiff.le -theorem le_sup_sdiff : y ≤ x ⊔ (y \ x) := -by { rw [sup_sdiff_self_right], exact le_sup_right } - -theorem le_sdiff_sup : y ≤ (y \ x) ⊔ x := -by { rw [sup_comm], exact le_sup_sdiff } - @[simp] theorem inf_sdiff_self_right : x ⊓ (y \ x) = ⊥ := calc x ⊓ (y \ x) = ((x ⊓ y) ⊔ (x \ y)) ⊓ (y \ x) : by rw sup_inf_sdiff ... = (x ⊓ y) ⊓ (y \ x) ⊔ (x \ y) ⊓ (y \ x) : by rw inf_sup_right ... = ⊥ : by rw [@inf_comm _ _ x y, inf_inf_sdiff, sdiff_inf_sdiff, bot_sup_eq] @[simp] theorem inf_sdiff_self_left : (y \ x) ⊓ x = ⊥ := by rw [inf_comm, inf_sdiff_self_right] +@[priority 100] -- see Note [lower instance priority] +instance generalized_boolean_algebra.to_generalized_coheyting_algebra : + generalized_coheyting_algebra α := +{ sdiff := (\), + sdiff_le_iff := λ y x z, ⟨λ h, le_of_inf_le_sup_le + (le_of_eq + (calc y ⊓ (y \ x) = y \ x : inf_sdiff_right + ... = (x ⊓ (y \ x)) ⊔ (z ⊓ (y \ x)) + : by rw [inf_eq_right.2 h, inf_sdiff_self_right, bot_sup_eq] + ... = (x ⊔ z) ⊓ (y \ x) : inf_sup_right.symm)) + (calc y ⊔ y \ x = y : sup_sdiff_left + ... ≤ y ⊔ (x ⊔ z) : le_sup_left + ... = ((y \ x) ⊔ x) ⊔ z : by rw [←sup_assoc, ←@sup_sdiff_self_left _ x y] + ... = x ⊔ z ⊔ y \ x : by ac_refl), + λ h, le_of_inf_le_sup_le + (calc y \ x ⊓ x = ⊥ : inf_sdiff_self_left + ... ≤ z ⊓ x : bot_le) + (calc y \ x ⊔ x = y ⊔ x : sup_sdiff_self_left + ... ≤ (x ⊔ z) ⊔ x : sup_le_sup_right h x + ... ≤ z ⊔ x : by rw [sup_assoc, sup_comm, sup_assoc, sup_idem])⟩, + ..‹generalized_boolean_algebra α›, ..generalized_boolean_algebra.to_order_bot } + theorem disjoint_sdiff_self_left : disjoint (y \ x) x := inf_sdiff_self_left.le theorem disjoint_sdiff_self_right : disjoint x (y \ x) := inf_sdiff_self_right.le -lemma disjoint.disjoint_sdiff_left (h : disjoint x y) : disjoint (x \ z) y := h.mono_left sdiff_le -lemma disjoint.disjoint_sdiff_right (h : disjoint x y) : disjoint x (y \ z) := h.mono_right sdiff_le - /- TODO: we could make an alternative constructor for `generalized_boolean_algebra` using `disjoint x (y \ x)` and `x ⊔ (y \ x) = y` as axioms. -/ theorem disjoint.sdiff_eq_of_sup_eq (hi : disjoint x z) (hs : x ⊔ z = y) : y \ x = z := @@ -268,41 +274,16 @@ sdiff_unique ... = ⊥ : by rw [inf_inf_sdiff, bot_inf_eq, bot_sup_eq, @inf_comm _ _ (y \ z), inf_inf_sdiff, inf_bot_eq]) --- cf. `is_compl.inf_sup` -lemma sdiff_inf : y \ (x ⊓ z) = y \ x ⊔ y \ z := -sdiff_unique - (calc y ⊓ (x ⊓ z) ⊔ (y \ x ⊔ y \ z) = - (z ⊓ (y ⊓ x)) ⊔ (y \ x ⊔ y \ z) : by ac_refl - ... = (z ⊔ (y \ x ⊔ y \ z)) ⊓ ((y ⊓ x) ⊔ (y \ x ⊔ y \ z)) : by rw sup_inf_right - ... = (y \ x ⊔ (y \ z ⊔ z)) ⊓ (y ⊓ x ⊔ (y \ x ⊔ y \ z)) : by ac_refl - ... = (y ⊔ z) ⊓ ((y ⊓ x) ⊔ (y \ x ⊔ y \ z)) : - by rw [sup_sdiff_self_left, ←sup_assoc, sup_sdiff_right] - ... = (y ⊔ z) ⊓ y : by rw [←sup_assoc, sup_inf_sdiff, sup_sdiff_left] - ... = y : by rw [inf_comm, inf_sup_self]) - (calc y ⊓ (x ⊓ z) ⊓ ((y \ x) ⊔ (y \ z)) = - (y ⊓ (x ⊓ z) ⊓ (y \ x)) ⊔ (y ⊓ (x ⊓ z) ⊓ (y \ z)) : by rw inf_sup_left - ... = z ⊓ (y ⊓ x ⊓ (y \ x)) ⊔ z ⊓ (y ⊓ x) ⊓ (y \ z) : by ac_refl - ... = z ⊓ (y ⊓ x) ⊓ (y \ z) : by rw [inf_inf_sdiff, inf_bot_eq, bot_sup_eq] - ... = x ⊓ ((y ⊓ z) ⊓ (y \ z)) : by ac_refl - ... = ⊥ : by rw [inf_inf_sdiff, inf_bot_eq]) - -@[simp] lemma sdiff_inf_self_right : y \ (x ⊓ y) = y \ x := -by rw [sdiff_inf, sdiff_self, sup_bot_eq] -@[simp] lemma sdiff_inf_self_left : y \ (y ⊓ x) = y \ x := by rw [inf_comm, sdiff_inf_self_right] - lemma sdiff_eq_sdiff_iff_inf_eq_inf : y \ x = y \ z ↔ y ⊓ x = y ⊓ z := ⟨λ h, eq_of_inf_eq_sup_eq (by rw [inf_inf_sdiff, h, inf_inf_sdiff]) (by rw [sup_inf_sdiff, h, sup_inf_sdiff]), - λ h, by rw [←sdiff_inf_self_right, ←@sdiff_inf_self_right _ z y, inf_comm, h, inf_comm]⟩ + λ h, by rw [←sdiff_inf_self_right, ←sdiff_inf_self_right z y, inf_comm, h, inf_comm]⟩ theorem disjoint.sdiff_eq_left (h : disjoint x y) : x \ y = x := by conv_rhs { rw [←sup_inf_sdiff x y, h.eq_bot, bot_sup_eq] } theorem disjoint.sdiff_eq_right (h : disjoint x y) : y \ x = y := h.symm.sdiff_eq_left --- cf. `is_compl_bot_top` -@[simp] theorem sdiff_bot : x \ ⊥ = x := disjoint_bot_right.sdiff_eq_left - theorem sdiff_eq_self_iff_disjoint : x \ y = x ↔ disjoint y x := calc x \ y = x ↔ x \ y = x \ ⊥ : by rw sdiff_bot ... ↔ x ⊓ y = x ⊓ ⊥ : sdiff_eq_sdiff_iff_inf_eq_inf @@ -319,34 +300,6 @@ begin rw [←h, inf_eq_right.mpr hx], end --- cf. `is_compl.antitone` -lemma sdiff_le_sdiff_left (h : z ≤ x) : w \ x ≤ w \ z := -le_of_inf_le_sup_le - (calc (w \ x) ⊓ (w ⊓ z) ≤ (w \ x) ⊓ (w ⊓ x) : inf_le_inf le_rfl (inf_le_inf le_rfl h) - ... = ⊥ : by rw [inf_comm, inf_inf_sdiff] - ... ≤ (w \ z) ⊓ (w ⊓ z) : bot_le) - (calc w \ x ⊔ (w ⊓ z) ≤ w \ x ⊔ (w ⊓ x) : sup_le_sup le_rfl (inf_le_inf le_rfl h) - ... ≤ w : by rw [sup_comm, sup_inf_sdiff] - ... = (w \ z) ⊔ (w ⊓ z) : by rw [sup_comm, sup_inf_sdiff]) - -lemma sdiff_le_iff : y \ x ≤ z ↔ y ≤ x ⊔ z := -⟨λ h, le_of_inf_le_sup_le - (le_of_eq - (calc y ⊓ (y \ x) = y \ x : inf_sdiff_right - ... = (x ⊓ (y \ x)) ⊔ (z ⊓ (y \ x)) : - by rw [inf_eq_right.2 h, inf_sdiff_self_right, bot_sup_eq] - ... = (x ⊔ z) ⊓ (y \ x) : inf_sup_right.symm)) - (calc y ⊔ y \ x = y : sup_sdiff_left - ... ≤ y ⊔ (x ⊔ z) : le_sup_left - ... = ((y \ x) ⊔ x) ⊔ z : by rw [←sup_assoc, ←@sup_sdiff_self_left _ x y] - ... = x ⊔ z ⊔ y \ x : by ac_refl), - λ h, le_of_inf_le_sup_le - (calc y \ x ⊓ x = ⊥ : inf_sdiff_self_left - ... ≤ z ⊓ x : bot_le) - (calc y \ x ⊔ x = y ⊔ x : sup_sdiff_self_left - ... ≤ (x ⊔ z) ⊔ x : sup_le_sup_right h x - ... ≤ z ⊔ x : by rw [sup_assoc, sup_comm, sup_assoc, sup_idem])⟩ - lemma sdiff_sdiff_le : x \ (x \ y) ≤ y := sdiff_le_iff.2 le_sdiff_sup lemma sdiff_triangle (x y z : α) : x \ z ≤ x \ y ⊔ y \ z := @@ -358,30 +311,6 @@ end @[simp] lemma le_sdiff_iff : x ≤ y \ x ↔ x = ⊥ := ⟨λ h, disjoint_self.1 (disjoint_sdiff_self_right.mono_right h), λ h, h.le.trans bot_le⟩ -@[simp] lemma sdiff_eq_bot_iff : y \ x = ⊥ ↔ y ≤ x := -by rw [←le_bot_iff, sdiff_le_iff, sup_bot_eq] - -lemma sdiff_le_comm : x \ y ≤ z ↔ x \ z ≤ y := -by rw [sdiff_le_iff, sup_comm, sdiff_le_iff] - -lemma sdiff_le_sdiff_right (h : w ≤ y) : w \ x ≤ y \ x := -le_of_inf_le_sup_le - (calc (w \ x) ⊓ (w ⊓ x) = ⊥ : by rw [inf_comm, inf_inf_sdiff] - ... ≤ (y \ x) ⊓ (w ⊓ x) : bot_le) - (calc w \ x ⊔ (w ⊓ x) = w : by rw [sup_comm, sup_inf_sdiff] - ... ≤ (y ⊓ (y \ x)) ⊔ w : le_sup_right - ... = (y ⊓ (y \ x)) ⊔ (y ⊓ w) : by rw inf_eq_right.2 h - ... = y ⊓ ((y \ x) ⊔ w) : by rw inf_sup_left - ... = ((y \ x) ⊔ (y ⊓ x)) ⊓ ((y \ x) ⊔ w) : - by rw [@sup_comm _ _ (y \ x) (y ⊓ x), sup_inf_sdiff] - ... = (y \ x) ⊔ ((y ⊓ x) ⊓ w) : by rw ←sup_inf_left - ... = (y \ x) ⊔ ((w ⊓ y) ⊓ x) : by ac_refl - ... = (y \ x) ⊔ (w ⊓ x) : by rw inf_eq_left.2 h) - -theorem sdiff_le_sdiff (h₁ : w ≤ y) (h₂ : z ≤ x) : w \ x ≤ y \ z := -calc w \ x ≤ w \ z : sdiff_le_sdiff_left h₂ - ... ≤ y \ z : sdiff_le_sdiff_right h₁ - lemma sdiff_lt_sdiff_right (h : x < y) (hz : z ≤ x) : x \ z < y \ z := (sdiff_le_sdiff_right h.le).lt_of_not_le $ λ h', h.not_le $ le_sdiff_sup.trans $ sup_le_of_le_sdiff_right h' hz @@ -579,10 +508,9 @@ end generalized_boolean_algebra /-! ### Boolean algebras -/ - /-- A Boolean algebra is a bounded distributive lattice with a complement operator `ᶜ` such that `x ⊓ xᶜ = ⊥` and `x ⊔ xᶜ = ⊤`. For convenience, it must also provide a set difference operation `\` -satisfying `x \ y = x ⊓ yᶜ`. +and a Heyting implication `⇨` satisfying `x \ y = x ⊓ yᶜ` and `x ⇨ y = y ⊔ xᶜ`. This is a generalization of (classical) logic of propositions, or the powerset lattice. @@ -591,16 +519,18 @@ to be present at define-time, the `extends` mechanism does not work with them. Instead, we extend using the underlying `has_bot` and `has_top` data typeclasses, and replicate the order axioms of those classes here. A "forgetful" instance back to `bounded_order` is provided. -/ -class boolean_algebra (α : Type u) extends distrib_lattice α, has_compl α, has_sdiff α, +class boolean_algebra (α : Type u) extends distrib_lattice α, has_compl α, has_sdiff α, has_himp α, has_top α, has_bot α := (inf_compl_le_bot : ∀x:α, x ⊓ xᶜ ≤ ⊥) (top_le_sup_compl : ∀x:α, ⊤ ≤ x ⊔ xᶜ) (le_top : ∀ a : α, a ≤ ⊤) (bot_le : ∀ a : α, ⊥ ≤ a) (sdiff := λ x y, x ⊓ yᶜ) +(himp := λ x y, y ⊔ xᶜ) (sdiff_eq : ∀ x y : α, x \ y = x ⊓ yᶜ . obviously) +(himp_eq : ∀ x y : α, x ⇨ y = y ⊔ xᶜ . obviously) -@[priority 100] -- see Note [lower instance priority] +@[priority 100] -- see Note [lower instance priority] instance boolean_algebra.to_bounded_order [h : boolean_algebra α] : bounded_order α := { ..h } @@ -617,13 +547,36 @@ def generalized_boolean_algebra.to_boolean_algebra [generalized_boolean_algebra section boolean_algebra variables [boolean_algebra α] -@[simp] lemma inf_compl_eq_bot : x ⊓ xᶜ = ⊥ := bot_unique $ boolean_algebra.inf_compl_le_bot x +@[simp] lemma inf_compl_eq_bot' : x ⊓ xᶜ = ⊥ := bot_unique $ boolean_algebra.inf_compl_le_bot x @[simp] lemma sup_compl_eq_top : x ⊔ xᶜ = ⊤ := top_unique $ boolean_algebra.top_le_sup_compl x -@[simp] lemma compl_inf_eq_bot : xᶜ ⊓ x = ⊥ := inf_comm.trans inf_compl_eq_bot @[simp] lemma compl_sup_eq_top : xᶜ ⊔ x = ⊤ := sup_comm.trans sup_compl_eq_top -theorem is_compl_compl : is_compl x xᶜ := -is_compl.of_eq inf_compl_eq_bot sup_compl_eq_top +lemma is_compl_compl : is_compl x xᶜ := is_compl.of_eq inf_compl_eq_bot' sup_compl_eq_top + +lemma sdiff_eq : x \ y = x ⊓ yᶜ := boolean_algebra.sdiff_eq x y +lemma himp_eq : x ⇨ y = y ⊔ xᶜ := boolean_algebra.himp_eq x y + +@[simp] lemma top_sdiff : ⊤ \ x = xᶜ := by rw [sdiff_eq, top_inf_eq] + +@[priority 100] +instance boolean_algebra.to_is_complemented : is_complemented α := ⟨λ x, ⟨xᶜ, is_compl_compl⟩⟩ + +@[priority 100] -- see Note [lower instance priority] +instance boolean_algebra.to_generalized_boolean_algebra : generalized_boolean_algebra α := +{ sup_inf_sdiff := λ a b, by rw [sdiff_eq, ←inf_sup_left, sup_compl_eq_top, inf_top_eq], + inf_inf_sdiff := λ a b, by { rw [sdiff_eq, ←inf_inf_distrib_left, inf_compl_eq_bot', inf_bot_eq], + congr }, + ..‹boolean_algebra α› } + +@[priority 100] -- See note [lower instance priority] +instance boolean_algebra.to_biheyting_algebra : biheyting_algebra α := +{ hnot := compl, + le_himp_iff := λ a b c, by rw [himp_eq, is_compl_compl.le_sup_right_iff_inf_left_le], + himp_bot := λ _, himp_eq.trans bot_sup_eq, + top_sdiff := λ a, by rw [sdiff_eq, top_inf_eq], + ..‹boolean_algebra α›, ..generalized_boolean_algebra.to_generalized_coheyting_algebra } + +@[simp] lemma hnot_eq_compl : ¬x = xᶜ := rfl theorem is_compl.eq_compl (h : is_compl x y) : x = yᶜ := h.left_unique is_compl_compl.symm @@ -643,18 +596,9 @@ by rw [eq_comm, compl_eq_iff_is_compl, eq_compl_iff_is_compl] theorem eq_compl_comm : x = yᶜ ↔ y = xᶜ := by rw [eq_comm, compl_eq_iff_is_compl, eq_compl_iff_is_compl] -theorem disjoint_compl_right : disjoint x xᶜ := is_compl_compl.disjoint -theorem disjoint_compl_left : disjoint xᶜ x := disjoint_compl_right.symm - theorem compl_unique (i : x ⊓ y = ⊥) (s : x ⊔ y = ⊤) : xᶜ = y := (is_compl.of_eq i s).compl_eq -@[simp] theorem compl_top : ⊤ᶜ = (⊥:α) := -is_compl_top_bot.compl_eq - -@[simp] theorem compl_bot : ⊥ᶜ = (⊤:α) := -is_compl_bot_top.compl_eq - @[simp] theorem compl_compl (x : α) : xᶜᶜ = x := is_compl_compl.symm.compl_eq @@ -686,12 +630,6 @@ is_compl_top_bot.compl_eq_iff @[simp] theorem compl_inf : (x ⊓ y)ᶜ = xᶜ ⊔ yᶜ := (is_compl_compl.inf_sup is_compl_compl).compl_eq -@[simp] theorem compl_sup : (x ⊔ y)ᶜ = xᶜ ⊓ yᶜ := -(is_compl_compl.sup_inf is_compl_compl).compl_eq - -theorem compl_le_compl (h : y ≤ x) : xᶜ ≤ yᶜ := -is_compl_compl.antitone is_compl_compl h - @[simp] theorem compl_le_compl_iff_le : yᶜ ≤ xᶜ ↔ x ≤ y := ⟨assume h, by have h := compl_le_compl h; simp at h; assumption, compl_le_compl⟩ @@ -708,71 +646,50 @@ theorem le_compl_iff_le_compl : y ≤ xᶜ ↔ x ≤ yᶜ := theorem compl_le_iff_compl_le : xᶜ ≤ y ↔ yᶜ ≤ x := ⟨compl_le_of_compl_le, compl_le_of_compl_le⟩ -lemma le_compl_iff_disjoint_right : x ≤ yᶜ ↔ disjoint x y := is_compl_compl.le_right_iff -lemma le_compl_iff_disjoint_left : y ≤ xᶜ ↔ disjoint x y := -le_compl_iff_disjoint_right.trans disjoint.comm - -lemma disjoint_compl_left_iff : disjoint xᶜ y ↔ y ≤ x := -by rw [←le_compl_iff_disjoint_left, compl_compl] - -lemma disjoint_compl_right_iff : disjoint x yᶜ ↔ x ≤ y := -by rw [←le_compl_iff_disjoint_right, compl_compl] - -alias le_compl_iff_disjoint_right ↔ _ disjoint.le_compl_right -alias le_compl_iff_disjoint_left ↔ _ disjoint.le_compl_left -alias disjoint_compl_left_iff ↔ _ has_le.le.disjoint_compl_left -alias disjoint_compl_right_iff ↔ _ has_le.le.disjoint_compl_right - -lemma sdiff_eq : x \ y = x ⊓ yᶜ := boolean_algebra.sdiff_eq x y - -@[priority 100] -- see Note [lower instance priority] -instance boolean_algebra.to_generalized_boolean_algebra : generalized_boolean_algebra α := -{ sup_inf_sdiff := λ a b, by rw [sdiff_eq, ←inf_sup_left, sup_compl_eq_top, inf_top_eq], - inf_inf_sdiff := λ a b, by { rw [sdiff_eq, ←inf_inf_distrib_left, inf_compl_eq_bot, inf_bot_eq], - congr }, - ..‹boolean_algebra α› } - -@[priority 100] -instance boolean_algebra.to_is_complemented : is_complemented α := ⟨λ x, ⟨xᶜ, is_compl_compl⟩⟩ +@[simp] lemma sdiff_compl : x \ yᶜ = x ⊓ y := by rw [sdiff_eq, compl_compl] ---TODO@Yaël: Once we have co-Heyting algebras, we won't need to use the default value for `sdiff` instance : boolean_algebra αᵒᵈ := { compl := λ a, to_dual (of_dual a)ᶜ, - inf_compl_le_bot := λ _, sup_compl_eq_top.ge, - top_le_sup_compl := λ _, inf_compl_eq_bot.ge, + sdiff := λ a b, to_dual (of_dual b ⇨ of_dual a), + himp := λ a b, to_dual (of_dual b \ of_dual a), + inf_compl_le_bot := λ a, @codisjoint_hnot_right _ _ (of_dual a), + top_le_sup_compl := λ a, @disjoint_compl_right _ _ (of_dual a), + sdiff_eq := λ _ _, himp_eq, + himp_eq := λ _ _, sdiff_eq, ..order_dual.distrib_lattice α, ..order_dual.bounded_order α } -@[simp] lemma of_dual_compl (a : αᵒᵈ) : of_dual aᶜ = (of_dual a)ᶜ := rfl -@[simp] lemma to_dual_compl (a : α) : to_dual aᶜ = (to_dual a)ᶜ := rfl - -@[simp] theorem sdiff_compl : x \ yᶜ = x ⊓ y := by rw [sdiff_eq, compl_compl] - -@[simp] theorem top_sdiff : ⊤ \ x = xᶜ := by rw [sdiff_eq, top_inf_eq] -@[simp] theorem sdiff_top : x \ ⊤ = ⊥ := by rw [sdiff_eq, compl_top, inf_bot_eq] - @[simp] lemma sup_inf_inf_compl : (x ⊓ y) ⊔ (x ⊓ yᶜ) = x := by rw [← sdiff_eq, sup_inf_sdiff _ _] @[simp] lemma compl_sdiff : (x \ y)ᶜ = xᶜ ⊔ y := by rw [sdiff_eq, compl_inf, compl_compl] +lemma disjoint_compl_left_iff : disjoint xᶜ y ↔ y ≤ x := +by rw [←le_compl_iff_disjoint_left, compl_compl] + +lemma disjoint_compl_right_iff : disjoint x yᶜ ↔ x ≤ y := +by rw [←le_compl_iff_disjoint_right, compl_compl] + +alias disjoint_compl_left_iff ↔ _ has_le.le.disjoint_compl_left +alias disjoint_compl_right_iff ↔ _ has_le.le.disjoint_compl_right + end boolean_algebra instance Prop.boolean_algebra : boolean_algebra Prop := -{ inf_compl_le_bot := λ p ⟨Hp, Hpc⟩, Hpc Hp, +{ compl := not, + himp_eq := λ p q, propext imp_iff_or_not, + inf_compl_le_bot := λ p ⟨Hp, Hpc⟩, Hpc Hp, top_le_sup_compl := λ p H, classical.em p, - .. Prop.has_compl, - .. Prop.distrib_lattice, - .. Prop.bounded_order } + .. Prop.heyting_algebra, ..generalized_heyting_algebra.to_distrib_lattice } instance pi.boolean_algebra {ι : Type u} {α : ι → Type v} [∀ i, boolean_algebra (α i)] : boolean_algebra (Π i, α i) := { sdiff_eq := λ x y, funext $ λ i, sdiff_eq, + himp_eq := λ x y, funext $ λ i, himp_eq, inf_compl_le_bot := λ _ _, boolean_algebra.inf_compl_le_bot _, top_le_sup_compl := λ _ _, boolean_algebra.top_le_sup_compl _, .. pi.has_sdiff, - .. pi.has_compl, - .. pi.bounded_order, + .. pi.heyting_algebra, .. pi.distrib_lattice } instance : boolean_algebra bool := @@ -803,21 +720,10 @@ protected def function.injective.generalized_boolean_algebra [has_sup α] [has_i (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_bot : f ⊥ = ⊥) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : generalized_boolean_algebra α := -{ sdiff := (\), - bot := ⊥, - sup_inf_sdiff := λ a b, hf $ (map_sup _ _).trans begin - rw map_sdiff, - convert sup_inf_sdiff _ _, - exact map_inf _ _, - end, - inf_inf_sdiff := λ a b, hf $ (map_inf _ _).trans begin - rw map_sdiff, - convert inf_inf_sdiff _ _, - exact map_inf _ _, - end, - le_sup_inf := λ a b c, (map_inf _ _).le.trans $ by { convert le_sup_inf, exact map_sup _ _, - exact map_sup _ _, convert map_sup _ _, exact (map_inf _ _).symm }, - ..hf.lattice f map_sup map_inf } +{ sup_inf_sdiff := λ a b, hf $ by erw [map_sup, map_sdiff, map_inf, sup_inf_sdiff], + inf_inf_sdiff := λ a b, hf $ by erw [map_inf, map_sdiff, map_inf, inf_inf_sdiff, map_bot], + ..hf.generalized_coheyting_algebra f map_sup map_inf map_bot map_sdiff, + ..hf.distrib_lattice f map_sup map_inf } /-- Pullback a `boolean_algebra` along an injection. -/ @[reducible] -- See note [reducible non-instances] @@ -839,24 +745,7 @@ protected def function.injective.boolean_algebra [has_sup α] [has_inf α] [has_ end lift -namespace punit -variables (a b : punit.{u+1}) - instance : boolean_algebra punit := by refine_struct -{ top := star, - bot := star, - sup := λ _ _, star, - inf := λ _ _, star, - compl := λ _, star, - sdiff := λ _ _, star, ..punit.linear_order }; +{ ..punit.biheyting_algebra }; intros; trivial <|> exact subsingleton.elim _ _ - -@[simp] lemma top_eq : (⊤ : punit) = star := rfl -@[simp] lemma bot_eq : (⊥ : punit) = star := rfl -@[simp] lemma sup_eq : a ⊔ b = star := rfl -@[simp] lemma inf_eq : a ⊓ b = star := rfl -@[simp] lemma compl_eq : aᶜ = star := rfl -@[simp] lemma sdiff_eq : a \ b = star := rfl - -end punit diff --git a/src/order/heyting/basic.lean b/src/order/heyting/basic.lean new file mode 100644 index 0000000000000..0b602af952915 --- /dev/null +++ b/src/order/heyting/basic.lean @@ -0,0 +1,725 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.bounded_order + +/-! +# Heyting algebras + +This file defines Heyting, co-Heyting and bi-Heyting algebras. + +An Heyting algebra is a bounded distributive lattice with an implication operation `⇨` such that +`a ≤ b ⇨ c ↔ a ⊓ b ≤ c`. It also comes with a pseudo-complement `ᶜ`, such that `aᶜ = a ⇨ ⊥`. + +Co-Heyting algebras are dual to Heyting algebras. They have a difference `\` and a negation `¬` +such that `a \ b ≤ c ↔ a ≤ b ⊔ c` and `¬a = ⊤ \ a`. + +Bi-Heyting algebras are Heyting algebras that are also co-Heyting algebras. + +From a logic standpoint, Heyting algebras precisely model intuitionistic logic, whereas boolean +algebras model classical logic. + +Heyting algebras are the order theoretic equivalent of cartesian-closed categories. + +## Main declarations + +* `generalized_heyting_algebra`: Heyting algebra without a top element (nor negation). +* `generalized_coheyting_algebra`: Co-Heyting algebra without a bottom element (nor complement). +* `heyting_algebra`: Heyting algebra. +* `coheyting_algebra`: Co-Heyting algebra. +* `biheyting_algebra`: bi-Heyting algebra. + +## Notation + +* `⇨`: Heyting implication +* `\`: Difference +* `¬`: Heyting negation +* `ᶜ`: (Pseudo-)complement + +## References + +* [Francis Borceux, *Handbook of Categorical Algebra III*][borceux-vol3] + +## Tags + +Heyting, Brouwer, algebra, implication, negation, intuitionistic +-/ + +set_option old_structure_cmd true + +open function order_dual + +universes u +variables {ι α β : Type*} + +/-! ### Notation -/ + +/-- Syntax typeclass for Heyting implication `⇨`. -/ +@[notation_class] class has_himp (α : Type*) := (himp : α → α → α) + +/-- Syntax typeclass for Heyting negation `¬`. + +The difference between `has_hnot` and `has_compl` is that the former belongs to Heyting algebras, +while the latter belongs to co-Heyting algebras. They are both pseudo-complements, but `compl` +underestimates while `hnot` overestimates. In boolean algebras, they are equal. See `hnot_eq_compl`. +-/ +@[notation_class] class has_hnot (α : Type*) := (hnot : α → α) + +export has_himp (himp) has_sdiff (sdiff) has_hnot (hnot) + +infixr ` ⇨ `:60 := himp +prefix `¬`:72 := hnot + +instance [has_himp α] [has_himp β] : has_himp (α × β) := ⟨λ a b, (a.1 ⇨ b.1, a.2 ⇨ b.2)⟩ +instance [has_hnot α] [has_hnot β] : has_hnot (α × β) := ⟨λ a, (¬a.1, ¬a.2)⟩ +instance [has_sdiff α] [has_sdiff β] : has_sdiff (α × β) := ⟨λ a b, (a.1 \ b.1, a.2 \ b.2)⟩ +instance [has_compl α] [has_compl β] : has_compl (α × β) := ⟨λ a, (a.1ᶜ, a.2ᶜ)⟩ + +@[simp] lemma fst_himp [has_himp α] [has_himp β] (a b : α × β) : (a ⇨ b).1 = a.1 ⇨ b.1 := rfl +@[simp] lemma snd_himp [has_himp α] [has_himp β] (a b : α × β) : (a ⇨ b).2 = a.2 ⇨ b.2 := rfl +@[simp] lemma fst_hnot [has_hnot α] [has_hnot β] (a : α × β) : (¬a).1 = ¬a.1 := rfl +@[simp] lemma snd_hnot [has_hnot α] [has_hnot β] (a : α × β) : (¬a).2 = ¬a.2 := rfl +@[simp] lemma fst_sdiff [has_sdiff α] [has_sdiff β] (a b : α × β) : (a \ b).1 = a.1 \ b.1 := rfl +@[simp] lemma snd_sdiff [has_sdiff α] [has_sdiff β] (a b : α × β) : (a \ b).2 = a.2 \ b.2 := rfl +@[simp] lemma fst_compl [has_compl α] [has_compl β] (a : α × β) : aᶜ.1 = a.1ᶜ := rfl +@[simp] lemma snd_compl [has_compl α] [has_compl β] (a : α × β) : aᶜ.2 = a.2ᶜ := rfl + +namespace pi +variables {π : ι → Type*} + +instance [Π i, has_himp (π i)] : has_himp (Π i, π i) := ⟨λ a b i, a i ⇨ b i⟩ +instance [Π i, has_hnot (π i)] : has_hnot (Π i, π i) := ⟨λ a i, ¬a i⟩ + +lemma himp_def [Π i, has_himp (π i)] (a b : Π i, π i) : (a ⇨ b) = λ i, a i ⇨ b i := rfl +lemma hnot_def [Π i, has_hnot (π i)] (a : Π i, π i) : ¬a = λ i, ¬a i := rfl + +@[simp] lemma himp_apply [Π i, has_himp (π i)] (a b : Π i, π i) (i : ι) : (a ⇨ b) i = a i ⇨ b i := +rfl +@[simp] lemma hnot_apply [Π i, has_hnot (π i)] (a : Π i, π i) (i : ι) : (¬a) i = ¬a i := rfl + +end pi + +/-- A generalized Heyting algebra is a lattice with an additional binary operation `⇨` called +Heyting implication such that `a ⇨` is right adjoint to `a ⊓`. + + This generalizes `heyting_algebra` by not requiring a bottom element. -/ +class generalized_heyting_algebra (α : Type*) extends lattice α, has_top α, has_himp α := +(le_top : ∀ a : α, a ≤ ⊤) +(le_himp_iff (a b c : α) : a ≤ b ⇨ c ↔ a ⊓ b ≤ c) + +/-- A generalized co-Heyting algebra is a lattice with an additional binary difference operation `\` +such that `\ a` is right adjoint to `⊔ a`. + +This generalizes `coheyting_algebra` by not requiring a top element. -/ +class generalized_coheyting_algebra (α : Type*) extends lattice α, has_bot α, has_sdiff α := +(bot_le : ∀ a : α, ⊥ ≤ a) +(sdiff_le_iff (a b c : α) : a \ b ≤ c ↔ a ≤ b ⊔ c) + +/-- A Heyting algebra is a bounded lattice with an additional binary operation `⇨` called Heyting +implication such that `a ⇨` is right adjoint to `a ⊓`. -/ +class heyting_algebra (α : Type*) extends generalized_heyting_algebra α, has_bot α, has_compl α := +(bot_le : ∀ a : α, ⊥ ≤ a) +(himp_bot (a : α) : a ⇨ ⊥ = aᶜ) + +/-- A co-Heyting algebra is a bounded lattice with an additional binary difference operation `\` +such that `\ a` is right adjoint to `⊔ a`. -/ +class coheyting_algebra (α : Type*) + extends generalized_coheyting_algebra α, has_top α, has_hnot α := +(le_top : ∀ a : α, a ≤ ⊤) +(top_sdiff (a : α) : ⊤ \ a = ¬a) + +/-- A bi-Heyting algebra is a Heyting algebra that is also a co-Heyting algebra. -/ +class biheyting_algebra (α : Type*) extends heyting_algebra α, has_sdiff α, has_hnot α := +(sdiff_le_iff (a b c : α) : a \ b ≤ c ↔ a ≤ b ⊔ c) +(top_sdiff (a : α) : ⊤ \ a = ¬a) + +@[priority 100] -- See note [lower instance priority] +instance generalized_heyting_algebra.to_order_top [generalized_heyting_algebra α] : order_top α := +{ ..‹generalized_heyting_algebra α› } + +@[priority 100] -- See note [lower instance priority] +instance generalized_coheyting_algebra.to_order_bot [generalized_coheyting_algebra α] : + order_bot α := +{ ..‹generalized_coheyting_algebra α› } + +@[priority 100] -- See note [lower instance priority] +instance heyting_algebra.to_bounded_order [heyting_algebra α] : bounded_order α := +{ ..‹heyting_algebra α› } + +@[priority 100] -- See note [lower instance priority] +instance coheyting_algebra.to_bounded_order [coheyting_algebra α] : bounded_order α := +{ ..‹coheyting_algebra α› } + +@[priority 100] -- See note [lower instance priority] +instance biheyting_algebra.to_coheyting_algebra [biheyting_algebra α] : coheyting_algebra α := +{ ..‹biheyting_algebra α› } + +/-- Construct a Heyting algebra from the lattice structure and Heyting implication alone. -/ +@[reducible] -- See note [reducible non-instances] +def heyting_algebra.of_himp [distrib_lattice α] [bounded_order α] (himp : α → α → α) + (le_himp_iff : ∀ a b c, a ≤ himp b c ↔ a ⊓ b ≤ c) : heyting_algebra α := +{ himp := himp, + compl := λ a, himp a ⊥, + le_himp_iff := le_himp_iff, + himp_bot := λ a, rfl, + ..‹distrib_lattice α›, ..‹bounded_order α› } + +/-- Construct a Heyting algebra from the lattice structure and complement operator alone. -/ +@[reducible] -- See note [reducible non-instances] +def heyting_algebra.of_compl [distrib_lattice α] [bounded_order α] (compl : α → α) + (le_himp_iff : ∀ a b c, a ≤ compl b ⊔ c ↔ a ⊓ b ≤ c) : heyting_algebra α := +{ himp := λ a, (⊔) (compl a), + compl := compl, + le_himp_iff := le_himp_iff, + himp_bot := λ a, sup_bot_eq, + ..‹distrib_lattice α›, ..‹bounded_order α› } + +/-- Construct a co-Heyting algebra from the lattice structure and the difference alone. -/ +@[reducible] -- See note [reducible non-instances] +def coheyting_algebra.of_sdiff [distrib_lattice α] [bounded_order α] (sdiff : α → α → α) + (sdiff_le_iff : ∀ a b c, sdiff a b ≤ c ↔ a ≤ b ⊔ c) : coheyting_algebra α := +{ sdiff := sdiff, + hnot := λ a, sdiff ⊤ a, + sdiff_le_iff := sdiff_le_iff, + top_sdiff := λ a, rfl, + ..‹distrib_lattice α›, ..‹bounded_order α› } + +/-- Construct a co-Heyting algebra from the difference and Heyting negation alone. -/ +@[reducible] -- See note [reducible non-instances] +def coheyting_algebra.of_hnot [distrib_lattice α] [bounded_order α] (hnot : α → α) + (sdiff_le_iff : ∀ a b c, a ⊓ hnot b ≤ c ↔ a ≤ b ⊔ c) : coheyting_algebra α := +{ sdiff := λ a b, (a ⊓ hnot b), + hnot := hnot, + sdiff_le_iff := sdiff_le_iff, + top_sdiff := λ a, top_inf_eq, + ..‹distrib_lattice α›, ..‹bounded_order α› } + +section generalized_heyting_algebra +variables [generalized_heyting_algebra α] {a b c d : α} + +/- In this section, we'll give interpretations of these results in the Heyting algebra model of +intuitionistic logic,- where `≤` can be interpreted as "validates", `⇨` as "implies", `⊓` as "and", +`⊔` as "or", `⊥` as "false" and `⊤` as "true". Note that we confuse `→` and `⊢` because those are +the same in this logic. + +See also `Prop.heyting_algebra`. -/ + +-- `p → q → r ↔ p ∧ q → r` +@[simp] lemma le_himp_iff : a ≤ b ⇨ c ↔ a ⊓ b ≤ c := generalized_heyting_algebra.le_himp_iff _ _ _ + +-- `p → q → r ↔ q → p → r` +lemma le_himp_comm : a ≤ b ⇨ c ↔ b ≤ a ⇨ c := by rw [le_himp_iff, le_himp_iff, inf_comm] + +-- `p → q → p` +lemma le_himp : a ≤ b ⇨ a := le_himp_iff.2 inf_le_left + +-- `p → p → q ↔ p → q` +@[simp] lemma le_himp_iff_left : a ≤ a ⇨ b ↔ a ≤ b := by rw [le_himp_iff, inf_idem] + +-- `p → p` +@[simp] lemma himp_self : a ⇨ a = ⊤ := top_le_iff.1 $ le_himp_iff.2 inf_le_right + +-- `(p → q) ∧ p → q` +lemma himp_inf_le : (a ⇨ b) ⊓ a ≤ b := le_himp_iff.1 le_rfl + +-- `p ∧ (p → q) → q` +lemma inf_himp_le : a ⊓ (a ⇨ b) ≤ b := by rw [inf_comm, ←le_himp_iff] + +-- `p ∧ (p → q) ↔ p ∧ q` +@[simp] lemma inf_himp (a b : α) : a ⊓ (a ⇨ b) = a ⊓ b := +le_antisymm (le_inf inf_le_left $ by rw [inf_comm, ←le_himp_iff]) $ inf_le_inf_left _ le_himp + +-- `(p → q) ∧ p ↔ q ∧ p` +@[simp] lemma himp_inf_self (a b : α) : (a ⇨ b) ⊓ a = b ⊓ a := by rw [inf_comm, inf_himp, inf_comm] + +/-- The **deduction theorem** in the Heyting algebra model of intuitionistic logic: +an implication holds iff the conclusion follows from the hypothesis. -/ +@[simp] lemma himp_eq_top_iff : a ⇨ b = ⊤ ↔ a ≤ b := by rw [←top_le_iff, le_himp_iff, top_inf_eq] + +-- `p → true`, `true → p ↔ p` +@[simp] lemma himp_top : a ⇨ ⊤ = ⊤ := himp_eq_top_iff.2 le_top +@[simp] lemma top_himp : ⊤ ⇨ a = a := eq_of_forall_le_iff $ λ b, by rw [le_himp_iff, inf_top_eq] + +-- `p → q → r ↔ p ∧ q → r` +lemma himp_himp (a b c : α) : a ⇨ b ⇨ c = a ⊓ b ⇨ c := +eq_of_forall_le_iff $ λ d, by simp_rw [le_himp_iff, inf_assoc] + +-- `(q → r) → (p → q) → q → r` +@[simp] lemma himp_le_himp_himp : b ⇨ c ≤ (a ⇨ b) ⇨ a ⇨ c := +begin + rw [le_himp_iff, le_himp_iff, inf_assoc, himp_inf_self, ←inf_assoc, himp_inf_self, inf_assoc], + exact inf_le_left, +end + +-- `p → q → r ↔ q → p → r` +lemma himp_left_comm (a b c : α) : a ⇨ b ⇨ c = b ⇨ a ⇨ c := by simp_rw [himp_himp, inf_comm] + +lemma himp_inf_distrib (a b c : α) : a ⇨ b ⊓ c = (a ⇨ b) ⊓ (a ⇨ c) := +eq_of_forall_le_iff $ λ d, by simp_rw [le_himp_iff, le_inf_iff, le_himp_iff] + +lemma sup_himp_distrib (a b c : α) : a ⊔ b ⇨ c = (a ⇨ c) ⊓ (b ⇨ c) := +eq_of_forall_le_iff $ λ d, by { rw [le_inf_iff, le_himp_comm, sup_le_iff], simp_rw le_himp_comm } + +lemma himp_le_himp_left (h : a ≤ b) : c ⇨ a ≤ c ⇨ b := le_himp_iff.2 $ himp_inf_le.trans h + +lemma himp_le_himp_right (h : a ≤ b) : b ⇨ c ≤ a ⇨ c := +le_himp_iff.2 $ (inf_le_inf_left _ h).trans himp_inf_le + +lemma himp_le_himp (hab : a ≤ b) (hcd : c ≤ d) : b ⇨ c ≤ a ⇨ d := +(himp_le_himp_right hab).trans $ himp_le_himp_left hcd + +@[simp] lemma sup_himp_self_left (a b : α) : (a ⊔ b) ⇨ a = b ⇨ a := +by rw [sup_himp_distrib, himp_self, top_inf_eq] + +@[simp] lemma sup_himp_self_right (a b : α) : (a ⊔ b) ⇨ b = a ⇨ b := +by rw [sup_himp_distrib, himp_self, inf_top_eq] + +@[priority 100] -- See note [lower instance priority] +instance generalized_heyting_algebra.to_distrib_lattice : distrib_lattice α := +distrib_lattice.of_inf_sup_le $ λ a b c, + by simp_rw [@inf_comm _ _ a, ←le_himp_iff, sup_le_iff, le_himp_iff, ←sup_le_iff] + +instance : generalized_coheyting_algebra αᵒᵈ := +{ sdiff := λ a b, to_dual (of_dual b ⇨ of_dual a), + sdiff_le_iff := λ a b c, by { rw sup_comm, exact le_himp_iff }, + ..order_dual.lattice α, ..order_dual.order_bot α } + +instance prod.generalized_heyting_algebra [generalized_heyting_algebra β] : + generalized_heyting_algebra (α × β) := +{ le_himp_iff := λ a b c, and_congr le_himp_iff le_himp_iff, + ..prod.lattice α β, ..prod.order_top α β, ..prod.has_himp, ..prod.has_compl } + +instance pi.generalized_heyting_algebra {α : ι → Type*} [Π i, generalized_heyting_algebra (α i)] : + generalized_heyting_algebra (Π i, α i) := +by { pi_instance, exact λ a b c, forall_congr (λ i, le_himp_iff) } + +end generalized_heyting_algebra + +section generalized_coheyting_algebra +variables [generalized_coheyting_algebra α] {a b c d : α} + +@[simp] lemma sdiff_le_iff : a \ b ≤ c ↔ a ≤ b ⊔ c := +generalized_coheyting_algebra.sdiff_le_iff _ _ _ + +lemma sdiff_le_comm : a \ b ≤ c ↔ a \ c ≤ b := by rw [sdiff_le_iff, sdiff_le_iff, sup_comm] + +lemma sdiff_le : a \ b ≤ a := sdiff_le_iff.2 le_sup_right + +lemma disjoint.disjoint_sdiff_left (h : disjoint a b) : disjoint (a \ c) b := h.mono_left sdiff_le +lemma disjoint.disjoint_sdiff_right (h : disjoint a b) : disjoint a (b \ c) := h.mono_right sdiff_le + +@[simp] lemma sdiff_le_iff_left : a \ b ≤ b ↔ a ≤ b := by rw [sdiff_le_iff, sup_idem] + +@[simp] lemma sdiff_self : a \ a = ⊥ := le_bot_iff.1 $ sdiff_le_iff.2 le_sup_left + +lemma le_sup_sdiff : a ≤ b ⊔ a \ b := sdiff_le_iff.1 le_rfl +lemma le_sdiff_sup : a ≤ a \ b ⊔ b := by rw [sup_comm, ←sdiff_le_iff] + +@[simp] lemma sup_sdiff_self (a b : α) : a ⊔ b \ a = a ⊔ b := +le_antisymm (sup_le_sup_left sdiff_le _) (sup_le le_sup_left le_sup_sdiff) + +@[simp] lemma sdiff_sup_self (a b : α) : b \ a ⊔ a = b ⊔ a := +by rw [sup_comm, sup_sdiff_self, sup_comm] + +@[simp] lemma sdiff_eq_bot_iff : a \ b = ⊥ ↔ a ≤ b := by rw [←le_bot_iff, sdiff_le_iff, sup_bot_eq] + +@[simp] lemma sdiff_bot : a \ ⊥ = a := eq_of_forall_ge_iff $ λ b, by rw [sdiff_le_iff, bot_sup_eq] +@[simp] lemma bot_sdiff : ⊥ \ a = ⊥ := sdiff_eq_bot_iff.2 bot_le + +lemma sdiff_sdiff (a b c : α) : a \ b \ c = a \ (b ⊔ c) := +eq_of_forall_ge_iff $ λ d, by simp_rw [sdiff_le_iff, sup_assoc] + +@[simp] lemma sdiff_sdiff_le_sdiff : a \ b \ (a \ c) ≤ c \ b := +begin + rw [sdiff_le_iff, sdiff_le_iff, sup_left_comm, sup_sdiff_self, sup_left_comm, sdiff_sup_self, + sup_left_comm], + exact le_sup_left, +end + +lemma sdiff_right_comm (a b c : α) : a \ b \ c = a \ c \ b := by simp_rw [sdiff_sdiff, sup_comm] + +lemma sup_sdiff_distrib (a b c : α) : (a ⊔ b) \ c = a \ c ⊔ b \ c := +eq_of_forall_ge_iff $ λ d, by simp_rw [sdiff_le_iff, sup_le_iff, sdiff_le_iff] + +lemma sdiff_inf_distrib (a b c : α) : a \ (b ⊓ c) = a \ b ⊔ a \ c := +eq_of_forall_ge_iff $ λ d, by { rw [sup_le_iff, sdiff_le_comm, le_inf_iff], simp_rw sdiff_le_comm } + +lemma sdiff_le_sdiff_right (h : a ≤ b) : a \ c ≤ b \ c := sdiff_le_iff.2 $ h.trans $ le_sup_sdiff + +lemma sdiff_le_sdiff_left (h : a ≤ b) : c \ b ≤ c \ a := +sdiff_le_iff.2 $ le_sup_sdiff.trans $ sup_le_sup_right h _ + +lemma sdiff_le_sdiff (hab : a ≤ b) (hcd : c ≤ d) : a \ d ≤ b \ c := +(sdiff_le_sdiff_right hab).trans $ sdiff_le_sdiff_left hcd + +-- cf. `is_compl.inf_sup` +lemma sdiff_inf : a \ (b ⊓ c) = a \ b ⊔ a \ c := sdiff_inf_distrib _ _ _ + +@[simp] lemma sdiff_inf_self_left (a b : α) : a \ (a ⊓ b) = a \ b := +by rw [sdiff_inf, sdiff_self, bot_sup_eq] + +@[simp] lemma sdiff_inf_self_right (a b : α) : b \ (a ⊓ b) = b \ a := +by rw [sdiff_inf, sdiff_self, sup_bot_eq] + +@[priority 100] -- See note [lower instance priority] +instance generalized_coheyting_algebra.to_distrib_lattice : distrib_lattice α := +{ le_sup_inf := λ a b c, by simp_rw [←sdiff_le_iff, le_inf_iff, sdiff_le_iff, ←le_inf_iff], + ..‹generalized_coheyting_algebra α› } + +instance : generalized_heyting_algebra αᵒᵈ := +{ himp := λ a b, to_dual (of_dual b \ of_dual a), + le_himp_iff := λ a b c, by { rw inf_comm, exact sdiff_le_iff }, + ..order_dual.lattice α, ..order_dual.order_top α } + +instance prod.generalized_coheyting_algebra [generalized_coheyting_algebra β] : + generalized_coheyting_algebra (α × β) := +{ sdiff_le_iff := λ a b c, and_congr sdiff_le_iff sdiff_le_iff, + ..prod.lattice α β, ..prod.order_bot α β, ..prod.has_sdiff, ..prod.has_hnot } + +instance pi.generalized_coheyting_algebra {α : ι → Type*} + [Π i, generalized_coheyting_algebra (α i)] : generalized_coheyting_algebra (Π i, α i) := +by { pi_instance, exact λ a b c, forall_congr (λ i, sdiff_le_iff) } + +end generalized_coheyting_algebra + +section heyting_algebra +variables [heyting_algebra α] {a b c : α} + +@[simp] lemma himp_bot (a : α) : a ⇨ ⊥ = aᶜ := heyting_algebra.himp_bot _ +@[simp] lemma bot_himp (a : α) : ⊥ ⇨ a = ⊤ := himp_eq_top_iff.2 bot_le + +lemma compl_sup_distrib (a b : α) : (a ⊔ b)ᶜ = aᶜ ⊓ bᶜ := by simp_rw [←himp_bot, sup_himp_distrib] +@[simp] lemma compl_sup : (a ⊔ b)ᶜ = aᶜ ⊓ bᶜ := compl_sup_distrib _ _ + +lemma compl_le_himp : aᶜ ≤ a ⇨ b := (himp_bot _).ge.trans $ himp_le_himp_left bot_le + +lemma compl_sup_le_himp : aᶜ ⊔ b ≤ a ⇨ b := sup_le compl_le_himp le_himp +lemma sup_compl_le_himp : b ⊔ aᶜ ≤ a ⇨ b := sup_le le_himp compl_le_himp + +-- `p → ¬ p ↔ ¬ p` +@[simp] lemma himp_compl (a : α) : a ⇨ aᶜ = aᶜ := by rw [←himp_bot, himp_himp, inf_idem] + +-- `p → ¬ q ↔ q → ¬ p` +lemma himp_compl_comm (a b : α) : a ⇨ bᶜ = b ⇨ aᶜ := by simp_rw [←himp_bot, himp_left_comm] + +lemma le_compl_iff_disjoint_right : a ≤ bᶜ ↔ disjoint a b := +by rw [←himp_bot, le_himp_iff, disjoint] + +lemma le_compl_iff_disjoint_left : a ≤ bᶜ ↔ disjoint b a := +le_compl_iff_disjoint_right.trans disjoint.comm + +lemma le_compl_comm : a ≤ bᶜ ↔ b ≤ aᶜ := +by rw [le_compl_iff_disjoint_right, le_compl_iff_disjoint_left] + +alias le_compl_iff_disjoint_right ↔ _ disjoint.le_compl_right +alias le_compl_iff_disjoint_left ↔ _ disjoint.le_compl_left + +lemma disjoint_compl_left : disjoint aᶜ a := le_himp_iff.1 (himp_bot _).ge +lemma disjoint_compl_right : disjoint a aᶜ := disjoint_compl_left.symm + +@[simp] lemma inf_compl_self (a : α) : a ⊓ aᶜ = ⊥ := disjoint_compl_right.eq_bot +@[simp] lemma compl_inf_self (a : α) : aᶜ ⊓ a = ⊥ := disjoint_compl_left.eq_bot +lemma inf_compl_eq_bot : a ⊓ aᶜ = ⊥ := inf_compl_self _ +lemma compl_inf_eq_bot : aᶜ ⊓ a = ⊥ := compl_inf_self _ + +@[simp] lemma compl_top : (⊤ : α)ᶜ = ⊥ := +eq_of_forall_le_iff $ λ a, by rw [le_compl_iff_disjoint_right, disjoint_top, le_bot_iff] + +@[simp] lemma compl_bot : (⊥ : α)ᶜ = ⊤ := by rw [←himp_bot, himp_self] + +lemma le_compl_compl : a ≤ aᶜᶜ := disjoint_compl_right.le_compl_right + +lemma compl_anti : antitone (compl : α → α) := λ a b h, le_compl_comm.1 $ h.trans le_compl_compl + +lemma compl_le_compl (h : a ≤ b) : bᶜ ≤ aᶜ := compl_anti h + +@[simp] lemma compl_compl_compl (a : α) : aᶜᶜᶜ = aᶜ := +(compl_anti le_compl_compl).antisymm le_compl_compl + +@[simp] lemma disjoint_compl_compl_left_iff : disjoint aᶜᶜ b ↔ disjoint a b := +by simp_rw [←le_compl_iff_disjoint_left, compl_compl_compl] + +@[simp] lemma disjoint_compl_compl_right_iff : disjoint a bᶜᶜ ↔ disjoint a b := +by simp_rw [←le_compl_iff_disjoint_right, compl_compl_compl] + +lemma compl_sup_compl_le : aᶜ ⊔ bᶜ ≤ (a ⊓ b)ᶜ := +sup_le (compl_anti inf_le_left) $ compl_anti inf_le_right + +lemma compl_compl_inf_distrib (a b : α) : (a ⊓ b)ᶜᶜ = aᶜᶜ ⊓ bᶜᶜ := +begin + refine ((compl_anti compl_sup_compl_le).trans (compl_sup_distrib _ _).le).antisymm _, + rw [le_compl_iff_disjoint_right, disjoint_assoc, disjoint_compl_compl_left_iff, + disjoint_left_comm, disjoint_compl_compl_left_iff, ←disjoint_assoc, inf_comm], + exact disjoint_compl_right, +end + +lemma compl_compl_himp_distrib (a b : α) : (a ⇨ b)ᶜᶜ = aᶜᶜ ⇨ bᶜᶜ := +begin + refine le_antisymm _ _, + { rw [le_himp_iff, ←compl_compl_inf_distrib], + exact compl_anti (compl_anti himp_inf_le) }, + { refine le_compl_comm.1 ((compl_anti compl_sup_le_himp).trans _), + rw [compl_sup_distrib, le_compl_iff_disjoint_right, disjoint_right_comm, + ←le_compl_iff_disjoint_right], + exact inf_himp_le } +end + +instance : coheyting_algebra αᵒᵈ := +{ hnot := to_dual ∘ compl ∘ of_dual, + sdiff := λ a b, to_dual (of_dual b ⇨ of_dual a), + sdiff_le_iff := λ a b c, by { rw sup_comm, exact le_himp_iff }, + top_sdiff := himp_bot, + ..order_dual.lattice α, ..order_dual.bounded_order α } + +@[simp] lemma of_dual_hnot (a : αᵒᵈ) : of_dual ¬a = (of_dual a)ᶜ := rfl +@[simp] lemma to_dual_compl (a : α) : to_dual aᶜ = ¬to_dual a := rfl + +instance prod.heyting_algebra [heyting_algebra β] : heyting_algebra (α × β) := +{ himp_bot := λ a, prod.ext (himp_bot a.1) (himp_bot a.2), + ..prod.generalized_heyting_algebra, ..prod.bounded_order α β, ..prod.has_compl } + +instance pi.heyting_algebra {α : ι → Type*} [Π i, heyting_algebra (α i)] : + heyting_algebra (Π i, α i) := +by { pi_instance, exact λ a b c, forall_congr (λ i, le_himp_iff) } + +end heyting_algebra + +section coheyting_algebra +variables [coheyting_algebra α] {a b c : α} + +@[simp] lemma top_sdiff' (a : α) : ⊤ \ a = ¬a := coheyting_algebra.top_sdiff _ +@[simp] lemma sdiff_top (a : α) : a \ ⊤ = ⊥ := sdiff_eq_bot_iff.2 le_top + +lemma hnot_inf_distrib (a b : α) : ¬ (a ⊓ b) = ¬a ⊔ ¬b := +by simp_rw [←top_sdiff', sdiff_inf_distrib] + +lemma sdiff_le_hnot : a \ b ≤ ¬b := (sdiff_le_sdiff_right le_top).trans_eq $ top_sdiff' _ + +lemma sdiff_le_inf_hnot : a \ b ≤ a ⊓ ¬b := le_inf sdiff_le sdiff_le_hnot + +@[priority 100] -- See note [lower instance priority] +instance coheyting_algebra.to_distrib_lattice : distrib_lattice α := +{ le_sup_inf := λ a b c, by simp_rw [←sdiff_le_iff, le_inf_iff, sdiff_le_iff, ←le_inf_iff], + ..‹coheyting_algebra α› } + +@[simp] lemma hnot_sdiff (a : α) : ¬a \ a = ¬a := by rw [←top_sdiff', sdiff_sdiff, sup_idem] + +lemma hnot_sdiff_comm (a b : α) : ¬a \ b = ¬b \ a := by simp_rw [←top_sdiff', sdiff_right_comm] + +lemma hnot_le_iff_codisjoint_right : ¬a ≤ b ↔ codisjoint a b := +by rw [←top_sdiff', sdiff_le_iff, codisjoint] + +lemma hnot_le_iff_codisjoint_left : ¬a ≤ b ↔ codisjoint b a := +hnot_le_iff_codisjoint_right.trans codisjoint.comm + +lemma hnot_le_comm : ¬a ≤ b ↔ ¬b ≤ a := +by rw [hnot_le_iff_codisjoint_right, hnot_le_iff_codisjoint_left] + +alias hnot_le_iff_codisjoint_right ↔ _ codisjoint.hnot_le_right +alias hnot_le_iff_codisjoint_left ↔ _ codisjoint.hnot_le_left + +lemma codisjoint_hnot_right : codisjoint a (¬a) := sdiff_le_iff.1 (top_sdiff' _).le +lemma codisjoint_hnot_left : codisjoint (¬a) a := codisjoint_hnot_right.symm + +@[simp] lemma sup_hnot_self (a : α) : a ⊔ ¬a = ⊤ := codisjoint_hnot_right.eq_top +@[simp] lemma hnot_sup_self (a : α) : ¬a ⊔ a = ⊤ := codisjoint_hnot_left.eq_top + +@[simp] lemma hnot_bot : ¬(⊥ : α) = ⊤ := +eq_of_forall_ge_iff $ λ a, by rw [hnot_le_iff_codisjoint_left, codisjoint_bot, top_le_iff] + +@[simp] lemma hnot_top : ¬(⊤ : α) = ⊥ := by rw [←top_sdiff', sdiff_self] + +lemma hnot_hnot_le : ¬¬a ≤ a := codisjoint_hnot_right.hnot_le_left + +lemma hnot_anti : antitone (hnot : α → α) := λ a b h, hnot_le_comm.1 $ hnot_hnot_le.trans h + +lemma hnot_le_hnot (h : a ≤ b) : ¬b ≤ ¬a := hnot_anti h + +@[simp] lemma hnot_hnot_hnot (a : α) : ¬¬¬a = ¬a := hnot_hnot_le.antisymm $ hnot_anti hnot_hnot_le + +@[simp] lemma codisjoint_hnot_hnot_left_iff : codisjoint (¬¬a) b ↔ codisjoint a b := +by simp_rw [←hnot_le_iff_codisjoint_right, hnot_hnot_hnot] + +@[simp] lemma codisjoint_hnot_hnot_right_iff : codisjoint a (¬¬b) ↔ codisjoint a b := +by simp_rw [←hnot_le_iff_codisjoint_left, hnot_hnot_hnot] + +lemma le_hnot_inf_hnot : ¬ (a ⊔ b) ≤ ¬a ⊓ ¬b := +le_inf (hnot_anti le_sup_left) $ hnot_anti le_sup_right + +lemma hnot_hnot_sup_distrib (a b : α) : ¬¬(a ⊔ b) = ¬¬a ⊔ ¬¬b := +begin + refine ((hnot_inf_distrib _ _).ge.trans $ hnot_anti le_hnot_inf_hnot).antisymm' _, + rw [hnot_le_iff_codisjoint_left, codisjoint_assoc, codisjoint_hnot_hnot_left_iff, + codisjoint_left_comm, codisjoint_hnot_hnot_left_iff, ←codisjoint_assoc, sup_comm], + exact codisjoint_hnot_right, +end + +lemma hnot_hnot_sdiff_distrib (a b : α) : ¬¬(a \ b) = ¬¬a \ ¬¬b := +begin + refine le_antisymm _ _, + { refine hnot_le_comm.1 ((hnot_anti sdiff_le_inf_hnot).trans' _), + rw [hnot_inf_distrib, hnot_le_iff_codisjoint_right, codisjoint_left_comm, + ←hnot_le_iff_codisjoint_right], + exact le_sdiff_sup }, + { rw [sdiff_le_iff, ←hnot_hnot_sup_distrib], + exact hnot_anti (hnot_anti le_sup_sdiff) } +end + +instance : heyting_algebra αᵒᵈ := +{ compl := to_dual ∘ hnot ∘ of_dual, + himp := λ a b, to_dual (of_dual b \ of_dual a), + le_himp_iff := λ a b c, by { rw inf_comm, exact sdiff_le_iff }, + himp_bot := top_sdiff', + ..order_dual.lattice α, ..order_dual.bounded_order α } + +@[simp] lemma of_dual_compl (a : αᵒᵈ) : of_dual aᶜ = ¬of_dual a := rfl +@[simp] lemma of_dual_himp (a b : αᵒᵈ) : of_dual (a ⇨ b) = of_dual b \ of_dual a := rfl +@[simp] lemma to_dual_hnot (a : α) : to_dual ¬a = (to_dual a)ᶜ := rfl +@[simp] lemma to_dual_sdiff (a b : α) : to_dual (a \ b) = to_dual b ⇨ to_dual a := rfl + +instance prod.coheyting_algebra [coheyting_algebra β] : coheyting_algebra (α × β) := +{ sdiff_le_iff := λ a b c, and_congr sdiff_le_iff sdiff_le_iff, + top_sdiff := λ a, prod.ext (top_sdiff' a.1) (top_sdiff' a.2), + ..prod.lattice α β, ..prod.bounded_order α β, ..prod.has_sdiff, ..prod.has_hnot } + +instance pi.coheyting_algebra {α : ι → Type*} [Π i, coheyting_algebra (α i)] : + coheyting_algebra (Π i, α i) := +by { pi_instance, exact λ a b c, forall_congr (λ i, sdiff_le_iff) } + +end coheyting_algebra + +section biheyting_algebra +variables [biheyting_algebra α] {a : α} + +lemma compl_le_hnot : aᶜ ≤ ¬a := +(disjoint_compl_left : disjoint _ a).le_of_codisjoint codisjoint_hnot_right + +end biheyting_algebra + +/-- Propositions form a Heyting algebra with implication as Heyting implication and negation as +complement. -/ +instance Prop.heyting_algebra : heyting_algebra Prop := +{ himp := (→), + le_himp_iff := λ p q r, and_imp.symm, + himp_bot := λ p, rfl, + ..Prop.has_compl, ..Prop.distrib_lattice, ..Prop.bounded_order } + +@[simp] lemma himp_iff_imp (p q : Prop) : p ⇨ q ↔ p → q := iff.rfl +@[simp] lemma compl_iff_not (p : Prop) : pᶜ ↔ ¬ p := iff.rfl + +/-- A bounded linear order is a bi-Heyting algebra by setting +* `a ⇨ b = ⊤` if `a ≤ b` and `a ⇨ b = b` otherwise. +* `a \ b = ⊥` if `a ≤ b` and `a \ b = a` otherwise. -/ +@[reducible] -- See note [reducible non-instances] +def linear_order.to_biheyting_algebra [linear_order α] [bounded_order α] : biheyting_algebra α := +{ himp := λ a b, if a ≤ b then ⊤ else b, + compl := λ a, if a = ⊥ then ⊤ else ⊥, + le_himp_iff := λ a b c, begin + change _ ≤ ite _ _ _ ↔ _, + split_ifs, + { exact iff_of_true le_top (inf_le_of_right_le h) }, + { rw [inf_le_iff, or_iff_left h] } + end, + himp_bot := λ a, if_congr le_bot_iff rfl rfl, + sdiff := λ a b, if a ≤ b then ⊥ else a, + hnot := λ a, if a = ⊤ then ⊥ else ⊤, + sdiff_le_iff := λ a b c, begin + change ite _ _ _ ≤ _ ↔ _, + split_ifs, + { exact iff_of_true bot_le (le_sup_of_le_left h) }, + { rw [le_sup_iff, or_iff_right h] } + end, + top_sdiff := λ a, if_congr top_le_iff rfl rfl, + ..linear_order.to_lattice, ..‹bounded_order α› } + +section lift + +/-- Pullback a `generalized_heyting_algebra` along an injection. -/ +@[reducible] -- See note [reducible non-instances] +protected def function.injective.generalized_heyting_algebra [has_sup α] [has_inf α] [has_top α] + [has_himp α] [generalized_heyting_algebra β] (f : α → β) (hf : injective f) + (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) + (map_top : f ⊤ = ⊤) (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) : + generalized_heyting_algebra α := +{ le_top := λ a, by { change f _ ≤ _, rw map_top, exact le_top }, + le_himp_iff := λ a b c, by { change f _ ≤ _ ↔ f _ ≤ _, erw [map_himp, map_inf, le_himp_iff] }, + ..hf.lattice f map_sup map_inf, ..‹has_top α›, ..‹has_himp α› } + +/-- Pullback a `generalized_coheyting_algebra` along an injection. -/ +@[reducible] -- See note [reducible non-instances] +protected def function.injective.generalized_coheyting_algebra [has_sup α] [has_inf α] [has_bot α] + [has_sdiff α] [generalized_coheyting_algebra β] (f : α → β) (hf : injective f) + (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) + (map_bot : f ⊥ = ⊥) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : + generalized_coheyting_algebra α := +{ bot_le := λ a, by { change f _ ≤ _, rw map_bot, exact bot_le }, + sdiff_le_iff := λ a b c, by { change f _ ≤ _ ↔ f _ ≤ _, erw [map_sdiff, map_sup, sdiff_le_iff] }, + ..hf.lattice f map_sup map_inf, ..‹has_bot α›, ..‹has_sdiff α› } + +/-- Pullback a `heyting_algebra` along an injection. -/ +@[reducible] -- See note [reducible non-instances] +protected def function.injective.heyting_algebra [has_sup α] [has_inf α] [has_top α] [has_bot α] + [has_compl α] [has_himp α] [heyting_algebra β] (f : α → β) (hf : injective f) + (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) (map_compl : ∀ a, f aᶜ = (f a)ᶜ) + (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) : + heyting_algebra α := +{ bot_le := λ a, by { change f _ ≤ _, rw map_bot, exact bot_le }, + himp_bot := λ a, hf $ by erw [map_himp, map_compl, map_bot, himp_bot], + ..hf.generalized_heyting_algebra f map_sup map_inf map_top map_himp, + ..‹has_bot α›, ..‹has_compl α› } + +/-- Pullback a `coheyting_algebra` along an injection. -/ +@[reducible] -- See note [reducible non-instances] +protected def function.injective.coheyting_algebra [has_sup α] [has_inf α] [has_top α] [has_bot α] + [has_hnot α] [has_sdiff α] [coheyting_algebra β] (f : α → β) (hf : injective f) + (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) + (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) (map_hnot : ∀ a, f ¬a = ¬f a) + (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : + coheyting_algebra α := +{ le_top := λ a, by { change f _ ≤ _, rw map_top, exact le_top }, + top_sdiff := λ a, hf $ by erw [map_sdiff, map_hnot, map_top, top_sdiff'], + ..hf.generalized_coheyting_algebra f map_sup map_inf map_bot map_sdiff, + ..‹has_top α›, ..‹has_hnot α› } + +/-- Pullback a `biheyting_algebra` along an injection. -/ +@[reducible] -- See note [reducible non-instances] +protected def function.injective.biheyting_algebra [has_sup α] [has_inf α] [has_top α] [has_bot α] + [has_compl α] [has_hnot α] [has_himp α] [has_sdiff α] [biheyting_algebra β] (f : α → β) + (hf : injective f) (map_sup : ∀ a b, f (a ⊔ b) = f a ⊔ f b) + (map_inf : ∀ a b, f (a ⊓ b) = f a ⊓ f b) (map_top : f ⊤ = ⊤) (map_bot : f ⊥ = ⊥) + (map_compl : ∀ a, f aᶜ = (f a)ᶜ) (map_hnot : ∀ a, f ¬a = ¬f a) + (map_himp : ∀ a b, f (a ⇨ b) = f a ⇨ f b) (map_sdiff : ∀ a b, f (a \ b) = f a \ f b) : + biheyting_algebra α := +{ ..hf.heyting_algebra f map_sup map_inf map_top map_bot map_compl map_himp, + ..hf.coheyting_algebra f map_sup map_inf map_top map_bot map_hnot map_sdiff } + +end lift + +namespace punit +variables (a b : punit.{u+1}) + +instance : biheyting_algebra punit := +by refine_struct +{ top := star, + bot := star, + sup := λ _ _, star, + inf := λ _ _, star, + compl := λ _, star, + sdiff := λ _ _, star, + hnot := λ _, star, + himp := λ _ _, star, ..punit.linear_order }; + intros; trivial <|> exact subsingleton.elim _ _ + +@[simp] lemma top_eq : (⊤ : punit) = star := rfl +@[simp] lemma bot_eq : (⊥ : punit) = star := rfl +@[simp] lemma sup_eq : a ⊔ b = star := rfl +@[simp] lemma inf_eq : a ⊓ b = star := rfl +@[simp] lemma compl_eq : aᶜ = star := rfl +@[simp] lemma sdiff_eq : a \ b = star := rfl +@[simp, nolint simp_nf] lemma hnot_eq : ¬a = star := rfl -- eligible for `dsimp` +@[simp] lemma himp_eq : a ⇨ b = star := rfl + +end punit From 0be7214f67198d1035c34ec5435679a27b1cde77 Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Wed, 24 Aug 2022 18:34:12 +0000 Subject: [PATCH 026/828] feat(category_theory/currying): bifunctor version of whiskering_right (#15504) This cannot go in `category_theory/whiskering` because this file is imported indirectly by `category_theory/functor/currying`. Hopefully it is ok to put it in `category_theory/functor/currying`. - [x] depends on: #15445 Co-authored-by: antoinelab01 <66086247+antoinelab01@users.noreply.github.com> --- src/category_theory/functor/currying.lean | 19 +++++++++++++++---- 1 file changed, 15 insertions(+), 4 deletions(-) diff --git a/src/category_theory/functor/currying.lean b/src/category_theory/functor/currying.lean index b62d53fbcbad9..26a95ef824fa7 100644 --- a/src/category_theory/functor/currying.lean +++ b/src/category_theory/functor/currying.lean @@ -15,11 +15,12 @@ and verify that they provide an equivalence of categories -/ namespace category_theory -universes v₁ v₂ v₃ u₁ u₂ u₃ +universes v₁ v₂ v₃ v₄ u₁ u₂ u₃ u₄ -variables {C : Type u₁} [category.{v₁} C] - {D : Type u₂} [category.{v₂} D] - {E : Type u₃} [category.{v₃} E] +variables {B : Type u₁} [category.{v₁} B] + {C : Type u₂} [category.{v₂} C] + {D : Type u₃} [category.{v₃} D] + {E : Type u₄} [category.{v₄} E] /-- The uncurrying functor, taking a functor `C ⥤ (D ⥤ E)` and producing a functor `(C × D) ⥤ E`. @@ -99,4 +100,14 @@ swapping the factors followed by the uncurrying of `F`. -/ def uncurry_obj_flip (F : C ⥤ D ⥤ E) : uncurry.obj F.flip ≅ prod.swap _ _ ⋙ uncurry.obj F := nat_iso.of_components (λ p, iso.refl _) (by tidy) +variables (B C D E) + +/-- +A version of `category_theory.whiskering_right` for bifunctors, obtained by uncurrying, +applying `whiskering_right` and currying back +-/ +@[simps] def whiskering_right₂ : (C ⥤ D ⥤ E) ⥤ ((B ⥤ C) ⥤ (B ⥤ D) ⥤ (B ⥤ E)) := +uncurry ⋙ (whiskering_right _ _ _) ⋙ +((whiskering_left _ _ _).obj (prod_functor_to_functor_prod _ _ _)) ⋙ curry + end category_theory From fc5b2d2ed3099815903ad8f4419a84623634970c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 24 Aug 2022 18:34:13 +0000 Subject: [PATCH 027/828] feat(algebraic_topology/dold_kan): the normalized Moore complex is a direct factor of the alternating face map complex (#16071) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this PR, we show that when the category `A` is abelian, there is an isomorphism `N₁_iso_to_karoubi_normalized A` between the functor `N₁ : simplicial_object A ⥤ karoubi (chain_complex A ℕ)` defined in `functor_n.lean` and the composition of `normalized_Moore_complex A` with the inclusion `chain_complex A ℕ ⥤ karoubi (chain_complex A ℕ)`. (In particular, the normalized Moore complex is a direct factor of the alternating face map complex.) Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- src/algebraic_topology/Moore_complex.lean | 7 + .../dold_kan/normalized.lean | 145 ++++++++++++++++++ src/algebraic_topology/dold_kan/p_infty.lean | 3 + 3 files changed, 155 insertions(+) create mode 100644 src/algebraic_topology/dold_kan/normalized.lean diff --git a/src/algebraic_topology/Moore_complex.lean b/src/algebraic_topology/Moore_complex.lean index 911de401dca0c..d9760eb05bb79 100644 --- a/src/algebraic_topology/Moore_complex.lean +++ b/src/algebraic_topology/Moore_complex.lean @@ -157,4 +157,11 @@ def normalized_Moore_complex : simplicial_object C ⥤ chain_complex C ℕ := map_id' := λ X, by { ext n, cases n; { dsimp, simp, }, }, map_comp' := λ X Y Z f g, by { ext n, cases n; simp, }, } +variable {C} + +@[simp] +lemma normalized_Moore_complex_obj_d (X : simplicial_object C) (n : ℕ) : + ((normalized_Moore_complex C).obj X).d (n+1) n = normalized_Moore_complex.obj_d X n := +by apply chain_complex.of_d + end algebraic_topology diff --git a/src/algebraic_topology/dold_kan/normalized.lean b/src/algebraic_topology/dold_kan/normalized.lean new file mode 100644 index 0000000000000..e6e55b3f70a4d --- /dev/null +++ b/src/algebraic_topology/dold_kan/normalized.lean @@ -0,0 +1,145 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.dold_kan.functor_n + +/-! + +# Comparison with the normalized Moore complex functor + +TODO (@joelriou) continue adding the various files referenced below + +In this file, we show that when the category `A` is abelian, +there is an isomorphism `N₁_iso_normalized_Moore_complex_comp_to_karoubi` between +the functor `N₁ : simplicial_object A ⥤ karoubi (chain_complex A ℕ)` +defined in `functor_n.lean` and the composition of +`normalized_Moore_complex A` with the inclusion +`chain_complex A ℕ ⥤ karoubi (chain_complex A ℕ)`. + +This isomorphism shall be used in `equivalence.lean` in order to obtain +the Dold-Kan equivalence +`category_theory.abelian.dold_kan.equivalence : simplicial_object A ≌ chain_complex A ℕ` +with a functor (definitionally) equal to `normalized_Moore_complex A`. + +-/ + +open category_theory category_theory.category category_theory.limits + category_theory.subobject category_theory.idempotents +open_locale dold_kan + +noncomputable theory + +namespace algebraic_topology + +namespace dold_kan + +universe v + +variables {A : Type*} [category A] [abelian A] {X : simplicial_object A} + +lemma higher_faces_vanish.inclusion_of_Moore_complex_map (n : ℕ) : + higher_faces_vanish (n+1) ((inclusion_of_Moore_complex_map X).f (n+1)) := λ j hj, +begin + dsimp [inclusion_of_Moore_complex_map], + rw [← factor_thru_arrow _ _ (finset_inf_arrow_factors finset.univ + _ j (by simp only [finset.mem_univ])), assoc, kernel_subobject_arrow_comp, comp_zero], +end + +lemma factors_normalized_Moore_complex_P_infty (n : ℕ) : + subobject.factors (normalized_Moore_complex.obj_X X n) (P_infty.f n) := +begin + cases n, + { apply top_factors, }, + { rw [P_infty_f, normalized_Moore_complex.obj_X, finset_inf_factors], + intros i hi, + apply kernel_subobject_factors, + exact (higher_faces_vanish.of_P (n+1) n) i (le_add_self), } +end + +/-- P_infty factors through the normalized Moore complex -/ +@[simps] +def P_infty_to_normalized_Moore_complex (X : simplicial_object A) : K[X] ⟶ N[X] := +chain_complex.of_hom _ _ _ _ _ _ + (λ n, factor_thru _ _ (factors_normalized_Moore_complex_P_infty n)) + (λ n, begin + rw [← cancel_mono (normalized_Moore_complex.obj_X X n).arrow, assoc, assoc, + factor_thru_arrow, ← inclusion_of_Moore_complex_map_f, + ← normalized_Moore_complex_obj_d, ← (inclusion_of_Moore_complex_map X).comm' (n+1) n rfl, + inclusion_of_Moore_complex_map_f, factor_thru_arrow_assoc, + ← alternating_face_map_complex_obj_d], + exact P_infty.comm' (n+1) n rfl, + end) + +@[simp, reassoc] +lemma P_infty_to_normalized_Moore_complex_comp_inclusion_of_Moore_complex_map + (X : simplicial_object A) : + P_infty_to_normalized_Moore_complex X ≫ inclusion_of_Moore_complex_map X = P_infty := by tidy + +@[simp, reassoc] +lemma P_infty_to_normalized_Moore_complex_naturality {X Y : simplicial_object A} (f : X ⟶ Y) : + alternating_face_map_complex.map f ≫ P_infty_to_normalized_Moore_complex Y = + P_infty_to_normalized_Moore_complex X ≫ normalized_Moore_complex.map f := by tidy + +@[simp, reassoc] +lemma P_infty_comp_P_infty_to_normalized_Moore_complex (X : simplicial_object A) : + P_infty ≫ P_infty_to_normalized_Moore_complex X = P_infty_to_normalized_Moore_complex X := +by tidy + +@[simp, reassoc] +lemma inclusion_of_Moore_complex_map_comp_P_infty (X : simplicial_object A) : + inclusion_of_Moore_complex_map X ≫ P_infty = inclusion_of_Moore_complex_map X := +begin + ext n, + cases n, + { dsimp, simp only [comp_id], }, + { exact (higher_faces_vanish.inclusion_of_Moore_complex_map n).comp_P_eq_self, }, +end + +instance : mono (inclusion_of_Moore_complex_map X) := +⟨λ Y f₁ f₂ hf, by { ext n, exact homological_complex.congr_hom hf n, }⟩ + +/-- `inclusion_of_Moore_complex_map X` is a split mono. -/ +def split_mono_inclusion_of_Moore_complex_map (X : simplicial_object A) : + split_mono (inclusion_of_Moore_complex_map X) := +{ retraction := P_infty_to_normalized_Moore_complex X, + id' := by simp only [← cancel_mono (inclusion_of_Moore_complex_map X), assoc, id_comp, + P_infty_to_normalized_Moore_complex_comp_inclusion_of_Moore_complex_map, + inclusion_of_Moore_complex_map_comp_P_infty], } + +variable (A) + +/-- When the category `A` is abelian, +the functor `N₁ : simplicial_object A ⥤ karoubi (chain_complex A ℕ)` defined +using `P_infty` identifies to the composition of the normalized Moore complex functor +and the inclusion in the Karoubi envelope. -/ +def N₁_iso_normalized_Moore_complex_comp_to_karoubi : + N₁ ≅ (normalized_Moore_complex A ⋙ to_karoubi _) := +{ hom := + { app := λ X, + { f := P_infty_to_normalized_Moore_complex X, + comm := by tidy, }, }, + inv := + { app := λ X, + { f := inclusion_of_Moore_complex_map X, + comm := by tidy, }, }, + hom_inv_id' := begin + ext X : 3, + simp only [P_infty_to_normalized_Moore_complex_comp_inclusion_of_Moore_complex_map, + nat_trans.comp_app, karoubi.comp, N₁_obj_p, nat_trans.id_app, karoubi.id_eq], + end, + inv_hom_id' := begin + ext X : 3, + simp only [← cancel_mono (inclusion_of_Moore_complex_map X), + nat_trans.comp_app, karoubi.comp, assoc, nat_trans.id_app, karoubi.id_eq, + P_infty_to_normalized_Moore_complex_comp_inclusion_of_Moore_complex_map, + inclusion_of_Moore_complex_map_comp_P_infty], + dsimp only [functor.comp_obj, to_karoubi], + rw id_comp, + end } + +end dold_kan + +end algebraic_topology diff --git a/src/algebraic_topology/dold_kan/p_infty.lean b/src/algebraic_topology/dold_kan/p_infty.lean index d63259f68d3e9..bfdcca8a76d36 100644 --- a/src/algebraic_topology/dold_kan/p_infty.lean +++ b/src/algebraic_topology/dold_kan/p_infty.lean @@ -61,6 +61,9 @@ def P_infty : K[X] ⟶ K[X] := chain_complex.of_hom _ _ _ _ _ _ (λ n, by simpa only [← P_is_eventually_constant (show n ≤ n, by refl), alternating_face_map_complex.obj_d_eq] using (P (n+1)).comm (n+1) n) +@[simp] +lemma P_infty_f_0 : (P_infty.f 0 : X _[0] ⟶ X _[0]) = 𝟙 _ := rfl + lemma P_infty_f (n : ℕ) : (P_infty.f n : X _[n] ⟶ X _[n] ) = (P n).f n := rfl @[simp, reassoc] From 81afa2cb52ec7efd3ac6d9c88d7f1b5d9d5b7d7f Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 24 Aug 2022 18:34:15 +0000 Subject: [PATCH 028/828] feat(data/polynomial/module): Some api for `polynomial_module`. (#16126) --- src/data/polynomial/module.lean | 121 +++++++++++++++++++++++++++++++- 1 file changed, 118 insertions(+), 3 deletions(-) diff --git a/src/data/polynomial/module.lean b/src/data/polynomial/module.lean index c2d4f4f2744f3..a38b84fdf0c30 100644 --- a/src/data/polynomial/module.lean +++ b/src/data/polynomial/module.lean @@ -93,9 +93,21 @@ finsupp.induction_linear f h0 hadd hsingle instance polynomial_module : module R[X] (polynomial_module R M) := module_polynomial_of_endo (finsupp.lmap_domain _ _ nat.succ) -instance (M : Type u) [add_comm_group M] [module R M] : - is_scalar_tower R R[X] (polynomial_module R M) := -module_polynomial_of_endo.is_scalar_tower _ +instance (M : Type u) [add_comm_group M] [module R M] [module S M] [is_scalar_tower S R M] : + is_scalar_tower S R (polynomial_module R M) := +finsupp.is_scalar_tower _ _ + +instance is_scalar_tower' (M : Type u) [add_comm_group M] [module R M] [module S M] + [is_scalar_tower S R M] : + is_scalar_tower S R[X] (polynomial_module R M) := +begin + haveI : is_scalar_tower R R[X] (polynomial_module R M) := + module_polynomial_of_endo.is_scalar_tower _, + constructor, + intros x y z, + rw [← @is_scalar_tower.algebra_map_smul S R, ← @is_scalar_tower.algebra_map_smul S R, + smul_assoc], +end @[simp] lemma monomial_smul_single (i : ℕ) (r : R) (j : ℕ) (m : M) : @@ -177,4 +189,107 @@ def equiv_polynomial {S : Type*} [comm_ring S] [algebra R S] : polynomial_module R S ≃ₗ[R] S[X] := { map_smul' := λ r x, rfl, ..(polynomial.to_finsupp_iso S).symm } +variables (R' : Type*) {M' : Type*} [comm_ring R'] [add_comm_group M'] [module R' M'] +variables [algebra R R'] [module R M'] [is_scalar_tower R R' M'] + +/-- The image of a polynomial under a linear map. -/ +noncomputable +def map (f : M →ₗ[R] M') : polynomial_module R M →ₗ[R] polynomial_module R' M' := +finsupp.map_range.linear_map f + +@[simp] +lemma map_single (f : M →ₗ[R] M') (i : ℕ) (m : M) : + map R' f (single R i m) = single R' i (f m) := +finsupp.map_range_single + +lemma map_smul (f : M →ₗ[R] M') (p : R[X]) (q : polynomial_module R M) : + map R' f (p • q) = p.map (algebra_map R R') • map R' f q := +begin + apply induction_linear q, + { rw [smul_zero, map_zero, smul_zero] }, + { intros f g e₁ e₂, rw [smul_add, map_add, e₁, e₂, map_add, smul_add] }, + intros i m, + apply polynomial.induction_on' p, + { intros p q e₁ e₂, rw [add_smul, map_add, e₁, e₂, polynomial.map_add, add_smul] }, + { intros j s, + rw [monomial_smul_single, map_single, polynomial.map_monomial, map_single, + monomial_smul_single, f.map_smul, algebra_map_smul] } +end + +/-- Evaulate a polynomial `p : polynomial_module R M` at `r : R`. -/ +@[simps (lemmas_only)] +def eval (r : R) : polynomial_module R M →ₗ[R] M := +{ to_fun := λ p, p.sum (λ i m, r ^ i • m), + map_add' := λ x y, finsupp.sum_add_index' (λ _, smul_zero _) (λ _ _ _, smul_add _ _ _), + map_smul' := λ s m, begin + refine (finsupp.sum_smul_index' _).trans _, + { exact λ i, smul_zero _ }, + { simp_rw [← smul_comm s, ← finsupp.smul_sum], refl } + end } + +@[simp] +lemma eval_single (r : R) (i : ℕ) (m : M) : eval r (single R i m) = r ^ i • m := +finsupp.sum_single_index (smul_zero _) + +lemma eval_smul (p : R[X]) (q : polynomial_module R M) (r : R) : + eval r (p • q) = p.eval r • eval r q := +begin + apply induction_linear q, + { rw [smul_zero, map_zero, smul_zero] }, + { intros f g e₁ e₂, rw [smul_add, map_add, e₁, e₂, map_add, smul_add] }, + intros i m, + apply polynomial.induction_on' p, + { intros p q e₁ e₂, rw [add_smul, map_add, polynomial.eval_add, e₁, e₂, add_smul] }, + { intros j s, + rw [monomial_smul_single, eval_single, polynomial.eval_monomial, eval_single, + smul_comm, ← smul_smul, pow_add, mul_smul] } +end + +@[simp] +lemma eval_map (f : M →ₗ[R] M') (q : polynomial_module R M) (r : R) : + eval (algebra_map R R' r) (map R' f q) = f (eval r q) := +begin + apply induction_linear q, + { simp_rw map_zero }, + { intros f g e₁ e₂, simp_rw [map_add, e₁, e₂] }, + { intros i m, + rw [map_single, eval_single, eval_single, f.map_smul, ← map_pow, algebra_map_smul] } +end + +@[simp] +lemma eval_map' (f : M →ₗ[R] M) (q : polynomial_module R M) (r : R) : + eval r (map R f q) = f (eval r q) := +eval_map R f q r + +/-- `comp p q` is the composition of `p : R[X]` and `q : M[X]` as `q(p(x))`. -/ +@[simps] noncomputable +def comp (p : R[X]) : polynomial_module R M →ₗ[R] polynomial_module R M := +((eval p).restrict_scalars R).comp (map R[X] (lsingle R 0)) + +lemma comp_single (p : R[X]) (i : ℕ) (m : M) : comp p (single R i m) = p ^ i • single R 0 m := +begin + rw comp_apply, + erw [map_single, eval_single], + refl +end + +lemma comp_eval (p : R[X]) (q : polynomial_module R M) (r : R) : + eval r (comp p q) = eval (p.eval r) q := +begin + rw ← linear_map.comp_apply, + apply induction_linear q, + { rw [map_zero, map_zero] }, + { intros _ _ e₁ e₂, rw [map_add, map_add, e₁, e₂] }, + { intros i m, + rw [linear_map.comp_apply, comp_single, eval_single, eval_smul, eval_single, pow_zero, + one_smul, polynomial.eval_pow] } +end + +lemma comp_smul (p p' : R[X]) (q : polynomial_module R M) : + comp p (p' • q) = p'.comp p • comp p q := +begin + rw [comp_apply, map_smul, eval_smul, polynomial.comp, polynomial.eval_map, comp_apply], + refl +end + end polynomial_module From f5afe205c23d7e9f065a04390cc337dfdb7e6e65 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 24 Aug 2022 18:34:16 +0000 Subject: [PATCH 029/828] split(analysis/normed/group/seminorm): Split off `analysis.seminorm` (#16152) Move `group_seminorm` and `add_group_seminorm` to a new file `analysis.normed.group.seminorm`. Move `norm_add_group_seminorm` to `analysis.normed.group.basic`. Remove the `nonneg` field from `add_group_seminorm` and `group_seminorm` because it is redundant. --- src/analysis/normed/group/basic.lean | 10 + src/analysis/normed/group/seminorm.lean | 285 ++++++++++++++++++++++ src/analysis/seminorm.lean | 309 +----------------------- 3 files changed, 298 insertions(+), 306 deletions(-) create mode 100644 src/analysis/normed/group/seminorm.lean diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 5458632c8ed25..37e31726e714b 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Johannes Hölzl -/ import algebra.module.ulift +import analysis.normed.group.seminorm import order.liminf_limsup import topology.algebra.uniform_group import topology.metric_space.algebra @@ -385,6 +386,15 @@ ne_zero_of_norm_ne_zero $ by rwa norm_eq_of_mem_sphere x lemma ne_zero_of_mem_unit_sphere (x : sphere (0:E) 1) : (x:E) ≠ 0 := ne_zero_of_mem_sphere one_ne_zero _ +variables (E) + +/-- The norm of a seminormed group as an additive group seminorm. -/ +def norm_add_group_seminorm : add_group_seminorm E := ⟨norm, norm_zero, norm_add_le, norm_neg⟩ + +@[simp] lemma coe_norm_add_group_seminorm : ⇑(norm_add_group_seminorm E) = norm := rfl + +variables {E} + namespace isometric -- TODO This material is superseded by similar constructions such as -- `affine_isometry_equiv.const_vadd`; deduplicate diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean new file mode 100644 index 0000000000000..3442f5374e264 --- /dev/null +++ b/src/analysis/normed/group/seminorm.lean @@ -0,0 +1,285 @@ +/- +Copyright (c) 2022 María Inés de Frutos-Fernández, Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: María Inés de Frutos-Fernández, Yaël Dillies +-/ +import data.real.nnreal + +/-! +# Group seminorms + +This file defines seminorms in a group. A group seminorm is a function to the reals which is +positive-semidefinite and subadditive. + +## Main declarations + +* `add_group_seminorm`: A function `f` from an additive group `G` to the reals that preserves zero, + takes nonnegative values, is subadditive and such that `f (-x) = f x` for all `x`. +* `group_seminorm`: A function `f` from a group `G` to the reals that sends one to zero, takes + nonnegative values, is submultiplicative and such that `f x⁻¹ = f x` for all `x`. + +## References + +* [H. H. Schaefer, *Topological Vector Spaces*][schaefer1966] + +## Tags + +seminorm +-/ + +set_option old_structure_cmd true + +open set +open_locale nnreal + +variables {ι R R' E F G : Type*} + +/-- A seminorm on an additive group `G` is a function `f : G → ℝ` that preserves zero, is +subadditive and such that `f (-x) = f x` for all `x`. -/ +structure add_group_seminorm (G : Type*) [add_group G] extends zero_hom G ℝ := +(add_le' : ∀ r s, to_fun (r + s) ≤ to_fun r + to_fun s) +(neg' : ∀ r, to_fun (-r) = to_fun r) + +/-- A seminorm on a group `G` is a function `f : G → ℝ` that sends one to zero, is submultiplicative +and such that `f x⁻¹ = f x` for all `x`. -/ +@[to_additive] +structure group_seminorm (G : Type*) [group G] := +(to_fun : G → ℝ) +(map_one' : to_fun 1 = 0) +(mul_le' : ∀ x y, to_fun (x * y) ≤ to_fun x + to_fun y) +(inv' : ∀ x, to_fun x⁻¹ = to_fun x) + +attribute [nolint doc_blame] add_group_seminorm.to_zero_hom + +namespace group_seminorm +section group +variables [group E] [group F] [group G] + +@[to_additive] instance fun_like : fun_like (group_seminorm E) E (λ _, ℝ) := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ +@[to_additive "Helper instance for when there's too many metavariables to apply +`fun_like.has_coe_to_fun`. "] +instance : has_coe_to_fun (group_seminorm E) (λ _, E → ℝ) := ⟨to_fun⟩ + +@[ext, to_additive] lemma ext {p q : group_seminorm E} : (∀ x, (p : E → ℝ) x = q x) → p = q := +fun_like.ext p q + +variables (p q : group_seminorm E) (x y : E) (r : ℝ) + +@[simp, to_additive] protected lemma map_one : p 1 = 0 := p.map_one' +@[to_additive] protected lemma mul_le : p (x * y) ≤ p x + p y := p.mul_le' _ _ +@[simp, to_additive] protected lemma inv : p x⁻¹ = p x := p.inv' _ + +@[to_additive] protected lemma div_le : p (x / y) ≤ p x + p y := +by { rw [div_eq_mul_inv, ←p.inv y], exact p.mul_le _ _ } + +@[to_additive] protected lemma nonneg : 0 ≤ p x := +nonneg_of_mul_nonneg_right (by { rw [two_mul, ←p.map_one, ←div_self' x], apply p.div_le }) two_pos + +@[to_additive] lemma div_rev : p (x / y) = p (y / x) := by rw [←inv_div, p.inv] + +@[to_additive] instance : partial_order (group_seminorm E) := +partial_order.lift _ fun_like.coe_injective + +@[to_additive] lemma le_def : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl +@[to_additive] lemma lt_def : p < q ↔ (p : E → ℝ) < q := iff.rfl + +variables {p q} + +@[simp, to_additive] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl +@[simp, to_additive] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl + +variables (p q) (f : F →* E) + +@[to_additive] instance : has_zero (group_seminorm E) := +⟨{ to_fun := 0, + map_one' := pi.zero_apply _, + mul_le' := λ _ _, (zero_add _).ge, + inv' := λ x, rfl}⟩ + +@[simp, to_additive] lemma coe_zero : ⇑(0 : group_seminorm E) = 0 := rfl +@[simp, to_additive] lemma zero_apply (x : E) : (0 : group_seminorm E) x = 0 := rfl + +@[to_additive] instance : inhabited (group_seminorm E) := ⟨0⟩ + +@[to_additive] instance : has_add (group_seminorm E) := +⟨λ p q, + { to_fun := λ x, p x + q x, + map_one' := by rw [p.map_one, q.map_one, zero_add], + mul_le' := λ _ _, (add_le_add (p.mul_le _ _) $ q.mul_le _ _).trans_eq $ + add_add_add_comm _ _ _ _, + inv' := λ x, by rw [p.inv, q.inv] }⟩ + +@[simp, to_additive] lemma coe_add : ⇑(p + q) = p + q := rfl +@[simp, to_additive] lemma add_apply (x : E) : (p + q) x = p x + q x := rfl + +-- TODO: define `has_Sup` too, from the skeleton at +-- https://github.com/leanprover-community/mathlib/pull/11329#issuecomment-1008915345 +@[to_additive] noncomputable instance : has_sup (group_seminorm E) := +⟨λ p q, + { to_fun := p ⊔ q, + map_one' := by rw [pi.sup_apply, ←p.map_one, sup_eq_left, p.map_one, q.map_one], + mul_le' := λ x y, sup_le + ((p.mul_le x y).trans $ add_le_add le_sup_left le_sup_left) + ((q.mul_le x y).trans $ add_le_add le_sup_right le_sup_right), + inv' := λ x, by rw [pi.sup_apply, pi.sup_apply, p.inv, q.inv] }⟩ + +@[simp, to_additive] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl +@[simp, to_additive] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl + +@[to_additive] noncomputable instance : semilattice_sup (group_seminorm E) := +fun_like.coe_injective.semilattice_sup _ coe_sup + +/-- Composition of a group seminorm with a monoid homomorphism as a group seminorm. -/ +@[to_additive "Composition of an additive group seminorm with an additive monoid homomorphism as an +additive group seminorm."] +def comp (p : group_seminorm E) (f : F →* E) : group_seminorm F := +{ to_fun := λ x, p (f x), + map_one' := by rw [f.map_one, p.map_one], + mul_le' := λ _ _, (congr_arg p $ f.map_mul _ _).trans_le $ p.mul_le _ _, + inv' := λ x, by rw [map_inv, p.inv] } + +@[simp, to_additive] lemma coe_comp : ⇑(p.comp f) = p ∘ f := rfl +@[simp, to_additive] lemma comp_apply (x : F) : (p.comp f) x = p (f x) := rfl +@[simp, to_additive] lemma comp_id : p.comp (monoid_hom.id _) = p := ext $ λ _, rfl +@[simp, to_additive] lemma comp_zero : p.comp (1 : F →* E) = 0 := ext $ λ _, p.map_one +@[simp, to_additive] lemma zero_comp : (0 : group_seminorm E).comp f = 0 := ext $ λ _, rfl + +@[to_additive] lemma comp_assoc (g : F →* E) (f : G →* F) : p.comp (g.comp f) = (p.comp g).comp f := +ext $ λ _, rfl + +@[to_additive] lemma add_comp (f : F →* E) : (p + q).comp f = p.comp f + q.comp f := ext $ λ _, rfl + +variables {p q} + +@[to_additive] lemma comp_mono (hp : p ≤ q) : p.comp f ≤ q.comp f := λ _, hp _ + +end group + +section comm_group +variables [comm_group E] [comm_group F] (p q : group_seminorm E) (x y : E) + +/-- The direct path from `1` to `y` is shorter than the path with `x` "inserted" in between. -/ +@[to_additive "The direct path from `0` to `y` is shorter than the path with `x` \"inserted\" in +between."] +lemma le_insert : p y ≤ p x + p (x / y) := (p.div_le _ _).trans_eq' $ by rw div_div_cancel + +/-- The direct path from `1` to `x` is shorter than the path with `y` "inserted" in between. -/ +@[to_additive "The direct path from `0` to `x` is shorter than the path with `y` \"inserted\" in +between."] +lemma le_insert' : p x ≤ p y + p (x / y) := by { rw div_rev, exact le_insert _ _ _ } + +@[to_additive] lemma comp_mul_le (f g : F →* E) : p.comp (f * g) ≤ p.comp f + p.comp g := +λ _, p.mul_le _ _ + +@[to_additive] private lemma mul_bdd_below_range_add {p q : group_seminorm E} {x : E} : + bdd_below (range $ λ y, p y + q (x / y)) := +⟨0, by { rintro _ ⟨x, rfl⟩, exact add_nonneg (p.nonneg _) (q.nonneg _) }⟩ + +@[to_additive] noncomputable instance : has_inf (group_seminorm E) := +⟨λ p q, + { to_fun := λ x, ⨅ y, p y + q (x / y), + map_one' := cinfi_eq_of_forall_ge_of_forall_gt_exists_lt + (λ x, add_nonneg (p.nonneg _) (q.nonneg _)) + (λ r hr, ⟨1, by rwa [div_one, p.map_one, q.map_one, add_zero]⟩), + mul_le' := λ x y, le_cinfi_add_cinfi $ λ u v, begin + refine cinfi_le_of_le mul_bdd_below_range_add (u * v) _, + rw [mul_div_mul_comm, add_add_add_comm], + exact add_le_add (p.mul_le _ _) (q.mul_le _ _), + end, + inv' := λ x, (inv_surjective.infi_comp _).symm.trans $ by simp_rw [p.inv, ←inv_div', q.inv] }⟩ + +@[simp, to_additive] lemma inf_apply : (p ⊓ q) x = ⨅ y, p y + q (x / y) := rfl + +@[to_additive] noncomputable instance : lattice (group_seminorm E) := +{ inf := (⊓), + inf_le_left := λ p q x, cinfi_le_of_le mul_bdd_below_range_add x $ + by rw [div_self', q.map_one, add_zero], + inf_le_right := λ p q x, cinfi_le_of_le mul_bdd_below_range_add (1 : E) $ + by simp only [div_one, p.map_one, zero_add], + le_inf := λ a b c hb hc x, le_cinfi $ λ u, (a.le_insert' _ _).trans $ add_le_add (hb _) (hc _), + ..group_seminorm.semilattice_sup } + +end comm_group +end group_seminorm + +namespace add_group_seminorm +variables [add_group E] (p : add_group_seminorm E) (x y : E) (r : ℝ) + +instance zero_hom_class : zero_hom_class (add_group_seminorm E) E ℝ := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_zero := λ f, f.map_zero' } + +/- TODO: All the following ought to be automated using `to_additive`. The problem is that it doesn't +see that `has_smul R ℝ` should be fixed because `ℝ` is fixed. -/ + +variables [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] + +/-- Any action on `ℝ` which factors through `ℝ≥0` applies to an `add_group_seminorm`. -/ +instance : has_smul R (add_group_seminorm E) := +⟨λ r p, + { to_fun := λ x, r • p x, + map_zero' := by simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul, + p.map_zero, mul_zero], + add_le' := λ _ _, begin + simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul], + exact (mul_le_mul_of_nonneg_left (p.add_le _ _) (nnreal.coe_nonneg _)).trans_eq + (mul_add _ _ _), + end, + neg' := λ x, by rw p.neg }⟩ + +@[simp] lemma coe_smul (r : R) (p : add_group_seminorm E) : ⇑(r • p) = r • p := rfl +@[simp] lemma smul_apply (r : R) (p : add_group_seminorm E) (x : E) : (r • p) x = r • p x := rfl + +instance [has_smul R' ℝ] [has_smul R' ℝ≥0] [is_scalar_tower R' ℝ≥0 ℝ] + [has_smul R R'] [is_scalar_tower R R' ℝ] : + is_scalar_tower R R' (add_group_seminorm E) := +⟨λ r a p, ext $ λ x, smul_assoc r a (p x)⟩ + +lemma smul_sup (r : R) (p q : add_group_seminorm E) : r • (p ⊔ q) = r • p ⊔ r • q := +have real.smul_max : ∀ x y : ℝ, r • max x y = max (r • x) (r • y), +from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ℝ≥0 r (_ : ℝ)] + using mul_max_of_nonneg x y (r • 1 : ℝ≥0).prop, +ext $ λ x, real.smul_max _ _ + +end add_group_seminorm + +namespace group_seminorm +variables [group E] [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] + +/-- Any action on `ℝ` which factors through `ℝ≥0` applies to an `add_group_seminorm`. -/ +@[to_additive add_group_seminorm.has_smul] instance : has_smul R (group_seminorm E) := +⟨λ r p, + { to_fun := λ x, r • p x, + map_one' := by simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul, + p.map_one, mul_zero], + mul_le' := λ _ _, begin + simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul], + exact (mul_le_mul_of_nonneg_left (p.mul_le _ _) $ nnreal.coe_nonneg _).trans_eq + (mul_add _ _ _), + end, + inv' := λ x, by rw p.inv }⟩ + +@[to_additive add_group_seminorm.is_scalar_tower] +instance [has_smul R' ℝ] [has_smul R' ℝ≥0] [is_scalar_tower R' ℝ≥0 ℝ] [has_smul R R'] + [is_scalar_tower R R' ℝ] : is_scalar_tower R R' (group_seminorm E) := +⟨λ r a p, ext $ λ x, smul_assoc r a $ p x⟩ + +@[simp, to_additive add_group_seminorm.coe_smul] +lemma coe_smul (r : R) (p : group_seminorm E) : ⇑(r • p) = r • p := rfl + +@[simp, to_additive add_group_seminorm.smul_apply] +lemma smul_apply (r : R) (p : group_seminorm E) (x : E) : (r • p) x = r • p x := rfl + +@[to_additive add_group_seminorm.smul_sup] +lemma smul_sup (r : R) (p q : group_seminorm E) : r • (p ⊔ q) = r • p ⊔ r • q := +have real.smul_max : ∀ x y : ℝ, r • max x y = max (r • x) (r • y), +from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ℝ≥0 r (_ : ℝ)] + using mul_max_of_nonneg x y (r • 1 : ℝ≥0).prop, +ext $ λ x, real.smul_max _ _ + +end group_seminorm diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 8c19d9266c509..4d7f039414e7e 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -3,7 +3,6 @@ Copyright (c) 2019 Jean Lo. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean Lo, Yaël Dillies, Moritz Doll -/ -import analysis.locally_convex.basic import data.real.pointwise import data.real.sqrt import topology.algebra.filter_basis @@ -20,11 +19,6 @@ convex if and only if its topology is induced by a family of seminorms. ## Main declarations -* `add_group_seminorm`: A function `f` from an additive group `G` to the reals that preserves zero, - takes nonnegative values, is subadditive and such that `f (-x) = f x` for all `x`. -* `group_seminorm`: A function `f` from a group `G` to the reals that sends one to zero, takes - nonnegative values, is submultiplicative and such that `f x⁻¹ = f x` for all `x`. - For a module over a normed ring: * `seminorm`: A function to the reals that is positive-semidefinite, absolutely homogeneous, and subadditive. @@ -46,277 +40,6 @@ open_locale big_operators nnreal pointwise topological_space variables {R R' 𝕜 E F G ι : Type*} -/-- A seminorm on an additive group `G` is a function `f : G → ℝ` that preserves zero, takes -nonnegative values, is subadditive and such that `f (-x) = f x` for all `x ∈ G`. -/ -structure add_group_seminorm (G : Type*) [add_group G] - extends zero_hom G ℝ := -(nonneg' : ∀ r, 0 ≤ to_fun r) -(add_le' : ∀ r s, to_fun (r + s) ≤ to_fun r + to_fun s) -(neg' : ∀ r, to_fun (- r) = to_fun r) - -/-- A seminorm on a group `G` is a function `f : G → ℝ` that preserves zero, takes nonnegative -values, is submultiplicative and such that `f x⁻¹ = f x` for all `x`. -/ -@[to_additive] structure group_seminorm (G : Type*) [group G] := -(to_fun : G → ℝ) -(map_one' : to_fun 1 = 0) -(nonneg' : ∀ x, 0 ≤ to_fun x) -(mul_le' : ∀ x y, to_fun (x * y) ≤ to_fun x + to_fun y) -(inv' : ∀ x, to_fun x⁻¹ = to_fun x) - -attribute [nolint doc_blame] add_group_seminorm.to_zero_hom - -namespace group_seminorm -section group -variables [group E] [group F] [group G] - -@[to_additive] instance fun_like : fun_like (group_seminorm E) E (λ _, ℝ) := -{ coe := λ f, f.to_fun, - coe_injective' := λ f g h, by cases f; cases g; congr' } - -/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ -@[to_additive "Helper instance for when there's too many metavariables to apply -`fun_like.has_coe_to_fun`. "] -instance : has_coe_to_fun (group_seminorm E) (λ _, E → ℝ) := ⟨to_fun⟩ - -@[ext, to_additive] lemma ext {p q : group_seminorm E} : (∀ x, (p : E → ℝ) x = q x) → p = q := -fun_like.ext p q - -variables (p q : group_seminorm E) (x y : E) (r : ℝ) - -@[to_additive] protected lemma nonneg : 0 ≤ p x := p.nonneg' _ -@[simp, to_additive] protected lemma map_one : p 1 = 0 := p.map_one' -@[to_additive] protected lemma mul_le : p (x * y) ≤ p x + p y := p.mul_le' _ _ -@[simp, to_additive] protected lemma inv : p x⁻¹ = p x := p.inv' _ - -@[to_additive] protected lemma div_le : p (x / y) ≤ p x + p y := -by { rw [div_eq_mul_inv, ←p.inv y], exact p.mul_le _ _ } - -@[to_additive] lemma div_rev : p (x / y) = p (y / x) := by rw [←inv_div, p.inv] - -@[to_additive] instance : partial_order (group_seminorm E) := -partial_order.lift _ fun_like.coe_injective - -@[to_additive] lemma le_def : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl -@[to_additive] lemma lt_def : p < q ↔ (p : E → ℝ) < q := iff.rfl - -variables {p q} - -@[simp, to_additive] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl -@[simp, to_additive] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl - -variables (p q) (f : F →* E) - -@[to_additive] instance : has_zero (group_seminorm E) := -⟨{ to_fun := 0, - nonneg' := λ r, le_rfl, - map_one' := pi.zero_apply _, - mul_le' := λ _ _, (zero_add _).ge, - inv' := λ x, rfl}⟩ - -@[simp, to_additive] lemma coe_zero : ⇑(0 : group_seminorm E) = 0 := rfl -@[simp, to_additive] lemma zero_apply (x : E) : (0 : group_seminorm E) x = 0 := rfl - -@[to_additive] instance : inhabited (group_seminorm E) := ⟨0⟩ - -@[to_additive] instance : has_add (group_seminorm E) := -{ add := λ p q, - { to_fun := λ x, p x + q x, - nonneg' := λ x, add_nonneg (p.nonneg _) (q.nonneg _), - map_one' := by rw [p.map_one, q.map_one, zero_add], - mul_le' := λ _ _, (add_le_add (p.mul_le _ _) $ q.mul_le _ _).trans_eq $ - add_add_add_comm _ _ _ _, - inv' := λ x, by rw [p.inv, q.inv] }} - -@[simp, to_additive] lemma coe_add : ⇑(p + q) = p + q := rfl -@[simp, to_additive] lemma add_apply (x : E) : (p + q) x = p x + q x := rfl - --- TODO: define `has_Sup` too, from the skeleton at --- https://github.com/leanprover-community/mathlib/pull/11329#issuecomment-1008915345 -@[to_additive] noncomputable instance : has_sup (group_seminorm E) := -{ sup := λ p q, - { to_fun := p ⊔ q, - nonneg' := λ x, le_sup_of_le_left $ p.nonneg _, - map_one' := by rw [pi.sup_apply, ←p.map_one, sup_eq_left, p.map_one, q.map_one], - mul_le' := λ x y, sup_le - ((p.mul_le x y).trans $ add_le_add le_sup_left le_sup_left) - ((q.mul_le x y).trans $ add_le_add le_sup_right le_sup_right), - inv' := λ x, by rw [pi.sup_apply, pi.sup_apply, p.inv, q.inv] }} - -@[simp, to_additive] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl -@[simp, to_additive] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl - -@[to_additive] noncomputable instance : semilattice_sup (group_seminorm E) := -fun_like.coe_injective.semilattice_sup _ coe_sup - -/-- Composition of a group seminorm with a monoid homomorphism as a group seminorm. -/ -@[to_additive "Composition of an additive group seminorm with an additive monoid homomorphism as an -additive group seminorm."] -def comp (p : group_seminorm E) (f : F →* E) : group_seminorm F := -{ to_fun := λ x, p (f x), - nonneg' := λ x, p.nonneg _, - map_one' := by rw [f.map_one, p.map_one], - mul_le' := λ _ _, (congr_arg p $ f.map_mul _ _).trans_le $ p.mul_le _ _, - inv' := λ x, by rw [map_inv, p.inv] } - -@[simp, to_additive] lemma coe_comp : ⇑(p.comp f) = p ∘ f := rfl -@[simp, to_additive] lemma comp_apply (x : F) : (p.comp f) x = p (f x) := rfl -@[simp, to_additive] lemma comp_id : p.comp (monoid_hom.id _) = p := ext $ λ _, rfl -@[simp, to_additive] lemma comp_zero : p.comp (1 : F →* E) = 0 := ext $ λ _, p.map_one -@[simp, to_additive] lemma zero_comp : (0 : group_seminorm E).comp f = 0 := ext $ λ _, rfl - -@[to_additive] lemma comp_assoc (g : F →* E) (f : G →* F) : p.comp (g.comp f) = (p.comp g).comp f := -ext $ λ _, rfl - -@[to_additive] lemma add_comp (f : F →* E) : (p + q).comp f = p.comp f + q.comp f := ext $ λ _, rfl - -variables {p q} - -@[to_additive] lemma comp_mono (hp : p ≤ q) : p.comp f ≤ q.comp f := λ _, hp _ - -end group - -section comm_group -variables [comm_group E] [comm_group F] (p q : group_seminorm E) (x y : E) - -/-- The direct path from `1` to `y` is shorter than the path with `x` "inserted" in between. -/ -@[to_additive "The direct path from `0` to `y` is shorter than the path with `x` \"inserted\" in -between."] -lemma le_insert : p y ≤ p x + p (x / y) := -calc p y = p (x / (x / y)) : by rw div_div_cancel - ... ≤ p x + p (x / y) : p.div_le _ _ - -/-- The direct path from 0 to x is shorter than the path with y "inserted" in between. -/ -@[to_additive "The direct path from 0 to x is shorter than the path with y \"inserted\" in -between."] -lemma le_insert' : p x ≤ p y + p (x / y) := by { rw div_rev, exact le_insert _ _ _ } - -@[to_additive] lemma comp_mul_le (f g : F →* E) : p.comp (f * g) ≤ p.comp f + p.comp g := -λ _, p.mul_le _ _ - -@[to_additive] private lemma mul_bdd_below_range_add {p q : group_seminorm E} {x : E} : - bdd_below (range (λ y, p y + q (x / y))) := -⟨0, by { rintro _ ⟨x, rfl⟩, exact add_nonneg (p.nonneg _) (q.nonneg _) }⟩ - -@[to_additive] noncomputable instance : has_inf (group_seminorm E) := -{ inf := λ p q, - { to_fun := λ x, ⨅ y, p y + q (x / y), - nonneg' := λ x, le_cinfi $ λ x, add_nonneg (p.nonneg _) (q.nonneg _), - map_one' := cinfi_eq_of_forall_ge_of_forall_gt_exists_lt - (λ x, add_nonneg (p.nonneg _) (q.nonneg _)) - (λ r hr, ⟨1, by rwa [div_one, p.map_one, q.map_one, add_zero]⟩), - mul_le' := λ x y, le_cinfi_add_cinfi $ λ u v, begin - apply cinfi_le_of_le mul_bdd_below_range_add (u * v), - dsimp, - rw [mul_div_mul_comm, add_add_add_comm], - exact add_le_add (p.mul_le _ _) (q.mul_le _ _), - end, - inv' := λ x, (inv_surjective.infi_comp _).symm.trans $ by simp_rw [p.inv, ←inv_div', q.inv] }} - -@[simp, to_additive] lemma inf_apply : (p ⊓ q) x = ⨅ y, p y + q (x / y) := rfl - -@[to_additive] noncomputable instance : lattice (group_seminorm E) := -{ inf := (⊓), - inf_le_left := λ p q x, cinfi_le_of_le mul_bdd_below_range_add x $ - by rw [div_self', q.map_one, add_zero], - inf_le_right := λ p q x, cinfi_le_of_le mul_bdd_below_range_add (1 : E) $ - by simp only [div_one, p.map_one, zero_add], - le_inf := λ a b c hb hc x, le_cinfi $ λ u, (a.le_insert' _ _).trans $ add_le_add (hb _) (hc _), - ..group_seminorm.semilattice_sup } - -end comm_group -end group_seminorm - -namespace add_group_seminorm -variables [add_group E] (p : add_group_seminorm E) (x y : E) (r : ℝ) - -instance zero_hom_class : zero_hom_class (add_group_seminorm E) E ℝ := -{ coe := λ f, f.to_fun, - coe_injective' := λ f g h, by cases f; cases g; congr', - map_zero := λ f, f.map_zero' } - -/- TODO: All the following ought to be automated using `to_additive`. The problem is that it doesn't -see that `has_smul R ℝ` should be fixed because `ℝ` is fixed. -/ - -/-- Any action on `ℝ` which factors through `ℝ≥0` applies to an `add_group_seminorm`. -/ -instance [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] : - has_smul R (add_group_seminorm E) := -{ smul := λ r p, - { to_fun := λ x, r • p x, - nonneg' := λ x, begin - simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul], - exact mul_nonneg (nnreal.coe_nonneg _) (p.nonneg _) - end, - map_zero' := by simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul, - p.map_zero, mul_zero], - add_le' := λ _ _, begin - simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul], - exact (mul_le_mul_of_nonneg_left (p.add_le _ _) (nnreal.coe_nonneg _)).trans_eq - (mul_add _ _ _), - end, - neg' := λ x, by rw p.neg }} - -instance [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] - [has_smul R' ℝ] [has_smul R' ℝ≥0] [is_scalar_tower R' ℝ≥0 ℝ] - [has_smul R R'] [is_scalar_tower R R' ℝ] : - is_scalar_tower R R' (add_group_seminorm E) := -{ smul_assoc := λ r a p, ext $ λ x, smul_assoc r a (p x) } - -@[simp] lemma coe_smul [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] - (r : R) (p : add_group_seminorm E) : ⇑(r • p) = r • p := rfl - -@[simp] lemma smul_apply [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] - (r : R) (p : add_group_seminorm E) (x : E) : (r • p) x = r • p x := rfl - -lemma smul_sup [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] - (r : R) (p q : add_group_seminorm E) : - r • (p ⊔ q) = r • p ⊔ r • q := -have real.smul_max : ∀ x y : ℝ, r • max x y = max (r • x) (r • y), -from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ℝ≥0 r (_ : ℝ)] - using mul_max_of_nonneg x y (r • 1 : ℝ≥0).prop, -ext $ λ x, real.smul_max _ _ - -end add_group_seminorm - -namespace group_seminorm -variables [group E] [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] - -/-- Any action on `ℝ` which factors through `ℝ≥0` applies to an `add_group_seminorm`. -/ -@[to_additive add_group_seminorm.has_smul] instance : has_smul R (group_seminorm E) := - ⟨λ r p, - { to_fun := λ x, r • p x, - nonneg' := λ x, begin - simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul], - exact mul_nonneg (nnreal.coe_nonneg _) (p.nonneg _) - end, - map_one' := by simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul, - p.map_one, mul_zero], - mul_le' := λ _ _, begin - simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul], - exact (mul_le_mul_of_nonneg_left (p.mul_le _ _) $ nnreal.coe_nonneg _).trans_eq - (mul_add _ _ _), - end, - inv' := λ x, by rw p.inv }⟩ - -@[to_additive add_group_seminorm.is_scalar_tower] -instance [has_smul R' ℝ] [has_smul R' ℝ≥0] [is_scalar_tower R' ℝ≥0 ℝ] [has_smul R R'] - [is_scalar_tower R R' ℝ] : is_scalar_tower R R' (group_seminorm E) := -⟨λ r a p, ext $ λ x, smul_assoc r a $ p x⟩ - -@[simp, to_additive add_group_seminorm.coe_smul] -lemma coe_smul (r : R) (p : group_seminorm E) : ⇑(r • p) = r • p := rfl - -@[simp, to_additive add_group_seminorm.smul_apply] -lemma smul_apply (r : R) (p : group_seminorm E) (x : E) : (r • p) x = r • p x := rfl - -@[to_additive add_group_seminorm.smul_sup] -lemma smul_sup (r : R) (p q : group_seminorm E) : r • (p ⊔ q) = r • p ⊔ r • q := -have real.smul_max : ∀ x y : ℝ, r • max x y = max (r • x) (r • y), -from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ℝ≥0 r (_ : ℝ)] - using mul_max_of_nonneg x y (r • 1 : ℝ≥0).prop, -ext $ λ x, real.smul_max _ _ - -end group_seminorm - /-- A seminorm on a module over a normed ring is a function to the reals that is positive semidefinite, positive homogeneous, and subadditive. -/ structure seminorm (𝕜 : Type*) (E : Type*) [semi_normed_ring 𝕜] [add_group E] [has_smul 𝕜 E] @@ -325,36 +48,16 @@ structure seminorm (𝕜 : Type*) (E : Type*) [semi_normed_ring 𝕜] [add_group attribute [nolint doc_blame] seminorm.to_add_group_seminorm -private lemma map_zero.of_smul {𝕜 : Type*} {E : Type*} [semi_normed_ring 𝕜] [add_group E] - [smul_with_zero 𝕜 E] {f : E → ℝ} (smul : ∀ (a : 𝕜) (x : E), f (a • x) = ∥a∥ * f x) : f 0 = 0 := -calc f 0 = f ((0 : 𝕜) • 0) : by rw zero_smul - ... = 0 : by rw [smul, norm_zero, zero_mul] - -private lemma neg.of_smul {𝕜 : Type*} {E : Type*} [semi_normed_ring 𝕜] [add_comm_group E] - [module 𝕜 E] {f : E → ℝ} (smul : ∀ (a : 𝕜) (x : E), f (a • x) = ∥a∥ * f x) (x : E) : - f (-x) = f x := -by rw [←neg_one_smul 𝕜, smul, norm_neg, ← smul, one_smul] - -private lemma nonneg.of {𝕜 : Type*} {E : Type*} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] - {f : E → ℝ} (add_le : ∀ (x y : E), f (x + y) ≤ f x + f y) - (smul : ∀ (a : 𝕜) (x : E), f (a • x) = ∥a∥ * f x) (x : E) : 0 ≤ f x := -have h: 0 ≤ 2 * f x, from -calc 0 = f (x + (- x)) : by rw [add_neg_self, map_zero.of_smul smul] -... ≤ f x + f (-x) : add_le _ _ -... = 2 * f x : by rw [neg.of_smul smul, two_mul], -nonneg_of_mul_nonneg_right h zero_lt_two - /-- Alternative constructor for a `seminorm` on an `add_comm_group E` that is a module over a `semi_norm_ring 𝕜`. -/ def seminorm.of {𝕜 : Type*} {E : Type*} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] (f : E → ℝ) (add_le : ∀ (x y : E), f (x + y) ≤ f x + f y) (smul : ∀ (a : 𝕜) (x : E), f (a • x) = ∥a∥ * f x) : seminorm 𝕜 E := { to_fun := f, - map_zero' := map_zero.of_smul smul, - nonneg' := nonneg.of add_le smul, + map_zero' := by rw [←zero_smul 𝕜 (0 : E), smul, norm_zero, zero_mul], add_le' := add_le, smul' := smul, - neg' := neg.of_smul smul } + neg' := λ x, by rw [←neg_one_smul 𝕜, smul, norm_neg, ← smul, one_smul] } namespace seminorm @@ -389,7 +92,7 @@ instance : inhabited (seminorm 𝕜 E) := ⟨0⟩ variables (p : seminorm 𝕜 E) (c : 𝕜) (x y : E) (r : ℝ) -protected lemma nonneg : 0 ≤ p x := p.nonneg' _ +protected lemma nonneg : 0 ≤ p x := p.to_add_group_seminorm.nonneg _ protected lemma map_zero : p 0 = 0 := p.map_zero' protected lemma smul : p (c • x) = ∥c∥ * p x := p.smul' _ _ protected lemma add_le : p (x + y) ≤ p x + p y := p.add_le' _ _ @@ -929,12 +632,6 @@ end seminorm section norm_seminorm variables (𝕜) (E) [normed_field 𝕜] [seminormed_add_comm_group E] [normed_space 𝕜 E] {r : ℝ} -/-- The norm of a seminormed group as an add_monoid seminorm. -/ -def norm_add_group_seminorm : add_group_seminorm E := -⟨norm, norm_zero, norm_nonneg, norm_add_le, norm_neg⟩ - -@[simp] lemma coe_norm_add_group_seminorm : ⇑(norm_add_group_seminorm E) = norm := rfl - /-- The norm of a seminormed group as a seminorm. -/ def norm_seminorm : seminorm 𝕜 E := { smul' := norm_smul, From 093c5036c7d80f381c16b74813d4ca1d4c3d7c64 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 24 Aug 2022 21:27:27 +0000 Subject: [PATCH 030/828] feat(category_theory/strong_epi): various results about strong_epi (#16182) In order to prepare for the proof of the existence of strong epi mono factorisations in `simplex_category` (which shall be used for the Dold-Kan correspondence), various results are obtained, mostly about strong epimorphisms (and strong monomorphisms). If two arrows are isomorphic, then one is a strong epi iff the other is. An equivalence of categories preserves and reflects strong epimorphisms. --- src/category_theory/arrow.lean | 22 ++++++++ .../concrete_category/basic.lean | 4 ++ src/category_theory/functor/epi_mono.lean | 53 +++++++++++++++++++ .../lifting_properties/basic.lean | 20 +++++++ .../limits/shapes/strong_epi.lean | 26 +++++++++ 5 files changed, 125 insertions(+) diff --git a/src/category_theory/arrow.lean b/src/category_theory/arrow.lean index e95092f6dfb3f..51dfed31145cc 100644 --- a/src/category_theory/arrow.lean +++ b/src/category_theory/arrow.lean @@ -106,6 +106,21 @@ abbreviation iso_mk' {W X Y Z : T} (f : W ⟶ X) (g : Y ⟶ Z) (e₁ : W ≅ Y) (e₂ : X ≅ Z) (h : e₁.hom ≫ g = f ≫ e₂.hom) : arrow.mk f ≅ arrow.mk g := arrow.iso_mk e₁ e₂ h +lemma hom.congr_left {f g : arrow T} {φ₁ φ₂ : f ⟶ g} (h : φ₁ = φ₂) : + φ₁.left = φ₂.left := by rw h +lemma hom.congr_right {f g : arrow T} {φ₁ φ₂ : f ⟶ g} (h : φ₁ = φ₂) : + φ₁.right = φ₂.right := by rw h + +lemma iso_w {f g : arrow T} (e : f ≅ g) : g.hom = e.inv.left ≫ f.hom ≫ e.hom.right := +begin + have eq := arrow.hom.congr_right e.inv_hom_id, + dsimp at eq, + erw [w_assoc, eq, category.comp_id], +end + +lemma iso_w' {W X Y Z : T} {f : W ⟶ X} {g : Y ⟶ Z} (e : arrow.mk f ≅ arrow.mk g) : + g = e.inv.left ≫ f ≫ e.hom.right := iso_w e + section variables {f g : arrow T} (sq : f ⟶ g) @@ -216,4 +231,11 @@ def map_arrow (F : C ⥤ D) : arrow C ⥤ arrow D := end functor +/-- The images of `f : arrow C` by two isomorphic functors `F : C ⥤ D` are +isomorphic arrows in `D`. -/ +def arrow.iso_of_nat_iso {C D : Type*} [category C] [category D] + {F G : C ⥤ D} (e : F ≅ G) (f : arrow C) : + F.map_arrow.obj f ≅ G.map_arrow.obj f := +arrow.iso_mk (e.app f.left) (e.app f.right) (by simp) + end category_theory diff --git a/src/category_theory/concrete_category/basic.lean b/src/category_theory/concrete_category/basic.lean index d9fc670e07014..f39794e77d776 100644 --- a/src/category_theory/concrete_category/basic.lean +++ b/src/category_theory/concrete_category/basic.lean @@ -134,6 +134,10 @@ lemma concrete_category.epi_of_surjective {X Y : C} (f : X ⟶ Y) (s : function. epi f := (forget C).epi_of_epi_map ((epi_iff_surjective f).2 s) +lemma concrete_category.bijective_of_is_iso {X Y : C} (f : X ⟶ Y) [is_iso f] : + function.bijective ((forget C).map f) := +by { rw ← is_iso_iff_bijective, apply_instance, } + @[simp] lemma concrete_category.has_coe_to_fun_Type {X Y : Type u} (f : X ⟶ Y) : coe_fn f = f := rfl diff --git a/src/category_theory/functor/epi_mono.lean b/src/category_theory/functor/epi_mono.lean index 54039ba411dde..363bd3e4d8602 100644 --- a/src/category_theory/functor/epi_mono.lean +++ b/src/category_theory/functor/epi_mono.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ import category_theory.epi_mono +import category_theory.limits.shapes.strong_epi +import category_theory.lifting_properties.adjunction /-! # Preservation and reflection of monomorphisms and epimorphisms @@ -199,6 +201,14 @@ def split_epi_equiv [full F] [faithful F] : split_epi f ≃ split_epi (F.map f) left_inv := by tidy, right_inv := by tidy, } +@[simp] +lemma is_split_epi_iff [full F] [faithful F] : is_split_epi (F.map f) ↔ is_split_epi f := +begin + split, + { intro h, exact is_split_epi.mk' ((split_epi_equiv F f).inv_fun h.exists_split_epi.some), }, + { intro h, exact is_split_epi.mk' ((split_epi_equiv F f).to_fun h.exists_split_epi.some), }, +end + /-- If `F` is a fully faithful functor, split monomorphisms are preserved and reflected by `F`. -/ def split_mono_equiv [full F] [faithful F] : split_mono f ≃ split_mono (F.map f) := { to_fun := λ f, f.map F, @@ -211,6 +221,14 @@ def split_mono_equiv [full F] [faithful F] : split_mono f ≃ split_mono (F.map left_inv := by tidy, right_inv := by tidy, } +@[simp] +lemma is_split_mono_iff [full F] [faithful F] : is_split_mono (F.map f) ↔ is_split_mono f := +begin + split, + { intro h, exact is_split_mono.mk' ((split_mono_equiv F f).inv_fun h.exists_split_mono.some), }, + { intro h, exact is_split_mono.mk' ((split_mono_equiv F f).to_fun h.exists_split_mono.some), }, +end + @[simp] lemma epi_map_iff_epi [hF₁ : preserves_epimorphisms F] [hF₂ : reflects_epimorphisms F] : epi (F.map f) ↔ epi f := @@ -234,3 +252,38 @@ end end end category_theory.functor + +namespace category_theory.adjunction + +variables {C D : Type*} [category C] [category D] {F : C ⥤ D} {F' : D ⥤ C} {A B : C} + +lemma strong_epi_map_of_strong_epi (adj : F ⊣ F') (f : A ⟶ B) + [h₁ : F'.preserves_monomorphisms] [h₂ : F.preserves_epimorphisms] [strong_epi f] : + strong_epi (F.map f) := +⟨infer_instance, λ X Y Z, by { introI, rw adj.has_lifting_property_iff, apply_instance, }⟩ + +instance strong_epi_map_of_is_equivalence [is_equivalence F] (f : A ⟶ B) [h : strong_epi f] : + strong_epi (F.map f) := +F.as_equivalence.to_adjunction.strong_epi_map_of_strong_epi f + +end category_theory.adjunction + +namespace category_theory.functor + +variables {C D : Type*} [category C] [category D] {F : C ⥤ D} {A B : C} (f : A ⟶ B) + +@[simp] +lemma strong_epi_map_iff_strong_epi_of_is_equivalence [is_equivalence F] : + strong_epi (F.map f) ↔ strong_epi f := +begin + split, + { introI, + have e : arrow.mk f ≅ arrow.mk (F.inv.map (F.map f)) := + arrow.iso_of_nat_iso F.as_equivalence.unit_iso (arrow.mk f), + rw strong_epi.iff_of_arrow_iso e, + apply_instance, }, + { introI, + apply_instance, }, +end + +end category_theory.functor diff --git a/src/category_theory/lifting_properties/basic.lean b/src/category_theory/lifting_properties/basic.lean index 9171c760b84d6..79a4919486576 100644 --- a/src/category_theory/lifting_properties/basic.lean +++ b/src/category_theory/lifting_properties/basic.lean @@ -104,6 +104,26 @@ instance of_comp_right [has_lifting_property i p] [has_lifting_property i p'] : fac_right' := by simp only [comm_sq.fac_right_assoc, comm_sq.fac_right], }, end⟩ +lemma of_arrow_iso_left {A B A' B' X Y : C} {i : A ⟶ B} {i' : A' ⟶ B'} + (e : arrow.mk i ≅ arrow.mk i') (p : X ⟶ Y) + [hip : has_lifting_property i p] : has_lifting_property i' p := +by { rw arrow.iso_w' e, apply_instance, } + +lemma of_arrow_iso_right {A B X Y X' Y' : C} (i : A ⟶ B) {p : X ⟶ Y} {p' : X' ⟶ Y'} + (e : arrow.mk p ≅ arrow.mk p') + [hip : has_lifting_property i p] : has_lifting_property i p' := +by { rw arrow.iso_w' e, apply_instance, } + +lemma iff_of_arrow_iso_left {A B A' B' X Y : C} {i : A ⟶ B} {i' : A' ⟶ B'} + (e : arrow.mk i ≅ arrow.mk i') (p : X ⟶ Y) : + has_lifting_property i p ↔ has_lifting_property i' p := +by { split; introI, exacts [of_arrow_iso_left e p, of_arrow_iso_left e.symm p], } + +lemma iff_of_arrow_iso_right {A B X Y X' Y' : C} (i : A ⟶ B) {p : X ⟶ Y} {p' : X' ⟶ Y'} + (e : arrow.mk p ≅ arrow.mk p') : + has_lifting_property i p ↔ has_lifting_property i p' := +by { split; introI, exacts [of_arrow_iso_right i e, of_arrow_iso_right i e.symm], } + end has_lifting_property end category_theory diff --git a/src/category_theory/limits/shapes/strong_epi.lean b/src/category_theory/limits/shapes/strong_epi.lean index 90efac5332580..ce67158f4b03e 100644 --- a/src/category_theory/limits/shapes/strong_epi.lean +++ b/src/category_theory/limits/shapes/strong_epi.lean @@ -119,6 +119,32 @@ lemma strong_mono_of_strong_mono [strong_mono (f ≫ g)] : strong_mono f := { mono := by apply_instance, rlp := λ X Y z hz, has_lifting_property.of_right_iso _ _, } +lemma strong_epi.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} + (e : arrow.mk f ≅ arrow.mk g) [h : strong_epi f] : strong_epi g := +{ epi := begin + rw arrow.iso_w' e, + haveI := epi_comp f e.hom.right, + apply epi_comp, + end, + llp := λ X Y z, by { introI, apply has_lifting_property.of_arrow_iso_left e z, }, } + +lemma strong_mono.of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} + (e : arrow.mk f ≅ arrow.mk g) [h : strong_mono f] : strong_mono g := +{ mono := begin + rw arrow.iso_w' e, + haveI := mono_comp f e.hom.right, + apply mono_comp, + end, + rlp := λ X Y z, by { introI, apply has_lifting_property.of_arrow_iso_right z e, }, } + +lemma strong_epi.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} + (e : arrow.mk f ≅ arrow.mk g) : strong_epi f ↔ strong_epi g := +by { split; introI, exacts [strong_epi.of_arrow_iso e, strong_epi.of_arrow_iso e.symm], } + +lemma strong_mono.iff_of_arrow_iso {A B A' B' : C} {f : A ⟶ B} {g : A' ⟶ B'} + (e : arrow.mk f ≅ arrow.mk g) : strong_mono f ↔ strong_mono g := +by { split; introI, exacts [strong_mono.of_arrow_iso e, strong_mono.of_arrow_iso e.symm], } + end /-- A strong epimorphism that is a monomorphism is an isomorphism. -/ From 62740866400028a5fdbbd58a6004a4872d058c1a Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 24 Aug 2022 21:27:28 +0000 Subject: [PATCH 031/828] refactor(algebra/order/{monoid_lemmas_zero_lt + ring} + crumbs) remove duplicate lemmas (#16189) This is one of the cleaning up stages of the `order` refactor. This PR removes some of the lemmas for ordered semirings that are now superfluous. To keep things simple, the new lemmas are in the `zero_lt` namespace and we export this namespace. Probably, this will not stay this way once the tide of unsuccessful builds washes away. The plan for this PR is to align names and order of explicit hypotheses in the `zero_lt` file with the lemmas in `ring`, in order to delete them. This will likely involve changing a few extra files, but I'm hoping to seriously contain the diff outside of these two files. I want to thank @negiizhao who has been a great help in taking over and pushing forward the refactor! --- archive/imo/imo1972_q5.lean | 2 +- src/algebra/order/monoid_lemmas_zero_lt.lean | 328 +++++++++--------- src/algebra/order/ring.lean | 154 +------- src/analysis/specific_limits/floor_pow.lean | 4 +- src/data/nat/basic.lean | 2 +- src/data/real/nnreal.lean | 2 +- src/measure_theory/integral/set_to_l1.lean | 4 +- .../liouville/liouville_constant.lean | 2 +- 8 files changed, 183 insertions(+), 315 deletions(-) diff --git a/archive/imo/imo1972_q5.lean b/archive/imo/imo1972_q5.lean index 4b83138a45def..5664210339a6c 100644 --- a/archive/imo/imo1972_q5.lean +++ b/archive/imo/imo1972_q5.lean @@ -112,7 +112,7 @@ begin { suffices : ∀ x, ∥f x∥ ≤ k / ∥g y∥, from csupr_le this, intro x, suffices : 2 * (∥f x∥ * ∥g y∥) ≤ 2 * k, - by { rwa [le_div_iff hgy, ←mul_le_mul_left zero_lt_two], apply_instance }, + by { rwa [le_div_iff hgy, ←mul_le_mul_left (zero_lt_two : (0 : ℝ) < 2)] }, calc 2 * (∥f x∥ * ∥g y∥) = ∥2 * f x * g y∥ : by simp [abs_mul, mul_assoc] ... = ∥f (x + y) + f (x - y)∥ : by rw hf1 diff --git a/src/algebra/order/monoid_lemmas_zero_lt.lean b/src/algebra/order/monoid_lemmas_zero_lt.lean index 43d34ab2c1309..83f6fcdb9471e 100644 --- a/src/algebra/order/monoid_lemmas_zero_lt.lean +++ b/src/algebra/order/monoid_lemmas_zero_lt.lean @@ -132,12 +132,12 @@ variables [has_mul α] [has_zero α] section has_lt variables [has_lt α] -lemma mul_lt_mul_left [pos_mul_strict_mono α] +lemma mul_lt_mul_left' [pos_mul_strict_mono α] (bc : b < c) (a0 : 0 < a) : a * b < a * c := @covariant_class.elim α>0 α (λ x y, x * y) (<) _ ⟨a, a0⟩ _ _ bc -lemma mul_lt_mul_right [mul_pos_strict_mono α] +lemma mul_lt_mul_right' [mul_pos_strict_mono α] (bc : b < c) (a0 : 0 < a) : b * a < c * a := @covariant_class.elim α>0 α (λ x y, y * x) (<) _ ⟨a, a0⟩ _ _ bc @@ -155,13 +155,13 @@ lemma lt_of_mul_lt_mul_right' [mul_pos_reflect_lt α] @contravariant_class.elim α>0 α (λ x y, y * x) (<) _ ⟨a, a0⟩ _ _ bc @[simp] -lemma mul_lt_mul_iff_left [pos_mul_strict_mono α] [pos_mul_reflect_lt α] +lemma mul_lt_mul_left [pos_mul_strict_mono α] [pos_mul_reflect_lt α] (a0 : 0 < a) : a * b < a * c ↔ b < c := @rel_iff_cov α>0 α (λ x y, x * y) (<) _ _ ⟨a, a0⟩ _ _ @[simp] -lemma mul_lt_mul_iff_right [mul_pos_strict_mono α] [mul_pos_reflect_lt α] +lemma mul_lt_mul_right [mul_pos_strict_mono α] [mul_pos_reflect_lt α] (a0 : 0 < a) : b * a < c * a ↔ b < c := @rel_iff_cov α>0 α (λ x y, y * x) (<) _ _ ⟨a, a0⟩ _ _ @@ -172,13 +172,13 @@ section has_lt_le variables [has_lt α] [has_le α] -- proven with `a0 : 0 ≤ a` as `mul_le_mul_left` -lemma mul_le_mul_left' [pos_mul_mono α] +lemma mul_le_mul_left_of_le_of_pos [pos_mul_mono α] (bc : b ≤ c) (a0 : 0 < a) : a * b ≤ a * c := @covariant_class.elim α>0 α (λ x y, x * y) (≤) _ ⟨a, a0⟩ _ _ bc -- proven with `a0 : 0 ≤ a` as `mul_le_mul_right` -lemma mul_le_mul_right' [mul_pos_mono α] +lemma mul_le_mul_right_of_le_of_pos [mul_pos_mono α] (bc : b ≤ c) (a0 : 0 < a) : b * a ≤ c * a := @covariant_class.elim α>0 α (λ x y, y * x) (≤) _ ⟨a, a0⟩ _ _ bc @@ -194,13 +194,13 @@ lemma le_of_mul_le_mul_right' [mul_pos_mono_rev α] @contravariant_class.elim α>0 α (λ x y, y * x) (≤) _ ⟨a, a0⟩ _ _ bc @[simp] -lemma mul_le_mul_iff_left [pos_mul_mono α] [pos_mul_mono_rev α] +lemma mul_le_mul_left [pos_mul_mono α] [pos_mul_mono_rev α] (a0 : 0 < a) : a * b ≤ a * c ↔ b ≤ c := @rel_iff_cov α>0 α (λ x y, x * y) (≤) _ _ ⟨a, a0⟩ _ _ @[simp] -lemma mul_le_mul_iff_right [mul_pos_mono α] [mul_pos_mono_rev α] +lemma mul_le_mul_right [mul_pos_mono α] [mul_pos_mono_rev α] (a0 : 0 < a) : b * a ≤ c * a ↔ b ≤ c := @rel_iff_cov α>0 α (λ x y, y * x) (≤) _ _ ⟨a, a0⟩ _ _ @@ -213,80 +213,80 @@ variables [preorder α] -- proven with `a0 : 0 ≤ a` `d0 : 0 ≤ d` as `mul_le_mul_of_le_of_le` lemma preorder.mul_le_mul_of_le_of_le [pos_mul_mono α] [mul_pos_mono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (a0 : 0 < a) (d0 : 0 < d) : a * c ≤ b * d := -(mul_le_mul_left' h₂ a0).trans (mul_le_mul_right' h₁ d0) +(mul_le_mul_left_of_le_of_pos h₂ a0).trans (mul_le_mul_right_of_le_of_pos h₁ d0) -- proven with `b0 : 0 ≤ b` `c0 : 0 ≤ c` as `mul_le_mul_of_le_of_le'` lemma preorder.mul_le_mul_of_le_of_le' [pos_mul_mono α] [mul_pos_mono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (b0 : 0 < b) (c0 : 0 < c) : a * c ≤ b * d := -(mul_le_mul_right' h₁ c0).trans (mul_le_mul_left' h₂ b0) +(mul_le_mul_right_of_le_of_pos h₁ c0).trans (mul_le_mul_left_of_le_of_pos h₂ b0) lemma mul_lt_mul_of_le_of_lt [pos_mul_strict_mono α] [mul_pos_mono α] (h₁ : a ≤ b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := -(mul_lt_mul_left h₂ a0).trans_le (mul_le_mul_right' h₁ d0) +(mul_lt_mul_left' h₂ a0).trans_le (mul_le_mul_right_of_le_of_pos h₁ d0) lemma mul_lt_mul_of_le_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] (h₁ : a ≤ b) (h₂ : c < d) (b0 : 0 < b) (c0 : 0 < c) : a * c < b * d := -(mul_le_mul_right' h₁ c0).trans_lt (mul_lt_mul_left h₂ b0) +(mul_le_mul_right_of_le_of_pos h₁ c0).trans_lt (mul_lt_mul_left' h₂ b0) lemma mul_lt_mul_of_lt_of_le [pos_mul_mono α] [mul_pos_strict_mono α] (h₁ : a < b) (h₂ : c ≤ d) (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := -(mul_le_mul_left' h₂ a0).trans_lt (mul_lt_mul_right h₁ d0) +(mul_le_mul_left_of_le_of_pos h₂ a0).trans_lt (mul_lt_mul_right' h₁ d0) lemma mul_lt_mul_of_lt_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] (h₁ : a < b) (h₂ : c ≤ d) (b0 : 0 < b) (c0 : 0 < c) : a * c < b * d := -(mul_lt_mul_right h₁ c0).trans_le (mul_le_mul_left' h₂ b0) +(mul_lt_mul_right' h₁ c0).trans_le (mul_le_mul_left_of_le_of_pos h₂ b0) lemma mul_lt_mul_of_lt_of_lt [pos_mul_strict_mono α] [mul_pos_strict_mono α] (h₁ : a < b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := -(mul_lt_mul_left h₂ a0).trans (mul_lt_mul_right h₁ d0) +(mul_lt_mul_left' h₂ a0).trans (mul_lt_mul_right' h₁ d0) lemma mul_lt_mul_of_lt_of_lt' [pos_mul_strict_mono α] [mul_pos_strict_mono α] (h₁ : a < b) (h₂ : c < d) (b0 : 0 < b) (c0 : 0 < c) : a * c < b * d := -(mul_lt_mul_right h₁ c0).trans (mul_lt_mul_left h₂ b0) +(mul_lt_mul_right' h₁ c0).trans (mul_lt_mul_left' h₂ b0) -- proven with `a0 : 0 ≤ a` as `mul_le_of_mul_le_left` lemma preorder.mul_le_of_mul_le_left [pos_mul_mono α] (h : a * b ≤ c) (hle : d ≤ b) (a0 : 0 < a) : a * d ≤ c := -(mul_le_mul_left' hle a0).trans h +(mul_le_mul_left_of_le_of_pos hle a0).trans h lemma mul_lt_of_mul_lt_left [pos_mul_mono α] (h : a * b < c) (hle : d ≤ b) (a0 : 0 < a) : a * d < c := -(mul_le_mul_left' hle a0).trans_lt h +(mul_le_mul_left_of_le_of_pos hle a0).trans_lt h -- proven with `b0 : 0 ≤ b` as `le_mul_of_le_mul_left` lemma preorder.le_mul_of_le_mul_left [pos_mul_mono α] (h : a ≤ b * c) (hle : c ≤ d) (b0 : 0 < b) : a ≤ b * d := -h.trans (mul_le_mul_left' hle b0) +h.trans (mul_le_mul_left_of_le_of_pos hle b0) lemma lt_mul_of_lt_mul_left [pos_mul_mono α] (h : a < b * c) (hle : c ≤ d) (b0 : 0 < b) : a < b * d := -h.trans_le (mul_le_mul_left' hle b0) +h.trans_le (mul_le_mul_left_of_le_of_pos hle b0) -- proven with `b0 : 0 ≤ b` as `mul_le_of_mul_le_right` lemma preorder.mul_le_of_mul_le_right [mul_pos_mono α] (h : a * b ≤ c) (hle : d ≤ a) (b0 : 0 < b) : d * b ≤ c := -(mul_le_mul_right' hle b0).trans h +(mul_le_mul_right_of_le_of_pos hle b0).trans h lemma mul_lt_of_mul_lt_right [mul_pos_mono α] (h : a * b < c) (hle : d ≤ a) (b0 : 0 < b) : d * b < c := -(mul_le_mul_right' hle b0).trans_lt h +(mul_le_mul_right_of_le_of_pos hle b0).trans_lt h -- proven with `c0 : 0 ≤ c` as `le_mul_of_le_mul_right` lemma preorder.le_mul_of_le_mul_right [mul_pos_mono α] (h : a ≤ b * c) (hle : b ≤ d) (c0 : 0 < c) : a ≤ d * c := -h.trans (mul_le_mul_right' hle c0) +h.trans (mul_le_mul_right_of_le_of_pos hle c0) lemma lt_mul_of_lt_mul_right [mul_pos_mono α] (h : a < b * c) (hle : b ≤ d) (c0 : 0 < c) : a < d * c := -h.trans_le (mul_le_mul_right' hle c0) +h.trans_le (mul_le_mul_right_of_le_of_pos hle c0) end preorder @@ -295,11 +295,11 @@ variables [partial_order α] @[priority 100] -- see Note [lower instance priority] instance pos_mul_strict_mono.to_pos_mul_mono [pos_mul_strict_mono α] : pos_mul_mono α := -⟨λ x a b h, h.eq_or_lt.elim (λ h', h' ▸ le_rfl) (λ h', (mul_lt_mul_left h' x.prop).le)⟩ +⟨λ x a b h, h.eq_or_lt.elim (λ h', h' ▸ le_rfl) (λ h', (mul_lt_mul_left' h' x.prop).le)⟩ @[priority 100] -- see Note [lower instance priority] instance mul_pos_strict_mono.to_mul_pos_mono [mul_pos_strict_mono α] : mul_pos_mono α := -⟨λ x a b h, h.eq_or_lt.elim (λ h', h' ▸ le_rfl) (λ h', (mul_lt_mul_right h' x.prop).le)⟩ +⟨λ x a b h, h.eq_or_lt.elim (λ h', h' ▸ le_rfl) (λ h', (mul_lt_mul_right' h' x.prop).le)⟩ @[priority 100] -- see Note [lower instance priority] instance pos_mul_mono_rev.to_pos_mul_reflect_lt [pos_mul_mono_rev α] : pos_mul_reflect_lt α := @@ -316,11 +316,11 @@ variables [linear_order α] @[priority 100] -- see Note [lower instance priority] instance pos_mul_strict_mono.to_pos_mul_mono_rev [pos_mul_strict_mono α] : pos_mul_mono_rev α := -⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt (mul_lt_mul_left h' x.prop)⟩ +⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt (mul_lt_mul_left' h' x.prop)⟩ @[priority 100] -- see Note [lower instance priority] instance mul_pos_strict_mono.to_mul_pos_mono_rev [mul_pos_strict_mono α] : mul_pos_mono_rev α := -⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt (mul_lt_mul_right h' x.prop)⟩ +⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt (mul_lt_mul_right' h' x.prop)⟩ lemma pos_mul_mono_rev.to_pos_mul_strict_mono [pos_mul_mono_rev α] : pos_mul_strict_mono α := ⟨λ x a b h, lt_of_not_le $ λ h', h.not_le (le_of_mul_le_mul_left' h' x.prop)⟩ @@ -343,10 +343,10 @@ lemma mul_pos_reflect_lt.to_mul_pos_mono [mul_pos_reflect_lt α] : mul_pos_mono ⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt (lt_of_mul_lt_mul_right' h' x.prop)⟩ lemma pos_mul_mono.to_pos_mul_reflect_lt [pos_mul_mono α] : pos_mul_reflect_lt α := -⟨λ x a b h, lt_of_not_le $ λ h', h.not_le (mul_le_mul_left' h' x.prop)⟩ +⟨λ x a b h, lt_of_not_le $ λ h', h.not_le (mul_le_mul_left_of_le_of_pos h' x.prop)⟩ lemma mul_pos_mono.to_mul_pos_reflect_lt [mul_pos_mono α] : mul_pos_reflect_lt α := -⟨λ x a b h, lt_of_not_le $ λ h', h.not_le (mul_le_mul_right' h' x.prop)⟩ +⟨λ x a b h, lt_of_not_le $ λ h', h.not_le (mul_le_mul_right_of_le_of_pos h' x.prop)⟩ lemma pos_mul_mono_iff_pos_mul_reflect_lt : pos_mul_mono α ↔ pos_mul_reflect_lt α := ⟨@pos_mul_mono.to_pos_mul_reflect_lt _ _ _ _, @pos_mul_reflect_lt.to_pos_mul_mono _ _ _ _⟩ @@ -368,67 +368,79 @@ variables [preorder α] lemma left.mul_pos [pos_mul_strict_mono α] (ha : 0 < a) (hb : 0 < b) : 0 < a * b := -have h : a * 0 < a * b, from mul_lt_mul_left hb ha, +have h : a * 0 < a * b, from mul_lt_mul_left' hb ha, by rwa [mul_zero] at h +alias left.mul_pos ← _root_.mul_pos + lemma mul_neg_of_pos_of_neg [pos_mul_strict_mono α] (ha : 0 < a) (hb : b < 0) : a * b < 0 := -have h : a * b < a * 0, from mul_lt_mul_left hb ha, +have h : a * b < a * 0, from mul_lt_mul_left' hb ha, by rwa [mul_zero] at h +@[simp] lemma zero_lt_mul_left [pos_mul_strict_mono α] [pos_mul_reflect_lt α] (h : 0 < c) : + 0 < c * b ↔ 0 < b := +by { convert mul_lt_mul_left h, simp } + /-- Assumes right covariance. -/ lemma right.mul_pos [mul_pos_strict_mono α] (ha : 0 < a) (hb : 0 < b) : 0 < a * b := -have h : 0 * b < a * b, from mul_lt_mul_right ha hb, +have h : 0 * b < a * b, from mul_lt_mul_right' ha hb, by rwa [zero_mul] at h lemma mul_neg_of_neg_of_pos [mul_pos_strict_mono α] (ha : a < 0) (hb : 0 < b) : a * b < 0 := -have h : a * b < 0 * b, from mul_lt_mul_right ha hb, +have h : a * b < 0 * b, from mul_lt_mul_right' ha hb, by rwa [zero_mul] at h +@[simp] lemma zero_lt_mul_right [mul_pos_strict_mono α] [mul_pos_reflect_lt α] (h : 0 < c) : + 0 < b * c ↔ 0 < b := +by { convert mul_lt_mul_right h, simp } + end preorder section partial_order variables [partial_order α] -lemma mul_le_mul_left [pos_mul_mono α] +lemma mul_le_mul_of_nonneg_left [pos_mul_mono α] (bc : b ≤ c) (a0 : 0 ≤ a) : a * b ≤ a * c := -a0.lt_or_eq.elim (mul_le_mul_left' bc) (λ h, by simp only [← h, zero_mul]) +a0.lt_or_eq.elim (mul_le_mul_left_of_le_of_pos bc) (λ h, by simp only [← h, zero_mul]) -lemma mul_le_mul_right [mul_pos_mono α] +lemma mul_le_mul_of_nonneg_right [mul_pos_mono α] (bc : b ≤ c) (a0 : 0 ≤ a) : b * a ≤ c * a := -a0.lt_or_eq.elim (mul_le_mul_right' bc) (λ h, by simp only [← h, mul_zero]) +a0.lt_or_eq.elim (mul_le_mul_right_of_le_of_pos bc) (λ h, by simp only [← h, mul_zero]) /-- Assumes left covariance. -/ lemma left.mul_nonneg [pos_mul_mono α] (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := -have h : a * 0 ≤ a * b, from mul_le_mul_left hb ha, +have h : a * 0 ≤ a * b, from mul_le_mul_of_nonneg_left hb ha, by rwa [mul_zero] at h +alias left.mul_nonneg ← mul_nonneg + lemma mul_nonpos_of_nonneg_of_nonpos [pos_mul_mono α] (ha : 0 ≤ a) (hb : b ≤ 0) : a * b ≤ 0 := -have h : a * b ≤ a * 0, from mul_le_mul_left hb ha, +have h : a * b ≤ a * 0, from mul_le_mul_of_nonneg_left hb ha, by rwa [mul_zero] at h /-- Assumes right covariance. -/ lemma right.mul_nonneg [mul_pos_mono α] (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := -have h : 0 * b ≤ a * b, from mul_le_mul_right ha hb, +have h : 0 * b ≤ a * b, from mul_le_mul_of_nonneg_right ha hb, by rwa [zero_mul] at h lemma mul_nonpos_of_nonpos_of_nonneg [mul_pos_mono α] (ha : a ≤ 0) (hb : 0 ≤ b) : a * b ≤ 0 := -have h : a * b ≤ 0 * b, from mul_le_mul_right ha hb, +have h : a * b ≤ 0 * b, from mul_le_mul_of_nonneg_right ha hb, by rwa [zero_mul] at h lemma lt_of_mul_lt_mul_left [pos_mul_reflect_lt α] @@ -445,12 +457,12 @@ lemma pos_of_mul_pos_right [pos_mul_reflect_lt α] (h : 0 < a * b) (ha : 0 ≤ a lt_of_mul_lt_mul_left ((mul_zero a).symm ▸ h : a * 0 < a * b) ha lemma lt_of_mul_lt_mul_right [mul_pos_reflect_lt α] - (bc : b * a < c * a) (a0 : 0 ≤ a) : - b < c := + (ab : a * c < b * c) (c0 : 0 ≤ c) : + a < b := begin - by_cases a₀ : a = 0, - { exact (lt_irrefl (0 : α) (by simpa only [a₀, mul_zero] using bc)).elim }, - { exact lt_of_mul_lt_mul_right' bc ((ne.symm a₀).le_iff_lt.mp a0) } + by_cases c₀ : c = 0, + { exact (lt_irrefl (0 : α) (by simpa only [c₀, mul_zero] using ab)).elim }, + { exact lt_of_mul_lt_mul_right' ab ((ne.symm c₀).le_iff_lt.mp c0) } end lemma pos_of_mul_pos_left [mul_pos_reflect_lt α] (h : 0 < a * b) (hb : 0 ≤ b) : @@ -463,38 +475,38 @@ lemma pos_iff_pos_of_mul_pos [pos_mul_reflect_lt α] [mul_pos_reflect_lt α] (ha lemma mul_le_mul_of_le_of_le [pos_mul_mono α] [mul_pos_mono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (a0 : 0 ≤ a) (d0 : 0 ≤ d) : a * c ≤ b * d := -(mul_le_mul_left h₂ a0).trans $ mul_le_mul_right h₁ d0 +(mul_le_mul_of_nonneg_left h₂ a0).trans $ mul_le_mul_of_nonneg_right h₁ d0 -lemma mul_le_mul_of_le_of_le' [pos_mul_mono α] [mul_pos_mono α] - (h₁ : a ≤ b) (h₂ : c ≤ d) (b0 : 0 ≤ b) (c0 : 0 ≤ c) : a * c ≤ b * d := -(mul_le_mul_right h₁ c0).trans $ mul_le_mul_left h₂ b0 +lemma mul_le_mul [pos_mul_mono α] [mul_pos_mono α] + (h₁ : a ≤ b) (h₂ : c ≤ d) (c0 : 0 ≤ c) (b0 : 0 ≤ b) : a * c ≤ b * d := +(mul_le_mul_of_nonneg_right h₁ c0).trans $ mul_le_mul_of_nonneg_left h₂ b0 lemma mul_le_of_mul_le_left [pos_mul_mono α] (h : a * b ≤ c) (hle : d ≤ b) (a0 : 0 ≤ a) : a * d ≤ c := -(mul_le_mul_left hle a0).trans h +(mul_le_mul_of_nonneg_left hle a0).trans h lemma le_mul_of_le_mul_left [pos_mul_mono α] (h : a ≤ b * c) (hle : c ≤ d) (b0 : 0 ≤ b) : a ≤ b * d := -h.trans (mul_le_mul_left hle b0) +h.trans (mul_le_mul_of_nonneg_left hle b0) lemma mul_le_of_mul_le_right [mul_pos_mono α] (h : a * b ≤ c) (hle : d ≤ a) (b0 : 0 ≤ b) : d * b ≤ c := -(mul_le_mul_right hle b0).trans h +(mul_le_mul_of_nonneg_right hle b0).trans h lemma le_mul_of_le_mul_right [mul_pos_mono α] (h : a ≤ b * c) (hle : b ≤ d) (c0 : 0 ≤ c) : a ≤ d * c := -h.trans (mul_le_mul_right hle c0) +h.trans (mul_le_mul_of_nonneg_right hle c0) -lemma mul_left_cancel_iff [pos_mul_mono_rev α] +lemma mul_left_cancel_iff_of_pos [pos_mul_mono_rev α] (a0 : 0 < a) : a * b = a * c ↔ b = c := ⟨λ h, (le_of_mul_le_mul_left' h.le a0).antisymm (le_of_mul_le_mul_left' h.ge a0), congr_arg _⟩ -lemma mul_right_cancel_iff [mul_pos_mono_rev α] +lemma mul_right_cancel_iff_of_pos [mul_pos_mono_rev α] (b0 : 0 < b) : a * b = c * b ↔ a = c := ⟨λ h, (le_of_mul_le_mul_right' h.le b0).antisymm (le_of_mul_le_mul_right' h.ge b0), congr_arg _⟩ @@ -506,9 +518,9 @@ lemma mul_eq_mul_iff_eq_and_eq [pos_mul_strict_mono α] [mul_pos_strict_mono α] begin refine ⟨λ h, _, λ h, congr_arg2 (*) h.1 h.2⟩, rcases hac.eq_or_lt with rfl | hac, - { exact ⟨rfl, (mul_left_cancel_iff a0).mp h⟩ }, + { exact ⟨rfl, (mul_left_cancel_iff_of_pos a0).mp h⟩ }, rcases eq_or_lt_of_le hbd with rfl | hbd, - { exact ⟨(mul_right_cancel_iff d0).mp h, rfl⟩ }, + { exact ⟨(mul_right_cancel_iff_of_pos d0).mp h, rfl⟩ }, exact ((mul_lt_mul_of_lt_of_lt hac hbd a0 d0).ne h).elim, end @@ -519,9 +531,9 @@ lemma mul_eq_mul_iff_eq_and_eq' [pos_mul_strict_mono α] [mul_pos_strict_mono α begin refine ⟨λ h, _, λ h, congr_arg2 (*) h.1 h.2⟩, rcases hac.eq_or_lt with rfl | hac, - { exact ⟨rfl, (mul_left_cancel_iff b0).mp h⟩ }, + { exact ⟨rfl, (mul_left_cancel_iff_of_pos b0).mp h⟩ }, rcases eq_or_lt_of_le hbd with rfl | hbd, - { exact ⟨(mul_right_cancel_iff c0).mp h, rfl⟩ }, + { exact ⟨(mul_right_cancel_iff_of_pos c0).mp h, rfl⟩ }, exact ((mul_lt_mul_of_lt_of_lt' hac hbd b0 c0).ne h).elim, end @@ -593,28 +605,28 @@ lemma le_mul_iff_one_le_right [pos_mul_mono α] [pos_mul_mono_rev α] (a0 : 0 < a) : a ≤ a * b ↔ 1 ≤ b := -iff.trans (by rw [mul_one]) (mul_le_mul_iff_left a0) +iff.trans (by rw [mul_one]) (mul_le_mul_left a0) @[simp] lemma lt_mul_iff_one_lt_right [pos_mul_strict_mono α] [pos_mul_reflect_lt α] (a0 : 0 < a) : a < a * b ↔ 1 < b := -iff.trans (by rw [mul_one]) (mul_lt_mul_iff_left a0) +iff.trans (by rw [mul_one]) (mul_lt_mul_left a0) @[simp] lemma mul_le_iff_le_one_right [pos_mul_mono α] [pos_mul_mono_rev α] (a0 : 0 < a) : a * b ≤ a ↔ b ≤ 1 := -iff.trans (by rw [mul_one]) (mul_le_mul_iff_left a0) +iff.trans (by rw [mul_one]) (mul_le_mul_left a0) @[simp] lemma mul_lt_iff_lt_one_right [pos_mul_strict_mono α] [pos_mul_reflect_lt α] (a0 : 0 < a) : a * b < a ↔ b < 1 := -iff.trans (by rw [mul_one]) (mul_lt_mul_iff_left a0) +iff.trans (by rw [mul_one]) (mul_lt_mul_left a0) /-! Lemmas of the form `a ≤ b * a ↔ 1 ≤ b` and `a * b ≤ b ↔ a ≤ 1`, which assume right covariance. -/ @@ -624,53 +636,53 @@ lemma le_mul_iff_one_le_left [mul_pos_mono α] [mul_pos_mono_rev α] (a0 : 0 < a) : a ≤ b * a ↔ 1 ≤ b := -iff.trans (by rw [one_mul]) (mul_le_mul_iff_right a0) +iff.trans (by rw [one_mul]) (mul_le_mul_right a0) @[simp] lemma lt_mul_iff_one_lt_left [mul_pos_strict_mono α] [mul_pos_reflect_lt α] (a0 : 0 < a) : a < b * a ↔ 1 < b := -iff.trans (by rw [one_mul]) (mul_lt_mul_iff_right a0) +iff.trans (by rw [one_mul]) (mul_lt_mul_right a0) @[simp] lemma mul_le_iff_le_one_left [mul_pos_mono α] [mul_pos_mono_rev α] (b0 : 0 < b) : a * b ≤ b ↔ a ≤ 1 := -iff.trans (by rw [one_mul]) (mul_le_mul_iff_right b0) +iff.trans (by rw [one_mul]) (mul_le_mul_right b0) @[simp] lemma mul_lt_iff_lt_one_left [mul_pos_strict_mono α] [mul_pos_reflect_lt α] (b0 : 0 < b) : a * b < b ↔ a < 1 := -iff.trans (by rw [one_mul]) (mul_lt_mul_iff_right b0) +iff.trans (by rw [one_mul]) (mul_lt_mul_right b0) /-! Lemmas of the form `b ≤ c → a ≤ 1 → b * a ≤ c`. -/ -- proven with `b0 : 0 ≤ b` as `mul_le_of_le_of_le_one` lemma preorder.mul_le_of_le_of_le_one [pos_mul_mono α] (bc : b ≤ c) (ha : a ≤ 1) (b0 : 0 < b) : b * a ≤ c := -calc b * a ≤ b * 1 : mul_le_mul_left' ha b0 +calc b * a ≤ b * 1 : mul_le_mul_left_of_le_of_pos ha b0 ... = b : mul_one b ... ≤ c : bc lemma mul_lt_of_le_of_lt_one [pos_mul_strict_mono α] (bc : b ≤ c) (ha : a < 1) (b0 : 0 < b) : b * a < c := -calc b * a < b * 1 : mul_lt_mul_left ha b0 +calc b * a < b * 1 : mul_lt_mul_left' ha b0 ... = b : mul_one b ... ≤ c : bc lemma mul_lt_of_lt_of_le_one [pos_mul_mono α] (bc : b < c) (ha : a ≤ 1) (b0 : 0 < b) : b * a < c := -calc b * a ≤ b * 1 : mul_le_mul_left' ha b0 +calc b * a ≤ b * 1 : mul_le_mul_left_of_le_of_pos ha b0 ... = b : mul_one b ... < c : bc lemma mul_lt_of_lt_of_lt_one [pos_mul_strict_mono α] (bc : b < c) (ha : a < 1) (b0 : 0 < b) : b * a < c := -calc b * a < b * 1 : mul_lt_mul_left ha b0 +calc b * a < b * 1 : mul_lt_mul_left' ha b0 ... = b : mul_one b ... < c : bc @@ -698,28 +710,28 @@ mul_lt_of_lt_of_lt_one ha hb a0 -- proven with `a0 : 0 ≤ a` and `c0 : 0 ≤ c` as `mul_le_of_le_of_le_one'` lemma preorder.mul_le_of_le_of_le_one' [pos_mul_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : a ≤ 1) (a0 : 0 < a) (c0 : 0 < c) : b * a ≤ c := -calc b * a ≤ c * a : mul_le_mul_right' bc a0 - ... ≤ c * 1 : mul_le_mul_left' ha c0 +calc b * a ≤ c * a : mul_le_mul_right_of_le_of_pos bc a0 + ... ≤ c * 1 : mul_le_mul_left_of_le_of_pos ha c0 ... = c : mul_one c -- proven with `c0 : 0 ≤ c` as `mul_lt_of_lt_of_le_one'` lemma preorder.mul_lt_of_lt_of_le_one' [pos_mul_mono α] [mul_pos_strict_mono α] (bc : b < c) (ha : a ≤ 1) (a0 : 0 < a) (c0 : 0 < c) : b * a < c := -calc b * a < c * a : mul_lt_mul_right bc a0 - ... ≤ c * 1 : mul_le_mul_left' ha c0 +calc b * a < c * a : mul_lt_mul_right' bc a0 + ... ≤ c * 1 : mul_le_mul_left_of_le_of_pos ha c0 ... = c : mul_one c -- proven with `a0 : 0 ≤ a` as `mul_lt_of_le_of_lt_one'` lemma preorder.mul_lt_of_le_of_lt_one' [pos_mul_strict_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : a < 1) (a0 : 0 < a) (c0 : 0 < c) : b * a < c := -calc b * a ≤ c * a : mul_le_mul_right' bc a0 - ... < c * 1 : mul_lt_mul_left ha c0 +calc b * a ≤ c * a : mul_le_mul_right_of_le_of_pos bc a0 + ... < c * 1 : mul_lt_mul_left' ha c0 ... = c : mul_one c lemma mul_lt_of_lt_of_lt_one' [pos_mul_strict_mono α] [mul_pos_strict_mono α] (bc : b < c) (ha : a < 1) (a0 : 0 < a) (c0 : 0 < c) : b * a < c := -calc b * a < c * a : mul_lt_mul_right bc a0 - ... < c * 1 : mul_lt_mul_left ha c0 +calc b * a < c * a : mul_lt_mul_right' bc a0 + ... < c * 1 : mul_lt_mul_left' ha c0 ... = c : mul_one c /-! Lemmas of the form `b ≤ c → 1 ≤ a → b ≤ c * a`. -/ @@ -729,25 +741,25 @@ lemma preorder.le_mul_of_le_of_one_le [pos_mul_mono α] (bc : b ≤ c) (ha : 1 ≤ a) (c0 : 0 < c) : b ≤ c * a := calc b ≤ c : bc ... = c * 1 : (mul_one c).symm - ... ≤ c * a : mul_le_mul_left' ha c0 + ... ≤ c * a : mul_le_mul_left_of_le_of_pos ha c0 lemma lt_mul_of_le_of_one_lt [pos_mul_strict_mono α] (bc : b ≤ c) (ha : 1 < a) (c0 : 0 < c) : b < c * a := calc b ≤ c : bc ... = c * 1 : (mul_one c).symm - ... < c * a : mul_lt_mul_left ha c0 + ... < c * a : mul_lt_mul_left' ha c0 lemma lt_mul_of_lt_of_one_le [pos_mul_mono α] (bc : b < c) (ha : 1 ≤ a) (c0 : 0 < c) : b < c * a := calc b < c : bc ... = c * 1 : (mul_one c).symm - ... ≤ c * a : mul_le_mul_left' ha c0 + ... ≤ c * a : mul_le_mul_left_of_le_of_pos ha c0 lemma lt_mul_of_lt_of_one_lt [pos_mul_strict_mono α] (bc : b < c) (ha : 1 < a) (c0 : 0 < c) : b < c * a := calc b < c : bc ... = c * 1 : (mul_one _).symm - ... < c * a : mul_lt_mul_left ha c0 + ... < c * a : mul_lt_mul_left' ha c0 -- proven with `a0 : 0 ≤ a` as `left.one_le_mul_of_le_of_le` /-- Assumes left covariance. -/ @@ -774,53 +786,53 @@ lt_mul_of_lt_of_one_lt ha hb a0 lemma preorder.le_mul_of_le_of_one_le' [pos_mul_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : 1 ≤ a) (a0 : 0 < a) (b0 : 0 < b) : b ≤ c * a := calc b = b * 1 : (mul_one b).symm - ... ≤ b * a : mul_le_mul_left' ha b0 - ... ≤ c * a : mul_le_mul_right' bc a0 + ... ≤ b * a : mul_le_mul_left_of_le_of_pos ha b0 + ... ≤ c * a : mul_le_mul_right_of_le_of_pos bc a0 -- proven with `a0 : 0 ≤ a` as `lt_mul_of_le_of_one_lt'` lemma preorder.lt_mul_of_le_of_one_lt' [pos_mul_strict_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : 1 < a) (a0 : 0 < a) (b0 : 0 < b) : b < c * a := calc b = b * 1 : (mul_one b).symm - ... < b * a : mul_lt_mul_left ha b0 - ... ≤ c * a : mul_le_mul_right' bc a0 + ... < b * a : mul_lt_mul_left' ha b0 + ... ≤ c * a : mul_le_mul_right_of_le_of_pos bc a0 -- proven with `b0 : 0 ≤ b` as `lt_mul_of_lt_of_one_le'` lemma preorder.lt_mul_of_lt_of_one_le' [pos_mul_mono α] [mul_pos_strict_mono α] (bc : b < c) (ha : 1 ≤ a) (a0 : 0 < a) (b0 : 0 < b) : b < c * a := calc b = b * 1 : (mul_one b).symm - ... ≤ b * a : mul_le_mul_left' ha b0 - ... < c * a : mul_lt_mul_right bc a0 + ... ≤ b * a : mul_le_mul_left_of_le_of_pos ha b0 + ... < c * a : mul_lt_mul_right' bc a0 lemma lt_mul_of_lt_of_one_lt' [pos_mul_strict_mono α] [mul_pos_strict_mono α] (bc : b < c) (ha : 1 < a) (a0 : 0 < a) (b0 : 0 < b) : b < c * a := calc b = b * 1 : (mul_one b).symm - ... < b * a : mul_lt_mul_left ha b0 - ... < c * a : mul_lt_mul_right bc a0 + ... < b * a : mul_lt_mul_left' ha b0 + ... < c * a : mul_lt_mul_right' bc a0 /-! Lemmas of the form `a ≤ 1 → b ≤ c → a * b ≤ c`. -/ -- proven with `b0 : 0 ≤ b` as `mul_le_of_le_one_of_le` lemma preorder.mul_le_of_le_one_of_le [mul_pos_mono α] (ha : a ≤ 1) (bc : b ≤ c) (b0 : 0 < b) : a * b ≤ c := -calc a * b ≤ 1 * b : mul_le_mul_right' ha b0 +calc a * b ≤ 1 * b : mul_le_mul_right_of_le_of_pos ha b0 ... = b : one_mul b ... ≤ c : bc lemma mul_lt_of_lt_one_of_le [mul_pos_strict_mono α] (ha : a < 1) (bc : b ≤ c) (b0 : 0 < b) : a * b < c := -calc a * b < 1 * b : mul_lt_mul_right ha b0 +calc a * b < 1 * b : mul_lt_mul_right' ha b0 ... = b : one_mul b ... ≤ c : bc lemma mul_lt_of_le_one_of_lt [mul_pos_mono α] (ha : a ≤ 1) (hb : b < c) (b0 : 0 < b) : a * b < c := -calc a * b ≤ 1 * b : mul_le_mul_right' ha b0 +calc a * b ≤ 1 * b : mul_le_mul_right_of_le_of_pos ha b0 ... = b : one_mul b ... < c : hb lemma mul_lt_of_lt_one_of_lt [mul_pos_strict_mono α] (ha : a < 1) (bc : b < c) (b0 : 0 < b) : a * b < c := -calc a * b < 1 * b : mul_lt_mul_right ha b0 +calc a * b < 1 * b : mul_lt_mul_right' ha b0 ... = b : one_mul b ... < c : bc @@ -848,28 +860,28 @@ mul_lt_of_lt_one_of_lt ha hb b0 -- proven with `a0 : 0 ≤ a` and `c0 : 0 ≤ c` as `mul_le_of_le_one_of_le'` lemma preorder.mul_le_of_le_one_of_le' [pos_mul_mono α] [mul_pos_mono α] (ha : a ≤ 1) (bc : b ≤ c) (a0 : 0 < a) (c0 : 0 < c) : a * b ≤ c := -calc a * b ≤ a * c : mul_le_mul_left' bc a0 - ... ≤ 1 * c : mul_le_mul_right' ha c0 +calc a * b ≤ a * c : mul_le_mul_left_of_le_of_pos bc a0 + ... ≤ 1 * c : mul_le_mul_right_of_le_of_pos ha c0 ... = c : one_mul c -- proven with `a0 : 0 ≤ a` as `mul_lt_of_lt_one_of_le'` lemma preorder.mul_lt_of_lt_one_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] (ha : a < 1) (bc : b ≤ c) (a0 : 0 < a) (c0 : 0 < c) : a * b < c := -calc a * b ≤ a * c : mul_le_mul_left' bc a0 - ... < 1 * c : mul_lt_mul_right ha c0 +calc a * b ≤ a * c : mul_le_mul_left_of_le_of_pos bc a0 + ... < 1 * c : mul_lt_mul_right' ha c0 ... = c : one_mul c -- proven with `c0 : 0 ≤ c` as `mul_lt_of_le_one_of_lt'` lemma preorder.mul_lt_of_le_one_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] (ha : a ≤ 1) (bc : b < c) (a0 : 0 < a) (c0 : 0 < c) : a * b < c := -calc a * b < a * c : mul_lt_mul_left bc a0 - ... ≤ 1 * c : mul_le_mul_right' ha c0 +calc a * b < a * c : mul_lt_mul_left' bc a0 + ... ≤ 1 * c : mul_le_mul_right_of_le_of_pos ha c0 ... = c : one_mul c lemma mul_lt_of_lt_one_of_lt' [pos_mul_strict_mono α] [mul_pos_strict_mono α] (ha : a < 1) (bc : b < c) (a0 : 0 < a) (c0 : 0 < c) : a * b < c := -calc a * b < a * c : mul_lt_mul_left bc a0 - ... < 1 * c : mul_lt_mul_right ha c0 +calc a * b < a * c : mul_lt_mul_left' bc a0 + ... < 1 * c : mul_lt_mul_right' ha c0 ... = c : one_mul c /-! Lemmas of the form `1 ≤ a → b ≤ c → b ≤ a * c`. -/ @@ -879,25 +891,25 @@ lemma preorder.le_mul_of_one_le_of_le [mul_pos_mono α] (ha : 1 ≤ a) (bc : b ≤ c) (c0 : 0 < c) : b ≤ a * c := calc b ≤ c : bc ... = 1 * c : (one_mul c).symm - ... ≤ a * c : mul_le_mul_right' ha c0 + ... ≤ a * c : mul_le_mul_right_of_le_of_pos ha c0 lemma lt_mul_of_one_lt_of_le [mul_pos_strict_mono α] (ha : 1 < a) (bc : b ≤ c) (c0 : 0 < c) : b < a * c := calc b ≤ c : bc ... = 1 * c : (one_mul c).symm - ... < a * c : mul_lt_mul_right ha c0 + ... < a * c : mul_lt_mul_right' ha c0 lemma lt_mul_of_one_le_of_lt [mul_pos_mono α] (ha : 1 ≤ a) (bc : b < c) (c0 : 0 < c) : b < a * c := calc b < c : bc ... = 1 * c : (one_mul c).symm - ... ≤ a * c : mul_le_mul_right' ha c0 + ... ≤ a * c : mul_le_mul_right_of_le_of_pos ha c0 lemma lt_mul_of_one_lt_of_lt [mul_pos_strict_mono α] (ha : 1 < a) (bc : b < c) (c0 : 0 < c) : b < a * c := calc b < c : bc ... = 1 * c : (one_mul c).symm - ... < a * c : mul_lt_mul_right ha c0 + ... < a * c : mul_lt_mul_right' ha c0 -- proven with `b0 : 0 ≤ b` as `right.one_le_mul_of_le_of_le` /-- Assumes right covariance. -/ @@ -924,28 +936,28 @@ lt_mul_of_one_lt_of_lt ha hb b0 lemma preorder.le_mul_of_one_le_of_le' [pos_mul_mono α] [mul_pos_mono α] (ha : 1 ≤ a) (bc : b ≤ c) (a0 : 0 < a) (b0 : 0 < b) : b ≤ a * c := calc b = 1 * b : (one_mul b).symm - ... ≤ a * b : mul_le_mul_right' ha b0 - ... ≤ a * c : mul_le_mul_left' bc a0 + ... ≤ a * b : mul_le_mul_right_of_le_of_pos ha b0 + ... ≤ a * c : mul_le_mul_left_of_le_of_pos bc a0 -- proven with `a0 : 0 ≤ a` as `lt_mul_of_one_lt_of_le'` lemma preorder.lt_mul_of_one_lt_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] (ha : 1 < a) (bc : b ≤ c) (a0 : 0 < a) (b0 : 0 < b) : b < a * c := calc b = 1 * b : (one_mul b).symm - ... < a * b : mul_lt_mul_right ha b0 - ... ≤ a * c : mul_le_mul_left' bc a0 + ... < a * b : mul_lt_mul_right' ha b0 + ... ≤ a * c : mul_le_mul_left_of_le_of_pos bc a0 -- proven with `b0 : 0 ≤ b` as `lt_mul_of_one_le_of_lt'` lemma preorder.lt_mul_of_one_le_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] (ha : 1 ≤ a) (bc : b < c) (a0 : 0 < a) (b0 : 0 < b) : b < a * c := calc b = 1 * b : (one_mul b).symm - ... ≤ a * b : mul_le_mul_right' ha b0 - ... < a * c : mul_lt_mul_left bc a0 + ... ≤ a * b : mul_le_mul_right_of_le_of_pos ha b0 + ... < a * c : mul_lt_mul_left' bc a0 lemma lt_mul_of_one_lt_of_lt' [pos_mul_strict_mono α] [mul_pos_strict_mono α] (ha : 1 < a) (bc : b < c) (a0 : 0 < a) (b0 : 0 < b) : b < a * c := calc b = 1 * b : (one_mul b).symm - ... < a * b : mul_lt_mul_right ha b0 - ... < a * c : mul_lt_mul_left bc a0 + ... < a * b : mul_lt_mul_right' ha b0 + ... < a * c : mul_lt_mul_left' bc a0 -- proven with `a0 : 0 ≤ a` as `mul_le_of_le_one_right` lemma preorder.mul_le_of_le_one_right [pos_mul_mono α] (h : b ≤ 1) (a0 : 0 < a) : @@ -967,19 +979,19 @@ lemma preorder.le_mul_of_one_le_left [mul_pos_mono α] (h : 1 ≤ a) (b0 : 0 < b b ≤ a * b := preorder.le_mul_of_one_le_of_le h le_rfl b0 -lemma mul_lt_of_lt_one_right [pos_mul_strict_mono α] (h : b < 1) (a0 : 0 < a) : +lemma mul_lt_of_lt_one_right [pos_mul_strict_mono α] (a0 : 0 < a) (h : b < 1) : a * b < a := mul_lt_of_le_of_lt_one le_rfl h a0 -lemma lt_mul_of_one_lt_right [pos_mul_strict_mono α] (h : 1 < b) (a0 : 0 < a) : +lemma lt_mul_of_one_lt_right [pos_mul_strict_mono α] (a0 : 0 < a) (h : 1 < b) : a < a * b := lt_mul_of_le_of_one_lt le_rfl h a0 -lemma mul_lt_of_lt_one_left [mul_pos_strict_mono α] (h : a < 1) (b0 : 0 < b) : +lemma mul_lt_of_lt_one_left [mul_pos_strict_mono α] (b0 : 0 < b) (h : a < 1) : a * b < b := mul_lt_of_lt_one_of_le h le_rfl b0 -lemma lt_mul_of_one_lt_left [mul_pos_strict_mono α] (h : 1 < a) (b0 : 0 < b) : +lemma lt_mul_of_one_lt_left [mul_pos_strict_mono α] (b0 : 0 < b) (h : 1 < a) : b < a * b := lt_mul_of_one_lt_of_le h le_rfl b0 @@ -1039,7 +1051,7 @@ begin by_cases h : a < 1, { use a, have : a*a < a*1, - exact mul_lt_mul_left h a0, + exact mul_lt_mul_left' h a0, rw mul_one at this, exact le_of_lt this }, { use 1, @@ -1069,22 +1081,22 @@ lemma left.mul_le_one_of_le_of_le [pos_mul_mono α] (ha : a ≤ 1) (hb : b ≤ 1) (a0 : 0 ≤ a) : a * b ≤ 1 := mul_le_of_le_of_le_one ha hb a0 -lemma mul_le_of_le_of_le_one' [pos_mul_mono α] [mul_pos_mono α] - (bc : b ≤ c) (ha : a ≤ 1) (a0 : 0 ≤ a) (c0 : 0 ≤ c) : b * a ≤ c := -calc b * a ≤ c * a : mul_le_mul_right bc a0 - ... ≤ c * 1 : mul_le_mul_left ha c0 +lemma mul_nonneg_le_one_le [pos_mul_mono α] [mul_pos_mono α] + (c0 : 0 ≤ c) (bc : b ≤ c) (a0 : 0 ≤ a) (ha : a ≤ 1) : b * a ≤ c := +calc b * a ≤ c * a : mul_le_mul_of_nonneg_right bc a0 + ... ≤ c * 1 : mul_le_mul_of_nonneg_left ha c0 ... = c : mul_one c lemma mul_lt_of_lt_of_le_one' [pos_mul_mono α] [mul_pos_strict_mono α] (bc : b < c) (ha : a ≤ 1) (a0 : 0 < a) (c0 : 0 ≤ c) : b * a < c := -calc b * a < c * a : mul_lt_mul_right bc a0 - ... ≤ c * 1 : mul_le_mul_left ha c0 +calc b * a < c * a : mul_lt_mul_right' bc a0 + ... ≤ c * 1 : mul_le_mul_of_nonneg_left ha c0 ... = c : mul_one c lemma mul_lt_of_le_of_lt_one' [pos_mul_strict_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : a < 1) (a0 : 0 ≤ a) (c0 : 0 < c) : b * a < c := -calc b * a ≤ c * a : mul_le_mul_right bc a0 - ... < c * 1 : mul_lt_mul_left ha c0 +calc b * a ≤ c * a : mul_le_mul_of_nonneg_right bc a0 + ... < c * 1 : mul_lt_mul_left' ha c0 ... = c : mul_one c /-! Lemmas of the form `b ≤ c → 1 ≤ a → b ≤ c * a`. -/ @@ -1102,20 +1114,20 @@ le_mul_of_le_of_one_le ha hb a0 lemma le_mul_of_le_of_one_le' [pos_mul_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : 1 ≤ a) (a0 : 0 ≤ a) (b0 : 0 ≤ b) : b ≤ c * a := calc b = b * 1 : (mul_one b).symm - ... ≤ b * a : mul_le_mul_left ha b0 - ... ≤ c * a : mul_le_mul_right bc a0 + ... ≤ b * a : mul_le_mul_of_nonneg_left ha b0 + ... ≤ c * a : mul_le_mul_of_nonneg_right bc a0 lemma lt_mul_of_le_of_one_lt' [pos_mul_strict_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : 1 < a) (a0 : 0 ≤ a) (b0 : 0 < b) : b < c * a := calc b = b * 1 : (mul_one b).symm - ... < b * a : mul_lt_mul_left ha b0 - ... ≤ c * a : mul_le_mul_right bc a0 + ... < b * a : mul_lt_mul_left' ha b0 + ... ≤ c * a : mul_le_mul_of_nonneg_right bc a0 lemma lt_mul_of_lt_of_one_le' [pos_mul_mono α] [mul_pos_strict_mono α] (bc : b < c) (ha : 1 ≤ a) (a0 : 0 < a) (b0 : 0 ≤ b) : b < c * a := calc b = b * 1 : (mul_one b).symm - ... ≤ b * a : mul_le_mul_left ha b0 - ... < c * a : mul_lt_mul_right bc a0 + ... ≤ b * a : mul_le_mul_of_nonneg_left ha b0 + ... < c * a : mul_lt_mul_right' bc a0 /-! Lemmas of the form `a ≤ 1 → b ≤ c → a * b ≤ c`. -/ @@ -1131,20 +1143,20 @@ preorder.mul_le_of_le_one_of_le ha hb b0 lemma mul_le_of_le_one_of_le' [pos_mul_mono α] [mul_pos_mono α] (ha : a ≤ 1) (bc : b ≤ c) (a0 : 0 ≤ a) (c0 : 0 ≤ c) : a * b ≤ c := -calc a * b ≤ a * c : mul_le_mul_left bc a0 - ... ≤ 1 * c : mul_le_mul_right ha c0 +calc a * b ≤ a * c : mul_le_mul_of_nonneg_left bc a0 + ... ≤ 1 * c : mul_le_mul_of_nonneg_right ha c0 ... = c : one_mul c lemma mul_lt_of_lt_one_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] (ha : a < 1) (bc : b ≤ c) (a0 : 0 ≤ a) (c0 : 0 < c) : a * b < c := -calc a * b ≤ a * c : mul_le_mul_left bc a0 - ... < 1 * c : mul_lt_mul_right ha c0 +calc a * b ≤ a * c : mul_le_mul_of_nonneg_left bc a0 + ... < 1 * c : mul_lt_mul_right' ha c0 ... = c : one_mul c lemma mul_lt_of_le_one_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] (ha : a ≤ 1) (bc : b < c) (a0 : 0 < a) (c0 : 0 ≤ c) : a * b < c := -calc a * b < a * c : mul_lt_mul_left bc a0 - ... ≤ 1 * c : mul_le_mul_right ha c0 +calc a * b < a * c : mul_lt_mul_left' bc a0 + ... ≤ 1 * c : mul_le_mul_of_nonneg_right ha c0 ... = c : one_mul c /-! Lemmas of the form `1 ≤ a → b ≤ c → b ≤ a * c`. -/ @@ -1162,50 +1174,50 @@ le_mul_of_one_le_of_le ha hb b0 lemma le_mul_of_one_le_of_le' [pos_mul_mono α] [mul_pos_mono α] (ha : 1 ≤ a) (bc : b ≤ c) (a0 : 0 ≤ a) (b0 : 0 ≤ b) : b ≤ a * c := calc b = 1 * b : (one_mul b).symm - ... ≤ a * b : mul_le_mul_right ha b0 - ... ≤ a * c : mul_le_mul_left bc a0 + ... ≤ a * b : mul_le_mul_of_nonneg_right ha b0 + ... ≤ a * c : mul_le_mul_of_nonneg_left bc a0 lemma lt_mul_of_one_lt_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] (ha : 1 < a) (bc : b ≤ c) (a0 : 0 ≤ a) (b0 : 0 < b) : b < a * c := calc b = 1 * b : (one_mul b).symm - ... < a * b : mul_lt_mul_right ha b0 - ... ≤ a * c : mul_le_mul_left bc a0 + ... < a * b : mul_lt_mul_right' ha b0 + ... ≤ a * c : mul_le_mul_of_nonneg_left bc a0 lemma lt_mul_of_one_le_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] (ha : 1 ≤ a) (bc : b < c) (a0 : 0 < a) (b0 : 0 ≤ b) : b < a * c := calc b = 1 * b : (one_mul b).symm - ... ≤ a * b : mul_le_mul_right ha b0 - ... < a * c : mul_lt_mul_left bc a0 + ... ≤ a * b : mul_le_mul_of_nonneg_right ha b0 + ... < a * c : mul_lt_mul_left' bc a0 -lemma mul_le_of_le_one_right [pos_mul_mono α] (h : b ≤ 1) (a0 : 0 ≤ a) : +lemma mul_le_of_le_one_right [pos_mul_mono α] (a0 : 0 ≤ a) (h : b ≤ 1) : a * b ≤ a := mul_le_of_le_of_le_one le_rfl h a0 -lemma le_mul_of_one_le_right [pos_mul_mono α] (h : 1 ≤ b) (a0 : 0 ≤ a) : +lemma le_mul_of_one_le_right [pos_mul_mono α] (a0 : 0 ≤ a) (h : 1 ≤ b) : a ≤ a * b := le_mul_of_le_of_one_le le_rfl h a0 -lemma mul_le_of_le_one_left [mul_pos_mono α] (h : a ≤ 1) (b0 : 0 ≤ b) : +lemma mul_le_of_le_one_left [mul_pos_mono α] (b0 : 0 ≤ b) (h : a ≤ 1) : a * b ≤ b := mul_le_of_le_one_of_le h le_rfl b0 -lemma le_mul_of_one_le_left [mul_pos_mono α] (h : 1 ≤ a) (b0 : 0 ≤ b) : +lemma le_mul_of_one_le_left [mul_pos_mono α] (b0 : 0 ≤ b) (h : 1 ≤ a) : b ≤ a * b := le_mul_of_one_le_of_le h le_rfl b0 -lemma le_of_mul_le_of_one_le_left [pos_mul_mono α] +lemma le_of_mul_le_of_one_le_of_nonneg_left [pos_mul_mono α] (h : a * b ≤ c) (hle : 1 ≤ b) (a0 : 0 ≤ a) : a ≤ c := a0.lt_or_eq.elim (preorder.le_of_mul_le_of_one_le_left h hle) (λ ha, by simpa only [← ha, zero_mul] using h) -lemma le_of_le_mul_of_le_one_left [pos_mul_mono α] +lemma le_of_le_mul_of_le_one_of_nonneg_left [pos_mul_mono α] (h : a ≤ b * c) (hle : c ≤ 1) (b0 : 0 ≤ b) : a ≤ b := b0.lt_or_eq.elim (preorder.le_of_le_mul_of_le_one_left h hle) (λ hb, by simpa only [← hb, zero_mul] using h) -lemma le_of_mul_le_of_one_le_right [mul_pos_mono α] +lemma le_of_mul_le_of_one_le_nonneg_right [mul_pos_mono α] (h : a * b ≤ c) (hle : 1 ≤ a) (b0 : 0 ≤ b) : b ≤ c := b0.lt_or_eq.elim (preorder.le_of_mul_le_of_one_le_right h hle) @@ -1238,13 +1250,13 @@ section partial_order variables [partial_order α] lemma pos_mul_mono.to_pos_mul_strict_mono [pos_mul_mono α] : pos_mul_strict_mono α := -⟨λ x a b h, lt_of_le_of_ne (mul_le_mul_left' h.le x.2) (h.ne ∘ mul_left_cancel₀ x.2.ne.symm)⟩ +⟨λ x a b h, (mul_le_mul_left_of_le_of_pos h.le x.2).lt_of_ne (h.ne ∘ mul_left_cancel₀ x.2.ne')⟩ lemma pos_mul_mono_iff_pos_mul_strict_mono : pos_mul_mono α ↔ pos_mul_strict_mono α := ⟨@pos_mul_mono.to_pos_mul_strict_mono α _ _, @zero_lt.pos_mul_strict_mono.to_pos_mul_mono α _ _ _⟩ lemma mul_pos_mono.to_mul_pos_strict_mono [mul_pos_mono α] : mul_pos_strict_mono α := -⟨λ x a b h, lt_of_le_of_ne (mul_le_mul_right' h.le x.2) (h.ne ∘ mul_right_cancel₀ x.2.ne.symm)⟩ +⟨λ x a b h, (mul_le_mul_right_of_le_of_pos h.le x.2).lt_of_ne (h.ne ∘ mul_right_cancel₀ x.2.ne')⟩ lemma mul_pos_mono_iff_mul_pos_strict_mono : mul_pos_mono α ↔ mul_pos_strict_mono α := ⟨@mul_pos_mono.to_mul_pos_strict_mono α _ _, @zero_lt.mul_pos_strict_mono.to_mul_pos_mono α _ _ _⟩ @@ -1295,3 +1307,5 @@ by simp ! only [mul_comm] end comm_semigroup_has_zero end zero_lt + +export zero_lt diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 25a42974aacfd..2df672b7e01a7 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -165,12 +165,6 @@ alias zero_lt_four ← four_pos end nontrivial -lemma mul_lt_of_lt_one_left (hb : 0 < b) (ha : a < 1) : a * b < b := -(mul_lt_mul_of_pos_right ha hb).trans_le (one_mul _).le - -lemma mul_lt_of_lt_one_right (ha : 0 < a) (hb : b < 1) : a * b < a := -(mul_lt_mul_of_pos_left hb ha).trans_le (mul_one _).le - -- See Note [decidable namespace] protected lemma decidable.mul_le_mul_of_nonneg_left [@decidable_rel α (≤)] (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c * a ≤ c * b := @@ -180,9 +174,6 @@ begin exact (mul_lt_mul_of_pos_left (h₁.lt_of_not_le ba) (h₂.lt_of_not_le c0)).le, end -lemma mul_le_mul_of_nonneg_left : a ≤ b → 0 ≤ c → c * a ≤ c * b := -by classical; exact decidable.mul_le_mul_of_nonneg_left - -- See Note [decidable namespace] protected lemma decidable.mul_le_mul_of_nonneg_right [@decidable_rel α (≤)] (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a * c ≤ b * c := @@ -192,9 +183,6 @@ begin exact (mul_lt_mul_of_pos_right (h₁.lt_of_not_le ba) (h₂.lt_of_not_le c0)).le, end -lemma mul_le_mul_of_nonneg_right : a ≤ b → 0 ≤ c → a * c ≤ b * c := -by classical; exact decidable.mul_le_mul_of_nonneg_right - -- TODO: there are four variations, depending on which variables we assume to be nonneg -- See Note [decidable namespace] protected lemma decidable.mul_le_mul [@decidable_rel α (≤)] @@ -203,27 +191,18 @@ calc a * b ≤ c * b : decidable.mul_le_mul_of_nonneg_right hac nn_b ... ≤ c * d : decidable.mul_le_mul_of_nonneg_left hbd nn_c -lemma mul_le_mul : a ≤ c → b ≤ d → 0 ≤ b → 0 ≤ c → a * b ≤ c * d := -by classical; exact decidable.mul_le_mul - -- See Note [decidable namespace] protected lemma decidable.mul_nonneg_le_one_le {α : Type*} [ordered_semiring α] [@decidable_rel α (≤)] {a b c : α} (h₁ : 0 ≤ c) (h₂ : a ≤ c) (h₃ : 0 ≤ b) (h₄ : b ≤ 1) : a * b ≤ c := by simpa only [mul_one] using decidable.mul_le_mul h₂ h₄ h₃ h₁ -lemma mul_nonneg_le_one_le {α : Type*} [ordered_semiring α] {a b c : α} : - 0 ≤ c → a ≤ c → 0 ≤ b → b ≤ 1 → a * b ≤ c := -by classical; exact decidable.mul_nonneg_le_one_le - -- See Note [decidable namespace] protected lemma decidable.mul_nonneg [@decidable_rel α (≤)] (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := have h : 0 * b ≤ a * b, from decidable.mul_le_mul_of_nonneg_right ha hb, by rwa [zero_mul] at h -lemma mul_nonneg : 0 ≤ a → 0 ≤ b → 0 ≤ a * b := by classical; exact decidable.mul_nonneg - @[simp] theorem pow_nonneg (H : 0 ≤ a) : ∀ (n : ℕ), 0 ≤ a ^ n | 0 := by { rw pow_zero, exact zero_le_one} | (n+1) := by { rw pow_succ, exact mul_nonneg H (pow_nonneg _) } @@ -234,18 +213,12 @@ protected lemma decidable.mul_nonpos_of_nonneg_of_nonpos [@decidable_rel α (≤ have h : a * b ≤ a * 0, from decidable.mul_le_mul_of_nonneg_left hb ha, by rwa mul_zero at h -lemma mul_nonpos_of_nonneg_of_nonpos : 0 ≤ a → b ≤ 0 → a * b ≤ 0 := - by classical; exact decidable.mul_nonpos_of_nonneg_of_nonpos - -- See Note [decidable namespace] protected lemma decidable.mul_nonpos_of_nonpos_of_nonneg [@decidable_rel α (≤)] (ha : a ≤ 0) (hb : 0 ≤ b) : a * b ≤ 0 := have h : a * b ≤ 0 * b, from decidable.mul_le_mul_of_nonneg_right ha hb, by rwa zero_mul at h -lemma mul_nonpos_of_nonpos_of_nonneg : a ≤ 0 → 0 ≤ b → a * b ≤ 0 := -by classical; exact decidable.mul_nonpos_of_nonpos_of_nonneg - -- See Note [decidable namespace] protected lemma decidable.mul_lt_mul [@decidable_rel α (≤)] (hac : a < c) (hbd : b ≤ d) (pos_b : 0 < b) (nn_c : 0 ≤ c) : a * b < c * d := @@ -266,22 +239,10 @@ calc lemma mul_lt_mul' : a ≤ c → b < d → 0 ≤ b → 0 < c → a * b < c * d := by classical; exact decidable.mul_lt_mul' -lemma mul_pos (ha : 0 < a) (hb : 0 < b) : 0 < a * b := -have h : 0 * b < a * b, from mul_lt_mul_of_pos_right ha hb, -by rwa zero_mul at h - @[simp] theorem pow_pos (H : 0 < a) : ∀ (n : ℕ), 0 < a ^ n | 0 := by { nontriviality, rw pow_zero, exact zero_lt_one } | (n+1) := by { rw pow_succ, exact mul_pos H (pow_pos _) } -lemma mul_neg_of_pos_of_neg (ha : 0 < a) (hb : b < 0) : a * b < 0 := -have h : a * b < a * 0, from mul_lt_mul_of_pos_left hb ha, -by rwa mul_zero at h - -lemma mul_neg_of_neg_of_pos (ha : a < 0) (hb : 0 < b) : a * b < 0 := -have h : a * b < 0 * b, from mul_lt_mul_of_pos_right ha hb, -by rwa zero_mul at h - -- See Note [decidable namespace] protected lemma decidable.mul_self_lt_mul_self [@decidable_rel α (≤)] (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b := @@ -323,47 +284,29 @@ protected lemma decidable.le_mul_of_one_le_right [@decidable_rel α (≤)] suffices b * 1 ≤ b * a, by rwa mul_one at this, decidable.mul_le_mul_of_nonneg_left h hb -lemma le_mul_of_one_le_right : 0 ≤ b → 1 ≤ a → b ≤ b * a := -by classical; exact decidable.le_mul_of_one_le_right - -- See Note [decidable namespace] protected lemma decidable.le_mul_of_one_le_left [@decidable_rel α (≤)] (hb : 0 ≤ b) (h : 1 ≤ a) : b ≤ a * b := suffices 1 * b ≤ a * b, by rwa one_mul at this, decidable.mul_le_mul_of_nonneg_right h hb -lemma le_mul_of_one_le_left : 0 ≤ b → 1 ≤ a → b ≤ a * b := -by classical; exact decidable.le_mul_of_one_le_left - -- See Note [decidable namespace] protected lemma decidable.lt_mul_of_one_lt_right [@decidable_rel α (≤)] (hb : 0 < b) (h : 1 < a) : b < b * a := suffices b * 1 < b * a, by rwa mul_one at this, decidable.mul_lt_mul' le_rfl h zero_le_one hb -lemma lt_mul_of_one_lt_right : 0 < b → 1 < a → b < b * a := -by classical; exact decidable.lt_mul_of_one_lt_right - -- See Note [decidable namespace] protected lemma decidable.lt_mul_of_one_lt_left [@decidable_rel α (≤)] (hb : 0 < b) (h : 1 < a) : b < a * b := suffices 1 * b < a * b, by rwa one_mul at this, decidable.mul_lt_mul h le_rfl hb (zero_le_one.trans h.le) -lemma lt_mul_of_one_lt_left : 0 < b → 1 < a → b < a * b := -by classical; exact decidable.lt_mul_of_one_lt_left - -lemma lt_two_mul_self [nontrivial α] (ha : 0 < a) : a < 2 * a := -lt_mul_of_one_lt_left ha one_lt_two - -lemma lt_mul_left (hn : 0 < a) (hm : 1 < b) : a < b * a := -by { convert mul_lt_mul_of_pos_right hm hn, rw one_mul } - -lemma lt_mul_right (hn : 0 < a) (hm : 1 < b) : a < a * b := -by { convert mul_lt_mul_of_pos_left hm hn, rw mul_one } +lemma lt_two_mul_self (ha : 0 < a) : a < 2 * a := +lt_mul_of_one_lt_left ha (@one_lt_two α _ (nontrivial_of_ne 0 a ha.ne)) lemma lt_mul_self (hn : 1 < a) : a < a * a := -lt_mul_left (hn.trans_le' zero_le_one) hn +lt_mul_of_one_lt_left (hn.trans_le' zero_le_one) hn -- See Note [decidable namespace] protected lemma decidable.add_le_mul_two_add [@decidable_rel α (≤)] {a b : α} @@ -566,18 +509,12 @@ protected lemma decidable.mul_le_of_le_one_right [@decidable_rel α (≤)] calc a * b ≤ a * 1 : decidable.mul_le_mul_of_nonneg_left hb1 ha ... = a : mul_one a -lemma mul_le_of_le_one_right : 0 ≤ a → b ≤ 1 → a * b ≤ a := -by classical; exact decidable.mul_le_of_le_one_right - -- See Note [decidable namespace] protected lemma decidable.mul_le_of_le_one_left [@decidable_rel α (≤)] (hb : 0 ≤ b) (ha1 : a ≤ 1) : a * b ≤ b := calc a * b ≤ 1 * b : decidable.mul_le_mul ha1 le_rfl hb zero_le_one ... = b : one_mul b -lemma mul_le_of_le_one_left : 0 ≤ b → a ≤ 1 → a * b ≤ b := -by classical; exact decidable.mul_le_of_le_one_left - -- See Note [decidable namespace] protected lemma decidable.mul_lt_one_of_nonneg_of_lt_one_left [@decidable_rel α (≤)] (ha0 : 0 ≤ a) (ha : a < 1) (hb : b ≤ 1) : a * b < 1 := @@ -682,29 +619,12 @@ local attribute [instance] linear_ordered_semiring.decidable_le -- with only a `linear_ordered_semiring` typeclass argument. lemma zero_lt_one' : 0 < (1 : α) := zero_lt_one -lemma lt_of_mul_lt_mul_left (h : c * a < c * b) (hc : 0 ≤ c) : a < b := -(decidable.monotone_mul_left_of_nonneg hc).reflect_lt h - -lemma lt_of_mul_lt_mul_right (h : a * c < b * c) (hc : 0 ≤ c) : a < b := -(decidable.monotone_mul_right_of_nonneg hc).reflect_lt h - lemma le_of_mul_le_mul_left (h : c * a ≤ c * b) (hc : 0 < c) : a ≤ b := (strict_mono_mul_left_of_pos hc).le_iff_le.1 h lemma le_of_mul_le_mul_right (h : a * c ≤ b * c) (hc : 0 < c) : a ≤ b := (strict_mono_mul_right_of_pos hc).le_iff_le.1 h -lemma pos_and_pos_or_neg_and_neg_of_mul_pos (hab : 0 < a * b) : - (0 < a ∧ 0 < b) ∨ (a < 0 ∧ b < 0) := -begin - rcases lt_trichotomy 0 a with (ha|rfl|ha), - { refine or.inl ⟨ha, lt_imp_lt_of_le_imp_le (λ hb, _) hab⟩, - exact decidable.mul_nonpos_of_nonneg_of_nonpos ha.le hb }, - { rw [zero_mul] at hab, exact hab.false.elim }, - { refine or.inr ⟨ha, lt_imp_lt_of_le_imp_le (λ hb, _) hab⟩, - exact decidable.mul_nonpos_of_nonpos_of_nonneg ha.le hb } -end - lemma nonneg_and_nonneg_or_nonpos_and_nonpos_of_mul_nnonneg (hab : 0 ≤ a * b) : (0 ≤ a ∧ 0 ≤ b) ∨ (a ≤ 0 ∧ b ≤ 0) := begin @@ -715,24 +635,6 @@ begin mul_neg_of_neg_of_pos ha (nab ha.le)] end -lemma pos_of_mul_pos_left (h : 0 < a * b) (hb : 0 ≤ b) : 0 < a := -((pos_and_pos_or_neg_and_neg_of_mul_pos h).resolve_right $ λ h, h.2.not_le hb).1 - -lemma pos_of_mul_pos_right (h : 0 < a * b) (ha : 0 ≤ a) : 0 < b := -((pos_and_pos_or_neg_and_neg_of_mul_pos h).resolve_right $ λ h, h.1.not_le ha).2 - -lemma pos_iff_pos_of_mul_pos (hab : 0 < a * b) : 0 < a ↔ 0 < b := -⟨pos_of_mul_pos_right hab ∘ le_of_lt, pos_of_mul_pos_left hab ∘ le_of_lt⟩ - -lemma neg_of_mul_pos_left (h : 0 < a * b) (ha : b ≤ 0) : a < 0 := -((pos_and_pos_or_neg_and_neg_of_mul_pos h).resolve_left $ λ h, h.2.not_le ha).1 - -lemma neg_of_mul_pos_right (h : 0 < a * b) (ha : a ≤ 0) : b < 0 := -((pos_and_pos_or_neg_and_neg_of_mul_pos h).resolve_left $ λ h, h.1.not_le ha).2 - -lemma neg_iff_neg_of_mul_pos (hab : 0 < a * b) : a < 0 ↔ b < 0 := -⟨neg_of_mul_pos_right hab ∘ le_of_lt, neg_of_mul_pos_left hab ∘ le_of_lt⟩ - lemma nonneg_of_mul_nonneg_left (h : 0 ≤ a * b) (hb : 0 < b) : 0 ≤ a := le_of_not_gt $ λ ha, (mul_neg_of_neg_of_pos ha hb).not_le h @@ -751,30 +653,12 @@ le_of_not_gt (assume ha : a > 0, (mul_pos ha hb).not_le h) lemma nonpos_of_mul_nonpos_right (h : a * b ≤ 0) (ha : 0 < a) : b ≤ 0 := le_of_not_gt (assume hb : b > 0, (mul_pos ha hb).not_le h) -@[simp] lemma mul_le_mul_left (h : 0 < c) : c * a ≤ c * b ↔ a ≤ b := -(strict_mono_mul_left_of_pos h).le_iff_le - -@[simp] lemma mul_le_mul_right (h : 0 < c) : a * c ≤ b * c ↔ a ≤ b := -(strict_mono_mul_right_of_pos h).le_iff_le - -@[simp] lemma mul_lt_mul_left (h : 0 < c) : c * a < c * b ↔ a < b := -(strict_mono_mul_left_of_pos h).lt_iff_lt - -@[simp] lemma mul_lt_mul_right (h : 0 < c) : a * c < b * c ↔ a < b := -(strict_mono_mul_right_of_pos h).lt_iff_lt - @[simp] lemma zero_le_mul_left (h : 0 < c) : 0 ≤ c * b ↔ 0 ≤ b := by { convert mul_le_mul_left h, simp } @[simp] lemma zero_le_mul_right (h : 0 < c) : 0 ≤ b * c ↔ 0 ≤ b := by { convert mul_le_mul_right h, simp } -@[simp] lemma zero_lt_mul_left (h : 0 < c) : 0 < c * b ↔ 0 < b := -by { convert mul_lt_mul_left h, simp } - -@[simp] lemma zero_lt_mul_right (h : 0 < c) : 0 < b * c ↔ 0 < b := -by { convert mul_lt_mul_right h, simp } - lemma add_le_mul_of_left_le_right (a2 : 2 ≤ a) (ab : a ≤ b) : a + b ≤ a * b := have 0 < b, from calc 0 < 2 : zero_lt_two @@ -828,22 +712,6 @@ by rw [bit0, ← two_mul, zero_lt_mul_left (zero_lt_two : 0 < (2:α))] end -lemma le_mul_iff_one_le_left (hb : 0 < b) : b ≤ a * b ↔ 1 ≤ a := -suffices 1 * b ≤ a * b ↔ 1 ≤ a, by rwa one_mul at this, -mul_le_mul_right hb - -lemma lt_mul_iff_one_lt_left (hb : 0 < b) : b < a * b ↔ 1 < a := -suffices 1 * b < a * b ↔ 1 < a, by rwa one_mul at this, -mul_lt_mul_right hb - -lemma le_mul_iff_one_le_right (hb : 0 < b) : b ≤ b * a ↔ 1 ≤ a := -suffices b * 1 ≤ b * a ↔ 1 ≤ a, by rwa mul_one at this, -mul_le_mul_left hb - -lemma lt_mul_iff_one_lt_right (hb : 0 < b) : b < b * a ↔ 1 < a := -suffices b * 1 < b * a ↔ 1 < a, by rwa mul_one at this, -mul_lt_mul_left hb - theorem mul_nonneg_iff_right_nonneg_of_pos (ha : 0 < a) : 0 ≤ a * b ↔ 0 ≤ b := by haveI := @linear_order.decidable_le α _; exact ⟨λ h, nonneg_of_mul_nonneg_right h ha, λ h, decidable.mul_nonneg ha.le h⟩ @@ -852,20 +720,6 @@ theorem mul_nonneg_iff_left_nonneg_of_pos (hb : 0 < b) : 0 ≤ a * b ↔ 0 ≤ a by haveI := @linear_order.decidable_le α _; exact ⟨λ h, nonneg_of_mul_nonneg_left h hb, λ h, decidable.mul_nonneg h hb.le⟩ -lemma mul_le_iff_le_one_left (hb : 0 < b) : a * b ≤ b ↔ a ≤ 1 := -⟨ λ h, le_of_not_lt (mt (lt_mul_iff_one_lt_left hb).2 h.not_lt), - λ h, le_of_not_lt (mt (lt_mul_iff_one_lt_left hb).1 h.not_lt) ⟩ - -lemma mul_lt_iff_lt_one_left (hb : 0 < b) : a * b < b ↔ a < 1 := -lt_iff_lt_of_le_iff_le $ le_mul_iff_one_le_left hb - -lemma mul_le_iff_le_one_right (hb : 0 < b) : b * a ≤ b ↔ a ≤ 1 := -⟨ λ h, le_of_not_lt (mt (lt_mul_iff_one_lt_right hb).2 h.not_lt), - λ h, le_of_not_lt (mt (lt_mul_iff_one_lt_right hb).1 h.not_lt) ⟩ - -lemma mul_lt_iff_lt_one_right (hb : 0 < b) : b * a < b ↔ a < 1 := -lt_iff_lt_of_le_iff_le $ le_mul_iff_one_le_right hb - lemma nonpos_of_mul_nonneg_left (h : 0 ≤ a * b) (hb : b < 0) : a ≤ 0 := le_of_not_gt (λ ha, absurd h (mul_neg_of_pos_of_neg ha hb).not_le) @@ -1882,7 +1736,7 @@ instance [canonically_ordered_comm_semiring α] [nontrivial α] : { induction a using with_bot.rec_bot_coe, { simp_rw [mul_bot x0.ne.symm, bot_le], }, induction b using with_bot.rec_bot_coe, { exact absurd h (bot_lt_coe a).not_le, }, { simp only [← coe_mul, coe_le_coe] at *, - exact zero_lt.mul_le_mul_left h (zero_le x), }, }, + exact mul_le_mul_left' h x, }, }, end ⟩ instance [canonically_ordered_comm_semiring α] [nontrivial α] : diff --git a/src/analysis/specific_limits/floor_pow.lean b/src/analysis/specific_limits/floor_pow.lean index ac785503dbfbc..22a9af21cf370 100644 --- a/src/analysis/specific_limits/floor_pow.lean +++ b/src/analysis/specific_limits/floor_pow.lean @@ -55,7 +55,7 @@ begin by simp only [cnpos.ne', ne.def, nat.cast_eq_zero, not_false_iff] with field_simps ... ≤ ε * c n : begin - apply mul_le_mul_of_nonneg_right _ (nat.cast_nonneg _), + refine mul_le_mul_of_nonneg_right _ (nat.cast_nonneg _), simp only [mul_one, real.norm_eq_abs, abs_one] at hn, exact le_trans (le_abs_self _) hn, end }, @@ -120,7 +120,7 @@ begin by simp only [cnpos.ne', ne.def, nat.cast_eq_zero, not_false_iff, neg_sub] with field_simps ... ≤ ε * c n : begin - apply mul_le_mul_of_nonneg_right _ (nat.cast_nonneg _), + refine mul_le_mul_of_nonneg_right _ (nat.cast_nonneg _), simp only [mul_one, real.norm_eq_abs, abs_one] at hn, exact le_trans (neg_le_abs_self _) hn, end }, diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index e4b00182ed6ed..e85a51c87b1e7 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -869,7 +869,7 @@ attribute [simp] nat.div_self protected lemma div_le_of_le_mul' {m n : ℕ} {k} (h : m ≤ k * n) : m / k ≤ n := (nat.eq_zero_or_pos k).elim (λ k0, by rw [k0, nat.div_zero]; apply zero_le) - (λ k0, (_root_.mul_le_mul_left k0).1 $ + (λ k0, (mul_le_mul_left k0).1 $ calc k * (m / k) ≤ m % k + k * (m / k) : nat.le_add_left _ _ ... = m : mod_add_div _ _ diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index c08c191aad4a2..9c5ec791b4735 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -599,7 +599,7 @@ by rw [← mul_lt_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel h, mul_comm] lemma mul_le_iff_le_inv {a b r : ℝ≥0} (hr : r ≠ 0) : r * a ≤ b ↔ a ≤ r⁻¹ * b := have 0 < r, from lt_of_le_of_ne (zero_le r) hr.symm, -by rw [← @mul_le_mul_left _ _ a _ r this, ← mul_assoc, mul_inv_cancel hr, one_mul] +by rw [← mul_le_mul_left (inv_pos.mpr this), ← mul_assoc, inv_mul_cancel hr, one_mul] lemma le_div_iff_mul_le {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ a * r ≤ b := by rw [div_eq_inv_mul, ← mul_le_iff_le_inv hr, mul_comm] diff --git a/src/measure_theory/integral/set_to_l1.lean b/src/measure_theory/integral/set_to_l1.lean index 01ba2d27259f4..a0548ef4a96e1 100644 --- a/src/measure_theory/integral/set_to_l1.lean +++ b/src/measure_theory/integral/set_to_l1.lean @@ -604,7 +604,7 @@ calc ∥f.set_to_simple_func T∥ refine finset.sum_le_sum (λ b hb, _), by_cases hb : ∥b∥ = 0, { rw hb, simp, }, - rw _root_.mul_le_mul_right _, + refine (zero_lt.mul_le_mul_right _).mpr _, swap, -- remove swap or inline the second subcase { exact hT_norm _ (simple_func.measurable_set_fiber _ _), }, { exact lt_of_le_of_ne (norm_nonneg _) (ne.symm hb), }, end @@ -621,7 +621,7 @@ calc ∥f.set_to_simple_func T∥ refine finset.sum_le_sum (λ b hb, _), by_cases hb : ∥b∥ = 0, { rw hb, simp, }, - rw _root_.mul_le_mul_right _, + refine (zero_lt.mul_le_mul_right _).mpr _, swap, -- remove swap or inline the second subcase { refine hT_norm _ (simple_func.measurable_set_fiber _ _) (simple_func.measure_preimage_lt_top_of_integrable _ hf _), rwa norm_eq_zero at hb, }, diff --git a/src/number_theory/liouville/liouville_constant.lean b/src/number_theory/liouville/liouville_constant.lean index 4ab08ec434271..09f200e7d362d 100644 --- a/src/number_theory/liouville/liouville_constant.lean +++ b/src/number_theory/liouville/liouville_constant.lean @@ -135,7 +135,7 @@ calc (1 - 1 / m)⁻¹ * (1 / m ^ (n + 1)!) ≤ 2 * (1 / m ^ (n + 1)!) : apply (div_le_div_iff _ _).mpr, conv_rhs { rw [one_mul, mul_add, pow_add, mul_one, pow_mul, mul_comm, ← pow_mul] }, -- the second factors coincide, so we prove the inequality of the first factors* - apply (mul_le_mul_right _).mpr, + refine (mul_le_mul_right _).mpr _, -- solve all the inequalities `0 < m ^ ??` any_goals { exact pow_pos (zero_lt_two.trans_le hm) _ }, -- `2 ≤ m ^ n!` is a consequence of monotonicity of exponentiation at `2 ≤ m`. From a51c2fd2773192ad887961562841649704da9db3 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 24 Aug 2022 21:27:29 +0000 Subject: [PATCH 032/828] =?UTF-8?q?feat(topology/separation):=20add=20`eq?= =?UTF-8?q?=5Fon=5Fclosure=E2=82=82`=20(#16206)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also use it to golf `submonoid.comm_monoid_topological_closure`. --- src/topology/algebra/monoid.lean | 23 ++++------------------- src/topology/separation.lean | 18 +++++++++++++++++- 2 files changed, 21 insertions(+), 20 deletions(-) diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 8a0f6eb83a0d9..51b7f069585b0 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -295,25 +295,10 @@ topological closure."] def submonoid.comm_monoid_topological_closure [t2_space M] (s : submonoid M) (hs : ∀ (x y : s), x * y = y * x) : comm_monoid s.topological_closure := { mul_comm := - begin - intros a b, - have h₁ : (s.topological_closure : set M) = closure s := rfl, - let f₁ := λ (x : M × M), x.1 * x.2, - let f₂ := λ (x : M × M), x.2 * x.1, - let S : set (M × M) := s ×ˢ s, - have h₃ : set.eq_on f₁ f₂ (closure S), - { refine set.eq_on.closure _ continuous_mul (by continuity), - intros x hx, - rw [set.mem_prod] at hx, - rcases hx with ⟨hx₁, hx₂⟩, - change ((⟨x.1, hx₁⟩ : s) : M) * (⟨x.2, hx₂⟩ : s) = (⟨x.2, hx₂⟩ : s) * (⟨x.1, hx₁⟩ : s), - exact_mod_cast hs _ _ }, - ext, - change f₁ ⟨a, b⟩ = f₂ ⟨a, b⟩, - refine h₃ _, - rw [closure_prod_eq, set.mem_prod], - exact ⟨by simp [←h₁], by simp [←h₁]⟩ - end, + have ∀ (x ∈ s) (y ∈ s), x * y = y * x, + from λ x hx y hy, congr_arg subtype.val (hs ⟨x, hx⟩ ⟨y, hy⟩), + λ ⟨x, hx⟩ ⟨y, hy⟩, subtype.ext $ + eq_on_closure₂ this continuous_mul (continuous_snd.mul continuous_fst) x hx y hy, ..s.topological_closure.to_monoid } @[to_additive exists_open_nhds_zero_half] diff --git a/src/topology/separation.lean b/src/topology/separation.lean index c7a98f7f4abab..83f4695af14c6 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -1021,7 +1021,7 @@ begin { exact ⟨_, _, is_open_range_sigma_mk, is_open_range_sigma_mk, ⟨x, rfl⟩, ⟨y, rfl⟩, by tidy⟩ } end -variables [topological_space β] +variables {γ : Type*} [topological_space β] [topological_space γ] lemma is_closed_eq [t2_space α] {f g : β → α} (hf : continuous f) (hg : continuous g) : is_closed {x:β | f x = g x} := @@ -1044,6 +1044,22 @@ lemma continuous.ext_on [t2_space α] {s : set β} (hs : dense s) {f g : β → f = g := funext $ λ x, h.closure hf hg (hs x) +lemma eq_on_closure₂' [t2_space α] {s : set β} {t : set γ} {f g : β → γ → α} + (h : ∀ (x ∈ s) (y ∈ t), f x y = g x y) + (hf₁ : ∀ x, continuous (f x)) (hf₂ : ∀ y, continuous (λ x, f x y)) + (hg₁ : ∀ x, continuous (g x)) (hg₂ : ∀ y, continuous (λ x, g x y)) : + ∀ (x ∈ closure s) (y ∈ closure t), f x y = g x y := +suffices closure s ⊆ ⋂ y ∈ closure t, {x | f x y = g x y}, by simpa only [subset_def, mem_Inter], +closure_minimal (λ x hx, mem_Inter₂.2 $ set.eq_on.closure (h x hx) (hf₁ _) (hg₁ _)) $ + is_closed_bInter $ λ y hy, is_closed_eq (hf₂ _) (hg₂ _) + +lemma eq_on_closure₂ [t2_space α] {s : set β} {t : set γ} {f g : β → γ → α} + (h : ∀ (x ∈ s) (y ∈ t), f x y = g x y) + (hf : continuous (uncurry f)) (hg : continuous (uncurry g)) : + ∀ (x ∈ closure s) (y ∈ closure t), f x y = g x y := +eq_on_closure₂' h (λ x, continuous_uncurry_left x hf) (λ x, continuous_uncurry_right x hf) + (λ y, continuous_uncurry_left y hg) (λ y, continuous_uncurry_right y hg) + /-- If `f x = g x` for all `x ∈ s` and `f`, `g` are continuous on `t`, `s ⊆ t ⊆ closure s`, then `f x = g x` for all `x ∈ t`. See also `set.eq_on.closure`. -/ lemma set.eq_on.of_subset_closure [t2_space α] {s t : set β} {f g : β → α} (h : eq_on f g s) From f1f166b2aac443714cf32fbe8c1e2d59b3e518cf Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 24 Aug 2022 21:27:30 +0000 Subject: [PATCH 033/828] =?UTF-8?q?feat(geometry/euclidean/oriented=5Fangl?= =?UTF-8?q?e):=20angles=20equal=20to=200=20or=20=CF=80=20(#16229)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add lemmas about various conditions for oriented angles between two vectors to equal 0 or π. --- src/geometry/euclidean/oriented_angle.lean | 109 +++++++++++++++++++++ 1 file changed, 109 insertions(+) diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 504bb7cc1b7aa..873352be8a174 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ +import analysis.complex.arg import analysis.inner_product_space.orientation import analysis.inner_product_space.pi_L2 import analysis.special_functions.complex.circle @@ -274,6 +275,76 @@ begin simp [h] end +/-- The oriented angle between two vectors is zero if and only if the angle with the vectors +swapped is zero. -/ +lemma oangle_eq_zero_iff_oangle_rev_eq_zero {x y : V} : hb.oangle x y = 0 ↔ hb.oangle y x = 0 := +by rw [oangle_rev, neg_eq_zero] + +/-- The oriented angle between two vectors is zero if and only if they are on the same ray. -/ +lemma oangle_eq_zero_iff_same_ray {x y : V} : hb.oangle x y = 0 ↔ same_ray ℝ x y := +by rw [oangle, complex.arg_coe_angle_eq_iff_eq_to_real, real.angle.to_real_zero, + ←complex.same_ray_iff_arg_div_eq_zero, ←linear_isometry_equiv.coe_to_linear_equiv, + same_ray_map_iff, same_ray_comm] + +/-- The oriented angle between two vectors is `π` if and only if the angle with the vectors +swapped is `π`. -/ +lemma oangle_eq_pi_iff_oangle_rev_eq_pi {x y : V} : hb.oangle x y = π ↔ hb.oangle y x = π := +by rw [oangle_rev, neg_eq_iff_neg_eq, eq_comm, real.angle.neg_coe_pi] + +/-- The oriented angle between two vectors is `π` if and only they are nonzero and the first is +on the same ray as the negation of the second. -/ +lemma oangle_eq_pi_iff_same_ray_neg {x y : V} : + hb.oangle x y = π ↔ x ≠ 0 ∧ y ≠ 0 ∧ same_ray ℝ x (-y) := +begin + rw [←hb.oangle_eq_zero_iff_same_ray], + split, + { intro h, + by_cases hx : x = 0, { simpa [hx, real.angle.pi_ne_zero.symm] using h }, + by_cases hy : y = 0, { simpa [hy, real.angle.pi_ne_zero.symm] using h }, + refine ⟨hx, hy, _⟩, + rw [hb.oangle_neg_right hx hy, h, real.angle.coe_pi_add_coe_pi] }, + { rintro ⟨hx, hy, h⟩, + rwa [hb.oangle_neg_right hx hy, ←real.angle.sub_coe_pi_eq_add_coe_pi, sub_eq_zero] at h } +end + +/-- The oriented angle between two vectors is zero or `π` if and only if those two vectors are +not linearly independent. -/ +lemma oangle_eq_zero_or_eq_pi_iff_not_linear_independent {x y : V} : + (hb.oangle x y = 0 ∨ hb.oangle x y = π) ↔ ¬ _root_.linear_independent ℝ ![x, y] := +by rw [oangle_eq_zero_iff_same_ray, oangle_eq_pi_iff_same_ray_neg, + same_ray_or_ne_zero_and_same_ray_neg_iff_not_linear_independent] + +/-- The oriented angle between two vectors is zero or `π` if and only if the first vector is zero +or the second is a multiple of the first. -/ +lemma oangle_eq_zero_or_eq_pi_iff_right_eq_smul {x y : V} : + (hb.oangle x y = 0 ∨ hb.oangle x y = π) ↔ (x = 0 ∨ ∃ r : ℝ, y = r • x) := +begin + rw [oangle_eq_zero_iff_same_ray, oangle_eq_pi_iff_same_ray_neg], + refine ⟨λ h, _, λ h, _⟩, + { rcases h with h|⟨-, -, h⟩, + { by_cases hx : x = 0, { simp [hx] }, + obtain ⟨r, -, rfl⟩ := h.exists_nonneg_left hx, + exact or.inr ⟨r, rfl⟩ }, + { by_cases hx : x = 0, { simp [hx] }, + obtain ⟨r, -, hy⟩ := h.exists_nonneg_left hx, + refine or.inr ⟨-r, _⟩, + simp [hy] } }, + { rcases h with rfl|⟨r, rfl⟩, { simp }, + by_cases hx : x = 0, { simp [hx] }, + rcases lt_trichotomy r 0 with hr|hr|hr, + { rw ←neg_smul, + exact or.inr ⟨hx, smul_ne_zero.2 ⟨hr.ne, hx⟩, + same_ray_pos_smul_right x (left.neg_pos_iff.2 hr)⟩ }, + { simp [hr] }, + { exact or.inl (same_ray_pos_smul_right x hr) } } +end + +/-- The oriented angle between two vectors is not zero or `π` if and only if those two vectors +are linearly independent. -/ +lemma oangle_ne_zero_and_ne_pi_iff_linear_independent {x y : V} : + (hb.oangle x y ≠ 0 ∧ hb.oangle x y ≠ π) ↔ _root_.linear_independent ℝ ![x, y] := +by rw [←not_or_distrib, ←not_iff_not, not_not, oangle_eq_zero_or_eq_pi_iff_not_linear_independent] + /-- Two vectors are equal if and only if they have equal norms and zero angle between them. -/ lemma eq_iff_norm_eq_and_oangle_eq_zero (x y : V) : x = y ↔ ∥x∥ = ∥y∥ ∧ hb.oangle x y = 0 := begin @@ -1105,6 +1176,44 @@ angle. -/ (2 : ℤ) • o.oangle (r₁ • x) (r₂ • x) = 0 := (ob).two_zsmul_oangle_smul_smul_self x +/-- The oriented angle between two vectors is zero if and only if the angle with the vectors +swapped is zero. -/ +lemma oangle_eq_zero_iff_oangle_rev_eq_zero {x y : V} : o.oangle x y = 0 ↔ o.oangle y x = 0 := +(ob).oangle_eq_zero_iff_oangle_rev_eq_zero + +/-- The oriented angle between two vectors is zero if and only if they are on the same ray. -/ +lemma oangle_eq_zero_iff_same_ray {x y : V} : o.oangle x y = 0 ↔ same_ray ℝ x y := +(ob).oangle_eq_zero_iff_same_ray + +/-- The oriented angle between two vectors is `π` if and only if the angle with the vectors +swapped is `π`. -/ +lemma oangle_eq_pi_iff_oangle_rev_eq_pi {x y : V} : o.oangle x y = π ↔ o.oangle y x = π := +(ob).oangle_eq_pi_iff_oangle_rev_eq_pi + +/-- The oriented angle between two vectors is `π` if and only they are nonzero and the first is +on the same ray as the negation of the second. -/ +lemma oangle_eq_pi_iff_same_ray_neg {x y : V} : + o.oangle x y = π ↔ x ≠ 0 ∧ y ≠ 0 ∧ same_ray ℝ x (-y) := +(ob).oangle_eq_pi_iff_same_ray_neg + +/-- The oriented angle between two vectors is zero or `π` if and only if those two vectors are +not linearly independent. -/ +lemma oangle_eq_zero_or_eq_pi_iff_not_linear_independent {x y : V} : + (o.oangle x y = 0 ∨ o.oangle x y = π) ↔ ¬ linear_independent ℝ ![x, y] := +(ob).oangle_eq_zero_or_eq_pi_iff_not_linear_independent + +/-- The oriented angle between two vectors is zero or `π` if and only if the first vector is zero +or the second is a multiple of the first. -/ +lemma oangle_eq_zero_or_eq_pi_iff_right_eq_smul {x y : V} : + (o.oangle x y = 0 ∨ o.oangle x y = π) ↔ (x = 0 ∨ ∃ r : ℝ, y = r • x) := +(ob).oangle_eq_zero_or_eq_pi_iff_right_eq_smul + +/-- The oriented angle between two vectors is not zero or `π` if and only if those two vectors +are linearly independent. -/ +lemma oangle_ne_zero_and_ne_pi_iff_linear_independent {x y : V} : + (o.oangle x y ≠ 0 ∧ o.oangle x y ≠ π) ↔ linear_independent ℝ ![x, y] := +(ob).oangle_ne_zero_and_ne_pi_iff_linear_independent + /-- Two vectors are equal if and only if they have equal norms and zero angle between them. -/ lemma eq_iff_norm_eq_and_oangle_eq_zero (x y : V) : x = y ↔ ∥x∥ = ∥y∥ ∧ o.oangle x y = 0 := (ob).eq_iff_norm_eq_and_oangle_eq_zero x y From 4b440bb0be61a13cb5b222c153b7d407b3bffa4c Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 25 Aug 2022 00:16:11 +0000 Subject: [PATCH 034/828] refactor(order/well_founded_set): golf, review API (#11303) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ## New lemmas ### About `set.well_founded_on` - `set.well_founded_on.mono`, `set.well_founded_on.mono_set`; - `set.well_founded_on.union`, `set.well_founded_on_union`, `finset.well_founded_on`, `set.finite.well_founded_on`, `set.subsingleton.well_founded_on`, `set.well_founded_on_empty`, `set.well_founded_on_singleton`, `set.well_founded_on_insert`, `set.well_founded_on.insert`; ### About `set.is_wf` - `set.is_wf_union`; - `set.finite.is_wf`, `finset.is_wf`, `set.is_wf_empty`, `set.is_wf_singleton`, `set.subsingleton.is_wf`, `set.is_wf_insert`, `set.is_wf.insert`; - `finset.is_wf_bUnion` ### About `set.partially_well_ordered_on` - `set.partially_well_ordered_on.union`, `set.partially_well_ordered_on_union`; - `set.finite.partially_well_ordered_on`, `finset.partially_well_ordered_on`, `set.partially_well_ordered_on_empty`, `set.partially_well_ordered_on_singleton`, `set.partially_well_ordered_on_insert` (an `iff`), `set.partially_well_ordered_on.insert`, `set.subsingleton.partially_well_ordered_on`; ### About `set.is_pwo` - `set.is_pwo.image_of_monotone_on`, `set.is_pwo_union`; - `finset.is_pwo`, `set.is_pwo_insert`; - `finset.is_pwo_bUnion`; ## Other API changes - Definitions lemmas now use `∀ n, f n ∈ s` instead of `range f ⊆ s`. - Many lemmas now use weaker TC assumptions (e.g., `preorder` instead of `partial_order`). - `set.is_wf_iff_no_descending_seq` now uses `(f : ℕ → α) (hf : strict_anti f)` instead of `f : order_dual ℕ ↪o α`. - The order of binders in the `hf` argument of `set.partially_well_ordered_on.image_of_monotone_on` was changed to match `monotone_on`. - `set.is_pwo.prod` now allows sets from different types. - `finset.is_wf_sup` and `finset.is_pwo_sup` are now `iff`s Co-authored-by: YaelDillies Co-authored-by: Yaël Dillies --- src/combinatorics/simple_graph/basic.lean | 4 +- src/data/fintype/basic.lean | 9 +- src/data/set/finite.lean | 11 +- src/data/set/pointwise.lean | 176 ++++- src/group_theory/submonoid/pointwise.lean | 14 + src/order/well_founded_set.lean | 893 ++++++++-------------- src/probability/cond_count.lean | 2 +- src/ring_theory/hahn_series.lean | 132 ++-- 8 files changed, 594 insertions(+), 647 deletions(-) diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index 4c4736d82a20b..e51ba9debd13d 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -987,7 +987,7 @@ begin { rw finset.insert_subset, split, { simpa, }, - { rw [neighbor_finset, set.to_finset_mono], + { rw [neighbor_finset, set.to_finset_subset], exact G.common_neighbors_subset_neighbor_set_left _ _ } } end @@ -997,7 +997,7 @@ begin simp only [common_neighbors_top_eq, ← set.to_finset_card, set.to_finset_diff], rw finset.card_sdiff, { simp [finset.card_univ, h], }, - { simp only [set.to_finset_mono, set.subset_univ] }, + { simp only [set.to_finset_subset, set.subset_univ] }, end end finite diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index e6686eadbe6b4..0c2591208b25c 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -607,6 +607,7 @@ arbitrary `fintype` instances, use `fintype.card_eq_zero_iff`. -/ end fintype namespace set +variables {s t : set α} /-- Construct a finset enumerating a set `s`, given a `fintype` instance. -/ def to_finset (s : set α) [fintype s] : finset α := @@ -644,13 +645,11 @@ set.ext $ λ _, mem_to_finset s.to_finset = t.to_finset ↔ s = t := ⟨λ h, by rw [←s.coe_to_finset, h, t.coe_to_finset], λ h, by simp [h]; congr⟩ -@[simp, mono] theorem to_finset_mono {s t : set α} [fintype s] [fintype t] : - s.to_finset ⊆ t.to_finset ↔ s ⊆ t := +@[simp, mono] lemma to_finset_subset [fintype s] [fintype t] : s.to_finset ⊆ t.to_finset ↔ s ⊆ t := by simp [finset.subset_iff, set.subset_def] -@[simp, mono] theorem to_finset_strict_mono {s t : set α} [fintype s] [fintype t] : - s.to_finset ⊂ t.to_finset ↔ s ⊂ t := -by simp only [finset.ssubset_def, to_finset_mono, ssubset_def] +@[simp, mono] lemma to_finset_ssubset [fintype s] [fintype t] : s.to_finset ⊂ t.to_finset ↔ s ⊂ t := +by simp only [finset.ssubset_def, to_finset_subset, ssubset_def] @[simp] theorem to_finset_disjoint_iff [decidable_eq α] {s t : set α} [fintype s] [fintype t] : disjoint s.to_finset t.to_finset ↔ disjoint s t := diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index b59a266730b64..516e835d046b2 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -122,6 +122,7 @@ lemma finite_or_infinite {s : set α} : s.finite ∨ s.infinite := em _ /-! ### Basic properties of `set.finite.to_finset` -/ section finite_to_finset +variables {s t : set α} @[simp] lemma finite.coe_to_finset {s : set α} (h : s.finite) : (h.to_finset : set α) = s := @set.coe_to_finset _ s h.fintype @@ -152,11 +153,11 @@ by rw [← finset.coe_subset, ht.coe_to_finset] h.to_finset = ∅ ↔ s = ∅ := by simp only [←finset.coe_inj, finite.coe_to_finset, finset.coe_empty] -@[simp, mono] lemma finite.to_finset_mono {s t : set α} {hs : s.finite} {ht : t.finite} : +@[simp, mono] lemma finite.to_finset_subset {hs : s.finite} {ht : t.finite} : hs.to_finset ⊆ ht.to_finset ↔ s ⊆ t := by simp only [← finset.coe_subset, finite.coe_to_finset] -@[simp, mono] lemma finite.to_finset_strict_mono {s t : set α} {hs : s.finite} {ht : t.finite} : +@[simp, mono] lemma finite.to_finset_ssubset {hs : s.finite} {ht : t.finite} : hs.to_finset ⊂ ht.to_finset ↔ s ⊂ t := by simp only [← finset.coe_ssubset, finite.coe_to_finset] @@ -934,11 +935,11 @@ theorem infinite.exists_lt_map_eq_of_maps_to [linear_order α] {s : set α} {t : let ⟨x, hx, y, hy, hxy, hf⟩ := hs.exists_ne_map_eq_of_maps_to hf ht in hxy.lt_or_lt.elim (λ hxy, ⟨x, hx, y, hy, hxy, hf⟩) (λ hyx, ⟨y, hy, x, hx, hyx, hf.symm⟩) -lemma finite.exists_lt_map_eq_of_range_subset [linear_order α] [infinite α] {t : set β} - {f : α → β} (hf : range f ⊆ t) (ht : t.finite) : +lemma finite.exists_lt_map_eq_of_forall_mem [linear_order α] [infinite α] {t : set β} + {f : α → β} (hf : ∀ a, f a ∈ t) (ht : t.finite) : ∃ a b, a < b ∧ f a = f b := begin - rw [range_subset_iff, ←maps_univ_to] at hf, + rw ←maps_univ_to at hf, obtain ⟨a, -, b, -, h⟩ := (@infinite_univ α _).exists_lt_map_eq_of_maps_to hf ht, exact ⟨a, b, h⟩, end diff --git a/src/data/set/pointwise.lean b/src/data/set/pointwise.lean index c3dffead1a0ce..a10504bc3d005 100644 --- a/src/data/set/pointwise.lean +++ b/src/data/set/pointwise.lean @@ -5,8 +5,8 @@ Authors: Johan Commelin, Floris van Doorn -/ import algebra.module.basic import data.fin.tuple.basic -import data.set.finite import group_theory.submonoid.basic +import order.well_founded_set /-! # Pointwise operations of sets @@ -35,6 +35,13 @@ the latter has `(2 : ℕ) • {1, 2} = {2, 3, 4}`. See note [pointwise nat actio Appropriate definitions and results are also transported to the additive theory via `to_additive`. +### Definitions for Hahn series + +* `set.add_antidiagonal s t a`, `set.mul_antidiagonal s t a`: Sets of pairs of elements of `s` and + `t` that add/multiply to `a`. +* `finset.add_antidiagonal`, `finset.mul_antidiagonal`: Finset versions of the above when `s` and + `t` are well-founded. + ## Implementation notes * The following expressions are considered in simp-normal form in a group: @@ -1469,3 +1476,170 @@ begin end end group + +namespace set +variables {s t : set α} + +@[to_additive] +lemma is_pwo.mul [ordered_cancel_comm_monoid α] (hs : s.is_pwo) (ht : t.is_pwo) : is_pwo (s * t) := +by { rw ←image_mul_prod, exact (hs.prod ht).image_of_monotone (monotone_fst.mul' monotone_snd) } + +variables [linear_ordered_cancel_comm_monoid α] + +@[to_additive] +lemma is_wf.mul (hs : s.is_wf) (ht : t.is_wf) : is_wf (s * t) := (hs.is_pwo.mul ht.is_pwo).is_wf + +@[to_additive] +lemma is_wf.min_mul (hs : s.is_wf) (ht : t.is_wf) (hsn : s.nonempty) (htn : t.nonempty) : + (hs.mul ht).min (hsn.mul htn) = hs.min hsn * ht.min htn := +begin + refine le_antisymm (is_wf.min_le _ _ (mem_mul.2 ⟨_, _, hs.min_mem _, ht.min_mem _, rfl⟩)) _, + rw is_wf.le_min_iff, + rintro _ ⟨x, y, hx, hy, rfl⟩, + exact mul_le_mul' (hs.min_le _ hx) (ht.min_le _ hy), +end + +end set + +/-! ### Multiplication antidiagonal -/ + +namespace set +section has_mul +variables [has_mul α] {s s₁ s₂ t t₁ t₂ : set α} {a : α} {x : α × α} + +/-- `set.mul_antidiagonal s t a` is the set of all pairs of an element in `s` and an element in `t` +that multiply to `a`. -/ +@[to_additive "`set.add_antidiagonal s t a` is the set of all pairs of an element in `s` and an +element in `t` that add to `a`."] +def mul_antidiagonal (s t : set α) (a : α) : set (α × α) := {x | x.1 ∈ s ∧ x.2 ∈ t ∧ x.1 * x.2 = a} + +@[simp, to_additive] +lemma mem_mul_antidiagonal : x ∈ mul_antidiagonal s t a ↔ x.1 ∈ s ∧ x.2 ∈ t ∧ x.1 * x.2 = a := +iff.rfl + +@[to_additive] lemma mul_antidiagonal_mono_left (h : s₁ ⊆ s₂) : + mul_antidiagonal s₁ t a ⊆ mul_antidiagonal s₂ t a := +λ x hx, ⟨h hx.1, hx.2.1, hx.2.2⟩ + +@[to_additive] lemma mul_antidiagonal_mono_right (h : t₁ ⊆ t₂) : + mul_antidiagonal s t₁ a ⊆ mul_antidiagonal s t₂ a := +λ x hx, ⟨hx.1, h hx.2.1, hx.2.2⟩ + +end has_mul + +@[simp, to_additive] lemma swap_mem_mul_antidiagonal [comm_semigroup α] {s t : set α} {a : α} + {x : α × α} : + x.swap ∈ set.mul_antidiagonal s t a ↔ x ∈ set.mul_antidiagonal t s a := +by simp [mul_comm, and.left_comm] + +namespace mul_antidiagonal + +section cancel_comm_monoid +variables [cancel_comm_monoid α] {s t : set α} {a : α} {x y : mul_antidiagonal s t a} + +@[to_additive] +lemma fst_eq_fst_iff_snd_eq_snd : (x : α × α).1 = (y : α × α).1 ↔ (x : α × α).2 = (y : α × α).2 := +⟨λ h, mul_left_cancel (y.prop.2.2.trans $ by { rw ←h, exact x.2.2.2.symm }).symm, + λ h, mul_right_cancel (y.prop.2.2.trans $ by { rw ←h, exact x.2.2.2.symm }).symm⟩ + +@[to_additive] lemma eq_of_fst_eq_fst (h : (x : α × α).fst = (y : α × α).fst) : x = y := +subtype.ext $ prod.ext h $ fst_eq_fst_iff_snd_eq_snd.1 h + +@[to_additive] lemma eq_of_snd_eq_snd (h : (x : α × α).snd = (y : α × α).snd) : x = y := +subtype.ext $ prod.ext (fst_eq_fst_iff_snd_eq_snd.2 h) h + +end cancel_comm_monoid + +section ordered_cancel_comm_monoid +variables [ordered_cancel_comm_monoid α] (s t : set α) (a : α) {x y : mul_antidiagonal s t a} + +@[to_additive] +lemma eq_of_fst_le_fst_of_snd_le_snd (h₁ : (x : α × α).1 ≤ (y : α × α).1) + (h₂ : (x : α × α).2 ≤ (y : α × α).2) : + x = y := +eq_of_fst_eq_fst $ h₁.eq_of_not_lt $ λ hlt, (mul_lt_mul_of_lt_of_le hlt h₂).ne $ + (mem_mul_antidiagonal.1 x.2).2.2.trans (mem_mul_antidiagonal.1 y.2).2.2.symm + +variables {s t} + +@[to_additive] +lemma finite_of_is_pwo (hs : s.is_pwo) (ht : t.is_pwo) (a) : (mul_antidiagonal s t a).finite := +begin + refine not_infinite.1 (λ h, _), + have h1 : (mul_antidiagonal s t a).partially_well_ordered_on (prod.fst ⁻¹'o (≤)), + from λ f hf, hs (prod.fst ∘ f) (λ n, (mem_mul_antidiagonal.1 (hf n)).1), + have h2 : (mul_antidiagonal s t a).partially_well_ordered_on (prod.snd ⁻¹'o (≤)), + from λ f hf, ht (prod.snd ∘ f) (λ n, (mem_mul_antidiagonal.1 (hf n)).2.1), + obtain ⟨g, hg⟩ := h1.exists_monotone_subseq (λ n, h.nat_embedding _ n) + (λ n, (h.nat_embedding _ n).2), + obtain ⟨m, n, mn, h2'⟩ := h2 (λ x, (h.nat_embedding _) (g x)) (λ n, (h.nat_embedding _ _).2), + refine mn.ne (g.injective $ (h.nat_embedding _).injective _), + exact eq_of_fst_le_fst_of_snd_le_snd _ _ _ (hg _ _ mn.le) h2', +end + +end ordered_cancel_comm_monoid + +@[to_additive] +lemma finite_of_is_wf [linear_ordered_cancel_comm_monoid α] {s t : set α} (hs : s.is_wf) + (ht : t.is_wf) (a) : + (mul_antidiagonal s t a).finite := +finite_of_is_pwo hs.is_pwo ht.is_pwo a + +end mul_antidiagonal +end set + +namespace finset +variables [ordered_cancel_comm_monoid α] {s t : set α} (hs : s.is_pwo) (ht : t.is_pwo) (a : α) + +/-- `finset.mul_antidiagonal_of_is_wf hs ht a` is the set of all pairs of an element in `s` and an +element in `t` that multiply to `a`, but its construction requires proofs that `s` and `t` are +well-ordered. -/ +@[to_additive "`finset.add_antidiagonal_of_is_wf hs ht a` is the set of all pairs of an element in +`s` and an element in `t` that add to `a`, but its construction requires proofs that `s` and `t` are +well-ordered."] +noncomputable def mul_antidiagonal : finset (α × α) := +(set.mul_antidiagonal.finite_of_is_pwo hs ht a).to_finset + +variables {hs ht a} {u : set α} {hu : u.is_pwo} {x : α × α} + +@[simp, to_additive] +lemma mem_mul_antidiagonal : x ∈ mul_antidiagonal hs ht a ↔ x.1 ∈ s ∧ x.2 ∈ t ∧ x.1 * x.2 = a := +by simp [mul_antidiagonal, and_rotate] + +@[to_additive] lemma mul_antidiagonal_mono_left (h : u ⊆ s) : + mul_antidiagonal hu ht a ⊆ mul_antidiagonal hs ht a := +finite.to_finset_subset.2 $ set.mul_antidiagonal_mono_left h + +@[to_additive] lemma mul_antidiagonal_mono_right (h : u ⊆ t) : + mul_antidiagonal hs hu a ⊆ mul_antidiagonal hs ht a := +finite.to_finset_subset.2 $ set.mul_antidiagonal_mono_right h + +@[simp, to_additive] lemma swap_mem_mul_antidiagonal : + x.swap ∈ finset.mul_antidiagonal hs ht a ↔ x ∈ finset.mul_antidiagonal ht hs a := +by simp [mul_comm, and.left_comm] + +@[to_additive] +lemma support_mul_antidiagonal_subset_mul : {a | (mul_antidiagonal hs ht a).nonempty} ⊆ s * t := +λ a ⟨b, hb⟩, by { rw mem_mul_antidiagonal at hb, exact ⟨b.1, b.2, hb⟩ } + +@[to_additive] +lemma is_pwo_support_mul_antidiagonal : {a | (mul_antidiagonal hs ht a).nonempty}.is_pwo := +(hs.mul ht).mono support_mul_antidiagonal_subset_mul + +@[to_additive] +lemma mul_antidiagonal_min_mul_min {α} [linear_ordered_cancel_comm_monoid α] {s t : set α} + (hs : s.is_wf) (ht : t.is_wf) (hns : s.nonempty) (hnt : t.nonempty) : + mul_antidiagonal hs.is_pwo ht.is_pwo ((hs.min hns) * (ht.min hnt)) = {(hs.min hns, ht.min hnt)} := +begin + ext ⟨a, b⟩, + simp only [mem_mul_antidiagonal, mem_singleton, prod.ext_iff], + split, + { rintro ⟨has, hat, hst⟩, + obtain rfl := (hs.min_le hns has).eq_of_not_lt + (λ hlt, (mul_lt_mul_of_lt_of_le hlt $ ht.min_le hnt hat).ne' hst), + exact ⟨rfl, mul_left_cancel hst⟩ }, + { rintro ⟨rfl, rfl⟩, + exact ⟨hs.min_mem _, ht.min_mem _, rfl⟩ } +end + +end finset diff --git a/src/group_theory/submonoid/pointwise.lean b/src/group_theory/submonoid/pointwise.lean index a91f171ad906e..4da829732cdca 100644 --- a/src/group_theory/submonoid/pointwise.lean +++ b/src/group_theory/submonoid/pointwise.lean @@ -486,3 +486,17 @@ lemma pow_subset_pow {s : add_submonoid R} {n : ℕ} : (↑s : set R)^n ⊆ ↑( end semiring end add_submonoid + +namespace set.is_pwo +variables [ordered_cancel_comm_monoid α] {s : set α} + +@[to_additive] +lemma submonoid_closure (hpos : ∀ x : α, x ∈ s → 1 ≤ x) (h : s.is_pwo) : + is_pwo ((submonoid.closure s) : set α) := +begin + rw submonoid.closure_eq_image_prod, + refine (h.partially_well_ordered_on_sublist_forall₂ (≤)).image_of_monotone_on _, + exact λ l1 hl1 l2 hl2 h12, h12.prod_le_prod' (λ x hx, hpos x $ hl2 x hx) +end + +end set.is_pwo diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index d80a49d86ad96..5ae940977844e 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -3,7 +3,6 @@ Copyright (c) 2021 Aaron Anderson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ -import data.set.pointwise import order.antichain import order.order_iso_nat import order.well_founded @@ -19,14 +18,8 @@ A well-founded subset of an ordered type is one on which the relation `<` is wel * `set.is_wf s` indicates that `<` is well-founded when restricted to `s`. * `set.partially_well_ordered_on s r` indicates that the relation `r` is partially well-ordered (also known as well quasi-ordered) when restricted to the set `s`. - * `set.is_pwo s` indicates that any infinite sequence of elements in `s` - contains an infinite monotone subsequence. Note that - -### Definitions for Hahn Series - * `set.add_antidiagonal s t a` and `set.mul_antidiagonal s t a` are the sets of pairs of elements - from `s` and `t` that add/multiply to `a`. - * `finset.add_antidiagonal` and `finset.mul_antidiagonal` are finite versions of - `set.add_antidiagonal` and `set.mul_antidiagonal` defined when `s` and `t` are well-founded. + * `set.is_pwo s` indicates that any infinite sequence of elements in `s` contains an infinite + monotone subsequence. Note that this is equivalent to containing only two comparable elements. ## Main Results * Higman's Lemma, `set.partially_well_ordered_on.partially_well_ordered_on_sublist_forall₂`, @@ -48,20 +41,25 @@ Prove that `s` is partial well ordered iff it has no infinite descending chain o * [Nash-Williams, *On Well-Quasi-Ordering Finite Trees*][Nash-Williams63] -/ -open_locale pointwise - -variables {α : Type*} +variables {ι α β : Type*} namespace set +/-! ### Relations well-founded on sets -/ + /-- `s.well_founded_on r` indicates that the relation `r` is well-founded when restricted to `s`. -/ -def well_founded_on (s : set α) (r : α → α → Prop) : Prop := -well_founded (λ (a : s) (b : s), r a b) +def well_founded_on (s : set α) (r : α → α → Prop) : Prop := well_founded $ λ a b : s, r a b @[simp] lemma well_founded_on_empty (r : α → α → Prop) : well_founded_on ∅ r := well_founded_of_empty _ -lemma well_founded_on_iff {s : set α} {r : α → α → Prop} : +section well_founded_on +variables {r r' : α → α → Prop} + +section any_rel +variables {s t : set α} {x y : α} + +lemma well_founded_on_iff : s.well_founded_on r ↔ well_founded (λ (a b : α), r a b ∧ a ∈ s ∧ b ∈ s) := begin have f : rel_embedding (λ (a : s) (b : s), r a b) (λ (a b : α), r a b ∧ a ∈ s ∧ b ∈ s) := @@ -77,43 +75,72 @@ begin exact ⟨m, mt, λ x xt ⟨xm, xs, ms⟩, hst ⟨m, ⟨ms, mt⟩⟩⟩ } end -lemma well_founded_on.induction {s : set α} {r : α → α → Prop} (hs : s.well_founded_on r) {x : α} - (hx : x ∈ s) {P : α → Prop} (hP : ∀ (y ∈ s), (∀ (z ∈ s), r z y → P z) → P y) : - P x := +namespace well_founded_on + +protected lemma induction (hs : s.well_founded_on r) (hx : x ∈ s) {P : α → Prop} + (hP : ∀ y ∈ s, (∀ z ∈ s, r z y → P z) → P y) : P x := begin let Q : s → Prop := λ y, P y, change Q ⟨x, hx⟩, refine well_founded.induction hs ⟨x, hx⟩ _, - rintros ⟨y, ys⟩ ih, - exact hP _ ys (λ z zs zy, ih ⟨z, zs⟩ zy), + simpa only [subtype.forall] end -instance is_strict_order.subset {s : set α} {r : α → α → Prop} [is_strict_order α r] : - is_strict_order α (λ (a b : α), r a b ∧ a ∈ s ∧ b ∈ s) := +protected lemma mono (h : t.well_founded_on r') (hle : r ≤ r') (hst : s ⊆ t) : + s.well_founded_on r := +begin + rw well_founded_on_iff at *, + refine subrelation.wf (λ x y xy, _) h, + exact ⟨hle _ _ xy.1, hst xy.2.1, hst xy.2.2⟩ +end + +lemma subset (h : t.well_founded_on r) (hst : s ⊆ t) : s.well_founded_on r := h.mono le_rfl hst + +end well_founded_on +end any_rel + +section is_strict_order +variables [is_strict_order α r] {s t : set α} + +instance is_strict_order.subset : is_strict_order α (λ (a b : α), r a b ∧ a ∈ s ∧ b ∈ s) := { to_is_irrefl := ⟨λ a con, irrefl_of r a con.1 ⟩, to_is_trans := ⟨λ a b c ab bc, ⟨trans_of r ab.1 bc.1, ab.2.1, bc.2.2⟩ ⟩ } -theorem well_founded_on_iff_no_descending_seq {s : set α} {r : α → α → Prop} [is_strict_order α r] : - s.well_founded_on r ↔ ∀ (f : ((>) : ℕ → ℕ → Prop) ↪r r), ¬ (range f) ⊆ s := +lemma well_founded_on_iff_no_descending_seq : + s.well_founded_on r ↔ ∀ (f : ((>) : ℕ → ℕ → Prop) ↪r r), ¬∀ n, f n ∈ s := +begin + simp only [well_founded_on_iff, rel_embedding.well_founded_iff_no_descending_seq, ← not_exists, + ← not_nonempty_iff, not_iff_not], + split, + { rintro ⟨⟨f, hf⟩⟩, + have H : ∀ n, f n ∈ s, from λ n, (hf.2 n.lt_succ_self).2.2, + refine ⟨⟨f, _⟩, H⟩, + simpa only [H, and_true] using @hf }, + { rintro ⟨⟨f, hf⟩, hfs : ∀ n, f n ∈ s⟩, + refine ⟨⟨f, _⟩⟩, + simpa only [hfs, and_true] using @hf } +end + +lemma well_founded_on.union (hs : s.well_founded_on r) (ht : t.well_founded_on r) : + (s ∪ t).well_founded_on r := begin - rw [well_founded_on_iff, rel_embedding.well_founded_iff_no_descending_seq], - refine ⟨λ h f con, begin - refine h.elim' ⟨⟨f, f.injective⟩, λ a b, _⟩, - simp only [con (mem_range_self a), con (mem_range_self b), and_true, gt_iff_lt, - function.embedding.coe_fn_mk, f.map_rel_iff] - end, λ h, ⟨λ con, _⟩⟩, - rcases con with ⟨f, hf⟩, - have hfs' : ∀ n : ℕ, f n ∈ s := λ n, (hf.2 n.lt_succ_self).2.2, - refine h ⟨f, λ a b, _⟩ (λ n hn, _), - { rw ← hf, - exact ⟨λ h, ⟨h, hfs' _, hfs' _⟩, λ h, h.1⟩ }, - { rcases set.mem_range.1 hn with ⟨m, hm⟩, - rw ← hm, - apply hfs' } + rw well_founded_on_iff_no_descending_seq at *, + rintro f hf, + rcases nat.exists_subseq_of_forall_mem_union f hf with ⟨g, hg|hg⟩, + exacts [hs (g.dual.lt_embedding.trans f) hg, ht (g.dual.lt_embedding.trans f) hg] end +@[simp] lemma well_founded_on_union : + (s ∪ t).well_founded_on r ↔ s.well_founded_on r ∧ t.well_founded_on r := +⟨λ h, ⟨h.subset $ subset_union_left _ _, h.subset $ subset_union_right _ _⟩, λ h, h.1.union h.2⟩ + +end is_strict_order +end well_founded_on + +/-! ### Sets well-founded w.r.t. the strict inequality -/ + section has_lt -variables [has_lt α] +variables [has_lt α] {s t : set α} /-- `s.is_wf` indicates that `<` is well-founded when restricted to `s`. -/ def is_wf (s : set α) : Prop := well_founded_on s (<) @@ -123,140 +150,123 @@ def is_wf (s : set α) : Prop := well_founded_on s (<) lemma is_wf_univ_iff : is_wf (univ : set α) ↔ well_founded ((<) : α → α → Prop) := by simp [is_wf, well_founded_on_iff] -variables {s t : set α} +theorem is_wf.mono (h : is_wf t) (st : s ⊆ t) : is_wf s := h.subset st -theorem is_wf.mono (h : is_wf t) (st : s ⊆ t) : is_wf s := -begin - rw [is_wf, well_founded_on_iff] at *, - refine subrelation.wf (λ x y xy, _) h, - exact ⟨xy.1, st xy.2.1, st xy.2.2⟩, -end end has_lt -section partial_order -variables [partial_order α] {s t : set α} {a : α} +section preorder +variables [preorder α] {s t : set α} {a : α} -theorem is_wf_iff_no_descending_seq : is_wf s ↔ ∀ f : ℕᵒᵈ ↪o α, ¬ (range f) ⊆ s := -begin - haveI : is_strict_order α (λ (a b : α), a < b ∧ a ∈ s ∧ b ∈ s) := - { to_is_irrefl := ⟨λ x con, lt_irrefl x con.1⟩, - to_is_trans := ⟨λ a b c ab bc, ⟨lt_trans ab.1 bc.1, ab.2.1, bc.2.2⟩⟩, }, - rw [is_wf, well_founded_on_iff_no_descending_seq], - exact ⟨λ h f, h f.lt_embedding, λ h f, h (order_embedding.of_strict_mono - f (λ _ _, f.map_rel_iff.2))⟩, -end +protected lemma is_wf.union (hs : is_wf s) (ht : is_wf t) : is_wf (s ∪ t) := hs.union ht -theorem is_wf.union (hs : is_wf s) (ht : is_wf t) : is_wf (s ∪ t) := -begin - classical, - rw [is_wf_iff_no_descending_seq] at *, - rintros f fst, - have h : (f ⁻¹' s).infinite ∨ (f ⁻¹' t).infinite, - { have h : (univ : set ℕ).infinite := infinite_univ, - have hpre : f ⁻¹' (s ∪ t) = set.univ, - { rw [← image_univ, image_subset_iff, univ_subset_iff] at fst, - exact fst }, - rw preimage_union at hpre, - rw ← hpre at h, - rw [set.infinite, set.infinite], - rw set.infinite at h, - contrapose! h, - exact finite.union h.1 h.2, }, - rw [← infinite_coe_iff, ← infinite_coe_iff] at h, - cases h with inf inf; haveI := inf, - { apply hs ((nat.order_embedding_of_set (f ⁻¹' s)).dual.trans f), - change range (function.comp f (nat.order_embedding_of_set (f ⁻¹' s))) ⊆ s, - rw [range_comp, image_subset_iff], - simp }, - { apply ht ((nat.order_embedding_of_set (f ⁻¹' t)).dual.trans f), - change range (function.comp f (nat.order_embedding_of_set (f ⁻¹' t))) ⊆ t, - rw [range_comp, image_subset_iff], - simp } -end -end partial_order +@[simp] lemma is_wf_union : is_wf (s ∪ t) ↔ is_wf s ∧ is_wf t := well_founded_on_union -end set +end preorder -namespace set +section preorder +variables [preorder α] {s t : set α} {a : α} + +theorem is_wf_iff_no_descending_seq : + is_wf s ↔ ∀ f : ℕ → α, strict_anti f → ¬(∀ n, f (order_dual.to_dual n) ∈ s) := +well_founded_on_iff_no_descending_seq.trans + ⟨λ H f hf, H ⟨⟨f, hf.injective⟩, λ a b, hf.lt_iff_lt⟩, λ H f, H f (λ _ _, f.map_rel_iff.2)⟩ + +end preorder + +/-! +### Partially well-ordered sets + +A set is partially well-ordered by a relation `r` when any infinite sequence contains two elements +where the first is related to the second by `r`. Equivalently, any antichain (see `is_antichain`) is +finite, see `set.partially_well_ordered_on_iff_finite_antichains`. +-/ /-- A subset is partially well-ordered by a relation `r` when any infinite sequence contains two elements where the first is related to the second by `r`. -/ -def partially_well_ordered_on (s) (r : α → α → Prop) : Prop := - ∀ (f : ℕ → α), range f ⊆ s → ∃ (m n : ℕ), m < n ∧ r (f m) (f n) - -@[simp] theorem partially_well_ordered_on_empty (r : α → α → Prop) : - partially_well_ordered_on (∅ : set α) r := -λ f h, begin - rw subset_empty_iff at h, - exact ((range_nonempty f).ne_empty h).elim -end +def partially_well_ordered_on (s : set α) (r : α → α → Prop) : Prop := +∀ f : ℕ → α, (∀ n, f n ∈ s) → ∃ m n : ℕ, m < n ∧ r (f m) (f n) -/-- A subset of a preorder is partially well-ordered when any infinite sequence contains - a monotone subsequence of length 2 (or equivalently, an infinite monotone subsequence). -/ -def is_pwo [preorder α] (s) : Prop := -partially_well_ordered_on s ((≤) : α → α → Prop) +section partially_well_ordered_on +variables {r : α → α → Prop} {r' : β → β → Prop} {f : α → β} {s : set α} {t : set α} {a : α} -@[simp] theorem is_pwo_empty [preorder α] : is_pwo (∅ : set α) := partially_well_ordered_on_empty _ - -theorem partially_well_ordered_on.mono {s t : set α} {r : α → α → Prop} - (ht : t.partially_well_ordered_on r) (hsub : s ⊆ t) : +lemma partially_well_ordered_on.mono (ht : t.partially_well_ordered_on r) (h : s ⊆ t) : s.partially_well_ordered_on r := -λ f hf, ht f (set.subset.trans hf hsub) +λ f hf, ht f $ λ n, h $ hf n -theorem is_pwo.mono [preorder α] {s t : set α} - (ht : t.is_pwo) (hsub : s ⊆ t) : - s.is_pwo := -partially_well_ordered_on.mono ht hsub +@[simp] lemma partially_well_ordered_on_empty (r : α → α → Prop) : partially_well_ordered_on ∅ r := +λ f hf, (hf 0).elim -theorem partially_well_ordered_on.image_of_monotone_on {s : set α} - {r : α → α → Prop} {β : Type*} {r' : β → β → Prop} - (hs : s.partially_well_ordered_on r) {f : α → β} - (hf : ∀ a1 a2 : α, a1 ∈ s → a2 ∈ s → r a1 a2 → r' (f a1) (f a2)) : +lemma partially_well_ordered_on.union (hs : s.partially_well_ordered_on r) + (ht : t.partially_well_ordered_on r) : + (s ∪ t).partially_well_ordered_on r := +begin + rintro f hf, + rcases nat.exists_subseq_of_forall_mem_union f hf with ⟨g, hgs|hgt⟩, + { rcases hs _ hgs with ⟨m, n, hlt, hr⟩, + exact ⟨g m, g n, g.strict_mono hlt, hr⟩ }, + { rcases ht _ hgt with ⟨m, n, hlt, hr⟩, + exact ⟨g m, g n, g.strict_mono hlt, hr⟩ } +end + +@[simp] lemma partially_well_ordered_on_union : + (s ∪ t).partially_well_ordered_on r ↔ + s.partially_well_ordered_on r ∧ t.partially_well_ordered_on r := +⟨λ h, ⟨h.mono $ subset_union_left _ _, h.mono $ subset_union_right _ _⟩, λ h, h.1.union h.2⟩ + +lemma partially_well_ordered_on.image_of_monotone_on (hs : s.partially_well_ordered_on r) + (hf : ∀ (a₁ ∈ s) (a₂ ∈ s), r a₁ a₂ → r' (f a₁) (f a₂)) : (f '' s).partially_well_ordered_on r' := -λ g hg, begin - have h := λ (n : ℕ), ((mem_image _ _ _).1 (hg (mem_range_self n))), - obtain ⟨m, n, hlt, hmn⟩ := hs (λ n, classical.some (h n)) _, - { refine ⟨m, n, hlt, _⟩, - rw [← (classical.some_spec (h m)).2, - ← (classical.some_spec (h n)).2], - exact hf _ _ (classical.some_spec (h m)).1 (classical.some_spec (h n)).1 hmn }, - { rintros _ ⟨n, rfl⟩, - exact (classical.some_spec (h n)).1 } +begin + intros g' hg', + choose g hgs heq using hg', + obtain rfl : f ∘ g = g', from funext heq, + obtain ⟨m, n, hlt, hmn⟩ := hs g hgs, + exact ⟨m, n, hlt, hf _ (hgs m) _ (hgs n) hmn⟩ end -lemma _root_.is_antichain.finite_of_partially_well_ordered_on {s : set α} {r : α → α → Prop} - (ha : is_antichain r s) (hp : s.partially_well_ordered_on r) : +lemma _root_.is_antichain.finite_of_partially_well_ordered_on (ha : is_antichain r s) + (hp : s.partially_well_ordered_on r) : s.finite := begin refine finite_or_infinite.resolve_right (λ hi, _), - obtain ⟨m, n, hmn, h⟩ := hp (λ n, hi.nat_embedding _ n) (range_subset_iff.2 $ - λ n, (hi.nat_embedding _ n).2), + obtain ⟨m, n, hmn, h⟩ := hp (λ n, hi.nat_embedding _ n) (λ n, (hi.nat_embedding _ n).2), exact hmn.ne ((hi.nat_embedding _).injective $ subtype.val_injective $ ha.eq (hi.nat_embedding _ m).2 (hi.nat_embedding _ n).2 h), end -lemma finite.partially_well_ordered_on {s : set α} {r : α → α → Prop} [is_refl α r] - (hs : s.finite) : - s.partially_well_ordered_on r := +section is_refl +variables [is_refl α r] + +protected lemma finite.partially_well_ordered_on (hs : s.finite) : s.partially_well_ordered_on r := begin intros f hf, - obtain ⟨m, n, hmn, h⟩ := hs.exists_lt_map_eq_of_range_subset hf, + obtain ⟨m, n, hmn, h⟩ := hs.exists_lt_map_eq_of_forall_mem hf, exact ⟨m, n, hmn, h.subst $ refl (f m)⟩, end -lemma _root_.is_antichain.partially_well_ordered_on_iff {s : set α} {r : α → α → Prop} [is_refl α r] - (hs : is_antichain r s) : +lemma _root_.is_antichain.partially_well_ordered_on_iff (hs : is_antichain r s) : s.partially_well_ordered_on r ↔ s.finite := ⟨hs.finite_of_partially_well_ordered_on, finite.partially_well_ordered_on⟩ -lemma partially_well_ordered_on_iff_finite_antichains {s : set α} {r : α → α → Prop} [is_refl α r] - [is_symm α r] : +@[simp] lemma partially_well_ordered_on_singleton (a : α) : partially_well_ordered_on {a} r := +(finite_singleton a).partially_well_ordered_on + +@[simp] lemma partially_well_ordered_on_insert : + partially_well_ordered_on (insert a s) r ↔ partially_well_ordered_on s r := +by simp only [← singleton_union, partially_well_ordered_on_union, + partially_well_ordered_on_singleton, true_and] + +protected lemma partially_well_ordered_on.insert (h : partially_well_ordered_on s r) (a : α) : + partially_well_ordered_on (insert a s) r := +partially_well_ordered_on_insert.2 h + +lemma partially_well_ordered_on_iff_finite_antichains [is_symm α r] : s.partially_well_ordered_on r ↔ ∀ t ⊆ s, is_antichain r t → t.finite := begin refine ⟨λ h t ht hrt, hrt.finite_of_partially_well_ordered_on (h.mono ht), _⟩, rintro hs f hf, by_contra' H, - refine set.infinite_range_of_injective (λ m n hmn, _) (hs _ hf _), + refine infinite_range_of_injective (λ m n hmn, _) (hs _ (range_subset_iff.2 hf) _), { obtain h | h | h := lt_trichotomy m n, { refine (H _ _ h _).elim, rw hmn, @@ -271,197 +281,209 @@ begin { exact mt symm (H _ _ h) } end -section partial_order -variables {s : set α} {t : set α} {r : α → α → Prop} +variables [is_trans α r] -theorem partially_well_ordered_on.exists_monotone_subseq [is_refl α r] [is_trans α r] - (h : s.partially_well_ordered_on r) (f : ℕ → α) (hf : range f ⊆ s) : - ∃ (g : ℕ ↪o ℕ), ∀ m n : ℕ, m ≤ n → r (f (g m)) (f (g n)) := +lemma partially_well_ordered_on.exists_monotone_subseq (h : s.partially_well_ordered_on r) + (f : ℕ → α) (hf : ∀ n, f n ∈ s) : + ∃ g : ℕ ↪o ℕ, ∀ m n : ℕ, m ≤ n → r (f (g m)) (f (g n)) := begin obtain ⟨g, h1 | h2⟩ := exists_increasing_or_nonincreasing_subseq r f, { refine ⟨g, λ m n hle, _⟩, - obtain hlt | heq := lt_or_eq_of_le hle, - { exact h1 m n hlt, }, - { rw [heq], - apply refl_of r } }, + obtain hlt | rfl := hle.lt_or_eq, + exacts [h1 m n hlt, refl_of r _] }, { exfalso, - obtain ⟨m, n, hlt, hle⟩ := h (f ∘ g) (subset.trans (range_comp_subset_range _ _) hf), + obtain ⟨m, n, hlt, hle⟩ := h (f ∘ g) (λ n, hf _), exact h2 m n hlt hle } end -theorem partially_well_ordered_on_iff_exists_monotone_subseq [is_refl α r] [is_trans α r] : +lemma partially_well_ordered_on_iff_exists_monotone_subseq : s.partially_well_ordered_on r ↔ - ∀ f : ℕ → α, range f ⊆ s → ∃ (g : ℕ ↪o ℕ), ∀ m n : ℕ, m ≤ n → r (f (g m)) (f (g n)) := + ∀ f : ℕ → α, (∀ n, f n ∈ s) → ∃ (g : ℕ ↪o ℕ), ∀ m n : ℕ, m ≤ n → r (f (g m)) (f (g n)) := begin classical, split; intros h f hf, { exact h.exists_monotone_subseq f hf }, { obtain ⟨g, gmon⟩ := h f hf, - refine ⟨g 0, g 1, g.lt_iff_lt.2 zero_lt_one, gmon _ _ zero_le_one⟩, } + exact ⟨g 0, g 1, g.lt_iff_lt.2 zero_lt_one, gmon _ _ zero_le_one⟩ } end -lemma partially_well_ordered_on.well_founded_on [is_partial_order α r] - (h : s.partially_well_ordered_on r) : - s.well_founded_on (λ a b, r a b ∧ a ≠ b) := +protected lemma partially_well_ordered_on.prod {t : set β} (hs : partially_well_ordered_on s r) + (ht : partially_well_ordered_on t r') : + partially_well_ordered_on (s ×ˢ t) (λ x y : α × β, r x.1 y.1 ∧ r' x.2 y.2) := begin - haveI : is_strict_order α (λ a b, r a b ∧ a ≠ b) := - { to_is_irrefl := ⟨λ a con, con.2 rfl⟩, - to_is_trans := ⟨λ a b c ab bc, ⟨trans ab.1 bc.1, - λ ac, ab.2 (antisymm ab.1 (ac.symm ▸ bc.1))⟩⟩ }, - rw well_founded_on_iff_no_descending_seq, - intros f con, - obtain ⟨m, n, hlt, hle⟩ := h f con, - exact (f.map_rel_iff.2 hlt).2 (antisymm hle (f.map_rel_iff.2 hlt).1).symm, + intros f hf, + obtain ⟨g₁, h₁⟩ := hs.exists_monotone_subseq (prod.fst ∘ f) (λ n, (hf n).1), + obtain ⟨m, n, hlt, hle⟩ := ht (prod.snd ∘ f ∘ g₁) (λ n, (hf _).2), + exact ⟨g₁ m, g₁ n, g₁.strict_mono hlt, h₁ _ _ hlt.le, hle⟩ end -variables [partial_order α] +end is_refl -lemma is_pwo.is_wf (h : s.is_pwo) : - s.is_wf := +lemma partially_well_ordered_on.well_founded_on [is_preorder α r] + (h : s.partially_well_ordered_on r) : + s.well_founded_on (λ a b, r a b ∧ ¬r b a) := begin - rw [is_wf], - convert h.well_founded_on, - ext x y, - rw lt_iff_le_and_ne, + letI : preorder α := { le := r, le_refl := refl_of r, le_trans := λ _ _ _, trans_of r }, + change s.well_founded_on (<), change s.partially_well_ordered_on (≤) at h, + rw well_founded_on_iff_no_descending_seq, + intros f hf, + obtain ⟨m, n, hlt, hle⟩ := h f hf, + exact (f.map_rel_iff.2 hlt).not_le hle, end -theorem is_pwo.exists_monotone_subseq - (h : s.is_pwo) (f : ℕ → α) (hf : range f ⊆ s) : +end partially_well_ordered_on + +section is_pwo +variables [preorder α] [preorder β] {s t : set α} + +/-- A subset of a preorder is partially well-ordered when any infinite sequence contains + a monotone subsequence of length 2 (or equivalently, an infinite monotone subsequence). -/ +def is_pwo (s : set α) : Prop := partially_well_ordered_on s (≤) + +lemma is_pwo.mono (ht : t.is_pwo) : s ⊆ t → s.is_pwo := ht.mono + +theorem is_pwo.exists_monotone_subseq (h : s.is_pwo) (f : ℕ → α) (hf : ∀ n, f n ∈ s) : ∃ (g : ℕ ↪o ℕ), monotone (f ∘ g) := h.exists_monotone_subseq f hf theorem is_pwo_iff_exists_monotone_subseq : - s.is_pwo ↔ - ∀ f : ℕ → α, range f ⊆ s → ∃ (g : ℕ ↪o ℕ), monotone (f ∘ g) := + s.is_pwo ↔ ∀ f : ℕ → α, (∀ n, f n ∈ s) → ∃ (g : ℕ ↪o ℕ), monotone (f ∘ g) := partially_well_ordered_on_iff_exists_monotone_subseq -lemma is_pwo.prod (hs : s.is_pwo) - (ht : t.is_pwo) : - (s ×ˢ t).is_pwo := -begin - classical, - rw is_pwo_iff_exists_monotone_subseq at *, - intros f hf, - obtain ⟨g1, h1⟩ := hs (prod.fst ∘ f) _, - swap, - { rw [range_comp, image_subset_iff], - refine subset.trans hf _, - rintros ⟨x1, x2⟩ hx, - simp only [mem_preimage, hx.1] }, - obtain ⟨g2, h2⟩ := ht (prod.snd ∘ f ∘ g1) _, - refine ⟨g2.trans g1, λ m n mn, _⟩, - swap, - { rw [range_comp, image_subset_iff], - refine subset.trans (range_comp_subset_range _ _) (subset.trans hf _), - rintros ⟨x1, x2⟩ hx, - simp only [mem_preimage, hx.2] }, - simp only [rel_embedding.coe_trans, function.comp_app], - exact ⟨h1 (g2.le_iff_le.2 mn), h2 mn⟩, -end +protected lemma is_pwo.is_wf (h : s.is_pwo) : s.is_wf := +by simpa only [← lt_iff_le_not_le] using h.well_founded_on -theorem is_pwo.image_of_monotone {β : Type*} [partial_order β] - (hs : s.is_pwo) {f : α → β} (hf : monotone f) : - is_pwo (f '' s) := -hs.image_of_monotone_on (λ _ _ _ _ ab, hf ab) +lemma is_pwo.prod {t : set β} (hs : s.is_pwo) (ht : t.is_pwo) : is_pwo (s ×ˢ t) := hs.prod ht -theorem is_pwo.union (hs : is_pwo s) (ht : is_pwo t) : is_pwo (s ∪ t) := -begin - classical, - rw [is_pwo_iff_exists_monotone_subseq] at *, - rintros f fst, - have h : (f ⁻¹' s).infinite ∨ (f ⁻¹' t).infinite, - { have h : (univ : set ℕ).infinite := infinite_univ, - have hpre : f ⁻¹' (s ∪ t) = set.univ, - { rw [← image_univ, image_subset_iff, univ_subset_iff] at fst, - exact fst }, - rw preimage_union at hpre, - rw ← hpre at h, - rw [set.infinite, set.infinite], - rw set.infinite at h, - contrapose! h, - exact finite.union h.1 h.2, }, - rw [← infinite_coe_iff, ← infinite_coe_iff] at h, - cases h with inf inf; haveI := inf, - { obtain ⟨g, hg⟩ := hs (f ∘ (nat.order_embedding_of_set (f ⁻¹' s))) _, - { rw [function.comp.assoc, ← rel_embedding.coe_trans] at hg, - exact ⟨_, hg⟩ }, - rw [range_comp, image_subset_iff], - simp }, - { obtain ⟨g, hg⟩ := ht (f ∘ (nat.order_embedding_of_set (f ⁻¹' t))) _, - { rw [function.comp.assoc, ← rel_embedding.coe_trans] at hg, - exact ⟨_, hg⟩ }, - rw [range_comp, image_subset_iff], - simp } -end +lemma is_pwo.image_of_monotone_on (hs : s.is_pwo) {f : α → β} (hf : monotone_on f s) : + is_pwo (f '' s) := +hs.image_of_monotone_on hf + +lemma is_pwo.image_of_monotone (hs : s.is_pwo) {f : α → β} (hf : monotone f) : is_pwo (f '' s) := +hs.image_of_monotone_on (hf.monotone_on _) + +protected lemma is_pwo.union (hs : is_pwo s) (ht : is_pwo t) : is_pwo (s ∪ t) := hs.union ht + +@[simp] lemma is_pwo_union : is_pwo (s ∪ t) ↔ is_pwo s ∧ is_pwo t := partially_well_ordered_on_union + +protected lemma finite.is_pwo (hs : s.finite) : is_pwo s := hs.partially_well_ordered_on +@[simp] lemma is_pwo_of_finite [finite α] : s.is_pwo := s.to_finite.is_pwo + +@[simp] lemma is_pwo_singleton (a : α) : is_pwo ({a} : set α) := (finite_singleton a).is_pwo + +@[simp] lemma is_pwo_empty : is_pwo (∅ : set α) := finite_empty.is_pwo + +protected lemma subsingleton.is_pwo (hs : s.subsingleton) : is_pwo s := hs.finite.is_pwo + +@[simp] lemma is_pwo_insert {a} : is_pwo (insert a s) ↔ is_pwo s := +by simp only [←singleton_union, is_pwo_union, is_pwo_singleton, true_and] + +protected lemma is_pwo.insert (h : is_pwo s) (a : α) : is_pwo (insert a s) := is_pwo_insert.2 h + +protected lemma finite.is_wf (hs : s.finite) : is_wf s := hs.is_pwo.is_wf +@[simp] lemma is_wf_singleton {a : α} : is_wf ({a} : set α) := (finite_singleton a).is_wf +protected lemma subsingleton.is_wf (hs : s.subsingleton) : is_wf s := hs.is_pwo.is_wf + +@[simp] lemma is_wf_insert {a} : is_wf (insert a s) ↔ is_wf s := +by simp only [←singleton_union, is_wf_union, is_wf_singleton, true_and] -end partial_order - -theorem is_wf.is_pwo [linear_order α] {s : set α} - (hs : s.is_wf) : s.is_pwo := -λ f hf, begin - rw [is_wf, well_founded_on_iff] at hs, - have hrange : (range f).nonempty := ⟨f 0, mem_range_self 0⟩, - let a := hs.min (range f) hrange, - obtain ⟨m, hm⟩ := hs.min_mem (range f) hrange, - refine ⟨m, m.succ, m.lt_succ_self, le_of_not_lt (λ con, _)⟩, - rw hm at con, - apply hs.not_lt_min (range f) hrange (mem_range_self m.succ) - ⟨con, hf (mem_range_self m.succ), hf _⟩, - rw ← hm, - apply mem_range_self, +lemma is_wf.insert (h : is_wf s) (a : α) : is_wf (insert a s) := is_wf_insert.2 h + +end is_pwo + +section well_founded_on +variables {r : α → α → Prop} [is_strict_order α r] {s : set α} {a : α} + +protected lemma finite.well_founded_on (hs : s.finite) : s.well_founded_on r := +by { letI := partial_order_of_SO r, exact hs.is_wf } + +@[simp] lemma well_founded_on_singleton : well_founded_on ({a} : set α) r := +(finite_singleton a).well_founded_on + +protected lemma subsingleton.well_founded_on (hs : s.subsingleton) : s.well_founded_on r := +hs.finite.well_founded_on + +@[simp] lemma well_founded_on_insert : well_founded_on (insert a s) r ↔ well_founded_on s r := +by simp only [←singleton_union, well_founded_on_union, well_founded_on_singleton, true_and] + +lemma well_founded_on.insert (h : well_founded_on s r) (a : α) : well_founded_on (insert a s) r := +well_founded_on_insert.2 h + +end well_founded_on + +section linear_order +variables [linear_order α] {s : set α} + +protected lemma is_wf.is_pwo (hs : s.is_wf) : s.is_pwo := +begin + intros f hf, + lift f to ℕ → s using hf, + have hrange : (range f).nonempty := range_nonempty _, + rcases hs.has_min (range f) (range_nonempty _) with ⟨_, ⟨m, rfl⟩, hm⟩, + simp only [forall_range_iff, not_lt] at hm, + exact ⟨m, m + 1, lt_add_one m, hm _⟩, end -theorem is_wf_iff_is_pwo [linear_order α] {s : set α} : - s.is_wf ↔ s.is_pwo := -⟨is_wf.is_pwo, is_pwo.is_wf⟩ +/-- In a linear order, the predicates `set.is_wf` and `set.is_pwo` are equivalent. -/ +lemma is_wf_iff_is_pwo : s.is_wf ↔ s.is_pwo := ⟨is_wf.is_pwo, is_pwo.is_wf⟩ +end linear_order end set namespace finset +variables {r : α → α → Prop} -@[simp] lemma partially_well_ordered_on {r : α → α → Prop} [is_refl α r] (s : finset α) : +@[simp] protected lemma partially_well_ordered_on [is_refl α r] (s : finset α) : (s : set α).partially_well_ordered_on r := s.finite_to_set.partially_well_ordered_on -@[simp] -theorem is_pwo [partial_order α] (f : finset α) : - set.is_pwo (↑f : set α) := -f.partially_well_ordered_on +@[simp] protected lemma is_pwo [preorder α] (s : finset α) : set.is_pwo (↑s : set α) := +s.partially_well_ordered_on -@[simp] -theorem well_founded_on {r : α → α → Prop} [is_strict_order α r] (f : finset α) : - set.well_founded_on (↑f : set α) r := -begin - rw [set.well_founded_on_iff_no_descending_seq], - intros g con, - apply set.infinite_of_injective_forall_mem g.injective (set.range_subset_iff.1 con), - exact f.finite_to_set, -end +@[simp] protected lemma is_wf [preorder α] (s : finset α) : set.is_wf (↑s : set α) := +s.finite_to_set.is_wf -@[simp] -theorem is_wf [partial_order α] (f : finset α) : set.is_wf (↑f : set α) := -f.is_pwo.is_wf +@[simp] protected lemma well_founded_on [is_strict_order α r] (s : finset α) : + set.well_founded_on (↑s : set α) r := +by { letI := partial_order_of_SO r, exact s.is_wf } -end finset +lemma well_founded_on_sup [is_strict_order α r] (s : finset ι) {f : ι → set α} : + (s.sup f).well_founded_on r ↔ ∀ i ∈ s, (f i).well_founded_on r := +finset.cons_induction_on s (by simp) $ λ a s ha hs, by simp [-sup_set_eq_bUnion, hs] -namespace set -variables [partial_order α] {s : set α} {a : α} +lemma partially_well_ordered_on_sup (s : finset ι) {f : ι → set α} : + (s.sup f).partially_well_ordered_on r ↔ ∀ i ∈ s, (f i).partially_well_ordered_on r := +finset.cons_induction_on s (by simp) $ λ a s ha hs, by simp [-sup_set_eq_bUnion, hs] -theorem finite.is_pwo (h : s.finite) : s.is_pwo := -begin - rw ← h.coe_to_finset, - exact h.to_finset.is_pwo, -end +lemma is_wf_sup [preorder α] (s : finset ι) {f : ι → set α} : + (s.sup f).is_wf ↔ ∀ i ∈ s, (f i).is_wf := +s.well_founded_on_sup -@[simp] -theorem fintype.is_pwo [fintype α] : s.is_pwo := s.to_finite.is_pwo +lemma is_pwo_sup [preorder α] (s : finset ι) {f : ι → set α} : + (s.sup f).is_pwo ↔ ∀ i ∈ s, (f i).is_pwo := +s.partially_well_ordered_on_sup -@[simp] -theorem is_pwo_singleton (a) : is_pwo ({a} : set α) := -(finite_singleton a).is_pwo +@[simp] lemma well_founded_on_bUnion [is_strict_order α r] (s : finset ι) {f : ι → set α} : + (⋃ i ∈ s, f i).well_founded_on r ↔ ∀ i ∈ s, (f i).well_founded_on r := +by simpa only [finset.sup_eq_supr] using s.well_founded_on_sup + +@[simp] lemma partially_well_ordered_on_bUnion (s : finset ι) {f : ι → set α} : + (⋃ i ∈ s, f i).partially_well_ordered_on r ↔ ∀ i ∈ s, (f i).partially_well_ordered_on r := +by simpa only [finset.sup_eq_supr] using s.partially_well_ordered_on_sup + +@[simp] lemma is_wf_bUnion [preorder α] (s : finset ι) {f : ι → set α} : + (⋃ i ∈ s, f i).is_wf ↔ ∀ i ∈ s, (f i).is_wf := +s.well_founded_on_bUnion -theorem is_pwo.insert (a) (hs : is_pwo s) : is_pwo (insert a s) := -by { rw ← union_singleton, exact hs.union (is_pwo_singleton a) } +@[simp] lemma is_pwo_bUnion [preorder α] (s : finset ι) {f : ι → set α} : + (⋃ i ∈ s, f i).is_pwo ↔ ∀ i ∈ s, (f i).is_pwo := +s.partially_well_ordered_on_bUnion + +end finset + +namespace set +section preorder +variables [preorder α] {s : set α} {a : α} /-- `is_wf.min` returns a minimal element of a nonempty well-founded set. -/ noncomputable def is_wf.min (hs : is_wf s) (hn : s.nonempty) : α := @@ -478,26 +500,15 @@ lemma is_wf_min_singleton (a) {hs : is_wf ({a} : set α)} {hn : ({a} : set α).n hs.min hn = a := eq_of_mem_singleton (is_wf.min_mem hs hn) -end set - -theorem finset.is_wf_sup {ι : Type*} [partial_order α] (f : finset ι) (g : ι → set α) - (hf : ∀ i : ι, i ∈ f → (g i).is_wf) : (f.sup g).is_wf := -finset.sup_induction set.is_pwo_empty.is_wf (λ a ha b hb, ha.union hb) hf - -theorem finset.is_pwo_sup {ι : Type*} [partial_order α] (f : finset ι) (g : ι → set α) - (hf : ∀ i : ι, i ∈ f → (g i).is_pwo) : (f.sup g).is_pwo := -finset.sup_induction set.is_pwo_empty (λ a ha b hb, ha.union hb) hf +end preorder -namespace set +section linear_order variables [linear_order α] {s t : set α} {a : α} -lemma is_wf.min_le - (hs : s.is_wf) (hn : s.nonempty) (ha : a ∈ s) : hs.min hn ≤ a := +lemma is_wf.min_le (hs : s.is_wf) (hn : s.nonempty) (ha : a ∈ s) : hs.min hn ≤ a := le_of_not_lt (hs.not_lt_min hn ha) -lemma is_wf.le_min_iff - (hs : s.is_wf) (hn : s.nonempty) : - a ≤ hs.min hn ↔ ∀ b, b ∈ s → a ≤ b := +lemma is_wf.le_min_iff (hs : s.is_wf) (hn : s.nonempty) : a ≤ hs.min hn ↔ ∀ b, b ∈ s → a ≤ b := ⟨λ ha b hb, le_trans ha (hs.min_le hn hb), λ h, h _ (hs.min_mem _)⟩ lemma is_wf.min_le_min_of_subset @@ -515,60 +526,28 @@ begin (union_nonempty.2 (or.intro_left _ hsn)))).imp (hs.min_le _) (ht.min_le _), end +end linear_order end set -namespace set - -variables {s : set α} {t : set α} +open set -@[to_additive] -theorem is_pwo.mul [ordered_cancel_comm_monoid α] (hs : s.is_pwo) (ht : t.is_pwo) : - is_pwo (s * t) := -begin - rw ← image_mul_prod, - exact (is_pwo.prod hs ht).image_of_monotone (λ _ _ h, mul_le_mul' h.1 h.2), -end - -variable [linear_ordered_cancel_comm_monoid α] - -@[to_additive] -theorem is_wf.mul (hs : s.is_wf) (ht : t.is_wf) : is_wf (s * t) := -(hs.is_pwo.mul ht.is_pwo).is_wf - -@[to_additive] -theorem is_wf.min_mul (hs : s.is_wf) (ht : t.is_wf) (hsn : s.nonempty) (htn : t.nonempty) : - (hs.mul ht).min (hsn.mul htn) = hs.min hsn * ht.min htn := -begin - refine le_antisymm (is_wf.min_le _ _ (mem_mul.2 ⟨_, _, hs.min_mem _, ht.min_mem _, rfl⟩)) _, - rw is_wf.le_min_iff, - rintros _ ⟨x, y, hx, hy, rfl⟩, - exact mul_le_mul' (hs.min_le _ hx) (ht.min_le _ hy), -end - -end set - -namespace set -namespace partially_well_ordered_on +namespace set.partially_well_ordered_on +variables {r : α → α → Prop} /-- In the context of partial well-orderings, a bad sequence is a nonincreasing sequence whose range is contained in a particular set `s`. One exists if and only if `s` is not partially well-ordered. -/ def is_bad_seq (r : α → α → Prop) (s : set α) (f : ℕ → α) : Prop := -set.range f ⊆ s ∧ ∀ (m n : ℕ), m < n → ¬ r (f m) (f n) +(∀ n, f n ∈ s) ∧ ∀ m n : ℕ, m < n → ¬ r (f m) (f n) lemma iff_forall_not_is_bad_seq (r : α → α → Prop) (s : set α) : - s.partially_well_ordered_on r ↔ - ∀ f, ¬ is_bad_seq r s f := -begin - rw [set.partially_well_ordered_on], - apply forall_congr (λ f, _), - simp [is_bad_seq] -end + s.partially_well_ordered_on r ↔ ∀ f, ¬ is_bad_seq r s f := +forall_congr $ λ f, by simp [is_bad_seq] /-- This indicates that every bad sequence `g` that agrees with `f` on the first `n` terms has `rk (f n) ≤ rk (g n)`. -/ def is_min_bad_seq (r : α → α → Prop) (rk : α → ℕ) (s : set α) (n : ℕ) (f : ℕ → α) : Prop := - ∀ g : ℕ → α, (∀ (m : ℕ), m < n → f m = g m) → rk (g n) < rk (f n) → ¬ is_bad_seq r s g +∀ g : ℕ → α, (∀ (m : ℕ), m < n → f m = g m) → rk (g n) < rk (f n) → ¬ is_bad_seq r s g /-- Given a bad sequence `f`, this constructs a bad sequence that agrees with `f` on the first `n` terms and is minimal at `n`. @@ -606,8 +585,7 @@ begin rw [ih, ((min_bad_seq_of_bad_seq r rk s (m + k).succ (fs (m + k)).1 (fs (m + k)).2.1).2.1 m (nat.lt_succ_iff.2 (nat.add_le_add_left k.zero_le m)))], refl }, - refine ⟨λ n, (fs n).1 n, ⟨set.range_subset_iff.2 (λ n, ((fs n).2).1.1 (mem_range_self n)), - λ m n mn, _⟩, λ n g hg1 hg2, _⟩, + refine ⟨λ n, (fs n).1 n, ⟨(λ n, ((fs n).2).1.1 n), λ m n mn, _⟩, λ n g hg1 hg2, _⟩, { dsimp, rw [← subtype.val_eq_coe, h m n (le_of_lt mn)], convert (fs n).2.1.2 m n mn }, @@ -615,7 +593,7 @@ begin rw ← h m n (le_of_lt mn) }, end -lemma iff_not_exists_is_min_bad_seq {r : α → α → Prop} (rk : α → ℕ) {s : set α} : +lemma iff_not_exists_is_min_bad_seq (rk : α → ℕ) {s : set α} : s.partially_well_ordered_on r ↔ ¬ ∃ f, is_bad_seq r s f ∧ ∀ n, is_min_bad_seq r rk s n f := begin rw [iff_forall_not_is_bad_seq, ← not_exists, not_congr], @@ -646,7 +624,7 @@ begin λ n con, (hf1).2 n n.succ n.lt_succ_self (con.symm ▸ list.sublist_forall₂.nil), obtain ⟨g, hg⟩ := h.exists_monotone_subseq (list.head ∘ f) _, swap, { simp only [set.range_subset_iff, function.comp_apply], - exact λ n, hf1.1 (set.mem_range_self n) _ (list.head_mem_self (hnil n)) }, + exact λ n, hf1.1 n _ (list.head_mem_self (hnil n)) }, have hf' := hf2 (g 0) (λ n, if n < g 0 then f n else list.tail (f (g (n - g 0)))) (λ m hm, (if_pos hm).symm) _, swap, { simp only [if_neg (lt_irrefl (g 0)), tsub_self], @@ -655,11 +633,11 @@ begin rw [is_bad_seq] at hf', push_neg at hf', obtain ⟨m, n, mn, hmn⟩ := hf' _, - swap, { rw set.range_subset_iff, - rintro n x hx, + swap, + { rintro n x hx, split_ifs at hx with hn hn, - { exact hf1.1 (set.mem_range_self _) _ hx }, - { refine hf1.1 (set.mem_range_self _) _ (list.tail_subset _ hx), } }, + { exact hf1.1 _ _ hx }, + { refine hf1.1 _ _ (list.tail_subset _ hx), } }, by_cases hn : n < g 0, { apply hf1.2 m n mn, rwa [if_pos hn, if_pos (mn.trans hn)] at hmn }, @@ -674,231 +652,36 @@ begin exact list.sublist_forall₂.cons (hg _ _ (le_of_lt mn)) hmn, } } end -end partially_well_ordered_on - -namespace is_pwo - -@[to_additive] -lemma submonoid_closure [ordered_cancel_comm_monoid α] {s : set α} (hpos : ∀ x : α, x ∈ s → 1 ≤ x) - (h : s.is_pwo) : is_pwo ((submonoid.closure s) : set α) := -begin - have hl : ((submonoid.closure s) : set α) ⊆ list.prod '' { l : list α | ∀ x, x ∈ l → x ∈ s }, - { intros x hx, - rw set_like.mem_coe at hx, - refine submonoid.closure_induction hx (λ x hx, ⟨_, λ y hy, _, list.prod_singleton⟩) - ⟨_, λ y hy, (list.not_mem_nil _ hy).elim, list.prod_nil⟩ _, - { rwa list.mem_singleton.1 hy }, - rintros _ _ ⟨l, hl, rfl⟩ ⟨l', hl', rfl⟩, - refine ⟨_, λ y hy, _, list.prod_append⟩, - cases list.mem_append.1 hy with hy hy, - { exact hl _ hy }, - { exact hl' _ hy } }, - apply ((h.partially_well_ordered_on_sublist_forall₂ (≤)).image_of_monotone_on _).mono hl, - exact λ l1 l2 hl1 hl2 h12, h12.prod_le_prod' (λ x hx, hpos x $ hl2 x hx) -end - -end is_pwo - -/-- `set.mul_antidiagonal s t a` is the set of all pairs of an element in `s` and an element in `t` - that multiply to `a`. -/ -@[to_additive "`set.add_antidiagonal s t a` is the set of all pairs of an element in `s` - and an element in `t` that add to `a`."] -def mul_antidiagonal [monoid α] (s t : set α) (a : α) : set (α × α) := -{ x | x.1 * x.2 = a ∧ x.1 ∈ s ∧ x.2 ∈ t } - -namespace mul_antidiagonal - -@[simp, to_additive] -lemma mem_mul_antidiagonal [monoid α] {s t : set α} {a : α} {x : α × α} : - x ∈ mul_antidiagonal s t a ↔ x.1 * x.2 = a ∧ x.1 ∈ s ∧ x.2 ∈ t := iff.refl _ - -section cancel_comm_monoid -variables [cancel_comm_monoid α] {s t : set α} {a : α} - -@[to_additive] -lemma fst_eq_fst_iff_snd_eq_snd {x y : (mul_antidiagonal s t a)} : - (x : α × α).fst = (y : α × α).fst ↔ (x : α × α).snd = (y : α × α).snd := -⟨λ h, begin - have hx := x.2.1, - rw [subtype.val_eq_coe, h] at hx, - apply mul_left_cancel (hx.trans y.2.1.symm), -end, λ h, begin - have hx := x.2.1, - rw [subtype.val_eq_coe, h] at hx, - apply mul_right_cancel (hx.trans y.2.1.symm), -end⟩ - -@[to_additive] -lemma eq_of_fst_eq_fst {x y : (mul_antidiagonal s t a)} - (h : (x : α × α).fst = (y : α × α).fst) : x = y := -subtype.ext (prod.ext h (mul_antidiagonal.fst_eq_fst_iff_snd_eq_snd.1 h)) - -@[to_additive] -lemma eq_of_snd_eq_snd {x y : (mul_antidiagonal s t a)} - (h : (x : α × α).snd = (y : α × α).snd) : x = y := -subtype.ext (prod.ext (mul_antidiagonal.fst_eq_fst_iff_snd_eq_snd.2 h) h) - -end cancel_comm_monoid - -section ordered_cancel_comm_monoid -variables [ordered_cancel_comm_monoid α] (s t : set α) (a : α) - -@[to_additive] -lemma eq_of_fst_le_fst_of_snd_le_snd {x y : (mul_antidiagonal s t a)} - (h1 : (x : α × α).fst ≤ (y : α × α).fst) (h2 : (x : α × α).snd ≤ (y : α × α).snd ) : - x = y := -begin - apply eq_of_fst_eq_fst, - cases eq_or_lt_of_le h1 with heq hlt, - { exact heq }, - exfalso, - exact ne_of_lt (mul_lt_mul_of_lt_of_le hlt h2) - ((mem_mul_antidiagonal.1 x.2).1.trans (mem_mul_antidiagonal.1 y.2).1.symm) -end - -variables {s} {t} - -@[to_additive] -theorem finite_of_is_pwo (hs : s.is_pwo) (ht : t.is_pwo) (a) : - (mul_antidiagonal s t a).finite := -begin - by_contra h, - rw [← set.infinite] at h, - have h1 : (mul_antidiagonal s t a).partially_well_ordered_on (prod.fst ⁻¹'o (≤)), - { intros f hf, - refine hs (prod.fst ∘ f) _, - rw range_comp, - rintros _ ⟨⟨x, y⟩, hxy, rfl⟩, - exact (mem_mul_antidiagonal.1 (hf hxy)).2.1 }, - have h2 : (mul_antidiagonal s t a).partially_well_ordered_on (prod.snd ⁻¹'o (≤)), - { intros f hf, - refine ht (prod.snd ∘ f) _, - rw range_comp, - rintros _ ⟨⟨x, y⟩, hxy, rfl⟩, - exact (mem_mul_antidiagonal.1 (hf hxy)).2.2 }, - obtain ⟨g, hg⟩ := h1.exists_monotone_subseq (λ x, h.nat_embedding _ x) _, - swap, { rintro _ ⟨k, rfl⟩, - exact ((infinite.nat_embedding (s.mul_antidiagonal t a) h) _).2 }, - obtain ⟨m, n, mn, h2'⟩ := h2 (λ x, (h.nat_embedding _) (g x)) _, - swap, { rintro _ ⟨k, rfl⟩, - exact ((infinite.nat_embedding (s.mul_antidiagonal t a) h) _).2, }, - apply ne_of_lt mn (g.injective ((h.nat_embedding _).injective _)), - exact eq_of_fst_le_fst_of_snd_le_snd _ _ _ (hg _ _ (le_of_lt mn)) h2', -end - -end ordered_cancel_comm_monoid - -@[to_additive] -theorem finite_of_is_wf [linear_ordered_cancel_comm_monoid α] {s t : set α} - (hs : s.is_wf) (ht : t.is_wf) (a) : - (mul_antidiagonal s t a).finite := -finite_of_is_pwo hs.is_pwo ht.is_pwo a - -end mul_antidiagonal -end set - -namespace finset - -variables [ordered_cancel_comm_monoid α] -variables {s t : set α} (hs : s.is_pwo) (ht : t.is_pwo) (a : α) - -/-- `finset.mul_antidiagonal_of_is_wf hs ht a` is the set of all pairs of an element in - `s` and an element in `t` that multiply to `a`, but its construction requires proofs - `hs` and `ht` that `s` and `t` are well-ordered. -/ -@[to_additive "`finset.add_antidiagonal_of_is_wf hs ht a` is the set of all pairs of an element in - `s` and an element in `t` that add to `a`, but its construction requires proofs - `hs` and `ht` that `s` and `t` are well-ordered."] -noncomputable def mul_antidiagonal : finset (α × α) := -(set.mul_antidiagonal.finite_of_is_pwo hs ht a).to_finset - -variables {hs} {ht} {u : set α} {hu : u.is_pwo} {a} {x : α × α} - -@[simp, to_additive] -lemma mem_mul_antidiagonal : - x ∈ mul_antidiagonal hs ht a ↔ x.1 * x.2 = a ∧ x.1 ∈ s ∧ x.2 ∈ t := -by simp [mul_antidiagonal] - -@[to_additive] -lemma mul_antidiagonal_mono_left (hus : u ⊆ s) : - (finset.mul_antidiagonal hu ht a) ⊆ (finset.mul_antidiagonal hs ht a) := -λ x hx, begin - rw mem_mul_antidiagonal at *, - exact ⟨hx.1, hus hx.2.1, hx.2.2⟩, -end - -@[to_additive] -lemma mul_antidiagonal_mono_right (hut : u ⊆ t) : - (finset.mul_antidiagonal hs hu a) ⊆ (finset.mul_antidiagonal hs ht a) := -λ x hx, begin - rw mem_mul_antidiagonal at *, - exact ⟨hx.1, hx.2.1, hut hx.2.2⟩, -end - -@[to_additive] -lemma support_mul_antidiagonal_subset_mul : - { a : α | (mul_antidiagonal hs ht a).nonempty } ⊆ s * t := -(λ x ⟨⟨a1, a2⟩, ha⟩, begin - obtain ⟨hmul, h1, h2⟩ := mem_mul_antidiagonal.1 ha, - exact ⟨a1, a2, h1, h2, hmul⟩, -end) - -@[to_additive] -theorem is_pwo_support_mul_antidiagonal : - { a : α | (mul_antidiagonal hs ht a).nonempty }.is_pwo := -(hs.mul ht).mono support_mul_antidiagonal_subset_mul - -@[to_additive] -theorem mul_antidiagonal_min_mul_min {α} [linear_ordered_cancel_comm_monoid α] {s t : set α} - (hs : s.is_wf) (ht : t.is_wf) - (hns : s.nonempty) (hnt : t.nonempty) : - mul_antidiagonal hs.is_pwo ht.is_pwo ((hs.min hns) * (ht.min hnt)) = - {(hs.min hns, ht.min hnt)} := -begin - ext ⟨a1, a2⟩, - rw [mem_mul_antidiagonal, finset.mem_singleton, prod.ext_iff], - split, - { rintro ⟨hast, has, hat⟩, - cases eq_or_lt_of_le (hs.min_le hns has) with heq hlt, - { refine ⟨heq.symm, _⟩, - rw heq at hast, - exact mul_left_cancel hast }, - { contrapose hast, - exact ne_of_gt (mul_lt_mul_of_lt_of_le hlt (ht.min_le hnt hat)) } }, - { rintro ⟨ha1, ha2⟩, - rw [ha1, ha2], - exact ⟨rfl, hs.min_mem _, ht.min_mem _⟩ } -end - -end finset +end set.partially_well_ordered_on -lemma well_founded.is_wf [has_lt α] (h : well_founded ((<) : α → α → Prop)) (s : set α) : - s.is_wf := -(set.is_wf_univ_iff.2 h).mono (set.subset_univ s) +lemma well_founded.is_wf [has_lt α] (h : well_founded ((<) : α → α → Prop)) (s : set α) : s.is_wf := +(set.is_wf_univ_iff.2 h).mono s.subset_univ /-- A version of **Dickson's lemma** any subset of functions `Π s : σ, α s` is partially well ordered, when `σ` is a `fintype` and each `α s` is a linear well order. This includes the classical case of Dickson's lemma that `ℕ ^ n` is a well partial order. -Some generalizations would be possible based on this proof, to include cases where the target -is partially well ordered, and also to consider the case of `partially_well_ordered_on` instead of -`is_pwo`. -/ -lemma pi.is_pwo {σ : Type*} {α : σ → Type*} [∀ s, linear_order (α s)] [∀ s, is_well_order (α s) (<)] - [fintype σ] (S : set (Π s : σ, α s)) : S.is_pwo := +Some generalizations would be possible based on this proof, to include cases where the target is +partially well ordered, and also to consider the case of `set.partially_well_ordered_on` instead of +`set.is_pwo`. -/ +lemma pi.is_pwo {α : ι → Type*} [Π i, linear_order (α i)] [∀ i, is_well_order (α i) (<)] [finite ι] + (s : set (Π i, α i)) : s.is_pwo := begin classical, - refine set.is_pwo.mono _ (set.subset_univ _), - rw set.is_pwo_iff_exists_monotone_subseq, - simp_rw [monotone, pi.le_def], - suffices : ∀ s : finset σ, ∀ (f : ℕ → (Π s, α s)), set.range f ⊆ set.univ → ∃ (g : ℕ ↪o ℕ), - ∀ ⦃a b : ℕ⦄, a ≤ b → ∀ (x : σ) (hs : x ∈ s), (f ∘ g) a x ≤ (f ∘ g) b x, - { simpa only [forall_true_left, finset.mem_univ] using this finset.univ, }, + casesI nonempty_fintype ι, + refine is_pwo.mono _ (subset_univ _), + rw is_pwo_iff_exists_monotone_subseq, + simp only [mem_univ, forall_const, function.comp_app], + suffices : ∀ s : finset ι, ∀ (f : ℕ → Π s, α s), ∃ g : ℕ ↪o ℕ, + ∀ ⦃a b : ℕ⦄, a ≤ b → ∀ (x : ι) (hs : x ∈ s), (f ∘ g) a x ≤ (f ∘ g) b x, + { simpa only [forall_true_left, finset.mem_univ] using this finset.univ }, apply' finset.induction, - { intros f hf, existsi rel_embedding.refl (≤), + { intros f, existsi rel_embedding.refl (≤), simp only [is_empty.forall_iff, implies_true_iff, forall_const, finset.not_mem_empty], }, - { intros x s hx ih f hf, - obtain ⟨g, hg⟩ := (is_well_founded.wf.is_wf (set.univ : set _)).is_pwo.exists_monotone_subseq - ((λ mo : Π s : σ, α s, mo x) ∘ f) (set.subset_univ _), - obtain ⟨g', hg'⟩ := ih (f ∘ g) (set.subset_univ _), + { intros x s hx ih f, + obtain ⟨g, hg⟩ := (is_well_founded.wf.is_wf univ).is_pwo.exists_monotone_subseq (λ n, f n x) + mem_univ, + obtain ⟨g', hg'⟩ := ih (f ∘ g), refine ⟨g'.trans g, λ a b hab, _⟩, simp only [finset.mem_insert, rel_embedding.coe_trans, function.comp_app, forall_eq_or_imp], - exact ⟨hg (order_hom_class.mono g' hab), hg' hab⟩, }, + exact ⟨hg (order_hom_class.mono g' hab), hg' hab⟩ } end diff --git a/src/probability/cond_count.lean b/src/probability/cond_count.lean index 97ec5c4acea2b..2a708cef40cb8 100644 --- a/src/probability/cond_count.lean +++ b/src/probability/cond_count.lean @@ -118,7 +118,7 @@ begin { exact this ▸ λ x hx, hx.2 }, rw ← @set.finite.to_finset_inj _ _ _ (hsf.inter_of_left _) hsf, exact finset.eq_of_subset_of_card_le - (set.finite.to_finset_mono.2 (s.inter_subset_left t)) h.symm.le + (set.finite.to_finset_subset.2 (s.inter_subset_left t)) h.symm.le end lemma cond_count_eq_zero_iff (hs : s.finite) : diff --git a/src/ring_theory/hahn_series.lean b/src/ring_theory/hahn_series.lean index 2a8c0a09b8ae7..a622d351842cf 100644 --- a/src/ring_theory/hahn_series.lean +++ b/src/ring_theory/hahn_series.lean @@ -524,9 +524,8 @@ begin rw mul_coeff, apply sum_subset_zero_on_sdiff (add_antidiagonal_mono_right hys) _ (λ _ _, rfl), intros b hb, - simp only [not_and, not_not, mem_sdiff, mem_add_antidiagonal, - ne.def, set.mem_set_of_eq, mem_support] at hb, - rw [(hb.2 hb.1.1 hb.1.2.1), mul_zero] + simp only [not_and, mem_sdiff, mem_add_antidiagonal, mem_support, not_imp_not] at hb, + rw [hb.2 hb.1.1 hb.1.2.2, mul_zero], end lemma mul_coeff_left' [non_unital_non_assoc_semiring R] {x y : hahn_series Γ R} {a : Γ} {s : set Γ} @@ -537,9 +536,8 @@ begin rw mul_coeff, apply sum_subset_zero_on_sdiff (add_antidiagonal_mono_left hxs) _ (λ _ _, rfl), intros b hb, - simp only [not_and, not_not, mem_sdiff, mem_add_antidiagonal, - ne.def, set.mem_set_of_eq, mem_support] at hb, - rw [not_not.1 (λ con, hb.1.2.2 (hb.2 hb.1.1 con)), zero_mul], + simp only [not_and', mem_sdiff, mem_add_antidiagonal, mem_support, not_ne_iff] at hb, + rw [hb.2 ⟨hb.1.2.1, hb.1.2.2⟩, zero_mul], end instance [non_unital_non_assoc_semiring R] : distrib (hahn_series Γ R) := @@ -583,7 +581,7 @@ begin ext ⟨a1, a2⟩, simp only [not_mem_empty, not_and, set.mem_singleton_iff, not_not, mem_add_antidiagonal, set.mem_set_of_eq, iff_false], - rintro h1 rfl h2, + rintro rfl h2 h1, rw add_comm at h1, rw ← add_right_cancel h1 at hx, exact h2 hx, }, @@ -593,12 +591,11 @@ begin simp only [set.mem_singleton_iff, prod.mk.inj_iff, mem_add_antidiagonal, mem_singleton, set.mem_set_of_eq], split, - { rintro ⟨h1, rfl, h2⟩, + { rintro ⟨rfl, h2, h1⟩, rw add_comm at h1, refine ⟨rfl, add_right_cancel h1⟩ }, { rintro ⟨rfl, rfl⟩, - refine ⟨add_comm _ _, _⟩, - simp [hx] } }, + exact ⟨rfl, by simp [hx], add_comm _ _⟩ } }, { simp } end @@ -615,7 +612,7 @@ begin ext ⟨a1, a2⟩, simp only [not_mem_empty, not_and, set.mem_singleton_iff, not_not, mem_add_antidiagonal, set.mem_set_of_eq, iff_false], - rintro h1 h2 rfl, + rintro h2 rfl h1, rw ← add_right_cancel h1 at hx, exact h2 hx, }, transitivity ∑ (ij : Γ × Γ) in {(a,b)}, x.coeff ij.fst * (single b r).coeff ij.snd, @@ -624,7 +621,7 @@ begin simp only [set.mem_singleton_iff, prod.mk.inj_iff, mem_add_antidiagonal, mem_singleton, set.mem_set_of_eq], split, - { rintro ⟨h1, h2, rfl⟩, + { rintro ⟨h2, rfl, h1⟩, refine ⟨add_right_cancel h1, rfl⟩ }, { rintro ⟨rfl, rfl⟩, simp [hx] } }, @@ -678,22 +675,21 @@ begin simp only [mul_coeff, add_coeff, sum_mul, mul_sum, sum_sigma'], refine sum_bij_ne_zero (λ a has ha0, ⟨⟨a.2.1, a.2.2 + a.1.2⟩, ⟨a.2.2, a.1.2⟩⟩) _ _ _ _, { rintros ⟨⟨i,j⟩, ⟨k,l⟩⟩ H1 H2, - simp only [true_and, set.image2_add, eq_self_iff_true, mem_add_antidiagonal, ne.def, + simp only [and_true, set.image2_add, eq_self_iff_true, mem_add_antidiagonal, ne.def, set.image_prod, mem_sigma, set.mem_set_of_eq] at H1 H2 ⊢, - obtain ⟨⟨rfl, ⟨H3, nz⟩⟩, ⟨rfl, nx, ny⟩⟩ := H1, - refine ⟨⟨(add_assoc _ _ _).symm, nx, set.add_mem_add ny nz⟩, ny, nz⟩ }, - { rintros ⟨⟨i1,j1⟩, ⟨k1,l1⟩⟩ ⟨⟨i2,j2⟩, ⟨k2,l2⟩⟩ H1 H2 H3 H4 H5, + obtain ⟨⟨H3, nz, rfl⟩, nx, ny, rfl⟩ := H1, + exact ⟨⟨nx, set.add_mem_add ny nz, (add_assoc _ _ _).symm⟩, ny, nz⟩ }, + { rintros ⟨⟨i1,j1⟩, k1,l1⟩ ⟨⟨i2,j2⟩, k2,l2⟩ H1 H2 H3 H4 H5, simp only [set.image2_add, prod.mk.inj_iff, mem_add_antidiagonal, ne.def, set.image_prod, mem_sigma, set.mem_set_of_eq, heq_iff_eq] at H1 H3 H5, obtain ⟨⟨rfl, H⟩, rfl, rfl⟩ := H5, - simp only [and_true, prod.mk.inj_iff, eq_self_iff_true, heq_iff_eq], - exact add_right_cancel (H1.1.1.trans H3.1.1.symm) }, + simp only [and_true, prod.mk.inj_iff, eq_self_iff_true, heq_iff_eq, ←H1.2.2.2, ←H3.2.2.2] }, { rintros ⟨⟨i,j⟩, ⟨k,l⟩⟩ H1 H2, simp only [exists_prop, set.image2_add, prod.mk.inj_iff, mem_add_antidiagonal, sigma.exists, ne.def, set.image_prod, mem_sigma, set.mem_set_of_eq, heq_iff_eq, prod.exists] at H1 H2 ⊢, - obtain ⟨⟨rfl, nx, H⟩, rfl, ny, nz⟩ := H1, - exact ⟨i + k, l, i, k, ⟨⟨add_assoc _ _ _, set.add_mem_add nx ny, nz⟩, rfl, nx, ny⟩, + obtain ⟨⟨nx, H, rfl⟩, ny, nz, rfl⟩ := H1, + exact ⟨i + k, l, i, k, ⟨⟨set.add_mem_add nx ny, nz, add_assoc _ _ _⟩, nx, ny, rfl⟩, λ con, H2 ((mul_assoc _ _ _).symm.trans con), ⟨rfl, rfl⟩, rfl, rfl⟩ }, { rintros ⟨⟨i,j⟩, ⟨k,l⟩⟩ H1 H2, simp [mul_assoc], } @@ -737,20 +733,9 @@ instance [non_unital_comm_semiring R] : non_unital_comm_semiring (hahn_series Γ { mul_comm := λ x y, begin ext, simp_rw [mul_coeff, mul_comm], - refine sum_bij (λ a ha, ⟨a.2, a.1⟩) _ (λ a ha, by simp) _ _, - { intros a ha, - simp only [mem_add_antidiagonal, ne.def, set.mem_set_of_eq] at ha ⊢, - obtain ⟨h1, h2, h3⟩ := ha, - refine ⟨_, h3, h2⟩, - rw [add_comm, h1], }, - { rintros ⟨a1, a2⟩ ⟨b1, b2⟩ ha hb hab, - rw prod.ext_iff at *, - refine ⟨hab.2, hab.1⟩, }, - { intros a ha, - refine ⟨a.swap, _, by simp⟩, - simp only [prod.fst_swap, mem_add_antidiagonal, prod.snd_swap, - ne.def, set.mem_set_of_eq] at ha ⊢, - exact ⟨(add_comm _ _).trans ha.1, ha.2.2, ha.2.1⟩ } + refine sum_bij (λ a ha, a.swap) (λ a ha, _) (λ a ha, rfl) (λ _ _ _ _, prod.swap_inj.1) + (λ a ha, ⟨a.swap, _, a.swap_swap.symm⟩); + rwa swap_mem_add_antidiagonal, end, .. hahn_series.non_unital_semiring } @@ -837,9 +822,9 @@ begin { rw [h, mul_single_coeff_add], simp }, { rw [single_coeff_of_ne h, mul_coeff, sum_eq_zero], - rintros ⟨y1, y2⟩ hy, - obtain ⟨rfl, hy1, hy2⟩ := mem_add_antidiagonal.1 hy, - rw [eq_of_mem_support_single hy1, eq_of_mem_support_single hy2] at h, + simp_rw mem_add_antidiagonal, + rintro ⟨y, z⟩ ⟨hy, hz, rfl⟩, + rw [eq_of_mem_support_single hy, eq_of_mem_support_single hz] at h, exact (h rfl).elim } end @@ -912,20 +897,19 @@ begin { simp }, apply sum_subset, { rintro ⟨i, j⟩ hij, - simp only [exists_prop, mem_map, prod.mk.inj_iff, - mem_add_antidiagonal, ne.def, function.embedding.coe_prod_map, mem_support, - prod.exists] at hij, - obtain ⟨i, j, ⟨rfl, hx, hy⟩, rfl, rfl⟩ := hij, + simp only [exists_prop, mem_map, prod.mk.inj_iff, mem_add_antidiagonal, + function.embedding.coe_prod_map, mem_support, prod.exists] at hij, + obtain ⟨i, j, ⟨hx, hy, rfl⟩, rfl, rfl⟩ := hij, simp [hx, hy, hf], }, { rintro ⟨_, _⟩ h1 h2, contrapose! h2, obtain ⟨i, hi, rfl⟩ := support_emb_domain_subset (ne_zero_and_ne_zero_of_mul h2).1, obtain ⟨j, hj, rfl⟩ := support_emb_domain_subset (ne_zero_and_ne_zero_of_mul h2).2, - simp only [exists_prop, mem_map, prod.mk.inj_iff, - mem_add_antidiagonal, ne.def, function.embedding.coe_prod_map, mem_support, - prod.exists], - simp only [mem_add_antidiagonal, emb_domain_coeff, ne.def, mem_support, ← hf] at h1, - exact ⟨i, j, ⟨f.injective h1.1, h1.2⟩, rfl⟩, } }, + simp only [exists_prop, mem_map, prod.mk.inj_iff, mem_add_antidiagonal, + function.embedding.coe_prod_map, mem_support, prod.exists], + simp only [mem_add_antidiagonal, emb_domain_coeff, mem_support, ←hf, + order_embedding.eq_iff_eq] at h1, + exact ⟨i, j, h1, rfl⟩ } }, { rw [emb_domain_notin_range hg, eq_comm], contrapose! hg, obtain ⟨_, _, hi, hj, rfl⟩ := support_mul_subset_add_support ((mem_support _ _).2 hg), @@ -1003,7 +987,7 @@ section semiring variables [semiring R] /-- The ring `hahn_series ℕ R` is isomorphic to `power_series R`. -/ -@[simps] def to_power_series : (hahn_series ℕ R) ≃+* power_series R := +@[simps] def to_power_series : hahn_series ℕ R ≃+* power_series R := { to_fun := λ f, power_series.mk f.coeff, inv_fun := λ f, ⟨λ n, power_series.coeff R n f, (nat.lt_wf.is_wf _).is_pwo⟩, left_inv := λ f, by { ext, simp }, @@ -1013,15 +997,12 @@ variables [semiring R] ext n, simp only [power_series.coeff_mul, power_series.coeff_mk, mul_coeff, is_pwo_support], classical, - refine sum_filter_ne_zero.symm.trans - ((sum_congr _ (λ _ _, rfl)).trans sum_filter_ne_zero), + refine sum_filter_ne_zero.symm.trans ((sum_congr _ $ λ _ _, rfl).trans sum_filter_ne_zero), ext m, - simp only [nat.mem_antidiagonal, and.congr_left_iff, mem_add_antidiagonal, ne.def, - and_iff_left_iff_imp, mem_filter, mem_support], - intros h1 h2, - contrapose h1, - rw ← decidable.or_iff_not_and_not at h1, - cases h1; simp [h1] + simp only [nat.mem_antidiagonal, mem_add_antidiagonal, and.congr_left_iff, mem_filter, + mem_support], + rintro h, + rw [and_iff_right (left_ne_zero_of_mul h), and_iff_right (right_ne_zero_of_mul h)], end } lemma coeff_to_power_series {f : hahn_series ℕ R} {n : ℕ} : @@ -1116,12 +1097,10 @@ After importing `algebra.order.pi` the ring `hahn_series (σ → ℕ) R` could b simp_rw [mul_coeff], refine sum_filter_ne_zero.symm.trans ((sum_congr _ (λ _ _, rfl)).trans sum_filter_ne_zero), ext m, - simp only [and.congr_left_iff, mem_add_antidiagonal, ne.def, and_iff_left_iff_imp, mem_filter, - mem_support, finsupp.mem_antidiagonal], - intros h1 h2, - contrapose h1, - rw ← decidable.or_iff_not_and_not at h1, - cases h1; simp [h1], + simp only [and.congr_left_iff, mem_add_antidiagonal, mem_filter, mem_support, + finsupp.mem_antidiagonal], + rintro h, + rw [and_iff_right (left_ne_zero_of_mul h), and_iff_right (right_ne_zero_of_mul h)], end } variables {σ : Type*} [fintype σ] @@ -1399,9 +1378,9 @@ instance : has_smul (hahn_series Γ R) (summable_family Γ R α) := { exact λ ij hij, function.support (λ a, (s a).coeff ij.2) }, { apply s.finite_co_support }, { obtain ⟨i, j, hi, hj, rfl⟩ := support_mul_subset_add_support ha, - simp only [exists_prop, set.mem_Union, mem_add_antidiagonal, - mul_coeff, ne.def, mem_support, is_pwo_support, prod.exists], - refine ⟨i, j, mem_coe.2 (mem_add_antidiagonal.2 ⟨rfl, hi, set.mem_Union.2 ⟨a, hj⟩⟩), hj⟩, } + simp only [exists_prop, set.mem_Union, mem_add_antidiagonal, mul_coeff, mem_support, + is_pwo_support, prod.exists], + exact ⟨i, j, mem_coe.2 (mem_add_antidiagonal.2 ⟨hi, set.mem_Union.2 ⟨a, hj⟩, rfl⟩), hj⟩ } end } } @[simp] @@ -1430,7 +1409,7 @@ begin { refine sum_subset (add_antidiagonal_mono_right (set.subset_Union _ a)) _, rintro ⟨i, j⟩ hU ha, rw mem_add_antidiagonal at *, - rw [not_not.1 (λ con, ha ⟨hU.1, hU.2.1, con⟩), mul_zero] }, + rw [not_not.1 (λ con, ha ⟨hU.1, con, hU.2.2⟩), mul_zero] }, { rintro ⟨i, j⟩ hij, refine (s.finite_co_support j).subset _, simp_rw [function.support_subset_iff', function.mem_support, not_not], @@ -1446,7 +1425,7 @@ begin simp [hx] }, { rintro ⟨i, j⟩ hU ha, rw mem_add_antidiagonal at *, - rw [← hsum_coeff, not_not.1 (λ con, ha ⟨hU.1, hU.2.1, con⟩), mul_zero] } } + rw [← hsum_coeff, not_not.1 (λ con, ha ⟨hU.1, con, hU.2.2⟩), mul_zero] } } end /-- The summation of a `summable_family` as a `linear_map`. -/ @@ -1468,14 +1447,12 @@ def of_finsupp (f : α →₀ (hahn_series Γ R)) : summable_family Γ R α := { to_fun := f, is_pwo_Union_support' := begin - apply (f.support.is_pwo_sup (λ a, (f a).support) (λ a ha, (f a).is_pwo_support)).mono, - intros g hg, - obtain ⟨a, ha⟩ := set.mem_Union.1 hg, + apply (f.support.is_pwo_bUnion.2 $ λ a ha, (f a).is_pwo_support).mono, + refine set.Union_subset_iff.2 (λ a g hg, _), have haf : a ∈ f.support, { rw [finsupp.mem_support_iff, ← support_nonempty_iff], - exact ⟨g, ha⟩ }, - have h : (λ i, (f i).support) a ≤ _ := le_sup haf, - exact h ha, + exact ⟨g, hg⟩ }, + exact set.mem_bUnion haf hg end, finite_co_support' := λ g, begin refine f.support.finite_to_set.subset (λ a ha, _), @@ -1568,19 +1545,18 @@ def powers (x : hahn_series Γ R) (hx : 0 < add_val Γ R x) : intros y ys hy, refine ((((add_antidiagonal x.is_pwo_support hpwo y).finite_to_set.bUnion (λ ij hij, hy ij.snd _ _)).image nat.succ).union (set.finite_singleton 0)).subset _, - { exact (mem_add_antidiagonal.1 (mem_coe.1 hij)).2.2 }, - { obtain ⟨rfl, hi, hj⟩ := mem_add_antidiagonal.1 (mem_coe.1 hij), + { exact (mem_add_antidiagonal.1 (mem_coe.1 hij)).2.1 }, + { obtain ⟨hi, hj, rfl⟩ := mem_add_antidiagonal.1 (mem_coe.1 hij), rw [← zero_add ij.snd, ← add_assoc, add_zero], exact add_lt_add_right (with_top.coe_lt_coe.1 (lt_of_lt_of_le hx (add_val_le_of_coeff_ne_zero hi))) _, }, - { intros n hn, - cases n, + { rintro (_ | n) hn, { exact set.mem_union_right _ (set.mem_singleton 0) }, { obtain ⟨i, j, hi, hj, rfl⟩ := support_mul_subset_add_support hn, refine set.mem_union_left _ ⟨n, set.mem_Union.2 ⟨⟨i, j⟩, set.mem_Union.2 ⟨_, hj⟩⟩, rfl⟩, - simp only [true_and, set.mem_Union, mem_add_antidiagonal, mem_coe, eq_self_iff_true, + simp only [and_true, set.mem_Union, mem_add_antidiagonal, mem_coe, eq_self_iff_true, ne.def, mem_support, set.mem_set_of_eq], - exact ⟨hi, ⟨n, hj⟩⟩ } } + exact ⟨hi, n, hj⟩ } } end } variables {x : hahn_series Γ R} (hx : 0 < add_val Γ R x) From 2ca1b5715d46deb06352c10510de735ad59c1bb2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 25 Aug 2022 00:16:12 +0000 Subject: [PATCH 035/828] chore(algebra/order/{module, smul}): Move to the correct spot (#16131) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Make `algebra.order.module` do what it says on the tin. Namely, move everything that wasn't about `module` to `algebra.order.smul` and generalize accordingly. As a bonus, add a shortcut instance for `ordered_smul 𝕜 (ι → 𝕜)` as this solves #16021. --- src/algebra/order/module.lean | 79 +--------------------- src/algebra/order/smul.lean | 119 +++++++++++++++++++++++----------- 2 files changed, 82 insertions(+), 116 deletions(-) diff --git a/src/algebra/order/module.lean b/src/algebra/order/module.lean index c8daf52710803..191c1b490e5e4 100644 --- a/src/algebra/order/module.lean +++ b/src/algebra/order/module.lean @@ -3,11 +3,7 @@ Copyright (c) 2020 Frédéric Dupuis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis, Yaël Dillies -/ -import algebra.module.pi -import algebra.module.prod -import algebra.order.pi import algebra.order.smul -import data.set.pointwise /-! # Ordered module @@ -27,14 +23,10 @@ open_locale pointwise variables {k M N : Type*} -namespace order_dual - instance [semiring k] [ordered_add_comm_monoid M] [module k M] : module k Mᵒᵈ := { add_smul := λ r s x, order_dual.rec (add_smul _ _) x, zero_smul := λ m, order_dual.rec (zero_smul _) m } -end order_dual - section semiring variables [ordered_semiring k] [ordered_add_comm_group M] [module k M] [ordered_smul k M] {a b : M} {c : k} @@ -190,57 +182,10 @@ variables (M) right_inv := smul_inv_smul₀ hc.ne, map_rel_iff' := λ b₁ b₂, smul_le_smul_iff_of_neg hc } -variables {M} [ordered_add_comm_group N] [module k N] [ordered_smul k N] - --- TODO: solve `prod.has_lt` and `prod.has_le` misalignment issue -instance prod.ordered_smul : ordered_smul k (M × N) := -ordered_smul.mk' $ λ (v u : M × N) (c : k) h hc, - ⟨smul_le_smul_of_nonneg h.1.1 hc.le, smul_le_smul_of_nonneg h.1.2 hc.le⟩ - -instance pi.smul_with_zero'' {ι : Type*} {M : ι → Type*} [Π i, ordered_add_comm_group (M i)] - [Π i, mul_action_with_zero k (M i)] : - smul_with_zero k (Π i : ι, M i) := by apply_instance - -instance pi.ordered_smul {ι : Type*} {M : ι → Type*} [Π i, ordered_add_comm_group (M i)] - [Π i, mul_action_with_zero k (M i)] [∀ i, ordered_smul k (M i)] : - ordered_smul k (Π i : ι, M i) := -begin - refine (ordered_smul.mk' $ λ v u c h hc i, _), - change c • v i ≤ c • u i, - exact smul_le_smul_of_nonneg (h.le i) hc.le, -end - --- Sometimes Lean fails to apply the dependent version to non-dependent functions, --- so we define another instance -instance pi.ordered_smul' {ι : Type*} {M : Type*} [ordered_add_comm_group M] - [mul_action_with_zero k M] [ordered_smul k M] : - ordered_smul k (ι → M) := -pi.ordered_smul - end field /-! ### Upper/lower bounds -/ -section ordered_semiring -variables [ordered_semiring k] [ordered_add_comm_monoid M] [smul_with_zero k M] [ordered_smul k M] - {s : set M} {c : k} - -lemma smul_lower_bounds_subset_lower_bounds_smul (hc : 0 ≤ c) : - c • lower_bounds s ⊆ lower_bounds (c • s) := -(monotone_smul_left hc).image_lower_bounds_subset_lower_bounds_image - -lemma smul_upper_bounds_subset_upper_bounds_smul (hc : 0 ≤ c) : - c • upper_bounds s ⊆ upper_bounds (c • s) := -(monotone_smul_left hc).image_upper_bounds_subset_upper_bounds_image - -lemma bdd_below.smul_of_nonneg (hs : bdd_below s) (hc : 0 ≤ c) : bdd_below (c • s) := -(monotone_smul_left hc).map_bdd_below hs - -lemma bdd_above.smul_of_nonneg (hs : bdd_above s) (hc : 0 ≤ c) : bdd_above (c • s) := -(monotone_smul_left hc).map_bdd_above hs - -end ordered_semiring - section ordered_ring variables [ordered_ring k] [ordered_add_comm_group M] [module k M] [ordered_smul k M] {s : set M} {c : k} @@ -262,27 +207,8 @@ lemma bdd_above.smul_of_nonpos (hc : c ≤ 0) (hs : bdd_above s) : bdd_below (c end ordered_ring section linear_ordered_field -variables [linear_ordered_field k] [ordered_add_comm_group M] - -section mul_action_with_zero -variables [mul_action_with_zero k M] [ordered_smul k M] {s t : set M} {c : k} - -@[simp] lemma lower_bounds_smul_of_pos (hc : 0 < c) : lower_bounds (c • s) = c • lower_bounds s := -(order_iso.smul_left _ hc).lower_bounds_image - -@[simp] lemma upper_bounds_smul_of_pos (hc : 0 < c) : upper_bounds (c • s) = c • upper_bounds s := -(order_iso.smul_left _ hc).upper_bounds_image - -@[simp] lemma bdd_below_smul_iff_of_pos (hc : 0 < c) : bdd_below (c • s) ↔ bdd_below s := -(order_iso.smul_left _ hc).bdd_below_image - -@[simp] lemma bdd_above_smul_iff_of_pos (hc : 0 < c) : bdd_above (c • s) ↔ bdd_above s := -(order_iso.smul_left _ hc).bdd_above_image - -end mul_action_with_zero - -section module -variables [module k M] [ordered_smul k M] {s t : set M} {c : k} +variables [linear_ordered_field k] [ordered_add_comm_group M] [module k M] [ordered_smul k M] + {s : set M} {c : k} @[simp] lemma lower_bounds_smul_of_neg (hc : c < 0) : lower_bounds (c • s) = c • upper_bounds s := (order_iso.smul_left_dual M hc).upper_bounds_image @@ -296,5 +222,4 @@ variables [module k M] [ordered_smul k M] {s t : set M} {c : k} @[simp] lemma bdd_above_smul_iff_of_neg (hc : c < 0) : bdd_above (c • s) ↔ bdd_below s := (order_iso.smul_left_dual M hc).bdd_below_image -end module end linear_ordered_field diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index a2c4e5f920245..21cceefe4e6aa 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -3,10 +3,11 @@ Copyright (c) 2020 Frédéric Dupuis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Frédéric Dupuis -/ - +import algebra.module.pi +import algebra.module.prod import algebra.order.field -import algebra.smul_with_zero -import group_theory.group_action.group +import algebra.order.pi +import data.set.pointwise /-! # Ordered scalar product @@ -35,6 +36,7 @@ In this file we define ordered module, ordered scalar, ordered smul, ordered action, ordered vector space -/ +open_locale pointwise /-- The ordered scalar product property is when an ordered additive commutative monoid @@ -46,9 +48,9 @@ class ordered_smul (R M : Type*) (smul_lt_smul_of_pos : ∀ {a b : M}, ∀ {c : R}, a < b → 0 < c → c • a < c • b) (lt_of_smul_lt_smul_of_pos : ∀ {a b : M}, ∀ {c : R}, c • a < c • b → 0 < c → a < b) -namespace order_dual +variables {ι 𝕜 R M N : Type*} -variables {R M : Type*} +namespace order_dual instance [has_zero R] [add_zero_class M] [h : smul_with_zero R M] : smul_with_zero R Mᵒᵈ := { zero_smul := λ m, order_dual.rec (zero_smul _) m, @@ -79,10 +81,8 @@ instance [ordered_semiring R] [ordered_add_comm_monoid M] [smul_with_zero R M] end order_dual section ordered_smul - -variables {R M : Type*} - [ordered_semiring R] [ordered_add_comm_monoid M] [smul_with_zero R M] [ordered_smul R M] - {a b : M} {c : R} +variables [ordered_semiring R] [ordered_add_comm_monoid M] [smul_with_zero R M] [ordered_smul R M] + {s : set M} {a b : M} {c : R} lemma smul_lt_smul_of_pos : a < b → 0 < c → c • a < c • b := ordered_smul.smul_lt_smul_of_pos @@ -125,46 +125,72 @@ lemma monotone_smul_left (hc : 0 ≤ c) : monotone (has_smul.smul c : M → M) : lemma strict_mono_smul_left (hc : 0 < c) : strict_mono (has_smul.smul c : M → M) := λ a b h, smul_lt_smul_of_pos h hc +lemma smul_lower_bounds_subset_lower_bounds_smul (hc : 0 ≤ c) : + c • lower_bounds s ⊆ lower_bounds (c • s) := +(monotone_smul_left hc).image_lower_bounds_subset_lower_bounds_image + +lemma smul_upper_bounds_subset_upper_bounds_smul (hc : 0 ≤ c) : + c • upper_bounds s ⊆ upper_bounds (c • s) := +(monotone_smul_left hc).image_upper_bounds_subset_upper_bounds_image + +lemma bdd_below.smul_of_nonneg (hs : bdd_below s) (hc : 0 ≤ c) : bdd_below (c • s) := +(monotone_smul_left hc).map_bdd_below hs + +lemma bdd_above.smul_of_nonneg (hs : bdd_above s) (hc : 0 ≤ c) : bdd_above (c • s) := +(monotone_smul_left hc).map_bdd_above hs + end ordered_smul -/-- If `R` is a linear ordered semifield, then it suffices to verify only the first axiom of -`ordered_smul`. Moreover, it suffices to verify that `a < b` and `0 < c` imply -`c • a ≤ c • b`. We have no semifields in `mathlib`, so we use the assumption `∀ c ≠ 0, is_unit c` -instead. -/ -lemma ordered_smul.mk'' {R M : Type*} [linear_ordered_semiring R] [ordered_add_comm_monoid M] - [mul_action_with_zero R M] (hR : ∀ {c : R}, c ≠ 0 → is_unit c) - (hlt : ∀ ⦃a b : M⦄ ⦃c : R⦄, a < b → 0 < c → c • a ≤ c • b) : - ordered_smul R M := +instance linear_ordered_semiring.to_ordered_smul {R : Type*} [linear_ordered_semiring R] : + ordered_smul R R := +{ smul_lt_smul_of_pos := ordered_semiring.mul_lt_mul_of_pos_left, + lt_of_smul_lt_smul_of_pos := λ _ _ _ h hc, lt_of_mul_lt_mul_left h hc.le } + +section linear_ordered_semifield +variables [linear_ordered_semifield 𝕜] + +section ordered_add_comm_monoid +variables [ordered_add_comm_monoid M] [ordered_add_comm_monoid N] [mul_action_with_zero 𝕜 M] + [mul_action_with_zero 𝕜 N] + +/-- To prove that a vector space over a linear ordered field is ordered, it suffices to verify only +the first axiom of `ordered_smul`. -/ +lemma ordered_smul.mk' (h : ∀ ⦃a b : M⦄ ⦃c : 𝕜⦄, a < b → 0 < c → c • a ≤ c • b) : + ordered_smul 𝕜 M := begin - have hlt' : ∀ ⦃a b : M⦄ ⦃c : R⦄, a < b → 0 < c → c • a < c • b, - { refine λ a b c hab hc, (hlt hab hc).lt_of_ne _, - rw [ne.def, (hR hc.ne').smul_left_cancel], + have hlt' : ∀ ⦃a b : M⦄ ⦃c : 𝕜⦄, a < b → 0 < c → c • a < c • b, + { refine λ a b c hab hc, (h hab hc).lt_of_ne _, + rw [ne.def, hc.ne'.is_unit.smul_left_cancel], exact hab.ne }, refine { smul_lt_smul_of_pos := hlt', .. }, - intros a b c h hc, - rcases (hR hc.ne') with ⟨c, rfl⟩, + intros a b c hab hc, + obtain ⟨c, rfl⟩ := hc.ne'.is_unit, rw [← inv_smul_smul c a, ← inv_smul_smul c b], - refine hlt' h (pos_of_mul_pos_right _ hc.le), + refine hlt' hab (pos_of_mul_pos_right _ hc.le), simp only [c.mul_inv, zero_lt_one] end -/-- If `R` is a linear ordered field, then it suffices to verify only the first axiom of -`ordered_smul`. -/ -lemma ordered_smul.mk' {k M : Type*} [linear_ordered_field k] [ordered_add_comm_monoid M] - [mul_action_with_zero k M] (hlt : ∀ ⦃a b : M⦄ ⦃c : k⦄, a < b → 0 < c → c • a ≤ c • b) : - ordered_smul k M := -ordered_smul.mk'' (λ c hc, is_unit.mk0 _ hc) hlt +instance [ordered_smul 𝕜 M] [ordered_smul 𝕜 N] : ordered_smul 𝕜 (M × N) := +ordered_smul.mk' $ λ a b c h hc, + ⟨smul_le_smul_of_nonneg h.1.1 hc.le, smul_le_smul_of_nonneg h.1.2 hc.le⟩ -instance linear_ordered_semiring.to_ordered_smul {R : Type*} [linear_ordered_semiring R] : - ordered_smul R R := -{ smul_lt_smul_of_pos := ordered_semiring.mul_lt_mul_of_pos_left, - lt_of_smul_lt_smul_of_pos := λ _ _ _ h hc, lt_of_mul_lt_mul_left h hc.le } +instance pi.ordered_smul {M : ι → Type*} [Π i, ordered_add_comm_monoid (M i)] + [Π i, mul_action_with_zero 𝕜 (M i)] [∀ i, ordered_smul 𝕜 (M i)] : + ordered_smul 𝕜 (Π i, M i) := +ordered_smul.mk' $ λ v u c h hc i, smul_le_smul_of_nonneg (h.le i) hc.le + +/- Sometimes Lean fails to apply the dependent version to non-dependent functions, so we define +another instance. -/ +instance pi.ordered_smul' [ordered_smul 𝕜 M] : ordered_smul 𝕜 (ι → M) := pi.ordered_smul -section field +/- Sometimes Lean fails to unify the module with the scalars, so we define another instance. -/ +instance pi.ordered_smul'' : ordered_smul 𝕜 (ι → 𝕜) := @pi.ordered_smul' ι 𝕜 𝕜 _ _ _ _ -variables {k M : Type*} [linear_ordered_field k] - [ordered_add_comm_group M] [mul_action_with_zero k M] [ordered_smul k M] - {a b : M} {c : k} +end ordered_add_comm_monoid + +section ordered_add_comm_group +variables [ordered_add_comm_group M] [mul_action_with_zero 𝕜 M] [ordered_smul 𝕜 M] {s : set M} + {a b : M} {c : 𝕜} lemma smul_le_smul_iff_of_pos (hc : 0 < c) : c • a ≤ c • b ↔ a ≤ b := ⟨λ h, inv_smul_smul₀ hc.ne' a ▸ inv_smul_smul₀ hc.ne' b ▸ @@ -190,11 +216,26 @@ calc a ≤ c • b ↔ c • c⁻¹ • a ≤ c • b : by rw [smul_inv_smul₀ variables (M) /-- Left scalar multiplication as an order isomorphism. -/ -@[simps] def order_iso.smul_left {c : k} (hc : 0 < c) : M ≃o M := +@[simps] def order_iso.smul_left (hc : 0 < c) : M ≃o M := { to_fun := λ b, c • b, inv_fun := λ b, c⁻¹ • b, left_inv := inv_smul_smul₀ hc.ne', right_inv := smul_inv_smul₀ hc.ne', map_rel_iff' := λ b₁ b₂, smul_le_smul_iff_of_pos hc } -end field +variables {M} + +@[simp] lemma lower_bounds_smul_of_pos (hc : 0 < c) : lower_bounds (c • s) = c • lower_bounds s := +(order_iso.smul_left _ hc).lower_bounds_image + +@[simp] lemma upper_bounds_smul_of_pos (hc : 0 < c) : upper_bounds (c • s) = c • upper_bounds s := +(order_iso.smul_left _ hc).upper_bounds_image + +@[simp] lemma bdd_below_smul_iff_of_pos (hc : 0 < c) : bdd_below (c • s) ↔ bdd_below s := +(order_iso.smul_left _ hc).bdd_below_image + +@[simp] lemma bdd_above_smul_iff_of_pos (hc : 0 < c) : bdd_above (c • s) ↔ bdd_above s := +(order_iso.smul_left _ hc).bdd_above_image + +end ordered_add_comm_group +end linear_ordered_semifield From bd12701ee20fc32d6421a2b3ac68a7030ebc24dd Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 25 Aug 2022 00:16:14 +0000 Subject: [PATCH 036/828] =?UTF-8?q?feat(algebra/group=5Fwith=5Fzero):=20ad?= =?UTF-8?q?d=20`eq=5Fon=5Finv=E2=82=80`=20(#16222)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * move some lemmas up in the import chain; * use `namespace is_unit`; * add `is_unit.eq_on_inv` and `eq_on_inv₀`; * rename `monoid_hom.eq_on_inv` to `eq_on_inv`, make `f` and `g` explicit. --- src/algebra/group/units.lean | 63 +++++++++++++++----------- src/algebra/group_with_zero/basic.lean | 15 +++++- src/algebra/hom/equiv.lean | 3 -- src/algebra/hom/group.lean | 17 +++++-- src/algebra/hom/units.lean | 18 ++------ src/data/int/cast.lean | 2 +- 6 files changed, 67 insertions(+), 51 deletions(-) diff --git a/src/algebra/group/units.lean b/src/algebra/group/units.lean index 3bbc12ec81102..23f4851d08a2b 100644 --- a/src/algebra/group/units.lean +++ b/src/algebra/group/units.lean @@ -221,6 +221,10 @@ by rw [←inv_mul_eq_one, inv_inv] @[to_additive] lemma inv_unique {u₁ u₂ : αˣ} (h : (↑u₁ : α) = ↑u₂) : (↑u₁⁻¹ : α) = ↑u₂⁻¹ := units.inv_eq_of_mul_eq_one_right $ by rw [h, u₂.mul_inv] +@[simp, to_additive] +lemma coe_inv {M : Type*} [division_monoid M] (u : units M) : ↑u⁻¹ = (u⁻¹ : M) := +eq.symm $ inv_eq_of_mul_eq_one_right u.mul_inv + end units /-- For `a, b` in a `comm_monoid` such that `a * b = 1`, makes a unit out of `a`. -/ @@ -420,66 +424,71 @@ is_unit_iff_exists_inv.2 ⟨y * z, by rwa ← mul_assoc⟩ (hu : is_unit (x * y)) : is_unit y := @is_unit_of_mul_is_unit_left _ _ y x $ by rwa mul_comm +namespace is_unit + @[simp, to_additive] -lemma is_unit.mul_iff [comm_monoid M] {x y : M} : is_unit (x * y) ↔ is_unit x ∧ is_unit y := +lemma mul_iff [comm_monoid M] {x y : M} : is_unit (x * y) ↔ is_unit x ∧ is_unit y := ⟨λ h, ⟨is_unit_of_mul_is_unit_left h, is_unit_of_mul_is_unit_right h⟩, λ h, is_unit.mul h.1 h.2⟩ +section monoid + +variables [monoid M] {a b c : M} + /-- The element of the group of units, corresponding to an element of a monoid which is a unit. When `α` is a `division_monoid`, use `is_unit.unit'` instead. -/ @[to_additive "The element of the additive group of additive units, corresponding to an element of an additive monoid which is an additive unit. When `α` is a `subtraction_monoid`, use `is_add_unit.add_unit'` instead."] -noncomputable def is_unit.unit [monoid M] {a : M} (h : is_unit a) : Mˣ := +protected noncomputable def unit (h : is_unit a) : Mˣ := (classical.some h).copy a (classical.some_spec h).symm _ rfl @[simp, to_additive] -lemma is_unit.unit_of_coe_units [monoid M] {a : Mˣ} (h : is_unit (a : M)) : h.unit = a := +lemma unit_of_coe_units {a : Mˣ} (h : is_unit (a : M)) : h.unit = a := units.ext $ rfl -@[simp, to_additive] -lemma is_unit.unit_spec [monoid M] {a : M} (h : is_unit a) : ↑h.unit = a := -rfl +@[simp, to_additive] lemma unit_spec (h : is_unit a) : ↑h.unit = a := rfl @[simp, to_additive] -lemma is_unit.coe_inv_mul [monoid M] {a : M} (h : is_unit a) : - ↑(h.unit)⁻¹ * a = 1 := -units.mul_inv _ +lemma coe_inv_mul (h : is_unit a) : ↑(h.unit)⁻¹ * a = 1 := units.mul_inv _ -@[simp, to_additive] -lemma is_unit.mul_coe_inv [monoid M] {a : M} (h : is_unit a) : - a * ↑(h.unit)⁻¹ = 1 := -begin - convert units.mul_inv _, - simp [h.unit_spec] -end +@[simp, to_additive] lemma mul_coe_inv (h : is_unit a) : a * ↑(h.unit)⁻¹ = 1 := +by convert h.unit.mul_inv /-- `is_unit x` is decidable if we can decide if `x` comes from `Mˣ`. -/ -instance [monoid M] (x : M) [h : decidable (∃ u : Mˣ, ↑u = x)] : decidable (is_unit x) := h +instance (x : M) [h : decidable (∃ u : Mˣ, ↑u = x)] : decidable (is_unit x) := h -section monoid -variables [monoid M] {a b c : M} - -@[to_additive] lemma is_unit.mul_left_inj (h : is_unit a) : b * a = c * a ↔ b = c := +@[to_additive] lemma mul_left_inj (h : is_unit a) : b * a = c * a ↔ b = c := let ⟨u, hu⟩ := h in hu ▸ u.mul_left_inj -@[to_additive] lemma is_unit.mul_right_inj (h : is_unit a) : a * b = a * c ↔ b = c := +@[to_additive] lemma mul_right_inj (h : is_unit a) : a * b = a * c ↔ b = c := let ⟨u, hu⟩ := h in hu ▸ u.mul_right_inj -@[to_additive] protected lemma is_unit.mul_left_cancel (h : is_unit a) : a * b = a * c → b = c := +@[to_additive] protected lemma mul_left_cancel (h : is_unit a) : a * b = a * c → b = c := h.mul_right_inj.1 -@[to_additive] protected lemma is_unit.mul_right_cancel (h : is_unit b) : a * b = c * b → a = c := +@[to_additive] protected lemma mul_right_cancel (h : is_unit b) : a * b = c * b → a = c := h.mul_left_inj.1 -@[to_additive] protected lemma is_unit.mul_right_injective (h : is_unit a) : injective ((*) a) := +@[to_additive] protected lemma mul_right_injective (h : is_unit a) : injective ((*) a) := λ _ _, h.mul_left_cancel -@[to_additive] protected lemma is_unit.mul_left_injective (h : is_unit b) : injective (* b) := +@[to_additive] protected lemma mul_left_injective (h : is_unit b) : injective (* b) := λ _ _, h.mul_right_cancel end monoid -end is_unit + +variables [division_monoid M] {a : M} + +@[simp, to_additive] protected lemma inv_mul_cancel : is_unit a → a⁻¹ * a = 1 := +by { rintro ⟨u, rfl⟩, rw [← units.coe_inv, units.inv_mul] } + +@[simp, to_additive] protected lemma mul_inv_cancel : is_unit a → a * a⁻¹ = 1 := +by { rintro ⟨u, rfl⟩, rw [← units.coe_inv, units.mul_inv] } + +end is_unit -- namespace + +end is_unit -- section section noncomputable_defs diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index 05c2f13ede01d..fbaa1bb4b5609 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -39,7 +39,7 @@ set_option old_structure_cmd true open_locale classical open function -variables {α M₀ G₀ M₀' G₀' F : Type*} +variables {α M₀ G₀ M₀' G₀' F F' : Type*} section @@ -976,7 +976,8 @@ end commute section monoid_with_zero variables [group_with_zero G₀] [monoid_with_zero M₀] [nontrivial M₀] - [monoid_with_zero_hom_class F G₀ M₀] (f : F) {a : G₀} + [monoid_with_zero M₀'] [monoid_with_zero_hom_class F G₀ M₀] + [monoid_with_zero_hom_class F' G₀ M₀'] (f : F) {a : G₀} include M₀ lemma map_ne_zero : f a ≠ 0 ↔ a ≠ 0 := @@ -984,6 +985,16 @@ lemma map_ne_zero : f a ≠ 0 ↔ a ≠ 0 := @[simp] lemma map_eq_zero : f a = 0 ↔ a = 0 := not_iff_not.1 (map_ne_zero f) +omit M₀ +include M₀' + +lemma eq_on_inv₀ (f g : F') (h : f a = g a) : f a⁻¹ = g a⁻¹ := +begin + rcases eq_or_ne a 0 with rfl|ha, + { rw [inv_zero, map_zero, map_zero] }, + { exact (is_unit.mk0 a ha).eq_on_inv f g h } +end + end monoid_with_zero section group_with_zero diff --git a/src/algebra/hom/equiv.lean b/src/algebra/hom/equiv.lean index 4fcb7c96a05d4..6c97f007cdda4 100644 --- a/src/algebra/hom/equiv.lean +++ b/src/algebra/hom/equiv.lean @@ -503,9 +503,6 @@ def to_units [group G] : G ≃* Gˣ := @[simp, to_additive] lemma coe_to_units [group G] (g : G) : (to_units g : G) = g := rfl -@[to_additive] -protected lemma group.is_unit {G} [group G] (x : G) : is_unit x := (to_units x).is_unit - namespace units variables [monoid M] [monoid N] [monoid P] diff --git a/src/algebra/hom/group.lean b/src/algebra/hom/group.lean index c9a8ad1a93052..c67dbb5c7e05c 100644 --- a/src/algebra/hom/group.lean +++ b/src/algebra/hom/group.lean @@ -356,6 +356,9 @@ attribute [to_additive_reorder 8, to_additive] map_zpow end mul_one +@[to_additive] lemma group.is_unit [group G] (g : G) : is_unit g := +⟨⟨g, g⁻¹, mul_inv_self g, inv_mul_self g⟩, rfl⟩ + section mul_zero_one variables [mul_zero_one_class M] [mul_zero_one_class N] @@ -1061,13 +1064,21 @@ by { ext, simp only [map_one, coe_comp, function.comp_app, one_apply] } g.comp (f₁ * f₂) = g.comp f₁ * g.comp f₂ := by { ext, simp only [mul_apply, function.comp_app, map_mul, coe_comp] } +/-- If two homomorphisms from a division monoid to a monoid are equal at a unit `x`, then they are +equal at `x⁻¹`. -/ +@[to_additive "If two homomorphisms from a subtraction monoid to an additive monoid are equal at an +additive unit `x`, then they are equal at `-x`."] +lemma _root_.is_unit.eq_on_inv {G N} [division_monoid G] [monoid N] [monoid_hom_class F G N] {x : G} + (hx : is_unit x) (f g : F) (h : f x = g x) : f x⁻¹ = g x⁻¹ := +left_inv_eq_right_inv (map_mul_eq_one f hx.inv_mul_cancel) $ + h.symm ▸ map_mul_eq_one g $ hx.mul_inv_cancel + /-- If two homomorphism from a group to a monoid are equal at `x`, then they are equal at `x⁻¹`. -/ @[to_additive "If two homomorphism from an additive group to an additive monoid are equal at `x`, then they are equal at `-x`." ] -lemma eq_on_inv {G} [group G] [monoid M] [monoid_hom_class F G M] {f g : F} {x : G} +lemma _root_.eq_on_inv {G} [group G] [monoid M] [monoid_hom_class F G M] (f g : F) {x : G} (h : f x = g x) : f x⁻¹ = g x⁻¹ := -left_inv_eq_right_inv (map_mul_eq_one f $ inv_mul_self x) $ - h.symm ▸ map_mul_eq_one g $ mul_inv_self x +(group.is_unit x).eq_on_inv f g h /-- Group homomorphisms preserve inverse. -/ @[to_additive "Additive group homomorphisms preserve negation."] diff --git a/src/algebra/hom/units.lean b/src/algebra/hom/units.lean index 95a912adabd39..de8d130e7ea24 100644 --- a/src/algebra/hom/units.lean +++ b/src/algebra/hom/units.lean @@ -67,9 +67,6 @@ lemma coe_pow (u : Mˣ) (n : ℕ) : ((u ^ n : Mˣ) : M) = u ^ n := section division_monoid variables [division_monoid α] -@[simp, norm_cast, to_additive] lemma coe_inv : ∀ u : αˣ, ↑u⁻¹ = (u⁻¹ : α) := -(units.coe_hom α).map_inv - @[simp, norm_cast, to_additive] lemma coe_div : ∀ u₁ u₂ : αˣ, ↑(u₁ / u₂) = (u₁ / u₂ : α) := (units.coe_hom α).map_div @@ -118,12 +115,9 @@ and `f.to_hom_units` is the corresponding monoid homomorphism from `G` to `Mˣ`. then its image lies in the `add_units` of `M`, and `f.to_hom_units` is the corresponding homomorphism from `G` to `add_units M`."] def to_hom_units {G M : Type*} [group G] [monoid M] (f : G →* M) : G →* Mˣ := -{ to_fun := λ g, - ⟨f g, f (g⁻¹), - by rw [← f.map_mul, mul_inv_self, f.map_one], - by rw [← f.map_mul, inv_mul_self, f.map_one]⟩, - map_one' := units.ext (f.map_one), - map_mul' := λ _ _, units.ext (f.map_mul _ _) } +units.lift_right f + (λ g, ⟨f g, f g⁻¹, map_mul_eq_one f (mul_inv_self _), map_mul_eq_one f (inv_mul_self _)⟩) + (λ g, rfl) @[simp] lemma coe_to_hom_units {G M : Type*} [group G] [monoid M] (f : G →* M) (g : G): (f.to_hom_units g : M) = f g := rfl @@ -162,12 +156,6 @@ end monoid section division_monoid variables [division_monoid α] {a b c : α} -@[simp, to_additive] protected lemma inv_mul_cancel : is_unit a → a⁻¹ * a = 1 := -by { rintro ⟨u, rfl⟩, rw [←units.coe_inv, units.inv_mul] } - -@[simp, to_additive] protected lemma mul_inv_cancel : is_unit a → a * a⁻¹ = 1 := -by { rintro ⟨u, rfl⟩, rw [←units.coe_inv, units.mul_inv] } - /-- The element of the group of units, corresponding to an element of a monoid which is a unit. As opposed to `is_unit.unit`, the inverse is computable and comes from the inversion on `α`. This is useful to transfer properties of inversion in `units α` to `α`. See also `to_units`. -/ diff --git a/src/data/int/cast.lean b/src/data/int/cast.lean index b9edd2c05d752..412152f5ae7a5 100644 --- a/src/data/int/cast.lean +++ b/src/data/int/cast.lean @@ -177,7 +177,7 @@ if `f 1 = g 1`. -/ @[ext] theorem ext_int [add_monoid A] {f g : ℤ →+ A} (h1 : f 1 = g 1) : f = g := have f.comp (int.of_nat_hom : ℕ →+ ℤ) = g.comp (int.of_nat_hom : ℕ →+ ℤ) := ext_nat' _ _ h1, have ∀ n : ℕ, f n = g n := ext_iff.1 this, -ext $ λ n, int.cases_on n this $ λ n, eq_on_neg (this $ n + 1) +ext $ λ n, int.cases_on n this $ λ n, eq_on_neg _ _ (this $ n + 1) variables [add_group_with_one A] From 42d640723588f3e71bce7a7857e407eb60ef9574 Mon Sep 17 00:00:00 2001 From: Apurva Nakade Date: Thu, 25 Aug 2022 03:05:46 +0000 Subject: [PATCH 037/828] feat(analysis/convex/cone): add `inner_dual_cone_of_inner_dual_cone_eq_self` for nonempty, closed, convex cones (#15637) We add the following results about convex cones: - instance `has_zero` - `inner_dual_cone_zero` - `inner_dual_cone_univ` - `pointed_of_nonempty_of_is_closed` - `hyperplane_separation_of_nonempty_of_is_closed_of_nmem` - `inner_dual_cone_of_inner_dual_cone_eq_self` References: - https://ti.inf.ethz.ch/ew/lehre/ApproxSDP09/notes/conelp.pdf - Stephen P. Boyd and Lieven Vandenberghe. Convex Optimization. Cambridge University Press. ISBN 978-0-521-83378-3. available at https://web.stanford.edu/~boyd/cvxbook/bv_cvxbook.pdf --- docs/references.bib | 15 ++++ src/analysis/convex/cone.lean | 124 +++++++++++++++++++++++++++++++++- 2 files changed, 137 insertions(+), 2 deletions(-) diff --git a/docs/references.bib b/docs/references.bib index 668baca390ad7..a425febd4dfc2 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -322,6 +322,15 @@ @Book{ bourbaki1987 url = {https://doi.org/10.1007/978-3-642-61715-7} } +@book{ boydVandenberghe2004, + author = {Stephen P. Boyd and Lieven Vandenberghe}, + title = {Convex Optimization}, + publisher = {Cambridge University Press}, + year = {2004}, + isbn = {978-0-521-83378-3}, + url = {https://web.stanford.edu/~boyd/cvxbook/bv_cvxbook.pdf} +} + @Book{ cabreragarciarodriguezpalacios2014, author = {Miguel {Cabrera Garc\'{\i}a} and \'Angel {Rodr\'{\i}guez Palacios}}, @@ -1885,6 +1894,12 @@ @book{ weidmann_linear pages = {xiii+402}, } +@misc{ welzl_garter, + author = {Emo Welzl and Bernd G\"{a}rtner}, + title = {Cone Programming}, + url = {https://ti.inf.ethz.ch/ew/lehre/ApproxSDP09/notes/conelp.pdf} +} + @TechReport{ zaanen1966, author = {Zaanen, A. C.}, title = {Lectures on "Riesz Spaces"}, diff --git a/src/analysis/convex/cone.lean b/src/analysis/convex/cone.lean index 174937d42e662..f8b1c2888089b 100644 --- a/src/analysis/convex/cone.lean +++ b/src/analysis/convex/cone.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Frédéric Dupuis -/ import analysis.convex.hull -import analysis.inner_product_space.basic +import analysis.inner_product_space.projection /-! # Convex cones @@ -16,7 +16,7 @@ images (`convex_cone.map`) and preimages (`convex_cone.comap`) under linear maps We define pointed, blunt, flat and salient cones, and prove the correspondence between convex cones and ordered modules. -We also define `convex.to_cone` to be the minimal cone that includes a given convex set. +We define `convex.to_cone` to be the minimal cone that includes a given convex set. We define `set.inner_dual_cone` to be the cone consisting of all points `y` such that for all points `x` in a given set `0 ≤ ⟪ x, y ⟫`. @@ -36,6 +36,18 @@ We prove two extension theorems: then `f` can be extended to the whole space to a linear map `g` such that `g x ≤ N x` for all `x` +We prove the following theorems: +* `convex_cone.hyperplane_separation_of_nonempty_of_is_closed_of_nmem`: + This variant of the + [hyperplane separation theorem](https://en.wikipedia.org/wiki/Hyperplane_separation_theorem) + states that given a nonempty, closed, convex cone `K` in a complete, real inner product space `H` + and a point `b` disjoint from it, there is a vector `y` which separates `b` from `K` in the sense + that for all points `x` in `K`, `0 ≤ ⟪x, y⟫_ℝ` and `⟪y, b⟫_ℝ < 0`. This is also a geometric + interpretation of the + [Farkas lemma](https://en.wikipedia.org/wiki/Farkas%27_lemma#Geometric_interpretation). +* `convex_cone.inner_dual_cone_of_inner_dual_cone_eq_self`: + The `inner_dual_cone` of the `inner_dual_cone` of a nonempty, closed, convex cone is itself. + ## Implementation notes While `convex 𝕜` is a predicate on sets, `convex_cone 𝕜 E` is a bundled convex cone. @@ -43,6 +55,8 @@ While `convex 𝕜` is a predicate on sets, `convex_cone 𝕜 E` is a bundled co ## References * https://en.wikipedia.org/wiki/Convex_cone +* [Stephen P. Boyd and Lieven Vandenberghe, *Convex Optimization*][boydVandenberghe2004] +* [Emo Welzl and Bernd Gärtner, *Cone Programming*][welzl_garter] -/ @@ -341,6 +355,18 @@ def to_ordered_add_comm_group (h₁ : S.pointed) (h₂ : S.salient) : ..show add_comm_group E, by apply_instance } end add_comm_group + +section module +variables [add_comm_monoid E] [module 𝕜 E] + +instance : has_zero (convex_cone 𝕜 E) := ⟨⟨0, λ _ _, by simp, λ _, by simp⟩⟩ + +@[simp] lemma mem_zero (x : E) : x ∈ (0 : convex_cone 𝕜 E) ↔ x = 0 := iff.rfl +@[simp] lemma coe_zero : ((0 : convex_cone 𝕜 E) : set E) = 0 := rfl + +lemma pointed_zero : (0 : convex_cone 𝕜 E).pointed := by rw [pointed, mem_zero] + +end module end ordered_semiring /-! ### Positive cone of an ordered module -/ @@ -647,6 +673,19 @@ def set.inner_dual_cone (s : set H) : convex_cone ℝ H := @[simp] lemma inner_dual_cone_empty : (∅ : set H).inner_dual_cone = ⊤ := eq_top_iff.mpr $ λ x hy y, false.elim +/-- Dual cone of the convex cone {0} is the total space. -/ +@[simp] lemma inner_dual_cone_zero : (0 : set H).inner_dual_cone = ⊤ := +eq_top_iff.mpr $ λ x hy y (hy : y = 0), hy.symm ▸ inner_zero_left.ge + +/-- Dual cone of the total space is the convex cone {0}. -/ +@[simp] lemma inner_dual_cone_univ : (univ : set H).inner_dual_cone = 0 := +begin + suffices : ∀ x : H, x ∈ (univ : set H).inner_dual_cone → x = 0, + { apply set_like.coe_injective, + exact eq_singleton_iff_unique_mem.mpr ⟨λ x hx, inner_zero_right.ge, this⟩ }, + exact λ x hx, by simpa [←real_inner_self_nonpos] using hx (-x) (mem_univ _), +end + lemma inner_dual_cone_le_inner_dual_cone (h : t ⊆ s) : s.inner_dual_cone ≤ t.inner_dual_cone := λ y hy x hx, hy x (h hx) @@ -706,4 +745,85 @@ begin exact is_closed_Ici.preimage (by continuity), end +lemma convex_cone.pointed_of_nonempty_of_is_closed (K : convex_cone ℝ H) + (ne : (K : set H).nonempty) (hc : is_closed (K : set H)) : K.pointed := +begin + obtain ⟨x, hx⟩ := ne, + let f : ℝ → H := (• x), + + -- f (0, ∞) is a subset of K + have fI : f '' set.Ioi 0 ⊆ (K : set H), + { rintro _ ⟨_, h, rfl⟩, + exact K.smul_mem (set.mem_Ioi.1 h) hx }, + + -- closure of f (0, ∞) is a subset of K + have clf : closure (f '' set.Ioi 0) ⊆ (K : set H) := hc.closure_subset_iff.2 fI, + + -- f is continuous at 0 from the right + have fc : continuous_within_at f (set.Ioi (0 : ℝ)) 0 := + (continuous_id.smul continuous_const).continuous_within_at, + + -- 0 belongs to the closure of the f (0, ∞) + have mem₀ := fc.mem_closure_image (by rw [closure_Ioi (0 : ℝ), mem_Ici]), + + -- as 0 ∈ closure f (0, ∞) and closure f (0, ∞) ⊆ K, 0 ∈ K. + have f₀ : f 0 = 0 := zero_smul ℝ x, + simpa only [f₀, convex_cone.pointed, ← set_like.mem_coe] using mem_of_subset_of_mem clf mem₀, +end + +section complete_space +variables [complete_space H] + +/-- This is a stronger version of the Hahn-Banach separation theorem for closed convex cones. This +is also the geometric interpretation of Farkas' lemma. -/ +theorem convex_cone.hyperplane_separation_of_nonempty_of_is_closed_of_nmem (K : convex_cone ℝ H) + (ne : (K : set H).nonempty) (hc : is_closed (K : set H)) {b : H} (disj : b ∉ K) : + ∃ (y : H), (∀ x : H, x ∈ K → 0 ≤ ⟪x, y⟫_ℝ) ∧ ⟪y, b⟫_ℝ < 0 := +begin + -- let `z` be the point in `K` closest to `b` + obtain ⟨z, hzK, infi⟩ := exists_norm_eq_infi_of_complete_convex ne hc.is_complete K.convex b, + + -- for any `w` in `K`, we have `⟪b - z, w - z⟫_ℝ ≤ 0` + have hinner := (norm_eq_infi_iff_real_inner_le_zero K.convex hzK).1 infi, + + -- set `y := z - b` + use z - b, + + split, + { -- the rest of the proof is a straightforward calculation + rintros x hxK, + specialize hinner _ (K.add_mem hxK hzK), + rwa [add_sub_cancel, real_inner_comm, ← neg_nonneg, neg_eq_neg_one_mul, + ← real_inner_smul_right, neg_smul, one_smul, neg_sub] at hinner }, + { -- as `K` is closed and non-empty, it is pointed + have hinner₀ := hinner 0 (K.pointed_of_nonempty_of_is_closed ne hc), + + -- the rest of the proof is a straightforward calculation + rw [zero_sub, inner_neg_right, right.neg_nonpos_iff] at hinner₀, + have hbz : b - z ≠ 0 := by { rw sub_ne_zero, contrapose! hzK, rwa ← hzK }, + rw [← neg_zero, lt_neg, ← neg_one_mul, ← real_inner_smul_left, smul_sub, neg_smul, one_smul, + neg_smul, neg_sub_neg, one_smul], + calc 0 < ⟪b - z, b - z⟫_ℝ : lt_of_not_le ((iff.not real_inner_self_nonpos).2 hbz) + ... = ⟪b - z, b - z⟫_ℝ + 0 : (add_zero _).symm + ... ≤ ⟪b - z, b - z⟫_ℝ + ⟪b - z, z⟫_ℝ : add_le_add rfl.ge hinner₀ + ... = ⟪b - z, b - z + z⟫_ℝ : inner_add_right.symm + ... = ⟪b - z, b⟫_ℝ : by rw sub_add_cancel }, +end + +/-- The inner dual of inner dual of a non-empty, closed convex cone is itself. -/ +theorem convex_cone.inner_dual_cone_of_inner_dual_cone_eq_self (K : convex_cone ℝ H) + (ne : (K : set H).nonempty) (hc : is_closed (K : set H)) : + ((K : set H).inner_dual_cone : set H).inner_dual_cone = K := +begin + ext x, + split, + { rw [mem_inner_dual_cone, ← set_like.mem_coe], + contrapose!, + exact K.hyperplane_separation_of_nonempty_of_is_closed_of_nmem ne hc }, + { rintro hxK y h, + specialize h x hxK, + rwa real_inner_comm }, +end + +end complete_space end dual From fec8ffec7fd192963d83ae6bcc5265318e4f5748 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 25 Aug 2022 03:05:47 +0000 Subject: [PATCH 038/828] =?UTF-8?q?chore(analysis/normed=5Fspace/star/*):?= =?UTF-8?q?=20migrate=20use=20of=20`a=20=E2=88=88=20self=5Fadjoint=20A`=20?= =?UTF-8?q?to=20`is=5Fself=5Fadjoint=20a`=20(#16212)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit After #15326, this PR migrates existing uses of `a ∈ self_adjoint A` to `is_self_adjoint a` so as to standardize. We also move several results into the `is_self_adjoint` namespace in order to take advantage of dot notation. --- src/algebra/star/self_adjoint.lean | 4 +- src/analysis/normed_space/star/basic.lean | 6 +-- .../normed_space/star/exponential.lean | 4 +- src/analysis/normed_space/star/spectrum.lean | 37 +++++++++---------- 4 files changed, 25 insertions(+), 26 deletions(-) diff --git a/src/algebra/star/self_adjoint.lean b/src/algebra/star/self_adjoint.lean index 75a1b2a0e43ba..d33bddca15b18 100644 --- a/src/algebra/star/self_adjoint.lean +++ b/src/algebra/star/self_adjoint.lean @@ -259,9 +259,9 @@ instance : has_pow (self_adjoint R) ℤ := @[simp, norm_cast] lemma coe_zpow (x : self_adjoint R) (z : ℤ) : ↑(x ^ z) = (x : R) ^ z := rfl -lemma rat_cast_mem : ∀ (x : ℚ), (x : R) ∈ self_adjoint R +lemma rat_cast_mem : ∀ (x : ℚ), is_self_adjoint (x : R) | ⟨a, b, h1, h2⟩ := - by rw [mem_iff, rat.cast_mk', star_mul', star_inv', star_nat_cast, star_int_cast] + by rw [is_self_adjoint, rat.cast_mk', star_mul', star_inv', star_nat_cast, star_int_cast] instance : has_rat_cast (self_adjoint R) := ⟨λ n, ⟨n, rat_cast_mem n⟩⟩ diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index 0ee496043638a..1ce247f5e8725 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -162,8 +162,8 @@ norm_mul_coe_unitary A ⟨U, hU⟩ end unital end cstar_ring -lemma nnnorm_pow_two_pow_of_self_adjoint [normed_ring E] [star_ring E] [cstar_ring E] - {x : E} (hx : x ∈ self_adjoint E) (n : ℕ) : ∥x ^ 2 ^ n∥₊ = ∥x∥₊ ^ (2 ^ n) := +lemma is_self_adjoint.nnnorm_pow_two_pow [normed_ring E] [star_ring E] + [cstar_ring E] {x : E} (hx : is_self_adjoint x) (n : ℕ) : ∥x ^ 2 ^ n∥₊ = ∥x∥₊ ^ (2 ^ n) := begin induction n with k hk, { simp only [pow_zero, pow_one] }, @@ -174,7 +174,7 @@ end lemma self_adjoint.nnnorm_pow_two_pow [normed_ring E] [star_ring E] [cstar_ring E] (x : self_adjoint E) (n : ℕ) : ∥x ^ 2 ^ n∥₊ = ∥x∥₊ ^ (2 ^ n) := -nnnorm_pow_two_pow_of_self_adjoint x.property _ +x.prop.nnnorm_pow_two_pow _ section starₗᵢ diff --git a/src/analysis/normed_space/star/exponential.lean b/src/analysis/normed_space/star/exponential.lean index f7a147543e781..1a23c9643dee9 100644 --- a/src/analysis/normed_space/star/exponential.lean +++ b/src/analysis/normed_space/star/exponential.lean @@ -28,7 +28,7 @@ variables {A : Type*} open complex -lemma self_adjoint.exp_i_smul_unitary {a : A} (ha : a ∈ self_adjoint A) : +lemma is_self_adjoint.exp_i_smul_unitary {a : A} (ha : is_self_adjoint a) : exp ℂ (I • a) ∈ unitary A := begin rw [unitary.mem_iff, star_exp], @@ -42,7 +42,7 @@ end over ℂ. -/ @[simps] noncomputable def self_adjoint.exp_unitary (a : self_adjoint A) : unitary A := -⟨exp ℂ (I • a), self_adjoint.exp_i_smul_unitary (a.property)⟩ +⟨exp ℂ (I • a), a.prop.exp_i_smul_unitary⟩ open self_adjoint diff --git a/src/analysis/normed_space/star/spectrum.lean b/src/analysis/normed_space/star/spectrum.lean index 4575aad083a64..3bf43cdea1cc4 100644 --- a/src/analysis/normed_space/star/spectrum.lean +++ b/src/analysis/normed_space/star/spectrum.lean @@ -51,26 +51,24 @@ variables {A : Type*} local notation `↑ₐ` := algebra_map ℂ A -lemma spectral_radius_eq_nnnorm_of_self_adjoint [norm_one_class A] {a : A} - (ha : a ∈ self_adjoint A) : +lemma is_self_adjoint.spectral_radius_eq_nnnorm [norm_one_class A] {a : A} + (ha : is_self_adjoint a) : spectral_radius ℂ a = ∥a∥₊ := begin have hconst : tendsto (λ n : ℕ, (∥a∥₊ : ℝ≥0∞)) at_top _ := tendsto_const_nhds, refine tendsto_nhds_unique _ hconst, convert (spectrum.pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius (a : A)).comp - (nat.tendsto_pow_at_top_at_top_of_one_lt (by linarith : 1 < 2)), + (nat.tendsto_pow_at_top_at_top_of_one_lt one_lt_two), refine funext (λ n, _), - rw [function.comp_app, nnnorm_pow_two_pow_of_self_adjoint ha, ennreal.coe_pow, ←rpow_nat_cast, + rw [function.comp_app, ha.nnnorm_pow_two_pow, ennreal.coe_pow, ←rpow_nat_cast, ←rpow_mul], simp, end -lemma spectral_radius_eq_nnnorm_of_star_normal [norm_one_class A] (a : A) [is_star_normal a] : +lemma is_star_normal.spectral_radius_eq_nnnorm [norm_one_class A] (a : A) [is_star_normal a] : spectral_radius ℂ a = ∥a∥₊ := begin refine (ennreal.pow_strict_mono two_ne_zero).injective _, - have ha : a⋆ * a ∈ self_adjoint A, - from self_adjoint.mem_iff.mpr (by simpa only [star_star] using (star_mul a⋆ a)), have heq : (λ n : ℕ, ((∥(a⋆ * a) ^ n∥₊ ^ (1 / n : ℝ)) : ℝ≥0∞)) = (λ x, x ^ 2) ∘ (λ n : ℕ, ((∥a ^ n∥₊ ^ (1 / n : ℝ)) : ℝ≥0∞)), { funext, @@ -80,12 +78,13 @@ begin (spectrum.pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius a), rw ←heq at h₂, convert tendsto_nhds_unique h₂ (pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius (a⋆ * a)), - rw [spectral_radius_eq_nnnorm_of_self_adjoint ha, sq, nnnorm_star_mul_self, coe_mul], + rw [(is_self_adjoint.star_mul_self a).spectral_radius_eq_nnnorm, sq, nnnorm_star_mul_self, + coe_mul], end /-- Any element of the spectrum of a selfadjoint is real. -/ -theorem self_adjoint.mem_spectrum_eq_re [star_module ℂ A] [nontrivial A] {a : A} - (ha : a ∈ self_adjoint A) {z : ℂ} (hz : z ∈ spectrum ℂ a) : z = z.re := +theorem is_self_adjoint.mem_spectrum_eq_re [star_module ℂ A] [nontrivial A] {a : A} + (ha : is_self_adjoint a) {z : ℂ} (hz : z ∈ spectrum ℂ a) : z = z.re := begin let Iu := units.mk0 I I_ne_zero, have : exp ℂ (I • z) ∈ spectrum ℂ (exp ℂ (I • a)), @@ -94,24 +93,24 @@ begin exact complex.ext (of_real_re _) (by simpa only [←complex.exp_eq_exp_ℂ, mem_sphere_zero_iff_norm, norm_eq_abs, abs_exp, real.exp_eq_one_iff, smul_eq_mul, I_mul, neg_eq_zero] - using spectrum.subset_circle_of_unitary (self_adjoint.exp_i_smul_unitary ha) this), + using spectrum.subset_circle_of_unitary ha.exp_i_smul_unitary this), end /-- Any element of the spectrum of a selfadjoint is real. -/ -theorem self_adjoint.mem_spectrum_eq_re' [star_module ℂ A] [nontrivial A] +theorem self_adjoint.mem_spectrum_eq_re [star_module ℂ A] [nontrivial A] (a : self_adjoint A) {z : ℂ} (hz : z ∈ spectrum ℂ (a : A)) : z = z.re := -self_adjoint.mem_spectrum_eq_re a.property hz +a.prop.mem_spectrum_eq_re hz /-- The spectrum of a selfadjoint is real -/ -theorem self_adjoint.coe_re_map_spectrum [star_module ℂ A] [nontrivial A] {a : A} - (ha : a ∈ self_adjoint A) : spectrum ℂ a = (coe ∘ re '' (spectrum ℂ a) : set ℂ) := -le_antisymm (λ z hz, ⟨z, hz, (self_adjoint.mem_spectrum_eq_re ha hz).symm⟩) (λ z, by +theorem is_self_adjoint.coe_re_map_spectrum [star_module ℂ A] [nontrivial A] {a : A} + (ha : is_self_adjoint a) : spectrum ℂ a = (coe ∘ re '' (spectrum ℂ a) : set ℂ) := +le_antisymm (λ z hz, ⟨z, hz, (ha.mem_spectrum_eq_re hz).symm⟩) (λ z, by { rintros ⟨z, hz, rfl⟩, - simpa only [(self_adjoint.mem_spectrum_eq_re ha hz).symm, function.comp_app] using hz }) + simpa only [(ha.mem_spectrum_eq_re hz).symm, function.comp_app] using hz }) /-- The spectrum of a selfadjoint is real -/ -theorem self_adjoint.coe_re_map_spectrum' [star_module ℂ A] [nontrivial A] (a : self_adjoint A) : +theorem self_adjoint.coe_re_map_spectrum [star_module ℂ A] [nontrivial A] (a : self_adjoint A) : spectrum ℂ (a : A) = (coe ∘ re '' (spectrum ℂ (a : A)) : set ℂ) := -self_adjoint.coe_re_map_spectrum a.property +a.property.coe_re_map_spectrum end complex_scalars From c5be877a3e8f77c486665389993501b9818e4731 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 25 Aug 2022 05:07:57 +0000 Subject: [PATCH 039/828] =?UTF-8?q?feat(tactic/tauto):=20add=20support=20f?= =?UTF-8?q?or=20`=C2=AC=5F=20=E2=89=A0=20=5F`=20(#16232)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is an attempt to fix [this issue](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/tauto!.20fails.20on.20ne). --- src/tactic/tauto.lean | 15 ++++++++------- test/tauto.lean | 8 ++++++++ 2 files changed, 16 insertions(+), 7 deletions(-) diff --git a/src/tactic/tauto.lean b/src/tactic/tauto.lean index 89ffec7d81452..a224e560f85d5 100644 --- a/src/tactic/tauto.lean +++ b/src/tactic/tauto.lean @@ -22,13 +22,14 @@ do hs ← local_context, do h ← get_local h.local_pp_name, e ← infer_type h, match e with - | `(¬ _ = _) := replace h.local_pp_name ``(mt iff.to_eq %%h) - | `(_ ≠ _) := replace h.local_pp_name ``(mt iff.to_eq %%h) - | `(_ = _) := replace h.local_pp_name ``(eq.to_iff %%h) - | `(¬ (_ ∧ _)) := replace h.local_pp_name ``(decidable.not_and_distrib'.mp %%h) <|> - replace h.local_pp_name ``(decidable.not_and_distrib.mp %%h) - | `(¬ (_ ∨ _)) := replace h.local_pp_name ``(not_or_distrib.mp %%h) - | `(¬ ¬ _) := replace h.local_pp_name ``(decidable.of_not_not %%h) + | `(¬ _ = _) := replace h.local_pp_name ``(mt iff.to_eq %%h) + | `(_ ≠ _) := replace h.local_pp_name ``(mt iff.to_eq %%h) + | `(_ = _) := replace h.local_pp_name ``(eq.to_iff %%h) + | `(¬ (_ ∧ _)) := replace h.local_pp_name ``(decidable.not_and_distrib'.mp %%h) <|> + replace h.local_pp_name ``(decidable.not_and_distrib.mp %%h) + | `(¬ (_ ∨ _)) := replace h.local_pp_name ``(not_or_distrib.mp %%h) + | `(¬ _ ≠ _) := replace h.local_pp_name ``(decidable.of_not_not %%h) + | `(¬ ¬ _) := replace h.local_pp_name ``(decidable.of_not_not %%h) | `(¬ (_ → (_ : Prop))) := replace h.local_pp_name ``(decidable.not_imp.mp %%h) | `(¬ (_ ↔ _)) := replace h.local_pp_name ``(decidable.not_iff.mp %%h) | `(_ ↔ _) := replace h.local_pp_name ``(decidable.iff_iff_and_or_not_and_not.mp %%h) <|> diff --git a/test/tauto.lean b/test/tauto.lean index 14434287cd890..23c8e69030152 100644 --- a/test/tauto.lean +++ b/test/tauto.lean @@ -100,3 +100,11 @@ begin end end closer + +/- Zulip discussion: +https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/tauto!.20fails.20on.20ne +-/ +example {x y : ℕ} (h : ¬x ≠ y) : x = y := +begin + tauto!, +end From b6931e1ca9c430465c510822420d3f5fbfd09d1a Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Thu, 25 Aug 2022 05:07:58 +0000 Subject: [PATCH 040/828] chore(scripts): update nolints.txt (#16244) I am happy to remove some nolints for you! --- scripts/nolints.txt | 6 ------ scripts/style-exceptions.txt | 2 +- 2 files changed, 1 insertion(+), 7 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 35d29446a863b..1cf093edd0e28 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -214,12 +214,6 @@ apply_nolint add_monoid_hom.pi_ext fintype_finite apply_nolint finset.noncomm_prod_union_of_disjoint to_additive_doc apply_nolint monoid_hom.pi_ext fintype_finite --- data/finset/pi_induction.lean -apply_nolint finset.induction_on_pi fintype_finite -apply_nolint finset.induction_on_pi_max fintype_finite -apply_nolint finset.induction_on_pi_min fintype_finite -apply_nolint finset.induction_on_pi_of_choice fintype_finite - -- data/finsupp/pwo.lean apply_nolint finsupp.is_pwo fintype_finite diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index 50c219e809d51..0ee7f31fc65e1 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -20,7 +20,7 @@ src/data/rbtree/main.lean : line 11 : ERR_MOD : Module docstring missing, or too src/data/rbtree/min_max.lean : line 7 : ERR_MOD : Module docstring missing, or too late src/data/seq/computation.lean : line 12 : ERR_MOD : Module docstring missing, or too late src/data/seq/parallel.lean : line 14 : ERR_MOD : Module docstring missing, or too late -src/data/seq/seq.lean : line 11 : ERR_MOD : Module docstring missing, or too late +src/data/seq/seq.lean : line 12 : ERR_MOD : Module docstring missing, or too late src/data/seq/wseq.lean : line 10 : ERR_MOD : Module docstring missing, or too late src/logic/relator.lean : line 11 : ERR_MOD : Module docstring missing, or too late src/meta/coinductive_predicates.lean : line 8 : ERR_MOD : Module docstring missing, or too late From 36e16da1642a2dac652ac1c2dfc5e8b6f9eee091 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 25 Aug 2022 08:22:00 +0000 Subject: [PATCH 041/828] docs(algebra/field/basic): fix typo (#16241) --- src/algebra/field/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/field/basic.lean b/src/algebra/field/basic.lean index 6a73e79985570..4315534eb93b8 100644 --- a/src/algebra/field/basic.lean +++ b/src/algebra/field/basic.lean @@ -80,7 +80,7 @@ class division_semiring (α : Type*) extends semiring α, group_with_zero α An instance of `division_ring K` includes maps `of_rat : ℚ → K` and `qsmul : ℚ → K → K`. If the division ring has positive characteristic p, we define `of_rat (1 / p) = 1 / 0 = 0` for consistency with our division by zero convention. -The fields `of_rat` and `qsmul are needed to implement the +The fields `of_rat` and `qsmul` are needed to implement the `algebra_rat [division_ring K] : algebra ℚ K` instance, since we need to control the specific definitions for some special cases of `K` (in particular `K = ℚ` itself). See also Note [forgetful inheritance]. From b62580375619170f369adb3010dff4643ad43e4f Mon Sep 17 00:00:00 2001 From: Kevin Wilson Date: Thu, 25 Aug 2022 11:08:07 +0000 Subject: [PATCH 042/828] feat(topology/uniform_space/uniform_convergence): tendsto_uniformly_on_filter (#15871) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Currently, mathlib supports several notions of uniform convergence (e.g., tendsto_uniformly_on). These are "global" versions of a more local notion, which we call tendsto_uniformly_on_filter. Specifically, as revealed by tendsto_prod_top_iff is that uniform convergence means convergence on a product filter. So why can't you have a more general filter than a principal filter? There's no reason you can't! Indeed, if you replace 𝓟 s with 𝓝 x you get a notion of "local uniform convergence" which is enough to prove, e.g., the derivative operator at a point commutes with the pointwise limit. --- src/order/filter/prod.lean | 36 ++ src/topology/algebra/uniform_group.lean | 22 +- .../uniform_space/uniform_convergence.lean | 350 +++++++++++++++--- 3 files changed, 352 insertions(+), 56 deletions(-) diff --git a/src/order/filter/prod.lean b/src/order/filter/prod.lean index 7c93895fa46c3..40dd60b283837 100644 --- a/src/order/filter/prod.lean +++ b/src/order/filter/prod.lean @@ -94,6 +94,10 @@ begin simp only [mem_univ, forall_true_left] end +lemma eventually_prod_principal_iff {p : α × β → Prop} {s : set β} : + (∀ᶠ (x : α × β) in (f ×ᶠ (𝓟 s)), p x) ↔ ∀ᶠ (x : α) in f, ∀ (y : β), y ∈ s → p (x, y) := +by { rw [eventually_iff, eventually_iff, mem_prod_principal], simp only [mem_set_of_eq], } + lemma comap_prod (f : α → β × γ) (b : filter β) (c : filter γ) : comap f (b ×ᶠ c) = (comap (prod.fst ∘ f) b) ⊓ (comap (prod.snd ∘ f) c) := by erw [comap_inf, filter.comap_comap, filter.comap_comap] @@ -166,6 +170,26 @@ begin apply (ht.and hs).mono (λ x hx, hst hx.1 hx.2), end +lemma eventually.diag_of_prod_left {f : filter α} {g : filter γ} + {p : (α × α) × γ → Prop} : + (∀ᶠ x in (f ×ᶠ f ×ᶠ g), p x) → + (∀ᶠ (x : α × γ) in (f ×ᶠ g), p ((x.1, x.1), x.2)) := +begin + intros h, + obtain ⟨t, ht, s, hs, hst⟩ := eventually_prod_iff.1 h, + refine (ht.diag_of_prod.prod_mk hs).mono (λ x hx, by simp only [hst hx.1 hx.2, prod.mk.eta]), +end + +lemma eventually.diag_of_prod_right {f : filter α} {g : filter γ} + {p : α × γ × γ → Prop} : + (∀ᶠ x in (f ×ᶠ (g ×ᶠ g)), p x) → + (∀ᶠ (x : α × γ) in (f ×ᶠ g), p (x.1, x.2, x.2)) := +begin + intros h, + obtain ⟨t, ht, s, hs, hst⟩ := eventually_prod_iff.1 h, + refine (ht.prod_mk hs.diag_of_prod).mono (λ x hx, by simp only [hst hx.1 hx.2, prod.mk.eta]), +end + lemma tendsto_diag : tendsto (λ i, (i, i)) f (f ×ᶠ f) := tendsto_iff_eventually.mpr (λ _ hpr, hpr.diag_of_prod) @@ -181,6 +205,14 @@ by { rw [filter.prod, comap_infi, inf_infi], simp only [filter.prod, eq_self_iff f₁ ×ᶠ g₁ ≤ f₂ ×ᶠ g₂ := inf_le_inf (comap_mono hf) (comap_mono hg) +lemma prod_mono_left (g : filter β) {f₁ f₂ : filter α} (hf : f₁ ≤ f₂) : + f₁ ×ᶠ g ≤ f₂ ×ᶠ g := +filter.prod_mono hf rfl.le + +lemma prod_mono_right (f : filter α) {g₁ g₂ : filter β} (hf : g₁ ≤ g₂) : + f ×ᶠ g₁ ≤ f ×ᶠ g₂ := +filter.prod_mono rfl.le hf + lemma {u v w x} prod_comap_comap_eq {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} {f₁ : filter α₁} {f₂ : filter α₂} {m₁ : β₁ → α₁} {m₂ : β₂ → α₂} : (comap m₁ f₁) ×ᶠ (comap m₂ f₂) = comap (λ p : β₁×β₂, (m₁ p.1, m₂ p.2)) (f₁ ×ᶠ f₂) := @@ -193,6 +225,10 @@ by simp only [filter.prod, comap_comap, (∘), inf_comm, prod.fst_swap, lemma prod_comm : f ×ᶠ g = map (λ p : β×α, (p.2, p.1)) (g ×ᶠ f) := by { rw [prod_comm', ← map_swap_eq_comap_swap], refl } +lemma eventually_swap_iff {p : (α × β) → Prop} : (∀ᶠ (x : α × β) in (f ×ᶠ g), p x) ↔ + ∀ᶠ (y : β × α) in (g ×ᶠ f), p y.swap := +by { rw [prod_comm, eventually_map], simpa, } + lemma prod_assoc (f : filter α) (g : filter β) (h : filter γ) : map (equiv.prod_assoc α β γ) ((f ×ᶠ g) ×ᶠ h) = f ×ᶠ (g ×ᶠ h) := by simp_rw [← comap_equiv_symm, filter.prod, comap_inf, comap_comap, inf_assoc, function.comp, diff --git a/src/topology/algebra/uniform_group.lean b/src/topology/algebra/uniform_group.lean index 53c720bdbcd1a..1833669c5c04b 100644 --- a/src/topology/algebra/uniform_group.lean +++ b/src/topology/algebra/uniform_group.lean @@ -330,7 +330,19 @@ uniform_continuous_inv.comp_cauchy_seq h by simp [← preimage_smul_inv, preimage] section uniform_convergence -variables {ι : Type*} {l : filter ι} {f f' : ι → β → α} {g g' : β → α} {s : set β} +variables {ι : Type*} {l : filter ι} {l' : filter β} {f f' : ι → β → α} {g g' : β → α} {s : set β} + +@[to_additive] lemma tendsto_uniformly_on_filter.mul (hf : tendsto_uniformly_on_filter f g l l') + (hf' : tendsto_uniformly_on_filter f' g' l l') : + tendsto_uniformly_on_filter (f * f') (g * g') l l' := +λ u hu, ((uniform_continuous_mul.comp_tendsto_uniformly_on_filter + (hf.prod hf')) u hu).diag_of_prod_left + +@[to_additive] lemma tendsto_uniformly_on_filter.div (hf : tendsto_uniformly_on_filter f g l l') + (hf' : tendsto_uniformly_on_filter f' g' l l') : + tendsto_uniformly_on_filter (f / f') (g / g') l l' := +λ u hu, ((uniform_continuous_div.comp_tendsto_uniformly_on_filter + (hf.prod hf')) u hu).diag_of_prod_left @[to_additive] lemma tendsto_uniformly_on.mul (hf : tendsto_uniformly_on f g l s) (hf' : tendsto_uniformly_on f' g' l s) : tendsto_uniformly_on (f * f') (g * g') l s := @@ -340,6 +352,14 @@ variables {ι : Type*} {l : filter ι} {f f' : ι → β → α} {g g' : β → (hf' : tendsto_uniformly_on f' g' l s) : tendsto_uniformly_on (f / f') (g / g') l s := λ u hu, ((uniform_continuous_div.comp_tendsto_uniformly_on (hf.prod hf')) u hu).diag_of_prod +@[to_additive] lemma tendsto_uniformly.mul (hf : tendsto_uniformly f g l) + (hf' : tendsto_uniformly f' g' l) : tendsto_uniformly (f * f') (g * g') l := +λ u hu, ((uniform_continuous_mul.comp_tendsto_uniformly (hf.prod hf')) u hu).diag_of_prod + +@[to_additive] lemma tendsto_uniformly.div (hf : tendsto_uniformly f g l) + (hf' : tendsto_uniformly f' g' l) : tendsto_uniformly (f / f') (g / g') l := +λ u hu, ((uniform_continuous_div.comp_tendsto_uniformly (hf.prod hf')) u hu).diag_of_prod + @[to_additive] lemma uniform_cauchy_seq_on.mul (hf : uniform_cauchy_seq_on f l s) (hf' : uniform_cauchy_seq_on f' l s) : uniform_cauchy_seq_on (f * f') l s := λ u hu, by simpa using ((uniform_continuous_mul.comp_uniform_cauchy_seq_on (hf.prod' hf')) u hu) diff --git a/src/topology/uniform_space/uniform_convergence.lean b/src/topology/uniform_space/uniform_convergence.lean index ebb7ea3e693c6..b351b64278cb5 100644 --- a/src/topology/uniform_space/uniform_convergence.lean +++ b/src/topology/uniform_space/uniform_convergence.lean @@ -41,6 +41,13 @@ convergence what a Cauchy sequence is to the usual notion of convergence. ## Implementation notes +We derive most of our initial results from an auxiliary definition `tendsto_uniformly_on_filter`. +This definition in and of itself can sometimes be useful, e.g., when studying the local behavior +of the `Fₙ` near a point, which would typically look like `tendsto_uniformly_on_filter F f p (𝓝 x)`. +Still, while this may be the "correct" definition (see +`tendsto_uniformly_on_iff_tendsto_uniformly_on_filter`), it is somewhat unwieldy to work with in +practice. Thus, we provide the more traditional definition in `tendsto_uniformly_on`. + Most results hold under weaker assumptions of locally uniform approximation. In a first section, we prove the results under these weaker assumptions. Then, we derive the results on uniform convergence from them. @@ -57,7 +64,8 @@ open set filter universes u v w variables {α β γ ι : Type*} [uniform_space β] -variables {F : ι → α → β} {f : α → β} {s s' : set α} {x : α} {p : filter ι} {g : ι → α} +variables {F : ι → α → β} {f : α → β} {s s' : set α} {x : α} {p : filter ι} {p' : filter α} + {g : ι → α} /-! ### Different notions of uniform convergence @@ -65,11 +73,39 @@ variables {F : ι → α → β} {f : α → β} {s s' : set α} {x : α} {p : f We define uniform convergence and locally uniform convergence, on a set or in the whole space. -/ +/-- A sequence of functions `Fₙ` converges uniformly on a filter `p'` to a limiting function `f` +with respect to the filter `p` if, for any entourage of the diagonal `u`, one has +`p ×ᶠ p'`-eventually `(f x, Fₙ x) ∈ u`. -/ +def tendsto_uniformly_on_filter (F : ι → α → β) (f : α → β) (p : filter ι) (p' : filter α) := +∀ u ∈ 𝓤 β, ∀ᶠ (n : ι × α) in (p ×ᶠ p'), (f n.snd, F n.fst n.snd) ∈ u + +/-- +A sequence of functions `Fₙ` converges uniformly on a filter `p'` to a limiting function `f` w.r.t. +filter `p` iff the function `(n, x) ↦ (f x, Fₙ x)` converges along `p ×ᶠ p'` to the uniformity. +In other words: one knows nothing about the behavior of `x` in this limit besides it being in `p'`. +-/ +lemma tendsto_uniformly_on_filter_iff_tendsto : + tendsto_uniformly_on_filter F f p p' ↔ + tendsto (λ q : ι × α, (f q.2, F q.1 q.2)) (p ×ᶠ p') (𝓤 β) := +forall₂_congr $ λ u u_in, by simp [mem_map, filter.eventually, mem_prod_iff, preimage] + /-- A sequence of functions `Fₙ` converges uniformly on a set `s` to a limiting function `f` with respect to the filter `p` if, for any entourage of the diagonal `u`, one has `p`-eventually `(f x, Fₙ x) ∈ u` for all `x ∈ s`. -/ def tendsto_uniformly_on (F : ι → α → β) (f : α → β) (p : filter ι) (s : set α) := -∀ u ∈ 𝓤 β, ∀ᶠ n in p, ∀ x ∈ s, (f x, F n x) ∈ u +∀ u ∈ 𝓤 β, ∀ᶠ n in p, ∀ (x : α), x ∈ s → (f x, F n x) ∈ u + +lemma tendsto_uniformly_on_iff_tendsto_uniformly_on_filter : + tendsto_uniformly_on F f p s ↔ tendsto_uniformly_on_filter F f p (𝓟 s) := +begin + simp only [tendsto_uniformly_on, tendsto_uniformly_on_filter], + apply forall₂_congr, + simp_rw [eventually_prod_principal_iff], + simp, +end + +alias tendsto_uniformly_on_iff_tendsto_uniformly_on_filter ↔ + tendsto_uniformly_on.tendsto_uniformly_on_filter tendsto_uniformly_on_filter.tendsto_uniformly_on /-- A sequence of functions `Fₙ` converges uniformly on a set `s` to a limiting function `f` w.r.t. @@ -78,17 +114,35 @@ In other words: one knows nothing about the behavior of `x` in this limit beside -/ lemma tendsto_uniformly_on_iff_tendsto {F : ι → α → β} {f : α → β} {p : filter ι} {s : set α} : tendsto_uniformly_on F f p s ↔ tendsto (λ q : ι × α, (f q.2, F q.1 q.2)) (p ×ᶠ 𝓟 s) (𝓤 β) := -forall₂_congr $ λ u u_in, by simp [mem_map, filter.eventually, mem_prod_principal] +by simp [tendsto_uniformly_on_iff_tendsto_uniformly_on_filter, + tendsto_uniformly_on_filter_iff_tendsto] /-- A sequence of functions `Fₙ` converges uniformly to a limiting function `f` with respect to a filter `p` if, for any entourage of the diagonal `u`, one has `p`-eventually `(f x, Fₙ x) ∈ u` for all `x`. -/ def tendsto_uniformly (F : ι → α → β) (f : α → β) (p : filter ι) := -∀ u ∈ 𝓤 β, ∀ᶠ n in p, ∀ x, (f x, F n x) ∈ u +∀ u ∈ 𝓤 β, ∀ᶠ n in p, ∀ (x : α), (f x, F n x) ∈ u + +lemma tendsto_uniformly_iff_tendsto_uniformly_on_filter : + tendsto_uniformly F f p ↔ tendsto_uniformly_on_filter F f p ⊤ := +begin + simp only [tendsto_uniformly, tendsto_uniformly_on_filter], + apply forall₂_congr, + simp_rw [← principal_univ, eventually_prod_principal_iff], + simp, +end + +lemma tendsto_uniformly.tendsto_uniformly_on_filter + (h : tendsto_uniformly F f p) : tendsto_uniformly_on_filter F f p ⊤ := +by rwa ← tendsto_uniformly_iff_tendsto_uniformly_on_filter lemma tendsto_uniformly_on_iff_tendsto_uniformly_comp_coe : tendsto_uniformly_on F f p s ↔ tendsto_uniformly (λ i (x : s), F i x) (f ∘ coe) p := -forall₂_congr $ λ V hV, by simp +begin + apply forall₂_congr, + intros u hu, + simp, +end /-- A sequence of functions `Fₙ` converges uniformly to a limiting function `f` w.r.t. @@ -97,43 +151,107 @@ In other words: one knows nothing about the behavior of `x` in this limit. -/ lemma tendsto_uniformly_iff_tendsto {F : ι → α → β} {f : α → β} {p : filter ι} : tendsto_uniformly F f p ↔ tendsto (λ q : ι × α, (f q.2, F q.1 q.2)) (p ×ᶠ ⊤) (𝓤 β) := -forall₂_congr $ λ u u_in, by simp [mem_map, filter.eventually, mem_prod_top] +by simp [tendsto_uniformly_iff_tendsto_uniformly_on_filter, tendsto_uniformly_on_filter_iff_tendsto] + +/-- Uniform converence implies pointwise convergence. -/ +lemma tendsto_uniformly_on_filter.tendsto_at (h : tendsto_uniformly_on_filter F f p p') + (hx : 𝓟 {x} ≤ p') : tendsto (λ n, F n x) p $ 𝓝 (f x) := +begin + refine uniform.tendsto_nhds_right.mpr (λ u hu, mem_map.mpr _), + filter_upwards [(h u hu).curry], + intros i h, + simpa using (h.filter_mono hx), +end + +/-- Uniform converence implies pointwise convergence. -/ +lemma tendsto_uniformly_on.tendsto_at (h : tendsto_uniformly_on F f p s) {x : α} (hx : x ∈ s) : + tendsto (λ n, F n x) p $ 𝓝 (f x) := +h.tendsto_uniformly_on_filter.tendsto_at + (le_principal_iff.mpr $ mem_principal.mpr $ singleton_subset_iff.mpr $ hx) /-- Uniform converence implies pointwise convergence. -/ lemma tendsto_uniformly.tendsto_at (h : tendsto_uniformly F f p) (x : α) : tendsto (λ n, F n x) p $ 𝓝 (f x) := -uniform.tendsto_nhds_right.mpr $ λ u hu, mem_map.mpr $ by { filter_upwards [h u hu], tauto, } +h.tendsto_uniformly_on_filter.tendsto_at le_top lemma tendsto_uniformly_on_univ : tendsto_uniformly_on F f p univ ↔ tendsto_uniformly F f p := by simp [tendsto_uniformly_on, tendsto_uniformly] +lemma tendsto_uniformly_on_filter.mono_left {p'' : filter ι} + (h : tendsto_uniformly_on_filter F f p p') (hp : p'' ≤ p) : + tendsto_uniformly_on_filter F f p'' p' := +λ u hu, (h u hu).filter_mono (p'.prod_mono_left hp) + +lemma tendsto_uniformly_on_filter.mono_right {p'' : filter α} + (h : tendsto_uniformly_on_filter F f p p') (hp : p'' ≤ p') : + tendsto_uniformly_on_filter F f p p'' := +λ u hu, (h u hu).filter_mono (p.prod_mono_right hp) + lemma tendsto_uniformly_on.mono {s' : set α} (h : tendsto_uniformly_on F f p s) (h' : s' ⊆ s) : tendsto_uniformly_on F f p s' := -λ u hu, (h u hu).mono (λ n hn x hx, hn x (h' hx)) +tendsto_uniformly_on_iff_tendsto_uniformly_on_filter.mpr + (h.tendsto_uniformly_on_filter.mono_right (le_principal_iff.mpr $ mem_principal.mpr h')) + +lemma tendsto_uniformly_on_filter.congr {F' : ι → α → β} + (hf : tendsto_uniformly_on_filter F f p p') + (hff' : ∀ᶠ (n : ι × α) in (p ×ᶠ p'), F n.fst n.snd = F' n.fst n.snd) : + tendsto_uniformly_on_filter F' f p p' := +begin + refine (λ u hu, ((hf u hu).and hff').mono (λ n h, _)), + rw ← h.right, + exact h.left, +end lemma tendsto_uniformly_on.congr {F' : ι → α → β} (hf : tendsto_uniformly_on F f p s) (hff' : ∀ᶠ n in p, set.eq_on (F n) (F' n) s) : tendsto_uniformly_on F' f p s := begin - refine (λ u hu, ((hf u hu).and hff').mono (λ n h x hx, _)), - rw ← h.right hx, - exact h.left x hx, + rw tendsto_uniformly_on_iff_tendsto_uniformly_on_filter at hf ⊢, + refine hf.congr _, + rw eventually_iff at hff' ⊢, + simp only [set.eq_on] at hff', + simp only [mem_prod_principal, hff', mem_set_of_eq], end protected lemma tendsto_uniformly.tendsto_uniformly_on (h : tendsto_uniformly F f p) : tendsto_uniformly_on F f p s := (tendsto_uniformly_on_univ.2 h).mono (subset_univ s) +/-- Composing on the right by a function preserves uniform convergence on a filter -/ +lemma tendsto_uniformly_on_filter.comp (h : tendsto_uniformly_on_filter F f p p') (g : γ → α) : + tendsto_uniformly_on_filter (λ n, F n ∘ g) (f ∘ g) p (p'.comap g) := +begin + intros u hu, + obtain ⟨pa, hpa, pb, hpb, hpapb⟩ := eventually_prod_iff.mp (h u hu), + rw eventually_prod_iff, + simp_rw eventually_comap, + exact ⟨pa, hpa, pb ∘ g, ⟨hpb.mono (λ x hx y hy, by simp only [hx, hy, function.comp_app]), + λ x hx y hy, hpapb hx hy⟩⟩, +end + /-- Composing on the right by a function preserves uniform convergence on a set -/ lemma tendsto_uniformly_on.comp (h : tendsto_uniformly_on F f p s) (g : γ → α) : tendsto_uniformly_on (λ n, F n ∘ g) (f ∘ g) p (g ⁻¹' s) := -λ u hu, (h u hu).mono (λ i hi, λ a, hi (g a)) +begin + rw tendsto_uniformly_on_iff_tendsto_uniformly_on_filter at h ⊢, + simpa [tendsto_uniformly_on, comap_principal] using (tendsto_uniformly_on_filter.comp h g), +end /-- Composing on the right by a function preserves uniform convergence -/ lemma tendsto_uniformly.comp (h : tendsto_uniformly F f p) (g : γ → α) : tendsto_uniformly (λ n, F n ∘ g) (f ∘ g) p := -λ u hu, (h u hu).mono (λ i hi, λ a, hi (g a)) +begin + rw tendsto_uniformly_iff_tendsto_uniformly_on_filter at h ⊢, + simpa [principal_univ, comap_principal] using (h.comp g), +end + +/-- Composing on the left by a uniformly continuous function preserves + uniform convergence on a filter -/ +lemma uniform_continuous.comp_tendsto_uniformly_on_filter [uniform_space γ] {g : β → γ} + (hg : uniform_continuous g) (h : tendsto_uniformly_on_filter F f p p') : + tendsto_uniformly_on_filter (λ i, g ∘ (F i)) (g ∘ f) p p' := +λ u hu, h _ (hg hu) /-- Composing on the left by a uniformly continuous function preserves uniform convergence on a set -/ @@ -148,17 +266,32 @@ lemma uniform_continuous.comp_tendsto_uniformly [uniform_space γ] {g : β → tendsto_uniformly (λ i, g ∘ (F i)) (g ∘ f) p := λ u hu, h _ (hg hu) +lemma tendsto_uniformly_on_filter.prod_map {ι' α' β' : Type*} [uniform_space β'] + {F' : ι' → α' → β'} {f' : α' → β'} {q : filter ι'} {q' : filter α'} + (h : tendsto_uniformly_on_filter F f p p') (h' : tendsto_uniformly_on_filter F' f' q q') : + tendsto_uniformly_on_filter (λ (i : ι × ι'), prod.map (F i.1) (F' i.2)) + (prod.map f f') (p.prod q) (p'.prod q') := +begin + intros u hu, + rw [uniformity_prod_eq_prod, mem_map, mem_prod_iff] at hu, + obtain ⟨v, hv, w, hw, hvw⟩ := hu, + apply (tendsto_swap4_prod.eventually ((h v hv).prod_mk (h' w hw))).mono, + simp only [prod_map, and_imp, prod.forall], + intros n n' x hxv hxw, + have hout : ((f x.fst, F n x.fst), (f' x.snd, F' n' x.snd)) ∈ + {x : (β × β) × β' × β' | ((x.fst.fst, x.snd.fst), x.fst.snd, x.snd.snd) ∈ u}, + { exact mem_of_mem_of_subset (set.mem_prod.mpr ⟨hxv, hxw⟩) hvw, }, + exact hout, +end + lemma tendsto_uniformly_on.prod_map {ι' α' β' : Type*} [uniform_space β'] {F' : ι' → α' → β'} {f' : α' → β'} {p' : filter ι'} {s' : set α'} (h : tendsto_uniformly_on F f p s) (h' : tendsto_uniformly_on F' f' p' s') : tendsto_uniformly_on (λ (i : ι × ι'), prod.map (F i.1) (F' i.2)) (prod.map f f') (p.prod p') (s ×ˢ s') := begin - intros u hu, - rw [uniformity_prod_eq_prod, mem_map, mem_prod_iff] at hu, - obtain ⟨v, hv, w, hw, hvw⟩ := hu, - exact mem_prod_iff.mpr ⟨_, h v hv, _, h' w hw, - λ i hi a ha, hvw (show (_, _) ∈ v ×ˢ w, from ⟨hi.1 a.1 ha.1, hi.2 a.2 ha.2⟩)⟩, + rw tendsto_uniformly_on_iff_tendsto_uniformly_on_filter at h h' ⊢, + simpa only [prod_principal_principal] using (h.prod_map h'), end lemma tendsto_uniformly.prod_map {ι' α' β' : Type*} [uniform_space β'] {F' : ι' → α' → β'} @@ -169,6 +302,13 @@ begin exact h.prod_map h', end +lemma tendsto_uniformly_on_filter.prod {ι' β' : Type*} [uniform_space β'] + {F' : ι' → α → β'} {f' : α → β'} {q : filter ι'} + (h : tendsto_uniformly_on_filter F f p p') (h' : tendsto_uniformly_on_filter F' f' q p') : + tendsto_uniformly_on_filter (λ (i : ι × ι') a, (F i.1 a, F' i.2 a)) + (λ a, (f a, f' a)) (p.prod q) p' := +λ u hu, ((h.prod_map h') u hu).diag_of_prod_right + lemma tendsto_uniformly_on.prod {ι' β' : Type*} [uniform_space β'] {F' : ι' → α → β'} {f' : α → β'} {p' : filter ι'} (h : tendsto_uniformly_on F f p s) (h' : tendsto_uniformly_on F' f' p' s) : tendsto_uniformly_on (λ (i : ι × ι') a, (F i.1 a, F' i.2 a)) (λ a, (f a, f' a)) (p.prod p') s := @@ -179,20 +319,30 @@ lemma tendsto_uniformly.prod {ι' β' : Type*} [uniform_space β'] {F' : ι' → tendsto_uniformly (λ (i : ι × ι') a, (F i.1 a, F' i.2 a)) (λ a, (f a, f' a)) (p.prod p') := (h.prod_map h').comp (λ a, (a, a)) +/-- Uniform convergence on a filter `p'` to a constant function is equivalent to convergence in +`p ×ᶠ p'`. -/ +lemma tendsto_prod_filter_iff {c : β} : + tendsto ↿F (p ×ᶠ p') (𝓝 c) ↔ tendsto_uniformly_on_filter F (λ _, c) p p' := +begin + simp_rw [tendsto, nhds_eq_comap_uniformity, map_le_iff_le_comap.symm, map_map, le_def, mem_map], + exact forall₂_congr (λ u hu, by simpa [eventually_iff]), +end + /-- Uniform convergence on a set `s` to a constant function is equivalent to convergence in `p ×ᶠ 𝓟 s`. -/ lemma tendsto_prod_principal_iff {c : β} : tendsto ↿F (p ×ᶠ 𝓟 s) (𝓝 c) ↔ tendsto_uniformly_on F (λ _, c) p s := begin - unfold tendsto, - simp_rw [nhds_eq_comap_uniformity, map_le_iff_le_comap.symm, map_map, le_def, mem_map, - mem_prod_principal], - simpa, + rw tendsto_uniformly_on_iff_tendsto_uniformly_on_filter, + exact tendsto_prod_filter_iff, end /-- Uniform convergence to a constant function is equivalent to convergence in `p ×ᶠ ⊤`. -/ lemma tendsto_prod_top_iff {c : β} : tendsto ↿F (p ×ᶠ ⊤) (𝓝 c) ↔ tendsto_uniformly F (λ _, c) p := -by rw [←principal_univ, ←tendsto_uniformly_on_univ, ←tendsto_prod_principal_iff] +begin + rw tendsto_uniformly_iff_tendsto_uniformly_on_filter, + exact tendsto_prod_filter_iff, +end /-- Uniform convergence on the empty set is vacuously true -/ lemma tendsto_uniformly_on_empty : @@ -202,16 +352,30 @@ lemma tendsto_uniformly_on_empty : /-- Uniform convergence on a singleton is equivalent to regular convergence -/ lemma tendsto_uniformly_on_singleton_iff_tendsto : tendsto_uniformly_on F f p {x} ↔ tendsto (λ n : ι, F n x) p (𝓝 (f x)) := -by simp_rw [uniform.tendsto_nhds_right, tendsto_uniformly_on, mem_singleton_iff, forall_eq, - tendsto_def, preimage, filter.eventually] +begin + simp_rw [tendsto_uniformly_on_iff_tendsto, uniform.tendsto_nhds_right, tendsto_def], + exact forall₂_congr (λ u hu, by simp [mem_prod_principal, preimage]), +end + +/-- If a sequence `g` converges to some `b`, then the sequence of constant functions +`λ n, λ a, g n` converges to the constant function `λ a, b` on any set `s` -/ +lemma filter.tendsto.tendsto_uniformly_on_filter_const + {g : ι → β} {b : β} (hg : tendsto g p (𝓝 b)) (p' : filter α) : + tendsto_uniformly_on_filter (λ n : ι, λ a : α, g n) (λ a : α, b) p p' := +begin + rw tendsto_uniformly_on_filter_iff_tendsto, + rw uniform.tendsto_nhds_right at hg, + exact (hg.comp (tendsto_fst.comp ((@tendsto_id ι p).prod_map (@tendsto_id α p')))).congr + (λ x, by simp), +end /-- If a sequence `g` converges to some `b`, then the sequence of constant functions `λ n, λ a, g n` converges to the constant function `λ a, b` on any set `s` -/ lemma filter.tendsto.tendsto_uniformly_on_const {g : ι → β} {b : β} (hg : tendsto g p (𝓝 b)) (s : set α) : tendsto_uniformly_on (λ n : ι, λ a : α, g n) (λ a : α, b) p s := -λ u hu, hg.eventually - (eventually_of_mem (mem_nhds_left b hu) (λ x hx y hy, hx) : ∀ᶠ x in 𝓝 b, ∀ y ∈ s, (b, x) ∈ u) +tendsto_uniformly_on_iff_tendsto_uniformly_on_filter.mpr + (hg.tendsto_uniformly_on_filter_const (𝓟 s)) lemma uniform_continuous_on.tendsto_uniformly [uniform_space α] [uniform_space γ] {x : α} {U : set α} (hU : U ∈ 𝓝 x) @@ -237,31 +401,57 @@ lemma uniform_continuous₂.tendsto_uniformly [uniform_space α] [uniform_space uniform_continuous_on.tendsto_uniformly univ_mem $ by rwa [univ_prod_univ, uniform_continuous_on_univ] +/-- A sequence is uniformly Cauchy if eventually all of its pairwise differences are +uniformly bounded -/ +def uniform_cauchy_seq_on_filter + (F : ι → α → β) (p : filter ι) (p' : filter α) : Prop := + ∀ u : set (β × β), u ∈ 𝓤 β → ∀ᶠ (m : (ι × ι) × α) in ((p ×ᶠ p) ×ᶠ p'), + (F m.fst.fst m.snd, F m.fst.snd m.snd) ∈ u + /-- A sequence is uniformly Cauchy if eventually all of its pairwise differences are uniformly bounded -/ def uniform_cauchy_seq_on (F : ι → α → β) (p : filter ι) (s : set α) : Prop := - ∀ u : set (β × β), u ∈ 𝓤 β → ∀ᶠ (m : ι × ι) in (p ×ᶠ p), - ∀ (x : α), x ∈ s → (F m.fst x, F m.snd x) ∈ u + ∀ u : set (β × β), u ∈ 𝓤 β → ∀ᶠ (m : ι × ι) in (p ×ᶠ p), ∀ (x : α), x ∈ s → + (F m.fst x, F m.snd x) ∈ u + +lemma uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter : + uniform_cauchy_seq_on F p s ↔ uniform_cauchy_seq_on_filter F p (𝓟 s) := +begin + simp only [uniform_cauchy_seq_on, uniform_cauchy_seq_on_filter], + refine forall₂_congr (λ u hu, _), + rw eventually_prod_principal_iff, +end + +lemma uniform_cauchy_seq_on.uniform_cauchy_seq_on_filter (hF : uniform_cauchy_seq_on F p s) : + uniform_cauchy_seq_on_filter F p (𝓟 s) := +by rwa ←uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter /-- A sequence that converges uniformly is also uniformly Cauchy -/ -lemma tendsto_uniformly_on.uniform_cauchy_seq_on (hF : tendsto_uniformly_on F f p s) : - uniform_cauchy_seq_on F p s := +lemma tendsto_uniformly_on_filter.uniform_cauchy_seq_on_filter + (hF : tendsto_uniformly_on_filter F f p p') : + uniform_cauchy_seq_on_filter F p p' := begin intros u hu, rcases comp_symm_of_uniformity hu with ⟨t, ht, htsymm, htmem⟩, - apply ((hF t ht).prod_mk (hF t ht)).mono, - intros n h x hx, - cases h with hl hr, - specialize hl x hx, - specialize hr x hx, + have := tendsto_swap4_prod.eventually ((hF t ht).prod_mk (hF t ht)), + apply this.diag_of_prod_right.mono, + simp only [and_imp, prod.forall], + intros n1 n2 x hl hr, exact set.mem_of_mem_of_subset (prod_mk_mem_comp_rel (htsymm hl) hr) htmem, end +/-- A sequence that converges uniformly is also uniformly Cauchy -/ +lemma tendsto_uniformly_on.uniform_cauchy_seq_on (hF : tendsto_uniformly_on F f p s) : + uniform_cauchy_seq_on F p s := +uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter.mpr + hF.tendsto_uniformly_on_filter.uniform_cauchy_seq_on_filter + /-- A uniformly Cauchy sequence converges uniformly to its limit -/ -lemma uniform_cauchy_seq_on.tendsto_uniformly_on_of_tendsto [ne_bot p] - (hF : uniform_cauchy_seq_on F p s) (hF' : ∀ x : α, x ∈ s → tendsto (λ n, F n x) p (nhds (f x))) : - tendsto_uniformly_on F f p s := +lemma uniform_cauchy_seq_on_filter.tendsto_uniformly_on_filter_of_tendsto [ne_bot p] + (hF : uniform_cauchy_seq_on_filter F p p') + (hF' : ∀ᶠ (x : α) in p', tendsto (λ n, F n x) p (𝓝 (f x))) : + tendsto_uniformly_on_filter F f p p' := begin -- Proof idea: |f_n(x) - f(x)| ≤ |f_n(x) - f_m(x)| + |f_m(x) - f(x)|. We choose `n` -- so that |f_n(x) - f_m(x)| is uniformly small across `s` whenever `m ≥ n`. Then for @@ -269,30 +459,79 @@ begin intros u hu, rcases comp_symm_of_uniformity hu with ⟨t, ht, htsymm, htmem⟩, - -- Choose n - apply (hF t ht).curry.mono, - - -- Work with a specific x - intros n hn x hx, + -- We will choose n, x, and m simultaneously. n and x come from hF. m comes from hF' + -- But we need to promote hF' to the full product filter to use it + have hmc : ∀ᶠ (x : (ι × ι) × α) in p ×ᶠ p ×ᶠ p', tendsto (λ (n : ι), F n x.snd) p (𝓝 (f x.snd)), + { rw eventually_prod_iff, + refine ⟨(λ x, true), by simp, _, hF', by simp⟩, }, + + -- To apply filter operations we'll need to do some order manipulation + rw filter.eventually_swap_iff, + have := tendsto_prod_assoc.eventually (tendsto_prod_swap.eventually ((hF t ht).and hmc)), + apply this.curry.mono, + simp only [equiv.prod_assoc_apply, eventually_and, eventually_const, prod.snd_swap, + prod.fst_swap, and_imp, prod.forall], + + -- Complete the proof + intros x n hx hm', refine set.mem_of_mem_of_subset (mem_comp_rel.mpr _) htmem, + rw uniform.tendsto_nhds_right at hm', + have := hx.and (hm' ht), + obtain ⟨m, hm⟩ := this.exists, + exact ⟨F m x, ⟨hm.2, htsymm hm.1⟩⟩, +end - -- Choose m - specialize hF' x hx, - rw uniform.tendsto_nhds_right at hF', - rcases (hn.and (hF'.eventually (eventually_mem_set.mpr ht))).exists with ⟨m, hm, hm'⟩, +/-- A uniformly Cauchy sequence converges uniformly to its limit -/ +lemma uniform_cauchy_seq_on.tendsto_uniformly_on_of_tendsto [ne_bot p] + (hF : uniform_cauchy_seq_on F p s) (hF' : ∀ x : α, x ∈ s → tendsto (λ n, F n x) p (𝓝 (f x))) : + tendsto_uniformly_on F f p s := +tendsto_uniformly_on_iff_tendsto_uniformly_on_filter.mpr + (hF.uniform_cauchy_seq_on_filter.tendsto_uniformly_on_filter_of_tendsto hF') + +lemma uniform_cauchy_seq_on_filter.mono_left {p'' : filter ι} + (hf : uniform_cauchy_seq_on_filter F p p') (hp : p'' ≤ p) : + uniform_cauchy_seq_on_filter F p'' p' := +begin + intros u hu, + have := (hf u hu).filter_mono (p'.prod_mono_left (filter.prod_mono hp hp)), + exact this.mono (by simp), +end - -- Finish the proof - exact ⟨F m x, ⟨hm', htsymm (hm x hx)⟩⟩, +lemma uniform_cauchy_seq_on_filter.mono_right {p'' : filter α} + (hf : uniform_cauchy_seq_on_filter F p p') (hp : p'' ≤ p') : + uniform_cauchy_seq_on_filter F p p'' := +begin + intros u hu, + have := (hf u hu).filter_mono ((p ×ᶠ p).prod_mono_right hp), + exact this.mono (by simp), end lemma uniform_cauchy_seq_on.mono {s' : set α} (hf : uniform_cauchy_seq_on F p s) (hss' : s' ⊆ s) : uniform_cauchy_seq_on F p s' := -λ u hu, (hf u hu).mono (λ x hx y hy, hx y (hss' hy)) +begin + rw uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter at hf ⊢, + exact hf.mono_right (le_principal_iff.mpr $mem_principal.mpr hss'), +end + +/-- Composing on the right by a function preserves uniform Cauchy sequences -/ +lemma uniform_cauchy_seq_on_filter.comp {γ : Type*} (hf : uniform_cauchy_seq_on_filter F p p') + (g : γ → α) : + uniform_cauchy_seq_on_filter (λ n, F n ∘ g) p (p'.comap g) := +begin + intros u hu, + obtain ⟨pa, hpa, pb, hpb, hpapb⟩ := eventually_prod_iff.mp (hf u hu), + rw eventually_prod_iff, + refine ⟨pa, hpa, pb ∘ g, _, λ x hx y hy, hpapb hx hy⟩, + exact eventually_comap.mpr (hpb.mono (λ x hx y hy, by simp only [hx, hy, function.comp_app])), +end /-- Composing on the right by a function preserves uniform Cauchy sequences -/ lemma uniform_cauchy_seq_on.comp {γ : Type*} (hf : uniform_cauchy_seq_on F p s) (g : γ → α) : uniform_cauchy_seq_on (λ n, F n ∘ g) p (g ⁻¹' s) := -λ u hu, (hf u hu).mono (λ x hx y hy, hx (g y) hy) +begin + rw uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter at hf ⊢, + simpa only [uniform_cauchy_seq_on, comap_principal] using (hf.comp g), +end /-- Composing on the left by a uniformly continuous function preserves uniform Cauchy sequences -/ @@ -429,11 +668,11 @@ end protected lemma tendsto_uniformly_on.tendsto_locally_uniformly_on (h : tendsto_uniformly_on F f p s) : tendsto_locally_uniformly_on F f p s := -λ u hu x hx, ⟨s, self_mem_nhds_within, h u hu⟩ +λ u hu x hx,⟨s, self_mem_nhds_within, by simpa using (h u hu)⟩ protected lemma tendsto_uniformly.tendsto_locally_uniformly (h : tendsto_uniformly F f p) : tendsto_locally_uniformly F f p := -λ u hu x, ⟨univ, univ_mem, by simpa using h u hu⟩ +λ u hu x, ⟨univ, univ_mem, by simpa using (h u hu)⟩ lemma tendsto_locally_uniformly_on.mono (h : tendsto_locally_uniformly_on F f p s) (h' : s' ⊆ s) : tendsto_locally_uniformly_on F f p s' := @@ -654,7 +893,8 @@ tends to `f x` if `f` is continuous at `x` within `s`. -/ lemma tendsto_uniformly_on.tendsto_comp (h : tendsto_uniformly_on F f p s) (hf : continuous_within_at f s x) (hg : tendsto g p (𝓝[s] x)) : tendsto (λ n, F n (g n)) p (𝓝 (f x)) := -tendsto_comp_of_locally_uniform_limit_within hf hg (λ u hu, ⟨s, self_mem_nhds_within, h u hu⟩) +tendsto_comp_of_locally_uniform_limit_within hf hg (λ u hu, + ⟨s, self_mem_nhds_within, h u hu⟩) /-- If `Fₙ` tends locally uniformly to `f`, and `gₙ` tends to `x`, then `Fₙ gₙ` tends to `f x`. -/ lemma tendsto_locally_uniformly.tendsto_comp (h : tendsto_locally_uniformly F f p) From 3d0bb8702b1b443741e830beab970a81f5beb2b2 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Thu, 25 Aug 2022 13:18:46 +0000 Subject: [PATCH 043/828] feat(data/zmod/quotient): Version of `order_eq_card_zpowers` in terms of `nat.card` without a `fintype`/`finite` assumption (#16175) This PR adds a version of `order_eq_card_zpowers` in terms of `nat.card` without a `fintype`/`finite` assumption. --- src/data/zmod/quotient.lean | 8 ++++++++ src/group_theory/order_of_element.lean | 3 ++- 2 files changed, 10 insertions(+), 1 deletion(-) diff --git a/src/data/zmod/quotient.lean b/src/data/zmod/quotient.lean index f1cfd89b5585e..23f0fbecb6f07 100644 --- a/src/data/zmod/quotient.lean +++ b/src/data/zmod/quotient.lean @@ -160,4 +160,12 @@ by rw [←fintype.of_equiv_card (orbit_zpowers_equiv a b), zmod.card] exact fintype.card_ne_zero, end⟩ +/-- See also `order_eq_card_zpowers`. -/ +@[to_additive add_order_eq_card_zmultiples' "See also `add_order_eq_card_zmultiples`."] +lemma _root_.order_eq_card_zpowers' : order_of a = nat.card (zpowers a) := +begin + have := nat.card_congr (mul_action.orbit_zpowers_equiv a (1 : α)), + rwa [nat.card_zmod, orbit_subgroup_one_eq_self, eq_comm] at this, +end + end mul_action diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 03e51ce62749d..0d3369fc35c3d 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -616,7 +616,8 @@ end variables [fintype G] -@[to_additive add_order_eq_card_zmultiples] +/-- See also `order_eq_card_zpowers'`. -/ +@[to_additive add_order_eq_card_zmultiples "See also `add_order_eq_card_zmultiples'`."] lemma order_eq_card_zpowers : order_of x = fintype.card (zpowers x) := (fintype.card_fin (order_of x)).symm.trans (fintype.card_eq.2 ⟨fin_equiv_zpowers x⟩) From 4ed5f0e1c3a47185ad888aee697381054f0c9189 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 25 Aug 2022 15:37:33 +0000 Subject: [PATCH 044/828] feat(algebra/order/smul): `positivity` extension for `smul` (#16162) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add `positivity_smul`, a `positivity` extension to deal with `•`. --- src/algebra/order/smul.lean | 30 ++++++++++++++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index 21cceefe4e6aa..9b3e7f1bfc7da 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -8,6 +8,7 @@ import algebra.module.prod import algebra.order.field import algebra.order.pi import data.set.pointwise +import tactic.positivity /-! # Ordered scalar product @@ -239,3 +240,32 @@ variables {M} end ordered_add_comm_group end linear_ordered_semifield + +namespace tactic +variables [ordered_semiring R] [ordered_add_comm_monoid M] [smul_with_zero R M] [ordered_smul R M] + {a : R} {b : M} + +private lemma smul_nonneg_of_pos_of_nonneg (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ a • b := +smul_nonneg ha.le hb + +private lemma smul_nonneg_of_nonneg_of_pos (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ a • b := +smul_nonneg ha hb.le + +open positivity + +/-- Extension for the `positivity` tactic: scalar multiplication is nonnegative if both sides are +nonnegative, and strictly positive if both sides are. -/ +@[positivity] +meta def positivity_smul : expr → tactic strictness +| `(%%a • %%b) := do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | positive pa, positive pb := positive <$> mk_app ``smul_pos [pa, pb] + | positive pa, nonnegative pb := nonnegative <$> mk_app ``smul_nonneg_of_pos_of_nonneg [pa, pb] + | nonnegative pa, positive pb := nonnegative <$> mk_app ``smul_nonneg_of_nonneg_of_pos [pa, pb] + | nonnegative pa, nonnegative pb := nonnegative <$> mk_app ``smul_nonneg [pa, pb] + end +| _ := failed + +end tactic From b059a4423c74a5a9850685b66500c8c65f2c57ed Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Thu, 25 Aug 2022 18:07:03 +0000 Subject: [PATCH 045/828] feat(algebra/order/floor): lemmas about `int.floor`, `int.ceil`, `int.fract`, `round` (#16158) --- src/algebra/order/floor.lean | 130 ++++++++++++++++++++++---- src/data/int/succ_pred.lean | 2 + src/data/rat/floor.lean | 2 +- src/order/basic.lean | 12 +++ src/topology/algebra/order/floor.lean | 18 ++-- 5 files changed, 137 insertions(+), 27 deletions(-) diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index aaa5276505d00..4c38e55e3fb12 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -428,6 +428,11 @@ lemma floor_le (a : α) : (⌊a⌋ : α) ≤ a := gc_coe_floor.l_u_le a lemma floor_nonneg : 0 ≤ ⌊a⌋ ↔ 0 ≤ a := by rw [le_floor, int.cast_zero] +@[simp] lemma floor_le_sub_one_iff : ⌊a⌋ ≤ z - 1 ↔ a < z := by rw [← floor_lt, le_sub_one_iff] + +@[simp] lemma floor_le_neg_one_iff : ⌊a⌋ ≤ -1 ↔ a < 0 := +by rw [← zero_sub (1 : ℤ), floor_le_sub_one_iff, cast_zero] + lemma floor_nonpos (ha : a ≤ 0) : ⌊a⌋ ≤ 0 := begin rw [← @cast_le α, int.cast_zero], @@ -436,17 +441,20 @@ end lemma lt_succ_floor (a : α) : a < ⌊a⌋.succ := floor_lt.1 $ int.lt_succ_self _ -lemma lt_floor_add_one (a : α) : a < ⌊a⌋ + 1 := +@[simp] lemma lt_floor_add_one (a : α) : a < ⌊a⌋ + 1 := by simpa only [int.succ, int.cast_add, int.cast_one] using lt_succ_floor a -lemma sub_one_lt_floor (a : α) : a - 1 < ⌊a⌋ := sub_lt_iff_lt_add.2 (lt_floor_add_one a) +@[simp] lemma sub_one_lt_floor (a : α) : a - 1 < ⌊a⌋ := sub_lt_iff_lt_add.2 (lt_floor_add_one a) -@[simp] lemma floor_coe (z : ℤ) : ⌊(z : α)⌋ = z := +@[simp] lemma floor_int_cast (z : ℤ) : ⌊(z : α)⌋ = z := eq_of_forall_le_iff $ λ a, by rw [le_floor, int.cast_le] -@[simp] lemma floor_zero : ⌊(0 : α)⌋ = 0 := by rw [← int.cast_zero, floor_coe] +@[simp] lemma floor_nat_cast (n : ℕ) : ⌊(n : α)⌋ = n := +eq_of_forall_le_iff $ λ a, by rw [le_floor, ← cast_coe_nat, cast_le] + +@[simp] lemma floor_zero : ⌊(0 : α)⌋ = 0 := by rw [← cast_zero, floor_int_cast] -@[simp] lemma floor_one : ⌊(1 : α)⌋ = 1 := by rw [← int.cast_one, floor_coe] +@[simp] lemma floor_one : ⌊(1 : α)⌋ = 1 := by rw [← cast_one, floor_int_cast] @[mono] lemma floor_mono : monotone (floor : α → ℤ) := gc_coe_floor.monotone_u @@ -520,7 +528,7 @@ by rw [add_comm, fract_add_int] @[simp] lemma fract_sub_self (a : α) : fract a - a = -⌊a⌋ := sub_sub_cancel_left _ _ -lemma fract_nonneg (a : α) : 0 ≤ fract a := sub_nonneg.2 $ floor_le _ +@[simp] lemma fract_nonneg (a : α) : 0 ≤ fract a := sub_nonneg.2 $ floor_le _ lemma fract_lt_one (a : α) : fract a < 1 := sub_lt.1 $ sub_one_lt_floor _ @@ -529,10 +537,17 @@ lemma fract_lt_one (a : α) : fract a < 1 := sub_lt.1 $ sub_one_lt_floor _ @[simp] lemma fract_one : fract (1 : α) = 0 := by simp [fract] -@[simp] lemma fract_coe (z : ℤ) : fract (z : α) = 0 := -by { unfold fract, rw floor_coe, exact sub_self _ } +lemma abs_fract : |int.fract a| = int.fract a := abs_eq_self.mpr $ fract_nonneg a + +@[simp] lemma abs_one_sub_fract : |1 - fract a| = 1 - fract a := +abs_eq_self.mpr $ sub_nonneg.mpr (fract_lt_one a).le + +@[simp] lemma fract_int_cast (z : ℤ) : fract (z : α) = 0 := +by { unfold fract, rw floor_int_cast, exact sub_self _ } -@[simp] lemma fract_floor (a : α) : fract (⌊a⌋ : α) = 0 := fract_coe _ +@[simp] lemma fract_nat_cast (n : ℕ) : fract (n : α) = 0 := by simp [fract] + +@[simp] lemma fract_floor (a : α) : fract (⌊a⌋ : α) = 0 := fract_int_cast _ @[simp] lemma floor_fract (a : α) : ⌊fract a⌋ = 0 := by rw [floor_eq_iff, int.cast_zero, zero_add]; exact ⟨fract_nonneg _, fract_lt_one _⟩ @@ -632,14 +647,22 @@ eq_of_forall_ge_iff (λ z, by rw [neg_le, ceil_le, le_floor, int.cast_neg, neg_l lemma lt_ceil : z < ⌈a⌉ ↔ (z : α) < a := lt_iff_lt_of_le_iff_le ceil_le +@[simp] lemma add_one_le_ceil_iff : z + 1 ≤ ⌈a⌉ ↔ (z : α) < a := by rw [← lt_ceil, add_one_le_iff] + +@[simp] lemma one_le_ceil_iff : 1 ≤ ⌈a⌉ ↔ 0 < a := +by rw [← zero_add (1 : ℤ), add_one_le_ceil_iff, cast_zero] + lemma ceil_le_floor_add_one (a : α) : ⌈a⌉ ≤ ⌊a⌋ + 1 := by { rw [ceil_le, int.cast_add, int.cast_one], exact (lt_floor_add_one a).le } lemma le_ceil (a : α) : a ≤ ⌈a⌉ := gc_ceil_coe.le_u_l a -@[simp] lemma ceil_coe (z : ℤ) : ⌈(z : α)⌉ = z := +@[simp] lemma ceil_int_cast (z : ℤ) : ⌈(z : α)⌉ = z := eq_of_forall_ge_iff $ λ a, by rw [ceil_le, int.cast_le] +@[simp] lemma ceil_nat_cast (n : ℕ) : ⌈(n : α)⌉ = n := +eq_of_forall_ge_iff $ λ a, by rw [ceil_le, ← cast_coe_nat, cast_le] + lemma ceil_mono : monotone (ceil : α → ℤ) := gc_ceil_coe.monotone_l @[simp] lemma ceil_add_int (a : α) (z : ℤ) : ⌈a + z⌉ = ⌈a⌉ + z := @@ -659,9 +682,9 @@ by { rw [← lt_ceil, ← int.cast_one, ceil_add_int], apply lt_add_one } @[simp] lemma ceil_pos : 0 < ⌈a⌉ ↔ 0 < a := by rw [lt_ceil, cast_zero] -@[simp] lemma ceil_zero : ⌈(0 : α)⌉ = 0 := by rw [← int.cast_zero, ceil_coe] +@[simp] lemma ceil_zero : ⌈(0 : α)⌉ = 0 := by rw [← cast_zero, ceil_int_cast] -@[simp] lemma ceil_one : ⌈(1 : α)⌉ = 1 := by rw [←int.cast_one, ceil_coe] +@[simp] lemma ceil_one : ⌈(1 : α)⌉ = 1 := by rw [← cast_one, ceil_int_cast] lemma ceil_nonneg (ha : 0 ≤ a) : 0 ≤ ⌈a⌉ := by exact_mod_cast ha.trans (le_ceil a) @@ -684,6 +707,23 @@ cast_lt.1 $ (floor_le a).trans_lt $ h.trans_le $ le_ceil b @[simp] lemma preimage_ceil_singleton (m : ℤ) : (ceil : α → ℤ) ⁻¹' {m} = Ioc (m - 1) m := ext $ λ x, ceil_eq_iff +lemma fract_eq_zero_or_add_one_sub_ceil (a : α) : fract a = 0 ∨ fract a = a + 1 - (⌈a⌉ : α) := +begin + cases eq_or_ne (fract a) 0 with ha ha, { exact or.inl ha, }, right, + suffices : (⌈a⌉ : α) = ⌊a⌋ + 1, { rw [this, ← self_sub_fract], abel, }, + norm_cast, + rw ceil_eq_iff, + refine ⟨_, _root_.le_of_lt $ by simp⟩, + rw [cast_add, cast_one, add_tsub_cancel_right, ← self_sub_fract a, sub_lt_self_iff], + exact ha.symm.lt_of_le (fract_nonneg a), +end + +lemma ceil_eq_add_one_sub_fract (ha : fract a ≠ 0) : (⌈a⌉ : α) = a + 1 - fract a := +by { rw (or_iff_right ha).mp (fract_eq_zero_or_add_one_sub_ceil a), abel, } + +lemma ceil_sub_self_eq (ha : fract a ≠ 0) : (⌈a⌉ : α) - a = 1 - fract a := +by { rw (or_iff_right ha).mp (fract_eq_zero_or_add_one_sub_ceil a), abel, } + /-! #### Intervals -/ @[simp] lemma preimage_Ioo {a b : α} : ((coe : ℤ → α) ⁻¹' (set.Ioo a b)) = set.Ioo ⌊a⌋ ⌈b⌉ := @@ -717,22 +757,76 @@ open int /-! ### Round -/ section round -variables [linear_ordered_field α] [floor_ring α] + +section linear_ordered_ring + +variables [linear_ordered_ring α] [floor_ring α] /-- `round` rounds a number to the nearest integer. `round (1 / 2) = 1` -/ -def round (x : α) : ℤ := ⌊x + 1 / 2⌋ +def round (x : α) : ℤ := if 2 * fract x < 1 then ⌊x⌋ else ⌈x⌉ + +@[simp] lemma round_zero : round (0 : α) = 0 := by simp [round] -@[simp] lemma round_zero : round (0 : α) = 0 := floor_eq_iff.2 (by norm_num) -@[simp] lemma round_one : round (1 : α) = 1 := floor_eq_iff.2 (by norm_num) +@[simp] lemma round_one : round (1 : α) = 1 := by simp [round] + +@[simp] lemma round_nat_cast (n : ℕ) : round (n : α) = n := by simp [round] + +@[simp] lemma round_int_cast (n : ℤ) : round (n : α) = n := by simp [round] + +lemma abs_sub_round_eq_min (x : α) : |x - round x| = min (fract x) (1 - fract x) := +begin + simp_rw [round, min_def', two_mul, ← lt_tsub_iff_left], + cases lt_or_ge (fract x) (1 - fract x) with hx hx, + { rw [if_pos hx, if_pos hx, self_sub_floor, abs_fract], }, + { have : 0 < fract x, + { replace hx : 0 < fract x + fract x := lt_of_lt_of_le zero_lt_one (tsub_le_iff_left.mp hx), + simpa only [← two_mul, zero_lt_mul_left, zero_lt_two] using hx, }, + rw [if_neg (not_lt.mpr hx), if_neg (not_lt.mpr hx), abs_sub_comm, ceil_sub_self_eq this.ne.symm, + abs_one_sub_fract], }, +end + +lemma abs_sub_round_le_abs_self (x : α) : |x - round x| ≤ |x| := +begin + rw [abs_sub_round_eq_min, min_le_iff], + rcases le_or_gt 0 x with hx | (hx : x < 0); [left, right], + { conv_rhs { rw [abs_eq_self.mpr hx, ← fract_add_floor x], }, + simpa only [le_add_iff_nonneg_right, cast_nonneg] using floor_nonneg.mpr hx, }, + { rw abs_eq_neg_self.mpr hx.le, + conv_rhs { rw ← fract_add_floor x, }, + simp only [neg_add_rev, le_add_neg_iff_add_le, sub_add_cancel], + norm_cast, + exact (le_neg.mp $ floor_le_neg_one_iff.mpr hx), }, +end + +end linear_ordered_ring + +section linear_ordered_field + +variables [linear_ordered_field α] [floor_ring α] + +lemma round_eq (x : α) : round x = ⌊x + 1 / 2⌋ := +begin + simp_rw [round, (by simp only [lt_div_iff', two_pos] : 2 * fract x < 1 ↔ fract x < 1 / 2)], + cases lt_or_ge (fract x) (1 / 2) with hx hx, + { conv_rhs { rw [← fract_add_floor x, add_assoc, add_left_comm, floor_int_add], }, + rw [if_pos hx, self_eq_add_right, floor_eq_iff, cast_zero, zero_add], + split; linarith [fract_nonneg x], }, + { have : ⌊fract x + 1 / 2⌋ = 1, { rw floor_eq_iff, split; norm_num; linarith [fract_lt_one x], }, + rw [if_neg (not_lt.mpr hx), ← fract_add_floor x, add_assoc, add_left_comm, floor_int_add, + ceil_add_int, add_comm _ ⌊x⌋, add_right_inj, ceil_eq_iff, this, cast_one, sub_self], + split; linarith [fract_lt_one x], }, +end lemma abs_sub_round (x : α) : |x - round x| ≤ 1 / 2 := begin - rw [round, abs_sub_le_iff], + rw [round_eq, abs_sub_le_iff], have := floor_le (x + 1 / 2), have := lt_floor_add_one (x + 1 / 2), split; linarith end +end linear_ordered_field + end round namespace nat @@ -787,7 +881,7 @@ variables [linear_ordered_field α] [linear_ordered_field β] [floor_ring α] [f include β lemma map_round (f : F) (hf : strict_mono f) (a : α) : round (f a) = round a := -by simp_rw [round, ←map_floor _ hf, map_add, one_div, map_inv₀, map_bit0, map_one] +by simp_rw [round_eq, ←map_floor _ hf, map_add, one_div, map_inv₀, map_bit0, map_one] end int diff --git a/src/data/int/succ_pred.lean b/src/data/int/succ_pred.lean index 1e8cf596bda21..7f67817ce0480 100644 --- a/src/data/int/succ_pred.lean +++ b/src/data/int/succ_pred.lean @@ -32,6 +32,8 @@ instance : pred_order ℤ := @[simp] lemma succ_eq_succ : order.succ = succ := rfl @[simp] lemma pred_eq_pred : order.pred = pred := rfl +lemma pos_iff_one_le {a : ℤ} : 0 < a ↔ 1 ≤ a := order.succ_le_iff.symm + lemma succ_iterate (a : ℤ) : ∀ n, succ^[n] a = a + n | 0 := (add_zero a).symm | (n + 1) := by { rw [function.iterate_succ', int.coe_nat_succ, ←add_assoc], diff --git a/src/data/rat/floor.lean b/src/data/rat/floor.lean index 4882ef3d0361e..4b76ae26228f2 100644 --- a/src/data/rat/floor.lean +++ b/src/data/rat/floor.lean @@ -65,7 +65,7 @@ by rw [←neg_inj, ←floor_neg, ←floor_neg, ← rat.cast_neg, rat.floor_cast] @[simp, norm_cast] lemma round_cast (x : ℚ) : round (x : α) = round x := have ((x + 1 / 2 : ℚ) : α) = x + 1 / 2, by simp, -by rw [round, round, ← this, floor_cast] +by rw [round_eq, round_eq, ← this, floor_cast] @[simp, norm_cast] lemma cast_fract (x : ℚ) : (↑(fract x) : α) = fract x := by simp only [fract, cast_sub, cast_coe_int, floor_cast] diff --git a/src/order/basic.lean b/src/order/basic.lean index 69560d816f3b0..1e4188049bb58 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -534,6 +534,18 @@ lemma max_rec (hx : y ≤ x → p x) (hy : x ≤ y → p y) : p (max x y) := @mi lemma min_rec' (p : α → Prop) (hx : p x) (hy : p y) : p (min x y) := min_rec (λ _, hx) (λ _, hy) lemma max_rec' (p : α → Prop) (hx : p x) (hy : p y) : p (max x y) := max_rec (λ _, hx) (λ _, hy) +lemma min_def' (x y : α) : min x y = if x < y then x else y := +begin + rw [min_comm, min_def, ← ite_not], + simp only [not_le], +end + +lemma max_def' (x y : α) : max x y = if y < x then x else y := +begin + rw [max_comm, max_def, ← ite_not], + simp only [not_le], +end + end min_max_rec /-! ### `has_sup` and `has_inf` -/ diff --git a/src/topology/algebra/order/floor.lean b/src/topology/algebra/order/floor.lean index 26adfe819ab00..c6c5cef0e8fde 100644 --- a/src/topology/algebra/order/floor.lean +++ b/src/topology/algebra/order/floor.lean @@ -29,16 +29,18 @@ open_locale topological_space variables {α β γ : Type*} [linear_ordered_ring α] [floor_ring α] lemma tendsto_floor_at_top : tendsto (floor : α → ℤ) at_top at_top := -floor_mono.tendsto_at_top_at_top $ λ b, ⟨(b + 1 : ℤ), by { rw floor_coe, exact (lt_add_one _).le }⟩ +floor_mono.tendsto_at_top_at_top $ λ b, ⟨(b + 1 : ℤ), + by { rw floor_int_cast, exact (lt_add_one _).le }⟩ lemma tendsto_floor_at_bot : tendsto (floor : α → ℤ) at_bot at_bot := -floor_mono.tendsto_at_bot_at_bot $ λ b, ⟨b, (floor_coe _).le⟩ +floor_mono.tendsto_at_bot_at_bot $ λ b, ⟨b, (floor_int_cast _).le⟩ lemma tendsto_ceil_at_top : tendsto (ceil : α → ℤ) at_top at_top := -ceil_mono.tendsto_at_top_at_top $ λ b, ⟨b, (ceil_coe _).ge⟩ +ceil_mono.tendsto_at_top_at_top $ λ b, ⟨b, (ceil_int_cast _).ge⟩ lemma tendsto_ceil_at_bot : tendsto (ceil : α → ℤ) at_bot at_bot := -ceil_mono.tendsto_at_bot_at_bot $ λ b, ⟨(b - 1 : ℤ), by { rw ceil_coe, exact (sub_one_lt _).le }⟩ +ceil_mono.tendsto_at_bot_at_bot $ λ b, ⟨(b - 1 : ℤ), + by { rw ceil_int_cast, exact (sub_one_lt _).le }⟩ variables [topological_space α] @@ -52,7 +54,7 @@ lemma tendsto_floor_right' [order_closed_topology α] (n : ℤ) : tendsto (λ x, floor x : α → α) (𝓝[≥] n) (𝓝 n) := begin rw ← nhds_within_Ico_eq_nhds_within_Ici (lt_add_one (n : α)), - simpa only [floor_coe] using + simpa only [floor_int_cast] using (continuous_on_floor n _ (left_mem_Ico.mpr $ lt_add_one (_ : α))).tendsto end @@ -60,7 +62,7 @@ lemma tendsto_ceil_left' [order_closed_topology α] (n : ℤ) : tendsto (λ x, ceil x : α → α) (𝓝[≤] n) (𝓝 n) := begin rw ← nhds_within_Ioc_eq_nhds_within_Iic (sub_one_lt (n : α)), - simpa only [ceil_coe] using + simpa only [ceil_int_cast] using (continuous_on_ceil _ _ (right_mem_Ioc.mpr $ sub_one_lt (_ : α))).tendsto end @@ -173,7 +175,7 @@ begin exact ⟨trivial, lt_or_le p.2 _⟩ }, refine continuous_within_at.mono _ this, refine continuous_within_at.union _ _, - { simp only [continuous_within_at, fract_coe, nhds_within_prod_eq, + { simp only [continuous_within_at, fract_int_cast, nhds_within_prod_eq, nhds_within_univ, id.def, comp_app, prod.map_mk], have : (uncurry f) (s, 0) = (uncurry f) (s, (1 : α)), by simp [uncurry, hf], @@ -183,7 +185,7 @@ begin rw nhds_within_Icc_eq_nhds_within_Iic (@zero_lt_one α _ _), exact tendsto_id.prod_map (tendsto_nhds_within_mono_right Iio_subset_Iic_self $ tendsto_fract_left _) }, - { simp only [continuous_within_at, fract_coe, nhds_within_prod_eq, + { simp only [continuous_within_at, fract_int_cast, nhds_within_prod_eq, nhds_within_univ, id.def, comp_app, prod.map_mk], refine (h _ ⟨⟨⟩, by exact_mod_cast left_mem_Icc.2 (zero_le_one' α)⟩).tendsto.comp _, rw [nhds_within_prod_eq, nhds_within_univ, From 3b0a50749fa89847de5d8ffcc6a9cd900f939958 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Fri, 26 Aug 2022 04:46:45 +0000 Subject: [PATCH 046/828] chore(scripts): update nolints.txt (#16253) I am happy to remove some nolints for you! --- scripts/nolints.txt | 4 ---- 1 file changed, 4 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 1cf093edd0e28..2beebcbc31830 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -658,10 +658,6 @@ apply_nolint is_cyclotomic_extension.number_field fintype_finite -- order/prime_ideal.lean apply_nolint order.ideal.is_prime.is_maximal fails_quickly --- order/well_founded_set.lean -apply_nolint pi.is_pwo fintype_finite -apply_nolint set.fintype.is_pwo fintype_finite - -- ring_theory/finiteness.lean apply_nolint algebra.finite_presentation.mv_polynomial fintype_finite apply_nolint algebra.finite_presentation.mv_polynomial_of_finite_presentation fintype_finite From 15c8fddadfb214a58fc40e60b42606bf17f23b87 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Fri, 26 Aug 2022 07:00:48 +0000 Subject: [PATCH 047/828] feat(data/{list, multiset}/basic): generalize `pmap_congr` (#16220) --- src/data/list/basic.lean | 10 +++++++--- src/data/multiset/basic.lean | 8 ++++---- 2 files changed, 11 insertions(+), 7 deletions(-) diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index de9795469bc8a..7f155655805e0 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -2674,9 +2674,13 @@ end by induction l; [refl, simp only [*, pmap, map]]; split; refl theorem pmap_congr {p q : α → Prop} {f : Π a, p a → β} {g : Π a, q a → β} - (l : list α) {H₁ H₂} (h : ∀ a h₁ h₂, f a h₁ = g a h₂) : + (l : list α) {H₁ H₂} (h : ∀ (a ∈ l) h₁ h₂, f a h₁ = g a h₂) : pmap f l H₁ = pmap g l H₂ := -by induction l with _ _ ih; [refl, rw [pmap, pmap, h, ih]] +begin + induction l with _ _ ih, + { refl, }, + { rw [pmap, pmap, h _ (mem_cons_self _ _), ih (λ a ha, h a (mem_cons_of_mem _ ha))], }, +end theorem map_pmap {p : α → Prop} (g : β → γ) (f : Π a, p a → β) (l H) : map g (pmap f l H) = pmap (λ a h, g (f a h)) l H := @@ -2688,7 +2692,7 @@ by induction l; [refl, simp only [*, pmap, map]]; split; refl theorem pmap_eq_map_attach {p : α → Prop} (f : Π a, p a → β) (l H) : pmap f l H = l.attach.map (λ x, f x.1 (H _ x.2)) := -by rw [attach, map_pmap]; exact pmap_congr l (λ a h₁ h₂, rfl) +by rw [attach, map_pmap]; exact pmap_congr l (λ _ _ _ _, rfl) theorem attach_map_val (l : list α) : l.attach.map subtype.val = l := by rw [attach, map_pmap]; exact (pmap_eq_map _ _ _ _).trans (map_id l) diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 8ae596e41fd72..54ef913df33d1 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -1065,9 +1065,9 @@ theorem pmap_eq_map (p : α → Prop) (f : α → β) (s : multiset α) : quot.induction_on s $ λ l H, congr_arg coe $ pmap_eq_map p f l H theorem pmap_congr {p q : α → Prop} {f : Π a, p a → β} {g : Π a, q a → β} - (s : multiset α) {H₁ H₂} (h : ∀ a h₁ h₂, f a h₁ = g a h₂) : - pmap f s H₁ = pmap g s H₂ := -quot.induction_on s (λ l H₁ H₂, congr_arg coe $ pmap_congr l h) H₁ H₂ + (s : multiset α) {H₁ H₂} : + (∀ (a ∈ s) h₁ h₂, f a h₁ = g a h₂) → pmap f s H₁ = pmap g s H₂ := +quot.induction_on s (λ l H₁ H₂ h, congr_arg coe $ pmap_congr l h) H₁ H₂ theorem map_pmap {p : α → Prop} (g : β → γ) (f : Π a, p a → β) (s) : ∀ H, map g (pmap f s H) = pmap (λ a h, g (f a h)) s H := @@ -1098,7 +1098,7 @@ quot.induction_on s (λ l H, length_pmap) H lemma attach_cons (a : α) (m : multiset α) : (a ::ₘ m).attach = ⟨a, mem_cons_self a m⟩ ::ₘ (m.attach.map $ λp, ⟨p.1, mem_cons_of_mem p.2⟩) := quotient.induction_on m $ assume l, congr_arg coe $ congr_arg (list.cons _) $ - by rw [list.map_pmap]; exact list.pmap_congr _ (assume a' h₁ h₂, subtype.eq rfl) + by rw [list.map_pmap]; exact list.pmap_congr _ (λ _ _ _ _, subtype.eq rfl) section decidable_pi_exists variables {m : multiset α} From 5ef6426143338092cd126421361e6cb22be01913 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 26 Aug 2022 20:25:38 +0000 Subject: [PATCH 048/828] =?UTF-8?q?feat(analysis/normed=5Fspace/star/spect?= =?UTF-8?q?rum):=20star=20algebra=20morphisms=20over=20=E2=84=82=20are=20n?= =?UTF-8?q?orm=20contractive=20(#16219)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - [x] depends on: #16177 - [x] depends on: #16212 --- src/analysis/normed_space/star/spectrum.lean | 37 ++++++++++++++++++++ 1 file changed, 37 insertions(+) diff --git a/src/analysis/normed_space/star/spectrum.lean b/src/analysis/normed_space/star/spectrum.lean index 3bf43cdea1cc4..69938599375c8 100644 --- a/src/analysis/normed_space/star/spectrum.lean +++ b/src/analysis/normed_space/star/spectrum.lean @@ -7,6 +7,7 @@ import analysis.normed_space.star.basic import analysis.normed_space.spectrum import algebra.star.module import analysis.normed_space.star.exponential +import algebra.star.star_alg_hom /-! # Spectral properties in C⋆-algebras In this file, we establish various propreties related to the spectrum of elements in C⋆-algebras. @@ -114,3 +115,39 @@ theorem self_adjoint.coe_re_map_spectrum [star_module ℂ A] [nontrivial A] (a : a.property.coe_re_map_spectrum end complex_scalars + +namespace star_alg_hom + +variables {F A B : Type*} +[normed_ring A] [normed_algebra ℂ A] [norm_one_class A] +[complete_space A] [star_ring A] [cstar_ring A] +[normed_ring B] [normed_algebra ℂ B] [norm_one_class B] +[complete_space B] [star_ring B] [cstar_ring B] +[hF : star_alg_hom_class F ℂ A B] (φ : F) +include hF + +/-- A star algebra homomorphism of complex C⋆-algebras is norm contractive. -/ +lemma nnnorm_apply_le (a : A) : ∥(φ a : B)∥₊ ≤ ∥a∥₊ := +begin + suffices : ∀ s : A, is_self_adjoint s → ∥φ s∥₊ ≤ ∥s∥₊, + { exact nonneg_le_nonneg_of_sq_le_sq zero_le' + (by simpa only [nnnorm_star_mul_self, map_star, map_mul] + using this _ (is_self_adjoint.star_mul_self a)) }, + { intros s hs, + simpa only [hs.spectral_radius_eq_nnnorm, (hs.star_hom_apply φ).spectral_radius_eq_nnnorm, + coe_le_coe] using (show spectral_radius ℂ (φ s) ≤ spectral_radius ℂ s, + from supr_le_supr_of_subset (alg_hom.spectrum_apply_subset φ s)) } +end + +/-- A star algebra homomorphism of complex C⋆-algebras is norm contractive. -/ +lemma norm_apply_le (a : A) : ∥(φ a : B)∥ ≤ ∥a∥ := nnnorm_apply_le φ a + +/-- Star algebra homomorphisms between C⋆-algebras are continuous linear maps. +See note [lower instance priority] -/ +@[priority 100] +noncomputable instance : continuous_linear_map_class F ℂ A B := +{ map_continuous := λ φ, add_monoid_hom_class.continuous_of_bound φ 1 + (by simpa only [one_mul] using nnnorm_apply_le φ), + .. alg_hom_class.linear_map_class } + +end star_alg_hom From 18361694c2341d7e9f80bed880a3efa15417c4ad Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 27 Aug 2022 00:58:35 +0000 Subject: [PATCH 049/828] feat(ring_theory/ideal): Two little `restrict_scalars` lemmas (#16248) I used these two lemmas in one of the definitions of the ideal norm. --- src/ring_theory/ideal/operations.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 2feafe45fcf7b..db1d940fa911a 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -1138,6 +1138,19 @@ begin rw smul_add, exact submodule.add_mem _ hx hy }, end +@[simp] lemma coe_restrict_scalars {R S : Type*} [comm_semiring R] [semiring S] [algebra R S] + (I : ideal S) : ((I.restrict_scalars R) : set S) = ↑I := +rfl + +/-- The smallest `S`-submodule that contains all `x ∈ I * y ∈ J` +is also the smallest `R`-submodule that does so. -/ +@[simp] lemma restrict_scalars_mul {R S : Type*} [comm_semiring R] [comm_semiring S] [algebra R S] + (I J : ideal S) : (I * J).restrict_scalars R = I.restrict_scalars R * J.restrict_scalars R := +le_antisymm (λ x hx, submodule.mul_induction_on hx + (λ x hx y hy, submodule.mul_mem_mul hx hy) + (λ x y, submodule.add_mem _)) + (submodule.mul_le.mpr (λ x hx y hy, ideal.mul_mem_mul hx hy)) + section surjective variables (hf : function.surjective f) include hf From c95908a611c41be7cb172868907dd120dbd4f6a1 Mon Sep 17 00:00:00 2001 From: kkytola Date: Sat, 27 Aug 2022 13:27:26 +0000 Subject: [PATCH 050/828] feat(measure_theory/measure/finite_measure_weak_convergence): Add portmanteau implications concerning open and closed sets. (#15321) Add equivalence of a limsup condition for closed sets and a liminf condition for open sets, which will later both be shown to be characterizing conditions of weak convergence of probability measures. Also add that either of these equivalent conditions implies that for any Borel set whose boundary carries zero probability in the candidate limit measure, we have that the measures of this set tend to its measure under the candidate limit measure. Co-authored-by: kkytola <39528102+kkytola@users.noreply.github.com> --- src/algebra/order/sub.lean | 3 + .../finite_measure_weak_convergence.lean | 261 +++++++++++++++++- 2 files changed, 261 insertions(+), 3 deletions(-) diff --git a/src/algebra/order/sub.lean b/src/algebra/order/sub.lean index d426fac5f8b78..bf808bdee5cea 100644 --- a/src/algebra/order/sub.lean +++ b/src/algebra/order/sub.lean @@ -131,6 +131,9 @@ tsub_le_iff_left.mpr $ le_add_tsub.trans $ add_le_add_right h _ lemma tsub_le_tsub (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c := (tsub_le_tsub_right hab _).trans $ tsub_le_tsub_left hcd _ +lemma antitone_const_tsub : antitone (λ x, c - x) := +λ x y hxy, tsub_le_tsub rfl.le hxy + /-- See `add_tsub_assoc_of_le` for the equality. -/ lemma add_tsub_le_assoc : a + b - c ≤ a + (b - c) := by { rw [tsub_le_iff_left, add_left_comm], exact add_le_add_left le_add_tsub a } diff --git a/src/measure_theory/measure/finite_measure_weak_convergence.lean b/src/measure_theory/measure/finite_measure_weak_convergence.lean index aed6f9a293539..7f16b36478e8a 100644 --- a/src/measure_theory/measure/finite_measure_weak_convergence.lean +++ b/src/measure_theory/measure/finite_measure_weak_convergence.lean @@ -60,7 +60,21 @@ TODO: for example metrizability or pseudo-emetrizability would be sufficient assumptions. The typeclass assumptions should be later adjusted in a way that takes into account use cases, but the proof will presumably remain essentially the same. + * `measure_theory.limsup_measure_closed_le_iff_liminf_measure_open_ge` proves the equivalence of + the limsup condition for closed sets and the liminf condition for open sets for probability + measures. + * `measure_theory.tendsto_measure_of_null_frontier` proves that the liminf condition for open + sets (which is equivalent to the limsup condition for closed sets) implies the convergence of + probabilities of sets whose boundary carries no mass under the limit measure. + * `measure_theory.probability_measure.tendsto_measure_of_null_frontier_of_tendsto` is a + combination of earlier implications, which shows that weak convergence of probability measures + implies the convergence of probabilities of sets whose boundary carries no mass + under the limit measure. * Prove the rest of the implications. + (Where formulations are currently only provided for probability measures, one can obtain the + finite measure formulations using the characterization of convergence of finite measures by + their total masses and their probability-normalized versions, i.e., by + `measure_theory.finite_measure.tendsto_normalize_iff_tendsto`.) ## Notations @@ -488,13 +502,13 @@ This formulation assumes: -/ lemma tendsto_lintegral_nn_filter_of_le_const {ι : Type*} {L : filter ι} [L.is_countably_generated] (μ : measure α) [is_finite_measure μ] {fs : ι → (α →ᵇ ℝ≥0)} {c : ℝ≥0} - (fs_le_const : ∀ᶠ i in L, ∀ᵐ (a : α) ∂(μ : measure α), fs i a ≤ c) {f : α → ℝ≥0} - (fs_lim : ∀ᵐ (a : α) ∂(μ : measure α), tendsto (λ i, fs i a) L (𝓝 (f a))) : + (fs_le_const : ∀ᶠ i in L, ∀ᵐ (a : α) ∂μ, fs i a ≤ c) {f : α → ℝ≥0} + (fs_lim : ∀ᵐ (a : α) ∂μ, tendsto (λ i, fs i a) L (𝓝 (f a))) : tendsto (λ i, (∫⁻ a, fs i a ∂μ)) L (𝓝 (∫⁻ a, (f a) ∂μ)) := begin simpa only using tendsto_lintegral_filter_of_dominated_convergence (λ _, c) (eventually_of_forall ((λ i, (ennreal.continuous_coe.comp (fs i).continuous).measurable))) - _ ((@lintegral_const_lt_top _ _ (μ : measure α) _ _ (@ennreal.coe_ne_top c)).ne) _, + _ ((@lintegral_const_lt_top _ _ μ _ _ (@ennreal.coe_ne_top c)).ne) _, { simpa only [ennreal.coe_le_coe] using fs_le_const, }, { simpa only [ennreal.tendsto_coe] using fs_lim, }, end @@ -1026,6 +1040,179 @@ end finite_measure --namespace end normalize_finite_measure -- section +section limsup_closed_le_and_le_liminf_open +/-! ### Portmanteau: limsup condition for closed sets iff liminf condition for open sets + +In this section we prove that for a sequence of Borel probability measures on a topological space +and its candidate limit measure, the following two conditions are equivalent: + (C) For any closed set `F` in `α` the limsup of the measures of `F` is at most the limit + measure of `F`. + (O) For any open set `G` in `α` the liminf of the measures of `G` is at least the limit + measure of `G`. +Either of these will later be shown to be equivalent to the weak convergence of the sequence +of measures. +-/ + +variables {α : Type*} [measurable_space α] + +lemma le_measure_compl_liminf_of_limsup_measure_le + {ι : Type*} {L : filter ι} {μ : measure α} {μs : ι → measure α} + [is_probability_measure μ] [∀ i, is_probability_measure (μs i)] + {E : set α} (E_mble : measurable_set E) (h : L.limsup (λ i, μs i E) ≤ μ E) : + μ Eᶜ ≤ L.liminf (λ i, μs i Eᶜ) := +begin + by_cases L_bot : L = ⊥, + { simp only [L_bot, le_top, + (show liminf ⊥ (λ i, μs i Eᶜ) = ⊤, by simp only [liminf, filter.map_bot, Liminf_bot])], }, + haveI : L.ne_bot, from {ne' := L_bot}, + have meas_Ec : μ Eᶜ = 1 - μ E, + { simpa only [measure_univ] using measure_compl E_mble (measure_lt_top μ E).ne, }, + have meas_i_Ec : ∀ i, μs i Eᶜ = 1 - μs i E, + { intro i, + simpa only [measure_univ] using measure_compl E_mble (measure_lt_top (μs i) E).ne, }, + simp_rw [meas_Ec, meas_i_Ec], + have obs : L.liminf (λ (i : ι), 1 - μs i E) = L.liminf ((λ x, 1 - x) ∘ (λ (i : ι), μs i E)), + by refl, + rw obs, + simp_rw ← antitone_const_tsub.map_limsup_of_continuous_at (λ i, μs i E) + (ennreal.continuous_sub_left ennreal.one_ne_top).continuous_at, + exact antitone_const_tsub h, +end + +lemma le_measure_liminf_of_limsup_measure_compl_le + {ι : Type*} {L : filter ι} {μ : measure α} {μs : ι → measure α} + [is_probability_measure μ] [∀ i, is_probability_measure (μs i)] + {E : set α} (E_mble : measurable_set E) (h : L.limsup (λ i, μs i Eᶜ) ≤ μ Eᶜ) : + μ E ≤ L.liminf (λ i, μs i E) := +compl_compl E ▸ (le_measure_compl_liminf_of_limsup_measure_le (measurable_set.compl E_mble) h) + +lemma limsup_measure_compl_le_of_le_liminf_measure + {ι : Type*} {L : filter ι} {μ : measure α} {μs : ι → measure α} + [is_probability_measure μ] [∀ i, is_probability_measure (μs i)] + {E : set α} (E_mble : measurable_set E) (h : μ E ≤ L.liminf (λ i, μs i E)) : + L.limsup (λ i, μs i Eᶜ) ≤ μ Eᶜ := +begin + by_cases L_bot : L = ⊥, + { simp only [L_bot, bot_le, + (show limsup ⊥ (λ i, μs i Eᶜ) = ⊥, by simp only [limsup, filter.map_bot, Limsup_bot])], }, + haveI : L.ne_bot, from {ne' := L_bot}, + have meas_Ec : μ Eᶜ = 1 - μ E, + { simpa only [measure_univ] using measure_compl E_mble (measure_lt_top μ E).ne, }, + have meas_i_Ec : ∀ i, μs i Eᶜ = 1 - μs i E, + { intro i, + simpa only [measure_univ] using measure_compl E_mble (measure_lt_top (μs i) E).ne, }, + simp_rw [meas_Ec, meas_i_Ec], + have obs : L.limsup (λ (i : ι), 1 - μs i E) = L.limsup ((λ x, 1 - x) ∘ (λ (i : ι), μs i E)), + by refl, + rw obs, + simp_rw ← antitone_const_tsub.map_liminf_of_continuous_at (λ i, μs i E) + (ennreal.continuous_sub_left ennreal.one_ne_top).continuous_at, + exact antitone_const_tsub h, +end + +lemma limsup_measure_le_of_le_liminf_measure_compl + {ι : Type*} {L : filter ι} {μ : measure α} {μs : ι → measure α} + [is_probability_measure μ] [∀ i, is_probability_measure (μs i)] + {E : set α} (E_mble : measurable_set E) (h : μ Eᶜ ≤ L.liminf (λ i, μs i Eᶜ)) : + L.limsup (λ i, μs i E) ≤ μ E := +compl_compl E ▸ (limsup_measure_compl_le_of_le_liminf_measure (measurable_set.compl E_mble) h) + +variables [topological_space α] [opens_measurable_space α] + +/-- One pair of implications of the portmanteau theorem: +For a sequence of Borel probability measures, the following two are equivalent: + +(C) The limsup of the measures of any closed set is at most the measure of the closed set +under a candidate limit measure. + +(O) The liminf of the measures of any open set is at least the measure of the open set +under a candidate limit measure. +-/ +lemma limsup_measure_closed_le_iff_liminf_measure_open_ge + {ι : Type*} {L : filter ι} {μ : measure α} {μs : ι → measure α} + [is_probability_measure μ] [∀ i, is_probability_measure (μs i)] : + (∀ F, is_closed F → L.limsup (λ i, μs i F) ≤ μ F) + ↔ (∀ G, is_open G → μ G ≤ L.liminf (λ i, μs i G)) := +begin + split, + { intros h G G_open, + exact le_measure_liminf_of_limsup_measure_compl_le + G_open.measurable_set (h Gᶜ (is_closed_compl_iff.mpr G_open)), }, + { intros h F F_closed, + exact limsup_measure_le_of_le_liminf_measure_compl + F_closed.measurable_set (h Fᶜ (is_open_compl_iff.mpr F_closed)), }, +end + +end limsup_closed_le_and_le_liminf_open -- section + +section tendsto_of_null_frontier +/-! ### Portmanteau: limit of measures of Borel sets whose boundary carries no mass in the limit + +In this section we prove that for a sequence of Borel probability measures on a topological space +and its candidate limit measure, either of the following equivalent conditions: + (C) For any closed set `F` in `α` the limsup of the measures of `F` is at most the limit + measure of `F` + (O) For any open set `G` in `α` the liminf of the measures of `G` is at least the limit + measure of `G` +implies that + (B) For any Borel set `E` in `α` whose boundary `∂E` carries no mass under the candidate limit + measure, we have that the limit of measures of `E` is the measure of `E` under the + candidate limit measure. +-/ + +variables {α : Type*} [measurable_space α] + +lemma tendsto_measure_of_le_liminf_measure_of_limsup_measure_le + {ι : Type*} {L : filter ι} {μ : measure α} {μs : ι → measure α} + {E₀ E E₁ : set α} (E₀_subset : E₀ ⊆ E) (subset_E₁ : E ⊆ E₁) (nulldiff : μ (E₁ \ E₀) = 0) + (h_E₀ : μ E₀ ≤ L.liminf (λ i, μs i E₀)) (h_E₁ : L.limsup (λ i, μs i E₁) ≤ μ E₁) : + L.tendsto (λ i, μs i E) (𝓝 (μ E)) := +begin + apply tendsto_of_le_liminf_of_limsup_le, + { have E₀_ae_eq_E : E₀ =ᵐ[μ] E, + from eventually_le.antisymm E₀_subset.eventually_le + (subset_E₁.eventually_le.trans (ae_le_set.mpr nulldiff)), + calc μ(E) + = μ(E₀) : measure_congr E₀_ae_eq_E.symm + ... ≤ L.liminf (λ i, μs i E₀) : h_E₀ + ... ≤ L.liminf (λ i, μs i E) : _, + { refine liminf_le_liminf (eventually_of_forall (λ _, measure_mono E₀_subset)) _, + apply_auto_param, }, }, + { have E_ae_eq_E₁ : E =ᵐ[μ] E₁, + from eventually_le.antisymm subset_E₁.eventually_le + ((ae_le_set.mpr nulldiff).trans E₀_subset.eventually_le), + calc L.limsup (λ i, μs i E) + ≤ L.limsup (λ i, μs i E₁) : _ + ... ≤ μ E₁ : h_E₁ + ... = μ E : measure_congr E_ae_eq_E₁.symm, + { refine limsup_le_limsup (eventually_of_forall (λ _, measure_mono subset_E₁)) _, + apply_auto_param, }, }, +end + +variables [topological_space α] [opens_measurable_space α] + +/-- One implication of the portmanteau theorem: +For a sequence of Borel probability measures, if the liminf of the measures of any open set is at +least the measure of the open set under a candidate limit measure, then for any set whose +boundary carries no probability mass under the candidate limit measure, then its measures under the +sequence converge to its measure under the candidate limit measure. +-/ +lemma tendsto_measure_of_null_frontier + {ι : Type*} {L : filter ι} {μ : measure α} {μs : ι → measure α} + [is_probability_measure μ] [∀ i, is_probability_measure (μs i)] + (h_opens : ∀ G, is_open G → μ G ≤ L.liminf (λ i, μs i G)) + {E : set α} (E_nullbdry : μ (frontier E) = 0) : + L.tendsto (λ i, μs i E) (𝓝 (μ E)) := +begin + have h_closeds : ∀ F, is_closed F → L.limsup (λ i, μs i F) ≤ μ F, + from limsup_measure_closed_le_iff_liminf_measure_open_ge.mpr h_opens, + exact tendsto_measure_of_le_liminf_measure_of_limsup_measure_le + interior_subset subset_closure E_nullbdry + (h_opens _ is_open_interior) (h_closeds _ is_closed_closure), +end + +end tendsto_of_null_frontier --section + section convergence_implies_limsup_closed_le /-! ### Portmanteau implication: weak convergence implies a limsup condition for closed sets @@ -1034,6 +1221,10 @@ pseudo-emetrizable, that the weak convergence of measures on `measure_theory.fin implies that for any closed set `F` in `α` the limsup of the measures of `F` is at most the limit measure of `F`. This is one implication of the portmanteau theorem characterizing weak convergence of measures. + +Combining with an earlier implication we also get that weak convergence implies that for any Borel +set `E` in `α` whose boundary `∂E` carries no mass under the limit measure, the limit of measures +of `E` is the measure of `E` under the limit measure. -/ variables {α : Type*} [measurable_space α] @@ -1144,6 +1335,70 @@ begin simp only [add_assoc, ennreal.add_halves, le_refl], end +/-- One implication of the portmanteau theorem: +Weak convergence of probability measures implies that the limsup of the measures of any closed +set is at most the measure of the closed set under the limit probability measure. +-/ +lemma probability_measure.limsup_measure_closed_le_of_tendsto + {α ι : Type*} {L : filter ι} + [measurable_space α] [pseudo_emetric_space α] [opens_measurable_space α] + {μ : probability_measure α} {μs : ι → probability_measure α} + (μs_lim : tendsto μs L (𝓝 μ)) {F : set α} (F_closed : is_closed F) : + L.limsup (λ i, (μs i : measure α) F) ≤ (μ : measure α) F := +by apply finite_measure.limsup_measure_closed_le_of_tendsto + ((probability_measure.tendsto_nhds_iff_to_finite_measures_tendsto_nhds L).mp μs_lim) + F_closed + +/-- One implication of the portmanteau theorem: +Weak convergence of probability measures implies that the liminf of the measures of any open set +is at least the measure of the open set under the limit probability measure. +-/ +lemma probability_measure.le_liminf_measure_open_of_tendsto + {α ι : Type*} {L : filter ι} + [measurable_space α] [pseudo_emetric_space α] [opens_measurable_space α] + {μ : probability_measure α} {μs : ι → probability_measure α} + (μs_lim : tendsto μs L (𝓝 μ)) {G : set α} (G_open : is_open G) : + (μ : measure α) G ≤ L.liminf (λ i, (μs i : measure α) G) := +begin + have h_closeds : ∀ F, is_closed F → L.limsup (λ i, (μs i : measure α) F) ≤ (μ : measure α) F, + from λ F F_closed, probability_measure.limsup_measure_closed_le_of_tendsto μs_lim F_closed, + exact le_measure_liminf_of_limsup_measure_compl_le + G_open.measurable_set (h_closeds _ (is_closed_compl_iff.mpr G_open)), +end + +lemma probability_measure.tendsto_measure_of_null_frontier_of_tendsto' + {α ι : Type*} {L : filter ι} + [measurable_space α] [pseudo_emetric_space α] [opens_measurable_space α] + {μ : probability_measure α} {μs : ι → probability_measure α} + (μs_lim : tendsto μs L (𝓝 μ)) {E : set α} (E_nullbdry : (μ : measure α) (frontier E) = 0) : + tendsto (λ i, (μs i : measure α) E) L (𝓝 ((μ : measure α) E)) := +begin + have h_opens : ∀ G, is_open G → (μ : measure α) G ≤ L.liminf (λ i, (μs i : measure α) G), + from λ G G_open, probability_measure.le_liminf_measure_open_of_tendsto μs_lim G_open, + exact tendsto_measure_of_null_frontier h_opens E_nullbdry, +end + +/-- One implication of the portmanteau theorem: +Weak convergence of probability measures implies that if the boundary of a Borel set +carries no probability mass under the limit measure, then the limit of the measures of the set +equals the measure of the set under the limit probability measure. + +A version with coercions to ordinary `ℝ≥0∞`-valued measures is +`measure_theory.probability_measure.tendsto_measure_of_null_frontier_of_tendsto'`. +-/ +lemma probability_measure.tendsto_measure_of_null_frontier_of_tendsto + {α ι : Type*} {L : filter ι} + [measurable_space α] [pseudo_emetric_space α] [opens_measurable_space α] + {μ : probability_measure α} {μs : ι → probability_measure α} + (μs_lim : tendsto μs L (𝓝 μ)) {E : set α} (E_nullbdry : μ (frontier E) = 0) : + tendsto (λ i, μs i E) L (𝓝 (μ E)) := +begin + have E_nullbdry' : (μ : measure α) (frontier E) = 0, + by rw [← probability_measure.ennreal_coe_fn_eq_coe_fn_to_measure, E_nullbdry, ennreal.coe_zero], + have key := probability_measure.tendsto_measure_of_null_frontier_of_tendsto' μs_lim E_nullbdry', + exact (ennreal.tendsto_to_nnreal (measure_ne_top ↑μ E)).comp key, +end + end convergence_implies_limsup_closed_le --section end measure_theory --namespace From fef8efdf78f223294c34a41875923ab1272322d4 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Sat, 27 Aug 2022 16:01:27 +0000 Subject: [PATCH 051/828] feat(category_theory): coseparators in structured_arrow (#15981) Hopefully the last preliminary PR for the Special Adjoint Functor Theorem. --- src/category_theory/generator.lean | 40 ++++++++++++++++++----- src/category_theory/structured_arrow.lean | 17 ++++++++++ 2 files changed, 49 insertions(+), 8 deletions(-) diff --git a/src/category_theory/generator.lean b/src/category_theory/generator.lean index e398a80921da1..aed1af711e8f1 100644 --- a/src/category_theory/generator.lean +++ b/src/category_theory/generator.lean @@ -51,12 +51,12 @@ We -/ -universes w v u +universes w v₁ v₂ u₁ u₂ open category_theory.limits opposite namespace category_theory -variables {C : Type u} [category.{v} C] +variables {C : Type u₁} [category.{v₁} C] {D : Type u₂} [category.{v₂} D] /-- We say that `𝒢` is a separating set if the functors `C(G, -)` for `G ∈ 𝒢` are collectively faithful, i.e., if `h ≫ f = h ≫ g` for all `h` with domain in `𝒢` implies `f = g`. -/ @@ -264,11 +264,11 @@ end category with a small coseparating set has an initial object. In fact, it follows from the Special Adjoint Functor Theorem that `C` is already cocomplete. -/ -lemma has_initial_of_is_cosepatating [well_powered C] [has_limits C] {𝒢 : set C} [small.{v} 𝒢] +lemma has_initial_of_is_coseparating [well_powered C] [has_limits C] {𝒢 : set C} [small.{v₁} 𝒢] (h𝒢 : is_coseparating 𝒢) : has_initial C := begin haveI := has_products_of_shape_of_small C 𝒢, - haveI := λ A, has_products_of_shape_of_small.{v} C (Σ G : 𝒢, A ⟶ (G : C)), + haveI := λ A, has_products_of_shape_of_small.{v₁} C (Σ G : 𝒢, A ⟶ (G : C)), letI := complete_lattice_of_complete_semilattice_Inf (subobject (pi_obj (coe : 𝒢 → C))), suffices : ∀ A : C, unique (((⊥ : subobject (pi_obj (coe : 𝒢 → C))) : C) ⟶ A), { exactI has_initial_of_unique ((⊥ : subobject (pi_obj (coe : 𝒢 → C))) : C) }, @@ -288,12 +288,12 @@ end category with a small separating set has a terminal object. In fact, it follows from the Special Adjoint Functor Theorem that `C` is already complete. -/ -lemma has_terminal_of_is_separating [well_powered Cᵒᵖ] [has_colimits C] {𝒢 : set C} [small.{v} 𝒢] +lemma has_terminal_of_is_separating [well_powered Cᵒᵖ] [has_colimits C] {𝒢 : set C} [small.{v₁} 𝒢] (h𝒢 : is_separating 𝒢) : has_terminal C := begin haveI : has_limits Cᵒᵖ := has_limits_op_of_has_colimits, - haveI : small.{v} 𝒢.op := small_of_injective (set.op_equiv_self 𝒢).injective, - haveI : has_initial Cᵒᵖ := has_initial_of_is_cosepatating ((is_coseparating_op_iff _).2 h𝒢), + haveI : small.{v₁} 𝒢.op := small_of_injective (set.op_equiv_self 𝒢).injective, + haveI : has_initial Cᵒᵖ := has_initial_of_is_coseparating ((is_coseparating_op_iff _).2 h𝒢), exact has_terminal_of_has_initial_op end @@ -327,13 +327,37 @@ calc P = P ⊓ Q : eq.symm $ inf_eq_of_is_detecting h𝒢 _ _ $ λ G hG f hf, (h end subobject /-- A category with pullbacks and a small detecting set is well-powered. -/ -lemma well_powered_of_is_detecting [has_pullbacks C] {𝒢 : set C} [small.{v} 𝒢] +lemma well_powered_of_is_detecting [has_pullbacks C] {𝒢 : set C} [small.{v₁} 𝒢] (h𝒢 : is_detecting 𝒢) : well_powered C := ⟨λ X, @small_of_injective _ _ _ (λ P : subobject X, { f : Σ G : 𝒢, G.1 ⟶ X | P.factors f.2 }) $ λ P Q h, subobject.eq_of_is_detecting h𝒢 _ _ (by simpa [set.ext_iff] using h)⟩ end well_powered +namespace structured_arrow +variables (S : D) (T : C ⥤ D) + +lemma is_coseparating_proj_preimage {𝒢 : set C} (h𝒢 : is_coseparating 𝒢) : + is_coseparating ((proj S T).obj ⁻¹' 𝒢) := +begin + refine λ X Y f g hfg, ext _ _ (h𝒢 _ _ (λ G hG h, _)), + exact congr_arg comma_morphism.right (hfg (mk (Y.hom ≫ T.map h)) hG (hom_mk h rfl)) +end + +end structured_arrow + +namespace costructured_arrow +variables (S : C ⥤ D) (T : D) + +lemma is_separating_proj_preimage {𝒢 : set C} (h𝒢 : is_separating 𝒢) : + is_separating ((proj S T).obj ⁻¹' 𝒢) := +begin + refine λ X Y f g hfg, ext _ _ (h𝒢 _ _ (λ G hG h, _)), + convert congr_arg comma_morphism.left (hfg (mk (S.map h ≫ X.hom)) hG (hom_mk h rfl)) +end + +end costructured_arrow + /-- We say that `G` is a separator if the functor `C(G, -)` is faithful. -/ def is_separator (G : C) : Prop := is_separating ({G} : set C) diff --git a/src/category_theory/structured_arrow.lean b/src/category_theory/structured_arrow.lean index b478571ddaa5c..2b289b795b53b 100644 --- a/src/category_theory/structured_arrow.lean +++ b/src/category_theory/structured_arrow.lean @@ -6,6 +6,7 @@ Authors: Adam Topaz, Scott Morrison import category_theory.punit import category_theory.comma import category_theory.limits.shapes.terminal +import category_theory.essentially_small /-! # The category of "structured arrows" @@ -168,6 +169,14 @@ comma.pre_right _ F G map := λ X Y f, { right := f.right, w' := by { simp [functor.comp_map, ←G.map_comp, ← f.w] } } } +instance small_proj_preimage_of_locally_small {𝒢 : set C} [small.{v₁} 𝒢] [locally_small.{v₁} D] : + small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := +begin + suffices : (proj S T).obj ⁻¹' 𝒢 = set.range (λ f : Σ G : 𝒢, S ⟶ T.obj G, mk f.2), + { rw this, apply_instance }, + exact set.ext (λ X, ⟨λ h, ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by tidy⟩) +end + end structured_arrow @@ -310,6 +319,14 @@ comma.pre_left F G _ map := λ X Y f, { left := f.left, w' := by { simp [functor.comp_map, ←G.map_comp, ← f.w] } } } +instance small_proj_preimage_of_locally_small {𝒢 : set C} [small.{v₁} 𝒢] [locally_small.{v₁} D] : + small.{v₁} ((proj S T).obj ⁻¹' 𝒢) := +begin + suffices : (proj S T).obj ⁻¹' 𝒢 = set.range (λ f : Σ G : 𝒢, S.obj G ⟶ T, mk f.2), + { rw this, apply_instance }, + exact set.ext (λ X, ⟨λ h, ⟨⟨⟨_, h⟩, X.hom⟩, (eq_mk _).symm⟩, by tidy⟩) +end + end costructured_arrow open opposite From c241a033734908a14bf6760e8407e19afd6f4a31 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sat, 27 Aug 2022 16:01:28 +0000 Subject: [PATCH 052/828] docs(ring_theory/valuation/valuation_subring): clarify value group docstring (#16272) The "value group" of a valuation is not a group! It's a group with zero. I've heard the phrase "value group" being used to describe the units of this group with zero before, so we should perhaps clarify in the docstring what we mean by this definition. --- src/ring_theory/valuation/valuation_subring.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index 7840bd4ef1253..dd880990215ec 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -100,7 +100,7 @@ instance : is_fraction_ring A K := eq_iff_exists := λ a b, ⟨ λ h, ⟨1, by { ext, simpa using h }⟩, λ ⟨c, h⟩, congr_arg coe ((mul_eq_mul_right_iff.1 h).resolve_right (non_zero_divisors.ne_zero c.2)) ⟩ } -/-- The value group of the valuation associated to `A`. -/ +/-- The value group of the valuation associated to `A`. Note: it is actually a group with zero. -/ @[derive linear_ordered_comm_group_with_zero] def value_group := valuation_ring.value_group A K From 6901db6025feb157f9d0ade9bd65609373942c59 Mon Sep 17 00:00:00 2001 From: Gabriel Ebner Date: Sat, 27 Aug 2022 18:04:24 +0000 Subject: [PATCH 053/828] chore(category_theory/limits/shapes/finite_products): avoid tidy (#16261) --- src/category_theory/limits/shapes/finite_products.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/category_theory/limits/shapes/finite_products.lean b/src/category_theory/limits/shapes/finite_products.lean index 0e6f79c7a08c2..56e23ddd3b4bf 100644 --- a/src/category_theory/limits/shapes/finite_products.lean +++ b/src/category_theory/limits/shapes/finite_products.lean @@ -47,7 +47,7 @@ instance has_fintype_products [has_finite_products C] (ι : Type w) [fintype ι] has_limits_of_shape (discrete ι) C := has_limits_of_shape_of_equivalence (discrete.equivalence - ((show ulift.{0} (fin (fintype.card ι)) ≃ fin (fintype.card ι), by tidy).trans + (equiv.ulift.{0}.trans (fintype.equiv_fin ι).symm)) /-- We can now write this for powers. -/ @@ -83,7 +83,7 @@ instance has_fintype_coproducts [has_finite_coproducts C] (ι : Type w) [fintype has_colimits_of_shape (discrete ι) C := has_colimits_of_shape_of_equivalence (discrete.equivalence - ((show ulift.{0} (fin (fintype.card ι)) ≃ fin (fintype.card ι), by tidy).trans + (equiv.ulift.{0}.trans (fintype.equiv_fin ι).symm)) /-- From ebb53dc3929b6c25462f59be3e29f82246c6bcda Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Sat, 27 Aug 2022 19:36:00 +0000 Subject: [PATCH 054/828] feat(data/rat/defs): add two lemmas for `pnat_denom` of 0 and 1 (#15864) Co-authored-by: Jon Eugster --- src/data/rat/defs.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/data/rat/defs.lean b/src/data/rat/defs.lean index 1f4b953429037..678af31d05455 100644 --- a/src/data/rat/defs.lean +++ b/src/data/rat/defs.lean @@ -907,6 +907,10 @@ by rw [pnat_denom, mk_pnat_eq, num_denom] lemma pnat_denom_eq_iff_denom_eq {x : ℚ} {n : ℕ+} : x.pnat_denom = n ↔ x.denom = ↑n := subtype.ext_iff +@[simp] lemma pnat_denom_one : (1 : ℚ).pnat_denom = 1 := rfl + +@[simp] lemma pnat_denom_zero : (0 : ℚ).pnat_denom = 1 := rfl + end pnat_denom end rat From 1c2d71f1833b5d5d4b2fe26e2d805047508e23b6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 27 Aug 2022 21:53:53 +0000 Subject: [PATCH 055/828] feat(algebra/order/field): Canonically linear ordered semifields (#15677) Define `canonically_linear_ordered_semifield`. The target is `nnreal` and the future `nnrat`. --- src/algebra/order/field.lean | 8 ++++ src/algebra/order/field_defs.lean | 15 +++++- src/algebra/order/nonneg.lean | 78 ++++++++++++++++++------------- 3 files changed, 67 insertions(+), 34 deletions(-) diff --git a/src/algebra/order/field.lean b/src/algebra/order/field.lean index 927c23a3f67fa..8840c4a17cd63 100644 --- a/src/algebra/order/field.lean +++ b/src/algebra/order/field.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Lewis, Leonardo de Moura, Mario Carneiro, Floris van Doorn -/ import algebra.order.field_defs +import algebra.order.with_zero /-! # Linear ordered (semi)fields @@ -852,3 +853,10 @@ theorem nat.cast_le_pow_div_sub (H : 1 < a) (n : ℕ) : (n : α) ≤ a ^ n / (a (sub_le_self _ zero_le_one) end + +section canonically_linear_ordered_semifield +variables [canonically_linear_ordered_semifield α] [has_sub α] [has_ordered_sub α] + +lemma tsub_div (a b c : α) : (a - b) / c = a / c - b / c := by simp_rw [div_eq_mul_inv, tsub_mul] + +end canonically_linear_ordered_semifield diff --git a/src/algebra/order/field_defs.lean b/src/algebra/order/field_defs.lean index a345e440e108c..c734f15d942f3 100644 --- a/src/algebra/order/field_defs.lean +++ b/src/algebra/order/field_defs.lean @@ -32,6 +32,8 @@ The lemmata are instead located there. set_option old_structure_cmd true +variables {α : Type*} + /-- A linear ordered semifield is a field with a linear order respecting the operations. -/ @[protect_proj] class linear_ordered_semifield (α : Type*) extends linear_ordered_semiring α, semifield α @@ -39,7 +41,18 @@ set_option old_structure_cmd true /-- A linear ordered field is a field with a linear order respecting the operations. -/ @[protect_proj] class linear_ordered_field (α : Type*) extends linear_ordered_comm_ring α, field α +/-- A canonically linear ordered field is a linear ordered field in which `a ≤ b` iff there exists +`c` with `b = a + c`. -/ +@[protect_proj] class canonically_linear_ordered_semifield (α : Type*) + extends canonically_ordered_comm_semiring α, linear_ordered_semifield α + @[priority 100] -- See note [lower instance priority] -instance linear_ordered_field.to_linear_ordered_semifield {α} [linear_ordered_field α] : +instance linear_ordered_field.to_linear_ordered_semifield [linear_ordered_field α] : linear_ordered_semifield α := { ..linear_ordered_ring.to_linear_ordered_semiring, ..‹linear_ordered_field α› } + +@[priority 100] -- See note [lower instance priority] +instance canonically_linear_ordered_semifield.to_linear_ordered_comm_group_with_zero + [canonically_linear_ordered_semifield α] : linear_ordered_comm_group_with_zero α := +{ mul_le_mul_left := λ a b h c, mul_le_mul_of_nonneg_left h $ zero_le _, + ..‹canonically_linear_ordered_semifield α› } diff --git a/src/algebra/order/nonneg.lean b/src/algebra/order/nonneg.lean index aeb62e116e86c..1fe76aca8d83f 100644 --- a/src/algebra/order/nonneg.lean +++ b/src/algebra/order/nonneg.lean @@ -4,9 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ import algebra.order.archimedean -import algebra.order.floor -import algebra.order.sub -import algebra.order.with_zero import order.lattice_intervals import order.complete_lattice_intervals @@ -227,36 +224,6 @@ def coe_ring_hom [ordered_semiring α] : {x : α // 0 ≤ x} →+* α := protected lemma coe_nat_cast [ordered_semiring α] (n : ℕ) : ((↑n : {x : α // 0 ≤ x}) : α) = n := map_nat_cast (coe_ring_hom : {x : α // 0 ≤ x} →+* α) n -instance has_inv [linear_ordered_field α] : has_inv {x : α // 0 ≤ x} := -{ inv := λ x, ⟨x⁻¹, inv_nonneg.mpr x.2⟩ } - -@[simp, norm_cast] -protected lemma coe_inv [linear_ordered_field α] (a : {x : α // 0 ≤ x}) : - ((a⁻¹ : {x : α // 0 ≤ x}) : α) = a⁻¹ := rfl - -@[simp] lemma inv_mk [linear_ordered_field α] {x : α} (hx : 0 ≤ x) : - (⟨x, hx⟩ : {x : α // 0 ≤ x})⁻¹ = ⟨x⁻¹, inv_nonneg.mpr hx⟩ := -rfl - -instance linear_ordered_comm_group_with_zero [linear_ordered_field α] : - linear_ordered_comm_group_with_zero {x : α // 0 ≤ x} := -{ inv_zero := by { ext, exact inv_zero }, - mul_inv_cancel := by { intros a ha, ext, refine mul_inv_cancel (mt (λ h, _) ha), ext, exact h }, - ..nonneg.nontrivial, - ..nonneg.has_inv, - ..nonneg.linear_ordered_comm_monoid_with_zero } - -instance has_div [linear_ordered_field α] : has_div {x : α // 0 ≤ x} := -{ div := λ x y, ⟨x / y, div_nonneg x.2 y.2⟩ } - -@[simp, norm_cast] -protected lemma coe_div [linear_ordered_field α] (a b : {x : α // 0 ≤ x}) : - ((a / b : {x : α // 0 ≤ x}) : α) = a / b := rfl - -@[simp] lemma mk_div_mk [linear_ordered_field α] {x y : α} (hx : 0 ≤ x) (hy : 0 ≤ y) : - (⟨x, hx⟩ : {x : α // 0 ≤ x}) / ⟨y, hy⟩ = ⟨x / y, div_nonneg hx hy⟩ := -rfl - instance canonically_ordered_add_monoid [ordered_ring α] : canonically_ordered_add_monoid {x : α // 0 ≤ x} := { le_self_add := λ a b, le_add_of_nonneg_right b.2, @@ -275,6 +242,51 @@ instance canonically_linear_ordered_add_monoid [linear_ordered_ring α] : canonically_linear_ordered_add_monoid {x : α // 0 ≤ x} := { ..subtype.linear_order _, ..nonneg.canonically_ordered_add_monoid } +section linear_ordered_semifield +variables [linear_ordered_semifield α] {x y : α} + +instance has_inv : has_inv {x : α // 0 ≤ x} := ⟨λ x, ⟨x⁻¹, inv_nonneg.mpr x.2⟩⟩ + +@[simp, norm_cast] +protected lemma coe_inv (a : {x : α // 0 ≤ x}) : ((a⁻¹ : {x : α // 0 ≤ x}) : α) = a⁻¹ := rfl + +@[simp] lemma inv_mk (hx : 0 ≤ x) : (⟨x, hx⟩ : {x : α // 0 ≤ x})⁻¹ = ⟨x⁻¹, inv_nonneg.mpr hx⟩ := rfl + +instance has_div : has_div {x : α // 0 ≤ x} := ⟨λ x y, ⟨x / y, div_nonneg x.2 y.2⟩⟩ + +@[simp, norm_cast] protected lemma coe_div (a b : {x : α // 0 ≤ x}) : + ((a / b : {x : α // 0 ≤ x}) : α) = a / b := rfl + +@[simp] lemma mk_div_mk (hx : 0 ≤ x) (hy : 0 ≤ y) : + (⟨x, hx⟩ : {x : α // 0 ≤ x}) / ⟨y, hy⟩ = ⟨x / y, div_nonneg hx hy⟩ := rfl + +instance has_zpow : has_pow {x : α // 0 ≤ x} ℤ := ⟨λ a n, ⟨a ^ n, zpow_nonneg a.2 _⟩⟩ + +@[simp, norm_cast] protected lemma coe_zpow (a : {x : α // 0 ≤ x}) (n : ℤ) : + ((a ^ n : {x : α // 0 ≤ x}) : α) = a ^ n := rfl + +@[simp] lemma mk_zpow (hx : 0 ≤ x) (n : ℤ) : + (⟨x, hx⟩ : {x : α // 0 ≤ x}) ^ n = ⟨x ^ n, zpow_nonneg hx n⟩ := rfl + +instance linear_ordered_semifield : linear_ordered_semifield {x : α // 0 ≤ x} := +subtype.coe_injective.linear_ordered_semifield _ nonneg.coe_zero nonneg.coe_one nonneg.coe_add + nonneg.coe_mul nonneg.coe_inv nonneg.coe_div (λ _ _, rfl) nonneg.coe_pow nonneg.coe_zpow + nonneg.coe_nat_cast (λ _ _, rfl) (λ _ _, rfl) + +end linear_ordered_semifield + +instance linear_ordered_comm_group_with_zero [linear_ordered_field α] : + linear_ordered_comm_group_with_zero {x : α // 0 ≤ x} := +{ inv_zero := by { ext, exact inv_zero }, + mul_inv_cancel := by { intros a ha, ext, refine mul_inv_cancel (mt (λ h, _) ha), ext, exact h }, + ..nonneg.nontrivial, + ..nonneg.has_inv, + ..nonneg.linear_ordered_comm_monoid_with_zero } + +instance canonically_linear_ordered_semifield [linear_ordered_field α] : + canonically_linear_ordered_semifield {x : α // 0 ≤ x} := +{ ..nonneg.linear_ordered_semifield, ..nonneg.canonically_ordered_comm_semiring } + instance floor_semiring [ordered_semiring α] [floor_semiring α] : floor_semiring {r : α // 0 ≤ r} := { floor := λ a, ⌊(a : α)⌋₊, ceil := λ a, ⌈(a : α)⌉₊, From 8dd800933aebd08d1adda7e989e0ca9e745aca24 Mon Sep 17 00:00:00 2001 From: Gabriel Ebner Date: Sat, 27 Aug 2022 21:53:55 +0000 Subject: [PATCH 056/828] fix(combinatorics/composition): remove non-terminal simp (#16259) --- src/combinatorics/composition.lean | 19 +++++++++---------- 1 file changed, 9 insertions(+), 10 deletions(-) diff --git a/src/combinatorics/composition.lean b/src/combinatorics/composition.lean index d7dcb7f1871a3..56fe48c566bfb 100644 --- a/src/combinatorics/composition.lean +++ b/src/combinatorics/composition.lean @@ -886,18 +886,17 @@ end c.to_composition.boundaries = c.boundaries := begin ext j, - simp [c.mem_boundaries_iff_exists_blocks_sum_take_eq, c.card_boundaries_eq_succ_length, - composition.boundary, fin.ext_iff, composition.size_up_to, exists_prop, finset.mem_univ, - take, exists_prop_of_true, finset.mem_image, composition_as_set.to_composition_blocks, - composition.boundaries], + simp only [c.mem_boundaries_iff_exists_blocks_sum_take_eq, composition.boundaries, + finset.mem_map], split, - { rintros ⟨i, hi⟩, - refine ⟨i.1, _, hi⟩, - convert i.2, - simp }, + { rintros ⟨i, _, hi⟩, + refine ⟨i.1, _, _⟩, + simpa [c.card_boundaries_eq_succ_length] using i.2, + simp [composition.boundary, composition.size_up_to, ← hi] }, { rintros ⟨i, i_lt, hi⟩, - have : i < c.to_composition.length + 1, by simpa using i_lt, - exact ⟨⟨i, this⟩, hi⟩ } + refine ⟨i, by simp, _⟩, + rw [c.card_boundaries_eq_succ_length] at i_lt, + simp [composition.boundary, nat.mod_eq_of_lt i_lt, composition.size_up_to, hi] } end @[simp] lemma composition.to_composition_as_set_boundaries (c : composition n) : From 937c692d73f5130c7fecd3fd32e81419f4e04eb7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 27 Aug 2022 21:53:56 +0000 Subject: [PATCH 057/828] chore(category_theory): fix name of mono_app_of_mono (#16275) This PR fixes the names of a few lemmas like `mono_app_of_mono`. --- src/category_theory/adjunction/evaluation.lean | 8 ++++---- src/category_theory/functor/category.lean | 4 ++-- src/category_theory/sites/subsheaf.lean | 2 +- 3 files changed, 7 insertions(+), 7 deletions(-) diff --git a/src/category_theory/adjunction/evaluation.lean b/src/category_theory/adjunction/evaluation.lean index 1bcf4ca7e93ef..735917af13f81 100644 --- a/src/category_theory/adjunction/evaluation.lean +++ b/src/category_theory/adjunction/evaluation.lean @@ -75,14 +75,14 @@ instance evaluation_is_right_adjoint (c : C) : is_right_adjoint ((evaluation _ D).obj c) := ⟨_, evaluation_adjunction_right _ _⟩ -lemma nat_trans.mono_iff_app_mono {F G : C ⥤ D} (η : F ⟶ G) : +lemma nat_trans.mono_iff_mono_app {F G : C ⥤ D} (η : F ⟶ G) : mono η ↔ (∀ c, mono (η.app c)) := begin split, { introsI h c, exact (infer_instance : mono (((evaluation _ _).obj c).map η)) }, { introsI _, - apply nat_trans.mono_app_of_mono } + apply nat_trans.mono_of_mono_app } end end @@ -140,14 +140,14 @@ instance evaluation_is_left_adjoint (c : C) : is_left_adjoint ((evaluation _ D).obj c) := ⟨_, evaluation_adjunction_left _ _⟩ -lemma nat_trans.epi_iff_app_epi {F G : C ⥤ D} (η : F ⟶ G) : +lemma nat_trans.epi_iff_epi_app {F G : C ⥤ D} (η : F ⟶ G) : epi η ↔ (∀ c, epi (η.app c)) := begin split, { introsI h c, exact (infer_instance : epi (((evaluation _ _).obj c).map η)) }, { introsI, - apply nat_trans.epi_app_of_epi } + apply nat_trans.epi_of_epi_app } end end diff --git a/src/category_theory/functor/category.lean b/src/category_theory/functor/category.lean index d19024ee0f9a5..77144a5621fc3 100644 --- a/src/category_theory/functor/category.lean +++ b/src/category_theory/functor/category.lean @@ -67,11 +67,11 @@ lemma naturality_app {F G : C ⥤ (D ⥤ E)} (T : F ⟶ G) (Z : D) {X Y : C} (f congr_fun (congr_arg app (T.naturality f)) Z /-- A natural transformation is a monomorphism if each component is. -/ -lemma mono_app_of_mono (α : F ⟶ G) [∀ (X : C), mono (α.app X)] : mono α := +lemma mono_of_mono_app (α : F ⟶ G) [∀ (X : C), mono (α.app X)] : mono α := ⟨λ H g h eq, by { ext X, rw [←cancel_mono (α.app X), ←comp_app, eq, comp_app] }⟩ /-- A natural transformation is an epimorphism if each component is. -/ -lemma epi_app_of_epi (α : F ⟶ G) [∀ (X : C), epi (α.app X)] : epi α := +lemma epi_of_epi_app (α : F ⟶ G) [∀ (X : C), epi (α.app X)] : epi α := ⟨λ H g h eq, by { ext X, rw [←cancel_epi (α.app X), ←comp_app, eq, comp_app] }⟩ /-- `hcomp α β` is the horizontal composition of natural transformations. -/ diff --git a/src/category_theory/sites/subsheaf.lean b/src/category_theory/sites/subsheaf.lean index 70eace616b728..c3081cd5705f4 100644 --- a/src/category_theory/sites/subsheaf.lean +++ b/src/category_theory/sites/subsheaf.lean @@ -324,7 +324,7 @@ begin rw is_iso_iff_bijective, split, { intros x y e, - have := (nat_trans.mono_iff_app_mono _ _).mp hf X, + have := (nat_trans.mono_iff_mono_app _ _).mp hf X, rw mono_iff_injective at this, exact this (congr_arg subtype.val e : _) }, { rintro ⟨_, ⟨x, rfl⟩⟩, exact ⟨x, rfl⟩ } From de62604b58cf61aad9188a41aa7d9f51cfe73245 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sun, 28 Aug 2022 08:57:09 +0000 Subject: [PATCH 058/828] feat(probability/integration): remove integrability assumption in indep_fun.integral_mul (#16167) The integral of a product of independent random variables is the product of the integrals. This is already in mathlib when the random variables are integrable, but it also holds when they are not integrable (as both sides are then zero). In this PR, we remove the integrability assumption from this statement (and weaken accordingly further results that depended on this one). --- src/probability/independence.lean | 4 ++ src/probability/integration.lean | 93 +++++++++++++++++++++++++++++-- src/probability/moments.lean | 58 +++++++++++++++---- src/probability/variance.lean | 6 +- 4 files changed, 143 insertions(+), 18 deletions(-) diff --git a/src/probability/independence.lean b/src/probability/independence.lean index e43fcb7cb62ee..12d65123b7b09 100644 --- a/src/probability/independence.lean +++ b/src/probability/independence.lean @@ -520,6 +520,10 @@ begin { rwa ← indep_set_iff_measure_inter_eq_mul (hf hs) (hg ht) μ, }, end +lemma indep_fun.symm {mβ : measurable_space β} {f g : Ω → β} (hfg : indep_fun f g μ) : + indep_fun g f μ := +hfg.symm + lemma indep_fun.ae_eq {mβ : measurable_space β} {f g f' g' : Ω → β} (hfg : indep_fun f g μ) (hf : f =ᵐ[μ] f') (hg : g =ᵐ[μ] g') : indep_fun f' g' μ := diff --git a/src/probability/integration.lean b/src/probability/integration.lean index 7a28eaab17b53..1612be45178b8 100644 --- a/src/probability/integration.lean +++ b/src/probability/integration.lean @@ -129,6 +129,11 @@ begin exact h_indep_fun.ae_eq h_meas_f.ae_eq_mk h_meas_g.ae_eq_mk end +lemma lintegral_mul_eq_lintegral_mul_lintegral_of_indep_fun'' + (h_meas_f : ae_measurable f μ) (h_meas_g : ae_measurable g μ) (h_indep_fun : indep_fun f g μ) : + ∫⁻ ω, f ω * g ω ∂μ = ∫⁻ ω, f ω ∂μ * ∫⁻ ω, g ω ∂μ := +lintegral_mul_eq_lintegral_mul_lintegral_of_indep_fun' h_meas_f h_meas_g h_indep_fun + /-- The product of two independent, integrable, real_valued random variables is integrable. -/ lemma indep_fun.integrable_mul {β : Type*} [measurable_space β] {X Y : Ω → β} [normed_division_ring β] [borel_space β] @@ -154,6 +159,56 @@ begin exact ennreal.mul_lt_top_iff.mpr (or.inl ⟨hX.2, hY.2⟩) end +/-- If the product of two independent real_valued random variables is integrable and +the second one is not almost everywhere zero, then the first one is integrable. -/ +lemma indep_fun.integrable_left_of_integrable_mul {β : Type*} [measurable_space β] {X Y : Ω → β} + [normed_division_ring β] [borel_space β] + (hXY : indep_fun X Y μ) (h'XY : integrable (X * Y) μ) + (hX : ae_strongly_measurable X μ) (hY : ae_strongly_measurable Y μ) (h'Y : ¬(Y =ᵐ[μ] 0)) : + integrable X μ := +begin + refine ⟨hX, _⟩, + have I : ∫⁻ ω, ∥Y ω∥₊ ∂μ ≠ 0, + { assume H, + have I : (λ ω, ↑∥Y ω∥₊) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hY.ennnorm).1 H, + apply h'Y, + filter_upwards [I] with ω hω, + simpa using hω }, + apply lt_top_iff_ne_top.2 (λ H, _), + have J : indep_fun (λ ω, ↑∥X ω∥₊) (λ ω, ↑∥Y ω∥₊) μ, + { have M : measurable (λ (x : β), (∥x∥₊ : ℝ≥0∞)) := measurable_nnnorm.coe_nnreal_ennreal, + apply indep_fun.comp hXY M M }, + have A : ∫⁻ ω, ∥X ω * Y ω∥₊ ∂μ < ∞ := h'XY.2, + simp only [nnnorm_mul, ennreal.coe_mul] at A, + rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indep_fun'' hX.ennnorm hY.ennnorm J, H] at A, + simpa [ennreal.top_mul, I] using A, +end + +/-- If the product of two independent real_valued random variables is integrable and the +first one is not almost everywhere zero, then the second one is integrable. -/ +lemma indep_fun.integrable_right_of_integrable_mul {β : Type*} [measurable_space β] {X Y : Ω → β} + [normed_division_ring β] [borel_space β] + (hXY : indep_fun X Y μ) (h'XY : integrable (X * Y) μ) + (hX : ae_strongly_measurable X μ) (hY : ae_strongly_measurable Y μ) (h'X : ¬(X =ᵐ[μ] 0)) : + integrable Y μ := +begin + refine ⟨hY, _⟩, + have I : ∫⁻ ω, ∥X ω∥₊ ∂μ ≠ 0, + { assume H, + have I : (λ ω, ↑∥X ω∥₊) =ᵐ[μ] 0 := (lintegral_eq_zero_iff' hX.ennnorm).1 H, + apply h'X, + filter_upwards [I] with ω hω, + simpa using hω }, + apply lt_top_iff_ne_top.2 (λ H, _), + have J : indep_fun (λ ω, ↑∥X ω∥₊) (λ ω, ↑∥Y ω∥₊) μ, + { have M : measurable (λ (x : β), (∥x∥₊ : ℝ≥0∞)) := measurable_nnnorm.coe_nnreal_ennreal, + apply indep_fun.comp hXY M M }, + have A : ∫⁻ ω, ∥X ω * Y ω∥₊ ∂μ < ∞ := h'XY.2, + simp only [nnnorm_mul, ennreal.coe_mul] at A, + rw [lintegral_mul_eq_lintegral_mul_lintegral_of_indep_fun'' hX.ennnorm hY.ennnorm J, H] at A, + simpa [ennreal.top_mul, I] using A, +end + /-- The (Bochner) integral of the product of two independent, nonnegative random variables is the product of their integrals. The proof is just plumbing around `lintegral_mul_eq_lintegral_mul_lintegral_of_indep_fun'`. -/ @@ -234,10 +289,40 @@ begin ring end -theorem indep_fun.integral_mul_of_integrable' (hXY : indep_fun X Y μ) - (hX : integrable X μ) (hY : integrable Y μ) : - integral μ (λ x, X x * Y x) = integral μ X * integral μ Y := -hXY.integral_mul_of_integrable hX hY +/-- The (Bochner) integral of the product of two independent random + variables is the product of their integrals. -/ +theorem indep_fun.integral_mul (hXY : indep_fun X Y μ) + (hX : ae_strongly_measurable X μ) (hY : ae_strongly_measurable Y μ) : + integral μ (X * Y) = integral μ X * integral μ Y := +begin + by_cases h'X : X =ᵐ[μ] 0, + { have h' : X * Y =ᵐ[μ] 0, + { filter_upwards [h'X] with ω hω, + simp [hω] }, + simp only [integral_congr_ae h'X, integral_congr_ae h', pi.zero_apply, integral_const, + algebra.id.smul_eq_mul, mul_zero, zero_mul] }, + by_cases h'Y : Y =ᵐ[μ] 0, + { have h' : X * Y =ᵐ[μ] 0, + { filter_upwards [h'Y] with ω hω, + simp [hω] }, + simp only [integral_congr_ae h'Y, integral_congr_ae h', pi.zero_apply, integral_const, + algebra.id.smul_eq_mul, mul_zero, zero_mul] }, + by_cases h : integrable (X * Y) μ, + { have HX : integrable X μ := hXY.integrable_left_of_integrable_mul h hX hY h'Y, + have HY : integrable Y μ := hXY.integrable_right_of_integrable_mul h hX hY h'X, + exact hXY.integral_mul_of_integrable HX HY }, + { have I : ¬(integrable X μ ∧ integrable Y μ), + { rintros ⟨HX, HY⟩, + exact h (hXY.integrable_mul HX HY) }, + rw not_and_distrib at I, + cases I; + simp [integral_undef, I, h] } +end + +theorem indep_fun.integral_mul' (hXY : indep_fun X Y μ) + (hX : ae_strongly_measurable X μ) (hY : ae_strongly_measurable Y μ) : + integral μ (λ ω, X ω * Y ω) = integral μ X * integral μ Y := +hXY.integral_mul hX hY /-- Independence of functions `f` and `g` into arbitrary types is characterized by the relation `E[(φ ∘ f) * (ψ ∘ g)] = E[φ ∘ f] * E[ψ ∘ g]` for all measurable `φ` and `ψ` with values in `ℝ` diff --git a/src/probability/moments.lean b/src/probability/moments.lean index 46efcf3c96c5f..bbbdc72197fdf 100644 --- a/src/probability/moments.lean +++ b/src/probability/moments.lean @@ -192,12 +192,24 @@ begin end lemma indep_fun.mgf_add {X Y : Ω → ℝ} (h_indep : indep_fun X Y μ) - (h_int_X : integrable (λ ω, exp (t * X ω)) μ) - (h_int_Y : integrable (λ ω, exp (t * Y ω)) μ) : + (hX : ae_strongly_measurable (λ ω, exp (t * X ω)) μ) + (hY : ae_strongly_measurable (λ ω, exp (t * Y ω)) μ) : mgf (X + Y) μ t = mgf X μ t * mgf Y μ t := begin simp_rw [mgf, pi.add_apply, mul_add, exp_add], - exact (h_indep.exp_mul t t).integral_mul_of_integrable' h_int_X h_int_Y, + exact (h_indep.exp_mul t t).integral_mul hX hY, +end + +lemma indep_fun.mgf_add' {X Y : Ω → ℝ} (h_indep : indep_fun X Y μ) + (hX : ae_strongly_measurable X μ) (hY : ae_strongly_measurable Y μ) : + mgf (X + Y) μ t = mgf X μ t * mgf Y μ t := +begin + have A : continuous (λ (x : ℝ), exp (t * x)), by continuity, + have h'X : ae_strongly_measurable (λ ω, exp (t * X ω)) μ := + A.ae_strongly_measurable.comp_ae_measurable hX.ae_measurable, + have h'Y : ae_strongly_measurable (λ ω, exp (t * Y ω)) μ := + A.ae_strongly_measurable.comp_ae_measurable hY.ae_measurable, + exact h_indep.mgf_add h'X h'Y end lemma indep_fun.cgf_add {X Y : Ω → ℝ} (h_indep : indep_fun X Y μ) @@ -207,10 +219,34 @@ lemma indep_fun.cgf_add {X Y : Ω → ℝ} (h_indep : indep_fun X Y μ) begin by_cases hμ : μ = 0, { simp [hμ], }, - simp only [cgf, h_indep.mgf_add h_int_X h_int_Y], + simp only [cgf, h_indep.mgf_add h_int_X.ae_strongly_measurable h_int_Y.ae_strongly_measurable], exact log_mul (mgf_pos' hμ h_int_X).ne' (mgf_pos' hμ h_int_Y).ne', end +lemma ae_strongly_measurable_exp_mul_add {X Y : Ω → ℝ} + (h_int_X : ae_strongly_measurable (λ ω, exp (t * X ω)) μ) + (h_int_Y : ae_strongly_measurable (λ ω, exp (t * Y ω)) μ) : + ae_strongly_measurable (λ ω, exp (t * (X + Y) ω)) μ := +begin + simp_rw [pi.add_apply, mul_add, exp_add], + exact ae_strongly_measurable.mul h_int_X h_int_Y, +end + +lemma ae_strongly_measurable_exp_mul_sum {X : ι → Ω → ℝ} {s : finset ι} + (h_int : ∀ i ∈ s, ae_strongly_measurable (λ ω, exp (t * X i ω)) μ) : + ae_strongly_measurable (λ ω, exp (t * (∑ i in s, X i) ω)) μ := +begin + classical, + induction s using finset.induction_on with i s hi_notin_s h_rec h_int, + { simp only [pi.zero_apply, sum_apply, sum_empty, mul_zero, exp_zero], + exact ae_strongly_measurable_const, }, + { have : ∀ (i : ι), i ∈ s → ae_strongly_measurable (λ (ω : Ω), exp (t * X i ω)) μ, + from λ i hi, h_int i (mem_insert_of_mem hi), + specialize h_rec this, + rw sum_insert hi_notin_s, + apply ae_strongly_measurable_exp_mul_add (h_int i (mem_insert_self _ _)) h_rec } +end + lemma indep_fun.integrable_exp_mul_add {X Y : Ω → ℝ} (h_indep : indep_fun X Y μ) (h_int_X : integrable (λ ω, exp (t * X ω)) μ) (h_int_Y : integrable (λ ω, exp (t * Y ω)) μ) : @@ -239,18 +275,18 @@ end lemma Indep_fun.mgf_sum [is_probability_measure μ] {X : ι → Ω → ℝ} (h_indep : Indep_fun (λ i, infer_instance) X μ) (h_meas : ∀ i, measurable (X i)) - {s : finset ι} (h_int : ∀ i ∈ s, integrable (λ ω, exp (t * X i ω)) μ) : + (s : finset ι) : mgf (∑ i in s, X i) μ t = ∏ i in s, mgf (X i) μ t := begin classical, induction s using finset.induction_on with i s hi_notin_s h_rec h_int, { simp only [sum_empty, mgf_zero_fun, measure_univ, ennreal.one_to_real, prod_empty], }, - { have h_int' : ∀ (i : ι), i ∈ s → integrable (λ (ω : Ω), exp (t * X i ω)) μ, - from λ i hi, h_int i (mem_insert_of_mem hi), + { have h_int' : ∀ (i : ι), ae_strongly_measurable (λ (ω : Ω), exp (t * X i ω)) μ, + from λ i, ((h_meas i).const_mul t).exp.ae_strongly_measurable, rw [sum_insert hi_notin_s, indep_fun.mgf_add - (h_indep.indep_fun_finset_sum_of_not_mem h_meas hi_notin_s).symm - (h_int i (mem_insert_self _ _)) (h_indep.integrable_exp_mul_sum h_meas h_int'), - h_rec h_int', prod_insert hi_notin_s], }, + (h_indep.indep_fun_finset_sum_of_not_mem h_meas hi_notin_s).symm (h_int' i) + (ae_strongly_measurable_exp_mul_sum (λ i hi, h_int' i)), + h_rec, prod_insert hi_notin_s] } end lemma Indep_fun.cgf_sum [is_probability_measure μ] @@ -260,7 +296,7 @@ lemma Indep_fun.cgf_sum [is_probability_measure μ] begin simp_rw cgf, rw ← log_prod _ _ (λ j hj, _), - { rw h_indep.mgf_sum h_meas h_int, }, + { rw h_indep.mgf_sum h_meas }, { exact (mgf_pos (h_int j hj)).ne', }, end diff --git a/src/probability/variance.lean b/src/probability/variance.lean index 55c010f4ed306..e884101ba30ec 100644 --- a/src/probability/variance.lean +++ b/src/probability/variance.lean @@ -224,11 +224,11 @@ begin (λ i hi, (mem_ℒp.integrable one_le_two (hs _ (mem_insert_of_mem hi)))), mul_sum, mul_sum, ← sum_sub_distrib], apply finset.sum_eq_zero (λ i hi, _), - rw [integral_mul_left, indep_fun.integral_mul_of_integrable', sub_self], + rw [integral_mul_left, indep_fun.integral_mul', sub_self], { apply h (mem_insert_self _ _) (mem_insert_of_mem hi), exact (λ hki, ks (hki.symm ▸ hi)) }, - { exact mem_ℒp.integrable one_le_two (hs _ (mem_insert_self _ _)) }, - { exact mem_ℒp.integrable one_le_two (hs _ (mem_insert_of_mem hi)) } + { exact mem_ℒp.ae_strongly_measurable (hs _ (mem_insert_self _ _)) }, + { exact mem_ℒp.ae_strongly_measurable (hs _ (mem_insert_of_mem hi)) } end ... = Var[X k] + ∑ i in s, Var[X i] : by rw IH (λ i hi, hs i (mem_insert_of_mem hi)) From 6de241f9a8c59ba3a7f3bb9fe92a18367c6f5544 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 29 Aug 2022 04:56:43 +0000 Subject: [PATCH 059/828] feat(data/rat/nnrat): Nonnegative rationals (#16283) Define `nnrat`, the subtype of nonnegative rational numbers, and show it forms a `canonically_linear_ordered_semifield`. --- src/data/rat/nnrat.lean | 301 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 301 insertions(+) create mode 100644 src/data/rat/nnrat.lean diff --git a/src/data/rat/nnrat.lean b/src/data/rat/nnrat.lean new file mode 100644 index 0000000000000..8b14701c39ca4 --- /dev/null +++ b/src/data/rat/nnrat.lean @@ -0,0 +1,301 @@ +/- +Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Bhavik Mehta +-/ +import algebra.algebra.basic +import algebra.order.nonneg + +/-! +# Nonnegative rationals + +This file defines the nonnegative rationals as a subtype of `rat` and provides its algebraic order +structure. + +We also define an instance `can_lift ℚ ℚ≥0`. This instance can be used by the `lift` tactic to +replace `x : ℚ` and `hx : 0 ≤ x` in the proof context with `x : ℚ≥0` while replacing all occurences +of `x` with `↑x`. This tactic also works for a function `f : α → ℚ` with a hypothesis +`hf : ∀ x, 0 ≤ f x`. + +## Notation + +`ℚ≥0` is notation for `nnrat` in locale `nnrat`. +-/ + +open function +open_locale big_operators + +/-- Nonnegative rational numbers. -/ +@[derive [canonically_ordered_comm_semiring, canonically_linear_ordered_semifield, + linear_ordered_comm_group_with_zero, has_sub, has_ordered_sub, + densely_ordered, archimedean, inhabited]] +def nnrat := {q : ℚ // 0 ≤ q} + +localized "notation ` ℚ≥0 ` := nnrat" in nnrat + +namespace nnrat +variables {α : Type*} {p q : ℚ≥0} + +instance : has_coe ℚ≥0 ℚ := ⟨subtype.val⟩ + +/- Simp lemma to put back `n.val` into the normal form given by the coercion. -/ +@[simp] lemma val_eq_coe (q : ℚ≥0) : q.val = q := rfl + +instance : can_lift ℚ ℚ≥0 := +{ coe := coe, + cond := λ q, 0 ≤ q, + prf := λ q hq, ⟨⟨q, hq⟩, rfl⟩ } + +@[ext] lemma ext : (p : ℚ) = (q : ℚ) → p = q := subtype.ext + +protected lemma coe_injective : injective (coe : ℚ≥0 → ℚ) := subtype.coe_injective + +@[simp, norm_cast] lemma coe_inj : (p : ℚ) = q ↔ p = q := subtype.coe_inj + +lemma ext_iff : p = q ↔ (p : ℚ) = q := subtype.ext_iff + +lemma ne_iff {x y : ℚ≥0} : (x : ℚ) ≠ (y : ℚ) ↔ x ≠ y := nnrat.coe_inj.not + +@[norm_cast] lemma coe_mk (q : ℚ) (hq) : ((⟨q, hq⟩ : ℚ≥0) : ℚ) = q := rfl + +/-- Reinterpret a rational number `q` as a non-negative rational number. Returns `0` if `q ≤ 0`. -/ +def _root_.rat.to_nnrat (q : ℚ) : ℚ≥0 := ⟨max q 0, le_max_right _ _⟩ + +lemma _root_.rat.coe_to_nnrat (q : ℚ) (hq : 0 ≤ q) : (q.to_nnrat : ℚ) = q := max_eq_left hq + +lemma _root_.rat.le_coe_to_nnrat (q : ℚ) : q ≤ q.to_nnrat := le_max_left _ _ + +open _root_.rat (to_nnrat) + +@[simp] lemma coe_nonneg (q : ℚ≥0) : (0 : ℚ) ≤ q := q.2 + +@[simp, norm_cast] lemma coe_zero : ((0 : ℚ≥0) : ℚ) = 0 := rfl +@[simp, norm_cast] lemma coe_one : ((1 : ℚ≥0) : ℚ) = 1 := rfl +@[simp, norm_cast] lemma coe_add (p q : ℚ≥0) : ((p + q : ℚ≥0) : ℚ) = p + q := rfl +@[simp, norm_cast] lemma coe_mul (p q : ℚ≥0) : ((p * q : ℚ≥0) : ℚ) = p * q := rfl +@[simp, norm_cast] lemma coe_inv (q : ℚ≥0) : ((q⁻¹ : ℚ≥0) : ℚ) = q⁻¹ := rfl +@[simp, norm_cast] lemma coe_div (p q : ℚ≥0) : ((p / q : ℚ≥0) : ℚ) = p / q := rfl +@[simp, norm_cast] lemma coe_bit0 (q : ℚ≥0) : ((bit0 q : ℚ≥0) : ℚ) = bit0 q := rfl +@[simp, norm_cast] lemma coe_bit1 (q : ℚ≥0) : ((bit1 q : ℚ≥0) : ℚ) = bit1 q := rfl +@[simp, norm_cast] lemma coe_sub (h : q ≤ p) : ((p - q : ℚ≥0) : ℚ) = p - q := +max_eq_left $ le_sub.2 $ by simp [show (q : ℚ) ≤ p, from h] + +@[simp] lemma coe_eq_zero : (q : ℚ) = 0 ↔ q = 0 := by norm_cast +lemma coe_ne_zero : (q : ℚ) ≠ 0 ↔ q ≠ 0 := coe_eq_zero.not + +@[simp, norm_cast] lemma coe_le_coe : (p : ℚ) ≤ q ↔ p ≤ q := iff.rfl +@[simp, norm_cast] lemma coe_lt_coe : (p : ℚ) < q ↔ p < q := iff.rfl +@[simp, norm_cast] lemma coe_pos : (0 : ℚ) < q ↔ 0 < q := iff.rfl + +lemma coe_mono : monotone (coe : ℚ≥0 → ℚ) := λ _ _, coe_le_coe.2 + +lemma to_nnrat_mono : monotone to_nnrat := λ x y h, max_le_max h le_rfl + +@[simp] lemma to_nnrat_coe (q : ℚ≥0) : to_nnrat q = q := ext $ max_eq_left q.2 + +@[simp] lemma to_nnrat_coe_nat (n : ℕ) : to_nnrat n = n := +ext $ by simp [rat.coe_to_nnrat] + +/-- `to_nnrat` and `coe : ℚ≥0 → ℚ` form a Galois insertion. -/ +protected def gi : galois_insertion to_nnrat coe := +galois_insertion.monotone_intro coe_mono to_nnrat_mono rat.le_coe_to_nnrat to_nnrat_coe + +/-- Coercion `ℚ≥0 → ℚ` as a `ring_hom`. -/ +def coe_hom : ℚ≥0 →+* ℚ := ⟨coe, coe_one, coe_mul, coe_zero, coe_add⟩ + +@[simp, norm_cast] lemma coe_nat_cast (n : ℕ) : (↑(↑n : ℚ≥0) : ℚ) = n := map_nat_cast coe_hom n + +@[simp] lemma mk_coe_nat (n : ℕ) : @eq ℚ≥0 (⟨(n : ℚ), n.cast_nonneg⟩ : ℚ≥0) n := +ext (coe_nat_cast n).symm + +/-- The rational numbers are an algebra over the non-negative rationals. -/ +instance : algebra ℚ≥0 ℚ := coe_hom.to_algebra + +/-- A `mul_action` over `ℚ` restricts to a `mul_action` over `ℚ≥0`. -/ +instance [mul_action ℚ α] : mul_action ℚ≥0 α := mul_action.comp_hom α coe_hom.to_monoid_hom + +/-- A `distrib_mul_action` over `ℚ` restricts to a `distrib_mul_action` over `ℚ≥0`. -/ +instance [add_comm_monoid α] [distrib_mul_action ℚ α] : distrib_mul_action ℚ≥0 α := +distrib_mul_action.comp_hom α coe_hom.to_monoid_hom + +/-- A `module` over `ℚ` restricts to a `module` over `ℚ≥0`. -/ +instance [add_comm_monoid α] [module ℚ α] : module ℚ≥0 α := module.comp_hom α coe_hom + +@[simp] lemma coe_coe_hom : ⇑coe_hom = coe := rfl + +@[simp, norm_cast] lemma coe_indicator (s : set α) (f : α → ℚ≥0) (a : α) : + ((s.indicator f a : ℚ≥0) : ℚ) = s.indicator (λ x, f x) a := +(coe_hom : ℚ≥0 →+ ℚ).map_indicator _ _ _ + +@[simp, norm_cast] lemma coe_pow (q : ℚ≥0) (n : ℕ) : (↑(q ^ n) : ℚ) = q ^ n := coe_hom.map_pow _ _ + +@[norm_cast] lemma coe_list_sum (l : list ℚ≥0) : (l.sum : ℚ) = (l.map coe).sum := +coe_hom.map_list_sum _ + +@[norm_cast] lemma coe_list_prod (l : list ℚ≥0) : (l.prod : ℚ) = (l.map coe).prod := +coe_hom.map_list_prod _ + +@[norm_cast] lemma coe_multiset_sum (s : multiset ℚ≥0) : (s.sum : ℚ) = (s.map coe).sum := +coe_hom.map_multiset_sum _ + +@[norm_cast] lemma coe_multiset_prod (s : multiset ℚ≥0) : (s.prod : ℚ) = (s.map coe).prod := +coe_hom.map_multiset_prod _ + +@[norm_cast] lemma coe_sum {s : finset α} {f : α → ℚ≥0} : ↑(∑ a in s, f a) = ∑ a in s, (f a : ℚ) := +coe_hom.map_sum _ _ + +lemma to_nnrat_sum_of_nonneg {s : finset α} {f : α → ℚ} (hf : ∀ a, a ∈ s → 0 ≤ f a) : + (∑ a in s, f a).to_nnrat = ∑ a in s, (f a).to_nnrat := +begin + rw [←coe_inj, coe_sum, rat.coe_to_nnrat _ (finset.sum_nonneg hf)], + exact finset.sum_congr rfl (λ x hxs, by rw rat.coe_to_nnrat _ (hf x hxs)), +end + +@[norm_cast] lemma coe_prod {s : finset α} {f : α → ℚ≥0} : ↑(∏ a in s, f a) = ∏ a in s, (f a : ℚ) := +coe_hom.map_prod _ _ + +lemma to_nnrat_prod_of_nonneg {s : finset α} {f : α → ℚ} (hf : ∀ a ∈ s, 0 ≤ f a) : + (∏ a in s, f a).to_nnrat = ∏ a in s, (f a).to_nnrat := +begin + rw [←coe_inj, coe_prod, rat.coe_to_nnrat _ (finset.prod_nonneg hf)], + exact finset.prod_congr rfl (λ x hxs, by rw rat.coe_to_nnrat _ (hf x hxs)), +end + +@[norm_cast] lemma nsmul_coe (q : ℚ≥0) (n : ℕ) : ↑(n • q) = n • (q : ℚ) := +coe_hom.to_add_monoid_hom.map_nsmul _ _ + +lemma bdd_above_coe {s : set ℚ≥0} : bdd_above (coe '' s : set ℚ) ↔ bdd_above s := +⟨λ ⟨b, hb⟩, ⟨to_nnrat b, λ ⟨y, hy⟩ hys, show y ≤ max b 0, from + (hb $ set.mem_image_of_mem _ hys).trans $ le_max_left _ _⟩, + λ ⟨b, hb⟩, ⟨b, λ y ⟨x, hx, eq⟩, eq ▸ hb hx⟩⟩ + +lemma bdd_below_coe (s : set ℚ≥0) : bdd_below ((coe : ℚ≥0 → ℚ) '' s) := ⟨0, λ r ⟨q, _, h⟩, h ▸ q.2⟩ + +@[simp, norm_cast] lemma coe_max (x y : ℚ≥0) : ((max x y : ℚ≥0) : ℚ) = max (x : ℚ) (y : ℚ) := +coe_mono.map_max + +@[simp, norm_cast] lemma coe_min (x y : ℚ≥0) : ((min x y : ℚ≥0) : ℚ) = min (x : ℚ) (y : ℚ) := +coe_mono.map_min + +lemma sub_def (p q : ℚ≥0) : p - q = to_nnrat (p - q) := rfl + +@[simp] lemma abs_coe (q : ℚ≥0) : |(q : ℚ)| = q := abs_of_nonneg q.2 + +end nnrat + +open nnrat + +namespace rat +variables {p q : ℚ} + +@[simp] lemma to_nnrat_zero : to_nnrat 0 = 0 := by simp [to_nnrat]; refl +@[simp] lemma to_nnrat_one : to_nnrat 1 = 1 := by simp [to_nnrat, max_eq_left zero_le_one] + +@[simp] lemma to_nnrat_pos : 0 < to_nnrat q ↔ 0 < q := by simp [to_nnrat, ←coe_lt_coe] + +@[simp] lemma to_nnrat_eq_zero : to_nnrat q = 0 ↔ q ≤ 0 := +by simpa [-to_nnrat_pos] using (@to_nnrat_pos q).not + +alias to_nnrat_eq_zero ↔ _ to_nnrat_of_nonpos + +@[simp] lemma to_nnrat_le_to_nnrat_iff (hp : 0 ≤ p) : to_nnrat q ≤ to_nnrat p ↔ q ≤ p := +by simp [←coe_le_coe, to_nnrat, hp] + +@[simp] lemma to_nnrat_lt_to_nnrat_iff' : to_nnrat q < to_nnrat p ↔ q < p ∧ 0 < p := +by { simp [←coe_lt_coe, to_nnrat, lt_irrefl], exact lt_trans' } + +lemma to_nnrat_lt_to_nnrat_iff (h : 0 < p) : to_nnrat q < to_nnrat p ↔ q < p := +to_nnrat_lt_to_nnrat_iff'.trans (and_iff_left h) + +lemma to_nnrat_lt_to_nnrat_iff_of_nonneg (hq : 0 ≤ q) : to_nnrat q < to_nnrat p ↔ q < p := +to_nnrat_lt_to_nnrat_iff'.trans ⟨and.left, λ h, ⟨h, hq.trans_lt h⟩⟩ + +@[simp] lemma to_nnrat_add (hq : 0 ≤ q) (hp : 0 ≤ p) : to_nnrat (q + p) = to_nnrat q + to_nnrat p := +nnrat.ext $ by simp [to_nnrat, hq, hp, add_nonneg] + +lemma to_nnrat_add_le : to_nnrat (q + p) ≤ to_nnrat q + to_nnrat p := +coe_le_coe.1 $ max_le (add_le_add (le_max_left _ _) (le_max_left _ _)) $ coe_nonneg _ + +lemma to_nnrat_le_iff_le_coe {p : ℚ≥0} : to_nnrat q ≤ p ↔ q ≤ ↑p := nnrat.gi.gc q p + +lemma le_to_nnrat_iff_coe_le {q : ℚ≥0} (hp : 0 ≤ p) : q ≤ to_nnrat p ↔ ↑q ≤ p := +by rw [←coe_le_coe, rat.coe_to_nnrat p hp] + +lemma le_to_nnrat_iff_coe_le' {q : ℚ≥0} (hq : 0 < q) : q ≤ to_nnrat p ↔ ↑q ≤ p := +(le_or_lt 0 p).elim le_to_nnrat_iff_coe_le $ λ hp, + by simp only [(hp.trans_le q.coe_nonneg).not_le, to_nnrat_eq_zero.2 hp.le, hq.not_le] + +lemma to_nnrat_lt_iff_lt_coe {p : ℚ≥0} (hq : 0 ≤ q) : to_nnrat q < p ↔ q < ↑p := +by rw [←coe_lt_coe, rat.coe_to_nnrat q hq] + +lemma lt_to_nnrat_iff_coe_lt {q : ℚ≥0} : q < to_nnrat p ↔ ↑q < p := nnrat.gi.gc.lt_iff_lt + +@[simp] lemma to_nnrat_bit0 (hq : 0 ≤ q) : to_nnrat (bit0 q) = bit0 (to_nnrat q) := +to_nnrat_add hq hq + +@[simp] lemma to_nnrat_bit1 (hq : 0 ≤ q) : to_nnrat (bit1 q) = bit1 (to_nnrat q) := +(to_nnrat_add (by simp [hq]) zero_le_one).trans $ by simp [to_nnrat_one, bit1, hq] + +lemma to_nnrat_mul (hp : 0 ≤ p) : to_nnrat (p * q) = to_nnrat p * to_nnrat q := +begin + cases le_total 0 q with hq hq, + { ext; simp [to_nnrat, hp, hq, max_eq_left, mul_nonneg] }, + { have hpq := mul_nonpos_of_nonneg_of_nonpos hp hq, + rw [to_nnrat_eq_zero.2 hq, to_nnrat_eq_zero.2 hpq, mul_zero] } +end + +lemma to_nnrat_inv (q : ℚ) : to_nnrat q⁻¹ = (to_nnrat q)⁻¹ := +begin + obtain hq | hq := le_total q 0, + { rw [to_nnrat_eq_zero.mpr hq, inv_zero, to_nnrat_eq_zero.mpr (inv_nonpos.mpr hq)] }, + { nth_rewrite 0 ←rat.coe_to_nnrat q hq, + rw [←coe_inv, to_nnrat_coe] } +end + +lemma to_nnrat_div (hp : 0 ≤ p) : to_nnrat (p / q) = to_nnrat p / to_nnrat q := +by rw [div_eq_mul_inv, div_eq_mul_inv, ←to_nnrat_inv, ←to_nnrat_mul hp] + +lemma to_nnrat_div' (hq : 0 ≤ q) : to_nnrat (p / q) = to_nnrat p / to_nnrat q := +by rw [div_eq_inv_mul, div_eq_inv_mul, to_nnrat_mul (inv_nonneg.2 hq), to_nnrat_inv] + +end rat + +/-- The absolute value on `ℚ` as a map to `ℚ≥0`. -/ +@[pp_nodot] def rat.nnabs (x : ℚ) : ℚ≥0 := ⟨abs x, abs_nonneg x⟩ + +@[norm_cast, simp] lemma rat.coe_nnabs (x : ℚ) : (rat.nnabs x : ℚ) = abs x := by simp [rat.nnabs] + +/-! ### Numerator and denominator -/ + +namespace nnrat +variables {p q : ℚ≥0} + +/-- The numerator of a nonnegative rational. -/ +def num (q : ℚ≥0) : ℕ := (q : ℚ).num.nat_abs + +/-- The denominator of a nonnegative rational. -/ +def denom (q : ℚ≥0) : ℕ := (q : ℚ).denom + +@[simp] lemma nat_abs_num_coe : (q : ℚ).num.nat_abs = q.num := rfl +@[simp] lemma denom_coe : (q : ℚ).denom = q.denom := rfl + +lemma ext_num_denom (hn : p.num = q.num) (hd : p.denom = q.denom) : p = q := +ext $ rat.ext ((int.nat_abs_inj_of_nonneg_of_nonneg + (rat.num_nonneg_iff_zero_le.2 p.2) $ rat.num_nonneg_iff_zero_le.2 q.2).1 hn) hd + +lemma ext_num_denom_iff : p = q ↔ p.num = q.num ∧ p.denom = q.denom := +⟨by { rintro rfl, exact ⟨rfl, rfl⟩ }, λ h, ext_num_denom h.1 h.2⟩ + +@[simp] lemma num_div_denom (q : ℚ≥0) : (q.num : ℚ≥0) / q.denom = q := +begin + ext1, + rw [coe_div, coe_nat_cast, coe_nat_cast, num, ←int.cast_coe_nat, + int.nat_abs_of_nonneg (rat.num_nonneg_iff_zero_le.2 q.prop)], + exact rat.num_div_denom q, +end + +/-- A recursor for nonnegative rationals in terms of numerators and denominators. -/ +protected def rec {α : ℚ≥0 → Sort*} (h : Π m n : ℕ, α (m / n)) (q : ℚ≥0) : α q := +(num_div_denom _).rec (h _ _) + +end nnrat From 7c4a46f3ad6674c4728eb73dd0f2535c423464e7 Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Mon, 29 Aug 2022 06:20:49 +0000 Subject: [PATCH 060/828] feat(measure_theory/function/conditional_expectation/basic): add some condexp lemmas (#16273) --- .../conditional_expectation/basic.lean | 32 +++++++++++++++++++ src/topology/algebra/monoid.lean | 11 +++++++ 2 files changed, 43 insertions(+) diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 2005c88e1d8c9..f98ee0b433525 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -2083,6 +2083,18 @@ begin ((condexp_ae_eq_condexp_L1 hm _).symm.add (condexp_ae_eq_condexp_L1 hm _).symm), end +lemma condexp_finset_sum {ι : Type*} {s : finset ι} {f : ι → α → F'} + (hf : ∀ i ∈ s, integrable (f i) μ) : + μ[∑ i in s, f i | m] =ᵐ[μ] ∑ i in s, μ[f i | m] := +begin + induction s using finset.induction_on with i s his heq hf, + { rw [finset.sum_empty, finset.sum_empty, condexp_zero] }, + { rw [finset.sum_insert his, finset.sum_insert his], + exact (condexp_add (hf i $ finset.mem_insert_self i s) $ integrable_finset_sum' _ + (λ j hmem, hf j $ finset.mem_insert_of_mem hmem)).trans + ((eventually_eq.refl _ _).add (heq $ λ j hmem, hf j $ finset.mem_insert_of_mem hmem)) } +end + lemma condexp_smul (c : 𝕜) (f : α → F') : μ[c • f | m] =ᵐ[μ] c • μ[f|m] := begin by_cases hm : m ≤ m0, @@ -2142,6 +2154,26 @@ begin ((condexp_L1_mono hf hg hfg).trans_eq (condexp_ae_eq_condexp_L1 hm _).symm), end +lemma condexp_nonneg {E} [normed_lattice_add_comm_group E] [complete_space E] [normed_space ℝ E] + [ordered_smul ℝ E] {f : α → E} (hf : 0 ≤ᵐ[μ] f) : + 0 ≤ᵐ[μ] μ[f | m] := +begin + by_cases hfint : integrable f μ, + { rw (condexp_zero.symm : (0 : α → E) = μ[0 | m]), + exact condexp_mono (integrable_zero _ _ _) hfint hf }, + { exact eventually_eq.le (condexp_undef hfint).symm } +end + +lemma condexp_nonpos {E} [normed_lattice_add_comm_group E] [complete_space E] [normed_space ℝ E] + [ordered_smul ℝ E] {f : α → E} (hf : f ≤ᵐ[μ] 0) : + μ[f | m] ≤ᵐ[μ] 0 := +begin + by_cases hfint : integrable f μ, + { rw (condexp_zero.symm : (0 : α → E) = μ[0 | m]), + exact condexp_mono hfint (integrable_zero _ _ _) hf }, + { exact eventually_eq.le (condexp_undef hfint) } +end + /-- **Lebesgue dominated convergence theorem**: sufficient conditions under which almost everywhere convergence of a sequence of functions implies the convergence of their image by `condexp_L1`. -/ diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 51b7f069585b0..f7123d4ceeb80 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro import algebra.big_operators.finprod import data.set.pointwise import topology.algebra.mul_action +import algebra.big_operators.pi /-! # Theory of topological monoids @@ -500,6 +501,16 @@ lemma continuous_finset_prod {f : ι → X → M} (s : finset ι) : (∀ i ∈ s, continuous (f i)) → continuous (λa, ∏ i in s, f i a) := continuous_multiset_prod _ +@[to_additive] lemma eventually_eq_prod {X M : Type*} [comm_monoid M] + {s : finset ι} {l : filter X} {f g : ι → X → M} (hs : ∀ i ∈ s, f i =ᶠ[l] g i) : + ∏ i in s, f i =ᶠ[l] ∏ i in s, g i := +begin + replace hs: ∀ᶠ x in l, ∀ i ∈ s, f i x = g i x, + { rwa eventually_all_finset }, + filter_upwards [hs] with x hx, + simp only [finset.prod_apply, finset.prod_congr rfl hx], +end + open function @[to_additive] From 24a233d4199bf664ed549ffb36813b76b45273e0 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 29 Aug 2022 08:18:19 +0000 Subject: [PATCH 061/828] feat(algebra/order/positive): new file (#14833) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define various algebraic structures on the set of positive numbers. I have two reasons to introduce these classes: - we can use `{x : ℝ // 0 < x}` in various filter bases; this may save us from explicit usage of `half_pos` or `div_pos` here or there; - I want to introduce a multiplicative action of `{x : ℝ // 0 < x}` on the upper half plane. --- src/algebra/order/positive/field.lean | 33 +++++++ src/algebra/order/positive/ring.lean | 118 ++++++++++++++++++++++++++ src/data/pnat/basic.lean | 43 ++-------- 3 files changed, 159 insertions(+), 35 deletions(-) create mode 100644 src/algebra/order/positive/field.lean create mode 100644 src/algebra/order/positive/ring.lean diff --git a/src/algebra/order/positive/field.lean b/src/algebra/order/positive/field.lean new file mode 100644 index 0000000000000..1cdc637185be3 --- /dev/null +++ b/src/algebra/order/positive/field.lean @@ -0,0 +1,33 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import algebra.order.positive.ring +import algebra.field_power + +/-! +# Algebraic structures on the set of positive numbers + +In this file we prove that the set of positive elements of a linear ordered field is a linear +ordered commutative group. +-/ + +variables {K : Type*} [linear_ordered_field K] + +namespace positive + +instance : has_inv {x : K // 0 < x} := ⟨λ x, ⟨x⁻¹, inv_pos.2 x.2⟩⟩ + +@[simp] lemma coe_inv (x : {x : K // 0 < x}) : ↑x⁻¹ = (x⁻¹ : K) := rfl + +instance : has_pow {x : K // 0 < x} ℤ := +⟨λ x n, ⟨x ^ n, zpow_pos_of_pos x.2 _⟩⟩ + +@[simp] lemma coe_zpow (x : {x : K // 0 < x}) (n : ℤ) : ↑(x ^ n) = (x ^ n : K) := rfl + +instance : linear_ordered_comm_group {x : K // 0 < x} := +{ mul_left_inv := λ a, subtype.ext $ inv_mul_cancel a.2.ne', + .. positive.subtype.has_inv, .. positive.subtype.linear_ordered_cancel_comm_monoid } + +end positive diff --git a/src/algebra/order/positive/ring.lean b/src/algebra/order/positive/ring.lean new file mode 100644 index 0000000000000..6fad786474e51 --- /dev/null +++ b/src/algebra/order/positive/ring.lean @@ -0,0 +1,118 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import algebra.order.ring + +/-! +# Algebraic structures on the set of positive numbers + +In this file we define various instances (`add_semigroup`, `ordered_comm_monoid` etc) on the +type `{x : R // 0 < x}`. In each case we try to require the weakest possible typeclass +assumptions on `R` but possibly, there is a room for improvements. +-/ +open function + +namespace positive + +variables {M R K : Type*} + +section add_basic + +variables [add_monoid M] [preorder M] [covariant_class M M (+) (<)] + +instance : has_add {x : M // 0 < x} := ⟨λ x y, ⟨x + y, add_pos x.2 y.2⟩⟩ + +@[simp, norm_cast] lemma coe_add (x y : {x : M // 0 < x}) : ↑(x + y) = (x + y : M) := rfl + +instance : add_semigroup {x : M // 0 < x} := subtype.coe_injective.add_semigroup _ coe_add + +instance {M : Type*} [add_comm_monoid M] [preorder M] [covariant_class M M (+) (<)] : + add_comm_semigroup {x : M // 0 < x} := +subtype.coe_injective.add_comm_semigroup _ coe_add + +instance {M : Type*} [add_left_cancel_monoid M] [preorder M] [covariant_class M M (+) (<)] : + add_left_cancel_semigroup {x : M // 0 < x} := +subtype.coe_injective.add_left_cancel_semigroup _ coe_add + +instance {M : Type*} [add_right_cancel_monoid M] [preorder M] [covariant_class M M (+) (<)] : + add_right_cancel_semigroup {x : M // 0 < x} := +subtype.coe_injective.add_right_cancel_semigroup _ coe_add + +instance covariant_class_add_lt : covariant_class {x : M // 0 < x} {x : M // 0 < x} (+) (<) := +⟨λ x y z hyz, subtype.coe_lt_coe.1 $ add_lt_add_left hyz _⟩ + +instance covariant_class_swap_add_lt [covariant_class M M (swap (+)) (<)] : + covariant_class {x : M // 0 < x} {x : M // 0 < x} (swap (+)) (<) := +⟨λ x y z hyz, subtype.coe_lt_coe.1 $ add_lt_add_right hyz _⟩ + +instance contravariant_class_add_lt [contravariant_class M M (+) (<)] : + contravariant_class {x : M // 0 < x} {x : M // 0 < x} (+) (<) := +⟨λ x y z h, subtype.coe_lt_coe.1 $ lt_of_add_lt_add_left h⟩ + +instance contravariant_class_swap_add_lt [contravariant_class M M (swap (+)) (<)] : + contravariant_class {x : M // 0 < x} {x : M // 0 < x} (swap (+)) (<) := +⟨λ x y z h, subtype.coe_lt_coe.1 $ lt_of_add_lt_add_right h⟩ + +instance contravariant_class_add_le [contravariant_class M M (+) (≤)] : + contravariant_class {x : M // 0 < x} {x : M // 0 < x} (+) (≤) := +⟨λ x y z h, subtype.coe_le_coe.1 $ le_of_add_le_add_left h⟩ + +instance contravariant_class_swap_add_le [contravariant_class M M (swap (+)) (≤)] : + contravariant_class {x : M // 0 < x} {x : M // 0 < x} (swap (+)) (≤) := +⟨λ x y z h, subtype.coe_le_coe.1 $ le_of_add_le_add_right h⟩ + +end add_basic + +instance covariant_class_add_le [add_monoid M] [partial_order M] [covariant_class M M (+) (<)] : + covariant_class {x : M // 0 < x} {x : M // 0 < x} (+) (≤) := +⟨λ x, strict_mono.monotone $ λ _ _ h, add_lt_add_left h _⟩ + +section mul + +variables [ordered_semiring R] + +instance : has_mul {x : R // 0 < x} := ⟨λ x y, ⟨x * y, mul_pos x.2 y.2⟩⟩ + +@[simp] lemma coe_mul (x y : {x : R // 0 < x}) : ↑(x * y) = (x * y : R) := rfl + +instance : has_pow {x : R // 0 < x} ℕ := ⟨λ x n, ⟨x ^ n, pow_pos x.2 n⟩⟩ + +@[simp] lemma coe_pow (x : {x : R // 0 < x}) (n : ℕ) : ↑(x ^ n) = (x ^ n : R) := rfl + +instance : semigroup {x : R // 0 < x} := subtype.coe_injective.semigroup coe coe_mul + +instance : distrib {x : R // 0 < x} := subtype.coe_injective.distrib _ coe_add coe_mul + +instance [nontrivial R] : has_one {x : R // 0 < x} := ⟨⟨1, one_pos⟩⟩ + +@[simp] lemma coe_one [nontrivial R] : ((1 : {x : R // 0 < x}) : R) = 1 := rfl + +instance [nontrivial R] : monoid {x : R // 0 < x} := +subtype.coe_injective.monoid _ coe_one coe_mul coe_pow + +end mul + +section mul_comm + +instance [ordered_comm_semiring R] [nontrivial R] : ordered_comm_monoid {x : R // 0 < x} := +{ mul_le_mul_left := λ x y hxy c, subtype.coe_le_coe.1 $ mul_le_mul_of_nonneg_left hxy c.2.le, + .. subtype.partial_order _, + .. subtype.coe_injective.comm_monoid (coe : {x : R // 0 < x} → R) coe_one coe_mul coe_pow } + +/-- If `R` is a nontrivial linear ordered commutative semiring, then `{x : R // 0 < x}` is a linear +ordered cancellative commutative monoid. We don't have a typeclass for linear ordered commutative +semirings, so we assume `[linear_ordered_semiring R] [is_commutative R (*)] instead. -/ +instance [linear_ordered_semiring R] [is_commutative R (*)] [nontrivial R] : + linear_ordered_cancel_comm_monoid {x : R // 0 < x} := +{ mul_left_cancel := λ a b c h, subtype.ext $ (strict_mono_mul_left_of_pos a.2).injective $ + by convert congr_arg subtype.val h, + le_of_mul_le_mul_left := λ a b c h, subtype.coe_le_coe.1 $ (mul_le_mul_left a.2).1 h, + .. subtype.linear_order _, + .. @positive.subtype.ordered_comm_monoid R + { mul_comm := is_commutative.comm, .. ‹linear_ordered_semiring R› } _ } + +end mul_comm + +end positive diff --git a/src/data/pnat/basic.lean b/src/data/pnat/basic.lean index b778079441a6d..71a4a55f13056 100644 --- a/src/data/pnat/basic.lean +++ b/src/data/pnat/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Neil Strickland -/ import data.nat.basic +import algebra.order.positive.ring /-! # The positive natural numbers @@ -14,7 +15,8 @@ This file defines the type `ℕ+` or `pnat`, the subtype of natural numbers that /-- `ℕ+` is the type of positive natural numbers. It is defined as a subtype, and the VM representation of `ℕ+` is the same as `ℕ` because the proof is not stored. -/ -@[derive linear_order] +@[derive [decidable_eq, add_left_cancel_semigroup, add_right_cancel_semigroup, add_comm_semigroup, + linear_ordered_cancel_comm_monoid, linear_order, has_add, has_mul, has_one, distrib]] def pnat := {n : ℕ // 0 < n} notation `ℕ+` := pnat @@ -100,8 +102,6 @@ open nat subtraction, division and powers. -/ -instance : decidable_eq ℕ+ := λ (a b : ℕ+), by apply_instance - @[simp] lemma mk_le_mk (n k : ℕ) (hn : 0 < n) (hk : 0 < k) : (⟨n, hn⟩ : ℕ+) ≤ ⟨k, hk⟩ ↔ n ≤ k := iff.rfl @@ -125,10 +125,6 @@ lemma coe_injective : function.injective (coe : ℕ+ → ℕ) := subtype.coe_inj @[simp] theorem mk_coe (n h) : ((⟨n, h⟩ : ℕ+) : ℕ) = n := rfl -instance : has_add ℕ+ := ⟨λ a b, ⟨(a + b : ℕ), add_pos a.pos b.pos⟩⟩ - -instance : add_comm_semigroup ℕ+ := coe_injective.add_comm_semigroup coe (λ _ _, rfl) - @[simp] theorem add_coe (m n : ℕ+) : ((m + n : ℕ+) : ℕ) = m + n := rfl /-- `pnat.coe` promoted to an `add_hom`, that is, a morphism which preserves addition. -/ @@ -136,11 +132,10 @@ def coe_add_hom : add_hom ℕ+ ℕ := { to_fun := coe, map_add' := add_coe } -instance : add_left_cancel_semigroup ℕ+ := -coe_injective.add_left_cancel_semigroup coe (λ _ _, rfl) - -instance : add_right_cancel_semigroup ℕ+ := -coe_injective.add_right_cancel_semigroup coe (λ _ _, rfl) +instance : covariant_class ℕ+ ℕ+ (+) (≤) := positive.covariant_class_add_le +instance : covariant_class ℕ+ ℕ+ (+) (<) := positive.covariant_class_add_lt +instance : contravariant_class ℕ+ ℕ+ (+) (≤) := positive.contravariant_class_add_le +instance : contravariant_class ℕ+ ℕ+ (+) (<) := positive.contravariant_class_add_lt /-- An equivalence between `ℕ+` and `ℕ` given by `pnat.nat_pred` and `nat.succ_pnat`. -/ @[simps { fully_applied := ff }] def _root_.equiv.pnat_equiv_nat : ℕ+ ≃ ℕ := @@ -157,22 +152,12 @@ coe_injective.add_right_cancel_semigroup coe (λ _ _, rfl) @[simp] lemma _root_.order_iso.pnat_iso_nat_symm_apply : ⇑order_iso.pnat_iso_nat.symm = nat.succ_pnat := rfl -@[priority 10] -instance : covariant_class ℕ+ ℕ+ ((+)) (≤) := -⟨by { rintro ⟨a, ha⟩ ⟨b, hb⟩ ⟨c, hc⟩, simp [←pnat.coe_le_coe] }⟩ - @[simp] theorem ne_zero (n : ℕ+) : (n : ℕ) ≠ 0 := n.2.ne' theorem to_pnat'_coe {n : ℕ} : 0 < n → (n.to_pnat' : ℕ) = n := succ_pred_eq_of_pos @[simp] theorem coe_to_pnat' (n : ℕ+) : (n : ℕ).to_pnat' = n := eq (to_pnat'_coe n.pos) -instance : has_mul ℕ+ := ⟨λ m n, ⟨m.1 * n.1, mul_pos m.2 n.2⟩⟩ -instance : has_one ℕ+ := ⟨succ_pnat 0⟩ -instance : has_pow ℕ+ ℕ := ⟨λ x n, ⟨x ^ n, pow_pos x.2 n⟩⟩ - -instance : comm_monoid ℕ+ := coe_injective.comm_monoid coe rfl (λ _ _, rfl) (λ _ _, rfl) - theorem lt_add_one_iff : ∀ {a b : ℕ+}, a < b + 1 ↔ a ≤ b := λ a b, nat.lt_add_one_iff @@ -223,8 +208,7 @@ def coe_monoid_hom : ℕ+ →* ℕ := @[simp] lemma coe_coe_monoid_hom : (coe_monoid_hom : ℕ+ → ℕ) = coe := rfl -@[simp] -lemma coe_eq_one_iff {m : ℕ+} : (m : ℕ) = 1 ↔ m = 1 := by rw [← one_coe, coe_inj] +@[simp] lemma coe_eq_one_iff {m : ℕ+} : (m : ℕ) = 1 ↔ m = 1 := subtype.coe_injective.eq_iff' one_coe @[simp] lemma le_one_iff {n : ℕ+} : n ≤ 1 ↔ n = 1 := le_bot_iff @@ -238,17 +222,6 @@ lemma lt_add_right (n m : ℕ+) : n < n + m := (lt_add_left n m).trans_eq (add_c @[simp] theorem pow_coe (m : ℕ+) (n : ℕ) : ((m ^ n : ℕ+) : ℕ) = (m : ℕ) ^ n := rfl -instance : ordered_cancel_comm_monoid ℕ+ := -{ mul_le_mul_left := by { intros, apply nat.mul_le_mul_left, assumption }, - le_of_mul_le_mul_left := by { intros a b c h, apply nat.le_of_mul_le_mul_left h a.property, }, - mul_left_cancel := λ a b c h, by - { replace h := congr_arg (coe : ℕ+ → ℕ) h, - exact eq ((nat.mul_right_inj a.pos).mp h)}, - .. pnat.comm_monoid, - .. pnat.linear_order } - -instance : distrib ℕ+ := coe_injective.distrib coe (λ _ _, rfl) (λ _ _, rfl) - /-- Subtraction a - b is defined in the obvious way when a > b, and by a - b = 1 if a ≤ b. -/ From f0514a8f5aaf33b62a63bdc086c88497e81bed3b Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 29 Aug 2022 08:18:20 +0000 Subject: [PATCH 062/828] feat(category_theory/preadditive): projective iff variants of Yoneda preserve epimorphisms (#15123) --- .../preadditive/injective.lean | 27 +++++++++++++++++ .../preadditive/projective.lean | 29 +++++++++++++++++++ 2 files changed, 56 insertions(+) diff --git a/src/category_theory/preadditive/injective.lean b/src/category_theory/preadditive/injective.lean index 93388285bfa9f..0f4b579d247b3 100644 --- a/src/category_theory/preadditive/injective.lean +++ b/src/category_theory/preadditive/injective.lean @@ -186,6 +186,33 @@ lemma injective_of_adjoint (adj : L ⊣ R) (J : D) [injective J] : injective $ R end adjunction +section preadditive +variables [preadditive C] + +lemma injective_iff_preserves_epimorphisms_preadditive_yoneda_obj (J : C) : + injective J ↔ (preadditive_yoneda.obj J).preserves_epimorphisms := +begin + rw injective_iff_preserves_epimorphisms_yoneda_obj, + refine ⟨λ (h : (preadditive_yoneda.obj J ⋙ (forget _)).preserves_epimorphisms), _, _⟩, + { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_yoneda.obj J) + (forget _) }, + { introI, + exact (infer_instance : (preadditive_yoneda.obj J ⋙ forget _).preserves_epimorphisms) } +end + +lemma injective_iff_preserves_epimorphisms_preadditive_yoneda_obj' (J : C) : + injective J ↔ (preadditive_yoneda_obj J).preserves_epimorphisms := +begin + rw injective_iff_preserves_epimorphisms_yoneda_obj, + refine ⟨λ (h : (preadditive_yoneda_obj J ⋙ (forget _)).preserves_epimorphisms), _, _⟩, + { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_yoneda_obj J) + (forget _) }, + { introI, + exact (infer_instance : (preadditive_yoneda_obj J ⋙ forget _).preserves_epimorphisms) } +end + +end preadditive + section enough_injectives variable [enough_injectives C] diff --git a/src/category_theory/preadditive/projective.lean b/src/category_theory/preadditive/projective.lean index 5fb7122a8b522..ae125a48878db 100644 --- a/src/category_theory/preadditive/projective.lean +++ b/src/category_theory/preadditive/projective.lean @@ -6,6 +6,8 @@ Authors: Markus Himmel, Scott Morrison import algebra.homology.exact import category_theory.types import category_theory.limits.shapes.biproducts +import category_theory.preadditive.yoneda +import algebra.category.Module.epi_mono /-! # Projective objects and categories with enough projectives @@ -134,6 +136,33 @@ lemma projective_iff_preserves_epimorphisms_coyoneda_obj (P : C) : λ h, ⟨λ E X f e he, by exactI (epi_iff_surjective _).1 (infer_instance : epi ((coyoneda.obj (op P)).map e)) f⟩⟩ +section preadditive +variables [preadditive C] + +lemma projective_iff_preserves_epimorphisms_preadditive_coyoneda_obj (P : C) : + projective P ↔ (preadditive_coyoneda.obj (op P)).preserves_epimorphisms := +begin + rw projective_iff_preserves_epimorphisms_coyoneda_obj, + refine ⟨λ (h : (preadditive_coyoneda.obj (op P) ⋙ (forget _)).preserves_epimorphisms), _, _⟩, + { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_coyoneda.obj (op P)) + (forget _) }, + { introI, + exact (infer_instance : (preadditive_coyoneda.obj (op P) ⋙ forget _).preserves_epimorphisms) } +end + +lemma projective_iff_preserves_epimorphisms_preadditive_coyoneda_obj' (P : C) : + projective P ↔ (preadditive_coyoneda_obj (op P)).preserves_epimorphisms := +begin + rw projective_iff_preserves_epimorphisms_coyoneda_obj, + refine ⟨λ (h : (preadditive_coyoneda_obj (op P) ⋙ (forget _)).preserves_epimorphisms), _, _⟩, + { exactI functor.preserves_epimorphisms_of_preserves_of_reflects (preadditive_coyoneda_obj (op P)) + (forget _) }, + { introI, + exact (infer_instance : (preadditive_coyoneda_obj (op P) ⋙ forget _).preserves_epimorphisms) } +end + +end preadditive + section enough_projectives variables [enough_projectives C] From ae93ce85a5691aabb8f6e5f55802f28c4129a9a0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Mon, 29 Aug 2022 08:18:21 +0000 Subject: [PATCH 063/828] feat(topology/algebra/polynomial): add coeff_le_of_roots_le (#15275) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is the proof that, if the roots of a polynomial are bounded, then its coefficients are bounded. More precisely, it is the following statement: ```lean lemma coeff_le_of_roots_le [field F] [normed_field K] {p : F[X]} {f : F →+* K} {B : ℝ} (i : ℕ) (h1 : p.monic) (h2 : splits f p) (h3 : ∀ z ∈ (map f p).roots, ∥z∥ ≤ B) : ∥ (map f p).coeff i ∥ ≤ B^(p.nat_degree - i) * p.nat_degree.choose i ``` From flt-regular --- src/topology/algebra/polynomial.lean | 76 ++++++++++++++++++++++++++++ 1 file changed, 76 insertions(+) diff --git a/src/topology/algebra/polynomial.lean b/src/topology/algebra/polynomial.lean index 6e546ece1fd33..054f74a6854c4 100644 --- a/src/topology/algebra/polynomial.lean +++ b/src/topology/algebra/polynomial.lean @@ -6,6 +6,8 @@ Authors: Robert Y. Lewis import analysis.normed_space.basic import data.polynomial.algebra_map import data.polynomial.inductions +import ring_theory.polynomial.vieta +import field_theory.splitting_field /-! # Polynomials and limits @@ -129,4 +131,78 @@ if hp0 : 0 < degree p then p.continuous.norm.exists_forall_le $ p.tendsto_norm_at_top hp0 tendsto_norm_cocompact_at_top else ⟨p.coeff 0, by rw [eq_C_of_degree_le_zero (le_of_not_gt hp0)]; simp⟩ +section roots + +open_locale polynomial +open_locale nnreal + +variables {F K : Type*} [field F] [normed_field K] + +open multiset + +lemma coeff_le_of_roots_le {p : F[X]} {f : F →+* K} {B : ℝ} (i : ℕ) + (h1 : p.monic) (h2 : splits f p) (h3 : ∀ z ∈ (map f p).roots, ∥z∥ ≤ B) : + ∥ (map f p).coeff i ∥ ≤ B^(p.nat_degree - i) * p.nat_degree.choose i := +begin + have hcd : card (map f p).roots = p.nat_degree := (nat_degree_eq_card_roots h2).symm, + by_cases hB : 0 ≤ B, + { by_cases hi : i ≤ p.nat_degree, + { rw [eq_prod_roots_of_splits h2, monic.def.mp h1, ring_hom.map_one, ring_hom.map_one, one_mul], + rw prod_X_sub_C_coeff, + swap, rwa hcd, + rw [norm_mul, (by norm_num : ∥(-1 : K) ^ (card (map f p).roots - i)∥= 1), one_mul], + apply le_trans (le_sum_of_subadditive norm _ _ _ ), + rotate, exact norm_zero, exact norm_add_le, + rw multiset.map_map, + suffices : ∀ r ∈ multiset.map (norm_hom ∘ prod) + (powerset_len (card (map f p).roots - i) (map f p).roots), r ≤ B^(p.nat_degree - i), + { convert sum_le_sum_sum _ this, + simp only [hi, hcd, multiset.map_const, card_map, card_powerset_len, nat.choose_symm, + sum_repeat, nsmul_eq_mul, mul_comm], }, + { intros r hr, + obtain ⟨t, ht⟩ := multiset.mem_map.mp hr, + have hbounds : ∀ x ∈ (multiset.map norm_hom t), 0 ≤ x ∧ x ≤ B, + { intros x hx, + obtain ⟨z, hz⟩ := multiset.mem_map.mp hx, + rw ←hz.right, + exact ⟨norm_nonneg z, + h3 z (mem_of_le (mem_powerset_len.mp ht.left).left hz.left)⟩, }, + lift B to ℝ≥0 using hB, + lift (multiset.map norm_hom t) to (multiset ℝ≥0) using (λ x hx, (hbounds x hx).left) + with normt hn, + rw (by rw_mod_cast [←ht.right, function.comp_apply, ←prod_hom t norm_hom, ←hn] : + r = normt.prod), + convert multiset.prod_le_pow_card normt _ _, + { rw (_ : card normt = card (multiset.map coe normt)), + rwa [hn, ←hcd, card_map, (mem_powerset_len.mp ht.left).right.symm], + rwa card_map, }, + { intros x hx, + have xmem : (x : ℝ) ∈ multiset.map coe normt := mem_map_of_mem _ hx, + exact (hbounds x xmem).right, }}}, + { push_neg at hi, + rw [nat.choose_eq_zero_of_lt hi, coeff_eq_zero_of_nat_degree_lt, norm_zero], + rw_mod_cast mul_zero, + { rwa monic.nat_degree_map h1, + apply_instance, }}}, + { push_neg at hB, + have noroots : (map f p).roots = 0, + { contrapose! hB, + obtain ⟨z, hz⟩ := exists_mem_of_ne_zero hB, + exact le_trans (norm_nonneg z) (h3 z hz), }, + suffices : p.nat_degree = 0, + { by_cases hi : i = 0, + { rw [this, hi, (monic.nat_degree_eq_zero_iff_eq_one h1).mp this], + simp only [polynomial.map_one, coeff_one_zero, norm_one, pow_zero, nat.choose_self, + nat.cast_one, mul_one], }, + { replace hi := zero_lt_iff.mpr hi, + rw ←this at hi, + rw [nat.choose_eq_zero_of_lt hi, coeff_eq_zero_of_nat_degree_lt, norm_zero], + rw_mod_cast mul_zero, + { rwa monic.nat_degree_map h1, + apply_instance, }}}, + rw [←hcd, noroots, card_zero], }, +end + +end roots + end polynomial From 9c0a5431b82849850bd2d72ac0ae87b758d821fa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 29 Aug 2022 08:18:23 +0000 Subject: [PATCH 064/828] feat(counterexamples/map_floor): `floor`/`ceil` are not preserved under order ring homs (#16198) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit #16025 proves that `⌊f a⌋ = ⌊a⌋` and `⌈f a⌉ = ⌈a⌉` for a **strictly** monotone ring hom `f`. This counterexample shows that this can't be relaxed to `f` monotone. --- counterexamples/map_floor.lean | 126 ++++++++++++++++++ src/data/polynomial/basic.lean | 8 +- .../polynomial/degree/trailing_degree.lean | 7 + 3 files changed, 140 insertions(+), 1 deletion(-) create mode 100644 counterexamples/map_floor.lean diff --git a/counterexamples/map_floor.lean b/counterexamples/map_floor.lean new file mode 100644 index 0000000000000..8c00769a84eaf --- /dev/null +++ b/counterexamples/map_floor.lean @@ -0,0 +1,126 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import algebra.order.hom.ring +import data.polynomial.reverse + +/-! +# Floors and ceils aren't preserved under ordered ring homomorphisms + +Intuitively, if `f : α → β` is an ordered ring homomorphism, then floors and ceils should be +preserved by `f` because: +* `f` preserves the naturals/integers in `α` and `β` because it's a ring hom. +* `f` preserves what's between `n` and `n + 1` because it's monotone. + +However, there is a catch. Potentially something whose floor was `n` could +get mapped to `n + 1`, and this has floor `n + 1`, not `n`. Note that this is at most an off by one +error. + +This pathology disappears if you require `f` to be strictly monotone or `α` to be archimedean. + +## The counterexample + +Consider `ℤ[ε]` (`int_with_epsilons`), the integers with infinitesimals adjoined. This is a linearly +ordered commutative floor ring (`int_with_epsilons.linear_ordered_comm_ring`, +`int_with_epsilons.floor_ring`). + +The map `f : ℤ[ε] → ℤ` that forgets about the epsilons (`int_with_epsilons.forget_epsilons`) is an +ordered ring homomorphism. + +But it does not preserve floors (nor ceils) as `⌊-ε⌋ = -1` while `⌊f (-ε)⌋ = ⌊0⌋ = 0` +(`int_with_epsilons.forget_epsilons_floor_lt`, `int_with_epsilons.lt_forget_epsilons_ceil`). +-/ + +noncomputable theory + +open function int polynomial +open_locale polynomial + +/-- The integers with infinitesimals adjoined. -/ +@[derive [comm_ring, nontrivial, inhabited]] def int_with_epsilon := ℤ[X] + +notation `ℤ[ε]` := int_with_epsilon + +notation `ε` := (X : ℤ[ε]) + +namespace int_with_epsilon + +instance : linear_order ℤ[ε] := linear_order.lift' (to_lex ∘ coeff) coeff_injective + +instance : ordered_add_comm_group ℤ[ε] := +by refine (to_lex.injective.comp coeff_injective).ordered_add_comm_group _ _ _ _ _ _ _; + refl <|> intros; ext; simp [←nsmul_eq_mul, ←zsmul_eq_mul] + +lemma pos_iff {p : ℤ[ε]} : 0 < p ↔ 0 < p.trailing_coeff := +begin + rw trailing_coeff, + refine ⟨_, λ h, + ⟨p.nat_trailing_degree, λ m hm, (coeff_eq_zero_of_lt_nat_trailing_degree hm).symm, h⟩⟩, + rintro ⟨n, hn⟩, + convert hn.2, + exact (nat_trailing_degree_le_of_ne_zero hn.2.ne').antisymm + (le_nat_trailing_degree (by { rintro rfl, cases hn.2.false }) $ λ m hm, (hn.1 _ hm).symm), +end + +instance : linear_ordered_comm_ring ℤ[ε] := +{ zero_le_one := or.inr ⟨0, by simp⟩, + mul_pos := λ p q, by { simp_rw [pos_iff, trailing_coeff_mul], exact mul_pos }, + ..int_with_epsilon.linear_order, ..int_with_epsilon.comm_ring, + ..int_with_epsilon.ordered_add_comm_group, ..int_with_epsilon.nontrivial } + +instance : floor_ring ℤ[ε] := +floor_ring.of_floor _ (λ p, if (p.coeff 0 : ℤ[ε]) ≤ p then p.coeff 0 else p.coeff 0 - 1) $ λ p q, + begin + simp_rw [←not_lt, not_iff_not], + split, + { split_ifs, + { rintro ⟨_ | n, hn⟩, + { refine (sub_one_lt _).trans _, + simpa using hn }, + { dsimp at hn, + simp [hn.1 _ n.zero_lt_succ] } }, + { exact λ h', cast_lt.1 ((not_lt.1 h).trans_lt h') } }, + { split_ifs, + { exact λ h', h.trans_le (cast_le.2 $ sub_one_lt_iff.1 h') }, + { exact λ h', ⟨0, by simpa using h'⟩ } } + end + +/-- The ordered ring homomorphisms from `ℤ[ε]` to `ℤ` that "forgets" the `ε`s. -/ +def forget_epsilons : ℤ[ε] →+*o ℤ := +{ to_fun := λ p, coeff p 0, + map_zero' := coeff_zero _, + map_one' := coeff_one_zero, + map_add' := λ _ _, coeff_add _ _ _, + map_mul' := mul_coeff_zero, + monotone' := monotone_iff_forall_lt.2 begin + rintro p q ⟨n, hn⟩, + cases n, + { exact hn.2.le }, + { exact (hn.1 _ n.zero_lt_succ).le } + end } + +@[simp] lemma forget_epsilons_apply (p : ℤ[ε]) : forget_epsilons p = coeff p 0 := rfl + +/-- The floor of `n - ε` is `n - 1` but its image under `forget_epsilons` is `n`, whose floor is +itself. -/ +lemma forget_epsilons_floor_lt (n : ℤ) : + forget_epsilons ⌊(n - ε : ℤ[ε])⌋ < ⌊forget_epsilons (n - ε)⌋ := +begin + suffices : ⌊(n - ε : ℤ[ε])⌋ = n - 1, + { simp [this] }, + have : (0 : ℤ[ε]) < ε := ⟨1, by simp⟩, + exact (if_neg $ by simp [this]).trans (by simp), +end + +/-- The ceil of `n + ε` is `n + 1` but its image under `forget_epsilons` is `n`, whose ceil is +itself. -/ +lemma lt_forget_epsilons_ceil (n : ℤ) : + ⌈forget_epsilons (n + ε)⌉ < forget_epsilons ⌈(n + ε : ℤ[ε])⌉ := +begin + rw [←neg_lt_neg_iff, ←map_neg, ←cast_neg, ←floor_neg, ←floor_neg, ←map_neg, neg_add', ←cast_neg], + exact forget_epsilons_floor_lt _, +end + +end int_with_epsilon diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index aa015f24edd1f..9a95be5d81596 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -57,7 +57,8 @@ structure polynomial (R : Type*) [semiring R] := of_finsupp :: (to_finsupp : add_monoid_algebra R ℕ) localized "notation R`[X]`:9000 := polynomial R" in polynomial -open finsupp add_monoid_algebra + +open add_monoid_algebra finsupp function open_locale big_operators polynomial namespace polynomial @@ -428,6 +429,11 @@ by rw [X_pow_mul, monomial_mul_X_pow] @[simp] def coeff : R[X] → ℕ → R | ⟨p⟩ := p +lemma coeff_injective : injective (coeff : R[X] → ℕ → R) := +by { rintro ⟨p⟩ ⟨q⟩, simp only [coeff, fun_like.coe_fn_eq, imp_self] } + +@[simp] lemma coeff_inj : p.coeff = q.coeff ↔ p = q := coeff_injective.eq_iff + lemma coeff_monomial : coeff (monomial n a) m = if n = m then a else 0 := by { simp only [←of_finsupp_single, coeff, linear_map.coe_mk], rw finsupp.single_apply } diff --git a/src/data/polynomial/degree/trailing_degree.lean b/src/data/polynomial/degree/trailing_degree.lean index 8499ebabe851e..bef7f53b802b2 100644 --- a/src/data/polynomial/degree/trailing_degree.lean +++ b/src/data/polynomial/degree/trailing_degree.lean @@ -242,6 +242,13 @@ begin exact mem_support_iff.mpr (trailing_coeff_nonzero_iff_nonzero.mpr h), }, end +lemma le_nat_trailing_degree (hp : p ≠ 0) (hn : ∀ m < n, p.coeff m = 0) : + n ≤ p.nat_trailing_degree := +begin + rw nat_trailing_degree_eq_support_min' hp, + exact finset.le_min' _ _ _ (λ m hm, not_lt.1 $ λ hmn, mem_support_iff.1 hm $ hn _ hmn), +end + lemma nat_trailing_degree_le_nat_degree (p : R[X]) : p.nat_trailing_degree ≤ p.nat_degree := begin From 301a6253ef026cb5938343e320d8222f54140289 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 29 Aug 2022 08:18:24 +0000 Subject: [PATCH 065/828] feat(ring_theory/polynomial/basic): generalize `is_domain` to `no_zero_divisors` (#16226) This generalization makes this work on `comm_semiring` instead of `comm_ring`. This also removes a duplicate proof of nontriviality. `mv_polynomial.eq_zero_or_eq_zero_of_mul_eq_zero` has been removed in favor of this instance. --- src/ring_theory/polynomial/basic.lean | 64 ++++++++++----------------- 1 file changed, 24 insertions(+), 40 deletions(-) diff --git a/src/ring_theory/polynomial/basic.lean b/src/ring_theory/polynomial/basic.lean index a7b083ecd4d12..428e8ce20bf52 100644 --- a/src/ring_theory/polynomial/basic.lean +++ b/src/ring_theory/polynomial/basic.lean @@ -992,68 +992,52 @@ instance is_noetherian_ring [fintype σ] [is_noetherian_ring R] : @is_noetherian_ring_of_ring_equiv (mv_polynomial (fin (fintype.card σ)) R) _ _ _ (rename_equiv R (fintype.equiv_fin σ).symm).to_ring_equiv is_noetherian_ring_fin -lemma is_domain_fin_zero (R : Type u) [comm_ring R] [is_domain R] : - is_domain (mv_polynomial (fin 0) R) := -ring_equiv.is_domain R - ((rename_equiv R fin_zero_equiv').to_ring_equiv.trans - (mv_polynomial.is_empty_ring_equiv R pempty)) - /-- Auxiliary lemma: Multivariate polynomials over an integral domain with variables indexed by `fin n` form an integral domain. This fact is proven inductively, and then used to prove the general case without any finiteness hypotheses. -See `mv_polynomial.is_domain` for the general case. -/ -lemma is_domain_fin (R : Type u) [comm_ring R] [is_domain R] : - ∀ (n : ℕ), is_domain (mv_polynomial (fin n) R) -| 0 := is_domain_fin_zero R -| (n+1) := - begin - haveI := is_domain_fin n, - exact ring_equiv.is_domain - (polynomial (mv_polynomial (fin n) R)) - (mv_polynomial.fin_succ_equiv _ n).to_ring_equiv +See `mv_polynomial.no_zero_divisors` for the general case. -/ +lemma no_zero_divisors_fin (R : Type u) [comm_semiring R] [no_zero_divisors R] : + ∀ (n : ℕ), no_zero_divisors (mv_polynomial (fin n) R) +| 0 := (mv_polynomial.is_empty_alg_equiv R _).injective.no_zero_divisors _ (map_zero _) (map_mul _) +| (n+1) := begin + haveI := no_zero_divisors_fin n, + exact (mv_polynomial.fin_succ_equiv R n).injective.no_zero_divisors _ (map_zero _) (map_mul _) end /-- Auxiliary definition: Multivariate polynomials in finitely many variables over an integral domain form an integral domain. -This fact is proven by transport of structure from the `mv_polynomial.is_domain_fin`, +This fact is proven by transport of structure from the `mv_polynomial.no_zero_divisors_fin`, and then used to prove the general case without finiteness hypotheses. -See `mv_polynomial.is_domain` for the general case. -/ -lemma is_domain_fintype (R : Type u) (σ : Type v) [comm_ring R] [fintype σ] - [is_domain R] : is_domain (mv_polynomial σ R) := -@ring_equiv.is_domain _ (mv_polynomial (fin $ fintype.card σ) R) _ _ - (mv_polynomial.is_domain_fin _ _) - (rename_equiv R (fintype.equiv_fin σ)).to_ring_equiv - -protected theorem eq_zero_or_eq_zero_of_mul_eq_zero - {R : Type u} [comm_ring R] [is_domain R] {σ : Type v} - (p q : mv_polynomial σ R) (h : p * q = 0) : p = 0 ∨ q = 0 := +See `mv_polynomial.no_zero_divisors` for the general case. -/ +lemma no_zero_divisors_of_finite (R : Type u) (σ : Type v) [comm_semiring R] [finite σ] + [no_zero_divisors R] : no_zero_divisors (mv_polynomial σ R) := begin + casesI nonempty_fintype σ, + haveI := no_zero_divisors_fin R (fintype.card σ), + exact (rename_equiv R (fintype.equiv_fin σ)).injective.no_zero_divisors _ (map_zero _) (map_mul _) +end + +instance {R : Type u} [comm_semiring R] [no_zero_divisors R] {σ : Type v} : + no_zero_divisors (mv_polynomial σ R) := +⟨λ p q h, begin obtain ⟨s, p, rfl⟩ := exists_finset_rename p, obtain ⟨t, q, rfl⟩ := exists_finset_rename q, have : rename (subtype.map id (finset.subset_union_left s t) : {x // x ∈ s} → {x // x ∈ s ∪ t}) p * rename (subtype.map id (finset.subset_union_right s t) : {x // x ∈ t} → {x // x ∈ s ∪ t}) q = 0, { apply rename_injective _ subtype.val_injective, simpa using h }, - letI := mv_polynomial.is_domain_fintype R {x // x ∈ (s ∪ t)}, + letI := mv_polynomial.no_zero_divisors_of_finite R {x // x ∈ (s ∪ t)}, rw mul_eq_zero at this, cases this; [left, right], all_goals { simpa using congr_arg (rename subtype.val) this } -end +end⟩ /-- The multivariate polynomial ring over an integral domain is an integral domain. -/ -instance {R : Type u} {σ : Type v} [comm_ring R] [is_domain R] : - is_domain (mv_polynomial σ R) := -{ eq_zero_or_eq_zero_of_mul_eq_zero := mv_polynomial.eq_zero_or_eq_zero_of_mul_eq_zero, - exists_pair_ne := ⟨0, 1, λ H, - begin - have : eval₂ (ring_hom.id _) (λ s, (0:R)) (0 : mv_polynomial σ R) = - eval₂ (ring_hom.id _) (λ s, (0:R)) (1 : mv_polynomial σ R), - { congr, exact H }, - simpa, - end⟩, - .. (by apply_instance : comm_ring (mv_polynomial σ R)) } +instance {R : Type u} {σ : Type v} [comm_ring R] [is_domain R] : is_domain (mv_polynomial σ R) := +{ .. mv_polynomial.no_zero_divisors, + .. add_monoid_algebra.nontrivial } lemma map_mv_polynomial_eq_eval₂ {S : Type*} [comm_ring S] [fintype σ] (ϕ : mv_polynomial σ R →+* S) (p : mv_polynomial σ R) : From 31297bdaf3803162f3fe86b99f688d3995d4f08b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 29 Aug 2022 08:18:25 +0000 Subject: [PATCH 066/828] refactor(algebra/order/with_zero): Generalize lemmas (#16240) Generalize the lemmas introduced in #15644 to monoids + covariant classes. --- src/algebra/order/monoid_lemmas.lean | 25 +++++++++++++++++++++++++ src/algebra/order/with_zero.lean | 20 ++++---------------- 2 files changed, 29 insertions(+), 16 deletions(-) diff --git a/src/algebra/order/monoid_lemmas.lean b/src/algebra/order/monoid_lemmas.lean index 441534f08ae2f..dede9ac014afc 100644 --- a/src/algebra/order/monoid_lemmas.lean +++ b/src/algebra/order/monoid_lemmas.lean @@ -758,6 +758,31 @@ iff.intro and.intro ‹a = 1› ‹b = 1›) (assume ⟨ha', hb'⟩, by rw [ha', hb', mul_one]) +section left +variables [covariant_class α α (*) (≤)] {a b : α} + +@[to_additive eq_zero_of_add_nonneg_left] +lemma eq_one_of_one_le_mul_left (ha : a ≤ 1) (hb : b ≤ 1) (hab : 1 ≤ a * b) : a = 1 := +ha.eq_of_not_lt $ λ h, hab.not_lt $ mul_lt_one_of_lt_of_le h hb + +@[to_additive] +lemma eq_one_of_mul_le_one_left (ha : 1 ≤ a) (hb : 1 ≤ b) (hab : a * b ≤ 1) : a = 1 := +ha.eq_of_not_gt $ λ h, hab.not_lt $ one_lt_mul_of_lt_of_le' h hb + +end left + +section right +variables [covariant_class α α (swap (*)) (≤)] {a b : α} + +@[to_additive eq_zero_of_add_nonneg_right] +lemma eq_one_of_one_le_mul_right (ha : a ≤ 1) (hb : b ≤ 1) (hab : 1 ≤ a * b) : b = 1 := +hb.eq_of_not_lt $ λ h, hab.not_lt $ right.mul_lt_one_of_le_of_lt ha h + +@[to_additive] +lemma eq_one_of_mul_le_one_right (ha : 1 ≤ a) (hb : 1 ≤ b) (hab : a * b ≤ 1) : b = 1 := +hb.eq_of_not_gt $ λ h, hab.not_lt $ right.one_lt_mul_of_le_of_lt ha h + +end right end partial_order section linear_order diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index e7b9859484e26..888b1a41efa30 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -111,8 +111,12 @@ variables [linear_ordered_comm_group_with_zero α] lemma zero_lt_one₀ : (0 : α) < 1 := lt_of_le_of_ne zero_le_one zero_ne_one +-- TODO: Do we really need the following two? + +/-- Alias of `mul_le_one'` for unification. -/ lemma mul_le_one₀ (ha : a ≤ 1) (hb : b ≤ 1) : a * b ≤ 1 := mul_le_one' ha hb +/-- Alias of `one_le_mul'` for unification. -/ lemma one_le_mul₀ (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := one_le_mul ha hb lemma le_of_le_mul_right (h : c ≠ 0) (hab : a * c ≤ b * c) : a ≤ b := @@ -202,22 +206,6 @@ by rw [div_eq_mul_inv, le_mul_inv_iff₀ hc] lemma div_le_iff₀ (hc : c ≠ 0) : a / c ≤ b ↔ a ≤ b*c := by rw [div_eq_mul_inv, mul_inv_le_iff₀ hc] -lemma eq_one_of_mul_eq_one_left (ha : a ≤ 1) (hb : b ≤ 1) (hab : a * b = 1) : a = 1 := -le_antisymm ha $ (inv_le_one₀ $ left_ne_zero_of_mul_eq_one hab).mp $ - eq_inv_of_mul_eq_one_right hab ▸ hb - -lemma eq_one_of_mul_eq_one_right (ha : a ≤ 1) (hb : b ≤ 1) (hab : a * b = 1) : b = 1 := -le_antisymm hb $ (inv_le_one₀ $ right_ne_zero_of_mul_eq_one hab).mp $ - eq_inv_of_mul_eq_one_left hab ▸ ha - -lemma eq_one_of_mul_eq_one_left' (ha : 1 ≤ a) (hb : 1 ≤ b) (hab : a * b = 1) : a = 1 := -le_antisymm - ((one_le_inv₀ $ left_ne_zero_of_mul_eq_one hab).mp $ eq_inv_of_mul_eq_one_right hab ▸ hb) ha - -lemma eq_one_of_mul_eq_one_right' (ha : 1 ≤ a) (hb : 1 ≤ b) (hab : a * b = 1) : b = 1 := -le_antisymm - ((one_le_inv₀ $ right_ne_zero_of_mul_eq_one hab).mp $ eq_inv_of_mul_eq_one_left hab ▸ ha) hb - /-- `equiv.mul_left₀` as an order_iso on a `linear_ordered_comm_group_with_zero.`. Note that `order_iso.mul_left₀` refers to the `linear_ordered_field` version. -/ From 3e2fb4c7e9dde300694e54f3f6ec7c6029a71f21 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 29 Aug 2022 08:18:26 +0000 Subject: [PATCH 067/828] feat(data/rat/cast): generalize `ext` lemmas (#16268) * generalize `monoid_with_zero_hom.ext_rat` to `monoid_with_zero`; * deduce `ext` lemma for `ring_hom`s from `monoid_with_zero_hom` version; * use hom classes. --- src/data/rat/cast.lean | 58 ++++++++----------- .../witt_vector/structure_polynomial.lean | 2 +- 2 files changed, 24 insertions(+), 36 deletions(-) diff --git a/src/data/rat/cast.lean b/src/data/rat/cast.lean index 325ca348c8e4e..b37b323ebb688 100644 --- a/src/data/rat/cast.lean +++ b/src/data/rat/cast.lean @@ -300,53 +300,41 @@ by rw [cast_def, map_div₀, map_int_cast, map_nat_cast, cast_def] @[simp] lemma eq_rat_cast {k} [division_ring k] [ring_hom_class F ℚ k] (f : F) (r : ℚ) : f r = r := by rw [← map_rat_cast f, rat.cast_id] -lemma ring_hom.ext_rat {R : Type*} [semiring R] (f g : ℚ →+* R) : f = g := -begin - ext r, - refine rat.num_denom_cases_on' r _, - intros a b b0, - let φ : ℤ →+* R := f.comp (int.cast_ring_hom ℚ), - let ψ : ℤ →+* R := g.comp (int.cast_ring_hom ℚ), - rw [rat.mk_eq_div, int.cast_coe_nat], - have b0' : (b:ℚ) ≠ 0 := nat.cast_ne_zero.2 b0, - have : ∀ n : ℤ, f n = g n := λ n, show φ n = ψ n, by rw [φ.ext_int ψ], - calc f (a * b⁻¹) - = f a * f b⁻¹ * (g (b:ℤ) * g b⁻¹) : - by rw [int.cast_coe_nat, ← g.map_mul, mul_inv_cancel b0', g.map_one, mul_one, f.map_mul] - ... = g a * f b⁻¹ * (f (b:ℤ) * g b⁻¹) : by rw [this a, ← this b] - ... = g (a * b⁻¹) : - by rw [int.cast_coe_nat, mul_assoc, ← mul_assoc (f b⁻¹), - ← f.map_mul, inv_mul_cancel b0', f.map_one, one_mul, g.map_mul] -end - -instance rat.subsingleton_ring_hom {R : Type*} [semiring R] : subsingleton (ℚ →+* R) := -⟨ring_hom.ext_rat⟩ - namespace monoid_with_zero_hom -variables {M : Type*} [group_with_zero M] +variables {M₀ : Type*} [monoid_with_zero M₀] [monoid_with_zero_hom_class F ℚ M₀] {f g : F} +include M₀ + +/-- If `f` and `g` agree on the integers then they are equal `φ`. -/ +theorem ext_rat' (h : ∀ m : ℤ, f m = g m) : f = g := +fun_like.ext f g $ λ r, by rw [← r.num_div_denom, div_eq_mul_inv, map_mul, map_mul, h, + ← int.cast_coe_nat, eq_on_inv₀ f g (h _)] /-- If `f` and `g` agree on the integers then they are equal `φ`. See note [partially-applied ext lemmas] for why `comp` is used here. -/ -@[ext] -theorem ext_rat {f g : ℚ →*₀ M} - (same_on_int : f.comp (int.cast_ring_hom ℚ).to_monoid_with_zero_hom = - g.comp (int.cast_ring_hom ℚ).to_monoid_with_zero_hom) : f = g := -begin - have same_on_int' : ∀ k : ℤ, f k = g k := congr_fun same_on_int, - ext x, - rw [← @rat.num_denom x, rat.mk_eq_div, map_div₀ f, map_div₀ g, - same_on_int' x.num, same_on_int' x.denom], -end +@[ext] theorem ext_rat {f g : ℚ →*₀ M₀} + (h : f.comp (int.cast_ring_hom ℚ : ℤ →*₀ ℚ) = g.comp (int.cast_ring_hom ℚ)) : f = g := +ext_rat' $ congr_fun h /-- Positive integer values of a morphism `φ` and its value on `-1` completely determine `φ`. -/ -theorem ext_rat_on_pnat {f g : ℚ →*₀ M} +theorem ext_rat_on_pnat (same_on_neg_one : f (-1) = g (-1)) (same_on_pnat : ∀ n : ℕ, 0 < n → f n = g n) : f = g := -ext_rat $ ext_int' (by simpa) ‹_› +ext_rat' $ fun_like.congr_fun $ show (f : ℚ →*₀ M₀).comp (int.cast_ring_hom ℚ : ℤ →*₀ ℚ) = + (g : ℚ →*₀ M₀).comp (int.cast_ring_hom ℚ : ℤ →*₀ ℚ), + from ext_int' (by simpa) (by simpa) end monoid_with_zero_hom +/-- Any two ring homomorphisms from `ℚ` to a semiring are equal. If the codomain is a division ring, +then this lemma follows from `eq_rat_cast`. -/ +lemma ring_hom.ext_rat {R : Type*} [semiring R] [ring_hom_class F ℚ R] (f g : F) : f = g := +monoid_with_zero_hom.ext_rat' $ ring_hom.congr_fun $ + ((f : ℚ →+* R).comp (int.cast_ring_hom ℚ)).ext_int ((g : ℚ →+* R).comp (int.cast_ring_hom ℚ)) + +instance rat.subsingleton_ring_hom {R : Type*} [semiring R] : subsingleton (ℚ →+* R) := +⟨ring_hom.ext_rat⟩ + namespace mul_opposite variables [division_ring α] diff --git a/src/ring_theory/witt_vector/structure_polynomial.lean b/src/ring_theory/witt_vector/structure_polynomial.lean index 68ae7f2962086..44075d444f732 100644 --- a/src/ring_theory/witt_vector/structure_polynomial.lean +++ b/src/ring_theory/witt_vector/structure_polynomial.lean @@ -133,7 +133,7 @@ theorem witt_structure_rat_prop (Φ : mv_polynomial idx ℚ) (n : ℕ) : calc bind₁ (witt_structure_rat p Φ) (W_ ℚ n) = bind₁ (λ k, bind₁ (λ i, (rename (prod.mk i)) (W_ ℚ k)) Φ) (bind₁ (X_in_terms_of_W p ℚ) (W_ ℚ n)) : - by { rw bind₁_bind₁, apply eval₂_hom_congr (ring_hom.ext_rat _ _) rfl rfl } + by { rw bind₁_bind₁, exact eval₂_hom_congr (ring_hom.ext_rat _ _) rfl rfl } ... = bind₁ (λ i, (rename (prod.mk i) (W_ ℚ n))) Φ : by rw [bind₁_X_in_terms_of_W_witt_polynomial p _ n, bind₁_X_right] From 39d5a98cf1ef0a1e2c2dd2b8ef13877401d9810a Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 29 Aug 2022 08:18:27 +0000 Subject: [PATCH 068/828] feat(linear_algebra/affine_space/affine_equiv): lemmas about coercion to `affine_map` (#16284) Add two more `rfl` (and `simp`) lemmas about the coercion from `affine_equiv` to `affine_map`. In both cases very similar lemmas are already present (in one case a version using `to_affine_map`, in one case a version for the coercion to a function), but apparently not these particular ones. --- src/linear_algebra/affine_space/affine_equiv.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/linear_algebra/affine_space/affine_equiv.lean b/src/linear_algebra/affine_space/affine_equiv.lean index 646f3b099e444..0f740a267e3e8 100644 --- a/src/linear_algebra/affine_space/affine_equiv.lean +++ b/src/linear_algebra/affine_space/affine_equiv.lean @@ -86,6 +86,8 @@ rfl @[simp] lemma linear_to_affine_map (e : P₁ ≃ᵃ[k] P₂) : e.to_affine_map.linear = e.linear := rfl +@[simp] lemma coe_linear (e : P₁ ≃ᵃ[k] P₂) : (e : P₁ →ᵃ[k] P₂).linear = e.linear := rfl + lemma to_affine_map_injective : injective (to_affine_map : (P₁ ≃ᵃ[k] P₂) → (P₁ →ᵃ[k] P₂)) := begin rintros ⟨e, el, h⟩ ⟨e', el', h'⟩ H, @@ -206,6 +208,10 @@ include V₂ V₃ @[simp] lemma coe_trans (e : P₁ ≃ᵃ[k] P₂) (e' : P₂ ≃ᵃ[k] P₃) : ⇑(e.trans e') = e' ∘ e := rfl +@[simp] lemma coe_trans_to_affine_map (e : P₁ ≃ᵃ[k] P₂) (e' : P₂ ≃ᵃ[k] P₃) : + (e.trans e' : P₁ →ᵃ[k] P₃) = (e' : P₂ →ᵃ[k] P₃).comp e := +rfl + @[simp] lemma trans_apply (e : P₁ ≃ᵃ[k] P₂) (e' : P₂ ≃ᵃ[k] P₃) (p : P₁) : e.trans e' p = e' (e p) := rfl From 211bf1abab2f253791c6b98c2631c68cfbcefc43 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 29 Aug 2022 08:18:28 +0000 Subject: [PATCH 069/828] feat(linear_algebra/affine_space/affine_subspace): `map_eq_bot_iff` (#16285) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add the lemma that the image of an affine subspace under an affine map is `⊥` if and only if that subspace is `⊥`. --- src/linear_algebra/affine_space/affine_subspace.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index b65614f1633c2..bad91a54e2b2c 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -1241,6 +1241,13 @@ def map (s : affine_subspace k P₁) : affine_subspace k P₂ := @[simp] lemma map_bot : (⊥ : affine_subspace k P₁).map f = ⊥ := coe_injective $ image_empty f +@[simp] lemma map_eq_bot_iff {s : affine_subspace k P₁} : s.map f = ⊥ ↔ s = ⊥ := +begin + refine ⟨λ h, _, λ h, _⟩, + { rwa [←coe_eq_bot_iff, coe_map, image_eq_empty, coe_eq_bot_iff] at h }, + { rw [h, map_bot] } +end + omit V₂ @[simp] lemma map_id (s : affine_subspace k P₁) : s.map (affine_map.id k P₁) = s := From 8ad70779a8f2bc26acc3271ef1c222a75aeabad7 Mon Sep 17 00:00:00 2001 From: Stuart Presnell Date: Mon, 29 Aug 2022 11:13:00 +0000 Subject: [PATCH 070/828] feat(data/nat/factorization/basic): lemmas about `n.factorization p = 0` (#16185) Adds some lemmas characterising conditions when `n.factorization p = 0`. This PR also rearranges the order of some lemmas to better group them together (and adds some section docstrings). Also swaps the names `factorization_eq_zero_iff` and `factorization_eq_zero_iff'`. --- src/data/nat/factorization/basic.lean | 93 ++++++++++++++++++--------- 1 file changed, 61 insertions(+), 32 deletions(-) diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 1ed144fcb8d43..2ba7d488a9ef9 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -119,9 +119,23 @@ prime.pos (prime_of_mem_factorization hp) lemma le_of_mem_factorization {n p : ℕ} (h : p ∈ n.factorization.support) : p ≤ n := le_of_mem_factors (factor_iff_mem_factorization.mp h) +/-! ## Lemmas characterising when `n.factorization p = 0` -/ + +lemma factorization_eq_zero_iff (n p : ℕ) : + n.factorization p = 0 ↔ ¬p.prime ∨ ¬p ∣ n ∨ n = 0 := +begin + rw [←not_mem_support_iff, support_factorization, mem_to_finset], + rcases eq_or_ne n 0 with rfl | hn, + { simp }, + { simp [hn, nat.mem_factors, not_and_distrib] }, +end + @[simp] lemma factorization_eq_zero_of_non_prime (n : ℕ) {p : ℕ} (hp : ¬p.prime) : n.factorization p = 0 := -not_mem_support_iff.1 (mt prime_of_mem_factorization hp) +by simp [factorization_eq_zero_iff, hp] + +lemma factorization_eq_zero_of_not_dvd {n p : ℕ} (h : ¬ p ∣ n) : n.factorization p = 0 := +by simp [factorization_eq_zero_iff, h] lemma factorization_eq_zero_of_lt {n p : ℕ} (h : n < p) : n.factorization p = 0 := finsupp.not_mem_support_iff.mp (mt le_of_mem_factorization (not_le_of_lt h)) @@ -139,21 +153,29 @@ lemma prime.factorization_pos_of_dvd {n p : ℕ} (hp : p.prime) (hn : n ≠ 0) ( 0 < n.factorization p := by rwa [←factors_count_eq, count_pos, mem_factors_iff_dvd hn hp] +lemma factorization_eq_zero_of_remainder {p r : ℕ} (i : ℕ) (hr : ¬ p ∣ r) : + (p * i + r).factorization p = 0 := +by { apply factorization_eq_zero_of_not_dvd, rwa ←nat.dvd_add_iff_right (dvd.intro i rfl) } + +lemma factorization_eq_zero_iff_remainder {p r : ℕ} (i : ℕ) (pp : p.prime) (hr0 : r ≠ 0) : + (¬ p ∣ r) ↔ (p * i + r).factorization p = 0 := +begin + refine ⟨factorization_eq_zero_of_remainder i, λ h, _⟩, + rw factorization_eq_zero_iff at h, + contrapose! h, + refine ⟨pp, _, _⟩, + { rwa ←nat.dvd_add_iff_right ((dvd.intro i rfl)) }, + { contrapose! hr0, exact (_root_.add_eq_zero_iff.mp hr0).2 }, +end + /-- The only numbers with empty prime factorization are `0` and `1` -/ -lemma factorization_eq_zero_iff (n : ℕ) : n.factorization = 0 ↔ n = 0 ∨ n = 1 := +lemma factorization_eq_zero_iff' (n : ℕ) : n.factorization = 0 ↔ n = 0 ∨ n = 1 := begin rw factorization_eq_factors_multiset n, simp [factorization, add_equiv.map_eq_zero_iff, multiset.coe_eq_zero], end -lemma factorization_eq_zero_iff' (n p : ℕ) : - n.factorization p = 0 ↔ ¬p.prime ∨ ¬p ∣ n ∨ n = 0 := -begin - rw [←not_mem_support_iff, support_factorization, mem_to_finset], - rcases eq_or_ne n 0 with rfl | hn, - { simp }, - { simp [hn, nat.mem_factors, not_and_distrib] }, -end +/-! ## Lemmas about factorizations of products and powers -/ /-- For nonzero `a` and `b`, the power of `p` in `a * b` is the sum of the powers in `a` and `b` -/ @[simp] lemma factorization_mul {a b : ℕ} (ha : a ≠ 0) (hb : b ≠ 0) : @@ -169,6 +191,26 @@ begin exact mem_factors_mul ha hb end +/-- If a product over `n.factorization` doesn't use the multiplicities of the prime factors +then it's equal to the corresponding product over `n.factors.to_finset` -/ +lemma prod_factorization_eq_prod_factors {n : ℕ} {β : Type*} [comm_monoid β] (f : ℕ → β) : + n.factorization.prod (λ p k, f p) = ∏ p in n.factors.to_finset, (f p) := +by { apply prod_congr support_factorization, simp } + +/-- For any `p : ℕ` and any function `g : α → ℕ` that's non-zero on `S : finset α`, +the power of `p` in `S.prod g` equals the sum over `x ∈ S` of the powers of `p` in `g x`. +Generalises `factorization_mul`, which is the special case where `S.card = 2` and `g = id`. -/ +lemma factorization_prod {α : Type*} {S : finset α} {g : α → ℕ} (hS : ∀ x ∈ S, g x ≠ 0) : + (S.prod g).factorization = S.sum (λ x, (g x).factorization) := +begin + classical, + ext p, + apply finset.induction_on' S, { simp }, + { intros x T hxS hTS hxT IH, + have hT : T.prod g ≠ 0 := prod_ne_zero_iff.mpr (λ x hx, hS x (hTS hx)), + simp [prod_insert hxT, sum_insert hxT, ←IH, factorization_mul (hS x hxS) hT] } +end + /-- For any `p`, the power of `p` in `n^k` is `k` times the power in `n` -/ @[simp] lemma factorization_pow (n k : ℕ) : factorization (n^k) = k • n.factorization := @@ -178,6 +220,8 @@ begin rw [pow_succ, factorization_mul hn (pow_ne_zero _ hn), ih, succ_eq_one_add, add_smul, one_smul], end +/-! ## Lemmas about factorizations of primes and prime powers -/ + /-- The only prime factor of prime `p` is `p` itself, with multiplicity `1` -/ @[simp] lemma prime.factorization {p : ℕ} (hp : prime p) : p.factorization = single p 1 := @@ -187,7 +231,7 @@ begin refl, end -/-- The only prime factor of prime `p` is `p` itself, with multiplicity `1` -/ +/-- The multiplicity of prime `p` in `p` is `1` -/ @[simp] lemma prime.factorization_self {p : ℕ} (hp : prime p) : p.factorization p = 1 := by simp [hp] @@ -201,25 +245,10 @@ lemma eq_pow_of_factorization_eq_single {n p k : ℕ} (hn : n ≠ 0) (h : n.factorization = finsupp.single p k) : n = p ^ k := by { rw [←nat.factorization_prod_pow_eq_self hn, h], simp } -/-- If a product over `n.factorization` doesn't use the multiplicities of the prime factors -then it's equal to the corresponding product over `n.factors.to_finset` -/ -lemma prod_factorization_eq_prod_factors {n : ℕ} {β : Type*} [comm_monoid β] (f : ℕ → β) : - n.factorization.prod (λ p k, f p) = ∏ p in n.factors.to_finset, (f p) := -by { apply prod_congr support_factorization, simp } - -/-- For any `p : ℕ` and any function `g : α → ℕ` that's non-zero on `S : finset α`, -the power of `p` in `S.prod g` equals the sum over `x ∈ S` of the powers of `p` in `g x`. -Generalises `factorization_mul`, which is the special case where `S.card = 2` and `g = id`. -/ -lemma factorization_prod {α : Type*} {S : finset α} {g : α → ℕ} (hS : ∀ x ∈ S, g x ≠ 0) : - (S.prod g).factorization = S.sum (λ x, (g x).factorization) := -begin - classical, - ext p, - apply finset.induction_on' S, { simp }, - { intros x T hxS hTS hxT IH, - have hT : T.prod g ≠ 0 := prod_ne_zero_iff.mpr (λ x hx, hS x (hTS hx)), - simp [prod_insert hxT, sum_insert hxT, ←IH, factorization_mul (hS x hxS) hT] } -end +/-- The only prime factor of prime `p` is `p` itself. -/ +lemma prime.eq_of_factorization_pos {p q : ℕ} (hp : prime p) (h : p.factorization q ≠ 0) : + p = q := +by simpa [hp.factorization, single_apply] using h /-! ### Equivalence between `ℕ+` and `ℕ →₀ ℕ` with support in the primes. -/ @@ -433,7 +462,7 @@ begin by_cases pp : p.prime, swap, { simp [pp] }, ext q, rcases eq_or_ne q p with rfl | hqp, - { simp only [finsupp.erase_same, factorization_eq_zero_iff', not_dvd_ord_compl pp hn], + { simp only [finsupp.erase_same, factorization_eq_zero_iff, not_dvd_ord_compl pp hn], simp }, { rw [finsupp.erase_ne hqp, factorization_div (ord_proj_dvd n p)], simp [pp.factorization, hqp.symm] }, @@ -448,7 +477,7 @@ begin rw [←(factorization_le_iff_dvd hd0 (ord_compl_pos p hn0).ne'), factorization_ord_compl], intro q, rcases eq_or_ne q p with rfl | hqp, - { simp [factorization_eq_zero_iff', hpd] }, + { simp [factorization_eq_zero_iff, hpd] }, { simp [hqp, (factorization_le_iff_dvd hd0 hn0).2 hdn q] }, end From e4b9af2635651b67fa4fdfbb002aa67a8e27aa6d Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 29 Aug 2022 11:13:01 +0000 Subject: [PATCH 071/828] feat(algebra/char_p/basic + data/zmod/basic): the characteristic is the order of one and zmod related lemmas (#16277) [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F) --- src/algebra/char_p/basic.lean | 3 +++ src/data/zmod/basic.lean | 19 +++++++++++++++++++ 2 files changed, 22 insertions(+) diff --git a/src/algebra/char_p/basic.lean b/src/algebra/char_p/basic.lean index adaa870d5c620..697516969afa9 100644 --- a/src/algebra/char_p/basic.lean +++ b/src/algebra/char_p/basic.lean @@ -42,6 +42,9 @@ theorem char_p.cast_eq_zero [add_monoid_with_one R] (p : ℕ) [char_p R p] : (fintype.card R : R) = 0 := by rw [← nsmul_one, card_nsmul_eq_zero] +lemma char_p.add_order_of_one (R) [semiring R] : char_p R (add_order_of (1 : R)) := +⟨λ n, by rw [← nat.smul_one_eq_coe, add_order_of_dvd_iff_nsmul_eq_zero]⟩ + lemma char_p.int_cast_eq_zero_iff [add_group_with_one R] (p : ℕ) [char_p R p] (a : ℤ) : (a : R) = 0 ↔ (p:ℤ) ∣ a := diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index 76ef1f9022cf9..cdd04c40094c9 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -81,6 +81,25 @@ instance (n : ℕ) : char_p (zmod n) n := rw [val_nat_cast, val_zero, nat.dvd_iff_mod_eq_zero], end } +@[simp] lemma add_order_of_one (n : ℕ) : add_order_of (1 : zmod n) = n := +char_p.eq _ (char_p.add_order_of_one _) (zmod.char_p n) + +/-- This lemma works in the case in which `zmod n` is not infinite, i.e. `n ≠ 0`. The version +where `a ≠ 0` is `add_order_of_coe'`. -/ +@[simp] lemma add_order_of_coe (a : ℕ) {n : ℕ} (n0 : n ≠ 0) : + add_order_of (a : zmod n) = n / n.gcd a := +begin + cases a, + simp [nat.pos_of_ne_zero n0], + rw [← nat.smul_one_eq_coe, add_order_of_nsmul' _ a.succ_ne_zero, zmod.add_order_of_one], +end + +/-- This lemma works in the case in which `a ≠ 0`. The version where + `zmod n` is not infinite, i.e. `n ≠ 0`, is `add_order_of_coe`. -/ +@[simp] lemma add_order_of_coe' {a : ℕ} (n : ℕ) (a0 : a ≠ 0) : + add_order_of (a : zmod n) = n / n.gcd a := +by rw [← nat.smul_one_eq_coe, add_order_of_nsmul' _ a0, zmod.add_order_of_one] + /-- We have that `ring_char (zmod n) = n`. -/ lemma ring_char_zmod_n (n : ℕ) : ring_char (zmod n) = n := by { rw ring_char.eq_iff, exact zmod.char_p n, } From 24b08ce66ce674bec1f7836fce389ebed6d80acb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 29 Aug 2022 13:10:47 +0000 Subject: [PATCH 072/828] feat(order/heyting/regular): Heyting-regular elements (#16033) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define Heyting-regular elements of an Heyting lattice, namely those `a` such that `aᶜᶜ = a`, and prove that they form a boolean algebra. --- src/order/heyting/regular.lean | 157 +++++++++++++++++++++++++++++++++ 1 file changed, 157 insertions(+) create mode 100644 src/order/heyting/regular.lean diff --git a/src/order/heyting/regular.lean b/src/order/heyting/regular.lean new file mode 100644 index 0000000000000..41684eb983ffc --- /dev/null +++ b/src/order/heyting/regular.lean @@ -0,0 +1,157 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.galois_connection + +/-! +# Heyting regular elements + +This file defines Heyting regular elements, elements of an Heyting algebra that are their own double +complement, and proves that they form a boolean algebra. + +From a logic standpoint, this means that we can perform classical logic within intuitionistic logic +by simply double-negating all propositions. This is practical for synthetic computability theory. + +## Main declarations + +* `is_regular`: `a` is Heyting-regular if `aᶜᶜ = a`. +* `regular`: The subtype of Heyting-regular elements. +* `regular.boolean_algebra`: Heyting-regular elements form a boolean algebra. + +## References + +* [Francis Borceux, *Handbook of Categorical Algebra III*][borceux-vol3] +-/ + +open function + +variables {α : Type*} + +namespace heyting +section has_compl +variables [has_compl α] {a : α} + +/-- An element of an Heyting algebra is regular if its double complement is itself. -/ +def is_regular (a : α) : Prop := aᶜᶜ = a + +protected lemma is_regular.eq : is_regular a → aᶜᶜ = a := id + +instance is_regular.decidable_pred [decidable_eq α] : @decidable_pred α is_regular := +λ _, ‹decidable_eq α› _ _ + +end has_compl + +section heyting_algebra +variables [heyting_algebra α] {a b : α} + +lemma is_regular_bot : is_regular (⊥ : α) := by rw [is_regular, compl_bot, compl_top] +lemma is_regular_top : is_regular (⊤ : α) := by rw [is_regular, compl_top, compl_bot] + +lemma is_regular.inf (ha : is_regular a) (hb : is_regular b) : is_regular (a ⊓ b) := +by rw [is_regular, compl_compl_inf_distrib, ha.eq, hb.eq] + +lemma is_regular.himp (ha : is_regular a) (hb : is_regular b) : is_regular (a ⇨ b) := +by rw [is_regular, compl_compl_himp_distrib, ha.eq, hb.eq] + +lemma is_regular_compl (a : α) : is_regular aᶜ := compl_compl_compl _ + +protected lemma is_regular.disjoint_compl_left_iff (ha : is_regular a) : disjoint aᶜ b ↔ b ≤ a := +by rw [←le_compl_iff_disjoint_left, ha.eq] + +protected lemma is_regular.disjoint_compl_right_iff (hb : is_regular b) : disjoint a bᶜ ↔ a ≤ b := +by rw [←le_compl_iff_disjoint_right, hb.eq] + +/-- A Heyting algebra with regular excluded middle is a boolean algebra. -/ +@[reducible] -- See note [reducible non-instances] +def _root_.boolean_algebra.of_regular (h : ∀ a : α, is_regular (a ⊔ aᶜ)) : boolean_algebra α := +have ∀ a : α, is_compl a aᶜ := λ a, ⟨disjoint_compl_right, codisjoint_iff.2 $ + by erw [←(h a).eq, compl_sup, inf_compl_eq_bot, compl_bot]⟩, +{ himp_eq := λ a b, eq_of_forall_le_iff $ λ c, + le_himp_iff.trans ((this _).le_sup_right_iff_inf_left_le).symm, + inf_compl_le_bot := λ a, (this _).1, + top_le_sup_compl := λ a, (this _).2, + ..‹heyting_algebra α›, ..generalized_heyting_algebra.to_distrib_lattice } + +variables (α) + +/-- The boolean algebra of Heyting regular elements. -/ +def regular : Type* := {a : α // is_regular a} + +variables {α} + +namespace regular + +instance : has_coe (regular α) α := coe_subtype + +lemma coe_injective : injective (coe : regular α → α) := subtype.coe_injective +@[simp] lemma coe_inj {a b : regular α} : (a : α) = b ↔ a = b := subtype.coe_inj + +instance : has_top (regular α) := ⟨⟨⊤, is_regular_top⟩⟩ +instance : has_bot (regular α) := ⟨⟨⊥, is_regular_bot⟩⟩ +instance : has_inf (regular α) := ⟨λ a b, ⟨a ⊓ b, a.2.inf b.2⟩⟩ +instance : has_himp (regular α) := ⟨λ a b, ⟨a ⇨ b, a.2.himp b.2⟩⟩ +instance : has_compl (regular α) := ⟨λ a, ⟨aᶜ, is_regular_compl _⟩⟩ + +@[simp, norm_cast] lemma coe_top : ((⊤ : regular α) : α) = ⊤ := rfl +@[simp, norm_cast] lemma coe_bot : ((⊥ : regular α) : α) = ⊥ := rfl +@[simp, norm_cast] lemma coe_inf (a b : regular α) : (↑(a ⊓ b) : α) = a ⊓ b := rfl +@[simp, norm_cast] lemma coe_himp (a b : regular α) : (↑(a ⇨ b) : α) = a ⇨ b := rfl +@[simp, norm_cast] lemma coe_compl (a : regular α) : (↑(aᶜ) : α) = aᶜ := rfl + +instance : inhabited (regular α) := ⟨⊥⟩ +instance : semilattice_inf (regular α) := coe_injective.semilattice_inf _ coe_inf +instance : bounded_order (regular α) := bounded_order.lift coe (λ _ _, id) coe_top coe_bot + +@[simp, norm_cast] lemma coe_le_coe {a b : regular α} : (a : α) ≤ b ↔ a ≤ b := iff.rfl +@[simp, norm_cast] lemma coe_lt_coe {a b : regular α} : (a : α) < b ↔ a < b := iff.rfl + +/-- **Regularization** of `a`. The smallest regular element greater than `a`. -/ +def to_regular : α →o regular α := +⟨λ a, ⟨aᶜᶜ, is_regular_compl _⟩, λ a b h, coe_le_coe.1 $ compl_le_compl $ compl_le_compl h⟩ + +@[simp, norm_cast] lemma coe_to_regular (a : α) : (to_regular a : α) = aᶜᶜ := rfl +@[simp] lemma to_regular_coe (a : regular α) : to_regular (a : α) = a := coe_injective a.2 + +/-- The Galois insertion between `regular.to_regular` and `coe`. -/ +def gi : galois_insertion to_regular (coe : regular α → α) := +{ choice := λ a ha, ⟨a, ha.antisymm le_compl_compl⟩, + gc := λ a b, coe_le_coe.symm.trans $ + ⟨le_compl_compl.trans, λ h, (compl_anti $ compl_anti h).trans_eq b.2⟩, + le_l_u := λ _, le_compl_compl, + choice_eq := λ a ha, coe_injective $ le_compl_compl.antisymm ha } + +instance : lattice (regular α) := gi.lift_lattice + +@[simp, norm_cast] lemma coe_sup (a b : regular α) : (↑(a ⊔ b) : α) = (a ⊔ b)ᶜᶜ := rfl + +instance : boolean_algebra (regular α) := +{ le_sup_inf := λ a b c, coe_le_coe.1 $ by { dsimp, rw [sup_inf_left, compl_compl_inf_distrib] }, + inf_compl_le_bot := λ a, coe_le_coe.1 disjoint_compl_right, + top_le_sup_compl := λ a, coe_le_coe.1 $ + by { dsimp, rw [compl_sup, inf_compl_eq_bot, compl_bot], refl }, + himp_eq := λ a b, coe_injective begin + dsimp, + rw [compl_sup, a.prop.eq], + refine eq_of_forall_le_iff (λ c, le_himp_iff.trans _), + rw [le_compl_iff_disjoint_right, disjoint_left_comm, b.prop.disjoint_compl_left_iff], + end, + ..regular.lattice, ..regular.bounded_order, ..regular.has_himp, + ..regular.has_compl } + +@[simp, norm_cast] lemma coe_sdiff (a b : regular α) : (↑(a \ b) : α) = a ⊓ bᶜ := rfl + +end regular +end heyting_algebra + +variables [boolean_algebra α] + +lemma is_regular_of_boolean : ∀ a : α, is_regular a := compl_compl + +/-- A decidable proposition is intuitionistically Heyting-regular. -/ +@[nolint decidable_classical] +lemma is_regular_of_decidable (p : Prop) [decidable p] : is_regular p := +propext $ decidable.not_not_iff _ + +end heyting From daec7d610330e8ae491fb09dc3bad1f9529e85b2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 29 Aug 2022 13:10:48 +0000 Subject: [PATCH 073/828] chore(order/bounded_order): Rename `is_complemented` to `complemented_lattice` (#16263) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is to make space for the predicate `is_complemented : α → Prop := ∃ b, is_compl a b`. --- src/linear_algebra/basis.lean | 2 +- src/order/atoms.lean | 14 ++++++++------ src/order/boolean_algebra.lean | 3 ++- src/order/bounded_order.lean | 14 +++++++------- src/order/compactly_generated.lean | 17 +++++++++-------- src/order/hom/basic.lean | 10 +++++----- src/order/modular_lattice.lean | 10 +++++----- src/representation_theory/maschke.lean | 2 +- src/ring_theory/simple_module.lean | 6 +++--- 9 files changed, 41 insertions(+), 37 deletions(-) diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index 6086acf42aaa6..f32ebb284f918 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -1307,7 +1307,7 @@ lemma submodule.exists_is_compl (p : submodule K V) : ∃ q : submodule K V, is_ let ⟨f, hf⟩ := p.subtype.exists_left_inverse_of_injective p.ker_subtype in ⟨f.ker, linear_map.is_compl_of_proj $ linear_map.ext_iff.1 hf⟩ -instance module.submodule.is_complemented : is_complemented (submodule K V) := +instance module.submodule.complemented_lattice : complemented_lattice (submodule K V) := ⟨submodule.exists_is_compl⟩ lemma linear_map.exists_right_inverse_of_surjective (f : V →ₗ[K] V') diff --git a/src/order/atoms.lean b/src/order/atoms.lean index a4281d76cac69..7435fec154fd2 100644 --- a/src/order/atoms.lean +++ b/src/order/atoms.lean @@ -636,9 +636,10 @@ lemma is_coatom_iff_is_atom : is_coatom a ↔ is_atom b := hc.symm.is_atom_iff_i end is_compl -variables [is_complemented α] +variables [complemented_lattice α] -lemma is_coatomic_of_is_atomic_of_is_complemented_of_is_modular [is_atomic α] : is_coatomic α := +lemma is_coatomic_of_is_atomic_of_complemented_lattice_of_is_modular [is_atomic α] : + is_coatomic α := ⟨λ x, begin rcases exists_is_compl x with ⟨y, xy⟩, apply (eq_bot_or_exists_atom_le y).imp _ _, @@ -651,12 +652,13 @@ lemma is_coatomic_of_is_atomic_of_is_complemented_of_is_modular [is_atomic α] : apply ha.Iic } end⟩ -lemma is_atomic_of_is_coatomic_of_is_complemented_of_is_modular [is_coatomic α] : is_atomic α := -is_coatomic_dual_iff_is_atomic.1 is_coatomic_of_is_atomic_of_is_complemented_of_is_modular +lemma is_atomic_of_is_coatomic_of_complemented_lattice_of_is_modular [is_coatomic α] : + is_atomic α := +is_coatomic_dual_iff_is_atomic.1 is_coatomic_of_is_atomic_of_complemented_lattice_of_is_modular theorem is_atomic_iff_is_coatomic : is_atomic α ↔ is_coatomic α := -⟨λ h, @is_coatomic_of_is_atomic_of_is_complemented_of_is_modular _ _ _ _ _ h, - λ h, @is_atomic_of_is_coatomic_of_is_complemented_of_is_modular _ _ _ _ _ h⟩ +⟨λ h, @is_coatomic_of_is_atomic_of_complemented_lattice_of_is_modular _ _ _ _ _ h, + λ h, @is_atomic_of_is_coatomic_of_complemented_lattice_of_is_modular _ _ _ _ _ h⟩ end is_modular_lattice diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index 265df62ad5559..aadd18dcdade3 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -559,7 +559,8 @@ lemma himp_eq : x ⇨ y = y ⊔ xᶜ := boolean_algebra.himp_eq x y @[simp] lemma top_sdiff : ⊤ \ x = xᶜ := by rw [sdiff_eq, top_inf_eq] @[priority 100] -instance boolean_algebra.to_is_complemented : is_complemented α := ⟨λ x, ⟨xᶜ, is_compl_compl⟩⟩ +instance boolean_algebra.to_complemented_lattice : complemented_lattice α := +⟨λ x, ⟨xᶜ, is_compl_compl⟩⟩ @[priority 100] -- see Note [lower instance priority] instance boolean_algebra.to_generalized_boolean_algebra : generalized_boolean_algebra α := diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index 92f3ecad3af54..95e91ab0e5085 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -24,7 +24,7 @@ instances for `Prop` and `fun`. * `with_ α`: Equips `option α` with the order on `α` plus `none` as the top/bottom element. * `is_compl x y`: In a bounded lattice, predicate for "`x` is a complement of `y`". Note that in a non distributive lattice, an element can have several complements. -* `is_complemented α`: Typeclass stating that any element of a lattice has a complement. +* `complemented_lattice α`: Typeclass stating that any element of a lattice has a complement. ## Common lattices @@ -1724,18 +1724,18 @@ end /-- A complemented bounded lattice is one where every element has a (not necessarily unique) complement. -/ -class is_complemented (α) [lattice α] [bounded_order α] : Prop := +class complemented_lattice (α) [lattice α] [bounded_order α] : Prop := (exists_is_compl : ∀ (a : α), ∃ (b : α), is_compl a b) -export is_complemented (exists_is_compl) +export complemented_lattice (exists_is_compl) -namespace is_complemented -variables [lattice α] [bounded_order α] [is_complemented α] +namespace complemented_lattice +variables [lattice α] [bounded_order α] [complemented_lattice α] -instance : is_complemented αᵒᵈ := +instance : complemented_lattice αᵒᵈ := ⟨λ a, let ⟨b, hb⟩ := exists_is_compl (show α, from a) in ⟨b, hb.dual⟩⟩ -end is_complemented +end complemented_lattice end is_compl diff --git a/src/order/compactly_generated.lean b/src/order/compactly_generated.lean index 6eb7b15458e1d..10087eaa9b514 100644 --- a/src/order/compactly_generated.lean +++ b/src/order/compactly_generated.lean @@ -437,7 +437,7 @@ section variables [is_modular_lattice α] [is_compactly_generated α] @[priority 100] -instance is_atomic_of_is_complemented [is_complemented α] : is_atomic α := +instance is_atomic_of_complemented_lattice [complemented_lattice α] : is_atomic α := ⟨λ b, begin by_cases h : {c : α | complete_lattice.is_compact_element c ∧ c ≤ b} ⊆ {⊥}, { left, @@ -459,7 +459,7 @@ end⟩ /-- See Lemma 5.1, Călugăreanu -/ @[priority 100] -instance is_atomistic_of_is_complemented [is_complemented α] : is_atomistic α := +instance is_atomistic_of_complemented_lattice [complemented_lattice α] : is_atomistic α := ⟨λ b, ⟨{a | is_atom a ∧ a ≤ b}, begin symmetry, have hle : Sup {a : α | is_atom a ∧ a ≤ b} ≤ b := (Sup_le $ λ _, and.right), @@ -475,7 +475,8 @@ instance is_atomistic_of_is_complemented [is_complemented α] : is_atomistic α end, λ _, and.left⟩⟩ /-- See Theorem 6.6, Călugăreanu -/ -theorem is_complemented_of_Sup_atoms_eq_top (h : Sup {a : α | is_atom a} = ⊤) : is_complemented α := +theorem complemented_lattice_of_Sup_atoms_eq_top (h : Sup {a : α | is_atom a} = ⊤) : + complemented_lattice α := ⟨λ b, begin obtain ⟨s, ⟨s_ind, b_inf_Sup_s, s_atoms⟩, s_max⟩ := zorn_subset {s : set α | complete_lattice.set_independent s ∧ b ⊓ Sup s = ⊥ ∧ ∀ a ∈ s, is_atom a} _, @@ -525,14 +526,14 @@ theorem is_complemented_of_Sup_atoms_eq_top (h : Sup {a : α | is_atom a} = ⊤) end⟩ /-- See Theorem 6.6, Călugăreanu -/ -theorem is_complemented_of_is_atomistic [is_atomistic α] : is_complemented α := -is_complemented_of_Sup_atoms_eq_top Sup_atoms_eq_top +theorem complemented_lattice_of_is_atomistic [is_atomistic α] : complemented_lattice α := +complemented_lattice_of_Sup_atoms_eq_top Sup_atoms_eq_top -theorem is_complemented_iff_is_atomistic : is_complemented α ↔ is_atomistic α := +theorem complemented_lattice_iff_is_atomistic : complemented_lattice α ↔ is_atomistic α := begin split; introsI, - { exact is_atomistic_of_is_complemented }, - { exact is_complemented_of_is_atomistic } + { exact is_atomistic_of_complemented_lattice }, + { exact complemented_lattice_of_is_atomistic } end end diff --git a/src/order/hom/basic.lean b/src/order/hom/basic.lean index 5607f8afbb337..4779aea5d9daf 100644 --- a/src/order/hom/basic.lean +++ b/src/order/hom/basic.lean @@ -925,17 +925,17 @@ theorem order_iso.is_compl_iff {x y : α} : is_compl x y ↔ is_compl (f x) (f y) := ⟨f.is_compl, λ h, f.symm_apply_apply x ▸ f.symm_apply_apply y ▸ f.symm.is_compl h⟩ -lemma order_iso.is_complemented - [is_complemented α] : is_complemented β := +lemma order_iso.complemented_lattice + [complemented_lattice α] : complemented_lattice β := ⟨λ x, begin obtain ⟨y, hy⟩ := exists_is_compl (f.symm x), rw ← f.symm_apply_apply y at hy, refine ⟨f y, f.symm.is_compl_iff.2 hy⟩, end⟩ -theorem order_iso.is_complemented_iff : - is_complemented α ↔ is_complemented β := -⟨by { introI, exact f.is_complemented }, by { introI, exact f.symm.is_complemented }⟩ +theorem order_iso.complemented_lattice_iff : + complemented_lattice α ↔ complemented_lattice β := +⟨by { introI, exact f.complemented_lattice }, by { introI, exact f.symm.complemented_lattice }⟩ end bounded_order end lattice_isos diff --git a/src/order/modular_lattice.lean b/src/order/modular_lattice.lean index 2be73b526f6ca..278327a0625cf 100644 --- a/src/order/modular_lattice.lean +++ b/src/order/modular_lattice.lean @@ -310,10 +310,10 @@ instance is_modular_lattice_Iic : is_modular_lattice (set.Iic a) := instance is_modular_lattice_Ici : is_modular_lattice (set.Ici a) := ⟨λ x y z xz, (sup_inf_le_assoc_of_le (y : α) xz : (↑x ⊔ ↑y) ⊓ ↑z ≤ ↑x ⊔ ↑y ⊓ ↑z)⟩ -section is_complemented -variables [bounded_order α] [is_complemented α] +section complemented_lattice +variables [bounded_order α] [complemented_lattice α] -instance is_complemented_Iic : is_complemented (set.Iic a) := +instance complemented_lattice_Iic : complemented_lattice (set.Iic a) := ⟨λ ⟨x, hx⟩, let ⟨y, hy⟩ := exists_is_compl x in ⟨⟨y ⊓ a, set.mem_Iic.2 inf_le_right⟩, begin split, @@ -324,7 +324,7 @@ instance is_complemented_Iic : is_complemented (set.Iic a) := rw [← sup_inf_assoc_of_le _ (set.mem_Iic.1 hx), top_le_iff.1 hy.2, top_inf_eq] } end⟩⟩ -instance is_complemented_Ici : is_complemented (set.Ici a) := +instance complemented_lattice_Ici : complemented_lattice (set.Ici a) := ⟨λ ⟨x, hx⟩, let ⟨y, hy⟩ := exists_is_compl x in ⟨⟨y ⊔ a, set.mem_Ici.2 le_sup_right⟩, begin split, @@ -335,6 +335,6 @@ instance is_complemented_Ici : is_complemented (set.Ici a) := exact le_trans hy.2 le_sup_left } end⟩⟩ -end is_complemented +end complemented_lattice end is_modular_lattice diff --git a/src/representation_theory/maschke.lean b/src/representation_theory/maschke.lean index cb4e0aa7bf7c4..38d7d320ee236 100644 --- a/src/representation_theory/maschke.lean +++ b/src/representation_theory/maschke.lean @@ -184,7 +184,7 @@ let ⟨f, hf⟩ := monoid_algebra.exists_left_inverse_of_injective p.subtype p.k ⟨f.ker, linear_map.is_compl_of_proj $ linear_map.ext_iff.1 hf⟩ /-- This also implies an instance `is_semisimple_module (monoid_algebra k G) V`. -/ -instance is_complemented : is_complemented (submodule (monoid_algebra k G) V) := +instance complemented_lattice : complemented_lattice (submodule (monoid_algebra k G) V) := ⟨exists_is_compl⟩ end submodule diff --git a/src/ring_theory/simple_module.lean b/src/ring_theory/simple_module.lean index 1baeee6b8877b..9036f849b7885 100644 --- a/src/ring_theory/simple_module.lean +++ b/src/ring_theory/simple_module.lean @@ -34,7 +34,7 @@ abbreviation is_simple_module := (is_simple_order (submodule R M)) /-- A module is semisimple when every submodule has a complement, or equivalently, the module is a direct sum of simple modules. -/ -abbreviation is_semisimple_module := (is_complemented (submodule R M)) +abbreviation is_semisimple_module := (complemented_lattice (submodule R M)) -- Making this an instance causes the linter to complain of "dangerous instances" theorem is_simple_module.nontrivial [is_simple_module R M] : nontrivial M := @@ -67,7 +67,7 @@ end is_simple_module theorem is_semisimple_of_Sup_simples_eq_top (h : Sup {m : submodule R M | is_simple_module R m} = ⊤) : is_semisimple_module R M := -is_complemented_of_Sup_atoms_eq_top (by simp_rw [← h, is_simple_module_iff_is_atom]) +complemented_lattice_of_Sup_atoms_eq_top (by simp_rw [← h, is_simple_module_iff_is_atom]) namespace is_semisimple_module @@ -82,7 +82,7 @@ end instance is_semisimple_submodule {m : submodule R M} : is_semisimple_module R m := begin have f : submodule R m ≃o set.Iic m := submodule.map_subtype.rel_iso m, - exact f.is_complemented_iff.2 is_modular_lattice.is_complemented_Iic, + exact f.complemented_lattice_iff.2 is_modular_lattice.complemented_lattice_Iic, end end is_semisimple_module From 1b8a0c2932932c938c6ea5ecd672aeecbd2ec5c7 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 29 Aug 2022 13:10:49 +0000 Subject: [PATCH 074/828] chore(set_theory/cardinal/ordinal): follow your linter (#16265) Editing `set_theory.cardinal.ordinal`, I noticed that the linter was unhappy. So, I thought of pleasing it. I followed its instructions, but, admittedly, I do not know if this is a good change or not! I replaced some `fintype` assumptions by `finite` ones and changed one proof as a consequence. --- src/set_theory/cardinal/ordinal.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/src/set_theory/cardinal/ordinal.lean b/src/set_theory/cardinal/ordinal.lean index 6dffdf258ff18..b9a334642767b 100644 --- a/src/set_theory/cardinal/ordinal.lean +++ b/src/set_theory/cardinal/ordinal.lean @@ -888,11 +888,12 @@ lemma mk_compl_eq_mk_compl_infinite {α : Type*} [infinite α] {s t : set α} (h (ht : #t < #α) : #(sᶜ : set α) = #(tᶜ : set α) := by { rw [mk_compl_of_infinite s hs, mk_compl_of_infinite t ht] } -lemma mk_compl_eq_mk_compl_finite_lift {α : Type u} {β : Type v} [fintype α] +lemma mk_compl_eq_mk_compl_finite_lift {α : Type u} {β : Type v} [finite α] {s : set α} {t : set β} (h1 : lift.{max v w} (#α) = lift.{max u w} (#β)) (h2 : lift.{max v w} (#s) = lift.{max u w} (#t)) : lift.{max v w} (#(sᶜ : set α)) = lift.{max u w} (#(tᶜ : set β)) := begin + casesI nonempty_fintype α, rcases lift_mk_eq.1 h1 with ⟨e⟩, letI : fintype β := fintype.of_equiv α e, replace h1 : fintype.card α = fintype.card β := (fintype.of_equiv_card _).symm, classical, @@ -903,11 +904,11 @@ begin lift_nat_cast, nat.cast_inj, h1, h2] end -lemma mk_compl_eq_mk_compl_finite {α β : Type u} [fintype α] {s : set α} {t : set β} +lemma mk_compl_eq_mk_compl_finite {α β : Type u} [finite α] {s : set α} {t : set β} (h1 : #α = #β) (h : #s = #t) : #(sᶜ : set α) = #(tᶜ : set β) := by { rw ← lift_inj, apply mk_compl_eq_mk_compl_finite_lift; rwa [lift_inj] } -lemma mk_compl_eq_mk_compl_finite_same {α : Type*} [fintype α] {s t : set α} +lemma mk_compl_eq_mk_compl_finite_same {α : Type*} [finite α] {s t : set α} (h : #s = #t) : #(sᶜ : set α) = #(tᶜ : set α) := mk_compl_eq_mk_compl_finite rfl h @@ -924,7 +925,7 @@ begin refine ⟨h, _⟩, rintro ⟨x, hx⟩, simp [set.sum_compl_symm_apply_of_mem, hx] end -theorem extend_function_finite {α β : Type*} [fintype α] {s : set α} (f : s ↪ β) +theorem extend_function_finite {α β : Type*} [finite α] {s : set α} (f : s ↪ β) (h : nonempty (α ≃ β)) : ∃ (g : α ≃ β), ∀ x : s, g x = f x := begin apply extend_function f, From 7e20c8fc0b87eba49f283281286a57dd9eaa27aa Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Mon, 29 Aug 2022 16:05:59 +0000 Subject: [PATCH 075/828] feat(topology): composition of continuous functions is continuous w.r.t. the compact-open topologies (#15721) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Remove the `t2_space` assumption from `exists_compact_between` by generalizing the proof of `exists_compact_superset`, and move it from `topology/separation` to `topology/subset_properties`. * Use it to prove `continuous_map.continuous_prod` in `topology/continuous_map` stating that for topological spaces `X,Y,Z` with `Y` locally compact, the composition induces a continuous map from `C(X,Y) x C(Y,Z)` to `C(X,Z)` where all function spaces are endowed with the compact-open topology. The (statement and the) proof is taken from Prop. 9 of Chap. X, §3, №. 4 of Bourbaki's *Topologie Générale*. * Generalize `exists_open_between_and_is_compact_closure` from `t3_space` to `t2_space` using the generalized `exists_compact_between`. This has been briefly discussed in [this Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/continuity.20of.20continuous.20composition) Co-authored-by: Oliver Nash --- src/topology/compact_open.lean | 22 +++++++++++++ src/topology/separation.lean | 50 ++++++++--------------------- src/topology/subset_properties.lean | 25 +++++++++------ 3 files changed, 51 insertions(+), 46 deletions(-) diff --git a/src/topology/compact_open.lean b/src/topology/compact_open.lean index dc8278b0d546e..2c41791bd4e02 100644 --- a/src/topology/compact_open.lean +++ b/src/topology/compact_open.lean @@ -103,6 +103,28 @@ lemma continuous_comp_left : continuous (λ g, g.comp f : C(β, γ) → C(α, γ continuous_generated_from $ assume m ⟨s, hs, u, hu, hm⟩, by { rw [hm, image_gen f hs hu], exact continuous_map.is_open_gen (hs.image f.2) hu } +/-- Composition is a continuous map from `C(α, β) × C(β, γ)` to `C(α, γ)`, provided that `β` is + locally compact. This is Prop. 9 of Chap. X, §3, №. 4 of Bourbaki's *Topologie Générale*. -/ +lemma continuous_comp' [locally_compact_space β] : + continuous (λ x : C(α, β) × C(β, γ), x.2.comp x.1) := +continuous_generated_from begin + rintros M ⟨K, hK, U, hU, rfl⟩, + conv { congr, rw [compact_open.gen, preimage_set_of_eq], + congr, funext, rw [coe_comp, image_comp, image_subset_iff] }, + rw is_open_iff_forall_mem_open, + rintros ⟨φ₀, ψ₀⟩ H, + obtain ⟨L, hL, hKL, hLU⟩ := exists_compact_between (hK.image φ₀.2) (hU.preimage ψ₀.2) H, + use {φ : C(α, β) | φ '' K ⊆ interior L} ×ˢ {ψ : C(β, γ) | ψ '' L ⊆ U}, + use λ ⟨φ, ψ⟩ ⟨hφ, hψ⟩, subset_trans hφ (interior_subset.trans $ image_subset_iff.mp hψ), + use (continuous_map.is_open_gen hK is_open_interior).prod (continuous_map.is_open_gen hL hU), + exact mem_prod.mpr ⟨hKL, image_subset_iff.mpr hLU⟩, +end + +lemma continuous.comp' {X : Type*} [topological_space X] [locally_compact_space β] + {f : X → C(α, β)} {g : X → C(β, γ)} (hf : continuous f) (hg : continuous g) : + continuous (λ x, (g x).comp (f x)) := +continuous_comp'.comp (hf.prod_mk hg : continuous $ λ x, (f x, g x)) + end functorial section ev diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 83f4695af14c6..cf3cdedc41555 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -1228,6 +1228,20 @@ begin compact_closure_of_subset_compact hK' interior_subset⟩, end +/-- +In a locally compact T₂ space, given a compact set `K` inside an open set `U`, we can find a +open set `V` between these sets with compact closure: `K ⊆ V` and the closure of `V` is inside `U`. +-/ +lemma exists_open_between_and_is_compact_closure [locally_compact_space α] [t2_space α] + {K U : set α} (hK : is_compact K) (hU : is_open U) (hKU : K ⊆ U) : + ∃ V, is_open V ∧ K ⊆ V ∧ closure V ⊆ U ∧ is_compact (closure V) := +begin + rcases exists_compact_between hK hU hKU with ⟨V, hV, hKV, hVU⟩, + exact ⟨interior V, is_open_interior, hKV, + (closure_minimal interior_subset hV.is_closed).trans hVU, + compact_closure_of_subset_compact hV interior_subset⟩, +end + lemma is_preirreducible_iff_subsingleton [t2_space α] (S : set α) : is_preirreducible S ↔ S.subsingleton := begin @@ -1349,42 +1363,6 @@ begin tauto end -/-- -In a locally compact T₃ space, given a compact set `K` inside an open set `U`, we can find a -compact set `K'` between these sets: `K` is inside the interior of `K'` and `K' ⊆ U`. --/ -lemma exists_compact_between [locally_compact_space α] [t3_space α] - {K U : set α} (hK : is_compact K) (hU : is_open U) (hKU : K ⊆ U) : - ∃ K', is_compact K' ∧ K ⊆ interior K' ∧ K' ⊆ U := -begin - choose C hxC hCU hC using λ x : K, nhds_is_closed (hU.mem_nhds $ hKU x.2), - choose L hL hxL using λ x : K, exists_compact_mem_nhds (x : α), - have : K ⊆ ⋃ x, interior (L x) ∩ interior (C x), from - λ x hx, mem_Union.mpr ⟨⟨x, hx⟩, - ⟨mem_interior_iff_mem_nhds.mpr (hxL _), mem_interior_iff_mem_nhds.mpr (hxC _)⟩⟩, - rcases hK.elim_finite_subcover _ _ this with ⟨t, ht⟩, - { refine ⟨⋃ x ∈ t, L x ∩ C x, t.compact_bUnion (λ x _, (hL x).inter_right (hC x)), λ x hx, _, _⟩, - { obtain ⟨y, hyt, hy : x ∈ interior (L y) ∩ interior (C y)⟩ := mem_Union₂.mp (ht hx), - rw [← interior_inter] at hy, - refine interior_mono (subset_bUnion_of_mem hyt) hy }, - { simp_rw [Union_subset_iff], rintro x -, exact (inter_subset_right _ _).trans (hCU _) } }, - { exact λ _, is_open_interior.inter is_open_interior } -end - -/-- -In a locally compact regular space, given a compact set `K` inside an open set `U`, we can find a -open set `V` between these sets with compact closure: `K ⊆ V` and the closure of `V` is inside `U`. --/ -lemma exists_open_between_and_is_compact_closure [locally_compact_space α] [t3_space α] - {K U : set α} (hK : is_compact K) (hU : is_open U) (hKU : K ⊆ U) : - ∃ V, is_open V ∧ K ⊆ V ∧ closure V ⊆ U ∧ is_compact (closure V) := -begin - rcases exists_compact_between hK hU hKU with ⟨V, hV, hKV, hVU⟩, - refine ⟨interior V, is_open_interior, hKV, - (closure_minimal interior_subset hV.is_closed).trans hVU, - compact_closure_of_subset_compact hV interior_subset⟩, -end - end t3 section normality diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index 0d2f4d6c59084..8dbe23d765f0b 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -1097,19 +1097,24 @@ lemma exists_compact_mem_nhds [locally_compact_space α] (x : α) : let ⟨K, hKc, hx, H⟩ := exists_compact_subset is_open_univ (mem_univ x) in ⟨K, hKc, mem_interior_iff_mem_nhds.1 hx⟩ +/-- In a locally compact space, for every containement `K ⊆ U` of a compact set `K` in an open + set `U`, there is a compact neighborhood `L` such that `K ⊆ L ⊆ U`: equivalently, there is a + compact `L` such that `K ⊆ interior L` and `L ⊆ U`. -/ +lemma exists_compact_between [hα : locally_compact_space α] {K U : set α} (hK : is_compact K) + (hU : is_open U) (h_KU : K ⊆ U) : ∃ L, is_compact L ∧ K ⊆ interior L ∧ L ⊆ U := +begin + choose V hVc hxV hKV using λ x : K, exists_compact_subset hU (h_KU x.2), + have : K ⊆ ⋃ x, interior (V x), from λ x hx, mem_Union.2 ⟨⟨x, hx⟩, hxV _⟩, + rcases hK.elim_finite_subcover _ (λ x, @is_open_interior α _ (V x)) this with ⟨t, ht⟩, + refine ⟨_, t.compact_bUnion (λ x _, hVc x), λ x hx, _, set.Union₂_subset (λ i _, hKV i)⟩, + rcases mem_Union₂.1 (ht hx) with ⟨y, hyt, hy⟩, + exact interior_mono (subset_bUnion_of_mem hyt) hy, +end + /-- In a locally compact space, every compact set is contained in the interior of a compact set. -/ lemma exists_compact_superset [locally_compact_space α] {K : set α} (hK : is_compact K) : ∃ K', is_compact K' ∧ K ⊆ interior K' := -begin - choose U hUc hxU using λ x : K, exists_compact_mem_nhds (x : α), - have : K ⊆ ⋃ x, interior (U x), - from λ x hx, mem_Union.2 ⟨⟨x, hx⟩, mem_interior_iff_mem_nhds.2 (hxU _)⟩, - rcases hK.elim_finite_subcover _ _ this with ⟨t, ht⟩, - { refine ⟨_, t.compact_bUnion (λ x _, hUc x), λ x hx, _⟩, - rcases mem_Union₂.1 (ht hx) with ⟨y, hyt, hy⟩, - exact interior_mono (subset_bUnion_of_mem hyt) hy }, - { exact λ _, is_open_interior } -end +let ⟨L, hLc, hKL, _⟩ := exists_compact_between hK is_open_univ K.subset_univ in ⟨L, hLc, hKL⟩ protected lemma closed_embedding.locally_compact_space [locally_compact_space β] {f : α → β} (hf : closed_embedding f) : locally_compact_space α := From d032ed1dcb4e8b3fcc07c8fc92e0a7423b732b87 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Mon, 29 Aug 2022 16:06:00 +0000 Subject: [PATCH 076/828] feat(analysis/seminorm): Add construction of seminorm over normed fields (#15893) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit As for normed spaces the properties of seminorms over fields are strong enough so that it sufficies to assume that `∀ (r : 𝕜) x, f (r • x) ≤ ∥r∥ * f x`. This PR adds this construction. --- src/analysis/seminorm.lean | 26 ++++++++++++++++++++++++-- 1 file changed, 24 insertions(+), 2 deletions(-) diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 4d7f039414e7e..d5963af7f647b 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -48,10 +48,12 @@ structure seminorm (𝕜 : Type*) (E : Type*) [semi_normed_ring 𝕜] [add_group attribute [nolint doc_blame] seminorm.to_add_group_seminorm +section of + /-- Alternative constructor for a `seminorm` on an `add_comm_group E` that is a module over a `semi_norm_ring 𝕜`. -/ -def seminorm.of {𝕜 : Type*} {E : Type*} [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] - (f : E → ℝ) (add_le : ∀ (x y : E), f (x + y) ≤ f x + f y) +def seminorm.of [semi_normed_ring 𝕜] [add_comm_group E] [module 𝕜 E] (f : E → ℝ) + (add_le : ∀ (x y : E), f (x + y) ≤ f x + f y) (smul : ∀ (a : 𝕜) (x : E), f (a • x) = ∥a∥ * f x) : seminorm 𝕜 E := { to_fun := f, map_zero' := by rw [←zero_smul 𝕜 (0 : E), smul, norm_zero, zero_mul], @@ -59,6 +61,26 @@ def seminorm.of {𝕜 : Type*} {E : Type*} [semi_normed_ring 𝕜] [add_comm_gro smul' := smul, neg' := λ x, by rw [←neg_one_smul 𝕜, smul, norm_neg, ← smul, one_smul] } +/-- Alternative constructor for a `seminorm` over a normed field `𝕜` that only assumes `f 0 = 0` +and an inequality for the scalar multiplication. -/ +def seminorm.of_smul_le [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] (f : E → ℝ) + (map_zero : f 0 = 0) (add_le : ∀ x y, f (x + y) ≤ f x + f y) + (smul_le : ∀ (r : 𝕜) x, f (r • x) ≤ ∥r∥ * f x) : seminorm 𝕜 E := +seminorm.of f add_le + (λ r x, begin + refine le_antisymm (smul_le r x) _, + by_cases r = 0, + { simp [h, map_zero] }, + rw ←mul_le_mul_left (inv_pos.mpr (norm_pos_iff.mpr h)), + rw inv_mul_cancel_left₀ (norm_ne_zero_iff.mpr h), + specialize smul_le r⁻¹ (r • x), + rw norm_inv at smul_le, + convert smul_le, + simp [h], + end) + +end of + namespace seminorm section semi_normed_ring From b490ab4ac99873a41a6fc5731cb548728d6df9d5 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 29 Aug 2022 18:29:10 +0000 Subject: [PATCH 077/828] feat(analysis/normed_space/star/basic): add `cstar_ring` instances for `pi` and `prod` (#16254) --- src/analysis/normed_space/star/basic.lean | 41 +++++++++++++++++++++++ 1 file changed, 41 insertions(+) diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index 1ce247f5e8725..5051945464265 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -112,6 +112,47 @@ subtype.ext norm_star_mul_self end non_unital +section prod_pi + +variables {ι R₁ R₂ : Type*} {R : ι → Type*} +variables [non_unital_normed_ring R₁] [star_ring R₁] [cstar_ring R₁] +variables [non_unital_normed_ring R₂] [star_ring R₂] [cstar_ring R₂] +variables [Π i, non_unital_normed_ring (R i)] [Π i, star_ring (R i)] + +/-- This instance exists to short circuit type class resolution because of problems with +inference involving Π-types. -/ +instance _root_.pi.star_ring' : star_ring (Π i, R i) := infer_instance + +variables [fintype ι] [Π i, cstar_ring (R i)] + +instance _root_.prod.cstar_ring : cstar_ring (R₁ × R₂) := +{ norm_star_mul_self := λ x, + begin + unfold norm, + simp only [prod.fst_mul, prod.fst_star, prod.snd_mul, prod.snd_star, norm_star_mul_self, ←sq], + refine le_antisymm _ _, + { refine max_le _ _; + rw [sq_le_sq, abs_of_nonneg (norm_nonneg _)], + exact (le_max_left _ _).trans (le_abs_self _), + exact (le_max_right _ _).trans (le_abs_self _) }, + { rw le_max_iff, + rcases le_total (∥x.fst∥) (∥x.snd∥) with (h | h); + simp [h] } + end } + +instance _root_.pi.cstar_ring : cstar_ring (Π i, R i) := +{ norm_star_mul_self := λ x, + begin + simp only [norm, pi.mul_apply, pi.star_apply, nnnorm_star_mul_self, ←sq], + norm_cast, + exact (finset.comp_sup_eq_sup_comp_of_is_total (λ x : nnreal, x ^ 2) + (λ x y h, by simpa only [sq] using mul_le_mul' h h) (by simp)).symm, + end } + +instance _root_.pi.cstar_ring' : cstar_ring (ι → R₁) := pi.cstar_ring + +end prod_pi + section unital variables [normed_ring E] [star_ring E] [cstar_ring E] From 879155bff5af618b9062cbb2915347dafd749ad6 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Mon, 29 Aug 2022 20:00:59 +0000 Subject: [PATCH 078/828] feat(*): replace fact(0 < t) with ne_zero (#16145) This has almost fully made things easier, whilst also being easier on the typeclass system (`fact` can cause many slowdowns). It also allows us to make many things that were not instances before into instances. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- .../homogeneous_prime_not_prime.lean | 6 +- src/algebra/algebra/basic.lean | 7 +- src/algebra/char_p/basic.lean | 16 +++- src/algebra/char_p/invertible.lean | 7 +- src/algebra/invertible.lean | 4 + src/algebra/ne_zero.lean | 62 +++++++------ src/control/random.lean | 10 +-- src/data/bitvec/basic.lean | 2 +- src/data/fin/basic.lean | 25 ++---- src/data/nat/basic.lean | 13 --- src/data/nat/totient.lean | 18 ++-- src/data/pnat/basic.lean | 3 - src/data/zmod/basic.lean | 90 ++++++++++--------- src/data/zmod/defs.lean | 5 +- src/data/zmod/quotient.lean | 2 +- src/field_theory/finite/basic.lean | 2 +- src/field_theory/is_alg_closed/basic.lean | 2 +- src/group_theory/finite_abelian.lean | 4 +- .../specific_groups/dihedral.lean | 17 ++-- .../specific_groups/quaternion.lean | 35 +++----- src/number_theory/cyclotomic/gal.lean | 2 - .../legendre_symbol/add_character.lean | 4 - src/number_theory/lucas_lehmer.lean | 3 - src/number_theory/lucas_primality.lean | 2 +- src/number_theory/padics/ring_homs.lean | 15 ++-- src/number_theory/ramification_inertia.lean | 8 +- src/number_theory/sum_four_squares.lean | 12 +-- src/ring_theory/roots_of_unity.lean | 6 +- .../witt_vector/witt_polynomial.lean | 8 +- src/tactic/norm_fin.lean | 2 +- src/testing/slim_check/sampleable.lean | 2 +- 31 files changed, 186 insertions(+), 208 deletions(-) diff --git a/counterexamples/homogeneous_prime_not_prime.lean b/counterexamples/homogeneous_prime_not_prime.lean index 493fc2b3edbdd..9532317504c1c 100644 --- a/counterexamples/homogeneous_prime_not_prime.lean +++ b/counterexamples/homogeneous_prime_not_prime.lean @@ -118,14 +118,18 @@ instance : graded_algebra (grading R) := /-- The counterexample is the ideal `I = span {(2, 2)}`. -/ def I : ideal (R × R) := ideal.span {((2, 2) : (R × R))}. +set_option class.instance_max_depth 33 + lemma I_not_prime : ¬ I.is_prime := begin rintro ⟨rid1, rid2⟩, apply rid1, clear rid1, revert rid2, simp only [I, ideal.mem_span_singleton, ideal.eq_top_iff_one], - dec_trivial, + dec_trivial, -- this is what we change the max instance depth for, it's only 1 above the default end +set_option class.instance_max_depth 32 + lemma I_is_homogeneous : I.is_homogeneous (grading R) := begin rw ideal.is_homogeneous.iff_exists, diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index dc3e00e9dfdf8..198f5cf791ba0 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Yury Kudryashov -/ import algebra.module.basic +import algebra.module.ulift +import algebra.ne_zero import algebra.ring.aut import algebra.ring.ulift -import algebra.module.ulift import linear_algebra.span import tactic.abel @@ -1450,6 +1451,10 @@ suffices function.injective (λ (c : R), c • (1 : A)), by { convert this, ext, rw [algebra.smul_def, mul_one] }, smul_left_injective R one_ne_zero +lemma _root_.ne_zero.of_no_zero_smul_divisors (n : ℕ) [comm_ring R] [ne_zero (n : R)] [ring A] + [nontrivial A] [algebra R A] [no_zero_smul_divisors R A] : ne_zero (n : A) := +ne_zero.nat_of_injective $ no_zero_smul_divisors.algebra_map_injective R A + variables {R A} lemma iff_algebra_map_injective [comm_ring R] [ring A] [is_domain A] [algebra R A] : no_zero_smul_divisors R A ↔ function.injective (algebra_map R A) := diff --git a/src/algebra/char_p/basic.lean b/src/algebra/char_p/basic.lean index 697516969afa9..aa3017f32bdee 100644 --- a/src/algebra/char_p/basic.lean +++ b/src/algebra/char_p/basic.lean @@ -427,8 +427,8 @@ match p, hc with | (m+2), hc := or.inl (@char_is_prime_of_two_le R _ _ (m+2) hc (nat.le_add_left 2 m)) end -lemma char_is_prime_of_pos (p : ℕ) [h : fact (0 < p)] [char_p R p] : fact p.prime := -⟨(char_p.char_is_prime_or_zero R _).resolve_right (pos_iff_ne_zero.1 h.1)⟩ +lemma char_is_prime_of_pos (p : ℕ) [ne_zero p] [char_p R p] : fact p.prime := +⟨(char_p.char_is_prime_or_zero R _).resolve_right $ ne_zero.ne p⟩ end nontrivial @@ -591,3 +591,15 @@ begin end end + +namespace ne_zero + +variables (R) [add_monoid_with_one R] {r : R} {n p : ℕ} {a : ℕ+} + +lemma of_not_dvd [char_p R p] (h : ¬ p ∣ n) : ne_zero (n : R) := +⟨(char_p.cast_eq_zero_iff R p n).not.mpr h⟩ + +lemma not_char_dvd (p : ℕ) [char_p R p] (k : ℕ) [h : ne_zero (k : R)] : ¬ p ∣ k := +by rwa [←char_p.cast_eq_zero_iff R p k, ←ne.def, ←ne_zero_iff] + +end ne_zero diff --git a/src/algebra/char_p/invertible.lean b/src/algebra/char_p/invertible.lean index e54ec5996ea8b..4144d91f24611 100644 --- a/src/algebra/char_p/invertible.lean +++ b/src/algebra/char_p/invertible.lean @@ -40,9 +40,10 @@ def invertible_of_char_p_not_dvd {p : ℕ} [char_p K p] {t : ℕ} (not_dvd : ¬(p ∣ t)) : invertible (t : K) := invertible_of_nonzero (λ h, not_dvd ((char_p.cast_eq_zero_iff K p t).mp h)) -instance invertible_of_pos [char_zero K] (n : ℕ) [h : fact (0 < n)] : - invertible (n : K) := -invertible_of_nonzero $ by simpa [pos_iff_ne_zero] using h.out +-- warning: this could potentially loop with `ne_zero.invertible` - if there is weird type-class +-- loops, watch out for that. +instance invertible_of_pos [char_zero K] (n : ℕ) [ne_zero n] : invertible (n : K) := +invertible_of_nonzero $ ne_zero.out end field diff --git a/src/algebra/invertible.lean b/src/algebra/invertible.lean index e7bf9715b3724..d10805f17ceb1 100644 --- a/src/algebra/invertible.lean +++ b/src/algebra/invertible.lean @@ -5,6 +5,7 @@ Authors: Anne Baanen -/ import algebra.group.units +import algebra.ne_zero import algebra.ring.basic /-! @@ -212,6 +213,9 @@ lemma nonzero_of_invertible [mul_zero_one_class α] (a : α) [nontrivial α] [in λ ha, zero_ne_one $ calc 0 = ⅟a * a : by simp [ha] ... = 1 : inv_of_mul_self a +@[priority 100] instance invertible.ne_zero [mul_zero_one_class α] [nontrivial α] (a : α) + [invertible a] : ne_zero a := ⟨nonzero_of_invertible a⟩ + section monoid_with_zero variable [monoid_with_zero α] diff --git a/src/algebra/ne_zero.lean b/src/algebra/ne_zero.lean index 40db44f08d73f..cb460646b7071 100644 --- a/src/algebra/ne_zero.lean +++ b/src/algebra/ne_zero.lean @@ -3,7 +3,10 @@ Copyright (c) 2021 Eric Rodriguez. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Rodriguez -/ -import data.zmod.basic + +import data.nat.cast +import data.pnat.basic +import algebra.group_power.ring /-! # `ne_zero` typeclass @@ -30,6 +33,9 @@ lemma ne_zero_iff {R : Type*} [has_zero R] {n : R} : ne_zero n ↔ n ≠ 0 := lemma not_ne_zero {R : Type*} [has_zero R] {n : R} : ¬ ne_zero n ↔ n = 0 := by simp [ne_zero_iff] +lemma eq_zero_or_ne_zero {α} [has_zero α] (a : α) : a = 0 ∨ ne_zero a := +(eq_or_ne a 0).imp_right ne_zero.mk + namespace ne_zero variables {R S M F : Type*} {r : R} {x y : M} {n p : ℕ} {a : ℕ+} @@ -37,14 +43,36 @@ variables {R S M F : Type*} {r : R} {x y : M} {n p : ℕ} {a : ℕ+} instance pnat : ne_zero (a : ℕ) := ⟨a.ne_zero⟩ instance succ : ne_zero (n + 1) := ⟨n.succ_ne_zero⟩ +instance one (R) [mul_zero_one_class R] [nontrivial R] : ne_zero (1 : R) := ⟨one_ne_zero⟩ + +lemma pos (r : R) [canonically_ordered_add_monoid R] [ne_zero r] : 0 < r := +(zero_le r).lt_of_ne $ ne_zero.out.symm + lemma of_pos [preorder M] [has_zero M] (h : 0 < x) : ne_zero x := ⟨h.ne'⟩ lemma of_gt [canonically_ordered_add_monoid M] (h : x < y) : ne_zero y := of_pos $ pos_of_gt h -instance char_zero [ne_zero n] [add_monoid_with_one M] [char_zero M] : ne_zero (n : M) := -⟨nat.cast_ne_zero.mpr $ ne_zero.ne n⟩ +-- 1 < p is still an often-used `fact`, due to `nat.prime` implying it, and it implying `nontrivial` +-- on `zmod`'s ring structure. We cannot just set this to be any `x < y`, else that becomes a +-- metavariable and it will hugely slow down typeclass inference. +@[priority 10] +instance of_gt' [canonically_ordered_add_monoid M] [has_one M] [fact (1 < y)] : ne_zero y := +of_gt $ fact.out $ 1 < y + +instance bit0 [canonically_ordered_add_monoid M] [ne_zero x] : ne_zero (bit0 x) := +of_pos $ bit0_pos $ ne_zero.pos x + +instance bit1 [canonically_ordered_comm_semiring M] [nontrivial M] : ne_zero (bit1 x) := +⟨mt (λ h, le_iff_exists_add'.2 ⟨_, h.symm⟩) canonically_ordered_comm_semiring.zero_lt_one.not_le⟩ + +instance pow [monoid_with_zero M] [no_zero_divisors M] [ne_zero x] : ne_zero (x ^ n) := +⟨pow_ne_zero n out⟩ + +instance mul [has_zero M] [has_mul M] [no_zero_divisors M] [ne_zero x] [ne_zero y] : + ne_zero (x * y) := +⟨mul_ne_zero out out⟩ -@[priority 100] instance invertible [mul_zero_one_class M] [nontrivial M] [invertible x] : - ne_zero x := ⟨nonzero_of_invertible x⟩ +instance char_zero [ne_zero n] [add_monoid_with_one M] [char_zero M] : ne_zero (n : M) := +⟨nat.cast_ne_zero.mpr out⟩ instance coe_trans [has_zero M] [has_coe R S] [has_coe_t S M] [h : ne_zero (r : M)] : ne_zero ((r : S) : M) := ⟨h.out⟩ @@ -70,36 +98,12 @@ lemma nat_of_injective [non_assoc_semiring M] [non_assoc_semiring R] [h : ne_zer [ring_hom_class F R M] {f : F} (hf : function.injective f) : ne_zero (n : M) := ⟨λ h, (ne_zero.ne' n R) $ hf $ by simpa⟩ -lemma pos (r : R) [canonically_ordered_add_monoid R] [ne_zero r] : 0 < r := -(zero_le r).lt_of_ne $ ne_zero.out.symm - variables (R M) -lemma of_not_dvd [add_monoid_with_one M] [char_p M p] (h : ¬ p ∣ n) : ne_zero (n : M) := -⟨(not_iff_not.mpr $ char_p.cast_eq_zero_iff M p n).mpr h⟩ - -lemma of_no_zero_smul_divisors (n : ℕ) [comm_ring R] [ne_zero (n : R)] [ring M] [nontrivial M] - [algebra R M] [no_zero_smul_divisors R M] : ne_zero (n : M) := -nat_of_injective $ no_zero_smul_divisors.algebra_map_injective R M - lemma of_ne_zero_coe [add_monoid_with_one R] [h : ne_zero (n : R)] : ne_zero n := ⟨by {casesI h, rintro rfl, by simpa using h}⟩ -lemma not_char_dvd [add_monoid_with_one R] (p : ℕ) [char_p R p] (k : ℕ) [h : ne_zero (k : R)] : - ¬ p ∣ k := -by rwa [←not_iff_not.mpr $ char_p.cast_eq_zero_iff R p k, ←ne.def, ←ne_zero_iff] - lemma pos_of_ne_zero_coe [add_monoid_with_one R] [ne_zero (n : R)] : 0 < n := (ne_zero.of_ne_zero_coe R).out.bot_lt end ne_zero - -lemma eq_zero_or_ne_zero {α} [has_zero α] (a : α) : a = 0 ∨ ne_zero a := -(eq_or_ne a 0).imp_right ne_zero.mk - -namespace zmod - -instance fintype' (n : ℕ) [ne_zero n] : fintype (zmod n) := -@zmod.fintype n ⟨ne_zero.pos n⟩ - -end zmod diff --git a/src/control/random.lean b/src/control/random.lean index d61173a5ddf32..26f1f24bbcc38 100644 --- a/src/control/random.lean +++ b/src/control/random.lean @@ -5,11 +5,11 @@ Authors: Simon Hudon -/ import control.monad.basic +import control.uliftable +import data.bitvec.basic import data.int.basic import data.stream.defs -import control.uliftable import tactic.norm_num -import data.bitvec.basic /-! @@ -224,7 +224,7 @@ variables {g : Type} [random_gen g] open nat namespace fin -variables {n : ℕ} [fact (0 < n)] +variables {n : ℕ} [ne_zero n] /-- generate a `fin` randomly -/ protected def random : rand_g g (fin n) := @@ -252,7 +252,7 @@ instance int_bounded_random : bounded_random ℤ := (int.coe_nat_le_coe_nat_of_le h₁) (le_of_eq $ int.of_nat_nat_abs_eq_of_nonneg (int.sub_nonneg_of_le hxy)) ⟩ } -instance fin_random (n : ℕ) [fact (0 < n)] : random (fin n) := +instance fin_random (n : ℕ) [ne_zero n] : random (fin n) := { random := λ g inst, @fin.random g inst _ _ } instance fin_bounded_random (n : ℕ) : bounded_random (fin n) := @@ -284,8 +284,6 @@ instance : bounded_random bool := subtype.map bool.of_nat (bool_of_nat_mem_Icc_of_mem_Icc_to_nat x y) <$> @bounded_random.random_r ℕ _ _ g _inst x.to_nat y.to_nat (bool.to_nat_le_to_nat p) } -open_locale fin_fact - /-- generate a random bit vector of length `n` -/ def bitvec.random (n : ℕ) : rand_g g (bitvec n) := bitvec.of_fin <$> rand.random (fin $ 2^n) diff --git a/src/data/bitvec/basic.lean b/src/data/bitvec/basic.lean index d3e65633c634c..354ec842d1a93 100644 --- a/src/data/bitvec/basic.lean +++ b/src/data/bitvec/basic.lean @@ -22,7 +22,7 @@ by rw [of_fin,to_nat_of_nat,nat.mod_eq_of_lt]; apply i.is_lt /-- convert `bitvec` to `fin` -/ def to_fin {n : ℕ} (i : bitvec n) : fin $ 2^n := -@fin.of_nat' _ ⟨pow_pos (by norm_num) _⟩ i.to_nat +fin.of_nat' i.to_nat lemma add_lsb_eq_twice_add_one {x b} : add_lsb x b = 2 * x + cond b 1 0 := diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index 69ba83256b4af..057bad0186779 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -3,9 +3,10 @@ Copyright (c) 2017 Robert Y. Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis, Keeley Hoek -/ -import tactic.apply_fun +import algebra.ne_zero import data.nat.cast import order.rel_iso +import tactic.apply_fun import tactic.localized /-! @@ -54,7 +55,7 @@ This file expands on the development in the core library. ### Other casts -* `fin.of_nat'`: given a positive number `n` (deduced from `[fact (0 < n)]`), `fin.of_nat' i` is +* `fin.of_nat'`: given a positive number `n` (deduced from `[ne_zero n]`), `fin.of_nat' i` is `i % n` interpreted as an element of `fin n`; * `fin.cast_lt i h` : embed `i` into a `fin` where `h` proves it belongs into; * `fin.pred_above (p : fin n) i` : embed `i : fin (n+1)` into `fin n` by subtracting one if `p < i`; @@ -77,22 +78,6 @@ open fin nat function /-- Elimination principle for the empty set `fin 0`, dependent version. -/ def fin_zero_elim {α : fin 0 → Sort u} (x : fin 0) : α x := x.elim0 -lemma fact.succ.pos {n} : fact (0 < succ n) := ⟨zero_lt_succ _⟩ - -lemma fact.bit0.pos {n} [h : fact (0 < n)] : fact (0 < bit0 n) := -⟨nat.zero_lt_bit0 $ ne_of_gt h.1⟩ - -lemma fact.bit1.pos {n} : fact (0 < bit1 n) := -⟨nat.zero_lt_bit1 _⟩ - -lemma fact.pow.pos {p n : ℕ} [h : fact $ 0 < p] : fact (0 < p ^ n) := -⟨pow_pos h.1 _⟩ - -localized "attribute [instance] fact.succ.pos" in fin_fact -localized "attribute [instance] fact.bit0.pos" in fin_fact -localized "attribute [instance] fact.bit1.pos" in fin_fact -localized "attribute [instance] fact.pow.pos" in fin_fact - namespace fin /-- A non-dependent variant of `elim0`. -/ @@ -323,7 +308,7 @@ section add -/ /-- Given a positive `n`, `fin.of_nat' i` is `i % n` as an element of `fin n`. -/ -def of_nat' [h : fact (0 < n)] (i : ℕ) : fin n := ⟨i%n, mod_lt _ h.1⟩ +def of_nat' [ne_zero n] (i : ℕ) : fin n := ⟨i%n, mod_lt _ $ ne_zero.pos n⟩ lemma one_val {n : ℕ} : (1 : fin (n+1)).val = 1 % (n+1) := rfl lemma coe_one' {n : ℕ} : ((1 : fin (n+1)) : ℕ) = 1 % (n+1) := rfl @@ -1668,7 +1653,7 @@ lemma coe_of_nat_eq_mod (m n : ℕ) : ((n : fin (succ m)) : ℕ) = n % succ m := by rw [← of_nat_eq_coe]; refl -@[simp] lemma coe_of_nat_eq_mod' (m n : ℕ) [I : fact (0 < m)] : +@[simp] lemma coe_of_nat_eq_mod' (m n : ℕ) [I : ne_zero m] : (@fin.of_nat' _ I n : ℕ) = n % m := rfl diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index e85a51c87b1e7..9db3673f1a36d 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -138,19 +138,6 @@ attribute [simp] nat.not_lt_zero nat.succ_ne_zero nat.succ_ne_self nat.bit0_ne_one nat.one_ne_bit0 nat.bit0_ne_bit1 nat.bit1_ne_bit0 -/-! -Inject some simple facts into the type class system. -This `fact` should not be confused with the factorial function `nat.fact`! --/ -section facts - -instance succ_pos'' (n : ℕ) : fact (0 < n.succ) := ⟨n.succ_pos⟩ - -instance pos_of_one_lt (n : ℕ) [h : fact (1 < n)] : fact (0 < n) := -⟨lt_trans zero_lt_one h.1⟩ - -end facts - variables {m n k : ℕ} namespace nat diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index ded01f637e46c..e4d214bfdfcda 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -89,13 +89,13 @@ open zmod /-- Note this takes an explicit `fintype ((zmod n)ˣ)` argument to avoid trouble with instance diamonds. -/ -@[simp] lemma _root_.zmod.card_units_eq_totient (n : ℕ) [h : fact (0 < n)] [fintype ((zmod n)ˣ)] : +@[simp] lemma _root_.zmod.card_units_eq_totient (n : ℕ) [ne_zero n] [fintype ((zmod n)ˣ)] : fintype.card ((zmod n)ˣ) = φ n := calc fintype.card ((zmod n)ˣ) = fintype.card {x : zmod n // x.val.coprime n} : fintype.card_congr zmod.units_equiv_coprime ... = φ n : begin - unfreezingI { obtain ⟨m, rfl⟩ : ∃ m, n = m + 1 := exists_eq_succ_of_ne_zero h.out.ne' }, + unfreezingI { obtain ⟨m, rfl⟩ : ∃ m, n = m + 1 := exists_eq_succ_of_ne_zero ne_zero.out }, simp only [totient, finset.card_eq_sum_ones, fintype.card_subtype, finset.sum_filter, ← fin.sum_univ_eq_sum_range, @nat.coprime_comm (m + 1)], refl @@ -104,6 +104,7 @@ end lemma totient_even {n : ℕ} (hn : 2 < n) : even n.totient := begin haveI : fact (1 < n) := ⟨one_lt_two.trans hn⟩, + haveI : ne_zero n := ne_zero.of_gt hn, suffices : 2 = order_of (-1 : (zmod n)ˣ), { rw [← zmod.card_units_eq_totient, even_iff_two_dvd, this], exact order_of_dvd_card_univ }, rw [←order_of_units, units.coe_neg_one, order_of_neg_one, ring_char.eq (zmod n) n, if_neg hn.ne'], @@ -115,9 +116,9 @@ if hmn0 : m * n = 0 simp only [totient_zero, mul_zero, zero_mul, h] else begin - haveI : fact (0 < (m * n)) := ⟨nat.pos_of_ne_zero hmn0⟩, - haveI : fact (0 < m) := ⟨nat.pos_of_ne_zero $ left_ne_zero_of_mul hmn0⟩, - haveI : fact (0 < n) := ⟨nat.pos_of_ne_zero $ right_ne_zero_of_mul hmn0⟩, + haveI : ne_zero (m * n) := ⟨hmn0⟩, + haveI : ne_zero m := ⟨left_ne_zero_of_mul hmn0⟩, + haveI : ne_zero n := ⟨right_ne_zero_of_mul hmn0⟩, simp only [← zmod.card_units_eq_totient], rw [fintype.card_congr (units.map_equiv (zmod.chinese_remainder h).to_mul_equiv).to_equiv, fintype.card_congr (@mul_equiv.prod_units (zmod m) (zmod n) _ _).to_equiv, @@ -226,7 +227,7 @@ end lemma card_units_zmod_lt_sub_one {p : ℕ} (hp : 1 < p) [fintype ((zmod p)ˣ)] : fintype.card ((zmod p)ˣ) ≤ p - 1 := begin - haveI : fact (0 < p) := ⟨zero_lt_one.trans hp⟩, + haveI : ne_zero p := ⟨(pos_of_gt hp).ne'⟩, rw zmod.card_units_eq_totient p, exact nat.le_pred_of_lt (nat.totient_lt p hp), end @@ -234,14 +235,13 @@ end lemma prime_iff_card_units (p : ℕ) [fintype ((zmod p)ˣ)] : p.prime ↔ fintype.card ((zmod p)ˣ) = p - 1 := begin - by_cases hp : p = 0, + casesI eq_zero_or_ne_zero p with hp hp, { substI hp, simp only [zmod, not_prime_zero, false_iff, zero_tsub], -- the substI created an non-defeq but subsingleton instance diamond; resolve it suffices : fintype.card ℤˣ ≠ 0, { convert this }, simp }, - haveI : fact (0 < p) := ⟨nat.pos_of_ne_zero hp⟩, - rw [zmod.card_units_eq_totient, nat.totient_eq_iff_prime (fact.out (0 < p))], + rw [zmod.card_units_eq_totient, nat.totient_eq_iff_prime $ ne_zero.pos p], end @[simp] lemma totient_two : φ 2 = 1 := diff --git a/src/data/pnat/basic.lean b/src/data/pnat/basic.lean index 71a4a55f13056..231e05a3e7eab 100644 --- a/src/data/pnat/basic.lean +++ b/src/data/pnat/basic.lean @@ -114,9 +114,6 @@ open nat @[simp] theorem pos (n : ℕ+) : 0 < (n : ℕ) := n.2 --- see note [fact non_instances] -lemma fact_pos (n : ℕ+) : fact (0 < ↑n) := ⟨n.pos⟩ - theorem eq {m n : ℕ+} : (m : ℕ) = n → m = n := subtype.eq @[simp] lemma coe_inj {m n : ℕ+} : (m : ℕ) = n ↔ m = n := set_coe.ext_iff diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index cdd04c40094c9..fb7dd8c76d543 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -42,14 +42,14 @@ def val : Π {n : ℕ}, zmod n → ℕ | 0 := int.nat_abs | (n+1) := (coe : fin (n + 1) → ℕ) -lemma val_lt {n : ℕ} [fact (0 < n)] (a : zmod n) : a.val < n := +lemma val_lt {n : ℕ} [ne_zero n] (a : zmod n) : a.val < n := begin casesI n, - { exfalso, exact nat.not_lt_zero 0 (fact.out _) }, + { cases ne_zero.ne 0 rfl }, exact fin.is_lt a end -lemma val_le {n : ℕ} [fact (0 < n)] (a : zmod n) : a.val ≤ n := +lemma val_le {n : ℕ} [ne_zero n] (a : zmod n) : a.val ≤ n := a.val_lt.le @[simp] lemma val_zero : ∀ {n}, (0 : zmod n).val = 0 @@ -130,6 +130,13 @@ def cast : Π {n : ℕ}, zmod n → R @[simp] lemma cast_zero : ((0 : zmod n) : R) = 0 := by cases n; simp +lemma cast_eq_val [ne_zero n] (a : zmod n) : (a : R) = a.val := +begin + casesI n, + { cases ne_zero.ne 0 rfl }, + refl, +end + variables {S : Type*} [add_group_with_one S] @[simp] lemma _root_.prod.fst_zmod_cast (a : zmod n) : (a : R × S).fst = a := @@ -142,17 +149,17 @@ end /-- So-named because the coercion is `nat.cast` into `zmod`. For `nat.cast` into an arbitrary ring, see `zmod.nat_cast_val`. -/ -lemma nat_cast_zmod_val {n : ℕ} [fact (0 < n)] (a : zmod n) : (a.val : zmod n) = a := +lemma nat_cast_zmod_val {n : ℕ} [ne_zero n] (a : zmod n) : (a.val : zmod n) = a := begin casesI n, - { exfalso, exact nat.not_lt_zero 0 (fact.out _) }, + { cases ne_zero.ne 0 rfl }, { apply fin.coe_coe_eq_self } end -lemma nat_cast_right_inverse [fact (0 < n)] : function.right_inverse val (coe : ℕ → zmod n) := +lemma nat_cast_right_inverse [ne_zero n] : function.right_inverse val (coe : ℕ → zmod n) := nat_cast_zmod_val -lemma nat_cast_zmod_surjective [fact (0 < n)] : function.surjective (coe : ℕ → zmod n) := +lemma nat_cast_zmod_surjective [ne_zero n] : function.surjective (coe : ℕ → zmod n) := nat_cast_right_inverse.surjective /-- So-named because the outer coercion is `int.cast` into `zmod`. For `int.cast` into an arbitrary @@ -181,11 +188,11 @@ lemma cast_id' : (coe : zmod n → zmod n) = id := funext (cast_id n) variables (R) [ring R] /-- The coercions are respectively `nat.cast` and `zmod.cast`. -/ -@[simp] lemma nat_cast_comp_val [fact (0 < n)] : +@[simp] lemma nat_cast_comp_val [ne_zero n] : (coe : ℕ → R) ∘ (val : zmod n → ℕ) = coe := begin casesI n, - { exfalso, exact nat.not_lt_zero 0 (fact.out _) }, + { cases ne_zero.ne 0 rfl }, refl end @@ -199,7 +206,7 @@ end variables {R} -@[simp] lemma nat_cast_val [fact (0 < n)] (i : zmod n) : (i.val : R) = i := +@[simp] lemma nat_cast_val [ne_zero n] (i : zmod n) : (i.val : R) = i := congr_fun (nat_cast_comp_val R) i @[simp] lemma int_cast_cast (i : zmod n) : ((i : ℤ) : R) = i := @@ -331,9 +338,8 @@ end lemma cast_hom_bijective [fintype R] (h : fintype.card R = n) : function.bijective (zmod.cast_hom (dvd_refl n) R) := begin - haveI : fact (0 < n) := + haveI : ne_zero n := ⟨begin - rw [pos_iff_ne_zero], intro hn, rw hn at h, exact (fintype.card_eq_zero_iff.mp h).elim' 0 @@ -395,7 +401,7 @@ end lemma nat_coe_zmod_eq_zero_iff_dvd (a b : ℕ) : (a : zmod b) = 0 ↔ b ∣ a := by rw [← nat.cast_zero, zmod.nat_coe_eq_nat_coe_iff, nat.modeq_zero_iff_dvd] -lemma val_int_cast {n : ℕ} (a : ℤ) [fact (0 < n)] : ↑(a : zmod n).val = a % n := +lemma val_int_cast {n : ℕ} (a : ℤ) [ne_zero n] : ↑(a : zmod n).val = a % n := begin have hle : (0 : ℤ) ≤ ↑(a : zmod n).val := int.coe_nat_nonneg _, have hlt : ↑(a : zmod n).val < (n : ℤ) := int.coe_nat_lt.mpr (zmod.val_lt a), @@ -439,7 +445,7 @@ begin { exact hk } } }, end -lemma nat_coe_zmod_eq_iff (p : ℕ) (n : ℕ) (z : zmod p) [fact (0 < p)] : +lemma nat_coe_zmod_eq_iff (p : ℕ) (n : ℕ) (z : zmod p) [ne_zero p] : ↑n = z ↔ ∃ k, n = z.val + p * k := begin split, @@ -450,7 +456,7 @@ begin rw [nat.cast_add, nat_cast_zmod_val, nat.cast_mul, nat_cast_self, zero_mul, add_zero] } end -lemma int_coe_zmod_eq_iff (p : ℕ) (n : ℤ) (z : zmod p) [fact (0 < p)] : +lemma int_coe_zmod_eq_iff (p : ℕ) (n : ℤ) (z : zmod p) [ne_zero p] : ↑n = z ↔ ∃ k, n = z.val + p * k := begin split, @@ -486,11 +492,11 @@ local attribute [semireducible] int.nonneg | (n : ℕ) h := by simp only [int.cast_coe_nat, int.to_nat_coe_nat] | -[1+n] h := false.elim h -lemma val_injective (n : ℕ) [fact (0 < n)] : +lemma val_injective (n : ℕ) [ne_zero n] : function.injective (zmod.val : zmod n → ℕ) := begin casesI n, - { exfalso, exact nat.not_lt_zero 0 (fact.out _) }, + { cases ne_zero.ne 0 rfl }, assume a b h, ext, exact h @@ -502,10 +508,10 @@ by rw [← nat.cast_one, val_nat_cast] lemma val_one (n : ℕ) [fact (1 < n)] : (1 : zmod n).val = 1 := by { rw val_one_eq_one_mod, exact nat.mod_eq_of_lt (fact.out _) } -lemma val_add {n : ℕ} [fact (0 < n)] (a b : zmod n) : (a + b).val = (a.val + b.val) % n := +lemma val_add {n : ℕ} [ne_zero n] (a b : zmod n) : (a + b).val = (a.val + b.val) % n := begin casesI n, - { exfalso, exact nat.not_lt_zero 0 (fact.out _) }, + { cases ne_zero.ne 0 rfl }, { apply fin.val_add } end @@ -522,6 +528,8 @@ instance nontrivial (n : ℕ) [fact (1 < n)] : nontrivial (zmod n) := ... = (1 : zmod n).val : congr_arg zmod.val h ... = 1 : val_one n ⟩⟩ +instance nontrivial' : nontrivial (zmod 0) := int.nontrivial + /-- The inversion on `zmod n`. It is setup in such a way that `a * a⁻¹` is equal to `gcd a.val n`. In particular, if `a` is coprime to `n`, and hence a unit, `a * a⁻¹ = 1`. -/ @@ -610,10 +618,10 @@ lemma inv_mul_of_unit {n : ℕ} (a : zmod n) (h : is_unit a) : a⁻¹ * a = 1 := by rw [mul_comm, mul_inv_of_unit a h] +-- TODO: this equivalence is true for `zmod 0 = ℤ`, but needs to use different functions. /-- Equivalence between the units of `zmod n` and the subtype of terms `x : zmod n` for which `x.val` is comprime to `n` -/ -def units_equiv_coprime {n : ℕ} [fact (0 < n)] : - (zmod n)ˣ ≃ {x : zmod n // nat.coprime x.val n} := +def units_equiv_coprime {n : ℕ} [ne_zero n] : (zmod n)ˣ ≃ {x : zmod n // nat.coprime x.val n} := { to_fun := λ x, ⟨x, val_coe_unit_coprime x⟩, inv_fun := λ x, unit_of_coprime x.1.val x.2, left_inv := λ ⟨_, _, _, _⟩, units.ext (nat_cast_zmod_val _), @@ -644,9 +652,9 @@ have inv : function.left_inverse inv_fun to_fun ∧ function.right_inverse inv_f end else begin - haveI : fact (0 < (m * n)) := ⟨nat.pos_of_ne_zero hmn0⟩, - haveI : fact (0 < m) := ⟨nat.pos_of_ne_zero $ left_ne_zero_of_mul hmn0⟩, - haveI : fact (0 < n) := ⟨nat.pos_of_ne_zero $ right_ne_zero_of_mul hmn0⟩, + haveI : ne_zero (m * n) := ⟨hmn0⟩, + haveI : ne_zero m := ⟨left_ne_zero_of_mul hmn0⟩, + haveI : ne_zero n := ⟨right_ne_zero_of_mul hmn0⟩, have left_inv : function.left_inverse inv_fun to_fun, { intro x, dsimp only [dvd_mul_left, dvd_mul_right, zmod.cast_hom_apply, coe_coe, inv_fun, to_fun], @@ -667,26 +675,20 @@ have inv : function.left_inverse inv_fun to_fun ∧ function.right_inverse inv_f left_inv := inv.1, right_inv := inv.2 } +-- todo: this can be made a `unique` instance. instance subsingleton_units : subsingleton ((zmod 2)ˣ) := -⟨λ x y, begin - ext1, - cases x with x xi hx1 hx2, - cases y with y yi hy1 hy2, - revert hx1 hx2 hy1 hy2, - fin_cases x; fin_cases y; simp -end⟩ +⟨dec_trivial⟩ lemma le_div_two_iff_lt_neg (n : ℕ) [hn : fact ((n : ℕ) % 2 = 1)] {x : zmod n} (hx0 : x ≠ 0) : x.val ≤ (n / 2 : ℕ) ↔ (n / 2 : ℕ) < (-x).val := begin - haveI npos : fact (0 < n) := ⟨by - { apply (nat.eq_zero_or_pos n).resolve_left, - unfreezingI { rintro rfl }, + haveI npos : ne_zero n := ⟨by + { unfreezingI { rintro rfl }, simpa [fact_iff] using hn, }⟩, have hn2 : (n : ℕ) / 2 < n := nat.div_lt_of_lt_mul - ((lt_mul_iff_one_lt_left npos.1).2 dec_trivial), + ((lt_mul_iff_one_lt_left $ ne_zero.pos n).2 dec_trivial), have hn2' : (n : ℕ) - n / 2 = n / 2 + 1, - { conv {to_lhs, congr, rw [← nat.succ_sub_one n, nat.succ_sub npos.1]}, + { conv {to_lhs, congr, rw [← nat.succ_sub_one n, nat.succ_sub $ ne_zero.pos n]}, rw [← nat.two_mul_odd_div_two hn.1, two_mul, ← nat.succ_add, add_tsub_cancel_right], }, have hxn : (n : ℕ) - x.val < n, { rw [tsub_lt_iff_tsub_lt x.val_le le_rfl, tsub_self], @@ -723,18 +725,18 @@ end lemma val_cast_of_lt {n : ℕ} {a : ℕ} (h : a < n) : (a : zmod n).val = a := by rw [val_nat_cast, nat.mod_eq_of_lt h] -lemma neg_val' {n : ℕ} [fact (0 < n)] (a : zmod n) : (-a).val = (n - a.val) % n := +lemma neg_val' {n : ℕ} [ne_zero n] (a : zmod n) : (-a).val = (n - a.val) % n := calc (-a).val = val (-a) % n : by rw nat.mod_eq_of_lt ((-a).val_lt) ... = (n - val a) % n : nat.modeq.add_right_cancel' _ (by rw [nat.modeq, ←val_add, add_left_neg, tsub_add_cancel_of_le a.val_le, nat.mod_self, val_zero]) -lemma neg_val {n : ℕ} [fact (0 < n)] (a : zmod n) : (-a).val = if a = 0 then 0 else n - a.val := +lemma neg_val {n : ℕ} [ne_zero n] (a : zmod n) : (-a).val = if a = 0 then 0 else n - a.val := begin rw neg_val', by_cases h : a = 0, { rw [if_pos h, h, val_zero, tsub_zero, nat.mod_self] }, rw if_neg h, apply nat.mod_eq_of_lt, - apply nat.sub_lt (fact.out (0 < n)), + apply nat.sub_lt (ne_zero.pos n), contrapose! h, rwa [le_zero_iff, val_eq_zero] at h, end @@ -747,11 +749,11 @@ def val_min_abs : Π {n : ℕ}, zmod n → ℤ @[simp] lemma val_min_abs_def_zero (x : zmod 0) : val_min_abs x = x := rfl -lemma val_min_abs_def_pos {n : ℕ} [fact (0 < n)] (x : zmod n) : +lemma val_min_abs_def_pos {n : ℕ} [ne_zero n] (x : zmod n) : val_min_abs x = if x.val ≤ n / 2 then x.val else x.val - n := begin casesI n, - { exfalso, exact nat.not_lt_zero 0 (fact.out (0 < 0)) }, + { cases ne_zero.ne 0 rfl }, { refl } end @@ -766,7 +768,7 @@ begin sub_zero] } end -lemma nat_abs_val_min_abs_le {n : ℕ} [fact (0 < n)] (x : zmod n) : x.val_min_abs.nat_abs ≤ n / 2 := +lemma nat_abs_val_min_abs_le {n : ℕ} [ne_zero n] (x : zmod n) : x.val_min_abs.nat_abs ≤ n / 2 := begin rw zmod.val_min_abs_def_pos, split_ifs with h, { exact h }, @@ -800,7 +802,7 @@ begin { rintro rfl, rw val_min_abs_zero } end -lemma nat_cast_nat_abs_val_min_abs {n : ℕ} [fact (0 < n)] (a : zmod n) : +lemma nat_cast_nat_abs_val_min_abs {n : ℕ} [ne_zero n] (a : zmod n) : (a.val_min_abs.nat_abs : zmod n) = if a.val ≤ (n : ℕ) / 2 then a else -a := begin have : (a.val : ℤ) - n ≤ 0, @@ -840,7 +842,7 @@ begin { rw [hn1, add_comm, nat.succ_le_iff] } end -lemma val_eq_ite_val_min_abs {n : ℕ} [fact (0 < n)] (a : zmod n) : +lemma val_eq_ite_val_min_abs {n : ℕ} [ne_zero n] (a : zmod n) : (a.val : ℤ) = a.val_min_abs + if a.val ≤ n / 2 then 0 else n := by { rw [zmod.val_min_abs_def_pos], split_ifs; simp only [add_zero, sub_add_cancel] } diff --git a/src/data/zmod/defs.lean b/src/data/zmod/defs.lean index 92a5c05f6814c..1e77307364842 100644 --- a/src/data/zmod/defs.lean +++ b/src/data/zmod/defs.lean @@ -5,6 +5,7 @@ Authors: Eric Rodriguez -/ import data.int.modeq +import algebra.ne_zero /-! # Definition of `zmod n` + basic results. @@ -81,8 +82,8 @@ instance zmod.has_repr : Π (n : ℕ), has_repr (zmod n) namespace zmod -instance fintype : Π (n : ℕ) [fact (0 < n)], fintype (zmod n) -| 0 h := (lt_irrefl _ h.1).elim +instance fintype : Π (n : ℕ) [ne_zero n], fintype (zmod n) +| 0 h := by exactI (ne_zero.ne 0 rfl).elim | (n+1) _ := fin.fintype (n+1) instance infinite : infinite (zmod 0) := diff --git a/src/data/zmod/quotient.lean b/src/data/zmod/quotient.lean index 23f0fbecb6f07..2f757275d72cb 100644 --- a/src/data/zmod/quotient.lean +++ b/src/data/zmod/quotient.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 algebra.ne_zero +import data.zmod.basic import group_theory.group_action.quotient import ring_theory.int.basic diff --git a/src/field_theory/finite/basic.lean b/src/field_theory/finite/basic.lean index 0785b25ff1d87..1096c12d080a8 100644 --- a/src/field_theory/finite/basic.lean +++ b/src/field_theory/finite/basic.lean @@ -337,7 +337,7 @@ end zmod namespace char_p lemma sq_add_sq (R : Type*) [comm_ring R] [is_domain R] - (p : ℕ) [fact (0 < p)] [char_p R p] (x : ℤ) : + (p : ℕ) [ne_zero p] [char_p R p] (x : ℤ) : ∃ a b : ℕ, (a^2 + b^2 : R) = x := begin haveI := char_is_prime_of_pos R p, diff --git a/src/field_theory/is_alg_closed/basic.lean b/src/field_theory/is_alg_closed/basic.lean index 9b4b51071aed6..6f9df61bb0b2a 100644 --- a/src/field_theory/is_alg_closed/basic.lean +++ b/src/field_theory/is_alg_closed/basic.lean @@ -342,7 +342,7 @@ omit hS @[priority 100] noncomputable instance perfect_ring (p : ℕ) [fact p.prime] [char_p k p] [is_alg_closed k] : perfect_ring k p := -perfect_ring.of_surjective k p $ λ x, is_alg_closed.exists_pow_nat_eq _ $ fact.out _ +perfect_ring.of_surjective k p $ λ x, is_alg_closed.exists_pow_nat_eq _ $ ne_zero.pos p /-- Algebraically closed fields are infinite since `Xⁿ⁺¹ - 1` is separable when `#K = n` -/ @[priority 500] diff --git a/src/group_theory/finite_abelian.lean b/src/group_theory/finite_abelian.lean index 4f0a5364c0624..0a844cd0e78bc 100644 --- a/src/group_theory/finite_abelian.lean +++ b/src/group_theory/finite_abelian.lean @@ -29,8 +29,8 @@ lemma finite_of_fg_torsion [add_comm_group M] [module ℤ M] [module.finite ℤ (hM : module.is_torsion ℤ M) : _root_.finite M := begin rcases module.equiv_direct_sum_of_is_torsion hM with ⟨ι, _, p, h, e, ⟨l⟩⟩, - haveI : ∀ i : ι, fact $ 0 < (p i ^ e i).nat_abs := - λ i, fact.mk $ int.nat_abs_pos_of_ne_zero $ pow_ne_zero (e i) (h i).ne_zero, + haveI : ∀ i : ι, ne_zero (p i ^ e i).nat_abs := + λ i, ⟨int.nat_abs_ne_zero_of_ne_zero $ pow_ne_zero (e i) (h i).ne_zero⟩, haveI : ∀ i : ι, _root_.finite $ ℤ ⧸ submodule.span ℤ {p i ^ e i} := λ i, finite.of_equiv _ (p i ^ e i).quotient_span_equiv_zmod.symm.to_equiv, haveI : _root_.finite ⨁ i, ℤ ⧸ (submodule.span ℤ {p i ^ e i} : submodule ℤ ℤ) := diff --git a/src/group_theory/specific_groups/dihedral.lean b/src/group_theory/specific_groups/dihedral.lean index e6c44a0b85319..15cc944e65216 100644 --- a/src/group_theory/specific_groups/dihedral.lean +++ b/src/group_theory/specific_groups/dihedral.lean @@ -107,14 +107,14 @@ private def fintype_helper : (zmod n ⊕ zmod n) ≃ dihedral_group n := /-- If `0 < n`, then `dihedral_group n` is a finite group. -/ -instance [fact (0 < n)] : fintype (dihedral_group n) := fintype.of_equiv _ fintype_helper +instance [ne_zero n] : fintype (dihedral_group n) := fintype.of_equiv _ fintype_helper instance : nontrivial (dihedral_group n) := ⟨⟨r 0, sr 0, dec_trivial⟩⟩ /-- If `0 < n`, then `dihedral_group n` has `2n` elements. -/ -lemma card [fact (0 < n)] : fintype.card (dihedral_group n) = 2 * n := +lemma card [ne_zero n] : fintype.card (dihedral_group n) = 2 * n := by rw [← fintype.card_eq.mpr ⟨fintype_helper⟩, fintype.card_sum, zmod.card, two_mul] @[simp] lemma r_one_pow (k : ℕ) : (r 1 : dihedral_group n) ^ k = r k := @@ -152,14 +152,15 @@ If `0 < n`, then `r 1` has order `n`. -/ @[simp] lemma order_of_r_one : order_of (r 1 : dihedral_group n) = n := begin - rcases n.eq_zero_or_pos with rfl | hn, + rcases eq_zero_or_ne_zero n with rfl | hn, { rw order_of_eq_zero_iff', intros n hn, rw [r_one_pow, one_def], apply mt r.inj, simpa using hn.ne' }, - { haveI := fact.mk hn, - apply (nat.le_of_dvd hn $ order_of_dvd_of_pow_eq_one $ @r_one_pow_n n).lt_or_eq.resolve_left, + { resetI, + apply (nat.le_of_dvd (ne_zero.pos n) $ order_of_dvd_of_pow_eq_one $ @r_one_pow_n n) + .lt_or_eq.resolve_left, intro h, have h1 : (r 1 : dihedral_group n)^(order_of (r 1)) = 1, { exact pow_order_of_eq_one _ }, @@ -172,7 +173,7 @@ end /-- If `0 < n`, then `i : zmod n` has order `n / gcd n i`. -/ -lemma order_of_r [fact (0 < n)] (i : zmod n) : order_of (r i) = n / nat.gcd n i.val := +lemma order_of_r [ne_zero n] (i : zmod n) : order_of (r i) = n / nat.gcd n i.val := begin conv_lhs { rw ←zmod.nat_cast_zmod_val i }, rw [←r_one_pow, order_of_pow, order_of_r_one] @@ -180,9 +181,9 @@ end lemma exponent : monoid.exponent (dihedral_group n) = lcm n 2 := begin - rcases n.eq_zero_or_pos with rfl | hn, + rcases eq_zero_or_ne_zero n with rfl | hn, { exact monoid.exponent_eq_zero_of_order_zero order_of_r_one }, - haveI := fact.mk hn, + resetI, apply nat.dvd_antisymm, { apply monoid.exponent_dvd_of_forall_pow_eq_one, rintro (m | m), diff --git a/src/group_theory/specific_groups/quaternion.lean b/src/group_theory/specific_groups/quaternion.lean index 00cbba859b363..fac7aa55f9907 100644 --- a/src/group_theory/specific_groups/quaternion.lean +++ b/src/group_theory/specific_groups/quaternion.lean @@ -148,24 +148,17 @@ def quaternion_group_zero_equiv_dihedral_group_zero : quaternion_group 0 ≃* di right_inv := by rintro (k | k); refl, map_mul' := by { rintros (k | k) (l | l); { dsimp, simp, }, } } -/-- Some of the lemmas on `zmod m` require that `m` is positive, as `m = 2 * n` is the case relevant -in this file but we don't want to write `[fact (0 < 2 * n)]` we make this lemma a local instance. -/ -private lemma succ_mul_pos_fact {m : ℕ} [hn : fact (0 < n)] : fact (0 < (nat.succ m) * n) := -⟨nat.succ_mul_pos m hn.1⟩ - -local attribute [instance] succ_mul_pos_fact - /-- If `0 < n`, then `quaternion_group n` is a finite group. -/ -instance [fact (0 < n)] : fintype (quaternion_group n) := fintype.of_equiv _ fintype_helper +instance [ne_zero n] : fintype (quaternion_group n) := fintype.of_equiv _ fintype_helper instance : nontrivial (quaternion_group n) := ⟨⟨a 0, xa 0, dec_trivial⟩⟩ /-- If `0 < n`, then `quaternion_group n` has `4n` elements. -/ -lemma card [fact (0 < n)] : fintype.card (quaternion_group n) = 4 * n := +lemma card [ne_zero n] : fintype.card (quaternion_group n) = 4 * n := begin rw [← fintype.card_eq.mpr ⟨fintype_helper⟩, fintype.card_sum, zmod.card, two_mul], ring @@ -205,7 +198,7 @@ end /-- If `0 < n`, then `xa i` has order 4. -/ -@[simp] lemma order_of_xa [hpos : fact (0 < n)] (i : zmod (2 * n)) : order_of (xa i) = 4 := +@[simp] lemma order_of_xa [ne_zero n] (i : zmod (2 * n)) : order_of (xa i) = 4 := begin change _ = 2^2, haveI : fact(nat.prime 2) := fact.mk (nat.prime_two), @@ -216,7 +209,7 @@ begin apply_fun zmod.val at h', apply_fun ( / n) at h', simp only [zmod.val_nat_cast, zmod.val_zero, nat.zero_div, nat.mod_mul_left_div_self, - nat.div_self hpos.1] at h', + nat.div_self (ne_zero.pos n)] at h', norm_num at h' }, { norm_num } end @@ -234,15 +227,15 @@ If `0 < n`, then `a 1` has order `2 * n`. -/ @[simp] lemma order_of_a_one : order_of (a 1 : quaternion_group n) = 2 * n := begin - rcases n.eq_zero_or_pos with rfl | hn, - { simp_rw [mul_zero, order_of_eq_zero_iff'], - intros n hn, + casesI eq_zero_or_ne_zero n with hn hn, + { subst hn, + simp_rw [mul_zero, order_of_eq_zero_iff'], + intros n h, rw [one_def, a_one_pow], apply mt a.inj, haveI : char_zero (zmod (2 * 0)) := zmod.char_zero, - simpa using hn.ne' }, - haveI := fact.mk hn, - apply (nat.le_of_dvd (nat.succ_mul_pos _ hn) + simpa using h.ne' }, + apply (nat.le_of_dvd (ne_zero.pos _) (order_of_dvd_of_pow_eq_one (@a_one_pow_n n))).lt_or_eq.resolve_left, intro h, have h1 : (a 1 : quaternion_group n)^(order_of (a 1)) = 1 := pow_order_of_eq_one _, @@ -255,7 +248,7 @@ end /-- If `0 < n`, then `a i` has order `(2 * n) / gcd (2 * n) i`. -/ -lemma order_of_a [fact (0 < n)] (i : zmod (2 * n)) : +lemma order_of_a [ne_zero n] (i : zmod (2 * n)) : order_of (a i) = (2 * n) / nat.gcd (2 * n) i.val := begin conv_lhs { rw ← zmod.nat_cast_zmod_val i }, @@ -266,10 +259,10 @@ lemma exponent : monoid.exponent (quaternion_group n) = 2 * lcm n 2 := begin rw [←normalize_eq 2, ←lcm_mul_left, normalize_eq], norm_num, - rcases n.eq_zero_or_pos with rfl | hn, - { simp only [lcm_zero_left, mul_zero], + casesI eq_zero_or_ne_zero n with hn hn, + { subst hn, + simp only [lcm_zero_left, mul_zero], exact monoid.exponent_eq_zero_of_order_zero order_of_a_one }, - haveI := fact.mk hn, apply nat.dvd_antisymm, { apply monoid.exponent_dvd_of_forall_pow_eq_one, rintro (m | m), diff --git a/src/number_theory/cyclotomic/gal.lean b/src/number_theory/cyclotomic/gal.lean index 42b0de42698bf..528db295c59e2 100644 --- a/src/number_theory/cyclotomic/gal.lean +++ b/src/number_theory/cyclotomic/gal.lean @@ -37,8 +37,6 @@ it is always a subgroup, and if the `n`th cyclotomic polynomial is irreducible, -/ -local attribute [instance] pnat.fact_pos - variables {n : ℕ+} (K : Type*) [field K] {L : Type*} {μ : L} open polynomial is_cyclotomic_extension diff --git a/src/number_theory/legendre_symbol/add_character.lean b/src/number_theory/legendre_symbol/add_character.lean index 8d0904f7138cf..68737dcbd10bf 100644 --- a/src/number_theory/legendre_symbol/add_character.lean +++ b/src/number_theory/legendre_symbol/add_character.lean @@ -241,10 +241,6 @@ structure primitive_add_char (R : Type u) [comm_ring R] [fintype R] (R' : Type v variables {C : Type v} [comm_ring C] --- For `n : ℕ+`, automatically generate instance `fact (0 < (n : ℕ))`. --- This is needed for the API for `zmod n` (until that gets refactored to use `ne_zero`). -local attribute [instance] pnat.fact_pos - section zmod_char_def open multiplicative -- so we can write simply `to_add`, which we need here again diff --git a/src/number_theory/lucas_lehmer.lean b/src/number_theory/lucas_lehmer.lean index 7dcc83f26b148..788f96290d363 100644 --- a/src/number_theory/lucas_lehmer.lean +++ b/src/number_theory/lucas_lehmer.lean @@ -160,9 +160,6 @@ def lucas_lehmer_test (p : ℕ) : Prop := lucas_lehmer_residue p = 0 /-- `q` is defined as the minimum factor of `mersenne p`, bundled as an `ℕ+`. -/ def q (p : ℕ) : ℕ+ := ⟨nat.min_fac (mersenne p), nat.min_fac_pos (mersenne p)⟩ -local attribute [instance] -lemma fact_pnat_pos (q : ℕ+) : fact (0 < (q : ℕ)) := ⟨q.2⟩ - /-- We construct the ring `X q` as ℤ/qℤ + √3 ℤ/qℤ. -/ -- It would be nice to define this as (ℤ/qℤ)[x] / (x^2 - 3), -- obtaining the ring structure for free, diff --git a/src/number_theory/lucas_primality.lean b/src/number_theory/lucas_primality.lean index b7d0084fffc26..8ac50c5181f02 100644 --- a/src/number_theory/lucas_primality.lean +++ b/src/number_theory/lucas_primality.lean @@ -47,7 +47,7 @@ begin have order_of_a : order_of a = p-1, { apply order_of_eq_of_pow_and_pow_div_prime _ ha hd, exact tsub_pos_of_lt hp1, }, - haveI fhp0 : fact (0 < p) := ⟨h0.bot_lt⟩, + haveI : ne_zero p := ⟨h0⟩, rw nat.prime_iff_card_units, -- Prove cardinality of `units` of `zmod p` is both `≤ p-1` and `≥ p-1` refine le_antisymm (nat.card_units_zmod_lt_sub_one hp1) _, diff --git a/src/number_theory/padics/ring_homs.lean b/src/number_theory/padics/ring_homs.lean index b6006ffc1ef20..3172279bbab27 100644 --- a/src/number_theory/padics/ring_homs.lean +++ b/src/number_theory/padics/ring_homs.lean @@ -44,7 +44,7 @@ open nat local_ring padic namespace padic_int -variables {p : ℕ} [hp_prime : fact (p.prime)] +variables {p : ℕ} [hp_prime : fact p.prime] include hp_prime section ring_homs @@ -496,8 +496,6 @@ begin int.cast_sub], dsimp [nth_hom], rw [← f_compat, ring_hom.comp_apply], - haveI : fact (p ^ i > 0) := ⟨pow_pos hp_prime.1.pos _⟩, - haveI : fact (p ^ j > 0) := ⟨pow_pos hp_prime.1.pos _⟩, simp only [zmod.cast_id, zmod.cast_hom_apply, sub_self, zmod.nat_cast_val, zmod.int_cast_cast], end @@ -520,14 +518,16 @@ The `n`th value of the sequence is `((f n r).val : ℚ)`. -/ def nth_hom_seq (r : R) : padic_seq p := ⟨λ n, nth_hom f r n, is_cau_seq_nth_hom f_compat r⟩ +-- this lemma ran into issues after changing to `ne_zero` and I'm not sure why. lemma nth_hom_seq_one : nth_hom_seq f_compat 1 ≈ 1 := begin intros ε hε, change _ < _ at hε, use 1, intros j hj, - haveI : fact (1 < p^j) := ⟨nat.one_lt_pow _ _ (by linarith) hp_prime.1.one_lt⟩, - simp [nth_hom_seq, nth_hom, zmod.val_one, hε], + haveI : fact (1 < p ^ j) := ⟨nat.one_lt_pow _ _ (by linarith) hp_prime.1.one_lt⟩, + suffices : ((1 : zmod (p ^ j)) : ℚ) = 1, by simp [nth_hom_seq, nth_hom, this, hε], + rw [zmod.cast_eq_val, zmod.val_one, nat.cast_one] end lemma nth_hom_seq_add (r s : R) : @@ -542,8 +542,6 @@ begin rw [← int.cast_add, ← int.cast_sub, ← padic_norm.dvd_iff_norm_le, ← zmod.int_coe_zmod_eq_zero_iff_dvd], dsimp [nth_hom], - haveI : fact (p ^ n > 0) := ⟨pow_pos hp_prime.1.pos _⟩, - haveI : fact (p ^ j > 0) := ⟨pow_pos hp_prime.1.pos _⟩, simp only [zmod.nat_cast_val, ring_hom.map_add, int.cast_sub, zmod.int_cast_cast, int.cast_add], rw [zmod.cast_add (show p^n ∣ p^j, from pow_dvd_pow _ hj), sub_self], { apply_instance }, @@ -561,8 +559,6 @@ begin rw [← int.cast_mul, ← int.cast_sub, ← padic_norm.dvd_iff_norm_le, ← zmod.int_coe_zmod_eq_zero_iff_dvd], dsimp [nth_hom], - haveI : fact (p ^ n > 0) := ⟨pow_pos hp_prime.1.pos _⟩, - haveI : fact (p ^ j > 0) := ⟨pow_pos hp_prime.1.pos _⟩, simp only [zmod.nat_cast_val, ring_hom.map_mul, int.cast_sub, zmod.int_cast_cast, int.cast_mul], rw [zmod.cast_mul (show p^n ∣ p^j, from pow_dvd_pow _ hj), sub_self], { apply_instance }, @@ -641,7 +637,6 @@ See also `padic_int.lift_unique`. lemma lift_spec (n : ℕ) : (to_zmod_pow n).comp (lift f_compat) = f n := begin ext r, - haveI : fact (0 < p ^ n) := ⟨pow_pos hp_prime.1.pos n⟩, rw [ring_hom.comp_apply, ← zmod.nat_cast_zmod_val (f n r), ← map_nat_cast $ to_zmod_pow n, ← sub_eq_zero, ← ring_hom.map_sub, ← ring_hom.mem_ker, ker_to_zmod_pow], apply lift_sub_val_mem_span, diff --git a/src/number_theory/ramification_inertia.lean b/src/number_theory/ramification_inertia.lean index 34e1b5a5596cb..5a54db5a86d19 100644 --- a/src/number_theory/ramification_inertia.lean +++ b/src/number_theory/ramification_inertia.lean @@ -457,7 +457,7 @@ quotient.algebra_quotient_of_le_comap (ideal.map_le_iff_le_comap.mp le_pow_ramif algebra_map (R ⧸ p) (S ⧸ P ^ e) (ideal.quotient.mk p x) = ideal.quotient.mk _ (f x) := rfl -variables [hfp : fact (ramification_idx f p P ≠ 0)] +variables [hfp : ne_zero (ramification_idx f p P)] include hfp /-- If `P` lies over `p`, then `R / p` has a canonical map to `S / P`. @@ -645,7 +645,7 @@ lemma dim_prime_pow_ramification_idx [is_domain S] [is_dedekind_domain S] [p.is_ e • @module.rank (R ⧸ p) (S ⧸ P) _ _ (@algebra.to_module _ _ _ _ $ @@quotient.algebra_quotient_of_ramification_idx_ne_zero _ _ _ _ _ ⟨he⟩) := begin - letI : fact (e ≠ 0) := ⟨he⟩, + letI : ne_zero e := ⟨he⟩, have := dim_pow_quot f p P hP0 0 (nat.zero_le e), rw [pow_zero, nat.sub_zero, ideal.one_eq_top, ideal.map_top] at this, exact (dim_top (R ⧸ p) _).symm.trans this @@ -659,7 +659,7 @@ lemma finrank_prime_pow_ramification_idx [is_domain S] [is_dedekind_domain S] e * @finrank (R ⧸ p) (S ⧸ P) _ _ (@algebra.to_module _ _ _ _ $ @@quotient.algebra_quotient_of_ramification_idx_ne_zero _ _ _ _ _ ⟨he⟩) := begin - letI : fact (e ≠ 0) := ⟨he⟩, + letI : ne_zero e := ⟨he⟩, letI : algebra (R ⧸ p) (S ⧸ P) := quotient.algebra_quotient_of_ramification_idx_ne_zero f p P, letI := ideal.quotient.field p, have hdim := dim_prime_pow_ramification_idx _ _ _ hP0 he, @@ -698,7 +698,7 @@ is_dedekind_domain.ramification_idx_ne_zero (ideal.le_of_dvd (dvd_of_mem_factors (multiset.mem_to_finset.mp P.2))) instance factors.fact_ramification_idx_ne_zero (P : (factors (map (algebra_map R S) p)).to_finset) : - fact (ramification_idx (algebra_map R S) p P ≠ 0) := + ne_zero (ramification_idx (algebra_map R S) p P) := ⟨factors.ramification_idx_ne_zero p P⟩ local attribute [instance] quotient.algebra_quotient_of_ramification_idx_ne_zero diff --git a/src/number_theory/sum_four_squares.lean b/src/number_theory/sum_four_squares.lean index c9f5476093a62..5c36439501b47 100644 --- a/src/number_theory/sum_four_squares.lean +++ b/src/number_theory/sum_four_squares.lean @@ -105,7 +105,7 @@ let ⟨x, hx⟩ := h01 in let ⟨y, hy⟩ := h23 in simpa [finset.sum_eq_multiset_sum, fin4univ, multiset.sum_cons, f, add_assoc] end⟩ -private lemma prime_sum_four_squares (p : ℕ) [hp : _root_.fact p.prime] : +private lemma prime_sum_four_squares (p : ℕ) [hp : fact p.prime] : ∃ a b c d : ℤ, a^2 + b^2 + c^2 + d^2 = p := have hm : ∃ m < p, 0 < m ∧ ∃ a b c d : ℤ, a^2 + b^2 + c^2 + d^2 = m * p, from let ⟨a, b, k, hk⟩ := exists_sq_add_sq_add_one_eq_k p in @@ -116,7 +116,7 @@ have hm : ∃ m < p, 0 < m ∧ ∃ a b c d : ℤ, a^2 + b^2 + c^2 + d^2 = m * p, a, b, 1, 0, by simpa [sq] using hk.1⟩, let m := nat.find hm in let ⟨a, b, c, d, (habcd : a^2 + b^2 + c^2 + d^2 = m * p)⟩ := (nat.find_spec hm).snd.2 in -by haveI hm0 : _root_.fact (0 < m) := ⟨(nat.find_spec hm).snd.1⟩; exact +by haveI hm0 : ne_zero m := ne_zero.of_pos (nat.find_spec hm).snd.1; exact have hmp : m < p, from (nat.find_spec hm).fst, m.mod_two_eq_zero_or_one.elim (λ hm2 : m % 2 = 0, @@ -167,7 +167,7 @@ m.mod_two_eq_zero_or_one.elim ⟨mc, hmc⟩ := habcd0.2.2.1, ⟨md, hmd⟩ := habcd0.2.2.2 in have hmdvdp : m ∣ p, from int.coe_nat_dvd.1 ⟨ma^2 + mb^2 + mc^2 + md^2, - (mul_right_inj' (show (m : ℤ) ≠ 0, from int.coe_nat_ne_zero_iff_pos.2 hm0.1)).1 $ + (mul_right_inj' (show (m : ℤ) ≠ 0, from int.coe_nat_ne_zero.2 hm0.1)).1 $ by { rw [← habcd, hma, hmb, hmc, hmd], ring }⟩, (hp.1.eq_one_or_self_of_dvd _ hmdvdp).elim hm1 (λ hmeqp, by simpa [lt_irrefl, hmeqp] using hmp)), @@ -183,14 +183,14 @@ m.mod_two_eq_zero_or_one.elim have hn_nonneg : 0 ≤ n, from nonneg_of_mul_nonneg_right (by { erw [← hn], repeat {try {refine add_nonneg _ _}, try {exact sq_nonneg _}} }) - (int.coe_nat_pos.2 hm0.1), + (int.coe_nat_pos.2 $ ne_zero.pos m), have hnm : n.nat_abs < m, from int.coe_nat_lt.1 (lt_of_mul_lt_mul_left (by { rw [int.nat_abs_of_nonneg hn_nonneg, ← hn, ← sq], exact hwxyzlt }) (int.coe_nat_nonneg m)), have hstuv : s^2 + t^2 + u^2 + v^2 = n.nat_abs * p, from (mul_right_inj' (show (m^2 : ℤ) ≠ 0, from pow_ne_zero 2 - (int.coe_nat_ne_zero_iff_pos.2 hm0.1))).1 $ + (int.coe_nat_ne_zero.2 hm0.1))).1 $ calc (m : ℤ)^2 * (s^2 + t^2 + u^2 + v^2) = ((m : ℕ) * s)^2 + ((m : ℕ) * t)^2 + ((m : ℕ) * u)^2 + ((m : ℕ) * v)^2 : by { simp [mul_pow], ring } @@ -204,7 +204,7 @@ lemma sum_four_squares : ∀ n : ℕ, ∃ a b c d : ℕ, a^2 + b^2 + c^2 + d^2 = | 0 := ⟨0, 0, 0, 0, rfl⟩ | 1 := ⟨1, 0, 0, 0, rfl⟩ | n@(k+2) := -have hm : _root_.fact (min_fac (k+2)).prime := ⟨min_fac_prime dec_trivial⟩, +have hm : fact (min_fac (k+2)).prime := ⟨min_fac_prime dec_trivial⟩, have n / min_fac n < n := factors_lemma, let ⟨a, b, c, d, h₁⟩ := show ∃ a b c d : ℤ, a^2 + b^2 + c^2 + d^2 = min_fac n, by exactI prime_sum_four_squares (min_fac (k+2)) in diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity.lean index 0f01a413cbe08..0bab3118e50aa 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity.lean @@ -562,8 +562,8 @@ begin obtain ⟨m, hm⟩ := multiplicity.exists_eq_pow_mul_and_not_dvd hfin, by_cases hp : p ∣ n, { obtain ⟨k, hk⟩ := nat.exists_eq_succ_of_ne_zero (multiplicity.pos_of_dvd hfin hp).ne', - haveI hpri : fact p.prime := - @char_p.char_is_prime_of_pos R _ _ _ p ⟨nat.pos_of_dvd_of_pos hp n.pos⟩ _, + haveI : ne_zero p := ne_zero.of_pos (nat.pos_of_dvd_of_pos hp n.pos), + haveI hpri : fact p.prime := char_p.char_is_prime_of_pos R p, have := hζ.pow_eq_one, rw [hm.1, hk, pow_succ, mul_assoc, pow_mul', ← frobenius_def, ← frobenius_one p] at this, exfalso, @@ -663,7 +663,6 @@ lemma zpowers_eq {k : ℕ+} {ζ : Rˣ} (h : is_primitive_root ζ k) : subgroup.zpowers ζ = roots_of_unity k R := begin apply set_like.coe_injective, - haveI : fact (0 < (k : ℕ)) := ⟨k.pos⟩, haveI F : fintype (subgroup.zpowers ζ) := fintype.of_equiv _ (h.zmod_equiv_zpowers).to_equiv, refine @set.eq_of_subset_of_card_le Rˣ (subgroup.zpowers ζ) (roots_of_unity k R) F (roots_of_unity.fintype R k) @@ -728,7 +727,6 @@ end lemma card_roots_of_unity' {n : ℕ+} (h : is_primitive_root ζ n) : fintype.card (roots_of_unity n R) = n := begin - haveI : fact (0 < ↑n) := ⟨n.pos⟩, let e := h.zmod_equiv_zpowers, haveI F : fintype (subgroup.zpowers ζ) := fintype.of_equiv _ e.to_equiv, calc fintype.card (roots_of_unity n R) diff --git a/src/ring_theory/witt_vector/witt_polynomial.lean b/src/ring_theory/witt_vector/witt_polynomial.lean index eebf93e8d0dfb..fdca5dec0355c 100644 --- a/src/ring_theory/witt_vector/witt_polynomial.lean +++ b/src/ring_theory/witt_vector/witt_polynomial.lean @@ -150,8 +150,8 @@ begin end section p_prime --- in fact, `0 < p` would be sufficient -variables [hp : fact p.prime] + +variables [hp : ne_zero p] include hp lemma witt_polynomial_vars [char_zero R] (n : ℕ) : @@ -159,9 +159,9 @@ lemma witt_polynomial_vars [char_zero R] (n : ℕ) : begin have : ∀ i, (monomial (finsupp.single i (p ^ (n - i))) (p ^ i : R)).vars = {i}, { intro i, - refine vars_monomial_single i (pow_ne_zero _ hp.1.ne_zero) _, + refine vars_monomial_single i (pow_ne_zero _ hp.1) _, rw [← nat.cast_pow, nat.cast_ne_zero], - exact pow_ne_zero i hp.1.ne_zero }, + exact pow_ne_zero i hp.1 }, rw [witt_polynomial, vars_sum_of_disjoint], { simp only [this, bUnion_singleton_eq_self], }, { simp only [this], diff --git a/src/tactic/norm_fin.lean b/src/tactic/norm_fin.lean index b2a2013719555..a224738290e5c 100644 --- a/src/tactic/norm_fin.lean +++ b/src/tactic/norm_fin.lean @@ -41,7 +41,7 @@ def normalize_fin_lt (n : ℕ) (a : fin n) (b : ℕ) := a.1 = b theorem normalize_fin_lt.coe {n} {a : fin n} {b : ℕ} (h : normalize_fin_lt n a b) : ↑a = b := h -theorem normalize_fin_iff {n} [fact (0 < n)] {a b} : +theorem normalize_fin_iff {n : ℕ} [ne_zero n] {a b} : normalize_fin n a b ↔ a = fin.of_nat' b := iff.symm (fin.eq_iff_veq _ _) diff --git a/src/testing/slim_check/sampleable.lean b/src/testing/slim_check/sampleable.lean index 25db1e8fa4359..27710eef65f21 100644 --- a/src/testing/slim_check/sampleable.lean +++ b/src/testing/slim_check/sampleable.lean @@ -281,7 +281,7 @@ well_founded.fix has_well_founded.wf $ λ x f_rec, y ← (shrink x).find (λ a, p a), f_rec y y.property <|> some y.val . -instance fin.sampleable {n} [fact $ 0 < n] : sampleable (fin n) := +instance fin.sampleable {n : ℕ} [ne_zero n] : sampleable (fin n) := sampleable.lift ℕ fin.of_nat' subtype.val $ λ i, (mod_le _ _ : i % n ≤ i) From e6c6dfd508deb9f9148f2dd0b509eba41eaec883 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Mon, 29 Aug 2022 20:01:00 +0000 Subject: [PATCH 079/828] chore(analysis/special_functions): correct typo (#16280) --- src/analysis/special_functions/pow.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 5b6eacd69bc38..014d93414a158 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -1069,7 +1069,7 @@ lemma tendsto_log_div_rpow_nhds_zero {r : ℝ} (hr : r < 0) : tendsto (λ x, log x / x ^ r) (𝓝[>] 0) (𝓝 0) := (is_o_log_rpow_nhds_zero hr).tendsto_div_nhds_zero -lemma tensdto_log_mul_rpow_nhds_zero {r : ℝ} (hr : 0 < r) : +lemma tendsto_log_mul_rpow_nhds_zero {r : ℝ} (hr : 0 < r) : tendsto (λ x, log x * x ^ r) (𝓝[>] 0) (𝓝 0) := (tendsto_log_div_rpow_nhds_zero $ neg_lt_zero.2 hr).congr' $ eventually_mem_nhds_within.mono $ λ x hx, by rw [rpow_neg hx.out.le, div_inv_eq_mul] From 17ab30b2f1d096e41d84dc9b522c6f14be5f7743 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Tue, 30 Aug 2022 00:38:43 +0000 Subject: [PATCH 080/828] feat(field_theory/normal): Define the normal closure (#16144) This PR adds the definition of the normal closure, and proves that the normal closure is finite-dimensional and normal. --- src/field_theory/normal.lean | 63 ++++++++++++++++++++++++++++++++++++ 1 file changed, 63 insertions(+) diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index 297b7620d247f..5e104283fa2ec 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -219,6 +219,12 @@ begin exact polynomial.splits_comp_of_splits _ (inclusion hE).to_ring_hom this, end +variables {F K} {L : Type*} [field L] [algebra F L] [algebra K L] [is_scalar_tower F K L] + +@[simp] lemma restrict_scalars_normal {E : intermediate_field K L} : + normal F (E.restrict_scalars F) ↔ normal F E := +iff.rfl + end intermediate_field variables {F} {K} {K₁ K₂ K₃:Type*} [field K₁] [field K₂] [field K₃] @@ -375,3 +381,60 @@ begin end end lift + +section normal_closure + +open intermediate_field + +variables (F K) (L : Type*) [field L] [algebra F L] [algebra K L] [is_scalar_tower F K L] + +/-- The normal closure of `K` in `L`. -/ +noncomputable! def normal_closure : intermediate_field K L := +{ algebra_map_mem' := λ r, le_supr (λ f : K →ₐ[F] L, f.field_range) + (is_scalar_tower.to_alg_hom F K L) ⟨r, rfl⟩, + .. (⨆ f : K →ₐ[F] L, f.field_range).to_subfield } + +namespace normal_closure + +lemma restrict_scalars_eq_supr_adjoin [h : normal F L] : + (normal_closure F K L).restrict_scalars F = ⨆ x : K, adjoin F ((minpoly F x).root_set L) := +begin + refine le_antisymm (supr_le _) (supr_le (λ x, adjoin_le_iff.mpr (λ y hy, _))), + { rintros f _ ⟨x, rfl⟩, + refine le_supr (λ x, adjoin F ((minpoly F x).root_set L)) x + (subset_adjoin F ((minpoly F x).root_set L) _), + rw [polynomial.mem_root_set, alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom, + polynomial.aeval_alg_hom_apply, minpoly.aeval, map_zero], + exact minpoly.ne_zero ((is_integral_algebra_map_iff (algebra_map K L).injective).mp + (h.is_integral (algebra_map K L x))) }, + { rw [polynomial.root_set, finset.mem_coe, multiset.mem_to_finset] at hy, + let g := (alg_hom_adjoin_integral_equiv F ((is_integral_algebra_map_iff + (algebra_map K L).injective).mp (h.is_integral (algebra_map K L x)))).symm ⟨y, hy⟩, + refine le_supr (λ f : K →ₐ[F] L, f.field_range) + ((g.lift_normal L).comp (is_scalar_tower.to_alg_hom F K L)) + ⟨x, (g.lift_normal_commutes L (adjoin_simple.gen F x)).trans _⟩, + rw [algebra.id.map_eq_id, ring_hom.id_apply], + apply power_basis.lift_gen }, +end + +instance normal [h : normal F L] : normal F (normal_closure F K L) := +let ϕ := algebra_map K L in begin + rw [←intermediate_field.restrict_scalars_normal, restrict_scalars_eq_supr_adjoin], + apply intermediate_field.normal_supr F L _, + intro x, + apply normal.of_is_splitting_field (minpoly F x), + exact adjoin_root_set_is_splitting_field ((minpoly.eq_of_algebra_map_eq ϕ.injective + ((is_integral_algebra_map_iff ϕ.injective).mp (h.is_integral (ϕ x))) rfl).symm ▸ h.splits _), +end + +instance is_finite_dimensional [finite_dimensional F K] : + finite_dimensional F (normal_closure F K L) := +begin + haveI : ∀ f : K →ₐ[F] L, finite_dimensional F f.field_range := + λ f, f.to_linear_map.finite_dimensional_range, + apply intermediate_field.finite_dimensional_supr_of_finite, +end + +end normal_closure + +end normal_closure From 20f94541a6e7ddf3916a6a49cb96ff54b5effb57 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Tue, 30 Aug 2022 01:42:25 +0000 Subject: [PATCH 081/828] feat(ring_theory/unique_factorization_domain): add `associated_iff_normalized_factors_eq_normalized_factors` (#16221) --- src/ring_theory/unique_factorization_domain.lean | 15 +++++++++++---- 1 file changed, 11 insertions(+), 4 deletions(-) diff --git a/src/ring_theory/unique_factorization_domain.lean b/src/ring_theory/unique_factorization_domain.lean index d8f2fe739ec38..268c6a5c17543 100644 --- a/src/ring_theory/unique_factorization_domain.lean +++ b/src/ring_theory/unique_factorization_domain.lean @@ -604,6 +604,15 @@ begin apply multiset.prod_dvd_prod_of_le } end +lemma associated_iff_normalized_factors_eq_normalized_factors {x y : α} (hx : x ≠ 0) (hy : y ≠ 0) : + x ~ᵤ y ↔ normalized_factors x = normalized_factors y := +begin + refine ⟨λ h, _, + λ h, (normalized_factors_prod hx).symm.trans (trans (by rw h) (normalized_factors_prod hy))⟩, + apply le_antisymm; rw [← dvd_iff_normalized_factors_le_normalized_factors], + all_goals { simp [*, h.dvd, h.symm.dvd], }, +end + theorem normalized_factors_of_irreducible_pow {p : α} (hp : irreducible p) (k : ℕ) : normalized_factors (p ^ k) = multiset.repeat (normalize p) k := by rw [normalized_factors_pow, normalized_factors_irreducible hp, multiset.nsmul_singleton] @@ -1672,10 +1681,8 @@ by { ext, simp [factorization] } lemma associated_of_factorization_eq (a b: α) (ha: a ≠ 0) (hb: b ≠ 0) (h: factorization a = factorization b) : associated a b := begin - simp only [factorization, add_equiv.apply_eq_iff_eq] at h, - have ha' := normalized_factors_prod ha, - rw h at ha', - exact associated.trans ha'.symm (normalized_factors_prod hb), + simp_rw [factorization, add_equiv.apply_eq_iff_eq] at h, + rwa [associated_iff_normalized_factors_eq_normalized_factors ha hb], end end finsupp From 3683b34a79d34b3088e0a87e5412c377bb3bd28f Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Tue, 30 Aug 2022 03:46:59 +0000 Subject: [PATCH 082/828] chore(scripts): update nolints.txt (#16305) I am happy to remove some nolints for you! --- scripts/nolints.txt | 7 ------- 1 file changed, 7 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 2beebcbc31830..d94ea9858ab94 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -692,7 +692,6 @@ apply_nolint mv_polynomial.is_prime.vanishing_ideal_zero_locus fintype_finite apply_nolint mv_polynomial.vanishing_ideal_zero_locus_eq_radical fintype_finite -- ring_theory/polynomial/basic.lean -apply_nolint mv_polynomial.is_domain_fintype fintype_finite apply_nolint mv_polynomial.is_noetherian_ring fintype_finite apply_nolint mv_polynomial.map_mv_polynomial_eq_eval₂ fintype_finite @@ -703,12 +702,6 @@ apply_nolint algebra.trace_trace_of_basis fintype_finite -- ring_theory/witt_vector/basic.lean apply_nolint witt_vector.comm_ring check_reducibility --- set_theory/cardinal/ordinal.lean -apply_nolint cardinal.extend_function_finite fintype_finite -apply_nolint cardinal.mk_compl_eq_mk_compl_finite fintype_finite -apply_nolint cardinal.mk_compl_eq_mk_compl_finite_lift fintype_finite -apply_nolint cardinal.mk_compl_eq_mk_compl_finite_same fintype_finite - -- set_theory/lists.lean apply_nolint finsets doc_blame From ed23151629ccb637b04b750d0ace801e3fed5d4a Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Tue, 30 Aug 2022 06:55:36 +0000 Subject: [PATCH 083/828] feat(data/list/rdrop): drop or take from the right (#15475) --- src/data/list/rdrop.lean | 205 +++++++++++++++++++++++++++++++++++++++ 1 file changed, 205 insertions(+) create mode 100644 src/data/list/rdrop.lean diff --git a/src/data/list/rdrop.lean b/src/data/list/rdrop.lean new file mode 100644 index 0000000000000..be776a0524591 --- /dev/null +++ b/src/data/list/rdrop.lean @@ -0,0 +1,205 @@ +/- +Copyright (c) 2022 Yakov Pechersky. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yakov Pechersky +-/ + +import data.list.basic +import data.list.infix + +/-! + +# Dropping or taking from lists on the right + +Taking or removing element from the tail end of a list + +## Main defintions + +- `rdrop n`: drop `n : ℕ` elements from the tail +- `rtake n`: take `n : ℕ` elements from the tail +- `drop_while p`: remove all the elements that satisfy a decidable `p : α → Prop` from the tail of + a list until hitting the first non-satisfying element +- `take_while p`: take all the elements that satisfy a decidable `p : α → Prop` from the tail of + a list until hitting the first non-satisfying element + +## Implementation detail + +The two predicate-based methods operate by performing the regular "from-left" operation on +`list.reverse`, followed by another `list.reverse`, so they are not the most performant. +The other two rely on `list.length l` so they still traverse the list twice. One could construct +another function that takes a `L : ℕ` and use `L - n`. Under a proof condition that +`L = l.length`, the function would do the right thing. + +-/ + +variables {α : Type*} (p : α → Prop) [decidable_pred p] (l : list α) (n : ℕ) + +namespace list + +/-- Drop `n` elements from the tail end of a list. -/ +def rdrop : list α := l.take (l.length - n) + +@[simp] lemma rdrop_nil : rdrop ([] : list α) n = [] := by simp [rdrop] +@[simp] lemma rdrop_zero : rdrop l 0 = l := by simp [rdrop] + +lemma rdrop_eq_reverse_drop_reverse : l.rdrop n = reverse (l.reverse.drop n) := +begin + rw rdrop, + induction l using list.reverse_rec_on with xs x IH generalizing n, + { simp }, + { cases n, + { simp [take_append] }, + { simp [take_append_eq_append_take, IH] } } +end + +@[simp] lemma rdrop_concat_succ (x : α) : rdrop (l ++ [x]) (n + 1) = rdrop l n := +by simp [rdrop_eq_reverse_drop_reverse] + +/-- Take `n` elements from the tail end of a list. -/ +def rtake : list α := l.drop (l.length - n) + +@[simp] lemma rtake_nil : rtake ([] : list α) n = [] := by simp [rtake] +@[simp] lemma rtake_zero : rtake l 0 = [] := by simp [rtake] + +lemma rtake_eq_reverse_take_reverse : l.rtake n = reverse (l.reverse.take n) := +begin + rw rtake, + induction l using list.reverse_rec_on with xs x IH generalizing n, + { simp }, + { cases n, + { simp }, + { simp [drop_append_eq_append_drop, IH] } } +end + +@[simp] lemma rtake_concat_succ (x : α) : rtake (l ++ [x]) (n + 1) = rtake l n ++ [x] := +by simp [rtake_eq_reverse_take_reverse] + +/-- Drop elements from the tail end of a list that satisfy `p : α → Prop`. +Implemented naively via `list.reverse` -/ +def rdrop_while : list α := reverse (l.reverse.drop_while p) + +@[simp] lemma rdrop_while_nil : rdrop_while p ([] : list α) = [] := +by simp [rdrop_while, drop_while] +lemma rdrop_while_concat (x : α) : + rdrop_while p (l ++ [x]) = if p x then rdrop_while p l else l ++ [x] := +begin + simp only [rdrop_while, drop_while, reverse_append, reverse_singleton, singleton_append], + split_ifs with h h; + simp [h] +end +@[simp] lemma rdrop_while_concat_pos (x : α) (h : p x) : + rdrop_while p (l ++ [x]) = rdrop_while p l := +by rw [rdrop_while_concat, if_pos h] +@[simp] lemma rdrop_while_concat_neg (x : α) (h : ¬ p x) : + rdrop_while p (l ++ [x]) = l ++ [x] := +by rw [rdrop_while_concat, if_neg h] + +lemma rdrop_while_singleton (x : α) : + rdrop_while p [x] = if p x then [] else [x] := +by rw [←nil_append [x], rdrop_while_concat, rdrop_while_nil] + +lemma rdrop_while_last_not (hl : (l.rdrop_while p) ≠ []): + ¬ p ((rdrop_while p l).last hl) := +begin + simp_rw rdrop_while, + rw last_reverse, + exact drop_while_nth_le_zero_not _ _ _ +end + +lemma rdrop_while_prefix : l.rdrop_while p <+: l := +begin + rw [←reverse_suffix, rdrop_while, reverse_reverse], + exact drop_while_suffix _ +end + +variables {p} {l} + +@[simp] lemma rdrop_while_eq_nil_iff : rdrop_while p l = [] ↔ ∀ x ∈ l, p x := +by simp [rdrop_while] + +-- it is in this file because it requires `list.infix` +@[simp] lemma drop_while_eq_self_iff : + drop_while p l = l ↔ ∀ (hl : 0 < l.length), ¬ p (l.nth_le 0 hl) := +begin + induction l with hd tl IH, + { simp }, + { rw drop_while, + split_ifs, + { simp only [h, length, nth_le, nat.succ_pos', not_true, forall_true_left, iff_false], + intro H, + refine (cons_ne_self hd tl) (sublist.antisymm _ (sublist_cons _ _)), + rw ←H, + exact (drop_while_suffix _).sublist }, + { simp [h] } } +end + +@[simp] lemma rdrop_while_eq_self_iff : rdrop_while p l = l ↔ ∀ (hl : l ≠ []), ¬ p (l.last hl) := +begin + simp only [rdrop_while, reverse_eq_iff, length_reverse, ne.def, drop_while_eq_self_iff, + last_eq_nth_le, ←length_eq_zero, pos_iff_ne_zero], + refine forall_congr _, + intro h, + rw [nth_le_reverse'], + { simp }, + { rw [←ne.def, ←pos_iff_ne_zero] at h, + simp [tsub_lt_iff_right (nat.succ_le_of_lt h)] } +end + +variables (p) (l) + +lemma drop_while_idempotent : drop_while p (drop_while p l) = drop_while p l := +drop_while_eq_self_iff.mpr (drop_while_nth_le_zero_not _ _) + +lemma rdrop_while_idempotent : rdrop_while p (rdrop_while p l) = rdrop_while p l := +rdrop_while_eq_self_iff.mpr (rdrop_while_last_not _ _) + +/-- Take elements from the tail end of a list that satisfy `p : α → Prop`. +Implemented naively via `list.reverse` -/ +def rtake_while : list α := reverse (l.reverse.take_while p) + +@[simp] lemma rtake_while_nil : rtake_while p ([] : list α) = [] := +by simp [rtake_while, take_while] +lemma rtake_while_concat (x : α) : + rtake_while p (l ++ [x]) = if p x then rtake_while p l ++ [x] else [] := +begin + simp only [rtake_while, take_while, reverse_append, reverse_singleton, singleton_append], + split_ifs with h h; + simp [h] +end +@[simp] lemma rtake_while_concat_pos (x : α) (h : p x) : + rtake_while p (l ++ [x]) = rtake_while p l ++ [x] := +by rw [rtake_while_concat, if_pos h] +@[simp] lemma rtake_while_concat_neg (x : α) (h : ¬ p x) : + rtake_while p (l ++ [x]) = [] := +by rw [rtake_while_concat, if_neg h] + +lemma rtake_while_suffix : l.rtake_while p <:+ l := +begin + rw [←reverse_prefix, rtake_while, reverse_reverse], + exact take_while_prefix _ +end + +variables {p} {l} + +@[simp] lemma rtake_while_eq_self_iff : rtake_while p l = l ↔ ∀ x ∈ l, p x := +by simp [rtake_while, reverse_eq_iff] + +@[simp] lemma rtake_while_eq_nil_iff : rtake_while p l = [] ↔ ∀ (hl : l ≠ []), ¬ p (l.last hl) := +begin + induction l using list.reverse_rec_on; + simp [rtake_while] +end + +lemma mem_rtake_while_imp {x : α} (hx : x ∈ rtake_while p l) : p x := +begin + suffices : x ∈ take_while p l.reverse, + { exact mem_take_while_imp this }, + rwa [←mem_reverse, ←rtake_while] +end + +variables (p) (l) + +lemma rtake_while_idempotent : rtake_while p (rtake_while p l) = rtake_while p l := +rtake_while_eq_self_iff.mpr (λ _, mem_rtake_while_imp) + +end list From c82fab81e34d8ab0fbc90e657cfa9668e5c885cd Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Tue, 30 Aug 2022 08:38:06 +0000 Subject: [PATCH 084/828] feat(category_theory/adjunction): the Special Adjoint Functor Theorem (#15987) --- .../adjunction/adjoint_functor_theorems.lean | 39 +++++++++++++++++-- 1 file changed, 36 insertions(+), 3 deletions(-) diff --git a/src/category_theory/adjunction/adjoint_functor_theorems.lean b/src/category_theory/adjunction/adjoint_functor_theorems.lean index 370acdc8ad9b7..ac21e8514d3bb 100644 --- a/src/category_theory/adjunction/adjoint_functor_theorems.lean +++ b/src/category_theory/adjunction/adjoint_functor_theorems.lean @@ -5,11 +5,13 @@ Authors: Bhavik Mehta -/ import category_theory.adjunction.basic import category_theory.adjunction.comma +import category_theory.generator import category_theory.limits.constructions.weakly_initial import category_theory.limits.preserves.basic import category_theory.limits.creates import category_theory.limits.comma import category_theory.punit +import category_theory.subobject.comma /-! # Adjoint functor theorem @@ -26,8 +28,13 @@ We define the *solution set condition* for the functor `G : D ⥤ C` to mean, fo `A : C`, there is a set-indexed family ${f_i : A ⟶ G (B_i)}$ such that any morphism `A ⟶ G X` factors through one of the `f_i`. +This file also proves the special adjoint functor theorem, in the form: +* If `G : D ⥤ C` preserves limits and `D` is complete, well-powered and has a small coseparating + set, then `G` has a left adjoint: `is_right_adjoint_of_preserves_limits_of_is_coseparating` + + -/ -universes v u +universes v u u' namespace category_theory open limits @@ -48,9 +55,8 @@ def solution_set_condition {D : Type u} [category.{v} D] (G : D ⥤ C) : Prop := ∀ (A : C), ∃ (ι : Type v) (B : ι → D) (f : Π (i : ι), A ⟶ G.obj (B i)), ∀ X (h : A ⟶ G.obj X), ∃ (i : ι) (g : B i ⟶ X), f i ≫ G.map g = h -variables {D : Type u} [category.{v} D] - section general_adjoint_functor_theorem +variables {D : Type u} [category.{v} D] variables (G : D ⥤ C) @@ -88,4 +94,31 @@ end end general_adjoint_functor_theorem +section special_adjoint_functor_theorem +variables {D : Type u'} [category.{v} D] + +/-- +The special adjoint functor theorem: if `G : D ⥤ C` preserves limits and `D` is complete, +well-powered and has a small coseparating set, then `G` has a left adjoint. +-/ +noncomputable def is_right_adjoint_of_preserves_limits_of_is_coseparating [has_limits D] + [well_powered D] {𝒢 : set D} [small.{v} 𝒢] (h𝒢 : is_coseparating 𝒢) (G : D ⥤ C) + [preserves_limits G] : is_right_adjoint G := +have ∀ A, has_initial (structured_arrow A G), + from λ A, has_initial_of_is_coseparating (structured_arrow.is_coseparating_proj_preimage A G h𝒢), +by exactI is_right_adjoint_of_structured_arrow_initials _ + +/-- +The special adjoint functor theorem: if `F : C ⥤ D` preserves colimits and `C` is cocomplete, +well-copowered and has a small separating set, then `F` has a right adjoint. +-/ +noncomputable def is_left_adjoint_of_preserves_colimits_of_is_separatig [has_colimits C] + [well_powered Cᵒᵖ] {𝒢 : set C} [small.{v} 𝒢] (h𝒢 : is_separating 𝒢) (F : C ⥤ D) + [preserves_colimits F] : is_left_adjoint F := +have ∀ A, has_terminal (costructured_arrow F A), + from λ A, has_terminal_of_is_separating (costructured_arrow.is_separating_proj_preimage F A h𝒢), +by exactI is_left_adjoint_of_costructured_arrow_terminals _ + +end special_adjoint_functor_theorem + end category_theory From 86e936232dc0f78d82fbf2e918b36bd74107176b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 30 Aug 2022 09:26:20 +0000 Subject: [PATCH 085/828] feat(algebraic_topology): the category simplicial_object.split (#16274) This PR introduces the category of simplicial objects equipped with a splitting. --- .../split_simplicial_object.lean | 122 +++++++++++++++++- 1 file changed, 121 insertions(+), 1 deletion(-) diff --git a/src/algebraic_topology/split_simplicial_object.lean b/src/algebraic_topology/split_simplicial_object.lean index 892ae93101879..eef5946db01dd 100644 --- a/src/algebraic_topology/split_simplicial_object.lean +++ b/src/algebraic_topology/split_simplicial_object.lean @@ -14,7 +14,8 @@ import category_theory.limits.shapes.finite_products In this file, we introduce the notion of split simplicial object. If `C` is a category that has finite coproducts, a splitting `s : splitting X` of a simplical object `X` in `C` consists -of the datum of a sequence of objects `s.N : ℕ → C` and a +of the datum of a sequence of objects `s.N : ℕ → C` (which +we shall refer to as "nondegenerate simplices") and a sequence of morphisms `s.ι n : s.N n → X _[n]` that have the property that a certain canonical map identifies `X _[n]` with the coproduct of objects `s.N i` indexed by all possible @@ -23,6 +24,9 @@ assume that the morphisms `s.ι n` are monomorphisms: in the most common categories, this would be a consequence of the axioms.) +Simplicial objects equipped with a splitting form a category +`simplicial_object.split C`. + ## References * [Stacks: Splitting simplicial objects] https://stacks.math.columbia.edu/tag/017O @@ -223,6 +227,122 @@ begin erw [colimit.ι_desc, cofan.mk_ι_app], end +/-- A simplicial object that is isomorphic to a split simplicial object is split. -/ +@[simps] +def of_iso (e : X ≅ Y) : splitting Y := +{ N := s.N, + ι := λ n, s.ι n ≫ e.hom.app (op [n]), + map_is_iso' := λ Δ, begin + convert (infer_instance : is_iso ((s.iso Δ).hom ≫ e.hom.app Δ)), + tidy, + end, } + end splitting +variable (C) + +/-- The category `simplicial_object.split C` is the category of simplicial objects +in `C` equipped with a splitting, and morphisms are morphisms of simplicial objects +which are compatible with the splittings. -/ +@[ext, nolint has_nonempty_instance] +structure split := (X : simplicial_object C) (s : splitting X) + +namespace split + +variable {C} + +/-- The object in `simplicial_object.split C` attached to a splitting `s : splitting X` +of a simplicial object `X`. -/ +@[simps] +def mk' {X : simplicial_object C} (s : splitting X) : split C := ⟨X, s⟩ + +/-- Morphisms in `simplicial_object.split C` are morphisms of simplicial objects that +are compatible with the splittings. -/ +@[nolint has_nonempty_instance] +structure hom (S₁ S₂ : split C) := +(F : S₁.X ⟶ S₂.X) +(f : Π (n : ℕ), S₁.s.N n ⟶ S₂.s.N n) +(comm' : ∀ (n : ℕ), S₁.s.ι n ≫ F.app (op [n]) = f n ≫ S₂.s.ι n) + +@[ext] +lemma hom.ext {S₁ S₂ : split C} (Φ₁ Φ₂ : hom S₁ S₂) (h : ∀ (n : ℕ), Φ₁.f n = Φ₂.f n) : + Φ₁ = Φ₂ := +begin + rcases Φ₁ with ⟨F₁, f₁, c₁⟩, + rcases Φ₂ with ⟨F₂, f₂, c₂⟩, + have h' : f₁ = f₂ := by { ext, apply h, }, + subst h', + simp only [eq_self_iff_true, and_true], + apply S₁.s.hom_ext, + intro n, + dsimp, + rw [c₁, c₂], +end + +restate_axiom hom.comm' +attribute [simp, reassoc] hom.comm + +end split + +instance : category (split C) := +{ hom := split.hom, + id := λ S, { F := 𝟙 _, f := λ n, 𝟙 _, comm' := by tidy, }, + comp := λ S₁ S₂ S₃ Φ₁₂ Φ₂₃, + { F := Φ₁₂.F ≫ Φ₂₃.F, f := λ n, Φ₁₂.f n ≫ Φ₂₃.f n, comm' := by tidy, }, } + +variable {C} + +namespace split + +lemma congr_F {S₁ S₂ : split C} {Φ₁ Φ₂ : S₁ ⟶ S₂} (h : Φ₁ = Φ₂) : Φ₁.F = Φ₂.F := by rw h +lemma congr_f {S₁ S₂ : split C} {Φ₁ Φ₂ : S₁ ⟶ S₂} (h : Φ₁ = Φ₂) (n : ℕ) : + Φ₁.f n = Φ₂.f n := by rw h + +@[simp] +lemma id_F (S : split C) : (𝟙 S : S ⟶ S).F = 𝟙 (S.X) := rfl + +@[simp] +lemma id_f (S : split C) (n : ℕ) : (𝟙 S : S ⟶ S).f n = 𝟙 (S.s.N n) := rfl + +@[simp] +lemma comp_F {S₁ S₂ S₃ : split C} (Φ₁₂ : S₁ ⟶ S₂) (Φ₂₃ : S₂ ⟶ S₃) : + (Φ₁₂ ≫ Φ₂₃).F = Φ₁₂.F ≫ Φ₂₃.F := rfl + +@[simp] +lemma comp_f {S₁ S₂ S₃ : split C} (Φ₁₂ : S₁ ⟶ S₂) (Φ₂₃ : S₂ ⟶ S₃) (n : ℕ) : + (Φ₁₂ ≫ Φ₂₃).f n = Φ₁₂.f n ≫ Φ₂₃.f n := rfl + +@[simp, reassoc] +lemma ι_summand_naturality_symm {S₁ S₂ : split C} (Φ : S₁ ⟶ S₂) + {Δ : simplex_categoryᵒᵖ} (A : splitting.index_set Δ) : + S₁.s.ι_summand A ≫ Φ.F.app Δ = Φ.f A.1.unop.len ≫ S₂.s.ι_summand A := +by rw [S₁.s.ι_summand_eq, S₂.s.ι_summand_eq, assoc, Φ.F.naturality, ← Φ.comm_assoc] + +variable (C) + +/-- The functor `simplicial_object.split C ⥤ simplicial_object C` which forgets +the splitting. -/ +@[simps] +def forget : split C ⥤ simplicial_object C := +{ obj := λ S, S.X, + map := λ S₁ S₂ Φ, Φ.F, } + +/-- The functor `simplicial_object.split C ⥤ C` which sends a simplicial object equipped +with a splitting to its nondegenerate `n`-simplices. -/ +@[simps] +def eval_N (n : ℕ) : split C ⥤ C := +{ obj := λ S, S.s.N n, + map := λ S₁ S₂ Φ, Φ.f n, } + +/-- The inclusion of each summand in the coproduct decomposition of simplices +in split simplicial objects is a natural transformation of functors +`simplicial_object.split C ⥤ C` -/ +@[simps] +def nat_trans_ι_summand {Δ : simplex_categoryᵒᵖ} (A : splitting.index_set Δ) : + eval_N C A.1.unop.len ⟶ forget C ⋙ (evaluation simplex_categoryᵒᵖ C).obj Δ := +{ app := λ S, S.s.ι_summand A, + naturality' := λ S₁ S₂ Φ, (ι_summand_naturality_symm Φ A).symm, } + +end split + end simplicial_object From fbfe26dccf631eec689f33d47e609b3e5a27834a Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 30 Aug 2022 09:26:21 +0000 Subject: [PATCH 086/828] feat(linear_algebra/finite_dimensional): `dim_add_le_dim_add_dim` (#16301) Add a `finrank` version of a lemma that already exists for `module.rank`. The proof is exactly the same as the proof of the `module.rank` version, and, as with the previous `dim_sup_add_dim_inf_eq` lemma that it uses, so is the name (but in the `submodule` namespace in the case of the `finrank` versions). --- src/linear_algebra/finite_dimensional.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index f1d2b9c1c1e57..b9ba86648c0d7 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -861,6 +861,11 @@ begin exact key end +lemma dim_add_le_dim_add_dim (s t : submodule K V) + [finite_dimensional K s] [finite_dimensional K t] : + finrank K (s ⊔ t : submodule K V) ≤ finrank K s + finrank K t := +by { rw [← dim_sup_add_dim_inf_eq], exact self_le_add_right _ _ } + lemma eq_top_of_disjoint [finite_dimensional K V] (s t : submodule K V) (hdim : finrank K s + finrank K t = finrank K V) (hdisjoint : disjoint s t) : s ⊔ t = ⊤ := From 3d7987cda72abc473c7cdbbb075170e9ac620042 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Tue, 30 Aug 2022 11:08:03 +0000 Subject: [PATCH 087/828] chore(*): bump to lean 3.47.0 (#16252) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A major change is that notations now require names when they are shadowing another identical notation, even if it is a local notation. Also, because localized notations can be imported in a variety of contexts, there are some new best practices for them: * localized notations should always have a `(name := ...)`. Notation names are unrelated to declaration names, but the declaration name is a reasonable base for the notation name. * localized notations should never use `_` in the notation, because this gets desugared to a unique metavariable index, meaning that the notation will not be recognized as a duplicate of itself if `open_locale` is used when the notation is already available. Instead, you should use the `hole!` notation, which unfolds to `_`. Another major change is that projection notation (`x.foo`) now always instantiates implicit arguments with metavariables, which is consistent with Lean 4. To simulate the older behavior, one can use either strict implicit arguments for the structure field (e.g. `∀ {{n}}, p n` instead of `∀ {n}, p n`) or, depending on specifics, writing `λ _, x.foo` to ensure the implicit argument is preserved as an argument. Co-authored-by: Kyle Miller --- archive/imo/imo1998_q2.lean | 6 ++-- leanpkg.toml | 2 +- src/algebra/algebra/subalgebra/basic.lean | 4 +-- src/algebra/big_operators/basic.lean | 18 +++++----- src/algebra/big_operators/finprod.lean | 6 ++-- src/algebra/category/Module/basic.lean | 6 ++-- src/algebra/direct_sum/basic.lean | 3 +- src/algebra/direct_sum/ring.lean | 4 +-- src/algebra/dual_number.lean | 4 +-- src/algebra/hom/freiman.lean | 4 +-- src/algebra/hom/group_action.lean | 2 +- src/algebra/lie/quotient.lean | 2 +- src/algebra/lie/subalgebra.lean | 4 +-- src/algebra/lie/submodule.lean | 8 ++--- src/algebra/module/submodule/basic.lean | 2 +- src/algebra/monoid_algebra/basic.lean | 4 +-- src/algebra/order/hom/ring.lean | 4 +-- src/algebra/quandle.lean | 6 ++-- src/algebra/quaternion.lean | 5 +-- src/algebra/star/basic.lean | 2 +- .../dold_kan/notations.lean | 6 ++-- .../fundamental_groupoid/basic.lean | 12 +++---- src/algebraic_topology/simplex_category.lean | 2 +- src/algebraic_topology/simplicial_object.lean | 13 +++---- src/algebraic_topology/simplicial_set.lean | 7 ++-- .../asymptotics/asymptotic_equivalent.lean | 3 +- .../box_integral/partition/additive.lean | 6 ++-- src/analysis/complex/isometry.lean | 2 +- .../complex/upper_half_plane/basic.lean | 2 +- src/analysis/convex/segment.lean | 2 +- src/analysis/convolution.lean | 12 ++++--- src/analysis/inner_product_space/adjoint.lean | 2 +- src/analysis/inner_product_space/basic.lean | 6 ++-- .../normed_space/lattice_ordered_group.lean | 2 +- src/analysis/quaternion.lean | 2 +- .../trigonometric/basic.lean | 2 +- src/category_theory/bicategory/basic.lean | 15 +++++--- src/category_theory/bicategory/coherence.lean | 2 +- .../bicategory/coherence_tactic.lean | 8 ++--- src/category_theory/bicategory/free.lean | 20 +++++------ src/category_theory/bicategory/functor.lean | 18 +++++----- src/category_theory/closed/monoidal.lean | 2 +- src/category_theory/enriched/basic.lean | 2 +- src/category_theory/equivalence.lean | 5 +-- src/category_theory/monad/limits.lean | 4 +-- src/category_theory/monoidal/category.lean | 8 ++--- src/category_theory/monoidal/functor.lean | 3 +- src/category_theory/monoidal/rigid/basic.lean | 8 ++--- src/category_theory/types.lean | 2 +- src/combinatorics/configuration.lean | 4 +-- .../set_family/compression/down.lean | 2 +- .../set_family/compression/uv.lean | 2 +- src/combinatorics/set_family/shadow.lean | 4 +-- src/combinatorics/simple_graph/prod.lean | 12 ++++--- src/combinatorics/simple_graph/subgraph.lean | 4 +-- src/control/functor/multivariate.lean | 2 +- src/control/traversable/basic.lean | 5 ++- src/data/finmap.lean | 2 +- src/data/finset/fold.lean | 2 +- src/data/finset/prod.lean | 2 +- src/data/finset/slice.lean | 2 +- src/data/fintype/card_embedding.lean | 4 +-- src/data/list/basic.lean | 4 +-- src/data/list/defs.lean | 2 +- src/data/list/func.lean | 3 +- src/data/list/perm.lean | 6 ++-- src/data/matrix/basic.lean | 8 ++--- src/data/matrix/dmatrix.lean | 2 +- src/data/matrix/hadamard.lean | 2 +- src/data/matrix/kronecker.lean | 11 +++--- src/data/multiset/bind.lean | 2 +- src/data/multiset/fold.lean | 2 +- src/data/nat/basic.lean | 2 +- src/data/nat/factorial/basic.lean | 2 +- src/data/nat/totient.lean | 2 +- src/data/num/bitwise.lean | 6 ++-- src/data/polynomial/basic.lean | 2 +- src/data/quot.lean | 2 +- src/data/rat/defs.lean | 2 +- src/data/rbtree/insert.lean | 2 +- src/data/real/ennreal.lean | 4 +-- src/data/real/golden_ratio.lean | 4 +-- src/data/real/hyperreal.lean | 4 +-- src/data/real/nnreal.lean | 2 +- src/data/real/pi/leibniz.lean | 2 +- src/data/rel.lean | 2 +- src/data/seq/computation.lean | 4 +-- src/data/seq/seq.lean | 2 +- src/data/seq/wseq.lean | 2 +- .../set/intervals/unordered_interval.lean | 2 +- src/data/set/prod.lean | 2 +- src/data/stream/defs.lean | 2 +- src/data/sym/basic.lean | 2 +- src/data/typevec.lean | 12 +++---- src/data/vector3.lean | 5 +-- src/deprecated/ring.lean | 3 +- src/deprecated/subgroup.lean | 18 +++++----- src/deprecated/submonoid.lean | 4 +-- src/deprecated/subring.lean | 6 ++-- src/dynamics/omega_limit.lean | 6 ++-- src/field_theory/intermediate_field.lean | 4 +-- src/field_theory/krull_topology.lean | 8 ++--- src/geometry/euclidean/basic.lean | 2 +- src/geometry/manifold/algebra/monoid.lean | 4 +-- src/geometry/manifold/charted_space.lean | 8 +++-- src/geometry/manifold/cont_mdiff_map.lean | 4 +-- src/geometry/manifold/derivation_bundle.lean | 6 ++-- src/geometry/manifold/diffeomorph.lean | 10 +++--- src/geometry/manifold/instances/real.lean | 4 +-- .../smooth_manifold_with_corners.lean | 8 +++-- src/group_theory/coset.lean | 8 ++--- src/group_theory/subgroup/basic.lean | 12 +++---- src/group_theory/submonoid/basic.lean | 2 +- src/group_theory/submonoid/operations.lean | 8 ++--- src/group_theory/subsemigroup/basic.lean | 2 +- src/group_theory/subsemigroup/operations.lean | 8 ++--- src/group_theory/sylow.lean | 4 +-- src/linear_algebra/affine_space/basic.lean | 2 +- src/linear_algebra/cross_product.lean | 2 +- src/linear_algebra/pi_tensor_product.lean | 4 +-- src/linear_algebra/special_linear_group.lean | 3 +- src/linear_algebra/tensor_power.lean | 4 +-- src/linear_algebra/tensor_product.lean | 6 ++-- src/logic/equiv/local_equiv.lean | 8 ++--- src/logic/function/basic.lean | 2 +- .../conditional_expectation/basic.lean | 3 +- src/measure_theory/function/lp_space.lean | 6 ++-- .../function/strongly_measurable.lean | 4 +-- src/measure_theory/integral/bochner.lean | 9 ++--- src/measure_theory/measurable_space_def.lean | 6 ++-- src/measure_theory/measure/hausdorff.lean | 3 +- src/measure_theory/measure/measure_space.lean | 3 +- .../measure/mutually_singular.lean | 3 +- .../measure/vector_measure.lean | 12 ++++--- src/meta/coinductive_predicates.lean | 2 +- src/model_theory/basic.lean | 9 +++-- src/model_theory/elementary_maps.lean | 16 ++++----- src/model_theory/language_map.lean | 13 +++---- src/model_theory/satisfiability.lean | 3 +- src/model_theory/semantics.lean | 10 +++--- src/model_theory/skolem.lean | 2 +- src/model_theory/syntax.lean | 21 ++++++++---- src/number_theory/arithmetic_function.lean | 20 ++++++----- src/number_theory/cyclotomic/basic.lean | 6 ++-- src/number_theory/dioph.lean | 34 +++++++++---------- src/number_theory/modular.lean | 4 +-- .../modular_forms/slash_actions.lean | 3 +- src/number_theory/number_field.lean | 3 +- src/number_theory/prime_counting.lean | 4 +-- src/number_theory/von_mangoldt.lean | 3 +- src/order/filter/basic.lean | 2 +- src/order/filter/prod.lean | 2 +- src/order/hom/basic.lean | 2 +- src/order/initial_seg.lean | 6 ++-- src/order/rel_iso.lean | 10 +++--- src/probability/conditional_probability.lean | 6 ++-- src/probability/notation.lean | 20 ++++++----- src/probability/variance.lean | 5 ++- src/ring_theory/ideal/basic.lean | 6 ++-- src/ring_theory/ideal/operations.lean | 2 +- src/ring_theory/localization/num_denom.lean | 3 +- src/ring_theory/noetherian.lean | 3 +- src/ring_theory/non_zero_divisors.lean | 3 +- src/ring_theory/polynomial/eisenstein.lean | 2 +- src/ring_theory/principal_ideal_domain.lean | 2 +- src/ring_theory/subring/basic.lean | 2 +- src/ring_theory/valuation/basic.lean | 6 ++-- src/ring_theory/witt_vector/isocrystal.lean | 19 ++++++----- .../witt_vector/witt_polynomial.lean | 4 +-- src/set_theory/cardinal/basic.lean | 10 +++--- src/set_theory/cardinal/cofinality.lean | 6 ++-- src/set_theory/cardinal/continuum.lean | 2 +- src/set_theory/game/pgame.lean | 8 ++--- src/set_theory/ordinal/arithmetic.lean | 4 +-- src/set_theory/ordinal/basic.lean | 2 +- src/set_theory/ordinal/fixed_point.lean | 2 +- src/set_theory/ordinal/natural_ops.lean | 2 +- src/set_theory/ordinal/notation.lean | 6 ++-- src/set_theory/ordinal/principal.lean | 2 +- src/set_theory/surreal/dyadic.lean | 2 +- src/set_theory/zfc/basic.lean | 6 ++-- src/tactic/converter/apply_congr.lean | 2 +- src/tactic/core.lean | 4 +-- src/tactic/ext.lean | 4 +-- src/tactic/generalize_proofs.lean | 2 +- src/tactic/interval_cases.lean | 2 +- src/tactic/localized.lean | 26 +++++++++++--- src/tactic/monotonicity/interactive.lean | 4 +-- src/tactic/omega/int/form.lean | 10 +++--- src/tactic/omega/int/preterm.lean | 6 ++-- src/tactic/omega/misc.lean | 2 +- src/tactic/omega/nat/form.lean | 10 +++--- src/tactic/omega/nat/preterm.lean | 8 ++--- src/tactic/push_neg.lean | 4 +-- src/tactic/rcases.lean | 2 -- src/tactic/reassoc_axiom.lean | 2 -- src/tactic/ring.lean | 2 +- src/tactic/ring2.lean | 2 +- src/tactic/ring_exp.lean | 4 +-- src/tactic/simpa.lean | 2 +- src/tactic/tauto.lean | 2 +- src/topology/alexandroff.lean | 2 +- src/topology/basic.lean | 20 +++++++---- src/topology/continuous_function/bounded.lean | 3 +- .../continuous_function/zero_at_infty.lean | 7 ++-- src/topology/homotopy/equiv.lean | 3 +- src/topology/nhds_set.lean | 2 +- src/topology/uniform_space/basic.lean | 4 +-- src/topology/uniform_space/separation.lean | 2 +- src/topology/unit_interval.lean | 4 +-- test/induction.lean | 4 +-- test/localized/localized.lean | 6 ++-- test/noncomm_ring.lean | 4 +-- test/simps.lean | 20 +++++------ test/to_additive.lean | 2 +- 215 files changed, 618 insertions(+), 528 deletions(-) diff --git a/archive/imo/imo1998_q2.lean b/archive/imo/imo1998_q2.lean index 51a73ca85113b..3c897525b8abe 100644 --- a/archive/imo/imo1998_q2.lean +++ b/archive/imo/imo1998_q2.lean @@ -190,16 +190,14 @@ end end -local notation x `/` y := (x : ℚ) / y - lemma clear_denominators {a b k : ℕ} (ha : 0 < a) (hb : 0 < b) : - (b - 1) / (2 * b) ≤ k / a ↔ (b - 1) * a ≤ k * (2 * b) := + (b - 1 : ℚ) / (2 * b) ≤ k / a ↔ (b - 1) * a ≤ k * (2 * b) := by rw div_le_div_iff; norm_cast; simp [ha, hb] theorem imo1998_q2 [fintype J] [fintype C] (a b k : ℕ) (hC : fintype.card C = a) (hJ : fintype.card J = b) (ha : 0 < a) (hb : odd b) (hk : ∀ (p : judge_pair J), p.distinct → (agreed_contestants r p).card ≤ k) : - (b - 1) / (2 * b) ≤ k / a := + (b - 1 : ℚ) / (2 * b) ≤ k / a := begin rw clear_denominators ha hb.pos, obtain ⟨z, hz⟩ := hb, rw hz at hJ, rw hz, diff --git a/leanpkg.toml b/leanpkg.toml index 44c125e476141..009b32dd0d769 100644 --- a/leanpkg.toml +++ b/leanpkg.toml @@ -1,7 +1,7 @@ [package] name = "mathlib" version = "0.1" -lean_version = "leanprover-community/lean:3.46.0" +lean_version = "leanprover-community/lean:3.47.0" path = "src" [dependencies] diff --git a/src/algebra/algebra/subalgebra/basic.lean b/src/algebra/algebra/subalgebra/basic.lean index b9ebea1b8d3ff..1d0f93eba1835 100644 --- a/src/algebra/algebra/subalgebra/basic.lean +++ b/src/algebra/algebra/subalgebra/basic.lean @@ -68,8 +68,8 @@ to_subsemiring_injective.eq_iff equalities. -/ protected def copy (S : subalgebra R A) (s : set A) (hs : s = ↑S) : subalgebra R A := { carrier := s, - add_mem' := hs.symm ▸ S.add_mem', - mul_mem' := hs.symm ▸ S.mul_mem', + add_mem' := λ _ _, hs.symm ▸ S.add_mem', + mul_mem' := λ _ _, hs.symm ▸ S.mul_mem', algebra_map_mem' := hs.symm ▸ S.algebra_map_mem' } @[simp] lemma coe_copy (S : subalgebra R A) (s : set A) (hs : s = ↑S) : diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index ffc7e8ceaa609..a77d695ab33be 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -82,15 +82,15 @@ In practice, this means that parentheses should be placed as follows: -/ library_note "operator precedence of big operators" -localized "notation `∑` binders `, ` r:(scoped:67 f, finset.sum finset.univ f) := r" - in big_operators -localized "notation `∏` binders `, ` r:(scoped:67 f, finset.prod finset.univ f) := r" - in big_operators - -localized "notation `∑` binders ` in ` s `, ` r:(scoped:67 f, finset.sum s f) := r" - in big_operators -localized "notation `∏` binders ` in ` s `, ` r:(scoped:67 f, finset.prod s f) := r" - in big_operators +localized "notation (name := finset.sum_univ) + `∑` binders `, ` r:(scoped:67 f, finset.sum finset.univ f) := r" in big_operators +localized "notation (name := finset.prod_univ) + `∏` binders `, ` r:(scoped:67 f, finset.prod finset.univ f) := r" in big_operators + +localized "notation (name := finset.sum) + `∑` binders ` in ` s `, ` r:(scoped:67 f, finset.sum s f) := r" in big_operators +localized "notation (name := finset.prod) + `∏` binders ` in ` s `, ` r:(scoped:67 f, finset.prod s f) := r" in big_operators open_locale big_operators diff --git a/src/algebra/big_operators/finprod.lean b/src/algebra/big_operators/finprod.lean index b38c4c8b552f3..8a9a4ac7ecc96 100644 --- a/src/algebra/big_operators/finprod.lean +++ b/src/algebra/big_operators/finprod.lean @@ -97,9 +97,11 @@ if h : (mul_support (f ∘ plift.down)).finite then ∏ i in h.to_finset, f i.do end -localized "notation `∑ᶠ` binders `, ` r:(scoped:67 f, finsum f) := r" in big_operators +localized "notation (name := finsum) + `∑ᶠ` binders `, ` r:(scoped:67 f, finsum f) := r" in big_operators -localized "notation `∏ᶠ` binders `, ` r:(scoped:67 f, finprod f) := r" in big_operators +localized "notation (name := finprod) + `∏ᶠ` binders `, ` r:(scoped:67 f, finprod f) := r" in big_operators @[to_additive] lemma finprod_eq_prod_plift_of_mul_support_to_finset_subset {f : α → M} (hf : (mul_support (f ∘ plift.down)).finite) {s : finset (plift α)} diff --git a/src/algebra/category/Module/basic.lean b/src/algebra/category/Module/basic.lean index c4a47c36b69f0..f331e9d117c94 100644 --- a/src/algebra/category/Module/basic.lean +++ b/src/algebra/category/Module/basic.lean @@ -163,19 +163,19 @@ variables {X₁ X₂ : Type v} def Module.as_hom [add_comm_group X₁] [module R X₁] [add_comm_group X₂] [module R X₂] : (X₁ →ₗ[R] X₂) → (Module.of R X₁ ⟶ Module.of R X₂) := id -localized "notation `↟` f : 1024 := Module.as_hom f" in Module +localized "notation (name := Module.as_hom) `↟` f : 1024 := Module.as_hom f" in Module /-- Reinterpreting a linear map in the category of `R`-modules. -/ def Module.as_hom_right [add_comm_group X₁] [module R X₁] {X₂ : Module.{v} R} : (X₁ →ₗ[R] X₂) → (Module.of R X₁ ⟶ X₂) := id -localized "notation `↾` f : 1024 := Module.as_hom_right f" in Module +localized "notation (name := Module.as_hom_right) `↾` f : 1024 := Module.as_hom_right f" in Module /-- Reinterpreting a linear map in the category of `R`-modules. -/ def Module.as_hom_left {X₁ : Module.{v} R} [add_comm_group X₂] [module R X₂] : (X₁ →ₗ[R] X₂) → (X₁ ⟶ Module.of R X₂) := id -localized "notation `↿` f : 1024 := Module.as_hom_left f" in Module +localized "notation (name := Module.as_hom_left) `↿` f : 1024 := Module.as_hom_left f" in Module /-- Build an isomorphism in the category `Module R` from a `linear_equiv` between `module`s. -/ @[simps] diff --git a/src/algebra/direct_sum/basic.lean b/src/algebra/direct_sum/basic.lean index 44ecf3f72a33b..7169e3051b440 100644 --- a/src/algebra/direct_sum/basic.lean +++ b/src/algebra/direct_sum/basic.lean @@ -36,7 +36,8 @@ def direct_sum [Π i, add_comm_monoid (β i)] : Type* := Π₀ i, β i instance [Π i, add_comm_monoid (β i)] : has_coe_to_fun (direct_sum ι β) (λ _, Π i : ι, β i) := dfinsupp.has_coe_to_fun -localized "notation `⨁` binders `, ` r:(scoped f, direct_sum _ f) := r" in direct_sum +localized "notation (name := direct_sum) + `⨁` binders `, ` r:(scoped f, direct_sum _ f) := r" in direct_sum namespace direct_sum diff --git a/src/algebra/direct_sum/ring.lean b/src/algebra/direct_sum/ring.lean index 99e73c97e609f..b8c27d87566bc 100644 --- a/src/algebra/direct_sum/ring.lean +++ b/src/algebra/direct_sum/ring.lean @@ -566,7 +566,7 @@ def lift_ring_hom : f (graded_monoid.ghas_one.one) = 1 ∧ ∀ {i j} (ai : A i) (aj : A j), f (graded_monoid.ghas_mul.mul ai aj) = f ai * f aj} ≃ ((⨁ i, A i) →+* R) := -{ to_fun := λ f, to_semiring f.1 f.2.1 f.2.2, +{ to_fun := λ f, to_semiring (λ _, f.1) f.2.1 (λ _ _, f.2.2), inv_fun := λ F, ⟨λ i, (F : (⨁ i, A i) →+ R).comp (of _ i), begin simp only [add_monoid_hom.comp_apply, ring_hom.coe_add_monoid_hom], @@ -578,7 +578,7 @@ def lift_ring_hom : end⟩, left_inv := λ f, begin ext xi xv, - exact to_add_monoid_of f.1 xi xv, + exact to_add_monoid_of (λ _, f.1) xi xv, end, right_inv := λ F, begin apply ring_hom.coe_add_monoid_hom_injective, diff --git a/src/algebra/dual_number.lean b/src/algebra/dual_number.lean index 9e1fee266c46f..9aa9098f64a4c 100644 --- a/src/algebra/dual_number.lean +++ b/src/algebra/dual_number.lean @@ -43,8 +43,8 @@ abbreviation dual_number (R : Type*) : Type* := triv_sq_zero_ext R R /-- The unit element $ε$ that squares to zero. -/ def dual_number.eps [has_zero R] [has_one R] : dual_number R := triv_sq_zero_ext.inr 1 -localized "notation `ε` := dual_number.eps" in dual_number -localized "postfix `[ε]`:1025 := dual_number" in dual_number +localized "notation (name := dual_number.eps) `ε` := dual_number.eps" in dual_number +localized "postfix (name := dual_number) `[ε]`:1025 := dual_number" in dual_number open_locale dual_number diff --git a/src/algebra/hom/freiman.lean b/src/algebra/hom/freiman.lean index bc7ffa6628599..367d5689379f4 100644 --- a/src/algebra/hom/freiman.lean +++ b/src/algebra/hom/freiman.lean @@ -66,8 +66,8 @@ structure freiman_hom (A : set α) (β : Type*) [comm_monoid α] [comm_monoid β (hs : s.card = n) (ht : t.card = n) (h : s.prod = t.prod) : (s.map to_fun).prod = (t.map to_fun).prod) -notation A ` →+[`:25 n:25 `] `:0 β:0 := add_freiman_hom A β n -notation A ` →*[`:25 n:25 `] `:0 β:0 := freiman_hom A β n +notation (name := add_freiman_hom) A ` →+[`:25 n:25 `] `:0 β:0 := add_freiman_hom A β n +notation (name := freiman_hom) A ` →*[`:25 n:25 `] `:0 β:0 := freiman_hom A β n /-- `add_freiman_hom_class F s β n` states that `F` is a type of `n`-ary sums-preserving morphisms. You should extend this class when you extend `add_freiman_hom`. -/ diff --git a/src/algebra/hom/group_action.lean b/src/algebra/hom/group_action.lean index cd9109774eeb2..f0864ee100e08 100644 --- a/src/algebra/hom/group_action.lean +++ b/src/algebra/hom/group_action.lean @@ -59,7 +59,7 @@ structure mul_action_hom := (to_fun : X → Y) (map_smul' : ∀ (m : M') (x : X), to_fun (m • x) = m • to_fun x) -notation X ` →[`:25 M:25 `] `:0 Y:0 := mul_action_hom M X Y +notation (name := mul_action_hom) X ` →[`:25 M:25 `] `:0 Y:0 := mul_action_hom M X Y /-- `smul_hom_class F M X Y` states that `F` is a type of morphisms preserving scalar multiplication by `M`. diff --git a/src/algebra/lie/quotient.lean b/src/algebra/lie/quotient.lean index 95bce9a6887ca..ad05865053818 100644 --- a/src/algebra/lie/quotient.lean +++ b/src/algebra/lie/quotient.lean @@ -64,7 +64,7 @@ lemma is_quotient_mk (m : M) : quotient.mk' m = (mk m : M ⧸ N) := rfl /-- Given a Lie module `M` over a Lie algebra `L`, together with a Lie submodule `N ⊆ M`, there is a natural linear map from `L` to the endomorphisms of `M` leaving `N` invariant. -/ def lie_submodule_invariant : L →ₗ[R] submodule.compatible_maps N.to_submodule N.to_submodule := -linear_map.cod_restrict _ (lie_module.to_endomorphism R L M) N.lie_mem +linear_map.cod_restrict _ (lie_module.to_endomorphism R L M) $ λ _ _, N.lie_mem variables (N) diff --git a/src/algebra/lie/subalgebra.lean b/src/algebra/lie/subalgebra.lean index dc2c5a405795d..b994a6914b853 100644 --- a/src/algebra/lie/subalgebra.lean +++ b/src/algebra/lie/subalgebra.lean @@ -57,7 +57,7 @@ instance : set_like (lie_subalgebra R L) L := coe_injective' := λ L' L'' h, by { rcases L' with ⟨⟨⟩⟩, rcases L'' with ⟨⟨⟩⟩, congr' } } instance : add_subgroup_class (lie_subalgebra R L) L := -{ add_mem := λ L', L'.add_mem', +{ add_mem := λ L' _ _, L'.add_mem', zero_mem := λ L', L'.zero_mem', neg_mem := λ L' x hx, show -x ∈ (L' : submodule R L), from neg_mem hx } @@ -258,7 +258,7 @@ lemma submodule.exists_lie_subalgebra_coe_eq_iff (p : submodule R L) : (∃ (K : lie_subalgebra R L), ↑K = p) ↔ ∀ (x y : L), x ∈ p → y ∈ p → ⁅x, y⁆ ∈ p := begin split, - { rintros ⟨K, rfl⟩, exact K.lie_mem', }, + { rintros ⟨K, rfl⟩ _ _, exact K.lie_mem', }, { intros h, use { lie_mem' := h, ..p }, exact lie_subalgebra.coe_to_submodule_mk p _, }, end diff --git a/src/algebra/lie/submodule.lean b/src/algebra/lie/submodule.lean index 41df037476c8d..003c520455125 100644 --- a/src/algebra/lie/submodule.lean +++ b/src/algebra/lie/submodule.lean @@ -54,7 +54,7 @@ instance : set_like (lie_submodule R L M) M := coe_injective' := λ N O h, by cases N; cases O; congr' } instance : add_subgroup_class (lie_submodule R L M) M := -{ add_mem := λ N, N.add_mem', +{ add_mem := λ N _ _, N.add_mem', zero_mem := λ N, N.zero_mem', neg_mem := λ N x hx, show -x ∈ N.to_submodule, from neg_mem hx } @@ -108,9 +108,9 @@ equalities. -/ protected def copy (s : set M) (hs : s = ↑N) : lie_submodule R L M := { carrier := s, zero_mem' := hs.symm ▸ N.zero_mem', - add_mem' := hs.symm ▸ N.add_mem', + add_mem' := λ _ _, hs.symm ▸ N.add_mem', smul_mem' := hs.symm ▸ N.smul_mem', - lie_mem := hs.symm ▸ N.lie_mem, } + lie_mem := λ _ _, hs.symm ▸ N.lie_mem, } @[simp] lemma coe_copy (S : lie_submodule R L M) (s : set M) (hs : s = ↑S) : (S.copy s hs : set M) = s := rfl @@ -208,7 +208,7 @@ lemma submodule.exists_lie_submodule_coe_eq_iff (p : submodule R M) : (∃ (N : lie_submodule R L M), ↑N = p) ↔ ∀ (x : L) (m : M), m ∈ p → ⁅x, m⁆ ∈ p := begin split, - { rintros ⟨N, rfl⟩, exact N.lie_mem, }, + { rintros ⟨N, rfl⟩ _ _, exact N.lie_mem, }, { intros h, use { lie_mem := h, ..p }, exact lie_submodule.coe_to_submodule_mk p _, }, end diff --git a/src/algebra/module/submodule/basic.lean b/src/algebra/module/submodule/basic.lean index 106569d968336..7c376204bdf8f 100644 --- a/src/algebra/module/submodule/basic.lean +++ b/src/algebra/module/submodule/basic.lean @@ -77,7 +77,7 @@ equalities. -/ protected def copy (p : submodule R M) (s : set M) (hs : s = ↑p) : submodule R M := { carrier := s, zero_mem' := hs.symm ▸ p.zero_mem', - add_mem' := hs.symm ▸ p.add_mem', + add_mem' := λ _ _, hs.symm ▸ p.add_mem', smul_mem' := hs.symm ▸ p.smul_mem' } @[simp] lemma coe_copy (S : submodule R M) (s : set M) (hs : s = ↑S) : diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index be59f29d52cd2..e472769810eff 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -893,13 +893,13 @@ def submodule_of_smul_mem (W : submodule k V) (h : ∀ (g : G) (v : V), v ∈ W submodule (monoid_algebra k G) V := { carrier := W, zero_mem' := W.zero_mem', - add_mem' := W.add_mem', + add_mem' := λ _ _, W.add_mem', smul_mem' := begin intros f v hv, rw [←finsupp.sum_single f, finsupp.sum, finset.sum_smul], simp_rw [←smul_of, smul_assoc], exact submodule.sum_smul_mem W _ (λ g _, h g v hv) - end } + end } end submodule diff --git a/src/algebra/order/hom/ring.lean b/src/algebra/order/hom/ring.lean index f5cdb1c1f7c13..2f86a97b9bb63 100644 --- a/src/algebra/order/hom/ring.lean +++ b/src/algebra/order/hom/ring.lean @@ -215,7 +215,7 @@ variables [has_mul α] [has_add α] [has_le α] [has_mul β] [has_add β] [has_l [has_add γ] [has_le γ] [has_mul δ] [has_add δ] [has_le δ] /-- Reinterpret an ordered ring isomorphism as an order isomorphism. -/ -def to_order_iso (f : α ≃+*o β) : α ≃o β := ⟨f.to_ring_equiv.to_equiv, f.map_le_map_iff'⟩ +def to_order_iso (f : α ≃+*o β) : α ≃o β := ⟨f.to_ring_equiv.to_equiv, λ _ _, f.map_le_map_iff'⟩ instance : order_ring_iso_class (α ≃+*o β) α β := { coe := λ f, f.to_fun, @@ -223,7 +223,7 @@ instance : order_ring_iso_class (α ≃+*o β) α β := coe_injective' := λ f g h₁ h₂, by { obtain ⟨⟨_, _⟩, _⟩ := f, obtain ⟨⟨_, _⟩, _⟩ := g, congr' }, map_add := λ f, f.map_add', map_mul := λ f, f.map_mul', - map_le_map_iff := λ f, f.map_le_map_iff', + map_le_map_iff := λ f _ _, f.map_le_map_iff', left_inv := λ f, f.left_inv, right_inv := λ f, f.right_inv } diff --git a/src/algebra/quandle.lean b/src/algebra/quandle.lean index 75534effef233..92aaa5c0fb566 100644 --- a/src/algebra/quandle.lean +++ b/src/algebra/quandle.lean @@ -110,9 +110,9 @@ class rack (α : Type u) extends shelf α := (left_inv : ∀ x, function.left_inverse (inv_act x) (act x)) (right_inv : ∀ x, function.right_inverse (inv_act x) (act x)) -localized "infixr ` ◃ `:65 := shelf.act" in quandles -localized "infixr ` ◃⁻¹ `:65 := rack.inv_act" in quandles -localized "infixr ` →◃ `:25 := shelf_hom" in quandles +localized "infixr (name := shelf.act) ` ◃ `:65 := shelf.act" in quandles +localized "infixr (name := rack.inv_act) ` ◃⁻¹ `:65 := rack.inv_act" in quandles +localized "infixr (name := shelf_hom) ` →◃ `:25 := shelf_hom" in quandles open_locale quandles diff --git a/src/algebra/quaternion.lean b/src/algebra/quaternion.lean index ccead14a9e35b..62f97a483b852 100644 --- a/src/algebra/quaternion.lean +++ b/src/algebra/quaternion.lean @@ -52,7 +52,8 @@ Implemented as a structure with four fields: `re`, `im_i`, `im_j`, and `im_k`. - structure quaternion_algebra (R : Type*) (a b : R) := mk {} :: (re : R) (im_i : R) (im_j : R) (im_k : R) -localized "notation `ℍ[` R`,` a`,` b `]` := quaternion_algebra R a b" in quaternion +localized "notation (name := quaternion_algebra) `ℍ[` R`,` a`,` b `]` := + quaternion_algebra R a b" in quaternion namespace quaternion_algebra @@ -341,7 +342,7 @@ end quaternion_algebra `re`, `im_i`, `im_j`, and `im_k`. -/ def quaternion (R : Type*) [has_one R] [has_neg R] := quaternion_algebra R (-1) (-1) -localized "notation `ℍ[` R `]` := quaternion R" in quaternion +localized "notation (name := quaternion) `ℍ[` R `]` := quaternion R" in quaternion /-- The equivalence between the quaternions over R and R × R × R × R. -/ def quaternion.equiv_prod (R : Type*) [has_one R] [has_neg R] : ℍ[R] ≃ R × R × R × R := diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index c90cc3dd7990f..07ab3af09d291 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -284,7 +284,7 @@ case for `(↑star_ring_aut : R →* R)`. -/ def star_ring_end [comm_semiring R] [star_ring R] : R →+* R := @star_ring_aut R _ _ variables {R} -localized "notation `conj` := star_ring_end _" in complex_conjugate +localized "notation (name := star_ring_end) `conj` := star_ring_end hole!" in complex_conjugate /-- This is not a simp lemma, since we usually want simp to keep `star_ring_end` bundled. For example, for complex conjugation, we don't want simp to turn `conj x` diff --git a/src/algebraic_topology/dold_kan/notations.lean b/src/algebraic_topology/dold_kan/notations.lean index e7194e9cf2bd8..0bfe1a14e6162 100644 --- a/src/algebraic_topology/dold_kan/notations.lean +++ b/src/algebraic_topology/dold_kan/notations.lean @@ -16,5 +16,7 @@ as `N[X]` for the normalized subcomplex in the case `C` is an abelian category. -/ -localized "notation `K[`X`]` := algebraic_topology.alternating_face_map_complex.obj X" in dold_kan -localized "notation `N[`X`]` := algebraic_topology.normalized_Moore_complex.obj X" in dold_kan +localized "notation (name := alternating_face_map_complex) `K[`X`]` := + algebraic_topology.alternating_face_map_complex.obj X" in dold_kan +localized "notation (name := normalized_Moore_complex) `N[`X`]` := + algebraic_topology.normalized_Moore_complex.obj X" in dold_kan diff --git a/src/algebraic_topology/fundamental_groupoid/basic.lean b/src/algebraic_topology/fundamental_groupoid/basic.lean index 0fd3afef68d48..9c3fea201bd58 100644 --- a/src/algebraic_topology/fundamental_groupoid/basic.lean +++ b/src/algebraic_topology/fundamental_groupoid/basic.lean @@ -323,12 +323,12 @@ def fundamental_groupoid_functor : Top ⥤ category_theory.Groupoid := refl, end } -localized "notation `π` := fundamental_groupoid.fundamental_groupoid_functor" - in fundamental_groupoid -localized "notation `πₓ` := fundamental_groupoid.fundamental_groupoid_functor.obj" - in fundamental_groupoid -localized "notation `πₘ` := fundamental_groupoid.fundamental_groupoid_functor.map" - in fundamental_groupoid +localized "notation (name := fundamental_groupoid_functor) + `π` := fundamental_groupoid.fundamental_groupoid_functor" in fundamental_groupoid +localized "notation (name := fundamental_groupoid_functor.obj) + `πₓ` := fundamental_groupoid.fundamental_groupoid_functor.obj" in fundamental_groupoid +localized "notation (name := fundamental_groupoid_functor.map) + `πₘ` := fundamental_groupoid.fundamental_groupoid_functor.map" in fundamental_groupoid lemma map_eq {X Y : Top} {x₀ x₁ : X} (f : C(X, Y)) (p : path.homotopic.quotient x₀ x₁) : (πₘ f).map p = p.map_fn f := rfl diff --git a/src/algebraic_topology/simplex_category.lean b/src/algebraic_topology/simplex_category.lean index fd819e375bcd5..62fddc97ee704 100644 --- a/src/algebraic_topology/simplex_category.lean +++ b/src/algebraic_topology/simplex_category.lean @@ -51,7 +51,7 @@ local attribute [semireducible] simplex_category /-- Interpet a natural number as an object of the simplex category. -/ def mk (n : ℕ) : simplex_category := n -localized "notation `[`n`]` := simplex_category.mk n" in simplicial +localized "notation (name := simplex_category.mk) `[`n`]` := simplex_category.mk n" in simplicial -- TODO: Make `len` irreducible. /-- The length of an object of `simplex_category`. -/ diff --git a/src/algebraic_topology/simplicial_object.lean b/src/algebraic_topology/simplicial_object.lean index be384bc4a9b86..3bbdcf4a7e6ad 100644 --- a/src/algebraic_topology/simplicial_object.lean +++ b/src/algebraic_topology/simplicial_object.lean @@ -36,9 +36,8 @@ def simplicial_object := simplex_categoryᵒᵖ ⥤ C namespace simplicial_object -localized - "notation X `_[`:1000 n `]` := - (X : category_theory.simplicial_object _).obj (opposite.op (simplex_category.mk n))" +localized "notation (name := simplicial_object.at) X `_[`:1000 n `]` := + (X : category_theory.simplicial_object hole!).obj (opposite.op (simplex_category.mk n))" in simplicial instance {J : Type v} [small_category J] [has_limits_of_shape J C] : @@ -237,7 +236,7 @@ end augmented open_locale simplicial -/-- Aaugment a simplicial object with an object. -/ +/-- Augment a simplicial object with an object. -/ @[simps] def augment (X : simplicial_object C) (X₀ : C) (f : X _[0] ⟶ X₀) (w : ∀ (i : simplex_category) (g₁ g₂ : [0] ⟶ i), @@ -266,10 +265,8 @@ def cosimplicial_object := simplex_category ⥤ C namespace cosimplicial_object -localized - "notation X `_[`:1000 n `]` := - (X : category_theory.cosimplicial_object _).obj (simplex_category.mk n)" - in simplicial +localized "notation (name := cosimplicial_object.at) X `_[`:1000 n `]` := + (X : category_theory.cosimplicial_object hole!).obj (simplex_category.mk n)" in simplicial instance {J : Type v} [small_category J] [has_limits_of_shape J C] : has_limits_of_shape J (cosimplicial_object C) := by {dsimp [cosimplicial_object], apply_instance} diff --git a/src/algebraic_topology/simplicial_set.lean b/src/algebraic_topology/simplicial_set.lean index 5d5aca42921e6..36c6280682286 100644 --- a/src/algebraic_topology/simplicial_set.lean +++ b/src/algebraic_topology/simplicial_set.lean @@ -48,7 +48,8 @@ namespace sSet is the Yoneda embedding of `n`. -/ def standard_simplex : simplex_category ⥤ sSet := yoneda -localized "notation `Δ[`n`]` := sSet.standard_simplex.obj (simplex_category.mk n)" in simplicial +localized "notation (name := standard_simplex) `Δ[`n`]` := + sSet.standard_simplex.obj (simplex_category.mk n)" in simplicial instance : inhabited sSet := ⟨Δ[0]⟩ @@ -68,7 +69,7 @@ def boundary (n : ℕ) : sSet := map := λ m₁ m₂ f α, ⟨f.unop ≫ (α : Δ[n].obj m₁), by { intro h, apply α.property, exact function.surjective.of_comp h }⟩ } -localized "notation `∂Δ[`n`]` := sSet.boundary n" in simplicial +localized "notation (name := sSet.boundary) `∂Δ[`n`]` := sSet.boundary n" in simplicial /-- The inclusion of the boundary of the `n`-th standard simplex into that standard simplex. -/ def boundary_inclusion (n : ℕ) : @@ -91,7 +92,7 @@ def horn (n : ℕ) (i : fin (n+1)) : sSet := exact set.range_comp_subset_range _ _ hj, end⟩ } -localized "notation `Λ[`n`, `i`]` := sSet.horn (n : ℕ) i" in simplicial +localized "notation (name := sSet.horn) `Λ[`n`, `i`]` := sSet.horn (n : ℕ) i" in simplicial /-- The inclusion of the `i`-th horn of the `n`-th standard simplex into that standard simplex. -/ def horn_inclusion (n : ℕ) (i : fin (n+1)) : diff --git a/src/analysis/asymptotics/asymptotic_equivalent.lean b/src/analysis/asymptotics/asymptotic_equivalent.lean index b7d397a74d148..d3574393f76be 100644 --- a/src/analysis/asymptotics/asymptotic_equivalent.lean +++ b/src/analysis/asymptotics/asymptotic_equivalent.lean @@ -67,7 +67,8 @@ variables {α β : Type*} [normed_add_comm_group β] `u x - v x = o(v x)` as x converges along `l`. -/ def is_equivalent (l : filter α) (u v : α → β) := (u - v) =o[l] v -localized "notation u ` ~[`:50 l:50 `] `:0 v:50 := asymptotics.is_equivalent l u v" in asymptotics +localized "notation (name := asymptotics.is_equivalent) + u ` ~[`:50 l:50 `] `:0 v:50 := asymptotics.is_equivalent l u v" in asymptotics variables {u v w : α → β} {l : filter α} diff --git a/src/analysis/box_integral/partition/additive.lean b/src/analysis/box_integral/partition/additive.lean index e16b9a790d099..bb787d52dc58c 100644 --- a/src/analysis/box_integral/partition/additive.lean +++ b/src/analysis/box_integral/partition/additive.lean @@ -45,8 +45,10 @@ structure box_additive_map (ι M : Type*) [add_comm_monoid M] (I : with_top (box (sum_partition_boxes' : ∀ J : box ι, ↑J ≤ I → ∀ π : prepartition J, π.is_partition → ∑ Ji in π.boxes, to_fun Ji = to_fun J) -localized "notation ι ` →ᵇᵃ `:25 M := box_integral.box_additive_map ι M ⊤" in box_integral -localized "notation ι ` →ᵇᵃ[`:25 I `] ` M := box_integral.box_additive_map ι M I" in box_integral +localized "notation (name := box_integral.box_additive_map.top) + ι ` →ᵇᵃ `:25 M := box_integral.box_additive_map ι M ⊤" in box_integral +localized "notation (name := box_integral.box_additive_map) + ι ` →ᵇᵃ[`:25 I `] ` M := box_integral.box_additive_map ι M I" in box_integral namespace box_additive_map diff --git a/src/analysis/complex/isometry.lean b/src/analysis/complex/isometry.lean index db8eb79a5c681..88de13fee8aa5 100644 --- a/src/analysis/complex/isometry.lean +++ b/src/analysis/complex/isometry.lean @@ -29,7 +29,7 @@ noncomputable theory open complex open_locale complex_conjugate -local notation `|` x `|` := complex.abs x +local notation (name := complex.abs) `|` x `|` := complex.abs x /-- An element of the unit circle defines a `linear_isometry_equiv` from `ℂ` to itself, by rotation. -/ diff --git a/src/analysis/complex/upper_half_plane/basic.lean b/src/analysis/complex/upper_half_plane/basic.lean index 44f6c4323fc41..cac94c99a3f31 100644 --- a/src/analysis/complex/upper_half_plane/basic.lean +++ b/src/analysis/complex/upper_half_plane/basic.lean @@ -42,7 +42,7 @@ local notation `GL(` n `, ` R `)`⁺ := matrix.GL_pos (fin n) R @[derive [λ α, has_coe α ℂ]] def upper_half_plane := {point : ℂ // 0 < point.im} -localized "notation `ℍ` := upper_half_plane" in upper_half_plane +localized "notation (name := upper_half_plane) `ℍ` := upper_half_plane" in upper_half_plane namespace upper_half_plane diff --git a/src/analysis/convex/segment.lean b/src/analysis/convex/segment.lean index 56bd3606980c2..15a02d5fbdf9d 100644 --- a/src/analysis/convex/segment.lean +++ b/src/analysis/convex/segment.lean @@ -48,7 +48,7 @@ the base semiring has some element between `0` and `1`. -/ def open_segment (x y : E) : set E := {z : E | ∃ (a b : 𝕜) (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1), a • x + b • y = z} -localized "notation `[` x ` -[` 𝕜 `] ` y `]` := segment 𝕜 x y" in convex +localized "notation (name := segment) `[` x ` -[` 𝕜 `] ` y `]` := segment 𝕜 x y" in convex lemma segment_eq_image₂ (x y : E) : [x -[𝕜] y] = (λ p : 𝕜 × 𝕜, p.1 • x + p.2 • y) '' {p | 0 ≤ p.1 ∧ 0 ≤ p.2 ∧ p.1 + p.2 = 1} := diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 67f117ad07fca..626e60bb616c5 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -356,11 +356,13 @@ noncomputable def convolution [has_sub G] (f : G → E) (g : G → E') (L : E (μ : measure G . volume_tac) : G → F := λ x, ∫ t, L (f t) (g (x - t)) ∂μ -localized "notation f ` ⋆[`:67 L:67 `, ` μ:67 `] `:0 g:66 := convolution f g L μ" in convolution -localized "notation f ` ⋆[`:67 L:67 `]`:0 g:66 := convolution f g L - measure_theory.measure_space.volume" in convolution -localized "notation f ` ⋆ `:67 g:66 := convolution f g (continuous_linear_map.lsmul ℝ ℝ) - measure_theory.measure_space.volume" in convolution +localized "notation (name := convolution) f ` ⋆[`:67 L:67 `, ` μ:67 `] `:0 g:66 := + convolution f g L μ" in convolution +localized "notation (name := convolution.volume) f ` ⋆[`:67 L:67 `]`:0 g:66 := + convolution f g L measure_theory.measure_space.volume" in convolution +localized "notation (name := convolution.lsmul) f ` ⋆ `:67 g:66 := + convolution f g (continuous_linear_map.lsmul ℝ ℝ) measure_theory.measure_space.volume" + in convolution lemma convolution_def [has_sub G] : (f ⋆[L, μ] g) x = ∫ t, L (f t) (g (x - t)) ∂μ := rfl diff --git a/src/analysis/inner_product_space/adjoint.lean b/src/analysis/inner_product_space/adjoint.lean index b6721d77b1977..e03af80401fc1 100644 --- a/src/analysis/inner_product_space/adjoint.lean +++ b/src/analysis/inner_product_space/adjoint.lean @@ -100,7 +100,7 @@ linear_isometry_equiv.of_surjective ..adjoint_aux } (λ A, ⟨adjoint_aux A, adjoint_aux_adjoint_aux A⟩) -localized "postfix `†`:1000 := continuous_linear_map.adjoint" in inner_product +localized "postfix (name := adjoint) `†`:1000 := continuous_linear_map.adjoint" in inner_product /-- The fundamental property of the adjoint. -/ lemma adjoint_inner_left (A : E →L[𝕜] F) (x : E) (y : F) : ⟪A† y, x⟫ = ⟪y, A x⟫ := diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index ed43423e98b00..cf780211e6a01 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -85,8 +85,10 @@ notation `⟪`x`, `y`⟫_ℂ` := @inner ℂ _ _ x y section notations -localized "notation `⟪`x`, `y`⟫` := @inner ℝ _ _ x y" in real_inner_product_space -localized "notation `⟪`x`, `y`⟫` := @inner ℂ _ _ x y" in complex_inner_product_space +localized "notation (name := inner.real) + `⟪`x`, `y`⟫` := @inner ℝ _ _ x y" in real_inner_product_space +localized "notation (name := inner.complex) + `⟪`x`, `y`⟫` := @inner ℂ _ _ x y" in complex_inner_product_space end notations diff --git a/src/analysis/normed_space/lattice_ordered_group.lean b/src/analysis/normed_space/lattice_ordered_group.lean index 737aa3b69e7c1..ecce0fa34c0c1 100644 --- a/src/analysis/normed_space/lattice_ordered_group.lean +++ b/src/analysis/normed_space/lattice_ordered_group.lean @@ -33,7 +33,7 @@ normed, lattice, ordered, group Motivated by the theory of Banach Lattices, this section introduces normed lattice ordered groups. -/ -local notation `|`a`|` := abs a +local notation (name := abs) `|`a`|` := abs a /-- Let `α` be a normed commutative group equipped with a partial order covariant with addition, with diff --git a/src/analysis/quaternion.lean b/src/analysis/quaternion.lean index 2b77f49fb5427..6ff353e009e1d 100644 --- a/src/analysis/quaternion.lean +++ b/src/analysis/quaternion.lean @@ -26,7 +26,7 @@ The following notation is available with `open_locale quaternion`: quaternion, normed ring, normed space, normed algebra -/ -localized "notation `ℍ` := quaternion ℝ" in quaternion +localized "notation (name := quaternion.real) `ℍ` := quaternion ℝ" in quaternion open_locale real_inner_product_space noncomputable theory diff --git a/src/analysis/special_functions/trigonometric/basic.lean b/src/analysis/special_functions/trigonometric/basic.lean index 9696d63e56795..6969a50c07909 100644 --- a/src/analysis/special_functions/trigonometric/basic.lean +++ b/src/analysis/special_functions/trigonometric/basic.lean @@ -97,7 +97,7 @@ intermediate_value_Icc' (by norm_num) continuous_on_cos which one can derive all its properties. For explicit bounds on π, see `data.real.pi.bounds`. -/ protected noncomputable def pi : ℝ := 2 * classical.some exists_cos_eq_zero -localized "notation `π` := real.pi" in real +localized "notation (name := real.pi) `π` := real.pi" in real @[simp] lemma cos_pi_div_two : cos (π / 2) = 0 := by rw [real.pi, mul_div_cancel_left _ (@two_ne_zero' ℝ _ _)]; diff --git a/src/category_theory/bicategory/basic.lean b/src/category_theory/bicategory/basic.lean index ae6620c0000f1..c69c681a711bf 100644 --- a/src/category_theory/bicategory/basic.lean +++ b/src/category_theory/bicategory/basic.lean @@ -107,11 +107,16 @@ class bicategory (B : Type u) extends category_struct.{v} B := (α_ f (𝟙 b) g).hom ≫ f ◁ (λ_ g).hom = (ρ_ f).hom ▷ g . obviously) -- The precedence of the whiskerings is higher than that of the composition `≫`. -localized "infixr ` ◁ `:81 := bicategory.whisker_left" in bicategory -localized "infixl ` ▷ `:81 := bicategory.whisker_right" in bicategory -localized "notation `α_` := bicategory.associator" in bicategory -localized "notation `λ_` := bicategory.left_unitor" in bicategory -localized "notation `ρ_` := bicategory.right_unitor" in bicategory +localized "infixr (name := bicategory.whisker_left) ` ◁ `:81 := bicategory.whisker_left" + in bicategory +localized "infixl (name := bicategory.whisker_right) ` ▷ `:81 := bicategory.whisker_right" + in bicategory +localized "notation (name := bicategory.associator) `α_` := bicategory.associator" + in bicategory +localized "notation (name := bicategory.left_unitor) `λ_` := bicategory.left_unitor" + in bicategory +localized "notation (name := bicategory.right_unitor) `ρ_` := bicategory.right_unitor" + in bicategory namespace bicategory diff --git a/src/category_theory/bicategory/coherence.lean b/src/category_theory/bicategory/coherence.lean index 17e06361f01b2..9b8829296af70 100644 --- a/src/category_theory/bicategory/coherence.lean +++ b/src/category_theory/bicategory/coherence.lean @@ -68,7 +68,7 @@ def preinclusion (B : Type u) [quiver.{v+1} B] : prelax_functor (locally_discrete (paths B)) (free_bicategory B) := { obj := id, map := λ a b, (inclusion_path a b).obj, - map₂ := λ a b, (inclusion_path a b).map } + map₂ := λ a b f g η, (inclusion_path a b).map η } @[simp] lemma preinclusion_obj (a : B) : diff --git a/src/category_theory/bicategory/coherence_tactic.lean b/src/category_theory/bicategory/coherence_tactic.lean index 7b65cfecc8849..f999bdc5e4b95 100644 --- a/src/category_theory/bicategory/coherence_tactic.lean +++ b/src/category_theory/bicategory/coherence_tactic.lean @@ -161,8 +161,8 @@ def bicategorical_comp {f g h i : a ⟶ b} [lift_hom g] [lift_hom h] [bicategorical_coherence g h] (η : f ⟶ g) (θ : h ⟶ i) : f ⟶ i := η ≫ bicategorical_coherence.hom g h ≫ θ -localized "infixr ` ⊗≫ `:80 := category_theory.bicategory.bicategorical_comp" - in bicategory -- type as \ot \gg +localized "infixr (name := bicategorical_comp) ` ⊗≫ `:80 := + category_theory.bicategory.bicategorical_comp" in bicategory -- type as \ot \gg /-- Compose two isomorphisms in a bicategorical category, inserting unitors and associators between as necessary. -/ @@ -170,8 +170,8 @@ def bicategorical_iso_comp {f g h i : a ⟶ b} [lift_hom g] [lift_hom h] [bicategorical_coherence g h] (η : f ≅ g) (θ : h ≅ i) : f ≅ i := η ≪≫ as_iso (bicategorical_coherence.hom g h) ≪≫ θ -localized "infixr ` ≪⊗≫ `:80 := category_theory.bicategory.bicategorical_iso_comp" - in bicategory -- type as \ot \gg +localized "infixr (name := bicategorical_iso_comp) ` ≪⊗≫ `:80 := + category_theory.bicategory.bicategorical_iso_comp" in bicategory -- type as \ot \gg example {f' : a ⟶ d} {f : a ⟶ b} {g : b ⟶ c} {h : c ⟶ d} {h' : a ⟶ d} (η : f' ⟶ f ≫ (g ≫ h)) (θ : (f ≫ g) ≫ h ⟶ h') : f' ⟶ h' := η ⊗≫ θ diff --git a/src/category_theory/bicategory/free.lean b/src/category_theory/bicategory/free.lean index 6528cd11f2bf2..23ace75977b00 100644 --- a/src/category_theory/bicategory/free.lean +++ b/src/category_theory/bicategory/free.lean @@ -65,16 +65,16 @@ section variables {B} -- The following notations are only used in the definition of `rel` to simplify the notation. -local infixr ` ≫ ` := hom₂.vcomp -local notation `𝟙` := hom₂.id -local notation f ` ◁ ` η := hom₂.whisker_left f η -local notation η ` ▷ ` h := hom₂.whisker_right h η -local notation `α_` := hom₂.associator -local notation `λ_` := hom₂.left_unitor -local notation `ρ_` := hom₂.right_unitor -local notation `α⁻¹_` := hom₂.associator_inv -local notation `λ⁻¹_` := hom₂.left_unitor_inv -local notation `ρ⁻¹_` := hom₂.right_unitor_inv +local infixr (name := vcomp) ` ≫ ` := hom₂.vcomp +local notation (name := id) `𝟙` := hom₂.id +local notation (name := whisker_left) f ` ◁ ` η := hom₂.whisker_left f η +local notation (name := whisker_right) η ` ▷ ` h := hom₂.whisker_right h η +local notation (name := associator) `α_` := hom₂.associator +local notation (name := left_unitor) `λ_` := hom₂.left_unitor +local notation (name := right_unitor) `ρ_` := hom₂.right_unitor +local notation (name := associator_inv) `α⁻¹_` := hom₂.associator_inv +local notation (name := left_unitor_inv) `λ⁻¹_` := hom₂.left_unitor_inv +local notation (name := right_unitor_inv) `ρ⁻¹_` := hom₂.right_unitor_inv /-- Relations between 2-morphisms in the free bicategory. -/ inductive rel : Π {a b : B} {f g : hom a b}, hom₂ f g → hom₂ f g → Prop diff --git a/src/category_theory/bicategory/functor.lean b/src/category_theory/bicategory/functor.lean index ad5757a9bb35b..6d5e10ff3110e 100644 --- a/src/category_theory/bicategory/functor.lean +++ b/src/category_theory/bicategory/functor.lean @@ -90,7 +90,7 @@ variables (F : prelax_functor B C) @[simp] lemma to_prefunctor_eq_coe : F.to_prefunctor = F := rfl @[simp] lemma to_prefunctor_obj : (F : prefunctor B C).obj = F.obj := rfl -@[simp] lemma to_prefunctor_map : (F : prefunctor B C).map = F.map := rfl +@[simp] lemma to_prefunctor_map : @prefunctor.map B _ C _ F = @map _ _ _ _ _ _ F := rfl /-- The identity prelax functor. -/ @[simps] @@ -153,7 +153,7 @@ structure oplax_functor (B : Type u₁) [bicategory.{w₁ v₁} B] (C : Type u (map₂_comp' : ∀ {a b : B} {f g h : a ⟶ b} (η : f ⟶ g) (θ : g ⟶ h), map₂ (η ≫ θ) = map₂ η ≫ map₂ θ . obviously) (map₂_associator' : ∀ {a b c d : B} (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d), - oplax_functor.map₂_associator_aux obj (λ a b, map) (λ a b f g, map₂) (λ a b c, map_comp) f g h + oplax_functor.map₂_associator_aux obj (λ _ _, map) (λ a b f g, map₂) (λ a b c, map_comp) f g h . obviously) (map₂_left_unitor' : ∀ {a b : B} (f : a ⟶ b), map₂ (λ_ f).hom = map_comp (𝟙 a) f ≫ map_id a ▷ map f ≫ (λ_ (map f)).hom . obviously) @@ -187,8 +187,8 @@ variables (F : oplax_functor B C) @[simp] lemma to_prelax_eq_coe : F.to_prelax_functor = F := rfl @[simp] lemma to_prelax_functor_obj : (F : prelax_functor B C).obj = F.obj := rfl -@[simp] lemma to_prelax_functor_map : (F : prelax_functor B C).map = F.map := rfl -@[simp] lemma to_prelax_functor_map₂ : (F : prelax_functor B C).map₂ = F.map₂ := rfl +@[simp] lemma to_prelax_functor_map : @prelax_functor.map B _ _ C _ _ F = @map _ _ _ _ F := rfl +@[simp] lemma to_prelax_functor_map₂ : @prelax_functor.map₂ B _ _ C _ _ F = @map₂ _ _ _ _ F := rfl /-- Function between 1-morphisms as a functor. -/ @[simps] @@ -334,8 +334,8 @@ variables (F : pseudofunctor B C) @[simp] lemma to_prelax_functor_eq_coe : F.to_prelax_functor = F := rfl @[simp] lemma to_prelax_functor_obj : (F : prelax_functor B C).obj = F.obj := rfl -@[simp] lemma to_prelax_functor_map : (F : prelax_functor B C).map = F.map := rfl -@[simp] lemma to_prelax_functor_map₂ : (F : prelax_functor B C).map₂ = F.map₂ := rfl +@[simp] lemma to_prelax_functor_map : @prelax_functor.map B _ _ C _ _ F = @map _ _ _ _ F := rfl +@[simp] lemma to_prelax_functor_map₂ : @prelax_functor.map₂ B _ _ C _ _ F = @map₂ _ _ _ _ F := rfl /-- The oplax functor associated with a pseudofunctor. -/ def to_oplax : oplax_functor B C := @@ -347,8 +347,8 @@ instance has_coe_to_oplax : has_coe (pseudofunctor B C) (oplax_functor B C) := @[simp] lemma to_oplax_eq_coe : F.to_oplax = F := rfl @[simp] lemma to_oplax_obj : (F : oplax_functor B C).obj = F.obj := rfl -@[simp] lemma to_oplax_map : (F : oplax_functor B C).map = F.map := rfl -@[simp] lemma to_oplax_map₂ : (F : oplax_functor B C).map₂ = F.map₂ := rfl +@[simp] lemma to_oplax_map : @oplax_functor.map B _ C _ F = @map _ _ _ _ F := rfl +@[simp] lemma to_oplax_map₂ : @oplax_functor.map₂ B _ C _ F = @map₂ _ _ _ _ F := rfl @[simp] lemma to_oplax_map_id (a : B) : (F : oplax_functor B C).map_id a = (F.map_id a).hom := rfl @[simp] lemma to_oplax_map_comp {a b c : B} (f : a ⟶ b) (g : b ⟶ c) : (F : oplax_functor B C).map_comp f g = (F.map_comp f g).hom := rfl @@ -381,7 +381,7 @@ Construct a pseudofunctor from an oplax functor whose `map_id` and `map_comp` ar @[simps] def mk_of_oplax (F : oplax_functor B C) (F' : F.pseudo_core) : pseudofunctor B C := { map_id := F'.map_id_iso, - map_comp := F'.map_comp_iso, + map_comp := λ _ _ _, F'.map_comp_iso, map₂_whisker_left' := λ a b c f g h η, by { dsimp, rw [F'.map_comp_iso_hom f g, ←F.map_comp_naturality_right_assoc, diff --git a/src/category_theory/closed/monoidal.lean b/src/category_theory/closed/monoidal.lean index 1a8c4194dfdce..433aec0d14dce 100644 --- a/src/category_theory/closed/monoidal.lean +++ b/src/category_theory/closed/monoidal.lean @@ -107,7 +107,7 @@ lemma coev_naturality {X Y : C} (f : X ⟶ Y) : f ≫ (coev A).app Y = (coev A).app X ≫ (ihom A).map ((𝟙 A) ⊗ f) := (coev A).naturality f -notation A ` ⟶[`C`] ` B:10 := (@ihom C _ _ A _).obj B +notation (name := ihom) A ` ⟶[`C`] ` B:10 := (@ihom C _ _ A _).obj B @[simp, reassoc] lemma ev_coev : ((𝟙 A) ⊗ ((coev A).app B)) ≫ (ev A).app (A ⊗ B) = 𝟙 (A ⊗ B) := diff --git a/src/category_theory/enriched/basic.lean b/src/category_theory/enriched/basic.lean index dcb15e7119b03..8644f3bafb15a 100644 --- a/src/category_theory/enriched/basic.lean +++ b/src/category_theory/enriched/basic.lean @@ -55,7 +55,7 @@ class enriched_category (C : Type u₁) := Π W X Y Z, (α_ _ _ _).inv ≫ (comp W X Y ⊗ 𝟙 _) ≫ comp W Y Z = (𝟙 _ ⊗ comp X Y Z) ≫ comp W X Z . obviously) -notation X ` ⟶[`V`] ` Y:10 := (enriched_category.hom X Y : V) +notation (name := enriched_category.hom) X ` ⟶[`V`] ` Y:10 := (enriched_category.hom X Y : V) variables (V) {C : Type u₁} [enriched_category V C] diff --git a/src/category_theory/equivalence.lean b/src/category_theory/equivalence.lean index fec19f77119b8..49fa891c5602a 100644 --- a/src/category_theory/equivalence.lean +++ b/src/category_theory/equivalence.lean @@ -135,8 +135,9 @@ by { erw [←iso.hom_comp_eq_id (e.functor.map_iso (e.unit_iso.app X)), functor_ @[simp] lemma unit_inverse_comp (e : C ≌ D) (Y : D) : e.unit.app (e.inverse.obj Y) ≫ e.inverse.map (e.counit.app Y) = 𝟙 (e.inverse.obj Y) := begin - rw [←id_comp (e.inverse.map _), ←map_id e.inverse, ←counit_inv_functor_comp, map_comp, - ←iso.hom_inv_id_assoc (e.unit_iso.app _) (e.inverse.map (e.functor.map _)), + rw [←id_comp (e.inverse.map _), ←map_id e.inverse, ←counit_inv_functor_comp, map_comp], + dsimp, + rw [←iso.hom_inv_id_assoc (e.unit_iso.app _) (e.inverse.map (e.functor.map _)), app_hom, app_inv], slice_lhs 2 3 { erw [e.unit.naturality] }, slice_lhs 1 2 { erw [e.unit.naturality] }, diff --git a/src/category_theory/monad/limits.lean b/src/category_theory/monad/limits.lean index fe9acefab8e64..aa3cbf51c4909 100644 --- a/src/category_theory/monad/limits.lean +++ b/src/category_theory/monad/limits.lean @@ -339,8 +339,8 @@ lemma has_colimits_of_shape_of_reflective (R : D ⥤ C) { has_colimit := λ F, begin let c := (left_adjoint R).map_cocone (colimit.cocone (F ⋙ R)), - let h := (adjunction.of_right_adjoint R).left_adjoint_preserves_colimits.1, - letI := @h J _, + letI : preserves_colimits_of_shape J _ := + (adjunction.of_right_adjoint R).left_adjoint_preserves_colimits.1, let t : is_colimit c := is_colimit_of_preserves (left_adjoint R) (colimit.is_colimit _), apply has_colimit.mk ⟨_, (is_colimit.precompose_inv_equiv _ _).symm t⟩, apply (iso_whisker_left F (as_iso (adjunction.of_right_adjoint R).counit) : _) ≪≫ F.right_unitor, diff --git a/src/category_theory/monoidal/category.lean b/src/category_theory/monoidal/category.lean index b1bb884e221f4..08f64db141167 100644 --- a/src/category_theory/monoidal/category.lean +++ b/src/category_theory/monoidal/category.lean @@ -67,7 +67,7 @@ See . class monoidal_category (C : Type u) [𝒞 : category.{v} C] := -- curried tensor product of objects: (tensor_obj : C → C → C) -(infixr ` ⊗ `:70 := tensor_obj) -- This notation is only temporary +(infixr (name := tensor_obj) ` ⊗ `:70 := tensor_obj) -- This notation is only temporary -- curried tensor product of morphisms: (tensor_hom : Π {X₁ Y₁ X₂ Y₂ : C}, (X₁ ⟶ Y₁) → (X₂ ⟶ Y₂) → ((X₁ ⊗ X₂) ⟶ (Y₁ ⊗ Y₂))) @@ -124,8 +124,8 @@ attribute [simp, reassoc] monoidal_category.triangle open monoidal_category -infixr ` ⊗ `:70 := tensor_obj -infixr ` ⊗ `:70 := tensor_hom +infixr (name := tensor_obj) ` ⊗ `:70 := tensor_obj +infixr (name := tensor_hom) ` ⊗ `:70 := tensor_hom notation `𝟙_` := tensor_unit notation `α_` := associator @@ -142,7 +142,7 @@ def tensor_iso {C : Type u} {X Y X' Y' : C} [category.{v} C] [monoidal_category. hom_inv_id' := by rw [←tensor_comp, iso.hom_inv_id, iso.hom_inv_id, ←tensor_id], inv_hom_id' := by rw [←tensor_comp, iso.inv_hom_id, iso.inv_hom_id, ←tensor_id] } -infixr ` ⊗ `:70 := tensor_iso +infixr (name := tensor_iso) ` ⊗ `:70 := tensor_iso namespace monoidal_category diff --git a/src/category_theory/monoidal/functor.lean b/src/category_theory/monoidal/functor.lean index 6f3630b5a3abb..d43a48e845e3c 100644 --- a/src/category_theory/monoidal/functor.lean +++ b/src/category_theory/monoidal/functor.lean @@ -359,7 +359,8 @@ def comp : monoidal_functor.{v₁ v₃} C E := μ_is_iso := by { dsimp, apply_instance }, .. (F.to_lax_monoidal_functor).comp (G.to_lax_monoidal_functor) }. -infixr ` ⊗⋙ `:80 := comp -- We overload notation; potentially dangerous, but it seems to work. +-- We overload notation; potentially dangerous, but it seems to work. +infixr (name := monoidal_functor.comp) ` ⊗⋙ `:80 := comp end monoidal_functor diff --git a/src/category_theory/monoidal/rigid/basic.lean b/src/category_theory/monoidal/rigid/basic.lean index be984afca9497..a8c467d0f0362 100644 --- a/src/category_theory/monoidal/rigid/basic.lean +++ b/src/category_theory/monoidal/rigid/basic.lean @@ -108,8 +108,8 @@ attribute [instance] has_left_dual.exact open exact_pairing has_right_dual has_left_dual monoidal_category -prefix `ᘁ`:1025 := left_dual -postfix `ᘁ`:1025 := right_dual +prefix (name := left_dual) `ᘁ`:1025 := left_dual +postfix (name := right_dual) `ᘁ`:1025 := right_dual instance has_right_dual_unit : has_right_dual (𝟙_ C) := { right_dual := 𝟙_ C } @@ -139,8 +139,8 @@ def left_adjoint_mate {X Y : C} [has_left_dual X] [has_left_dual Y] (f : X ⟶ Y (λ_ _).inv ≫ (η_ (ᘁX) X ⊗ 𝟙 _) ≫ ((𝟙 _ ⊗ f) ⊗ 𝟙 _) ≫ (α_ _ _ _).hom ≫ (𝟙 _ ⊗ ε_ _ _) ≫ (ρ_ _).hom -notation f `ᘁ` := right_adjoint_mate f -notation `ᘁ` f := left_adjoint_mate f +notation (name := right_adjoint_mate) f `ᘁ` := right_adjoint_mate f +notation (name := left_adjoint_mate) `ᘁ` f := left_adjoint_mate f @[simp] lemma right_adjoint_mate_id {X : C} [has_right_dual X] : (𝟙 X)ᘁ = 𝟙 (Xᘁ) := diff --git a/src/category_theory/types.lean b/src/category_theory/types.lean index 41c59d98a3a8b..1f16fcbe2ad27 100644 --- a/src/category_theory/types.lean +++ b/src/category_theory/types.lean @@ -63,7 +63,7 @@ congr_fun f.inv_hom_id y -- Unfortunately without this wrapper we can't use `category_theory` idioms, such as `is_iso f`. abbreviation as_hom {α β : Type u} (f : α → β) : α ⟶ β := f -- If you don't mind some notation you can use fewer keystrokes: -localized "notation `↾` f : 200 := category_theory.as_hom f" +localized "notation (name := category_theory.as_hom) `↾` f : 200 := category_theory.as_hom f" in category_theory.Type -- type as \upr in VScode section -- We verify the expected type checking behaviour of `as_hom`. diff --git a/src/combinatorics/configuration.lean b/src/combinatorics/configuration.lean index 353471a704eb0..1c3de6a1590dc 100644 --- a/src/combinatorics/configuration.lean +++ b/src/combinatorics/configuration.lean @@ -301,8 +301,8 @@ end in noncomputable def has_points.has_lines [has_points P L] [fintype P] [fintype L] (h : fintype.card P = fintype.card L) : has_lines P L := let this := @has_lines.has_points (dual L) (dual P) _ _ _ _ h.symm in -{ mk_line := this.mk_point, - mk_line_ax := this.mk_point_ax } +{ mk_line := λ _ _, this.mk_point, + mk_line_ax := λ _ _, this.mk_point_ax } variables (P L) diff --git a/src/combinatorics/set_family/compression/down.lean b/src/combinatorics/set_family/compression/down.lean index b9470340b6671..df9d79b4c574a 100644 --- a/src/combinatorics/set_family/compression/down.lean +++ b/src/combinatorics/set_family/compression/down.lean @@ -130,7 +130,7 @@ def compression (a : α) (𝒜 : finset (finset α)) : finset (finset α) := (𝒜.filter $ λ s, erase s a ∈ 𝒜).disj_union ((𝒜.image $ λ s, erase s a).filter $ λ s, s ∉ 𝒜) $ λ s h₁ h₂, (mem_filter.1 h₂).2 (mem_filter.1 h₁).1 -localized "notation `𝓓 ` := down.compression" in finset_family +localized "notation (name := down.compression) `𝓓 ` := down.compression" in finset_family /-- `a` is in the down-compressed family iff it's in the original and its compression is in the original, or it's not in the original but it's the compression of something in the original. -/ diff --git a/src/combinatorics/set_family/compression/uv.lean b/src/combinatorics/set_family/compression/uv.lean index 4c1dd364e0b6d..453c48e94f14e 100644 --- a/src/combinatorics/set_family/compression/uv.lean +++ b/src/combinatorics/set_family/compression/uv.lean @@ -85,7 +85,7 @@ reduce the cardinality, so we keep all elements whose compression is already pre def compression (u v : α) (s : finset α) := s.filter (λ a, compress u v a ∈ s) ∪ (s.image $ compress u v).filter (λ a, a ∉ s) -localized "notation `𝓒 ` := uv.compression" in finset_family +localized "notation (name := uv.compression) `𝓒 ` := uv.compression" in finset_family /-- `is_compressed u v s` expresses that `s` is UV-compressed. -/ def is_compressed (u v : α) (s : finset α) := 𝓒 u v s = s diff --git a/src/combinatorics/set_family/shadow.lean b/src/combinatorics/set_family/shadow.lean index d08ac501c6671..7471273e79fb1 100644 --- a/src/combinatorics/set_family/shadow.lean +++ b/src/combinatorics/set_family/shadow.lean @@ -53,7 +53,7 @@ variables [decidable_eq α] {𝒜 : finset (finset α)} {s t : finset α} {a : elements from any set in `𝒜`. -/ def shadow (𝒜 : finset (finset α)) : finset (finset α) := 𝒜.sup (λ s, s.image (erase s)) -localized "notation `∂ `:90 := finset.shadow" in finset_family +localized "notation (name := finset.shadow) `∂ `:90 := finset.shadow" in finset_family /-- The shadow of the empty set is empty. -/ @[simp] lemma shadow_empty : ∂ (∅ : finset (finset α)) = ∅ := rfl @@ -160,7 +160,7 @@ variables [decidable_eq α] [fintype α] {𝒜 : finset (finset α)} {s t : fins def up_shadow (𝒜 : finset (finset α)) : finset (finset α) := 𝒜.sup $ λ s, sᶜ.image $ λ a, insert a s -localized "notation `∂⁺ `:90 := finset.up_shadow" in finset_family +localized "notation (name := finset.up_shadow) `∂⁺ `:90 := finset.up_shadow" in finset_family /-- The upper shadow of the empty set is empty. -/ @[simp] lemma up_shadow_empty : ∂⁺ (∅ : finset (finset α)) = ∅ := rfl diff --git a/src/combinatorics/simple_graph/prod.lean b/src/combinatorics/simple_graph/prod.lean index 33d59b673e90b..8c1eaa9c9c0f1 100644 --- a/src/combinatorics/simple_graph/prod.lean +++ b/src/combinatorics/simple_graph/prod.lean @@ -85,21 +85,23 @@ variables (G) {H} protected def box_prod_right (a : α) : H.walk b₁ b₂ → (G □ H).walk (a, b₁) (a, b₂) := walk.map (G.box_prod_right H a).to_hom -variables {G} [decidable_eq α] [decidable_eq β] [decidable_rel G.adj] [decidable_rel H.adj] +variables {G} /-- Project a walk on `G □ H` to a walk on `G` by discarding the moves in the direction of `H`. -/ -def of_box_prod_left : Π {x y : α × β}, (G □ H).walk x y → G.walk x.1 y.1 +def of_box_prod_left [decidable_eq β] [decidable_rel G.adj] : + Π {x y : α × β}, (G □ H).walk x y → G.walk x.1 y.1 | _ _ nil := nil | x z (cons h w) := or.by_cases h (λ hG, w.of_box_prod_left.cons hG.1) (λ hH, show G.walk x.1 z.1, by rw hH.2; exact w.of_box_prod_left) /-- Project a walk on `G □ H` to a walk on `H` by discarding the moves in the direction of `G`. -/ -def of_box_prod_right : Π {x y : α × β}, (G □ H).walk x y → H.walk x.2 y.2 +def of_box_prod_right [decidable_eq α] [decidable_rel H.adj] : + Π {x y : α × β}, (G □ H).walk x y → H.walk x.2 y.2 | _ _ nil := nil | x z (cons h w) := (or.symm h).by_cases (λ hH, w.of_box_prod_right.cons hH.1) (λ hG, show H.walk x.2 z.2, by rw hG.2; exact w.of_box_prod_right) -@[simp] lemma of_box_prod_left_box_prod_left : +@[simp] lemma of_box_prod_left_box_prod_left [decidable_eq β] [decidable_rel G.adj] : ∀ {a₁ a₂ : α} (w : G.walk a₁ a₂), (w.box_prod_left H b).of_box_prod_left = w | _ _ nil := rfl | _ _ (cons' x y z h w) := begin @@ -108,7 +110,7 @@ def of_box_prod_right : Π {x y : α × β}, (G □ H).walk x y → H.walk x.2 y exacts [rfl, ⟨h, rfl⟩], end -@[simp] lemma of_box_prod_left_box_prod_right : +@[simp] lemma of_box_prod_left_box_prod_right [decidable_eq α] [decidable_rel G.adj] : ∀ {b₁ b₂ : α} (w : G.walk b₁ b₂), (w.box_prod_right G a).of_box_prod_right = w | _ _ nil := rfl | _ _ (cons' x y z h w) := begin diff --git a/src/combinatorics/simple_graph/subgraph.lean b/src/combinatorics/simple_graph/subgraph.lean index f3aca1f2e626c..8ca062543d1c2 100644 --- a/src/combinatorics/simple_graph/subgraph.lean +++ b/src/combinatorics/simple_graph/subgraph.lean @@ -186,8 +186,8 @@ def copy (G' : subgraph G) subgraph G := { verts := V'', adj := adj', - adj_sub := hadj.symm ▸ G'.adj_sub, - edge_vert := hV.symm ▸ hadj.symm ▸ G'.edge_vert, + adj_sub := λ _ _, hadj.symm ▸ G'.adj_sub, + edge_vert := λ _ _, hV.symm ▸ hadj.symm ▸ G'.edge_vert, symm := hadj.symm ▸ G'.symm } lemma copy_eq (G' : subgraph G) diff --git a/src/control/functor/multivariate.lean b/src/control/functor/multivariate.lean index 0acfc191fe0c8..da4c8fa925150 100644 --- a/src/control/functor/multivariate.lean +++ b/src/control/functor/multivariate.lean @@ -28,7 +28,7 @@ and the category of Type -/ class mvfunctor {n : ℕ} (F : typevec n → Type*) := (map : Π {α β : typevec n}, (α ⟹ β) → (F α → F β)) -localized "infixr ` <$$> `:100 := mvfunctor.map" in mvfunctor +localized "infixr (name := mvfunctor.map) ` <$$> `:100 := mvfunctor.map" in mvfunctor variables {n : ℕ} diff --git a/src/control/traversable/basic.lean b/src/control/traversable/basic.lean index fd30e7352d29d..e11e66dba7c9b 100644 --- a/src/control/traversable/basic.lean +++ b/src/control/traversable/basic.lean @@ -115,11 +115,10 @@ section preserves variables (η : applicative_transformation F G) @[functor_norm] -lemma preserves_pure : ∀ {α} (x : α), η (pure x) = pure x := η.preserves_pure' +lemma preserves_pure {α} : ∀ (x : α), η (pure x) = pure x := η.preserves_pure' @[functor_norm] -lemma preserves_seq : - ∀ {α β : Type u} (x : F (α → β)) (y : F α), η (x <*> y) = η x <*> η y := +lemma preserves_seq {α β : Type u} : ∀ (x : F (α → β)) (y : F α), η (x <*> y) = η x <*> η y := η.preserves_seq' @[functor_norm] diff --git a/src/data/finmap.lean b/src/data/finmap.lean index 7b064236817dc..4e4c242f50af9 100644 --- a/src/data/finmap.lean +++ b/src/data/finmap.lean @@ -45,7 +45,7 @@ structure finmap (β : α → Type v) : Type (max u v) := /-- The quotient map from `alist` to `finmap`. -/ def alist.to_finmap (s : alist β) : finmap β := ⟨s.entries, s.nodupkeys⟩ -local notation `⟦`:max a `⟧`:0 := alist.to_finmap a +local notation (name := to_finmap) `⟦`:max a `⟧`:0 := alist.to_finmap a theorem alist.to_finmap_eq {s₁ s₂ : alist β} : ⟦s₁⟧ = ⟦s₂⟧ ↔ s₁.entries ~ s₂.entries := diff --git a/src/data/finset/fold.lean b/src/data/finset/fold.lean index 4c76ff85c76e2..a6db4aa26f38e 100644 --- a/src/data/finset/fold.lean +++ b/src/data/finset/fold.lean @@ -18,7 +18,7 @@ variables {α β γ : Type*} /-! ### fold -/ section fold variables (op : β → β → β) [hc : is_commutative β op] [ha : is_associative β op] -local notation a * b := op a b +local notation (name := op) a * b := op a b include hc ha /-- `fold op b f s` folds the commutative associative operation `op` over the diff --git a/src/data/finset/prod.lean b/src/data/finset/prod.lean index 885a9af98772b..5e8f04e86d7cb 100644 --- a/src/data/finset/prod.lean +++ b/src/data/finset/prod.lean @@ -34,7 +34,7 @@ variables {s s' : finset α} {t t' : finset β} {a : α} {b : β} protected def product (s : finset α) (t : finset β) : finset (α × β) := ⟨_, s.nodup.product t.nodup⟩ /- This notation binds more strongly than (pre)images, unions and intersections. -/ -infixr ` ×ˢ `:82 := finset.product +infixr (name := finset.product) ` ×ˢ `:82 := finset.product @[simp] lemma product_val : (s ×ˢ t).1 = s.1 ×ˢ t.1 := rfl diff --git a/src/data/finset/slice.lean b/src/data/finset/slice.lean index 05879414a878f..7f780f844b2b4 100644 --- a/src/data/finset/slice.lean +++ b/src/data/finset/slice.lean @@ -100,7 +100,7 @@ variables {𝒜 : finset (finset α)} {A A₁ A₂ : finset α} {r r₁ r₂ : /-- The `r`-th slice of a set family is the subset of its elements which have cardinality `r`. -/ def slice (𝒜 : finset (finset α)) (r : ℕ) : finset (finset α) := 𝒜.filter (λ i, i.card = r) -localized "infix ` # `:90 := finset.slice" in finset_family +localized "infix (name := finset.slice) ` # `:90 := finset.slice" in finset_family /-- `A` is in the `r`-th slice of `𝒜` iff it's in `𝒜` and has cardinality `r`. -/ lemma mem_slice : A ∈ 𝒜 # r ↔ A ∈ 𝒜 ∧ A.card = r := mem_filter diff --git a/src/data/fintype/card_embedding.lean b/src/data/fintype/card_embedding.lean index f5f3a810df06b..6bd0a76a67166 100644 --- a/src/data/fintype/card_embedding.lean +++ b/src/data/fintype/card_embedding.lean @@ -13,8 +13,8 @@ import logic.equiv.embedding This file establishes the cardinality of `α ↪ β` in full generality. -/ -local notation `|` x `|` := finset.card x -local notation `‖` x `‖` := fintype.card x +local notation (name := finset.card) `|` x `|` := finset.card x +local notation (name := fintype.card) `‖` x `‖` := fintype.card x open function open_locale nat big_operators diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 7f155655805e0..9d8b10d2692cf 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -2459,8 +2459,8 @@ end foldl_eq_foldlr' section variables {op : α → α → α} [ha : is_associative α op] [hc : is_commutative α op] -local notation a * b := op a b -local notation l <*> a := foldl op a l +local notation (name := op) a * b := op a b +local notation (name := foldl) l <*> a := foldl op a l include ha diff --git a/src/data/list/defs.lean b/src/data/list/defs.lean index 6c8c0a3cf96a0..641e6ad01e818 100644 --- a/src/data/list/defs.lean +++ b/src/data/list/defs.lean @@ -541,7 +541,7 @@ def product (l₁ : list α) (l₂ : list β) : list (α × β) := l₁.bind $ λ a, l₂.map $ prod.mk a /- This notation binds more strongly than (pre)images, unions and intersections. -/ -infixr ` ×ˢ `:82 := list.product +infixr (name := list.product) ` ×ˢ `:82 := list.product /-- `sigma l₁ l₂` is the list of dependent pairs `(a, b)` where `a ∈ l₁` and `b ∈ l₂ a`. diff --git a/src/data/list/func.lean b/src/data/list/func.lean index acc07aa763c6d..4cbef33c6264e 100644 --- a/src/data/list/func.lean +++ b/src/data/list/func.lean @@ -50,7 +50,8 @@ elements | (h::as) (k+1) := h::(set as k) | [] (k+1) := default::(set ([] : list α) k) -localized "notation as ` {` m ` ↦ ` a `}` := list.func.set a as m" in list.func +localized "notation (name := list.func.set) + as ` {` m ` ↦ ` a `}` := list.func.set a as m" in list.func /-- Get element of a list by index. If the index is out of range, return the default element -/ @[simp] def get : ℕ → list α → α diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index 88df2a51d0e59..cacfad8a95820 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -38,7 +38,7 @@ inductive perm : list α → list α → Prop open perm (swap) -infix ` ~ `:50 := perm +infix (name := list.perm) ` ~ `:50 := perm @[refl] protected theorem perm.refl : ∀ (l : list α), l ~ l | [] := perm.nil @@ -457,8 +457,8 @@ end section variables {op : α → α → α} [is_associative α op] [is_commutative α op] -local notation a * b := op a b -local notation l <*> a := foldl op a l +local notation (name := op) a * b := op a b +local notation (name := foldl) l <*> a := foldl op a l lemma perm.fold_op_eq {l₁ l₂ : list α} {a : α} (h : l₁ ~ l₂) : l₁ <*> a = l₂ <*> a := h.foldl_eq (right_comm _ is_commutative.comm is_associative.assoc) _ diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index b9cd3f40a67b2..08d0d9237e2dc 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -119,13 +119,13 @@ lemma map_injective {f : α → β} (hf : function.injective f) : def transpose (M : matrix m n α) : matrix n m α | x y := M y x -localized "postfix `ᵀ`:1500 := matrix.transpose" in matrix +localized "postfix (name := matrix.transpose) `ᵀ`:1500 := matrix.transpose" in matrix /-- The conjugate transpose of a matrix defined in term of `star`. -/ def conj_transpose [has_star α] (M : matrix m n α) : matrix n m α := M.transpose.map star -localized "postfix `ᴴ`:1500 := matrix.conj_transpose" in matrix +localized "postfix (name := matrix.conj_transpose) `ᴴ`:1500 := matrix.conj_transpose" in matrix /-- `matrix.col u` is the column matrix whose entries are given by `u`. -/ def col (w : m → α) : matrix m unit α @@ -433,7 +433,7 @@ def dot_product [has_mul α] [add_comm_monoid α] (v w : m → α) : α := /- The precedence of 72 comes immediately after ` • ` for `has_smul.smul`, so that `r₁ • a ⬝ᵥ r₂ • b` is parsed as `(r₁ • a) ⬝ᵥ (r₂ • b)` here. -/ -localized "infix ` ⬝ᵥ `:72 := matrix.dot_product" in matrix +localized "infix (name := matrix.dot_product) ` ⬝ᵥ `:72 := matrix.dot_product" in matrix lemma dot_product_assoc [non_unital_semiring α] (u : m → α) (w : n → α) (v : matrix m n α) : @@ -549,7 +549,7 @@ protected def mul [fintype m] [has_mul α] [add_comm_monoid α] (M : matrix l m α) (N : matrix m n α) : matrix l n α := λ i k, (λ j, M i j) ⬝ᵥ (λ j, N j k) -localized "infixl ` ⬝ `:75 := matrix.mul" in matrix +localized "infixl (name := matrix.mul) ` ⬝ `:75 := matrix.mul" in matrix theorem mul_apply [fintype m] [has_mul α] [add_comm_monoid α] {M : matrix l m α} {N : matrix m n α} {i k} : (M ⬝ N) i k = ∑ j, M i j * N j k := rfl diff --git a/src/data/matrix/dmatrix.lean b/src/data/matrix/dmatrix.lean index a8a0854578708..dacbb08124504 100644 --- a/src/data/matrix/dmatrix.lean +++ b/src/data/matrix/dmatrix.lean @@ -53,7 +53,7 @@ by { ext, simp, } def transpose (M : dmatrix m n α) : dmatrix n m (λ j i, α i j) | x y := M y x -localized "postfix `ᵀ`:1500 := dmatrix.transpose" in dmatrix +localized "postfix (name := dmatrix.transpose) `ᵀ`:1500 := dmatrix.transpose" in dmatrix /-- `dmatrix.col u` is the column matrix whose entries are given by `u`. -/ def col {α : m → Type v} (w : Π i, α i) : dmatrix m unit (λ i j, α i) diff --git a/src/data/matrix/hadamard.lean b/src/data/matrix/hadamard.lean index 60d731fce4ae6..b93f500f0de78 100644 --- a/src/data/matrix/hadamard.lean +++ b/src/data/matrix/hadamard.lean @@ -41,7 +41,7 @@ open_locale matrix big_operators def hadamard [has_mul α] (A : matrix m n α) (B : matrix m n α) : matrix m n α | i j := A i j * B i j -localized "infix ` ⊙ `:100 := matrix.hadamard" in matrix +localized "infix (name := matrix.hadamard) ` ⊙ `:100 := matrix.hadamard" in matrix section basic_properties diff --git a/src/data/matrix/kronecker.lean b/src/data/matrix/kronecker.lean index 4b3c9b9eab590..0137a0896b01e 100644 --- a/src/data/matrix/kronecker.lean +++ b/src/data/matrix/kronecker.lean @@ -197,7 +197,8 @@ open_locale matrix @[simp] def kronecker [has_mul α] : matrix l m α → matrix n p α → matrix (l × n) (m × p) α := kronecker_map (*) -localized "infix ` ⊗ₖ `:100 := matrix.kronecker_map (*)" in kronecker +localized "infix (name := matrix.kronecker_map.mul) + ` ⊗ₖ `:100 := matrix.kronecker_map (*)" in kronecker @[simp] lemma kronecker_apply [has_mul α] (A : matrix l m α) (B : matrix n p α) (i₁ i₂ j₁ j₂) : @@ -276,10 +277,10 @@ Prefer the notation `⊗ₖₜ` rather than this definition. -/ matrix l m α → matrix n p β → matrix (l × n) (m × p) (α ⊗[R] β) := kronecker_map (⊗ₜ) -localized "infix ` ⊗ₖₜ `:100 := matrix.kronecker_map (⊗ₜ)" in kronecker -localized - "notation x ` ⊗ₖₜ[`:100 R `] `:0 y:100 := matrix.kronecker_map (tensor_product.tmul R) x y" - in kronecker +localized "infix (name := matrix.kronecker_map.tmul) + ` ⊗ₖₜ `:100 := matrix.kronecker_map (⊗ₜ)" in kronecker +localized "notation (name := matrix.kronecker_map.tmul') + x ` ⊗ₖₜ[`:100 R `] `:0 y:100 := matrix.kronecker_map (tensor_product.tmul R) x y" in kronecker @[simp] lemma kronecker_tmul_apply (A : matrix l m α) (B : matrix n p β) (i₁ i₂ j₁ j₂) : diff --git a/src/data/multiset/bind.lean b/src/data/multiset/bind.lean index 9191a8111b743..1942de94367d9 100644 --- a/src/data/multiset/bind.lean +++ b/src/data/multiset/bind.lean @@ -151,7 +151,7 @@ variables (a : α) (b : β) (s : multiset α) (t : multiset β) def product (s : multiset α) (t : multiset β) : multiset (α × β) := s.bind $ λ a, t.map $ prod.mk a /- This notation binds more strongly than (pre)images, unions and intersections. -/ -infixr ` ×ˢ `:82 := multiset.product +infixr (name := multiset.product) ` ×ˢ `:82 := multiset.product @[simp] lemma coe_product (l₁ : list α) (l₂ : list β) : @product α β l₁ l₂ = l₁.product l₂ := by { rw [product, list.product, ←coe_bind], simp } diff --git a/src/data/multiset/fold.lean b/src/data/multiset/fold.lean index 33a0279a5e87c..ff8c7fc44d233 100644 --- a/src/data/multiset/fold.lean +++ b/src/data/multiset/fold.lean @@ -16,7 +16,7 @@ variables {α β : Type*} /-! ### fold -/ section fold variables (op : α → α → α) [hc : is_commutative α op] [ha : is_associative α op] -local notation a * b := op a b +local notation (name := op) a * b := op a b include hc ha /-- `fold op b s` folds a commutative associative operation `op` over diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index 9db3673f1a36d..de379321d77c9 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -633,7 +633,7 @@ def le_rec_on {C : ℕ → Sort u} {n : ℕ} : Π {m : ℕ}, n ≤ m → (Π {k} theorem le_rec_on_self {C : ℕ → Sort u} {n} {h : n ≤ n} {next} (x : C n) : (le_rec_on h next x : C n) = x := -by cases n; unfold le_rec_on or.by_cases; rw [dif_neg n.not_succ_le_self, dif_pos rfl] +by cases n; unfold le_rec_on or.by_cases; rw [dif_neg n.not_succ_le_self] theorem le_rec_on_succ {C : ℕ → Sort u} {n m} (h1 : n ≤ m) {h2 : n ≤ m+1} {next} (x : C n) : (le_rec_on h2 @next x : C (m+1)) = next (le_rec_on h1 @next x : C m) := diff --git a/src/data/nat/factorial/basic.lean b/src/data/nat/factorial/basic.lean index 7c5f330db94f8..5760aa0df8186 100644 --- a/src/data/nat/factorial/basic.lean +++ b/src/data/nat/factorial/basic.lean @@ -26,7 +26,7 @@ namespace nat | 0 := 1 | (succ n) := succ n * factorial n -localized "notation n `!`:10000 := nat.factorial n" in nat +localized "notation (name := nat.factorial) n `!`:10000 := nat.factorial n" in nat section factorial diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index e4d214bfdfcda..b1fccbdf88a80 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -27,7 +27,7 @@ namespace nat coprime with `n`. -/ def totient (n : ℕ) : ℕ := ((range n).filter n.coprime).card -localized "notation `φ` := nat.totient" in nat +localized "notation (name := nat.totient) `φ` := nat.totient" in nat @[simp] theorem totient_zero : φ 0 = 0 := rfl diff --git a/src/data/num/bitwise.lean b/src/data/num/bitwise.lean index 7f3d7c9324e8f..63b576b02521a 100644 --- a/src/data/num/bitwise.lean +++ b/src/data/num/bitwise.lean @@ -193,7 +193,7 @@ The `snum` representation uses a bit string, essentially a list of 0 (`ff`) and and the negation of the MSB is sign-extended to all higher bits. -/ namespace nzsnum - notation a :: b := bit a b + notation (name := nznum.bit) a :: b := bit a b /-- Sign of a `nzsnum`. -/ def sign : nzsnum → bool @@ -237,14 +237,14 @@ namespace snum @[pattern] def not : snum → snum | (zero z) := zero (bnot z) | (nz p) := ~p - prefix ~ := not + prefix (name := snum.not) ~ := not /-- Add a bit at the end of a `snum`. This mimics `nzsnum.bit`. -/ @[pattern] def bit : bool → snum → snum | b (zero z) := if b = z then zero b else msb b | b (nz p) := p.bit b - notation a :: b := bit a b + notation (name := snum.bit) a :: b := bit a b /-- Add an inactive bit at the end of a `snum`. This mimics `znum.bit0`. -/ def bit0 : snum → snum := bit ff diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 9a95be5d81596..3517ef332e485 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -56,7 +56,7 @@ The embedding from `R` is called `C`. -/ structure polynomial (R : Type*) [semiring R] := of_finsupp :: (to_finsupp : add_monoid_algebra R ℕ) -localized "notation R`[X]`:9000 := polynomial R" in polynomial +localized "notation (name := polynomial) R`[X]`:9000 := polynomial R" in polynomial open add_monoid_algebra finsupp function open_locale big_operators polynomial diff --git a/src/data/quot.lean b/src/data/quot.lean index d3107d8d2fd5f..3c164299526da 100644 --- a/src/data/quot.lean +++ b/src/data/quot.lean @@ -29,7 +29,7 @@ end setoid namespace quot variables {ra : α → α → Prop} {rb : β → β → Prop} {φ : quot ra → quot rb → Sort*} -local notation `⟦`:max a `⟧` := quot.mk _ a +local notation (name := mk) `⟦`:max a `⟧` := quot.mk _ a instance (r : α → α → Prop) [inhabited α] : inhabited (quot r) := ⟨⟦default⟧⟩ diff --git a/src/data/rat/defs.lean b/src/data/rat/defs.lean index 678af31d05455..7032e48f3b27f 100644 --- a/src/data/rat/defs.lean +++ b/src/data/rat/defs.lean @@ -103,7 +103,7 @@ def mk : ℤ → ℤ → ℚ | n (d : ℕ) := mk_nat n d | n -[1+ d] := mk_pnat (-n) d.succ_pnat -localized "infix ` /. `:70 := rat.mk" in rat +localized "infix (name := rat.mk) ` /. `:70 := rat.mk" in rat theorem mk_pnat_eq (n d h) : mk_pnat n ⟨d, h⟩ = n /. d := by change n /. d with dite _ _ _; simp [ne_of_gt h] diff --git a/src/data/rbtree/insert.lean b/src/data/rbtree/insert.lean index 9103dd71a7b32..22bbd2934e4a9 100644 --- a/src/data/rbtree/insert.lean +++ b/src/data/rbtree/insert.lean @@ -188,7 +188,7 @@ parameters {α : Type u} (lt : α → α → Prop) local attribute [simp] mem balance1_node balance2_node -local infix `∈` := mem lt +local infix (name := mem) `∈` := mem lt lemma mem_balance1_node_of_mem_left {x s} (v) (t : rbnode α) : x ∈ s → x ∈ balance1_node s v t := begin diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 3d0809ea6634b..c518cfdd1cc10 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -81,8 +81,8 @@ variables {α : Type*} {β : Type*} linear_ordered_add_comm_monoid_with_top]] def ennreal := with_top ℝ≥0 -localized "notation `ℝ≥0∞` := ennreal" in ennreal -localized "notation `∞` := (⊤ : ennreal)" in ennreal +localized "notation (name := ennreal) `ℝ≥0∞` := ennreal" in ennreal +localized "notation (name := ennreal.top) `∞` := (⊤ : ennreal)" in ennreal namespace ennreal variables {a b c d : ℝ≥0∞} {r p q : ℝ≥0} diff --git a/src/data/real/golden_ratio.lean b/src/data/real/golden_ratio.lean index e3e6e5209b3ec..f08b613f981a8 100644 --- a/src/data/real/golden_ratio.lean +++ b/src/data/real/golden_ratio.lean @@ -28,8 +28,8 @@ noncomputable theory /-- The conjugate of the golden ratio `ψ := (1 - √5)/2`. -/ @[reducible] def golden_conj := (1 - real.sqrt 5)/2 -localized "notation `φ` := golden_ratio" in real -localized "notation `ψ` := golden_conj" in real +localized "notation (name := golden_ratio) `φ` := golden_ratio" in real +localized "notation (name := golden_conj) `ψ` := golden_conj" in real /-- The inverse of the golden ratio is the opposite of its conjugate. -/ lemma inv_gold : φ⁻¹ = -ψ := diff --git a/src/data/real/hyperreal.lean b/src/data/real/hyperreal.lean index 056a5dc833097..de5b5d7c49da3 100644 --- a/src/data/real/hyperreal.lean +++ b/src/data/real/hyperreal.lean @@ -62,8 +62,8 @@ noncomputable def epsilon : ℝ* := of_seq $ λ n, n⁻¹ /-- A sample infinite hyperreal-/ noncomputable def omega : ℝ* := of_seq coe -localized "notation `ε` := hyperreal.epsilon" in hyperreal -localized "notation `ω` := hyperreal.omega" in hyperreal +localized "notation (name := hyperreal.epsilon) `ε` := hyperreal.epsilon" in hyperreal +localized "notation (name := hyperreal.omega) `ω` := hyperreal.omega" in hyperreal lemma epsilon_eq_inv_omega : ε = ω⁻¹ := rfl diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 9c5ec791b4735..85b7f462c783d 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -63,7 +63,7 @@ open_locale classical big_operators linear_ordered_semiring, ordered_comm_semiring, canonically_ordered_comm_semiring, has_sub, has_ordered_sub, has_div, inhabited]] def nnreal := {r : ℝ // 0 ≤ r} -localized "notation ` ℝ≥0 ` := nnreal" in nnreal +localized "notation (name := nnreal) ` ℝ≥0 ` := nnreal" in nnreal namespace nnreal diff --git a/src/data/real/pi/leibniz.lean b/src/data/real/pi/leibniz.lean index 4a803ae34bf6f..5ba42465a6d9d 100644 --- a/src/data/real/pi/leibniz.lean +++ b/src/data/real/pi/leibniz.lean @@ -11,7 +11,7 @@ namespace real open filter set open_locale classical big_operators topological_space real -local notation `|`x`|` := abs x +local notation (name := abs) `|`x`|` := abs x /-- This theorem establishes **Leibniz's series for `π`**: The alternating sum of the reciprocals of the odd numbers is `π/4`. Note that this is a conditionally rather than absolutely convergent diff --git a/src/data/rel.lean b/src/data/rel.lean index 5cecf7d511ef2..8d8a9a4194989 100644 --- a/src/data/rel.lean +++ b/src/data/rel.lean @@ -62,7 +62,7 @@ lemma dom_inv : r.inv.dom = r.codom := by { ext x y, reflexivity} def comp (r : rel α β) (s : rel β γ) : rel α γ := λ x z, ∃ y, r x y ∧ s y z -local infixr ` ∘ ` := rel.comp +local infixr (name := rel.comp) ` ∘ ` := rel.comp lemma comp_assoc (r : rel α β) (s : rel β γ) (t : rel γ δ) : (r ∘ s) ∘ t = r ∘ s ∘ t := diff --git a/src/data/seq/computation.lean b/src/data/seq/computation.lean index 819eb5174e804..e1324f9536a29 100644 --- a/src/data/seq/computation.lean +++ b/src/data/seq/computation.lean @@ -22,7 +22,7 @@ coinductive computation (α : Type u) : Type u An element of `computation α` is an infinite sequence of `option α` such that if `f n = some a` for some `n` then it is constantly `some a` after that. -/ def computation (α : Type u) : Type u := -{f : stream (option α) // ∀ {n a}, f n = some a → f (n + 1) = some a} +{f : stream (option α) // ∀ {{n a}}, f n = some a → f (n + 1) = some a} namespace computation variables {α : Type u} {β : Type v} {γ : Type w} @@ -53,7 +53,7 @@ def head (c : computation α) : option α := c.1.head /-- `tail c` is the remainder of computation, either `c` if `c = return a` or `c'` if `c = think c'`. -/ def tail (c : computation α) : computation α := -⟨c.1.tail, λ n a, let t := c.2 in t⟩ +⟨c.1.tail, λ n a h, c.2 h⟩ /-- `empty α` is the computation that never returns, an infinite sequence of `think`s. -/ diff --git a/src/data/seq/seq.lean b/src/data/seq/seq.lean index 36f0179570d69..bcb2a44a7315f 100644 --- a/src/data/seq/seq.lean +++ b/src/data/seq/seq.lean @@ -237,7 +237,7 @@ end section bisim variable (R : seq α → seq α → Prop) - local infix ` ~ `:50 := R + local infix (name := R) ` ~ `:50 := R def bisim_o : option (seq1 α) → option (seq1 α) → Prop | none none := true diff --git a/src/data/seq/wseq.lean b/src/data/seq/wseq.lean index 842bfc06f6f78..540787c0c45c5 100644 --- a/src/data/seq/wseq.lean +++ b/src/data/seq/wseq.lean @@ -381,7 +381,7 @@ theorem lift_rel_destruct_iff {R : α → β → Prop} {s : wseq α} {t : wseq intros s t, apply or.inl end⟩⟩ -infix ` ~ `:50 := equiv +infix (name := equiv) ` ~ `:50 := equiv theorem destruct_congr {s t : wseq α} : s ~ t → computation.lift_rel (bisim_o (~)) (destruct s) (destruct t) := diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index ed25f4e8ecb7d..eac3814af5b4a 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -37,7 +37,7 @@ variables {α : Type u} [linear_order α] {a a₁ a₂ b b₁ b₂ c x : α} /-- `interval a b` is the set of elements lying between `a` and `b`, with `a` and `b` included. -/ def interval (a b : α) := Icc (min a b) (max a b) -localized "notation `[`a `, ` b `]` := set.interval a b" in interval +localized "notation (name := set.interval) `[`a `, ` b `]` := set.interval a b" in interval @[simp] lemma interval_of_le (h : a ≤ b) : [a, b] = Icc a b := by rw [interval, min_eq_left h, max_eq_right h] diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index 0e2aa8f6d631d..9bdeb92e401a6 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -32,7 +32,7 @@ variables {α β γ δ : Type*} {s s₁ s₂ : set α} {t t₁ t₂ : set β} {a def prod (s : set α) (t : set β) : set (α × β) := {p | p.1 ∈ s ∧ p.2 ∈ t} /- This notation binds more strongly than (pre)images, unions and intersections. -/ -infixr ` ×ˢ `:82 := set.prod +infixr (name := set.prod) ` ×ˢ `:82 := set.prod lemma prod_eq (s : set α) (t : set β) : s ×ˢ t = prod.fst ⁻¹' s ∩ prod.snd ⁻¹' t := rfl diff --git a/src/data/stream/defs.lean b/src/data/stream/defs.lean index 67ae6bcd57d83..874a07eb4fab0 100644 --- a/src/data/stream/defs.lean +++ b/src/data/stream/defs.lean @@ -26,7 +26,7 @@ def cons (a : α) (s : stream α) : stream α | 0 := a | (n + 1) := s n -notation h :: t := cons h t +notation (name := stream.cons) h :: t := cons h t /-- Head of a stream: `stream.head s = stream.nth 0 s`. -/ def head (s : stream α) : α := diff --git a/src/data/sym/basic.lean b/src/data/sym/basic.lean index 6401ec58408ae..3ea65790af6e9 100644 --- a/src/data/sym/basic.lean +++ b/src/data/sym/basic.lean @@ -162,7 +162,7 @@ This is `cons` but for the alternative `sym'` definition. def cons' {α : Type*} {n : ℕ} : α → sym' α n → sym' α (nat.succ n) := λ a, quotient.map (vector.cons a) (λ ⟨l₁, h₁⟩ ⟨l₂, h₂⟩ h, list.perm.cons _ h) -notation a :: b := cons' a b +notation (name := sym.cons') a :: b := cons' a b /-- Multisets of cardinality n are equivalent to length-n vectors up to permutations. diff --git a/src/data/typevec.lean b/src/data/typevec.lean index 778f801331ba4..70b5a41af691f 100644 --- a/src/data/typevec.lean +++ b/src/data/typevec.lean @@ -47,7 +47,7 @@ variable {n : ℕ} /-- arrow in the category of `typevec` -/ def arrow (α β : typevec n) := Π i : fin2 n, α i → β i -localized "infixl ` ⟹ `:40 := typevec.arrow" in mvfunctor +localized "infixl (name := typevec.arrow) ` ⟹ `:40 := typevec.arrow" in mvfunctor instance arrow.inhabited (α β : typevec n) [Π i, inhabited (β i)] : inhabited (α ⟹ β) := ⟨ λ _ _, default ⟩ @@ -59,7 +59,7 @@ def id {α : typevec n} : α ⟹ α := λ i x, x def comp {α β γ : typevec n} (g : β ⟹ γ) (f : α ⟹ β) : α ⟹ γ := λ i x, g i (f i x) -localized "infixr ` ⊚ `:80 := typevec.comp" in mvfunctor -- type as \oo +localized "infixr (name := typevec.comp) ` ⊚ `:80 := typevec.comp" in mvfunctor -- type as \oo @[simp] theorem id_comp {α β : typevec n} (f : α ⟹ β) : id ⊚ f = f := rfl @@ -77,7 +77,7 @@ def append1 (α : typevec n) (β : Type*) : typevec (n+1) | (fin2.fs i) := α i | fin2.fz := β -infixl ` ::: `:67 := append1 +infixl (name := typevec.append1) ` ::: `:67 := append1 /-- retain only a `n-length` prefix of the argument -/ def drop (α : typevec.{u} (n+1)) : typevec n := λ i, α i.fs @@ -120,7 +120,7 @@ and target types / typevecs -/ def append_fun {α α' : typevec n} {β β' : Type*} (f : α ⟹ α') (g : β → β') : append1 α β ⟹ append1 α' β' := split_fun f g -infixl ` ::: ` := append_fun +infixl (name := typevec.append_fun) ` ::: ` := append_fun /-- split off the prefix of an arrow -/ def drop_fun {α β : typevec (n+1)} (f : α ⟹ β) : drop α ⟹ drop β := @@ -334,7 +334,7 @@ def prod : Π {n} (α β : typevec.{u} n), typevec n | 0 α β := fin2.elim0 | (n+1) α β := prod (drop α) (drop β) ::: (last α × last β) -localized "infix ` ⊗ `:45 := typevec.prod" in mvfunctor +localized "infix (name := typevec.prod) ` ⊗ `:45 := typevec.prod" in mvfunctor /-- `const x α` is an arrow that ignores its source and constructs a `typevec` that contains nothing but `x` -/ @@ -442,7 +442,7 @@ protected def prod.map : Π {n} {α α' β β' : typevec.{u} n}, (α ⟹ β) → @prod.map _ (drop α) (drop α') (drop β) (drop β') (drop_fun x) (drop_fun y) _ a | (succ n) α α' β β' x y fin2.fz a := (x _ a.1,y _ a.2) -localized "infix ` ⊗' `:45 := typevec.prod.map" in mvfunctor +localized "infix (name := typevec.prod.map) ` ⊗' `:45 := typevec.prod.map" in mvfunctor theorem fst_prod_mk {α α' β β' : typevec n} (f : α ⟹ β) (g : α' ⟹ β') : typevec.prod.fst ⊚ (f ⊗' g) = f ⊚ typevec.prod.fst := diff --git a/src/data/vector3.lean b/src/data/vector3.lean index 823a146c1225f..925a9cb462d77 100644 --- a/src/data/vector3.lean +++ b/src/data/vector3.lean @@ -37,8 +37,9 @@ namespace vector3 /- We do not want to make the following notation global, because then these expressions will be overloaded, and only the expected type will be able to disambiguate the meaning. Worse: Lean will try to insert a coercion from `vector3 α _` to `list α`, if a list is expected. -/ -localized "notation `[` l:(foldr `, ` (h t, vector3.cons h t) vector3.nil `]`) := l" in vector3 -notation a :: b := cons a b +localized "notation (name := vector.list) + `[` l:(foldr `, ` (h t, vector3.cons h t) vector3.nil `]`) := l" in vector3 +notation (name := vector.cons) a :: b := cons a b @[simp] lemma cons_fz (a : α) (v : vector3 α n) : (a :: v) fz = a := rfl @[simp] lemma cons_fs (a : α) (v : vector3 α n) (i) : (a :: v) (fs i) = v i := rfl diff --git a/src/deprecated/ring.lean b/src/deprecated/ring.lean index 7ef81059d0450..9c9cbdda46206 100644 --- a/src/deprecated/ring.lean +++ b/src/deprecated/ring.lean @@ -107,7 +107,8 @@ lemma comp (hf : is_ring_hom f) {γ} [ring γ] {g : β → γ} (hg : is_ring_hom lemma to_is_semiring_hom (hf : is_ring_hom f) : is_semiring_hom f := { map_zero := map_zero hf, ..‹is_ring_hom f› } -lemma to_is_add_group_hom (hf : is_ring_hom f) : is_add_group_hom f := { map_add := hf.map_add } +lemma to_is_add_group_hom (hf : is_ring_hom f) : is_add_group_hom f := +{ map_add := λ _ _, hf.map_add } end is_ring_hom diff --git a/src/deprecated/subgroup.lean b/src/deprecated/subgroup.lean index 161b407ac9cc5..2cd64ba456c5d 100644 --- a/src/deprecated/subgroup.lean +++ b/src/deprecated/subgroup.lean @@ -51,8 +51,8 @@ by simpa only [div_eq_mul_inv] using hs.mul_mem hx (hs.inv_mem hy) lemma additive.is_add_subgroup {s : set G} (hs : is_subgroup s) : @is_add_subgroup (additive G) _ s := -@is_add_subgroup.mk (additive G) _ _ (additive.is_add_submonoid hs.to_is_submonoid) - hs.inv_mem +@is_add_subgroup.mk (additive G) _ _ (additive.is_add_submonoid hs.to_is_submonoid) $ + λ _, hs.inv_mem theorem additive.is_add_subgroup_iff {s : set G} : @is_add_subgroup (additive G) _ s ↔ is_subgroup s := @@ -61,8 +61,8 @@ theorem additive.is_add_subgroup_iff lemma multiplicative.is_subgroup {s : set A} (hs : is_add_subgroup s) : @is_subgroup (multiplicative A) _ s := -@is_subgroup.mk (multiplicative A) _ _ (multiplicative.is_submonoid hs.to_is_add_submonoid) - hs.neg_mem +@is_subgroup.mk (multiplicative A) _ _ (multiplicative.is_submonoid hs.to_is_add_submonoid) $ + λ _, hs.neg_mem theorem multiplicative.is_subgroup_iff {s : set A} : @is_subgroup (multiplicative A) _ s ↔ is_add_subgroup s := @@ -486,7 +486,7 @@ theorem closure_eq_mclosure {s : set G} : closure s = monoid.closure (s ∪ has_ set.subset.antisymm (@closure_subset _ _ _ (monoid.closure (s ∪ has_inv.inv ⁻¹' s)) { one_mem := (monoid.closure.is_submonoid _).one_mem, - mul_mem := (monoid.closure.is_submonoid _).mul_mem, + mul_mem := λ _ _, (monoid.closure.is_submonoid _).mul_mem, inv_mem := λ x hx, monoid.in_closure.rec_on hx (λ x hx, or.cases_on hx (λ hx, monoid.subset_closure $ or.inr $ show x⁻¹⁻¹ ∈ s, from (inv_inv x).symm ▸ hx) @@ -597,14 +597,14 @@ end group def subgroup.of [group G] {s : set G} (h : is_subgroup s) : subgroup G := { carrier := s, one_mem' := h.1.1, - mul_mem' := h.1.2, - inv_mem' := h.2 } + mul_mem' := λ _ _, h.1.2, + inv_mem' := λ _, h.2 } @[to_additive] lemma subgroup.is_subgroup [group G] (K : subgroup G) : is_subgroup (K : set G) := { one_mem := K.one_mem', - mul_mem := K.mul_mem', - inv_mem := K.inv_mem' } + mul_mem := λ _ _, K.mul_mem', + inv_mem := λ _, K.inv_mem' } -- this will never fire if it's an instance @[to_additive] diff --git a/src/deprecated/submonoid.lean b/src/deprecated/submonoid.lean index 80c787b1e988f..7750164c2a044 100644 --- a/src/deprecated/submonoid.lean +++ b/src/deprecated/submonoid.lean @@ -320,7 +320,7 @@ end monoid /-- Create a bundled submonoid from a set `s` and `[is_submonoid s]`. -/ @[to_additive "Create a bundled additive submonoid from a set `s` and `[is_add_submonoid s]`."] -def submonoid.of {s : set M} (h : is_submonoid s) : submonoid M := ⟨s, h.2, h.1⟩ +def submonoid.of {s : set M} (h : is_submonoid s) : submonoid M := ⟨s, λ _ _, h.2, h.1⟩ @[to_additive] -lemma submonoid.is_submonoid (S : submonoid M) : is_submonoid (S : set M) := ⟨S.3, S.2⟩ +lemma submonoid.is_submonoid (S : submonoid M) : is_submonoid (S : set M) := ⟨S.3, λ _ _, S.2⟩ diff --git a/src/deprecated/subring.lean b/src/deprecated/subring.lean index 7e21b25f66011..3c88103c3f8b8 100644 --- a/src/deprecated/subring.lean +++ b/src/deprecated/subring.lean @@ -39,10 +39,10 @@ structure is_subring (S : set R) extends is_add_subgroup S, is_submonoid S : Pro def is_subring.subring {S : set R} (hs : is_subring S) : subring R := { carrier := S, one_mem' := hs.one_mem, - mul_mem' := hs.mul_mem, + mul_mem' := λ _ _, hs.mul_mem, zero_mem' := hs.zero_mem, - add_mem' := hs.add_mem, - neg_mem' := hs.neg_mem } + add_mem' := λ _ _, hs.add_mem, + neg_mem' := λ _, hs.neg_mem } namespace ring_hom diff --git a/src/dynamics/omega_limit.lean b/src/dynamics/omega_limit.lean index 96876df8dda8c..ccdbced232442 100644 --- a/src/dynamics/omega_limit.lean +++ b/src/dynamics/omega_limit.lean @@ -45,10 +45,10 @@ variables {τ : Type*} {α : Type*} {β : Type*} {ι : Type*} def omega_limit [topological_space β] (f : filter τ) (ϕ : τ → α → β) (s : set α) : set β := ⋂ u ∈ f, closure (image2 ϕ u s) -localized "notation `ω` := omega_limit" in omega_limit +localized "notation (name := omega_limit) `ω` := omega_limit" in omega_limit -localized "notation `ω⁺` := omega_limit filter.at_top" in omega_limit -localized "notation `ω⁻` := omega_limit filter.at_bot" in omega_limit +localized "notation (name := omega_limit.at_top) `ω⁺` := omega_limit filter.at_top" in omega_limit +localized "notation (name := omega_limit.at_bot) `ω⁻` := omega_limit filter.at_bot" in omega_limit variables [topological_space β] variables (f : filter τ) (ϕ : τ → α → β) (s s₁ s₂: set α) diff --git a/src/field_theory/intermediate_field.lean b/src/field_theory/intermediate_field.lean index 020f86eed9569..a7bb7bfd6fbd1 100644 --- a/src/field_theory/intermediate_field.lean +++ b/src/field_theory/intermediate_field.lean @@ -61,10 +61,10 @@ instance : set_like (intermediate_field K L) L := ⟨λ S, S.to_subalgebra.carrier, by { rintros ⟨⟨⟩⟩ ⟨⟨⟩⟩ ⟨h⟩, congr, }⟩ instance : subfield_class (intermediate_field K L) L := -{ add_mem := λ s, s.add_mem', +{ add_mem := λ s _ _, s.add_mem', zero_mem := λ s, s.zero_mem', neg_mem := neg_mem', - mul_mem := λ s, s.mul_mem', + mul_mem := λ s _ _, s.mul_mem', one_mem := λ s, s.one_mem', inv_mem := inv_mem' } diff --git a/src/field_theory/krull_topology.lean b/src/field_theory/krull_topology.lean index 74bea4b0c9ccb..2dbbace10155d 100644 --- a/src/field_theory/krull_topology.lean +++ b/src/field_theory/krull_topology.lean @@ -163,7 +163,7 @@ def gal_group_basis (K L : Type*) [field K] [field L] [algebra K L] : end⟩, inv' := λ U hU, ⟨U, hU, begin rcases hU with ⟨H, hH, rfl⟩, - exact H.inv_mem', + exact λ _, H.inv_mem', end⟩, conj' := begin @@ -176,13 +176,13 @@ def gal_group_basis (K L : Type*) [field K] [field L] [algebra K L] : change σ * g * σ⁻¹ ∈ E.fixing_subgroup, rw intermediate_field.mem_fixing_subgroup_iff, intros x hx, - change σ(g(σ⁻¹ x)) = x, + change σ (g (σ⁻¹ x)) = x, have h_in_F : σ⁻¹ x ∈ F := ⟨x, hx, by {dsimp, rw ← alg_equiv.inv_fun_eq_symm, refl }⟩, - have h_g_fix : g (σ⁻¹ x) = (σ⁻¹ x), + have h_g_fix : g (σ⁻¹ x) = σ⁻¹ x, { rw [subgroup.mem_carrier, intermediate_field.mem_fixing_subgroup_iff F g] at hg, exact hg (σ⁻¹ x) h_in_F }, rw h_g_fix, - change σ(σ⁻¹ x) = x, + change σ (σ⁻¹ x) = x, exact alg_equiv.apply_symm_apply σ x, end } diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index d67d199e4960f..d4e407a43b7b5 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -394,7 +394,7 @@ include V notation. -/ def angle (p1 p2 p3 : P) : ℝ := angle (p1 -ᵥ p2 : V) (p3 -ᵥ p2) -localized "notation `∠` := euclidean_geometry.angle" in euclidean_geometry +localized "notation (name := angle) `∠` := euclidean_geometry.angle" in euclidean_geometry lemma continuous_at_angle {x : P × P × P} (hx12 : x.1 ≠ x.2.1) (hx32 : x.2.2 ≠ x.2.1) : continuous_at (λ y : P × P × P, ∠ y.1 y.2.1 y.2.2) x := diff --git a/src/geometry/manifold/algebra/monoid.lean b/src/geometry/manifold/algebra/monoid.lean index 1daf60fb0b289..64c5c72065b9b 100644 --- a/src/geometry/manifold/algebra/monoid.lean +++ b/src/geometry/manifold/algebra/monoid.lean @@ -156,10 +156,10 @@ names. -/ def smooth_right_mul : C^∞⟮I, G; I, G⟯ := ⟨(right_mul g), smooth_mul_right⟩ /- Left multiplication. The abbreviation is `MIL`. -/ -localized "notation `𝑳` := smooth_left_mul" in lie_group +localized "notation (name := smooth_left_mul) `𝑳` := smooth_left_mul" in lie_group /- Right multiplication. The abbreviation is `MIR`. -/ -localized "notation `𝑹` := smooth_right_mul" in lie_group +localized "notation (name := smooth_right_mul) `𝑹` := smooth_right_mul" in lie_group open_locale lie_group diff --git a/src/geometry/manifold/charted_space.lean b/src/geometry/manifold/charted_space.lean index d7733b339616e..9e7705a5ed866 100644 --- a/src/geometry/manifold/charted_space.lean +++ b/src/geometry/manifold/charted_space.lean @@ -118,8 +118,10 @@ variables {H : Type u} {H' : Type*} {M : Type*} {M' : Type*} {M'' : Type*} `local_homeomorph.trans` and `local_equiv.trans`. Note that, as is usual for equivs, the composition is from left to right, hence the direction of the arrow. -/ -localized "infixr ` ≫ₕ `:100 := local_homeomorph.trans" in manifold -localized "infixr ` ≫ `:100 := local_equiv.trans" in manifold +localized "infixr (name := local_homeomorph.trans) + ` ≫ₕ `:100 := local_homeomorph.trans" in manifold +localized "infixr (name := local_equiv.trans) + ` ≫ `:100 := local_equiv.trans" in manifold open set local_homeomorph @@ -774,7 +776,7 @@ has_groupoid.compatible G he he' lemma has_groupoid_of_le {G₁ G₂ : structure_groupoid H} (h : has_groupoid M G₁) (hle : G₁ ≤ G₂) : has_groupoid M G₂ := -⟨ λ e e' he he', hle ((h.compatible : _) he he') ⟩ +⟨λ e e' he he', hle (h.compatible he he')⟩ lemma has_groupoid_of_pregroupoid (PG : pregroupoid H) (h : ∀{e e' : local_homeomorph M H}, e ∈ atlas H M → e' ∈ atlas H M diff --git a/src/geometry/manifold/cont_mdiff_map.lean b/src/geometry/manifold/cont_mdiff_map.lean index 69f5cf0955d30..3f168262d5c0c 100644 --- a/src/geometry/manifold/cont_mdiff_map.lean +++ b/src/geometry/manifold/cont_mdiff_map.lean @@ -37,9 +37,9 @@ structure cont_mdiff_map := /-- Bundled smooth maps. -/ @[reducible] def smooth_map := cont_mdiff_map I I' M M' ⊤ -localized "notation `C^` n `⟮` I `, ` M `; ` I' `, ` M' `⟯` := +localized "notation (name := cont_mdiff_map) `C^` n `⟮` I `, ` M `; ` I' `, ` M' `⟯` := cont_mdiff_map I I' M M' n" in manifold -localized "notation `C^` n `⟮` I `, ` M `; ` k `⟯` := +localized "notation (name := cont_mdiff_map.self) `C^` n `⟮` I `, ` M `; ` k `⟯` := cont_mdiff_map I (model_with_corners_self k k) M k n" in manifold open_locale manifold diff --git a/src/geometry/manifold/derivation_bundle.lean b/src/geometry/manifold/derivation_bundle.lean index e110fe02e4ee6..0bc52955e9860 100644 --- a/src/geometry/manifold/derivation_bundle.lean +++ b/src/geometry/manifold/derivation_bundle.lean @@ -35,7 +35,7 @@ instance smooth_functions_tower : is_scalar_tower 𝕜 C^∞⟮I, M; 𝕜⟯ C^ which is defined as `f • r = f(x) * r`. -/ @[nolint unused_arguments] def pointed_smooth_map (x : M) := C^n⟮I, M; 𝕜⟯ -localized "notation `C^` n `⟮` I `,` M `;` 𝕜 `⟯⟨` x `⟩` := +localized "notation (name := pointed_smooth_map) `C^` n `⟮` I `,` M `;` 𝕜 `⟯⟨` x `⟩` := pointed_smooth_map 𝕜 I M n x" in derivation variables {𝕜 M} @@ -120,10 +120,10 @@ def fdifferential (f : C^∞⟮I, M; I', M'⟯) (x : M) : hfdifferential (rfl : f x = f x) /- Standard notation for the differential. The abbreviation is `MId`. -/ -localized "notation `𝒅` := fdifferential" in manifold +localized "notation (name := fdifferential) `𝒅` := fdifferential" in manifold /- Standard notation for the differential. The abbreviation is `MId`. -/ -localized "notation `𝒅ₕ` := hfdifferential" in manifold +localized "notation (name := hfdifferential) `𝒅ₕ` := hfdifferential" in manifold @[simp] lemma apply_fdifferential (f : C^∞⟮I, M; I', M'⟯) {x : M} (v : point_derivation I x) (g : C^∞⟮I', M'; 𝕜⟯) : 𝒅f x v g = v (g.comp f) := rfl diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index e452c81399888..a9968eaaf5ffc 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -77,11 +77,13 @@ structure diffeomorph extends M ≃ M' := end defs -localized "notation M ` ≃ₘ^` n:1000 `⟮`:50 I `,` J `⟯ ` N := diffeomorph I J M N n" in manifold -localized "notation M ` ≃ₘ⟮` I `,` J `⟯ ` N := diffeomorph I J M N ⊤" in manifold -localized "notation E ` ≃ₘ^` n:1000 `[`:50 𝕜 `] ` E' := +localized "notation (name := diffeomorph) M ` ≃ₘ^` n:1000 `⟮`:50 I `,` J `⟯ ` N := + diffeomorph I J M N n" in manifold +localized "notation (name := diffeomorph.top) M ` ≃ₘ⟮` I `,` J `⟯ ` N := + diffeomorph I J M N ⊤" in manifold +localized "notation (name := diffeomorph.self) E ` ≃ₘ^` n:1000 `[`:50 𝕜 `] ` E' := diffeomorph (model_with_corners_self 𝕜 E) (model_with_corners_self 𝕜 E') E E' n" in manifold -localized "notation E ` ≃ₘ[` 𝕜 `] ` E' := +localized "notation (name := diffeomorph.self.top) E ` ≃ₘ[` 𝕜 `] ` E' := diffeomorph (model_with_corners_self 𝕜 E) (model_with_corners_self 𝕜 E') E E' ⊤" in manifold namespace diffeomorph diff --git a/src/geometry/manifold/instances/real.lean b/src/geometry/manifold/instances/real.lean index eb90335fedeeb..4c835a09792d2 100644 --- a/src/geometry/manifold/instances/real.lean +++ b/src/geometry/manifold/instances/real.lean @@ -123,10 +123,10 @@ def model_with_corners_euclidean_quadrant (n : ℕ) : continuous_inv_fun := continuous.subtype_mk (continuous_pi $ λ i, (continuous_id.max continuous_const).comp (continuous_apply i)) _ } -localized "notation `𝓡 `n := +localized "notation (name := model_with_corners_self.euclidean) `𝓡 `n := (model_with_corners_self ℝ (euclidean_space ℝ (fin n)) : model_with_corners ℝ (euclidean_space ℝ (fin n)) (euclidean_space ℝ (fin n)))" in manifold -localized "notation `𝓡∂ `n := +localized "notation (name := model_with_corners_euclidean_half_space.euclidean) `𝓡∂ `n := (model_with_corners_euclidean_half_space n : model_with_corners ℝ (euclidean_space ℝ (fin n)) (euclidean_half_space n))" in manifold diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index 2c8fff7140c7d..a6c6ba9631dab 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -116,7 +116,7 @@ universes u v w u' v' w' open set filter function open_locale manifold filter topological_space -localized "notation `∞` := (⊤ : ℕ∞)" in manifold +localized "notation (name := with_top.nat.top) `∞` := (⊤ : ℕ∞)" in manifold /-! ### Models with corners. -/ @@ -144,9 +144,11 @@ def model_with_corners_self (𝕜 : Type*) [nontrivially_normed_field 𝕜] continuous_to_fun := continuous_id, continuous_inv_fun := continuous_id } -localized "notation `𝓘(` 𝕜 `, ` E `)` := model_with_corners_self 𝕜 E" in manifold +localized "notation (name := model_with_corners_self) `𝓘(` 𝕜 `, ` E `)` := + model_with_corners_self 𝕜 E" in manifold -localized "notation `𝓘(` 𝕜 `)` := model_with_corners_self 𝕜 𝕜" in manifold +localized "notation (name := model_with_corners_self.self) `𝓘(` 𝕜 `)` := + model_with_corners_self 𝕜 𝕜" in manifold section variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] diff --git a/src/group_theory/coset.lean b/src/group_theory/coset.lean index d3c3cdb98fa7d..c0441368fb8ee 100644 --- a/src/group_theory/coset.lean +++ b/src/group_theory/coset.lean @@ -54,10 +54,10 @@ def left_coset [has_mul α] (a : α) (s : set α) : set α := (λ x, a * x) '' s and a subset `s : set α`"] def right_coset [has_mul α] (s : set α) (a : α) : set α := (λ x, x * a) '' s -localized "infix ` *l `:70 := left_coset" in coset -localized "infix ` +l `:70 := left_add_coset" in coset -localized "infix ` *r `:70 := right_coset" in coset -localized "infix ` +r `:70 := right_add_coset" in coset +localized "infix (name := left_coset) ` *l `:70 := left_coset" in coset +localized "infix (name := left_add_coset) ` +l `:70 := left_add_coset" in coset +localized "infix (name := right_coset) ` *r `:70 := right_coset" in coset +localized "infix (name := right_add_coset) ` +r `:70 := right_add_coset" in coset section coset_mul variable [has_mul α] diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 7e18b8d0da9f7..2a263e1b0cf51 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -373,10 +373,10 @@ section mul_add @[simps] def subgroup.to_add_subgroup : subgroup G ≃o add_subgroup (additive G) := { to_fun := λ S, - { neg_mem' := S.inv_mem', + { neg_mem' := λ _, S.inv_mem', ..S.to_submonoid.to_add_submonoid }, inv_fun := λ S, - { inv_mem' := S.neg_mem', + { inv_mem' := λ _, S.neg_mem', ..S.to_add_submonoid.to_submonoid' }, left_inv := λ x, by cases x; refl, right_inv := λ x, by cases x; refl, @@ -391,10 +391,10 @@ subgroup.to_add_subgroup.symm @[simps] def add_subgroup.to_subgroup : add_subgroup A ≃o subgroup (multiplicative A) := { to_fun := λ S, - { inv_mem' := S.neg_mem', + { inv_mem' := λ _, S.neg_mem', ..S.to_add_submonoid.to_submonoid }, inv_fun := λ S, - { neg_mem' := S.inv_mem', + { neg_mem' := λ _, S.inv_mem', ..S.to_submonoid.to_add_submonoid' }, left_inv := λ x, by cases x; refl, right_inv := λ x, by cases x; refl, @@ -418,8 +418,8 @@ Useful to fix definitional equalities"] protected def copy (K : subgroup G) (s : set G) (hs : s = K) : subgroup G := { carrier := s, one_mem' := hs.symm ▸ K.one_mem', - mul_mem' := hs.symm ▸ K.mul_mem', - inv_mem' := hs.symm ▸ K.inv_mem' } + mul_mem' := λ _ _, hs.symm ▸ K.mul_mem', + inv_mem' := λ _, hs.symm ▸ K.inv_mem' } @[simp, to_additive] lemma coe_copy (K : subgroup G) (s : set G) (hs : s = ↑K) : (K.copy s hs : set G) = s := rfl diff --git a/src/group_theory/submonoid/basic.lean b/src/group_theory/submonoid/basic.lean index 99493198c8db4..442244b297116 100644 --- a/src/group_theory/submonoid/basic.lean +++ b/src/group_theory/submonoid/basic.lean @@ -168,7 +168,7 @@ theorem ext {S T : submonoid M} protected def copy (S : submonoid M) (s : set M) (hs : s = S) : submonoid M := { carrier := s, one_mem' := hs.symm ▸ S.one_mem', - mul_mem' := hs.symm ▸ S.mul_mem' } + mul_mem' := λ _ _, hs.symm ▸ S.mul_mem' } variable {S : submonoid M} diff --git a/src/group_theory/submonoid/operations.lean b/src/group_theory/submonoid/operations.lean index 3179802f3c295..2aa27c94f47dc 100644 --- a/src/group_theory/submonoid/operations.lean +++ b/src/group_theory/submonoid/operations.lean @@ -75,11 +75,11 @@ def submonoid.to_add_submonoid : submonoid M ≃o add_submonoid (additive M) := { to_fun := λ S, { carrier := additive.to_mul ⁻¹' S, zero_mem' := S.one_mem', - add_mem' := S.mul_mem' }, + add_mem' := λ _ _, S.mul_mem' }, inv_fun := λ S, { carrier := additive.of_mul ⁻¹' S, one_mem' := S.zero_mem', - mul_mem' := S.add_mem' }, + mul_mem' := λ _ _, S.add_mem' }, left_inv := λ x, by cases x; refl, right_inv := λ x, by cases x; refl, map_rel_iff' := λ a b, iff.rfl, } @@ -115,11 +115,11 @@ def add_submonoid.to_submonoid : add_submonoid A ≃o submonoid (multiplicative { to_fun := λ S, { carrier := multiplicative.to_add ⁻¹' S, one_mem' := S.zero_mem', - mul_mem' := S.add_mem' }, + mul_mem' := λ _ _, S.add_mem' }, inv_fun := λ S, { carrier := multiplicative.of_add ⁻¹' S, zero_mem' := S.one_mem', - add_mem' := S.mul_mem' }, + add_mem' := λ _ _, S.mul_mem' }, left_inv := λ x, by cases x; refl, right_inv := λ x, by cases x; refl, map_rel_iff' := λ a b, iff.rfl, } diff --git a/src/group_theory/subsemigroup/basic.lean b/src/group_theory/subsemigroup/basic.lean index a454a84451711..82c6b7322538f 100644 --- a/src/group_theory/subsemigroup/basic.lean +++ b/src/group_theory/subsemigroup/basic.lean @@ -118,7 +118,7 @@ theorem ext {S T : subsemigroup M} it."] protected def copy (S : subsemigroup M) (s : set M) (hs : s = S) : subsemigroup M := { carrier := s, - mul_mem' := hs.symm ▸ S.mul_mem' } + mul_mem' := λ _ _, hs.symm ▸ S.mul_mem' } variable {S : subsemigroup M} diff --git a/src/group_theory/subsemigroup/operations.lean b/src/group_theory/subsemigroup/operations.lean index f7d869856605f..db85872d922ed 100644 --- a/src/group_theory/subsemigroup/operations.lean +++ b/src/group_theory/subsemigroup/operations.lean @@ -75,10 +75,10 @@ variables [has_mul M] def subsemigroup.to_add_subsemigroup : subsemigroup M ≃o add_subsemigroup (additive M) := { to_fun := λ S, { carrier := additive.to_mul ⁻¹' S, - add_mem' := S.mul_mem' }, + add_mem' := λ _ _, S.mul_mem' }, inv_fun := λ S, { carrier := additive.of_mul ⁻¹' S, - mul_mem' := S.add_mem' }, + mul_mem' := λ _ _, S.add_mem' }, left_inv := λ x, by cases x; refl, right_inv := λ x, by cases x; refl, map_rel_iff' := λ a b, iff.rfl, } @@ -115,10 +115,10 @@ multiplicative subsemigroups of `multiplicative A`. -/ def add_subsemigroup.to_subsemigroup : add_subsemigroup A ≃o subsemigroup (multiplicative A) := { to_fun := λ S, { carrier := multiplicative.to_add ⁻¹' S, - mul_mem' := S.add_mem' }, + mul_mem' := λ _ _, S.add_mem' }, inv_fun := λ S, { carrier := multiplicative.of_add ⁻¹' S, - add_mem' := S.mul_mem' }, + add_mem' := λ _ _, S.mul_mem' }, left_inv := λ x, by cases x; refl, right_inv := λ x, by cases x; refl, map_rel_iff' := λ a b, iff.rfl, } diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index 304130b9e3113..279e1f5a59608 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -68,9 +68,9 @@ instance : set_like (sylow p G) G := coe_injective' := λ P Q h, ext (set_like.coe_injective h) } instance : subgroup_class (sylow p G) G := -{ mul_mem := λ s, s.mul_mem', +{ mul_mem := λ s _ _, s.mul_mem', one_mem := λ s, s.one_mem', - inv_mem := λ s, s.inv_mem' } + inv_mem := λ s _, s.inv_mem' } variables (P : sylow p G) diff --git a/src/linear_algebra/affine_space/basic.lean b/src/linear_algebra/affine_space/basic.lean index dc2f712962af6..9df96f95457ad 100644 --- a/src/linear_algebra/affine_space/basic.lean +++ b/src/linear_algebra/affine_space/basic.lean @@ -37,4 +37,4 @@ Some key definitions are not yet present. coordinates, with appropriate proofs of existence when `k` is a field. -/ -localized "notation `affine_space` := add_torsor" in affine +localized "notation (name := add_torsor) `affine_space` := add_torsor" in affine diff --git a/src/linear_algebra/cross_product.lean b/src/linear_algebra/cross_product.lean index 66e126dd97a93..91eac76bb2ce3 100644 --- a/src/linear_algebra/cross_product.lean +++ b/src/linear_algebra/cross_product.lean @@ -58,7 +58,7 @@ begin simp [smul_vec3 (_ : R) (_ : R), mul_comm, mul_assoc, mul_left_comm, mul_add, sub_eq_add_neg] }, end -localized "infixl ` ×₃ `: 74 := cross_product" in matrix +localized "infixl (name := cross_product) ` ×₃ `: 74 := cross_product" in matrix lemma cross_apply (a b : fin 3 → R) : a ×₃ b = ![a 1 * b 2 - a 2 * b 1, diff --git a/src/linear_algebra/pi_tensor_product.lean b/src/linear_algebra/pi_tensor_product.lean index a637590fc86ca..451d7d9408fbd 100644 --- a/src/linear_algebra/pi_tensor_product.lean +++ b/src/linear_algebra/pi_tensor_product.lean @@ -100,8 +100,8 @@ def pi_tensor_product : Type* := variables {R} /- This enables the notation `⨂[R] i : ι, s i` for the pi tensor product, given `s : ι → Type*`. -/ -localized "notation `⨂[`:100 R `] ` binders `, ` r:(scoped:67 f, pi_tensor_product R f) := r" - in tensor_product +localized "notation (name := pi_tensor_product) + `⨂[`:100 R `] ` binders `, ` r:(scoped:67 f, pi_tensor_product R f) := r" in tensor_product open_locale tensor_product diff --git a/src/linear_algebra/special_linear_group.lean b/src/linear_algebra/special_linear_group.lean index 8db2dd55e160b..f085db85e748c 100644 --- a/src/linear_algebra/special_linear_group.lean +++ b/src/linear_algebra/special_linear_group.lean @@ -62,7 +62,8 @@ def special_linear_group := { A : matrix n n R // A.det = 1 } end -localized "notation `SL(` n `,` R `)`:= matrix.special_linear_group (fin n) R" in matrix_groups +localized "notation (name := special_linear_group.fin) + `SL(` n `,` R `)`:= matrix.special_linear_group (fin n) R" in matrix_groups namespace special_linear_group diff --git a/src/linear_algebra/tensor_power.lean b/src/linear_algebra/tensor_power.lean index f200c4298cd2d..360f289e3870e 100644 --- a/src/linear_algebra/tensor_power.lean +++ b/src/linear_algebra/tensor_power.lean @@ -43,8 +43,8 @@ open_locale tensor_product variables {R : Type*} {M : Type*} [comm_semiring R] [add_comm_monoid M] [module R M] -localized "notation `⨂[`:100 R `]^`:80 n:max := tensor_power R n" - in tensor_product +localized "notation (name := tensor_power) + `⨂[`:100 R `]^`:80 n:max := tensor_power R n" in tensor_product namespace tensor_power open_locale tensor_product direct_sum diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index 6b89a08161cca..c5f504961e7b4 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -78,8 +78,10 @@ def tensor_product : Type* := variables {R} -localized "infix ` ⊗ `:100 := tensor_product _" in tensor_product -localized "notation M ` ⊗[`:100 R `] `:0 N:100 := tensor_product R M N" in tensor_product +localized "infix (name := tensor_product.infer) + ` ⊗ `:100 := tensor_product hole!" in tensor_product +localized "notation (name := tensor_product) + M ` ⊗[`:100 R `] `:0 N:100 := tensor_product R M N" in tensor_product namespace tensor_product diff --git a/src/logic/equiv/local_equiv.lean b/src/logic/equiv/local_equiv.lean index 9ce17ae02c71f..ffdc7f1f2075f 100644 --- a/src/logic/equiv/local_equiv.lean +++ b/src/logic/equiv/local_equiv.lean @@ -126,10 +126,10 @@ structure local_equiv (α : Type*) (β : Type*) := (inv_fun : β → α) (source : set α) (target : set β) -(map_source' : ∀{x}, x ∈ source → to_fun x ∈ target) -(map_target' : ∀{x}, x ∈ target → inv_fun x ∈ source) -(left_inv' : ∀{x}, x ∈ source → inv_fun (to_fun x) = x) -(right_inv' : ∀{x}, x ∈ target → to_fun (inv_fun x) = x) +(map_source' : ∀ {{x}}, x ∈ source → to_fun x ∈ target) +(map_target' : ∀ {{x}}, x ∈ target → inv_fun x ∈ source) +(left_inv' : ∀ {{x}}, x ∈ source → inv_fun (to_fun x) = x) +(right_inv' : ∀ {{x}}, x ∈ target → to_fun (inv_fun x) = x) namespace local_equiv diff --git a/src/logic/function/basic.lean b/src/logic/function/basic.lean index c80b3d209d686..4a8e328352e8e 100644 --- a/src/logic/function/basic.lean +++ b/src/logic/function/basic.lean @@ -647,7 +647,7 @@ class has_uncurry (α : Type*) (β : out_param Type*) (γ : out_param Type*) := for bundled maps.-/ add_decl_doc has_uncurry.uncurry -notation `↿`:max x:max := has_uncurry.uncurry x +notation (name := uncurry) `↿`:max x:max := has_uncurry.uncurry x instance has_uncurry_base : has_uncurry (α → β) α β := ⟨id⟩ diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index f98ee0b433525..797a53a9cfd41 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -1919,7 +1919,8 @@ if hm : m ≤ m0 else 0 -- We define notation `μ[f|m]` for the conditional expectation of `f` with respect to `m`. -localized "notation μ `[` f `|` m `]` := measure_theory.condexp m μ f" in measure_theory +localized "notation (name := measure_theory.condexp) + μ `[` f `|` m `]` := measure_theory.condexp m μ f" in measure_theory lemma condexp_of_not_le (hm_not : ¬ m ≤ m0) : μ[f|m] = 0 := by rw [condexp, dif_neg hm_not] diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index efad68ace2274..a4c06311a8ed9 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1405,8 +1405,10 @@ def Lp {α} (E : Type*) {m : measurable_space α} [normed_add_comm_group E] neg_mem' := λ f hf, by rwa [set.mem_set_of_eq, snorm_congr_ae (ae_eq_fun.coe_fn_neg _), snorm_neg] } -localized "notation α ` →₁[`:25 μ `] ` E := measure_theory.Lp E 1 μ" in measure_theory -localized "notation α ` →₂[`:25 μ `] ` E := measure_theory.Lp E 2 μ" in measure_theory +localized "notation (name := measure_theory.L1) + α ` →₁[`:25 μ `] ` E := measure_theory.Lp E 1 μ" in measure_theory +localized "notation (name := measure_theory.L2) + α ` →₂[`:25 μ `] ` E := measure_theory.Lp E 2 μ" in measure_theory namespace mem_ℒp diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index df0bec642aab7..4fefccb4f2bab 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -89,8 +89,8 @@ variable [topological_space β] def strongly_measurable [measurable_space α] (f : α → β) : Prop := ∃ fs : ℕ → α →ₛ β, ∀ x, tendsto (λ n, fs n x) at_top (𝓝 (f x)) -localized "notation `strongly_measurable[` m `]` := @measure_theory.strongly_measurable _ _ _ m" -in measure_theory +localized "notation (name := strongly_measurable_of) + `strongly_measurable[` m `]` := @measure_theory.strongly_measurable _ _ _ m" in measure_theory /-- A function is `fin_strongly_measurable` with respect to a measure if it is the limit of simple functions with support with finite measure. -/ diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 76badaf315264..b7be4c344544b 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -513,7 +513,7 @@ def integral_clm : (α →₁ₛ[μ] E) →L[ℝ] E := integral_clm' α E ℝ μ variables {α E μ 𝕜} -local notation `Integral` := integral_clm α E μ +local notation (name := simple_func.integral_clm) `Integral` := integral_clm α E μ open continuous_linear_map @@ -585,7 +585,7 @@ end simple_func_integral end simple_func open simple_func -local notation `Integral` := @integral_clm α E _ _ _ _ _ μ _ +local notation (name := simple_func.integral_clm) `Integral` := @integral_clm α E _ _ _ _ _ μ _ variables [normed_space ℝ E] [nontrivially_normed_field 𝕜] [normed_space 𝕜 E] @@ -639,8 +639,9 @@ map_sub integral_clm f g lemma integral_smul (c : 𝕜) (f : α →₁[μ] E) : integral (c • f) = c • integral f := show (integral_clm' 𝕜) (c • f) = c • (integral_clm' 𝕜) f, from map_smul (integral_clm' 𝕜) c f -local notation `Integral` := @integral_clm α E _ _ μ _ _ -local notation `sIntegral` := @simple_func.integral_clm α E _ _ μ _ +local notation (name := integral_clm) `Integral` := @integral_clm α E _ _ μ _ _ +local notation (name := simple_func.integral_clm') `sIntegral` := + @simple_func.integral_clm α E _ _ μ _ lemma norm_Integral_le_one : ∥Integral∥ ≤ 1 := norm_set_to_L1_le (dominated_fin_meas_additive_weighted_smul μ) zero_le_one diff --git a/src/measure_theory/measurable_space_def.lean b/src/measure_theory/measurable_space_def.lean index 7f0a406e3fdb7..dd5fcad47e9a3 100644 --- a/src/measure_theory/measurable_space_def.lean +++ b/src/measure_theory/measurable_space_def.lean @@ -62,7 +62,8 @@ section /-- `measurable_set s` means that `s` is measurable (in the ambient measure space on `α`) -/ def measurable_set [measurable_space α] : set α → Prop := ‹measurable_space α›.measurable_set' -localized "notation `measurable_set[` m `]` := @measurable_set _ m" in measure_theory +localized "notation (name := measurable_set_of) + `measurable_set[` m `]` := @measurable_set hole! m" in measure_theory @[simp] lemma measurable_set.empty [measurable_space α] : measurable_set (∅ : set α) := ‹measurable_space α›.measurable_set_empty @@ -458,7 +459,8 @@ open measurable_space def measurable [measurable_space α] [measurable_space β] (f : α → β) : Prop := ∀ ⦃t : set β⦄, measurable_set t → measurable_set (f ⁻¹' t) -localized "notation `measurable[` m `]` := @measurable _ _ m _" in measure_theory +localized "notation (name := measurable_of) + `measurable[` m `]` := @measurable hole! hole! m hole!" in measure_theory lemma measurable_id {ma : measurable_space α} : measurable (@id α) := λ t, id diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index c0cf8d401d8c6..8370d00a788fa 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -553,7 +553,8 @@ end /-- Hausdorff measure on an (e)metric space. -/ def hausdorff_measure (d : ℝ) : measure X := mk_metric (λ r, r ^ d) -localized "notation `μH[` d `]` := measure_theory.measure.hausdorff_measure d" in measure_theory +localized "notation (name := hausdorff_measure) + `μH[` d `]` := measure_theory.measure.hausdorff_measure d" in measure_theory lemma le_hausdorff_measure (d : ℝ) (μ : measure X) (ε : ℝ≥0∞) (h₀ : 0 < ε) (h : ∀ s : set X, diam s ≤ ε → μ s ≤ diam s ^ d) : diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 529820cdfca5f..9a1b595ebae31 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -1767,7 +1767,8 @@ end count def absolutely_continuous {m0 : measurable_space α} (μ ν : measure α) : Prop := ∀ ⦃s : set α⦄, ν s = 0 → μ s = 0 -localized "infix ` ≪ `:50 := measure_theory.measure.absolutely_continuous" in measure_theory +localized "infix (name := measure.absolutely_continuous) + ` ≪ `:50 := measure_theory.measure.absolutely_continuous" in measure_theory lemma absolutely_continuous_of_le (h : μ ≤ ν) : μ ≪ ν := λ s hs, nonpos_iff_eq_zero.1 $ hs ▸ le_iff'.1 h s diff --git a/src/measure_theory/measure/mutually_singular.lean b/src/measure_theory/measure/mutually_singular.lean index 2df9de47e0ecd..c1d7dcc9a56e4 100644 --- a/src/measure_theory/measure/mutually_singular.lean +++ b/src/measure_theory/measure/mutually_singular.lean @@ -35,7 +35,8 @@ such that `μ s = 0` and `ν sᶜ = 0`. -/ def mutually_singular {m0 : measurable_space α} (μ ν : measure α) : Prop := ∃ (s : set α), measurable_set s ∧ μ s = 0 ∧ ν sᶜ = 0 -localized "infix ` ⊥ₘ `:60 := measure_theory.measure.mutually_singular" in measure_theory +localized "infix (name := measure.mutually_singular) + ` ⊥ₘ `:60 := measure_theory.measure.mutually_singular" in measure_theory namespace mutually_singular diff --git a/src/measure_theory/measure/vector_measure.lean b/src/measure_theory/measure/vector_measure.lean index 5ac7b7bfe7577..02547d722f6ea 100644 --- a/src/measure_theory/measure/vector_measure.lean +++ b/src/measure_theory/measure/vector_measure.lean @@ -775,9 +775,9 @@ end end -localized "notation v ` ≤[`:50 i:50 `] `:0 w:50 := -measure_theory.vector_measure.restrict v i ≤ measure_theory.vector_measure.restrict w i" -in measure_theory +localized "notation (name := vector_measure.restrict) v ` ≤[`:50 i:50 `] `:0 w:50 := + measure_theory.vector_measure.restrict v i ≤ measure_theory.vector_measure.restrict w i" + in measure_theory section @@ -1007,7 +1007,8 @@ def absolutely_continuous (v : vector_measure α M) (w : vector_measure α N) := ∀ ⦃s : set α⦄, w s = 0 → v s = 0 -localized "infix ` ≪ᵥ `:50 := measure_theory.vector_measure.absolutely_continuous" +localized "infix (name := vector_measure.absolutely_continuous) + ` ≪ᵥ `:50 := measure_theory.vector_measure.absolutely_continuous" in measure_theory open_locale measure_theory @@ -1099,7 +1100,8 @@ to use. This is equivalent to the definition which requires measurability. To pr def mutually_singular (v : vector_measure α M) (w : vector_measure α N) : Prop := ∃ (s : set α), measurable_set s ∧ (∀ t ⊆ s, v t = 0) ∧ (∀ t ⊆ sᶜ, w t = 0) -localized "infix ` ⊥ᵥ `:60 := measure_theory.vector_measure.mutually_singular" in measure_theory +localized "infix (name := vector_measure.mutually_singular) + ` ⊥ᵥ `:60 := measure_theory.vector_measure.mutually_singular" in measure_theory namespace mutually_singular diff --git a/src/meta/coinductive_predicates.lean b/src/meta/coinductive_predicates.lean index 1d9941f26d960..e84ad42a8491f 100644 --- a/src/meta/coinductive_predicates.lean +++ b/src/meta/coinductive_predicates.lean @@ -517,7 +517,7 @@ do namespace interactive open interactive interactive.types expr lean.parser local postfix `?`:9001 := optional -local postfix *:9001 := many +local postfix (name := parser.many) *:9001 := many meta def coinduction (corec_name : parse ident) (ns : parse with_ident_list) diff --git a/src/model_theory/basic.lean b/src/model_theory/basic.lean index 2c3df278de426..33dd3d5c7da4c 100644 --- a/src/model_theory/basic.lean +++ b/src/model_theory/basic.lean @@ -267,7 +267,8 @@ structure hom := (map_fun' : ∀{n} (f : L.functions n) x, to_fun (fun_map f x) = fun_map f (to_fun ∘ x) . obviously) (map_rel' : ∀{n} (r : L.relations n) x, rel_map r x → rel_map r (to_fun ∘ x) . obviously) -localized "notation A ` →[`:25 L `] ` B := first_order.language.hom L A B" in first_order +localized "notation (name := language.hom) A ` →[`:25 L `] ` B := + first_order.language.hom L A B" in first_order /-- An embedding of first-order structures is an embedding that commutes with the interpretations of functions and relations. -/ @@ -275,7 +276,8 @@ localized "notation A ` →[`:25 L `] ` B := first_order.language.hom L A B" in (map_fun' : ∀{n} (f : L.functions n) x, to_fun (fun_map f x) = fun_map f (to_fun ∘ x) . obviously) (map_rel' : ∀{n} (r : L.relations n) x, rel_map r (to_fun ∘ x) ↔ rel_map r x . obviously) -localized "notation A ` ↪[`:25 L `] ` B := first_order.language.embedding L A B" in first_order +localized "notation (name := language.embedding) A ` ↪[`:25 L `] ` B := + first_order.language.embedding L A B" in first_order /-- An equivalence of first-order structures is an equivalence that commutes with the interpretations of functions and relations. -/ @@ -283,7 +285,8 @@ structure equiv extends M ≃ N := (map_fun' : ∀{n} (f : L.functions n) x, to_fun (fun_map f x) = fun_map f (to_fun ∘ x) . obviously) (map_rel' : ∀{n} (r : L.relations n) x, rel_map r (to_fun ∘ x) ↔ rel_map r x . obviously) -localized "notation A ` ≃[`:25 L `] ` B := first_order.language.equiv L A B" in first_order +localized "notation (name := language.equiv) A ` ≃[`:25 L `] ` B := + first_order.language.equiv L A B" in first_order variables {L M N} {P : Type*} [L.Structure P] {Q : Type*} [L.Structure Q] diff --git a/src/model_theory/elementary_maps.lean b/src/model_theory/elementary_maps.lean index 16d9ded3ce40b..ada6f2c3ce971 100644 --- a/src/model_theory/elementary_maps.lean +++ b/src/model_theory/elementary_maps.lean @@ -38,11 +38,11 @@ variables [L.Structure M] [L.Structure N] [L.Structure P] [L.Structure Q] realizations of formulas. -/ structure elementary_embedding := (to_fun : M → N) -(map_formula' : ∀{n} (φ : L.formula (fin n)) (x : fin n → M), +(map_formula' : ∀ {{n}} (φ : L.formula (fin n)) (x : fin n → M), φ.realize (to_fun ∘ x) ↔ φ.realize x . obviously) -localized "notation A ` ↪ₑ[`:25 L `] ` B := first_order.language.elementary_embedding L A B" - in first_order +localized "notation (name := elementary_embedding) + A ` ↪ₑ[`:25 L `] ` B := first_order.language.elementary_embedding L A B" in first_order variables {L} {M} {N} @@ -290,18 +290,18 @@ end /-- A substructure is elementary when every formula applied to a tuple in the subtructure agrees with its value in the overall structure. -/ def is_elementary (S : L.substructure M) : Prop := -∀{n} (φ : L.formula (fin n)) (x : fin n → S), φ.realize ((coe : _ → M) ∘ x) ↔ φ.realize x +∀ {{n}} (φ : L.formula (fin n)) (x : fin n → S), φ.realize ((coe : _ → M) ∘ x) ↔ φ.realize x end substructure -variables (L) (M) +variables (L M) /-- An elementary substructure is one in which every formula applied to a tuple in the subtructure agrees with its value in the overall structure. -/ structure elementary_substructure := (to_substructure : L.substructure M) (is_elementary' : to_substructure.is_elementary) -variables {L} {M} +variables {L M} namespace elementary_substructure @@ -323,7 +323,7 @@ substructure.induced_Structure /-- The natural embedding of an `L.substructure` of `M` into `M`. -/ def subtype (S : L.elementary_substructure M) : S ↪ₑ[L] M := { to_fun := coe, - map_formula' := λ n, S.is_elementary } + map_formula' := S.is_elementary } @[simp] theorem coe_subtype {S : L.elementary_substructure M} : ⇑S.subtype = coe := rfl @@ -372,7 +372,7 @@ theorem is_elementary_of_exists (S : L.substructure M) φ.realize default (fin.snoc (coe ∘ x) a : _ → M) → ∃ b : S, φ.realize default (fin.snoc (coe ∘ x) b : _ → M)) : L.elementary_substructure M := -⟨S, λ _, S.is_elementary_of_exists htv⟩ +⟨S, S.is_elementary_of_exists htv⟩ end substructure diff --git a/src/model_theory/language_map.lean b/src/model_theory/language_map.lean index 8bbbe065e63ab..e0eb3c7b02482 100644 --- a/src/model_theory/language_map.lean +++ b/src/model_theory/language_map.lean @@ -38,8 +38,8 @@ variables (L : language.{u v}) (L' : language.{u' v'}) {M : Type w} [L.Structure /-- A language homomorphism maps the symbols of one language to symbols of another. -/ structure Lhom := -(on_function : ∀{n}, L.functions n → L'.functions n) -(on_relation : ∀{n}, L.relations n → L'.relations n) +(on_function : ∀ ⦃n⦄, L.functions n → L'.functions n) +(on_relation : ∀ ⦃n⦄, L.relations n → L'.relations n) infix ` →ᴸ `:10 := Lhom -- \^L @@ -108,7 +108,7 @@ Lhom.funext (funext (λ n, nat.cases_on n (funext h0) (λ n, nat.cases_on n (fun @[simps] def comp (g : L' →ᴸ L'') (f : L →ᴸ L') : L →ᴸ L'' := ⟨λ n F, g.1 (f.1 F), λ _ R, g.2 (f.2 R)⟩ -local infix ` ∘ `:60 := Lhom.comp +local infix (name := Lhom.comp) ` ∘ `:60 := Lhom.comp @[simp] lemma id_comp (F : L →ᴸ L') : (Lhom.id L') ∘ F = F := by {cases F, refl} @@ -168,8 +168,8 @@ end sum_map /-- A language homomorphism is injective when all the maps between symbol types are. -/ protected structure injective : Prop := -(on_function {n} : function.injective (on_function ϕ : L.functions n → L'.functions n)) -(on_relation {n} : function.injective (on_relation ϕ : L.relations n → L'.relations n)) +(on_function {n} : function.injective (λ f : L.functions n, on_function ϕ f)) +(on_relation {n} : function.injective (λ R : L.relations n, on_relation ϕ R)) /-- A language homomorphism is an expansion on a structure if it commutes with the interpretation of all symbols on that structure. -/ @@ -334,7 +334,8 @@ variables (α : Type w') /-- Extends a language with a constant for each element of a parameter set in `M`. -/ def with_constants : language.{(max u w') v} := L.sum (constants_on α) -localized "notation L`[[`:95 α`]]`:90 := L.with_constants α" in first_order +localized "notation (name := language.with_constants) + L`[[`:95 α`]]`:90 := L.with_constants α" in first_order @[simp] lemma card_with_constants : (L[[α]]).card = cardinal.lift.{w'} L.card + cardinal.lift.{max u v} (# α) := diff --git a/src/model_theory/satisfiability.lean b/src/model_theory/satisfiability.lean index cf31df0d3ff24..1930002c68f48 100644 --- a/src/model_theory/satisfiability.lean +++ b/src/model_theory/satisfiability.lean @@ -260,7 +260,8 @@ variable (T) def models_bounded_formula (φ : L.bounded_formula α n) : Prop := ∀ (M : Model.{u v (max u v)} T) (v : α → M) (xs : fin n → M), φ.realize v xs -infix ` ⊨ `:51 := models_bounded_formula -- input using \|= or \vDash, but not using \models +-- input using \|= or \vDash, but not using \models +infix (name := models_bounded_formula) ` ⊨ `:51 := models_bounded_formula variable {T} diff --git a/src/model_theory/semantics.lean b/src/model_theory/semantics.lean index c4ed8a4946d18..4a946f2219829 100644 --- a/src/model_theory/semantics.lean +++ b/src/model_theory/semantics.lean @@ -652,7 +652,8 @@ variable (M) def sentence.realize (φ : L.sentence) : Prop := φ.realize (default : _ → M) -infix ` ⊨ `:51 := sentence.realize -- input using \|= or \vDash, but not using \models +-- input using \|= or \vDash, but not using \models +infix (name := sentence.realize) ` ⊨ `:51 := sentence.realize @[simp] lemma sentence.realize_not {φ : L.sentence} : M ⊨ φ.not ↔ ¬ M ⊨ φ := @@ -673,8 +674,8 @@ variable (N) /-- Two structures are elementarily equivalent when they satisfy the same sentences. -/ def elementarily_equivalent : Prop := L.complete_theory M = L.complete_theory N -localized "notation A ` ≅[`:25 L `] ` B:50 := first_order.language.elementarily_equivalent L A B" - in first_order +localized "notation (name := elementarily_equivalent) A ` ≅[`:25 L `] ` B:50 := + first_order.language.elementarily_equivalent L A B" in first_order variables {L} {M} {N} @@ -689,7 +690,8 @@ variables (M) class Theory.model (T : L.Theory) : Prop := (realize_of_mem : ∀ φ ∈ T, M ⊨ φ) -infix ` ⊨ `:51 := Theory.model -- input using \|= or \vDash, but not using \models +-- input using \|= or \vDash, but not using \models +infix (name := Theory.model) ` ⊨ `:51 := Theory.model variables {M} (T : L.Theory) diff --git a/src/model_theory/skolem.lean b/src/model_theory/skolem.lean index 8557251d5c8c0..33d7de5234f3e 100644 --- a/src/model_theory/skolem.lean +++ b/src/model_theory/skolem.lean @@ -82,7 +82,7 @@ end /-- Any `L.sum L.skolem₁`-substructure is an elementary `L`-substructure. -/ noncomputable def elementary_skolem₁_reduct (S : (L.sum L.skolem₁).substructure M) : L.elementary_substructure M := -⟨Lhom.sum_inl.substructure_reduct S, λ _, S.skolem₁_reduct_is_elementary⟩ +⟨Lhom.sum_inl.substructure_reduct S, S.skolem₁_reduct_is_elementary⟩ lemma coe_sort_elementary_skolem₁_reduct (S : (L.sum L.skolem₁).substructure M) : diff --git a/src/model_theory/syntax.lean b/src/model_theory/syntax.lean index afa5ab848bd72..471f52d01f344 100644 --- a/src/model_theory/syntax.lean +++ b/src/model_theory/syntax.lean @@ -212,7 +212,8 @@ relabel (sum.map id (λ i, if ↑i < m then fin.cast_add n' i else fin.add_nat n end term -localized "prefix `&`:max := first_order.language.term.var ∘ sum.inr" in first_order +localized "prefix (name := language.term.var) `&`:max := + first_order.language.term.var ∘ sum.inr" in first_order namespace Lhom @@ -845,15 +846,21 @@ rfl end Lequiv -localized "infix ` =' `:88 := first_order.language.term.bd_equal" in first_order +localized "infix (name := term.bd_equal) + ` =' `:88 := first_order.language.term.bd_equal" in first_order -- input \~- or \simeq -localized "infixr ` ⟹ `:62 := first_order.language.bounded_formula.imp" in first_order +localized "infixr (name := bounded_formula.imp) + ` ⟹ `:62 := first_order.language.bounded_formula.imp" in first_order -- input \==> -localized "prefix `∀'`:110 := first_order.language.bounded_formula.all" in first_order -localized "prefix `∼`:max := first_order.language.bounded_formula.not" in first_order +localized "prefix (name := bounded_formula.all) + `∀'`:110 := first_order.language.bounded_formula.all" in first_order +localized "prefix (name := bounded_formula.not) + `∼`:max := first_order.language.bounded_formula.not" in first_order -- input \~, the ASCII character ~ has too low precedence -localized "infix ` ⇔ `:61 := first_order.language.bounded_formula.iff" in first_order -- input \<=> -localized "prefix `∃'`:110 := first_order.language.bounded_formula.ex" in first_order -- input \ex +localized "infix (name := bounded_formula.iff) + ` ⇔ `:61 := first_order.language.bounded_formula.iff" in first_order -- input \<=> +localized "prefix (name := bounded_formula.ex) + `∃'`:110 := first_order.language.bounded_formula.ex" in first_order -- input \ex namespace formula diff --git a/src/number_theory/arithmetic_function.lean b/src/number_theory/arithmetic_function.lean index b3eb74b4ab231..80539f7af0fc3 100644 --- a/src/number_theory/arithmetic_function.lean +++ b/src/number_theory/arithmetic_function.lean @@ -327,7 +327,8 @@ section zeta def zeta : arithmetic_function ℕ := ⟨λ x, ite (x = 0) 0 1, rfl⟩ -localized "notation `ζ` := nat.arithmetic_function.zeta" in arithmetic_function +localized "notation (name := arithmetic_function.zeta) + `ζ` := nat.arithmetic_function.zeta" in arithmetic_function @[simp] lemma zeta_apply {x : ℕ} : ζ x = if (x = 0) then 0 else 1 := rfl @@ -604,9 +605,8 @@ end⟩ /-- For any multiplicative function `f` and any `n > 0`, we can evaluate `f n` by evaluating `f` at `p ^ k` over the factorization of `n` -/ lemma multiplicative_factorization [comm_monoid_with_zero R] (f : arithmetic_function R) - (hf : f.is_multiplicative) : - ∀ {n : ℕ}, n ≠ 0 → f n = n.factorization.prod (λ p k, f (p ^ k)) := -λ n hn, multiplicative_factorization f hf.2 hf.1 hn + (hf : f.is_multiplicative) {n : ℕ} (hn : n ≠ 0) : f n = n.factorization.prod (λ p k, f (p ^ k)) := +multiplicative_factorization f (λ _ _, hf.2) hf.1 hn /-- A recapitulation of the definition of multiplicative that is simpler for proofs -/ lemma iff_ne_zero [monoid_with_zero R] {f : arithmetic_function R} : @@ -668,7 +668,8 @@ lemma pow_zero_eq_zeta : pow 0 = ζ := by { ext n, simp } def sigma (k : ℕ) : arithmetic_function ℕ := ⟨λ n, ∑ d in divisors n, d ^ k, by simp⟩ -localized "notation `σ` := nat.arithmetic_function.sigma" in arithmetic_function +localized "notation (name := arithmetic_function.sigma) + `σ` := nat.arithmetic_function.sigma" in arithmetic_function lemma sigma_apply {k n : ℕ} : σ k n = ∑ d in divisors n, d ^ k := rfl @@ -732,7 +733,8 @@ end def card_factors : arithmetic_function ℕ := ⟨λ n, n.factors.length, by simp⟩ -localized "notation `Ω` := nat.arithmetic_function.card_factors" in arithmetic_function +localized "notation (name := card_factors) + `Ω` := nat.arithmetic_function.card_factors" in arithmetic_function lemma card_factors_apply {n : ℕ} : Ω n = n.factors.length := rfl @@ -779,7 +781,8 @@ by rw [card_factors_apply, hp.factors_pow, list.length_repeat] def card_distinct_factors : arithmetic_function ℕ := ⟨λ n, n.factors.dedup.length, by simp⟩ -localized "notation `ω` := nat.arithmetic_function.card_distinct_factors" in arithmetic_function +localized "notation (name := card_distinct_factors) + `ω` := nat.arithmetic_function.card_distinct_factors" in arithmetic_function lemma card_distinct_factors_zero : ω 0 = 0 := by simp @@ -812,7 +815,8 @@ by rw [←pow_one p, card_distinct_factors_apply_prime_pow hp one_ne_zero] def moebius : arithmetic_function ℤ := ⟨λ n, if squarefree n then (-1) ^ (card_factors n) else 0, by simp⟩ -localized "notation `μ` := nat.arithmetic_function.moebius" in arithmetic_function +localized "notation (name := moebius) + `μ` := nat.arithmetic_function.moebius" in arithmetic_function @[simp] lemma moebius_apply_of_squarefree {n : ℕ} (h : squarefree n) : μ n = (-1) ^ card_factors n := diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index f323dcc6e0948..ac5a90622bbb6 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -86,7 +86,7 @@ section basic lemma iff_adjoin_eq_top : is_cyclotomic_extension S A B ↔ (∀ (a : ℕ+), a ∈ S → ∃ r : B, is_primitive_root r a) ∧ (adjoin A { b : B | ∃ a : ℕ+, a ∈ S ∧ b ^ (a : ℕ) = 1 } = ⊤) := -⟨λ h, ⟨h.exists_prim_root, algebra.eq_top_iff.2 h.adjoin_roots⟩, +⟨λ h, ⟨λ _, h.exists_prim_root, algebra.eq_top_iff.2 h.adjoin_roots⟩, λ h, ⟨h.1, algebra.eq_top_iff.1 h.2⟩⟩ /-- A reformulation of `is_cyclotomic_extension` in the case `S` is a singleton. -/ @@ -504,8 +504,8 @@ instance is_cyclotomic_extension [ne_zero ((n : ℕ) : A)] : subst a, haveI := ne_zero.of_no_zero_smul_divisors A K n, haveI := ne_zero.of_no_zero_smul_divisors A (cyclotomic_field n K) n, - obtain ⟨μ, hμ⟩ := let h := (cyclotomic_field.is_cyclotomic_extension n K).exists_prim_root - in h $ mem_singleton n, + obtain ⟨μ, hμ⟩ := + (cyclotomic_field.is_cyclotomic_extension n K).exists_prim_root (mem_singleton n), refine ⟨⟨μ, subset_adjoin _⟩, _⟩, { apply (is_root_of_unity_iff n.pos (cyclotomic_field n K)).mpr, refine ⟨n, nat.mem_divisors_self _ n.ne_zero, _⟩, diff --git a/src/number_theory/dioph.lean b/src/number_theory/dioph.lean index e6e2bea2e78a0..a08aa848b1497 100644 --- a/src/number_theory/dioph.lean +++ b/src/number_theory/dioph.lean @@ -395,19 +395,19 @@ dioph_comp ((dioph_fn_vec _).1 df) ((λ v, v none) :: λ i v, g i (v ∘ some)) by simp; exact ⟨proj_dioph none, (vector_allp_iff_forall _ _).2 $ λ i, reindex_dioph_fn _ $ (vector_allp_iff_forall _ _).1 dg _⟩ -localized "notation x ` D∧ `:35 y := dioph.inter x y" in dioph -localized "notation x ` D∨ `:35 y := dioph.union x y" in dioph +localized "notation (name := dioph.inter) x ` D∧ `:35 y := dioph.inter x y" in dioph +localized "notation (name := dioph.union) x ` D∨ `:35 y := dioph.union x y" in dioph -localized "notation `D∃`:30 := dioph.vec_ex1_dioph" in dioph +localized "notation (name := dioph.vec_ex1_dioph) `D∃`:30 := dioph.vec_ex1_dioph" in dioph -localized "prefix `&`:max := fin2.of_nat'" in dioph +localized "prefix (name := fin2.of_nat') `&`:max := fin2.of_nat'" in dioph theorem proj_dioph_of_nat {n : ℕ} (m : ℕ) [is_lt m n] : dioph_fn (λ v : vector3 ℕ n, v &m) := proj_dioph &m -localized "prefix `D&`:100 := dioph.proj_dioph_of_nat" in dioph +localized "prefix (name := proj_dioph_of_nat) `D&`:100 := dioph.proj_dioph_of_nat" in dioph theorem const_dioph (n : ℕ) : dioph_fn (const (α → ℕ) n) := abs_poly_dioph (poly.const n) -localized "prefix `D.`:100 := dioph.const_dioph" in dioph +localized "prefix (name := const_dioph) `D.`:100 := dioph.const_dioph" in dioph variables {f g : (α → ℕ) → ℕ} (df : dioph_fn f) (dg : dioph_fn g) include df dg @@ -424,26 +424,26 @@ lemma eq_dioph : dioph (λ v, f v = g v) := dioph_comp2 df dg $ of_no_dummies _ (poly.proj &0 - poly.proj &1) (λ v, (int.coe_nat_eq_coe_nat_iff _ _).symm.trans ⟨@sub_eq_zero_of_eq ℤ _ (v &0) (v &1), eq_of_sub_eq_zero⟩) -localized "infix ` D= `:50 := dioph.eq_dioph" in dioph +localized "infix (name := eq_dioph) ` D= `:50 := dioph.eq_dioph" in dioph lemma add_dioph : dioph_fn (λ v, f v + g v) := dioph_fn_comp2 df dg $ abs_poly_dioph (poly.proj &0 + poly.proj &1) -localized "infix ` D+ `:80 := dioph.add_dioph" in dioph +localized "infix (name := add_dioph) ` D+ `:80 := dioph.add_dioph" in dioph lemma mul_dioph : dioph_fn (λ v, f v * g v) := dioph_fn_comp2 df dg $ abs_poly_dioph (poly.proj &0 * poly.proj &1) -localized "infix ` D* `:90 := dioph.mul_dioph" in dioph +localized "infix (name := mul_dioph) ` D* `:90 := dioph.mul_dioph" in dioph lemma le_dioph : dioph {v | f v ≤ g v} := dioph_comp2 df dg $ ext (D∃2 $ D&1 D+ D&0 D= D&2) (λ v, ⟨λ ⟨x, hx⟩, le.intro hx, le.dest⟩) -localized "infix ` D≤ `:50 := dioph.le_dioph" in dioph +localized "infix (name := le_dioph) ` D≤ `:50 := dioph.le_dioph" in dioph lemma lt_dioph : dioph {v | f v < g v} := df D+ (D. 1) D≤ dg -localized "infix ` D< `:50 := dioph.lt_dioph" in dioph +localized "infix (name := lt_dioph) ` D< `:50 := dioph.lt_dioph" in dioph lemma ne_dioph : dioph {v | f v ≠ g v} := ext (df D< dg D∨ dg D< df) $ λ v, by { dsimp, exact lt_or_lt_iff_ne } -localized "infix ` D≠ `:50 := dioph.ne_dioph" in dioph +localized "infix (name := ne_dioph) ` D≠ `:50 := dioph.ne_dioph" in dioph lemma sub_dioph : dioph_fn (λ v, f v - g v) := dioph_fn_comp2 df dg $ (dioph_fn_vec _).2 $ @@ -459,11 +459,11 @@ end, begin { exact or.inr ⟨yz, tsub_eq_zero_iff_le.mpr yz⟩ }, { exact or.inl (tsub_add_cancel_of_le zy).symm }, end⟩ -localized "infix ` D- `:80 := dioph.sub_dioph" in dioph +localized "infix (name := sub_dioph) ` D- `:80 := dioph.sub_dioph" in dioph lemma dvd_dioph : dioph (λ v, f v ∣ g v) := dioph_comp (D∃2 $ D&2 D= D&1 D* D&0) [f, g] (by exact ⟨df, dg⟩) -localized "infix ` D∣ `:50 := dioph.dvd_dioph" in dioph +localized "infix (name := dvd_dioph) ` D∣ `:50 := dioph.dvd_dioph" in dioph lemma mod_dioph : dioph_fn (λ v, f v % g v) := have dioph (λ v : vector3 ℕ 3, (v &2 = 0 ∨ v &0 < v &2) ∧ ∃ (x : ℕ), v &0 + v &2 * x = v &1), @@ -475,11 +475,11 @@ show ((y = 0 ∨ z < y) ∧ ∃ c, z + y * c = x) ↔ x % y = z, from λ e, by rw ← e; exact ⟨or_iff_not_imp_left.2 $ λ h, mod_lt _ (nat.pos_of_ne_zero h), x / y, mod_add_div _ _⟩⟩ -localized "infix ` D% `:80 := dioph.mod_dioph" in dioph +localized "infix (name := mod_dioph) ` D% `:80 := dioph.mod_dioph" in dioph lemma modeq_dioph {h : (α → ℕ) → ℕ} (dh : dioph_fn h) : dioph (λ v, f v ≡ g v [MOD h v]) := df D% dh D= dg D% dh -localized "notation `D≡` := dioph.modeq_dioph" in dioph +localized "notation (name := modeq_dioph) ` D≡ ` := dioph.modeq_dioph" in dioph lemma div_dioph : dioph_fn (λ v, f v / g v) := have dioph (λ v : vector3 ℕ 3, v &2 = 0 ∧ v &0 = 0 ∨ v &0 * v &2 ≤ v &1 ∧ v &1 < (v &0 + 1) * v &2), @@ -492,7 +492,7 @@ by refine iff.trans _ eq_comm; exact y.eq_zero_or_pos.elim (λ ypos, iff.trans ⟨λ o, o.resolve_left $ λ ⟨h1, _⟩, ne_of_gt ypos h1, or.inr⟩ (le_antisymm_iff.trans $ and_congr (nat.le_div_iff_mul_le ypos) $ iff.trans ⟨lt_succ_of_le, le_of_lt_succ⟩ (div_lt_iff_lt_mul ypos)).symm) -localized "infix ` D/ `:80 := dioph.div_dioph" in dioph +localized "infix (name := div_dioph) ` D/ `:80 := dioph.div_dioph" in dioph omit df dg open pell diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index 0f32ee0bba55b..79919a734a557 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -409,9 +409,9 @@ def fd : set ℍ := def fdo : set ℍ := {z | 1 < (z : ℂ).norm_sq ∧ |z.re| < (1 : ℝ) / 2} -localized "notation `𝒟` := modular_group.fd" in modular +localized "notation (name := modular_group.fd) `𝒟` := modular_group.fd" in modular -localized "notation `𝒟ᵒ` := modular_group.fdo" in modular +localized "notation (name := modular_group.fdo) `𝒟ᵒ` := modular_group.fdo" in modular lemma abs_two_mul_re_lt_one_of_mem_fdo (h : z ∈ 𝒟ᵒ) : |2 * z.re| < 1 := begin diff --git a/src/number_theory/modular_forms/slash_actions.lean b/src/number_theory/modular_forms/slash_actions.lean index afb5f4d58ef8f..0f73ab5dd46f6 100644 --- a/src/number_theory/modular_forms/slash_actions.lean +++ b/src/number_theory/modular_forms/slash_actions.lean @@ -56,7 +56,8 @@ def slash : ℤ → GL(2, ℝ)⁺ → (ℍ → ℂ) → (ℍ → ℂ) := λ k γ variables {Γ : subgroup SL(2,ℤ)} {k: ℤ} (f : (ℍ → ℂ)) -localized "notation f ` ∣[`:100 k `]`:0 γ :100 := slash k γ f" in modular_forms +localized "notation (name := modular_forms.slash) f ` ∣[`:100 k `]`:0 γ :100 := + modular_forms.slash k γ f" in modular_forms lemma slash_right_action (k : ℤ) (A B : GL(2, ℝ)⁺) (f : ℍ → ℂ) : (f ∣[k] A) ∣[k] B = f ∣[k] (A * B) := diff --git a/src/number_theory/number_field.lean b/src/number_theory/number_field.lean index 7d73c1a9b9b9f..0317ff7dc52c9 100644 --- a/src/number_theory/number_field.lean +++ b/src/number_theory/number_field.lean @@ -66,7 +66,8 @@ omit nf is the integral closure of ℤ in the number field. -/ def ring_of_integers := integral_closure ℤ K -localized "notation `𝓞` := number_field.ring_of_integers" in number_field +localized "notation (name := ring_of_integers) + `𝓞` := number_field.ring_of_integers" in number_field lemma mem_ring_of_integers (x : K) : x ∈ 𝓞 K ↔ is_integral ℤ x := iff.rfl diff --git a/src/number_theory/prime_counting.lean b/src/number_theory/prime_counting.lean index 1054b077f7add..80d226d27dfa8 100644 --- a/src/number_theory/prime_counting.lean +++ b/src/number_theory/prime_counting.lean @@ -48,8 +48,8 @@ def prime_counting' : ℕ → ℕ := nat.count prime /-- The prime counting function: Returns the number of primes less than or equal to the input. -/ def prime_counting (n : ℕ) : ℕ := prime_counting' (n + 1) -localized "notation `π` := nat.prime_counting" in nat -localized "notation `π'` := nat.prime_counting'" in nat +localized "notation (name := prime_counting) `π` := nat.prime_counting" in nat +localized "notation (name := prime_counting') `π'` := nat.prime_counting'" in nat lemma monotone_prime_counting' : monotone prime_counting' := count_monotone prime diff --git a/src/number_theory/von_mangoldt.lean b/src/number_theory/von_mangoldt.lean index 063d85d981246..4f7744af3a2ff 100644 --- a/src/number_theory/von_mangoldt.lean +++ b/src/number_theory/von_mangoldt.lean @@ -57,7 +57,8 @@ In the `arithmetic_function` locale, we have the notation `Λ` for this function noncomputable def von_mangoldt : arithmetic_function ℝ := ⟨λ n, if is_prime_pow n then real.log (min_fac n) else 0, if_neg not_is_prime_pow_zero⟩ -localized "notation `Λ` := nat.arithmetic_function.von_mangoldt" in arithmetic_function +localized "notation (name := von_mangoldt) + `Λ` := nat.arithmetic_function.von_mangoldt" in arithmetic_function lemma von_mangoldt_apply {n : ℕ} : Λ n = if is_prime_pow n then real.log (min_fac n) else 0 := rfl diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 85569c1cc42f0..7e654bca61aec 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -242,7 +242,7 @@ def principal (s : set α) : filter α := sets_of_superset := λ x y hx, subset.trans hx, inter_sets := λ x y, subset_inter } -localized "notation `𝓟` := filter.principal" in filter +localized "notation (name := filter.principal) `𝓟` := filter.principal" in filter instance : inhabited (filter α) := ⟨𝓟 ∅⟩ diff --git a/src/order/filter/prod.lean b/src/order/filter/prod.lean index 40dd60b283837..c674f22486f75 100644 --- a/src/order/filter/prod.lean +++ b/src/order/filter/prod.lean @@ -49,7 +49,7 @@ of elements of the component filters. -/ protected def prod (f : filter α) (g : filter β) : filter (α × β) := f.comap prod.fst ⊓ g.comap prod.snd -localized "infix ` ×ᶠ `:60 := filter.prod" in filter +localized "infix (name := filter.prod) ` ×ᶠ `:60 := filter.prod" in filter lemma prod_mem_prod {s : set α} {t : set β} {f : filter α} {g : filter β} (hs : s ∈ f) (ht : t ∈ g) : s ×ˢ t ∈ f ×ᶠ g := diff --git a/src/order/hom/basic.lean b/src/order/hom/basic.lean index 4779aea5d9daf..074068651b763 100644 --- a/src/order/hom/basic.lean +++ b/src/order/hom/basic.lean @@ -499,7 +499,7 @@ instance : order_iso_class (α ≃o β) α β := left_inv := λ f, f.left_inv, right_inv := λ f, f.right_inv, coe_injective' := λ f g h₁ h₂, by { obtain ⟨⟨_, _⟩, _⟩ := f, obtain ⟨⟨_, _⟩, _⟩ := g, congr' }, - map_le_map_iff := λ f, f.map_rel_iff' } + map_le_map_iff := λ f _ _, f.map_rel_iff' } @[simp] lemma to_fun_eq_coe {f : α ≃o β} : f.to_fun = f := rfl diff --git a/src/order/initial_seg.lean b/src/order/initial_seg.lean index 27b3920a62acd..25d3a21cfd946 100644 --- a/src/order/initial_seg.lean +++ b/src/order/initial_seg.lean @@ -48,7 +48,7 @@ range of `f`. -/ structure initial_seg {α β : Type*} (r : α → α → Prop) (s : β → β → Prop) extends r ↪r s := (init : ∀ a b, s b (to_rel_embedding a) → ∃ a', to_rel_embedding a' = b) -localized "infix ` ≼i `:25 := initial_seg" in initial_seg +localized "infix (name := initial_seg) ` ≼i `:25 := initial_seg" in initial_seg namespace initial_seg @@ -118,7 +118,7 @@ initial_seg.eq (f.trans g) (initial_seg.refl _) is a well-order then `α` and `β` are order-isomorphic. -/ def antisymm [is_well_order β s] (f : r ≼i s) (g : s ≼i r) : r ≃r s := by haveI := f.to_rel_embedding.is_well_order; exact -⟨⟨f, g, antisymm.aux f g, antisymm.aux g f⟩, f.map_rel_iff'⟩ +⟨⟨f, g, antisymm.aux f g, antisymm.aux g f⟩, λ _ _, f.map_rel_iff'⟩ @[simp] theorem antisymm_to_fun [is_well_order β s] (f : r ≼i s) (g : s ≼i r) : (antisymm f g : α → β) = f := rfl @@ -174,7 +174,7 @@ structure principal_seg {α β : Type*} (r : α → α → Prop) (s : β → β (top : β) (down' : ∀ b, s b top ↔ ∃ a, to_rel_embedding a = b) -localized "infix ` ≺i `:25 := principal_seg" in initial_seg +localized "infix (name := principal_seg) ` ≺i `:25 := principal_seg" in initial_seg namespace principal_seg diff --git a/src/order/rel_iso.lean b/src/order/rel_iso.lean index 17eb6b62f5a5d..8b326a1d484c9 100644 --- a/src/order/rel_iso.lean +++ b/src/order/rel_iso.lean @@ -108,7 +108,7 @@ instance : has_coe_to_fun (r →r s) (λ _, α → β) := ⟨λ o, o.to_fun⟩ initialize_simps_projections rel_hom (to_fun → apply) -protected theorem map_rel (f : r →r s) : ∀ {a b}, r a b → s (f a) (f b) := f.map_rel' +protected theorem map_rel (f : r →r s) {a b} : r a b → s (f a) (f b) := f.map_rel' @[simp] theorem coe_fn_mk (f : α → β) (o) : (@rel_hom.mk _ _ r s f o : α → β) = f := rfl @@ -215,7 +215,7 @@ theorem injective (f : r ↪r s) : injective f := f.inj' @[simp] theorem inj (f : r ↪r s) {a b} : f a = f b ↔ a = b := f.injective.eq_iff -theorem map_rel_iff (f : r ↪r s) : ∀ {a b}, s (f a) (f b) ↔ r a b := f.map_rel_iff' +theorem map_rel_iff (f : r ↪r s) {a b} : s (f a) (f b) ↔ r a b := f.map_rel_iff' @[simp] theorem coe_fn_mk (f : α ↪ β) (o) : (@rel_embedding.mk _ _ r s f o : α → β) = f := rfl @@ -435,7 +435,7 @@ namespace rel_iso but often it is easier to write `f.to_rel_embedding` than to write explicitly `r` and `s` in the target type. -/ def to_rel_embedding (f : r ≃r s) : r ↪r s := -⟨f.to_equiv.to_embedding, f.map_rel_iff'⟩ +⟨f.to_equiv.to_embedding, λ _ _, f.map_rel_iff'⟩ theorem to_equiv_injective : injective (to_equiv : (r ≃r s) → α ≃ β) | ⟨e₁, o₁⟩ ⟨e₂, o₂⟩ h := by { congr, exact h } @@ -454,7 +454,7 @@ instance : rel_hom_class (r ≃r s) r s := @[simp] lemma coe_coe_fn (f : r ≃r s) : ((f : r ↪r s) : α → β) = f := rfl -theorem map_rel_iff (f : r ≃r s) : ∀ {a b}, s (f a) (f b) ↔ r a b := f.map_rel_iff' +theorem map_rel_iff (f : r ≃r s) {a b} : s (f a) (f b) ↔ r a b := f.map_rel_iff' @[simp] theorem coe_fn_mk (f : α ≃ β) (o : ∀ ⦃a b⦄, s (f a) (f b) ↔ r a b) : (rel_iso.mk f o : α → β) = f := rfl @@ -648,7 +648,7 @@ end subrel /-- Restrict the codomain of a relation embedding. -/ def rel_embedding.cod_restrict (p : set β) (f : r ↪r s) (H : ∀ a, f a ∈ p) : r ↪r subrel s p := -⟨f.to_embedding.cod_restrict p H, f.map_rel_iff'⟩ +⟨f.to_embedding.cod_restrict p H, λ _ _, f.map_rel_iff'⟩ @[simp] theorem rel_embedding.cod_restrict_apply (p) (f : r ↪r s) (H a) : rel_embedding.cod_restrict p f H a = ⟨f a, H a⟩ := rfl diff --git a/src/probability/conditional_probability.lean b/src/probability/conditional_probability.lean index dfa312ffd476d..83f94ecabc3ca 100644 --- a/src/probability/conditional_probability.lean +++ b/src/probability/conditional_probability.lean @@ -73,8 +73,10 @@ def cond (s : set Ω) : measure Ω := end definitions -localized "notation μ `[` s `|` t `]` := probability_theory.cond μ t s" in probability_theory -localized "notation μ `[|`:60 t`]` := probability_theory.cond μ t" in probability_theory +localized "notation (name := probability_theory.cond) + μ `[` s `|` t `]` := probability_theory.cond μ t s" in probability_theory +localized "notation (name := probability_theory.cond_fn) + μ `[|`:60 t`]` := probability_theory.cond μ t" in probability_theory /-- The conditional probability measure of any finite measure on any set of positive measure is a probability measure. -/ diff --git a/src/probability/notation.lean b/src/probability/notation.lean index 882efd388369c..4c4d31f9632c7 100644 --- a/src/probability/notation.lean +++ b/src/probability/notation.lean @@ -27,19 +27,21 @@ We note that the notation `∂P/∂Q` applies to three different cases, namely, open measure_theory -- We define notations `𝔼[f|m]` for the conditional expectation of `f` with respect to `m`. -localized "notation `𝔼[` X `|` m `]` := +localized "notation (name := condexp.volume) `𝔼[` X `|` m `]` := measure_theory.condexp m measure_theory.measure_space.volume X" in probability_theory -localized "notation P `[` X `]` := ∫ x, X x ∂P" in probability_theory +localized "notation (name := condexp.probability) + P `[` X `]` := ∫ x, X x ∂P" in probability_theory -localized "notation `𝔼[` X `]` := ∫ a, X a" in probability_theory +localized "notation (name := expected_value) `𝔼[` X `]` := ∫ a, X a" in probability_theory -localized "notation X ` =ₐₛ `:50 Y:50 := X =ᵐ[measure_theory.measure_space.volume] Y" - in probability_theory +localized "notation (name := eq_ae_volume) + X ` =ₐₛ `:50 Y:50 := X =ᵐ[measure_theory.measure_space.volume] Y" in probability_theory -localized "notation X ` ≤ₐₛ `:50 Y:50 := X ≤ᵐ[measure_theory.measure_space.volume] Y" - in probability_theory +localized "notation (name := le_ae_volume) + X ` ≤ₐₛ `:50 Y:50 := X ≤ᵐ[measure_theory.measure_space.volume] Y" in probability_theory -localized "notation `∂` P `/∂`:50 Q:50 := P.rn_deriv Q" in probability_theory +localized "notation (name := rn_deriv) `∂` P `/∂`:50 Q:50 := P.rn_deriv Q" in probability_theory -localized "notation `ℙ` := measure_theory.measure_space.volume" in probability_theory +localized "notation (name := measure_space.volume) + `ℙ` := measure_theory.measure_space.volume" in probability_theory diff --git a/src/probability/variance.lean b/src/probability/variance.lean index e884101ba30ec..cb59ffa0df42e 100644 --- a/src/probability/variance.lean +++ b/src/probability/variance.lean @@ -66,9 +66,8 @@ begin { simp only [algebra.smul_def, map_pow], } end -localized -"notation `Var[` X `]` := probability_theory.variance X measure_theory.measure_space.volume" -in probability_theory +localized "notation (name := probability_theory.variance) `Var[` X `]` := + probability_theory.variance X measure_theory.measure_space.volume" in probability_theory variables {Ω : Type*} [measure_space Ω] [is_probability_measure (volume : measure Ω)] diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index b085195f21f20..15a1f9efd1d43 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -151,12 +151,12 @@ class is_prime (I : ideal α) : Prop := theorem is_prime_iff {I : ideal α} : is_prime I ↔ I ≠ ⊤ ∧ ∀ {x y : α}, x * y ∈ I → x ∈ I ∨ y ∈ I := -⟨λ h, ⟨h.1, h.2⟩, λ h, ⟨h.1, h.2⟩⟩ +⟨λ h, ⟨h.1, λ _ _, h.2⟩, λ h, ⟨h.1, λ _ _, h.2⟩⟩ theorem is_prime.ne_top {I : ideal α} (hI : I.is_prime) : I ≠ ⊤ := hI.1 -theorem is_prime.mem_or_mem {I : ideal α} (hI : I.is_prime) : - ∀ {x y : α}, x * y ∈ I → x ∈ I ∨ y ∈ I := hI.2 +theorem is_prime.mem_or_mem {I : ideal α} (hI : I.is_prime) {x y : α} : + x * y ∈ I → x ∈ I ∨ y ∈ I := hI.2 theorem is_prime.mem_or_mem_of_mul_eq_zero {I : ideal α} (hI : I.is_prime) {x y : α} (h : x * y = 0) : x ∈ I ∨ y ∈ I := diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index db1d940fa911a..b2ce5384f0c12 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -1311,7 +1311,7 @@ def rel_iso_of_bijective : ideal S ≃o ideal R := left_inv := (rel_iso_of_surjective f hf.right).left_inv, right_inv := λ J, subtype.ext_iff.1 ((rel_iso_of_surjective f hf.right).right_inv ⟨J, comap_bot_le_of_injective f hf.left⟩), - map_rel_iff' := (rel_iso_of_surjective f hf.right).map_rel_iff' } + map_rel_iff' := λ _ _, (rel_iso_of_surjective f hf.right).map_rel_iff' } lemma comap_le_iff_le_map {I : ideal R} {K : ideal S} : comap f K ≤ I ↔ K ≤ map f I := ⟨λ h, le_map_of_comap_le_of_surjective f hf.right h, diff --git a/src/ring_theory/localization/num_denom.lean b/src/ring_theory/localization/num_denom.lean index 2da5967b481dd..db61a20f6d821 100644 --- a/src/ring_theory/localization/num_denom.lean +++ b/src/ring_theory/localization/num_denom.lean @@ -54,8 +54,7 @@ classical.some (exists_reduced_fraction A x) noncomputable def denom (x : K) : non_zero_divisors A := classical.some (classical.some_spec (exists_reduced_fraction A x)) -lemma num_denom_reduced (x : K) : - ∀ {d}, d ∣ num A x → d ∣ denom A x → is_unit d := +lemma num_denom_reduced (x : K) {d} : d ∣ num A x → d ∣ denom A x → is_unit d := (classical.some_spec (classical.some_spec (exists_reduced_fraction A x))).1 @[simp] lemma mk'_num_denom (x : K) : mk' K (num A x) (denom A x) = x := diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index 9233e86bc33c6..4419ad7d2f4a4 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -544,8 +544,7 @@ begin { simp only [or.by_cases, dif_pos] }, { ext ⟨i, his⟩, have : ¬i = a, { rintro rfl, exact has his }, - dsimp only [or.by_cases], change i ∈ s at his, - rw [dif_neg this, dif_pos his] } }, + simp only [or.by_cases, this, not_false_iff, dif_neg] } }, { intro f, ext ⟨i, hi⟩, rcases finset.mem_insert.1 hi with rfl | h, { simp only [or.by_cases, dif_pos], }, diff --git a/src/ring_theory/non_zero_divisors.lean b/src/ring_theory/non_zero_divisors.lean index eb3efd03c51e4..f7577bfa68dfd 100644 --- a/src/ring_theory/non_zero_divisors.lean +++ b/src/ring_theory/non_zero_divisors.lean @@ -30,7 +30,8 @@ def non_zero_divisors (R : Type*) [monoid_with_zero R] : submonoid R := have z * x₁ * x₂ = 0, by rwa mul_assoc, hx₁ z $ hx₂ (z * x₁) this } -localized "notation R`⁰`:9000 := non_zero_divisors R" in non_zero_divisors +localized "notation (name := non_zero_divisors) + R`⁰`:9000 := non_zero_divisors R" in non_zero_divisors variables {M M' M₁ R R' F : Type*} [monoid_with_zero M] [monoid_with_zero M'] [comm_monoid_with_zero M₁] [ring R] [comm_ring R'] diff --git a/src/ring_theory/polynomial/eisenstein.lean b/src/ring_theory/polynomial/eisenstein.lean index cf28fd70a6b87..096d48fa57251 100644 --- a/src/ring_theory/polynomial/eisenstein.lean +++ b/src/ring_theory/polynomial/eisenstein.lean @@ -211,7 +211,7 @@ lemma _root_.polynomial.monic.is_eisenstein_at_of_mem_of_not_mem (hf : f.monic) include hf -lemma is_weakly_eisenstein_at : is_weakly_eisenstein_at f 𝓟 := ⟨hf.mem⟩ +lemma is_weakly_eisenstein_at : is_weakly_eisenstein_at f 𝓟 := ⟨λ _, hf.mem⟩ lemma coeff_mem {n : ℕ} (hn : n ≠ f.nat_degree) : f.coeff n ∈ 𝓟 := begin diff --git a/src/ring_theory/principal_ideal_domain.lean b/src/ring_theory/principal_ideal_domain.lean index 8283f90ce82a8..c0603f90d6d8f 100644 --- a/src/ring_theory/principal_ideal_domain.lean +++ b/src/ring_theory/principal_ideal_domain.lean @@ -104,7 +104,7 @@ lemma prime_generator_of_is_prime (S : ideal R) [submodule.is_principal S] [is_p prime (generator S) := ⟨λ h, ne_bot ((eq_bot_iff_generator_eq_zero S).2 h), λ h, is_prime.ne_top (S.eq_top_of_is_unit_mem (generator_mem S) h), - by simpa only [← mem_iff_generator_dvd S] using is_prime.2⟩ + λ _ _, by simpa only [← mem_iff_generator_dvd S] using is_prime.2⟩ -- Note that the converse may not hold if `ϕ` is not injective. lemma generator_map_dvd_of_mem {N : submodule R M} diff --git a/src/ring_theory/subring/basic.lean b/src/ring_theory/subring/basic.lean index cac9f06f6c9b0..58fb30a99bd6c 100644 --- a/src/ring_theory/subring/basic.lean +++ b/src/ring_theory/subring/basic.lean @@ -211,7 +211,7 @@ iff.rfl equalities. -/ protected def copy (S : subring R) (s : set R) (hs : s = ↑S) : subring R := { carrier := s, - neg_mem' := hs.symm ▸ S.neg_mem', + neg_mem' := λ _, hs.symm ▸ S.neg_mem', ..S.to_subsemiring.copy s hs } @[simp] lemma coe_copy (S : subring R) (s : set R) (hs : s = ↑S) : diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index aa8f40d5db680..bade5ec815bf0 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -810,7 +810,9 @@ end add_valuation section valuation_notation -localized "notation `ℕₘ₀` := with_zero (multiplicative ℕ)" in discrete_valuation -localized "notation `ℤₘ₀` := with_zero (multiplicative ℤ)" in discrete_valuation +localized "notation (name := nat.multiplicative_zero) + `ℕₘ₀` := with_zero (multiplicative ℕ)" in discrete_valuation +localized "notation (name := int.multiplicative_zero) + `ℤₘ₀` := with_zero (multiplicative ℤ)" in discrete_valuation end valuation_notation diff --git a/src/ring_theory/witt_vector/isocrystal.lean b/src/ring_theory/witt_vector/isocrystal.lean index 7323a6d8f3312..853f192542547 100644 --- a/src/ring_theory/witt_vector/isocrystal.lean +++ b/src/ring_theory/witt_vector/isocrystal.lean @@ -58,7 +58,8 @@ namespace witt_vector variables (p : ℕ) [fact p.prime] variables (k : Type*) [comm_ring k] -localized "notation `K(` p`,` k`)` := fraction_ring (witt_vector p k)" in isocrystal +localized "notation (name := witt_vector.fraction_ring) + `K(` p`,` k`)` := fraction_ring (witt_vector p k)" in isocrystal section perfect_ring variables [is_domain k] [char_p k p] [perfect_ring k p] @@ -72,8 +73,8 @@ is_fraction_ring.field_equiv_of_ring_equiv (frobenius_equiv p k) /-- The Frobenius automorphism of `k` induces an endomorphism of `K`. For notation purposes. -/ def fraction_ring.frobenius_ring_hom : K(p, k) →+* K(p, k) := fraction_ring.frobenius p k -localized "notation `φ(` p`,` k`)` := witt_vector.fraction_ring.frobenius_ring_hom p k" - in isocrystal +localized "notation (name := witt_vector.frobenius_ring_hom) + `φ(` p`,` k`)` := witt_vector.fraction_ring.frobenius_ring_hom p k" in isocrystal instance inv_pair₁ : ring_hom_inv_pair (φ(p, k)) _ := ring_hom_inv_pair.of_ring_equiv (fraction_ring.frobenius p k) @@ -82,9 +83,9 @@ instance inv_pair₂ : ring_hom_inv_pair ((fraction_ring.frobenius p k).symm : K(p, k) →+* K(p, k)) _ := ring_hom_inv_pair.of_ring_equiv (fraction_ring.frobenius p k).symm -localized "notation M ` →ᶠˡ[`:50 p `,` k `] ` M₂ := +localized "notation (name := frobenius_ring_hom.linear_map) M ` →ᶠˡ[`:50 p `,` k `] ` M₂ := linear_map (witt_vector.fraction_ring.frobenius_ring_hom p k) M M₂" in isocrystal -localized "notation M ` ≃ᶠˡ[`:50 p `,` k `] ` M₂ := +localized "notation (name := frobenius_ring_hom.linear_equiv) M ` ≃ᶠˡ[`:50 p `,` k `] ` M₂ := linear_equiv (witt_vector.fraction_ring.frobenius_ring_hom p k) M M₂" in isocrystal /-! ### Isocrystals -/ @@ -119,10 +120,10 @@ structure isocrystal_hom extends V →ₗ[K(p, k)] V₂ := structure isocrystal_equiv extends V ≃ₗ[K(p, k)] V₂ := ( frob_equivariant : ∀ x : V, Φ(p, k) (to_linear_equiv x) = to_linear_equiv (Φ(p, k) x) ) -localized "notation M ` →ᶠⁱ[`:50 p `,` k `] ` M₂ := witt_vector.isocrystal_hom p k M M₂" - in isocrystal -localized "notation M ` ≃ᶠⁱ[`:50 p `,` k `] ` M₂ := witt_vector.isocrystal_equiv p k M M₂" - in isocrystal +localized "notation (name := isocrystal_hom) + M ` →ᶠⁱ[`:50 p `,` k `] ` M₂ := witt_vector.isocrystal_hom p k M M₂" in isocrystal +localized "notation (name := isocrystal_equiv) + M ` ≃ᶠⁱ[`:50 p `,` k `] ` M₂ := witt_vector.isocrystal_equiv p k M M₂" in isocrystal end perfect_ring diff --git a/src/ring_theory/witt_vector/witt_polynomial.lean b/src/ring_theory/witt_vector/witt_polynomial.lean index fdca5dec0355c..8a4a5d7e6fc46 100644 --- a/src/ring_theory/witt_vector/witt_polynomial.lean +++ b/src/ring_theory/witt_vector/witt_polynomial.lean @@ -88,9 +88,9 @@ end This allows us to simply write `W n` or `W_ ℤ n`. -/ -- Notation with ring of coefficients explicit -localized "notation `W_` := witt_polynomial p" in witt +localized "notation (name := witt_polynomial) `W_` := witt_polynomial p" in witt -- Notation with ring of coefficients implicit -localized "notation `W` := witt_polynomial p _" in witt +localized "notation (name := witt_polynomial.infer) `W` := witt_polynomial p hole!" in witt open_locale witt open mv_polynomial diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index 9907e6c93ae41..fe6ffafe84351 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -59,7 +59,7 @@ We define cardinal numbers as a quotient of types under the equivalence relation * There is an instance `has_pow cardinal`, but this will only fire if Lean already knows that both the base and the exponent live in the same universe. As a workaround, you can add ``` - local infixr ^ := @has_pow.pow cardinal cardinal cardinal.has_pow + local infixr (name := cardinal.pow) ^ := @has_pow.pow cardinal cardinal cardinal.has_pow ``` to a file. This notation will work even if Lean doesn't know yet that the base and the exponent live in the same universe (but no exponents in other types can be used). @@ -102,7 +102,7 @@ namespace cardinal /-- The cardinal number of a type -/ def mk : Type u → cardinal := quotient.mk -localized "notation `#` := cardinal.mk" in cardinal +localized "notation (name := cardinal.mk) `#` := cardinal.mk" in cardinal instance can_lift_cardinal_Type : can_lift cardinal.{u} (Type u) := ⟨mk, λ c, true, λ c _, quot.induction_on c $ λ α, ⟨α, rfl⟩⟩ @@ -336,8 +336,8 @@ induction_on₂ a b $ λ α β, mk_congr $ equiv.prod_comm α β instance : has_pow cardinal.{u} cardinal.{u} := ⟨map₂ (λ α β, β → α) (λ α β γ δ e₁ e₂, e₂.arrow_congr e₁)⟩ -local infixr ^ := @has_pow.pow cardinal cardinal cardinal.has_pow -local infixr ` ^ℕ `:80 := @has_pow.pow cardinal ℕ monoid.has_pow +local infixr (name := cardinal.pow) ^ := @has_pow.pow cardinal cardinal cardinal.has_pow +local infixr (name := cardinal.pow.nat) ` ^ℕ `:80 := @has_pow.pow cardinal ℕ monoid.has_pow theorem power_def (α β) : #α ^ #β = #(β → α) := rfl @@ -821,7 +821,7 @@ lift_supr_le_lift_supr hf hf' h /-- `ℵ₀` is the smallest infinite cardinal. -/ def aleph_0 : cardinal.{u} := lift (#ℕ) -localized "notation `ℵ₀` := cardinal.aleph_0" in cardinal +localized "notation (name := cardinal.aleph_0) `ℵ₀` := cardinal.aleph_0" in cardinal lemma mk_nat : #ℕ = ℵ₀ := (lift_id _).symm diff --git a/src/set_theory/cardinal/cofinality.lean b/src/set_theory/cardinal/cofinality.lean index 5aaafe6425b09..595e1b1fcf498 100644 --- a/src/set_theory/cardinal/cofinality.lean +++ b/src/set_theory/cardinal/cofinality.lean @@ -445,8 +445,8 @@ variables {a o : ordinal.{u}} {f : Π b < o, ordinal.{u}} protected theorem cof_eq (hf : is_fundamental_sequence a o f) : a.cof.ord = o := hf.1.antisymm' $ by { rw ←hf.2.2, exact (ord_le_ord.2 (cof_blsub_le f)).trans (ord_card_le o) } -protected theorem strict_mono (hf : is_fundamental_sequence a o f) : - ∀ {i j} (hi) (hj), i < j → f i hi < f j hj := +protected theorem strict_mono (hf : is_fundamental_sequence a o f) {i j} : + ∀ hi hj, i < j → f i hi < f j hj := hf.2.1 theorem blsub_eq (hf : is_fundamental_sequence a o f) : blsub.{u u} o f = a := @@ -710,7 +710,7 @@ end ordinal namespace cardinal open ordinal -local infixr ^ := @pow cardinal.{u} cardinal cardinal.has_pow +local infixr (name := cardinal.pow) ^ := @pow cardinal.{u} cardinal cardinal.has_pow /-- A cardinal is a limit if it is not zero or a successor cardinal. Note that `ℵ₀` is a limit cardinal by this definition. -/ diff --git a/src/set_theory/cardinal/continuum.lean b/src/set_theory/cardinal/continuum.lean index 550db497dc7e7..acd366fac23b5 100644 --- a/src/set_theory/cardinal/continuum.lean +++ b/src/set_theory/cardinal/continuum.lean @@ -25,7 +25,7 @@ open_locale cardinal /-- Cardinality of continuum. -/ def continuum : cardinal.{u} := 2 ^ aleph_0.{u} -localized "notation `𝔠` := cardinal.continuum" in cardinal +localized "notation (name := cardinal.continuum) `𝔠` := cardinal.continuum" in cardinal @[simp] lemma two_power_aleph_0 : 2 ^ aleph_0.{u} = continuum.{u} := rfl diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 7b893d6bdeefa..e9bf0c246ede7 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -281,7 +281,7 @@ instance : has_le pgame := ⟨λ x y, (le_lf x y).1⟩ If `0 ⧏ x`, then Left can win `x` as the first player. -/ def lf (x y : pgame) : Prop := (le_lf x y).2 -localized "infix ` ⧏ `:50 := pgame.lf" in pgame +localized "infix (name := pgame.lf) ` ⧏ `:50 := pgame.lf" in pgame /-- Definition of `x ≤ y` on pre-games built using the constructor. -/ @[simp] theorem mk_le_mk {xl xr xL xR yl yr yL yR} : @@ -510,7 +510,7 @@ classical.some_spec $ (zero_le.1 h) j If `x ≈ 0`, then the second player can always win `x`. -/ def equiv (x y : pgame) : Prop := x ≤ y ∧ y ≤ x -localized "infix ` ≈ ` := pgame.equiv" in pgame +localized "infix (name := pgame.equiv) ` ≈ ` := pgame.equiv" in pgame instance : is_equiv _ (≈) := { refl := λ x, ⟨le_rfl, le_rfl⟩, @@ -608,7 +608,7 @@ end If `x ∥ 0`, then the first player can always win `x`. -/ def fuzzy (x y : pgame) : Prop := x ⧏ y ∧ y ⧏ x -localized "infix ` ∥ `:50 := pgame.fuzzy" in pgame +localized "infix (name := pgame.fuzzy) ` ∥ `:50 := pgame.fuzzy" in pgame @[symm] theorem fuzzy.swap {x y : pgame} : x ∥ y → y ∥ x := and.swap instance : is_symm _ (∥) := ⟨λ x y, fuzzy.swap⟩ @@ -708,7 +708,7 @@ inductive relabelling : pgame.{u} → pgame.{u} → Type (u+1) (∀ j, relabelling (x.move_right j) (y.move_right (R j))) → relabelling x y -localized "infix ` ≡r `:50 := pgame.relabelling" in pgame +localized "infix (name := pgame.relabelling) ` ≡r `:50 := pgame.relabelling" in pgame namespace relabelling variables {x y : pgame.{u}} diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 5bfc553e3f012..4a49974afe067 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -1790,7 +1790,7 @@ end instance : has_pow ordinal ordinal := ⟨λ a b, if a = 0 then 1 - b else limit_rec_on b 1 (λ _ IH, IH * a) (λ b _, bsup.{u u} b)⟩ -local infixr ^ := @pow ordinal ordinal ordinal.has_pow +local infixr (name := ordinal.pow) ^ := @pow ordinal ordinal ordinal.has_pow theorem opow_def (a b : ordinal) : a ^ b = if a = 0 then 1 - b else limit_rec_on b 1 (λ _ IH, IH * a) (λ b _, bsup.{u u} b) := @@ -2351,7 +2351,7 @@ begin { exact (mul_is_normal ho).apply_omega } end -local infixr ^ := @pow ordinal ordinal ordinal.has_pow +local infixr (name := ordinal.pow) ^ := @pow ordinal ordinal ordinal.has_pow theorem sup_opow_nat {o : ordinal} (ho : 0 < o) : sup (λ n : ℕ, o ^ n) = o ^ ω := begin rcases lt_or_eq_of_le (one_le_iff_pos.2 ho) with ho₁ | rfl, diff --git a/src/set_theory/ordinal/basic.lean b/src/set_theory/ordinal/basic.lean index 87aae5edf6b63..6180ef4ff8fcb 100644 --- a/src/set_theory/ordinal/basic.lean +++ b/src/set_theory/ordinal/basic.lean @@ -572,7 +572,7 @@ def lift.initial_seg : @initial_seg ordinal.{u} ordinal.{max u v} (<) (<) := /-- `ω` is the first infinite ordinal, defined as the order type of `ℕ`. -/ def omega : ordinal.{u} := lift $ @type ℕ (<) _ -localized "notation `ω` := ordinal.omega" in ordinal +localized "notation (name := ordinal.omega) `ω` := ordinal.omega" in ordinal /-- Note that the presence of this lemma makes `simp [omega]` form a loop. -/ @[simp] theorem type_nat_lt : @type ℕ (<) _ = ω := (lift_id _).symm diff --git a/src/set_theory/ordinal/fixed_point.lean b/src/set_theory/ordinal/fixed_point.lean index 5d54671f85016..a8c1e5039b5b9 100644 --- a/src/set_theory/ordinal/fixed_point.lean +++ b/src/set_theory/ordinal/fixed_point.lean @@ -495,7 +495,7 @@ end /-! ### Fixed points of multiplication -/ -local infixr ^ := @pow ordinal ordinal ordinal.has_pow +local infixr (name := ordinal.pow) ^ := @pow ordinal ordinal ordinal.has_pow @[simp] theorem nfp_mul_one {a : ordinal} (ha : 0 < a) : nfp ((*) a) 1 = a ^ omega := begin rw [←sup_iterate_eq_nfp, ←sup_opow_nat], diff --git a/src/set_theory/ordinal/natural_ops.lean b/src/set_theory/ordinal/natural_ops.lean index c57041816c21d..d44b0193ea09c 100644 --- a/src/set_theory/ordinal/natural_ops.lean +++ b/src/set_theory/ordinal/natural_ops.lean @@ -120,7 +120,7 @@ noncomputable def nadd : ordinal → ordinal → ordinal (blsub.{u u} b $ λ b' h, nadd a b') using_well_founded { dec_tac := `[solve_by_elim [psigma.lex.left, psigma.lex.right]] } -localized "infix ` ♯ `:65 := ordinal.nadd" in natural_ops +localized "infix (name := ordinal.nadd) ` ♯ `:65 := ordinal.nadd" in natural_ops theorem nadd_def (a b : ordinal) : a ♯ b = max (blsub.{u u} a $ λ a' h, a' ♯ b) diff --git a/src/set_theory/ordinal/notation.lean b/src/set_theory/ordinal/notation.lean index 1d81c754c5b62..1e06926c5a76a 100644 --- a/src/set_theory/ordinal/notation.lean +++ b/src/set_theory/ordinal/notation.lean @@ -695,7 +695,7 @@ begin end section -local infixr ^ := @pow ordinal.{0} ordinal ordinal.has_pow +local infixr (name := ordinal.pow) ^ := @pow ordinal.{0} ordinal ordinal.has_pow theorem repr_opow_aux₂ {a0 a'} [N0 : NF a0] [Na' : NF a'] (m : ℕ) (d : ω ∣ repr a') @@ -828,7 +828,7 @@ begin exact ⟨i, h'.trans_le (le_add_right _ _)⟩ end -local infixr ^ := @pow ordinal ordinal ordinal.has_pow +local infixr (name := ordinal.pow) ^ := @pow ordinal ordinal ordinal.has_pow private theorem exists_lt_omega_opow' {α} {o b : ordinal} (hb : 1 < b) (ho : o.is_limit) {f : α → ordinal} (H : ∀ ⦃a⦄, a < o → ∃ i, a < f i) ⦃a⦄ (h : a < b ^ o) : ∃ i, a < b ^ f i := @@ -951,7 +951,7 @@ begin end section -local infixr ^ := pow +local infixr (name := pow) ^ := pow @[simp] theorem fast_growing_two : fast_growing 2 = (λ n, 2 ^ n * n) := begin rw [@fast_growing_succ 2 1 rfl], funext i, rw [fast_growing_one], diff --git a/src/set_theory/ordinal/principal.lean b/src/set_theory/ordinal/principal.lean index f3a07a8808b1a..672c33df46ba9 100644 --- a/src/set_theory/ordinal/principal.lean +++ b/src/set_theory/ordinal/principal.lean @@ -32,7 +32,7 @@ noncomputable theory open order namespace ordinal -local infixr ^ := @pow ordinal ordinal ordinal.has_pow +local infixr (name := ordinal.pow) ^ := @pow ordinal ordinal ordinal.has_pow /-! ### Principal ordinals -/ diff --git a/src/set_theory/surreal/dyadic.lean b/src/set_theory/surreal/dyadic.lean index 4a7c24ed6d7b8..3c8884e57e637 100644 --- a/src/set_theory/surreal/dyadic.lean +++ b/src/set_theory/surreal/dyadic.lean @@ -26,7 +26,7 @@ rational numbers to construct an ordered field embedding of ℝ into `surreal`. universes u -local infix ` ≈ ` := pgame.equiv +local infix (name := pgame.equiv) ` ≈ ` := pgame.equiv namespace pgame diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 7b7600f04ced7..3447accd70398 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -281,7 +281,7 @@ def powerset (x : pSet) : pSet := ⟨set x.type, λ p, ⟨{a // p a}, λ y, x.fu /-- The pre-set union operator -/ def sUnion (a : pSet) : pSet := ⟨Σ x, (a.func x).type, λ ⟨x, y⟩, (a.func x).func y⟩ -prefix `⋃₀ `:110 := pSet.sUnion +prefix (name := pSet.sUnion) `⋃₀ `:110 := pSet.sUnion @[simp] theorem mem_sUnion : Π {x y : pSet.{u}}, y ∈ ⋃₀ x ↔ ∃ z ∈ x, y ∈ z | ⟨α, A⟩ y := @@ -621,7 +621,7 @@ resp.eval 1 ⟨pSet.sUnion, λ ⟨α, A⟩ ⟨β, B⟩ ⟨αβ, βα⟩, ⟨sUnion_lem A B αβ, λ a, exists.elim (sUnion_lem B A (λ b, exists.elim (βα b) (λ c hc, ⟨c, pSet.equiv.symm hc⟩)) a) (λ b hb, ⟨b, pSet.equiv.symm hb⟩)⟩⟩ -prefix `⋃₀ `:110 := Set.sUnion +prefix (name := Set.sUnion) `⋃₀ `:110 := Set.sUnion @[simp] theorem mem_sUnion {x y : Set.{u}} : y ∈ ⋃₀ x ↔ ∃ z ∈ x, y ∈ z := quotient.induction_on₂ x y (λ x y, iff.trans mem_sUnion @@ -873,7 +873,7 @@ def powerset (x : Class) : Class := Cong_to_Class (set.powerset x) /-- The union of a class is the class of all members of ZFC sets in the class -/ def sUnion (x : Class) : Class := ⋃₀ (Class_to_Cong x) -prefix `⋃₀ `:110 := Class.sUnion +prefix (name := Class.sUnion) `⋃₀ `:110 := Class.sUnion theorem of_Set.inj {x y : Set.{u}} (h : (x : Class.{u}) = y) : x = y := Set.ext $ λ z, by { change (x : Class.{u}) z ↔ (y : Class.{u}) z, rw h } diff --git a/src/tactic/converter/apply_congr.lean b/src/tactic/converter/apply_congr.lean index 5074c3dd99213..cf77f977cfb2d 100644 --- a/src/tactic/converter/apply_congr.lean +++ b/src/tactic/converter/apply_congr.lean @@ -20,7 +20,7 @@ open tactic namespace conv.interactive open interactive interactive.types lean.parser -local postfix `?`:9001 := optional +local postfix (name := parser.optional) `?`:9001 := optional /-- diff --git a/src/tactic/core.lean b/src/tactic/core.lean index e3e86b4369906..7524e2a649cc3 100644 --- a/src/tactic/core.lean +++ b/src/tactic/core.lean @@ -1962,8 +1962,8 @@ open _root_.lean open _root_.lean.parser open _root_.interactive _root_.interactive.types -local postfix `?`:9001 := optional -local postfix *:9001 := many . +local postfix (name := parser.optional) `?`:9001 := optional +local postfix (name := parser.many) *:9001 := many . " /-- `finally tac finalizer` runs `tac` first, then runs `finalizer` even if diff --git a/src/tactic/ext.lean b/src/tactic/ext.lean index 643af928ea7fe..f7e422bc6ed89 100644 --- a/src/tactic/ext.lean +++ b/src/tactic/ext.lean @@ -445,8 +445,8 @@ do ⟨_, σ⟩ ← state_t.run (ext_core cfg) {patts := xs, fuel := fuel}, when trace $ tactic.trace $ "Try this: " ++ ", ".intercalate σ.trace_msg, pure σ.patts -local postfix `?`:9001 := optional -local postfix *:9001 := many +local postfix (name := parser.optional) `?`:9001 := optional +local postfix (name := parser.many) *:9001 := many /-- `ext1 id` selects and apply one extensionality lemma (with attribute diff --git a/src/tactic/generalize_proofs.lean b/src/tactic/generalize_proofs.lean index 205f46294d8a6..1f380aecab555 100644 --- a/src/tactic/generalize_proofs.lean +++ b/src/tactic/generalize_proofs.lean @@ -75,7 +75,7 @@ do intros_dep, collect_proofs_in t [] (ns, hs), intron n <|> (intros $> ()) -local postfix *:9001 := many +local postfix (name := parser.many) *:9001 := many namespace interactive /-- Generalize proofs in the goal, naming them with the provided list. diff --git a/src/tactic/interval_cases.lean b/src/tactic/interval_cases.lean index 424fb3c2fe9dc..d923e76a6d4bc 100644 --- a/src/tactic/interval_cases.lean +++ b/src/tactic/interval_cases.lean @@ -206,7 +206,7 @@ setup_tactic_parser namespace interactive -local postfix `?`:9001 := optional +local postfix (name := parser.optional) `?`:9001 := optional /-- `interval_cases n` searches for upper and lower bounds on a variable `n`, diff --git a/src/tactic/localized.lean b/src/tactic/localized.lean index 04b6968244af2..c6d2b77820d26 100644 --- a/src/tactic/localized.lean +++ b/src/tactic/localized.lean @@ -64,11 +64,11 @@ locale. * Declare notation which is localized to a locale using: ```lean -localized "infix ` ⊹ `:60 := my_add" in my.add +localized "infix (name := my_add) ` ⊹ `:60 := my_add" in my.add ``` * After this command it will be available in the same section/namespace/file, just as if you wrote - `local infix ` ⊹ `:60 := my_add` + `local infix (name := my_add) ` ⊹ `:60 := my_add` * You can open it in other places. The following command will declare the notation again as local notation in that section/namespace/file: @@ -100,6 +100,20 @@ run_cmd do * Warning: You have to give full names of all declarations used in localized notation, so that the localized notation also works when the appropriate namespaces are not opened. + +* Note: In mathlib, you should always provide names for localized notations using the + `(name := ...)` parameter. This prevents issues if the localized notation overrides + an existing notation when it gets opened. + +* Warning: Due to limitations in the implementation, you cannot use `_` in localized notations. + (Otherwise `open_locale foo` will fail if `foo` is already opened or partially opened.) + Instead, you should use the `hole!` notation as a drop-in replacement. For example: +```lean +-- BAD +-- localized "infix (name := my_add) ` ⊹[` R `] ` := my_add _ R" in foo +-- GOOD +localized "infix (name := my_add) ` ⊹[` R `] ` := my_add hole! R" in foo +``` -/ add_tactic_doc { name := "localized notation", @@ -111,6 +125,10 @@ add_tactic_doc meta def print_localized_commands (ns : list name) : tactic unit := do cmds ← get_localized ns, cmds.mmap' trace +-- This should be used instead of `_` inside localized commands, +-- because otherwise `open_locale` will fail if some of the notations are already available. +notation `hole!` := _ + -- you can run `open_locale classical` to get the decidability of all propositions, and downgrade -- the priority of decidability instances that make Lean run through all the algebraic hierarchy -- whenever it wants to solve a decidability question @@ -118,5 +136,5 @@ localized "attribute [instance, priority 9] classical.prop_decidable" in classic localized "attribute [instance, priority 8] eq.decidable" in classical -localized "postfix `?`:9001 := optional" in parser -localized "postfix *:9001 := lean.parser.many" in parser +localized "postfix (name := parser.optional) `?`:9001 := optional" in parser +localized "postfix (name := parser.many) *:9001 := lean.parser.many" in parser diff --git a/src/tactic/monotonicity/interactive.lean b/src/tactic/monotonicity/interactive.lean index 7d0091697c6eb..b7bf3eaf8adec 100644 --- a/src/tactic/monotonicity/interactive.lean +++ b/src/tactic/monotonicity/interactive.lean @@ -16,8 +16,8 @@ open lean lean.parser interactive open interactive.types open tactic -local postfix `?`:9001 := optional -local postfix *:9001 := many +local postfix (name := parser.optional) `?`:9001 := optional +local postfix (name := parser.many) *:9001 := many meta inductive mono_function (elab : bool := tt) | non_assoc : expr elab → list (expr elab) → list (expr elab) → mono_function diff --git a/src/tactic/omega/int/form.lean b/src/tactic/omega/int/form.lean index 45ea52ad6b193..7f839dcb4438d 100644 --- a/src/tactic/omega/int/form.lean +++ b/src/tactic/omega/int/form.lean @@ -30,11 +30,11 @@ inductive preform | or : preform → preform → preform | and : preform → preform → preform -localized "notation x ` =* ` y := omega.int.preform.eq x y" in omega.int -localized "notation x ` ≤* ` y := omega.int.preform.le x y" in omega.int -localized "notation `¬* ` p := omega.int.preform.not p" in omega.int -localized "notation p ` ∨* ` q := omega.int.preform.or p q" in omega.int -localized "notation p ` ∧* ` q := omega.int.preform.and p q" in omega.int +localized "notation (name := preform.eq) x ` =* ` y := omega.int.preform.eq x y" in omega.int +localized "notation (name := preform.le) x ` ≤* ` y := omega.int.preform.le x y" in omega.int +localized "notation (name := preform.not) `¬* ` p := omega.int.preform.not p" in omega.int +localized "notation (name := preform.or) p ` ∨* ` q := omega.int.preform.or p q" in omega.int +localized "notation (name := preform.and) p ` ∧* ` q := omega.int.preform.and p q" in omega.int namespace preform diff --git a/src/tactic/omega/int/preterm.lean b/src/tactic/omega/int/preterm.lean index 1f86ddf74e3f8..dafc9fc939d60 100644 --- a/src/tactic/omega/int/preterm.lean +++ b/src/tactic/omega/int/preterm.lean @@ -32,9 +32,9 @@ inductive preterm : Type | var : int → nat → preterm | add : preterm → preterm → preterm -localized "notation `&` k := omega.int.preterm.cst k" in omega.int -localized "infix ` ** ` : 300 := omega.int.preterm.var" in omega.int -localized "notation t `+*` s := omega.int.preterm.add t s" in omega.int +localized "notation (name := preterm.cst) `&` k := omega.int.preterm.cst k" in omega.int +localized "infix (name := preterm.var) ` ** ` : 300 := omega.int.preterm.var" in omega.int +localized "notation (name := preterm.add) t `+*` s := omega.int.preterm.add t s" in omega.int namespace preterm diff --git a/src/tactic/omega/misc.lean b/src/tactic/omega/misc.lean index b6ebe964c9c8a..adb61ec932185 100644 --- a/src/tactic/omega/misc.lean +++ b/src/tactic/omega/misc.lean @@ -31,7 +31,7 @@ lemma pred_mono_2' {c : Prop → Prop → Prop} {a1 a2 b1 b2 : Prop} : def update (m : nat) (a : α) (v : nat → α) : nat → α | n := if n = m then a else v n -localized "notation v ` ⟨` m ` ↦ ` a `⟩` := omega.update m a v" in omega +localized "notation (name := omega.update) v ` ⟨` m ` ↦ ` a `⟩` := omega.update m a v" in omega lemma update_eq (m : nat) (a : α) (v : nat → α) : (v ⟨m ↦ a⟩) m = a := by simp only [update, if_pos rfl] diff --git a/src/tactic/omega/nat/form.lean b/src/tactic/omega/nat/form.lean index 9c08991fe0c51..cd4bfb5b6e4f6 100644 --- a/src/tactic/omega/nat/form.lean +++ b/src/tactic/omega/nat/form.lean @@ -31,11 +31,11 @@ inductive preform | or : preform → preform → preform | and : preform → preform → preform -localized "notation x ` =* ` y := omega.nat.preform.eq x y" in omega.nat -localized "notation x ` ≤* ` y := omega.nat.preform.le x y" in omega.nat -localized "notation `¬* ` p := omega.nat.preform.not p" in omega.nat -localized "notation p ` ∨* ` q := omega.nat.preform.or p q" in omega.nat -localized "notation p ` ∧* ` q := omega.nat.preform.and p q" in omega.nat +localized "notation (name := preform.eq) x ` =* ` y := omega.nat.preform.eq x y" in omega.nat +localized "notation (name := preform.le) x ` ≤* ` y := omega.nat.preform.le x y" in omega.nat +localized "notation (name := preform.not) `¬* ` p := omega.nat.preform.not p" in omega.nat +localized "notation (name := preform.or) p ` ∨* ` q := omega.nat.preform.or p q" in omega.nat +localized "notation (name := preform.and) p ` ∧* ` q := omega.nat.preform.and p q" in omega.nat namespace preform diff --git a/src/tactic/omega/nat/preterm.lean b/src/tactic/omega/nat/preterm.lean index 2343c6361aded..a0ef3c187fda2 100644 --- a/src/tactic/omega/nat/preterm.lean +++ b/src/tactic/omega/nat/preterm.lean @@ -36,10 +36,10 @@ inductive preterm : Type | add : preterm → preterm → preterm | sub : preterm → preterm → preterm -localized "notation `&` k := omega.nat.preterm.cst k" in omega.nat -localized "infix ` ** ` : 300 := omega.nat.preterm.var" in omega.nat -localized "notation t ` +* ` s := omega.nat.preterm.add t s" in omega.nat -localized "notation t ` -* ` s := omega.nat.preterm.sub t s" in omega.nat +localized "notation (name := preterm.cst) `&` k := omega.nat.preterm.cst k" in omega.nat +localized "infix (name := preterm.var) ` ** ` : 300 := omega.nat.preterm.var" in omega.nat +localized "notation (name := preterm.add) t ` +* ` s := omega.nat.preterm.add t s" in omega.nat +localized "notation (name := preterm.sub) t ` -* ` s := omega.nat.preterm.sub t s" in omega.nat namespace preterm diff --git a/src/tactic/push_neg.lean b/src/tactic/push_neg.lean index 3d81811a0fc87..0feea63b832a0 100644 --- a/src/tactic/push_neg.lean +++ b/src/tactic/push_neg.lean @@ -114,8 +114,8 @@ end push_neg open interactive (parse loc.ns loc.wildcard) open interactive.types (location texpr) open lean.parser (tk ident many) interactive.loc -local postfix `?`:9001 := optional -local postfix *:9001 := many +local postfix (name := parser.optional) `?`:9001 := optional +local postfix (name := parser.many) *:9001 := many open push_neg /-- diff --git a/src/tactic/rcases.lean b/src/tactic/rcases.lean index 68263d166d7cb..07a6c8707d507 100644 --- a/src/tactic/rcases.lean +++ b/src/tactic/rcases.lean @@ -878,8 +878,6 @@ add_tactic_doc tags := ["induction"], inherit_description_from := `tactic.interactive.rintro } -setup_tactic_parser - /-- Parses `patt? (: expr)? (:= expr)?`, the arguments for `obtain`. (This is almost the same as `rcases_patt_parse`, but it allows the pattern part to be empty.) -/ diff --git a/src/tactic/reassoc_axiom.lean b/src/tactic/reassoc_axiom.lean index 3f07919e3d73f..100b4a130be59 100644 --- a/src/tactic/reassoc_axiom.lean +++ b/src/tactic/reassoc_axiom.lean @@ -172,8 +172,6 @@ add_tactic_doc namespace interactive -setup_tactic_parser - /-- `reassoc h`, for assumption `h : x ≫ y = z`, creates a new assumption `h : ∀ {W} (f : Z ⟶ W), x ≫ y ≫ f = z ≫ f`. `reassoc! h`, does the same but deletes the initial `h` assumption. diff --git a/src/tactic/ring.lean b/src/tactic/ring.lean index 8aff0fdd568ab..db7ece203857b 100644 --- a/src/tactic/ring.lean +++ b/src/tactic/ring.lean @@ -708,7 +708,7 @@ open conv interactive open tactic tactic.interactive (ring.mode ring1) open tactic.ring (normalize normalize_mode.horner) -local postfix `?`:9001 := optional +local postfix (name := parser.optional) `?`:9001 := optional /-- Normalises expressions in commutative (semi-)rings inside of a `conv` block using the tactic `ring`. diff --git a/src/tactic/ring2.lean b/src/tactic/ring2.lean index e035886136b18..044ea34dd20a9 100644 --- a/src/tactic/ring2.lean +++ b/src/tactic/ring2.lean @@ -475,7 +475,7 @@ namespace interactive open interactive interactive.types lean.parser open tactic.ring2 -local postfix `?`:9001 := optional +local postfix (name := parser.optional) `?`:9001 := optional /-- `ring2` solves equations in the language of rings. diff --git a/src/tactic/ring_exp.lean b/src/tactic/ring_exp.lean index 42f2f55f91f72..c4c6e06c08079 100644 --- a/src/tactic/ring_exp.lean +++ b/src/tactic/ring_exp.lean @@ -1477,7 +1477,7 @@ end tactic.ring_exp namespace tactic.interactive open interactive interactive.types lean.parser tactic tactic.ring_exp -local postfix `?`:9001 := optional +local postfix (name := parser.optional) `?`:9001 := optional /-- Tactic for solving equations of *commutative* (semi)rings, @@ -1543,7 +1543,7 @@ open conv interactive open tactic tactic.interactive (ring_exp_eq) open tactic.ring_exp (normalize) -local postfix `?`:9001 := optional +local postfix (name := parser.optional) `?`:9001 := optional /-- Normalises expressions in commutative (semi-)rings inside of a `conv` block using the tactic diff --git a/src/tactic/simpa.lean b/src/tactic/simpa.lean index 2f90df9cf89d1..0d0ab4a381357 100644 --- a/src/tactic/simpa.lean +++ b/src/tactic/simpa.lean @@ -12,7 +12,7 @@ namespace tactic namespace interactive open expr lean.parser -local postfix `?`:9001 := optional +local postfix (name := parser.optional) `?`:9001 := optional /-- This is a "finishing" tactic modification of `simp`. It has two forms. diff --git a/src/tactic/tauto.lean b/src/tactic/tauto.lean index a224e560f85d5..7bd627e1dec77 100644 --- a/src/tactic/tauto.lean +++ b/src/tactic/tauto.lean @@ -238,7 +238,7 @@ meta def tautology (cfg : tauto_cfg := {}) : tactic unit := focus1 $ repeat (first basic_tauto_tacs); cfg.closer, done namespace interactive -local postfix `?`:9001 := optional +local postfix (name := parser.optional) `?`:9001 := optional setup_tactic_parser /-- diff --git a/src/topology/alexandroff.lean b/src/topology/alexandroff.lean index 6dacac3369e96..dac07a9dacc39 100644 --- a/src/topology/alexandroff.lean +++ b/src/topology/alexandroff.lean @@ -55,7 +55,7 @@ namespace alexandroff /-- The point at infinity -/ def infty : alexandroff X := none -localized "notation `∞` := alexandroff.infty" in alexandroff +localized "notation (name := alexandroff.infty) `∞` := alexandroff.infty" in alexandroff instance : has_coe_t X (alexandroff X) := ⟨option.some⟩ diff --git a/src/topology/basic.lean b/src/topology/basic.lean index ee11d73197e6e..4f5fb87d0eb77 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -641,18 +641,24 @@ neighborhoods of `a` forms a filter, the neighborhood filter at `a`, is here def infimum over the principal filters of all open sets containing `a`. -/ @[irreducible] def nhds (a : α) : filter α := (⨅ s ∈ {s : set α | a ∈ s ∧ is_open s}, 𝓟 s) -localized "notation `𝓝` := nhds" in topological_space +localized "notation (name := nhds) `𝓝` := nhds" in topological_space /-- The "neighborhood within" filter. Elements of `𝓝[s] a` are sets containing the intersection of `s` and a neighborhood of `a`. -/ def nhds_within (a : α) (s : set α) : filter α := 𝓝 a ⊓ 𝓟 s -localized "notation `𝓝[` s `] ` x:100 := nhds_within x s" in topological_space -localized "notation `𝓝[≠] ` x:100 := nhds_within x {x}ᶜ" in topological_space -localized "notation `𝓝[≥] ` x:100 := nhds_within x (set.Ici x)" in topological_space -localized "notation `𝓝[≤] ` x:100 := nhds_within x (set.Iic x)" in topological_space -localized "notation `𝓝[>] ` x:100 := nhds_within x (set.Ioi x)" in topological_space -localized "notation `𝓝[<] ` x:100 := nhds_within x (set.Iio x)" in topological_space +localized "notation (name := nhds_within) + `𝓝[` s `] ` x:100 := nhds_within x s" in topological_space +localized "notation (name := nhds_within.ne) + `𝓝[≠] ` x:100 := nhds_within x {x}ᶜ" in topological_space +localized "notation (name := nhds_within.ge) + `𝓝[≥] ` x:100 := nhds_within x (set.Ici x)" in topological_space +localized "notation (name := nhds_within.le) + `𝓝[≤] ` x:100 := nhds_within x (set.Iic x)" in topological_space +localized "notation (name := nhds_within.gt) + `𝓝[>] ` x:100 := nhds_within x (set.Ioi x)" in topological_space +localized "notation (name := nhds_within.lt) + `𝓝[<] ` x:100 := nhds_within x (set.Iio x)" in topological_space lemma nhds_def (a : α) : 𝓝 a = (⨅ s ∈ {s : set α | a ∈ s ∧ is_open s}, 𝓟 s) := by rw nhds diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index ca045fbf9df68..f5328761a9e29 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -37,7 +37,8 @@ structure bounded_continuous_function (α : Type u) (β : Type v) Type (max u v) := (map_bounded' : ∃ C, ∀ x y, dist (to_fun x) (to_fun y) ≤ C) -localized "infixr ` →ᵇ `:25 := bounded_continuous_function" in bounded_continuous_function +localized "infixr (name := bounded_continuous_function) + ` →ᵇ `:25 := bounded_continuous_function" in bounded_continuous_function /-- `bounded_continuous_map_class F α β` states that `F` is a type of bounded continuous maps. diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index 96c93b4f6a7fd..c41cd5e5196c4 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -39,9 +39,10 @@ structure zero_at_infty_continuous_map (α : Type u) (β : Type v) Type (max u v) := (zero_at_infty' : tendsto to_fun (cocompact α) (𝓝 0)) -localized "notation [priority 2000] `C₀(` α `, ` β `)` := zero_at_infty_continuous_map α β" - in zero_at_infty -localized "notation α ` →C₀ ` β := zero_at_infty_continuous_map α β" in zero_at_infty +localized "notation [priority 2000] (name := zero_at_infty_continuous_map) + `C₀(` α `, ` β `)` := zero_at_infty_continuous_map α β" in zero_at_infty +localized "notation (name := zero_at_infty_continuous_map.arrow) + α ` →C₀ ` β := zero_at_infty_continuous_map α β" in zero_at_infty /-- `zero_at_infty_continuous_map_class F α β` states that `F` is a type of continuous maps which vanish at infinity. diff --git a/src/topology/homotopy/equiv.lean b/src/topology/homotopy/equiv.lean index 1e353c25bcae3..641495cefeb98 100644 --- a/src/topology/homotopy/equiv.lean +++ b/src/topology/homotopy/equiv.lean @@ -44,7 +44,8 @@ structure homotopy_equiv (X : Type u) (Y : Type v) [topological_space X] [topolo (left_inv : (inv_fun.comp to_fun).homotopic (continuous_map.id X)) (right_inv : (to_fun.comp inv_fun).homotopic (continuous_map.id Y)) -localized "infix ` ≃ₕ `:25 := continuous_map.homotopy_equiv" in continuous_map +localized "infix (name := continuous_map.homotopy_equiv) + ` ≃ₕ `:25 := continuous_map.homotopy_equiv" in continuous_map namespace homotopy_equiv diff --git a/src/topology/nhds_set.lean b/src/topology/nhds_set.lean index 9b4dded88b23a..907437b3152ed 100644 --- a/src/topology/nhds_set.lean +++ b/src/topology/nhds_set.lean @@ -33,7 +33,7 @@ variables {α β : Type*} [topological_space α] [topological_space β] def nhds_set (s : set α) : filter α := Sup (nhds '' s) -localized "notation `𝓝ˢ` := nhds_set" in topological_space +localized "notation (name := nhds_set) `𝓝ˢ` := nhds_set" in topological_space lemma mem_nhds_set_iff_forall : s ∈ 𝓝ˢ t ↔ ∀ (x : α), x ∈ t → s ∈ 𝓝 x := by simp_rw [nhds_set, filter.mem_Sup, ball_image_iff] diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index c01f3f0b6213d..c63b90d52d037 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -132,7 +132,7 @@ by simp [subset_def]; exact forall_congr (λ a, by simp) /-- The composition of relations -/ def comp_rel {α : Type u} (r₁ r₂ : set (α×α)) := {p : α × α | ∃z:α, (p.1, z) ∈ r₁ ∧ (z, p.2) ∈ r₂} -localized "infix ` ○ `:55 := comp_rel" in uniformity +localized "infix (name := uniformity.comp_rel) ` ○ `:55 := comp_rel" in uniformity @[simp] theorem mem_comp_rel {r₁ r₂ : set (α×α)} {x y : α} : (x, y) ∈ r₁ ○ r₂ ↔ ∃ z, (x, z) ∈ r₁ ∧ (z, y) ∈ r₂ := iff.rfl @@ -317,7 +317,7 @@ variables [uniform_space α] def uniformity (α : Type u) [uniform_space α] : filter (α × α) := (@uniform_space.to_core α _).uniformity -localized "notation `𝓤` := uniformity" in uniformity +localized "notation (name := uniformity) `𝓤` := uniformity" in uniformity lemma is_open_uniformity {s : set α} : is_open s ↔ (∀x∈s, { p : α × α | p.1 = x → p.2 ∈ s } ∈ 𝓤 α) := diff --git a/src/topology/uniform_space/separation.lean b/src/topology/uniform_space/separation.lean index 4c5393d07fdac..8ea80cca1bf87 100644 --- a/src/topology/uniform_space/separation.lean +++ b/src/topology/uniform_space/separation.lean @@ -87,7 +87,7 @@ variables [uniform_space α] [uniform_space β] [uniform_space γ] protected def separation_rel (α : Type u) [u : uniform_space α] := ⋂₀ (𝓤 α).sets -localized "notation `𝓢` := separation_rel" in uniformity +localized "notation (name := separation_rel) `𝓢` := separation_rel" in uniformity lemma separated_equiv : equivalence (λx y, (x, y) ∈ 𝓢 α) := ⟨assume x, assume s, refl_mem_uniformity, diff --git a/src/topology/unit_interval.lean b/src/topology/unit_interval.lean index 5c7e1f8d7e72e..0618eb9204244 100644 --- a/src/topology/unit_interval.lean +++ b/src/topology/unit_interval.lean @@ -27,7 +27,7 @@ open set int set.Icc /-- The unit interval `[0,1]` in ℝ. -/ abbreviation unit_interval : set ℝ := set.Icc 0 1 -localized "notation `I` := unit_interval" in unit_interval +localized "notation (name := unit_interval) `I` := unit_interval" in unit_interval namespace unit_interval @@ -74,7 +74,7 @@ subtype.coe_le_coe.mp $ (mul_le_mul_of_nonneg_right x.2.2 y.2.1).trans_eq $ one_ /-- Unit interval central symmetry. -/ def symm : I → I := λ t, ⟨1 - t, mem_iff_one_sub_mem.mp t.prop⟩ -localized "notation `σ` := unit_interval.symm" in unit_interval +localized "notation (name := unit_interval.symm) `σ` := unit_interval.symm" in unit_interval @[simp] lemma symm_zero : σ 0 = 1 := subtype.ext $ by simp [symm] diff --git a/test/induction.lean b/test/induction.lean index 8cbb43218df8c..ddd471dc6d49e 100644 --- a/test/induction.lean +++ b/test/induction.lean @@ -1080,8 +1080,8 @@ begin exact small_step.while, } end -infixr ` ⇒ ` := small_step -infixr ` ⇒* ` : 100 := star small_step +infixr (name := small_step) ` ⇒ ` := small_step +infixr (name := small_step.star) ` ⇒* ` : 100 := star small_step /- More lemmas about big-step and small-step semantics. These are taken from the diff --git a/test/localized/localized.lean b/test/localized/localized.lean index 70c188719f12b..07fad09ff60b7 100644 --- a/test/localized/localized.lean +++ b/test/localized/localized.lean @@ -10,9 +10,9 @@ example : 2 ↓ 3 = 8 := rfl example : 2 ⊖ 3 = 8 := rfl example {n m : ℕ} (h : n < m) : n ≤ m := by { success_if_fail { simp [h] }, exact le_of_lt h } section -localized "infix ` ⊹ `:59 := nat.add" in nat -localized "infix ` ↓ `:59 := nat.mul" in nat -localized "infix ` ⊖ `:59 := nat.mul" in nat.mul +localized "infix (name := plus) ` ⊹ `:59 := nat.add" in nat +localized "infix (name := down) ` ↓ `:59 := nat.mul" in nat +localized "infix (name := minus) ` ⊖ `:59 := nat.mul" in nat.mul localized "attribute [simp] le_of_lt" in le example : 2 ⊹ 3 = 5 := rfl example : 2 ↓ 3 = 6 := rfl diff --git a/test/noncomm_ring.lean b/test/noncomm_ring.lean index ccd4677adcb40..47e487c84a3e4 100644 --- a/test/noncomm_ring.lean +++ b/test/noncomm_ring.lean @@ -1,7 +1,7 @@ import tactic.noncomm_ring -local notation `⁅`a`,` b`⁆` := a * b - b * a -local infix ` ⚬ `:70 := λ a b, a * b + b * a +local notation (name := commutator) `⁅`a`,` b`⁆` := a * b - b * a +local infix (name := op) ` ⚬ `:70 := λ a b, a * b + b * a variables {R : Type*} [ring R] variables (a b c : R) diff --git a/test/simps.lean b/test/simps.lean index 1a4da9efb90ed..6ab8a76833210 100644 --- a/test/simps.lean +++ b/test/simps.lean @@ -15,7 +15,7 @@ structure equiv' (α : Sort*) (β : Sort*) := (left_inv : left_inverse inv_fun to_fun) (right_inv : right_inverse inv_fun to_fun) -local infix ` ≃ `:25 := equiv' +local infix (name := equiv') ` ≃ `:25 := equiv' /- Since `prod` and `pprod` are a special case for `@[simps]`, we define a new structure to test the basic functionality.-/ @@ -421,7 +421,7 @@ example {α β} [semigroup α] [semigroup β] (x y : α × β) : (x * y).1 = x.1 structure Semigroup := (G : Type*) (op : G → G → G) - (infix * := op) + (infix (name := op) * := op) (op_assoc : ∀ (x y z : G), (x * y) * z = x * (y * z)) namespace Group @@ -481,7 +481,7 @@ structure equiv (α : Sort*) (β : Sort*) := (to_fun : α → β) (inv_fun : β → α) -local infix ` ≃ `:25 := manual_coercion.equiv +local infix (name := equiv) ` ≃ `:25 := manual_coercion.equiv variables {α β γ : Sort*} @@ -507,7 +507,7 @@ structure equiv (α : Sort*) (β : Sort*) := (to_fun : α → β) (inv_fun : β → α) -local infix ` ≃ `:25 := faulty_manual_coercion.equiv +local infix (name := equiv) ` ≃ `:25 := faulty_manual_coercion.equiv variables {α β γ : Sort*} @@ -530,7 +530,7 @@ structure equiv (α : Sort*) (β : Sort*) := (to_fun : α → β) (inv_fun : β → α) -local infix ` ≃ `:25 := manual_initialize.equiv +local infix (name := equiv) ` ≃ `:25 := manual_initialize.equiv instance : has_coe_to_fun (α ≃ β) (λ _, α → β) := ⟨equiv.to_fun⟩ @@ -558,7 +558,7 @@ structure equiv (α : Sort u) (β : Sort v) := (to_fun : α → β) (inv_fun : β → α) -local infix ` ≃ `:25 := faulty_universes.equiv +local infix (name := equiv) ` ≃ `:25 := faulty_universes.equiv instance : has_coe_to_fun (α ≃ β) (λ _, α → β) := ⟨equiv.to_fun⟩ @@ -588,7 +588,7 @@ structure equiv (α : Sort u) (β : Sort v) := (to_fun : α → β) (inv_fun : β → α) -local infix ` ≃ `:25 := manual_universes.equiv +local infix (name := equiv) ` ≃ `:25 := manual_universes.equiv instance : has_coe_to_fun (α ≃ β) (λ _, α → β) := ⟨equiv.to_fun⟩ @@ -609,7 +609,7 @@ structure equiv (α : Sort*) (β : Sort*) := (to_fun : α → β) (inv_fun : β → α) -local infix ` ≃ `:25 := manual_projection_names.equiv +local infix (name := equiv) ` ≃ `:25 := manual_projection_names.equiv variables {α β γ : Sort*} @@ -649,7 +649,7 @@ structure equiv (α : Sort*) (β : Sort*) := (to_fun : α → β) (inv_fun : β → α) -local infix ` ≃ `:25 := prefix_projection_names.equiv +local infix (name := equiv) ` ≃ `:25 := prefix_projection_names.equiv variables {α β γ : Sort*} @@ -726,7 +726,7 @@ structure equiv (α : Sort*) (β : Sort*) := (to_fun : α → β) (inv_fun : β → α) -local infix ` ≃ `:25 := nested_non_fully_applied.equiv +local infix (name := equiv) ` ≃ `:25 := nested_non_fully_applied.equiv variables {α β γ : Sort*} diff --git a/test/to_additive.lean b/test/to_additive.lean index 6f62149311fcf..ca99623d2f5bf 100644 --- a/test/to_additive.lean +++ b/test/to_additive.lean @@ -22,7 +22,7 @@ attribute [to_additive test.my_has_smul.smul] my_has_pow.pow -- set_option pp.notation false @[priority 10000] -local infix ` ^ `:80 := my_has_pow.pow +local infix (name := pow) ` ^ `:80 := my_has_pow.pow @[to_additive bar1] def foo1 {α} [my_has_pow α ℕ] (x : α) (n : ℕ) : α := @my_has_pow.pow α ℕ _ x n From d003c55042c3cd08aefd1ae9a42ef89441cdaaf3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 30 Aug 2022 13:54:31 +0000 Subject: [PATCH 088/828] chore(analysis, measure_theory): Fix lint (#16216) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Satisfy the `fintype_finite` and `to_additive_doc` linters. Further, perform the similar weakening from `encodable` to `countable`, even though that's not yet linted for. The advantage of this is that `finite α → countable α` in known to TC inference, while `fintype α → encodable α` is not (because it's noncanonical data). As a result, some lemmas that were proved for both `fintype` and `encodable` are now deduplicated. --- src/analysis/box_integral/box/basic.lean | 2 +- src/analysis/box_integral/integrability.lean | 4 +- .../box_integral/partition/additive.lean | 2 +- .../box_integral/partition/measure.lean | 25 ++++---- .../box_integral/partition/split.lean | 3 +- src/analysis/calculus/fderiv_measurable.lean | 8 +-- .../calculus/lagrange_multipliers.lean | 3 +- src/analysis/calculus/tangent_cone.lean | 8 +-- .../normed_space/add_torsor_bases.lean | 2 +- .../normed_space/finite_dimension.lean | 14 ++-- src/analysis/specific_limits/basic.lean | 28 ++++---- src/data/set/countable.lean | 3 +- src/data/set/lattice.lean | 5 ++ src/group_theory/subgroup/basic.lean | 3 + src/linear_algebra/determinant.lean | 4 +- src/linear_algebra/finite_dimensional.lean | 22 +++---- src/linear_algebra/pi.lean | 5 +- src/logic/encodable/basic.lean | 3 + src/logic/encodable/lattice.lean | 11 ---- .../constructions/borel_space.lean | 64 +++++++------------ src/measure_theory/constructions/pi.lean | 25 ++++---- src/measure_theory/constructions/polish.lean | 10 +-- src/measure_theory/constructions/prod.lean | 4 +- src/measure_theory/covering/besicovitch.lean | 2 +- .../covering/differentiation.lean | 2 +- .../decomposition/signed_hahn.lean | 6 +- .../function/ae_measurable_order.lean | 2 +- .../function/ae_measurable_sequence.lean | 12 ++-- src/measure_theory/function/egorov.lean | 12 ++-- src/measure_theory/function/ess_sup.lean | 2 +- src/measure_theory/function/floor.lean | 8 +-- src/measure_theory/function/lp_space.lean | 4 +- .../function/simple_func_dense.lean | 2 +- .../function/strongly_measurable.lean | 8 +-- src/measure_theory/group/arithmetic.lean | 12 ++-- .../group/fundamental_domain.lean | 2 +- src/measure_theory/group/prod.lean | 22 +++++-- src/measure_theory/integral/bochner.lean | 2 +- .../integral/integrable_on.lean | 9 +-- .../integral/integral_eq_improper.lean | 6 +- .../integral/interval_integral.lean | 2 +- src/measure_theory/integral/lebesgue.lean | 19 +++--- src/measure_theory/integral/set_integral.lean | 10 +-- .../integral/vitali_caratheodory.lean | 2 +- src/measure_theory/measurable_space.lean | 35 +++------- src/measure_theory/measurable_space_def.lean | 31 ++------- src/measure_theory/measure/ae_disjoint.lean | 6 +- src/measure_theory/measure/ae_measurable.lean | 8 +-- src/measure_theory/measure/haar.lean | 23 ++++--- src/measure_theory/measure/haar_lebesgue.lean | 3 +- src/measure_theory/measure/haar_quotient.lean | 2 +- src/measure_theory/measure/hausdorff.lean | 14 ++-- src/measure_theory/measure/lebesgue.lean | 2 +- src/measure_theory/measure/measure_space.lean | 55 ++++++++-------- .../measure/measure_space_def.lean | 24 +++---- .../measure/mutually_singular.lean | 4 +- .../measure/null_measurable.lean | 32 +++------- src/measure_theory/measure/outer_measure.lean | 28 ++++---- src/measure_theory/measure/regular.lean | 4 +- src/measure_theory/measure/stieltjes.lean | 4 +- .../measure/vector_measure.lean | 12 ++-- src/measure_theory/pi_system.lean | 6 +- src/measure_theory/tactic.lean | 2 - src/order/complete_lattice.lean | 13 ++++ src/order/filter/at_top_bot.lean | 8 +-- src/order/filter/bases.lean | 3 +- src/order/filter/countable_Inter.lean | 42 ++++++------ src/order/filter/ennreal.lean | 2 +- src/probability/hitting_time.lean | 8 +-- src/probability/independence.lean | 4 +- src/probability/stopping.lean | 28 ++++---- src/ring_theory/noetherian.lean | 11 ++-- src/ring_theory/norm.lean | 2 +- src/ring_theory/trace.lean | 2 +- src/topology/algebra/infinite_sum.lean | 17 +++-- src/topology/bases.lean | 9 ++- src/topology/metric_space/basic.lean | 2 +- src/topology/metric_space/polish.lean | 7 +- 78 files changed, 412 insertions(+), 445 deletions(-) diff --git a/src/analysis/box_integral/box/basic.lean b/src/analysis/box_integral/box/basic.lean index e86a4ba47d00c..edbc1d4c03659 100644 --- a/src/analysis/box_integral/box/basic.lean +++ b/src/analysis/box_integral/box/basic.lean @@ -335,7 +335,7 @@ lemma Ioo_subset_coe (I : box ι) : I.Ioo ⊆ I := λ x hx i, Ioo_subset_Ioc_sel protected lemma Ioo_subset_Icc (I : box ι) : I.Ioo ⊆ I.Icc := I.Ioo_subset_coe.trans coe_subset_Icc -lemma Union_Ioo_of_tendsto [fintype ι] {I : box ι} {J : ℕ → box ι} (hJ : monotone J) +lemma Union_Ioo_of_tendsto [finite ι] {I : box ι} {J : ℕ → box ι} (hJ : monotone J) (hl : tendsto (lower ∘ J) at_top (𝓝 I.lower)) (hu : tendsto (upper ∘ J) at_top (𝓝 I.upper)) : (⋃ n, (J n).Ioo) = I.Ioo := have hl' : ∀ i, antitone (λ n, (J n).lower i), diff --git a/src/analysis/box_integral/integrability.lean b/src/analysis/box_integral/integrability.lean index 4573252fca6dd..99f0439fc7155 100644 --- a/src/analysis/box_integral/integrability.lean +++ b/src/analysis/box_integral/integrability.lean @@ -102,7 +102,7 @@ begin measure less than `ε / 2 ^ n / (n + 1)`. Then the norm of the integral sum is less than `ε`. -/ refine has_integral_iff.2 (λ ε ε0, _), lift ε to ℝ≥0 using ε0.lt.le, rw [gt_iff_lt, nnreal.coe_pos] at ε0, - rcases nnreal.exists_pos_sum_of_encodable ε0.ne' ℕ with ⟨δ, δ0, c, hδc, hcε⟩, + rcases nnreal.exists_pos_sum_of_countable ε0.ne' ℕ with ⟨δ, δ0, c, hδc, hcε⟩, haveI := fact.mk (I.measure_coe_lt_top μ), change μ.restrict I {x | f x ≠ 0} = 0 at hf, set N : (ι → ℝ) → ℕ := λ x, ⌈∥f x∥⌉₊, @@ -237,7 +237,7 @@ begin exact ((eventually_ge_at_top N₀).and $ this $ closed_ball_mem_nhds _ ε0).exists }, choose Nx hNx hNxε, /- We also choose a convergent series with `∑' i : ℕ, δ i < ε`. -/ - rcases nnreal.exists_pos_sum_of_encodable ε0.ne' ℕ with ⟨δ, δ0, c, hδc, hcε⟩, + rcases nnreal.exists_pos_sum_of_countable ε0.ne' ℕ with ⟨δ, δ0, c, hδc, hcε⟩, /- Since each simple function `fᵢ` is integrable, there exists `rᵢ : ℝⁿ → (0, ∞)` such that the integral sum of `f` over any tagged prepartition is `δᵢ`-close to the sum of integrals of `fᵢ` over the boxes of this prepartition. For each `x`, we choose `r (Nx x)` as the radius diff --git a/src/analysis/box_integral/partition/additive.lean b/src/analysis/box_integral/partition/additive.lean index bb787d52dc58c..e93a40af01726 100644 --- a/src/analysis/box_integral/partition/additive.lean +++ b/src/analysis/box_integral/partition/additive.lean @@ -134,7 +134,7 @@ map. -/ /-- If `f` is a box additive function on subboxes of `I` and `π₁`, `π₂` are two prepartitions of `I` that cover the same part of `I`, then `∑ J in π₁.boxes, f J = ∑ J in π₂.boxes, f J`. -/ -lemma sum_boxes_congr [fintype ι] (f : ι →ᵇᵃ[I₀] M) (hI : ↑I ≤ I₀) {π₁ π₂ : prepartition I} +lemma sum_boxes_congr [finite ι] (f : ι →ᵇᵃ[I₀] M) (hI : ↑I ≤ I₀) {π₁ π₂ : prepartition I} (h : π₁.Union = π₂.Union) : ∑ J in π₁.boxes, f J = ∑ J in π₂.boxes, f J := begin diff --git a/src/analysis/box_integral/partition/measure.lean b/src/analysis/box_integral/partition/measure.lean index 886f829a265d6..b7f96b106aefc 100644 --- a/src/analysis/box_integral/partition/measure.lean +++ b/src/analysis/box_integral/partition/measure.lean @@ -34,28 +34,27 @@ namespace box_integral open measure_theory namespace box +variables (I : box ι) -lemma measure_Icc_lt_top (I : box ι) (μ : measure (ι → ℝ)) [is_locally_finite_measure μ] : - μ I.Icc < ∞ := +lemma measure_Icc_lt_top (μ : measure (ι → ℝ)) [is_locally_finite_measure μ] : μ I.Icc < ∞ := show μ (Icc I.lower I.upper) < ∞, from I.is_compact_Icc.measure_lt_top -lemma measure_coe_lt_top (I : box ι) (μ : measure (ι → ℝ)) [is_locally_finite_measure μ] : - μ I < ∞ := +lemma measure_coe_lt_top (μ : measure (ι → ℝ)) [is_locally_finite_measure μ] : μ I < ∞ := (measure_mono $ coe_subset_Icc).trans_lt (I.measure_Icc_lt_top μ) -variables [fintype ι] (I : box ι) +section countable +variables [countable ι] lemma measurable_set_coe : measurable_set (I : set (ι → ℝ)) := -begin - rw [coe_eq_pi], - haveI := fintype.to_encodable ι, - exact measurable_set.univ_pi (λ i, measurable_set_Ioc) -end +by { rw coe_eq_pi, exact measurable_set.univ_pi (λ i, measurable_set_Ioc) } lemma measurable_set_Icc : measurable_set I.Icc := measurable_set_Icc -lemma measurable_set_Ioo : measurable_set I.Ioo := -(measurable_set_pi (set.to_finite _).countable).2 $ or.inl $ λ i hi, measurable_set_Ioo +lemma measurable_set_Ioo : measurable_set I.Ioo := measurable_set.univ_pi $ λ i, measurable_set_Ioo + +end countable + +variables [fintype ι] lemma coe_ae_eq_Icc : (I : set (ι → ℝ)) =ᵐ[volume] I.Icc := by { rw coe_eq_pi, exact measure.univ_pi_Ioc_ae_eq_Icc } @@ -65,7 +64,7 @@ measure.univ_pi_Ioo_ae_eq_Icc end box -lemma prepartition.measure_Union_to_real [fintype ι] {I : box ι} (π : prepartition I) +lemma prepartition.measure_Union_to_real [finite ι] {I : box ι} (π : prepartition I) (μ : measure (ι → ℝ)) [is_locally_finite_measure μ] : (μ π.Union).to_real = ∑ J in π.boxes, (μ J).to_real := begin diff --git a/src/analysis/box_integral/partition/split.lean b/src/analysis/box_integral/partition/split.lean index 4ff176158bffc..7a3dd3c13e81e 100644 --- a/src/analysis/box_integral/partition/split.lean +++ b/src/analysis/box_integral/partition/split.lean @@ -258,7 +258,7 @@ end section fintype -variable [fintype ι] +variable [finite ι] /-- Let `s` be a finite set of boxes in `ℝⁿ = ι → ℝ`. Then there exists a finite set `t₀` of hyperplanes (namely, the set of all hyperfaces of boxes in `s`) such that for any `t ⊇ t₀` @@ -269,6 +269,7 @@ lemma eventually_not_disjoint_imp_le_of_mem_split_many (s : finset (box ι)) : ∀ᶠ t : finset (ι × ℝ) in at_top, ∀ (I : box ι) (J ∈ s) (J' ∈ split_many I t), ¬disjoint (J : with_bot (box ι)) J' → J' ≤ J := begin + casesI nonempty_fintype ι, refine eventually_at_top.2 ⟨s.bUnion (λ J, finset.univ.bUnion (λ i, {(i, J.lower i), (i, J.upper i)})), λ t ht I J hJ J' hJ', not_disjoint_imp_le_of_subset_of_mem_split_many (λ i, _) hJ'⟩, diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index 192e6c4a9a955..4b00960f73588 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -378,8 +378,8 @@ is Borel-measurable. -/ theorem measurable_set_of_differentiable_at_of_is_complete {K : set (E →L[𝕜] F)} (hK : is_complete K) : measurable_set {x | differentiable_at 𝕜 f x ∧ fderiv 𝕜 f x ∈ K} := -by simp [differentiable_set_eq_D K hK, D, is_open_B.measurable_set, measurable_set.Inter_Prop, - measurable_set.Inter, measurable_set.Union] +by simp [differentiable_set_eq_D K hK, D, is_open_B.measurable_set, measurable_set.Inter, + measurable_set.Union] variable [complete_space F] @@ -728,8 +728,8 @@ set, is Borel-measurable. -/ theorem measurable_set_of_differentiable_within_at_Ici_of_is_complete {K : set F} (hK : is_complete K) : measurable_set {x | differentiable_within_at ℝ f (Ici x) x ∧ deriv_within f (Ici x) x ∈ K} := -by simp [differentiable_set_eq_D K hK, D, measurable_set_B, measurable_set.Inter_Prop, - measurable_set.Inter, measurable_set.Union] +by simp [differentiable_set_eq_D K hK, D, measurable_set_B, measurable_set.Inter, + measurable_set.Union] variable [complete_space F] diff --git a/src/analysis/calculus/lagrange_multipliers.lean b/src/analysis/calculus/lagrange_multipliers.lean index 5c06edbf3880e..0e001fa1211cb 100644 --- a/src/analysis/calculus/lagrange_multipliers.lean +++ b/src/analysis/calculus/lagrange_multipliers.lean @@ -133,13 +133,14 @@ Then the derivatives `f' i : E → L[ℝ] ℝ` and `φ' : E →L[ℝ] ℝ` are l See also `is_local_extr_on.exists_multipliers_of_has_strict_fderiv_at` for a version that that states existence of Lagrange multipliers `Λ` and `Λ₀` instead of using `¬linear_independent ℝ _` -/ -lemma is_local_extr_on.linear_dependent_of_has_strict_fderiv_at {ι : Type*} [fintype ι] +lemma is_local_extr_on.linear_dependent_of_has_strict_fderiv_at {ι : Type*} [finite ι] {f : ι → E → ℝ} {f' : ι → E →L[ℝ] ℝ} (hextr : is_local_extr_on φ {x | ∀ i, f i x = f i x₀} x₀) (hf' : ∀ i, has_strict_fderiv_at (f i) (f' i) x₀) (hφ' : has_strict_fderiv_at φ φ' x₀) : ¬linear_independent ℝ (option.elim φ' f' : option ι → E →L[ℝ] ℝ) := begin + casesI nonempty_fintype ι, rw [fintype.linear_independent_iff], push_neg, rcases hextr.exists_multipliers_of_has_strict_fderiv_at hf' hφ' with ⟨Λ, Λ₀, hΛ, hΛf⟩, refine ⟨option.elim Λ₀ Λ, _, _⟩, diff --git a/src/analysis/calculus/tangent_cone.lean b/src/analysis/calculus/tangent_cone.lean index f030aa993bdfd..824d445ab946d 100644 --- a/src/analysis/calculus/tangent_cone.lean +++ b/src/analysis/calculus/tangent_cone.lean @@ -324,7 +324,7 @@ begin exact (hs.1.prod ht.1).mono this end -lemma unique_diff_within_at.univ_pi (ι : Type*) [fintype ι] (E : ι → Type*) +lemma unique_diff_within_at.univ_pi (ι : Type*) [finite ι] (E : ι → Type*) [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] (s : Π i, set (E i)) (x : Π i, E i) (h : ∀ i, unique_diff_within_at 𝕜 (s i) (x i)) : unique_diff_within_at 𝕜 (set.pi univ s) x := @@ -338,7 +338,7 @@ begin exact λ i, (maps_to_tangent_cone_pi $ λ j hj, (h j).2).mono subset.rfl submodule.subset_span end -lemma unique_diff_within_at.pi (ι : Type*) [fintype ι] (E : ι → Type*) +lemma unique_diff_within_at.pi (ι : Type*) [finite ι] (E : ι → Type*) [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] (s : Π i, set (E i)) (x : Π i, E i) (I : set ι) (h : ∀ i ∈ I, unique_diff_within_at 𝕜 (s i) (x i)) : @@ -357,7 +357,7 @@ lemma unique_diff_on.prod {t : set F} (hs : unique_diff_on 𝕜 s) (ht : unique_ /-- The finite product of a family of sets of unique differentiability is a set of unique differentiability. -/ -lemma unique_diff_on.pi (ι : Type*) [fintype ι] (E : ι → Type*) +lemma unique_diff_on.pi (ι : Type*) [finite ι] (E : ι → Type*) [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] (s : Π i, set (E i)) (I : set ι) (h : ∀ i ∈ I, unique_diff_on 𝕜 (s i)) : unique_diff_on 𝕜 (set.pi I s) := @@ -365,7 +365,7 @@ lemma unique_diff_on.pi (ι : Type*) [fintype ι] (E : ι → Type*) /-- The finite product of a family of sets of unique differentiability is a set of unique differentiability. -/ -lemma unique_diff_on.univ_pi (ι : Type*) [fintype ι] (E : ι → Type*) +lemma unique_diff_on.univ_pi (ι : Type*) [finite ι] (E : ι → Type*) [Π i, normed_add_comm_group (E i)] [Π i, normed_space 𝕜 (E i)] (s : Π i, set (E i)) (h : ∀ i, unique_diff_on 𝕜 (s i)) : unique_diff_on 𝕜 (set.pi univ s) := diff --git a/src/analysis/normed_space/add_torsor_bases.lean b/src/analysis/normed_space/add_torsor_bases.lean index 86479d528ff06..18a1c19ef60ad 100644 --- a/src/analysis/normed_space/add_torsor_bases.lean +++ b/src/analysis/normed_space/add_torsor_bases.lean @@ -54,7 +54,7 @@ to this basis. TODO Restate this result for affine spaces (instead of vector spaces) once the definition of convexity is generalised to this setting. -/ -lemma interior_convex_hull_aff_basis {ι E : Type*} [fintype ι] [normed_add_comm_group E] +lemma interior_convex_hull_aff_basis {ι E : Type*} [finite ι] [normed_add_comm_group E] [normed_space ℝ E] (b : affine_basis ι ℝ E) : interior (convex_hull ℝ (range b.points)) = { x | ∀ i, 0 < b.coord i x } := begin diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 5d6cd07bd40a2..5b662e08238c3 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -154,7 +154,7 @@ begin change continuous (λ (f : E →L[𝕜] E), (f : E →ₗ[𝕜] E).det), by_cases h : ∃ (s : finset E), nonempty (basis ↥s 𝕜 E), { rcases h with ⟨s, ⟨b⟩⟩, - haveI : finite_dimensional 𝕜 E := finite_dimensional.of_finset_basis b, + haveI : finite_dimensional 𝕜 E := finite_dimensional.of_fintype_basis b, simp_rw linear_map.det_eq_det_to_matrix_of_finset b, refine continuous.matrix_det _, exact ((linear_map.to_matrix b b).to_linear_map.comp @@ -216,9 +216,10 @@ begin exact ⟨_, e.nnnorm_symm_pos, e.antilipschitz⟩ } end -protected lemma linear_independent.eventually {ι} [fintype ι] {f : ι → E} +protected lemma linear_independent.eventually {ι} [finite ι] {f : ι → E} (hf : linear_independent 𝕜 f) : ∀ᶠ g in 𝓝 f, linear_independent 𝕜 g := begin + casesI nonempty_fintype ι, simp only [fintype.linear_independent_iff'] at hf ⊢, rcases linear_map.exists_antilipschitz_with _ hf with ⟨K, K0, hK⟩, have : tendsto (λ g : ι → E, ∑ i, ∥g i - f i∥) (𝓝 f) (𝓝 $ ∑ i, ∥f i - f i∥), @@ -237,7 +238,7 @@ begin exact mul_le_mul_of_nonneg_left (norm_le_pi_norm (v - u) i) (norm_nonneg _) end -lemma is_open_set_of_linear_independent {ι : Type*} [fintype ι] : +lemma is_open_set_of_linear_independent {ι : Type*} [finite ι] : is_open {f : ι → E | linear_independent 𝕜 f} := is_open_iff_mem_nhds.2 $ λ f, linear_independent.eventually @@ -331,14 +332,15 @@ lemma basis.op_norm_le {ι : Type*} [fintype ι] (v : basis ι 𝕜 E) {u : E by simpa using nnreal.coe_le_coe.mpr (v.op_nnnorm_le ⟨M, hM⟩ hu) /-- A weaker version of `basis.op_nnnorm_le` that abstracts away the value of `C`. -/ -lemma basis.exists_op_nnnorm_le {ι : Type*} [fintype ι] (v : basis ι 𝕜 E) : +lemma basis.exists_op_nnnorm_le {ι : Type*} [finite ι] (v : basis ι 𝕜 E) : ∃ C > (0 : ℝ≥0), ∀ {u : E →L[𝕜] F} (M : ℝ≥0), (∀ i, ∥u (v i)∥₊ ≤ M) → ∥u∥₊ ≤ C*M := -⟨ max (fintype.card ι • ∥v.equiv_funL.to_continuous_linear_map∥₊) 1, +by casesI nonempty_fintype ι; exact + ⟨max (fintype.card ι • ∥v.equiv_funL.to_continuous_linear_map∥₊) 1, zero_lt_one.trans_le (le_max_right _ _), λ u M hu, (v.op_nnnorm_le M hu).trans $ mul_le_mul_of_nonneg_right (le_max_left _ _) (zero_le M)⟩ /-- A weaker version of `basis.op_norm_le` that abstracts away the value of `C`. -/ -lemma basis.exists_op_norm_le {ι : Type*} [fintype ι] (v : basis ι 𝕜 E) : +lemma basis.exists_op_norm_le {ι : Type*} [finite ι] (v : basis ι 𝕜 E) : ∃ C > (0 : ℝ), ∀ {u : E →L[𝕜] F} {M : ℝ}, 0 ≤ M → (∀ i, ∥u (v i)∥ ≤ M) → ∥u∥ ≤ C*M := let ⟨C, hC, h⟩ := v.exists_op_nnnorm_le in ⟨C, hC, λ u, subtype.forall'.mpr h⟩ diff --git a/src/analysis/specific_limits/basic.lean b/src/analysis/specific_limits/basic.lean index db98954d6e8a0..0a964e07c6ab9 100644 --- a/src/analysis/specific_limits/basic.lean +++ b/src/analysis/specific_limits/basic.lean @@ -399,7 +399,7 @@ begin exact pow_pos (zero_lt_one.trans hm) _ end -/-! ### Positive sequences with small sums on encodable types -/ +/-! ### Positive sequences with small sums on countable types -/ /-- For any positive `ε`, define on an encodable type a positive sequence with sum less than `ε` -/ def pos_sum_of_encodable {ε : ℝ} (hε : 0 < ε) @@ -439,38 +439,40 @@ end namespace nnreal -theorem exists_pos_sum_of_encodable {ε : ℝ≥0} (hε : ε ≠ 0) (ι) [encodable ι] : +theorem exists_pos_sum_of_countable {ε : ℝ≥0} (hε : ε ≠ 0) (ι) [countable ι] : ∃ ε' : ι → ℝ≥0, (∀ i, 0 < ε' i) ∧ ∃c, has_sum ε' c ∧ c < ε := -let ⟨a, a0, aε⟩ := exists_between (pos_iff_ne_zero.2 hε) in -let ⟨ε', hε', c, hc, hcε⟩ := pos_sum_of_encodable a0 ι in -⟨ λi, ⟨ε' i, le_of_lt $ hε' i⟩, assume i, nnreal.coe_lt_coe.1 $ hε' i, - ⟨c, has_sum_le (assume i, le_of_lt $ hε' i) has_sum_zero hc ⟩, nnreal.has_sum_coe.1 hc, - lt_of_le_of_lt (nnreal.coe_le_coe.1 hcε) aε ⟩ +begin + casesI nonempty_encodable ι, + obtain ⟨a, a0, aε⟩ := exists_between (pos_iff_ne_zero.2 hε), + obtain ⟨ε', hε', c, hc, hcε⟩ := pos_sum_of_encodable a0 ι, + exact ⟨λ i, ⟨ε' i, (hε' i).le⟩, λ i, nnreal.coe_lt_coe.1 $ hε' i, ⟨c, has_sum_le (λ i, (hε' i).le) + has_sum_zero hc⟩, nnreal.has_sum_coe.1 hc, aε.trans_le' $ nnreal.coe_le_coe.1 hcε⟩, +end end nnreal namespace ennreal -theorem exists_pos_sum_of_encodable {ε : ℝ≥0∞} (hε : ε ≠ 0) (ι) [encodable ι] : +theorem exists_pos_sum_of_countable {ε : ℝ≥0∞} (hε : ε ≠ 0) (ι) [countable ι] : ∃ ε' : ι → ℝ≥0, (∀ i, 0 < ε' i) ∧ ∑' i, (ε' i : ℝ≥0∞) < ε := begin rcases exists_between (pos_iff_ne_zero.2 hε) with ⟨r, h0r, hrε⟩, rcases lt_iff_exists_coe.1 hrε with ⟨x, rfl, hx⟩, - rcases nnreal.exists_pos_sum_of_encodable (coe_pos.1 h0r).ne' ι with ⟨ε', hp, c, hc, hcr⟩, + rcases nnreal.exists_pos_sum_of_countable (coe_pos.1 h0r).ne' ι with ⟨ε', hp, c, hc, hcr⟩, exact ⟨ε', hp, (ennreal.tsum_coe_eq hc).symm ▸ lt_trans (coe_lt_coe.2 hcr) hrε⟩ end -theorem exists_pos_sum_of_encodable' {ε : ℝ≥0∞} (hε : ε ≠ 0) (ι) [encodable ι] : +theorem exists_pos_sum_of_countable' {ε : ℝ≥0∞} (hε : ε ≠ 0) (ι) [countable ι] : ∃ ε' : ι → ℝ≥0∞, (∀ i, 0 < ε' i) ∧ (∑' i, ε' i) < ε := -let ⟨δ, δpos, hδ⟩ := exists_pos_sum_of_encodable hε ι in +let ⟨δ, δpos, hδ⟩ := exists_pos_sum_of_countable hε ι in ⟨λ i, δ i, λ i, ennreal.coe_pos.2 (δpos i), hδ⟩ -theorem exists_pos_tsum_mul_lt_of_encodable {ε : ℝ≥0∞} (hε : ε ≠ 0) {ι} [encodable ι] +theorem exists_pos_tsum_mul_lt_of_countable {ε : ℝ≥0∞} (hε : ε ≠ 0) {ι} [countable ι] (w : ι → ℝ≥0∞) (hw : ∀ i, w i ≠ ∞) : ∃ δ : ι → ℝ≥0, (∀ i, 0 < δ i) ∧ ∑' i, (w i * δ i : ℝ≥0∞) < ε := begin lift w to ι → ℝ≥0 using hw, - rcases exists_pos_sum_of_encodable hε ι with ⟨δ', Hpos, Hsum⟩, + rcases exists_pos_sum_of_countable hε ι with ⟨δ', Hpos, Hsum⟩, have : ∀ i, 0 < max 1 (w i), from λ i, zero_lt_one.trans_le (le_max_left _ _), refine ⟨λ i, δ' i / max 1 (w i), λ i, nnreal.div_pos (Hpos _) (this i), _⟩, refine lt_of_le_of_lt (ennreal.tsum_le_tsum $ λ i, _) Hsum, diff --git a/src/data/set/countable.lean b/src/data/set/countable.lean index db06ee87cbebb..3882fc93f9cfe 100644 --- a/src/data/set/countable.lean +++ b/src/data/set/countable.lean @@ -28,7 +28,8 @@ constructive analogue of countability. (For the most part, theorems about -/ protected def countable (s : set α) : Prop := nonempty (encodable s) -@[simp] lemma countable_coe_iff {s : set α} : countable s ↔ s.countable := nonempty_encodable.symm +@[simp] lemma countable_coe_iff {s : set α} : countable s ↔ s.countable := +encodable.nonempty_encodable.symm /-- Prove `set.countable` from a `countable` instance on the subtype. -/ lemma to_countable (s : set α) [countable s] : s.countable := countable_coe_iff.mp ‹_› diff --git a/src/data/set/lattice.lean b/src/data/set/lattice.lean index a29d0eb26cd68..04d2d57d8a31c 100644 --- a/src/data/set/lattice.lean +++ b/src/data/set/lattice.lean @@ -205,6 +205,11 @@ supr_congr_Prop pq f (pq : p ↔ q) (f : ∀x, f₁ (pq.mpr x) = f₂ x) : Inter f₁ = Inter f₂ := infi_congr_Prop pq f +lemma Union_plift_up (f : plift ι → set α) : (⋃ i, f (plift.up i)) = ⋃ i, f i := supr_plift_up _ +lemma Union_plift_down (f : ι → set α) : (⋃ i, f (plift.down i)) = ⋃ i, f i := supr_plift_down _ +lemma Inter_plift_up (f : plift ι → set α) : (⋂ i, f (plift.up i)) = ⋂ i, f i := infi_plift_up _ +lemma Inter_plift_down (f : ι → set α) : (⋂ i, f (plift.down i)) = ⋂ i, f i := infi_plift_down _ + lemma Union_eq_if {p : Prop} [decidable p] (s : set α) : (⋃ h : p, s) = if p then s else ∅ := supr_eq_if _ diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 2a263e1b0cf51..d61ef72b57b89 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -3182,6 +3182,9 @@ mul_opposite.op_equiv.subtype_equiv $ λ _, iff.rfl @[to_additive] instance (H : subgroup G) [encodable H] : encodable H.opposite := encodable.of_equiv H H.opposite_equiv.symm +@[to_additive] instance (H : subgroup G) [countable H] : countable H.opposite := +countable.of_equiv H H.opposite_equiv + @[to_additive] lemma smul_opposite_mul {H : subgroup G} (x g : G) (h : H.opposite) : h • (g * x) = g * (h • x) := begin diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index 2011e369bfe35..c4ef584883b0e 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -238,7 +238,7 @@ linear_map.det.map_one begin by_cases H : ∃ (s : finset M), nonempty (basis s 𝕜 M), { haveI : finite_dimensional 𝕜 M, - { rcases H with ⟨s, ⟨hs⟩⟩, exact finite_dimensional.of_finset_basis hs }, + { rcases H with ⟨s, ⟨hs⟩⟩, exact finite_dimensional.of_fintype_basis hs }, simp only [← det_to_matrix (finite_dimensional.fin_basis 𝕜 M), linear_equiv.map_smul, fintype.card_fin, det_smul] }, { classical, @@ -295,7 +295,7 @@ lemma finite_dimensional_of_det_ne_one {𝕜 : Type*} [field 𝕜] [module 𝕜 (f : M →ₗ[𝕜] M) (hf : f.det ≠ 1) : finite_dimensional 𝕜 M := begin by_cases H : ∃ (s : finset M), nonempty (basis s 𝕜 M), - { rcases H with ⟨s, ⟨hs⟩⟩, exact finite_dimensional.of_finset_basis hs }, + { rcases H with ⟨s, ⟨hs⟩⟩, exact finite_dimensional.of_fintype_basis hs }, { classical, simp [linear_map.coe_det, H] at hf, exact hf.elim } diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index b9ba86648c0d7..6cd251e837685 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -33,8 +33,6 @@ that all these points of view are equivalent, with the following lemmas is `fin` - `of_fintype_basis` states that the existence of a basis indexed by a finite type implies finite-dimensionality -- `of_finset_basis` states that the existence of a basis indexed by a - `finset` implies finite-dimensionality - `of_finite_basis` states that the existence of a basis indexed by a finite set implies finite-dimensionality - `is_noetherian.iff_fg` states that the space is finite-dimensional if and only if @@ -105,10 +103,10 @@ module.finite.of_surjective f w variables (K V) -instance finite_dimensional_pi {ι} [fintype ι] : finite_dimensional K (ι → K) := +instance finite_dimensional_pi {ι : Type*} [_root_.finite ι] : finite_dimensional K (ι → K) := iff_fg.1 is_noetherian_pi -instance finite_dimensional_pi' {ι} [fintype ι] (M : ι → Type*) +instance finite_dimensional_pi' {ι : Type*} [_root_.finite ι] (M : ι → Type*) [∀ i, add_comm_group (M i)] [∀ i, module K (M i)] [I : ∀ i, finite_dimensional K (M i)] : finite_dimensional K (Π i, M i) := begin @@ -126,9 +124,8 @@ by { casesI nonempty_fintype K, haveI := fintype_of_fintype K V, apply_instance variables {K V} /-- If a vector space has a finite basis, then it is finite-dimensional. -/ -lemma of_fintype_basis {ι : Type w} [fintype ι] (h : basis ι K V) : - finite_dimensional K V := -⟨⟨finset.univ.image h, by { convert h.span_eq, simp } ⟩⟩ +lemma of_fintype_basis {ι : Type w} [_root_.finite ι] (h : basis ι K V) : finite_dimensional K V := +by { casesI nonempty_fintype ι, exact ⟨⟨finset.univ.image h, by { convert h.span_eq, simp } ⟩⟩ } /-- If a vector space is `finite_dimensional`, all bases are indexed by a finite type -/ noncomputable @@ -152,11 +149,6 @@ lemma of_finite_basis {ι : Type w} {s : set ι} (h : basis s K V) (hs : set.fin finite_dimensional K V := by haveI := hs.fintype; exact of_fintype_basis h -/-- If a vector space has a finite basis, then it is finite-dimensional, finset style. -/ -lemma of_finset_basis {ι : Type w} {s : finset ι} (h : basis s K V) : - finite_dimensional K V := -of_finite_basis h s.finite_to_set - /-- A subspace of a finite-dimensional space is also finite-dimensional. -/ instance finite_dimensional_submodule [finite_dimensional K V] (S : submodule K V) : finite_dimensional K S := @@ -795,9 +787,10 @@ end /-- The submodule generated by a supremum of finite dimensional submodules, indexed by a finite type is finite-dimensional. -/ -instance finite_dimensional_supr {ι : Type*} [fintype ι] (S : ι → submodule K V) +instance finite_dimensional_supr {ι : Type*} [_root_.finite ι] (S : ι → submodule K V) [Π i, finite_dimensional K (S i)] : finite_dimensional K ↥(⨆ i, S i) := begin + casesI nonempty_fintype ι, rw ←finset.sup_univ_eq_supr, exact submodule.finite_dimensional_finset_sup _ _, end @@ -912,9 +905,10 @@ end linear_equiv section variables [division_ring K] [add_comm_group V] [module K V] -instance finite_dimensional_finsupp {ι : Type*} [fintype ι] [h : finite_dimensional K V] : +instance finite_dimensional_finsupp {ι : Type*} [_root_.finite ι] [h : finite_dimensional K V] : finite_dimensional K (ι →₀ V) := begin + casesI nonempty_fintype ι, letI : is_noetherian K V := is_noetherian.iff_fg.2 infer_instance, exact (finsupp.linear_equiv_fun_on_fintype K V ι).symm.finite_dimensional end diff --git a/src/linear_algebra/pi.lean b/src/linear_algebra/pi.lean index 32500b6b6f376..5d2567d628bb6 100644 --- a/src/linear_algebra/pi.lean +++ b/src/linear_algebra/pi.lean @@ -124,7 +124,7 @@ variables {R φ} section ext -variables [fintype ι] [decidable_eq ι] [add_comm_monoid M] [module R M] +variables [finite ι] [decidable_eq ι] [add_comm_monoid M] [module R M] {f g : (Π i, φ i) →ₗ[R] M} lemma pi_ext (h : ∀ i x, f (pi.single i x) = g (pi.single i x)) : @@ -236,9 +236,10 @@ by { ext x, simp } lemma infi_comap_proj : (⨅ i, comap (proj i : (Πi, φ i) →ₗ[R] φ i) (p i)) = pi set.univ p := by { ext x, simp } -lemma supr_map_single [decidable_eq ι] [fintype ι] : +lemma supr_map_single [decidable_eq ι] [finite ι] : (⨆ i, map (linear_map.single i : φ i →ₗ[R] (Πi, φ i)) (p i)) = pi set.univ p := begin + casesI nonempty_fintype ι, refine (supr_le $ λ i, _).antisymm _, { rintro _ ⟨x, hx : x ∈ p i, rfl⟩ j -, rcases em (j = i) with rfl|hj; simp * }, diff --git a/src/logic/encodable/basic.lean b/src/logic/encodable/basic.lean index 32e9d8311ed7e..19b88c9d95fd5 100644 --- a/src/logic/encodable/basic.lean +++ b/src/logic/encodable/basic.lean @@ -333,6 +333,9 @@ nonempty.some $ let ⟨f, hf⟩ := exists_injective_nat α in ⟨of_inj f hf⟩ end encodable +lemma nonempty_encodable (α : Type*) [countable α] : nonempty (encodable α) := +⟨encodable.of_countable _⟩ + instance : countable ℕ+ := subtype.countable -- short-circuit instance search section ulower diff --git a/src/logic/encodable/lattice.lean b/src/logic/encodable/lattice.lean index 64e4e66934549..08591a50afbb4 100644 --- a/src/logic/encodable/lattice.lean +++ b/src/logic/encodable/lattice.lean @@ -51,14 +51,3 @@ begin end end encodable - -namespace finset - -lemma nonempty_encodable {α} (t : finset α) : nonempty $ encodable {i // i ∈ t} := -begin - classical, induction t using finset.induction with x t hx ih, - { refine ⟨⟨λ _, 0, λ _, none, λ ⟨x,y⟩, y.rec _⟩⟩ }, - { cases ih with ih, exactI ⟨encodable.of_equiv _ (finset.subtype_insert_equiv_option hx)⟩ } -end - -end finset diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 1f9cb59f3c2a4..aff33ca9e319e 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -60,7 +60,7 @@ lemma borel_eq_top_of_discrete [topological_space α] [discrete_topology α] : borel α = ⊤ := top_le_iff.1 $ λ s hs, generate_measurable.basic s (is_open_discrete s) -lemma borel_eq_top_of_encodable [topological_space α] [t1_space α] [encodable α] : +lemma borel_eq_top_of_countable [topological_space α] [t1_space α] [countable α] : borel α = ⊤ := begin refine (top_le_iff.1 $ λ s hs, bUnion_of_singleton s ▸ _), @@ -339,7 +339,7 @@ instance opens_measurable_space.to_measurable_singleton_class [t1_space α] : measurable_singleton_class α := ⟨λ x, is_closed_singleton.measurable_set⟩ -instance pi.opens_measurable_space_encodable {ι : Type*} {π : ι → Type*} [encodable ι] +instance pi.opens_measurable_space {ι : Type*} {π : ι → Type*} [countable ι] [t' : Π i, topological_space (π i)] [Π i, measurable_space (π i)] [∀ i, second_countable_topology (π i)] [∀ i, opens_measurable_space (π i)] : @@ -358,13 +358,6 @@ begin exact generate_open.basic _ (hi a ha) end -instance pi.opens_measurable_space_fintype {ι : Type*} {π : ι → Type*} [fintype ι] - [t' : Π i, topological_space (π i)] - [Π i, measurable_space (π i)] [∀ i, second_countable_topology (π i)] - [∀ i, opens_measurable_space (π i)] : - opens_measurable_space (Π i, π i) := -by { letI := fintype.to_encodable ι, apply_instance } - instance prod.opens_measurable_space [second_countable_topology α] [second_countable_topology β] : opens_measurable_space (α × β) := begin @@ -537,7 +530,7 @@ begin { rintro ⟨l, -, u, -, -, hua, -, hyu⟩, exact hyu.trans_le hua } }, { refine measurable_set.bUnion hc (λ a ha, measurable_set.bUnion hc $ λ b hb, _), - refine measurable_set.Union_Prop (λ hab, measurable_set.Union_Prop $ λ hb', _), + refine measurable_set.Union (λ hab, measurable_set.Union $ λ hb', _), exact generate_measurable.basic _ ⟨a, hts ha, b, hts hb, hab, mem_singleton _⟩ } }, { simp only [not_forall, not_nonempty_iff_eq_empty] at ha, replace ha : a ∈ s := hIoo ha.some a ha.some_spec.fst ha.some_spec.snd, @@ -547,7 +540,7 @@ begin mem_Ici, mem_Iio], intros x hx, rcases htd.exists_le' (λ b hb, htb _ hb (hbot b hb)) x with ⟨z, hzt, hzx⟩, exact ⟨z, hzt, hzx.trans_lt hx, hzx⟩ }, - { refine measurable_set.bUnion hc (λ x hx, measurable_set.Union_Prop $ λ hlt, _), + { refine measurable_set.bUnion hc (λ x hx, measurable_set.Union $ λ hlt, _), exact generate_measurable.basic _ ⟨x, hts hx, a, ha, hlt, mem_singleton _⟩ } } end @@ -929,17 +922,8 @@ begin { exact comap_le_iff_le_map.mpr continuous_snd.borel_measurable } end -instance pi.borel_space_fintype_encodable {ι : Type*} {π : ι → Type*} [encodable ι] - [t' : Π i, topological_space (π i)] - [Π i, measurable_space (π i)] [∀ i, second_countable_topology (π i)] - [∀ i, borel_space (π i)] : - borel_space (Π i, π i) := -⟨le_antisymm pi_le_borel_pi opens_measurable_space.borel_le⟩ - -instance pi.borel_space_fintype {ι : Type*} {π : ι → Type*} [fintype ι] - [t' : Π i, topological_space (π i)] - [Π i, measurable_space (π i)] [∀ i, second_countable_topology (π i)] - [∀ i, borel_space (π i)] : +instance pi.borel_space {ι : Type*} {π : ι → Type*} [countable ι] [Π i, topological_space (π i)] + [Π i, measurable_space (π i)] [∀ i, second_countable_topology (π i)] [∀ i, borel_space (π i)] : borel_space (Π i, π i) := ⟨le_antisymm pi_le_borel_pi opens_measurable_space.borel_le⟩ @@ -1000,7 +984,7 @@ begin assumption end -lemma measurable.is_lub {ι} [encodable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, measurable (f i)) +lemma measurable.is_lub {ι} [countable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, measurable (f i)) (hg : ∀ b, is_lub {a | ∃ i, f i b = a} (g b)) : measurable g := begin @@ -1013,7 +997,7 @@ begin end private lemma ae_measurable.is_lub_of_nonempty {ι} (hι : nonempty ι) - {μ : measure δ} [encodable ι] {f : ι → δ → α} {g : δ → α} + {μ : measure δ} [countable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, ae_measurable (f i) μ) (hg : ∀ᵐ b ∂μ, is_lub {a | ∃ i, f i b = a} (g b)) : ae_measurable g μ := begin @@ -1039,7 +1023,7 @@ begin (ae_seq.measure_compl_ae_seq_set_eq_zero hf hg)).symm, end -lemma ae_measurable.is_lub {ι} {μ : measure δ} [encodable ι] {f : ι → δ → α} {g : δ → α} +lemma ae_measurable.is_lub {ι} {μ : measure δ} [countable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, ae_measurable (f i) μ) (hg : ∀ᵐ b ∂μ, is_lub {a | ∃ i, f i b = a} (g b)) : ae_measurable g μ := begin @@ -1057,7 +1041,7 @@ begin exact ⟨hg.exists.some, hg.mono (λ y hy, is_lub.unique hy hg.exists.some_spec)⟩, end -lemma measurable.is_glb {ι} [encodable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, measurable (f i)) +lemma measurable.is_glb {ι} [countable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, measurable (f i)) (hg : ∀ b, is_glb {a | ∃ i, f i b = a} (g b)) : measurable g := begin @@ -1070,7 +1054,7 @@ begin end private lemma ae_measurable.is_glb_of_nonempty {ι} (hι : nonempty ι) - {μ : measure δ} [encodable ι] {f : ι → δ → α} {g : δ → α} + {μ : measure δ} [countable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, ae_measurable (f i) μ) (hg : ∀ᵐ b ∂μ, is_glb {a | ∃ i, f i b = a} (g b)) : ae_measurable g μ := begin @@ -1096,7 +1080,7 @@ begin (ae_seq.measure_compl_ae_seq_set_eq_zero hf hg)).symm, end -lemma ae_measurable.is_glb {ι} {μ : measure δ} [encodable ι] {f : ι → δ → α} {g : δ → α} +lemma ae_measurable.is_glb {ι} {μ : measure δ} [countable ι] {f : ι → δ → α} {g : δ → α} (hf : ∀ i, ae_measurable (f i) μ) (hg : ∀ᵐ b ∂μ, is_glb {a | ∃ i, f i b = a} (g b)) : ae_measurable g μ := begin @@ -1207,23 +1191,23 @@ section complete_linear_order variables [complete_linear_order α] [order_topology α] [second_countable_topology α] @[measurability] -lemma measurable_supr {ι} [encodable ι] {f : ι → δ → α} (hf : ∀ i, measurable (f i)) : +lemma measurable_supr {ι} [countable ι] {f : ι → δ → α} (hf : ∀ i, measurable (f i)) : measurable (λ b, ⨆ i, f i b) := measurable.is_lub hf $ λ b, is_lub_supr @[measurability] -lemma ae_measurable_supr {ι} {μ : measure δ} [encodable ι] {f : ι → δ → α} +lemma ae_measurable_supr {ι} {μ : measure δ} [countable ι] {f : ι → δ → α} (hf : ∀ i, ae_measurable (f i) μ) : ae_measurable (λ b, ⨆ i, f i b) μ := ae_measurable.is_lub hf $ (ae_of_all μ (λ b, is_lub_supr)) @[measurability] -lemma measurable_infi {ι} [encodable ι] {f : ι → δ → α} (hf : ∀ i, measurable (f i)) : +lemma measurable_infi {ι} [countable ι] {f : ι → δ → α} (hf : ∀ i, measurable (f i)) : measurable (λ b, ⨅ i, f i b) := measurable.is_glb hf $ λ b, is_glb_infi @[measurability] -lemma ae_measurable_infi {ι} {μ : measure δ} [encodable ι] {f : ι → δ → α} +lemma ae_measurable_infi {ι} {μ : measure δ} [countable ι] {f : ι → δ → α} (hf : ∀ i, ae_measurable (f i) μ) : ae_measurable (λ b, ⨅ i, f i b) μ := ae_measurable.is_glb hf $ (ae_of_all μ (λ b, is_glb_infi)) @@ -1334,7 +1318,7 @@ instance unit.borel_space : borel_space unit := ⟨borel_eq_top_of_discrete.symm instance bool.borel_space : borel_space bool := ⟨borel_eq_top_of_discrete.symm⟩ instance nat.borel_space : borel_space ℕ := ⟨borel_eq_top_of_discrete.symm⟩ instance int.borel_space : borel_space ℤ := ⟨borel_eq_top_of_discrete.symm⟩ -instance rat.borel_space : borel_space ℚ := ⟨borel_eq_top_of_encodable.symm⟩ +instance rat.borel_space : borel_space ℚ := ⟨borel_eq_top_of_countable.symm⟩ @[priority 900] instance is_R_or_C.measurable_space {𝕜 : Type*} [is_R_or_C 𝕜] : measurable_space 𝕜 := borel 𝕜 @@ -1767,13 +1751,13 @@ ennreal.measurable_to_real.comp_ae_measurable hf /-- note: `ℝ≥0∞` can probably be generalized in a future version of this lemma. -/ @[measurability] -lemma measurable.ennreal_tsum {ι} [encodable ι] {f : ι → α → ℝ≥0∞} (h : ∀ i, measurable (f i)) : +lemma measurable.ennreal_tsum {ι} [countable ι] {f : ι → α → ℝ≥0∞} (h : ∀ i, measurable (f i)) : measurable (λ x, ∑' i, f i x) := by { simp_rw [ennreal.tsum_eq_supr_sum], apply measurable_supr, exact λ s, s.measurable_sum (λ i _, h i) } @[measurability] -lemma measurable.ennreal_tsum' {ι} [encodable ι] {f : ι → α → ℝ≥0∞} (h : ∀ i, measurable (f i)) : +lemma measurable.ennreal_tsum' {ι} [countable ι] {f : ι → α → ℝ≥0∞} (h : ∀ i, measurable (f i)) : measurable (∑' i, f i) := begin convert measurable.ennreal_tsum h, @@ -1782,7 +1766,7 @@ begin end @[measurability] -lemma measurable.nnreal_tsum {ι} [encodable ι] {f : ι → α → ℝ≥0} (h : ∀ i, measurable (f i)) : +lemma measurable.nnreal_tsum {ι} [countable ι] {f : ι → α → ℝ≥0} (h : ∀ i, measurable (f i)) : measurable (λ x, ∑' i, f i x) := begin simp_rw [nnreal.tsum_eq_to_nnreal_tsum], @@ -1790,7 +1774,7 @@ begin end @[measurability] -lemma ae_measurable.ennreal_tsum {ι} [encodable ι] {f : ι → α → ℝ≥0∞} {μ : measure α} +lemma ae_measurable.ennreal_tsum {ι} [countable ι] {f : ι → α → ℝ≥0∞} {μ : measure α} (h : ∀ i, ae_measurable (f i) μ) : ae_measurable (λ x, ∑' i, f i x) μ := by { simp_rw [ennreal.tsum_eq_supr_sum], apply ae_measurable_supr, @@ -1960,7 +1944,7 @@ lemma measurable_of_tendsto_metrizable {f : ℕ → α → β} {g : α → β} measurable g := measurable_of_tendsto_metrizable' at_top hf lim -lemma ae_measurable_of_tendsto_metrizable_ae {ι : Type*} +lemma ae_measurable_of_tendsto_metrizable_ae {ι} {μ : measure α} {f : ι → α → β} {g : α → β} (u : filter ι) [hu : ne_bot u] [is_countably_generated u] (hf : ∀ n, ae_measurable (f n) μ) (h_tendsto : ∀ᵐ x ∂μ, tendsto (λ n, f n x) u (𝓝 (g x))) : @@ -1977,7 +1961,7 @@ begin { simp_rw [ae_seq, ae_seq_lim], split_ifs with hx, { simp_rw ae_seq.mk_eq_fun_of_mem_ae_seq_set h'f hx, - exact @ae_seq.fun_prop_of_mem_ae_seq_set α β _ _ _ _ _ _ h'f x hx, }, + exact @ae_seq.fun_prop_of_mem_ae_seq_set _ α β _ _ _ _ _ h'f x hx, }, { exact tendsto_const_nhds } }, { exact (ite_ae_eq_of_measure_compl_zero g (λ x, (⟨f (v 0) x⟩ : nonempty β).some) (ae_seq_set h'f p) (ae_seq.measure_compl_ae_seq_set_eq_zero h'f hp)).symm }, @@ -2014,7 +1998,7 @@ lemma measurable_of_tendsto_metrizable_ae {μ : measure α} [μ.is_complete] {f ae_measurable_iff_measurable.mp (ae_measurable_of_tendsto_metrizable_ae' (λ i, (hf i).ae_measurable) h_ae_tendsto) -lemma measurable_limit_of_tendsto_metrizable_ae {ι} [encodable ι] [nonempty ι] {μ : measure α} +lemma measurable_limit_of_tendsto_metrizable_ae {ι} [countable ι] [nonempty ι] {μ : measure α} {f : ι → α → β} {L : filter ι} [L.is_countably_generated] (hf : ∀ n, ae_measurable (f n) μ) (h_ae_tendsto : ∀ᵐ x ∂μ, ∃ l : β, tendsto (λ n, f n x) L (𝓝 l)) : ∃ (f_lim : α → β) (hf_lim_meas : measurable f_lim), diff --git a/src/measure_theory/constructions/pi.lean b/src/measure_theory/constructions/pi.lean index 7e8ab748c0393..e576954c67f23 100644 --- a/src/measure_theory/constructions/pi.lean +++ b/src/measure_theory/constructions/pi.lean @@ -33,12 +33,12 @@ where `pi univ s` is the product of the sets `{s i | i : ι}`. We then show that this induces a product of measures, called `measure_theory.measure.pi`. For a collection of σ-finite measures `μ` and a collection of measurable sets `s` we show that `measure.pi μ (pi univ s) = ∏ i, m i (s i)`. To do this, we follow the following steps: -* We know that there is some ordering on `ι`, given by an element of `[encodable ι]`. +* We know that there is some ordering on `ι`, given by an element of `[countable ι]`. * Using this, we have an equivalence `measurable_equiv.pi_measurable_equiv_tprod` between `Π ι, α i` and an iterated product of `α i`, called `list.tprod α l` for some list `l`. * On this iterated product we can easily define a product measure `measure_theory.measure.tprod` by iterating `measure_theory.measure.prod` -* Using the previous two steps we construct `measure_theory.measure.pi'` on `Π ι, α i` for encodable +* Using the previous two steps we construct `measure_theory.measure.pi'` on `Π ι, α i` for countable `ι`. * We know that `measure_theory.measure.pi'` sends products of sets to products of measures, and since `measure_theory.measure.pi` is the maximal such measure (or at least, it comes from an outer @@ -75,7 +75,8 @@ lemma is_pi_system_pi [Π i, measurable_space (α i)] : is_pi_system (pi univ '' pi univ (λ i, {s : set (α i) | measurable_set s})) := is_pi_system.pi (λ i, is_pi_system_measurable_set) -variables [fintype ι] [fintype ι'] +section finite +variables [finite ι] [finite ι'] /-- Boxes of countably spanning sets are countably spanning. -/ lemma is_countably_spanning.pi {C : Π i, set (set (α i))} @@ -83,7 +84,7 @@ lemma is_countably_spanning.pi {C : Π i, set (set (α i))} is_countably_spanning (pi univ '' pi univ C) := begin choose s h1s h2s using hC, - haveI := fintype.to_encodable ι, + casesI nonempty_encodable (ι → ℕ), let e : ℕ → (ι → ℕ) := λ n, (decode (ι → ℕ) n).iget, refine ⟨λ n, pi univ (λ i, s i (e n i)), λ n, mem_image_of_mem _ (λ i _, h1s i _), _⟩, simp_rw [(surjective_decode_iget (ι → ℕ)).Union_comp (λ x, pi univ (λ i, s i (x i))), @@ -96,7 +97,7 @@ lemma generate_from_pi_eq {C : Π i, set (set (α i))} (hC : ∀ i, is_countably_spanning (C i)) : @measurable_space.pi _ _ (λ i, generate_from (C i)) = generate_from (pi univ '' pi univ C) := begin - haveI := fintype.to_encodable ι, + casesI nonempty_encodable ι, apply le_antisymm, { refine supr_le _, intro i, rw [comap_generate_from], apply generate_from_le, rintro _ ⟨s, hs, rfl⟩, dsimp, @@ -132,9 +133,11 @@ lemma generate_from_pi [Π i, measurable_space (α i)] : measurable_space.pi := generate_from_eq_pi (λ i, generate_from_measurable_set) (λ i, is_countably_spanning_measurable_set) +end finite + namespace measure_theory -variables {m : Π i, outer_measure (α i)} +variables [fintype ι] {m : Π i, outer_measure (α i)} /-- An upper bound for the measure in a finite product space. It is defined to by taking the image of the set under all projections, and taking the product @@ -275,12 +278,12 @@ lemma pi_pi_aux [∀ i, sigma_finite (μ i)] (s : Π i, set (α i)) (hs : ∀ i, measure.pi μ (pi univ s) = ∏ i, μ i (s i) := begin refine le_antisymm _ _, - { rw [measure.pi, to_measure_apply _ _ (measurable_set.pi_fintype (λ i _, hs i))], + { rw [measure.pi, to_measure_apply _ _ (measurable_set.pi countable_univ (λ i _, hs i))], apply outer_measure.pi_pi_le }, { haveI : encodable ι := fintype.to_encodable ι, rw [← pi'_pi μ s], - simp_rw [← pi'_pi μ s, measure.pi, - to_measure_apply _ _ (measurable_set.pi_fintype (λ i _, hs i)), ← to_outer_measure_apply], + simp_rw [← pi'_pi μ s, measure.pi, to_measure_apply _ _ (measurable_set.pi countable_univ + (λ i _, hs i)), ← to_outer_measure_apply], suffices : (pi' μ).to_outer_measure ≤ outer_measure.pi (λ i, (μ i).to_outer_measure), { exact this _ }, clear hs s, @@ -513,7 +516,7 @@ begin refine ⟨λ x, (measure.pi_eq (λ s hs, _)).symm⟩, have h : has_mul.mul x ⁻¹' (pi univ s) = set.pi univ (λ i, (λ y, x i * y) ⁻¹' s i), { ext, simp }, - simp_rw [measure.map_apply (measurable_const_mul x) (measurable_set.univ_pi_fintype hs), h, + simp_rw [measure.map_apply (measurable_const_mul x) (measurable_set.univ_pi hs), h, pi_pi, measure_preimage_mul] end @@ -523,7 +526,7 @@ begin refine ⟨(measure.pi_eq (λ s hs, _)).symm⟩, have A : has_inv.inv ⁻¹' (pi univ s) = set.pi univ (λ i, has_inv.inv ⁻¹' s i), { ext, simp }, - simp_rw [measure.inv, measure.map_apply measurable_inv (measurable_set.univ_pi_fintype hs), A, + simp_rw [measure.inv, measure.map_apply measurable_inv (measurable_set.univ_pi hs), A, pi_pi, measure_preimage_inv] end diff --git a/src/measure_theory/constructions/polish.lean b/src/measure_theory/constructions/polish.lean index 9558e51250858..94562f03a2885 100644 --- a/src/measure_theory/constructions/polish.lean +++ b/src/measure_theory/constructions/polish.lean @@ -136,7 +136,7 @@ lemma analytic_set.image_of_continuous {β : Type*} [topological_space β] hs.image_of_continuous_on hf.continuous_on /-- A countable intersection of analytic sets is analytic. -/ -theorem analytic_set.Inter [hι : nonempty ι] [encodable ι] [t2_space α] +theorem analytic_set.Inter [hι : nonempty ι] [countable ι] [t2_space α] {s : ι → set α} (hs : ∀ n, analytic_set (s n)) : analytic_set (⋂ n, s n) := begin @@ -181,7 +181,7 @@ begin end /-- A countable union of analytic sets is analytic. -/ -theorem analytic_set.Union [encodable ι] {s : ι → set α} (hs : ∀ n, analytic_set (s n)) : +theorem analytic_set.Union [countable ι] {s : ι → set α} (hs : ∀ n, analytic_set (s n)) : analytic_set (⋃ n, s n) := begin /- For the proof, write each `s n` as the continuous image under a map `f n` of a @@ -270,7 +270,7 @@ This is mostly interesting for Borel-separable sets. -/ def measurably_separable {α : Type*} [measurable_space α] (s t : set α) : Prop := ∃ u, s ⊆ u ∧ disjoint t u ∧ measurable_set u -lemma measurably_separable.Union [encodable ι] +lemma measurably_separable.Union [countable ι] {α : Type*} [measurable_space α] {s t : ι → set α} (h : ∀ m n, measurably_separable (s m) (t n)) : measurably_separable (⋃ n, s n) (⋃ m, t m) := @@ -475,11 +475,11 @@ begin { assume b, refine is_closed_closure.measurable_set.inter _, refine measurable_set.Inter (λ s, _), - exact measurable_set.Inter_Prop (λ hs, (q_meas _).diff (q_meas _)) }, + exact measurable_set.Inter (λ hs, (q_meas _).diff (q_meas _)) }, have F_meas : ∀ n, measurable_set (F n), { assume n, refine measurable_set.Union (λ s, _), - exact measurable_set.Union_Prop (λ hs, E_meas _) }, + exact measurable_set.Union (λ hs, E_meas _) }, rw this, exact measurable_set.Inter (λ n, F_meas n) }, -- we check both inclusions. diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod.lean index 8fc239fc3e4eb..cff97ef13873a 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod.lean @@ -566,14 +566,14 @@ end lemma dirac_prod_dirac {x : α} {y : β} : (dirac x).prod (dirac y) = dirac (x, y) := by rw [prod_dirac, map_dirac measurable_prod_mk_right] -lemma prod_sum {ι : Type*} [fintype ι] (ν : ι → measure β) [∀ i, sigma_finite (ν i)] : +lemma prod_sum {ι : Type*} [finite ι] (ν : ι → measure β) [∀ i, sigma_finite (ν i)] : μ.prod (sum ν) = sum (λ i, μ.prod (ν i)) := begin refine prod_eq (λ s t hs ht, _), simp_rw [sum_apply _ (hs.prod ht), sum_apply _ ht, prod_prod, ennreal.tsum_mul_left] end -lemma sum_prod {ι : Type*} [fintype ι] (μ : ι → measure α) [∀ i, sigma_finite (μ i)] : +lemma sum_prod {ι : Type*} [finite ι] (μ : ι → measure α) [∀ i, sigma_finite (μ i)] : (sum μ).prod ν = sum (λ i, (μ i).prod ν) := begin refine prod_eq (λ s t hs ht, _), diff --git a/src/measure_theory/covering/besicovitch.lean b/src/measure_theory/covering/besicovitch.lean index 838ed37d00371..34435f1851700 100644 --- a/src/measure_theory/covering/besicovitch.lean +++ b/src/measure_theory/covering/besicovitch.lean @@ -602,7 +602,7 @@ begin { apply measurable_set.inter _ omeas, haveI : encodable (u i) := (u_count i).to_encodable, exact measurable_set.Union - (λ b, measurable_set.Union_Prop (λ hb, measurable_set_closed_ball)) }, + (λ b, measurable_set.Union (λ hb, measurable_set_closed_ball)) }, calc μ o = 1/(N+1) * μ s + N/(N+1) * μ s : by { rw [μo, ← add_mul, ennreal.div_add_div_same, add_comm, ennreal.div_self, one_mul]; simp } diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index 997b9240f317f..db35ceea73f92 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -240,7 +240,7 @@ begin exact ennreal.mul_le_of_le_div ha.le } }, have B : ∀ᵐ x ∂μ, ∀ (c ∈ w) (d ∈ w), (c < d) → ¬((∃ᶠ a in v.filter_at x, ρ a / μ a < c) ∧ (∃ᶠ a in v.filter_at x, d < ρ a / μ a)), - by simpa only [ae_ball_iff w_count, ae_imp_iff], + by simpa only [ae_ball_iff w_count, ae_all_iff], filter_upwards [B], assume x hx, exact tendsto_of_no_upcrossings w_dense hx, diff --git a/src/measure_theory/decomposition/signed_hahn.lean b/src/measure_theory/decomposition/signed_hahn.lean index c54a22e05ebb4..5d8d72f2f2856 100644 --- a/src/measure_theory/decomposition/signed_hahn.lean +++ b/src/measure_theory/decomposition/signed_hahn.lean @@ -242,7 +242,7 @@ begin set k := nat.find hn with hk₁, have hk₂ : s ≤[i \ ⋃ l < k, restrict_nonpos_seq s i l] 0 := nat.find_spec hn, have hmeas : measurable_set (⋃ (l : ℕ) (H : l < k), restrict_nonpos_seq s i l) := - (measurable_set.Union $ λ _, measurable_set.Union_Prop + (measurable_set.Union $ λ _, measurable_set.Union (λ _, restrict_nonpos_seq_measurable_set _)), refine ⟨i \ ⋃ l < k, restrict_nonpos_seq s i l, hi₁.diff hmeas, set.diff_subset _ _, hk₂, _⟩, rw [of_diff hmeas hi₁, s.of_disjoint_Union_nat], @@ -250,7 +250,7 @@ begin { intros l hl, refine le_of_lt (measure_of_restrict_nonpos_seq h _ _), refine mt (restrict_le_zero_subset _ (hi₁.diff _) (set.subset.refl _)) (nat.find_min hn hl), - exact (measurable_set.Union $ λ _, measurable_set.Union_Prop + exact (measurable_set.Union $ λ _, measurable_set.Union (λ _, restrict_nonpos_seq_measurable_set _)) }, suffices : 0 ≤ ∑' (l : ℕ), s (⋃ (H : l < k), restrict_nonpos_seq s i l), { rw sub_neg, @@ -264,7 +264,7 @@ begin { convert le_of_eq s.empty.symm, ext, simp only [exists_prop, set.mem_empty_eq, set.mem_Union, not_and, iff_false], exact λ h', false.elim (h h') } }, - { intro, exact measurable_set.Union_Prop (λ _, restrict_nonpos_seq_measurable_set _) }, + { intro, exact measurable_set.Union (λ _, restrict_nonpos_seq_measurable_set _) }, { intros a b hab x hx, simp only [exists_prop, set.mem_Union, set.mem_inter_eq, set.inf_eq_inter] at hx, exact let ⟨⟨_, hx₁⟩, _, hx₂⟩ := hx in restrict_nonpos_seq_disjoint a b hab ⟨hx₁, hx₂⟩ }, diff --git a/src/measure_theory/function/ae_measurable_order.lean b/src/measure_theory/function/ae_measurable_order.lean index 11bec09202c34..db4ba98bf7821 100644 --- a/src/measure_theory/function/ae_measurable_order.lean +++ b/src/measure_theory/function/ae_measurable_order.lean @@ -61,7 +61,7 @@ begin refine (measure_Union_le _).trans _, apply ennreal.tsum_le_tsum (λ p, _), apply measure_Union_le _, - exact (s_count.mono (inter_subset_left _ _)).to_encodable, + exact (s_count.mono (inter_subset_left _ _)).to_subtype, end ... ≤ ∑' (p : s) (q : s ∩ Ioi p), μ (u p q ∩ v p q) : begin apply ennreal.tsum_le_tsum (λ p, _), diff --git a/src/measure_theory/function/ae_measurable_sequence.lean b/src/measure_theory/function/ae_measurable_sequence.lean index 659ab55518ae6..cef9c24c6810e 100644 --- a/src/measure_theory/function/ae_measurable_sequence.lean +++ b/src/measure_theory/function/ae_measurable_sequence.lean @@ -23,7 +23,7 @@ and a measurable set `ae_seq_set hf p`, such that open measure_theory open_locale classical -variables {α β γ ι : Type*} [measurable_space α] [measurable_space β] +variables {ι : Sort*} {α β γ : Type*} [measurable_space α] [measurable_space β] {f : ι → α → β} {μ : measure α} {p : α → (ι → β) → Prop} /-- If we have the additional hypothesis `∀ᵐ x ∂μ, p x (λ n, f n x)`, this is a measurable set @@ -99,7 +99,7 @@ lemma measurable (hf : ∀ i, ae_measurable (f i) μ) (p : α → (ι → β) measurable.ite ae_seq_set_measurable_set (hf i).measurable_mk $ measurable_const' $ λ x y, rfl -lemma measure_compl_ae_seq_set_eq_zero [encodable ι] (hf : ∀ i, ae_measurable (f i) μ) +lemma measure_compl_ae_seq_set_eq_zero [countable ι] (hf : ∀ i, ae_measurable (f i) μ) (hp : ∀ᵐ x ∂μ, p x (λ n, f n x)) : μ (ae_seq_set hf p)ᶜ = 0 := begin @@ -109,7 +109,7 @@ begin exact filter.eventually.and hf_eq hp, end -lemma ae_seq_eq_mk_ae [encodable ι] (hf : ∀ i, ae_measurable (f i) μ) +lemma ae_seq_eq_mk_ae [countable ι] (hf : ∀ i, ae_measurable (f i) μ) (hp : ∀ᵐ x ∂μ, p x (λ n, f n x)) : ∀ᵐ (a : α) ∂μ, ∀ (i : ι), ae_seq hf p i a = (hf i).mk (f i) a := begin @@ -119,7 +119,7 @@ begin (le_of_eq (measure_compl_ae_seq_set_eq_zero hf hp))) (zero_le _), end -lemma ae_seq_eq_fun_ae [encodable ι] (hf : ∀ i, ae_measurable (f i) μ) +lemma ae_seq_eq_fun_ae [countable ι] (hf : ∀ i, ae_measurable (f i) μ) (hp : ∀ᵐ x ∂μ, p x (λ n, f n x)) : ∀ᵐ (a : α) ∂μ, ∀ (i : ι), ae_seq hf p i a = f i a := begin @@ -128,12 +128,12 @@ begin exact measure_mono_null h_ss (measure_compl_ae_seq_set_eq_zero hf hp), end -lemma ae_seq_n_eq_fun_n_ae [encodable ι] (hf : ∀ i, ae_measurable (f i) μ) +lemma ae_seq_n_eq_fun_n_ae [countable ι] (hf : ∀ i, ae_measurable (f i) μ) (hp : ∀ᵐ x ∂μ, p x (λ n, f n x)) (n : ι) : ae_seq hf p n =ᵐ[μ] f n:= ae_all_iff.mp (ae_seq_eq_fun_ae hf hp) n -lemma supr [complete_lattice β] [encodable ι] +lemma supr [complete_lattice β] [countable ι] (hf : ∀ i, ae_measurable (f i) μ) (hp : ∀ᵐ x ∂μ, p x (λ n, f n x)) : (⨆ n, ae_seq hf p n) =ᵐ[μ] ⨆ n, f n := begin diff --git a/src/measure_theory/function/egorov.lean b/src/measure_theory/function/egorov.lean index a982920ffed55..9dbdad09ed008 100644 --- a/src/measure_theory/function/egorov.lean +++ b/src/measure_theory/function/egorov.lean @@ -64,13 +64,13 @@ begin exact ⟨n, hn₁, hn₂.le⟩ end -lemma not_convergent_seq_measurable_set [preorder ι] [encodable ι] +lemma not_convergent_seq_measurable_set [preorder ι] [countable ι] (hf : ∀ n, strongly_measurable[m] (f n)) (hg : strongly_measurable g) : measurable_set (not_convergent_seq f g n j) := -measurable_set.Union (λ k, measurable_set.Union_Prop $ λ hk, +measurable_set.Union (λ k, measurable_set.Union $ λ hk, strongly_measurable.measurable_set_lt strongly_measurable_const $ (hf k).dist hg) -lemma measure_not_convergent_seq_tendsto_zero [semilattice_sup ι] [encodable ι] +lemma measure_not_convergent_seq_tendsto_zero [semilattice_sup ι] [countable ι] (hf : ∀ n, strongly_measurable (f n)) (hg : strongly_measurable g) (hsm : measurable_set s) (hs : μ s ≠ ∞) (hfg : ∀ᵐ x ∂μ, x ∈ s → tendsto (λ n, f n x) at_top (𝓝 (g x))) (n : ℕ) : @@ -87,7 +87,7 @@ begin ⟨h.some, (lt_of_le_of_lt (measure_mono $ inter_subset_left _ _) (lt_top_iff_ne_top.2 hs)).ne⟩, end -variables [semilattice_sup ι] [nonempty ι] [encodable ι] +variables [semilattice_sup ι] [nonempty ι] [countable ι] lemma exists_not_convergent_seq_lt (hε : 0 < ε) (hf : ∀ n, strongly_measurable (f n)) (hg : strongly_measurable g) @@ -191,14 +191,14 @@ end end egorov -variables [semilattice_sup ι] [nonempty ι] [encodable ι] +variables [semilattice_sup ι] [nonempty ι] [countable ι] {γ : Type*} [topological_space γ] {f : ι → α → β} {g : α → β} {s : set α} /-- **Egorov's theorem**: If `f : ι → α → β` is a sequence of strongly measurable functions that converges to `g : α → β` almost everywhere on a measurable set `s` of finite measure, then for all `ε > 0`, there exists a subset `t ⊆ s` such that `μ t ≤ ε` and `f` converges to `g` -uniformly on `s \ t`. We require the index type `ι` to be encodable, and usually `ι = ℕ`. +uniformly on `s \ t`. We require the index type `ι` to be countable, and usually `ι = ℕ`. In other words, a sequence of almost everywhere convergent functions converges uniformly except on an arbitrarily small set. -/ diff --git a/src/measure_theory/function/ess_sup.lean b/src/measure_theory/function/ess_sup.lean index b27f519531389..1737c1e06ffb2 100644 --- a/src/measure_theory/function/ess_sup.lean +++ b/src/measure_theory/function/ess_sup.lean @@ -255,7 +255,7 @@ limsup_const_mul lemma ess_sup_add_le (f g : α → ℝ≥0∞) : ess_sup (f + g) μ ≤ ess_sup f μ + ess_sup g μ := limsup_add_le f g -lemma ess_sup_liminf_le {ι} [encodable ι] [linear_order ι] (f : ι → α → ℝ≥0∞) : +lemma ess_sup_liminf_le {ι} [countable ι] [linear_order ι] (f : ι → α → ℝ≥0∞) : ess_sup (λ x, at_top.liminf (λ n, f n x)) μ ≤ at_top.liminf (λ n, ess_sup (λ x, f n x) μ) := by { simp_rw ess_sup, exact ennreal.limsup_liminf_le_liminf_limsup (λ a b, f b a), } diff --git a/src/measure_theory/function/floor.lean b/src/measure_theory/function/floor.lean index a14bb132cedf7..9d1199f835b65 100644 --- a/src/measure_theory/function/floor.lean +++ b/src/measure_theory/function/floor.lean @@ -21,7 +21,7 @@ variables {α R : Type*} [measurable_space α] [linear_ordered_ring R] [floor_ri lemma int.measurable_floor [opens_measurable_space R] : measurable (int.floor : R → ℤ) := -measurable_to_encodable $ λ x, by simpa only [int.preimage_floor_singleton] +measurable_to_countable $ λ x, by simpa only [int.preimage_floor_singleton] using measurable_set_Ico @[measurability] lemma measurable.floor [opens_measurable_space R] @@ -30,7 +30,7 @@ int.measurable_floor.comp hf lemma int.measurable_ceil [opens_measurable_space R] : measurable (int.ceil : R → ℤ) := -measurable_to_encodable $ λ x, +measurable_to_countable $ λ x, by simpa only [int.preimage_ceil_singleton] using measurable_set_Ioc @[measurability] lemma measurable.ceil [opens_measurable_space R] @@ -64,13 +64,13 @@ variables {α R : Type*} [measurable_space α] [linear_ordered_semiring R] [floo {f : α → R} lemma nat.measurable_floor : measurable (nat.floor : R → ℕ) := -measurable_to_encodable $ λ n, by cases eq_or_ne ⌊n⌋₊ 0; simp [*, nat.preimage_floor_of_ne_zero] +measurable_to_countable $ λ n, by cases eq_or_ne ⌊n⌋₊ 0; simp [*, nat.preimage_floor_of_ne_zero] @[measurability] lemma measurable.nat_floor (hf : measurable f) : measurable (λ x, ⌊f x⌋₊) := nat.measurable_floor.comp hf lemma nat.measurable_ceil : measurable (nat.ceil : R → ℕ) := -measurable_to_encodable $ λ n, by cases eq_or_ne ⌈n⌉₊ 0; simp [*, nat.preimage_ceil_of_ne_zero] +measurable_to_countable $ λ n, by cases eq_or_ne ⌈n⌉₊ 0; simp [*, nat.preimage_ceil_of_ne_zero] @[measurability] lemma measurable.nat_ceil (hf : measurable f) : measurable (λ x, ⌈f x⌉₊) := nat.measurable_ceil.comp hf diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index a4c06311a8ed9..481188eaba02c 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -2296,7 +2296,7 @@ begin exact (continuous_nnnorm.tendsto (f_lim x)).comp hx, end -lemma snorm_exponent_top_lim_le_liminf_snorm_exponent_top {ι} [nonempty ι] [encodable ι] +lemma snorm_exponent_top_lim_le_liminf_snorm_exponent_top {ι} [nonempty ι] [countable ι] [linear_order ι] {f : ι → α → F} {f_lim : α → F} (h_lim : ∀ᵐ (x : α) ∂μ, tendsto (λ n, f n x) at_top (𝓝 (f_lim x))) : snorm f_lim ∞ μ ≤ at_top.liminf (λ n, snorm (f n) ∞ μ) := @@ -2545,7 +2545,7 @@ begin by_cases hp_top : p = ∞, { simp_rw [hp_top] at *, have h_cau_ae : ∀ᵐ x ∂μ, ∀ N n m, N ≤ n → N ≤ m → (∥(f n - f m) x∥₊ : ℝ≥0∞) < B N, - { simp_rw [ae_all_iff, ae_imp_iff], + { simp_rw ae_all_iff, exact λ N n m hnN hmN, ae_lt_of_ess_sup_lt (h_cau N n m hnN hmN), }, simp_rw [snorm_exponent_top, snorm_ess_sup] at h_cau, refine h_cau_ae.mono (λ x hx, cauchy_seq_tendsto_of_complete _), diff --git a/src/measure_theory/function/simple_func_dense.lean b/src/measure_theory/function/simple_func_dense.lean index 0eb0884373226..43fbd7037f1a0 100644 --- a/src/measure_theory/function/simple_func_dense.lean +++ b/src/measure_theory/function/simple_func_dense.lean @@ -54,7 +54,7 @@ points `e 0`, ..., `e N`. If more than one point are at the same distance from ` noncomputable def nearest_pt_ind (e : ℕ → α) : ℕ → α →ₛ ℕ | 0 := const α 0 | (N + 1) := piecewise (⋂ k ≤ N, {x | edist (e (N + 1)) x < edist (e k) x}) - (measurable_set.Inter $ λ k, measurable_set.Inter_Prop $ λ hk, + (measurable_set.Inter $ λ k, measurable_set.Inter $ λ hk, measurable_set_lt measurable_edist_right measurable_edist_right) (const α $ N + 1) (nearest_pt_ind N) diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index 4fefccb4f2bab..cf30dc9951d47 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -77,7 +77,7 @@ class second_countable_topology_either second_countable_topology_either α β := { out := or.inr (by apply_instance) } -variables {α β γ ι : Type*} [encodable ι] +variables {α β γ ι : Type*} [countable ι] namespace measure_theory local infixr ` →ₛ `:25 := simple_func @@ -1848,12 +1848,11 @@ begin { rw tendsto_pi_nhds, exact λ p, ht_sf p.fst p.snd, }, refine measurable_of_tendsto_metrizable (λ n, _) h_tendsto, - haveI : encodable (t_sf n).range, from fintype.to_encodable ↥(t_sf n).range, have h_meas : measurable (λ (p : (t_sf n).range × α), u ↑p.fst p.snd), { have : (λ (p : ↥((t_sf n).range) × α), u ↑(p.fst) p.snd) = (λ (p : α × ((t_sf n).range)), u ↑(p.snd) p.fst) ∘ prod.swap := rfl, rw [this, @measurable_swap_iff α ↥((t_sf n).range) β m], - exact measurable_from_prod_encodable (λ j, h j), }, + exact measurable_from_prod_countable (λ j, h j), }, have : (λ p : ι × α, u (t_sf n p.fst) p.snd) = (λ p : ↥(t_sf n).range × α, u p.fst p.snd) ∘ (λ p : ι × α, (⟨t_sf n p.fst, simple_func.mem_range_self _ _⟩, p.snd)) := rfl, @@ -1880,13 +1879,12 @@ begin { rw tendsto_pi_nhds, exact λ p, ht_sf p.fst p.snd, }, refine strongly_measurable_of_tendsto _ (λ n, _) h_tendsto, - haveI : encodable (t_sf n).range, from fintype.to_encodable ↥(t_sf n).range, have h_str_meas : strongly_measurable (λ (p : (t_sf n).range × α), u ↑p.fst p.snd), { refine strongly_measurable_iff_measurable_separable.2 ⟨_, _⟩, { have : (λ (p : ↥((t_sf n).range) × α), u ↑(p.fst) p.snd) = (λ (p : α × ((t_sf n).range)), u ↑(p.snd) p.fst) ∘ prod.swap := rfl, rw [this, measurable_swap_iff], - exact measurable_from_prod_encodable (λ j, (h j).measurable), }, + exact measurable_from_prod_countable (λ j, (h j).measurable), }, { have : is_separable (⋃ (i : (t_sf n).range), range (u i)) := is_separable_Union (λ i, (h i).is_separable_range), apply this.mono, diff --git a/src/measure_theory/group/arithmetic.lean b/src/measure_theory/group/arithmetic.lean index 9fd1d770b65aa..4445b7fd061a3 100644 --- a/src/measure_theory/group/arithmetic.lean +++ b/src/measure_theory/group/arithmetic.lean @@ -174,7 +174,7 @@ export has_measurable_pow (measurable_pow) /-- `monoid.has_pow` is measurable. -/ instance monoid.has_measurable_pow (M : Type*) [monoid M] [measurable_space M] [has_measurable_mul₂ M] : has_measurable_pow M ℕ := -⟨measurable_from_prod_encodable $ λ n, begin +⟨measurable_from_prod_countable $ λ n, begin induction n with n ih, { simp only [pow_zero, ←pi.one_def, measurable_one] }, { simp only [pow_succ], exact measurable_id.mul ih } @@ -336,8 +336,8 @@ begin simp_rw [set.mem_set_of_eq, pi.sub_apply, sub_eq_zero], end -lemma measurable_set_eq_fun_of_encodable {m : measurable_space α} {E} [measurable_space E] - [measurable_singleton_class E] [encodable E] {f g : α → E} +lemma measurable_set_eq_fun_of_countable {m : measurable_space α} {E} [measurable_space E] + [measurable_singleton_class E] [countable E] {f g : α → E} (hf : measurable f) (hg : measurable g) : measurable_set {x | f x = g x} := begin @@ -433,7 +433,7 @@ end inv instance div_inv_monoid.has_measurable_zpow (G : Type u) [div_inv_monoid G] [measurable_space G] [has_measurable_mul₂ G] [has_measurable_inv G] : has_measurable_pow G ℤ := -⟨measurable_from_prod_encodable $ λ n, begin +⟨measurable_from_prod_countable $ λ n, begin cases n with n n, { simp_rw zpow_of_nat, exact measurable_id.pow_const _ }, { simp_rw zpow_neg_succ_of_nat, exact (measurable_id.pow_const (n + 1)).inv } @@ -575,7 +575,7 @@ instance add_monoid.has_measurable_smul_nat₂ (M : Type*) [add_monoid M] [measu ⟨begin suffices : measurable (λ p : M × ℕ, p.2 • p.1), { apply this.comp measurable_swap, }, - refine measurable_from_prod_encodable (λ n, _), + refine measurable_from_prod_countable (λ n, _), induction n with n ih, { simp only [zero_smul, ←pi.zero_def, measurable_zero] }, { simp only [succ_nsmul], exact measurable_id.add ih } @@ -587,7 +587,7 @@ instance sub_neg_monoid.has_measurable_smul_int₂ (M : Type*) [sub_neg_monoid M ⟨begin suffices : measurable (λ p : M × ℤ, p.2 • p.1), { apply this.comp measurable_swap, }, - refine measurable_from_prod_encodable (λ n, _), + refine measurable_from_prod_countable (λ n, _), induction n with n n ih, { simp only [of_nat_zsmul], exact measurable_const_smul _, }, { simp only [zsmul_neg_succ_of_nat], exact (measurable_const_smul _).neg } diff --git a/src/measure_theory/group/fundamental_domain.lean b/src/measure_theory/group/fundamental_domain.lean index e67cab39d1f8f..a7ef2ec2fb7d8 100644 --- a/src/measure_theory/group/fundamental_domain.lean +++ b/src/measure_theory/group/fundamental_domain.lean @@ -142,7 +142,7 @@ h.image_of_equiv (measurable_equiv.smul g) (measure_preserving_smul _ _) h.image_of_equiv (measurable_equiv.smul g) (measure_preserving_smul _ _) (equiv.refl _) $ smul_comm g -variables [encodable G] {ν : measure α} +variables [countable G] {ν : measure α} @[to_additive] lemma sum_restrict_of_ac (h : is_fundamental_domain G s μ) (hν : ν ≪ μ) : sum (λ g : G, ν.restrict (g • s)) = ν := diff --git a/src/measure_theory/group/prod.lean b/src/measure_theory/group/prod.lean index 296ab73148c1a..c10d86f822fde 100644 --- a/src/measure_theory/group/prod.lean +++ b/src/measure_theory/group/prod.lean @@ -57,7 +57,7 @@ open measure /-- A shear mapping preserves the measure `μ.prod ν`. This condition is part of the definition of a measurable group in [Halmos, §59]. There, the map in this lemma is called `S`. -/ -@[to_additive map_prod_sum_eq /-" An additive shear mapping preserves the measure `μ.prod ν`. "-/] +@[to_additive map_prod_add_eq /-" An additive shear mapping preserves the measure `μ.prod ν`. "-/] lemma map_prod_mul_eq [is_mul_left_invariant ν] : map (λ z : G × G, (z.1, z.1 * z.2)) (μ.prod ν) = μ.prod ν := ((measure_preserving.id μ).skew_product measurable_mul @@ -89,7 +89,8 @@ variables [has_measurable_inv G] /-- The function we are mapping along is `S⁻¹` in [Halmos, §59], where `S` is the map in `map_prod_mul_eq`. -/ -@[to_additive map_prod_neg_add_eq] +@[to_additive map_prod_neg_add_eq "The function we are mapping along is `-S` in [Halmos, §59], where +`S` is the map in `map_prod_add_eq`."] lemma map_prod_inv_mul_eq [is_mul_left_invariant ν] : map (λ z : G × G, (z.1, z.1⁻¹ * z.2)) (μ.prod ν) = μ.prod ν := (measurable_equiv.shear_mul_right G).map_apply_eq_iff_map_symm_apply_eq.mp $ map_prod_mul_eq μ ν @@ -108,7 +109,8 @@ variables [is_mul_left_invariant μ] /-- The function we are mapping along is `S⁻¹R` in [Halmos, §59], where `S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ -@[to_additive map_prod_neg_add_eq_swap] +@[to_additive map_prod_neg_add_eq_swap "The function we are mapping along is `-S + R` in +[Halmos, §59], where `S` is the map in `map_prod_add_eq` and `R` is `prod.swap`."] lemma map_prod_inv_mul_eq_swap : map (λ z : G × G, (z.2, z.2⁻¹ * z.1)) (μ.prod ν) = ν.prod μ := begin rw [← prod_swap], @@ -119,7 +121,8 @@ end /-- The function we are mapping along is `S⁻¹RSR` in [Halmos, §59], where `S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ -@[to_additive map_prod_add_neg_eq] +@[to_additive map_prod_add_neg_eq "The function we are mapping along is `-S + R + S + R ` in +[Halmos, §59], where `S` is the map in `map_prod_add_eq` and `R` is `prod.swap`."] lemma map_prod_mul_inv_eq [is_mul_left_invariant ν] : map (λ z : G × G, (z.2 * z.1, z.1⁻¹)) (μ.prod ν) = μ.prod ν := begin @@ -233,7 +236,7 @@ begin end /-- This is the computation performed in the proof of [Halmos, §60 Th. A]. -/ -@[to_additive] +@[to_additive "This is the computation performed in the proof of [Halmos, §60 Th. A]."] lemma measure_mul_lintegral_eq [is_mul_left_invariant ν] (Em : measurable_set E) (f : G → ℝ≥0∞) (hf : measurable f) : μ E * ∫⁻ y, f y ∂ν = ∫⁻ x, ν ((λ z, z * x) ⁻¹' E) * f (x⁻¹) ∂μ := @@ -304,7 +307,14 @@ end `0 < ν(Ex⁻¹) < ∞`. The first inequality follows from §59, Th. D, but the second inequality is not justified. We prove this inequality for almost all `x` in `measure_theory.ae_measure_preimage_mul_right_lt_top_of_ne_zero`. -/ -@[to_additive] +@[to_additive "A technical lemma relating two different measures. This is basically +[Halmos, §60 Th. A]. Note that if `f` is the characteristic function of a measurable set `F` this +states that `μ F = c * μ E` for a constant `c` that does not depend on `μ`. + +Note: There is a gap in the last step of the proof in [Halmos]. In the last line, the equality +`g(-x) + ν(E - x) = f(x)` holds if we can prove that `0 < ν(E - x) < ∞`. The first inequality +follows from §59, Th. D, but the second inequality is not justified. We prove this inequality for +almost all `x` in `measure_theory.ae_measure_preimage_add_right_lt_top_of_ne_zero`."] lemma measure_lintegral_div_measure [is_mul_left_invariant ν] (Em : measurable_set E) (h2E : ν E ≠ 0) (h3E : ν E ≠ ∞) (f : G → ℝ≥0∞) (hf : measurable f) : diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index b7be4c344544b..1aa4f04689360 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -864,7 +864,7 @@ tendsto_set_to_fun_filter_of_dominated_convergence (dominated_fin_meas_additive_ bound hF_meas h_bound bound_integrable h_lim /-- Lebesgue dominated convergence theorem for series. -/ -lemma has_sum_integral_of_dominated_convergence {ι} [encodable ι] +lemma has_sum_integral_of_dominated_convergence {ι} [countable ι] {F : ι → α → E} {f : α → E} (bound : ι → α → ℝ) (hF_meas : ∀ n, ae_strongly_measurable (F n) μ) (h_bound : ∀ n, ∀ᵐ a ∂μ, ∥F n a∥ ≤ bound n a) diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index f7c9c136d28ed..95a616ed92867 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -157,7 +157,7 @@ begin simp, end -@[simp] lemma integrable_on_finite_Union {s : set β} (hs : s.finite) +@[simp] lemma integrable_on_finite_bUnion {s : set β} (hs : s.finite) {t : β → set α} : integrable_on f (⋃ i ∈ s, t i) μ ↔ ∀ i ∈ s, integrable_on f (t i) μ := begin apply hs.induction_on, @@ -167,11 +167,12 @@ end @[simp] lemma integrable_on_finset_Union {s : finset β} {t : β → set α} : integrable_on f (⋃ i ∈ s, t i) μ ↔ ∀ i ∈ s, integrable_on f (t i) μ := -integrable_on_finite_Union s.finite_to_set +integrable_on_finite_bUnion s.finite_to_set -@[simp] lemma integrable_on_fintype_Union [fintype β] {t : β → set α} : +@[simp] lemma integrable_on_finite_Union [finite β] {t : β → set α} : integrable_on f (⋃ i, t i) μ ↔ ∀ i, integrable_on f (t i) μ := -by simpa using @integrable_on_finset_Union _ _ _ _ _ f μ finset.univ t +by { casesI nonempty_fintype β, + simpa using @integrable_on_finset_Union _ _ _ _ _ f μ finset.univ t } lemma integrable_on.add_measure (hμ : integrable_on f s μ) (hν : integrable_on f s ν) : integrable_on f s (μ + ν) := diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index bb5663185543c..4e08222f9edca 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -303,9 +303,9 @@ lemma ae_cover.comp_tendsto {α ι ι' : Type*} [measurable_space α] {μ : meas { ae_eventually_mem := hφ.ae_eventually_mem.mono (λ x hx, hu.eventually hx), measurable := λ i, hφ.measurable (u i) } -section ae_cover_Union_Inter_encodable +section ae_cover_Union_Inter_countable -variables {α ι : Type*} [encodable ι] +variables {α ι : Type*} [countable ι] [measurable_space α] {μ : measure α} lemma ae_cover.bUnion_Iic_ae_cover [preorder ι] {φ : ι → set α} (hφ : ae_cover μ at_top φ) : @@ -327,7 +327,7 @@ lemma ae_cover.bInter_Ici_ae_cover [semilattice_sup ι] [nonempty ι] {φ : ι end, measurable := λ i, measurable_set.bInter (to_countable _) (λ n _, hφ.measurable n) } -end ae_cover_Union_Inter_encodable +end ae_cover_Union_Inter_countable section lintegral diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index ee45f7dcc4c88..1397c54b61a15 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -960,7 +960,7 @@ begin end /-- Lebesgue dominated convergence theorem for series. -/ -lemma has_sum_integral_of_dominated_convergence {ι} [encodable ι] +lemma has_sum_integral_of_dominated_convergence {ι} [countable ι] {F : ι → ℝ → E} (bound : ι → ℝ → ℝ) (hF_meas : ∀ n, ae_strongly_measurable (F n) (μ.restrict (Ι a b))) (h_bound : ∀ n, ∀ᵐ t ∂μ, t ∈ Ι a b → ∥F n t∥ ≤ bound n t) diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index c6221acf215c9..885c6d67a1a82 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1926,11 +1926,12 @@ end section open encodable -/-- Monotone convergence for a suprema over a directed family and indexed by an encodable type -/ -theorem lintegral_supr_directed [encodable β] {f : β → α → ℝ≥0∞} +/-- Monotone convergence for a supremum over a directed family and indexed by a countable type -/ +theorem lintegral_supr_directed [countable β] {f : β → α → ℝ≥0∞} (hf : ∀b, measurable (f b)) (h_directed : directed (≤) f) : ∫⁻ a, ⨆b, f b a ∂μ = ⨆b, ∫⁻ a, f b a ∂μ := begin + casesI nonempty_encodable β, casesI is_empty_or_nonempty β, { simp [supr_of_empty] }, inhabit β, have : ∀a, (⨆ b, f b a) = (⨆ n, f (h_directed.sequence f n) a), @@ -1952,7 +1953,7 @@ end end -lemma lintegral_tsum [encodable β] {f : β → α → ℝ≥0∞} (hf : ∀i, measurable (f i)) : +lemma lintegral_tsum [countable β] {f : β → α → ℝ≥0∞} (hf : ∀i, measurable (f i)) : ∫⁻ a, ∑' i, f i a ∂μ = ∑' i, ∫⁻ a, f i a ∂μ := begin simp only [ennreal.tsum_eq_supr_sum], @@ -1968,12 +1969,12 @@ end open measure -lemma lintegral_Union₀ [encodable β] {s : β → set α} (hm : ∀ i, null_measurable_set (s i) μ) +lemma lintegral_Union₀ [countable β] {s : β → set α} (hm : ∀ i, null_measurable_set (s i) μ) (hd : pairwise (ae_disjoint μ on s)) (f : α → ℝ≥0∞) : ∫⁻ a in ⋃ i, s i, f a ∂μ = ∑' i, ∫⁻ a in s i, f a ∂μ := by simp only [measure.restrict_Union_ae hd hm, lintegral_sum_measure] -lemma lintegral_Union [encodable β] {s : β → set α} (hm : ∀ i, measurable_set (s i)) +lemma lintegral_Union [countable β] {s : β → set α} (hm : ∀ i, measurable_set (s i)) (hd : pairwise (disjoint on s)) (f : α → ℝ≥0∞) : ∫⁻ a in ⋃ i, s i, f a ∂μ = ∑' i, ∫⁻ a in s i, f a ∂μ := lintegral_Union₀ (λ i, (hm i).null_measurable_set) hd.ae_disjoint f @@ -2003,7 +2004,7 @@ lemma lintegral_bUnion_finset {s : finset β} {t : β → set α} ∫⁻ a in ⋃ b ∈ s, t b, f a ∂μ = ∑ b in s, ∫⁻ a in t b, f a ∂μ := lintegral_bUnion_finset₀ hd.ae_disjoint (λ b hb, (hm b hb).null_measurable_set) f -lemma lintegral_Union_le [encodable β] (s : β → set α) (f : α → ℝ≥0∞) : +lemma lintegral_Union_le [countable β] (s : β → set α) (f : α → ℝ≥0∞) : ∫⁻ a in ⋃ i, s i, f a ∂μ ≤ ∑' i, ∫⁻ a in s i, f a ∂μ := begin rw [← lintegral_sum_measure], @@ -2165,7 +2166,7 @@ section countable ### Lebesgue integral over finite and countable types and sets -/ -lemma lintegral_encodable [encodable α] [measurable_singleton_class α] (f : α → ℝ≥0∞) : +lemma lintegral_countable' [countable α] [measurable_singleton_class α] (f : α → ℝ≥0∞) : ∫⁻ a, f a ∂μ = ∑' a, f a * μ {a} := begin conv_lhs { rw [← sum_smul_dirac μ, lintegral_sum_measure] }, @@ -2694,13 +2695,13 @@ begin have : ∀ n, μ (s n) < ∞, from λ n, (measure_mono $ disjointed_subset _ _).trans_lt (measure_spanning_sets_lt_top μ n), obtain ⟨δ, δpos, δsum⟩ : ∃ δ : ℕ → ℝ≥0, (∀ i, 0 < δ i) ∧ ∑' i, μ (s i) * δ i < ε, - from ennreal.exists_pos_tsum_mul_lt_of_encodable ε0 _ (λ n, (this n).ne), + from ennreal.exists_pos_tsum_mul_lt_of_countable ε0 _ (λ n, (this n).ne), set N : α → ℕ := spanning_sets_index μ, have hN_meas : measurable N := measurable_spanning_sets_index μ, have hNs : ∀ n, N ⁻¹' {n} = s n := preimage_spanning_sets_index_singleton μ, refine ⟨δ ∘ N, λ x, δpos _, measurable_from_nat.comp hN_meas, _⟩, simpa [lintegral_comp measurable_from_nat.coe_nnreal_ennreal hN_meas, hNs, - lintegral_encodable, measurable_spanning_sets_index, mul_comm] using δsum, + lintegral_countable', measurable_spanning_sets_index, mul_comm] using δsum, end lemma lintegral_trim {μ : measure α} (hm : m ≤ m0) {f : α → ℝ≥0∞} (hf : measurable[m] f) : diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 37eb88792a7d1..0fa648f478874 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -171,7 +171,7 @@ by rw [← set.indicator_add_compl_eq_piecewise, integral_add' (hf.indicator hs) (hg.indicator hs.compl), integral_indicator hs, integral_indicator hs.compl] -lemma tendsto_set_integral_of_monotone {ι : Type*} [encodable ι] [semilattice_sup ι] +lemma tendsto_set_integral_of_monotone {ι : Type*} [countable ι] [semilattice_sup ι] {s : ι → set α} {f : α → E} (hsm : ∀ i, measurable_set (s i)) (h_mono : monotone s) (hfi : integrable_on f (⋃ n, s n) μ) : tendsto (λ i, ∫ a in s i, f a ∂μ) at_top (𝓝 (∫ a in (⋃ n, s n), f a ∂μ)) := @@ -195,7 +195,7 @@ begin (hi.2.trans_lt $ ennreal.add_lt_top.2 ⟨hfi', ennreal.coe_lt_top⟩).ne] end -lemma has_sum_integral_Union_ae {ι : Type*} [encodable ι] {s : ι → set α} {f : α → E} +lemma has_sum_integral_Union_ae {ι : Type*} [countable ι] {s : ι → set α} {f : α → E} (hm : ∀ i, null_measurable_set (s i) μ) (hd : pairwise (ae_disjoint μ on s)) (hfi : integrable_on f (⋃ i, s i) μ) : has_sum (λ n, ∫ a in s n, f a ∂ μ) (∫ a in ⋃ n, s n, f a ∂μ) := @@ -204,19 +204,19 @@ begin exact has_sum_integral_measure hfi end -lemma has_sum_integral_Union {ι : Type*} [encodable ι] {s : ι → set α} {f : α → E} +lemma has_sum_integral_Union {ι : Type*} [countable ι] {s : ι → set α} {f : α → E} (hm : ∀ i, measurable_set (s i)) (hd : pairwise (disjoint on s)) (hfi : integrable_on f (⋃ i, s i) μ) : has_sum (λ n, ∫ a in s n, f a ∂ μ) (∫ a in ⋃ n, s n, f a ∂μ) := has_sum_integral_Union_ae (λ i, (hm i).null_measurable_set) (hd.mono (λ i j h, h.ae_disjoint)) hfi -lemma integral_Union {ι : Type*} [encodable ι] {s : ι → set α} {f : α → E} +lemma integral_Union {ι : Type*} [countable ι] {s : ι → set α} {f : α → E} (hm : ∀ i, measurable_set (s i)) (hd : pairwise (disjoint on s)) (hfi : integrable_on f (⋃ i, s i) μ) : (∫ a in (⋃ n, s n), f a ∂μ) = ∑' n, ∫ a in s n, f a ∂ μ := (has_sum.tsum_eq (has_sum_integral_Union hm hd hfi)).symm -lemma integral_Union_ae {ι : Type*} [encodable ι] {s : ι → set α} {f : α → E} +lemma integral_Union_ae {ι : Type*} [countable ι] {s : ι → set α} {f : α → E} (hm : ∀ i, null_measurable_set (s i) μ) (hd : pairwise (ae_disjoint μ on s)) (hfi : integrable_on f (⋃ i, s i) μ) : (∫ a in (⋃ n, s n), f a ∂μ) = ∑' n, ∫ a in s n, f a ∂ μ := diff --git a/src/measure_theory/integral/vitali_caratheodory.lean b/src/measure_theory/integral/vitali_caratheodory.lean index 2a07f64a0eab1..6b3e57b920466 100644 --- a/src/measure_theory/integral/vitali_caratheodory.lean +++ b/src/measure_theory/integral/vitali_caratheodory.lean @@ -154,7 +154,7 @@ lemma exists_le_lower_semicontinuous_lintegral_ge (f : α → ℝ≥0∞) (hf : measurable f) {ε : ℝ≥0∞} (εpos : ε ≠ 0) : ∃ g : α → ℝ≥0∞, (∀ x, f x ≤ g x) ∧ lower_semicontinuous g ∧ (∫⁻ x, g x ∂μ ≤ ∫⁻ x, f x ∂μ + ε) := begin - rcases ennreal.exists_pos_sum_of_encodable' εpos ℕ with ⟨δ, δpos, hδ⟩, + rcases ennreal.exists_pos_sum_of_countable' εpos ℕ with ⟨δ, δpos, hδ⟩, have : ∀ n, ∃ g : α → ℝ≥0, (∀ x, simple_func.eapprox_diff f n x ≤ g x) ∧ lower_semicontinuous g ∧ (∫⁻ x, g x ∂μ ≤ ∫⁻ x, simple_func.eapprox_diff f n x ∂μ + δ n) := λ n, simple_func.exists_le_lower_semicontinuous_lintegral_ge μ diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index 386c3000b628f..b7c4d0ab4c33a 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -202,8 +202,7 @@ begin { convert measurable_const, exact funext (λ x, hf x h.some) } end -lemma measurable_of_fintype [fintype α] [measurable_singleton_class α] (f : α → β) : - measurable f := +lemma measurable_of_finite [finite α] [measurable_singleton_class α] (f : α → β) : measurable f := λ s hs, (f ⁻¹' s).to_finite.measurable_set end typeclass_measurable_space @@ -287,13 +286,13 @@ instance : measurable_singleton_class ℕ := ⟨λ _, trivial⟩ instance : measurable_singleton_class ℤ := ⟨λ _, trivial⟩ instance : measurable_singleton_class ℚ := ⟨λ _, trivial⟩ -lemma measurable_to_encodable [measurable_space α] [encodable α] [measurable_space β] {f : β → α} +lemma measurable_to_countable [measurable_space α] [countable α] [measurable_space β] {f : β → α} (h : ∀ y, measurable_set (f ⁻¹' {f y})) : measurable f := begin assume s hs, rw [← bUnion_preimage_singleton], - refine measurable_set.Union (λ y, measurable_set.Union_Prop $ λ hy, _), + refine measurable_set.Union (λ y, measurable_set.Union $ λ hy, _), by_cases hyf : y ∈ range f, { rcases hyf with ⟨y, rfl⟩, apply h }, @@ -310,7 +309,7 @@ variables [measurable_space α] measurable_from_top lemma measurable_to_nat {f : α → ℕ} : (∀ y, measurable_set (f ⁻¹' {f y})) → measurable f := -measurable_to_encodable +measurable_to_countable lemma measurable_find_greatest' {p : α → ℕ → Prop} [∀ x, decidable_pred (p x)] {N : ℕ} (hN : ∀ k ≤ N, measurable_set {x | nat.find_greatest (p x) N = k}) : @@ -324,7 +323,7 @@ begin refine measurable_find_greatest' (λ k hk, _), simp only [nat.find_greatest_eq_iff, set_of_and, set_of_forall, ← compl_set_of], repeat { apply_rules [measurable_set.inter, measurable_set.const, measurable_set.Inter, - measurable_set.Inter_Prop, measurable_set.compl, hN]; try { intros } } + measurable_set.compl, hN]; try { intros } } end lemma measurable_find {p : α → ℕ → Prop} [∀ x, decidable_pred (p x)] @@ -459,7 +458,7 @@ lemma measurable_of_measurable_on_compl_finite [measurable_singleton_class α] begin letI : fintype s := finite.fintype hs, exact measurable_of_restrict_of_restrict_compl hs.measurable_set - (measurable_of_fintype _) hf + (measurable_of_finite _) hf end lemma measurable_of_measurable_on_compl_singleton [measurable_singleton_class α] @@ -574,7 +573,7 @@ lemma measurable_set_swap_iff {s : set (α × β)} : measurable_set (prod.swap ⁻¹' s) ↔ measurable_set s := ⟨λ hs, by { convert measurable_swap hs, ext ⟨x, y⟩, refl }, λ hs, measurable_swap hs⟩ -lemma measurable_from_prod_encodable [encodable β] [measurable_singleton_class β] +lemma measurable_from_prod_countable [countable β] [measurable_singleton_class β] {mγ : measurable_space γ} {f : α × β → γ} (hf : ∀ y, measurable (λ x, f (x, y))) : measurable f := begin @@ -593,7 +592,7 @@ lemma measurable.find {m : measurable_space α} (hf : ∀ n, measurable (f n)) (hp : ∀ n, measurable_set {x | p n x}) (h : ∀ x, ∃ n, p n x) : measurable (λ x, f (nat.find (h x)) x) := begin - have : measurable (λ (p : α × ℕ), f p.2 p.1) := measurable_from_prod_encodable (λ n, hf n), + have : measurable (λ (p : α × ℕ), f p.2 p.1) := measurable_from_prod_countable (λ n, hf n), exact this.comp (measurable.prod_mk measurable_id (measurable_find h hp)), end @@ -677,7 +676,7 @@ lemma measurable_set.pi {s : set δ} {t : Π i : δ, set (π i)} (hs : s.countab measurable_set (s.pi t) := by { rw [pi_def], exact measurable_set.bInter hs (λ i hi, measurable_pi_apply _ (ht i hi)) } -lemma measurable_set.univ_pi [encodable δ] {t : Π i : δ, set (π i)} +lemma measurable_set.univ_pi [countable δ] {t : Π i : δ, set (π i)} (ht : ∀ i, measurable_set (t i)) : measurable_set (pi univ t) := measurable_set.pi (to_countable _) (λ i _, ht i) @@ -698,7 +697,6 @@ begin { simp [measurable_set_pi_of_nonempty hs, h, ← not_nonempty_iff_eq_empty] } end -section variable (π) @[measurability] @@ -727,21 +725,6 @@ begin simp only [pi_equiv_pi_subtype_prod_apply, measurable_pi_apply] } end -end - -section fintype - -local attribute [instance] fintype.to_encodable - -lemma measurable_set.pi_fintype [fintype δ] {s : set δ} {t : Π i, set (π i)} - (ht : ∀ i ∈ s, measurable_set (t i)) : measurable_set (pi s t) := -measurable_set.pi (to_countable _) ht - -lemma measurable_set.univ_pi_fintype [fintype δ] {t : Π i, set (π i)} - (ht : ∀ i, measurable_set (t i)) : measurable_set (pi univ t) := -measurable_set.pi_fintype (λ i _, ht i) - -end fintype end pi instance tprod.measurable_space (π : δ → Type*) [∀ x, measurable_space (π x)] : diff --git a/src/measure_theory/measurable_space_def.lean b/src/measure_theory/measurable_space_def.lean index dd5fcad47e9a3..65c9533c049ac 100644 --- a/src/measure_theory/measurable_space_def.lean +++ b/src/measure_theory/measurable_space_def.lean @@ -96,11 +96,12 @@ lemma measurable_set.bUnion_decode₂ [encodable β] ⦃f : β → set α⦄ (h (n : ℕ) : measurable_set (⋃ b ∈ decode₂ β n, f b) := encodable.Union_decode₂_cases measurable_set.empty h -lemma measurable_set.Union [encodable β] ⦃f : β → set α⦄ (h : ∀ b, measurable_set (f b)) : +lemma measurable_set.Union [countable ι] ⦃f : ι → set α⦄ (h : ∀ b, measurable_set (f b)) : measurable_set (⋃ b, f b) := begin - rw ← encodable.Union_decode₂, - exact ‹measurable_space α›.measurable_set_Union _ (measurable_set.bUnion_decode₂ h) + casesI nonempty_encodable (plift ι), + rw [←Union_plift_down, ←encodable.Union_decode₂], + exact ‹measurable_space α›.measurable_set_Union _ (measurable_set.bUnion_decode₂ $ λ _, h _), end lemma measurable_set.bUnion {f : β → set α} {s : set β} (hs : s.countable) @@ -130,29 +131,11 @@ lemma set.finite.measurable_set_sUnion {s : set (set α)} (hs : s.finite) measurable_set (⋃₀ s) := measurable_set.sUnion hs.countable h -lemma measurable_set.Union_Prop {p : Prop} {f : p → set α} (hf : ∀ b, measurable_set (f b)) : - measurable_set (⋃ b, f b) := -by { by_cases p; simp [h, hf, measurable_set.empty] } - -lemma measurable_set.Inter [encodable β] {f : β → set α} (h : ∀ b, measurable_set (f b)) : +lemma measurable_set.Inter [countable ι] {f : ι → set α} (h : ∀ b, measurable_set (f b)) : measurable_set (⋂ b, f b) := measurable_set.compl_iff.1 $ by { rw compl_Inter, exact measurable_set.Union (λ b, (h b).compl) } -section fintype - -local attribute [instance] fintype.to_encodable - -lemma measurable_set.Union_fintype [fintype β] {f : β → set α} (h : ∀ b, measurable_set (f b)) : - measurable_set (⋃ b, f b) := -measurable_set.Union h - -lemma measurable_set.Inter_fintype [fintype β] {f : β → set α} (h : ∀ b, measurable_set (f b)) : - measurable_set (⋂ b, f b) := -measurable_set.Inter h - -end fintype - lemma measurable_set.bInter {f : β → set α} {s : set β} (hs : s.countable) (h : ∀ b ∈ s, measurable_set (f b)) : measurable_set (⋂ b ∈ s, f b) := measurable_set.compl_iff.1 $ @@ -174,10 +157,6 @@ lemma set.finite.measurable_set_sInter {s : set (set α)} (hs : s.finite) (h : ∀ t ∈ s, measurable_set t) : measurable_set (⋂₀ s) := measurable_set.sInter hs.countable h -lemma measurable_set.Inter_Prop {p : Prop} {f : p → set α} (hf : ∀ b, measurable_set (f b)) : - measurable_set (⋂ b, f b) := -by { by_cases p; simp [h, hf, measurable_set.univ] } - @[simp] lemma measurable_set.union {s₁ s₂ : set α} (h₁ : measurable_set s₁) (h₂ : measurable_set s₂) : measurable_set (s₁ ∪ s₂) := diff --git a/src/measure_theory/measure/ae_disjoint.lean b/src/measure_theory/measure/ae_disjoint.lean index 8a806614f013b..992a83907378f 100644 --- a/src/measure_theory/measure/ae_disjoint.lean +++ b/src/measure_theory/measure/ae_disjoint.lean @@ -26,7 +26,7 @@ variables {μ} {s t u v : set α} /-- If `s : ι → set α` is a countable family of pairwise a.e. disjoint sets, then there exists a family of measurable null sets `t i` such that `s i \ t i` are pairwise disjoint. -/ -lemma exists_null_pairwise_disjoint_diff [encodable ι] {s : ι → set α} +lemma exists_null_pairwise_disjoint_diff [countable ι] {s : ι → set α} (hd : pairwise (ae_disjoint μ on s)) : ∃ t : ι → set α, (∀ i, measurable_set (t i)) ∧ (∀ i, μ (t i) = 0) ∧ pairwise (disjoint on (λ i, s i \ t i)) := @@ -71,11 +71,11 @@ measure_mono_null_ae (hu.inter hv) h lemma mono (h : ae_disjoint μ s t) (hu : u ⊆ s) (hv : v ⊆ t) : ae_disjoint μ u v := h.mono_ae hu.eventually_le hv.eventually_le -@[simp] lemma Union_left_iff [encodable ι] {s : ι → set α} : +@[simp] lemma Union_left_iff [countable ι] {s : ι → set α} : ae_disjoint μ (⋃ i, s i) t ↔ ∀ i, ae_disjoint μ (s i) t := by simp only [ae_disjoint, Union_inter, measure_Union_null_iff] -@[simp] lemma Union_right_iff [encodable ι] {t : ι → set α} : +@[simp] lemma Union_right_iff [countable ι] {t : ι → set α} : ae_disjoint μ s (⋃ i, t i) ↔ ∀ i, ae_disjoint μ s (t i) := by simp only [ae_disjoint, inter_Union, measure_Union_null_iff] diff --git a/src/measure_theory/measure/ae_measurable.lean b/src/measure_theory/measure/ae_measurable.lean index 23d39122c84a1..99b01758f417e 100644 --- a/src/measure_theory/measure/ae_measurable.lean +++ b/src/measure_theory/measure/ae_measurable.lean @@ -58,7 +58,7 @@ lemma ae_inf_principal_eq_mk {s} (h : ae_measurable f (μ.restrict s)) : le_ae_restrict h.ae_eq_mk @[measurability] -lemma sum_measure [encodable ι] {μ : ι → measure α} (h : ∀ i, ae_measurable f (μ i)) : +lemma sum_measure [countable ι] {μ : ι → measure α} (h : ∀ i, ae_measurable f (μ i)) : ae_measurable f (sum μ) := begin nontriviality β, inhabit β, @@ -87,7 +87,7 @@ begin exact λ h, hx (mem_Inter.1 h i) } end -@[simp] lemma _root_.ae_measurable_sum_measure_iff [encodable ι] {μ : ι → measure α} : +@[simp] lemma _root_.ae_measurable_sum_measure_iff [countable ι] {μ : ι → measure α} : ae_measurable f (sum μ) ↔ ∀ i, ae_measurable f (μ i) := ⟨λ h i, h.mono_measure (le_sum _ _), sum_measure⟩ @@ -101,11 +101,11 @@ lemma add_measure {f : α → β} (hμ : ae_measurable f μ) (hν : ae_measurabl ae_measurable_add_measure_iff.2 ⟨hμ, hν⟩ @[measurability] -protected lemma Union [encodable ι] {s : ι → set α} (h : ∀ i, ae_measurable f (μ.restrict (s i))) : +protected lemma Union [countable ι] {s : ι → set α} (h : ∀ i, ae_measurable f (μ.restrict (s i))) : ae_measurable f (μ.restrict (⋃ i, s i)) := (sum_measure h).mono_measure $ restrict_Union_le -@[simp] lemma _root_.ae_measurable_Union_iff [encodable ι] {s : ι → set α} : +@[simp] lemma _root_.ae_measurable_Union_iff [countable ι] {s : ι → set α} : ae_measurable f (μ.restrict (⋃ i, s i)) ↔ ∀ i, ae_measurable f (μ.restrict (s i)) := ⟨λ h i, h.mono_measure $ restrict_mono (subset_Union _ _) le_rfl, ae_measurable.Union⟩ diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar.lean index 7f81eef302559..0c07712eb7659 100644 --- a/src/measure_theory/measure/haar.lean +++ b/src/measure_theory/measure/haar.lean @@ -134,7 +134,8 @@ variables [topological_group G] /-- If `K` is compact and `V` has nonempty interior, then the index `(K : V)` is well-defined, there is a finite set `t` satisfying the desired properties. -/ -@[to_additive add_index_defined] +@[to_additive add_index_defined "If `K` is compact and `V` has nonempty interior, then the index +`(K : V)` is well-defined, there is a finite set `t` satisfying the desired properties."] lemma index_defined {K V : set G} (hK : is_compact K) (hV : (interior V).nonempty) : ∃ n : ℕ, n ∈ finset.card '' {t : finset G | K ⊆ ⋃ g ∈ t, (λ h, g * h) ⁻¹' V } := by { rcases compact_covered_by_mul_left_translates hK hV with ⟨t, ht⟩, exact ⟨t.card, t, ht, rfl⟩ } @@ -478,12 +479,12 @@ lemma haar_content_apply (K₀ : positive_compacts G) (K : compacts G) : haar_content K₀ K = show nnreal, from ⟨chaar K₀ K, chaar_nonneg _ _⟩ := rfl /-- The variant of `chaar_self` for `haar_content` -/ -@[to_additive] +@[to_additive "The variant of `add_chaar_self` for `add_haar_content`."] lemma haar_content_self {K₀ : positive_compacts G} : haar_content K₀ K₀.to_compacts = 1 := by { simp_rw [← ennreal.coe_one, haar_content_apply, ennreal.coe_eq_coe, chaar_self], refl } /-- The variant of `is_left_invariant_chaar` for `haar_content` -/ -@[to_additive] +@[to_additive "The variant of `is_left_invariant_add_chaar` for `add_haar_content`"] lemma is_left_invariant_haar_content {K₀ : positive_compacts G} (g : G) (K : compacts G) : haar_content K₀ (K.map _ $ continuous_mul_left g) = haar_content K₀ K := by simpa only [ennreal.coe_eq_coe, ←nnreal.coe_eq, haar_content_apply] @@ -545,7 +546,7 @@ begin end /-- The Haar measure is regular. -/ -@[to_additive] +@[to_additive "The additive Haar measure is regular."] instance regular_haar_measure {K₀ : positive_compacts G} : (haar_measure K₀).regular := begin @@ -556,14 +557,15 @@ begin end /-- The Haar measure is sigma-finite in a second countable group. -/ -@[to_additive] +@[to_additive "The additive Haar measure is sigma-finite in a second countable group."] instance sigma_finite_haar_measure [second_countable_topology G] {K₀ : positive_compacts G} : sigma_finite (haar_measure K₀) := by { haveI : locally_compact_space G := K₀.locally_compact_space_of_group, apply_instance, } /-- The Haar measure is a Haar measure, i.e., it is invariant and gives finite mass to compact sets and positive mass to nonempty open sets. -/ -@[to_additive] +@[to_additive "The additive Haar measure is an additive Haar measure, i.e., it is invariant and +gives finite mass to compact sets and positive mass to nonempty open sets."] instance is_haar_measure_haar_measure (K₀ : positive_compacts G) : is_haar_measure (haar_measure K₀) := begin @@ -586,7 +588,9 @@ variables [second_countable_topology G] is a scalar multiple of the Haar measure. This is slightly weaker than assuming that `μ` is a Haar measure (in particular we don't require `μ ≠ 0`). -/ -@[to_additive] +@[to_additive "The additive Haar measure is unique up to scaling. More precisely: every σ-finite +left invariant measure is a scalar multiple of the additive Haar measure. This is slightly weaker +than assuming that `μ` is an additive Haar measure (in particular we don't require `μ ≠ 0`)."] theorem haar_measure_unique (μ : measure G) [sigma_finite μ] [is_mul_left_invariant μ] (K₀ : positive_compacts G) : μ = μ K₀ • haar_measure K₀ := (measure_eq_div_smul μ (haar_measure K₀) K₀.compact.measurable_set @@ -599,7 +603,8 @@ haar_measure_unique μ K₀ /-- To show that an invariant σ-finite measure is regular it is sufficient to show that it is finite on some compact set with non-empty interior. -/ -@[to_additive] +@[to_additive "To show that an invariant σ-finite measure is regular it is sufficient to show that +it is finite on some compact set with non-empty interior."] theorem regular_of_is_mul_left_invariant {μ : measure G} [sigma_finite μ] [is_mul_left_invariant μ] {K : set G} (hK : is_compact K) (h2K : (interior K).nonempty) (hμK : μ K ≠ ∞) : regular μ := @@ -690,7 +695,7 @@ end end second_countable /-- Any Haar measure is invariant under inversion in a commutative group. -/ -@[to_additive] +@[to_additive "Any additive Haar measure is invariant under negation in a commutative group."] lemma map_haar_inv {G : Type*} [comm_group G] [topological_space G] [topological_group G] [t2_space G] [measurable_space G] [borel_space G] [locally_compact_space G] [second_countable_topology G] diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index a29416200a9bd..b42d1140e8211 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -188,10 +188,11 @@ linear equiv maps Haar measure to Haar measure. -/ lemma map_linear_map_add_haar_pi_eq_smul_add_haar - {ι : Type*} [fintype ι] {f : (ι → ℝ) →ₗ[ℝ] (ι → ℝ)} (hf : f.det ≠ 0) + {ι : Type*} [finite ι] {f : (ι → ℝ) →ₗ[ℝ] (ι → ℝ)} (hf : f.det ≠ 0) (μ : measure (ι → ℝ)) [is_add_haar_measure μ] : measure.map f μ = ennreal.of_real (abs (f.det)⁻¹) • μ := begin + casesI nonempty_fintype ι, /- We have already proved the result for the Lebesgue product measure, using matrices. We deduce it for any Haar measure by uniqueness (up to scalar multiplication). -/ have := add_haar_measure_unique μ (pi_Icc01 ι), diff --git a/src/measure_theory/measure/haar_quotient.lean b/src/measure_theory/measure/haar_quotient.lean index d6a684326c9c7..84cbc237b70ba 100644 --- a/src/measure_theory/measure/haar_quotient.lean +++ b/src/measure_theory/measure/haar_quotient.lean @@ -63,7 +63,7 @@ instance quotient_group.has_measurable_smul [measurable_space (G ⧸ Γ)] [borel variables {𝓕 : set G} (h𝓕 : is_fundamental_domain Γ.opposite 𝓕 μ) include h𝓕 -variables [encodable Γ] [measurable_space (G ⧸ Γ)] [borel_space (G ⧸ Γ)] +variables [countable Γ] [measurable_space (G ⧸ Γ)] [borel_space (G ⧸ Γ)] /-- The pushforward to the coset space `G ⧸ Γ` of the restriction of a both left- and right- invariant measure on `G` to a fundamental domain `𝓕` is a `G`-invariant measure on `G ⧸ Γ`. -/ diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index 8370d00a788fa..fb4aa3800f7ec 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -509,13 +509,14 @@ end /-- To bound the Hausdorff measure (or, more generally, for a measure defined using `measure_theory.measure.mk_metric`) of a set, one may use coverings with maximum diameter tending to -`0`, indexed by any sequence of encodable types. -/ -lemma mk_metric_le_liminf_tsum {β : Type*} {ι : β → Type*} [∀ n, encodable (ι n)] (s : set X) +`0`, indexed by any sequence of countable types. -/ +lemma mk_metric_le_liminf_tsum {β : Type*} {ι : β → Type*} [∀ n, countable (ι n)] (s : set X) {l : filter β} (r : β → ℝ≥0∞) (hr : tendsto r l (𝓝 0)) (t : Π (n : β), ι n → set X) (ht : ∀ᶠ n in l, ∀ i, diam (t n i) ≤ r n) (hst : ∀ᶠ n in l, s ⊆ ⋃ i, t n i) (m : ℝ≥0∞ → ℝ≥0∞) : mk_metric m s ≤ liminf l (λ n, ∑' i, m (diam (t n i))) := begin + haveI : Π n, encodable (ι n) := λ n, encodable.of_countable _, simp only [mk_metric_apply], refine supr₂_le (λ ε hε, _), refine le_of_forall_le_of_dense (λ c hc, _), @@ -541,10 +542,7 @@ lemma mk_metric_le_liminf_sum {β : Type*} {ι : β → Type*} [hι : ∀ n, fin (ht : ∀ᶠ n in l, ∀ i, diam (t n i) ≤ r n) (hst : ∀ᶠ n in l, s ⊆ ⋃ i, t n i) (m : ℝ≥0∞ → ℝ≥0∞) : mk_metric m s ≤ liminf l (λ n, ∑ i, m (diam (t n i))) := -begin - haveI : ∀ n, encodable (ι n), from λ n, fintype.to_encodable _, - simpa only [tsum_fintype] using mk_metric_le_liminf_tsum s r hr t ht hst m, -end +by simpa only [tsum_fintype] using mk_metric_le_liminf_tsum s r hr t ht hst m /-! ### Hausdorff measure and Hausdorff dimension @@ -568,8 +566,8 @@ lemma hausdorff_measure_apply (d : ℝ) (s : set X) : mk_metric_apply _ _ /-- To bound the Hausdorff measure of a set, one may use coverings with maximum diameter tending -to `0`, indexed by any sequence of encodable types. -/ -lemma hausdorff_measure_le_liminf_tsum {β : Type*} {ι : β → Type*} [hι : ∀ n, encodable (ι n)] +to `0`, indexed by any sequence of countable types. -/ +lemma hausdorff_measure_le_liminf_tsum {β : Type*} {ι : β → Type*} [hι : ∀ n, countable (ι n)] (d : ℝ) (s : set X) {l : filter β} (r : β → ℝ≥0∞) (hr : tendsto r l (𝓝 0)) (t : Π (n : β), ι n → set X) (ht : ∀ᶠ n in l, ∀ i, diam (t n i) ≤ r n) (hst : ∀ᶠ n in l, s ⊆ ⋃ i, t n i) : diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index 0249bbb9a3c6c..d853b83c1f69b 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -276,7 +276,7 @@ lemma smul_map_diagonal_volume_pi [decidable_eq ι] {D : ι → ℝ} (h : det (d begin refine (measure.pi_eq (λ s hs, _)).symm, simp only [det_diagonal, measure.coe_smul, algebra.id.smul_eq_mul, pi.smul_apply], - rw [measure.map_apply _ (measurable_set.univ_pi_fintype hs)], + rw [measure.map_apply _ (measurable_set.univ_pi hs)], swap, { exact continuous.measurable (linear_map.continuous_on_pi _) }, have : (matrix.to_lin' (diagonal D)) ⁻¹' (set.pi set.univ (λ (i : ι), s i)) = set.pi set.univ (λ (i : ι), ((*) (D i)) ⁻¹' (s i)), diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 9a1b595ebae31..de48daf41be0b 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -258,7 +258,7 @@ have A : μ t = μ s, from h₂.antisymm (measure_mono h₁), have B : μ s ≠ ∞, from A ▸ ht, h₁.eventually_le.antisymm $ ae_le_set.2 $ by rw [measure_diff h₁ hsm B, A, tsub_self] -lemma measure_Union_congr_of_subset [encodable β] {s : β → set α} {t : β → set α} +lemma measure_Union_congr_of_subset [countable β] {s : β → set α} {t : β → set α} (hsub : ∀ b, s b ⊆ t b) (h_le : ∀ b, μ (t b) ≤ μ (s b)) : μ (⋃ b, s b) = μ (⋃ b, t b) := begin @@ -294,7 +294,7 @@ begin exact measure_Union_congr_of_subset (bool.forall_bool.2 ⟨ht, hs⟩) (bool.forall_bool.2 ⟨htμ, hsμ⟩) end -@[simp] lemma measure_Union_to_measurable [encodable β] (s : β → set α) : +@[simp] lemma measure_Union_to_measurable [countable β] (s : β → set α) : μ (⋃ b, to_measurable μ (s b)) = μ (⋃ b, s b) := eq.symm $ measure_Union_congr_of_subset (λ b, subset_to_measurable _ _) (λ b, (measure_to_measurable _).le) @@ -376,9 +376,10 @@ end /-- Continuity from below: the measure of the union of a directed sequence of (not necessarily -measurable) sets is the supremum of the measures. -/ -lemma measure_Union_eq_supr [encodable ι] {s : ι → set α} (hd : directed (⊆) s) : +lemma measure_Union_eq_supr [countable ι] {s : ι → set α} (hd : directed (⊆) s) : μ (⋃ i, s i) = ⨆ i, μ (s i) := begin + casesI nonempty_encodable ι, -- WLOG, `ι = ℕ` generalize ht : function.extend encodable.encode s ⊥ = t, replace hd : directed (⊆) t := ht ▸ hd.extend_bot encodable.encode_injective, @@ -419,7 +420,7 @@ end /-- Continuity from above: the measure of the intersection of a decreasing sequence of measurable sets is the infimum of the measures. -/ -lemma measure_Inter_eq_infi [encodable ι] {s : ι → set α} +lemma measure_Inter_eq_infi [countable ι] {s : ι → set α} (h : ∀ i, measurable_set (s i)) (hd : directed (⊇) s) (hfin : ∃ i, μ (s i) ≠ ∞) : μ (⋂ i, s i) = (⨅ i, μ (s i)) := begin @@ -442,7 +443,7 @@ end /-- Continuity from below: the measure of the union of an increasing sequence of measurable sets is the limit of the measures. -/ -lemma tendsto_measure_Union [semilattice_sup ι] [encodable ι] {s : ι → set α} (hm : monotone s) : +lemma tendsto_measure_Union [semilattice_sup ι] [countable ι] {s : ι → set α} (hm : monotone s) : tendsto (μ ∘ s) at_top (𝓝 (μ (⋃ n, s n))) := begin rw measure_Union_eq_supr (directed_of_sup hm), @@ -451,7 +452,7 @@ end /-- Continuity from above: the measure of the intersection of a decreasing sequence of measurable sets is the limit of the measures. -/ -lemma tendsto_measure_Inter [encodable ι] [semilattice_sup ι] {s : ι → set α} +lemma tendsto_measure_Inter [countable ι] [semilattice_sup ι] {s : ι → set α} (hs : ∀ n, measurable_set (s n)) (hm : antitone s) (hf : ∃ i, μ (s i) ≠ ∞) : tendsto (μ ∘ s) at_top (𝓝 (μ (⋂ n, s n))) := begin @@ -1292,7 +1293,7 @@ begin apply measure_union_le end -lemma restrict_Union_apply_ae [encodable ι] {s : ι → set α} +lemma restrict_Union_apply_ae [countable ι] {s : ι → set α} (hd : pairwise (ae_disjoint μ on s)) (hm : ∀ i, null_measurable_set (s i) μ) {t : set α} (ht : measurable_set t) : μ.restrict (⋃ i, s i) t = ∑' i, μ.restrict (s i) t := @@ -1302,12 +1303,12 @@ begin (λ i, (ht.null_measurable_set.inter (hm i))) end -lemma restrict_Union_apply [encodable ι] {s : ι → set α} (hd : pairwise (disjoint on s)) +lemma restrict_Union_apply [countable ι] {s : ι → set α} (hd : pairwise (disjoint on s)) (hm : ∀ i, measurable_set (s i)) {t : set α} (ht : measurable_set t) : μ.restrict (⋃ i, s i) t = ∑' i, μ.restrict (s i) t := restrict_Union_apply_ae hd.ae_disjoint (λ i, (hm i).null_measurable_set) ht -lemma restrict_Union_apply_eq_supr [encodable ι] {s : ι → set α} +lemma restrict_Union_apply_eq_supr [countable ι] {s : ι → set α} (hd : directed (⊆) s) {t : set α} (ht : measurable_set t) : μ.restrict (⋃ i, s i) t = ⨆ i, μ.restrict (s i) t := begin @@ -1377,7 +1378,7 @@ begin rw [restrict_union_congr, ← hs] end -lemma restrict_Union_congr [encodable ι] {s : ι → set α} : +lemma restrict_Union_congr [countable ι] {s : ι → set α} : μ.restrict (⋃ i, s i) = ν.restrict (⋃ i, s i) ↔ ∀ i, μ.restrict (s i) = ν.restrict (s i) := begin @@ -1426,7 +1427,7 @@ end /-- Two measures are equal if they have equal restrictions on a spanning collection of sets (formulated using `Union`). -/ -lemma ext_iff_of_Union_eq_univ [encodable ι] {s : ι → set α} (hs : (⋃ i, s i) = univ) : +lemma ext_iff_of_Union_eq_univ [countable ι] {s : ι → set α} (hs : (⋃ i, s i) = univ) : μ = ν ↔ ∀ i, μ.restrict (s i) = ν.restrict (s i) := by rw [← restrict_Union_congr, hs, restrict_univ, restrict_univ] @@ -1570,7 +1571,7 @@ to_measure_apply _ _ hs lemma le_sum (μ : ι → measure α) (i : ι) : μ i ≤ sum μ := λ s hs, by simp only [sum_apply μ hs, ennreal.le_tsum i] -@[simp] lemma sum_apply_eq_zero [encodable ι] {μ : ι → measure α} {s : set α} : +@[simp] lemma sum_apply_eq_zero [countable ι] {μ : ι → measure α} {s : set α} : sum μ s = 0 ↔ ∀ i, μ i s = 0 := begin refine ⟨λ h i, nonpos_iff_eq_zero.1 $ h ▸ le_iff'.1 (le_sum μ i) _, λ h, nonpos_iff_eq_zero.1 _⟩, @@ -1583,7 +1584,7 @@ lemma sum_apply_eq_zero' {μ : ι → measure α} {s : set α} (hs : measurable_ sum μ s = 0 ↔ ∀ i, μ i s = 0 := by simp [hs] -lemma ae_sum_iff [encodable ι] {μ : ι → measure α} {p : α → Prop} : +lemma ae_sum_iff [countable ι] {μ : ι → measure α} {p : α → Prop} : (∀ᵐ x ∂(sum μ), p x) ↔ ∀ i, ∀ᵐ x ∂(μ i), p x := sum_apply_eq_zero @@ -1598,7 +1599,7 @@ by { ext1 s hs, simp only [sum_apply, finset_sum_apply, hs, tsum_fintype] } sum (λ i : s, μ i) = ∑ i in s, μ i := by rw [sum_fintype, finset.sum_coe_sort s μ] -@[simp] lemma ae_sum_eq [encodable ι] (μ : ι → measure α) : (sum μ).ae = ⨆ i, (μ i).ae := +@[simp] lemma ae_sum_eq [countable ι] (μ : ι → measure α) : (sum μ).ae = ⨆ i, (μ i).ae := filter.ext $ λ s, ae_sum_iff.trans mem_supr.symm @[simp] lemma sum_bool (f : bool → measure α) : sum f = f tt + f ff := @@ -1631,9 +1632,9 @@ begin tsum_add ennreal.summable ennreal.summable], end -/-- If `f` is a map with encodable codomain, then `μ.map f` is the sum of Dirac measures -/ -lemma map_eq_sum [encodable β] [measurable_singleton_class β] - (μ : measure α) (f : α → β) (hf : measurable f) : +/-- If `f` is a map with countable codomain, then `μ.map f` is a sum of Dirac measures. -/ +lemma map_eq_sum [countable β] [measurable_singleton_class β] (μ : measure α) (f : α → β) + (hf : measurable f) : μ.map f = sum (λ b : β, μ (f ⁻¹' {b}) • dirac b) := begin ext1 s hs, @@ -1642,25 +1643,25 @@ begin tsum_subtype s (λ b, μ (f ⁻¹' {b})), ← indicator_mul_right s (λ b, μ (f ⁻¹' {b}))] end -/-- A measure on an encodable type is a sum of dirac measures. -/ -@[simp] lemma sum_smul_dirac [encodable α] [measurable_singleton_class α] (μ : measure α) : +/-- A measure on a countable type is a sum of Dirac measures. -/ +@[simp] lemma sum_smul_dirac [countable α] [measurable_singleton_class α] (μ : measure α) : sum (λ a, μ {a} • dirac a) = μ := by simpa using (map_eq_sum μ id measurable_id).symm omit m0 end sum -lemma restrict_Union_ae [encodable ι] {s : ι → set α} (hd : pairwise (ae_disjoint μ on s)) +lemma restrict_Union_ae [countable ι] {s : ι → set α} (hd : pairwise (ae_disjoint μ on s)) (hm : ∀ i, null_measurable_set (s i) μ) : μ.restrict (⋃ i, s i) = sum (λ i, μ.restrict (s i)) := ext $ λ t ht, by simp only [sum_apply _ ht, restrict_Union_apply_ae hd hm ht] -lemma restrict_Union [encodable ι] {s : ι → set α} (hd : pairwise (disjoint on s)) +lemma restrict_Union [countable ι] {s : ι → set α} (hd : pairwise (disjoint on s)) (hm : ∀ i, measurable_set (s i)) : μ.restrict (⋃ i, s i) = sum (λ i, μ.restrict (s i)) := restrict_Union_ae hd.ae_disjoint (λ i, (hm i).null_measurable_set) -lemma restrict_Union_le [encodable ι] {s : ι → set α} : +lemma restrict_Union_le [countable ι] {s : ι → set α} : μ.restrict (⋃ i, s i) ≤ sum (λ i, μ.restrict (s i)) := begin intros t ht, @@ -1973,7 +1974,7 @@ begin { simp [map_of_not_ae_measurable h] } end -@[simp] lemma ae_restrict_Union_eq [encodable ι] (s : ι → set α) : +@[simp] lemma ae_restrict_Union_eq [countable ι] (s : ι → set α) : (μ.restrict (⋃ i, s i)).ae = ⨆ i, (μ.restrict (s i)).ae := le_antisymm (ae_sum_eq (λ i, μ.restrict (s i)) ▸ ae_mono restrict_Union_le) $ supr_le $ λ i, ae_mono $ restrict_mono (subset_Union s i) le_rfl @@ -2537,7 +2538,7 @@ monotone_accumulate lemma measurable_spanning_sets (μ : measure α) [sigma_finite μ] (i : ℕ) : measurable_set (spanning_sets μ i) := -measurable_set.Union $ λ j, measurable_set.Union_Prop $ +measurable_set.Union $ λ j, measurable_set.Union $ λ hij, μ.to_finite_spanning_sets_in.set_mem j lemma measure_spanning_sets_lt_top (μ : measure α) [sigma_finite μ] (i : ℕ) : @@ -2745,10 +2746,10 @@ begin exact (measure_mono $ inter_subset_left _ _).trans_lt (measure_spanning_sets_lt_top μ i) end -instance sum.sigma_finite {ι} [fintype ι] (μ : ι → measure α) [∀ i, sigma_finite (μ i)] : +instance sum.sigma_finite {ι} [finite ι] (μ : ι → measure α) [∀ i, sigma_finite (μ i)] : sigma_finite (sum μ) := begin - haveI : encodable ι := fintype.to_encodable ι, + casesI nonempty_fintype ι, have : ∀ n, measurable_set (⋂ (i : ι), spanning_sets (μ i) n) := λ n, measurable_set.Inter (λ i, measurable_spanning_sets (μ i) n), refine ⟨⟨⟨λ n, ⋂ i, spanning_sets (μ i) n, λ _, trivial, λ n, _, _⟩⟩⟩, @@ -2903,7 +2904,7 @@ lemma is_locally_finite_measure_of_is_finite_measure_on_compacts [topological_sp exact ⟨K, K_mem, K_compact.measure_lt_top⟩, end⟩ -lemma exists_pos_measure_of_cover [encodable ι] {U : ι → set α} (hU : (⋃ i, U i) = univ) +lemma exists_pos_measure_of_cover [countable ι] {U : ι → set α} (hU : (⋃ i, U i) = univ) (hμ : μ ≠ 0) : ∃ i, 0 < μ (U i) := begin contrapose! hμ with H, diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index f47853e491359..d406543fa6104 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -174,7 +174,7 @@ by simpa only [← measure_eq_trim] using μ.to_outer_measure.exists_measurable_ /-- For every set `s` and a countable collection of measures `μ i` there exists a measurable superset `t ⊇ s` such that each measure `μ i` takes the same value on `s` and `t`. -/ -lemma exists_measurable_superset_forall_eq {ι} [encodable ι] (μ : ι → measure α) (s : set α) : +lemma exists_measurable_superset_forall_eq {ι} [countable ι] (μ : ι → measure α) (s : set α) : ∃ t, s ⊆ t ∧ measurable_set t ∧ ∀ i, μ i t = μ i s := by simpa only [← measure_eq_trim] using outer_measure.exists_measurable_superset_forall_eq_trim (λ i, (μ i).to_outer_measure) s @@ -192,16 +192,12 @@ lemma exists_measurable_superset_iff_measure_eq_zero : (∃ t, s ⊆ t ∧ measurable_set t ∧ μ t = 0) ↔ μ s = 0 := ⟨λ ⟨t, hst, _, ht⟩, measure_mono_null hst ht, exists_measurable_superset_of_null⟩ -theorem measure_Union_le [encodable β] (s : β → set α) : μ (⋃ i, s i) ≤ ∑' i, μ (s i) := +theorem measure_Union_le [countable β] (s : β → set α) : μ (⋃ i, s i) ≤ ∑' i, μ (s i) := μ.to_outer_measure.Union _ lemma measure_bUnion_le {s : set β} (hs : s.countable) (f : β → set α) : μ (⋃ b ∈ s, f b) ≤ ∑' p : s, μ (f p) := -begin - haveI := hs.to_encodable, - rw [bUnion_eq_Union], - apply measure_Union_le -end +by { haveI := hs.to_subtype, rw bUnion_eq_Union, apply measure_Union_le } lemma measure_bUnion_finset_le (s : finset β) (f : β → set α) : μ (⋃ b ∈ s, f b) ≤ ∑ p in s, μ (f p) := @@ -222,11 +218,10 @@ begin apply ennreal.sum_lt_top, simpa only [finite.mem_to_finset] end -lemma measure_Union_null [encodable β] {s : β → set α} : - (∀ i, μ (s i) = 0) → μ (⋃ i, s i) = 0 := +lemma measure_Union_null [countable β] {s : β → set α} : (∀ i, μ (s i) = 0) → μ (⋃ i, s i) = 0 := μ.to_outer_measure.Union_null -@[simp] lemma measure_Union_null_iff [encodable ι] {s : ι → set α} : +@[simp] lemma measure_Union_null_iff [countable ι] {s : ι → set α} : μ (⋃ i, s i) = 0 ↔ ∀ i, μ (s i) = 0 := μ.to_outer_measure.Union_null_iff @@ -265,7 +260,7 @@ lemma measure_union_ne_top (hs : μ s ≠ ∞) (ht : μ t ≠ ∞) : μ (s ∪ t not_iff_not.1 $ by simp only [← lt_top_iff_ne_top, ← ne.def, not_or_distrib, measure_union_lt_top_iff] -lemma exists_measure_pos_of_not_measure_Union_null [encodable β] {s : β → set α} +lemma exists_measure_pos_of_not_measure_Union_null [countable β] {s : β → set α} (hs : μ (⋃ n, s n) ≠ 0) : ∃ n, 0 < μ (s n) := begin contrapose! hs, @@ -328,11 +323,8 @@ instance : countable_Inter_filter μ.ae := exact (measure_bUnion_null_iff hSc).2 hS end⟩ -lemma ae_imp_iff {p : α → Prop} {q : Prop} : (∀ᵐ x ∂μ, q → p x) ↔ (q → ∀ᵐ x ∂μ, p x) := -filter.eventually_imp_distrib_left - -lemma ae_all_iff [encodable ι] {p : α → ι → Prop} : - (∀ᵐ a ∂ μ, ∀ i, p a i) ↔ (∀ i, ∀ᵐ a ∂ μ, p a i) := +lemma ae_all_iff {ι : Sort*} [countable ι] {p : α → ι → Prop} : + (∀ᵐ a ∂ μ, ∀ i, p a i) ↔ ∀ i, ∀ᵐ a ∂ μ, p a i := eventually_countable_forall lemma ae_ball_iff {S : set ι} (hS : S.countable) {p : Π (x : α) (i ∈ S), Prop} : diff --git a/src/measure_theory/measure/mutually_singular.lean b/src/measure_theory/measure/mutually_singular.lean index c1d7dcc9a56e4..33b5e4ce06461 100644 --- a/src/measure_theory/measure/mutually_singular.lean +++ b/src/measure_theory/measure/mutually_singular.lean @@ -63,7 +63,7 @@ let ⟨s, hs, h₁, h₂⟩ := h in ⟨s, hs, hμ h₁, hν h₂⟩ lemma mono (h : μ₁ ⊥ₘ ν₁) (hμ : μ₂ ≤ μ₁) (hν : ν₂ ≤ ν₁) : μ₂ ⊥ₘ ν₂ := h.mono_ac hμ.absolutely_continuous hν.absolutely_continuous -@[simp] lemma sum_left {ι : Type*} [encodable ι] {μ : ι → measure α} : +@[simp] lemma sum_left {ι : Type*} [countable ι] {μ : ι → measure α} : (sum μ) ⊥ₘ ν ↔ ∀ i, μ i ⊥ₘ ν := begin refine ⟨λ h i, h.mono (le_sum _ _) le_rfl, λ H, _⟩, @@ -74,7 +74,7 @@ begin { rwa [compl_Inter, measure_Union_null_iff], } end -@[simp] lemma sum_right {ι : Type*} [encodable ι] {ν : ι → measure α} : +@[simp] lemma sum_right {ι : Type*} [countable ι] {ν : ι → measure α} : μ ⊥ₘ sum ν ↔ ∀ i, μ ⊥ₘ ν i := comm.trans $ sum_left.trans $ forall_congr $ λ i, comm diff --git a/src/measure_theory/measure/null_measurable.lean b/src/measure_theory/measure/null_measurable.lean index 5fac6e285d8ab..b55766e7ac003 100644 --- a/src/measure_theory/measure/null_measurable.lean +++ b/src/measure_theory/measure/null_measurable.lean @@ -109,8 +109,9 @@ protected lemma congr (hs : null_measurable_set s μ) (h : s =ᵐ[μ] t) : null_measurable_set t μ := let ⟨s', hm, hs'⟩ := hs in ⟨s', hm, h.symm.trans hs'⟩ -protected lemma Union [encodable ι] {s : ι → set α} - (h : ∀ i, null_measurable_set (s i) μ) : null_measurable_set (⋃ i, s i) μ := +protected lemma Union {ι : Sort*} [countable ι] {s : ι → set α} + (h : ∀ i, null_measurable_set (s i) μ) : + null_measurable_set (⋃ i, s i) μ := measurable_set.Union h protected lemma bUnion_decode₂ [encodable ι] ⦃f : ι → set α⦄ (h : ∀ i, null_measurable_set (f i) μ) @@ -125,15 +126,8 @@ protected lemma sUnion {s : set (set α)} (hs : s.countable) (h : ∀ t ∈ s, n null_measurable_set (⋃₀ s) μ := by { rw sUnion_eq_bUnion, exact measurable_set.bUnion hs h } -lemma Union_Prop {p : Prop} {f : p → set α} (hf : ∀ i, null_measurable_set (f i) μ) : - null_measurable_set (⋃ i, f i) μ := -measurable_set.Union_Prop hf - -lemma Union_fintype [fintype ι] {f : ι → set α} (h : ∀ b, null_measurable_set (f b) μ) : - null_measurable_set (⋃ b, f b) μ := -measurable_set.Union_fintype h - -protected lemma Inter [encodable ι] {f : ι → set α} (h : ∀ i, null_measurable_set (f i) μ) : +protected lemma Inter {ι : Sort*} [countable ι] {f : ι → set α} + (h : ∀ i, null_measurable_set (f i) μ) : null_measurable_set (⋂ i, f i) μ := measurable_set.Inter h @@ -145,14 +139,6 @@ protected lemma sInter {s : set (set α)} (hs : s.countable) (h : ∀ t ∈ s, n null_measurable_set (⋂₀ s) μ := measurable_set.sInter hs h -lemma Inter_Prop {p : Prop} {f : p → set α} (hf : ∀ b, null_measurable_set (f b) μ) : - null_measurable_set (⋂ b, f b) μ := -measurable_set.Inter_Prop hf - -lemma Inter_fintype [fintype ι] {f : ι → set α} (h : ∀ b, null_measurable_set (f b) μ) : - null_measurable_set (⋂ b, f b) μ := -measurable_set.Inter_fintype h - @[simp] protected lemma union (hs : null_measurable_set s μ) (ht : null_measurable_set t μ) : null_measurable_set (s ∪ t) μ := hs.union ht @@ -215,7 +201,7 @@ end null_measurable_set /-- If `sᵢ` is a countable family of (null) measurable pairwise `μ`-a.e. disjoint sets, then there exists a subordinate family `tᵢ ⊆ sᵢ` of measurable pairwise disjoint sets such that `tᵢ =ᵐ[μ] sᵢ`. -/ -lemma exists_subordinate_pairwise_disjoint [encodable ι] {s : ι → set α} +lemma exists_subordinate_pairwise_disjoint [countable ι] {s : ι → set α} (h : ∀ i, null_measurable_set (s i) μ) (hd : pairwise (ae_disjoint μ on s)) : ∃ t : ι → set α, (∀ i, t i ⊆ s i) ∧ (∀ i, s i =ᵐ[μ] t i) ∧ (∀ i, measurable_set (t i)) ∧ pairwise (disjoint on t) := @@ -228,7 +214,7 @@ begin λ i j h, h.mono (diff_subset_diff_left (ht_sub i)) (diff_subset_diff_left (ht_sub j))⟩ end -lemma measure_Union {m0 : measurable_space α} {μ : measure α} [encodable ι] {f : ι → set α} +lemma measure_Union {m0 : measurable_space α} {μ : measure α} [countable ι] {f : ι → set α} (hn : pairwise (disjoint on f)) (h : ∀ i, measurable_set (f i)) : μ (⋃ i, f i) = ∑' i, μ (f i) := begin @@ -239,8 +225,8 @@ begin { exact μ.m_Union } end -lemma measure_Union₀ [encodable ι] {f : ι → set α} - (hd : pairwise (ae_disjoint μ on f)) (h : ∀ i, null_measurable_set (f i) μ) : +lemma measure_Union₀ [countable ι] {f : ι → set α} (hd : pairwise (ae_disjoint μ on f)) + (h : ∀ i, null_measurable_set (f i) μ) : μ (⋃ i, f i) = ∑' i, μ (f i) := begin rcases exists_subordinate_pairwise_disjoint h hd with ⟨t, ht_sub, ht_eq, htm, htd⟩, diff --git a/src/measure_theory/measure/outer_measure.lean b/src/measure_theory/measure/outer_measure.lean index 1b0882ece5ff1..577bd61ffb9c6 100644 --- a/src/measure_theory/measure/outer_measure.lean +++ b/src/measure_theory/measure/outer_measure.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Mario Carneiro -/ import analysis.specific_limits.basic import measure_theory.pi_system +import data.countable.basic import data.fin.vec_notation import topology.algebra.infinite_sum @@ -52,7 +53,7 @@ outer measure, Carathéodory-measurable, Carathéodory's criterion noncomputable theory -open set finset function filter encodable topological_space (second_countable_topology) +open set finset function filter topological_space (second_countable_topology) open_locale classical big_operators nnreal topological_space ennreal measure_theory namespace measure_theory @@ -86,16 +87,15 @@ lemma pos_of_subset_ne_zero (m : outer_measure α) {a b : set α} (hs : a ⊆ b) 0 < m b := (lt_of_lt_of_le (pos_iff_ne_zero.mpr hnz) (m.mono hs)) -protected theorem Union (m : outer_measure α) - {β} [encodable β] (s : β → set α) : +protected theorem Union (m : outer_measure α) {β} [countable β] (s : β → set α) : m (⋃ i, s i) ≤ ∑' i, m (s i) := rel_supr_tsum m m.empty (≤) m.Union_nat s -lemma Union_null [encodable β] (m : outer_measure α) {s : β → set α} (h : ∀ i, m (s i) = 0) : +lemma Union_null [countable β] (m : outer_measure α) {s : β → set α} (h : ∀ i, m (s i) = 0) : m (⋃ i, s i) = 0 := by simpa [h] using m.Union s -@[simp] lemma Union_null_iff [encodable β] (m : outer_measure α) {s : β → set α} : +@[simp] lemma Union_null_iff [countable β] (m : outer_measure α) {s : β → set α} : m (⋃ i, s i) = 0 ↔ ∀ i, m (s i) = 0 := ⟨λ h i, m.mono_null (subset_Union _ _) h, m.Union_null⟩ @@ -521,7 +521,7 @@ let μ := λs, ⨅{f : ℕ → set α} (h : s ⊆ ⋃i, f i), ∑'i, m (f i) in infi_mono' $ assume hb, ⟨hs.trans hb, le_rfl⟩, Union_nat := assume s, ennreal.le_of_forall_pos_le_add $ begin assume ε hε (hb : ∑'i, μ (s i) < ∞), - rcases ennreal.exists_pos_sum_of_encodable (ennreal.coe_pos.2 hε).ne' ℕ with ⟨ε', hε', hl⟩, + rcases ennreal.exists_pos_sum_of_countable (ennreal.coe_pos.2 hε).ne' ℕ with ⟨ε', hε', hl⟩, refine le_trans _ (add_le_add_left (le_of_lt hl) _), rw ← ennreal.tsum_add, choose f hf using show @@ -1121,10 +1121,11 @@ end mono section unions include P0 m0 PU mU -lemma extend_Union {β} [encodable β] {f : β → set α} - (hd : pairwise (disjoint on f)) (hm : ∀i, P (f i)) : +lemma extend_Union {β} [countable β] {f : β → set α} (hd : pairwise (disjoint on f)) + (hm : ∀ i, P (f i)) : extend m (⋃i, f i) = ∑'i, extend m (f i) := begin + casesI nonempty_encodable β, rw [← encodable.Union_decode₂, ← tsum_Union_decode₂], { exact extend_Union_nat PU (λ n, encodable.Union_decode₂_cases P0 hm) @@ -1365,7 +1366,7 @@ end /-- If `μ i` is a countable family of outer measures, then for every set `s` there exists a measurable set `t ⊇ s` such that `μ i t = (μ i).trim s` for all `i`. -/ -lemma exists_measurable_superset_forall_eq_trim {ι} [encodable ι] (μ : ι → outer_measure α) +lemma exists_measurable_superset_forall_eq_trim {ι} [countable ι] (μ : ι → outer_measure α) (s : set α) : ∃ t, s ⊆ t ∧ measurable_set t ∧ ∀ i, μ i t = (μ i).trim s := begin choose t hst ht hμt using λ i, (μ i).exists_measurable_superset_eq_trim s, @@ -1410,12 +1411,13 @@ ext $ λ s, (trim_binop (sup_apply m₁ m₂) s).trans (sup_apply _ _ _).symm /-- `trim` sends the supremum of a countable family of outer measures to the supremum of the trimmed measures. -/ -lemma trim_supr {ι} [encodable ι] (μ : ι → outer_measure α) : - trim (⨆ i, μ i) = ⨆ i, trim (μ i) := +lemma trim_supr {ι} [countable ι] (μ : ι → outer_measure α) : trim (⨆ i, μ i) = ⨆ i, trim (μ i) := begin + simp_rw [←@supr_plift_down _ ι], ext1 s, - rcases exists_measurable_superset_forall_eq_trim (option.elim (supr μ) μ) s - with ⟨t, hst, ht, hμt⟩, + haveI : countable (option $ plift ι) := @option.countable (plift ι) _, + obtain ⟨t, hst, ht, hμt⟩ := exists_measurable_superset_forall_eq_trim + (option.elim (⨆ i, μ (plift.down i)) (μ ∘ plift.down)) s, simp only [option.forall, option.elim] at hμt, simp only [supr_apply, ← hμt.1, ← hμt.2] end diff --git a/src/measure_theory/measure/regular.lean b/src/measure_theory/measure/regular.lean index cd7069955c47a..97c39ce5cf98d 100644 --- a/src/measure_theory/measure/regular.lean +++ b/src/measure_theory/measure/regular.lean @@ -317,7 +317,7 @@ begin λ n, (inter_subset_right _ _).trans (disjointed_subset _ _), (disjoint_disjointed s.set).mono (λ k l hkl, hkl.mono inf_le_right inf_le_right), _⟩, rw [← inter_Union, Union_disjointed, s.spanning, inter_univ] }, - rcases ennreal.exists_pos_sum_of_encodable' (tsub_pos_iff_lt.2 hr).ne' ℕ with ⟨δ, δ0, hδε⟩, + rcases ennreal.exists_pos_sum_of_countable' (tsub_pos_iff_lt.2 hr).ne' ℕ with ⟨δ, δ0, hδε⟩, rw [lt_tsub_iff_right, add_comm] at hδε, have : ∀ n, ∃ U ⊇ A n, is_open U ∧ μ U < μ (A n) + δ n, { intro n, @@ -394,7 +394,7 @@ begin simp only [measure_compl_le_add_iff, *, hUo.measurable_set, hFc.measurable_set, true_and] }, -- check for disjoint unions { intros s hsd hsm H ε ε0, have ε0' : ε / 2 ≠ 0, from (ennreal.half_pos ε0).ne', - rcases ennreal.exists_pos_sum_of_encodable' ε0' ℕ with ⟨δ, δ0, hδε⟩, + rcases ennreal.exists_pos_sum_of_countable' ε0' ℕ with ⟨δ, δ0, hδε⟩, choose F hFs U hsU hFc hUo hF hU using λ n, H n (δ n) (δ0 n).ne', -- the approximating closed set is constructed by considering finitely many sets `s i`, which -- cover all the measure up to `ε/2`, approximating each of these by a closed set `F i`, and diff --git a/src/measure_theory/measure/stieltjes.lean b/src/measure_theory/measure/stieltjes.lean index d8e43bc743a1c..77600cdcf4d72 100644 --- a/src/measure_theory/measure/stieltjes.lean +++ b/src/measure_theory/measure/stieltjes.lean @@ -174,7 +174,7 @@ begin (le_infi₂ $ λ s hs, ennreal.le_of_forall_pos_le_add $ λ ε εpos h, _), let δ := ε / 2, have δpos : 0 < (δ : ℝ≥0∞), by simpa using εpos.ne', - rcases ennreal.exists_pos_sum_of_encodable δpos.ne' ℕ with ⟨ε', ε'0, hε⟩, + rcases ennreal.exists_pos_sum_of_countable δpos.ne' ℕ with ⟨ε', ε'0, hε⟩, obtain ⟨a', ha', aa'⟩ : ∃ a', f a' - f a < δ ∧ a < a', { have A : continuous_within_at (λ r, f r - f a) (Ioi a) a, { refine continuous_within_at.sub _ continuous_within_at_const, @@ -239,7 +239,7 @@ begin rw outer_measure.trim_eq_infi, refine le_infi (λ t, le_infi $ λ ht, ennreal.le_of_forall_pos_le_add $ λ ε ε0 h, _), - rcases ennreal.exists_pos_sum_of_encodable + rcases ennreal.exists_pos_sum_of_countable (ennreal.coe_pos.2 ε0).ne' ℕ with ⟨ε', ε'0, hε⟩, refine le_trans _ (add_le_add_left (le_of_lt hε) _), rw ← ennreal.tsum_add, diff --git a/src/measure_theory/measure/vector_measure.lean b/src/measure_theory/measure/vector_measure.lean index 02547d722f6ea..89a8d7b274b66 100644 --- a/src/measure_theory/measure/vector_measure.lean +++ b/src/measure_theory/measure/vector_measure.lean @@ -123,13 +123,14 @@ end variables [t2_space M] {v : vector_measure α M} {f : ℕ → set α} -lemma has_sum_of_disjoint_Union [encodable β] {f : β → set α} +lemma has_sum_of_disjoint_Union [countable β] {f : β → set α} (hf₁ : ∀ i, measurable_set (f i)) (hf₂ : pairwise (disjoint on f)) : has_sum (λ i, v (f i)) (v (⋃ i, f i)) := begin + casesI nonempty_encodable β, set g := λ i : ℕ, ⋃ (b : β) (H : b ∈ encodable.decode₂ β i), f b with hg, have hg₁ : ∀ i, measurable_set (g i), - { exact λ _, measurable_set.Union (λ b, measurable_set.Union_Prop $ λ _, hf₁ b) }, + { exact λ _, measurable_set.Union (λ b, measurable_set.Union $ λ _, hf₁ b) }, have hg₂ : pairwise (disjoint on g), { exact encodable.Union_decode₂_disjoint_on hf₂ }, have := v.of_disjoint_Union_nat hg₁ hg₂, @@ -157,7 +158,7 @@ begin exact false.elim ((hx i) ((encodable.decode₂_is_partial_inv _ _).1 hi)) } } end -lemma of_disjoint_Union [encodable β] {f : β → set α} +lemma of_disjoint_Union [countable β] {f : β → set α} (hf₁ : ∀ i, measurable_set (f i)) (hf₂ : pairwise (disjoint on f)) : v (⋃ i, f i) = ∑' i, v (f i) := (has_sum_of_disjoint_Union hf₁ hf₂).tsum_eq.symm @@ -882,10 +883,11 @@ begin { exact λ n, ha₁.inter (measurable_set.disjointed hf₁ n) } end -lemma restrict_le_restrict_encodable_Union [encodable β] {f : β → set α} +lemma restrict_le_restrict_countable_Union [countable β] {f : β → set α} (hf₁ : ∀ b, measurable_set (f b)) (hf₂ : ∀ b, v ≤[f b] w) : v ≤[⋃ b, f b] w := begin + casesI nonempty_encodable β, rw ← encodable.Union_decode₂, refine restrict_le_restrict_Union v w _ _, { intro n, measurability }, @@ -901,7 +903,7 @@ lemma restrict_le_restrict_union v ≤[i ∪ j] w := begin rw union_eq_Union, - refine restrict_le_restrict_encodable_Union v w _ _, + refine restrict_le_restrict_countable_Union v w _ _, { measurability }, { rintro (_ | _); simpa } end diff --git a/src/measure_theory/pi_system.lean b/src/measure_theory/pi_system.lean index c5177384cceeb..0af61def42504 100644 --- a/src/measure_theory/pi_system.lean +++ b/src/measure_theory/pi_system.lean @@ -471,9 +471,9 @@ lemma has_compl_iff {a} : d.has aᶜ ↔ d.has a := lemma has_univ : d.has univ := by simpa using d.has_compl d.has_empty -theorem has_Union {β} [encodable β] {f : β → set α} - (hd : pairwise (disjoint on f)) (h : ∀ i, d.has (f i)) : d.has (⋃ i, f i) := -by { rw ← encodable.Union_decode₂, exact +lemma has_Union {β} [countable β] {f : β → set α} (hd : pairwise (disjoint on f)) + (h : ∀ i, d.has (f i)) : d.has (⋃ i, f i) := +by { casesI nonempty_encodable β, rw ← encodable.Union_decode₂, exact d.has_Union_nat (encodable.Union_decode₂_disjoint_on hd) (λ n, encodable.Union_decode₂_cases d.has_empty h) } diff --git a/src/measure_theory/tactic.lean b/src/measure_theory/tactic.lean index 195407aa7b0ee..55bccc36e39fb 100644 --- a/src/measure_theory/tactic.lean +++ b/src/measure_theory/tactic.lean @@ -47,8 +47,6 @@ attribute [measurability] subsingleton.measurable_set measurable_set.Union measurable_set.Inter - measurable_set.Union_Prop - measurable_set.Inter_Prop measurable_set.union measurable_set.inter measurable_set.diff diff --git a/src/order/complete_lattice.lean b/src/order/complete_lattice.lean index 842518e2e4437..565b68593117f 100644 --- a/src/order/complete_lattice.lean +++ b/src/order/complete_lattice.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import data.bool.set +import data.ulift import data.nat.basic import order.bounds @@ -439,6 +440,12 @@ e.surjective.supr_congr _ h (f : ∀ x, f₁ (pq.mpr x) = f₂ x) : supr f₁ = supr f₂ := by { obtain rfl := propext pq, congr' with x, apply f } +lemma supr_plift_up (f : plift ι → α) : (⨆ i, f (plift.up i)) = ⨆ i, f i := +plift.up_surjective.supr_congr _ $ λ _, rfl + +lemma supr_plift_down (f : ι → α) : (⨆ i, f (plift.down i)) = ⨆ i, f i := +plift.down_surjective.supr_congr _ $ λ _, rfl + lemma supr_range' (g : β → α) (f : ι → β) : (⨆ b : range f, g b) = ⨆ i, g (f i) := by rw [supr, supr, ← image_eq_range, ← range_comp] @@ -475,6 +482,12 @@ protected lemma equiv.infi_congr {g : ι' → α} (e : ι ≃ ι') (h : ∀ x, g (pq : p ↔ q) (f : ∀ x, f₁ (pq.mpr x) = f₂ x) : infi f₁ = infi f₂ := @supr_congr_Prop αᵒᵈ _ p q f₁ f₂ pq f +lemma infi_plift_up (f : plift ι → α) : (⨅ i, f (plift.up i)) = ⨅ i, f i := +plift.up_surjective.infi_congr _ $ λ _, rfl + +lemma infi_plift_down (f : ι → α) : (⨅ i, f (plift.down i)) = ⨅ i, f i := +plift.down_surjective.infi_congr _ $ λ _, rfl + lemma infi_range' (g : β → α) (f : ι → β) : (⨅ b : range f, g b) = ⨅ i, g (f i) := @supr_range' αᵒᵈ _ _ _ _ _ diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index 2399f09b1b54a..620f0ef5b43f0 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -189,23 +189,23 @@ lemma at_top_basis_Ioi [nonempty α] [semilattice_sup α] [no_max_order α] : at_top_basis.to_has_basis (λ a ha, ⟨a, ha, Ioi_subset_Ici_self⟩) $ λ a ha, (exists_gt a).imp $ λ b hb, ⟨ha, Ici_subset_Ioi.2 hb⟩ -lemma at_top_countable_basis [nonempty α] [semilattice_sup α] [encodable α] : +lemma at_top_countable_basis [nonempty α] [semilattice_sup α] [countable α] : has_countable_basis (at_top : filter α) (λ _, true) Ici := { countable := to_countable _, .. at_top_basis } -lemma at_bot_countable_basis [nonempty α] [semilattice_inf α] [encodable α] : +lemma at_bot_countable_basis [nonempty α] [semilattice_inf α] [countable α] : has_countable_basis (at_bot : filter α) (λ _, true) Iic := { countable := to_countable _, .. at_bot_basis } @[priority 200] -instance at_top.is_countably_generated [preorder α] [encodable α] : +instance at_top.is_countably_generated [preorder α] [countable α] : (at_top : filter $ α).is_countably_generated := is_countably_generated_seq _ @[priority 200] -instance at_bot.is_countably_generated [preorder α] [encodable α] : +instance at_bot.is_countably_generated [preorder α] [countable α] : (at_bot : filter $ α).is_countably_generated := is_countably_generated_seq _ diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index 9dad20ab3b656..99daf52276edb 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -965,7 +965,7 @@ filter.sup.is_countably_generated _ _ end is_countably_generated -lemma is_countably_generated_seq [encodable β] (x : β → set α) : +lemma is_countably_generated_seq [countable β] (x : β → set α) : is_countably_generated (⨅ i, 𝓟 $ x i) := begin use [range x, countable_range x], @@ -1009,7 +1009,6 @@ begin rw [← plift.down_surjective.infi_comp], refine has_countable_basis.is_countably_generated ⟨has_basis_infi (λ n, (hs _).to_has_basis), _⟩, - haveI := encodable.of_countable (plift ι), refine (countable_range $ sigma.map (coe : finset (plift ι) → set (plift ι)) (λ _, id)).mono _, rintro ⟨I, f⟩ ⟨hI, -⟩, lift I to finset (plift ι) using hI, diff --git a/src/order/filter/countable_Inter.lean b/src/order/filter/countable_Inter.lean index 778ba22111cb6..29aa9e318bbe1 100644 --- a/src/order/filter/countable_Inter.lean +++ b/src/order/filter/countable_Inter.lean @@ -27,7 +27,7 @@ filter, countable open set filter open_locale filter -variables {ι α β : Type*} +variables {ι : Sort*} {α β : Type*} /-- A filter `l` has the countable intersection property if for any countable collection of sets `s ∈ l` their intersection belongs to `l` as well. -/ @@ -42,11 +42,10 @@ lemma countable_sInter_mem {S : set (set α)} (hSc : S.countable) : ⟨λ hS s hs, mem_of_superset hS (sInter_subset_of_mem hs), countable_Inter_filter.countable_sInter_mem' hSc⟩ -lemma countable_Inter_mem [encodable ι] {s : ι → set α} : - (⋂ i, s i) ∈ l ↔ ∀ i, s i ∈ l := +lemma countable_Inter_mem [countable ι] {s : ι → set α} : (⋂ i, s i) ∈ l ↔ ∀ i, s i ∈ l := sInter_range s ▸ (countable_sInter_mem (countable_range _)).trans forall_range_iff -lemma countable_bInter_mem {S : set ι} (hS : S.countable) {s : Π i ∈ S, set α} : +lemma countable_bInter_mem {ι : Type*} {S : set ι} (hS : S.countable) {s : Π i ∈ S, set α} : (⋂ i ∈ S, s i ‹_›) ∈ l ↔ ∀ i ∈ S, s i ‹_› ∈ l := begin rw [bInter_eq_Inter], @@ -54,58 +53,63 @@ begin exact countable_Inter_mem.trans subtype.forall end -lemma eventually_countable_forall [encodable ι] {p : α → ι → Prop} : +lemma eventually_countable_forall [countable ι] {p : α → ι → Prop} : (∀ᶠ x in l, ∀ i, p x i) ↔ ∀ i, ∀ᶠ x in l, p x i := by simpa only [filter.eventually, set_of_forall] using @countable_Inter_mem _ _ l _ _ (λ i, {x | p x i}) -lemma eventually_countable_ball {S : set ι} (hS : S.countable) {p : Π (x : α) (i ∈ S), Prop} : +lemma eventually_countable_ball {ι : Type*} {S : set ι} (hS : S.countable) + {p : Π (x : α) (i ∈ S), Prop} : (∀ᶠ x in l, ∀ i ∈ S, p x i ‹_›) ↔ ∀ i ∈ S, ∀ᶠ x in l, p x i ‹_› := by simpa only [filter.eventually, set_of_forall] - using @countable_bInter_mem _ _ l _ _ hS (λ i hi, {x | p x i hi}) + using @countable_bInter_mem _ l _ _ _ hS (λ i hi, {x | p x i hi}) -lemma eventually_le.countable_Union [encodable ι] {s t : ι → set α} (h : ∀ i, s i ≤ᶠ[l] t i) : +lemma eventually_le.countable_Union [countable ι] {s t : ι → set α} (h : ∀ i, s i ≤ᶠ[l] t i) : (⋃ i, s i) ≤ᶠ[l] ⋃ i, t i := (eventually_countable_forall.2 h).mono $ λ x hst hs, mem_Union.2 $ (mem_Union.1 hs).imp hst -lemma eventually_eq.countable_Union [encodable ι] {s t : ι → set α} (h : ∀ i, s i =ᶠ[l] t i) : +lemma eventually_eq.countable_Union [countable ι] {s t : ι → set α} (h : ∀ i, s i =ᶠ[l] t i) : (⋃ i, s i) =ᶠ[l] ⋃ i, t i := (eventually_le.countable_Union (λ i, (h i).le)).antisymm (eventually_le.countable_Union (λ i, (h i).symm.le)) -lemma eventually_le.countable_bUnion {S : set ι} (hS : S.countable) {s t : Π i ∈ S, set α} - (h : ∀ i ∈ S, s i ‹_› ≤ᶠ[l] t i ‹_›) : (⋃ i ∈ S, s i ‹_›) ≤ᶠ[l] ⋃ i ∈ S, t i ‹_› := +lemma eventually_le.countable_bUnion {ι : Type*} {S : set ι} (hS : S.countable) + {s t : Π i ∈ S, set α} (h : ∀ i ∈ S, s i ‹_› ≤ᶠ[l] t i ‹_›) : + (⋃ i ∈ S, s i ‹_›) ≤ᶠ[l] ⋃ i ∈ S, t i ‹_› := begin simp only [bUnion_eq_Union], haveI := hS.to_encodable, exact eventually_le.countable_Union (λ i, h i i.2) end -lemma eventually_eq.countable_bUnion {S : set ι} (hS : S.countable) {s t : Π i ∈ S, set α} - (h : ∀ i ∈ S, s i ‹_› =ᶠ[l] t i ‹_›) : (⋃ i ∈ S, s i ‹_›) =ᶠ[l] ⋃ i ∈ S, t i ‹_› := +lemma eventually_eq.countable_bUnion {ι : Type*} {S : set ι} (hS : S.countable) + {s t : Π i ∈ S, set α} (h : ∀ i ∈ S, s i ‹_› =ᶠ[l] t i ‹_›) : + (⋃ i ∈ S, s i ‹_›) =ᶠ[l] ⋃ i ∈ S, t i ‹_› := (eventually_le.countable_bUnion hS (λ i hi, (h i hi).le)).antisymm (eventually_le.countable_bUnion hS (λ i hi, (h i hi).symm.le)) -lemma eventually_le.countable_Inter [encodable ι] {s t : ι → set α} (h : ∀ i, s i ≤ᶠ[l] t i) : +lemma eventually_le.countable_Inter [countable ι] {s t : ι → set α} (h : ∀ i, s i ≤ᶠ[l] t i) : (⋂ i, s i) ≤ᶠ[l] ⋂ i, t i := (eventually_countable_forall.2 h).mono $ λ x hst hs, mem_Inter.2 $ λ i, hst _ (mem_Inter.1 hs i) -lemma eventually_eq.countable_Inter [encodable ι] {s t : ι → set α} (h : ∀ i, s i =ᶠ[l] t i) : +lemma eventually_eq.countable_Inter [countable ι] {s t : ι → set α} (h : ∀ i, s i =ᶠ[l] t i) : (⋂ i, s i) =ᶠ[l] ⋂ i, t i := (eventually_le.countable_Inter (λ i, (h i).le)).antisymm (eventually_le.countable_Inter (λ i, (h i).symm.le)) -lemma eventually_le.countable_bInter {S : set ι} (hS : S.countable) {s t : Π i ∈ S, set α} - (h : ∀ i ∈ S, s i ‹_› ≤ᶠ[l] t i ‹_›) : (⋂ i ∈ S, s i ‹_›) ≤ᶠ[l] ⋂ i ∈ S, t i ‹_› := +lemma eventually_le.countable_bInter {ι : Type*} {S : set ι} (hS : S.countable) + {s t : Π i ∈ S, set α} (h : ∀ i ∈ S, s i ‹_› ≤ᶠ[l] t i ‹_›) : + (⋂ i ∈ S, s i ‹_›) ≤ᶠ[l] ⋂ i ∈ S, t i ‹_› := begin simp only [bInter_eq_Inter], haveI := hS.to_encodable, exact eventually_le.countable_Inter (λ i, h i i.2) end -lemma eventually_eq.countable_bInter {S : set ι} (hS : S.countable) {s t : Π i ∈ S, set α} - (h : ∀ i ∈ S, s i ‹_› =ᶠ[l] t i ‹_›) : (⋂ i ∈ S, s i ‹_›) =ᶠ[l] ⋂ i ∈ S, t i ‹_› := +lemma eventually_eq.countable_bInter {ι : Type*} {S : set ι} (hS : S.countable) + {s t : Π i ∈ S, set α} (h : ∀ i ∈ S, s i ‹_› =ᶠ[l] t i ‹_›) : + (⋂ i ∈ S, s i ‹_›) =ᶠ[l] ⋂ i ∈ S, t i ‹_› := (eventually_le.countable_bInter hS (λ i hi, (h i hi).le)).antisymm (eventually_le.countable_bInter hS (λ i hi, (h i hi).symm.le)) diff --git a/src/order/filter/ennreal.lean b/src/order/filter/ennreal.lean index 2673f8643da8e..79274175fe617 100644 --- a/src/order/filter/ennreal.lean +++ b/src/order/filter/ennreal.lean @@ -100,7 +100,7 @@ lemma limsup_add_le [countable_Inter_filter f] (u v : α → ℝ≥0∞) : Inf_le ((eventually_le_limsup u).mp ((eventually_le_limsup v).mono (λ _ hxg hxf, add_le_add hxf hxg))) -lemma limsup_liminf_le_liminf_limsup {β} [encodable β] {f : filter α} [countable_Inter_filter f] +lemma limsup_liminf_le_liminf_limsup {β} [countable β] {f : filter α} [countable_Inter_filter f] {g : filter β} (u : α → β → ℝ≥0∞) : f.limsup (λ (a : α), g.liminf (λ (b : β), u a b)) ≤ g.liminf (λ b, f.limsup (λ a, u a b)) := begin diff --git a/src/probability/hitting_time.lean b/src/probability/hitting_time.lean index 89b76ec997164..7fe6710886ab6 100644 --- a/src/probability/hitting_time.lean +++ b/src/probability/hitting_time.lean @@ -179,7 +179,7 @@ end inequalities /-- A discrete hitting time is a stopping time. -/ lemma hitting_is_stopping_time - [conditionally_complete_linear_order ι] [is_well_order ι (<)] [encodable ι] + [conditionally_complete_linear_order ι] [is_well_order ι (<)] [countable ι] [topological_space β] [pseudo_metrizable_space β] [measurable_space β] [borel_space β] {f : filtration ι m} {u : ι → Ω → β} {s : set β} {n n' : ι} (hu : adapted f u) (hs : measurable_set s) : @@ -194,7 +194,7 @@ begin rw [set.mem_set_of_eq, hitting_le_iff_of_lt _ hi], simp only [set.mem_Icc, exists_prop, set.mem_Union, set.mem_preimage], }, rw h_set_eq_Union, - exact measurable_set.Union (λ j, measurable_set.Union_Prop $ + exact measurable_set.Union (λ j, measurable_set.Union $ λ hj, f.mono hj.2 _ ((hu j).measurable hs)) } end @@ -212,7 +212,7 @@ end /-- The hitting time of a discrete process with the starting time indexed by a stopping time is a stopping time. -/ lemma is_stopping_time_hitting_is_stopping_time - [conditionally_complete_linear_order ι] [is_well_order ι (<)] [encodable ι] + [conditionally_complete_linear_order ι] [is_well_order ι (<)] [countable ι] [topological_space ι] [order_topology ι] [first_countable_topology ι] [topological_space β] [pseudo_metrizable_space β] [measurable_space β] [borel_space β] {f : filtration ι m} {u : ι → Ω → β} {τ : Ω → ι} (hτ : is_stopping_time f τ) @@ -232,7 +232,7 @@ begin rintro m hm rfl, exact lt_of_lt_of_le hm (le_hitting (hτbdd _) _) }, rw [h₁, h₂, set.union_empty], - exact measurable_set.Union (λ i, measurable_set.Union_Prop + exact measurable_set.Union (λ i, measurable_set.Union (λ hi, (f.mono hi _ (hτ.measurable_set_eq i)).inter (hitting_is_stopping_time hf hs n))), end diff --git a/src/probability/independence.lean b/src/probability/independence.lean index 12d65123b7b09..e8262b7feb4d0 100644 --- a/src/probability/independence.lean +++ b/src/probability/independence.lean @@ -564,7 +564,7 @@ begin { rw [generate_from_pi.symm, comap_generate_from], { congr' with s, simp only [set.mem_image, set.mem_set_of_eq, exists_prop], }, - { exact finset.fintype_coe_sort S, }, }, + { apply_instance } }, let πTβ := (set.pi (set.univ : set T) '' (set.pi (set.univ : set T) (λ i, {s : set (β i) | measurable_set[m i] s}))), let πT := {s : set Ω | ∃ t ∈ πTβ, (λ a (i : T), f i a) ⁻¹' t = s}, @@ -573,7 +573,7 @@ begin { rw [generate_from_pi.symm, comap_generate_from], { congr' with s, simp only [set.mem_image, set.mem_set_of_eq, exists_prop], }, - { exact finset.fintype_coe_sort T, }, }, + { apply_instance } }, -- To prove independence, we prove independence of the generating π-systems. refine indep_sets.indep (measurable.comap_le (measurable_pi_iff.mpr (λ i, hf_meas i))) diff --git a/src/probability/stopping.lean b/src/probability/stopping.lean index b45d42c26a784..1bf78c0fe21ce 100644 --- a/src/probability/stopping.lean +++ b/src/probability/stopping.lean @@ -521,7 +521,7 @@ begin exact @measurable_set.empty _ (f i), }, end -protected lemma measurable_set_eq_of_encodable [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_eq_of_encodable [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[f i] {ω | τ ω = i} := hτ.measurable_set_eq_of_countable (set.to_countable _) i @@ -535,7 +535,7 @@ begin exact (hτ.measurable_set_le i).diff (hτ.measurable_set_eq_of_countable h_countable i), end -protected lemma measurable_set_lt_of_encodable [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_lt_of_encodable [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[f i] {ω | τ ω < i} := hτ.measurable_set_lt_of_countable (set.to_countable _) i @@ -550,7 +550,7 @@ begin end protected lemma measurable_set_ge_of_encodable {ι} [linear_order ι] {τ : Ω → ι} {f : filtration ι m} - [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : + [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[f i] {ω | i ≤ τ ω} := hτ.measurable_set_ge_of_countable (set.to_countable _) i @@ -648,9 +648,9 @@ end topological_space end linear_order -section encodable +section countable -lemma is_stopping_time_of_measurable_set_eq [preorder ι] [encodable ι] +lemma is_stopping_time_of_measurable_set_eq [preorder ι] [countable ι] {f : filtration ι m} {τ : Ω → ι} (hτ : ∀ i, measurable_set[f i] {ω | τ ω = i}) : is_stopping_time f τ := begin @@ -660,7 +660,7 @@ begin exact f.mono hk _ (hτ k), end -end encodable +end countable end measurable_set @@ -720,14 +720,14 @@ begin linarith }, end --- generalize to certain encodable type? +-- generalize to certain countable type? lemma add {f : filtration ℕ m} {τ π : Ω → ℕ} (hτ : is_stopping_time f τ) (hπ : is_stopping_time f π) : is_stopping_time f (τ + π) := begin intro i, rw (_ : {ω | (τ + π) ω ≤ i} = ⋃ k ≤ i, {ω | π ω = k} ∩ {ω | τ ω + k ≤ i}), - { exact measurable_set.Union (λ k, measurable_set.Union_Prop + { exact measurable_set.Union (λ k, measurable_set.Union (λ hk, (hπ.measurable_set_eq_le hk).inter (hτ.add_const_nat i))) }, ext ω, simp only [pi.add_apply, set.mem_set_of_eq, set.mem_Union, set.mem_inter_eq, exists_prop], @@ -780,7 +780,7 @@ begin exact le_trans (hle _) hle' }, end -lemma measurable_space_le_of_encodable [encodable ι] (hτ : is_stopping_time f τ) : +lemma measurable_space_le_of_encodable [countable ι] (hτ : is_stopping_time f τ) : hτ.measurable_space ≤ m := begin intros s hs, @@ -956,7 +956,7 @@ begin exact hτ.measurable_set_eq_of_countable h_countable i, end -protected lemma measurable_set_eq_of_encodable' [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_eq_of_encodable' [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[hτ.measurable_space] {ω | τ ω = i} := hτ.measurable_set_eq_of_countable' (set.to_countable _) i @@ -972,7 +972,7 @@ begin exact (hτ.measurable_set_eq_of_countable' h_countable i).union (hτ.measurable_set_gt' i), end -protected lemma measurable_set_ge_of_encodable' [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_ge_of_encodable' [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[hτ.measurable_space] {ω | i ≤ τ ω} := hτ.measurable_set_ge_of_countable' (set.to_countable _) i @@ -987,7 +987,7 @@ begin exact (hτ.measurable_set_le' i).diff (hτ.measurable_set_eq_of_countable' h_countable i), end -protected lemma measurable_set_lt_of_encodable' [encodable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_lt_of_encodable' [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[hτ.measurable_space] {ω | τ ω < i} := hτ.measurable_set_lt_of_countable' (set.to_countable _) i @@ -1161,7 +1161,7 @@ begin { exact (hπ.min_const j).measurable_of_le (λ _, min_le_right _ _), }, end -lemma measurable_set_eq_stopping_time_of_encodable [encodable ι] +lemma measurable_set_eq_stopping_time_of_countable [countable ι] [topological_space ι] [measurable_space ι] [borel_space ι] [order_topology ι] [measurable_singleton_class ι] [second_countable_topology ι] (hτ : is_stopping_time f τ) (hπ : is_stopping_time f π) : @@ -1182,7 +1182,7 @@ begin rw this, refine measurable_set.inter (measurable_set.inter _ (hτ.measurable_set_le j)) (hπ.measurable_set_le j), - apply measurable_set_eq_fun_of_encodable, + apply measurable_set_eq_fun_of_countable, { exact (hτ.min_const j).measurable_of_le (λ _, min_le_right _ _), }, { exact (hπ.min_const j).measurable_of_le (λ _, min_le_right _ _), }, end diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index 4419ad7d2f4a4..9066a23fb6f53 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -172,9 +172,9 @@ lemma fg_bsupr {ι : Type*} (s : finset ι) (N : ι → submodule R M) (h : ∀ (⨆ i ∈ s, N i).fg := by simpa only [finset.sup_eq_supr] using fg_finset_sup s N h -lemma fg_supr {ι : Type*} [fintype ι] (N : ι → submodule R M) (h : ∀ i, (N i).fg) : +lemma fg_supr {ι : Type*} [finite ι] (N : ι → submodule R M) (h : ∀ i, (N i).fg) : (supr N).fg := -by simpa using fg_bsupr finset.univ N (λ i hi, h i) +by { casesI nonempty_fintype ι, simpa using fg_bsupr finset.univ N (λ i hi, h i) } variables {P : Type*} [add_comm_monoid P] [module R P] variables (f : M →ₗ[R] P) @@ -211,7 +211,7 @@ fg_def.2 ⟨linear_map.inl R M P '' tb ∪ linear_map.inr R M P '' tc, (htb.1.image _).union (htc.1.image _), by rw [linear_map.span_inl_union_inr, htb.2, htc.2]⟩ -theorem fg_pi {ι : Type*} {M : ι → Type*} [fintype ι] [Π i, add_comm_monoid (M i)] +theorem fg_pi {ι : Type*} {M : ι → Type*} [finite ι] [Π i, add_comm_monoid (M i)] [Π i, module R (M i)] {p : Π i, submodule R (M i)} (hsb : ∀ i, (p i).fg) : (submodule.pi set.univ p).fg := begin @@ -511,9 +511,10 @@ from λ x ⟨hx1, hx2⟩, ⟨x.1, prod.ext rfl $ eq.symm $ linear_map.mem_ker.1 submodule.map_comap_eq_self this ▸ (noetherian _).map _⟩ instance is_noetherian_pi {R ι : Type*} {M : ι → Type*} [ring R] - [Π i, add_comm_group (M i)] [Π i, module R (M i)] [fintype ι] + [Π i, add_comm_group (M i)] [Π i, module R (M i)] [finite ι] [∀ i, is_noetherian R (M i)] : is_noetherian R (Π i, M i) := begin + casesI nonempty_fintype ι, haveI := classical.dec_eq ι, suffices on_finset : ∀ s : finset ι, is_noetherian R (Π i : s, M i), { let coe_e := equiv.subtype_univ_equiv finset.mem_univ, @@ -555,7 +556,7 @@ end /-- A version of `is_noetherian_pi` for non-dependent functions. We need this instance because sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to prove that `ι → ℝ` is finite dimensional over `ℝ`). -/ -instance is_noetherian_pi' {R ι M : Type*} [ring R] [add_comm_group M] [module R M] [fintype ι] +instance is_noetherian_pi' {R ι M : Type*} [ring R] [add_comm_group M] [module R M] [finite ι] [is_noetherian R M] : is_noetherian R (ι → M) := is_noetherian_pi diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index 7fb06fda724de..5a78147d2b4c8 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -196,7 +196,7 @@ begin contrapose! hx, obtain ⟨s, ⟨b⟩⟩ := hx, refine is_integral_of_mem_of_fg (K⟮x⟯).to_subalgebra _ x _, - { exact (submodule.fg_iff_finite_dimensional _).mpr (of_finset_basis b) }, + { exact (submodule.fg_iff_finite_dimensional _).mpr (of_fintype_basis b) }, { exact intermediate_field.subset_adjoin K _ (set.mem_singleton x) } end diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index eb9117f5572ce..a30e20a42ef2e 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -232,7 +232,7 @@ begin contrapose! hx, obtain ⟨s, ⟨b⟩⟩ := hx, refine is_integral_of_mem_of_fg (K⟮x⟯).to_subalgebra _ x _, - { exact (submodule.fg_iff_finite_dimensional _).mpr (finite_dimensional.of_finset_basis b) }, + { exact (submodule.fg_iff_finite_dimensional _).mpr (finite_dimensional.of_fintype_basis b) }, { exact subset_adjoin K _ (set.mem_singleton x) } end diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 5b34e3e3454fc..635e1ba677e82 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -574,8 +574,9 @@ end end has_continuous_star -section encodable open encodable + +section encodable variable [encodable γ] /-- You can compute a sum over an encodably type by summing over the natural numbers and @@ -607,25 +608,29 @@ theorem tsum_Union_decode₂ (m : set β → α) (m0 : m ∅ = 0) (s : γ → set β) : ∑' i, m (⋃ b ∈ decode₂ γ i, s b) = ∑' b, m (s b) := tsum_supr_decode₂ m m0 s +end encodable + /-! Some properties about measure-like functions. These could also be functions defined on complete sublattices of sets, with the property that they are countably sub-additive. `R` will probably be instantiated with `(≤)` in all applications. -/ -/-- If a function is countably sub-additive then it is sub-additive on encodable types -/ +section countable +variables [countable γ] + +/-- If a function is countably sub-additive then it is sub-additive on countable types -/ theorem rel_supr_tsum [complete_lattice β] (m : β → α) (m0 : m ⊥ = 0) (R : α → α → Prop) (m_supr : ∀(s : ℕ → β), R (m (⨆ i, s i)) ∑' i, m (s i)) (s : γ → β) : R (m (⨆ b : γ, s b)) ∑' b : γ, m (s b) := -by { rw [← supr_decode₂, ← tsum_supr_decode₂ _ m0 s], exact m_supr _ } +by { casesI nonempty_encodable γ, rw [←supr_decode₂, ←tsum_supr_decode₂ _ m0 s], exact m_supr _ } /-- If a function is countably sub-additive then it is sub-additive on finite sets -/ theorem rel_supr_sum [complete_lattice β] (m : β → α) (m0 : m ⊥ = 0) (R : α → α → Prop) (m_supr : ∀(s : ℕ → β), R (m (⨆ i, s i)) (∑' i, m (s i))) (s : δ → β) (t : finset δ) : R (m (⨆ d ∈ t, s d)) (∑ d in t, m (s d)) := -by { cases t.nonempty_encodable, rw [supr_subtype'], convert rel_supr_tsum m m0 R m_supr _, - rw [← finset.tsum_subtype], assumption } +by { rw [supr_subtype', ←finset.tsum_subtype], exact rel_supr_tsum m m0 R m_supr _ } /-- If a function is countably sub-additive then it is binary sub-additive -/ theorem rel_sup_add [complete_lattice β] (m : β → α) (m0 : m ⊥ = 0) @@ -637,7 +642,7 @@ begin { rw [tsum_fintype, fintype.sum_bool, cond, cond] } end -end encodable +end countable variables [has_continuous_add α] diff --git a/src/topology/bases.lean b/src/topology/bases.lean index 093b3ee0a425b..de9eec02d31e6 100644 --- a/src/topology/bases.lean +++ b/src/topology/bases.lean @@ -291,10 +291,10 @@ def dense_seq [separable_space α] [nonempty α] : ℕ → α := classical.some variable {α} @[priority 100] -instance encodable.to_separable_space [encodable α] : separable_space α := +instance countable.to_separable_space [countable α] : separable_space α := { exists_countable_dense := ⟨set.univ, set.countable_univ, dense_univ⟩ } -lemma separable_space_of_dense_range {ι : Type*} [encodable ι] (u : ι → α) (hu : dense_range u) : +lemma separable_space_of_dense_range {ι : Type*} [countable ι] (u : ι → α) (hu : dense_range u) : separable_space α := ⟨⟨range u, countable_range u, hu⟩⟩ @@ -352,7 +352,7 @@ begin exact ⟨c, c_count, by simpa using closure_mono hs⟩, end -lemma is_separable_Union {ι : Type*} [encodable ι] {s : ι → set α} (hs : ∀ i, is_separable (s i)) : +lemma is_separable_Union {ι : Type*} [countable ι] {s : ι → set α} (hs : ∀ i, is_separable (s i)) : is_separable (⋃ i, s i) := begin choose c hc h'c using hs, @@ -639,7 +639,6 @@ instance {ι : Type*} {π : ι → Type*} [countable ι] [t : ∀a, topological_space (π a)] [∀a, second_countable_topology (π a)] : second_countable_topology (∀a, π a) := begin - haveI := encodable.of_countable ι, have : t = (λa, generate_from (countable_basis (π a))), from funext (assume a, (is_basis_countable_basis (π a)).eq_generate_from), rw [this, pi_generate_from_eq], @@ -753,7 +752,7 @@ begin end /-- A countable disjoint union of second countable spaces is second countable. -/ -instance [encodable ι] [∀ i, second_countable_topology (E i)] : +instance [countable ι] [∀ i, second_countable_topology (E i)] : second_countable_topology (Σ i, E i) := begin let b := (⋃ (i : ι), (λ u, ((sigma.mk i) '' u : set (Σ i, E i))) '' (countable_basis (E i))), diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 6a270b2213e17..3f567456db3aa 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1682,7 +1682,7 @@ lemma _root_.topological_space.is_separable.separable_space {s : set α} (hs : i begin classical, rcases eq_empty_or_nonempty s with rfl|⟨⟨x₀, x₀s⟩⟩, - { haveI : encodable (∅ : set α) := fintype.to_encodable ↥∅, exact encodable.to_separable_space }, + { apply_instance }, rcases hs with ⟨c, hc, h'c⟩, haveI : encodable c := hc.to_encodable, obtain ⟨u, -, u_pos, u_lim⟩ : ∃ (u : ℕ → ℝ), strict_anti u ∧ (∀ (n : ℕ), 0 < u n) ∧ diff --git a/src/topology/metric_space/polish.lean b/src/topology/metric_space/polish.lean index d30db5f669224..4898fcf3f297c 100644 --- a/src/topology/metric_space/polish.lean +++ b/src/topology/metric_space/polish.lean @@ -104,10 +104,11 @@ instance t2_space (α : Type*) [topological_space α] [polish_space α] : t2_spa by { letI := upgrade_polish_space α, apply_instance } /-- A countable product of Polish spaces is Polish. -/ -instance pi_countable {ι : Type*} [encodable ι] {E : ι → Type*} +instance pi_countable {ι : Type*} [countable ι] {E : ι → Type*} [∀ i, topological_space (E i)] [∀ i, polish_space (E i)] : polish_space (Π i, E i) := begin + casesI nonempty_encodable ι, letI := λ i, upgrade_polish_space (E i), letI : metric_space (Π i, E i) := pi_countable.metric_space, apply_instance, @@ -119,7 +120,7 @@ instance nat_fun [topological_space α] [polish_space α] : by apply_instance /-- A countable disjoint union of Polish spaces is Polish. -/ -instance sigma {ι : Type*} [encodable ι] +instance sigma {ι : Type*} [countable ι] {E : ι → Type*} [∀ n, topological_space (E n)] [∀ n, polish_space (E n)] : polish_space (Σ n, E n) := begin @@ -185,7 +186,7 @@ def aux_copy (α : Type*) {ι : Type*} (i : ι) : Type* := α /-- Given a Polish space, and countably many finer Polish topologies, there exists another Polish topology which is finer than all of them. -/ -lemma exists_polish_space_forall_le {ι : Type*} [encodable ι] +lemma exists_polish_space_forall_le {ι : Type*} [countable ι] [t : topological_space α] [p : polish_space α] (m : ι → topological_space α) (hm : ∀ n, m n ≤ t) (h'm : ∀ n, @polish_space α (m n)) : ∃ (t' : topological_space α), (∀ n, t' ≤ m n) ∧ (t' ≤ t) ∧ @polish_space α t' := From 1da31fac88633c1b1062ec54cb2e07307a611f9e Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Tue, 30 Aug 2022 16:44:21 +0000 Subject: [PATCH 089/828] feat(data/polynomial/derivative): add more lemmas (#16139) Co-authored-by: Jujian Zhang --- src/data/polynomial/derivative.lean | 69 +++++++++++++++++++++++++++++ 1 file changed, 69 insertions(+) diff --git a/src/data/polynomial/derivative.lean b/src/data/polynomial/derivative.lean index d5f6156eba941..5a5131831fd1f 100644 --- a/src/data/polynomial/derivative.lean +++ b/src/data/polynomial/derivative.lean @@ -257,6 +257,10 @@ begin { simp only [ih, function.iterate_succ, polynomial.derivative_map, function.comp_app], }, end +lemma derivative_nat_cast_mul {n : ℕ} {f : R[X]} : + (↑n * f).derivative = n * f.derivative := +by simp + @[simp] lemma iterate_derivative_nat_cast_mul {n k : ℕ} {f : R[X]} : derivative^[k] (n * f) = n * (derivative^[k] f) := by induction k with k ih generalizing f; simp* @@ -386,6 +390,48 @@ theorem derivative_pow (p : R[X]) (n : ℕ) : nat.cases_on n (by rw [pow_zero, derivative_one, nat.cast_zero, zero_mul, zero_mul]) $ λ n, by rw [p.derivative_pow_succ n, n.succ_sub_one, n.cast_succ] +theorem dvd_iterate_derivative_pow (f : R[X]) (n : ℕ) {m : ℕ} (c : R) (hm : m ≠ 0) : + (n : R) ∣ eval c (derivative^[m] (f ^ n)) := +begin + obtain ⟨m, rfl⟩ := nat.exists_eq_succ_of_ne_zero hm, + rw [function.iterate_succ_apply, derivative_pow, mul_assoc, iterate_derivative_nat_cast_mul, + eval_mul, eval_nat_cast], + exact dvd_mul_right _ _, +end + +lemma iterate_derivative_X_pow_eq_nat_cast_mul (n k : ℕ) : + (derivative^[k] (X ^ n : R[X])) = ↑(nat.desc_factorial n k) * X ^ (n - k) := +begin + induction k with k ih, + { rw [function.iterate_zero_apply, tsub_zero, nat.desc_factorial_zero, nat.cast_one, one_mul] }, + { rw [function.iterate_succ_apply', ih, derivative_nat_cast_mul, derivative_X_pow, + nat.succ_eq_add_one, nat.desc_factorial_succ, nat.sub_sub, nat.cast_mul, ←mul_assoc, + mul_comm ↑(nat.desc_factorial _ _)] }, +end + +lemma iterate_derivative_X_pow_eq_C_mul (n k : ℕ) : + (derivative^[k] (X^n : R[X])) = C ↑(nat.desc_factorial n k) * X ^ (n - k) := +by rw [iterate_derivative_X_pow_eq_nat_cast_mul n k, C_eq_nat_cast] + +lemma iterate_derivative_X_pow_eq_smul (n : ℕ) (k : ℕ) : + (derivative^[k] (X^n : R[X])) = (nat.desc_factorial n k : R) • X ^ (n - k) := +by rw [iterate_derivative_X_pow_eq_C_mul n k, smul_eq_C_mul] + +lemma derivative_X_add_pow (c : R) (m : ℕ) : ((X + C c) ^ m).derivative = m * (X + C c) ^ (m - 1) := +by rw [derivative_pow, derivative_add, derivative_X, derivative_C, add_zero, mul_one] + +lemma iterate_derivative_X_add_pow (n k : ℕ) (c : R) : + (derivative^[k] ((X + C c) ^ n)) = + ↑(∏ i in finset.range k, (n - i)) * (X + C c) ^ (n - k) := +begin + induction k with k IH, + { rw [function.iterate_zero_apply, finset.range_zero, finset.prod_empty, nat.cast_one, one_mul, + tsub_zero] }, + { simp only [function.iterate_succ_apply', IH, derivative_mul, zero_mul, derivative_nat_cast, + zero_add, finset.prod_range_succ, C_eq_nat_cast, nat.sub_sub, ←mul_assoc, + derivative_X_add_pow, nat.succ_eq_add_one, nat.cast_mul] }, +end + lemma derivative_comp (p q : R[X]) : (p.comp q).derivative = q.derivative * p.derivative.comp q := begin @@ -448,6 +494,20 @@ linear_map.map_sub derivative f g derivative^[k] (f - g) = (derivative^[k] f) - (derivative^[k] g) := by induction k with k ih generalizing f g; simp* +@[simp] lemma derivative_int_cast {n : ℤ} : derivative (n : R[X]) = 0 := +begin + rw ← C_eq_int_cast n, + exact derivative_C, +end + +lemma derivative_int_cast_mul {n : ℤ} {f : R[X]} : + (↑n * f).derivative = n * f.derivative := +by simp + +@[simp] lemma iterate_derivative_int_cast_mul {n : ℤ} {k : ℕ} {f : R[X]} : + derivative^[k] (↑n * f) = n * (derivative^[k] f) := +by induction k with k ih generalizing f; simp* + end ring section comm_ring @@ -474,6 +534,15 @@ begin simpa using (eval_ring_hom r).map_multiset_prod (multiset.map (λ a, X - C a) (S.erase r)), end +lemma derivative_X_sub_pow (c : R) (m : ℕ) : + ((X - C c) ^ m).derivative = m * (X - C c) ^ (m - 1) := +by rw [derivative_pow, derivative_sub, derivative_X, derivative_C, sub_zero, mul_one] + +lemma iterate_derivative_X_sub_pow (n k : ℕ) (c : R) : + (derivative^[k] ((X - C c) ^ n)) = + (↑(∏ i in finset.range k, (n - i))) * (X - C c) ^ (n - k) := +by simp_rw [sub_eq_add_neg, ←C_neg, iterate_derivative_X_add_pow] + end comm_ring end derivative From 002aa5175af68467ffa39ad6ca79c2fb76bf88ba Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Tue, 30 Aug 2022 16:44:22 +0000 Subject: [PATCH 090/828] feat(probability/hitting_time): add some hitting time lemmas (#16296) --- src/probability/hitting_time.lean | 56 +++++++++++++++++++++++++++++++ src/probability/stopping.lean | 16 +++++++++ 2 files changed, 72 insertions(+) diff --git a/src/probability/hitting_time.lean b/src/probability/hitting_time.lean index 7fe6710886ab6..62171ead4962e 100644 --- a/src/probability/hitting_time.lean +++ b/src/probability/hitting_time.lean @@ -51,6 +51,7 @@ section inequalities variables [conditionally_complete_linear_order ι] {u : ι → Ω → β} {s : set β} {n i : ι} {ω : Ω} +/-- This lemma is strictly weaker than `hitting_of_le`. -/ lemma hitting_of_lt {m : ι} (h : m < n) : hitting u s n m ω = m := begin simp_rw [hitting], @@ -73,6 +74,36 @@ begin { rw hitting_of_lt h_lt, }, end +lemma not_mem_of_lt_hitting {m k : ι} + (hk₁ : k < hitting u s n m ω) (hk₂ : n ≤ k) : + u k ω ∉ s := +begin + classical, + intro h, + have hexists : ∃ j ∈ set.Icc n m, u j ω ∈ s, + refine ⟨k, ⟨hk₂, le_trans hk₁.le $ hitting_le _⟩, h⟩, + refine not_le.2 hk₁ _, + simp_rw [hitting, if_pos hexists], + exact cInf_le bdd_below_Icc.inter_of_left ⟨⟨hk₂, le_trans hk₁.le $ hitting_le _⟩, h⟩, +end + +lemma hitting_eq_end_iff {m : ι} : + hitting u s n m ω = m ↔ (∃ j ∈ set.Icc n m, u j ω ∈ s) → + Inf (set.Icc n m ∩ {i : ι | u i ω ∈ s}) = m := +by rw [hitting, ite_eq_right_iff] + +lemma hitting_of_le {m : ι} (hmn : m ≤ n) : + hitting u s n m ω = m := +begin + obtain (rfl | h) := le_iff_eq_or_lt.1 hmn, + { simp only [hitting, set.Icc_self, ite_eq_right_iff, set.mem_Icc, exists_prop, + forall_exists_index, and_imp], + intros i hi₁ hi₂ hi, + rw [set.inter_eq_left_iff_subset.2, cInf_singleton], + exact set.singleton_subset_iff.2 (le_antisymm hi₂ hi₁ ▸ hi) }, + { exact hitting_of_lt h } +end + lemma le_hitting {m : ι} (hnm : n ≤ m) (ω : Ω) : n ≤ hitting u s n m ω := begin simp only [hitting], @@ -109,6 +140,16 @@ begin exact h_mem.2, end +lemma hitting_mem_set_of_hitting_lt [is_well_order ι (<)] {m : ι} + (hl : hitting u s n m ω < m) : + u (hitting u s n m ω) ω ∈ s := +begin + by_cases h : ∃ j ∈ set.Icc n m, u j ω ∈ s, + { exact hitting_mem_set h }, + { simp_rw [hitting, if_neg h] at hl, + exact false.elim (hl.ne rfl) } +end + lemma hitting_le_of_mem {m : ι} (hin : n ≤ i) (him : i ≤ m) (his : u i ω ∈ s) : hitting u s n m ω ≤ i := begin @@ -175,6 +216,21 @@ begin exact ⟨j, ⟨hj₁.1, hj₁.2.trans h⟩, hj₂⟩, end +lemma hitting_mono {m₁ m₂ : ι} (hm : m₁ ≤ m₂) : + hitting u s n m₁ ω ≤ hitting u s n m₂ ω := +begin + by_cases h : ∃ j ∈ set.Icc n m₁, u j ω ∈ s, + { exact (hitting_eq_hitting_of_exists hm h).le }, + { simp_rw [hitting, if_neg h], + split_ifs with h', + { obtain ⟨j, hj₁, hj₂⟩ := h', + refine le_cInf ⟨j, hj₁, hj₂⟩ _, + by_contra hneg, push_neg at hneg, + obtain ⟨i, hi₁, hi₂⟩ := hneg, + exact h ⟨i, ⟨hi₁.1.1, hi₂.le⟩, hi₁.2⟩ }, + { exact hm } } +end + end inequalities /-- A discrete hitting time is a stopping time. -/ diff --git a/src/probability/stopping.lean b/src/probability/stopping.lean index 1bf78c0fe21ce..8823a96a56d0b 100644 --- a/src/probability/stopping.lean +++ b/src/probability/stopping.lean @@ -1420,6 +1420,22 @@ begin exact hneq.symm } }, end +lemma stopped_process_eq' (n : ℕ) : + stopped_process u τ n = + set.indicator {a | n + 1 ≤ τ a} (u n) + + ∑ i in finset.range (n + 1), set.indicator {a | τ a = i} (u i) := +begin + have : {a | n ≤ τ a}.indicator (u n) = + {a | n + 1 ≤ τ a}.indicator (u n) + {a | τ a = n}.indicator (u n), + { ext x, + rw [add_comm, pi.add_apply, ← set.indicator_union_of_not_mem_inter], + { simp_rw [@eq_comm _ _ n, @le_iff_eq_or_lt _ _ n, nat.succ_le_iff], + refl }, + { rintro ⟨h₁, h₂⟩, + exact (nat.succ_le_iff.1 h₂).ne h₁.symm } }, + rw [stopped_process_eq, this, finset.sum_range_succ_comm, ← add_assoc], +end + end add_comm_monoid section normed_add_comm_group From fe5e4ce6c72d96d77ad40ac832a6e7f8040990bc Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 31 Aug 2022 01:06:43 +0000 Subject: [PATCH 091/828] feat(category_theory/limits): (co)limits in full subcategories (#16188) --- src/category_theory/limits/creates.lean | 93 +++++++++---- .../limits/full_subcategory.lean | 128 ++++++++++++++++++ 2 files changed, 196 insertions(+), 25 deletions(-) create mode 100644 src/category_theory/limits/full_subcategory.lean diff --git a/src/category_theory/limits/creates.lean b/src/category_theory/limits/creates.lean index 5afd8b5e0267a..d76fccd769758 100644 --- a/src/category_theory/limits/creates.lean +++ b/src/category_theory/limits/creates.lean @@ -244,6 +244,21 @@ def creates_limit_of_reflects_iso {K : J ⥤ C} {F : C ⥤ D} [reflects_isomorph exact is_limit.of_iso_limit hd' (as_iso f).symm, end } } +/-- +When `F` is fully faithful, to show that `F` creates the limit for `K` it suffices to exhibit a lift +of a limit cone for `K ⋙ F`. +-/ +-- Notice however that even if the isomorphism is `iso.refl _`, +-- this construction will insert additional identity morphisms in the cone maps, +-- so the constructed limits may not be ideal, definitionally. +def creates_limit_of_fully_faithful_of_lift' {K : J ⥤ C} {F : C ⥤ D} [full F] [faithful F] + {l : cone (K ⋙ F)} (hl : is_limit l) (c : cone K) (i : F.map_cone c ≅ l) : creates_limit K F := +creates_limit_of_reflects_iso (λ c' t, +{ lifted_cone := c, + valid_lift := i ≪≫ is_limit.unique_up_to_iso hl t, + makes_limit := is_limit.of_faithful F (is_limit.of_iso_limit hl i.symm) _ + (λ s, F.image_preimage _) }) + /-- When `F` is fully faithful, and `has_limit (K ⋙ F)`, to show that `F` creates the limit for `K` it suffices to exhibit a lift of the chosen limit cone for `K ⋙ F`. @@ -254,11 +269,23 @@ it suffices to exhibit a lift of the chosen limit cone for `K ⋙ F`. def creates_limit_of_fully_faithful_of_lift {K : J ⥤ C} {F : C ⥤ D} [full F] [faithful F] [has_limit (K ⋙ F)] (c : cone K) (i : F.map_cone c ≅ limit.cone (K ⋙ F)) : creates_limit K F := -creates_limit_of_reflects_iso (λ c' t, -{ lifted_cone := c, - valid_lift := i.trans (is_limit.unique_up_to_iso (limit.is_limit _) t), - makes_limit := is_limit.of_faithful F (is_limit.of_iso_limit (limit.is_limit _) i.symm) - (λ s, F.preimage _) (λ s, F.image_preimage _) }) +creates_limit_of_fully_faithful_of_lift' (limit.is_limit _) c i + +/-- +When `F` is fully faithful, to show that `F` creates the limit for `K` it suffices to show that a +limit point is in the essential image of `F`. +-/ +-- Notice however that even if the isomorphism is `iso.refl _`, +-- this construction will insert additional identity morphisms in the cone maps, +-- so the constructed limits may not be ideal, definitionally. +def creates_limit_of_fully_faithful_of_iso' {K : J ⥤ C} {F : C ⥤ D} [full F] [faithful F] + {l : cone (K ⋙ F)} (hl : is_limit l) (X : C) (i : F.obj X ≅ l.X) : creates_limit K F := +creates_limit_of_fully_faithful_of_lift' hl +({ X := X, + π := + { app := λ j, F.preimage (i.hom ≫ l.π.app j), + naturality' := λ Y Z f, F.map_injective $ by { dsimp, simpa using (l.w f).symm } } }) +(cones.ext i (λ j, by simp only [functor.image_preimage, functor.map_cone_π_app])) /-- When `F` is fully faithful, and `has_limit (K ⋙ F)`, to show that `F` creates the limit for `K` @@ -270,12 +297,7 @@ it suffices to show that the chosen limit point is in the essential image of `F` def creates_limit_of_fully_faithful_of_iso {K : J ⥤ C} {F : C ⥤ D} [full F] [faithful F] [has_limit (K ⋙ F)] (X : C) (i : F.obj X ≅ limit (K ⋙ F)) : creates_limit K F := -creates_limit_of_fully_faithful_of_lift -({ X := X, - π := - { app := λ j, F.preimage (i.hom ≫ limit.π (K ⋙ F) j), - naturality' := λ Y Z f, F.map_injective (by { dsimp, simp, erw limit.w (K ⋙ F), }) }} : cone K) -(by { fapply cones.ext, exact i, tidy, }) +creates_limit_of_fully_faithful_of_iso' (limit.is_limit _) X i /-- `F` preserves the limit of `K` if it creates the limit and `K ⋙ F` has the limit. -/ @[priority 100] -- see Note [lower instance priority] @@ -324,6 +346,22 @@ def creates_colimit_of_reflects_iso {K : J ⥤ C} {F : C ⥤ D} [reflects_isomor exact is_colimit.of_iso_colimit hd' (as_iso f), end } } +/-- +When `F` is fully faithful, to show that `F` creates the colimit for `K` it suffices to exhibit a +lift of a colimit cocone for `K ⋙ F`. +-/ +-- Notice however that even if the isomorphism is `iso.refl _`, +-- this construction will insert additional identity morphisms in the cocone maps, +-- so the constructed colimits may not be ideal, definitionally. +def creates_colimit_of_fully_faithful_of_lift' {K : J ⥤ C} {F : C ⥤ D} [full F] [faithful F] + {l : cocone (K ⋙ F)} (hl : is_colimit l) (c : cocone K) (i : F.map_cocone c ≅ l) : + creates_colimit K F := +creates_colimit_of_reflects_iso (λ c' t, +{ lifted_cocone := c, + valid_lift := i ≪≫ is_colimit.unique_up_to_iso hl t, + makes_colimit := is_colimit.of_faithful F (is_colimit.of_iso_colimit hl i.symm) _ + (λ s, F.image_preimage _) }) + /-- When `F` is fully faithful, and `has_colimit (K ⋙ F)`, to show that `F` creates the colimit for `K` it suffices to exhibit a lift of the chosen colimit cocone for `K ⋙ F`. @@ -334,12 +372,24 @@ it suffices to exhibit a lift of the chosen colimit cocone for `K ⋙ F`. def creates_colimit_of_fully_faithful_of_lift {K : J ⥤ C} {F : C ⥤ D} [full F] [faithful F] [has_colimit (K ⋙ F)] (c : cocone K) (i : F.map_cocone c ≅ colimit.cocone (K ⋙ F)) : creates_colimit K F := -creates_colimit_of_reflects_iso (λ c' t, -{ lifted_cocone := c, - valid_lift := i.trans (is_colimit.unique_up_to_iso (colimit.is_colimit _) t), - makes_colimit := is_colimit.of_faithful F - (is_colimit.of_iso_colimit (colimit.is_colimit _) i.symm) - (λ s, F.preimage _) (λ s, F.image_preimage _) }) +creates_colimit_of_fully_faithful_of_lift' (colimit.is_colimit _) c i + +/-- +When `F` is fully faithful, to show that `F` creates the colimit for `K` it suffices to show that +a colimit point is in the essential image of `F`. +-/ +-- Notice however that even if the isomorphism is `iso.refl _`, +-- this construction will insert additional identity morphisms in the cocone maps, +-- so the constructed colimits may not be ideal, definitionally. +def creates_colimit_of_fully_faithful_of_iso' {K : J ⥤ C} {F : C ⥤ D} [full F] [faithful F] + {l : cocone (K ⋙ F)} (hl : is_colimit l) (X : C) (i : F.obj X ≅ l.X) : creates_colimit K F := +creates_colimit_of_fully_faithful_of_lift' hl +({ X := X, + ι := + { app := λ j, F.preimage (l.ι.app j ≫ i.inv), + naturality' := λ Y Z f, F.map_injective $ + by { dsimp, simpa [← cancel_mono i.hom] using (l.w f) } } }) +(cocones.ext i (λ j, by simp)) /-- When `F` is fully faithful, and `has_colimit (K ⋙ F)`, to show that `F` creates the colimit for `K` @@ -351,14 +401,7 @@ it suffices to show that the chosen colimit point is in the essential image of ` def creates_colimit_of_fully_faithful_of_iso {K : J ⥤ C} {F : C ⥤ D} [full F] [faithful F] [has_colimit (K ⋙ F)] (X : C) (i : F.obj X ≅ colimit (K ⋙ F)) : creates_colimit K F := -creates_colimit_of_fully_faithful_of_lift -({ X := X, - ι := - { app := λ j, F.preimage (colimit.ι (K ⋙ F) j ≫ i.inv : _), - naturality' := λ Y Z f, F.map_injective - (by { erw category.comp_id, simp only [functor.map_comp, functor.image_preimage], - erw colimit.w_assoc (K ⋙ F) }) }} : cocone K) -(by { fapply cocones.ext, exact i, tidy, }) +creates_colimit_of_fully_faithful_of_iso' (colimit.is_colimit _) X i /-- `F` preserves the colimit of `K` if it creates the colimit and `K ⋙ F` has the colimit. -/ diff --git a/src/category_theory/limits/full_subcategory.lean b/src/category_theory/limits/full_subcategory.lean new file mode 100644 index 0000000000000..7691455df253a --- /dev/null +++ b/src/category_theory/limits/full_subcategory.lean @@ -0,0 +1,128 @@ +/- +Copyright (c) 2022 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import category_theory.limits.creates + +/-! +# Limits in full subcategories + +We introduce the notion of a property closed under taking limits and show that if `P` is closed +under taking limits, then limits in `full_subcategory P` can be constructed from limits in `C`. +More precisely, the inclusion creates such limits. + +-/ + +noncomputable theory + +universes w' w v u + +open category_theory + +namespace category_theory.limits + +/-- We say that a property is closed under limits of shape `J` if whenever all objects in a + `J`-shaped diagram have the property, any limit of this diagram also has the property. -/ +def closed_under_limits_of_shape {C : Type u} [category.{v} C] (J : Type w) [category.{w'} J] + (P : C → Prop) : Prop := +∀ ⦃F : J ⥤ C⦄ ⦃c : cone F⦄ (hc : is_limit c), (∀ j, P (F.obj j)) → P c.X + +/-- We say that a property is closed under colimits of shape `J` if whenever all objects in a + `J`-shaped diagram have the property, any colimit of this diagram also has the property. -/ +def closed_under_colimits_of_shape {C : Type u} [category.{v} C] (J : Type w) [category.{w'} J] + (P : C → Prop) : Prop := +∀ ⦃F : J ⥤ C⦄ ⦃c : cocone F⦄ (hc : is_colimit c), (∀ j, P (F.obj j)) → P c.X + +section +variables {C : Type u} [category.{v} C] {J : Type w} [category.{w'} J] {P : C → Prop} + +lemma closed_under_limits_of_shape.limit (h : closed_under_limits_of_shape J P) {F : J ⥤ C} + [has_limit F] : (∀ j, P (F.obj j)) → P (limit F) := +h (limit.is_limit _) + +lemma closed_under_colimits_of_shape.colimit (h : closed_under_colimits_of_shape J P) {F : J ⥤ C} + [has_colimit F] : (∀ j, P (F.obj j)) → P (colimit F) := +h (colimit.is_colimit _) + +end + +section +variables {J : Type w} [category.{w'} J] {C : Type u} [category.{v} C] {P : C → Prop} + +/-- If a `J`-shaped diagram in `full_subcategory P` has a limit cone in `C` whose cone point lives + in the full subcategory, then this defines a limit in the full subcategory. -/ +def creates_limit_full_subcategory_inclusion' (F : J ⥤ full_subcategory P) + {c : cone (F ⋙ full_subcategory_inclusion P)} (hc : is_limit c) (h : P c.X) : + creates_limit F (full_subcategory_inclusion P) := +creates_limit_of_fully_faithful_of_iso' hc ⟨_, h⟩ (iso.refl _) + +/-- If a `J`-shaped diagram in `full_subcategory P` has a limit in `C` whose cone point lives in the + full subcategory, then this defines a limit in the full subcategory. -/ +def creates_limit_full_subcategory_inclusion (F : J ⥤ full_subcategory P) + [has_limit (F ⋙ full_subcategory_inclusion P)] + (h : P (limit (F ⋙ full_subcategory_inclusion P))) : + creates_limit F (full_subcategory_inclusion P) := +creates_limit_full_subcategory_inclusion' F (limit.is_limit _) h + +/-- If a `J`-shaped diagram in `full_subcategory P` has a colimit cocone in `C` whose cocone point + lives in the full subcategory, then this defines a colimit in the full subcategory. -/ +def creates_colimit_full_subcategory_inclusion' (F : J ⥤ full_subcategory P) + {c : cocone (F ⋙ full_subcategory_inclusion P)} (hc : is_colimit c) (h : P c.X) : + creates_colimit F (full_subcategory_inclusion P) := +creates_colimit_of_fully_faithful_of_iso' hc ⟨_, h⟩ (iso.refl _) + +/-- If a `J`-shaped diagram in `full_subcategory P` has a colimit in `C` whose cocone point lives in + the full subcategory, then this defines a colimit in the full subcategory. -/ +def creates_colimit_full_subcategory_inclusion (F : J ⥤ full_subcategory P) + [has_colimit (F ⋙ full_subcategory_inclusion P)] + (h : P (colimit (F ⋙ full_subcategory_inclusion P))) : + creates_colimit F (full_subcategory_inclusion P) := +creates_colimit_full_subcategory_inclusion' F (colimit.is_colimit _) h + +/-- If `P` is closed under limits of shape `J`, then the inclusion creates such limits. -/ +def creates_limit_full_subcategory_inclusion_of_closed (h : closed_under_limits_of_shape J P) + (F : J ⥤ full_subcategory P) [has_limit (F ⋙ full_subcategory_inclusion P)] : + creates_limit F (full_subcategory_inclusion P) := +creates_limit_full_subcategory_inclusion F (h.limit (λ j, (F.obj j).property)) + +/-- If `P` is closed under limits of shape `J`, then the inclusion creates such limits. -/ +def creates_limits_of_shape_full_subcategory_inclusion (h : closed_under_limits_of_shape J P) + [has_limits_of_shape J C] : creates_limits_of_shape J (full_subcategory_inclusion P) := +{ creates_limit := λ F, creates_limit_full_subcategory_inclusion_of_closed h F } + +lemma has_limit_of_closed_under_limits (h : closed_under_limits_of_shape J P) + (F : J ⥤ full_subcategory P) [has_limit (F ⋙ full_subcategory_inclusion P)] : has_limit F := +have creates_limit F (full_subcategory_inclusion P), + from creates_limit_full_subcategory_inclusion_of_closed h F, +by exactI has_limit_of_created F (full_subcategory_inclusion P) + +lemma has_limits_of_shape_of_closed_under_limits (h : closed_under_limits_of_shape J P) + [has_limits_of_shape J C] : has_limits_of_shape J (full_subcategory P) := +{ has_limit := λ F, has_limit_of_closed_under_limits h F } + +/-- If `P` is closed under colimits of shape `J`, then the inclusion creates such colimits. -/ +def creates_colimit_full_subcategory_inclusion_of_closed (h : closed_under_colimits_of_shape J P) + (F : J ⥤ full_subcategory P) [has_colimit (F ⋙ full_subcategory_inclusion P)] : + creates_colimit F (full_subcategory_inclusion P) := +creates_colimit_full_subcategory_inclusion F (h.colimit (λ j, (F.obj j).property)) + +/-- If `P` is closed under colimits of shape `J`, then the inclusion creates such colimits. -/ +def creates_colimits_of_shape_full_subcategory_inclusion + (h : closed_under_colimits_of_shape J P) [has_colimits_of_shape J C] : + creates_colimits_of_shape J (full_subcategory_inclusion P) := +{ creates_colimit := λ F, creates_colimit_full_subcategory_inclusion_of_closed h F } + +lemma has_colimit_of_closed_under_colimits (h : closed_under_colimits_of_shape J P) + (F : J ⥤ full_subcategory P) [has_colimit (F ⋙ full_subcategory_inclusion P)] : has_colimit F := +have creates_colimit F (full_subcategory_inclusion P), + from creates_colimit_full_subcategory_inclusion_of_closed h F, +by exactI has_colimit_of_created F (full_subcategory_inclusion P) + +lemma has_colimits_of_shape_of_closed_under_colimits (h : closed_under_colimits_of_shape J P) + [has_colimits_of_shape J C] : has_colimits_of_shape J (full_subcategory P) := +{ has_colimit := λ F, has_colimit_of_closed_under_colimits h F } + +end + +end category_theory.limits From 13618c1c5af6573d6974330af2a9010d6275e680 Mon Sep 17 00:00:00 2001 From: Mauricio Collares Date: Wed, 31 Aug 2022 02:14:46 +0000 Subject: [PATCH 092/828] chore(.github/workflows): increase olean generation timeout (#15784) We are currently experiencing timeouts when saving some oleans (leanprover-community/lean#749). Since Lean does not error out when deterministic timeouts happen during olean generation (leanprover-community/lean#750) and we already run Lean twice when compiling mathlib, Ben Toner suggested the following clever hack: Increase the timeout, but only on the second lean run. This effectively means anything that was reported as a timeout before is still reported as a timeout (by the first call), while olean generation gets more heartbeats. Co-authored-by: Ben Toner --- .github/workflows/bors.yml | 2 +- .github/workflows/build.yml | 2 +- .github/workflows/build.yml.in | 2 +- .github/workflows/build_fork.yml | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index e33ab41da731a..d1605c462fc23 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -82,7 +82,7 @@ jobs: leanpkg configure echo "::set-output name=started::true" lean --json -T150000 --make src | python3 scripts/detect_errors.py - lean --json -T150000 --make src | python3 scripts/detect_errors.py + lean --json -T400000 --make src | python3 scripts/detect_errors.py - name: push release to azure if: always() && github.repository == 'leanprover-community/mathlib' && steps.build.outputs.started == 'true' diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 64fae8f218664..a48a77a703d74 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -90,7 +90,7 @@ jobs: leanpkg configure echo "::set-output name=started::true" lean --json -T150000 --make src | python3 scripts/detect_errors.py - lean --json -T150000 --make src | python3 scripts/detect_errors.py + lean --json -T400000 --make src | python3 scripts/detect_errors.py - name: push release to azure if: always() && github.repository == 'leanprover-community/mathlib' && steps.build.outputs.started == 'true' diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index da0de295d9afa..395e1e1a2a553 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -68,7 +68,7 @@ jobs: leanpkg configure echo "::set-output name=started::true" lean --json -T150000 --make src | python3 scripts/detect_errors.py - lean --json -T150000 --make src | python3 scripts/detect_errors.py + lean --json -T400000 --make src | python3 scripts/detect_errors.py - name: push release to azure if: always() && github.repository == 'leanprover-community/mathlib' && steps.build.outputs.started == 'true' diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 7d7d00c9831d9..509d135f52d01 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -88,7 +88,7 @@ jobs: leanpkg configure echo "::set-output name=started::true" lean --json -T150000 --make src | python3 scripts/detect_errors.py - lean --json -T150000 --make src | python3 scripts/detect_errors.py + lean --json -T400000 --make src | python3 scripts/detect_errors.py - name: push release to azure if: always() && github.repository == 'leanprover-community/mathlib' && steps.build.outputs.started == 'true' From 5083718c6aab04d68ad8ce61cc58d90724e790a6 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Wed, 31 Aug 2022 04:06:01 +0000 Subject: [PATCH 093/828] chore(scripts): update nolints.txt (#16323) I am happy to remove some nolints for you! --- scripts/nolints.txt | 111 -------------------------------------------- 1 file changed, 111 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index d94ea9858ab94..d62ba5820a7c8 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -9,43 +9,6 @@ apply_nolint monoid_hom.ext_iff to_additive_doc apply_nolint mul_hom.ext_iff to_additive_doc apply_nolint one_hom.comp_assoc to_additive_doc --- analysis/box_integral/box/basic.lean -apply_nolint box_integral.box.Union_Ioo_of_tendsto fintype_finite - --- analysis/box_integral/partition/additive.lean -apply_nolint box_integral.box_additive_map.sum_boxes_congr fintype_finite - --- analysis/box_integral/partition/measure.lean -apply_nolint box_integral.box.measurable_set_Icc fintype_finite -apply_nolint box_integral.box.measurable_set_Ioo fintype_finite -apply_nolint box_integral.box.measurable_set_coe fintype_finite -apply_nolint box_integral.prepartition.measure_Union_to_real fintype_finite - --- analysis/box_integral/partition/split.lean -apply_nolint box_integral.prepartition.eventually_not_disjoint_imp_le_of_mem_split_many fintype_finite -apply_nolint box_integral.prepartition.eventually_split_many_inf_eq_filter fintype_finite -apply_nolint box_integral.prepartition.exists_Union_eq_diff fintype_finite -apply_nolint box_integral.prepartition.exists_split_many_inf_eq_filter_of_finite fintype_finite -apply_nolint box_integral.prepartition.is_partition.exists_split_many_le fintype_finite - --- analysis/calculus/lagrange_multipliers.lean -apply_nolint is_local_extr_on.linear_dependent_of_has_strict_fderiv_at fintype_finite - --- analysis/calculus/tangent_cone.lean -apply_nolint unique_diff_on.pi fintype_finite -apply_nolint unique_diff_on.univ_pi fintype_finite -apply_nolint unique_diff_within_at.pi fintype_finite -apply_nolint unique_diff_within_at.univ_pi fintype_finite - --- analysis/normed_space/add_torsor_bases.lean -apply_nolint interior_convex_hull_aff_basis fintype_finite - --- analysis/normed_space/finite_dimension.lean -apply_nolint basis.exists_op_nnnorm_le fintype_finite -apply_nolint basis.exists_op_norm_le fintype_finite -apply_nolint is_open_set_of_linear_independent fintype_finite -apply_nolint linear_independent.eventually fintype_finite - -- category_theory/limits/filtered_colimit_commutes_finite_limit.lean apply_nolint category_theory.limits.colimit_limit_to_limit_colimit_injective fintype_finite apply_nolint category_theory.limits.colimit_limit_to_limit_colimit_is_iso fails_quickly @@ -492,13 +455,6 @@ apply_nolint basis.eval_range fintype_finite apply_nolint basis.to_dual_range fintype_finite apply_nolint basis.total_coord fintype_finite --- linear_algebra/finite_dimensional.lean -apply_nolint finite_dimensional.finite_dimensional_pi fintype_finite -apply_nolint finite_dimensional.finite_dimensional_pi' fintype_finite -apply_nolint finite_dimensional.of_fintype_basis fintype_finite -apply_nolint finite_dimensional_finsupp fintype_finite -apply_nolint submodule.finite_dimensional_supr fintype_finite - -- linear_algebra/finsupp_vector_space.lean apply_nolint cardinal_lt_aleph_0_of_finite_dimensional fintype_finite @@ -537,13 +493,6 @@ apply_nolint module.free.multilinear_map fintype_finite -- linear_algebra/orientation.lean apply_nolint basis.map_orientation_eq_det_inv_smul fintype_finite --- linear_algebra/pi.lean -apply_nolint linear_map.pi_ext fintype_finite -apply_nolint linear_map.pi_ext' fintype_finite -apply_nolint linear_map.pi_ext'_iff fintype_finite -apply_nolint linear_map.pi_ext_iff fintype_finite -apply_nolint submodule.supr_map_single fintype_finite - -- linear_algebra/std_basis.lean apply_nolint linear_map.supr_range_std_basis fintype_finite @@ -559,60 +508,6 @@ apply_nolint relator.lift_fun doc_blame apply_nolint relator.right_total doc_blame apply_nolint relator.right_unique doc_blame --- measure_theory/constructions/borel_space.lean -apply_nolint pi.borel_space_fintype fintype_finite -apply_nolint pi.opens_measurable_space_fintype fintype_finite - --- measure_theory/constructions/pi.lean -apply_nolint generate_from_eq_pi fintype_finite -apply_nolint generate_from_pi fintype_finite -apply_nolint generate_from_pi_eq fintype_finite -apply_nolint is_countably_spanning.pi fintype_finite - --- measure_theory/constructions/prod.lean -apply_nolint measure_theory.measure.prod_sum fintype_finite -apply_nolint measure_theory.measure.sum_prod fintype_finite - --- measure_theory/group/prod.lean -apply_nolint measure_theory.map_prod_inv_mul_eq to_additive_doc -apply_nolint measure_theory.map_prod_inv_mul_eq_swap to_additive_doc -apply_nolint measure_theory.map_prod_mul_inv_eq to_additive_doc -apply_nolint measure_theory.measure_lintegral_div_measure to_additive_doc -apply_nolint measure_theory.measure_mul_lintegral_eq to_additive_doc - --- measure_theory/integral/integrable_on.lean -apply_nolint measure_theory.integrable_on_fintype_Union fintype_finite - --- measure_theory/measurable_space.lean -apply_nolint measurable_of_fintype fintype_finite -apply_nolint measurable_set.pi_fintype fintype_finite -apply_nolint measurable_set.univ_pi_fintype fintype_finite - --- measure_theory/measurable_space_def.lean -apply_nolint measurable_set.Inter_fintype fintype_finite -apply_nolint measurable_set.Union_fintype fintype_finite - --- measure_theory/measure/haar.lean -apply_nolint measure_theory.measure.haar.haar_content_self to_additive_doc -apply_nolint measure_theory.measure.haar.index_defined to_additive_doc -apply_nolint measure_theory.measure.haar.is_left_invariant_haar_content to_additive_doc -apply_nolint measure_theory.measure.haar_measure_unique to_additive_doc -apply_nolint measure_theory.measure.is_haar_measure_haar_measure to_additive_doc -apply_nolint measure_theory.measure.map_haar_inv to_additive_doc -apply_nolint measure_theory.measure.regular_haar_measure to_additive_doc -apply_nolint measure_theory.measure.regular_of_is_mul_left_invariant to_additive_doc -apply_nolint measure_theory.measure.sigma_finite_haar_measure to_additive_doc - --- measure_theory/measure/haar_lebesgue.lean -apply_nolint measure_theory.measure.map_linear_map_add_haar_pi_eq_smul_add_haar fintype_finite - --- measure_theory/measure/measure_space.lean -apply_nolint measure_theory.sum.sigma_finite fintype_finite - --- measure_theory/measure/null_measurable.lean -apply_nolint measure_theory.null_measurable_set.Inter_fintype fintype_finite -apply_nolint measure_theory.null_measurable_set.Union_fintype fintype_finite - -- meta/coinductive_predicates.lean apply_nolint monotonicity doc_blame apply_nolint tactic.add_coinductive_predicate doc_blame @@ -676,12 +571,6 @@ apply_nolint ideal.mv_polynomial.mv_polynomial.ideal.is_jacobson fintype_finite -- ring_theory/localization/integer.lean apply_nolint is_localization.exist_integer_multiples_of_fintype fintype_finite --- ring_theory/noetherian.lean -apply_nolint is_noetherian_pi fintype_finite -apply_nolint is_noetherian_pi' fintype_finite -apply_nolint submodule.fg_pi fintype_finite -apply_nolint submodule.fg_supr fintype_finite - -- ring_theory/norm.lean apply_nolint algebra.norm_eq_zero_iff_of_basis fintype_finite apply_nolint algebra.norm_ne_zero_iff_of_basis fintype_finite From f5da0822e83865ec895d35881b68cd74571e1d33 Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Wed, 31 Aug 2022 05:14:16 +0000 Subject: [PATCH 094/828] feat(order/filter/basic): add `eventually_le.add_le_add` (#16295) --- src/order/filter/basic.lean | 11 ++++++++++- 1 file changed, 10 insertions(+), 1 deletion(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 7e654bca61aec..195385af2f5dc 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1436,11 +1436,20 @@ h.mono $ λ x, mt (s \ s' : set α) ≤ᶠ[l] (t \ t' : set α) := h.inter h'.compl -lemma eventually_le.mul_le_mul [ordered_semiring β] {l : filter α} {f₁ f₂ g₁ g₂ : α → β} +lemma eventually_le.mul_le_mul + [mul_zero_class β] [partial_order β] [pos_mul_mono β] [mul_pos_mono β] + {l : filter α} {f₁ f₂ g₁ g₂ : α → β} (hf : f₁ ≤ᶠ[l] f₂) (hg : g₁ ≤ᶠ[l] g₂) (hg₀ : 0 ≤ᶠ[l] g₁) (hf₀ : 0 ≤ᶠ[l] f₂) : f₁ * g₁ ≤ᶠ[l] f₂ * g₂ := by filter_upwards [hf, hg, hg₀, hf₀] with x using mul_le_mul +@[to_additive eventually_le.add_le_add] +lemma eventually_le.mul_le_mul' [has_mul β] [preorder β] + [covariant_class β β (*) (≤)] [covariant_class β β (swap (*)) (≤)] + {l : filter α} {f₁ f₂ g₁ g₂ : α → β} (hf : f₁ ≤ᶠ[l] f₂) (hg : g₁ ≤ᶠ[l] g₂) : + f₁ * g₁ ≤ᶠ[l] f₂ * g₂ := +by filter_upwards [hf, hg] with x hfx hgx using mul_le_mul' hfx hgx + lemma eventually_le.mul_nonneg [ordered_semiring β] {l : filter α} {f g : α → β} (hf : 0 ≤ᶠ[l] f) (hg : 0 ≤ᶠ[l] g) : 0 ≤ᶠ[l] f * g := From 489c1c5706a424fbaaf545a6c83e9baace83235e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 31 Aug 2022 06:55:29 +0000 Subject: [PATCH 095/828] feat(analysis/complex/basic): add `complex.ring_hom_eq_id_or_conj_of_continuous` (#16169) We prove that the only continuous ring homorphism from $\mathbb{C}$ to $\mathbb{C}$ are the identity and the complex conjugation. We deduce the result from a first lemma that gives the same conclusion for the $\mathbb{R}$-algebra homomorphisms from $\mathbb{C}$ to $\mathbb{C}$. --- src/analysis/complex/basic.lean | 11 +++++++++++ src/data/complex/module.lean | 8 ++++++++ 2 files changed, 19 insertions(+) diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index b7f9f9f6324fe..13ccea2e8b15c 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -236,6 +236,17 @@ instance : has_continuous_star ℂ := ⟨conj_lie.continuous⟩ @[continuity] lemma continuous_conj : continuous (conj : ℂ → ℂ) := continuous_star +/-- The only continuous ring homomorphisms from `ℂ` to `ℂ` are the identity and the complex +conjugation. -/ +lemma ring_hom_eq_id_or_conj_of_continuous {f : ℂ →+* ℂ} (hf : continuous f) : + f = ring_hom.id ℂ ∨ f = conj := +begin + refine (real_alg_hom_eq_id_or_conj $ alg_hom.mk' f $ λ x z, congr_fun _ x).imp (λ h, _) (λ h, _), + { refine rat.dense_embedding_coe_real.dense.equalizer (by continuity) (by continuity) _, + ext1, simp only [real_smul, function.comp_app, map_rat_cast, of_real_rat_cast, map_mul], }, + all_goals { convert congr_arg alg_hom.to_ring_hom h, ext1, refl, }, +end + /-- Continuous linear equiv version of the conj function, from `ℂ` to `ℂ`. -/ def conj_cle : ℂ ≃L[ℝ] ℂ := conj_lie diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index 0c13030e78d2e..c9e8450b855eb 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -237,6 +237,14 @@ begin fin_cases i; fin_cases j; simp end +/-- The identity and the complex conjugation are the only two `ℝ`-algebra homomorphisms of `ℂ`. -/ +lemma real_alg_hom_eq_id_or_conj (f : ℂ →ₐ[ℝ] ℂ) : f = alg_hom.id ℝ ℂ ∨ f = conj_ae := +begin + refine (eq_or_eq_neg_of_sq_eq_sq (f I) I $ by rw [← map_pow, I_sq, map_neg, map_one]).imp _ _; + refine λ h, alg_hom_ext _, + exacts [h, conj_I.symm ▸ h], +end + section lift variables {A : Type*} [ring A] [algebra ℝ A] From 20cfd34181ff9a9aec60617b2ac945fc9d866ad0 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 31 Aug 2022 06:55:31 +0000 Subject: [PATCH 096/828] refactor(algebra/monoid_algebra/{ basic + support } + import dust): move lemmas to a new "support" file (#16322) This PR moves some lemmas from `algebra/monoid_algebra/basic` to the new file `algebra/monoid_algebra/support`. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/splitting.20support.20from.20algebra.2Fmonoid_algebra.2Fbasic) --- src/algebra/monoid_algebra/basic.lean | 77 +------------------ src/algebra/monoid_algebra/degree.lean | 2 +- src/algebra/monoid_algebra/grading.lean | 2 +- src/algebra/monoid_algebra/support.lean | 98 +++++++++++++++++++++++++ src/data/mv_polynomial/basic.lean | 2 +- 5 files changed, 102 insertions(+), 79 deletions(-) create mode 100644 src/algebra/monoid_algebra/support.lean diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index e472769810eff..c045905d7fe03 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -338,11 +338,6 @@ calc (f * g) x = (∑ a₁ in f.support, ∑ a₂ in g.support, F (a₁, a₂)) { rw [hp hps h1, mul_zero] } end -lemma support_mul [has_mul G] [decidable_eq G] (a b : monoid_algebra k G) : - (a * b).support ⊆ a.support.bUnion (λa₁, b.support.bUnion $ λa₂, {a₁ * a₂}) := -subset.trans support_sum $ bUnion_mono $ assume a₁ _, - subset.trans support_sum $ bUnion_mono $ assume a₂ _, support_single_subset - @[simp] lemma single_mul_single [has_mul G] {a₁ a₂ : G} {b₁ b₂ : k} : (single a₁ b₁ : monoid_algebra k G) * single a₂ b₂ = single (a₁ * a₂) (b₁ * b₂) := (sum_single_index (by simp only [zero_mul, single_zero, sum_zero])).trans @@ -426,20 +421,6 @@ lemma mul_single_one_apply [mul_one_class G] (f : monoid_algebra k G) (r : k) (x (f * single 1 r) x = f x * r := f.mul_single_apply_aux $ λ a, by rw [mul_one] -lemma support_mul_single [right_cancel_semigroup G] - (f : monoid_algebra k G) (r : k) (hr : ∀ y, y * r = 0 ↔ y = 0) (x : G) : - (f * single x r).support = f.support.map (mul_right_embedding x) := -begin - ext y, simp only [mem_support_iff, mem_map, exists_prop, mul_right_embedding_apply], - by_cases H : ∃ a, a * x = y, - { rcases H with ⟨a, rfl⟩, - rw [mul_single_apply_aux f (λ _, mul_left_inj x)], - simp [hr] }, - { push_neg at H, - classical, - simp [mul_apply, H] } -end - lemma single_mul_apply_aux [has_mul G] (f : monoid_algebra k G) {r : k} {x y z : G} (H : ∀ a, x * a = y ↔ a = z) : (single x r * f) y = r * f z := @@ -455,20 +436,6 @@ lemma single_one_mul_apply [mul_one_class G] (f : monoid_algebra k G) (r : k) (x (single 1 r * f) x = r * f x := f.single_mul_apply_aux $ λ a, by rw [one_mul] -lemma support_single_mul [left_cancel_semigroup G] - (f : monoid_algebra k G) (r : k) (hr : ∀ y, r * y = 0 ↔ y = 0) (x : G) : - (single x r * f).support = f.support.map (mul_left_embedding x) := -begin - ext y, simp only [mem_support_iff, mem_map, exists_prop, mul_left_embedding_apply], - by_cases H : ∃ a, x * a = y, - { rcases H with ⟨a, rfl⟩, - rw [single_mul_apply_aux f (λ _, mul_right_inj x)], - simp [hr] }, - { push_neg at H, - classical, - simp [mul_apply, H] } -end - lemma lift_nc_smul [mul_one_class G] {R : Type*} [semiring R] (f : k →+* R) (g : G →* R) (c : k) (φ : monoid_algebra k G) : lift_nc (f : k →+ R) g (c • φ) = f c * lift_nc (f : k →+ R) g φ := @@ -841,17 +808,6 @@ calc (f * g) x = sum g (λ a b, (f * single a b) x) : end -section span - -variables [semiring k] [mul_one_class G] - -/-- An element of `monoid_algebra R M` is in the subalgebra generated by its support. -/ -lemma mem_span_support (f : monoid_algebra k G) : - f ∈ submodule.span k (of k G '' (f.support : set G)) := -by rw [of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single, finsupp.mem_supported] - -end span - section opposite open finsupp mul_opposite @@ -1153,7 +1109,7 @@ finsupp.is_central_scalar G k because we've never discussed actions of additive groups. -/ end derived_instances - +. section misc_theorems variables [semiring k] @@ -1167,10 +1123,6 @@ lemma mul_apply_antidiagonal [has_add G] (f g : add_monoid_algebra k G) (x : G) (f * g) x = ∑ p in s, (f p.1 * g p.2) := @monoid_algebra.mul_apply_antidiagonal k (multiplicative G) _ _ _ _ _ s @hs -lemma support_mul [decidable_eq G] [has_add G] (a b : add_monoid_algebra k G) : - (a * b).support ⊆ a.support.bUnion (λa₁, b.support.bUnion $ λa₂, {a₁ + a₂}) := -@monoid_algebra.support_mul k (multiplicative G) _ _ _ _ _ - lemma single_mul_single [has_add G] {a₁ a₂ : G} {b₁ b₂ : k} : (single a₁ b₁ * single a₂ b₂ : add_monoid_algebra k G) = single (a₁ + a₂) (b₁ * b₂) := @monoid_algebra.single_mul_single k (multiplicative G) _ _ _ _ _ _ @@ -1276,16 +1228,6 @@ lemma single_mul_apply [add_group G] (r : k) (x : G) (f : add_monoid_algebra k G (single x r * f : add_monoid_algebra k G) y = r * f (- x + y) := @monoid_algebra.single_mul_apply k (multiplicative G) _ _ _ _ _ _ -lemma support_mul_single [add_right_cancel_semigroup G] - (f : add_monoid_algebra k G) (r : k) (hr : ∀ y, y * r = 0 ↔ y = 0) (x : G) : - (f * single x r : add_monoid_algebra k G).support = f.support.map (add_right_embedding x) := -@monoid_algebra.support_mul_single k (multiplicative G) _ _ _ _ hr _ - -lemma support_single_mul [add_left_cancel_semigroup G] - (f : add_monoid_algebra k G) (r : k) (hr : ∀ y, r * y = 0 ↔ y = 0) (x : G) : - (single x r * f : add_monoid_algebra k G).support = f.support.map (add_left_embedding x) := -@monoid_algebra.support_single_mul k (multiplicative G) _ _ _ _ hr _ - lemma lift_nc_smul {R : Type*} [add_zero_class G] [semiring R] (f : k →+* R) (g : multiplicative G →* R) (c : k) (φ : monoid_algebra k G) : lift_nc (f : k →+ R) g (c • φ) = f c * lift_nc (f : k →+ R) g φ := @@ -1314,23 +1256,6 @@ def map_domain_ring_hom (k : Type*) [semiring k] {H F : Type*} [add_monoid G] [a end misc_theorems -section span - -variables [semiring k] - -/-- An element of `add_monoid_algebra R M` is in the submodule generated by its support. -/ -lemma mem_span_support [add_zero_class G] (f : add_monoid_algebra k G) : - f ∈ submodule.span k (of k G '' (f.support : set G)) := -by rw [of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single, finsupp.mem_supported] - -/-- An element of `add_monoid_algebra R M` is in the subalgebra generated by its support, using -unbundled inclusion. -/ -lemma mem_span_support' (f : add_monoid_algebra k G) : - f ∈ submodule.span k (of' k G '' (f.support : set G)) := -by rw [of', ← finsupp.supported_eq_span_single, finsupp.mem_supported] - -end span - end add_monoid_algebra /-! diff --git a/src/algebra/monoid_algebra/degree.lean b/src/algebra/monoid_algebra/degree.lean index a15a552b8fba1..831bc4ae4777e 100644 --- a/src/algebra/monoid_algebra/degree.lean +++ b/src/algebra/monoid_algebra/degree.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ -import algebra.monoid_algebra.basic +import algebra.monoid_algebra.support /-! # Lemmas about the `sup` and `inf` of the support of `add_monoid_algebra` diff --git a/src/algebra/monoid_algebra/grading.lean b/src/algebra/monoid_algebra/grading.lean index 5313961cfe040..00881606744f4 100644 --- a/src/algebra/monoid_algebra/grading.lean +++ b/src/algebra/monoid_algebra/grading.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import linear_algebra.finsupp -import algebra.monoid_algebra.basic +import algebra.monoid_algebra.support import algebra.direct_sum.internal import ring_theory.graded_algebra.basic diff --git a/src/algebra/monoid_algebra/support.lean b/src/algebra/monoid_algebra/support.lean new file mode 100644 index 0000000000000..5ed1a6baba907 --- /dev/null +++ b/src/algebra/monoid_algebra/support.lean @@ -0,0 +1,98 @@ +/- +Copyright (c) 2022 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa +-/ +import algebra.monoid_algebra.basic + +/-! +# Lemmas about the support of a finitely supported function +-/ + +universes u₁ u₂ u₃ +namespace monoid_algebra + +open finset finsupp +variables {k : Type u₁} {G : Type u₂} {R : Type u₃} [semiring k] + +lemma support_mul [has_mul G] [decidable_eq G] (a b : monoid_algebra k G) : + (a * b).support ⊆ a.support.bUnion (λa₁, b.support.bUnion $ λa₂, {a₁ * a₂}) := +subset.trans support_sum $ bUnion_mono $ assume a₁ _, + subset.trans support_sum $ bUnion_mono $ assume a₂ _, support_single_subset + +lemma support_mul_single [right_cancel_semigroup G] + (f : monoid_algebra k G) (r : k) (hr : ∀ y, y * r = 0 ↔ y = 0) (x : G) : + (f * single x r).support = f.support.map (mul_right_embedding x) := +begin + ext y, simp only [mem_support_iff, mem_map, exists_prop, mul_right_embedding_apply], + by_cases H : ∃ a, a * x = y, + { rcases H with ⟨a, rfl⟩, + rw [mul_single_apply_aux f (λ _, mul_left_inj x)], + simp [hr] }, + { push_neg at H, + classical, + simp [mul_apply, H] } +end + +lemma support_single_mul [left_cancel_semigroup G] + (f : monoid_algebra k G) (r : k) (hr : ∀ y, r * y = 0 ↔ y = 0) (x : G) : + (single x r * f : monoid_algebra k G).support = f.support.map (mul_left_embedding x) := +begin + ext y, simp only [mem_support_iff, mem_map, exists_prop, mul_left_embedding_apply], + by_cases H : ∃ a, x * a = y, + { rcases H with ⟨a, rfl⟩, + rw [single_mul_apply_aux f (λ _, mul_right_inj x)], + simp [hr] }, + { push_neg at H, + classical, + simp [mul_apply, H] } +end +section span + +variables [mul_one_class G] + +/-- An element of `monoid_algebra R M` is in the subalgebra generated by its support. -/ +lemma mem_span_support (f : monoid_algebra k G) : + f ∈ submodule.span k (of k G '' (f.support : set G)) := +by rw [of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single, finsupp.mem_supported] + +end span + + +end monoid_algebra + +namespace add_monoid_algebra + +open finset finsupp mul_opposite +variables {k : Type u₁} {G : Type u₂} {R : Type u₃} [semiring k] + +lemma support_mul [decidable_eq G] [has_add G] (a b : add_monoid_algebra k G) : + (a * b).support ⊆ a.support.bUnion (λa₁, b.support.bUnion $ λa₂, {a₁ + a₂}) := +@monoid_algebra.support_mul k (multiplicative G) _ _ _ _ _ + +lemma support_mul_single [add_right_cancel_semigroup G] + (f : add_monoid_algebra k G) (r : k) (hr : ∀ y, y * r = 0 ↔ y = 0) (x : G) : + (f * single x r : add_monoid_algebra k G).support = f.support.map (add_right_embedding x) := +@monoid_algebra.support_mul_single k (multiplicative G) _ _ _ _ hr _ + +lemma support_single_mul [add_left_cancel_semigroup G] + (f : add_monoid_algebra k G) (r : k) (hr : ∀ y, r * y = 0 ↔ y = 0) (x : G) : + (single x r * f : add_monoid_algebra k G).support = f.support.map (add_left_embedding x) := +@monoid_algebra.support_single_mul k (multiplicative G) _ _ _ _ hr _ + +section span + +/-- An element of `add_monoid_algebra R M` is in the submodule generated by its support. -/ +lemma mem_span_support [add_zero_class G] (f : add_monoid_algebra k G) : + f ∈ submodule.span k (of k G '' (f.support : set G)) := +by rw [of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single, finsupp.mem_supported] + +/-- An element of `add_monoid_algebra R M` is in the subalgebra generated by its support, using +unbundled inclusion. -/ +lemma mem_span_support' (f : add_monoid_algebra k G) : + f ∈ submodule.span k (of' k G '' (f.support : set G)) := +by rw [of', ← finsupp.supported_eq_span_single, finsupp.mem_supported] + +end span + +end add_monoid_algebra diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index fc988106a70be..2ad8f7d923525 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -6,7 +6,7 @@ Authors: Johannes Hölzl, Johan Commelin, Mario Carneiro import ring_theory.adjoin.basic import data.finsupp.antidiagonal -import algebra.monoid_algebra.basic +import algebra.monoid_algebra.support import order.symm_diff /-! From 70c486388ff7a1d26d5576a966937c7ca9c30183 Mon Sep 17 00:00:00 2001 From: Gabriel Ebner Date: Wed, 31 Aug 2022 08:44:46 +0000 Subject: [PATCH 097/828] feat(data/finset/locally_finite): more on `filter (< c)` (#16260) --- src/data/finset/locally_finite.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index 2434c17409cd4..84879db5c71a8 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -192,6 +192,19 @@ begin exact and_iff_right_of_imp (λ h, hac.trans h.1), end +lemma Icc_filter_lt_of_lt_right {a b c : α} [decidable_pred (< c)] (h : b < c) : + (Icc a b).filter (< c) = Icc a b := +(finset.filter_eq_self _).2 (λ x hx, lt_of_le_of_lt (mem_Icc.1 hx).2 h) + +lemma Ioc_filter_lt_of_lt_right {a b c : α} [decidable_pred (< c)] (h : b < c) : + (Ioc a b).filter (< c) = Ioc a b := +(finset.filter_eq_self _).2 (λ x hx, lt_of_le_of_lt (mem_Ioc.1 hx).2 h) + +lemma Iic_filter_lt_of_lt_right {α} [preorder α] [locally_finite_order_bot α] + {a c : α} [decidable_pred (< c)] (h : a < c) : + (Iic a).filter (< c) = Iic a := +(finset.filter_eq_self _).2 (λ x hx, lt_of_le_of_lt (mem_Iic.1 hx) h) + variables (a b) [fintype α] lemma filter_lt_lt_eq_Ioo [decidable_pred (λ j, a < j ∧ j < b)] : @@ -436,6 +449,13 @@ begin { rw [Ico_filter_le_of_le_left h, max_eq_left h] } end +@[simp] lemma Ioo_filter_lt (a b c : α) : (Ioo a b).filter (< c) = Ioo a (min b c) := +by { ext, simp [and_assoc] } + +@[simp] lemma Iio_filter_lt {α} [linear_order α] [locally_finite_order_bot α] (a b : α) : + (Iio a).filter (< b) = Iio (min a b) := +by { ext, simp [and_assoc] } + @[simp] lemma Ico_diff_Ico_left (a b c : α) : (Ico a b) \ (Ico a c) = Ico (max a c) b := begin cases le_total a c, From 386acbac7b14df6ec64dabd9a578be5c2ee796f3 Mon Sep 17 00:00:00 2001 From: Gabriel Ebner Date: Wed, 31 Aug 2022 08:44:48 +0000 Subject: [PATCH 098/828] feat(order/locally_finite): transfer locally_finite_order across order_iso (#16264) --- src/order/locally_finite.lean | 37 +++++++++++++++++++++++++++++++++++ 1 file changed, 37 insertions(+) diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index a7d3bcd28f630..b9df19815ebb0 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -784,6 +784,43 @@ lemma Ioo_coe_coe : Ioo (a : with_bot α) b = (Ioo a b).map embedding.coe_option end with_bot +namespace order_iso +variables [preorder α] [preorder β] + +/-! #### Transfer locally finite orders across order isomorphisms -/ + +/-- Transfer `locally_finite_order` across an `order_iso`. -/ +@[reducible] -- See note [reducible non-instances] +def locally_finite_order [locally_finite_order β] (f : α ≃o β) : locally_finite_order α := +{ finset_Icc := λ a b, (Icc (f a) (f b)).map f.symm.to_equiv.to_embedding, + finset_Ico := λ a b, (Ico (f a) (f b)).map f.symm.to_equiv.to_embedding, + finset_Ioc := λ a b, (Ioc (f a) (f b)).map f.symm.to_equiv.to_embedding, + finset_Ioo := λ a b, (Ioo (f a) (f b)).map f.symm.to_equiv.to_embedding, + finset_mem_Icc := by simp, + finset_mem_Ico := by simp, + finset_mem_Ioc := by simp, + finset_mem_Ioo := by simp } + +/-- Transfer `locally_finite_order_top` across an `order_iso`. -/ +@[reducible] -- See note [reducible non-instances] +def locally_finite_order_top [locally_finite_order_top β] + (f : α ≃o β) : locally_finite_order_top α := +{ finset_Ici := λ a, (Ici (f a)).map f.symm.to_equiv.to_embedding, + finset_Ioi := λ a, (Ioi (f a)).map f.symm.to_equiv.to_embedding, + finset_mem_Ici := by simp, + finset_mem_Ioi := by simp } + +/-- Transfer `locally_finite_order_bot` across an `order_iso`. -/ +@[reducible] -- See note [reducible non-instances] +def locally_finite_order_bot [locally_finite_order_bot β] + (f : α ≃o β) : locally_finite_order_bot α := +{ finset_Iic := λ a, (Iic (f a)).map f.symm.to_equiv.to_embedding, + finset_Iio := λ a, (Iio (f a)).map f.symm.to_equiv.to_embedding, + finset_mem_Iic := by simp, + finset_mem_Iio := by simp } + +end order_iso + /-! #### Subtype of a locally finite order -/ variables [preorder α] (p : α → Prop) [decidable_pred p] From 6a68f86a6a796fffcb99598635c1b3046588fcc9 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 31 Aug 2022 08:44:49 +0000 Subject: [PATCH 099/828] chore(topology/algebra/order/basic): deduplicate (#16287) Primed and non-primed versions of 2 lemmas were almost defeq. Merge them. --- src/analysis/calculus/extend_deriv.lean | 4 +-- src/topology/algebra/order/basic.lean | 38 ++++--------------------- 2 files changed, 7 insertions(+), 35 deletions(-) diff --git a/src/analysis/calculus/extend_deriv.lean b/src/analysis/calculus/extend_deriv.lean index 0d87c20369832..73f30665aba7b 100644 --- a/src/analysis/calculus/extend_deriv.lean +++ b/src/analysis/calculus/extend_deriv.lean @@ -131,7 +131,7 @@ begin have : has_deriv_within_at f e (Icc a b) a, { rw [has_deriv_within_at_iff_has_fderiv_within_at, ← t_closure], exact has_fderiv_at_boundary_of_tendsto_fderiv t_diff t_conv t_open t_cont t_diff' }, - exact this.nhds_within (mem_nhds_within_Ici_iff_exists_Icc_subset.2 ⟨b, ab, subset.refl _⟩) + exact this.nhds_within (Icc_mem_nhds_within_Ici $ left_mem_Ico.2 ab) end /-- If a function is differentiable on the left of a point `a : ℝ`, continuous at `a`, and @@ -170,7 +170,7 @@ begin have : has_deriv_within_at f e (Icc b a) a, { rw [has_deriv_within_at_iff_has_fderiv_within_at, ← t_closure], exact has_fderiv_at_boundary_of_tendsto_fderiv t_diff t_conv t_open t_cont t_diff' }, - exact this.nhds_within (mem_nhds_within_Iic_iff_exists_Icc_subset.2 ⟨b, ba, subset.refl _⟩) + exact this.nhds_within (Icc_mem_nhds_within_Iic $ right_mem_Ioc.2 ba) end /-- If a real function `f` has a derivative `g` everywhere but at a point, and `f` and `g` are diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index 9ec7b4ba4999b..622487c98d35f 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -1527,8 +1527,8 @@ let ⟨u', hu'⟩ := exists_gt a in mem_nhds_within_Ici_iff_exists_Ico_subset' h /-- A set is a neighborhood of `a` within `[a, +∞)` if and only if it contains an interval `[a, u]` with `a < u`. -/ -lemma mem_nhds_within_Ici_iff_exists_Icc_subset' [no_max_order α] [densely_ordered α] - {a : α} {s : set α} : s ∈ 𝓝[≥] a ↔ ∃u ∈ Ioi a, Icc a u ⊆ s := +lemma mem_nhds_within_Ici_iff_exists_Icc_subset [no_max_order α] [densely_ordered α] + {a : α} {s : set α} : s ∈ 𝓝[≥] a ↔ ∃ u, a < u ∧ Icc a u ⊆ s := begin rw mem_nhds_within_Ici_iff_exists_Ico_subset, split, @@ -1573,42 +1573,14 @@ let ⟨l', hl'⟩ := exists_lt a in mem_nhds_within_Iic_iff_exists_Ioc_subset' h /-- A set is a neighborhood of `a` within `(-∞, a]` if and only if it contains an interval `[l, a]` with `l < a`. -/ -lemma mem_nhds_within_Iic_iff_exists_Icc_subset' [no_min_order α] [densely_ordered α] - {a : α} {s : set α} : s ∈ 𝓝[≤] a ↔ ∃l ∈ Iio a, Icc l a ⊆ s := +lemma mem_nhds_within_Iic_iff_exists_Icc_subset [no_min_order α] [densely_ordered α] + {a : α} {s : set α} : s ∈ 𝓝[≤] a ↔ ∃ l, l < a ∧ Icc l a ⊆ s := begin - convert @mem_nhds_within_Ici_iff_exists_Icc_subset' αᵒᵈ _ _ _ _ _ _ _, + convert @mem_nhds_within_Ici_iff_exists_Icc_subset αᵒᵈ _ _ _ _ _ _ _, simp_rw (show ∀ u : αᵒᵈ, @Icc αᵒᵈ _ a u = @Icc α _ u a, from λ u, dual_Icc), refl, end -/-- A set is a neighborhood of `a` within `[a, +∞)` if and only if it contains an interval `[a, u]` -with `a < u`. -/ -lemma mem_nhds_within_Ici_iff_exists_Icc_subset [no_max_order α] [densely_ordered α] - {a : α} {s : set α} : s ∈ 𝓝[≥] a ↔ ∃u, a < u ∧ Icc a u ⊆ s := -begin - rw mem_nhds_within_Ici_iff_exists_Ico_subset, - split, - { rintros ⟨u, au, as⟩, - rcases exists_between au with ⟨v, hv⟩, - exact ⟨v, hv.1, λx hx, as ⟨hx.1, lt_of_le_of_lt hx.2 hv.2⟩⟩ }, - { rintros ⟨u, au, as⟩, - exact ⟨u, au, subset.trans Ico_subset_Icc_self as⟩ } -end - -/-- A set is a neighborhood of `a` within `(-∞, a]` if and only if it contains an interval `[l, a]` -with `l < a`. -/ -lemma mem_nhds_within_Iic_iff_exists_Icc_subset [no_min_order α] [densely_ordered α] - {a : α} {s : set α} : s ∈ 𝓝[≤] a ↔ ∃l, l < a ∧ Icc l a ⊆ s := -begin - rw mem_nhds_within_Iic_iff_exists_Ioc_subset, - split, - { rintros ⟨l, la, as⟩, - rcases exists_between la with ⟨v, hv⟩, - refine ⟨v, hv.2, λx hx, as ⟨lt_of_lt_of_le hv.1 hx.1, hx.2⟩⟩, }, - { rintros ⟨l, la, as⟩, - exact ⟨l, la, subset.trans Ioc_subset_Icc_self as⟩ } -end - end order_topology end linear_order From b35461b88242ba6353ffeb31e5efebe6476c37cd Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Wed, 31 Aug 2022 08:44:50 +0000 Subject: [PATCH 100/828] feat(logic/basic): `function.mt` and `function.mtr` (#16315) Add modus tollens and reversed modus tollens to the `function` namespace so that they are available for dot notation on implications. --- src/logic/basic.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/logic/basic.lean b/src/logic/basic.lean index 0a712ccba10a8..7cafe15c710c1 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -298,6 +298,9 @@ if ha : a then by simp only [ha, true_and, true_implies_iff] @[simp] theorem and_or_imp : (a ∧ b) ∨ (a → c) ↔ a → (b ∨ c) := decidable.and_or_imp +/-- Provide modus tollens (`mt`) as dot notation for implications. -/ +protected lemma function.mt : (a → b) → ¬ b → ¬ a := mt + /-! ### Declarations about `not` -/ /-- Ex falso for negation. From `¬ a` and `a` anything follows. This is the same as `absurd` with @@ -581,6 +584,9 @@ protected theorem decidable.not_imp_not [decidable a] : (¬ a → ¬ b) ↔ (b theorem not_imp_not : (¬ a → ¬ b) ↔ (b → a) := decidable.not_imp_not +/-- Provide the reverse of modus tollens (`mt`) as dot notation for implications. -/ +protected theorem function.mtr : (¬ a → ¬ b) → (b → a) := not_imp_not.mp + -- See Note [decidable namespace] protected lemma decidable.or_congr_left [decidable c] (h : ¬ c → (a ↔ b)) : a ∨ c ↔ b ∨ c := by { rw [decidable.or_iff_not_imp_right, decidable.or_iff_not_imp_right], exact imp_congr_right h } From 17ff82d9ff0f238068d861e644d62b2fed95534a Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Wed, 31 Aug 2022 08:44:51 +0000 Subject: [PATCH 101/828] feat(topology/instances/ennreal): `ennreal.to_nnreal` applied to a `tsum` (#16320) If a function is finite for all inputs, then `to_nnreal` applied to the sum is the sum of `to_nnreal` applied to the values. --- src/topology/instances/ennreal.lean | 17 +++++++---------- 1 file changed, 7 insertions(+), 10 deletions(-) diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 1d6b77812a7a7..78c2d96d4ff96 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -1127,17 +1127,14 @@ end nnreal namespace ennreal -lemma tsum_to_real_eq - {f : α → ℝ≥0∞} (hf : ∀ a, f a ≠ ∞) : +lemma tsum_to_nnreal_eq {f : α → ℝ≥0∞} (hf : ∀ a, f a ≠ ∞) : + (∑' a, f a).to_nnreal = ∑' a, (f a).to_nnreal := +(congr_arg ennreal.to_nnreal (tsum_congr $ λ x, (coe_to_nnreal (hf x)).symm)).trans + nnreal.tsum_eq_to_nnreal_tsum.symm + +lemma tsum_to_real_eq {f : α → ℝ≥0∞} (hf : ∀ a, f a ≠ ∞) : (∑' a, f a).to_real = ∑' a, (f a).to_real := -begin - lift f to α → ℝ≥0 using hf, - have : (∑' (a : α), (f a : ℝ≥0∞)).to_real = - ((∑' (a : α), (f a : ℝ≥0∞)).to_nnreal : ℝ≥0∞).to_real, - { rw [ennreal.coe_to_real], refl }, - rw [this, ← nnreal.tsum_eq_to_nnreal_tsum, ennreal.coe_to_real], - exact nnreal.coe_tsum -end +by simp only [ennreal.to_real, tsum_to_nnreal_eq hf, nnreal.coe_tsum] lemma tendsto_sum_nat_add (f : ℕ → ℝ≥0∞) (hf : ∑' i, f i ≠ ∞) : tendsto (λ i, ∑' k, f (k + i)) at_top (𝓝 0) := From f84386d781c8750bcb28e0303d5199a7cdd46098 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 31 Aug 2022 08:44:53 +0000 Subject: [PATCH 102/828] chore(algebra/regular/basic): clean up variables and slight golf (#16321) Re-organize variables in the file, so that typeclass assumptions appear earlier. Also, golf two proofs and remove two unneeded `@`. --- src/algebra/regular/basic.lean | 36 +++++++++++++--------------------- 1 file changed, 14 insertions(+), 22 deletions(-) diff --git a/src/algebra/regular/basic.lean b/src/algebra/regular/basic.lean index 86bd53f1901de..231870d1342e7 100644 --- a/src/algebra/regular/basic.lean +++ b/src/algebra/regular/basic.lean @@ -24,24 +24,24 @@ by adding one further `0`. The final goal is to develop part of the API to prove, eventually, results about non-zero-divisors. -/ -variables {R : Type*} {a b : R} +variables {R : Type*} section has_mul -variable [has_mul R] +variables [has_mul R] /-- A left-regular element is an element `c` such that multiplication on the left by `c` is injective. -/ @[to_additive "An add-left-regular element is an element `c` such that addition on the left by `c` is injective. -/ "] -def is_left_regular (c : R) := function.injective ((*) c) +def is_left_regular (c : R) := ((*) c).injective /-- A right-regular element is an element `c` such that multiplication on the right by `c` is injective. -/ @[to_additive "An add-right-regular element is an element `c` such that addition on the right by `c` is injective."] -def is_right_regular (c : R) := function.injective (* c) +def is_right_regular (c : R) := (* c).injective /-- An add-regular element is an element `c` such that addition by `c` both on the left and on the right is injective. -/ @@ -74,7 +74,7 @@ end has_mul section semigroup -variable [semigroup R] +variables [semigroup R] {a b : R} /-- In a semigroup, the product of left-regular elements is left-regular. -/ @[to_additive "In an additive semigroup, the sum of add-left-regular elements is add-left.regular."] @@ -114,7 +114,7 @@ lemma is_right_regular.of_mul (ab : is_right_regular (b * a)) : begin refine λ x y xy, ab (_ : x * (b * a) = y * (b * a)), rw [← mul_assoc, ← mul_assoc], - exact congr_fun (congr_arg has_mul.mul xy) a, + exact congr_fun (congr_arg (*) xy) a, end /-- An element is right-regular if and only if multiplying it on the right with a right-regular @@ -154,7 +154,7 @@ end semigroup section mul_zero_class -variables [mul_zero_class R] +variables [mul_zero_class R] {a b : R} /-- The element `0` is left-regular if and only if `R` is trivial. -/ lemma is_left_regular.subsingleton (h : is_left_regular (0 : R)) : subsingleton R := @@ -170,11 +170,7 @@ h.left.subsingleton /-- The element `0` is left-regular if and only if `R` is trivial. -/ lemma is_left_regular_zero_iff_subsingleton : is_left_regular (0 : R) ↔ subsingleton R := -begin - refine ⟨λ h, h.subsingleton, _⟩, - intros H a b h, - exact @subsingleton.elim _ H a b -end +⟨λ h, h.subsingleton, λ H a b h, @subsingleton.elim _ H a b⟩ /-- In a non-trivial `mul_zero_class`, the `0` element is not left-regular. -/ lemma not_is_left_regular_zero_iff : ¬ is_left_regular (0 : R) ↔ nontrivial R := @@ -186,11 +182,7 @@ end /-- The element `0` is right-regular if and only if `R` is trivial. -/ lemma is_right_regular_zero_iff_subsingleton : is_right_regular (0 : R) ↔ subsingleton R := -begin - refine ⟨λ h, h.subsingleton, _⟩, - intros H a b h, - exact @subsingleton.elim _ H a b -end +⟨λ h, h.subsingleton, λ H a b h, @subsingleton.elim _ H a b⟩ /-- In a non-trivial `mul_zero_class`, the `0` element is not right-regular. -/ lemma not_is_right_regular_zero_iff : ¬ is_right_regular (0 : R) ↔ nontrivial R := @@ -255,7 +247,7 @@ end mul_one_class section comm_semigroup -variable [comm_semigroup R] +variables [comm_semigroup R] {a b : R} /-- A product is regular if and only if the factors are. -/ @[to_additive "A sum is add-regular if and only if the summands are."] @@ -269,17 +261,17 @@ end comm_semigroup section monoid -variables [monoid R] +variables [monoid R] {a b : R} /-- An element admitting a left inverse is left-regular. -/ @[to_additive "An element admitting a left additive opposite is add-left-regular."] lemma is_left_regular_of_mul_eq_one (h : b * a = 1) : is_left_regular a := -@is_left_regular.of_mul R _ a _ (by { rw h, exact is_regular_one.left }) +@is_left_regular.of_mul R _ _ _ (by { rw h, exact is_regular_one.left }) /-- An element admitting a right inverse is right-regular. -/ @[to_additive "An element admitting a right additive opposite is add-right-regular."] lemma is_right_regular_of_mul_eq_one (h : a * b = 1) : is_right_regular a := -@is_right_regular.of_mul R _ a _ (by { rw h, exact is_regular_one.right }) +is_right_regular.of_mul (by { rw h, exact is_regular_one.right }) /-- If `R` is a monoid, an element in `Rˣ` is regular. -/ @[to_additive "If `R` is an additive monoid, an element in `add_units R` is add-regular."] @@ -351,7 +343,7 @@ end cancel_monoid section cancel_monoid_with_zero -variables [cancel_monoid_with_zero R] +variables [cancel_monoid_with_zero R] {a : R} /-- Non-zero elements of an integral domain are regular. -/ lemma is_regular_of_ne_zero (a0 : a ≠ 0) : is_regular a := From 9e274549d1bd48b4fa00b228c8c270aa337db7a0 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 31 Aug 2022 11:42:37 +0000 Subject: [PATCH 103/828] feat(algebra/category/Module/change_of_rings): extension of scalars (#15673) Given a ring homomorphism $f : R \to S$ between commutative rings, the extension of scalars functor from $R$-module to $S$-module. In #15564 it will proven that extension of scalars $\dashv$ restriction of scalars Co-authored-by: Jujian Zhang --- .../category/Module/change_of_rings.lean | 98 +++++++++++++++++++ 1 file changed, 98 insertions(+) diff --git a/src/algebra/category/Module/change_of_rings.lean b/src/algebra/category/Module/change_of_rings.lean index fc8d97972b2a5..747bcec4a2fb2 100644 --- a/src/algebra/category/Module/change_of_rings.lean +++ b/src/algebra/category/Module/change_of_rings.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import algebra.category.Module.basic +import ring_theory.tensor_product /-! # Change Of Rings @@ -13,6 +14,16 @@ import algebra.category.Module.basic * `category_theory.Module.restrict_scalars`: given rings `R, S` and a ring homomorphism `R ⟶ S`, then `restrict_scalars : Module S ⥤ Module R` is defined by `M ↦ M` where `M : S-module` is seen as `R-module` by `r • m := f r • m` and `S`-linear map `l : M ⟶ M'` is `R`-linear as well. + +* `category_theory.Module.extend_scalars`: given **commutative** rings `R, S` and ring homomorphism + `f : R ⟶ S`, then `extend_scalars : Module R ⥤ Module S` is defined by `M ↦ S ⨂ M` where the + module structure is defined by `s • (s' ⊗ m) := (s * s') ⊗ m` and `R`-linear map `l : M ⟶ M'` + is sent to `S`-linear map `s ⊗ m ↦ s ⊗ l m : S ⨂ M ⟶ S ⨂ M'`. + +## List of notations +Let `R, S` be rings and `f : R →+* S` +* if `M` is an `R`-module, `s : S` and `m : M`, then `s ⊗ₜ[R, f] m` is the pure tensor + `s ⊗ m : S ⊗[R, f] M`. -/ @@ -62,4 +73,91 @@ def restrict_scalars {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R @[simp] lemma restrict_scalars.smul_def' {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) {M : Module.{v} S} (r : R) (m : M) : (r • m : (restrict_scalars f).obj M) = (f r • m : M) := rfl +@[priority 100] +instance smul_comm_class_mk {R : Type u₁} {S : Type u₂} [ring R] [comm_ring S] (f : R →+* S) + (M : Type v) [add_comm_group M] [module S M] : + @smul_comm_class R S M ((restrict_scalars.obj' f (Module.mk M)).is_module.to_has_smul) _ := +{ smul_comm := λ r s m, (by simp [←mul_smul, mul_comm] : f r • s • m = s • f r • m) } + +namespace extend_scalars + +open tensor_product + +variables {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] (f : R →+* S) + +section unbundled + +variables (M : Type v) [add_comm_monoid M] [module R M] +-- This notation is necessary because we need to reason about `s ⊗ₜ m` where `s : S` and `m : M`; +-- without this notation, one need to work with `s : (restrict_scalars f).obj ⟨S⟩`. +localized "notation s `⊗ₜ[` R `,` f `]` m := @tensor_product.tmul R _ _ _ _ _ + (module.comp_hom _ f) _ s m" in change_of_rings + +end unbundled + +open_locale change_of_rings + +variables (M : Module.{v} R) + +/-- +Extension of scalars turn an `R`-module into `S`-module by M ↦ S ⨂ M +-/ +def obj' : Module S := +⟨tensor_product R ((restrict_scalars f).obj ⟨S⟩) M⟩ + +/-- +Extension of scalars is a functor where an `R`-module `M` is sent to `S ⊗ M` and +`l : M1 ⟶ M2` is sent to `s ⊗ m ↦ s ⊗ l m` +-/ +def map' {M1 M2 : Module.{v} R} (l : M1 ⟶ M2) : (obj' f M1) ⟶ (obj' f M2) := +@linear_map.base_change R S M1 M2 _ _ ((algebra_map S _).comp f).to_algebra _ _ _ _ l + +lemma map'_id {M : Module.{v} R} : map' f (𝟙 M) = 𝟙 _ := +linear_map.ext $ λ (x : obj' f M), +begin + dsimp only [map', Module.id_apply], + induction x using tensor_product.induction_on with _ _ m s ihx ihy, + { simp only [map_zero], }, + { rw [linear_map.base_change_tmul, Module.id_apply], }, + { rw [map_add, ihx, ihy] }, +end + +lemma map'_comp {M₁ M₂ M₃ : Module.{v} R} (l₁₂ : M₁ ⟶ M₂) (l₂₃ : M₂ ⟶ M₃) : + map' f (l₁₂ ≫ l₂₃) = map' f l₁₂ ≫ map' f l₂₃ := +linear_map.ext $ λ (x : obj' f M₁), +begin + dsimp only [map'], + induction x using tensor_product.induction_on with _ _ x y ihx ihy, + { refl, }, + { refl, }, + { simp only [map_add, ihx, ihy], }, +end + +end extend_scalars + +/-- +Extension of scalars is a functor where an `R`-module `M` is sent to `S ⊗ M` and +`l : M1 ⟶ M2` is sent to `s ⊗ m ↦ s ⊗ l m` +-/ +def extend_scalars {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] (f : R →+* S) : + Module.{v} R ⥤ Module.{max v u₂} S := +{ obj := λ M, extend_scalars.obj' f M, + map := λ M1 M2 l, extend_scalars.map' f l, + map_id' := λ _, extend_scalars.map'_id f, + map_comp' := λ _ _ _, extend_scalars.map'_comp f } + +namespace extend_scalars + +open_locale change_of_rings + +variables {R : Type u₁} {S : Type u₂} [comm_ring R] [comm_ring S] (f : R →+* S) + +@[simp] protected lemma smul_tmul {M : Module.{v} R} (s s' : S) (m : M) : + s • (s' ⊗ₜ[R, f] m : (extend_scalars f).obj M) = (s * s') ⊗ₜ[R, f] m := rfl + +@[simp] lemma map_tmul {M M' : Module.{v} R} (g : M ⟶ M') (s : S) (m : M) : + (extend_scalars f).map g (s ⊗ₜ[R, f] m) = s ⊗ₜ[R, f] g m := rfl + +end extend_scalars + end category_theory.Module From f951e201d416fb50cc7826171d80aa510ec20747 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 31 Aug 2022 11:42:38 +0000 Subject: [PATCH 104/828] feat(algebraic_topology/dold_kan): the normalized Moore complex is homotopy equivalent to the alternating face map complex (#16246) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this PR, when the category `A` is abelian, we obtain the homotopy equivalence `homotopy_equiv_normalized_Moore_complex_alternating_face_map_complex` between the normalized Moore complex and the alternating face map complex of a simplicial object in `A`. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- .../dold_kan/homotopy_equivalence.lean | 93 +++++++++++++++++++ 1 file changed, 93 insertions(+) create mode 100644 src/algebraic_topology/dold_kan/homotopy_equivalence.lean diff --git a/src/algebraic_topology/dold_kan/homotopy_equivalence.lean b/src/algebraic_topology/dold_kan/homotopy_equivalence.lean new file mode 100644 index 0000000000000..da71ea0ff7381 --- /dev/null +++ b/src/algebraic_topology/dold_kan/homotopy_equivalence.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.dold_kan.normalized + +/-! + +# The normalized Moore complex and the alternating face map complex are homotopy equivalent + +In this file, when the category `A` is abelian, we obtain the homotopy equivalence +`homotopy_equiv_normalized_Moore_complex_alternating_face_map_complex` between the +normalized Moore complex and the alternating face map complex of a simplicial object in `A`. + +-/ + +open category_theory category_theory.category category_theory.limits + category_theory.preadditive +open_locale simplicial dold_kan + +noncomputable theory + +namespace algebraic_topology + +namespace dold_kan + +variables {C : Type*} [category C] [preadditive C] (X : simplicial_object C) + +/-- Inductive construction of homotopies from `P q` to `𝟙 _` -/ +noncomputable def homotopy_P_to_id : Π (q : ℕ), + homotopy (P q : K[X] ⟶ _) (𝟙 _) +| 0 := homotopy.refl _ +| (q+1) := begin + refine homotopy.trans (homotopy.of_eq _) + (homotopy.trans + (homotopy.add (homotopy_P_to_id q) (homotopy.comp_left (homotopy_Hσ_to_zero q) (P q))) + (homotopy.of_eq _)), + { unfold P, simp only [comp_add, comp_id], }, + { simp only [add_zero, comp_zero], }, + end + +/-- The complement projection `Q q` to `P q` is homotopic to zero. -/ +def homotopy_Q_to_zero (q : ℕ) : homotopy (Q q : K[X] ⟶ _) 0 := +homotopy.equiv_sub_zero.to_fun (homotopy_P_to_id X q).symm + +lemma homotopy_P_to_id_eventually_constant {q n : ℕ} (hqn : n Date: Wed, 31 Aug 2022 11:42:39 +0000 Subject: [PATCH 105/828] =?UTF-8?q?feat(algebra/order/smul):=20`ordered=5F?= =?UTF-8?q?smul`=20instances=20for=20`=E2=84=95`=20and=20`=E2=84=A4`=20(#1?= =?UTF-8?q?6247)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A linear ordered monoid/group has ordered scalar multiplication by `ℕ`/`ℤ`. --- src/algebra/order/smul.lean | 30 ++++++++++++++++++++++++++++-- test/positivity.lean | 4 ++++ 2 files changed, 32 insertions(+), 2 deletions(-) diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index 9b3e7f1bfc7da..a0e0eb0d9b53f 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -142,10 +142,36 @@ lemma bdd_above.smul_of_nonneg (hs : bdd_above s) (hc : 0 ≤ c) : bdd_above (c end ordered_smul +/-- To prove that a linear ordered monoid is an ordered module, it suffices to verify only the first +axiom of `ordered_smul`. -/ +lemma ordered_smul.mk'' [ordered_semiring 𝕜] [linear_ordered_add_comm_monoid M] [smul_with_zero 𝕜 M] + (h : ∀ ⦃c : 𝕜⦄, 0 < c → strict_mono (λ a : M, c • a)) : + ordered_smul 𝕜 M := +{ smul_lt_smul_of_pos := λ a b c hab hc, h hc hab, + lt_of_smul_lt_smul_of_pos := λ a b c hab hc, (h hc).lt_iff_lt.1 hab } + +instance nat.ordered_smul [linear_ordered_cancel_add_comm_monoid M] : ordered_smul ℕ M := +ordered_smul.mk'' $ λ n hn a b hab, begin + cases n, + { cases hn }, + induction n with n ih, + { simp only [one_nsmul, hab], }, + { simp only [succ_nsmul _ n.succ, add_lt_add hab (ih n.succ_pos)] } +end + +instance int.ordered_smul [linear_ordered_add_comm_group M] : ordered_smul ℤ M := +ordered_smul.mk'' $ λ n hn, begin + cases n, + { simp only [int.of_nat_eq_coe, int.coe_nat_pos, coe_nat_zsmul] at ⊢ hn, + exact strict_mono_smul_left hn }, + { cases (int.neg_succ_not_pos _).1 hn } +end + +-- TODO: `linear_ordered_field M → ordered_smul ℚ M` + instance linear_ordered_semiring.to_ordered_smul {R : Type*} [linear_ordered_semiring R] : ordered_smul R R := -{ smul_lt_smul_of_pos := ordered_semiring.mul_lt_mul_of_pos_left, - lt_of_smul_lt_smul_of_pos := λ _ _ _ h hc, lt_of_mul_lt_mul_left h hc.le } +ordered_smul.mk'' $ λ c, strict_mono_mul_left_of_pos section linear_ordered_semifield variables [linear_ordered_semifield 𝕜] diff --git a/test/positivity.lean b/test/positivity.lean index 0e7bd9fd059f5..8c3baf504fe01 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -1,3 +1,4 @@ +import algebra.order.smul import analysis.normed.group.basic import data.real.sqrt import tactic.positivity @@ -92,6 +93,9 @@ example : 0 ≤ max (0:ℤ) (-3) := by positivity example : 0 ≤ max (-3 : ℤ) 5 := by positivity +example {α β : Type*} [ordered_semiring α] [ordered_add_comm_monoid β] [smul_with_zero α β] + [ordered_smul α β] {a : α} (ha : 0 < a) {b : β} (hb : 0 < b) : 0 ≤ a • b := by positivity + example {V : Type*} [normed_add_comm_group V] (x : V) : 0 ≤ ∥x∥ := by positivity example {X : Type*} [metric_space X] (x y : X) : 0 ≤ dist x y := by positivity From 2d30ffcfa296fb74b71f54080a049120220a807f Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Wed, 31 Aug 2022 11:42:40 +0000 Subject: [PATCH 106/828] feat(algebra/gcd_monoid): `simp` lemmas for `associated (normalize x) y` (#16249) The lemmas in `gcd_monoid` package `associated_normalize` and `normalize_associated` into a form suitable for `simp`. Useful e.g. for working with `normalized_factors`. The lemmas in `associated` are the `simp`-normal forms of the `normalize` lemmas, since `normalize x` is not itself in `simp`-normal form. --- src/algebra/associated.lean | 42 +++++++++++++++++++++++++++++++ src/algebra/gcd_monoid/basic.lean | 8 ++++++ 2 files changed, 50 insertions(+) diff --git a/src/algebra/associated.lean b/src/algebra/associated.lean index 6e0df0f4ff04c..f27c8be38caae 100644 --- a/src/algebra/associated.lean +++ b/src/algebra/associated.lean @@ -292,6 +292,9 @@ instance [monoid α] : is_refl α associated := ⟨associated.refl⟩ | x _ ⟨u, rfl⟩ := ⟨u⁻¹, by rw [mul_assoc, units.mul_inv, mul_one]⟩ instance [monoid α] : is_symm α associated := ⟨λ a b, associated.symm⟩ +protected theorem comm [monoid α] {x y : α} : x ~ᵤ y ↔ y ~ᵤ x := +⟨associated.symm, associated.symm⟩ + @[trans] protected theorem trans [monoid α] : ∀{x y z : α}, x ~ᵤ y → y ~ᵤ z → x ~ᵤ z | x _ _ ⟨u, rfl⟩ ⟨v, rfl⟩ := ⟨u * v, by rw [units.coe_mul, mul_assoc]⟩ instance [monoid α] : is_trans α associated := ⟨λ a b c, associated.trans⟩ @@ -342,6 +345,45 @@ lemma associated_unit_mul_right {β : Type*} [comm_monoid β] (a u : β) (hu : i associated a (u * a) := (associated_unit_mul_left a u hu).symm +lemma associated_mul_is_unit_left_iff {β : Type*} [monoid β] {a u b : β} (hu : is_unit u) : + associated (a * u) b ↔ associated a b := +⟨trans (associated_mul_unit_right _ _ hu), trans (associated_mul_unit_left _ _ hu)⟩ + +lemma associated_is_unit_mul_left_iff {β : Type*} [comm_monoid β] {u a b : β} (hu : is_unit u) : + associated (u * a) b ↔ associated a b := +begin + rw mul_comm, + exact associated_mul_is_unit_left_iff hu +end + +lemma associated_mul_is_unit_right_iff {β : Type*} [monoid β] {a b u : β} (hu : is_unit u) : + associated a (b * u) ↔ associated a b := +associated.comm.trans $ (associated_mul_is_unit_left_iff hu).trans associated.comm + +lemma associated_is_unit_mul_right_iff {β : Type*} [comm_monoid β] {a u b : β} (hu : is_unit u) : + associated a (u * b) ↔ associated a b := +associated.comm.trans $ (associated_is_unit_mul_left_iff hu).trans associated.comm + +@[simp] +lemma associated_mul_unit_left_iff {β : Type*} [monoid β] {a b : β} {u : units β} : + associated (a * u) b ↔ associated a b := +associated_mul_is_unit_left_iff u.is_unit + +@[simp] +lemma associated_unit_mul_left_iff {β : Type*} [comm_monoid β] {a b : β} {u : units β} : + associated (↑u * a) b ↔ associated a b := +associated_is_unit_mul_left_iff u.is_unit + +@[simp] +lemma associated_mul_unit_right_iff {β : Type*} [monoid β] {a b : β} {u : units β} : + associated a (b * u) ↔ associated a b := +associated_mul_is_unit_right_iff u.is_unit + +@[simp] +lemma associated_unit_mul_right_iff {β : Type*} [comm_monoid β] {a b : β} {u : units β} : + associated a (↑u * b) ↔ associated a b := +associated_is_unit_mul_right_iff u.is_unit + lemma associated.mul_mul [comm_monoid α] {a₁ a₂ b₁ b₂ : α} : a₁ ~ᵤ b₁ → a₂ ~ᵤ b₂ → (a₁ * a₂) ~ᵤ (b₁ * b₂) | ⟨c₁, h₁⟩ ⟨c₂, h₂⟩ := ⟨c₁ * c₂, by simp [h₁.symm, h₂.symm, mul_assoc, mul_comm, mul_left_comm]⟩ diff --git a/src/algebra/gcd_monoid/basic.lean b/src/algebra/gcd_monoid/basic.lean index 28ce75282b94f..b9d1989f192fc 100644 --- a/src/algebra/gcd_monoid/basic.lean +++ b/src/algebra/gcd_monoid/basic.lean @@ -101,6 +101,14 @@ theorem associated_normalize (x : α) : associated x (normalize x) := theorem normalize_associated (x : α) : associated (normalize x) x := (associated_normalize _).symm +lemma associated_normalize_iff {x y : α} : + associated x (normalize y) ↔ associated x y := +⟨λ h, h.trans (normalize_associated y), λ h, h.trans (associated_normalize y)⟩ + +lemma normalize_associated_iff {x y : α} : + associated (normalize x) y ↔ associated x y := +⟨λ h, (associated_normalize _).trans h, λ h, (normalize_associated _).trans h⟩ + lemma associates.mk_normalize (x : α) : associates.mk (normalize x) = associates.mk x := associates.mk_eq_mk_iff_associated.2 (normalize_associated _) From d44692a8858b1b86990494f9240a4a608a46d381 Mon Sep 17 00:00:00 2001 From: Apurva Nakade Date: Wed, 31 Aug 2022 11:42:42 +0000 Subject: [PATCH 107/828] refactor(analysis/convex/cone): move cone.lean to cone/basic.lean (#16270) Part of #16266 I intend to create another file for `proper_cone`, which for cyclic import reasons cannot be included in cone.lean. So I want to move cone.lean to cone/basic.lean and create cone/proper.lean in a separate PR. --- src/analysis/convex/{cone.lean => cone/basic.lean} | 1 - src/analysis/normed_space/hahn_banach/extension.lean | 2 +- src/analysis/normed_space/hahn_banach/separation.lean | 2 +- 3 files changed, 2 insertions(+), 3 deletions(-) rename src/analysis/convex/{cone.lean => cone/basic.lean} (99%) diff --git a/src/analysis/convex/cone.lean b/src/analysis/convex/cone/basic.lean similarity index 99% rename from src/analysis/convex/cone.lean rename to src/analysis/convex/cone/basic.lean index f8b1c2888089b..92aea1081b836 100644 --- a/src/analysis/convex/cone.lean +++ b/src/analysis/convex/cone/basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Yury Kudryashov All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Frédéric Dupuis -/ -import analysis.convex.hull import analysis.inner_product_space.projection /-! diff --git a/src/analysis/normed_space/hahn_banach/extension.lean b/src/analysis/normed_space/hahn_banach/extension.lean index 103c5f6d68263..1210370615d58 100644 --- a/src/analysis/normed_space/hahn_banach/extension.lean +++ b/src/analysis/normed_space/hahn_banach/extension.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Yury Kudryashov All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Heather Macbeth -/ -import analysis.convex.cone +import analysis.convex.cone.basic import analysis.normed_space.is_R_or_C import analysis.normed_space.extend diff --git a/src/analysis/normed_space/hahn_banach/separation.lean b/src/analysis/normed_space/hahn_banach/separation.lean index 934aa3335a837..490875ab9ab28 100644 --- a/src/analysis/normed_space/hahn_banach/separation.lean +++ b/src/analysis/normed_space/hahn_banach/separation.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Bhavik Mehta All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta, Yaël Dillies -/ -import analysis.convex.cone +import analysis.convex.cone.basic import analysis.convex.gauge /-! From 1e91718fd194a04d40bd9fd9c6d60274e5839ffb Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 31 Aug 2022 11:42:43 +0000 Subject: [PATCH 108/828] feat(logic/equiv/option): equivalence with subtypes (#16302) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add an equivalence between equivalences ```lean /-- Equivalences between `option α` and `β` that send `none` to `x` are equivalent to equivalences between `α` and `{y : β // y ≠ x}`. -/ def option_subtype [decidable_eq β] (x : β) : {e : option α ≃ β // e none = x} ≃ (α ≃ {y : β // y ≠ x}) := ``` which can be used to give a much simpler definition of an equivalence in #16278. --- src/logic/equiv/option.lean | 83 +++++++++++++++++++++++++++++++++++++ 1 file changed, 83 insertions(+) diff --git a/src/logic/equiv/option.lean b/src/logic/equiv/option.lean index 02cf1fffc60cd..4e5b76afdd771 100644 --- a/src/logic/equiv/option.lean +++ b/src/logic/equiv/option.lean @@ -19,6 +19,8 @@ We define namespace equiv +open option + variables {α β γ : Type*} section option_congr @@ -127,4 +129,85 @@ end remove_none lemma option_congr_injective : function.injective (option_congr : α ≃ β → option α ≃ option β) := function.left_inverse.injective remove_none_option_congr +/-- Equivalences between `option α` and `β` that send `none` to `x` are equivalent to +equivalences between `α` and `{y : β // y ≠ x}`. -/ +def option_subtype [decidable_eq β] (x : β) : + {e : option α ≃ β // e none = x} ≃ (α ≃ {y : β // y ≠ x}) := +{ to_fun := λ e, + { to_fun := λ a, ⟨e a, ((equiv_like.injective _).ne_iff' e.property).2 (some_ne_none _)⟩, + inv_fun := λ b, get (ne_none_iff_is_some.1 (((equiv_like.injective _).ne_iff' + (((apply_eq_iff_eq_symm_apply _).1 e.property).symm)).2 b.property)), + left_inv := λ a, begin + rw [←some_inj, some_get, ←coe_def], + exact symm_apply_apply (e : option α ≃ β) a + end, + right_inv := λ b, begin + ext, + simp, + exact apply_symm_apply _ _ + end }, + inv_fun := λ e, + ⟨{ to_fun := λ a, cases_on' a x (coe ∘ e), + inv_fun := λ b, if h : b = x then none else e.symm ⟨b, h⟩, + left_inv := λ a, begin + cases a, { simp }, + simp only [cases_on'_some, function.comp_app, subtype.coe_eta, symm_apply_apply, + dite_eq_ite], + exact if_neg (e a).property + end, + right_inv := λ b, begin + by_cases h : b = x; + simp [h] + end}, + rfl⟩, + left_inv := λ e, begin + ext a, + cases a, + { simpa using e.property.symm }, + { simpa } + end, + right_inv := λ e, begin + ext a, + refl + end } + +@[simp] lemma option_subtype_apply_apply [decidable_eq β] (x : β) + (e : {e : option α ≃ β // e none = x}) (a : α) (h) : + option_subtype x e a = ⟨(e : option α ≃ β) a, h⟩ := +rfl + +@[simp] lemma coe_option_subtype_apply_apply [decidable_eq β] (x : β) + (e : {e : option α ≃ β // e none = x}) (a : α) : + ↑(option_subtype x e a) = (e : option α ≃ β) a := +rfl + +@[simp] lemma option_subtype_apply_symm_apply [decidable_eq β] (x : β) + (e : {e : option α ≃ β // e none = x}) (b : {y : β // y ≠ x}) : + ↑((option_subtype x e).symm b) = (e : option α ≃ β).symm b := +begin + dsimp only [option_subtype], + simp +end + +@[simp] lemma option_subtype_symm_apply_apply_coe [decidable_eq β] (x : β) + (e : α ≃ {y : β // y ≠ x}) (a : α) : (option_subtype x).symm e a = e a := +rfl + +@[simp] lemma option_subtype_symm_apply_apply_some [decidable_eq β] (x : β) + (e : α ≃ {y : β // y ≠ x}) (a : α) : (option_subtype x).symm e (some a) = e a := +rfl + +@[simp] lemma option_subtype_symm_apply_apply_none [decidable_eq β] (x : β) + (e : α ≃ {y : β // y ≠ x}) : (option_subtype x).symm e none = x := +rfl + +@[simp] lemma option_subtype_symm_apply_symm_apply [decidable_eq β] (x : β) + (e : α ≃ {y : β // y ≠ x}) (b : {y : β // y ≠ x}) : + ((option_subtype x).symm e : option α ≃ β).symm b = e.symm b := +begin + simp only [option_subtype, coe_fn_symm_mk, subtype.coe_mk, subtype.coe_eta, dite_eq_ite, + ite_eq_right_iff], + exact λ h, false.elim (b.property h), +end + end equiv From 0719076497ea04e9c594c34d9cc3779470af65eb Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 31 Aug 2022 11:42:44 +0000 Subject: [PATCH 109/828] chore(topology/instances/rat): add `namespace rat` (#16306) --- src/topology/instances/rat.lean | 20 +++++++++++--------- 1 file changed, 11 insertions(+), 9 deletions(-) diff --git a/src/topology/instances/rat.lean b/src/topology/instances/rat.lean index a1b57c42926f5..21933125a4015 100644 --- a/src/topology/instances/rat.lean +++ b/src/topology/instances/rat.lean @@ -58,15 +58,17 @@ uniform_embedding_bot_of_pairwise_le_dist zero_lt_one $ by simpa using int.pairw lemma int.closed_embedding_coe_rat : closed_embedding (coe : ℤ → ℚ) := closed_embedding_of_pairwise_le_dist zero_lt_one $ by simpa using int.pairwise_one_le_dist +namespace rat + instance : noncompact_space ℚ := int.closed_embedding_coe_rat.noncompact_space -- TODO(Mario): Find a way to use rat_add_continuous_lemma -theorem rat.uniform_continuous_add : uniform_continuous (λp : ℚ × ℚ, p.1 + p.2) := +theorem uniform_continuous_add : uniform_continuous (λp : ℚ × ℚ, p.1 + p.2) := rat.uniform_embedding_coe_real.to_uniform_inducing.uniform_continuous_iff.2 $ by simp only [(∘), rat.cast_add]; exact real.uniform_continuous_add.comp (rat.uniform_continuous_coe_real.prod_map rat.uniform_continuous_coe_real) -theorem rat.uniform_continuous_neg : uniform_continuous (@has_neg.neg ℚ _) := +theorem uniform_continuous_neg : uniform_continuous (@has_neg.neg ℚ _) := metric.uniform_continuous_iff.2 $ λ ε ε0, ⟨_, ε0, λ a b h, by rw dist_comm at h; simpa [rat.dist_eq] using h⟩ @@ -78,20 +80,20 @@ instance : topological_add_group ℚ := by apply_instance instance : order_topology ℚ := induced_order_topology _ (λ x y, rat.cast_lt) (@exists_rat_btwn _ _ _) -lemma rat.uniform_continuous_abs : uniform_continuous (abs : ℚ → ℚ) := +lemma uniform_continuous_abs : uniform_continuous (abs : ℚ → ℚ) := metric.uniform_continuous_iff.2 $ λ ε ε0, ⟨ε, ε0, λ a b h, lt_of_le_of_lt (by simpa [rat.dist_eq] using abs_abs_sub_abs_le_abs_sub _ _) h⟩ -lemma rat.continuous_mul : continuous (λp : ℚ × ℚ, p.1 * p.2) := +lemma continuous_mul : continuous (λp : ℚ × ℚ, p.1 * p.2) := rat.embedding_coe_real.continuous_iff.2 $ by simp [(∘)]; exact real.continuous_mul.comp ((rat.continuous_coe_real.prod_map rat.continuous_coe_real)) instance : topological_ring ℚ := { continuous_mul := rat.continuous_mul, ..rat.topological_add_group } -lemma rat.totally_bounded_Icc (a b : ℚ) : totally_bounded (Icc a b) := -begin - have := totally_bounded_preimage rat.uniform_embedding_coe_real (totally_bounded_Icc a b), - rwa (set.ext (λ q, _) : Icc _ _ = _), simp -end +lemma totally_bounded_Icc (a b : ℚ) : totally_bounded (Icc a b) := +by simpa only [preimage_cast_Icc] + using totally_bounded_preimage rat.uniform_embedding_coe_real (totally_bounded_Icc a b) + +end rat From 865184b48979b41ccba15dc5b39d7133db03b1fe Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 31 Aug 2022 11:42:45 +0000 Subject: [PATCH 110/828] feat(linear_algebra/affine_space/finite_dimensional): `insert` lemmas (#16308) Add various lemmas about adding a point to a finite-dimensional affine subspace and the dimension of the span of the result. --- .../affine_space/finite_dimensional.lean | 91 +++++++++++++++++++ 1 file changed, 91 insertions(+) diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index ec31a1f70fd5f..d793821c4ce5f 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -265,8 +265,42 @@ begin exact hi.affine_span_eq_of_le_of_card_eq_finrank_add_one le_top hc, }, end +/-- The `vector_span` of adding a point to a finite-dimensional subspace is finite-dimensional. -/ +lemma finite_dimensional_vector_span_insert (s : affine_subspace k P) + [finite_dimensional k s.direction] (p : P) : + finite_dimensional k (vector_span k (insert p (s : set P))) := +begin + rw [←direction_affine_span, ←affine_span_insert_affine_span], + rcases (s : set P).eq_empty_or_nonempty with hs | ⟨p₀, hp₀⟩, + { rw coe_eq_bot_iff at hs, + rw [hs, bot_coe, span_empty, bot_coe, direction_affine_span], + convert finite_dimensional_bot _ _; + simp }, + { rw [affine_span_coe, direction_affine_span_insert hp₀], + apply_instance } +end + +/-- The direction of the affine span of adding a point to a finite-dimensional subspace is +finite-dimensional. -/ +lemma finite_dimensional_direction_affine_span_insert (s : affine_subspace k P) + [finite_dimensional k s.direction] (p : P) : + finite_dimensional k (affine_span k (insert p (s : set P))).direction := +(direction_affine_span k (insert p (s : set P))).symm ▸ finite_dimensional_vector_span_insert s p + variables (k) +/-- The `vector_span` of adding a point to a set with a finite-dimensional `vector_span` is +finite-dimensional. -/ +lemma finite_dimensional_vector_span_insert_set (s : set P) + [finite_dimensional k (vector_span k s)] (p : P) : + finite_dimensional k (vector_span k (insert p s)) := +begin + haveI : finite_dimensional k (affine_span k s).direction := + (direction_affine_span k s).symm ▸ infer_instance, + rw [←direction_affine_span, ←affine_span_insert_affine_span, direction_affine_span], + exact finite_dimensional_vector_span_insert (affine_span k s) p +end + /-- A set of points is collinear if their `vector_span` has dimension at most `1`. -/ def collinear (s : set P) : Prop := module.rank k (vector_span k s) ≤ 1 @@ -384,3 +418,60 @@ by rw [collinear_iff_finrank_le_one, finrank_vector_span_le_iff_not_affine_independent k p (fintype.card_fin 3)] end affine_space' + +section field + +variables {k : Type*} {V : Type*} {P : Type*} +include V + +open affine_subspace finite_dimensional module + +variables [field k] [add_comm_group V] [module k V] [affine_space V P] + +/-- Adding a point to a finite-dimensional subspace increases the dimension by at most one. -/ +lemma finrank_vector_span_insert_le (s : affine_subspace k P) (p : P) : + finrank k (vector_span k (insert p (s : set P))) ≤ finrank k s.direction + 1 := +begin + by_cases hf : finite_dimensional k s.direction, swap, + { have hf' : ¬finite_dimensional k (vector_span k (insert p (s : set P))), + { intro h, + have h' : s.direction ≤ vector_span k (insert p (s : set P)), + { conv_lhs { rw [←affine_span_coe s, direction_affine_span] }, + exact vector_span_mono k (set.subset_insert _ _) }, + exactI hf (submodule.finite_dimensional_of_le h') }, + rw [finrank_of_infinite_dimensional hf, finrank_of_infinite_dimensional hf', zero_add], + exact zero_le_one }, + haveI := hf, + rw [←direction_affine_span, ←affine_span_insert_affine_span], + rcases (s : set P).eq_empty_or_nonempty with hs | ⟨p₀, hp₀⟩, + { rw coe_eq_bot_iff at hs, + rw [hs, bot_coe, span_empty, bot_coe, direction_affine_span, direction_bot, finrank_bot, + zero_add], + convert zero_le_one' ℕ, + rw ←finrank_bot k V, + convert rfl; + simp }, + { rw [affine_span_coe, direction_affine_span_insert hp₀, add_comm], + refine (submodule.dim_add_le_dim_add_dim _ _).trans (add_le_add_right _ _), + refine finrank_le_one ⟨p -ᵥ p₀, submodule.mem_span_singleton_self _⟩ (λ v, _), + have h := v.property, + rw submodule.mem_span_singleton at h, + rcases h with ⟨c, hc⟩, + refine ⟨c, _⟩, + ext, + exact hc } +end + +variables (k) + +/-- Adding a point to a set with a finite-dimensional span increases the dimension by at most +one. -/ +lemma finrank_vector_span_insert_le_set (s : set P) (p : P) : + finrank k (vector_span k (insert p s)) ≤ finrank k (vector_span k s) + 1 := +begin + rw [←direction_affine_span, ←affine_span_insert_affine_span, direction_affine_span], + refine (finrank_vector_span_insert_le _ _).trans (add_le_add_right _ _), + rw direction_affine_span +end + +end field From 361aa777b4d262212c31d7c4a245ccb23645c156 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 31 Aug 2022 14:18:55 +0000 Subject: [PATCH 111/828] feat(category_theory/adjunction): complete well-powered category with a coseparating set is cocomplete (#15988) This corollary of the special adjoint functor theorem immediately implies that Grothendieck categories are complete, which, according to the Wikipedia article on Grothendieck categories, is "a rather deep result". --- .../adjunction/adjoint_functor_theorems.lean | 29 +++++++++++++++---- src/category_theory/generator.lean | 6 ++-- 2 files changed, 27 insertions(+), 8 deletions(-) diff --git a/src/category_theory/adjunction/adjoint_functor_theorems.lean b/src/category_theory/adjunction/adjoint_functor_theorems.lean index ac21e8514d3bb..b13c7016fb27c 100644 --- a/src/category_theory/adjunction/adjoint_functor_theorems.lean +++ b/src/category_theory/adjunction/adjoint_functor_theorems.lean @@ -3,14 +3,10 @@ Copyright (c) 2021 Bhavik Mehta. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bhavik Mehta -/ -import category_theory.adjunction.basic -import category_theory.adjunction.comma import category_theory.generator +import category_theory.limits.cone_category import category_theory.limits.constructions.weakly_initial -import category_theory.limits.preserves.basic -import category_theory.limits.creates -import category_theory.limits.comma -import category_theory.punit +import category_theory.limits.functor_category import category_theory.subobject.comma /-! @@ -32,6 +28,9 @@ This file also proves the special adjoint functor theorem, in the form: * If `G : D ⥤ C` preserves limits and `D` is complete, well-powered and has a small coseparating set, then `G` has a left adjoint: `is_right_adjoint_of_preserves_limits_of_is_coseparating` +Finally, we prove the following corollary of the special adjoint functor theorem: +* If `C` is complete, well-powered and has a small coseparating set, then it is cocomplete: + `has_colimits_of_has_limits_of_is_coseparating` -/ universes v u u' @@ -121,4 +120,22 @@ by exactI is_left_adjoint_of_costructured_arrow_terminals _ end special_adjoint_functor_theorem +namespace limits + +/-- A consequence of the special adjoint functor theorem: if `C` is complete, well-powered and + has a small coseparating set, then it is cocomplete. -/ +lemma has_colimits_of_has_limits_of_is_coseparating [has_limits C] [well_powered C] + {𝒢 : set C} [small.{v} 𝒢] (h𝒢 : is_coseparating 𝒢) : has_colimits C := +{ has_colimits_of_shape := λ J hJ, by exactI has_colimits_of_shape_iff_is_right_adjoint_const.2 + ⟨is_right_adjoint_of_preserves_limits_of_is_coseparating h𝒢 _⟩ } + +/-- A consequence of the special adjoint functor theorem: if `C` is cocomplete, well-copowered and + has a small separating set, then it is complete. -/ +lemma has_limits_of_has_colimits_of_is_separating [has_colimits C] [well_powered Cᵒᵖ] + {𝒢 : set C} [small.{v} 𝒢] (h𝒢 : is_separating 𝒢) : has_limits C := +{ has_limits_of_shape := λ J hJ, by exactI has_limits_of_shape_iff_is_left_adjoint_const.2 + ⟨is_left_adjoint_of_preserves_colimits_of_is_separatig h𝒢 _⟩ } + +end limits + end category_theory diff --git a/src/category_theory/generator.lean b/src/category_theory/generator.lean index aed1af711e8f1..23fa91a45b432 100644 --- a/src/category_theory/generator.lean +++ b/src/category_theory/generator.lean @@ -263,7 +263,8 @@ end /-- An ingredient of the proof of the Special Adjoint Functor Theorem: a complete well-powered category with a small coseparating set has an initial object. - In fact, it follows from the Special Adjoint Functor Theorem that `C` is already cocomplete. -/ + In fact, it follows from the Special Adjoint Functor Theorem that `C` is already cocomplete, + see `has_colimits_of_has_limits_of_is_coseparating`. -/ lemma has_initial_of_is_coseparating [well_powered C] [has_limits C] {𝒢 : set C} [small.{v₁} 𝒢] (h𝒢 : is_coseparating 𝒢) : has_initial C := begin @@ -287,7 +288,8 @@ end /-- An ingredient of the proof of the Special Adjoint Functor Theorem: a cocomplete well-copowered category with a small separating set has a terminal object. - In fact, it follows from the Special Adjoint Functor Theorem that `C` is already complete. -/ + In fact, it follows from the Special Adjoint Functor Theorem that `C` is already complete, see + `has_limits_of_has_colimits_of_is_separating`. -/ lemma has_terminal_of_is_separating [well_powered Cᵒᵖ] [has_colimits C] {𝒢 : set C} [small.{v₁} 𝒢] (h𝒢 : is_separating 𝒢) : has_terminal C := begin From 27d83241f983b5eb73b4a8be9107c825e6db6770 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 31 Aug 2022 14:18:56 +0000 Subject: [PATCH 112/828] feat(topology/sheaves/sheaf_condition): Description of sections on the union of two open sets. (#16093) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/algebra/category/Ring/constructions.lean | 31 +++++++++++ .../pairwise_intersections.lean | 55 +++++++++++++++---- .../sheaf_condition/unique_gluing.lean | 15 +++++ 3 files changed, 89 insertions(+), 12 deletions(-) diff --git a/src/algebra/category/Ring/constructions.lean b/src/algebra/category/Ring/constructions.lean index d577bca1f0632..784d5c3326f39 100644 --- a/src/algebra/category/Ring/constructions.lean +++ b/src/algebra/category/Ring/constructions.lean @@ -212,4 +212,35 @@ end end equalizer +section pullback + +/-- +In the category of `CommRing`, the pullback of `f : A ⟶ C` and `g : B ⟶ C` is the `eq_locus` of +the two maps `A × B ⟶ C`. This is the constructed pullback cone. +-/ +def pullback_cone {A B C : CommRing.{u}} (f : A ⟶ C) (g : B ⟶ C) : pullback_cone f g := +pullback_cone.mk + (CommRing.of_hom $ (ring_hom.fst A B).comp + (ring_hom.eq_locus (f.comp (ring_hom.fst A B)) (g.comp (ring_hom.snd A B))).subtype) + (CommRing.of_hom $ (ring_hom.snd A B).comp + (ring_hom.eq_locus (f.comp (ring_hom.fst A B)) (g.comp (ring_hom.snd A B))).subtype) + (by { ext ⟨x, e⟩, simpa [CommRing.of_hom] using e }) + +/-- The constructed pullback cone is indeed the limit. -/ +def pullback_cone_is_limit {A B C : CommRing.{u}} (f : A ⟶ C) (g : B ⟶ C) : + is_limit (pullback_cone f g) := +begin + fapply pullback_cone.is_limit.mk, + { intro s, + apply (s.fst.prod s.snd).cod_restrict, + intro x, exact congr_arg (λ f : s.X →+* C, f x) s.condition }, + { intro s, ext x, refl }, + { intro s, ext x, refl }, + { intros s m e₁ e₂, ext, + { exact (congr_arg (λ f : s.X →+* A, f x) e₁ : _) }, + { exact (congr_arg (λ f : s.X →+* B, f x) e₂ : _) } } +end + +end pullback + end CommRing diff --git a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean index c9f8c083ec7e5..75a18fd7b8c1f 100644 --- a/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean +++ b/src/topology/sheaves/sheaf_condition/pairwise_intersections.lean @@ -8,6 +8,7 @@ import topology.sheaves.sheaf_condition.sites import category_theory.limits.preserves.basic import category_theory.category.pairwise import category_theory.limits.constructions.binary_products +import algebra.category.Ring.constructions /-! # Equivalent formulations of the sheaf condition @@ -401,33 +402,33 @@ variables {X : Top.{v}} {C : Type u} [category.{v} C] variables (F : X.sheaf C) (U V : opens X) open category_theory.limits -/-- For a sheaf `F`, `F(U ∪ V)` is the pullback of `F(U) ⟶ F(U ∩ V)` and `F(V) ⟶ F(U ∩ V)`. +/-- For a sheaf `F`, `F(U ⊔ V)` is the pullback of `F(U) ⟶ F(U ⊓ V)` and `F(V) ⟶ F(U ⊓ V)`. This is the pullback cone. -/ def inter_union_pullback_cone : pullback_cone - (F.1.map (hom_of_le inf_le_left : U ∩ V ⟶ _).op) (F.1.map (hom_of_le inf_le_right).op) := + (F.1.map (hom_of_le inf_le_left : U ⊓ V ⟶ _).op) (F.1.map (hom_of_le inf_le_right).op) := pullback_cone.mk (F.1.map (hom_of_le le_sup_left).op) (F.1.map (hom_of_le le_sup_right).op) (by { rw [← F.1.map_comp, ← F.1.map_comp], congr }) @[simp] lemma inter_union_pullback_cone_X : - (inter_union_pullback_cone F U V).X = F.1.obj (op $ U ∪ V) := rfl + (inter_union_pullback_cone F U V).X = F.1.obj (op $ U ⊔ V) := rfl @[simp] lemma inter_union_pullback_cone_fst : (inter_union_pullback_cone F U V).fst = F.1.map (hom_of_le le_sup_left).op := rfl @[simp] lemma inter_union_pullback_cone_snd : (inter_union_pullback_cone F U V).snd = F.1.map (hom_of_le le_sup_right).op := rfl variable (s : pullback_cone - (F.1.map (hom_of_le inf_le_left : U ∩ V ⟶ _).op) (F.1.map (hom_of_le inf_le_right).op)) + (F.1.map (hom_of_le inf_le_left : U ⊓ V ⟶ _).op) (F.1.map (hom_of_le inf_le_right).op)) variable [has_products.{v} C] /-- (Implementation). -Every cone over `F(U) ⟶ F(U ∩ V)` and `F(V) ⟶ F(U ∩ V)` factors through `F(U ∪ V)`. +Every cone over `F(U) ⟶ F(U ⊓ V)` and `F(V) ⟶ F(U ⊓ V)` factors through `F(U ⊔ V)`. TODO: generalize to `C` without products. -/ -def inter_union_pullback_cone_lift : s.X ⟶ F.1.obj (op (U ∪ V)) := +def inter_union_pullback_cone_lift : s.X ⟶ F.1.obj (op (U ⊔ V)) := begin let ι : ulift.{v} walking_pair → opens X := λ j, walking_pair.cases_on j.down U V, - have hι : U ∪ V = supr ι, + have hι : U ⊔ V = supr ι, { ext, rw [opens.coe_supr, set.mem_Union], split, @@ -447,7 +448,7 @@ begin let g : j ⟶ i := f.unop, have : f = g.op := rfl, clear_value g, subst this, rcases i with (⟨⟨(_|_)⟩⟩|⟨⟨(_|_)⟩,⟨_⟩⟩); rcases j with (⟨⟨(_|_)⟩⟩|⟨⟨(_|_)⟩,⟨_⟩⟩); rcases g; dsimp; simp only [category.id_comp, s.condition, category_theory.functor.map_id, category.comp_id], - { rw [← cancel_mono (F.1.map (eq_to_hom $ inf_comm : U ∩ V ⟶ _).op), category.assoc, + { rw [← cancel_mono (F.1.map (eq_to_hom $ inf_comm : U ⊓ V ⟶ _).op), category.assoc, category.assoc], erw [← F.1.map_comp, ← F.1.map_comp], convert s.condition.symm }, @@ -470,11 +471,11 @@ begin (op $ pairwise.single (ulift.up walking_pair.right)) end -/-- For a sheaf `F`, `F(U ∪ V)` is the pullback of `F(U) ⟶ F(U ∩ V)` and `F(V) ⟶ F(U ∩ V)`. -/ +/-- For a sheaf `F`, `F(U ⊔ V)` is the pullback of `F(U) ⟶ F(U ⊓ V)` and `F(V) ⟶ F(U ⊓ V)`. -/ def is_limit_pullback_cone : is_limit (inter_union_pullback_cone F U V) := begin let ι : ulift.{v} walking_pair → opens X := λ ⟨j⟩, walking_pair.cases_on j U V, - have hι : U ∪ V = supr ι, + have hι : U ⊔ V = supr ι, { ext, rw [opens.coe_supr, set.mem_Union], split, @@ -511,11 +512,41 @@ begin apply inter_union_pullback_cone_lift_right } } end -/-- If `U, V` are disjoint, then `F(U ∪ V) = F(U) × F(V)`. -/ -def is_product_of_disjoint (h : U ∩ V = ⊥) : is_limit +/-- If `U, V` are disjoint, then `F(U ⊔ V) = F(U) × F(V)`. -/ +def is_product_of_disjoint (h : U ⊓ V = ⊥) : is_limit (binary_fan.mk (F.1.map (hom_of_le le_sup_left : _ ⟶ U ⊔ V).op) (F.1.map (hom_of_le le_sup_right : _ ⟶ U ⊔ V).op)) := is_product_of_is_terminal_is_pullback _ _ _ _ (F.is_terminal_of_eq_empty h) (is_limit_pullback_cone F U V) +/-- `F(U ⊔ V)` is isomorphic to the `eq_locus` of the two maps `F(U) × F(V) ⟶ F(U ⊓ V)`. -/ +def obj_sup_iso_prod_eq_locus {X : Top} (F : X.sheaf CommRing) + (U V : opens X) : + F.1.obj (op $ U ⊔ V) ≅ CommRing.of (ring_hom.eq_locus _ _) := +(F.is_limit_pullback_cone U V).cone_point_unique_up_to_iso (CommRing.pullback_cone_is_limit _ _) + +lemma obj_sup_iso_prod_eq_locus_hom_fst {X : Top} (F : X.sheaf CommRing) + (U V : opens X) (x) : + ((F.obj_sup_iso_prod_eq_locus U V).hom x).1.fst = F.1.map (hom_of_le le_sup_left).op x := +concrete_category.congr_hom ((F.is_limit_pullback_cone U V).cone_point_unique_up_to_iso_hom_comp + (CommRing.pullback_cone_is_limit _ _) walking_cospan.left) x + +lemma obj_sup_iso_prod_eq_locus_hom_snd {X : Top} (F : X.sheaf CommRing) + (U V : opens X) (x) : + ((F.obj_sup_iso_prod_eq_locus U V).hom x).1.snd = F.1.map (hom_of_le le_sup_right).op x := +concrete_category.congr_hom ((F.is_limit_pullback_cone U V).cone_point_unique_up_to_iso_hom_comp + (CommRing.pullback_cone_is_limit _ _) walking_cospan.right) x + +lemma obj_sup_iso_prod_eq_locus_inv_fst {X : Top} (F : X.sheaf CommRing) + (U V : opens X) (x) : + F.1.map (hom_of_le le_sup_left).op ((F.obj_sup_iso_prod_eq_locus U V).inv x) = x.1.1 := +concrete_category.congr_hom ((F.is_limit_pullback_cone U V).cone_point_unique_up_to_iso_inv_comp + (CommRing.pullback_cone_is_limit _ _) walking_cospan.left) x + +lemma obj_sup_iso_prod_eq_locus_inv_snd {X : Top} (F : X.sheaf CommRing) + (U V : opens X) (x) : + F.1.map (hom_of_le le_sup_right).op ((F.obj_sup_iso_prod_eq_locus U V).inv x) = x.1.2 := +concrete_category.congr_hom ((F.is_limit_pullback_cone U V).cone_point_unique_up_to_iso_inv_comp + (CommRing.pullback_cone_is_limit _ _) walking_cospan.right) x + end Top.sheaf diff --git a/src/topology/sheaves/sheaf_condition/unique_gluing.lean b/src/topology/sheaves/sheaf_condition/unique_gluing.lean index 7e9b68806d3b1..2a57389de3247 100644 --- a/src/topology/sheaves/sheaf_condition/unique_gluing.lean +++ b/src/topology/sheaves/sheaf_condition/unique_gluing.lean @@ -301,6 +301,21 @@ begin convert h i, end +lemma eq_of_locally_eq₂ {U₁ U₂ V : opens X} + (i₁ : U₁ ⟶ V) (i₂ : U₂ ⟶ V) (hcover : V ≤ U₁ ⊔ U₂) + (s t : F.1.obj (op V)) + (h₁ : F.1.map i₁.op s = F.1.map i₁.op t) + (h₂ : F.1.map i₂.op s = F.1.map i₂.op t) : s = t := +begin + classical, + fapply F.eq_of_locally_eq' (λ t : ulift bool, if t.1 then U₁ else U₂), + { exact λ i, if h : i.1 then (eq_to_hom (if_pos h)) ≫ i₁ else (eq_to_hom (if_neg h)) ≫ i₂ }, + { refine le_trans hcover _, rw sup_le_iff, split, + { convert le_supr (λ t : ulift bool, if t.1 then U₁ else U₂) (ulift.up true) }, + { convert le_supr (λ t : ulift bool, if t.1 then U₁ else U₂) (ulift.up false) } }, + { rintro ⟨_|_⟩; simp [h₁, h₂] } +end + end end sheaf From 76be8c7392c7bad3ce467a7295561a30e8b42f00 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Wed, 31 Aug 2022 14:18:57 +0000 Subject: [PATCH 113/828] chore(algebra/order/absolute_value): superficial tidying (#16190) This weakens the requirements of many lemmas in this file, preparing for further changes. `is_absolute_value` is unbundled, which is an approach that is currently not favoured in mathlib. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/algebra/order/absolute_value.lean | 193 +++++++++++++++++--------- 1 file changed, 127 insertions(+), 66 deletions(-) diff --git a/src/algebra/order/absolute_value.lean b/src/algebra/order/absolute_value.lean index acd90dc3626c6..def9cda1ba8af 100644 --- a/src/algebra/order/absolute_value.lean +++ b/src/algebra/order/absolute_value.lean @@ -35,6 +35,8 @@ initialize_simps_projections absolute_value (to_mul_hom_to_fun → apply) section ordered_semiring +section semiring + variables {R S : Type*} [semiring R] [ordered_semiring S] (abv : absolute_value R S) instance mul_hom_class : mul_hom_class (absolute_value R S) R S := @@ -59,45 +61,40 @@ lt_of_le_of_ne (abv.nonneg x) (ne.symm $ mt abv.eq_zero.mp hx) protected theorem ne_zero {x : R} (hx : x ≠ 0) : abv x ≠ 0 := (abv.pos hx).ne' +theorem map_one_of_is_regular (h : is_left_regular (abv 1)) : abv 1 = 1 := +h $ by simp [←abv.map_mul] + @[simp] protected theorem map_zero : abv 0 = 0 := abv.eq_zero.2 rfl -end ordered_semiring +end semiring -section ordered_ring +section ring -variables {R S : Type*} [ring R] [ordered_ring S] (abv : absolute_value R S) +variables {R S : Type*} [ring R] [ordered_semiring S] (abv : absolute_value R S) protected lemma sub_le (a b c : R) : abv (a - c) ≤ abv (a - b) + abv (b - c) := by simpa [sub_eq_add_neg, add_assoc] using abv.add_le (a - b) (b - c) -protected lemma le_sub (a b : R) : abv a - abv b ≤ abv (a - b) := -sub_le_iff_le_add.2 $ by simpa using abv.add_le (a - b) b - @[simp] lemma map_sub_eq_zero_iff (a b : R) : abv (a - b) = 0 ↔ a = b := abv.eq_zero.trans sub_eq_zero -end ordered_ring +end ring -section linear_ordered_ring +end ordered_semiring -variables {R S : Type*} [semiring R] [linear_ordered_ring S] (abv : absolute_value R S) +section ordered_ring -/-- `absolute_value.abs` is `abs` as a bundled `absolute_value`. -/ -@[simps] -protected def abs : absolute_value S S := -{ to_fun := abs, - nonneg' := abs_nonneg, - eq_zero' := λ _, abs_eq_zero, - add_le' := abs_add, - map_mul' := abs_mul } +section semiring -instance : inhabited (absolute_value S S) := ⟨absolute_value.abs⟩ +section is_domain -variables [nontrivial R] +-- all of these are true for `no_zero_divisors S`; but it doesn't work smoothly with the +-- `is_domain`/`cancel_monoid_with_zero` API +variables {R S : Type*} [semiring R] [ordered_ring S] (abv : absolute_value R S) +variables [is_domain S] [nontrivial R] @[simp] protected theorem map_one : abv 1 = 1 := -(mul_right_inj' $ abv.ne_zero one_ne_zero).1 $ -by rw [← abv.map_mul, mul_one, mul_one] +abv.map_one_of_is_regular ((is_regular_of_ne_zero $ abv.ne_zero one_ne_zero).left) instance : monoid_with_zero_hom_class (absolute_value R S) R S := { map_zero := λ f, f.map_zero, @@ -117,13 +114,25 @@ def to_monoid_hom : R →* S := abv @[simp] protected lemma map_pow (a : R) (n : ℕ) : abv (a ^ n) = abv a ^ n := abv.to_monoid_hom.map_pow a n -end linear_ordered_ring +end is_domain -section linear_ordered_comm_ring +end semiring section ring -variables {R S : Type*} [ring R] [linear_ordered_comm_ring S] (abv : absolute_value R S) +variables {R S : Type*} [ring R] [ordered_ring S] (abv : absolute_value R S) + +protected lemma le_sub (a b : R) : abv a - abv b ≤ abv (a - b) := +sub_le_iff_le_add.2 $ by simpa using abv.add_le (a - b) b + +end ring + +end ordered_ring + +section ordered_comm_ring + +variables {R S : Type*} [ring R] [ordered_comm_ring S] (abv : absolute_value R S) +variables [no_zero_divisors S] @[simp] protected theorem map_neg (a : R) : abv (-a) = abv a := begin @@ -136,18 +145,37 @@ end protected theorem map_sub (a b : R) : abv (a - b) = abv (b - a) := by rw [← neg_sub, abv.map_neg] +end ordered_comm_ring + +section linear_ordered_ring + +variables {R S : Type*} [semiring R] [linear_ordered_ring S] (abv : absolute_value R S) + +/-- `absolute_value.abs` is `abs` as a bundled `absolute_value`. -/ +@[simps] +protected def abs : absolute_value S S := +{ to_fun := abs, + nonneg' := abs_nonneg, + eq_zero' := λ _, abs_eq_zero, + add_le' := abs_add, + map_mul' := abs_mul } + +instance : inhabited (absolute_value S S) := ⟨absolute_value.abs⟩ + +end linear_ordered_ring + +section linear_ordered_comm_ring + +variables {R S : Type*} [ring R] [linear_ordered_comm_ring S] (abv : absolute_value R S) + lemma abs_abv_sub_le_abv_sub (a b : R) : abs (abv a - abv b) ≤ abv (a - b) := abs_sub_le_iff.2 ⟨abv.le_sub _ _, by rw abv.map_sub; apply abv.le_sub⟩ -end ring - end linear_ordered_comm_ring end absolute_value -section is_absolute_value - /-- A function `f` is an absolute value if it is nonnegative, zero only at 0, additive, and multiplicative. @@ -168,7 +196,7 @@ variables {S : Type*} [ordered_semiring S] variables {R : Type*} [semiring R] (abv : R → S) [is_absolute_value abv] /-- A bundled absolute value is an absolute value. -/ -instance absolute_value.is_absolute_value +instance _root_.absolute_value.is_absolute_value (abv : absolute_value R S) : is_absolute_value abv := { abv_nonneg := abv.nonneg, abv_eq_zero := λ _, abv.eq_zero, @@ -184,85 +212,118 @@ def to_absolute_value : absolute_value R S := nonneg' := abv_nonneg abv, map_mul' := abv_mul abv } -theorem abv_zero : abv 0 = 0 := (abv_eq_zero abv).2 rfl - -theorem abv_pos {a : R} : 0 < abv a ↔ a ≠ 0 := -by rw [lt_iff_le_and_ne, ne, eq_comm]; simp [abv_eq_zero abv, abv_nonneg abv] +theorem abv_zero : abv 0 = 0 := (to_absolute_value abv).map_zero +theorem abv_pos {a : R} : 0 < abv a ↔ a ≠ 0 := (to_absolute_value abv).pos_iff end ordered_semiring section linear_ordered_ring variables {S : Type*} [linear_ordered_ring S] -variables {R : Type*} [semiring R] (abv : R → S) [is_absolute_value abv] -instance abs_is_absolute_value {S} [linear_ordered_ring S] : - is_absolute_value (abs : S → S) := -{ abv_nonneg := abs_nonneg, - abv_eq_zero := λ _, abs_eq_zero, - abv_add := abs_add, - abv_mul := abs_mul } +instance abs_is_absolute_value : is_absolute_value (abs : S → S) := + absolute_value.abs.is_absolute_value end linear_ordered_ring -section linear_ordered_comm_ring +section ordered_ring -variables {S : Type*} [linear_ordered_comm_ring S] +variables {S : Type*} [ordered_ring S] section semiring + variables {R : Type*} [semiring R] (abv : R → S) [is_absolute_value abv] -theorem abv_one [nontrivial R] : abv 1 = 1 := -(mul_right_inj' $ mt (abv_eq_zero abv).1 one_ne_zero).1 $ -by rw [← abv_mul abv, mul_one, mul_one] +variables [is_domain S] + +theorem abv_one [nontrivial R] : abv 1 = 1 := (to_absolute_value abv).map_one /-- `abv` as a `monoid_with_zero_hom`. -/ -def abv_hom [nontrivial R] : R →*₀ S := ⟨abv, abv_zero abv, abv_one abv, abv_mul abv⟩ +def abv_hom [nontrivial R] : R →*₀ S := (to_absolute_value abv).to_monoid_with_zero_hom lemma abv_pow [nontrivial R] (abv : R → S) [is_absolute_value abv] (a : R) (n : ℕ) : abv (a ^ n) = abv a ^ n := -(abv_hom abv).to_monoid_hom.map_pow a n +(to_absolute_value abv).map_pow a n end semiring -end linear_ordered_comm_ring +section ring -section linear_ordered_field -variables {S : Type*} [linear_ordered_field S] +variables {R : Type*} [ring R] (abv : R → S) [is_absolute_value abv] + +lemma abv_sub_le (a b c : R) : abv (a - c) ≤ abv (a - b) + abv (b - c) := +by simpa [sub_eq_add_neg, add_assoc] using abv_add abv (a - b) (b - c) + +lemma sub_abv_le_abv_sub (a b : R) : abv a - abv b ≤ abv (a - b) := +(to_absolute_value abv).le_sub a b + +end ring + +end ordered_ring + +section ordered_comm_ring + +variables {S : Type*} [ordered_comm_ring S] section ring + variables {R : Type*} [ring R] (abv : R → S) [is_absolute_value abv] +variables [no_zero_divisors S] + theorem abv_neg (a : R) : abv (-a) = abv a := -by rw [← mul_self_inj_of_nonneg (abv_nonneg abv _) (abv_nonneg abv _), -← abv_mul abv, ← abv_mul abv]; simp +(to_absolute_value abv).map_neg a theorem abv_sub (a b : R) : abv (a - b) = abv (b - a) := -by rw [← neg_sub, abv_neg abv] +(to_absolute_value abv).map_sub a b -lemma abv_sub_le (a b c : R) : abv (a - c) ≤ abv (a - b) + abv (b - c) := -by simpa [sub_eq_add_neg, add_assoc] using abv_add abv (a - b) (b - c) +end ring -lemma sub_abv_le_abv_sub (a b : R) : abv a - abv b ≤ abv (a - b) := -sub_le_iff_le_add.2 $ by simpa using abv_add abv (a - b) b +end ordered_comm_ring + +section linear_ordered_comm_ring + +variables {S : Type*} [linear_ordered_comm_ring S] + +section ring + +variables {R : Type*} [ring R] (abv : R → S) [is_absolute_value abv] lemma abs_abv_sub_le_abv_sub (a b : R) : abs (abv a - abv b) ≤ abv (a - b) := -abs_sub_le_iff.2 ⟨sub_abv_le_abv_sub abv _ _, - by rw abv_sub abv; apply sub_abv_le_abv_sub abv⟩ +(to_absolute_value abv).abs_abv_sub_le_abv_sub a b + end ring -section field -variables {R : Type*} [division_ring R] (abv : R → S) [is_absolute_value abv] +end linear_ordered_comm_ring -theorem abv_inv (a : R) : abv a⁻¹ = (abv a)⁻¹ := map_inv₀ (abv_hom abv) a -theorem abv_div (a b : R) : abv (a / b) = abv a / abv b := map_div₀ (abv_hom abv) a b +section linear_ordered_field -end field +variables {S : Type*} [linear_ordered_semifield S] -end linear_ordered_field +section semiring -end is_absolute_value +variables {R : Type*} [semiring R] [nontrivial R] (abv : R → S) [is_absolute_value abv] + +lemma abv_one' : abv 1 = 1 := +(to_absolute_value abv).map_one_of_is_regular + $ (is_regular_of_ne_zero $ (to_absolute_value abv).ne_zero one_ne_zero).left + +/-- An absolute value as a monoid with zero homomorphism, assuming the target is a semifield. -/ +def abv_hom' : R →*₀ S := ⟨abv, abv_zero abv, abv_one' abv, abv_mul abv⟩ + +end semiring + +section division_semiring + +variables {R : Type*} [division_semiring R] (abv : R → S) [is_absolute_value abv] + +theorem abv_inv (a : R) : abv a⁻¹ = (abv a)⁻¹ := map_inv₀ (abv_hom' abv) a +theorem abv_div (a b : R) : abv (a / b) = abv a / abv b := map_div₀ (abv_hom' abv) a b + +end division_semiring + +end linear_ordered_field end is_absolute_value From 5e194d481a4c83b11bf65ad3053f1ff68027f124 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 31 Aug 2022 14:18:58 +0000 Subject: [PATCH 114/828] feat(algebra/invertible): `map_inv_of` and some other basic results (#16202) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The titular lemma states that under suitable conditions, `f (⅟r) = ⅟(f r)`. This also provides some lemmas about left inverses, which are motivated primarily by proving `is_unit (algebra_map R (exterior_algebra R M) r) ↔ is_unit r`. --- src/algebra/hom/units.lean | 13 +++++++- src/algebra/invertible.lean | 31 ++++++++++++++++++- .../exterior_algebra/basic.lean | 9 ++++++ 3 files changed, 51 insertions(+), 2 deletions(-) diff --git a/src/algebra/hom/units.lean b/src/algebra/hom/units.lean index de8d130e7ea24..40746ceff609e 100644 --- a/src/algebra/hom/units.lean +++ b/src/algebra/hom/units.lean @@ -125,7 +125,7 @@ units.lift_right f end monoid_hom namespace is_unit -variables {F α M N : Type*} +variables {F G α M N : Type*} section monoid variables [monoid M] [monoid N] @@ -133,6 +133,17 @@ variables [monoid M] [monoid N] @[to_additive] lemma map [monoid_hom_class F M N] (f : F) {x : M} (h : is_unit x) : is_unit (f x) := by rcases h with ⟨y, rfl⟩; exact (units.map (f : M →* N) y).is_unit +@[to_additive] lemma of_left_inverse [monoid_hom_class F M N] [monoid_hom_class G N M] + {f : F} {x : M} (g : G) (hfg : function.left_inverse g f) (h : is_unit (f x)) : + is_unit x := +by simpa only [hfg x] using h.map g + +@[to_additive] lemma _root_.is_unit_map_of_left_inverse + [monoid_hom_class F M N] [monoid_hom_class G N M] + {f : F} {x : M} (g : G) (hfg : function.left_inverse g f) : + is_unit (f x) ↔ is_unit x := +⟨of_left_inverse g hfg, map _⟩ + /-- If a homomorphism `f : M →* N` sends each element to an `is_unit`, then it can be lifted to `f : M →* Nˣ`. See also `units.lift_right` for a computable version. -/ @[to_additive "If a homomorphism `f : M →+ N` sends each element to an `is_add_unit`, then it can be diff --git a/src/algebra/invertible.lean b/src/algebra/invertible.lean index d10805f17ceb1..8a980c994d890 100644 --- a/src/algebra/invertible.lean +++ b/src/algebra/invertible.lean @@ -107,7 +107,8 @@ instance [monoid α] (a : α) : subsingleton (invertible a) := ⟨ λ ⟨b, hba, hab⟩ ⟨c, hca, hac⟩, by { congr, exact left_inv_eq_right_inv hba hac } ⟩ /-- If `r` is invertible and `s = r`, then `s` is invertible. -/ -def invertible.copy [monoid α] {r : α} (hr : invertible r) (s : α) (hs : s = r) : invertible s := +def invertible.copy [mul_one_class α] {r : α} (hr : invertible r) (s : α) (hs : s = r) : + invertible s := { inv_of := ⅟r, inv_of_mul_self := by rw [hs, inv_of_mul_self], mul_inv_of_self := by rw [hs, mul_inv_of_self] } @@ -272,3 +273,31 @@ def invertible.map {R : Type*} {S : Type*} {F : Type*} [mul_one_class R] [mul_on { inv_of := f (⅟r), inv_of_mul_self := by rw [←map_mul, inv_of_mul_self, map_one], mul_inv_of_self := by rw [←map_mul, mul_inv_of_self, map_one] } + +/-- Note that the `invertible (f r)` argument can be satisfied by using `letI := invertible.map f r` +before applying this lemma. -/ +lemma map_inv_of {R : Type*} {S : Type*} {F : Type*} [mul_one_class R] [monoid S] + [monoid_hom_class F R S] (f : F) (r : R) [invertible r] [invertible (f r)] : + f (⅟r) = ⅟(f r) := +by { letI := invertible.map f r, convert rfl } + +/-- A monoid hom with a left-inverse that is also a monoid hom is invertible. + +The inverse is computed as `g (⅟(f r))` -/ +@[simps {attrs := []}] +def invertible.of_left_inverse {R : Type*} {S : Type*} {F G : Type*} + [mul_one_class R] [mul_one_class S] [monoid_hom_class F R S] [monoid_hom_class G S R] + (f : F) (g : G) (r : R) (h : function.left_inverse g f) [invertible (f r)] : + invertible r := +(invertible.map g (f r)).copy _ (h r).symm + +/-- Invertibility on either side of a monoid hom with a left-inverse is equivalent. -/ +@[simps] +def invertible_equiv_of_left_inverse {R : Type*} {S : Type*} {F G : Type*} + [monoid R] [monoid S] [monoid_hom_class F R S] [monoid_hom_class G S R] + (f : F) (g : G) (r : R) (h : function.left_inverse g f) : + invertible (f r) ≃ invertible r := +{ to_fun := λ _, by exactI invertible.of_left_inverse f _ _ h, + inv_fun := λ _, by exactI invertible.map f _, + left_inv := λ x, subsingleton.elim _ _, + right_inv := λ x, subsingleton.elim _ _ } diff --git a/src/linear_algebra/exterior_algebra/basic.lean b/src/linear_algebra/exterior_algebra/basic.lean index 3170456caee94..d5c712c549edb 100644 --- a/src/linear_algebra/exterior_algebra/basic.lean +++ b/src/linear_algebra/exterior_algebra/basic.lean @@ -149,6 +149,15 @@ map_eq_zero_iff (algebra_map _ _) (algebra_map_left_inverse _).injective @[simp] lemma algebra_map_eq_one_iff (x : R) : algebra_map R (exterior_algebra R M) x = 1 ↔ x = 1 := map_eq_one_iff (algebra_map _ _) (algebra_map_left_inverse _).injective +lemma is_unit_algebra_map (r : R) : is_unit (algebra_map R (exterior_algebra R M) r) ↔ is_unit r := +is_unit_map_of_left_inverse _ (algebra_map_left_inverse M) + +/-- Invertibility in the exterior algebra is the same as invertibility of the base ring. -/ +@[simps] +def invertible_algebra_map_equiv (r : R) : + invertible (algebra_map R (exterior_algebra R M) r) ≃ invertible r := +invertible_equiv_of_left_inverse _ _ _ (algebra_map_left_inverse M) + variables {M} /-- The canonical map from `exterior_algebra R M` into `triv_sq_zero_ext R M` that sends From 5a3cd16bb6c108757ce6474f4458cd045fe93264 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 31 Aug 2022 14:19:00 +0000 Subject: [PATCH 115/828] feat(order/filter/small_sets): yet another squeeze theorem (#16286) * Add a very general version of the squeeze theorem. * Use it to golf the proof of the usual squeeze theorem. * Add docstrings. --- src/order/filter/small_sets.lean | 11 +++++++++++ src/topology/algebra/order/basic.lean | 18 ++++++------------ 2 files changed, 17 insertions(+), 12 deletions(-) diff --git a/src/order/filter/small_sets.lean b/src/order/filter/small_sets.lean index 96269fff46113..ddb4e5fd9a6a9 100644 --- a/src/order/filter/small_sets.lean +++ b/src/order/filter/small_sets.lean @@ -106,6 +106,17 @@ begin exact λ u hu, (ht u hu).mp (hst.mono $ λ a hst ht, subset.trans hst ht) end +/-- Generalized **squeeze theorem** (also known as **sandwich theorem**). If `s : α → set β` is a +family of sets that tends to `filter.small_sets lb` along `la` and `f : α → β` is a function such +that `f x ∈ s x` eventually along `la`, then `f` tends to `lb` along `la`. + +If `s x` is the closed interval `[g x, h x]` for some functions `g`, `h` that tend to the same limit +`𝓝 y`, then we obtain the standard squeeze theorem, see +`tendsto_of_tendsto_of_tendsto_of_le_of_le'`. -/ +lemma tendsto.of_small_sets {s : α → set β} {f : α → β} (hs : tendsto s la lb.small_sets) + (hf : ∀ᶠ x in la, f x ∈ s x) : tendsto f la lb := +λ t ht, hf.mp $ (tendsto_small_sets_iff.mp hs t ht).mono $ λ x h₁ h₂, h₁ h₂ + @[simp] lemma eventually_small_sets_eventually {p : α → Prop} : (∀ᶠ s in l.small_sets, ∀ᶠ x in l', x ∈ s → p x) ↔ ∀ᶠ x in l ⊓ l', p x := calc _ ↔ ∃ s ∈ l, ∀ᶠ x in l', x ∈ s → p x : diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index 622487c98d35f..4ae424f91eb66 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -762,22 +762,16 @@ tendsto_Ixx_class_of_subset (λ _ _, Ioc_subset_Icc_self) instance tendsto_Ioo_class_nhds (a : α) : tendsto_Ixx_class Ioo (𝓝 a) (𝓝 a) := tendsto_Ixx_class_of_subset (λ _ _, Ioo_subset_Icc_self) -/-- Also known as squeeze or sandwich theorem. This version assumes that inequalities hold -eventually for the filter. -/ +/-- **Squeeze theorem** (also known as **sandwich theorem**). This version assumes that inequalities +hold eventually for the filter. -/ lemma tendsto_of_tendsto_of_tendsto_of_le_of_le' {f g h : β → α} {b : filter β} {a : α} (hg : tendsto g b (𝓝 a)) (hh : tendsto h b (𝓝 a)) (hgf : ∀ᶠ b in b, g b ≤ f b) (hfh : ∀ᶠ b in b, f b ≤ h b) : tendsto f b (𝓝 a) := -tendsto_order.2 - ⟨assume a' h', - have ∀ᶠ b in b, a' < g b, from (tendsto_order.1 hg).left a' h', - by filter_upwards [this, hgf] with _ using lt_of_lt_of_le, - assume a' h', - have ∀ᶠ b in b, h b < a', from (tendsto_order.1 hh).right a' h', - by filter_upwards [this, hfh] with a h₁ h₂ using lt_of_le_of_lt h₂ h₁⟩ - -/-- Also known as squeeze or sandwich theorem. This version assumes that inequalities hold -everywhere. -/ +(hg.Icc hh).of_small_sets $ hgf.and hfh + +/-- **Squeeze theorem** (also known as **sandwich theorem**). This version assumes that inequalities +hold everywhere. -/ lemma tendsto_of_tendsto_of_tendsto_of_le_of_le {f g h : β → α} {b : filter β} {a : α} (hg : tendsto g b (𝓝 a)) (hh : tendsto h b (𝓝 a)) (hgf : g ≤ f) (hfh : f ≤ h) : tendsto f b (𝓝 a) := From 39014ec913961a73600b7bee785481b9acee70b4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 31 Aug 2022 14:19:01 +0000 Subject: [PATCH 116/828] feat(probability/independence): add indep_bot_left and indep_bot_right (#16309) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove that for any `m`, `indep m ⊥ μ`, and prove the corresponding statement with bot on the left. Also declare two types as variables on top of the file and remove them from many lemmas. --- src/probability/independence.lean | 78 +++++++++++++++++++------------ 1 file changed, 48 insertions(+), 30 deletions(-) diff --git a/src/probability/independence.lean b/src/probability/independence.lean index e8262b7feb4d0..889236043b328 100644 --- a/src/probability/independence.lean +++ b/src/probability/independence.lean @@ -67,19 +67,21 @@ open_locale big_operators classical measure_theory namespace probability_theory +variables {Ω ι : Type*} + section definitions /-- A family of sets of sets `π : ι → set (set Ω)` is independent with respect to a measure `μ` if for any finite set of indices `s = {i_1, ..., i_n}`, for any sets `f i_1 ∈ π i_1, ..., f i_n ∈ π i_n`, then `μ (⋂ i in s, f i) = ∏ i in s, μ (f i) `. It will be used for families of pi_systems. -/ -def Indep_sets {Ω ι} [measurable_space Ω] (π : ι → set (set Ω)) (μ : measure Ω . volume_tac) : +def Indep_sets [measurable_space Ω] (π : ι → set (set Ω)) (μ : measure Ω . volume_tac) : Prop := ∀ (s : finset ι) {f : ι → set Ω} (H : ∀ i, i ∈ s → f i ∈ π i), μ (⋂ i ∈ s, f i) = ∏ i in s, μ (f i) /-- Two sets of sets `s₁, s₂` are independent with respect to a measure `μ` if for any sets `t₁ ∈ p₁, t₂ ∈ s₂`, then `μ (t₁ ∩ t₂) = μ (t₁) * μ (t₂)` -/ -def indep_sets {Ω} [measurable_space Ω] (s1 s2 : set (set Ω)) (μ : measure Ω . volume_tac) : Prop := +def indep_sets [measurable_space Ω] (s1 s2 : set (set Ω)) (μ : measure Ω . volume_tac) : Prop := ∀ t1 t2 : set Ω, t1 ∈ s1 → t2 ∈ s2 → μ (t1 ∩ t2) = μ t1 * μ t2 /-- A family of measurable space structures (i.e. of σ-algebras) is independent with respect to a @@ -87,39 +89,39 @@ measure `μ` (typically defined on a finer σ-algebra) if the family of sets of define is independent. `m : ι → measurable_space Ω` is independent with respect to measure `μ` if for any finite set of indices `s = {i_1, ..., i_n}`, for any sets `f i_1 ∈ m i_1, ..., f i_n ∈ m i_n`, then `μ (⋂ i in s, f i) = ∏ i in s, μ (f i) `. -/ -def Indep {Ω ι} (m : ι → measurable_space Ω) [measurable_space Ω] (μ : measure Ω . volume_tac) : +def Indep (m : ι → measurable_space Ω) [measurable_space Ω] (μ : measure Ω . volume_tac) : Prop := Indep_sets (λ x, {s | measurable_set[m x] s}) μ /-- Two measurable space structures (or σ-algebras) `m₁, m₂` are independent with respect to a measure `μ` (defined on a third σ-algebra) if for any sets `t₁ ∈ m₁, t₂ ∈ m₂`, `μ (t₁ ∩ t₂) = μ (t₁) * μ (t₂)` -/ -def indep {Ω} (m₁ m₂ : measurable_space Ω) [measurable_space Ω] (μ : measure Ω . volume_tac) : +def indep (m₁ m₂ : measurable_space Ω) [measurable_space Ω] (μ : measure Ω . volume_tac) : Prop := indep_sets {s | measurable_set[m₁] s} {s | measurable_set[m₂] s} μ /-- A family of sets is independent if the family of measurable space structures they generate is independent. For a set `s`, the generated measurable space has measurable sets `∅, s, sᶜ, univ`. -/ -def Indep_set {Ω ι} [measurable_space Ω] (s : ι → set Ω) (μ : measure Ω . volume_tac) : Prop := +def Indep_set [measurable_space Ω] (s : ι → set Ω) (μ : measure Ω . volume_tac) : Prop := Indep (λ i, generate_from {s i}) μ /-- Two sets are independent if the two measurable space structures they generate are independent. For a set `s`, the generated measurable space structure has measurable sets `∅, s, sᶜ, univ`. -/ -def indep_set {Ω} [measurable_space Ω] (s t : set Ω) (μ : measure Ω . volume_tac) : Prop := +def indep_set [measurable_space Ω] (s t : set Ω) (μ : measure Ω . volume_tac) : Prop := indep (generate_from {s}) (generate_from {t}) μ /-- A family of functions defined on the same space `Ω` and taking values in possibly different spaces, each with a measurable space structure, is independent if the family of measurable space structures they generate on `Ω` is independent. For a function `g` with codomain having measurable space structure `m`, the generated measurable space structure is `measurable_space.comap g m`. -/ -def Indep_fun {Ω ι} [measurable_space Ω] {β : ι → Type*} (m : Π (x : ι), measurable_space (β x)) +def Indep_fun [measurable_space Ω] {β : ι → Type*} (m : Π (x : ι), measurable_space (β x)) (f : Π (x : ι), Ω → β x) (μ : measure Ω . volume_tac) : Prop := Indep (λ x, measurable_space.comap (f x) (m x)) μ /-- Two functions are independent if the two measurable space structures they generate are independent. For a function `f` with codomain having measurable space structure `m`, the generated measurable space structure is `measurable_space.comap f m`. -/ -def indep_fun {Ω β γ} [measurable_space Ω] [mβ : measurable_space β] [mγ : measurable_space γ] +def indep_fun {β γ} [measurable_space Ω] [mβ : measurable_space β] [mγ : measurable_space γ] (f : Ω → β) (g : Ω → γ) (μ : measure Ω . volume_tac) : Prop := indep (measurable_space.comap f mβ) (measurable_space.comap g mγ) μ @@ -127,37 +129,53 @@ end definitions section indep -lemma indep_sets.symm {Ω} {s₁ s₂ : set (set Ω)} [measurable_space Ω] {μ : measure Ω} +lemma indep_sets.symm {s₁ s₂ : set (set Ω)} [measurable_space Ω] {μ : measure Ω} (h : indep_sets s₁ s₂ μ) : indep_sets s₂ s₁ μ := by { intros t1 t2 ht1 ht2, rw [set.inter_comm, mul_comm], exact h t2 t1 ht2 ht1, } -lemma indep.symm {Ω} {m₁ m₂ : measurable_space Ω} [measurable_space Ω] {μ : measure Ω} +lemma indep.symm {m₁ m₂ : measurable_space Ω} [measurable_space Ω] {μ : measure Ω} (h : indep m₁ m₂ μ) : indep m₂ m₁ μ := indep_sets.symm h -lemma indep_sets_of_indep_sets_of_le_left {Ω} {s₁ s₂ s₃: set (set Ω)} [measurable_space Ω] +lemma indep_bot_right (m' : measurable_space Ω) {m : measurable_space Ω} + {μ : measure Ω} [is_probability_measure μ] : + indep m' ⊥ μ := +begin + intros s t hs ht, + rw [set.mem_set_of_eq, measurable_space.measurable_set_bot_iff] at ht, + cases ht, + { rw [ht, set.inter_empty, measure_empty, mul_zero], }, + { rw [ht, set.inter_univ, measure_univ, mul_one], }, +end + +lemma indep_bot_left (m' : measurable_space Ω) {m : measurable_space Ω} + {μ : measure Ω} [is_probability_measure μ] : + indep ⊥ m' μ := +(indep_bot_right m').symm + +lemma indep_sets_of_indep_sets_of_le_left {s₁ s₂ s₃: set (set Ω)} [measurable_space Ω] {μ : measure Ω} (h_indep : indep_sets s₁ s₂ μ) (h31 : s₃ ⊆ s₁) : indep_sets s₃ s₂ μ := λ t1 t2 ht1 ht2, h_indep t1 t2 (set.mem_of_subset_of_mem h31 ht1) ht2 -lemma indep_sets_of_indep_sets_of_le_right {Ω} {s₁ s₂ s₃: set (set Ω)} [measurable_space Ω] +lemma indep_sets_of_indep_sets_of_le_right {s₁ s₂ s₃: set (set Ω)} [measurable_space Ω] {μ : measure Ω} (h_indep : indep_sets s₁ s₂ μ) (h32 : s₃ ⊆ s₂) : indep_sets s₁ s₃ μ := λ t1 t2 ht1 ht2, h_indep t1 t2 ht1 (set.mem_of_subset_of_mem h32 ht2) -lemma indep_of_indep_of_le_left {Ω} {m₁ m₂ m₃: measurable_space Ω} [measurable_space Ω] +lemma indep_of_indep_of_le_left {m₁ m₂ m₃: measurable_space Ω} [measurable_space Ω] {μ : measure Ω} (h_indep : indep m₁ m₂ μ) (h31 : m₃ ≤ m₁) : indep m₃ m₂ μ := λ t1 t2 ht1 ht2, h_indep t1 t2 (h31 _ ht1) ht2 -lemma indep_of_indep_of_le_right {Ω} {m₁ m₂ m₃: measurable_space Ω} [measurable_space Ω] +lemma indep_of_indep_of_le_right {m₁ m₂ m₃: measurable_space Ω} [measurable_space Ω] {μ : measure Ω} (h_indep : indep m₁ m₂ μ) (h32 : m₃ ≤ m₂) : indep m₁ m₃ μ := λ t1 t2 ht1 ht2, h_indep t1 t2 ht1 (h32 _ ht2) -lemma indep_sets.union {Ω} [measurable_space Ω] {s₁ s₂ s' : set (set Ω)} {μ : measure Ω} +lemma indep_sets.union [measurable_space Ω] {s₁ s₂ s' : set (set Ω)} {μ : measure Ω} (h₁ : indep_sets s₁ s' μ) (h₂ : indep_sets s₂ s' μ) : indep_sets (s₁ ∪ s₂) s' μ := begin @@ -167,14 +185,14 @@ begin { exact h₂ t1 t2 ht1₂ ht2, }, end -@[simp] lemma indep_sets.union_iff {Ω} [measurable_space Ω] {s₁ s₂ s' : set (set Ω)} +@[simp] lemma indep_sets.union_iff [measurable_space Ω] {s₁ s₂ s' : set (set Ω)} {μ : measure Ω} : indep_sets (s₁ ∪ s₂) s' μ ↔ indep_sets s₁ s' μ ∧ indep_sets s₂ s' μ := ⟨λ h, ⟨indep_sets_of_indep_sets_of_le_left h (set.subset_union_left s₁ s₂), indep_sets_of_indep_sets_of_le_left h (set.subset_union_right s₁ s₂)⟩, λ h, indep_sets.union h.left h.right⟩ -lemma indep_sets.Union {Ω ι} [measurable_space Ω] {s : ι → set (set Ω)} {s' : set (set Ω)} +lemma indep_sets.Union [measurable_space Ω] {s : ι → set (set Ω)} {s' : set (set Ω)} {μ : measure Ω} (hyp : ∀ n, indep_sets (s n) s' μ) : indep_sets (⋃ n, s n) s' μ := begin @@ -184,17 +202,17 @@ begin exact hyp n t1 t2 ht1 ht2, end -lemma indep_sets.inter {Ω} [measurable_space Ω] {s₁ s' : set (set Ω)} (s₂ : set (set Ω)) +lemma indep_sets.inter [measurable_space Ω] {s₁ s' : set (set Ω)} (s₂ : set (set Ω)) {μ : measure Ω} (h₁ : indep_sets s₁ s' μ) : indep_sets (s₁ ∩ s₂) s' μ := λ t1 t2 ht1 ht2, h₁ t1 t2 ((set.mem_inter_iff _ _ _).mp ht1).left ht2 -lemma indep_sets.Inter {Ω ι} [measurable_space Ω] {s : ι → set (set Ω)} {s' : set (set Ω)} +lemma indep_sets.Inter [measurable_space Ω] {s : ι → set (set Ω)} {s' : set (set Ω)} {μ : measure Ω} (h : ∃ n, indep_sets (s n) s' μ) : indep_sets (⋂ n, s n) s' μ := by {intros t1 t2 ht1 ht2, cases h with n h, exact h t1 t2 (set.mem_Inter.mp ht1 n) ht2 } -lemma indep_sets_singleton_iff {Ω} [measurable_space Ω] {s t : set Ω} {μ : measure Ω} : +lemma indep_sets_singleton_iff [measurable_space Ω] {s t : set Ω} {μ : measure Ω} : indep_sets {s} {t} μ ↔ μ (s ∩ t) = μ s * μ t := ⟨λ h, h s t rfl rfl, λ h s1 t1 hs1 ht1, by rwa [set.mem_singleton_iff.mp hs1, set.mem_singleton_iff.mp ht1]⟩ @@ -204,7 +222,7 @@ end indep /-! ### Deducing `indep` from `Indep` -/ section from_Indep_to_indep -lemma Indep_sets.indep_sets {Ω ι} {s : ι → set (set Ω)} [measurable_space Ω] {μ : measure Ω} +lemma Indep_sets.indep_sets {s : ι → set (set Ω)} [measurable_space Ω] {μ : measure Ω} (h_indep : Indep_sets s μ) {i j : ι} (hij : i ≠ j) : indep_sets (s i) (s j) μ := begin @@ -229,7 +247,7 @@ begin rw [←h_inter, ←h_prod, h_indep {i, j} hf_m], end -lemma Indep.indep {Ω ι} {m : ι → measurable_space Ω} [measurable_space Ω] {μ : measure Ω} +lemma Indep.indep {m : ι → measurable_space Ω} [measurable_space Ω] {μ : measure Ω} (h_indep : Indep m μ) {i j : ι} (hij : i ≠ j) : indep (m i) (m j) μ := begin @@ -237,7 +255,7 @@ begin exact Indep_sets.indep_sets h_indep hij, end -lemma Indep_fun.indep_fun {Ω ι : Type*} {m₀ : measurable_space Ω} {μ : measure Ω} {β : ι → Type*} +lemma Indep_fun.indep_fun {m₀ : measurable_space Ω} {μ : measure Ω} {β : ι → Type*} {m : Π x, measurable_space (β x)} {f : Π i, Ω → β i} (hf_Indep : Indep_fun m f μ) {i j : ι} (hij : i ≠ j) : indep_fun (f i) (f j) μ := @@ -254,14 +272,14 @@ Independence of measurable spaces is equivalent to independence of generating π section from_measurable_spaces_to_sets_of_sets /-! ### Independence of measurable space structures implies independence of generating π-systems -/ -lemma Indep.Indep_sets {Ω ι} [measurable_space Ω] {μ : measure Ω} {m : ι → measurable_space Ω} +lemma Indep.Indep_sets [measurable_space Ω] {μ : measure Ω} {m : ι → measurable_space Ω} {s : ι → set (set Ω)} (hms : ∀ n, m n = generate_from (s n)) (h_indep : Indep m μ) : Indep_sets s μ := λ S f hfs, h_indep S $ λ x hxS, ((hms x).symm ▸ measurable_set_generate_from (hfs x hxS) : measurable_set[m x] (f x)) -lemma indep.indep_sets {Ω} [measurable_space Ω] {μ : measure Ω} {s1 s2 : set (set Ω)} +lemma indep.indep_sets [measurable_space Ω] {μ : measure Ω} {s1 s2 : set (set Ω)} (h_indep : indep (generate_from s1) (generate_from s2) μ) : indep_sets s1 s2 μ := λ t1 t2 ht1 ht2, h_indep t1 t2 (measurable_set_generate_from ht1) (measurable_set_generate_from ht2) @@ -271,7 +289,7 @@ end from_measurable_spaces_to_sets_of_sets section from_pi_systems_to_measurable_spaces /-! ### Independence of generating π-systems implies independence of measurable space structures -/ -private lemma indep_sets.indep_aux {Ω} {m2 : measurable_space Ω} +private lemma indep_sets.indep_aux {m2 : measurable_space Ω} {m : measurable_space Ω} {μ : measure Ω} [is_probability_measure μ] {p1 p2 : set (set Ω)} (h2 : m2 ≤ m) (hp2 : is_pi_system p2) (hpm2 : m2 = generate_from p2) (hyp : indep_sets p1 p2 μ) {t1 t2 : set Ω} (ht1 : t1 ∈ p1) (ht2m : measurable_set[m2] t2) : @@ -292,7 +310,7 @@ begin exact hyp t1 t ht1 ht, end -lemma indep_sets.indep {Ω} {m1 m2 : measurable_space Ω} {m : measurable_space Ω} +lemma indep_sets.indep {m1 m2 : measurable_space Ω} {m : measurable_space Ω} {μ : measure Ω} [is_probability_measure μ] {p1 p2 : set (set Ω)} (h1 : m1 ≤ m) (h2 : m2 ≤ m) (hp1 : is_pi_system p1) (hp2 : is_pi_system p2) (hpm1 : m1 = generate_from p1) (hpm2 : m2 = generate_from p2) (hyp : indep_sets p1 p2 μ) : @@ -314,7 +332,7 @@ begin exact indep_sets.indep_aux h2 hp2 hpm2 hyp ht ht2, end -variables {Ω ι : Type*} {m0 : measurable_space Ω} {μ : measure Ω} +variables {m0 : measurable_space Ω} {μ : measure Ω} lemma Indep_sets.pi_Union_Inter_singleton {π : ι → set (set Ω)} {a : ι} {S : finset ι} (hp_ind : Indep_sets π μ) (haS : a ∉ S) : @@ -437,7 +455,7 @@ We prove the following equivalences on `indep_set`, for measurable sets `s, t`. * `indep_set s t μ ↔ indep_sets {s} {t} μ`. -/ -variables {Ω : Type*} [measurable_space Ω] {s t : set Ω} (S T : set (set Ω)) +variables [measurable_space Ω] {s t : set Ω} (S T : set (set Ω)) lemma indep_set_iff_indep_sets_singleton (hs_meas : measurable_set s) (ht_meas : measurable_set t) (μ : measure Ω . volume_tac) [is_probability_measure μ] : @@ -466,7 +484,7 @@ section indep_fun -/ -variables {Ω β β' γ γ' : Type*} {mΩ : measurable_space Ω} {μ : measure Ω} {f : Ω → β} {g : Ω → β'} +variables {β β' γ γ' : Type*} {mΩ : measurable_space Ω} {μ : measure Ω} {f : Ω → β} {g : Ω → β'} lemma indep_fun_iff_measure_inter_preimage_eq_mul {mβ : measurable_space β} {mβ' : measurable_space β'} : From 5b7bc67670813e8f3fd1cbe272f750c0616441cf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 31 Aug 2022 17:11:05 +0000 Subject: [PATCH 117/828] refactor(group_theory/subgroup/basic): Make `subgroup.opposite` an `equiv` (#16271) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit There is currently no way to turn a `subgroup Gᵐᵒᵖ` into a `subgroup G`, so this PR bundles `subgroup.opposite` as an `equiv` so that now we can use `subgroup.opposite.symm`. --- src/group_theory/subgroup/basic.lean | 16 +++++++++++----- 1 file changed, 11 insertions(+), 5 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index d61ef72b57b89..606d4c97c01e1 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -3168,11 +3168,17 @@ namespace subgroup /-- A subgroup `H` of `G` determines a subgroup `H.opposite` of the opposite group `Gᵐᵒᵖ`. -/ @[to_additive "An additive subgroup `H` of `G` determines an additive subgroup `H.opposite` of the opposite additive group `Gᵃᵒᵖ`."] -def opposite (H : subgroup G) : subgroup Gᵐᵒᵖ := -{ carrier := mul_opposite.unop ⁻¹' (H : set G), - one_mem' := H.one_mem, - mul_mem' := λ a b ha hb, H.mul_mem hb ha, - inv_mem' := λ a, H.inv_mem } +def opposite : subgroup G ≃ subgroup Gᵐᵒᵖ := +{ to_fun := λ H, { carrier := mul_opposite.unop ⁻¹' (H : set G), + one_mem' := H.one_mem, + mul_mem' := λ a b ha hb, H.mul_mem hb ha, + inv_mem' := λ a, H.inv_mem }, + inv_fun := λ H, { carrier := mul_opposite.op ⁻¹' (H : set Gᵐᵒᵖ), + one_mem' := H.one_mem, + mul_mem' := λ a b ha hb, H.mul_mem hb ha, + inv_mem' := λ a, H.inv_mem }, + left_inv := λ H, set_like.coe_injective rfl, + right_inv := λ H, set_like.coe_injective rfl } /-- Bijection between a subgroup `H` and its opposite. -/ @[to_additive "Bijection between an additive subgroup `H` and its opposite.", simps] From 371b4143b0298311cf76de600deb109816645e0f Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 31 Aug 2022 17:11:06 +0000 Subject: [PATCH 118/828] chore(algebra/divisibility): golf (#16327) --- src/algebra/divisibility.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/algebra/divisibility.lean b/src/algebra/divisibility.lean index 45a0b884cde27..aaf3eddfce219 100644 --- a/src/algebra/divisibility.lean +++ b/src/algebra/divisibility.lean @@ -259,8 +259,7 @@ section comm_monoid variables [comm_monoid α] theorem is_unit_iff_dvd_one {x : α} : is_unit x ↔ x ∣ 1 := -⟨by rintro ⟨u, rfl⟩; exact ⟨_, u.mul_inv.symm⟩, - λ ⟨y, h⟩, ⟨⟨x, y, h.symm, by rw [h, mul_comm]⟩, rfl⟩⟩ +⟨is_unit.dvd, λ ⟨y, h⟩, ⟨⟨x, y, h.symm, by rw [h, mul_comm]⟩, rfl⟩⟩ theorem is_unit_iff_forall_dvd {x : α} : is_unit x ↔ ∀ y, x ∣ y := From 87d8c4b33fa7d8eead1092abfcf7c21fc008786b Mon Sep 17 00:00:00 2001 From: mcdoll Date: Wed, 31 Aug 2022 20:06:29 +0000 Subject: [PATCH 119/828] doc(measure_theory/integral): integrability docstrings (#16269) This PR adds a few docstrings to make the conditions and more readable. --- .../integral/integral_eq_improper.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/measure_theory/integral/integral_eq_improper.lean b/src/measure_theory/integral/integral_eq_improper.lean index 4e08222f9edca..fe3dcabd1f97e 100644 --- a/src/measure_theory/integral/integral_eq_improper.lean +++ b/src/measure_theory/integral/integral_eq_improper.lean @@ -521,6 +521,10 @@ begin rwa ←interval_integral.integral_of_le (hai.trans hbi) end +/-- If `f` is integrable on intervals `Ioc (a i) (b i)`, +where `a i` tends to -∞ and `b i` tends to ∞, and +`∫ x in a i .. b i, ∥f x∥ ∂μ` converges to `I : ℝ` along a filter `l`, +then `f` is integrable on the interval (-∞, ∞) -/ lemma integrable_of_interval_integral_norm_tendsto (I : ℝ) (hfi : ∀ i, integrable_on f (Ioc (a i) (b i)) μ) (ha : tendsto a l at_bot) (hb : tendsto b l at_top) @@ -545,6 +549,10 @@ begin exact id end +/-- If `f` is integrable on intervals `Ioc (a i) b`, +where `a i` tends to -∞, and +`∫ x in a i .. b, ∥f x∥ ∂μ` converges to `I : ℝ` along a filter `l`, +then `f` is integrable on the interval (-∞, b) -/ lemma integrable_on_Iic_of_interval_integral_norm_tendsto (I b : ℝ) (hfi : ∀ i, integrable_on f (Ioc (a i) b) μ) (ha : tendsto a l at_bot) (h : tendsto (λ i, ∫ x in a i .. b, ∥f x∥ ∂μ) l (𝓝 I)) : @@ -569,6 +577,10 @@ begin exact id end +/-- If `f` is integrable on intervals `Ioc a (b i)`, +where `b i` tends to ∞, and +`∫ x in a .. b i, ∥f x∥ ∂μ` converges to `I : ℝ` along a filter `l`, +then `f` is integrable on the interval (a, ∞) -/ lemma integrable_on_Ioi_of_interval_integral_norm_tendsto (I a : ℝ) (hfi : ∀ i, integrable_on f (Ioc a (b i)) μ) (hb : tendsto b l at_top) (h : tendsto (λ i, ∫ x in a .. b i, ∥f x∥ ∂μ) l (𝓝 $ I)) : From acbe099ced8be9c9754d62860110295cde0d7181 Mon Sep 17 00:00:00 2001 From: Gabriel Ebner Date: Wed, 31 Aug 2022 20:06:30 +0000 Subject: [PATCH 120/828] feat(*): bump to lean 3.48.0 (#16292) `fin n` is a structure now. --- archive/imo/imo1987_q1.lean | 2 +- leanpkg.toml | 2 +- src/algebra/big_operators/fin.lean | 2 +- src/algebra/char_zero/quotient.lean | 2 +- src/algebraic_topology/simplex_category.lean | 12 +-- .../split_simplicial_object.lean | 2 +- .../topological_simplex.lean | 3 +- .../complex/upper_half_plane/basic.lean | 4 +- src/analysis/convex/stone_separation.lean | 2 +- src/combinatorics/composition.lean | 2 +- src/combinatorics/simple_graph/coloring.lean | 4 +- src/computability/primrec.lean | 2 +- src/data/buffer/basic.lean | 2 +- src/data/countable/defs.lean | 3 +- src/data/fin/basic.lean | 94 ++++++++++++++----- src/data/fin/interval.lean | 56 +++++------ src/data/fin/tuple/basic.lean | 2 +- src/data/finset/basic.lean | 19 ++++ src/data/finset/sort.lean | 4 +- src/data/fintype/card.lean | 7 +- src/data/list/range.lean | 6 +- src/field_theory/finite/polynomial.lean | 2 +- src/group_theory/order_of_element.lean | 7 +- src/linear_algebra/matrix/block.lean | 2 +- src/linear_algebra/matrix/transvection.lean | 6 +- src/logic/encodable/basic.lean | 3 +- src/logic/equiv/fin.lean | 2 +- .../class_number/admissible_abs.lean | 2 +- src/number_theory/class_number/finite.lean | 2 +- src/ring_theory/polynomial/basic.lean | 2 +- src/ring_theory/witt_vector/truncated.lean | 2 +- src/testing/slim_check/sampleable.lean | 4 +- .../metric_space/gromov_hausdorff.lean | 10 +- 33 files changed, 173 insertions(+), 103 deletions(-) diff --git a/archive/imo/imo1987_q1.lean b/archive/imo/imo1987_q1.lean index 9dd606ffe31bb..a70d2de162618 100644 --- a/archive/imo/imo1987_q1.lean +++ b/archive/imo/imo1987_q1.lean @@ -66,7 +66,7 @@ def fixed_points_equiv' : inv_fun := λ p, ⟨⟨card (fixed_points p.1.2), (card_subtype_le _).trans_lt (nat.lt_succ_self _)⟩, ⟨p.1.2, rfl⟩, ⟨p.1.1, p.2⟩⟩, - left_inv := λ ⟨⟨k, hk⟩, ⟨σ, hσ⟩, ⟨x, hx⟩⟩, by { simp only [mem_fiber, subtype.coe_mk] at hσ, + left_inv := λ ⟨⟨k, hk⟩, ⟨σ, hσ⟩, ⟨x, hx⟩⟩, by { simp only [mem_fiber, fin.coe_mk] at hσ, subst k, refl }, right_inv := λ ⟨⟨x, σ⟩, h⟩, rfl } diff --git a/leanpkg.toml b/leanpkg.toml index 009b32dd0d769..a20d0a6c9ff2c 100644 --- a/leanpkg.toml +++ b/leanpkg.toml @@ -1,7 +1,7 @@ [package] name = "mathlib" version = "0.1" -lean_version = "leanprover-community/lean:3.47.0" +lean_version = "leanprover-community/lean:3.48.0" path = "src" [dependencies] diff --git a/src/algebra/big_operators/fin.lean b/src/algebra/big_operators/fin.lean index 8645a7154fb64..13040f2bf89b1 100644 --- a/src/algebra/big_operators/fin.lean +++ b/src/algebra/big_operators/fin.lean @@ -222,7 +222,7 @@ begin rw prod_take_succ _ _ this, have A : ((finset.univ : finset (fin n)).filter (λ j, j.val < i + 1)) = ((finset.univ : finset (fin n)).filter (λ j, j.val < i)) ∪ {(⟨i, h⟩ : fin n)}, - by { ext j, simp [nat.lt_succ_iff_lt_or_eq, fin.ext_iff, - add_comm] }, + by { ext ⟨_, _⟩, simp [nat.lt_succ_iff_lt_or_eq] }, have B : _root_.disjoint (finset.filter (λ (j : fin n), j.val < i) finset.univ) (singleton (⟨i, h⟩ : fin n)), by simp, rw [A, finset.prod_union B, IH], diff --git a/src/algebra/char_zero/quotient.lean b/src/algebra/char_zero/quotient.lean index ef5a288643127..3588403521e3d 100644 --- a/src/algebra/char_zero/quotient.lean +++ b/src/algebra/char_zero/quotient.lean @@ -31,7 +31,7 @@ begin refine ⟨⟨(k % z).to_nat, _⟩, k / z, _⟩, { rw [←int.coe_nat_lt, int.to_nat_of_nonneg (int.mod_nonneg _ hz)], exact (int.mod_lt _ hz).trans_eq (int.abs_eq_nat_abs _) }, - rw [subtype.coe_mk, int.to_nat_of_nonneg (int.mod_nonneg _ hz), int.div_add_mod] }, + rw [fin.coe_mk, int.to_nat_of_nonneg (int.mod_nonneg _ hz), int.div_add_mod] }, { rintro ⟨k, n, h⟩, exact ⟨_, h⟩, }, end diff --git a/src/algebraic_topology/simplex_category.lean b/src/algebraic_topology/simplex_category.lean index 62fddc97ee704..b1e224344c08e 100644 --- a/src/algebraic_topology/simplex_category.lean +++ b/src/algebraic_topology/simplex_category.lean @@ -193,7 +193,7 @@ begin rcases i with ⟨i, _⟩, rcases j with ⟨j, _⟩, rcases k with ⟨k, _⟩, - simp only [subtype.mk_le_mk, fin.cast_succ_mk] at H, + simp only [fin.mk_le_mk, fin.cast_succ_mk] at H, dsimp, split_ifs, -- Most of the goals can now be handled by `linarith`, @@ -241,7 +241,7 @@ begin rcases i with ⟨i, _⟩, rcases j with ⟨j, _⟩, rcases k with ⟨k, _⟩, - simp only [subtype.mk_lt_mk, fin.cast_succ_mk] at H, + simp only [fin.mk_lt_mk, fin.cast_succ_mk] at H, suffices : ite (_ < ite (k < i + 1) _ _) _ _ = ite _ (ite (j < k) (k - 1) k) (ite (j < k) (k - 1) k + 1), { simpa [apply_dite fin.cast_succ, fin.pred_above] with push_cast, }, @@ -249,20 +249,20 @@ begin -- Most of the goals can now be handled by `linarith`, -- but we have to deal with three of them by hand. swap 2, - { simp only [subtype.mk_lt_mk] at h_1, + { simp only [fin.mk_lt_mk] at h_1, simp only [not_lt] at h_2, simp only [self_eq_add_right, one_ne_zero], exact lt_irrefl (k - 1) (lt_of_lt_of_le (nat.pred_lt (ne_of_lt (lt_of_le_of_lt (zero_le _) h_1)).symm) (le_trans (nat.le_of_lt_succ h) h_2)) }, swap 4, - { simp only [subtype.mk_lt_mk] at h_1, + { simp only [fin.mk_lt_mk] at h_1, simp only [not_lt] at h, simp only [nat.add_succ_sub_one, add_zero], exfalso, exact lt_irrefl _ (lt_of_le_of_lt (nat.le_pred_of_lt (nat.lt_of_succ_le h)) h_3), }, swap 4, - { simp only [subtype.mk_lt_mk] at h_1, + { simp only [fin.mk_lt_mk] at h_1, simp only [not_lt] at h_3, simp only [nat.add_succ_sub_one, add_zero], exact (nat.succ_pred_eq_of_pos (lt_of_le_of_lt (zero_le _) h_2)).symm, }, @@ -281,7 +281,7 @@ begin rcases i with ⟨i, _⟩, rcases j with ⟨j, _⟩, rcases k with ⟨k, _⟩, - simp only [subtype.mk_le_mk] at H, + simp only [fin.mk_le_mk] at H, -- At this point `simp with push_cast` makes good progress, but neither `simp?` nor `squeeze_simp` -- return usable sets of lemmas. -- To avoid using a non-terminal simp, we make a `suffices` statement indicating the shape diff --git a/src/algebraic_topology/split_simplicial_object.lean b/src/algebraic_topology/split_simplicial_object.lean index eef5946db01dd..30b18af9c349e 100644 --- a/src/algebraic_topology/split_simplicial_object.lean +++ b/src/algebraic_topology/split_simplicial_object.lean @@ -89,7 +89,7 @@ begin induction Δ₁ using opposite.rec, induction Δ₂ using opposite.rec, simp only at h₁, - have h₂ : Δ₁ = Δ₂ := by { ext1, simpa only [subtype.mk_eq_mk] using h₁.1, }, + have h₂ : Δ₁ = Δ₂ := by { ext1, simpa only [fin.mk_eq_mk] using h₁.1, }, subst h₂, refine ext _ _ rfl _, ext : 2, diff --git a/src/algebraic_topology/topological_simplex.lean b/src/algebraic_topology/topological_simplex.lean index 920922530903d..9b937cb513264 100644 --- a/src/algebraic_topology/topological_simplex.lean +++ b/src/algebraic_topology/topological_simplex.lean @@ -40,8 +40,7 @@ lemma to_Top_obj.ext {x : simplex_category} (f g : x.to_Top_obj) : def to_Top_map {x y : simplex_category} (f : x ⟶ y) : x.to_Top_obj → y.to_Top_obj := λ g, ⟨λ i, ∑ j in (finset.univ.filter (λ k, f k = i)), g j, begin - dsimp [to_Top_obj], - simp only [finset.filter_congr_decidable, finset.sum_congr], + simp only [finset.filter_congr_decidable, finset.sum_congr, to_Top_obj, set.mem_set_of], rw ← finset.sum_bUnion, convert g.2, { rw finset.eq_univ_iff_forall, diff --git a/src/analysis/complex/upper_half_plane/basic.lean b/src/analysis/complex/upper_half_plane/basic.lean index cac94c99a3f31..2832a2fbd3a0b 100644 --- a/src/analysis/complex/upper_half_plane/basic.lean +++ b/src/analysis/complex/upper_half_plane/basic.lean @@ -154,7 +154,7 @@ begin change _ = (_ * (_ / _) + _) * _, field_simp [denom_ne_zero, -denom, -num], simp only [matrix.mul, dot_product, fin.sum_univ_succ, denom, num, coe_coe, subgroup.coe_mul, - general_linear_group.coe_mul, fintype.univ_of_subsingleton, fin.mk_eq_subtype_mk, fin.mk_zero, + general_linear_group.coe_mul, fintype.univ_of_subsingleton, fin.mk_zero, finset.sum_singleton, fin.succ_zero_eq_one, complex.of_real_add, complex.of_real_mul], ring end @@ -167,7 +167,7 @@ begin rw denom_cocycle, field_simp [denom_ne_zero, -denom, -num], simp only [matrix.mul, dot_product, fin.sum_univ_succ, num, denom, coe_coe, subgroup.coe_mul, - general_linear_group.coe_mul, fintype.univ_of_subsingleton, fin.mk_eq_subtype_mk, fin.mk_zero, + general_linear_group.coe_mul, fintype.univ_of_subsingleton, fin.mk_zero, finset.sum_singleton, fin.succ_zero_eq_one, complex.of_real_add, complex.of_real_mul], ring end diff --git a/src/analysis/convex/stone_separation.lean b/src/analysis/convex/stone_separation.lean index 58752f53d5d71..905b8e6ea26f4 100644 --- a/src/analysis/convex/stone_separation.lean +++ b/src/analysis/convex/stone_separation.lean @@ -74,7 +74,7 @@ begin congr' 3, rw [←mul_smul, ←mul_rotate, mul_right_comm, mul_smul, ←mul_smul _ av, mul_rotate, mul_smul _ bz, ←smul_add], - simp only [list.map, list.pmap, nat.add_def, add_zero, fin.mk_eq_subtype_mk, fin.mk_bit0, + simp only [list.map, list.pmap, nat.add_def, add_zero, fin.mk_bit0, fin.mk_one, list.foldr_cons, list.foldr_nil], refl, end diff --git a/src/combinatorics/composition.lean b/src/combinatorics/composition.lean index 56fe48c566bfb..b39773ebcc5cb 100644 --- a/src/combinatorics/composition.lean +++ b/src/combinatorics/composition.lean @@ -719,7 +719,7 @@ def composition_as_set_equiv (n : ℕ) : composition_as_set n ≃ finset (fin (n erw [set.mem_set_of_eq], simp only [this, false_or, add_right_inj, add_eq_zero_iff, one_ne_zero, false_and, fin.coe_mk], split, - { rintros ⟨j, js, hj⟩, convert js, exact (fin.ext_iff _ _).2 hj }, + { rintros ⟨j, js, hj⟩, convert js, exact fin.ext_iff.2 hj }, { assume h, exact ⟨i, h, rfl⟩ } end } diff --git a/src/combinatorics/simple_graph/coloring.lean b/src/combinatorics/simple_graph/coloring.lean index d6544452730ad..9440e95fe8aa3 100644 --- a/src/combinatorics/simple_graph/coloring.lean +++ b/src/combinatorics/simple_graph/coloring.lean @@ -213,7 +213,7 @@ begin split, { rintro hc, have C : G.coloring (fin n) := hc.to_coloring (by simp), - let f := embedding.complete_graph (fin.coe_embedding n).to_embedding, + let f := embedding.complete_graph fin.coe_embedding, use f.to_hom.comp C, intro v, cases C with color valid, @@ -222,7 +222,7 @@ begin refine ⟨coloring.mk _ _⟩, { exact λ v, ⟨C v, Cf v⟩, }, { rintro v w hvw, - simp only [subtype.mk_eq_mk, ne.def], + simp only [fin.mk_eq_mk, ne.def], exact C.valid hvw, } } end diff --git a/src/computability/primrec.lean b/src/computability/primrec.lean index f67742e7b230c..7529faa862a0e 100644 --- a/src/computability/primrec.lean +++ b/src/computability/primrec.lean @@ -1035,7 +1035,7 @@ def subtype {p : α → Prop} [decidable_pred p] instance fin {n} : primcodable (fin n) := @of_equiv _ _ (subtype $ nat_lt.comp primrec.id (const n)) - (equiv.refl _) + fin.equiv_subtype instance vector {n} : primcodable (vector α n) := subtype ((@primrec.eq _ _ nat.decidable_eq).comp list_length (const _)) diff --git a/src/data/buffer/basic.lean b/src/data/buffer/basic.lean index 4ad7589694ea5..c5f0b084a399a 100644 --- a/src/data/buffer/basic.lean +++ b/src/data/buffer/basic.lean @@ -189,7 +189,7 @@ lemma nth_le_to_list (b : buffer α) {i : ℕ} (h) : nth_le_to_list' _ _ _ lemma read_eq_nth_le_to_list (b : buffer α) (i) : - b.read i = b.to_list.nth_le i (by simpa using i.is_lt) := + b.read i = b.to_list.nth_le i (by simp) := by simp [nth_le_to_list] lemma read_singleton (c : α) : [c].to_buffer.read ⟨0, by simp⟩ = c := diff --git a/src/data/countable/defs.lean b/src/data/countable/defs.lean index 8cb56fdef833b..d845f635ef8d0 100644 --- a/src/data/countable/defs.lean +++ b/src/data/countable/defs.lean @@ -67,7 +67,8 @@ instance subsingleton.to_countable [subsingleton α] : countable α := @[priority 500] instance [countable α] {p : α → Prop} : countable {x // p x} := subtype.val_injective.countable -instance {n : ℕ} : countable (fin n) := subtype.countable +instance {n : ℕ} : countable (fin n) := +function.injective.countable (@fin.eq_of_veq n) @[priority 100] instance finite.to_countable [finite α] : countable α := diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index 057bad0186779..de7bea16471e7 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -42,7 +42,9 @@ This file expands on the development in the core library. ### Order embeddings and an order isomorphism -* `fin.coe_embedding` : coercion to natural numbers as an `order_embedding`; +* `fin.order_iso_subtype` : coercion to `{ i // i < n }` as an `order_iso`; +* `fin.coe_embedding` : coercion to natural numbers as an `embedding`; +* `fin.coe_order_embedding` : coercion to natural numbers as an `order_embedding`; * `fin.succ_embedding` : `fin.succ` as an `order_embedding`; * `fin.cast_le h` : embed `fin n` into `fin m`, `h : n ≤ m`; * `fin.cast eq` : order isomorphism between `fin n` and fin m` provided that `n = m`, @@ -85,11 +87,23 @@ def elim0' {α : Sort*} (x : fin 0) : α := x.elim0 variables {n m : ℕ} {a b : fin n} -instance fin_to_nat (n : ℕ) : has_coe (fin n) nat := ⟨subtype.val⟩ +instance fin_to_nat (n : ℕ) : has_coe (fin n) nat := ⟨fin.val⟩ + +lemma val_injective : function.injective (@fin.val n) := @fin.eq_of_veq n + +protected lemma prop (a : fin n) : a.val < n := a.2 + +@[simp] lemma is_lt (a : fin n) : (a : nat) < n := a.2 lemma pos_iff_nonempty {n : ℕ} : 0 < n ↔ nonempty (fin n) := ⟨λ h, ⟨⟨0, h⟩⟩, λ ⟨i⟩, lt_of_le_of_lt (nat.zero_le _) i.2⟩ +/-- Equivalence between `fin n` and `{ i // i < n }`. -/ +@[simps apply symm_apply] +def equiv_subtype : fin n ≃ { i // i < n } := +{ to_fun := λ a, ⟨a.1, a.2⟩, inv_fun := λ a, ⟨a.1, a.2⟩, + left_inv := λ ⟨_, _⟩, rfl, right_inv := λ ⟨_, _⟩, rfl } + section coe /-! @@ -102,10 +116,13 @@ by cases a; refl @[ext] lemma ext {a b : fin n} (h : (a : ℕ) = b) : a = b := eq_of_veq h -lemma ext_iff (a b : fin n) : a = b ↔ (a : ℕ) = b := +lemma ext_iff {a b : fin n} : a = b ↔ (a : ℕ) = b := iff.intro (congr_arg _) fin.eq_of_veq -lemma coe_injective {n : ℕ} : injective (coe : fin n → ℕ) := subtype.coe_injective +lemma coe_injective {n : ℕ} : injective (coe : fin n → ℕ) := fin.val_injective + +lemma coe_eq_coe (a b : fin n) : (a : ℕ) = b ↔ a = b := +ext_iff.symm lemma eq_iff_veq (a b : fin n) : a = b ↔ a.1 = b.1 := ⟨veq_of_eq, eq_of_veq⟩ @@ -113,11 +130,13 @@ lemma eq_iff_veq (a b : fin n) : a = b ↔ a.1 = b.1 := lemma ne_iff_vne (a b : fin n) : a ≠ b ↔ a.1 ≠ b.1 := ⟨vne_of_ne, ne_of_vne⟩ -@[simp] lemma mk_eq_subtype_mk (a : ℕ) (h : a < n) : mk a h = ⟨a, h⟩ := rfl +@[simp, nolint simp_nf] -- built-in reduction doesn't always work +theorem mk_eq_mk {a h a' h'} : @mk n a h = @mk n a' h' ↔ a = a' := +ext_iff protected lemma mk.inj_iff {n a b : ℕ} {ha : a < n} {hb : b < n} : (⟨a, ha⟩ : fin n) = ⟨b, hb⟩ ↔ a = b := -subtype.mk_eq_mk +eq_iff_veq _ _ lemma mk_val {m n : ℕ} (h : m < n) : (⟨m, h⟩ : fin n).val = m := rfl @@ -137,11 +156,11 @@ lemma coe_eq_val (a : fin n) : (a : ℕ) = a.val := rfl then they coincide (in the heq sense). -/ protected lemma heq_fun_iff {α : Sort*} {k l : ℕ} (h : k = l) {f : fin k → α} {g : fin l → α} : f == g ↔ (∀ (i : fin k), f i = g ⟨(i : ℕ), h ▸ i.2⟩) := -by { induction h, simp [heq_iff_eq, function.funext_iff] } +by { subst h, simp [function.funext_iff] } protected lemma heq_ext_iff {k l : ℕ} (h : k = l) {i : fin k} {j : fin l} : i == j ↔ (i : ℕ) = (j : ℕ) := -by { induction h, simp [ext_iff] } +by { subst h, simp [coe_eq_coe] } lemma exists_iff {p : fin n → Prop} : (∃ i, p i) ↔ ∃ i h, p ⟨i, h⟩ := ⟨λ h, exists.elim h (λ ⟨i, hi⟩ hpi, ⟨i, hi, hpi⟩), @@ -158,10 +177,10 @@ section order ### order -/ -lemma is_lt (i : fin n) : (i : ℕ) < n := i.2 - lemma is_le (i : fin (n + 1)) : (i : ℕ) ≤ n := le_of_lt_succ i.is_lt +@[simp] lemma is_le' : (a : ℕ) ≤ n := le_of_lt a.is_lt + lemma lt_iff_coe_lt_coe : a < b ↔ (a : ℕ) < b := iff.rfl lemma le_iff_coe_le_coe : a ≤ b ↔ (a : ℕ) ≤ b := iff.rfl @@ -178,19 +197,44 @@ iff.rfl @[norm_cast, simp] lemma coe_fin_le {n : ℕ} {a b : fin n} : (a : ℕ) ≤ (b : ℕ) ↔ a ≤ b := iff.rfl -instance {n : ℕ} : linear_order (fin n) := subtype.linear_order _ +instance {n : ℕ} : linear_order (fin n) := +@linear_order.lift (fin n) _ _ ⟨λ x y, ⟨max x y, max_rec' (< n) x.2 y.2⟩⟩ + ⟨λ x y, ⟨min x y, min_rec' (< n) x.2 y.2⟩⟩ fin.val fin.val_injective (λ _ _, rfl) (λ _ _, rfl) + +@[simp] lemma mk_le_mk {x y : nat} {hx} {hy} : (⟨x, hx⟩ : fin n) ≤ ⟨y, hy⟩ ↔ x ≤ y := iff.rfl -instance {n : ℕ} : partial_order (fin n) := subtype.partial_order _ +@[simp] lemma mk_lt_mk {x y : nat} {hx} {hy} : (⟨x, hx⟩ : fin n) < ⟨y, hy⟩ ↔ x < y := iff.rfl + +@[simp] lemma min_coe : min (a : ℕ) n = a := by simp + +@[simp] lemma max_coe : max (a : ℕ) n = n := by simp + +instance {n : ℕ} : partial_order (fin n) := by apply_instance lemma coe_strict_mono : strict_mono (coe : fin n → ℕ) := λ _ _, id -/-- The inclusion map `fin n → ℕ` is a relation embedding. -/ -def coe_embedding (n) : (fin n) ↪o ℕ := -⟨⟨coe, @fin.eq_of_veq _⟩, λ a b, iff.rfl⟩ +/-- The equivalence `fin n ≃ { i // i < n }` is an order isomorphism. -/ +@[simps apply symm_apply] +def order_iso_subtype : fin n ≃o { i // i < n } := +equiv_subtype.to_order_iso (by simp [monotone]) (by simp [monotone]) + +/-- The inclusion map `fin n → ℕ` is an embedding. -/ +@[simps apply] +def coe_embedding : fin n ↪ ℕ := +⟨coe, coe_injective⟩ + +@[simp] lemma equiv_subtype_symm_trans_val_embedding : + equiv_subtype.symm.to_embedding.trans coe_embedding = embedding.subtype (< n) := +rfl + +/-- The inclusion map `fin n → ℕ` is an order embedding. -/ +@[simps apply] +def coe_order_embedding (n) : (fin n) ↪o ℕ := +⟨coe_embedding, λ a b, iff.rfl⟩ /-- The ordering on `fin n` is a well order. -/ instance fin.lt.is_well_order (n) : is_well_order (fin n) (<) := -(coe_embedding n).is_well_order +(coe_order_embedding n).is_well_order /-- Use the ordering on `fin n` for checking recursive definitions. @@ -264,7 +308,7 @@ map. In this lemma we state that for each `i : fin n` we have `(e i : ℕ) = (i @[simp] lemma coe_order_iso_apply (e : fin n ≃o fin m) (i : fin n) : (e i : ℕ) = i := begin rcases i with ⟨i, hi⟩, - rw [subtype.coe_mk], + rw [fin.coe_mk], induction i using nat.strong_induction_on with i h, refine le_antisymm (forall_lt_iff_le.1 $ λ j hj, _) (forall_lt_iff_le.1 $ λ j hj, _), { have := e.symm.lt_iff_lt.2 (mk_lt_of_lt_coe hj), @@ -513,7 +557,7 @@ lemma succ_injective (n : ℕ) : injective (@fin.succ n) := (succ_injective n).eq_iff lemma succ_ne_zero {n} : ∀ k : fin n, fin.succ k ≠ 0 -| ⟨k, hk⟩ heq := nat.succ_ne_zero k $ (ext_iff _ _).1 heq +| ⟨k, hk⟩ heq := nat.succ_ne_zero k $ ext_iff.1 heq @[simp] lemma succ_zero_eq_one : fin.succ (0 : fin (n + 1)) = 1 := rfl @@ -717,7 +761,7 @@ lemma cast_succ_pos {i : fin (n + 1)} (h : 0 < i) : 0 < cast_succ i := by simpa [lt_iff_coe_lt_coe] using h @[simp] lemma cast_succ_eq_zero_iff (a : fin (n + 1)) : a.cast_succ = 0 ↔ a = 0 := -subtype.ext_iff.trans $ (subtype.ext_iff.trans $ by exact iff.rfl).symm +fin.ext_iff.trans $ (fin.ext_iff.trans $ by exact iff.rfl).symm lemma cast_succ_ne_zero_iff (a : fin (n + 1)) : a.cast_succ ≠ 0 ↔ a ≠ 0 := not_iff_not.mpr $ cast_succ_eq_zero_iff a @@ -1216,14 +1260,14 @@ begin { simp [h] }, rw [sub_eq_add_neg, coe_add_eq_ite, coe_neg_one, if_pos, add_comm, add_tsub_add_eq_tsub_left], rw [add_comm ↑a, add_le_add_iff_left, nat.one_le_iff_ne_zero], - rwa subtype.ext_iff at h + rwa fin.ext_iff at h end lemma coe_sub_iff_le {n : ℕ} {a b : fin n} : (↑(a - b) : ℕ) = a - b ↔ b ≤ a := begin cases n, {exact fin_zero_elim a}, - rw [le_iff_coe_le_coe, fin.coe_sub, ←add_tsub_assoc_of_le b.is_lt.le], + rw [le_iff_coe_le_coe, fin.coe_sub, ←add_tsub_assoc_of_le b.is_lt.le a], cases le_or_lt (b : ℕ) a with h h, { simp [←tsub_add_eq_add_tsub h, h, nat.mod_eq_of_lt ((nat.sub_le _ _).trans_lt a.is_lt)] }, { rw [nat.mod_eq_of_lt, tsub_eq_zero_of_le h.le, tsub_eq_zero_iff_le, ←not_iff_not], @@ -1515,7 +1559,7 @@ begin end @[simp] lemma cast_pred_last : cast_pred (last (n + 1)) = last n := -by simp [eq_iff_veq, cast_pred, pred_above, cast_succ_lt_last] +eq_of_veq (by simp [cast_pred, pred_above, cast_succ_lt_last]) @[simp] lemma cast_pred_mk (n i : ℕ) (h : i < n + 1) : cast_pred ⟨i, lt_succ_of_lt h⟩ = ⟨i, h⟩ := @@ -1558,8 +1602,8 @@ begin swap 3, -- For some reason `simp` doesn't fire fully unless we discharge the third goal. { exact lt_of_le_of_ne H (ne.symm h), }, { simp, }, - { simp only [subtype.mk_eq_mk, ne.def, fin.cast_succ_mk] at h, - simp only [pred, subtype.mk_lt_mk, not_lt], + { simp only [fin.mk_eq_mk, ne.def, fin.cast_succ_mk] at h, + simp only [pred, fin.mk_lt_mk, not_lt], exact nat.le_pred_of_lt (nat.lt_of_le_and_ne H (ne.symm h)), }, }, end @@ -1693,7 +1737,7 @@ meta instance reflect : Π n, has_reflect (fin n) | (n + 1) := nat.mk_numeral `(fin n.succ) `(by apply_instance : has_zero (fin n.succ)) `(by apply_instance : has_one (fin n.succ)) - `(by apply_instance : has_add (fin n.succ)) ∘ subtype.val + `(by apply_instance : has_add (fin n.succ)) ∘ fin.val end end fin diff --git a/src/data/fin/interval.lean b/src/data/fin/interval.lean index da6c0ee9eeda9..af0d1b267af69 100644 --- a/src/data/fin/interval.lean +++ b/src/data/fin/interval.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import data.nat.interval +import data.finset.locally_finite /-! # Finite intervals in `fin n` @@ -18,8 +19,11 @@ open_locale big_operators variables (n : ℕ) -instance : locally_finite_order (fin n) := subtype.locally_finite_order _ -instance : locally_finite_order_bot (fin n) := subtype.locally_finite_order_bot _ +instance : locally_finite_order (fin n) := +order_iso.locally_finite_order fin.order_iso_subtype + +instance : locally_finite_order_bot (fin n) := +order_iso.locally_finite_order_bot fin.order_iso_subtype instance : Π n, locally_finite_order_top (fin n) | 0 := is_empty.to_locally_finite_order_top @@ -28,22 +32,22 @@ instance : Π n, locally_finite_order_top (fin n) namespace fin variables {n} (a b : fin n) -lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).subtype (λ x, x < n) := rfl -lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).subtype (λ x, x < n) := rfl -lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).subtype (λ x, x < n) := rfl -lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).subtype (λ x, x < n) := rfl +lemma Icc_eq_finset_subtype : Icc a b = (Icc (a : ℕ) b).fin n := rfl +lemma Ico_eq_finset_subtype : Ico a b = (Ico (a : ℕ) b).fin n := rfl +lemma Ioc_eq_finset_subtype : Ioc a b = (Ioc (a : ℕ) b).fin n := rfl +lemma Ioo_eq_finset_subtype : Ioo a b = (Ioo (a : ℕ) b).fin n := rfl -@[simp] lemma map_subtype_embedding_Icc : (Icc a b).map (embedding.subtype _) = Icc a b := -map_subtype_embedding_Icc _ _ _ (λ _ c x _ hx _, hx.trans_lt) +@[simp] lemma map_subtype_embedding_Icc : (Icc a b).map fin.coe_embedding = Icc a b := +by simp [Icc_eq_finset_subtype, finset.fin, finset.map_map, Icc_filter_lt_of_lt_right] -@[simp] lemma map_subtype_embedding_Ico : (Ico a b).map (embedding.subtype _) = Ico a b := -map_subtype_embedding_Ico _ _ _ (λ _ c x _ hx _, hx.trans_lt) +@[simp] lemma map_subtype_embedding_Ico : (Ico a b).map fin.coe_embedding = Ico a b := +by simp [Ico_eq_finset_subtype, finset.fin, finset.map_map] -@[simp] lemma map_subtype_embedding_Ioc : (Ioc a b).map (embedding.subtype _) = Ioc a b := -map_subtype_embedding_Ioc _ _ _ (λ _ c x _ hx _, hx.trans_lt) +@[simp] lemma map_subtype_embedding_Ioc : (Ioc a b).map fin.coe_embedding = Ioc a b := +by simp [Ioc_eq_finset_subtype, finset.fin, finset.map_map, Ioc_filter_lt_of_lt_right] -@[simp] lemma map_subtype_embedding_Ioo : (Ioo a b).map (embedding.subtype _) = Ioo a b := -map_subtype_embedding_Ioo _ _ _ (λ _ c x _ hx _, hx.trans_lt) +@[simp] lemma map_subtype_embedding_Ioo : (Ioo a b).map fin.coe_embedding = Ioo a b := +by simp [Ioo_eq_finset_subtype, finset.fin, finset.map_map] @[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := by rw [←nat.card_Icc, ←map_subtype_embedding_Icc, card_map] @@ -69,16 +73,12 @@ by rw [←card_Ioc, fintype.card_of_finset] @[simp] lemma card_fintype_Ioo : fintype.card (set.Ioo a b) = b - a - 1 := by rw [←card_Ioo, fintype.card_of_finset] -lemma Ici_eq_finset_subtype : Ici a = (Icc (a : ℕ) n).subtype (λ x, x < n) := -by { ext x, simp only [mem_subtype, mem_Ici, mem_Icc, coe_fin_le, iff_self_and], exact λ _, x.2.le } - -lemma Ioi_eq_finset_subtype : Ioi a = (Ioc (a : ℕ) n).subtype (λ x, x < n) := -by { ext x, simp only [mem_subtype, mem_Ioi, mem_Ioc, coe_fin_lt, iff_self_and], exact λ _, x.2.le } - -lemma Iic_eq_finset_subtype : Iic b = (Iic (b : ℕ)).subtype (λ x, x < n) := rfl -lemma Iio_eq_finset_subtype : Iio b = (Iio (b : ℕ)).subtype (λ x, x < n) := rfl +lemma Ici_eq_finset_subtype : Ici a = (Icc (a : ℕ) n).fin n := by { ext, simp } +lemma Ioi_eq_finset_subtype : Ioi a = (Ioc (a : ℕ) n).fin n := by { ext, simp } +lemma Iic_eq_finset_subtype : Iic b = (Iic (b : ℕ)).fin n := rfl +lemma Iio_eq_finset_subtype : Iio b = (Iio (b : ℕ)).fin n := rfl -@[simp] lemma map_subtype_embedding_Ici : (Ici a).map (embedding.subtype _) = Icc a (n - 1) := +@[simp] lemma map_subtype_embedding_Ici : (Ici a).map fin.coe_embedding = Icc a (n - 1) := begin ext x, simp only [exists_prop, embedding.coe_subtype, mem_Ici, mem_map, mem_Icc], @@ -90,7 +90,7 @@ begin { exact λ hx, ⟨⟨x, nat.lt_succ_iff.2 hx.2⟩, hx.1, rfl⟩ } end -@[simp] lemma map_subtype_embedding_Ioi : (Ioi a).map (embedding.subtype _) = Ioc a (n - 1) := +@[simp] lemma map_subtype_embedding_Ioi : (Ioi a).map fin.coe_embedding = Ioc a (n - 1) := begin ext x, simp only [exists_prop, embedding.coe_subtype, mem_Ioi, mem_map, mem_Ioc], @@ -102,11 +102,11 @@ begin { exact λ hx, ⟨⟨x, nat.lt_succ_iff.2 hx.2⟩, hx.1, rfl⟩ } end -@[simp] lemma map_subtype_embedding_Iic : (Iic b).map (embedding.subtype _) = Iic b := -map_subtype_embedding_Iic _ _ $ λ _ _, lt_of_le_of_lt +@[simp] lemma map_subtype_embedding_Iic : (Iic b).map fin.coe_embedding = Iic b := +by simp [Iic_eq_finset_subtype, finset.fin, finset.map_map, Iic_filter_lt_of_lt_right] -@[simp] lemma map_subtype_embedding_Iio : (Iio b).map (embedding.subtype _) = Iio b := -map_subtype_embedding_Iio _ _ $ λ _ _, lt_of_le_of_lt +@[simp] lemma map_subtype_embedding_Iio : (Iio b).map fin.coe_embedding = Iio b := +by simp [Iio_eq_finset_subtype, finset.fin, finset.map_map] @[simp] lemma card_Ici : (Ici a).card = n - a := by { cases n, { exact fin.elim0 a }, rw [←card_map, map_subtype_embedding_Ici, nat.card_Icc], refl } diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index cfc60e1bedab6..58987a34a94ed 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -704,7 +704,7 @@ end /-- `fin.sigma_eq_of_eq_comp_cast` as an `iff`. -/ lemma sigma_eq_iff_eq_comp_cast {α : Type*} {a b : Σ ii, fin ii → α} : a = b ↔ ∃ (h : a.fst = b.fst), a.snd = b.snd ∘ fin.cast h := -⟨λ h, h ▸ ⟨rfl, funext $ subtype.rec $ by exact λ i hi, rfl⟩, +⟨λ h, h ▸ ⟨rfl, funext $ fin.rec $ by exact λ i hi, rfl⟩, λ ⟨h, h'⟩, sigma_eq_of_eq_comp_cast _ h'⟩ end fin diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index a28db8d8bfc78..1066b69837c7c 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -2413,6 +2413,25 @@ begin convert property_of_mem_map_subtype s ha end +/-! ### Fin -/ + +/-- +Given a finset `s` of natural numbers and a bound `n`, +`s.fin n` is the finset of all elements of `s` less than `n`. +-/ +protected def fin (n : ℕ) (s : finset ℕ) : finset (fin n) := +(s.subtype _).map fin.equiv_subtype.symm.to_embedding + +@[simp] lemma mem_fin {n} {s : finset ℕ} : + ∀ a : fin n, a ∈ s.fin n ↔ (a : ℕ) ∈ s +| ⟨a, ha⟩ := by simp [finset.fin] + +@[mono] lemma fin_mono {n} : monotone (finset.fin n) := +λ s t h x, by simpa using @h x + +@[simp] lemma fin_map {n} {s : finset ℕ} : (s.fin n).map fin.coe_embedding = s.filter (< n) := +by simp [finset.fin, finset.map_map] + lemma subset_image_iff {s : set α} : ↑t ⊆ f '' s ↔ ∃ s' : finset α, ↑s' ⊆ s ∧ s'.image f = t := begin split, swap, diff --git a/src/data/finset/sort.lean b/src/data/finset/sort.lean index ff5ff884d8249..973cbe551ea8e 100644 --- a/src/data/finset/sort.lean +++ b/src/data/finset/sort.lean @@ -148,7 +148,7 @@ by simp [order_emb_of_fin, set.range_comp coe (s.order_iso_of_fin h)] /-- The bijection `order_emb_of_fin s h` sends `0` to the minimum of `s`. -/ lemma order_emb_of_fin_zero {s : finset α} {k : ℕ} (h : s.card = k) (hz : 0 < k) : order_emb_of_fin s h ⟨0, hz⟩ = s.min' (card_pos.mp (h.symm ▸ hz)) := -by simp only [order_emb_of_fin_apply, subtype.coe_mk, sorted_zero_eq_min'] +by simp only [order_emb_of_fin_apply, fin.coe_mk, sorted_zero_eq_min'] /-- The bijection `order_emb_of_fin s h` sends `k-1` to the maximum of `s`. -/ lemma order_emb_of_fin_last {s : finset α} {k : ℕ} (h : s.card = k) (hz : 0 < k) : @@ -186,7 +186,7 @@ and only if `i = j`. Since they can be defined on a priori not defeq types `fin s.order_emb_of_fin h i = s.order_emb_of_fin h' j ↔ (i : ℕ) = (j : ℕ) := begin substs k l, - exact (s.order_emb_of_fin rfl).eq_iff_eq.trans (fin.ext_iff _ _) + exact (s.order_emb_of_fin rfl).eq_iff_eq.trans fin.ext_iff end /-- Given a finset `s` of size at least `k` in a linear order `α`, the map `order_emb_of_card_le` diff --git a/src/data/fintype/card.lean b/src/data/fintype/card.lean index ab469b3a148e7..8d5532a2a5ee6 100644 --- a/src/data/fintype/card.lean +++ b/src/data/fintype/card.lean @@ -179,12 +179,17 @@ lemma equiv.prod_comp [fintype α] [fintype β] [comm_monoid γ] (e : α ≃ β) ∏ i, f (e i) = ∏ i, f i := e.bijective.prod_comp f +@[to_additive] +lemma equiv.prod_comp' [fintype α] [fintype β] [comm_monoid γ] (e : α ≃ β) (f : α → γ) (g : β → γ) + (h : ∀ i, f i = g (e i)) : ∏ i, f i = ∏ i, g i := +(show f = g ∘ e, from funext h).symm ▸ e.prod_comp _ + /-- It is equivalent to sum a function over `fin n` or `finset.range n`. -/ @[to_additive] lemma fin.prod_univ_eq_prod_range [comm_monoid α] (f : ℕ → α) (n : ℕ) : ∏ i : fin n, f i = ∏ i in range n, f i := calc (∏ i : fin n, f i) = ∏ i : {x // x ∈ range n}, f i : - (equiv.subtype_equiv_right $ λ m, (@mem_range n m).symm).prod_comp (f ∘ coe) + (fin.equiv_subtype.trans (equiv.subtype_equiv_right (by simp))).prod_comp' _ _ (by simp) ... = ∏ i in range n, f i : by rw [← attach_eq_univ, prod_attach] @[to_additive] diff --git a/src/data/list/range.lean b/src/data/list/range.lean index 365edb45ad940..fe806451fca54 100644 --- a/src/data/list/range.lean +++ b/src/data/list/range.lean @@ -219,14 +219,14 @@ by rw [← length_eq_zero, length_fin_range] @[simp] lemma map_coe_fin_range (n : ℕ) : (fin_range n).map coe = list.range n := begin - simp_rw [fin_range, map_pmap, fin.mk, subtype.coe_mk, pmap_eq_map], + simp_rw [fin_range, map_pmap, fin.coe_mk, pmap_eq_map], exact list.map_id _ end lemma fin_range_succ_eq_map (n : ℕ) : fin_range n.succ = 0 :: (fin_range n).map fin.succ := begin - apply map_injective_iff.mpr subtype.coe_injective, + apply map_injective_iff.mpr fin.coe_injective, rw [map_cons, map_coe_fin_range, range_succ_eq_map, fin.coe_zero, ←map_coe_fin_range, map_map, map_map, function.comp, function.comp], congr' 2 with x, @@ -280,7 +280,7 @@ option.some.inj $ by rw [← nth_le_nth _, nth_range (by simpa using H)] @[simp] lemma nth_le_fin_range {n : ℕ} {i : ℕ} (h) : (fin_range n).nth_le i h = ⟨i, length_fin_range n ▸ h⟩ := -by simp only [fin_range, nth_le_range, nth_le_pmap, fin.mk_eq_subtype_mk] +by simp only [fin_range, nth_le_range, nth_le_pmap] @[simp] lemma map_nth_le (l : list α) : (fin_range l.length).map (λ n, l.nth_le n n.2) = l := diff --git a/src/field_theory/finite/polynomial.lean b/src/field_theory/finite/polynomial.lean index 4dca5b3ad11e6..156caacf61b7b 100644 --- a/src/field_theory/finite/polynomial.lean +++ b/src/field_theory/finite/polynomial.lean @@ -193,7 +193,7 @@ calc module.rank K (R σ K) = ... = #(σ → {n // n < fintype.card K}) : (@equiv.subtype_pi_equiv_pi σ (λ_, ℕ) (λs n, n < fintype.card K)).cardinal_eq ... = #(σ → fin (fintype.card K)) : - (equiv.arrow_congr (equiv.refl σ) (equiv.refl _)).cardinal_eq + (equiv.arrow_congr (equiv.refl σ) fin.equiv_subtype.symm).cardinal_eq ... = #(σ → K) : (equiv.arrow_congr (equiv.refl σ) (fintype.equiv_fin K).symm).cardinal_eq ... = fintype.card (σ → K) : cardinal.mk_fintype _ diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 0d3369fc35c3d..38672cc20d204 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -503,9 +503,10 @@ classical.dec_pred _ `add_submonoid.multiples a`, sending `i` to `i • a`."] noncomputable def fin_equiv_powers [finite G] (x : G) : fin (order_of x) ≃ (submonoid.powers x : set G) := -equiv.of_bijective (λ n, ⟨x ^ ↑n, ⟨n, rfl⟩⟩) ⟨λ ⟨i, hi⟩ ⟨j, hj⟩ ij, - subtype.mk_eq_mk.2 (pow_injective_of_lt_order_of x hi hj (subtype.mk_eq_mk.1 ij)), - λ ⟨_, i, rfl⟩, ⟨⟨i % order_of x, mod_lt i (order_of_pos x)⟩, subtype.eq pow_eq_mod_order_of.symm⟩⟩ +equiv.of_bijective (λ n, ⟨x ^ ↑n, ⟨n, rfl⟩⟩) + ⟨λ ⟨i, hi⟩ ⟨j, hj⟩ ij, fin.ext (pow_injective_of_lt_order_of x hi hj (subtype.mk_eq_mk.1 ij)), + λ ⟨_, i, rfl⟩, ⟨⟨i % order_of x, mod_lt i (order_of_pos x)⟩, + subtype.eq pow_eq_mod_order_of.symm⟩⟩ @[simp, to_additive fin_equiv_multiples_apply] lemma fin_equiv_powers_apply [finite G] {x : G} {n : fin (order_of x)} : diff --git a/src/linear_algebra/matrix/block.lean b/src/linear_algebra/matrix/block.lean index 01289a61d0800..869f7e99c05df 100644 --- a/src/linear_algebra/matrix/block.lean +++ b/src/linear_algebra/matrix/block.lean @@ -83,7 +83,7 @@ begin rw [to_square_block_def', to_square_block_def], apply equiv_block_det, intro x, - apply (fin.ext_iff _ _).symm + apply fin.ext_iff.symm end /-- Let `b` map rows and columns of a square matrix `M` to `n` blocks. Then diff --git a/src/linear_algebra/matrix/transvection.lean b/src/linear_algebra/matrix/transvection.lean index 2789dbfeac6d6..96d969dcffe5a 100644 --- a/src/linear_algebra/matrix/transvection.lean +++ b/src/linear_algebra/matrix/transvection.lean @@ -362,7 +362,7 @@ begin simp only [matrix.mul_assoc, A, matrix.mul_eq_mul, list.prod_cons], by_cases h : n' = i, { have hni : n = i, - { cases i, simp only [subtype.mk_eq_mk] at h, simp [h] }, + { cases i, simp only [fin.mk_eq_mk] at h, simp [h] }, rw [h, transvection_mul_apply_same, IH, list_transvec_col_mul_last_row_drop _ _ hn, ← hni], field_simp [hM] }, { have hni : n ≠ i, @@ -411,7 +411,7 @@ begin if k ≤ i then M (inr star) (inl i) else 0, { have A : (list_transvec_row M).length = r, by simp [list_transvec_row], rw [← list.take_length (list_transvec_row M), A], - have : ¬ (r ≤ i), by simpa using i.2, + have : ¬ (r ≤ i), by simp, simpa only [this, ite_eq_right_iff] using H r le_rfl }, assume k hk, induction k with n IH, @@ -425,7 +425,7 @@ begin matrix.mul_eq_mul, list.prod_cons, list.prod_nil, option.to_list_some], by_cases h : n' = i, { have hni : n = i, - { cases i, simp only [subtype.mk_eq_mk] at h, simp only [h, coe_mk] }, + { cases i, simp only [fin.mk_eq_mk] at h, simp only [h, coe_mk] }, have : ¬ (n.succ ≤ i), by simp only [← hni, n.lt_succ_self, not_le], simp only [h, mul_transvection_apply_same, list.take, if_false, mul_list_transvec_row_last_col_take _ _ hnr.le, hni.le, this, if_true, IH hnr.le], diff --git a/src/logic/encodable/basic.lean b/src/logic/encodable/basic.lean index 19b88c9d95fd5..6b5e4ce62d351 100644 --- a/src/logic/encodable/basic.lean +++ b/src/logic/encodable/basic.lean @@ -7,6 +7,7 @@ import logic.equiv.nat import order.directed import data.countable.defs import order.rel_iso +import data.fin.basic /-! # Encodable types @@ -304,7 +305,7 @@ by cases a; refl end subtype instance _root_.fin.encodable (n) : encodable (fin n) := -subtype.encodable +of_equiv _ fin.equiv_subtype instance _root_.int.encodable : encodable ℤ := of_equiv _ equiv.int_equiv_nat diff --git a/src/logic/equiv/fin.lean b/src/logic/equiv/fin.lean index 348637dd9098c..596142c38943b 100644 --- a/src/logic/equiv/fin.lean +++ b/src/logic/equiv/fin.lean @@ -370,7 +370,7 @@ def fin_prod_fin_equiv : fin m × fin n ≃ fin (m * n) := values are retained. This is the `order_iso` version of `fin.cast_le`. -/ @[simps apply symm_apply] def fin.cast_le_order_iso {n m : ℕ} (h : n ≤ m) : fin n ≃o {i : fin m // (i : ℕ) < n} := -{ to_fun := λ i, ⟨fin.cast_le h i, by simpa using i.is_lt⟩, +{ to_fun := λ i, ⟨fin.cast_le h i, by simp⟩, inv_fun := λ i, ⟨i, i.prop⟩, left_inv := λ _, by simp, right_inv := λ _, by simp, diff --git a/src/number_theory/class_number/admissible_abs.lean b/src/number_theory/class_number/admissible_abs.lean index 91618303e24fb..53bff2817009c 100644 --- a/src/number_theory/class_number/admissible_abs.lean +++ b/src/number_theory/class_number/admissible_abs.lean @@ -43,7 +43,7 @@ begin exact int.mod_lt _ hb }, intros i₀ i₁ hi, have hi : (⌊↑(A i₀ % b) / abs b • ε⌋.nat_abs : ℤ) = ⌊↑(A i₁ % b) / abs b • ε⌋.nat_abs := - congr_arg (coe : ℕ → ℤ) (subtype.mk_eq_mk.mp hi), + congr_arg (coe : ℕ → ℤ) (fin.mk_eq_mk.mp hi), rw [nat_abs_of_nonneg (hfloor i₀), nat_abs_of_nonneg (hfloor i₁)] at hi, have hi := abs_sub_lt_one_of_floor_eq_floor hi, rw [abs_sub_comm, ← sub_div, abs_div, abs_of_nonneg hbε.le, div_lt_iff hbε, one_mul] at hi, diff --git a/src/number_theory/class_number/finite.lean b/src/number_theory/class_number/finite.lean index 3eb0d51674eee..332cfe561384c 100644 --- a/src/number_theory/class_number/finite.lean +++ b/src/number_theory/class_number/finite.lean @@ -158,7 +158,7 @@ variables [infinite R] /-- In the following results, we need a large set of distinct elements of `R`. -/ noncomputable def distinct_elems : fin (cardM bS adm).succ ↪ R := -function.embedding.trans (fin.coe_embedding _).to_embedding (infinite.nat_embedding R) +fin.coe_embedding.trans (infinite.nat_embedding R) variables [decidable_eq R] diff --git a/src/ring_theory/polynomial/basic.lean b/src/ring_theory/polynomial/basic.lean index 428e8ce20bf52..a6c95c9a9780c 100644 --- a/src/ring_theory/polynomial/basic.lean +++ b/src/ring_theory/polynomial/basic.lean @@ -127,7 +127,7 @@ def degree_lt_equiv (R) [semiring R] (n : ℕ) : degree_lt R n ≃ₗ[R] (fin n intro f, ext i, simp only [finset_sum_coeff, submodule.coe_mk], rw [finset.sum_eq_single i, coeff_monomial, if_pos rfl], - { rintro j - hji, rw [coeff_monomial, if_neg], rwa [← subtype.ext_iff] }, + { rintro j - hji, rw [coeff_monomial, if_neg], rwa [← fin.ext_iff] }, { intro h, exact (h (finset.mem_univ _)).elim } end } diff --git a/src/ring_theory/witt_vector/truncated.lean b/src/ring_theory/witt_vector/truncated.lean index 4aa9b539e664a..6bc9fee887a85 100644 --- a/src/ring_theory/witt_vector/truncated.lean +++ b/src/ring_theory/witt_vector/truncated.lean @@ -318,7 +318,7 @@ lemma mem_ker_truncate (x : 𝕎 R) : begin simp only [ring_hom.mem_ker, truncate, truncate_fun, ring_hom.coe_mk, truncated_witt_vector.ext_iff, truncated_witt_vector.coeff_mk, coeff_zero], - exact subtype.forall + exact fin.forall_iff end variables (p) diff --git a/src/testing/slim_check/sampleable.lean b/src/testing/slim_check/sampleable.lean index 27710eef65f21..c9755192c6b80 100644 --- a/src/testing/slim_check/sampleable.lean +++ b/src/testing/slim_check/sampleable.lean @@ -282,12 +282,12 @@ well_founded.fix has_well_founded.wf $ λ x f_rec, f_rec y y.property <|> some y.val . instance fin.sampleable {n : ℕ} [ne_zero n] : sampleable (fin n) := -sampleable.lift ℕ fin.of_nat' subtype.val $ +sampleable.lift ℕ fin.of_nat' fin.val $ λ i, (mod_le _ _ : i % n ≤ i) @[priority 100] instance fin.sampleable' {n} : sampleable (fin (succ n)) := -sampleable.lift ℕ fin.of_nat subtype.val $ +sampleable.lift ℕ fin.of_nat fin.val $ λ i, (mod_le _ _ : i % succ n ≤ i) instance pnat.sampleable : sampleable ℕ+ := diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index 61f22c5b4d64e..80d6d2b3311bc 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -675,12 +675,12 @@ begin have : (F p).2 ((E p) x) ((E p) y) = floor (ε⁻¹ * dist x y), by simp only [F, (E p).symm_apply_apply], have Ap : (F p).2 ⟨i, hip⟩ ⟨j, hjp⟩ = floor (ε⁻¹ * dist x y), - by { rw ← this, congr; apply (fin.ext_iff _ _).2; refl }, + by { rw ← this, congr; apply fin.ext_iff.2; refl }, -- Express `dist (Φ x) (Φ y)` in terms of `F q` have : (F q).2 ((E q) (Ψ x)) ((E q) (Ψ y)) = floor (ε⁻¹ * dist (Ψ x) (Ψ y)), by simp only [F, (E q).symm_apply_apply], have Aq : (F q).2 ⟨i, hiq⟩ ⟨j, hjq⟩ = floor (ε⁻¹ * dist (Ψ x) (Ψ y)), - by { rw ← this, congr; apply (fin.ext_iff _ _).2; [exact i', exact j'] }, + by { rw ← this, congr; apply fin.ext_iff.2; [exact i', exact j'] }, -- use the equality between `F p` and `F q` to deduce that the distances have equal -- integer parts have : (F p).2 ⟨i, hip⟩ ⟨j, hjp⟩ = (F q).2 ⟨i, hiq⟩ ⟨j, hjq⟩, @@ -763,7 +763,7 @@ begin refine ⟨_, _, (λ p, F p), _⟩, apply_instance, -- It remains to show that if `F p = F q`, then `p` and `q` are `ε`-close rintros ⟨p, pt⟩ ⟨q, qt⟩ hpq, - have Npq : N p = N q := (fin.ext_iff _ _).1 (sigma.mk.inj_iff.1 hpq).1, + have Npq : N p = N q := fin.ext_iff.1 (sigma.mk.inj_iff.1 hpq).1, let Ψ : s p → s q := λ x, (E q).symm (fin.cast Npq ((E p) x)), let Φ : s p → q.rep := λ x, Ψ x, have main : GH_dist p.rep (q.rep) ≤ ε + ε/2 + ε, @@ -817,7 +817,7 @@ begin -- Express `dist x y` in terms of `F p` have Ap : ((F p).2 ⟨i, hip⟩ ⟨j, hjp⟩).1 = ⌊ε⁻¹ * dist x y⌋₊ := calc ((F p).2 ⟨i, hip⟩ ⟨j, hjp⟩).1 = ((F p).2 ((E p) x) ((E p) y)).1 : - by { congr; apply (fin.ext_iff _ _).2; refl } + by { congr; apply fin.ext_iff.2; refl } ... = min M ⌊ε⁻¹ * dist x y⌋₊ : by simp only [F, (E p).symm_apply_apply] ... = ⌊ε⁻¹ * dist x y⌋₊ : @@ -831,7 +831,7 @@ begin -- Express `dist (Φ x) (Φ y)` in terms of `F q` have Aq : ((F q).2 ⟨i, hiq⟩ ⟨j, hjq⟩).1 = ⌊ε⁻¹ * dist (Ψ x) (Ψ y)⌋₊ := calc ((F q).2 ⟨i, hiq⟩ ⟨j, hjq⟩).1 = ((F q).2 ((E q) (Ψ x)) ((E q) (Ψ y))).1 : - by { congr; apply (fin.ext_iff _ _).2; [exact i', exact j'] } + by { congr; apply fin.ext_iff.2; [exact i', exact j'] } ... = min M ⌊ε⁻¹ * dist (Ψ x) (Ψ y)⌋₊ : by simp only [F, (E q).symm_apply_apply] ... = ⌊ε⁻¹ * dist (Ψ x) (Ψ y)⌋₊ : From 9b68ed5eeba8265959abbe1c16357047a4b2dd01 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Wed, 31 Aug 2022 20:06:31 +0000 Subject: [PATCH 121/828] refactor(topology/algebra/continuous_monoid_hom): Make variables consistent (#16304) This PR cleans up some redundant variables in `continuous_monoid_hom`. Co-authored-by: tb65536 --- .../algebra/continuous_monoid_hom.lean | 47 +++++++++---------- 1 file changed, 22 insertions(+), 25 deletions(-) diff --git a/src/topology/algebra/continuous_monoid_hom.lean b/src/topology/algebra/continuous_monoid_hom.lean index 3534df552d499..4150f256e0c8b 100644 --- a/src/topology/algebra/continuous_monoid_hom.lean +++ b/src/topology/algebra/continuous_monoid_hom.lean @@ -15,53 +15,51 @@ This file defines the space of continuous homomorphisms between two topological ## Main definitions * `continuous_monoid_hom A B`: The continuous homomorphisms `A →* B`. -* `continuous_add_monoid_hom α β`: The continuous additive homomorphisms `α →+ β`. +* `continuous_add_monoid_hom A B`: The continuous additive homomorphisms `A →+ B`. -/ open_locale pointwise open function -variables {F α β : Type*} (A B C D E : Type*) - [monoid A] [monoid B] [monoid C] [monoid D] [comm_group E] +variables (F A B C D E : Type*) [monoid A] [monoid B] [monoid C] [monoid D] [comm_group E] [topological_space A] [topological_space B] [topological_space C] [topological_space D] [topological_space E] [topological_group E] -/-- The type of continuous additive monoid homomorphisms from `α` to `β`. +/-- The type of continuous additive monoid homomorphisms from `A` to `B`. -When possible, instead of parametrizing results over `(f : continuous_add_monoid_hom α β)`, -you should parametrize over `(F : Type*) [continuous_add_monoid_hom_class F α β] (f : F)`. +When possible, instead of parametrizing results over `(f : continuous_add_monoid_hom A B)`, +you should parametrize over `(F : Type*) [continuous_add_monoid_hom_class F A B] (f : F)`. When you extend this structure, make sure to extend `continuous_add_monoid_hom_class`. -/ structure continuous_add_monoid_hom (A B : Type*) [add_monoid A] [add_monoid B] [topological_space A] [topological_space B] extends A →+ B := (continuous_to_fun : continuous to_fun) -/-- The type of continuous monoid homomorphisms from `α` to `β`. +/-- The type of continuous monoid homomorphisms from `A` to `B`. -When possible, instead of parametrizing results over `(f : continuous_monoid_hom α β)`, -you should parametrize over `(F : Type*) [continuous_monoid_hom_class F α β] (f : F)`. +When possible, instead of parametrizing results over `(f : continuous_monoid_hom A B)`, +you should parametrize over `(F : Type*) [continuous_monoid_hom_class F A B] (f : F)`. When you extend this structure, make sure to extend `continuous_add_monoid_hom_class`. -/ @[to_additive] structure continuous_monoid_hom extends A →* B := (continuous_to_fun : continuous to_fun) -/-- `continuous_add_monoid_hom_class F α β` states that `F` is a type of continuous additive monoid +/-- `continuous_add_monoid_hom_class F A B` states that `F` is a type of continuous additive monoid homomorphisms. You should also extend this typeclass when you extend `continuous_add_monoid_hom`. -/ -class continuous_add_monoid_hom_class (F α β : Type*) [add_monoid α] [add_monoid β] - [topological_space α] [topological_space β] extends add_monoid_hom_class F α β := +class continuous_add_monoid_hom_class (A B : Type*) [add_monoid A] [add_monoid B] + [topological_space A] [topological_space B] extends add_monoid_hom_class F A B := (map_continuous (f : F) : continuous f) -/-- `continuous_monoid_hom_class F α β` states that `F` is a type of continuous additive monoid +/-- `continuous_monoid_hom_class F A B` states that `F` is a type of continuous additive monoid homomorphisms. You should also extend this typeclass when you extend `continuous_monoid_hom`. -/ @[to_additive] -class continuous_monoid_hom_class (F α β : Type*) [monoid α] [monoid β] - [topological_space α] [topological_space β] extends monoid_hom_class F α β := +class continuous_monoid_hom_class extends monoid_hom_class F A B := (map_continuous (f : F) : continuous f) /-- Reinterpret a `continuous_monoid_hom` as a `monoid_hom`. -/ @@ -71,16 +69,15 @@ add_decl_doc continuous_monoid_hom.to_monoid_hom add_decl_doc continuous_add_monoid_hom.to_add_monoid_hom @[priority 100, to_additive] -- See note [lower instance priority] -instance continuous_monoid_hom_class.to_continuous_map_class [monoid α] [monoid β] - [topological_space α] [topological_space β] [continuous_monoid_hom_class F α β] : - continuous_map_class F α β := -{ .. ‹continuous_monoid_hom_class F α β› } +instance continuous_monoid_hom_class.to_continuous_map_class [continuous_monoid_hom_class F A B] : + continuous_map_class F A B := +{ .. ‹continuous_monoid_hom_class F A B› } namespace continuous_monoid_hom -variables {A B C D E} [monoid α] [monoid β] [topological_space α] [topological_space β] +variables {A B C D E} @[to_additive] -instance : continuous_monoid_hom_class (continuous_monoid_hom α β) α β := +instance : continuous_monoid_hom_class (continuous_monoid_hom A B) A B := { coe := λ f, f.to_fun, coe_injective' := λ f g h, by { obtain ⟨⟨_, _⟩, _⟩ := f, obtain ⟨⟨_, _⟩, _⟩ := g, congr' }, map_mul := λ f, f.map_mul', @@ -98,9 +95,9 @@ fun_like.ext _ _ h /-- Reinterpret a `continuous_monoid_hom` as a `continuous_map`. -/ @[to_additive "Reinterpret a `continuous_add_monoid_hom` as a `continuous_map`."] -def to_continuous_map (f : continuous_monoid_hom α β) : C(α, β) := { .. f} +def to_continuous_map (f : continuous_monoid_hom A B) : C(A, B) := { .. f} -@[to_additive] lemma to_continuous_map_injective : injective (to_continuous_map : _ → C(α, β)) := +@[to_additive] lemma to_continuous_map_injective : injective (to_continuous_map : _ → C(A, B)) := λ f g h, ext $ by convert fun_like.ext_iff.1 h /-- Construct a `continuous_monoid_hom` from a `continuous` `monoid_hom`. -/ @@ -244,6 +241,6 @@ let hi := is_inducing A E, hc := hi.continuous in end continuous_monoid_hom -/-- The Pontryagin dual of `G` is the group of continuous homomorphism `G → circle`. -/ +/-- The Pontryagin dual of `A` is the group of continuous homomorphism `A → circle`. -/ @[derive [topological_space, t2_space, comm_group, topological_group, inhabited]] -def pontryagin_dual (G : Type*) [monoid G] [topological_space G] := continuous_monoid_hom G circle +def pontryagin_dual := continuous_monoid_hom A circle From 616e6ef39709c8b1a7e9e226413ac7660a27871b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 31 Aug 2022 23:05:23 +0000 Subject: [PATCH 122/828] feat(data/sign): Allocating signs (#14335) A sum of integers of absolute value less than `n` can be broken up as a sum of less than `n` signs. --- src/data/sign.lean | 56 ++++++++++++++++++++++++++++++++++++++++++++-- 1 file changed, 54 insertions(+), 2 deletions(-) diff --git a/src/data/sign.lean b/src/data/sign.lean index 15cc885d10535..f43908af88240 100644 --- a/src/data/sign.lean +++ b/src/data/sign.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Eric Rodriguez. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Rodriguez -/ -import order.basic -import algebra.algebra.basic +import algebra.big_operators.basic +import data.fintype.card import tactic.derive_fintype /-! @@ -323,3 +323,55 @@ begin end end add_group + +namespace int + +lemma sign_eq_sign (n : ℤ) : n.sign = _root_.sign n := +begin + obtain ((_ | _) | _) := n, + { exact congr_arg coe sign_zero.symm }, + { exact congr_arg coe (sign_pos $ int.succ_coe_nat_pos _).symm }, + { exact congr_arg coe (_root_.sign_neg $ neg_succ_lt_zero _).symm } +end +end int + +open finset nat +open_locale big_operators + +private lemma exists_signed_sum_aux [decidable_eq α] (s : finset α) (f : α → ℤ) : + ∃ (β : Type u_1) (t : finset β) (sgn : β → sign_type) (g : β → α), (∀ b, g b ∈ s) ∧ + t.card = ∑ a in s, (f a).nat_abs ∧ + ∀ a ∈ s, (∑ b in t, if g b = a then (sgn b : ℤ) else 0) = f a := +begin + refine ⟨Σ a : {x // x ∈ s}, ℕ, finset.univ.sigma (λ a, range (f a).nat_abs), λ a, sign (f a.1), + λ a, a.1, λ a, a.1.prop, _, _⟩, + { simp [@sum_attach _ _ _ _ (λ a, (f a).nat_abs)] }, + { intros x hx, + simp [sum_sigma, hx, ← int.sign_eq_sign, int.sign_mul_nat_abs, mul_comm ((f _).nat_abs : ℤ), + @sum_attach _ _ _ _ (λ a, ∑ j in range (f a).nat_abs, if a = x then (f a).sign else 0)] } +end + +/-- We can decompose a sum of absolute value `n` into a sum of `n` signs. -/ +lemma exists_signed_sum [decidable_eq α] (s : finset α) (f : α → ℤ) : + ∃ (β : Type u_1) (_ : fintype β) (sgn : β → sign_type) (g : β → α), by exactI (∀ b, g b ∈ s) ∧ + fintype.card β = ∑ a in s, (f a).nat_abs ∧ + ∀ a ∈ s, (∑ b, if g b = a then (sgn b : ℤ) else 0) = f a := +let ⟨β, t, sgn, g, hg, ht, hf⟩ := exists_signed_sum_aux s f in + ⟨t, infer_instance, λ b, sgn b, λ b, g b, λ b, hg b, by simp [ht], λ a ha, + (@sum_attach _ _ t _ (λ b, ite (g b = a) (sgn b : ℤ) 0)).trans $ hf _ ha⟩ + +/-- We can decompose a sum of absolute value less than `n` into a sum of at most `n` signs. -/ +lemma exists_signed_sum' [nonempty α] [decidable_eq α] (s : finset α) (f : α → ℤ) (n : ℕ) + (h : ∑ i in s, (f i).nat_abs ≤ n) : + ∃ (β : Type u_1) (_ : fintype β) (sgn : β → sign_type) (g : β → α), by exactI + (∀ b, g b ∉ s → sgn b = 0) ∧ fintype.card β = n ∧ + ∀ a ∈ s, (∑ i, if g i = a then (sgn i : ℤ) else 0) = f a := +begin + obtain ⟨β, _, sgn, g, hg, hβ, hf⟩ := exists_signed_sum s f, + resetI, + refine ⟨β ⊕ fin (n - ∑ i in s, (f i).nat_abs), infer_instance, sum.elim sgn 0, + sum.elim g $ classical.arbitrary _, _, by simp [hβ, h], λ a ha, by simp [hf _ ha]⟩, + rintro (b | b) hb, + { cases hb (hg _) }, + { refl } +end From 370cc484a4c0f4252f96063bd6e979e23987e16c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 1 Sep 2022 05:10:37 +0000 Subject: [PATCH 123/828] chore(data/real/nnreal): Golf (#16307) Generalize various lemmas to `linear_ordered_semifield` or weaker so that they apply to `nnreal`. Golf the `nnreal` API using them. --- src/algebra/big_operators/ring.lean | 2 +- src/algebra/order/field.lean | 2 + src/algebra/order/field_defs.lean | 5 ++ src/algebra/order/group.lean | 16 ++---- src/algebra/order/monoid.lean | 19 +++++++ src/algebra/order/nonneg.lean | 3 ++ src/data/nat/choose/bounds.lean | 2 +- src/data/real/nnreal.lean | 77 +++++++++-------------------- src/order/filter/at_top_bot.lean | 37 ++++++++------ 9 files changed, 80 insertions(+), 83 deletions(-) diff --git a/src/algebra/big_operators/ring.lean b/src/algebra/big_operators/ring.lean index 7d5e7727e114e..312adc17b6e5f 100644 --- a/src/algebra/big_operators/ring.lean +++ b/src/algebra/big_operators/ring.lean @@ -70,7 +70,7 @@ by simp end semiring -lemma sum_div [division_ring β] {s : finset α} {f : α → β} {b : β} : +lemma sum_div [division_semiring β] {s : finset α} {f : α → β} {b : β} : (∑ x in s, f x) / b = ∑ x in s, f x / b := by simp only [div_eq_mul_inv, sum_mul] diff --git a/src/algebra/order/field.lean b/src/algebra/order/field.lean index 8840c4a17cd63..897665de9ccf1 100644 --- a/src/algebra/order/field.lean +++ b/src/algebra/order/field.lean @@ -451,6 +451,8 @@ end lemma one_half_lt_one : (1 / 2 : α) < 1 := half_lt_self zero_lt_one +lemma two_inv_lt_one : (2⁻¹ : α) < 1 := (one_div _).symm.trans_lt one_half_lt_one + lemma left_lt_add_div_two : a < (a + b) / 2 ↔ a < b := by simp [lt_div_iff, mul_two] lemma add_div_two_lt_right : (a + b) / 2 < b ↔ a < b := by simp [div_lt_iff, mul_two] diff --git a/src/algebra/order/field_defs.lean b/src/algebra/order/field_defs.lean index c734f15d942f3..4b7b97c0ea97f 100644 --- a/src/algebra/order/field_defs.lean +++ b/src/algebra/order/field_defs.lean @@ -56,3 +56,8 @@ instance canonically_linear_ordered_semifield.to_linear_ordered_comm_group_with_ [canonically_linear_ordered_semifield α] : linear_ordered_comm_group_with_zero α := { mul_le_mul_left := λ a b h c, mul_le_mul_of_nonneg_left h $ zero_le _, ..‹canonically_linear_ordered_semifield α› } + +@[priority 100] -- See note [lower instance priority] +instance canonically_linear_ordered_semifield.to_canonically_linear_ordered_add_monoid + [canonically_linear_ordered_semifield α] : canonically_linear_ordered_add_monoid α := +{ ..‹canonically_linear_ordered_semifield α› } diff --git a/src/algebra/order/group.lean b/src/algebra/order/group.lean index e41cf98dcf8ab..2b2c30317c7cd 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group.lean @@ -58,11 +58,9 @@ instance ordered_comm_group.to_ordered_cancel_comm_monoid (α : Type u) le_of_mul_le_mul_left := λ a b c, (mul_le_mul_iff_left a).mp, ..s } -@[priority 100, to_additive] -instance ordered_comm_group.has_exists_mul_of_le (α : Type u) - [ordered_comm_group α] : - has_exists_mul_of_le α := -⟨λ a b hab, ⟨b * a⁻¹, (mul_inv_cancel_comm_assoc a b).symm⟩⟩ +@[priority 100, to_additive] -- See note [lower instance priority] +instance group.has_exists_mul_of_le (α : Type u) [group α] [has_le α] : has_exists_mul_of_le α := +⟨λ a b hab, ⟨a⁻¹ * b, (mul_inv_cancel_left _ _).symm⟩⟩ @[to_additive] instance [ordered_comm_group α] : ordered_comm_group αᵒᵈ := @@ -818,15 +816,9 @@ end variable_names section densely_ordered variables [densely_ordered α] {a b c : α} -@[to_additive] -lemma le_of_forall_one_lt_le_mul (h : ∀ ε : α, 1 < ε → a ≤ b * ε) : a ≤ b := -le_of_forall_le_of_dense $ λ c hc, -calc a ≤ b * (b⁻¹ * c) : h _ (lt_inv_mul_iff_lt.mpr hc) - ... = c : mul_inv_cancel_left b c - @[to_additive] lemma le_of_forall_lt_one_mul_le (h : ∀ ε < 1, a * ε ≤ b) : a ≤ b := -@le_of_forall_one_lt_le_mul αᵒᵈ _ _ _ _ _ _ h +@le_of_forall_one_lt_le_mul αᵒᵈ _ _ _ _ _ _ _ _ h @[to_additive] lemma le_of_forall_one_lt_div_le (h : ∀ ε : α, 1 < ε → a / ε ≤ b) : a ≤ b := diff --git a/src/algebra/order/monoid.lean b/src/algebra/order/monoid.lean index f80005deee113..3dfe4abb65071 100644 --- a/src/algebra/order/monoid.lean +++ b/src/algebra/order/monoid.lean @@ -95,6 +95,25 @@ export has_exists_mul_of_le (exists_mul_of_le) export has_exists_add_of_le (exists_add_of_le) +section has_exists_mul_of_le +variables [linear_order α] [densely_ordered α] [monoid α] [has_exists_mul_of_le α] + [covariant_class α α (*) (<)] [contravariant_class α α (*) (<)] {a b : α} + +@[to_additive] +lemma le_of_forall_one_lt_le_mul (h : ∀ ε : α, 1 < ε → a ≤ b * ε) : a ≤ b := +le_of_forall_le_of_dense $ λ x hxb, by { obtain ⟨ε, rfl⟩ := exists_mul_of_le hxb.le, + exact h _ ((lt_mul_iff_one_lt_right' b).1 hxb) } + +@[to_additive] +lemma le_of_forall_one_lt_lt_mul' (h : ∀ ε : α, 1 < ε → a < b * ε) : a ≤ b := +le_of_forall_one_lt_le_mul $ λ ε hε, (h _ hε).le + +@[to_additive] +lemma le_iff_forall_one_lt_lt_mul' : a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε := +⟨λ h ε, lt_mul_of_le_of_one_lt h, le_of_forall_one_lt_lt_mul'⟩ + +end has_exists_mul_of_le + /-- A linearly ordered additive commutative monoid. -/ @[protect_proj, ancestor linear_order ordered_add_comm_monoid] class linear_ordered_add_comm_monoid (α : Type*) diff --git a/src/algebra/order/nonneg.lean b/src/algebra/order/nonneg.lean index 1fe76aca8d83f..c45acf1280d9a 100644 --- a/src/algebra/order/nonneg.lean +++ b/src/algebra/order/nonneg.lean @@ -202,6 +202,9 @@ by apply_instance instance comm_monoid_with_zero [ordered_comm_semiring α] : comm_monoid_with_zero {x : α // 0 ≤ x} := by apply_instance +instance semiring [ordered_semiring α] : semiring {x : α // 0 ≤ x} := infer_instance +instance comm_semiring [ordered_comm_semiring α] : comm_semiring {x : α // 0 ≤ x} := infer_instance + instance nontrivial [linear_ordered_semiring α] : nontrivial {x : α // 0 ≤ x} := ⟨ ⟨0, 1, λ h, zero_ne_one (congr_arg subtype.val h)⟩ ⟩ diff --git a/src/data/nat/choose/bounds.lean b/src/data/nat/choose/bounds.lean index 433ee226684ad..113353b97d516 100644 --- a/src/data/nat/choose/bounds.lean +++ b/src/data/nat/choose/bounds.lean @@ -23,7 +23,7 @@ bounds `n^r/r^r ≤ n.choose r ≤ e^r n^r/r^r` in the future. open_locale nat -variables {α : Type*} [linear_ordered_field α] +variables {α : Type*} [linear_ordered_semifield α] namespace nat diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 85b7f462c783d..12a5b2942c8f6 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -46,10 +46,6 @@ of `x` with `↑x`. This tactic also works for a function `f : α → ℝ` with ## Notations This file defines `ℝ≥0` as a localized notation for `nnreal`. - -## TODO - -`semifield` instance -/ open_locale classical big_operators @@ -57,9 +53,9 @@ open_locale classical big_operators /-- Nonnegative real numbers. -/ @[derive [ ordered_semiring, comm_monoid_with_zero, -- to ensure these instance are computable - floor_semiring, + floor_semiring, comm_semiring, semiring, semilattice_inf, densely_ordered, order_bot, - canonically_linear_ordered_add_monoid, linear_ordered_comm_group_with_zero, archimedean, + canonically_linear_ordered_semifield, linear_ordered_comm_group_with_zero, archimedean, linear_ordered_semiring, ordered_comm_semiring, canonically_ordered_comm_semiring, has_sub, has_ordered_sub, has_div, inhabited]] def nnreal := {r : ℝ // 0 ≤ r} @@ -130,7 +126,6 @@ protected lemma coe_two : ((2 : ℝ≥0) : ℝ) = 2 := rfl ((r₁ - r₂ : ℝ≥0) : ℝ) = r₁ - r₂ := max_eq_left $ le_sub.2 $ by simp [show (r₂ : ℝ) ≤ r₁, from h] --- TODO: setup semifield! @[simp, norm_cast] protected lemma coe_eq_zero (r : ℝ≥0) : ↑r = (0 : ℝ) ↔ r = 0 := by rw [← nnreal.coe_zero, nnreal.coe_eq] @@ -354,12 +349,9 @@ ordered_cancel_add_comm_monoid.to_contravariant_class_left ℝ≥0 instance covariant_mul : covariant_class ℝ≥0 ℝ≥0 (*) (≤) := ordered_comm_monoid.to_covariant_class_left ℝ≥0 +-- Why isn't `nnreal.contravariant_add` inferred? lemma le_of_forall_pos_le_add {a b : ℝ≥0} (h : ∀ε, 0 < ε → a ≤ b + ε) : a ≤ b := -le_of_forall_le_of_dense $ assume x hxb, -begin - rcases exists_add_of_le (le_of_lt hxb) with ⟨ε, rfl⟩, - exact h _ ((lt_add_iff_pos_right b).1 hxb) -end +@le_of_forall_pos_le_add _ _ _ _ _ _ nnreal.contravariant_add _ _ h lemma lt_iff_exists_rat_btwn (a b : ℝ≥0) : a < b ↔ (∃q:ℚ, 0 ≤ q ∧ a < real.to_nnreal q ∧ real.to_nnreal q < b) := @@ -502,11 +494,7 @@ namespace nnreal section mul lemma mul_eq_mul_left {a b c : ℝ≥0} (h : a ≠ 0) : (a * b = a * c ↔ b = c) := -begin - rw [← nnreal.eq_iff, ← nnreal.eq_iff, nnreal.coe_mul, nnreal.coe_mul], split, - { exact mul_left_cancel₀ (mt (@nnreal.eq_iff a 0).1 h) }, - { assume h, rw [h] } -end +by rw [mul_eq_mul_left_iff, or_iff_left h] lemma _root_.real.to_nnreal_mul {p q : ℝ} (hp : 0 ≤ p) : real.to_nnreal (p * q) = real.to_nnreal p * real.to_nnreal q := @@ -566,8 +554,7 @@ lemma coe_sub_def {r p : ℝ≥0} : ↑(r - p) = max (r - p : ℝ) 0 := rfl noncomputable example : has_ordered_sub ℝ≥0 := by apply_instance -lemma sub_div (a b c : ℝ≥0) : (a - b) / c = a / c - b / c := -by simp only [div_eq_mul_inv, tsub_mul] +lemma sub_div (a b c : ℝ≥0) : (a - b) / c = a / c - b / c := tsub_div _ _ _ end sub @@ -575,15 +562,13 @@ section inv lemma sum_div {ι} (s : finset ι) (f : ι → ℝ≥0) (b : ℝ≥0) : (∑ i in s, f i) / b = ∑ i in s, (f i / b) := -by simp only [div_eq_mul_inv, finset.sum_mul] +finset.sum_div -@[simp] lemma inv_pos {r : ℝ≥0} : 0 < r⁻¹ ↔ 0 < r := -by simp [pos_iff_ne_zero] +@[simp] lemma inv_pos {r : ℝ≥0} : 0 < r⁻¹ ↔ 0 < r := inv_pos -lemma div_pos {r p : ℝ≥0} (hr : 0 < r) (hp : 0 < p) : 0 < r / p := -by simpa only [div_eq_mul_inv] using mul_pos hr (inv_pos.2 hp) +lemma div_pos {r p : ℝ≥0} (hr : 0 < r) (hp : 0 < p) : 0 < r / p := div_pos hr hp -lemma div_self_le (r : ℝ≥0) : r / r ≤ 1 := div_self_le_one (r : ℝ) +lemma div_self_le (r : ℝ≥0) : r / r ≤ 1 := div_self_le_one r @[simp] lemma inv_le {r p : ℝ≥0} (h : r ≠ 0) : r⁻¹ ≤ p ↔ 1 ≤ r * p := by rw [← mul_le_mul_left (pos_iff_ne_zero.2 h), mul_inv_cancel h] @@ -601,11 +586,9 @@ lemma mul_le_iff_le_inv {a b r : ℝ≥0} (hr : r ≠ 0) : r * a ≤ b ↔ a ≤ have 0 < r, from lt_of_le_of_ne (zero_le r) hr.symm, by rw [← mul_le_mul_left (inv_pos.mpr this), ← mul_assoc, inv_mul_cancel hr, one_mul] -lemma le_div_iff_mul_le {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ a * r ≤ b := -by rw [div_eq_inv_mul, ← mul_le_iff_le_inv hr, mul_comm] +lemma le_div_iff_mul_le {a b r : ℝ≥0} (hr : r ≠ 0) : a ≤ b / r ↔ a * r ≤ b := le_div_iff₀ hr -lemma div_le_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a / r ≤ b ↔ a ≤ b * r := -@div_le_iff ℝ _ a r b $ pos_iff_ne_zero.2 hr +lemma div_le_iff {a b r : ℝ≥0} (hr : r ≠ 0) : a / r ≤ b ↔ a ≤ b * r := div_le_iff₀ hr lemma div_le_iff' {a b r : ℝ≥0} (hr : r ≠ 0) : a / r ≤ b ↔ a ≤ r * b := @div_le_iff' ℝ _ a r b $ pos_iff_ne_zero.2 hr @@ -653,8 +636,7 @@ end lemma div_le_div_left {a b c : ℝ≥0} (a0 : 0 < a) (b0 : 0 < b) (c0 : 0 < c) : a / b ≤ a / c ↔ c ≤ b := -by rw [nnreal.div_le_iff b0.ne.symm, div_mul_eq_mul_div, nnreal.le_div_iff_mul_le c0.ne.symm, - mul_le_mul_left a0] +div_le_div_left a0 b0 c0 lemma le_of_forall_lt_one_mul_le {x y : ℝ≥0} (h : ∀a<1, a * x ≤ y) : x ≤ y := le_of_forall_ge_of_dense $ assume a ha, @@ -664,21 +646,17 @@ le_of_forall_ge_of_dense $ assume a ha, have (a * x⁻¹) * x ≤ y, from h _ this, by rwa [mul_assoc, inv_mul_cancel hx, mul_one] at this -lemma div_add_div_same (a b c : ℝ≥0) : a / c + b / c = (a + b) / c := -eq.symm $ right_distrib a b (c⁻¹) +lemma div_add_div_same (a b c : ℝ≥0) : a / c + b / c = (a + b) / c := div_add_div_same _ _ _ -lemma half_pos {a : ℝ≥0} (h : 0 < a) : 0 < a / 2 := div_pos h zero_lt_two +lemma half_pos {a : ℝ≥0} (h : 0 < a) : 0 < a / 2 := half_pos h -lemma add_halves (a : ℝ≥0) : a / 2 + a / 2 = a := nnreal.eq (add_halves a) +lemma add_halves (a : ℝ≥0) : a / 2 + a / 2 = a := add_halves _ -lemma half_le_self (a : ℝ≥0) : a / 2 ≤ a := nnreal.coe_le_coe.mp $ half_le_self a.coe_nonneg +lemma half_le_self (a : ℝ≥0) : a / 2 ≤ a := half_le_self bot_le -lemma half_lt_self {a : ℝ≥0} (h : a ≠ 0) : a / 2 < a := -by rw [← nnreal.coe_lt_coe, nnreal.coe_div]; exact -half_lt_self (bot_lt_iff_ne_bot.2 h) +lemma half_lt_self {a : ℝ≥0} (h : a ≠ 0) : a / 2 < a := half_lt_self h.bot_lt -lemma two_inv_lt_one : (2⁻¹:ℝ≥0) < 1 := -by simpa using half_lt_self zero_ne_one.symm +lemma two_inv_lt_one : (2⁻¹:ℝ≥0) < 1 := two_inv_lt_one lemma div_lt_one_of_lt {a b : ℝ≥0} (h : a < b) : a / b < 1 := begin @@ -688,19 +666,15 @@ end @[field_simps] lemma div_add_div (a : ℝ≥0) {b : ℝ≥0} (c : ℝ≥0) {d : ℝ≥0} (hb : b ≠ 0) (hd : d ≠ 0) : a / b + c / d = (a * d + b * c) / (b * d) := -begin - rw ← nnreal.eq_iff, - simp only [nnreal.coe_add, nnreal.coe_div, nnreal.coe_mul], - exact div_add_div _ _ (coe_ne_zero.2 hb) (coe_ne_zero.2 hd) -end +div_add_div _ _ hb hd @[field_simps] lemma add_div' (a b c : ℝ≥0) (hc : c ≠ 0) : b + a / c = (b * c + a) / c := -by simpa using div_add_div b a one_ne_zero hc +add_div' _ _ _ hc @[field_simps] lemma div_add' (a b c : ℝ≥0) (hc : c ≠ 0) : a / c + b = (a + b * c) / c := -by rwa [add_comm, add_div', add_comm] +div_add' _ _ _ hc lemma _root_.real.to_nnreal_inv {x : ℝ} : real.to_nnreal x⁻¹ = (real.to_nnreal x)⁻¹ := @@ -723,8 +697,7 @@ by rw [div_eq_inv_mul, div_eq_inv_mul, real.to_nnreal_mul (inv_nonneg.2 hy), rea lemma inv_lt_one_iff {x : ℝ≥0} (hx : x ≠ 0) : x⁻¹ < 1 ↔ 1 < x := by rwa [← one_div, div_lt_iff hx, one_mul] -lemma inv_lt_one {x : ℝ≥0} (hx : 1 < x) : x⁻¹ < 1 := -(inv_lt_one_iff (zero_lt_one.trans hx).ne').2 hx +lemma inv_lt_one {x : ℝ≥0} (hx : 1 < x) : x⁻¹ < 1 := inv_lt_one hx lemma zpow_pos {x : ℝ≥0} (hx : x ≠ 0) (n : ℤ) : 0 < x ^ n := begin @@ -733,9 +706,7 @@ begin { simp [pow_pos hx.bot_lt _] } end -lemma inv_lt_inv_iff {x y : ℝ≥0} (hx : x ≠ 0) (hy : y ≠ 0) : - y⁻¹ < x⁻¹ ↔ x < y := -by rw [← one_div, div_lt_iff hy, ← div_eq_inv_mul, lt_div_iff hx, one_mul] +lemma inv_lt_inv_iff {x y : ℝ≥0} (hx : x ≠ 0) (hy : y ≠ 0) : y⁻¹ < x⁻¹ ↔ x < y := inv_lt_inv₀ hy hx lemma inv_lt_inv {x y : ℝ≥0} (hx : x ≠ 0) (h : x < y) : y⁻¹ < x⁻¹ := (inv_lt_inv_iff hx ((bot_le.trans_lt h).ne')).2 h diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index 620f0ef5b43f0..19d23a18ec819 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -788,9 +788,9 @@ lemma nonneg_of_eventually_pow_nonneg [linear_ordered_ring α] {a : α} (h : ∀ᶠ n in at_top, 0 ≤ a ^ (n : ℕ)) : 0 ≤ a := let ⟨n, hn⟩ := (tendsto_bit1_at_top.eventually h).exists in pow_bit1_nonneg_iff.1 hn -section linear_ordered_field +section linear_ordered_semifield -variables [linear_ordered_field α] {l : filter β} {f : β → α} {r : α} +variables [linear_ordered_semifield α] {l : filter β} {f : β → α} {r c : α} {n : ℕ} /-- If a function tends to infinity along a filter, then this function multiplied by a positive constant (on the left) also tends to infinity. For a version working in `ℕ` or `ℤ`, use @@ -812,6 +812,25 @@ lemma tendsto.at_top_div_const (hr : 0 < r) (hf : tendsto f l at_top) : tendsto (λx, f x / r) l at_top := by simpa only [div_eq_mul_inv] using hf.at_top_mul_const (inv_pos.2 hr) +lemma tendsto_const_mul_pow_at_top (hn : n ≠ 0) (hc : 0 < c) : + tendsto (λ x, c * x^n) at_top at_top := +tendsto.const_mul_at_top hc (tendsto_pow_at_top hn) + +lemma tendsto_const_mul_pow_at_top_iff : + tendsto (λ x, c * x^n) at_top at_top ↔ n ≠ 0 ∧ 0 < c := +begin + refine ⟨λ h, ⟨_, _⟩, λ h, tendsto_const_mul_pow_at_top h.1 h.2⟩, + { rintro rfl, + simpa only [pow_zero, not_tendsto_const_at_top] using h }, + { rcases ((h.eventually_gt_at_top 0).and (eventually_ge_at_top 0)).exists with ⟨k, hck, hk⟩, + exact pos_of_mul_pos_left hck (pow_nonneg hk _) }, +end + +end linear_ordered_semifield + +section linear_ordered_field +variables [linear_ordered_field α] {l : filter β} {f : β → α} {r : α} + /-- If a function tends to infinity along a filter, then this function multiplied by a negative constant (on the left) tends to negative infinity. -/ lemma tendsto.neg_const_mul_at_top (hr : r < 0) (hf : tendsto f l at_top) : @@ -857,20 +876,6 @@ lemma tendsto.at_bot_mul_neg_const (hr : r < 0) (hf : tendsto f l at_bot) : tendsto (λ x, f x * r) l at_top := by simpa only [mul_comm] using hf.neg_const_mul_at_bot hr -lemma tendsto_const_mul_pow_at_top {c : α} {n : ℕ} - (hn : n ≠ 0) (hc : 0 < c) : tendsto (λ x, c * x^n) at_top at_top := -tendsto.const_mul_at_top hc (tendsto_pow_at_top hn) - -lemma tendsto_const_mul_pow_at_top_iff {c : α} {n : ℕ} : - tendsto (λ x, c * x^n) at_top at_top ↔ n ≠ 0 ∧ 0 < c := -begin - refine ⟨λ h, ⟨_, _⟩, λ h, tendsto_const_mul_pow_at_top h.1 h.2⟩, - { rintro rfl, - simpa only [pow_zero, not_tendsto_const_at_top] using h }, - { rcases ((h.eventually_gt_at_top 0).and (eventually_ge_at_top 0)).exists with ⟨k, hck, hk⟩, - exact pos_of_mul_pos_left hck (pow_nonneg hk _) }, -end - lemma tendsto_neg_const_mul_pow_at_top {c : α} {n : ℕ} (hn : n ≠ 0) (hc : c < 0) : tendsto (λ x, c * x^n) at_top at_bot := tendsto.neg_const_mul_at_top hc (tendsto_pow_at_top hn) From 8d90fcdb89bcb549b41edc40418ac40e045ce523 Mon Sep 17 00:00:00 2001 From: Wrenna Robson Date: Thu, 1 Sep 2022 09:13:58 +0000 Subject: [PATCH 124/828] feat(data/finset/basic): Add `coe_pair` lemmas (#16235) Adds `coe_pair` lemmas to go with `coe_singleton` and `coe_eq_singleton` (and fixes a couple of style issues). --- src/data/finset/basic.lean | 17 ++++++++++------- 1 file changed, 10 insertions(+), 7 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 1066b69837c7c..0f7306821f6e0 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -446,9 +446,8 @@ singleton_injective.eq_iff @[simp, norm_cast] lemma coe_singleton (a : α) : (({a} : finset α) : set α) = {a} := by { ext, simp } -@[simp, norm_cast] lemma coe_eq_singleton {α : Type*} {s : finset α} {a : α} : - (s : set α) = {a} ↔ s = {a} := -by rw [←finset.coe_singleton, finset.coe_inj] +@[simp, norm_cast] lemma coe_eq_singleton {s : finset α} {a : α} : (s : set α) = {a} ↔ s = {a} := +by rw [←coe_singleton, coe_inj] lemma eq_singleton_iff_unique_mem {s : finset α} {a : α} : s = {a} ↔ a ∈ s ∧ ∀ x ∈ s, x = a := @@ -611,8 +610,7 @@ lemma mem_of_mem_insert_of_ne (h : b ∈ insert a s) : b ≠ a → b ∈ s := (m lemma eq_of_not_mem_of_mem_insert (ha : b ∈ insert a s) (hb : b ∉ s) : b = a := (mem_insert.1 ha).resolve_right hb -@[simp] theorem cons_eq_insert {α} [decidable_eq α] (a s h) : @cons α a s h = insert a s := -ext $ λ a, by simp +@[simp] theorem cons_eq_insert (a s h) : @cons α a s h = insert a s := ext $ λ a, by simp @[simp, norm_cast] lemma coe_insert (a : α) (s : finset α) : ↑(insert a s) = (insert a s : set α) := @@ -636,8 +634,13 @@ insert_eq_of_mem $ mem_singleton_self _ theorem insert.comm (a b : α) (s : finset α) : insert a (insert b s) = insert b (insert a s) := ext $ λ x, by simp only [mem_insert, or.left_comm] -theorem pair_comm (a b : α) : ({a, b} : finset α) = {b, a} := -insert.comm a b ∅ +@[simp, norm_cast] lemma coe_pair {a b : α} : + (({a, b} : finset α) : set α) = {a, b} := by { ext, simp } + +@[simp, norm_cast] lemma coe_eq_pair {s : finset α} {a b : α} : + (s : set α) = {a, b} ↔ s = {a, b} := by rw [←coe_pair, coe_inj] + +theorem pair_comm (a b : α) : ({a, b} : finset α) = {b, a} := insert.comm a b ∅ @[simp] theorem insert_idem (a : α) (s : finset α) : insert a (insert a s) = insert a s := ext $ λ x, by simp only [mem_insert, or.assoc.symm, or_self] From 44c94be80d1a9c267e6289a24299cd3ae0a2ded9 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 1 Sep 2022 11:45:52 +0000 Subject: [PATCH 125/828] feat(data/finsupp/lex): lexicographic orders on finsupps and covariant classes (#16127) This PR constructs a linear order on `finsupp`s where both source and target are linearly ordered. This is useful for #15983, where these linear orders help prove that (some) `add_monoid_algebra`s have no non-trivial zero-divisors. The PR also proves that the lexicographic linear order on finitely supported functions preserves a few covariant class assumptions. As always, comments and suggestions are really really appreciated! [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/linear.20order.20on.20finsupps) This PR supersedes #15984. --- src/data/finsupp/lex.lean | 147 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 147 insertions(+) create mode 100644 src/data/finsupp/lex.lean diff --git a/src/data/finsupp/lex.lean b/src/data/finsupp/lex.lean new file mode 100644 index 0000000000000..9178940be7cae --- /dev/null +++ b/src/data/finsupp/lex.lean @@ -0,0 +1,147 @@ +/- +Copyright (c) 2022 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa +-/ +import data.pi.lex +import data.finsupp.order +import data.finsupp.ne_locus + +/-! +# Lexicographic order on finitely supported functions + +This file defines the lexicographic order on `finsupp`. +-/ + +variables {α N : Type*} + +namespace finsupp + +section N_has_zero +variables [has_zero N] + +/-- `finsupp.lex r s` is the lexicographic relation on `α →₀ N`, where `α` is ordered by `r`, +and `N` is ordered by `s`. + +The type synonym `_root_.lex (α →₀ N)` has an order given by `finsupp.lex (<) (<)`. +-/ +protected def lex (r : α → α → Prop) (s : N → N → Prop) (x y : α →₀ N) : Prop := +pi.lex r (λ _, s) x y + +lemma _root_.pi.lex_eq_finsupp_lex {r : α → α → Prop} {s : N → N → Prop} (a b : α →₀ N) : + pi.lex r (λ _, s) (a : α → N) (b : α → N) = finsupp.lex r s a b := +rfl + +lemma lex_def {r : α → α → Prop} {s : N → N → Prop} {a b : α →₀ N} : + finsupp.lex r s a b ↔ ∃ j, (∀ d, r d j → a d = b d) ∧ s (a j) (b j) := iff.rfl + +instance [has_lt α] [has_lt N] : has_lt (lex (α →₀ N)) := +⟨λ f g, finsupp.lex (<) (<) (of_lex f) (of_lex g)⟩ + +instance lex.is_strict_order [linear_order α] [partial_order N] : + is_strict_order (lex (α →₀ N)) (<) := +let i : is_strict_order (lex (α → N)) (<) := pi.lex.is_strict_order in +{ irrefl := to_lex.surjective.forall.2 $ λ a, @irrefl _ _ i.to_is_irrefl a, + trans := to_lex.surjective.forall₃.2 $ λ a b c, @trans _ _ i.to_is_trans a b c } + +variables [linear_order α] + +/-- The partial order on `finsupp`s obtained by the lexicographic ordering. +See `finsupp.lex.linear_order` for a proof that this partial order is in fact linear. -/ +instance lex.partial_order [partial_order N] : partial_order (lex (α →₀ N)) := +partial_order.lift (λ x, to_lex ⇑(of_lex x)) finsupp.coe_fn_injective--fun_like.coe_injective + +variable [linear_order N] + +/-- Auxiliary helper to case split computably. There is no need for this to be public, as it +can be written with `or.by_cases` on `lt_trichotomy` once the instances below are constructed. -/ +private def lt_trichotomy_rec {P : lex (α →₀ N) → lex (α →₀ N) → Sort*} + (h_lt : Π {f g}, to_lex f < to_lex g → P (to_lex f) (to_lex g)) + (h_eq : Π {f g}, to_lex f = to_lex g → P (to_lex f) (to_lex g)) + (h_gt : Π {f g}, to_lex g < to_lex f → P (to_lex f) (to_lex g)) : + ∀ f g, P f g := +lex.rec $ λ f, lex.rec $ λ g, + match _, rfl : ∀ y, (f.ne_locus g).min = y → _ with + | ⊤, h := h_eq (finsupp.ne_locus_eq_empty.mp (finset.min_eq_top.mp h)) + | (wit : α), h := + have hne : f wit ≠ g wit := mem_ne_locus.mp (finset.mem_of_min h), + hne.lt_or_lt.by_cases + (λ hwit, h_lt ⟨wit, λ j hj, mem_ne_locus.not_left.mp (finset.not_mem_of_lt_min hj h), hwit⟩) + (λ hwit, h_gt ⟨wit, by exact λ j hj, begin + refine mem_ne_locus.not_left.mp (finset.not_mem_of_lt_min hj _), + rwa ne_locus_comm, + end, hwit⟩) + end + +@[irreducible] instance lex.decidable_le : @decidable_rel (lex (α →₀ N)) (≤) := +lt_trichotomy_rec + (λ f g h, is_true $ or.inr h) + (λ f g h, is_true $ or.inl $ congr_arg _ h) + (λ f g h, is_false $ λ h', (lt_irrefl _ (h.trans_le h')).elim) + +@[irreducible] instance lex.decidable_lt : @decidable_rel (lex (α →₀ N)) (<) := +lt_trichotomy_rec + (λ f g h, is_true h) + (λ f g h, is_false h.not_lt) + (λ f g h, is_false h.asymm) + +/-- The linear order on `finsupp`s obtained by the lexicographic ordering. -/ +instance lex.linear_order : linear_order (lex (α →₀ N)) := +{ le_total := lt_trichotomy_rec + (λ f g h, or.inl h.le) + (λ f g h, or.inl h.le) + (λ f g h, or.inr h.le), + decidable_lt := by apply_instance, + decidable_le := by apply_instance, + decidable_eq := by apply_instance, + ..lex.partial_order } + +lemma lex.le_of_forall_le {a b : lex (α →₀ N)} (h : ∀ i, of_lex a i ≤ of_lex b i) : a ≤ b := +le_of_not_lt (λ ⟨i, hi⟩, (h i).not_lt hi.2) + +lemma lex.le_of_of_lex_le {a b : lex (α →₀ N)} (h : of_lex a ≤ of_lex b) : a ≤ b := +lex.le_of_forall_le h + +lemma to_lex_monotone : monotone (@to_lex (α →₀ N)) := +λ _ _, lex.le_of_forall_le + +lemma lt_of_forall_lt_of_lt (a b : lex (α →₀ N)) (i : α) : + (∀ j < i, of_lex a j = of_lex b j) → of_lex a i < of_lex b i → a < b := +λ h1 h2, ⟨i, h1, h2⟩ + +end N_has_zero + +section covariants +variables [linear_order α] [add_monoid N] [linear_order N] + +/-! We are about to sneak in a hypothesis that might appear to be too strong. +We assume `covariant_class` with *strict* inequality `<` also when proving the one with the +*weak* inequality `≤`. This is actually necessary: addition on `lex (α →₀ N)` may fail to be +monotone, when it is "just" monotone on `N`. -/ +section left +variables [covariant_class N N (+) (<)] + +instance lex.covariant_class_lt_left : covariant_class (lex (α →₀ N)) (lex (α →₀ N)) (+) (<) := +⟨λ f g h ⟨a, lta, ha⟩, ⟨a, λ j ja, congr_arg ((+) _) (lta j ja), add_lt_add_left ha _⟩⟩ + +instance lex.covariant_class_le_left : covariant_class (lex (α →₀ N)) (lex (α →₀ N)) (+) (≤) := +has_add.to_covariant_class_left _ + +end left + +section right +variables [covariant_class N N (function.swap (+)) (<)] + +instance lex.covariant_class_lt_right : + covariant_class (lex (α →₀ N)) (lex (α →₀ N)) (function.swap (+)) (<) := +⟨λ f g h ⟨a, lta, ha⟩, ⟨a, λ j ja, congr_arg (+ (of_lex f j)) (lta j ja), add_lt_add_right ha _⟩⟩ + +instance lex.covariant_class_le_right : + covariant_class (lex (α →₀ N)) (lex (α →₀ N)) (function.swap (+)) (≤) := +has_add.to_covariant_class_right _ + +end right + +end covariants + +end finsupp From 394f6e63f63ebc49b2b723e62f89f1604aa4b87d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 1 Sep 2022 11:45:54 +0000 Subject: [PATCH 126/828] =?UTF-8?q?feat(measure=5Ftheory/integral/bochner)?= =?UTF-8?q?:=20H=C3=B6lder's=20inequality=20for=20real=20nonnegative=20fun?= =?UTF-8?q?ctions=20(#16291)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We already have several versions of Hölder's inequality for the Lebesgue integral and functions with values in nnreal or ennreal. This PR introduces two versions of this inequality for the Bochner integral, one of which applies to nonnegative real functions. --- src/measure_theory/function/ess_sup.lean | 3 + src/measure_theory/function/lp_space.lean | 107 ++++++++++++++++++++++ src/measure_theory/integral/bochner.lean | 74 +++++++++++++++ src/order/filter/ennreal.lean | 11 +++ 4 files changed, 195 insertions(+) diff --git a/src/measure_theory/function/ess_sup.lean b/src/measure_theory/function/ess_sup.lean index 1737c1e06ffb2..1f8e498fb7586 100644 --- a/src/measure_theory/function/ess_sup.lean +++ b/src/measure_theory/function/ess_sup.lean @@ -252,6 +252,9 @@ limsup_eq_zero_iff lemma ess_sup_const_mul {a : ℝ≥0∞} : ess_sup (λ (x : α), a * (f x)) μ = a * ess_sup f μ := limsup_const_mul +lemma ess_sup_mul_le (f g : α → ℝ≥0∞) : ess_sup (f * g) μ ≤ ess_sup f μ * ess_sup g μ := +limsup_mul_le f g + lemma ess_sup_add_le (f g : α → ℝ≥0∞) : ess_sup (f + g) μ ≤ ess_sup f μ + ess_sup g μ := limsup_add_le f g diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 481188eaba02c..48991e413c3b5 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1217,6 +1217,113 @@ begin hf.ennnorm, end +lemma snorm_smul_le_snorm_top_mul_snorm (p : ℝ≥0∞) + {f : α → E} (hf : ae_strongly_measurable f μ) (φ : α → 𝕜) : + snorm (φ • f) p μ ≤ snorm φ ∞ μ * snorm f p μ := +begin + by_cases hp_top : p = ∞, + { simp_rw [hp_top, snorm_exponent_top, snorm_ess_sup, pi.smul_apply', nnnorm_smul, + ennreal.coe_mul], + exact ennreal.ess_sup_mul_le _ _, }, + by_cases hp_zero : p = 0, + { simp only [hp_zero, snorm_exponent_zero, mul_zero, le_zero_iff], }, + simp_rw [snorm_eq_lintegral_rpow_nnnorm hp_zero hp_top, snorm_exponent_top, snorm_ess_sup], + calc (∫⁻ x, ↑∥(φ • f) x∥₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) + = (∫⁻ x, ↑∥φ x∥₊ ^ p.to_real * ↑∥f x∥₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + begin + congr, + ext1 x, + rw [pi.smul_apply', nnnorm_smul, ennreal.coe_mul, + ennreal.mul_rpow_of_nonneg _ _ (ennreal.to_real_nonneg)], + end + ... ≤ (∫⁻ x, (ess_sup (λ x, ↑∥φ x∥₊) μ) ^ p.to_real * ↑∥f x∥₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + begin + refine ennreal.rpow_le_rpow _ _, + swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, + refine lintegral_mono_ae _, + filter_upwards [@ennreal.ae_le_ess_sup _ _ μ (λ x, ↑∥φ x∥₊)] with x hx, + refine ennreal.mul_le_mul _ le_rfl, + exact ennreal.rpow_le_rpow hx ennreal.to_real_nonneg, + end + ... = ess_sup (λ x, ↑∥φ x∥₊) μ * (∫⁻ x, ↑∥f x∥₊ ^ p.to_real ∂μ) ^ (1 / p.to_real) : + begin + rw lintegral_const_mul'', + swap, { exact hf.nnnorm.ae_measurable.coe_nnreal_ennreal.pow ae_measurable_const, }, + rw ennreal.mul_rpow_of_nonneg, + swap, { rw one_div_nonneg, exact ennreal.to_real_nonneg, }, + rw [← ennreal.rpow_mul, one_div, mul_inv_cancel, ennreal.rpow_one], + rw [ne.def, ennreal.to_real_eq_zero_iff, auto.not_or_eq], + exact ⟨hp_zero, hp_top⟩, + end +end + +lemma snorm_smul_le_snorm_mul_snorm_top (p : ℝ≥0∞) + (f : α → E) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) : + snorm (φ • f) p μ ≤ snorm φ p μ * snorm f ∞ μ := +begin + rw ← snorm_norm, + simp_rw [pi.smul_apply', norm_smul], + have : (λ x, ∥φ x∥ * ∥f x∥) = (λ x, ∥f x∥) • (λ x, ∥φ x∥), + { rw [smul_eq_mul, mul_comm], refl, }, + rw this, + have h := snorm_smul_le_snorm_top_mul_snorm p hφ.norm (λ x, ∥f x∥), + refine h.trans_eq _, + simp_rw snorm_norm, + rw mul_comm, +end + +/-- Hölder's inequality, as an inequality on the `ℒp` seminorm of a scalar product `φ • f`. -/ +lemma snorm_smul_le_mul_snorm {p q r : ℝ≥0∞} + {f : α → E} (hf : ae_strongly_measurable f μ) {φ : α → 𝕜} (hφ : ae_strongly_measurable φ μ) + (hpqr : 1/p = 1/q + 1/r) : + snorm (φ • f) p μ ≤ snorm φ q μ * snorm f r μ := +begin + by_cases hp_zero : p = 0, + { simp [hp_zero], }, + have hq_ne_zero : q ≠ 0, + { intro hq_zero, + simp only [hq_zero, hp_zero, one_div, ennreal.inv_zero, ennreal.top_add, + ennreal.inv_eq_top] at hpqr, + exact hpqr, }, + have hr_ne_zero : r ≠ 0, + { intro hr_zero, + simp only [hr_zero, hp_zero, one_div, ennreal.inv_zero, ennreal.add_top, + ennreal.inv_eq_top] at hpqr, + exact hpqr, }, + by_cases hq_top : q = ∞, + { have hpr : p = r, + { simpa only [hq_top, one_div, ennreal.div_top, zero_add, inv_inj] using hpqr, }, + rw [← hpr, hq_top], + exact snorm_smul_le_snorm_top_mul_snorm p hf φ, }, + by_cases hr_top : r = ∞, + { have hpq : p = q, + { simpa only [hr_top, one_div, ennreal.div_top, add_zero, inv_inj] using hpqr, }, + rw [← hpq, hr_top], + exact snorm_smul_le_snorm_mul_snorm_top p f hφ, }, + have hpq : p < q, + { suffices : 1 / q < 1 / p, + { rwa [one_div, one_div, ennreal.inv_lt_inv] at this, }, + rw hpqr, + refine ennreal.lt_add_right _ _, + { simp only [hq_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, + { simp only [hr_top, one_div, ne.def, ennreal.inv_eq_zero, not_false_iff], }, }, + rw [snorm_eq_snorm' hp_zero (hpq.trans_le le_top).ne, snorm_eq_snorm' hq_ne_zero hq_top, + snorm_eq_snorm' hr_ne_zero hr_top], + refine snorm'_smul_le_mul_snorm' hf hφ _ _ _, + { exact ennreal.to_real_pos hp_zero (hpq.trans_le le_top).ne, }, + { exact ennreal.to_real_strict_mono hq_top hpq, }, + rw [← ennreal.one_to_real, ← ennreal.to_real_div, ← ennreal.to_real_div, ← ennreal.to_real_div, + hpqr, ennreal.to_real_add], + { simp only [hq_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, + { simp only [hr_ne_zero, one_div, ne.def, ennreal.inv_eq_top, not_false_iff], }, +end + +lemma mem_ℒp.smul {p q r : ℝ≥0∞} {f : α → E} {φ : α → 𝕜} + (hf : mem_ℒp f r μ) (hφ : mem_ℒp φ q μ) (hpqr : 1/p = 1/q + 1/r) : + mem_ℒp (φ • f) p μ := +⟨hφ.1.smul hf.1, (snorm_smul_le_mul_snorm hf.1 hφ.1 hpqr).trans_lt + (ennreal.mul_lt_top hφ.snorm_ne_top hf.snorm_ne_top)⟩ + end normed_space section monotonicity diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 1aa4f04689360..3270c6a05ec81 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -1422,6 +1422,80 @@ begin rw abs_of_nonneg (hf_nonneg x), }, end +/-- Hölder's inequality for the integral of a product of norms. The integral of the product of two +norms of functions is bounded by the product of their `ℒp` and `ℒq` seminorms when `p` and `q` are +conjugate exponents. -/ +theorem integral_mul_norm_le_Lp_mul_Lq {E} [normed_add_comm_group E] {f g : α → E} + {p q : ℝ} (hpq : p.is_conjugate_exponent q) + (hf : mem_ℒp f (ennreal.of_real p) μ) (hg : mem_ℒp g (ennreal.of_real q) μ) : + ∫ a, ∥f a∥ * ∥g a∥ ∂μ ≤ (∫ a, ∥f a∥ ^ p ∂μ) ^ (1/p) * (∫ a, ∥g a∥ ^ q ∂μ) ^ (1/q) := +begin + -- translate the Bochner integrals into Lebesgue integrals. + rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae, + integral_eq_lintegral_of_nonneg_ae], + rotate 1, + { exact eventually_of_forall (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _), }, + { exact (hg.1.norm.ae_measurable.pow ae_measurable_const).ae_strongly_measurable, }, + { exact eventually_of_forall (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _),}, + { exact (hf.1.norm.ae_measurable.pow ae_measurable_const).ae_strongly_measurable, }, + { exact eventually_of_forall (λ x, mul_nonneg (norm_nonneg _) (norm_nonneg _)), }, + { exact hf.1.norm.mul hg.1.norm, }, + rw [ennreal.to_real_rpow, ennreal.to_real_rpow, ← ennreal.to_real_mul], + -- replace norms by nnnorm + have h_left : ∫⁻ a, ennreal.of_real (∥f a∥ * ∥g a∥) ∂μ + = ∫⁻ a, ((λ x, (∥f x∥₊ : ℝ≥0∞)) * (λ x, ∥g x∥₊)) a ∂μ, + { simp_rw [pi.mul_apply, ← of_real_norm_eq_coe_nnnorm, ennreal.of_real_mul (norm_nonneg _)], }, + have h_right_f : ∫⁻ a, ennreal.of_real (∥f a∥ ^ p) ∂μ = ∫⁻ a, ∥f a∥₊ ^ p ∂μ, + { refine lintegral_congr (λ x, _), + rw [← of_real_norm_eq_coe_nnnorm, ennreal.of_real_rpow_of_nonneg (norm_nonneg _) hpq.nonneg], }, + have h_right_g : ∫⁻ a, ennreal.of_real (∥g a∥ ^ q) ∂μ = ∫⁻ a, ∥g a∥₊ ^ q ∂μ, + { refine lintegral_congr (λ x, _), + rw [← of_real_norm_eq_coe_nnnorm, + ennreal.of_real_rpow_of_nonneg (norm_nonneg _) hpq.symm.nonneg], }, + rw [h_left, h_right_f, h_right_g], + -- we can now apply `ennreal.lintegral_mul_le_Lp_mul_Lq` (up to the `to_real` application) + refine ennreal.to_real_mono _ _, + { refine ennreal.mul_ne_top _ _, + { convert hf.snorm_ne_top, + rw snorm_eq_lintegral_rpow_nnnorm, + { rw ennreal.to_real_of_real hpq.nonneg, }, + { rw [ne.def, ennreal.of_real_eq_zero, not_le], + exact hpq.pos, }, + { exact ennreal.coe_ne_top, }, }, + { convert hg.snorm_ne_top, + rw snorm_eq_lintegral_rpow_nnnorm, + { rw ennreal.to_real_of_real hpq.symm.nonneg, }, + { rw [ne.def, ennreal.of_real_eq_zero, not_le], + exact hpq.symm.pos, }, + { exact ennreal.coe_ne_top, }, }, }, + { exact ennreal.lintegral_mul_le_Lp_mul_Lq μ hpq hf.1.nnnorm.ae_measurable.coe_nnreal_ennreal + hg.1.nnnorm.ae_measurable.coe_nnreal_ennreal, }, +end + +/-- Hölder's inequality for functions `α → ℝ`. The integral of the product of two nonnegative +functions is bounded by the product of their `ℒp` and `ℒq` seminorms when `p` and `q` are conjugate +exponents. -/ +theorem integral_mul_le_Lp_mul_Lq_of_nonneg {p q : ℝ} + (hpq : p.is_conjugate_exponent q) {f g : α → ℝ} (hf_nonneg : 0 ≤ᵐ[μ] f) (hg_nonneg : 0 ≤ᵐ[μ] g) + (hf : mem_ℒp f (ennreal.of_real p) μ) (hg : mem_ℒp g (ennreal.of_real q) μ) : + ∫ a, f a * g a ∂μ ≤ (∫ a, (f a) ^ p ∂μ) ^ (1/p) * (∫ a, (g a) ^ q ∂μ) ^ (1/q) := +begin + have h_left : ∫ a, f a * g a ∂μ = ∫ a, ∥f a∥ * ∥g a∥ ∂μ, + { refine integral_congr_ae _, + filter_upwards [hf_nonneg, hg_nonneg] with x hxf hxg, + rw [real.norm_of_nonneg hxf, real.norm_of_nonneg hxg], }, + have h_right_f : ∫ a, (f a) ^ p ∂μ = ∫ a, ∥f a∥ ^ p ∂μ, + { refine integral_congr_ae _, + filter_upwards [hf_nonneg] with x hxf, + rw real.norm_of_nonneg hxf, }, + have h_right_g : ∫ a, (g a) ^ q ∂μ = ∫ a, ∥g a∥ ^ q ∂μ, + { refine integral_congr_ae _, + filter_upwards [hg_nonneg] with x hxg, + rw real.norm_of_nonneg hxg, }, + rw [h_left, h_right_f, h_right_g], + exact integral_mul_norm_le_Lp_mul_Lq hpq hf hg, +end + end properties mk_simp_attribute integral_simps "Simp set for integral rules." diff --git a/src/order/filter/ennreal.lean b/src/order/filter/ennreal.lean index 79274175fe617..a061290a9960b 100644 --- a/src/order/filter/ennreal.lean +++ b/src/order/filter/ennreal.lean @@ -95,6 +95,17 @@ begin simp only [h_top_le, hfu, if_false], }, end +lemma limsup_mul_le [countable_Inter_filter f] (u v : α → ℝ≥0∞) : + f.limsup (u * v) ≤ f.limsup u * f.limsup v := +calc f.limsup (u * v) ≤ f.limsup (λ x, (f.limsup u) * v x) : + begin + refine limsup_le_limsup _ _, + { filter_upwards [@eventually_le_limsup _ f _ u] with x hx, + exact ennreal.mul_le_mul hx le_rfl, }, + { is_bounded_default, }, + end +... = f.limsup u * f.limsup v : limsup_const_mul + lemma limsup_add_le [countable_Inter_filter f] (u v : α → ℝ≥0∞) : f.limsup (u + v) ≤ f.limsup u + f.limsup v := Inf_le ((eventually_le_limsup u).mp ((eventually_le_limsup v).mono From 525c5a81c5000785dd183c79da1c9365b1c27971 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Fri, 2 Sep 2022 04:12:12 +0000 Subject: [PATCH 127/828] feat(topology/fiber_bundle): Each fiber of a trivialization is homeomorphic to the specified fiber (#16079) This PR adds `topological_fiber_bundle.trivialization.preimage_singleton_homeomorph`, which states that each fiber of a trivialization is homeomorphic to the specified fiber. --- src/topology/fiber_bundle.lean | 51 ++++++++++++++++++++++++++++++ src/topology/homeomorph.lean | 6 ++++ src/topology/local_homeomorph.lean | 24 +++++++++----- 3 files changed, 73 insertions(+), 8 deletions(-) diff --git a/src/topology/fiber_bundle.lean b/src/topology/fiber_bundle.lean index 3ae152d4490e9..b3350fdf80eaf 100644 --- a/src/topology/fiber_bundle.lean +++ b/src/topology/fiber_bundle.lean @@ -361,6 +361,57 @@ lemma map_proj_nhds (ex : x ∈ e.source) : map proj (𝓝 x) = 𝓝 (proj x) := by rw [← e.coe_fst ex, ← map_congr (e.coe_fst_eventually_eq_proj ex), ← map_map, ← e.coe_coe, e.to_local_homeomorph.map_nhds_eq ex, map_fst_nhds] +lemma preimage_subset_source {s : set B} (hb : s ⊆ e.base_set) : proj ⁻¹' s ⊆ e.source := +λ p hp, e.mem_source.mpr (hb hp) + +lemma image_preimage_eq_prod_univ {s : set B} (hb : s ⊆ e.base_set) : + e '' (proj ⁻¹' s) = s ×ˢ univ := +subset.antisymm (image_subset_iff.mpr (λ p hp, + ⟨(e.proj_to_fun p (e.preimage_subset_source hb hp)).symm ▸ hp, trivial⟩)) (λ p hp, + let hp' : p ∈ e.target := e.mem_target.mpr (hb hp.1) in + ⟨e.inv_fun p, mem_preimage.mpr ((e.proj_symm_apply hp').symm ▸ hp.1), e.apply_symm_apply hp'⟩) + +/-- The preimage of a subset of the base set is homeomorphic to the product with the fiber. -/ +def preimage_homeomorph {s : set B} (hb : s ⊆ e.base_set) : proj ⁻¹' s ≃ₜ s × F := +(e.to_local_homeomorph.homeomorph_of_image_subset_source (e.preimage_subset_source hb) + (e.image_preimage_eq_prod_univ hb)).trans + ((homeomorph.set.prod s univ).trans ((homeomorph.refl s).prod_congr (homeomorph.set.univ F))) + +@[simp] lemma preimage_homeomorph_apply {s : set B} (hb : s ⊆ e.base_set) (p : proj ⁻¹' s) : + e.preimage_homeomorph hb p = (⟨proj p, p.2⟩, (e p).2) := +prod.ext (subtype.ext (e.proj_to_fun p (e.mem_source.mpr (hb p.2)))) rfl + +@[simp] lemma preimage_homeomorph_symm_apply {s : set B} (hb : s ⊆ e.base_set) (p : s × F) : + (e.preimage_homeomorph hb).symm p = ⟨e.symm (p.1, p.2), ((e.preimage_homeomorph hb).symm p).2⟩ := +rfl + +/-- The source is homeomorphic to the product of the base set with the fiber. -/ +def source_homeomorph_base_set_prod : e.source ≃ₜ e.base_set × F := +(homeomorph.set_congr e.source_eq).trans (e.preimage_homeomorph subset_rfl) + +@[simp] lemma source_homeomorph_base_set_prod_apply (p : e.source) : + e.source_homeomorph_base_set_prod p = (⟨proj p, e.mem_source.mp p.2⟩, (e p).2) := +e.preimage_homeomorph_apply subset_rfl ⟨p, e.mem_source.mp p.2⟩ + +@[simp] lemma source_homeomorph_base_set_prod_symm_apply (p : e.base_set × F) : + e.source_homeomorph_base_set_prod.symm p = + ⟨e.symm (p.1, p.2), (e.source_homeomorph_base_set_prod.symm p).2⟩ := +rfl + +/-- Each fiber of a trivialization is homeomorphic to the specified fiber. -/ +def preimage_singleton_homeomorph {b : B} (hb : b ∈ e.base_set) : proj ⁻¹' {b} ≃ₜ F := +(e.preimage_homeomorph (set.singleton_subset_iff.mpr hb)).trans (((homeomorph.homeomorph_of_unique + ({b} : set B) punit).prod_congr (homeomorph.refl F)).trans (homeomorph.punit_prod F)) + +@[simp] lemma preimage_singleton_homeomorph_apply {b : B} (hb : b ∈ e.base_set) + (p : proj ⁻¹' {b}) : e.preimage_singleton_homeomorph hb p = (e p).2 := +rfl + +@[simp] lemma preimage_singleton_homeomorph_symm_apply {b : B} (hb : b ∈ e.base_set) (p : F) : + (e.preimage_singleton_homeomorph hb).symm p = + ⟨e.symm (b, p), by rw [mem_preimage, e.proj_symm_apply' hb, mem_singleton_iff]⟩ := +rfl + /-- In the domain of a bundle trivialization, the projection is continuous-/ lemma continuous_at_proj (ex : x ∈ e.source) : continuous_at proj x := (e.map_proj_nhds ex).le diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 8531d7282888d..8e9afbe7afd0f 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -379,6 +379,12 @@ def punit_prod : punit × α ≃ₜ α := @[simp] lemma coe_punit_prod : ⇑(punit_prod α) = prod.snd := rfl +/-- If both `α` and `β` have a unique element, then `α ≃ₜ β`. -/ +@[simps] def _root_.homeomorph.homeomorph_of_unique [unique α] [unique β] : α ≃ₜ β := +{ continuous_to_fun := @continuous_const α β _ _ default, + continuous_inv_fun := @continuous_const β α _ _ default, + .. equiv.equiv_of_unique α β } + end /-- `ulift α` is homeomorphic to `α`. -/ diff --git a/src/topology/local_homeomorph.lean b/src/topology/local_homeomorph.lean index d54e294e7b397..d410a89b206f0 100644 --- a/src/topology/local_homeomorph.lean +++ b/src/topology/local_homeomorph.lean @@ -985,16 +985,24 @@ end end continuity +/-- The homeomorphism obtained by restricting a `local_homeomorph` to a subset of the source. -/ +@[simps] def homeomorph_of_image_subset_source + {s : set α} {t : set β} (hs : s ⊆ e.source) (ht : e '' s = t) : s ≃ₜ t := +{ to_fun := λ a, ⟨e a, (congr_arg ((∈) (e a)) ht).mp ⟨a, a.2, rfl⟩⟩, + inv_fun := λ b, ⟨e.symm b, let ⟨a, ha1, ha2⟩ := (congr_arg ((∈) ↑b) ht).mpr b.2 in + ha2 ▸ (e.left_inv (hs ha1)).symm ▸ ha1⟩, + left_inv := λ a, subtype.ext (e.left_inv (hs a.2)), + right_inv := λ b, let ⟨a, ha1, ha2⟩ := (congr_arg ((∈) ↑b) ht).mpr b.2 in + subtype.ext (e.right_inv (ha2 ▸ e.map_source (hs ha1))), + continuous_to_fun := (continuous_on_iff_continuous_restrict.mp + (e.continuous_on.mono hs)).subtype_mk _, + continuous_inv_fun := (continuous_on_iff_continuous_restrict.mp + (e.continuous_on_symm.mono (λ b hb, let ⟨a, ha1, ha2⟩ := show b ∈ e '' s, from ht.symm ▸ hb in + ha2 ▸ e.map_source (hs ha1)))).subtype_mk _ } + /-- A local homeomrphism defines a homeomorphism between its source and target. -/ def to_homeomorph_source_target : e.source ≃ₜ e.target := -{ to_fun := e.maps_to.restrict _ _ _, - inv_fun := e.symm_maps_to.restrict _ _ _, - left_inv := λ x, subtype.eq $ e.left_inv x.2, - right_inv := λ x, subtype.eq $ e.right_inv x.2, - continuous_to_fun := - (continuous_on_iff_continuous_restrict.1 e.continuous_on).subtype_mk _, - continuous_inv_fun := - (continuous_on_iff_continuous_restrict.1 e.symm.continuous_on).subtype_mk _ } +e.homeomorph_of_image_subset_source subset_rfl e.image_source_eq_target lemma second_countable_topology_source [second_countable_topology β] (e : local_homeomorph α β) : From 541b2880295e704dd563ddb328d2020fcd3b4455 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 2 Sep 2022 06:43:31 +0000 Subject: [PATCH 128/828] chore(topology/order): open `function`, don't open `classical` (#16288) --- src/topology/order.lean | 12 +++++------- 1 file changed, 5 insertions(+), 7 deletions(-) diff --git a/src/topology/order.lean b/src/topology/order.lean index a1377c4cfd675..b9f26affdd12c 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -41,7 +41,7 @@ finer, coarser, induced topology, coinduced topology -/ -open set filter classical +open function set filter open_locale classical topological_space filter universes u v w @@ -189,16 +189,14 @@ lemma generate_from_set_of_is_open (t : topological_space α) : (gi_generate_from α).l_u_eq t lemma left_inverse_generate_from : - function.left_inverse topological_space.generate_from - (λ t : topological_space α, {s | t.is_open s}) := + left_inverse topological_space.generate_from (λ t : topological_space α, {s | t.is_open s}) := (gi_generate_from α).left_inverse_l_u lemma generate_from_surjective : - function.surjective (topological_space.generate_from : set (set α) → topological_space α) := + surjective (topological_space.generate_from : set (set α) → topological_space α) := (gi_generate_from α).l_surjective -lemma set_of_is_open_injective : - function.injective (λ t : topological_space α, {s | t.is_open s}) := +lemma set_of_is_open_injective : injective (λ t : topological_space α, {s | t.is_open s}) := (gi_generate_from α).u_injective /-- The "temporary" order `tmp_order` on `topological_space α`, i.e. the inclusion order, is a @@ -738,7 +736,7 @@ tβ = tα.induced f ↔ ∀ b, 𝓝 b = comap f (𝓝 $ f b) := ⟨λ h a, h.symm ▸ nhds_induced f a, λ h, eq_of_nhds_eq_nhds $ λ x, by rw [h, nhds_induced]⟩ theorem map_nhds_induced_of_surjective [T : topological_space α] - {f : β → α} (hf : function.surjective f) (a : β) : + {f : β → α} (hf : surjective f) (a : β) : map f (@nhds β (topological_space.induced f T) a) = 𝓝 (f a) := by rw [nhds_induced, map_comap_of_surjective hf] From ddc748440a47e3f0b7ac24836d1de67d16e30054 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Fri, 2 Sep 2022 09:34:33 +0000 Subject: [PATCH 129/828] feat(geometry/euclidean/oriented_angle): signs of angles between linear combinations of vectors (#16243) Add a series of lemmas giving the sign of the oriented angle between two linear combinations of two vectors in terms of the sign of the angle between the original two vectors and the coefficients in the linear combination. --- src/geometry/euclidean/oriented_angle.lean | 224 +++++++++++++++++++++ 1 file changed, 224 insertions(+) diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 873352be8a174..8762d2e983886 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -1012,6 +1012,162 @@ begin simpa using ha } end +/-- Negating the first vector passed to `oangle` negates the sign of the angle. -/ +@[simp] lemma oangle_sign_neg_left (x y : V) : + (hb.oangle (-x) y).sign = -((hb.oangle x y).sign) := +begin + by_cases hx : x = 0, { simp [hx] }, + by_cases hy : y = 0, { simp [hy] }, + rw [hb.oangle_neg_left hx hy, real.angle.sign_add_pi] +end + +/-- Negating the second vector passed to `oangle` negates the sign of the angle. -/ +@[simp] lemma oangle_sign_neg_right (x y : V) : + (hb.oangle x (-y)).sign = -((hb.oangle x y).sign) := +begin + by_cases hx : x = 0, { simp [hx] }, + by_cases hy : y = 0, { simp [hy] }, + rw [hb.oangle_neg_right hx hy, real.angle.sign_add_pi] +end + +/-- Multiplying the first vector passed to `oangle` by a real multiplies the sign of the angle by +the sign of the real. -/ +@[simp] lemma oangle_sign_smul_left (x y : V) (r : ℝ) : + (hb.oangle (r • x) y).sign = sign r * (hb.oangle x y).sign := +begin + rcases lt_trichotomy r 0 with h|h|h; + simp [h] +end + +/-- Multiplying the second vector passed to `oangle` by a real multiplies the sign of the angle by +the sign of the real. -/ +@[simp] lemma oangle_sign_smul_right (x y : V) (r : ℝ) : + (hb.oangle x (r • y)).sign = sign r * (hb.oangle x y).sign := +begin + rcases lt_trichotomy r 0 with h|h|h; + simp [h] +end + +/-- Auxiliary lemma for the proof of `oangle_sign_smul_add_right`; not intended to be used +outside of that proof. -/ +lemma oangle_smul_add_right_eq_zero_or_eq_pi_iff {x y : V} (r : ℝ) : + (hb.oangle x (r • x + y) = 0 ∨ hb.oangle x (r • x + y) = π) ↔ + (hb.oangle x y = 0 ∨ hb.oangle x y = π) := +begin + simp_rw [oangle_eq_zero_or_eq_pi_iff_not_linear_independent, + fintype.not_linear_independent_iff, fin.sum_univ_two, fin.exists_fin_two], + refine ⟨λ h, _, λ h, _⟩, + { rcases h with ⟨m, h, hm⟩, + change m 0 • x + m 1 • (r • x + y) = 0 at h, + refine ⟨![m 0 + m 1 * r, m 1], _⟩, + change (m 0 + m 1 * r) • x + m 1 • y = 0 ∧ (m 0 + m 1 * r ≠ 0 ∨ m 1 ≠ 0), + rw [smul_add, smul_smul, ←add_assoc, ←add_smul] at h, + refine ⟨h, not_and_distrib.1 (λ h0, _)⟩, + obtain ⟨h0, h1⟩ := h0, + rw h1 at h0 hm, + rw [zero_mul, add_zero] at h0, + simpa [h0] using hm }, + { rcases h with ⟨m, h, hm⟩, + change m 0 • x + m 1 • y = 0 at h, + refine ⟨![m 0 - m 1 * r, m 1], _⟩, + change (m 0 - m 1 * r) • x + m 1 • (r • x + y) = 0 ∧ (m 0 - m 1 * r ≠ 0 ∨ m 1 ≠ 0), + rw [sub_smul, smul_add, smul_smul, ←add_assoc, sub_add_cancel], + refine ⟨h, not_and_distrib.1 (λ h0, _)⟩, + obtain ⟨h0, h1⟩ := h0, + rw h1 at h0 hm, + rw [zero_mul, sub_zero] at h0, + simpa [h0] using hm } +end + +/-- Adding a multiple of the first vector passed to `oangle` to the second vector does not change +the sign of the angle. -/ +@[simp] lemma oangle_sign_smul_add_right (x y : V) (r : ℝ) : + (hb.oangle x (r • x + y)).sign = (hb.oangle x y).sign := +begin + by_cases h : hb.oangle x y = 0 ∨ hb.oangle x y = π, + { rwa [real.angle.sign_eq_zero_iff.2 h, real.angle.sign_eq_zero_iff, + oangle_smul_add_right_eq_zero_or_eq_pi_iff] }, + have h' : ∀ r' : ℝ, hb.oangle x (r' • x + y) ≠ 0 ∧ hb.oangle x (r' • x + y) ≠ π, + { intro r', + rwa [←hb.oangle_smul_add_right_eq_zero_or_eq_pi_iff r', not_or_distrib] at h }, + let s : set (V × V) := (λ r' : ℝ, (x, r' • x + y)) '' set.univ, + have hc : is_connected s := is_connected_univ.image _ ((continuous_const.prod_mk + ((continuous_id.smul continuous_const).add continuous_const)).continuous_on), + have hf : continuous_on (λ z : V × V, hb.oangle z.1 z.2) s, + { refine continuous_at.continuous_on (λ z hz, hb.continuous_at_oangle _ _), + all_goals { simp_rw [s, set.mem_image] at hz, + obtain ⟨r', -, rfl⟩ := hz, + simp only [prod.fst, prod.snd], + intro hz }, + { simpa [hz] using (h' 0).1 }, + { simpa [hz] using (h' r').1 } }, + have hs : ∀ z : V × V, z ∈ s → hb.oangle z.1 z.2 ≠ 0 ∧ hb.oangle z.1 z.2 ≠ π, + { intros z hz, + simp_rw [s, set.mem_image] at hz, + obtain ⟨r', -, rfl⟩ := hz, + exact h' r' }, + have hx : (x, y) ∈ s, + { convert set.mem_image_of_mem (λ r' : ℝ, (x, r' • x + y)) (set.mem_univ 0), + simp }, + have hy : (x, r • x + y) ∈ s := set.mem_image_of_mem _ (set.mem_univ _), + convert real.angle.sign_eq_of_continuous_on hc hf hs hx hy +end + +/-- Adding a multiple of the second vector passed to `oangle` to the first vector does not change +the sign of the angle. -/ +@[simp] lemma oangle_sign_add_smul_left (x y : V) (r : ℝ) : + (hb.oangle (x + r • y) y).sign = (hb.oangle x y).sign := +by simp_rw [hb.oangle_rev y, real.angle.sign_neg, add_comm x, oangle_sign_smul_add_right] + +/-- Subtracting a multiple of the first vector passed to `oangle` from the second vector does +not change the sign of the angle. -/ +@[simp] lemma oangle_sign_sub_smul_right (x y : V) (r : ℝ) : + (hb.oangle x (y - r • x)).sign = (hb.oangle x y).sign := +by rw [sub_eq_add_neg, ←neg_smul, add_comm, oangle_sign_smul_add_right] + +/-- Subtracting a multiple of the second vector passed to `oangle` from the first vector does +not change the sign of the angle. -/ +@[simp] lemma oangle_sign_sub_smul_left (x y : V) (r : ℝ) : + (hb.oangle (x - r • y) y).sign = (hb.oangle x y).sign := +by rw [sub_eq_add_neg, ←neg_smul, oangle_sign_add_smul_left] + +/-- The sign of the angle between a vector, and a linear combination of that vector with a second +vector, is the sign of the factor by which the second vector is multiplied in that combination +multiplied by the sign of the angle between the two vectors. -/ +@[simp] lemma oangle_sign_smul_add_smul_right (x y : V) (r₁ r₂ : ℝ) : + (hb.oangle x (r₁ • x + r₂ • y)).sign = sign r₂ * (hb.oangle x y).sign := +begin + rw ←hb.oangle_sign_smul_add_right x (r₁ • x + r₂ • y) (-r₁), + simp +end + +/-- The sign of the angle between a linear combination of two vectors and the second vector is +the sign of the factor by which the first vector is multiplied in that combination multiplied by +the sign of the angle between the two vectors. -/ +@[simp] lemma oangle_sign_smul_add_smul_left (x y : V) (r₁ r₂ : ℝ) : + (hb.oangle (r₁ • x + r₂ • y) y).sign = sign r₁ * (hb.oangle x y).sign := +by simp_rw [hb.oangle_rev y, real.angle.sign_neg, add_comm (r₁ • x), + oangle_sign_smul_add_smul_right, mul_neg] + +/-- The sign of the angle between two linear combinations of two vectors is the sign of the +determinant of the factors in those combinations multiplied by the sign of the angle between the +two vectors. -/ +lemma oangle_sign_smul_add_smul_smul_add_smul (x y : V) (r₁ r₂ r₃ r₄ : ℝ) : + (hb.oangle (r₁ • x + r₂ • y) (r₃ • x + r₄ • y)).sign = + sign (r₁ * r₄ - r₂ * r₃) * (hb.oangle x y).sign := +begin + by_cases hr₁ : r₁ = 0, + { rw [hr₁, zero_smul, zero_mul, zero_add, zero_sub, left.sign_neg, oangle_sign_smul_left, + add_comm, oangle_sign_smul_add_smul_right, oangle_rev, real.angle.sign_neg, sign_mul, + mul_neg, mul_neg, neg_mul, mul_assoc] }, + { rw [←hb.oangle_sign_smul_add_right (r₁ • x + r₂ • y) (r₃ • x + r₄ • y) (-r₃ / r₁), + smul_add, smul_smul, smul_smul, div_mul_cancel _ hr₁, neg_smul, ←add_assoc, + add_comm (-(r₃ • x)), ←sub_eq_add_neg, sub_add_cancel, ←add_smul, + oangle_sign_smul_right, oangle_sign_smul_add_smul_left, ←mul_assoc, ←sign_mul, + add_mul, mul_assoc, mul_comm r₂ r₁, ←mul_assoc, div_mul_cancel _ hr₁, add_comm, + neg_mul, ←sub_eq_add_neg, mul_comm r₄, mul_comm r₃] } +end + end orthonormal namespace orientation @@ -1540,4 +1696,72 @@ lemma oangle_eq_pi_iff_angle_eq_pi {x y : V} : o.oangle x y = π ↔ inner_product_geometry.angle x y = π := (ob).oangle_eq_pi_iff_angle_eq_pi +/-- Negating the first vector passed to `oangle` negates the sign of the angle. -/ +@[simp] lemma oangle_sign_neg_left (x y : V) : + (o.oangle (-x) y).sign = -((o.oangle x y).sign) := +(ob).oangle_sign_neg_left x y + +/-- Negating the second vector passed to `oangle` negates the sign of the angle. -/ +@[simp] lemma oangle_sign_neg_right (x y : V) : + (o.oangle x (-y)).sign = -((o.oangle x y).sign) := +(ob).oangle_sign_neg_right x y + +/-- Multiplying the first vector passed to `oangle` by a real multiplies the sign of the angle by +the sign of the real. -/ +@[simp] lemma oangle_sign_smul_left (x y : V) (r : ℝ) : + (o.oangle (r • x) y).sign = sign r * (o.oangle x y).sign := +(ob).oangle_sign_smul_left x y r + +/-- Multiplying the second vector passed to `oangle` by a real multiplies the sign of the angle by +the sign of the real. -/ +@[simp] lemma oangle_sign_smul_right (x y : V) (r : ℝ) : + (o.oangle x (r • y)).sign = sign r * (o.oangle x y).sign := +(ob).oangle_sign_smul_right x y r + +/-- Adding a multiple of the first vector passed to `oangle` to the second vector does not change +the sign of the angle. -/ +@[simp] lemma oangle_sign_smul_add_right (x y : V) (r : ℝ) : + (o.oangle x (r • x + y)).sign = (o.oangle x y).sign := +(ob).oangle_sign_smul_add_right x y r + +/-- Adding a multiple of the second vector passed to `oangle` to the first vector does not change +the sign of the angle. -/ +@[simp] lemma oangle_sign_add_smul_left (x y : V) (r : ℝ) : + (o.oangle (x + r • y) y).sign = (o.oangle x y).sign := +(ob).oangle_sign_add_smul_left x y r + +/-- Subtracting a multiple of the first vector passed to `oangle` from the second vector does +not change the sign of the angle. -/ +@[simp] lemma oangle_sign_sub_smul_right (x y : V) (r : ℝ) : + (o.oangle x (y - r • x)).sign = (o.oangle x y).sign := +(ob).oangle_sign_sub_smul_right x y r + +/-- Subtracting a multiple of the second vector passed to `oangle` from the first vector does +not change the sign of the angle. -/ +@[simp] lemma oangle_sign_sub_smul_left (x y : V) (r : ℝ) : + (o.oangle (x - r • y) y).sign = (o.oangle x y).sign := +(ob).oangle_sign_sub_smul_left x y r + +/-- The sign of the angle between a vector, and a linear combination of that vector with a second +vector, is the sign of the factor by which the second vector is multiplied in that combination +multiplied by the sign of the angle between the two vectors. -/ +@[simp] lemma oangle_sign_smul_add_smul_right (x y : V) (r₁ r₂ : ℝ) : + (o.oangle x (r₁ • x + r₂ • y)).sign = sign r₂ * (o.oangle x y).sign := +(ob).oangle_sign_smul_add_smul_right x y r₁ r₂ + +/-- The sign of the angle between a linear combination of two vectors and the second vector is +the sign of the factor by which the first vector is multiplied in that combination multiplied by +the sign of the angle between the two vectors. -/ +@[simp] lemma oangle_sign_smul_add_smul_left (x y : V) (r₁ r₂ : ℝ) : + (o.oangle (r₁ • x + r₂ • y) y).sign = sign r₁ * (o.oangle x y).sign := +(ob).oangle_sign_smul_add_smul_left x y r₁ r₂ + +/-- The sign of the angle between two linear combinations of two vectors is the sign of the +determinant of the factors in those combinations multiplied by the sign of the angle between the +two vectors. -/ +lemma oangle_sign_smul_add_smul_smul_add_smul (x y : V) (r₁ r₂ r₃ r₄ : ℝ) : + (o.oangle (r₁ • x + r₂ • y) (r₃ • x + r₄ • y)).sign = + sign (r₁ * r₄ - r₂ * r₃) * (o.oangle x y).sign := +(ob).oangle_sign_smul_add_smul_smul_add_smul x y r₁ r₂ r₃ r₄ + end orientation From 86152105ecd765cba7d18d5a358e7b6bba470cec Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Fri, 2 Sep 2022 10:49:40 +0000 Subject: [PATCH 130/828] feat(data/rat/defs): inv_coe_{int,nat}_{num,denom} (#15863) Rename existing lemmas to `inv_coe_{int,nat}_{num,denom}_of_pos` since they took positivity assumptions Also provide `rat.inv_neg`. Co-authored-by: Yakov Pechersky --- src/data/rat/defs.lean | 41 +++++++++++++++++++++++++++++++++++------ 1 file changed, 35 insertions(+), 6 deletions(-) diff --git a/src/data/rat/defs.lean b/src/data/rat/defs.lean index 7032e48f3b27f..4d46b52bcd7b9 100644 --- a/src/data/rat/defs.lean +++ b/src/data/rat/defs.lean @@ -786,6 +786,13 @@ end casts lemma inv_def' {q : ℚ} : q⁻¹ = (q.denom : ℚ) / q.num := by { conv_lhs { rw ←(@num_denom q) }, cases q, simp [div_num_denom] } +protected lemma inv_neg (q : ℚ) : (-q)⁻¹ = -(q⁻¹) := +begin + simp only [inv_def'], + cases eq_or_ne (q.num : ℚ) 0 with hq hq; + simp [div_eq_iff, hq] +end + @[simp] lemma mul_denom_eq_num {q : ℚ} : q * q.denom = q.num := begin suffices : mk (q.num) ↑(q.denom) * mk ↑(q.denom) 1 = mk (q.num) 1, by @@ -859,7 +866,7 @@ begin coe_nat_div_self] end -lemma inv_coe_int_num {a : ℤ} (ha0 : 0 < a) : (a : ℚ)⁻¹.num = 1 := +lemma inv_coe_int_num_of_pos {a : ℤ} (ha0 : 0 < a) : (a : ℚ)⁻¹.num = 1 := begin rw [rat.inv_def', rat.coe_int_num, rat.coe_int_denom, nat.cast_one, ←int.cast_one], apply num_div_eq_of_coprime ha0, @@ -867,10 +874,10 @@ begin exact nat.coprime_one_left _, end -lemma inv_coe_nat_num {a : ℕ} (ha0 : 0 < a) : (a : ℚ)⁻¹.num = 1 := -inv_coe_int_num (by exact_mod_cast ha0 : 0 < (a : ℤ)) +lemma inv_coe_nat_num_of_pos {a : ℕ} (ha0 : 0 < a) : (a : ℚ)⁻¹.num = 1 := +inv_coe_int_num_of_pos (by exact_mod_cast ha0 : 0 < (a : ℤ)) -lemma inv_coe_int_denom {a : ℤ} (ha0 : 0 < a) : ((a : ℚ)⁻¹.denom : ℤ) = a := +lemma inv_coe_int_denom_of_pos {a : ℤ} (ha0 : 0 < a) : ((a : ℚ)⁻¹.denom : ℤ) = a := begin rw [rat.inv_def', rat.coe_int_num, rat.coe_int_denom, nat.cast_one, ←int.cast_one], apply denom_div_eq_of_coprime ha0, @@ -878,12 +885,34 @@ begin exact nat.coprime_one_left _, end -lemma inv_coe_nat_denom {a : ℕ} (ha0 : 0 < a) : (a : ℚ)⁻¹.denom = a := +lemma inv_coe_nat_denom_of_pos {a : ℕ} (ha0 : 0 < a) : (a : ℚ)⁻¹.denom = a := begin - rw [← int.coe_nat_eq_coe_nat_iff, ← int.cast_coe_nat a, inv_coe_int_denom], + rw [← int.coe_nat_eq_coe_nat_iff, ← int.cast_coe_nat a, inv_coe_int_denom_of_pos], rwa [← nat.cast_zero, nat.cast_lt] end +@[simp] lemma inv_coe_int_num (a : ℤ) : (a : ℚ)⁻¹.num = int.sign a := +begin + induction a using int.induction_on; + simp [←int.neg_succ_of_nat_coe', int.neg_succ_of_nat_coe, -neg_add_rev, rat.inv_neg, + int.coe_nat_add_one_out, -nat.cast_succ, inv_coe_nat_num_of_pos, -int.cast_neg_succ_of_nat, + @eq_comm ℤ 1, int.sign_eq_one_of_pos] +end + +@[simp] lemma inv_coe_nat_num (a : ℕ) : (a : ℚ)⁻¹.num = int.sign a := +inv_coe_int_num a + +@[simp] lemma inv_coe_int_denom (a : ℤ) : (a : ℚ)⁻¹.denom = if a = 0 then 1 else a.nat_abs := +begin + induction a using int.induction_on; + simp [←int.neg_succ_of_nat_coe', int.neg_succ_of_nat_coe, -neg_add_rev, rat.inv_neg, + int.coe_nat_add_one_out, -nat.cast_succ, inv_coe_nat_denom_of_pos, + -int.cast_neg_succ_of_nat] +end + +@[simp] lemma inv_coe_nat_denom (a : ℕ) : (a : ℚ)⁻¹.denom = if a = 0 then 1 else a := +by simpa using inv_coe_int_denom a + protected lemma «forall» {p : ℚ → Prop} : (∀ r, p r) ↔ ∀ a b : ℤ, p (a / b) := ⟨λ h _ _, h _, λ h q, (show q = q.num / q.denom, from by simp [rat.div_num_denom]).symm ▸ (h q.1 q.2)⟩ From 09fb38cc910fce386df04e045dcdebbc8cd7d789 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 2 Sep 2022 10:49:41 +0000 Subject: [PATCH 131/828] chore(analysis/box_integral): use GP instead of bot (#15939) This is how this integral is called in the original paper. --- .../box_integral/divergence_theorem.lean | 34 +++++++++-------- .../box_integral/partition/filter.lean | 37 ++++++++++++------- .../integral/divergence_theorem.lean | 8 ++-- 3 files changed, 46 insertions(+), 33 deletions(-) diff --git a/src/analysis/box_integral/divergence_theorem.lean b/src/analysis/box_integral/divergence_theorem.lean index 3b1d6577ad758..c23bc769cc82e 100644 --- a/src/analysis/box_integral/divergence_theorem.lean +++ b/src/analysis/box_integral/divergence_theorem.lean @@ -18,8 +18,8 @@ equal to the sum of integrals of `f` over the faces of `I` taken with appropriat To make the proof work, we had to ban tagged partitions with “long and thin” boxes. More precisely, we use the following generalization of one-dimensional Henstock-Kurzweil integral to functions -defined on a box in `ℝⁿ` (it corresponds to the value `⊥` of `box_integral.integration_params` in -the definition of `box_integral.has_integral`). +defined on a box in `ℝⁿ` (it corresponds to the value `box_integral.integration_params.GP = ⊥` of +`box_integral.integration_params` in the definition of `box_integral.has_integral`). We say that `f : ℝⁿ → E` has integral `y : E` over a box `I ⊆ ℝⁿ` if for an arbitrarily small positive `ε` and an arbitrarily large `c`, there exists a function `r : ℝⁿ → (0, ∞)` such that for @@ -40,6 +40,7 @@ Henstock-Kurzweil integral, integral, Stokes theorem, divergence theorem open_locale classical big_operators nnreal ennreal topological_space box_integral open continuous_linear_map (lsmul) filter set finset metric + box_integral.integration_params (GP GP_le) noncomputable theory universes u @@ -139,12 +140,12 @@ we allow `f` to be non-differentiable (but still continuous) at a countable set TODO: If `n > 0`, then the condition at `x ∈ s` can be replaced by a much weaker estimate but this requires either better integrability theorems, or usage of a filter depending on the countable set `s` (we need to ensure that none of the faces of a partition contain a point from `s`). -/ -lemma has_integral_bot_pderiv (f : ℝⁿ⁺¹ → E) (f' : ℝⁿ⁺¹ → ℝⁿ⁺¹ →L[ℝ] E) (s : set ℝⁿ⁺¹) +lemma has_integral_GP_pderiv (f : ℝⁿ⁺¹ → E) (f' : ℝⁿ⁺¹ → ℝⁿ⁺¹ →L[ℝ] E) (s : set ℝⁿ⁺¹) (hs : s.countable) (Hs : ∀ x ∈ s, continuous_within_at f I.Icc x) (Hd : ∀ x ∈ I.Icc \ s, has_fderiv_within_at f (f' x) I.Icc x) (i : fin (n + 1)) : - has_integral.{0 u u} I ⊥ (λ x, f' x (pi.single i 1)) box_additive_map.volume - (integral.{0 u u} (I.face i) ⊥ (λ x, f (i.insert_nth (I.upper i) x)) box_additive_map.volume - - integral.{0 u u} (I.face i) ⊥ (λ x, f (i.insert_nth (I.lower i) x)) + has_integral.{0 u u} I GP (λ x, f' x (pi.single i 1)) box_additive_map.volume + (integral.{0 u u} (I.face i) GP (λ x, f (i.insert_nth (I.upper i) x)) box_additive_map.volume - + integral.{0 u u} (I.face i) GP (λ x, f (i.insert_nth (I.lower i) x)) box_additive_map.volume) := begin /- Note that `f` is continuous on `I.Icc`, hence it is integrable on the faces of all boxes @@ -155,13 +156,14 @@ begin by_cases hxs : x ∈ s, exacts [Hs x hxs, (Hd x ⟨hx, hxs⟩).continuous_within_at] }, set fI : ℝ → box (fin n) → E := λ y J, - integral.{0 u u} J ⊥ (λ x, f (i.insert_nth y x)) box_additive_map.volume, + integral.{0 u u} J GP (λ x, f (i.insert_nth y x)) box_additive_map.volume, set fb : Icc (I.lower i) (I.upper i) → fin n →ᵇᵃ[↑(I.face i)] E := - λ x, (integrable_of_continuous_on ⊥ (box.continuous_on_face_Icc Hc x.2) volume).to_box_additive, + λ x, (integrable_of_continuous_on GP (box.continuous_on_face_Icc Hc x.2) + volume).to_box_additive, set F : fin (n + 1) →ᵇᵃ[I] E := box_additive_map.upper_sub_lower I i fI fb (λ x hx J, rfl), /- Thus our statement follows from some local estimates. -/ - change has_integral I ⊥ (λ x, f' x (pi.single i 1)) _ (F I), - refine has_integral_of_le_Henstock_of_forall_is_o bot_le _ _ _ s hs _ _, + change has_integral I GP (λ x, f' x (pi.single i 1)) _ (F I), + refine has_integral_of_le_Henstock_of_forall_is_o GP_le _ _ _ s hs _ _, { /- We use the volume as an upper estimate. -/ exact (volume : measure ℝⁿ⁺¹).to_box_additive.restrict _ le_top }, { exact λ J, ennreal.to_real_nonneg }, @@ -193,7 +195,7 @@ begin have Hl : J.lower i ∈ Icc (J.lower i) (J.upper i) := set.left_mem_Icc.2 (J.lower_le_upper i), have Hu : J.upper i ∈ Icc (J.lower i) (J.upper i) := set.right_mem_Icc.2 (J.lower_le_upper i), have Hi : ∀ x ∈ Icc (J.lower i) (J.upper i), - integrable.{0 u u} (J.face i) ⊥ (λ y, f (i.insert_nth x y)) box_additive_map.volume, + integrable.{0 u u} (J.face i) GP (λ y, f (i.insert_nth x y)) box_additive_map.volume, from λ x hx, integrable_of_continuous_on _ (box.continuous_on_face_Icc (Hc.mono $ box.le_iff_Icc.1 hJI) hx) volume, have hJδ' : J.Icc ⊆ closed_ball x δ ∩ I.Icc, @@ -250,20 +252,20 @@ the sum of integrals of `f` over the faces of `I` taken with appropriate signs. More precisely, we use a non-standard generalization of the Henstock-Kurzweil integral and we allow `f` to be non-differentiable (but still continuous) at a countable set of points. -/ -lemma has_integral_bot_divergence_of_forall_has_deriv_within_at +lemma has_integral_GP_divergence_of_forall_has_deriv_within_at (f : ℝⁿ⁺¹ → Eⁿ⁺¹) (f' : ℝⁿ⁺¹ → ℝⁿ⁺¹ →L[ℝ] Eⁿ⁺¹) (s : set ℝⁿ⁺¹) (hs : s.countable) (Hs : ∀ x ∈ s, continuous_within_at f I.Icc x) (Hd : ∀ x ∈ I.Icc \ s, has_fderiv_within_at f (f' x) I.Icc x) : - has_integral.{0 u u} I ⊥ (λ x, ∑ i, f' x (pi.single i 1) i) + has_integral.{0 u u} I GP (λ x, ∑ i, f' x (pi.single i 1) i) box_additive_map.volume - (∑ i, (integral.{0 u u} (I.face i) ⊥ (λ x, f (i.insert_nth (I.upper i) x) i) + (∑ i, (integral.{0 u u} (I.face i) GP (λ x, f (i.insert_nth (I.upper i) x) i) box_additive_map.volume - - integral.{0 u u} (I.face i) ⊥ (λ x, f (i.insert_nth (I.lower i) x) i) + integral.{0 u u} (I.face i) GP (λ x, f (i.insert_nth (I.lower i) x) i) box_additive_map.volume)) := begin refine has_integral_sum (λ i hi, _), clear hi, simp only [has_fderiv_within_at_pi', continuous_within_at_pi] at Hd Hs, - convert has_integral_bot_pderiv I _ _ s hs (λ x hx, Hs x hx i) (λ x hx, Hd x hx i) i + convert has_integral_GP_pderiv I _ _ s hs (λ x hx, Hs x hx i) (λ x hx, Hd x hx i) i end end box_integral diff --git a/src/analysis/box_integral/partition/filter.lean b/src/analysis/box_integral/partition/filter.lean index edd2fbd89e18d..8954ff270d186 100644 --- a/src/analysis/box_integral/partition/filter.lean +++ b/src/analysis/box_integral/partition/filter.lean @@ -41,7 +41,8 @@ The structure `box_integral.integration_params` has 3 boolean fields with the fo * `bDistortion`: the value `tt` means that `r` can depend on the maximal ratio of sides of the same box of a partition. Presence of this case make quite a few proofs harder but we can prove the - divergence theorem only for the filter `⊥ = {bRiemann := ff, bHenstock := tt, bDistortion := tt}`. + divergence theorem only for the filter + `box_integral.integration_params.GP = ⊥ = {bRiemann := ff, bHenstock := tt, bDistortion := tt}`. ### Well-known sets of parameters @@ -67,15 +68,15 @@ the library. that we allow tags to be outside of their boxes; the tags still have to be in the ambient closed box, and the partition still has to be subordinate to a function. -* `⊥` (`bRiemann = ff`, `bHenstock = tt`, `bDistortion = tt`): this is the least integration theory - in our list, i.e., all functions integrable in any other theory is integrable in this one as well. - This is a non-standard generalization of the Henstock-Kurzweil integral to higher dimension. - In dimension one, it generates the same filter as `Henstock`. In higher dimension, this - generalization defines an integration theory such that the divergence of any Fréchet - differentiable function `f` is integrable, and its integral is equal to the sum of integrals of - `f` over the faces of the box, taken with appropriate signs. +* `box_integral.integration_params.GP = ⊥` (`bRiemann = ff`, `bHenstock = tt`, `bDistortion = tt`): + this is the least integration theory in our list, i.e., all functions integrable in any other + theory is integrable in this one as well. This is a non-standard generalization of the + Henstock-Kurzweil integral to higher dimension. In dimension one, it generates the same filter as + `Henstock`. In higher dimension, this generalization defines an integration theory such that the + divergence of any Fréchet differentiable function `f` is integrable, and its integral is equal to + the sum of integrals of `f` over the faces of the box, taken with appropriate signs. - A function `f` is `⊥`-integrable if for any `ε > 0` and `c : ℝ≥0` there exists + A function `f` is `GP`-integrable if for any `ε > 0` and `c : ℝ≥0` there exists `r : (ι → ℝ) → {x : ℝ | 0 < x}` such that for any tagged partition `π` subordinate to `r`, if each tag belongs to the corresponding closed box and for each box `J ∈ π`, the maximal ratio of its sides is less than or equal to `c`, then the integral sum of `f` over `π` is `ε`-close to the @@ -183,7 +184,8 @@ used in the definition of a box-integrable function. * `bDistortion`: the value `tt` means that `r` can depend on the maximal ratio of sides of the same box of a partition. Presence of this case makes quite a few proofs harder but we can prove the - divergence theorem only for the filter `⊥ = {bRiemann := ff, bHenstock := tt, bDistortion := tt}`. + divergence theorem only for the filter + `box_integral.integration_params.GP = ⊥ = {bRiemann := ff, bHenstock := tt, bDistortion := tt}`. -/ @[ext] structure integration_params : Type := (bRiemann bHenstock bDistortion : bool) @@ -209,9 +211,10 @@ def iso_prod : integration_params ≃o bool × boolᵒᵈ × boolᵒᵈ := instance : bounded_order integration_params := iso_prod.symm.to_galois_insertion.lift_bounded_order -/-- The value `⊥` (`bRiemann = ff`, `bHenstock = tt`, `bDistortion = tt`) corresponds to a -generalization of the Henstock integral such that the Divergence theorem holds true without -additional integrability assumptions, see the module docstring for details. -/ +/-- The value +`box_integral.integration_params.GP = ⊥` (`bRiemann = ff`, `bHenstock = tt`, `bDistortion = tt`) +corresponds to a generalization of the Henstock integral such that the Divergence theorem holds true +without additional integrability assumptions, see the module docstring for details. -/ instance : inhabited integration_params := ⟨⊥⟩ instance : decidable_rel ((≤) : integration_params → integration_params → Prop) := @@ -239,10 +242,18 @@ discontinuous) positive function `r`; the tags may be outside of the correspondi (but still inside the ambient closed box `I.Icc`). -/ def McShane : integration_params := ⟨ff, ff, ff⟩ +/-- The `box_integral.integration_params` corresponding to the generalized Perron integral. In the +corresponding filter, we require that the tagged partition is subordinate to a (possibly, +discontinuous) positive function `r` and each tag belongs to the corresponding closed box. We also +require an upper estimate on the distortion of all boxes of the partition. -/ +def GP : integration_params := ⊥ + lemma Henstock_le_Riemann : Henstock ≤ Riemann := dec_trivial lemma Henstock_le_McShane : Henstock ≤ McShane := dec_trivial +lemma GP_le : GP ≤ l := bot_le + /-- The predicate corresponding to a base set of the filter defined by an `integration_params`. It says that diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index d931f9c9e463f..35585f5fc6dcb 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -69,13 +69,13 @@ section ### Divergence theorem for functions on `ℝⁿ⁺¹ = fin (n + 1) → ℝ`. In this section we use the divergence theorem for a Henstock-Kurzweil-like integral -`box_integral.has_integral_bot_divergence_of_forall_has_deriv_within_at` to prove the divergence +`box_integral.has_integral_GP_divergence_of_forall_has_deriv_within_at` to prove the divergence theorem for Bochner integral. The divergence theorem for Bochner integral `measure_theory.integral_divergence_of_has_fderiv_within_at_off_countable` assumes that the function itself is continuous on a closed box, differentiable at all but countably many points of its interior, and the divergence is integrable on the box. -This statement differs from `box_integral.has_integral_bot_divergence_of_forall_has_deriv_within_at` +This statement differs from `box_integral.has_integral_GP_divergence_of_forall_has_deriv_within_at` in several aspects. * We use Bochner integral instead of a Henstock-Kurzweil integral. This modification is done in @@ -93,7 +93,7 @@ in several aspects. /-- An auxiliary lemma for `measure_theory.integral_divergence_of_has_fderiv_within_at_off_countable`. This is exactly -`box_integral.has_integral_bot_divergence_of_forall_has_deriv_within_at` reformulated for the +`box_integral.has_integral_GP_divergence_of_forall_has_deriv_within_at` reformulated for the Bochner integral. -/ lemma integral_divergence_of_has_fderiv_within_at_off_countable_aux₁ (I : box (fin (n + 1))) (f : ℝⁿ⁺¹ → Eⁿ⁺¹) (f' : ℝⁿ⁺¹ → ℝⁿ⁺¹ →L[ℝ] Eⁿ⁺¹) (s : set ℝⁿ⁺¹) (hs : s.countable) @@ -106,7 +106,7 @@ lemma integral_divergence_of_has_fderiv_within_at_off_countable_aux₁ (I : box begin simp only [← set_integral_congr_set_ae (box.coe_ae_eq_Icc _)], have A := ((Hi.mono_set box.coe_subset_Icc).has_box_integral ⊥ rfl), - have B := has_integral_bot_divergence_of_forall_has_deriv_within_at I f f' (s ∩ I.Icc) + have B := has_integral_GP_divergence_of_forall_has_deriv_within_at I f f' (s ∩ I.Icc) (hs.mono (inter_subset_left _ _)) (λ x hx, Hc _ hx.2) (λ x hx, Hd _ ⟨hx.1, λ h, hx.2 ⟨h, hx.1⟩⟩), rw continuous_on_pi at Hc, From 3c9274879ed5c6c3a572a0db4d0b3ebf4e85a048 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Fri, 2 Sep 2022 10:49:43 +0000 Subject: [PATCH 132/828] feat(topology/separation): `embedding.discrete_topology` (#16092) This PR adds a short lemma `embedding.discrete_topology` and golfs the related lemma `discrete_topology_induced`. Co-authored-by: tb65536 --- src/topology/separation.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/topology/separation.lean b/src/topology/separation.lean index cf3cdedc41555..2ef12a338bee0 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -720,11 +720,11 @@ eq_of_nhds_eq_nhds (by simp [nhds_induced, ← set.image_singleton, hf.preimage_ is the discrete topology on `X`. -/ lemma discrete_topology_induced {X Y : Type*} [tY : topological_space Y] [discrete_topology Y] {f : X → Y} (hf : function.injective f) : @discrete_topology X (topological_space.induced f tY) := -begin - constructor, - rw discrete_topology.eq_bot Y, - exact induced_bot hf -end +by apply discrete_topology.mk; by rw [discrete_topology.eq_bot Y, induced_bot hf] + +lemma embedding.discrete_topology {X Y : Type*} [topological_space X] [tY : topological_space Y] + [discrete_topology Y] {f : X → Y} (hf : embedding f) : discrete_topology X := +⟨by rw [hf.induced, discrete_topology.eq_bot Y, induced_bot hf.inj]⟩ /-- Let `s, t ⊆ X` be two subsets of a topological space `X`. If `t ⊆ s` and the topology induced by `X`on `s` is discrete, then also the topology induces on `t` is discrete. -/ From d6789b8b5fe547c63f0656894ab63acbe9f04b19 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Fri, 2 Sep 2022 10:49:44 +0000 Subject: [PATCH 133/828] refactor(data/list/forall2): consistent names of relations (#16098) Co-authored-by: madvorak --- src/data/list/forall2.lean | 122 ++++++++++++++++++------------------- 1 file changed, 60 insertions(+), 62 deletions(-) diff --git a/src/data/list/forall2.lean b/src/data/list/forall2.lean index 3c04c6b97d798..4eea34cfad2d6 100644 --- a/src/data/list/forall2.lean +++ b/src/data/list/forall2.lean @@ -17,35 +17,35 @@ open nat function namespace list -variables {α β γ δ : Type*} {r : α → β → Prop} {p : γ → δ → Prop} +variables {α β γ δ : Type*} {R S : α → β → Prop} {P : γ → δ → Prop} {Rₐ : α → α → Prop} open relator mk_iff_of_inductive_prop list.forall₂ list.forall₂_iff -@[simp] theorem forall₂_cons {R : α → β → Prop} {a b l₁ l₂} : +@[simp] theorem forall₂_cons {a b l₁ l₂} : forall₂ R (a :: l₁) (b :: l₂) ↔ R a b ∧ forall₂ R l₁ l₂ := ⟨λ h, by cases h with h₁ h₂; split; assumption, λ ⟨h₁, h₂⟩, forall₂.cons h₁ h₂⟩ -theorem forall₂.imp {R S : α → β → Prop} +theorem forall₂.imp (H : ∀ a b, R a b → S a b) {l₁ l₂} (h : forall₂ R l₁ l₂) : forall₂ S l₁ l₂ := by induction h; constructor; solve_by_elim -lemma forall₂.mp {r q s : α → β → Prop} (h : ∀ a b, r a b → q a b → s a b) : - ∀ {l₁ l₂}, forall₂ r l₁ l₂ → forall₂ q l₁ l₂ → forall₂ s l₁ l₂ +lemma forall₂.mp {Q : α → β → Prop} (h : ∀ a b, Q a b → R a b → S a b) : + ∀ {l₁ l₂}, forall₂ Q l₁ l₂ → forall₂ R l₁ l₂ → forall₂ S l₁ l₂ | [] [] forall₂.nil forall₂.nil := forall₂.nil | (a :: l₁) (b :: l₂) (forall₂.cons hr hrs) (forall₂.cons hq hqs) := forall₂.cons (h a b hr hq) (forall₂.mp hrs hqs) -lemma forall₂.flip : ∀ {a b}, forall₂ (flip r) b a → forall₂ r a b +lemma forall₂.flip : ∀ {a b}, forall₂ (flip R) b a → forall₂ R a b | _ _ forall₂.nil := forall₂.nil | (a :: as) (b :: bs) (forall₂.cons h₁ h₂) := forall₂.cons h₁ h₂.flip -@[simp] lemma forall₂_same {r : α → α → Prop} : ∀ {l : list α}, forall₂ r l l ↔ ∀ x ∈ l, r x x +@[simp] lemma forall₂_same : ∀ {l : list α}, forall₂ Rₐ l l ↔ ∀ x ∈ l, Rₐ x x | [] := by simp | (a :: l) := by simp [@forall₂_same l] -lemma forall₂_refl {r} [is_refl α r] (l : list α) : forall₂ r l l := +lemma forall₂_refl [is_refl α Rₐ] (l : list α) : forall₂ Rₐ l l := forall₂_same.2 $ λ a h, refl _ @[simp] lemma forall₂_eq_eq_eq : forall₂ ((=) : α → α → Prop) = (=) := @@ -56,74 +56,74 @@ begin { rintro rfl, exact forall₂_refl _ } end -@[simp, priority 900] lemma forall₂_nil_left_iff {l} : forall₂ r nil l ↔ l = nil := +@[simp, priority 900] lemma forall₂_nil_left_iff {l} : forall₂ R nil l ↔ l = nil := ⟨λ H, by cases H; refl, by rintro rfl; exact forall₂.nil⟩ -@[simp, priority 900] lemma forall₂_nil_right_iff {l} : forall₂ r l nil ↔ l = nil := +@[simp, priority 900] lemma forall₂_nil_right_iff {l} : forall₂ R l nil ↔ l = nil := ⟨λ H, by cases H; refl, by rintro rfl; exact forall₂.nil⟩ lemma forall₂_cons_left_iff {a l u} : - forall₂ r (a :: l) u ↔ (∃b u', r a b ∧ forall₂ r l u' ∧ u = b :: u') := + forall₂ R (a :: l) u ↔ (∃b u', R a b ∧ forall₂ R l u' ∧ u = b :: u') := iff.intro (λ h, match u, h with (b :: u'), forall₂.cons h₁ h₂ := ⟨b, u', h₁, h₂, rfl⟩ end) (λ h, match u, h with _, ⟨b, u', h₁, h₂, rfl⟩ := forall₂.cons h₁ h₂ end) lemma forall₂_cons_right_iff {b l u} : - forall₂ r u (b :: l) ↔ (∃a u', r a b ∧ forall₂ r u' l ∧ u = a :: u') := + forall₂ R u (b :: l) ↔ (∃a u', R a b ∧ forall₂ R u' l ∧ u = a :: u') := iff.intro (λ h, match u, h with (b :: u'), forall₂.cons h₁ h₂ := ⟨b, u', h₁, h₂, rfl⟩ end) (λ h, match u, h with _, ⟨b, u', h₁, h₂, rfl⟩ := forall₂.cons h₁ h₂ end) -lemma forall₂_and_left {r : α → β → Prop} {p : α → Prop} : - ∀ l u, forall₂ (λa b, p a ∧ r a b) l u ↔ (∀ a∈l, p a) ∧ forall₂ r l u +lemma forall₂_and_left {p : α → Prop} : + ∀ l u, forall₂ (λa b, p a ∧ R a b) l u ↔ (∀ a∈l, p a) ∧ forall₂ R l u | [] u := by simp only [forall₂_nil_left_iff, forall_prop_of_false (not_mem_nil _), imp_true_iff, true_and] | (a :: l) u := by simp only [forall₂_and_left l, forall₂_cons_left_iff, forall_mem_cons, and_assoc, and_comm, and.left_comm, exists_and_distrib_left.symm] @[simp] lemma forall₂_map_left_iff {f : γ → α} : - ∀ {l u}, forall₂ r (map f l) u ↔ forall₂ (λc b, r (f c) b) l u + ∀ {l u}, forall₂ R (map f l) u ↔ forall₂ (λc b, R (f c) b) l u | [] _ := by simp only [map, forall₂_nil_left_iff] | (a :: l) _ := by simp only [map, forall₂_cons_left_iff, forall₂_map_left_iff] @[simp] lemma forall₂_map_right_iff {f : γ → β} : - ∀ {l u}, forall₂ r l (map f u) ↔ forall₂ (λa c, r a (f c)) l u + ∀ {l u}, forall₂ R l (map f u) ↔ forall₂ (λa c, R a (f c)) l u | _ [] := by simp only [map, forall₂_nil_right_iff] | _ (b :: u) := by simp only [map, forall₂_cons_right_iff, forall₂_map_right_iff] -lemma left_unique_forall₂' (hr : left_unique r) : - ∀ {a b c}, forall₂ r a c → forall₂ r b c → a = b +lemma left_unique_forall₂' (hr : left_unique R) : + ∀ {a b c}, forall₂ R a c → forall₂ R b c → a = b | a₀ nil a₁ forall₂.nil forall₂.nil := rfl | (a₀ :: l₀) (b :: l) (a₁ :: l₁) (forall₂.cons ha₀ h₀) (forall₂.cons ha₁ h₁) := hr ha₀ ha₁ ▸ left_unique_forall₂' h₀ h₁ ▸ rfl -lemma _root_.relator.left_unique.forall₂ (hr : left_unique r) : left_unique (forall₂ r) := +lemma _root_.relator.left_unique.forall₂ (hr : left_unique R) : left_unique (forall₂ R) := @left_unique_forall₂' _ _ _ hr -lemma right_unique_forall₂' (hr : right_unique r) : ∀ {a b c}, forall₂ r a b → forall₂ r a c → b = c +lemma right_unique_forall₂' (hr : right_unique R) : ∀ {a b c}, forall₂ R a b → forall₂ R a c → b = c | nil a₀ a₁ forall₂.nil forall₂.nil := rfl | (b :: l) (a₀ :: l₀) (a₁ :: l₁) (forall₂.cons ha₀ h₀) (forall₂.cons ha₁ h₁) := hr ha₀ ha₁ ▸ right_unique_forall₂' h₀ h₁ ▸ rfl -lemma _root_.relator.right_unique.forall₂ (hr : right_unique r) : right_unique (forall₂ r) := +lemma _root_.relator.right_unique.forall₂ (hr : right_unique R) : right_unique (forall₂ R) := @right_unique_forall₂' _ _ _ hr -lemma _root_.relator.bi_unique.forall₂ (hr : bi_unique r) : bi_unique (forall₂ r) := +lemma _root_.relator.bi_unique.forall₂ (hr : bi_unique R) : bi_unique (forall₂ R) := ⟨hr.left.forall₂, hr.right.forall₂⟩ -theorem forall₂.length_eq {R : α → β → Prop} : +theorem forall₂.length_eq : ∀ {l₁ l₂}, forall₂ R l₁ l₂ → length l₁ = length l₂ | _ _ forall₂.nil := rfl | _ _ (forall₂.cons h₁ h₂) := congr_arg succ (forall₂.length_eq h₂) theorem forall₂.nth_le : - ∀ {x : list α} {y : list β} (h : forall₂ r x y) ⦃i : ℕ⦄ (hx : i < x.length) (hy : i < y.length), - r (x.nth_le i hx) (y.nth_le i hy) + ∀ {x : list α} {y : list β} (h : forall₂ R x y) ⦃i : ℕ⦄ (hx : i < x.length) (hy : i < y.length), + R (x.nth_le i hx) (y.nth_le i hy) | (a₁ :: l₁) (a₂ :: l₂) (forall₂.cons ha hl) 0 hx hy := ha | (a₁ :: l₁) (a₂ :: l₂) (forall₂.cons ha hl) (succ i) hx hy := hl.nth_le _ _ lemma forall₂_of_length_eq_of_nth_le : ∀ {x : list α} {y : list β}, - x.length = y.length → (∀ i h₁ h₂, r (x.nth_le i h₁) (y.nth_le i h₂)) → forall₂ r x y + x.length = y.length → (∀ i h₁ h₂, R (x.nth_le i h₁) (y.nth_le i h₂)) → forall₂ R x y | [] [] hl h := forall₂.nil | (a₁ :: l₁) (a₂ :: l₂) hl h := forall₂.cons (h 0 (nat.zero_lt_succ _) (nat.zero_lt_succ _)) @@ -131,17 +131,17 @@ lemma forall₂_of_length_eq_of_nth_le : ∀ {x : list α} {y : list β}, λ i h₁ h₂, h i.succ (succ_lt_succ h₁) (succ_lt_succ h₂))) theorem forall₂_iff_nth_le {l₁ : list α} {l₂ : list β} : - forall₂ r l₁ l₂ ↔ l₁.length = l₂.length ∧ ∀ i h₁ h₂, r (l₁.nth_le i h₁) (l₂.nth_le i h₂) := + forall₂ R l₁ l₂ ↔ l₁.length = l₂.length ∧ ∀ i h₁ h₂, R (l₁.nth_le i h₁) (l₂.nth_le i h₂) := ⟨λ h, ⟨h.length_eq, h.nth_le⟩, and.rec forall₂_of_length_eq_of_nth_le⟩ -theorem forall₂_zip {R : α → β → Prop} : +theorem forall₂_zip : ∀ {l₁ l₂}, forall₂ R l₁ l₂ → ∀ {a b}, (a, b) ∈ zip l₁ l₂ → R a b | _ _ (forall₂.cons h₁ h₂) x y (or.inl rfl) := h₁ | _ _ (forall₂.cons h₁ h₂) x y (or.inr h₃) := forall₂_zip h₂ h₃ -theorem forall₂_iff_zip {R : α → β → Prop} {l₁ l₂} : forall₂ R l₁ l₂ ↔ +theorem forall₂_iff_zip {l₁ l₂} : forall₂ R l₁ l₂ ↔ length l₁ = length l₂ ∧ ∀ {a b}, (a, b) ∈ zip l₁ l₂ → R a b := -⟨λ h, ⟨h.length_eq, @forall₂_zip _ _ _ _ _ h⟩, +⟨λ h, ⟨forall₂.length_eq h, @forall₂_zip _ _ _ _ _ h⟩, λ h, begin cases h with h₁ h₂, induction l₁ with a l₁ IH generalizing l₂, @@ -150,43 +150,43 @@ theorem forall₂_iff_zip {R : α → β → Prop} {l₁ l₂} : forall₂ R l exact forall₂.cons (h₂ $ or.inl rfl) (IH h₁ $ λ a b h, h₂ $ or.inr h) } end⟩ -theorem forall₂_take {R : α → β → Prop} : +theorem forall₂_take : ∀ n {l₁ l₂}, forall₂ R l₁ l₂ → forall₂ R (take n l₁) (take n l₂) | 0 _ _ _ := by simp only [forall₂.nil, take] | (n+1) _ _ (forall₂.nil) := by simp only [forall₂.nil, take] | (n+1) _ _ (forall₂.cons h₁ h₂) := by simp [and.intro h₁ h₂, forall₂_take n] -theorem forall₂_drop {R : α → β → Prop} : +theorem forall₂_drop : ∀ n {l₁ l₂}, forall₂ R l₁ l₂ → forall₂ R (drop n l₁) (drop n l₂) | 0 _ _ h := by simp only [drop, h] | (n+1) _ _ (forall₂.nil) := by simp only [forall₂.nil, drop] | (n+1) _ _ (forall₂.cons h₁ h₂) := by simp [and.intro h₁ h₂, forall₂_drop n] -theorem forall₂_take_append {R : α → β → Prop} (l : list α) (l₁ : list β) (l₂ : list β) +theorem forall₂_take_append (l : list α) (l₁ : list β) (l₂ : list β) (h : forall₂ R l (l₁ ++ l₂)) : forall₂ R (list.take (length l₁) l) l₁ := have h': forall₂ R (take (length l₁) l) (take (length l₁) (l₁ ++ l₂)), from forall₂_take (length l₁) h, by rwa [take_left] at h' -theorem forall₂_drop_append {R : α → β → Prop} (l : list α) (l₁ : list β) (l₂ : list β) +theorem forall₂_drop_append (l : list α) (l₁ : list β) (l₂ : list β) (h : forall₂ R l (l₁ ++ l₂)) : forall₂ R (list.drop (length l₁) l) l₂ := have h': forall₂ R (drop (length l₁) l) (drop (length l₁) (l₁ ++ l₂)), from forall₂_drop (length l₁) h, by rwa [drop_left] at h' -lemma rel_mem (hr : bi_unique r) : (r ⇒ forall₂ r ⇒ iff) (∈) (∈) +lemma rel_mem (hr : bi_unique R) : (R ⇒ forall₂ R ⇒ iff) (∈) (∈) | a b h [] [] forall₂.nil := by simp only [not_mem_nil] | a b h (a' :: as) (b' :: bs) (forall₂.cons h₁ h₂) := rel_or (rel_eq hr h h₁) (rel_mem h h₂) -lemma rel_map : ((r ⇒ p) ⇒ forall₂ r ⇒ forall₂ p) map map +lemma rel_map : ((R ⇒ P) ⇒ forall₂ R ⇒ forall₂ P) map map | f g h [] [] forall₂.nil := forall₂.nil | f g h (a :: as) (b :: bs) (forall₂.cons h₁ h₂) := forall₂.cons (h h₁) (rel_map @h h₂) -lemma rel_append : (forall₂ r ⇒ forall₂ r ⇒ forall₂ r) append append +lemma rel_append : (forall₂ R ⇒ forall₂ R ⇒ forall₂ R) append append | [] [] h l₁ l₂ hl := hl | (a :: as) (b :: bs) (forall₂.cons h₁ h₂) l₁ l₂ hl := forall₂.cons h₁ (rel_append h₂ hl) -lemma rel_reverse : (forall₂ r ⇒ forall₂ r) reverse reverse +lemma rel_reverse : (forall₂ R ⇒ forall₂ R) reverse reverse | [] [] forall₂.nil := forall₂.nil | (a :: as) (b :: bs) (forall₂.cons h₁ h₂) := begin simp only [reverse_cons], @@ -194,29 +194,29 @@ lemma rel_reverse : (forall₂ r ⇒ forall₂ r) reverse reverse end @[simp] -lemma forall₂_reverse_iff {l₁ l₂} : forall₂ r (reverse l₁) (reverse l₂) ↔ forall₂ r l₁ l₂ := +lemma forall₂_reverse_iff {l₁ l₂} : forall₂ R (reverse l₁) (reverse l₂) ↔ forall₂ R l₁ l₂ := iff.intro (λ h, by { rw [← reverse_reverse l₁, ← reverse_reverse l₂], exact rel_reverse h }) (λ h, rel_reverse h) -lemma rel_join : (forall₂ (forall₂ r) ⇒ forall₂ r) join join +lemma rel_join : (forall₂ (forall₂ R) ⇒ forall₂ R) join join | [] [] forall₂.nil := forall₂.nil | (a :: as) (b :: bs) (forall₂.cons h₁ h₂) := rel_append h₁ (rel_join h₂) -lemma rel_bind : (forall₂ r ⇒ (r ⇒ forall₂ p) ⇒ forall₂ p) list.bind list.bind := +lemma rel_bind : (forall₂ R ⇒ (R ⇒ forall₂ P) ⇒ forall₂ P) list.bind list.bind := λ a b h₁ f g h₂, rel_join (rel_map @h₂ h₁) -lemma rel_foldl : ((p ⇒ r ⇒ p) ⇒ p ⇒ forall₂ r ⇒ p) foldl foldl +lemma rel_foldl : ((P ⇒ R ⇒ P) ⇒ P ⇒ forall₂ R ⇒ P) foldl foldl | f g hfg _ _ h _ _ forall₂.nil := h | f g hfg x y hxy _ _ (forall₂.cons hab hs) := rel_foldl @hfg (hfg hxy hab) hs -lemma rel_foldr : ((r ⇒ p ⇒ p) ⇒ p ⇒ forall₂ r ⇒ p) foldr foldr +lemma rel_foldr : ((R ⇒ P ⇒ P) ⇒ P ⇒ forall₂ R ⇒ P) foldr foldr | f g hfg _ _ h _ _ forall₂.nil := h | f g hfg x y hxy _ _ (forall₂.cons hab hs) := hfg hab (rel_foldr @hfg hxy hs) lemma rel_filter {p : α → Prop} {q : β → Prop} [decidable_pred p] [decidable_pred q] - (hpq : (r ⇒ (↔)) p q) : - (forall₂ r ⇒ forall₂ r) (filter p) (filter q) + (hpq : (R ⇒ (↔)) p q) : + (forall₂ R ⇒ forall₂ R) (filter p) (filter q) | _ _ forall₂.nil := forall₂.nil | (a :: as) (b :: bs) (forall₂.cons h₁ h₂) := begin @@ -228,7 +228,7 @@ lemma rel_filter {p : α → Prop} {q : β → Prop} [decidable_pred p] [decidab simp only [filter_cons_of_neg _ h, filter_cons_of_neg _ this, rel_filter h₂], }, end -lemma rel_filter_map : ((r ⇒ option.rel p) ⇒ forall₂ r ⇒ forall₂ p) filter_map filter_map +lemma rel_filter_map : ((R ⇒ option.rel P) ⇒ forall₂ R ⇒ forall₂ P) filter_map filter_map | f g hfg _ _ forall₂.nil := forall₂.nil | f g hfg (a :: as) (b :: bs) (forall₂.cons h₁ h₂) := by rw [filter_map_cons, filter_map_cons]; @@ -239,19 +239,19 @@ lemma rel_filter_map : ((r ⇒ option.rel p) ⇒ forall₂ r ⇒ forall₂ p) fi @[to_additive] lemma rel_prod [monoid α] [monoid β] - (h : r 1 1) (hf : (r ⇒ r ⇒ r) (*) (*)) : (forall₂ r ⇒ r) prod prod := + (h : R 1 1) (hf : (R ⇒ R ⇒ R) (*) (*)) : (forall₂ R ⇒ R) prod prod := rel_foldl hf h -/-- Given a relation `r`, `sublist_forall₂ r l₁ l₂` indicates that there is a sublist of `l₂` such +/-- Given a relation `R`, `sublist_forall₂ r l₁ l₂` indicates that there is a sublist of `l₂` such that `forall₂ r l₁ l₂`. -/ -inductive sublist_forall₂ (r : α → β → Prop) : list α → list β → Prop +inductive sublist_forall₂ (R : α → β → Prop) : list α → list β → Prop | nil {l} : sublist_forall₂ [] l -| cons {a₁ a₂ l₁ l₂} : r a₁ a₂ → sublist_forall₂ l₁ l₂ → +| cons {a₁ a₂ l₁ l₂} : R a₁ a₂ → sublist_forall₂ l₁ l₂ → sublist_forall₂ (a₁ :: l₁) (a₂ :: l₂) | cons_right {a l₁ l₂} : sublist_forall₂ l₁ l₂ → sublist_forall₂ l₁ (a :: l₂) lemma sublist_forall₂_iff {l₁ : list α} {l₂ : list β} : - sublist_forall₂ r l₁ l₂ ↔ ∃ l, forall₂ r l₁ l ∧ l <+ l₂ := + sublist_forall₂ R l₁ l₂ ↔ ∃ l, forall₂ R l₁ l ∧ l <+ l₂ := begin split; intro h, { induction h with _ a b l1 l2 rab rll ih b l1 l2 hl ih, @@ -270,14 +270,12 @@ begin exact sublist_forall₂.cons hr (ih hl) } } end -variable {ra : α → α → Prop} - -instance sublist_forall₂.is_refl [is_refl α ra] : - is_refl (list α) (sublist_forall₂ ra) := +instance sublist_forall₂.is_refl [is_refl α Rₐ] : + is_refl (list α) (sublist_forall₂ Rₐ) := ⟨λ l, sublist_forall₂_iff.2 ⟨l, forall₂_refl l, sublist.refl l⟩⟩ -instance sublist_forall₂.is_trans [is_trans α ra] : - is_trans (list α) (sublist_forall₂ ra) := +instance sublist_forall₂.is_trans [is_trans α Rₐ] : + is_trans (list α) (sublist_forall₂ Rₐ) := ⟨λ a b c, begin revert a b, induction c with _ _ ih, @@ -294,12 +292,12 @@ instance sublist_forall₂.is_trans [is_trans α ra] : { exact sublist_forall₂.cons_right (ih _ _ h1 btc), } } end⟩ -lemma sublist.sublist_forall₂ {l₁ l₂ : list α} (h : l₁ <+ l₂) (r : α → α → Prop) [is_refl α r] : - sublist_forall₂ r l₁ l₂ := +lemma sublist.sublist_forall₂ {l₁ l₂ : list α} (h : l₁ <+ l₂) [is_refl α Rₐ] : + sublist_forall₂ Rₐ l₁ l₂ := sublist_forall₂_iff.2 ⟨l₁, forall₂_refl l₁, h⟩ -lemma tail_sublist_forall₂_self [is_refl α ra] (l : list α) : - sublist_forall₂ ra l.tail l := -l.tail_sublist.sublist_forall₂ ra +lemma tail_sublist_forall₂_self [is_refl α Rₐ] (l : list α) : + sublist_forall₂ Rₐ l.tail l := +l.tail_sublist.sublist_forall₂ end list From 1555e3deca9121aa1d9bda92fe7c47f6a2863b4f Mon Sep 17 00:00:00 2001 From: tb65536 Date: Fri, 2 Sep 2022 10:49:45 +0000 Subject: [PATCH 134/828] feat(set_theory/cardinal/finite): `nat.card_pi` (#16143) This PR adds `nat.card_pi`. Co-authored-by: tb65536 --- src/set_theory/cardinal/basic.lean | 32 ++++++++++++++++++++++++++++- src/set_theory/cardinal/finite.lean | 4 ++++ 2 files changed, 35 insertions(+), 1 deletion(-) diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index fe6ffafe84351..2798b655518f7 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Floris van Doorn -/ +import data.fintype.card import data.finsupp.defs import data.nat.part_enat import data.set.countable @@ -75,7 +76,7 @@ Cantor's theorem, König's theorem, Konig's theorem -/ open function set order -open_locale classical +open_locale big_operators classical noncomputable theory @@ -716,6 +717,23 @@ begin exact mk_congr (equiv.ulift.trans $ equiv.Pi_congr_right $ λ i, equiv.ulift.symm) end +lemma prod_eq_of_fintype {α : Type u} [fintype α] (f : α → cardinal.{v}) : + prod f = cardinal.lift.{u} (∏ i, f i) := +begin + revert f, + refine fintype.induction_empty_option _ _ _ α, + { introsI α β hβ e h f, + letI := fintype.of_equiv β e.symm, + rw [←e.prod_comp f, ←h], + exact mk_congr (e.Pi_congr_left _).symm }, + { intro f, + rw [fintype.univ_pempty, finset.prod_empty, lift_one, cardinal.prod, mk_eq_one] }, + { intros α hα h f, + rw [cardinal.prod, mk_congr equiv.pi_option_equiv_prod, mk_prod, lift_umax', mk_out, + ←cardinal.prod, lift_prod, fintype.prod_option, lift_mul, ←h (λ a, f (some a))], + simp only [lift_id] }, +end + @[simp] theorem lift_Inf (s : set cardinal) : lift (Inf s) = Inf (lift '' s) := begin rcases eq_empty_or_nonempty s with rfl | hs, @@ -1179,6 +1197,18 @@ begin exact aleph_0_le_mul_iff'.2 (or.inr ⟨hx2, hy1⟩) }, end +/-- `cardinal.to_nat` as a `monoid_with_zero_hom`. -/ +@[simps] +def to_nat_hom : cardinal →*₀ ℕ := +{ to_fun := to_nat, + map_zero' := zero_to_nat, + map_one' := one_to_nat, + map_mul' := to_nat_mul } + +lemma to_nat_finset_prod (s : finset α) (f : α → cardinal) : + to_nat (∏ i in s, f i) = ∏ i in s, to_nat (f i) := +map_prod to_nat_hom _ _ + @[simp] lemma to_nat_add_of_lt_aleph_0 {a : cardinal.{u}} {b : cardinal.{v}} (ha : a < ℵ₀) (hb : b < ℵ₀) : ((lift.{v u} a) + (lift.{u v} b)).to_nat = a.to_nat + b.to_nat := begin diff --git a/src/set_theory/cardinal/finite.lean b/src/set_theory/cardinal/finite.lean index 0b79fc1eb4b05..7e1708e1a028e 100644 --- a/src/set_theory/cardinal/finite.lean +++ b/src/set_theory/cardinal/finite.lean @@ -19,6 +19,7 @@ import set_theory.cardinal.basic open cardinal noncomputable theory +open_locale big_operators variables {α β : Type*} @@ -77,6 +78,9 @@ card_congr equiv.ulift @[simp] lemma card_plift (α : Type*) : nat.card (plift α) = nat.card α := card_congr equiv.plift +lemma card_pi {β : α → Type*} [fintype α] : nat.card (Π a, β a) = ∏ a, nat.card (β a) := +by simp_rw [nat.card, mk_pi, prod_eq_of_fintype, to_nat_lift, to_nat_finset_prod] + @[simp] lemma card_zmod (n : ℕ) : nat.card (zmod n) = n := begin cases n, From 62f14c33bfdb4e29ffb18f25420d367b08a62f99 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 2 Sep 2022 10:49:46 +0000 Subject: [PATCH 135/828] feat(category_theory/limits/has_limits): `const` and `(co)lim` are adjoint (#16251) --- src/category_theory/limits/has_limits.lean | 32 ++++++++++++++++++++++ 1 file changed, 32 insertions(+) diff --git a/src/category_theory/limits/has_limits.lean b/src/category_theory/limits/has_limits.lean index d2d87212c65e4..b409eb045bf50 100644 --- a/src/category_theory/limits/has_limits.lean +++ b/src/category_theory/limits/has_limits.lean @@ -488,6 +488,22 @@ def lim_yoneda : lim ⋙ yoneda ⋙ (whiskering_right _ _ _).obj ulift_functor.{ nat_iso.of_components (λ F, nat_iso.of_components (λ W, limit.hom_iso F (unop W)) (by tidy)) (by tidy) +/--The constant functor and limit functor are adjoint to each other-/ +def const_lim_adj : (const J : C ⥤ (J ⥤ C)) ⊣ lim := +{ hom_equiv := λ c g, + { to_fun := λ f, limit.lift _ ⟨c, f⟩, + inv_fun := λ f, { app := λ j, f ≫ limit.π _ _ , naturality' := by tidy }, + left_inv := λ _, nat_trans.ext _ _ $ funext $ λ j, limit.lift_π _ _, + right_inv := λ α, limit.hom_ext $ λ j, limit.lift_π _ _ }, + unit := { app := λ c, limit.lift _ ⟨_, 𝟙 _⟩, naturality' := λ _ _ _, by tidy }, + counit := + { app := λ g, { app := limit.π _, naturality' := by tidy }, + naturality' := λ _ _ _, by tidy }, + hom_equiv_unit' := λ c g f, limit.hom_ext $ λ j, by simp, + hom_equiv_counit' := λ c g f, nat_trans.ext _ _ $ funext $ λ j, rfl } + +instance : is_right_adjoint (lim : (J ⥤ C) ⥤ C) := ⟨_, const_lim_adj⟩ + end lim_functor /-- @@ -972,6 +988,22 @@ def colim_coyoneda : colim.op ⋙ coyoneda ⋙ (whiskering_right _ _ _).obj ulif nat_iso.of_components (λ F, nat_iso.of_components (colimit.hom_iso (unop F)) (by tidy)) (by tidy) +/-- +The colimit functor and constant functor are adjoint to each other +-/ +def colim_const_adj : (colim : (J ⥤ C) ⥤ C) ⊣ const J := +{ hom_equiv := λ f c, + { to_fun := λ g, { app := λ _, colimit.ι _ _ ≫ g, naturality' := by tidy }, + inv_fun := λ g, colimit.desc _ ⟨_, g⟩, + left_inv := λ _, colimit.hom_ext $ λ j, colimit.ι_desc _ _, + right_inv := λ _, nat_trans.ext _ _ $ funext $ λ j, colimit.ι_desc _ _ }, + unit := { app := λ g, { app := colimit.ι _, naturality' := by tidy }, naturality' := by tidy }, + counit := { app := λ c, colimit.desc _ ⟨_, 𝟙 _⟩, naturality' := by tidy }, + hom_equiv_unit' := λ _ _ _, nat_trans.ext _ _ $ funext $ λ _ , rfl, + hom_equiv_counit' := λ _ _ _, colimit.hom_ext $ λ _, by simp } + +instance : is_left_adjoint (colim : (J ⥤ C) ⥤ C) := ⟨_, colim_const_adj⟩ + end colim_functor /-- From 8bedeafb7013df82cc02e6629c7e1b1c73e4311e Mon Sep 17 00:00:00 2001 From: negiizhao Date: Fri, 2 Sep 2022 10:49:48 +0000 Subject: [PATCH 136/828] feat(data/nat/enat): derive typeclasses about well founded (#16346) --- src/data/nat/enat.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/data/nat/enat.lean b/src/data/nat/enat.lean index 8f8e7b4060d50..7a2e3ea159366 100644 --- a/src/data/nat/enat.lean +++ b/src/data/nat/enat.lean @@ -17,7 +17,7 @@ about this type. @[derive [has_zero, add_comm_monoid_with_one, canonically_ordered_comm_semiring, nontrivial, linear_order, order_bot, order_top, has_bot, has_top, canonically_linear_ordered_add_monoid, has_sub, has_ordered_sub, complete_linear_order, linear_ordered_add_comm_monoid_with_top, - succ_order]] + succ_order, well_founded_lt, has_well_founded]] def enat : Type := with_top ℕ notation `ℕ∞` := enat @@ -25,6 +25,7 @@ notation `ℕ∞` := enat namespace enat instance : inhabited ℕ∞ := ⟨0⟩ +instance : is_well_order ℕ∞ (<) := { } variables {m n : ℕ∞} From 9556784a5b84697562e9c6acb40500d4a82e675a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 2 Sep 2022 10:49:49 +0000 Subject: [PATCH 137/828] chore(ring_theory/*): Fix lint (#16350) Satisfy the `fintype_finite` and `check_reducibility` linters. --- src/ring_theory/finiteness.lean | 12 +++++++----- src/ring_theory/ideal/quotient.lean | 12 +++++++----- src/ring_theory/jacobson.lean | 6 ++++-- src/ring_theory/localization/integer.lean | 5 +++-- src/ring_theory/norm.lean | 8 +++++--- src/ring_theory/nullstellensatz.lean | 3 ++- src/ring_theory/polynomial/basic.lean | 6 ++++-- src/ring_theory/trace.lean | 10 +++++----- src/ring_theory/witt_vector/basic.lean | 2 ++ 9 files changed, 39 insertions(+), 25 deletions(-) diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index eddc0aa982511..375ca0041060a 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -112,7 +112,7 @@ instance prod [hM : finite R M] [hN : finite R N] : finite R (M × N) := exact hM.1.prod hN.1 end⟩ -instance pi {ι : Type*} {M : ι → Type*} [fintype ι] [Π i, add_comm_monoid (M i)] +instance pi {ι : Type*} {M : ι → Type*} [_root_.finite ι] [Π i, add_comm_monoid (M i)] [Π i, module R (M i)] [h : ∀ i, finite R (M i)] : finite R (Π i, M i) := ⟨begin rw ← submodule.pi_top, @@ -180,8 +180,8 @@ lemma self : finite_type R R := ⟨⟨{1}, subsingleton.elim _ _⟩⟩ protected lemma polynomial : finite_type R R[X] := ⟨⟨{polynomial.X}, by { rw finset.coe_singleton, exact polynomial.adjoin_X }⟩⟩ -protected lemma mv_polynomial (ι : Type*) [fintype ι] : finite_type R (mv_polynomial ι R) := -⟨⟨finset.univ.image mv_polynomial.X, +protected lemma mv_polynomial (ι : Type*) [finite ι] : finite_type R (mv_polynomial ι R) := +by casesI nonempty_fintype ι; exact ⟨⟨finset.univ.image mv_polynomial.X, by {rw [finset.coe_image, finset.coe_univ, set.image_univ], exact mv_polynomial.adjoin_range_X}⟩⟩ lemma of_restrict_scalars_finite_type [algebra A B] [is_scalar_tower R A B] [hB : finite_type R B] : @@ -325,8 +325,9 @@ end variable (R) /-- The ring of polynomials in finitely many variables is finitely presented. -/ -protected lemma mv_polynomial (ι : Type u_2) [fintype ι] : +protected lemma mv_polynomial (ι : Type u_2) [finite ι] : finite_presentation R (mv_polynomial ι R) := +by casesI nonempty_fintype ι; exact let eqv := (mv_polynomial.rename_equiv R $ fintype.equiv_fin ι).symm in ⟨fintype.card ι, eqv, eqv.surjective, ((ring_hom.injective_iff_ker_eq_bot _).1 eqv.injective).symm ▸ submodule.fg_bot⟩ @@ -395,13 +396,14 @@ end /-- If `A` is a finitely presented `R`-algebra, then `mv_polynomial (fin n) A` is finitely presented as `R`-algebra. -/ lemma mv_polynomial_of_finite_presentation (hfp : finite_presentation R A) (ι : Type*) - [fintype ι] : finite_presentation R (mv_polynomial ι A) := + [finite ι] : finite_presentation R (mv_polynomial ι A) := begin rw iff_quotient_mv_polynomial' at hfp ⊢, classical, obtain ⟨ι', _, f, hf_surj, hf_ker⟩ := hfp, resetI, let g := (mv_polynomial.map_alg_hom f).comp (mv_polynomial.sum_alg_equiv R ι ι').to_alg_hom, + casesI nonempty_fintype (ι ⊕ ι'), refine ⟨ι ⊕ ι', by apply_instance, g, (mv_polynomial.map_surjective f.to_ring_hom hf_surj).comp (alg_equiv.surjective _), ideal.fg_ker_comp _ _ _ _ (alg_equiv.surjective _)⟩, diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index 1eafbe3513dfb..65c02a553dcd5 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -309,10 +309,11 @@ noncomputable def pi_quot_equiv : ((ι → R) ⧸ I.pi ι) ≃ₗ[(R ⧸ I)] (ι /-- If `f : R^n → R^m` is an `R`-linear map and `I ⊆ R` is an ideal, then the image of `I^n` is contained in `I^m`. -/ -lemma map_pi {ι} [fintype ι] {ι' : Type w} (x : ι → R) (hi : ∀ i, x i ∈ I) +lemma map_pi {ι : Type*} [finite ι] {ι' : Type w} (x : ι → R) (hi : ∀ i, x i ∈ I) (f : (ι → R) →ₗ[R] (ι' → R)) (i : ι') : f x i ∈ I := begin classical, + casesI nonempty_fintype ι, rw pi_eq_sum_univ x, simp only [finset.sum_apply, smul_eq_mul, linear_map.map_sum, pi.smul_apply, linear_map.map_smul], exact I.sum_mem (λ j hj, I.mul_mem_right _ (hi j)) @@ -347,10 +348,11 @@ begin rw quotient.eq_zero_iff_mem, exact hgj j hjs hji end -theorem exists_sub_mem [fintype ι] {f : ι → ideal R} - (hf : ∀ i j, i ≠ j → f i ⊔ f j = ⊤) (g : ι → R) : +theorem exists_sub_mem [finite ι] {f : ι → ideal R} (hf : ∀ i j, i ≠ j → f i ⊔ f j = ⊤) + (g : ι → R) : ∃ r : R, ∀ i, r - g i ∈ f i := begin + casesI nonempty_fintype ι, have : ∃ φ : ι → R, (∀ i, φ i - 1 ∈ f i) ∧ (∀ i j, i ≠ j → φ i ∈ f j), { have := exists_sub_one_mem_and_mem (finset.univ : finset ι) (λ i _ j _ hij, hf i j hij), choose φ hφ, @@ -379,7 +381,7 @@ quotient.lift (⨅ i, f i) exact quotient.eq_zero_iff_mem.2 (hr i) end -theorem quotient_inf_to_pi_quotient_bijective [fintype ι] {f : ι → ideal R} +theorem quotient_inf_to_pi_quotient_bijective [finite ι] {f : ι → ideal R} (hf : ∀ i j, i ≠ j → f i ⊔ f j = ⊤) : function.bijective (quotient_inf_to_pi_quotient f) := ⟨λ x y, quotient.induction_on₂' x y $ λ r s hrs, quotient.eq.2 $ @@ -389,7 +391,7 @@ theorem quotient_inf_to_pi_quotient_bijective [fintype ι] {f : ι → ideal R} ⟨quotient.mk _ r, funext $ λ i, quotient.out_eq' (g i) ▸ quotient.eq.2 (hr i)⟩⟩ /-- Chinese Remainder Theorem. Eisenbud Ex.2.6. Similar to Atiyah-Macdonald 1.10 and Stacks 00DT -/ -noncomputable def quotient_inf_ring_equiv_pi_quotient [fintype ι] (f : ι → ideal R) +noncomputable def quotient_inf_ring_equiv_pi_quotient [finite ι] (f : ι → ideal R) (hf : ∀ i j, i ≠ j → f i ⊔ f j = ⊤) : R ⧸ (⨅ i, f i) ≃+* Π i, R ⧸ f i := { .. equiv.of_bijective _ (quotient_inf_to_pi_quotient_bijective hf), diff --git a/src/ring_theory/jacobson.lean b/src/ring_theory/jacobson.lean index 380a4ce050589..fd5f5e0ae5ad9 100644 --- a/src/ring_theory/jacobson.lean +++ b/src/ring_theory/jacobson.lean @@ -583,9 +583,10 @@ lemma is_jacobson_mv_polynomial_fin {R : Type*} [comm_ring R] [H : is_jacobson R `Inf {P maximal | P ≥ I} = Inf {P prime | P ≥ I} = I.radical`. Fields are always Jacobson, and in that special case this is (most of) the classical Nullstellensatz, since `I(V(I))` is the intersection of maximal ideals containing `I`, which is then `I.radical` -/ -instance {R : Type*} [comm_ring R] {ι : Type*} [fintype ι] [is_jacobson R] : +instance is_jacobson {R : Type*} [comm_ring R] {ι : Type*} [finite ι] [is_jacobson R] : is_jacobson (mv_polynomial ι R) := begin + casesI nonempty_fintype ι, haveI := classical.dec_eq ι, let e := fintype.equiv_fin ι, rw is_jacobson_iso (rename_equiv R e).to_ring_equiv, @@ -628,9 +629,10 @@ end lemma comp_C_integral_of_surjective_of_jacobson {R : Type*} [comm_ring R] [is_jacobson R] - {σ : Type*} [fintype σ] {S : Type*} [field S] (f : mv_polynomial σ R →+* S) + {σ : Type*} [finite σ] {S : Type*} [field S] (f : mv_polynomial σ R →+* S) (hf : function.surjective f) : (f.comp C).is_integral := begin + casesI nonempty_fintype σ, have e := (fintype.equiv_fin σ).symm, let f' : mv_polynomial (fin _) R →+* S := f.comp (rename_equiv R e).to_ring_equiv.to_ring_hom, diff --git a/src/ring_theory/localization/integer.lean b/src/ring_theory/localization/integer.lean index a5fdd511ce77b..512ea6a21df5e 100644 --- a/src/ring_theory/localization/integer.lean +++ b/src/ring_theory/localization/integer.lean @@ -91,10 +91,11 @@ begin refl end -/-- We can clear the denominators of a `fintype`-indexed family of fractions. -/ -lemma exist_integer_multiples_of_fintype {ι : Type*} [fintype ι] (f : ι → S) : +/-- We can clear the denominators of a finite indexed family of fractions. -/ +lemma exist_integer_multiples_of_finite {ι : Type*} [finite ι] (f : ι → S) : ∃ (b : M), ∀ i, is_localization.is_integer R ((b : R) • f i) := begin + casesI nonempty_fintype ι, obtain ⟨b, hb⟩ := exist_integer_multiples M finset.univ f, exact ⟨b, λ i, hb i (finset.mem_univ _)⟩ end diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index 5a78147d2b4c8..a4c3f9778722d 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -45,7 +45,7 @@ variables {R S T : Type*} [comm_ring R] [comm_ring S] variables [algebra R S] variables {K L F : Type*} [field K] [field L] [field F] variables [algebra K L] [algebra K F] -variables {ι : Type w} [fintype ι] +variables {ι : Type w} open finite_dimensional open linear_map @@ -71,12 +71,12 @@ by { rw [norm_apply, linear_map.det], split_ifs with h, refl } variables {R} -- Can't be a `simp` lemma because it depends on a choice of basis -lemma norm_eq_matrix_det [decidable_eq ι] (b : basis ι R S) (s : S) : +lemma norm_eq_matrix_det [fintype ι] [decidable_eq ι] (b : basis ι R S) (s : S) : norm R s = matrix.det (algebra.left_mul_matrix b s) := by { rwa [norm_apply, ← linear_map.det_to_matrix b, ← to_matrix_lmul_eq], refl } /-- If `x` is in the base field `K`, then the norm is `x ^ [L : K]`. -/ -lemma norm_algebra_map_of_basis (b : basis ι R S) (x : R) : +lemma norm_algebra_map_of_basis [fintype ι] (b : basis ι R S) (x : R) : norm R (algebra_map R S x) = x ^ fintype.card ι := begin haveI := classical.dec_eq ι, @@ -129,10 +129,12 @@ end end eq_prod_roots section eq_zero_iff +variables [finite ι] lemma norm_eq_zero_iff_of_basis [is_domain R] [is_domain S] (b : basis ι R S) {x : S} : algebra.norm R x = 0 ↔ x = 0 := begin + casesI nonempty_fintype ι, have hι : nonempty ι := b.index_nonempty, letI := classical.dec_eq ι, rw algebra.norm_eq_matrix_det b, diff --git a/src/ring_theory/nullstellensatz.lean b/src/ring_theory/nullstellensatz.lean index 2384031ea8136..ab70ecadfb579 100644 --- a/src/ring_theory/nullstellensatz.lean +++ b/src/ring_theory/nullstellensatz.lean @@ -128,11 +128,12 @@ lemma point_to_point_zero_locus_le (I : ideal (mv_polynomial σ k)) : λ J hJ, let ⟨x, hx⟩ := hJ in (le_trans (le_vanishing_ideal_zero_locus I) (hx.2 ▸ vanishing_ideal_anti_mono (set.singleton_subset_iff.2 hx.1)) : I ≤ J.as_ideal) -variables [is_alg_closed k] [fintype σ] +variables [is_alg_closed k] [finite σ] lemma is_maximal_iff_eq_vanishing_ideal_singleton (I : ideal (mv_polynomial σ k)) : I.is_maximal ↔ ∃ (x : σ → k), I = vanishing_ideal {x} := begin + casesI nonempty_fintype σ, refine ⟨λ hI, _, λ h, let ⟨x, hx⟩ := h in hx.symm ▸ (mv_polynomial.vanishing_ideal_singleton_is_maximal)⟩, letI : I.is_maximal := hI, diff --git a/src/ring_theory/polynomial/basic.lean b/src/ring_theory/polynomial/basic.lean index a6c95c9a9780c..0108ac8987e6e 100644 --- a/src/ring_theory/polynomial/basic.lean +++ b/src/ring_theory/polynomial/basic.lean @@ -987,8 +987,9 @@ theorem is_noetherian_ring_fin [is_noetherian_ring R] : /-- The multivariate polynomial ring in finitely many variables over a noetherian ring is itself a noetherian ring. -/ -instance is_noetherian_ring [fintype σ] [is_noetherian_ring R] : +instance is_noetherian_ring [finite σ] [is_noetherian_ring R] : is_noetherian_ring (mv_polynomial σ R) := +by casesI nonempty_fintype σ; exact @is_noetherian_ring_of_ring_equiv (mv_polynomial (fin (fintype.card σ)) R) _ _ _ (rename_equiv R (fintype.equiv_fin σ).symm).to_ring_equiv is_noetherian_ring_fin @@ -1039,10 +1040,11 @@ instance {R : Type u} {σ : Type v} [comm_ring R] [is_domain R] : is_domain (mv_ { .. mv_polynomial.no_zero_divisors, .. add_monoid_algebra.nontrivial } -lemma map_mv_polynomial_eq_eval₂ {S : Type*} [comm_ring S] [fintype σ] +lemma map_mv_polynomial_eq_eval₂ {S : Type*} [comm_ring S] [finite σ] (ϕ : mv_polynomial σ R →+* S) (p : mv_polynomial σ R) : ϕ p = mv_polynomial.eval₂ (ϕ.comp mv_polynomial.C) (λ s, ϕ (mv_polynomial.X s)) p := begin + casesI nonempty_fintype σ, refine trans (congr_arg ϕ (mv_polynomial.as_sum p)) _, rw [mv_polynomial.eval₂_eq', ϕ.map_sum], congr, diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index a30e20a42ef2e..4ec7c28b8cceb 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -127,13 +127,14 @@ begin { simp [trace_eq_zero_of_not_exists_basis K H, finrank_eq_zero_of_not_exists_basis_finset H] } end -lemma trace_trace_of_basis [algebra S T] [is_scalar_tower R S T] - {ι κ : Type*} [fintype ι] [fintype κ] +lemma trace_trace_of_basis [algebra S T] [is_scalar_tower R S T] {ι κ : Type*} [finite ι] [finite κ] (b : basis ι R S) (c : basis κ S T) (x : T) : trace R S (trace S T x) = trace R T x := begin haveI := classical.dec_eq ι, haveI := classical.dec_eq κ, + casesI nonempty_fintype ι, + casesI nonempty_fintype κ, rw [trace_eq_matrix_trace (b.smul c), trace_eq_matrix_trace b, trace_eq_matrix_trace c, matrix.trace, matrix.trace, matrix.trace, ← finset.univ_product_univ, finset.sum_product], @@ -143,9 +144,8 @@ begin finset.sum_apply i _ (λ y, left_mul_matrix b (left_mul_matrix c x y y))] end -lemma trace_comp_trace_of_basis [algebra S T] [is_scalar_tower R S T] - {ι κ : Type*} [fintype ι] [fintype κ] - (b : basis ι R S) (c : basis κ S T) : +lemma trace_comp_trace_of_basis [algebra S T] [is_scalar_tower R S T] {ι κ : Type*} [finite ι] + [fintype κ] (b : basis ι R S) (c : basis κ S T) : (trace R S).comp ((trace S T).restrict_scalars R) = trace R T := by { ext, rw [linear_map.comp_apply, linear_map.restrict_scalars_apply, trace_trace_of_basis b c] } diff --git a/src/ring_theory/witt_vector/basic.lean b/src/ring_theory/witt_vector/basic.lean index a3def906bb3dc..dcb9ba16ff989 100644 --- a/src/ring_theory/witt_vector/basic.lean +++ b/src/ring_theory/witt_vector/basic.lean @@ -241,6 +241,8 @@ private def comm_ring_aux₂ : comm_ring (𝕎 (mv_polynomial R ℤ)) := (map_fun.zero _) (map_fun.one _) (map_fun.add _) (map_fun.mul _) (map_fun.neg _) (map_fun.sub _) (map_fun.nsmul _) (map_fun.zsmul _) (map_fun.pow _) (map_fun.nat_cast _) (map_fun.int_cast _) +attribute [reducible] comm_ring_aux₂ + /-- The commutative ring structure on `𝕎 R`. -/ instance : comm_ring (𝕎 R) := (map_fun.surjective _ $ counit_surjective _).comm_ring (map_fun $ mv_polynomial.counit _) From 447cbe9c30a5ed6127386a013b7939e5bdad079d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 2 Sep 2022 13:28:02 +0000 Subject: [PATCH 138/828] refactor({data,linear_algebra}/matrix/block): Generalize block indexing (#14035) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Currently many definitions around block matrices are specialized to either `ℕ` or `fin n` as the indexing type. This means that some definitions are duplicated and proofs feel unnatural. This PR generalizes them to an arbitrary type (with an order) and subsequently golfs proofs. Co-authored-by: Eric Co-authored-by: Alexander Bentkamp Co-authored-by: Eric Wieser --- src/data/matrix/block.lean | 30 +-- src/linear_algebra/matrix/block.lean | 310 ++++++++++++--------------- 2 files changed, 150 insertions(+), 190 deletions(-) diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index c5df48206bb56..1d2004bf7b3a2 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -152,30 +152,22 @@ def to_block (M : matrix m n α) (p : m → Prop) (q : n → Prop) : @[simp] lemma to_block_apply (M : matrix m n α) (p : m → Prop) (q : n → Prop) (i : {a // p a}) (j : {a // q a}) : to_block M p q i j = M ↑i ↑j := rfl -/-- Let `b` map rows and columns of a square matrix `M` to blocks. Then - `to_square_block M b k` is the block `k` matrix. -/ -def to_square_block (M : matrix m m α) {n : nat} (b : m → fin n) (k : fin n) : - matrix {a // b a = k} {a // b a = k} α := M.submatrix coe coe - -@[simp] lemma to_square_block_def (M : matrix m m α) {n : nat} (b : m → fin n) (k : fin n) : - to_square_block M b k = λ i j, M ↑i ↑j := rfl - -/-- Alternate version with `b : m → nat`. Let `b` map rows and columns of a square matrix `M` to - blocks. Then `to_square_block' M b k` is the block `k` matrix. -/ -def to_square_block' (M : matrix m m α) (b : m → nat) (k : nat) : - matrix {a // b a = k} {a // b a = k} α := M.submatrix coe coe - -@[simp] lemma to_square_block_def' (M : matrix m m α) (b : m → nat) (k : nat) : - to_square_block' M b k = λ i j, M ↑i ↑j := rfl - /-- Let `p` pick out certain rows and columns of a square matrix `M`. Then `to_square_block_prop M p` is the corresponding block matrix. -/ -def to_square_block_prop (M : matrix m m α) (p : m → Prop) : - matrix {a // p a} {a // p a} α := M.submatrix coe coe +def to_square_block_prop (M : matrix m m α) (p : m → Prop) : matrix {a // p a} {a // p a} α := +to_block M _ _ -@[simp] lemma to_square_block_prop_def (M : matrix m m α) (p : m → Prop) : +lemma to_square_block_prop_def (M : matrix m m α) (p : m → Prop) : to_square_block_prop M p = λ i j, M ↑i ↑j := rfl +/-- Let `b` map rows and columns of a square matrix `M` to blocks. Then + `to_square_block M b k` is the block `k` matrix. -/ +def to_square_block (M : matrix m m α) (b : m → β) (k : β) : + matrix {a // b a = k} {a // b a = k} α := to_square_block_prop M _ + +lemma to_square_block_def (M : matrix m m α) (b : m → β) (k : β) : + to_square_block M b k = λ i j, M ↑i ↑j := rfl + lemma from_blocks_smul [has_smul R α] (x : R) (A : matrix n l α) (B : matrix n m α) (C : matrix o l α) (D : matrix o m α) : x • (from_blocks A B C D) = from_blocks (x • A) (x • B) (x • C) (x • D) := diff --git a/src/linear_algebra/matrix/block.lean b/src/linear_algebra/matrix/block.lean index 869f7e99c05df..de0d47e2d99bd 100644 --- a/src/linear_algebra/matrix/block.lean +++ b/src/linear_algebra/matrix/block.lean @@ -16,7 +16,7 @@ matrices built out of blocks. ## Main definitions * `matrix.block_triangular` expresses that a `o` by `o` matrix is block triangular, - if the rows and columns are ordered according to some order `b : o → ℕ` + if the rows and columns are ordered according to some order `b : o → α` ## Main results * `matrix.det_of_block_triangular`: the determinant of a block triangular matrix @@ -30,20 +30,86 @@ matrix, diagonal, det, block triangular -/ -open_locale big_operators +open finset function order_dual +open_locale big_operators matrix universes v -variables {m n : Type*} [decidable_eq n] [fintype n] [decidable_eq m] [fintype m] -variables {R : Type v} [comm_ring R] +variables {α m n : Type*} +variables {R : Type v} [comm_ring R] {M : matrix m m R} {b : m → α} namespace matrix +section has_lt +variables [has_lt α] + +/-- Let `b` map rows and columns of a square matrix `M` to blocks indexed by `α`s. Then +`block_triangular M n b` says the matrix is block triangular. -/ +def block_triangular (M : matrix m m R) (b : m → α) : Prop := ∀ ⦃i j⦄, b j < b i → M i j = 0 + +@[simp] protected lemma block_triangular.submatrix {f : n → m} (h : M.block_triangular b) : + (M.submatrix f f).block_triangular (b ∘ f) := +λ i j hij, h hij + +lemma block_triangular_reindex_iff {b : n → α} {e : m ≃ n} : + (reindex e e M).block_triangular b ↔ M.block_triangular (b ∘ e) := +begin + refine ⟨λ h, _, λ h, _⟩, + { convert h.submatrix, + simp only [reindex_apply, submatrix_submatrix, submatrix_id_id, equiv.symm_comp_self] }, + { convert h.submatrix, + simp only [comp.assoc b e e.symm, equiv.self_comp_symm, comp.right_id] } +end + +protected lemma block_triangular.transpose : + M.block_triangular b → Mᵀ.block_triangular (to_dual ∘ b) := swap + +@[simp] protected lemma block_triangular_transpose_iff {b : m → αᵒᵈ} : + Mᵀ.block_triangular b ↔ M.block_triangular (of_dual ∘ b) := forall_swap + +end has_lt + +lemma upper_two_block_triangular [preorder α] + (A : matrix m m R) (B : matrix m n R) (D : matrix n n R) {a b : α} (hab : a < b) : + block_triangular (from_blocks A B 0 D) (sum.elim (λ i, a) (λ j, b)) := +begin + intros k1 k2 hk12, + have hor : ∀ (k : m ⊕ n), sum.elim (λ i, a) (λ j, b) k = a ∨ sum.elim (λ i, a) (λ j, b) k = b, + { simp }, + have hne : a ≠ b, from λ h, lt_irrefl _ (lt_of_lt_of_eq hab h.symm), + have ha : ∀ (k : m ⊕ n), sum.elim (λ i, a) (λ j, b) k = a → ∃ i, k = sum.inl i, + { simp [hne.symm] }, + have hb : ∀ (k : m ⊕ n), sum.elim (λ i, a) (λ j, b) k = b → ∃ j, k = sum.inr j, + { simp [hne] }, + cases (hor k1) with hk1 hk1; cases (hor k2) with hk2 hk2; rw [hk1, hk2] at hk12, + { exact false.elim (lt_irrefl a hk12), }, + { exact false.elim (lt_irrefl _ (lt_trans hab hk12)) }, + { obtain ⟨i, hi⟩ := hb k1 hk1, + obtain ⟨j, hj⟩ := ha k2 hk2, + rw [hi, hj], simp }, + { exact absurd hk12 (irrefl b) } +end + +/-! ### Determinant -/ + +variables [decidable_eq m] [fintype m] [decidable_eq n] [fintype n] + +lemma equiv_block_det (M : matrix m m R) {p q : m → Prop} [decidable_pred p] [decidable_pred q] + (e : ∀ x, q x ↔ p x) : (to_square_block_prop M p).det = (to_square_block_prop M q).det := +by convert matrix.det_reindex_self (equiv.subtype_equiv_right e) (to_square_block_prop M q) + +@[simp] lemma det_to_square_block_id (M : matrix m m R) (i : m) : + (M.to_square_block id i).det = M i i := +begin + letI : unique {a // id a = i} := ⟨⟨⟨i, rfl⟩⟩, λ j, subtype.ext j.property⟩, + exact (det_unique _).trans rfl, +end + lemma det_to_block (M : matrix m m R) (p : m → Prop) [decidable_pred p] : - M.det = (matrix.from_blocks (to_block M p p) (to_block M p (λ j, ¬p j)) - (to_block M (λ j, ¬p j) p) (to_block M (λ j, ¬p j) (λ j, ¬p j))).det := + M.det = (from_blocks (to_block M p p) (to_block M p $ λ j, ¬p j) + (to_block M (λ j, ¬p j) p) $ to_block M (λ j, ¬p j) $ λ j, ¬p j).det := begin - rw ← matrix.det_reindex_self (equiv.sum_compl p).symm M, + rw ←matrix.det_reindex_self (equiv.sum_compl p).symm M, rw [det_apply', det_apply'], congr, ext σ, congr, ext, generalize hy : σ x = y, @@ -54,16 +120,8 @@ begin matrix.submatrix_apply], end -lemma det_to_square_block (M : matrix m m R) {n : nat} (b : m → fin n) (k : fin n) : - (to_square_block M b k).det = (to_square_block_prop M (λ i, b i = k)).det := -by simp - -lemma det_to_square_block' (M : matrix m m R) (b : m → ℕ) (k : ℕ) : - (to_square_block' M b k).det = (to_square_block_prop M (λ i, b i = k)).det := -by simp - lemma two_block_triangular_det (M : matrix m m R) (p : m → Prop) [decidable_pred p] - (h : ∀ i (h1 : ¬p i) j (h2 : p j), M i j = 0) : + (h : ∀ i, ¬ p i → ∀ j, p j → M i j = 0) : M.det = (to_square_block_prop M p).det * (to_square_block_prop M (λ i, ¬p i)).det := begin rw det_to_block M p, @@ -73,172 +131,82 @@ begin exact h ↑i i.2 ↑j j.2 end -lemma equiv_block_det (M : matrix m m R) {p q : m → Prop} [decidable_pred p] [decidable_pred q] - (e : ∀x, q x ↔ p x) : (to_square_block_prop M p).det = (to_square_block_prop M q).det := -by convert matrix.det_reindex_self (equiv.subtype_equiv_right e) (to_square_block_prop M q) - -lemma to_square_block_det'' (M : matrix m m R) {n : nat} (b : m → fin n) (k : fin n) : - (to_square_block M b k).det = (to_square_block' M (λ i, ↑(b i)) ↑k).det := -begin - rw [to_square_block_def', to_square_block_def], - apply equiv_block_det, - intro x, - apply fin.ext_iff.symm -end - -/-- Let `b` map rows and columns of a square matrix `M` to `n` blocks. Then -`M.block_triangular' b` says the matrix is block triangular. -/ -def block_triangular' {o : Type*} (M : matrix o o R) {n : ℕ} - (b : o → fin n) : Prop := -∀ i j, b j < b i → M i j = 0 - -lemma upper_two_block_triangular' {m n : Type*} - (A : matrix m m R) (B : matrix m n R) (D : matrix n n R) : - (from_blocks A B 0 D).block_triangular' (sum.elim (λ i, (0 : fin 2)) (λ j, 1)) := -begin - intros k1 k2 hk12, - have h0 : ∀ (k : m ⊕ n), sum.elim (λ i, (0 : fin 2)) (λ j, 1) k = 0 → ∃ i, k = sum.inl i, - { simp }, - have h1 : ∀ (k : m ⊕ n), sum.elim (λ i, (0 : fin 2)) (λ j, 1) k = 1 → ∃ j, k = sum.inr j, - { simp }, - set mk1 := (sum.elim (λ i, (0 : fin 2)) (λ j, 1)) k1 with hmk1, - set mk2 := (sum.elim (λ i, (0 : fin 2)) (λ j, 1)) k2 with hmk2, - fin_cases mk1 using h; fin_cases mk2 using h_1; rw [h, h_1] at hk12, - { exact absurd hk12 (nat.not_lt_zero 0) }, - { exact absurd hk12 (by norm_num) }, - { rw hmk1 at h, - obtain ⟨i, hi⟩ := h1 k1 h, - rw hmk2 at h_1, - obtain ⟨j, hj⟩ := h0 k2 h_1, - rw [hi, hj], simp }, - { exact absurd hk12 (irrefl 1) } -end - -/-- Let `b` map rows and columns of a square matrix `M` to blocks indexed by `ℕ`s. Then - `M.block_triangular b` says the matrix is block triangular. -/ -def block_triangular {o : Type*} (M : matrix o o R) (b : o → ℕ) : Prop := -∀ i j, b j < b i → M i j = 0 - -lemma upper_two_block_triangular {m n : Type*} - (A : matrix m m R) (B : matrix m n R) (D : matrix n n R) : - (from_blocks A B 0 D).block_triangular (sum.elim (λ i, 0) (λ j, 1)) := -begin - intros k1 k2 hk12, - have h01 : ∀ (k : m ⊕ n), sum.elim (λ i, 0) (λ j, 1) k = 0 ∨ sum.elim (λ i, 0) (λ j, 1) k = 1, - { simp }, - have h0 : ∀ (k : m ⊕ n), sum.elim (λ i, 0) (λ j, 1) k = 0 → ∃ i, k = sum.inl i, { simp }, - have h1 : ∀ (k : m ⊕ n), sum.elim (λ i, 0) (λ j, 1) k = 1 → ∃ j, k = sum.inr j, { simp }, - cases (h01 k1) with hk1 hk1; cases (h01 k2) with hk2 hk2; rw [hk1, hk2] at hk12, - { exact absurd hk12 (nat.not_lt_zero 0) }, - { exact absurd hk12 (nat.not_lt_zero 1) }, - { obtain ⟨i, hi⟩ := h1 k1 hk1, - obtain ⟨j, hj⟩ := h0 k2 hk2, - rw [hi, hj], simp }, - { exact absurd hk12 (irrefl 1) } -end - -lemma det_of_block_triangular (M : matrix m m R) (b : m → ℕ) - (h : M.block_triangular b) : - ∀ (n : ℕ) (hn : ∀ i, b i < n), M.det = ∏ k in finset.range n, (to_square_block' M b k).det := +lemma two_block_triangular_det' (M : matrix m m R) (p : m → Prop) [decidable_pred p] + (h : ∀ i, p i → ∀ j, ¬ p j → M i j = 0) : + M.det = (to_square_block_prop M p).det * (to_square_block_prop M (λ i, ¬p i)).det := begin - intros n hn, - unfreezingI { induction n with n hi generalizing m M b }, - { rw finset.prod_range_zero, - apply det_eq_one_of_card_eq_zero, - apply fintype.card_eq_zero_iff.mpr, - exact ⟨λ i, nat.not_lt_zero (b i) (hn i)⟩ }, - { rw [finset.prod_range_succ_comm], - have h2 : (M.to_square_block_prop (λ (i : m), b i = n.succ)).det = - (M.to_square_block' b n.succ).det, - { dunfold to_square_block', dunfold to_square_block_prop, refl }, - rw two_block_triangular_det M (λ i, ¬(b i = n)), - { rw mul_comm, - apply congr (congr_arg has_mul.mul _), - { let m' := {a // ¬b a = n }, - let b' := (λ (i : m'), b ↑i), - have h' : - (M.to_square_block_prop (λ (i : m), ¬b i = n)).block_triangular b', - { intros i j, apply h ↑i ↑j }, - have hni : ∀ (i : {a // ¬b a = n}), b' i < n, - { exact λ i, (ne.le_iff_lt i.property).mp (nat.lt_succ_iff.mp (hn ↑i)) }, - have h1 := hi (M.to_square_block_prop (λ (i : m), ¬b i = n)) b' h' hni, - rw ←fin.prod_univ_eq_prod_range at h1 ⊢, - convert h1, - ext k, - simp only [to_square_block_def', to_square_block_def], - let he : {a // b' a = ↑k} ≃ {a // b a = ↑k}, - { have hc : ∀ (i : m), (λ a, b a = ↑k) i → (λ a, ¬b a = n) i, - { intros i hbi, rw hbi, exact ne_of_lt (fin.is_lt k) }, - exact equiv.subtype_subtype_equiv_subtype hc }, - exact matrix.det_reindex_self he (λ (i j : {a // b' a = ↑k}), M ↑i ↑j) }, - { rw det_to_square_block' M b n, - have hh : ∀ a, b a = n ↔ ¬(λ (i : m), ¬b i = n) a, - { intro i, simp only [not_not] }, - exact equiv_block_det M hh }}, - { intros i hi j hj, - apply (h i), simp only [not_not] at hi, - rw hi, - exact (ne.le_iff_lt hj).mp (nat.lt_succ_iff.mp (hn j)) }} + rw [M.two_block_triangular_det (λ i, ¬ p i), mul_comm], + simp_rw not_not, + congr' 1, + exact equiv_block_det _ (λ _, not_not.symm), + simpa only [not_not] using h, end -lemma det_of_block_triangular'' (M : matrix m m R) (b : m → ℕ) - (h : M.block_triangular b) : - M.det = ∏ k in finset.image b finset.univ, (to_square_block' M b k).det := +protected lemma block_triangular.det [decidable_eq α] [linear_order α] (hM : block_triangular M b) : + M.det = ∏ a in univ.image b, (M.to_square_block b a).det := begin - let n : ℕ := (Sup (finset.image b finset.univ : set ℕ)).succ, - have hn : ∀ i, b i < n, - { have hbi : ∀ i, b i ∈ finset.image b finset.univ, { simp }, - intro i, - dsimp only [n], - apply nat.lt_succ_iff.mpr, - exact le_cSup (finset.bdd_above _) (hbi i) }, - rw det_of_block_triangular M b h n hn, - refine (finset.prod_subset _ _).symm, - { intros a ha, apply finset.mem_range.mpr, - obtain ⟨i, ⟨hi, hbi⟩⟩ := finset.mem_image.mp ha, - rw ←hbi, - exact hn i }, - { intros k hk hbk, - apply det_eq_one_of_card_eq_zero, - apply fintype.card_eq_zero_iff.mpr, - constructor, - simp only [subtype.forall], - intros a hba, apply hbk, - apply finset.mem_image.mpr, - use a, - exact ⟨finset.mem_univ a, hba⟩ } + unfreezingI { induction hs : univ.image b using finset.strong_induction + with s ih generalizing m }, + subst hs, + by_cases h : univ.image b = ∅, + { haveI := univ_eq_empty_iff.1 (image_eq_empty.1 h), + simp [h] }, + { let k := (univ.image b).max' (nonempty_of_ne_empty h), + rw two_block_triangular_det' M (λ i, b i = k), + { have : univ.image b = insert k ((univ.image b).erase k), + { rw insert_erase, apply max'_mem }, + rw [this, prod_insert (not_mem_erase _ _)], + refine congr_arg _ _, + let b' := λ i : {a // b a ≠ k}, b ↑i, + have h' : block_triangular (M.to_square_block_prop (λ (i : m), b i ≠ k)) b', + { intros i j, apply hM }, + have hb' : image b' univ = (image b univ).erase k, + { apply subset_antisymm, + { rw image_subset_iff, + intros i _, + apply mem_erase_of_ne_of_mem i.2 (mem_image_of_mem _ (mem_univ _)) }, + { intros i hi, + rw mem_image, + rcases mem_image.1 (erase_subset _ _ hi) with ⟨a, _, ha⟩, + subst ha, + exact ⟨⟨a, ne_of_mem_erase hi⟩, mem_univ _, rfl⟩ } }, + rw ih ((univ.image b).erase k) (erase_ssubset (max'_mem _ _)) h' hb', + apply finset.prod_congr rfl, + intros l hl, + let he : {a // b' a = l} ≃ {a // b a = l}, + { have hc : ∀ (i : m), (λ a, b a = l) i → (λ a, b a ≠ k) i, + { intros i hbi, rw hbi, exact ne_of_mem_erase hl }, + exact equiv.subtype_subtype_equiv_subtype hc }, + simp only [to_square_block_def], + rw ← matrix.det_reindex_self he.symm (λ (i j : {a // b a = l}), M ↑i ↑j), + refine congr_arg _ _, + ext, + simp [to_square_block_def, to_square_block_prop_def] }, + { intros i hi j hj, + apply hM, + rw hi, + apply lt_of_le_of_ne _ hj, + exact finset.le_max' (univ.image b) _ (mem_image_of_mem _ (mem_univ _)) } } end -lemma det_of_block_triangular' (M : matrix m m R) {n : ℕ} (b : m → fin n) - (h : M.block_triangular' b) : - M.det = ∏ (k : fin n), (to_square_block M b k).det := +lemma block_triangular.det_fintype [decidable_eq α] [fintype α] [linear_order α] + (h : block_triangular M b) : + M.det = ∏ k : α, (M.to_square_block b k).det := begin - let b2 : m → ℕ := λ i, ↑(b i), - simp_rw to_square_block_det'', - rw fin.prod_univ_eq_prod_range (λ (k : ℕ), (M.to_square_block' b2 k).det) n, - apply det_of_block_triangular, - { intros i j hij, exact h i j (fin.coe_fin_lt.mp hij) }, - { intro i, exact fin.is_lt (b i) } + refine h.det.trans (prod_subset (subset_univ _) $ λ a _ ha, _), + have : is_empty {i // b i = a} := ⟨λ i, ha $ mem_image.2 ⟨i, mem_univ _, i.2⟩⟩, + exactI det_is_empty, end -lemma det_of_upper_triangular {n : ℕ} (M : matrix (fin n) (fin n) R) - (h : ∀ (i j : fin n), j < i → M i j = 0) : - M.det = ∏ i : (fin n), M i i := +lemma det_of_upper_triangular [linear_order m] (h : M.block_triangular id) : + M.det = ∏ i : m, M i i := begin - convert det_of_block_triangular' M id h, - ext i, - have h2 : ∀ (j : {a // id a = i}), j = ⟨i, rfl⟩ := - λ (j : {a // id a = i}), subtype.ext j.property, - haveI : unique {a // id a = i} := ⟨⟨⟨i, rfl⟩⟩, h2⟩, - simp [h2 default] + haveI : decidable_eq R := classical.dec_eq _, + simp_rw [h.det, image_id, det_to_square_block_id], end -lemma det_of_lower_triangular {n : ℕ} (M : matrix (fin n) (fin n) R) - (h : ∀ (i j : fin n), i < j → M i j = 0) : - M.det = ∏ i : (fin n), M i i := -begin - rw ← det_transpose, - exact det_of_upper_triangular _ (λ (i j : fin n) (hji : j < i), h j i hji) -end +lemma det_of_lower_triangular [linear_order m] (M : matrix m m R) (h : M.block_triangular to_dual) : + M.det = ∏ i : m, M i i := +by { rw ←det_transpose, exact det_of_upper_triangular h.transpose } end matrix From aab6d3c7a7b81e4da03928c3396a018791f402a3 Mon Sep 17 00:00:00 2001 From: Stuart Presnell Date: Fri, 2 Sep 2022 13:28:04 +0000 Subject: [PATCH 139/828] feat(data/nat/factorization/prime_pow): lemmas relating `is_prime_pow` and `ord_compl = 1` (#15692) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Adds lemmas showing that `n` is a prime power iff `ord_compl[p] n = 1` (for some prime `p`). `exists_ord_compl_eq_one_of_is_prime_pow (h : is_prime_pow n) : (∃ p : ℕ, p.prime ∧ ord_compl[p] n = 1)` `is_prime_pow_of_ord_compl_eq_one (hn : n ≠ 1) (pp : p.prime) (h : ord_compl[p] n = 1) : is_prime_pow n` --- src/data/nat/factorization/prime_pow.lean | 31 +++++++++++++++++++---- 1 file changed, 26 insertions(+), 5 deletions(-) diff --git a/src/data/nat/factorization/prime_pow.lean b/src/data/nat/factorization/prime_pow.lean index ba7ee36067bb9..5428dc4b658c9 100644 --- a/src/data/nat/factorization/prime_pow.lean +++ b/src/data/nat/factorization/prime_pow.lean @@ -60,6 +60,31 @@ lemma is_prime_pow_iff_card_support_factorization_eq_one {n : ℕ} : by simp_rw [is_prime_pow_iff_factorization_eq_single, finsupp.card_support_eq_one', exists_prop, pos_iff_ne_zero] +lemma is_prime_pow.exists_ord_compl_eq_one {n : ℕ} (h : is_prime_pow n) : + ∃ p : ℕ, p.prime ∧ ord_compl[p] n = 1 := +begin + rcases eq_or_ne n 0 with rfl | hn0, { cases not_is_prime_pow_zero h }, + rcases is_prime_pow_iff_factorization_eq_single.mp h with ⟨p, k, hk0, h1⟩, + rcases em' p.prime with pp | pp, + { refine absurd _ hk0.ne', simp [←nat.factorization_eq_zero_of_non_prime n pp, h1] }, + refine ⟨p, pp, _⟩, + refine nat.eq_of_factorization_eq (nat.ord_compl_pos p hn0).ne' (by simp) (λ q, _), + rw [nat.factorization_ord_compl n p, h1], + simp, +end + +lemma exists_ord_compl_eq_one_iff_is_prime_pow {n : ℕ} (hn : n ≠ 1) : + is_prime_pow n ↔ ∃ p : ℕ, p.prime ∧ ord_compl[p] n = 1 := +begin + refine ⟨λ h, is_prime_pow.exists_ord_compl_eq_one h, λ h, _⟩, + rcases h with ⟨p, pp, h⟩, + rw is_prime_pow_nat_iff, + rw [←nat.eq_of_dvd_of_div_eq_one (nat.ord_proj_dvd n p) h] at ⊢ hn, + refine ⟨p, n.factorization p, pp, _, by simp⟩, + contrapose! hn, + simp [le_zero_iff.1 hn], +end + /-- An equivalent definition for prime powers: `n` is a prime power iff there is a unique prime dividing it. -/ lemma is_prime_pow_iff_unique_prime_dvd {n : ℕ} : @@ -72,12 +97,8 @@ begin rintro q ⟨hq, hq'⟩, exact (nat.prime_dvd_prime_iff_eq hq hp).1 (hq.dvd_of_dvd_pow hq') }, rintro ⟨p, ⟨hp, hn⟩, hq⟩, - -- Take care of the n = 0 case rcases eq_or_ne n 0 with rfl | hn₀, - { obtain ⟨q, hq', hq''⟩ := nat.exists_infinite_primes (p + 1), - cases hq q ⟨hq'', by simp⟩, - simpa using hq' }, - -- So assume 0 < n + { cases (hq 2 ⟨nat.prime_two, dvd_zero 2⟩).trans (hq 3 ⟨nat.prime_three, dvd_zero 3⟩).symm }, refine ⟨p, n.factorization p, hp, hp.factorization_pos_of_dvd hn₀ hn, _⟩, simp only [and_imp] at hq, apply nat.dvd_antisymm (nat.ord_proj_dvd _ _), From 678f62f0b6c61a1402b193d2977f5052e6e4b425 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Fri, 2 Sep 2022 15:15:22 +0000 Subject: [PATCH 140/828] feat(data/polynomial): some lemmas about (nat)degree of bit0/1 in char_zero (#15726) --- src/data/polynomial/coeff.lean | 2 ++ src/data/polynomial/ring_division.lean | 30 ++++++++++++++++++++++++++ 2 files changed, 32 insertions(+) diff --git a/src/data/polynomial/coeff.lean b/src/data/polynomial/coeff.lean index 8a11187bef5e0..5e53ed55b5743 100644 --- a/src/data/polynomial/coeff.lean +++ b/src/data/polynomial/coeff.lean @@ -36,6 +36,8 @@ coeff_monomial lemma coeff_add (p q : R[X]) (n : ℕ) : coeff (p + q) n = coeff p n + coeff q n := by { rcases p, rcases q, simp_rw [←of_finsupp_add, coeff], exact finsupp.add_apply _ _ _ } +@[simp] lemma coeff_bit0 (p : R[X]) (n : ℕ) : coeff (bit0 p) n = bit0 (coeff p n) := by simp [bit0] + @[simp] lemma coeff_smul [monoid S] [distrib_mul_action S R] (r : S) (p : R[X]) (n : ℕ) : coeff (r • p) n = r • coeff p n := by { rcases p, simp_rw [←of_finsupp_smul, coeff], exact finsupp.smul_apply _ _ _ } diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index f2543c598ae9d..11dd463a3012b 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -160,6 +160,36 @@ begin rw [← nat_degree_mul hp₁ hq₂, ← nat_degree_mul hp₂ hq₁, h_eq] end +variables [char_zero R] + +@[simp] lemma degree_bit0_eq (p : R[X]) : degree (bit0 p) = degree p := +by rw [bit0_eq_two_mul, degree_mul, (by simp : (2 : polynomial R) = C 2), + @polynomial.degree_C R _ _ two_ne_zero', zero_add] + +@[simp] lemma nat_degree_bit0_eq (p : R[X]) : nat_degree (bit0 p) = nat_degree p := +nat_degree_eq_of_degree_eq $ degree_bit0_eq p + +@[simp] +lemma nat_degree_bit1_eq (p : R[X]) : nat_degree (bit1 p) = nat_degree p := +begin + rw bit1, + apply le_antisymm, + convert nat_degree_add_le _ _, + { simp, }, + by_cases h : p.nat_degree = 0, + { simp [h], }, + apply le_nat_degree_of_ne_zero, + intro hh, + apply h, + simp [*, coeff_one, if_neg (ne.symm h)] at *, +end + +lemma degree_bit1_eq {p : R[X]} (hp : 0 < degree p) : degree (bit1 p) = degree p := +begin + rw [bit1, degree_add_eq_left_of_degree_lt, degree_bit0_eq], + rwa [degree_one, degree_bit0_eq] +end + end no_zero_divisors section no_zero_divisors From 734c95f131c54174ec52b271b6845899c28bab91 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Fri, 2 Sep 2022 15:15:23 +0000 Subject: [PATCH 141/828] feat(combinatorics/simple_graph/connectivity): walk.copy function to relax dependent type issues (#15764) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Walks in a graph present a special challenge since the endpoints of a walk appear as type indices, leading to a constant struggle with dependent types and strict definitional equality. This commit introduces a `walk.copy` function to be able to change the definitions of a walk's endpoints. There are also many simp lemmas to push the copy function out of most expressions. An example of a lemma that `copy` lets us express: ```lean lemma map_eq_of_eq {f : G →g G'} (f' : G →g G') (h : f = f') : p.map f = (p.map f').copy (by rw h) (by rw h) ``` Thanks to @**Junyan Xu|224323** for the idea. --- .../simple_graph/connectivity.lean | 105 +++++++++++++++++- 1 file changed, 101 insertions(+), 4 deletions(-) diff --git a/src/combinatorics/simple_graph/connectivity.lean b/src/combinatorics/simple_graph/connectivity.lean index 874c368be0e18..4b5e0aec79048 100644 --- a/src/combinatorics/simple_graph/connectivity.lean +++ b/src/combinatorics/simple_graph/connectivity.lean @@ -62,10 +62,11 @@ walks, trails, paths, circuits, cycles open function -universes u v +universes u v w namespace simple_graph -variables {V : Type u} {V' : Type v} (G : simple_graph V) (G' : simple_graph V') +variables {V : Type u} {V' : Type v} {V'' : Type w} +variables (G : simple_graph V) (G' : simple_graph V') (G'' : simple_graph V'') /-- A walk is a sequence of adjacent vertices. For vertices `u v : V`, the type `walk u v` consists of all walks starting at `u` and ending at `v`. @@ -94,6 +95,35 @@ variables {G} @[pattern] abbreviation cons' (u v w : V) (h : G.adj u v) (p : G.walk v w) : G.walk u w := walk.cons h p +/-- Change the endpoints of a walk using equalities. This is helpful for relaxing +definitional equality constraints and to be able to state otherwise difficult-to-state +lemmas. While this is a simple wrapper around `eq.rec`, it gives a canonical way to write it. + +The simp-normal form is for the `copy` to be pushed outward. That way calculations can +occur within the "copy context." -/ +protected def copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : G.walk u' v' := +eq.rec (eq.rec p hv) hu + +@[simp] lemma copy_rfl_rfl {u v} (p : G.walk u v) : + p.copy rfl rfl = p := rfl + +@[simp] lemma copy_copy {u v u' v' u'' v''} (p : G.walk u v) + (hu : u = u') (hv : v = v') (hu' : u' = u'') (hv' : v' = v'') : + (p.copy hu hv).copy hu' hv' = p.copy (hu.trans hu') (hv.trans hv') := +by { subst_vars, refl } + +@[simp] lemma copy_nil {u u'} (hu : u = u') : (walk.nil : G.walk u u).copy hu hu = walk.nil := +by { subst_vars, refl } + +lemma copy_cons {u v w u' w'} (h : G.adj u v) (p : G.walk v w) (hu : u = u') (hw : w = w') : + (walk.cons h p).copy hu hw = walk.cons (by rwa ← hu) (p.copy rfl hw) := +by { subst_vars, refl } + +@[simp] +lemma cons_copy {u v w v' w'} (h : G.adj u v) (p : G.walk v' w') (hv : v' = v) (hw : w' = w) : + walk.cons h (p.copy hv hw) = (walk.cons (by rwa hv) p).copy rfl hw := +by { subst_vars, refl } + lemma exists_eq_cons_of_ne : Π {u v : V} (hne : u ≠ v) (p : G.walk u v), ∃ (w : V) (h : G.adj u w) (p' : G.walk w v), p = cons h p' | _ _ hne nil := (hne rfl).elim @@ -170,6 +200,10 @@ lemma append_assoc : Π {u v w x : V} (p : G.walk u v) (q : G.walk v w) (r : G.w | _ _ _ _ nil _ _ := rfl | _ _ _ _ (cons h p') q r := by { dunfold append, rw append_assoc, } +@[simp] lemma append_copy_copy {u v w u' v' w'} (p : G.walk u v) (q : G.walk v w) + (hu : u = u') (hv : v = v') (hw : w = w') : + (p.copy hu hv).append (q.copy hv hw) = (p.append q).copy hu hw := by { subst_vars, refl } + @[simp] lemma reverse_nil {u : V} : (nil : G.walk u u).reverse = nil := rfl lemma reverse_singleton {u v : V} (h : G.adj u v) : @@ -198,6 +232,9 @@ by simp [reverse] (cons h p).reverse = p.reverse.append (cons (G.symm h) nil) := by simp [reverse] +@[simp] lemma reverse_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : + (p.copy hu hv).reverse = p.reverse.copy hv hu := by { subst_vars, refl } + @[simp] lemma reverse_append {u v w : V} (p : G.walk u v) (q : G.walk v w) : (p.append q).reverse = q.reverse.append p.reverse := by simp [reverse] @@ -211,6 +248,10 @@ by simp [reverse] @[simp] lemma length_cons {u v w : V} (h : G.adj u v) (p : G.walk v w) : (cons h p).length = p.length + 1 := rfl +@[simp] lemma length_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : + (p.copy hu hv).length = p.length := +by { subst_vars, refl } + @[simp] lemma length_append : Π {u v w : V} (p : G.walk u v) (q : G.walk v w), (p.append q).length = p.length + q.length | _ _ _ nil _ := by simp @@ -258,6 +299,9 @@ def edges {u v : V} (p : G.walk u v) : list (sym2 V) := p.darts.map dart.edge @[simp] lemma support_cons {u v w : V} (h : G.adj u v) (p : G.walk v w) : (cons h p).support = u :: p.support := rfl +@[simp] lemma support_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : + (p.copy hu hv).support = p.support := by { subst_vars, refl } + lemma support_append {u v w : V} (p : G.walk u v) (p' : G.walk v w) : (p.append p').support = p.support ++ p'.support.tail := by induction p; cases p'; simp [*] @@ -366,6 +410,9 @@ lemma edges_subset_edge_set : Π {u v : V} (p : G.walk u v) ⦃e : sym2 V⦄ @[simp] lemma darts_cons {u v w : V} (h : G.adj u v) (p : G.walk v w) : (cons h p).darts = ⟨(u, v), h⟩ :: p.darts := rfl +@[simp] lemma darts_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : + (p.copy hu hv).darts = p.darts := by { subst_vars, refl } + @[simp] lemma darts_append {u v w : V} (p : G.walk u v) (p' : G.walk v w) : (p.append p').darts = p.darts ++ p'.darts := by induction p; simp [*] @@ -399,6 +446,9 @@ by simpa! using congr_arg list.init (map_fst_darts_append p) @[simp] lemma edges_cons {u v w : V} (h : G.adj u v) (p : G.walk v w) : (cons h p).edges = ⟦(u, v)⟧ :: p.edges := rfl +@[simp] lemma edges_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : + (p.copy hu hv).edges = p.edges := by { subst_vars, refl } + @[simp] lemma edges_append {u v w : V} (p : G.walk u v) (p' : G.walk v w) : (p.append p').edges = p.edges ++ p'.edges := by simp [edges] @@ -484,16 +534,32 @@ structure is_cycle {u : V} (p : G.walk u u) lemma is_trail_def {u v : V} (p : G.walk u v) : p.is_trail ↔ p.edges.nodup := ⟨is_trail.edges_nodup, λ h, ⟨h⟩⟩ +@[simp] lemma is_trail_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : + (p.copy hu hv).is_trail ↔ p.is_trail := by { subst_vars, refl } + lemma is_path.mk' {u v : V} {p : G.walk u v} (h : p.support.nodup) : is_path p := ⟨⟨edges_nodup_of_support_nodup h⟩, h⟩ lemma is_path_def {u v : V} (p : G.walk u v) : p.is_path ↔ p.support.nodup := ⟨is_path.support_nodup, is_path.mk'⟩ +@[simp] lemma is_path_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : + (p.copy hu hv).is_path ↔ p.is_path := by { subst_vars, refl } + +lemma is_circuit_def {u : V} (p : G.walk u u) : + p.is_circuit ↔ is_trail p ∧ p ≠ nil := +iff.intro (λ h, ⟨h.1, h.2⟩) (λ h, ⟨h.1, h.2⟩) + +@[simp] lemma is_circuit_copy {u u'} (p : G.walk u u) (hu : u = u') : + (p.copy hu hu).is_circuit ↔ p.is_circuit := by { subst_vars, refl } + lemma is_cycle_def {u : V} (p : G.walk u u) : p.is_cycle ↔ is_trail p ∧ p ≠ nil ∧ p.support.tail.nodup := iff.intro (λ h, ⟨h.1.1, h.1.2, h.2⟩) (λ h, ⟨⟨h.1, h.2.1⟩, h.2.2⟩) +@[simp] lemma is_cycle_copy {u u'} (p : G.walk u u) (hu : u = u') : + (p.copy hu hu).is_cycle ↔ p.is_cycle := by { subst_vars, refl } + @[simp] lemma is_trail.nil {u : V} : (nil : G.walk u u).is_trail := ⟨by simp [edges]⟩ @@ -646,6 +712,16 @@ begin { apply ih, } } } }, end +@[simp] lemma take_until_copy {u v w v' w'} (p : G.walk v w) + (hv : v = v') (hw : w = w') (h : u ∈ (p.copy hv hw).support) : + (p.copy hv hw).take_until u h = (p.take_until u (by { subst_vars, exact h })).copy hv rfl := +by { subst_vars, refl } + +@[simp] lemma drop_until_copy {u v w v' w'} (p : G.walk v w) + (hv : v = v') (hw : w = w') (h : u ∈ (p.copy hv hw).support) : + (p.copy hv hw).drop_until u h = (p.drop_until u (by { subst_vars, exact h })).copy rfl hw := +by { subst_vars, refl } + lemma support_take_until_subset {u v w : V} (p : G.walk v w) (h : u ∈ p.support) : (p.take_until u h).support ⊆ p.support := λ x hx, by { rw [← take_spec p h, mem_support_append_iff], exact or.inl hx } @@ -838,6 +914,9 @@ def bypass : Π {u v : V}, G.walk u v → G.walk u v then p'.drop_until u hs else cons ha p' +@[simp] lemma bypass_copy {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') : + (p.copy hu hv).bypass = p.bypass.copy hu hv := by { subst_vars, refl } + lemma bypass_is_path {u v : V} (p : G.walk u v) : p.bypass.is_path := begin induction p, @@ -912,20 +991,32 @@ end walk /-! ### Mapping paths -/ namespace walk -variables {G G'} +variables {G G' G''} /-- Given a graph homomorphism, map walks to walks. -/ protected def map (f : G →g G') : Π {u v : V}, G.walk u v → G'.walk (f u) (f v) | _ _ nil := nil | _ _ (cons h p) := cons (f.map_adj h) (map p) -variables (f : G →g G') {u v : V} (p : G.walk u v) +variables (f : G →g G') (f' : G' →g G'') {u v u' v' : V} (p : G.walk u v) @[simp] lemma map_nil : (nil : G.walk u u).map f = nil := rfl @[simp] lemma map_cons {w : V} (h : G.adj w u) : (cons h p).map f = cons (f.map_adj h) (p.map f) := rfl +@[simp] lemma map_copy (hu : u = u') (hv : v = v') : + (p.copy hu hv).map f = (p.map f).copy (by rw hu) (by rw hv) := by { subst_vars, refl } + +@[simp] lemma map_id (p : G.walk u v) : p.map hom.id = p := by { induction p; simp [*] } + +@[simp] lemma map_map : (p.map f).map f' = p.map (f'.comp f) := by { induction p; simp [*] } + +/-- Unlike categories, for graphs vertex equality is an important notion, so needing to be able to +to work with equality of graph homomorphisms is a necessary evil. -/ +lemma map_eq_of_eq {f : G →g G'} (f' : G →g G') (h : f = f') : + p.map f = (p.map f').copy (by rw h) (by rw h) := by { subst_vars, refl } + @[simp] lemma length_map : (p.map f).length = p.length := by induction p; simp [*] @@ -1053,6 +1144,12 @@ lemma is_path.to_delete_edges (s : set (sym2 V)) (p.to_delete_edges s hp).is_path := by { rw ← map_to_delete_edges_eq s hp at h, exact h.of_map } +@[simp] lemma to_delete_edges_copy (s : set (sym2 V)) + {u v u' v'} (p : G.walk u v) (hu : u = u') (hv : v = v') (h) : + (p.copy hu hv).to_delete_edges s h + = (p.to_delete_edges s (by { subst_vars, exact h })).copy hu hv := +by { subst_vars, refl } + end walk /-! ## `reachable` and `connected` -/ From 9191c40c89d6174750729c8e9b661ea326f8f815 Mon Sep 17 00:00:00 2001 From: Stuart Presnell Date: Fri, 2 Sep 2022 15:15:24 +0000 Subject: [PATCH 142/828] feat(data/nat/factorization/basic): lemmas on divisibility of `ord_proj` and `ord_compl` (#15793) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For `a b : ℕ`, we have `a ∣ b` iff * `∀ p : ℕ, ord_proj[p] a ∣ ord_proj[p] b` (for `a`, `b` positive) * `∀ p : ℕ, ord_compl[p] a ∣ ord_compl[p] b` --- src/data/nat/factorization/basic.lean | 49 +++++++++++++++++++++++++++ 1 file changed, 49 insertions(+) diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 2ba7d488a9ef9..8913a68033644 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -8,6 +8,7 @@ import data.finsupp.multiset import data.nat.prime import number_theory.padics.padic_val import data.nat.interval +import tactic.interval_cases /-! # Prime factorizations @@ -495,6 +496,54 @@ begin lt_self_iff_false] at hp end +lemma ord_proj_dvd_ord_proj_of_dvd {a b : ℕ} (hb0 : b ≠ 0) (hab : a ∣ b) (p : ℕ) : + ord_proj[p] a ∣ ord_proj[p] b := +begin + rcases em' p.prime with pp | pp, { simp [pp] }, + rcases eq_or_ne a 0 with rfl | ha0, { simp }, + rw pow_dvd_pow_iff_le_right pp.one_lt, + exact (factorization_le_iff_dvd ha0 hb0).2 hab p, +end + +lemma ord_proj_dvd_ord_proj_iff_dvd {a b : ℕ} (ha0 : a ≠ 0) (hb0 : b ≠ 0) : + (∀ p : ℕ, ord_proj[p] a ∣ ord_proj[p] b) ↔ (a ∣ b) := +begin + refine ⟨λ h, _, λ hab p, ord_proj_dvd_ord_proj_of_dvd hb0 hab p⟩, + rw ←factorization_le_iff_dvd ha0 hb0, + intro q, + rcases le_or_lt q 1 with hq_le | hq1, { interval_cases q; simp }, + exact (pow_dvd_pow_iff_le_right hq1).1 (h q), +end + +lemma ord_compl_dvd_ord_compl_of_dvd {a b : ℕ} (hab : a ∣ b) (p : ℕ) : + ord_compl[p] a ∣ ord_compl[p] b := +begin + rcases em' p.prime with pp | pp, { simp [pp, hab] }, + rcases eq_or_ne b 0 with rfl | hb0, { simp }, + rcases eq_or_ne a 0 with rfl | ha0, { cases hb0 (zero_dvd_iff.1 hab) }, + have ha := (nat.div_pos (ord_proj_le p ha0) (ord_proj_pos a p)).ne', + have hb := (nat.div_pos (ord_proj_le p hb0) (ord_proj_pos b p)).ne', + rw [←factorization_le_iff_dvd ha hb, factorization_ord_compl a p, factorization_ord_compl b p], + intro q, + rcases eq_or_ne q p with rfl | hqp, { simp }, + simp_rw erase_ne hqp, + exact (factorization_le_iff_dvd ha0 hb0).2 hab q, +end + +lemma ord_compl_dvd_ord_compl_iff_dvd (a b : ℕ) : + (∀ p : ℕ, ord_compl[p] a ∣ ord_compl[p] b) ↔ (a ∣ b) := +begin + refine ⟨λ h, _, λ hab p, ord_compl_dvd_ord_compl_of_dvd hab p⟩, + rcases eq_or_ne b 0 with rfl | hb0, { simp }, + by_cases pa : a.prime, swap, { simpa [pa] using h a }, + by_cases pb : b.prime, swap, { simpa [pb] using h b }, + rw prime_dvd_prime_iff_eq pa pb, + by_contradiction hab, + apply pa.ne_one, + rw [←nat.dvd_one, ←nat.mul_dvd_mul_iff_left hb0.bot_lt, mul_one], + simpa [prime.factorization_self pb, prime.factorization pa, hab] using h b, +end + lemma dvd_iff_prime_pow_dvd_dvd (n d : ℕ) : d ∣ n ↔ ∀ p k : ℕ, prime p → p ^ k ∣ d → p ^ k ∣ n := begin From 637412dfbc6987e2d4710e7dfb5b81b8731ac7c8 Mon Sep 17 00:00:00 2001 From: Vincent Beffara Date: Fri, 2 Sep 2022 15:15:25 +0000 Subject: [PATCH 143/828] feat(analysis/complex/schwarz): equality case in Schwarz lemma (#15861) This partially addresses one of the TODO items in `analysis/complex/schwarz` by showing the equality case in the main statement of the Schwarz lemma. --- src/analysis/complex/schwarz.lean | 33 ++++++++++++++++++++++++++++++- 1 file changed, 32 insertions(+), 1 deletion(-) diff --git a/src/analysis/complex/schwarz.lean b/src/analysis/complex/schwarz.lean index 36e97b00f016a..83e5125bb8ded 100644 --- a/src/analysis/complex/schwarz.lean +++ b/src/analysis/complex/schwarz.lean @@ -54,7 +54,7 @@ namespace complex section space variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {R R₁ R₂ : ℝ} {f : ℂ → E} - {c z : ℂ} + {c z z₀ : ℂ} /-- An auxiliary lemma for `complex.norm_dslope_le_div_of_maps_to_ball`. -/ lemma schwarz_aux {f : ℂ → ℂ} (hd : differentiable_on ℂ f (ball c R₁)) @@ -109,6 +109,37 @@ begin end end +/-- Equality case in the **Schwarz Lemma**: in the setup of `norm_dslope_le_div_of_maps_to_ball`, if +`∥dslope f c z₀∥ = R₂ / R₁` holds at a point in the ball then the map `f` is affine. -/ +lemma affine_of_maps_to_ball_of_exists_norm_dslope_eq_div [complete_space E] + [strict_convex_space ℝ E] (hd : differentiable_on ℂ f (ball c R₁)) + (h_maps : set.maps_to f (ball c R₁) (ball (f c) R₂)) (h_z₀ : z₀ ∈ ball c R₁) + (h_eq : ∥dslope f c z₀∥ = R₂ / R₁) : + set.eq_on f (λ z, f c + (z - c) • dslope f c z₀) (ball c R₁) := +begin + set g := dslope f c, + rintro z hz, + by_cases z = c, { simp [h] }, + have h_R₁ : 0 < R₁ := nonempty_ball.mp ⟨_, h_z₀⟩, + have g_le_div : ∀ z ∈ ball c R₁, ∥g z∥ ≤ R₂ / R₁, + from λ z hz, norm_dslope_le_div_of_maps_to_ball hd h_maps hz, + have g_max : is_max_on (norm ∘ g) (ball c R₁) z₀, + from is_max_on_iff.mpr (λ z hz, by simpa [h_eq] using g_le_div z hz), + have g_diff : differentiable_on ℂ g (ball c R₁), + from (differentiable_on_dslope (is_open_ball.mem_nhds (mem_ball_self h_R₁))).mpr hd, + have : g z = g z₀ := eq_on_of_is_preconnected_of_is_max_on_norm (convex_ball c R₁).is_preconnected + is_open_ball g_diff h_z₀ g_max hz, + simp [← this] +end + +lemma affine_of_maps_to_ball_of_exists_norm_dslope_eq_div' [complete_space E] + [strict_convex_space ℝ E] (hd : differentiable_on ℂ f (ball c R₁)) + (h_maps : set.maps_to f (ball c R₁) (ball (f c) R₂)) + (h_z₀ : ∃ z₀ ∈ ball c R₁, ∥dslope f c z₀∥ = R₂ / R₁) : + ∃ C : E, ∥C∥ = R₂ / R₁ ∧ set.eq_on f (λ z, f c + (z - c) • C) (ball c R₁) := +let ⟨z₀, h_z₀, h_eq⟩ := h_z₀ in + ⟨dslope f c z₀, h_eq, affine_of_maps_to_ball_of_exists_norm_dslope_eq_div hd h_maps h_z₀ h_eq⟩ + /-- The **Schwarz Lemma**: if `f : ℂ → E` sends an open disk with center `c` and a positive radius `R₁` to an open ball with center `f c` and radius `R₂`, then the absolute value of the derivative of `f` at `c` is at most the ratio `R₂ / R₁`. -/ From ebced6361ef6cceed7c2167d1584800f96e1f01b Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sat, 3 Sep 2022 04:24:01 +0000 Subject: [PATCH 144/828] chore(scripts): update nolints.txt (#16361) I am happy to remove some nolints for you! --- scripts/nolints.txt | 35 ----------------------------------- 1 file changed, 35 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index d62ba5820a7c8..35424dd9601a4 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -553,43 +553,8 @@ apply_nolint is_cyclotomic_extension.number_field fintype_finite -- order/prime_ideal.lean apply_nolint order.ideal.is_prime.is_maximal fails_quickly --- ring_theory/finiteness.lean -apply_nolint algebra.finite_presentation.mv_polynomial fintype_finite -apply_nolint algebra.finite_presentation.mv_polynomial_of_finite_presentation fintype_finite -apply_nolint algebra.finite_type.mv_polynomial fintype_finite -apply_nolint module.finite.pi fintype_finite - --- ring_theory/ideal/quotient.lean -apply_nolint ideal.exists_sub_mem fintype_finite -apply_nolint ideal.map_pi fintype_finite -apply_nolint ideal.quotient_inf_to_pi_quotient_bijective fintype_finite - --- ring_theory/jacobson.lean -apply_nolint ideal.mv_polynomial.comp_C_integral_of_surjective_of_jacobson fintype_finite -apply_nolint ideal.mv_polynomial.mv_polynomial.ideal.is_jacobson fintype_finite - --- ring_theory/localization/integer.lean -apply_nolint is_localization.exist_integer_multiples_of_fintype fintype_finite - --- ring_theory/norm.lean -apply_nolint algebra.norm_eq_zero_iff_of_basis fintype_finite -apply_nolint algebra.norm_ne_zero_iff_of_basis fintype_finite - --- ring_theory/nullstellensatz.lean -apply_nolint mv_polynomial.is_maximal_iff_eq_vanishing_ideal_singleton fintype_finite -apply_nolint mv_polynomial.is_prime.vanishing_ideal_zero_locus fintype_finite -apply_nolint mv_polynomial.vanishing_ideal_zero_locus_eq_radical fintype_finite - --- ring_theory/polynomial/basic.lean -apply_nolint mv_polynomial.is_noetherian_ring fintype_finite -apply_nolint mv_polynomial.map_mv_polynomial_eq_eval₂ fintype_finite - -- ring_theory/trace.lean apply_nolint algebra.trace_comp_trace_of_basis fintype_finite -apply_nolint algebra.trace_trace_of_basis fintype_finite - --- ring_theory/witt_vector/basic.lean -apply_nolint witt_vector.comm_ring check_reducibility -- set_theory/lists.lean apply_nolint finsets doc_blame From 08fd57f20f141758f2dad82ecaf0cfb5acbbb85c Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sat, 3 Sep 2022 07:46:07 +0000 Subject: [PATCH 145/828] chore(topology/connected): speed up sum.is_connected_iff (#16353) --- src/topology/connected.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/topology/connected.lean b/src/topology/connected.lean index 6c1bde7257c5a..bc34f21e5a63d 100644 --- a/src/topology/connected.lean +++ b/src/topology/connected.lean @@ -531,19 +531,19 @@ lemma sum.is_connected_iff [topological_space β] {s : set (α ⊕ β)} : (∃ t, is_connected t ∧ s = sum.inl '' t) ∨ ∃ t, is_connected t ∧ s = sum.inr '' t := begin refine ⟨λ hs, _, _⟩, - { let u : set (α ⊕ β):= range sum.inl, + { let u : set (α ⊕ β) := range sum.inl, let v : set (α ⊕ β) := range sum.inr, have hu : is_open u, exact is_open_range_inl, obtain ⟨x | x, hx⟩ := hs.nonempty, - { have h := is_preconnected.subset_left_of_subset_union + { have h : s ⊆ range sum.inl := is_preconnected.subset_left_of_subset_union is_open_range_inl is_open_range_inr is_compl_range_inl_range_inr.disjoint - (show s ⊆ range sum.inl ∪ range sum.inr, by simp) ⟨sum.inl x, hx, x, rfl⟩ hs.2, + (by simp) ⟨sum.inl x, hx, x, rfl⟩ hs.2, refine or.inl ⟨sum.inl ⁻¹' s, _, _⟩, { exact hs.preimage_of_open_map sum.inl_injective open_embedding_inl.is_open_map h }, { exact (set.image_preimage_eq_of_subset h).symm } }, - { have h := is_preconnected.subset_right_of_subset_union + { have h : s ⊆ range sum.inr := is_preconnected.subset_right_of_subset_union is_open_range_inl is_open_range_inr is_compl_range_inl_range_inr.disjoint - (show s ⊆ range sum.inl ∪ range sum.inr, by simp) ⟨sum.inr x, hx, x, rfl⟩ hs.2, + (by simp) ⟨sum.inr x, hx, x, rfl⟩ hs.2, refine or.inr ⟨sum.inr ⁻¹' s, _, _⟩, { exact hs.preimage_of_open_map sum.inr_injective open_embedding_inr.is_open_map h }, { exact (set.image_preimage_eq_of_subset h).symm } } }, From 51b175cf8a384ddf27651a75607896568abfcd29 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sat, 3 Sep 2022 10:46:20 +0000 Subject: [PATCH 146/828] feat(number_theory/padics/padic_norm): add int_eq_one_iff (#16074) Add the proof that the p-adic norm of an integer is 1 iff the integer is not a multiple of p and related API for `padic_norm`. --- src/number_theory/padics/padic_norm.lean | 91 +++++++++++++++++++++++- 1 file changed, 89 insertions(+), 2 deletions(-) diff --git a/src/number_theory/padics/padic_norm.lean b/src/number_theory/padics/padic_norm.lean index 848016d6de9f9..818567dbd0e01 100644 --- a/src/number_theory/padics/padic_norm.lean +++ b/src/number_theory/padics/padic_norm.lean @@ -46,7 +46,6 @@ if q = 0 then 0 else (↑p : ℚ) ^ (-(padic_val_rat p q)) namespace padic_norm -section padic_norm open padic_val_rat variables (p : ℕ) @@ -286,5 +285,93 @@ begin { exact_mod_cast hp.1.one_lt } } end -end padic_norm +/-- The `p`-adic norm of an integer `m` is one iff `p` doesn't divide `m`. -/ +lemma int_eq_one_iff (m : ℤ) : padic_norm p m = 1 ↔ ¬ (p : ℤ) ∣ m := +begin + nth_rewrite 1 ← pow_one p, + simp only [dvd_iff_norm_le, int.cast_coe_nat, nat.cast_one, zpow_neg, zpow_one, not_le], + split, + { intro h, + rw [h, inv_lt_one_iff_of_pos]; + norm_cast, + { exact nat.prime.one_lt (fact.out _), }, + { exact nat.prime.pos (fact.out _), }, }, + { simp only [padic_norm], + split_ifs, + { rw [inv_lt_zero, ← nat.cast_zero, nat.cast_lt], + intro h, exact (nat.not_lt_zero p h).elim, }, + { have : 1 < (p : ℚ) := by norm_cast; exact (nat.prime.one_lt (fact.out _ : nat.prime p)), + rw [← zpow_neg_one, zpow_lt_iff_lt this], + have : 0 ≤ padic_val_rat p m, simp only [of_int, nat.cast_nonneg], + intro h, + rw [← zpow_zero (p : ℚ), zpow_inj]; + linarith, } }, +end + +lemma int_lt_one_iff (m : ℤ) : padic_norm p m < 1 ↔ (p : ℤ) ∣ m := +begin + rw [← not_iff_not, ← int_eq_one_iff, eq_iff_le_not_lt], + simp only [padic_norm.of_int, true_and], +end + +lemma of_nat (m : ℕ) : padic_norm p m ≤ 1 := +padic_norm.of_int p (m : ℤ) + +/-- The `p`-adic norm of a natural `m` is one iff `p` doesn't divide `m`. -/ +lemma nat_eq_one_iff (m : ℕ) : padic_norm p m = 1 ↔ ¬ p ∣ m := +by simp only [←int.coe_nat_dvd, ←int_eq_one_iff, int.cast_coe_nat] + +lemma nat_lt_one_iff (m : ℕ) : padic_norm p m < 1 ↔ p ∣ m := +by simp only [←int.coe_nat_dvd, ←int_lt_one_iff, int.cast_coe_nat] + +open_locale big_operators + +lemma sum_lt {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} : + s.nonempty → (∀ i ∈ s, padic_norm p (F i) < t) → + padic_norm p (∑ i in s, F i) < t := +begin + classical, + refine s.induction_on (by { rintro ⟨-, ⟨⟩⟩, }) _, + rintro a S haS IH - ht, + by_cases hs : S.nonempty, + { rw finset.sum_insert haS, + exact lt_of_le_of_lt (padic_norm.nonarchimedean p) (max_lt + (ht a (finset.mem_insert_self a S)) + (IH hs (λ b hb, ht b (finset.mem_insert_of_mem hb)))), }, + { simp * at *, }, +end + +lemma sum_le {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} : + s.nonempty → (∀ i ∈ s, padic_norm p (F i) ≤ t) → + padic_norm p (∑ i in s, F i) ≤ t := +begin + classical, + refine s.induction_on (by { rintro ⟨-, ⟨⟩⟩, }) _, + rintro a S haS IH - ht, + by_cases hs : S.nonempty, + { rw finset.sum_insert haS, + exact (padic_norm.nonarchimedean p).trans (max_le + (ht a (finset.mem_insert_self a S)) + (IH hs (λ b hb, ht b (finset.mem_insert_of_mem hb)))), }, + { simp * at *, }, +end + +lemma sum_lt' {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} + (hF : ∀ i ∈ s, padic_norm p (F i) < t) (ht : 0 < t) : + padic_norm p (∑ i in s, F i) < t := +begin + obtain rfl | hs := finset.eq_empty_or_nonempty s, + { simp [ht], }, + { exact sum_lt hs hF, }, +end + +lemma sum_le' {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} + (hF : ∀ i ∈ s, padic_norm p (F i) ≤ t) (ht : 0 ≤ t) : + padic_norm p (∑ i in s, F i) ≤ t := +begin + obtain rfl | hs := finset.eq_empty_or_nonempty s, + { simp [ht], }, + { exact sum_le hs hF, }, +end + end padic_norm From 882fc1323cfd7c86aa3a39daaf4df3dd1cd4cba6 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Sat, 3 Sep 2022 10:46:21 +0000 Subject: [PATCH 147/828] refactor(group_theory/finiteness): Make `group.rank` noncomputable (#16256) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I've been working on some more API for `group.rank`, and the decidability hypothesis was making lemma statements really clunky. If you wanted to say `group.rank G ≤ group.rank G'`, then the decidability hypotheses would add two lines to the statement of the lemma. --- src/group_theory/finiteness.lean | 15 ++++++--------- src/group_theory/schreier.lean | 6 ++---- 2 files changed, 8 insertions(+), 13 deletions(-) diff --git a/src/group_theory/finiteness.lean b/src/group_theory/finiteness.lean index e75ebff295765..1ff57337b4b2f 100644 --- a/src/group_theory/finiteness.lean +++ b/src/group_theory/finiteness.lean @@ -282,19 +282,16 @@ variables (G) /-- The minimum number of generators of a group. -/ @[to_additive "The minimum number of generators of an additive group"] -def group.rank [h : group.fg G] - [decidable_pred (λ n, ∃ (S : finset G), S.card = n ∧ subgroup.closure (S : set G) = ⊤)] := -nat.find (group.fg_iff'.mp h) +noncomputable def group.rank [h : group.fg G] := +@nat.find _ (classical.dec_pred _) (group.fg_iff'.mp h) -@[to_additive] lemma group.rank_spec [h : group.fg G] - [decidable_pred (λ n, ∃ (S : finset G), S.card = n ∧ subgroup.closure (S : set G) = ⊤)] : +@[to_additive] lemma group.rank_spec [h : group.fg G] : ∃ S : finset G, S.card = group.rank G ∧ subgroup.closure (S : set G) = ⊤ := -nat.find_spec (group.fg_iff'.mp h) +@nat.find_spec _ (classical.dec_pred _) (group.fg_iff'.mp h) -@[to_additive] lemma group.rank_le [group.fg G] - [decidable_pred (λ n, ∃ (S : finset G), S.card = n ∧ subgroup.closure (S : set G) = ⊤)] +@[to_additive] lemma group.rank_le [h : group.fg G] {S : finset G} (hS : subgroup.closure (S : set G) = ⊤) : group.rank G ≤ S.card := -nat.find_le ⟨S, rfl, hS⟩ +@nat.find_le _ _ (classical.dec_pred _) (group.fg_iff'.mp h) ⟨S, rfl, hS⟩ end group diff --git a/src/group_theory/schreier.lean b/src/group_theory/schreier.lean index c787ff596e72f..bf5b797a3e06a 100644 --- a/src/group_theory/schreier.lean +++ b/src/group_theory/schreier.lean @@ -133,10 +133,8 @@ begin exact ⟨⟨T, hT⟩⟩, end -lemma rank_le_index_mul_rank [hG : group.fg G] {H : subgroup G} (hH : H.index ≠ 0) - [decidable_pred (λ n, ∃ (S : finset G), S.card = n ∧ subgroup.closure (S : set G) = ⊤)] - [decidable_pred (λ n, ∃ (S : finset H), S.card = n ∧ subgroup.closure (S : set H) = ⊤)] : - @group.rank H _ (fg_of_index_ne_zero hH) _ ≤ H.index * group.rank G := +lemma rank_le_index_mul_rank [hG : group.fg G] {H : subgroup G} (hH : H.index ≠ 0) : + @group.rank H _ (fg_of_index_ne_zero hH) ≤ H.index * group.rank G := begin haveI := fg_of_index_ne_zero hH, obtain ⟨S, hS₀, hS⟩ := group.rank_spec G, From 1995c7bbdbb0adb1b6d5acdc654f6cf46ed96cfa Mon Sep 17 00:00:00 2001 From: Mauricio Collares Date: Sat, 3 Sep 2022 13:10:34 +0000 Subject: [PATCH 148/828] chore(.github/workflows): revert first-run timeout to -T100000 (#16356) This reverts part of #15251, now that #15784 works around the olean issue. Co-authored-by: Junyan Xu Co-authored-by: Mauricio Collares <244239+collares@users.noreply.github.com> --- .github/workflows/bors.yml | 4 ++-- .github/workflows/build.yml | 4 ++-- .github/workflows/build.yml.in | 4 ++-- .github/workflows/build_fork.yml | 4 ++-- src/algebraic_topology/dold_kan/functor_n.lean | 4 +++- src/analysis/calculus/implicit.lean | 2 +- src/category_theory/category/Cat/limit.lean | 8 ++++++-- src/category_theory/limits/cones.lean | 12 +++++++----- src/category_theory/monad/basic.lean | 16 +++++++++++++--- src/linear_algebra/invariant_basis_number.lean | 6 ++++-- src/number_theory/cyclotomic/basic.lean | 4 ++-- src/ring_theory/ideal/quotient.lean | 12 ++++++------ src/ring_theory/roots_of_unity.lean | 6 +++++- 13 files changed, 55 insertions(+), 31 deletions(-) diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index d1605c462fc23..e43db8caf1548 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -81,7 +81,7 @@ jobs: set -o pipefail leanpkg configure echo "::set-output name=started::true" - lean --json -T150000 --make src | python3 scripts/detect_errors.py + lean --json -T100000 --make src | python3 scripts/detect_errors.py lean --json -T400000 --make src | python3 scripts/detect_errors.py - name: push release to azure @@ -197,7 +197,7 @@ jobs: - name: tests run: | set -o pipefail - lean --json -T150000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py + lean --json -T100000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py - uses: actions/setup-java@v2 if: ${{ ! 1 }} diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index a48a77a703d74..8296a49e4e11b 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -89,7 +89,7 @@ jobs: set -o pipefail leanpkg configure echo "::set-output name=started::true" - lean --json -T150000 --make src | python3 scripts/detect_errors.py + lean --json -T100000 --make src | python3 scripts/detect_errors.py lean --json -T400000 --make src | python3 scripts/detect_errors.py - name: push release to azure @@ -205,7 +205,7 @@ jobs: - name: tests run: | set -o pipefail - lean --json -T150000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py + lean --json -T100000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py - uses: actions/setup-java@v2 if: ${{ ! 1 }} diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index 395e1e1a2a553..ffed44267c589 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -67,7 +67,7 @@ jobs: set -o pipefail leanpkg configure echo "::set-output name=started::true" - lean --json -T150000 --make src | python3 scripts/detect_errors.py + lean --json -T100000 --make src | python3 scripts/detect_errors.py lean --json -T400000 --make src | python3 scripts/detect_errors.py - name: push release to azure @@ -183,7 +183,7 @@ jobs: - name: tests run: | set -o pipefail - lean --json -T150000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py + lean --json -T100000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py - uses: actions/setup-java@v2 if: ${{ ! IS_SELF_HOSTED }} diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index 509d135f52d01..3a093a8e654f4 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -87,7 +87,7 @@ jobs: set -o pipefail leanpkg configure echo "::set-output name=started::true" - lean --json -T150000 --make src | python3 scripts/detect_errors.py + lean --json -T100000 --make src | python3 scripts/detect_errors.py lean --json -T400000 --make src | python3 scripts/detect_errors.py - name: push release to azure @@ -203,7 +203,7 @@ jobs: - name: tests run: | set -o pipefail - lean --json -T150000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py + lean --json -T100000 --make docs archive roadmap test counterexamples | python3 scripts/detect_errors.py - uses: actions/setup-java@v2 if: ${{ ! 0 }} diff --git a/src/algebraic_topology/dold_kan/functor_n.lean b/src/algebraic_topology/dold_kan/functor_n.lean index 1a54418c218d9..2f0888cc7c0d4 100644 --- a/src/algebraic_topology/dold_kan/functor_n.lean +++ b/src/algebraic_topology/dold_kan/functor_n.lean @@ -55,7 +55,9 @@ def N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ) := idem := P_infty_idem, }, map := λ X Y f, { f := P_infty ≫ alternating_face_map_complex.map f, - comm := by tidy, }, } + comm := by { ext, simp }, }, + map_id' := λ X, by { ext, dsimp, simp }, + map_comp' := λ X Y Z f g, by { ext, simp } } /-- The extension of `N₁` to the Karoubi envelope of `simplicial_object C`. -/ @[simps] diff --git a/src/analysis/calculus/implicit.lean b/src/analysis/calculus/implicit.lean index 622726d32387c..80704de48ffb0 100644 --- a/src/analysis/calculus/implicit.lean +++ b/src/analysis/calculus/implicit.lean @@ -312,7 +312,7 @@ lemma to_implicit_function_of_complemented (hf : has_strict_fderiv_at f f' a) f'.ker.subtypeL 0 := by convert (implicit_function_data_of_complemented f f' hf hf' hker).implicit_function_has_strict_fderiv_at f'.ker.subtypeL _ _; - [skip, ext, ext]; simp [classical.some_spec hker] + simp only [implicit_function_data_of_complemented]; ext; simp [classical.some_spec hker] end complemented diff --git a/src/category_theory/category/Cat/limit.lean b/src/category_theory/category/Cat/limit.lean index cf29743be6132..778299a632e33 100644 --- a/src/category_theory/category/Cat/limit.lean +++ b/src/category_theory/category/Cat/limit.lean @@ -68,7 +68,9 @@ instance (F : J ⥤ Cat.{v v}) : category (limit (F ⋙ Cat.objects)) := rw [←congr_fun (limit.w (hom_diagram X Y) h) f, ←congr_fun (limit.w (hom_diagram Y Z) h) g], dsimp, simp, - end), } + end), + id_comp' := λ _ _ _, by { ext, simp only [category.id_comp, types.limit.π_mk'] }, + comp_id' := λ _ _ _, by { ext, simp only [types.limit.π_mk', category.comp_id] } } /-- Auxiliary definition: the limit category. -/ @[simps] @@ -110,7 +112,9 @@ def limit_cone_lift (F : J ⥤ Cat.{v v}) (s : cone F) : s.X ⟶ limit_cone_X F conv at this { congr, skip, dsimp, simp, }, erw [functor.congr_hom this f], dsimp, simp, }, - end, } + end, + map_id' := λ X, by simp, + map_comp' := λ X Y Z f g, by simp } @[simp] lemma limit_π_hom_diagram_eq_to_hom {F : J ⥤ Cat.{v v}} diff --git a/src/category_theory/limits/cones.lean b/src/category_theory/limits/cones.lean index b34b822e37d54..4ceb2a89461ee 100644 --- a/src/category_theory/limits/cones.lean +++ b/src/category_theory/limits/cones.lean @@ -715,18 +715,20 @@ def cocone_equivalence_op_cone_op : cocone F ≌ (cone F.op)ᵒᵖ := { obj := λ c, op (cocone.op c), map := λ X Y f, quiver.hom.op { hom := f.hom.op, - w' := λ j, by { apply quiver.hom.unop_inj, dsimp, simp, }, } }, + w' := λ j, by { apply quiver.hom.unop_inj, dsimp, apply cocone_morphism.w }, } }, inverse := { obj := λ c, cone.unop (unop c), map := λ X Y f, { hom := f.unop.hom.unop, - w' := λ j, by { apply quiver.hom.op_inj, dsimp, simp, }, } }, - unit_iso := nat_iso.of_components (λ c, cocones.ext (iso.refl _) (by tidy)) (by tidy), + w' := λ j, by { apply quiver.hom.op_inj, dsimp, apply cone_morphism.w }, } }, + unit_iso := nat_iso.of_components (λ c, + cocones.ext (iso.refl _) (by { dsimp, simp })) (λ X Y f, by { ext, simp }), counit_iso := nat_iso.of_components (λ c, by { induction c using opposite.rec, - dsimp, apply iso.op, exact cones.ext (iso.refl _) (by tidy), }) + dsimp, apply iso.op, exact cones.ext (iso.refl _) (by { dsimp, simp }), }) (λ X Y f, quiver.hom.unop_inj (cone_morphism.ext _ _ (by { dsimp, simp }))), - functor_unit_iso_comp' := λ c, begin apply quiver.hom.unop_inj, ext, dsimp, simp, end } + functor_unit_iso_comp' := (λ c, + by { apply quiver.hom.unop_inj, ext, dsimp, apply comp_id })} attribute [simps] cocone_equivalence_op_cone_op diff --git a/src/category_theory/monad/basic.lean b/src/category_theory/monad/basic.lean index cc9ac77b2929b..2516f941e3bc9 100644 --- a/src/category_theory/monad/basic.lean +++ b/src/category_theory/monad/basic.lean @@ -139,13 +139,23 @@ instance : category (monad C) := { hom := monad_hom, id := λ M, { to_nat_trans := 𝟙 (M : C ⥤ C) }, comp := λ _ _ _ f g, - { to_nat_trans := { app := λ X, f.app X ≫ g.app X } } } + { to_nat_trans := + { app := λ X, f.app X ≫ g.app X, + naturality' := λ X Y h, by rw [assoc, f.1.naturality_assoc, g.1.naturality] } }, + id_comp' := λ _ _ _, by {ext, apply id_comp}, + comp_id' := λ _ _ _, by {ext, apply comp_id}, + assoc' := λ _ _ _ _ _ _ _, by {ext, apply assoc} } instance : category (comonad C) := { hom := comonad_hom, id := λ M, { to_nat_trans := 𝟙 (M : C ⥤ C) }, - comp := λ M N L f g, - { to_nat_trans := { app := λ X, f.app X ≫ g.app X } } } + comp := λ _ _ _ f g, + { to_nat_trans := + { app := λ X, f.app X ≫ g.app X, + naturality' := λ X Y h, by rw [assoc, f.1.naturality_assoc, g.1.naturality] } }, + id_comp' := λ _ _ _, by {ext, apply id_comp}, + comp_id' := λ _ _ _, by {ext, apply comp_id}, + assoc' := λ _ _ _ _ _ _ _, by {ext, apply assoc} } instance {T : monad C} : inhabited (monad_hom T T) := ⟨𝟙 T⟩ diff --git a/src/linear_algebra/invariant_basis_number.lean b/src/linear_algebra/invariant_basis_number.lean index 1d08b949b58ff..e18738c0f7e00 100644 --- a/src/linear_algebra/invariant_basis_number.lean +++ b/src/linear_algebra/invariant_basis_number.lean @@ -252,8 +252,10 @@ private def induced_equiv [fintype ι'] (I : ideal R) (e : (ι → R) ≃ₗ[R] begin refine { to_fun := induced_map I e, inv_fun := induced_map I e.symm, .. }, all_goals { rintro ⟨a⟩ ⟨b⟩ <|> rintro ⟨a⟩, - change ideal.quotient.mk _ _ = ideal.quotient.mk _ _, - congr, simp } + convert_to ideal.quotient.mk _ _ = ideal.quotient.mk _ _, + congr, + simp only [map_add, linear_equiv.coe_coe, linear_equiv.map_smulₛₗ, ring_hom.id_apply, + linear_equiv.symm_apply_apply, linear_equiv.apply_symm_apply] } end end diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index ac5a90622bbb6..a01087be87cca 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -552,12 +552,12 @@ instance [ne_zero ((n : ℕ) : A)] : refine ⟨⟨a.1 * b.2 + b.1 * a.2, a.2 * b.2, mul_mem_non_zero_divisors.2 ⟨a.2.2, b.2.2⟩⟩, _⟩, rw [set_like.coe_mk, ring_hom.map_mul, add_mul, ← mul_assoc, ha, mul_comm ((algebra_map _ _) ↑a.2), ← mul_assoc, hb], - simp }, + simp only [map_add, map_mul] }, { rintro y z ⟨a, ha⟩ ⟨b, hb⟩, refine ⟨⟨a.1 * b.1, a.2 * b.2, mul_mem_non_zero_divisors.2 ⟨a.2.2, b.2.2⟩⟩, _⟩, rw [set_like.coe_mk, ring_hom.map_mul, mul_comm ((algebra_map _ _) ↑a.2), mul_assoc, ← mul_assoc z, hb, ← mul_comm ((algebra_map _ _) ↑a.2), ← mul_assoc, ha], - simp } + simp only [map_mul] } end, eq_iff_exists := λ x y, ⟨λ h, ⟨1, by rw adjoin_algebra_injective n A K h⟩, λ ⟨c, hc⟩, by rw mul_right_cancel₀ (non_zero_divisors.ne_zero c.prop) hc⟩ } diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index 65c02a553dcd5..41952cd53bd88 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -255,33 +255,33 @@ instance module_pi : module (R ⧸ I) ((ι → R) ⧸ I.pi ι) := end, one_smul := begin rintro ⟨a⟩, - change ideal.quotient.mk _ _ = ideal.quotient.mk _ _, + convert_to ideal.quotient.mk _ _ = ideal.quotient.mk _ _, congr' with i, exact one_mul (a i), end, mul_smul := begin rintro ⟨a⟩ ⟨b⟩ ⟨c⟩, - change ideal.quotient.mk _ _ = ideal.quotient.mk _ _, + convert_to ideal.quotient.mk _ _ = ideal.quotient.mk _ _, simp only [(•)], congr' with i, exact mul_assoc a b (c i), end, smul_add := begin rintro ⟨a⟩ ⟨b⟩ ⟨c⟩, - change ideal.quotient.mk _ _ = ideal.quotient.mk _ _, + convert_to ideal.quotient.mk _ _ = ideal.quotient.mk _ _, congr' with i, exact mul_add a (b i) (c i), end, smul_zero := begin rintro ⟨a⟩, - change ideal.quotient.mk _ _ = ideal.quotient.mk _ _, + convert_to ideal.quotient.mk _ _ = ideal.quotient.mk _ _, congr' with i, exact mul_zero a, end, add_smul := begin rintro ⟨a⟩ ⟨b⟩ ⟨c⟩, - change ideal.quotient.mk _ _ = ideal.quotient.mk _ _, + convert_to ideal.quotient.mk _ _ = ideal.quotient.mk _ _, congr' with i, exact add_mul a b (c i), end, zero_smul := begin rintro ⟨a⟩, - change ideal.quotient.mk _ _ = ideal.quotient.mk _ _, + convert_to ideal.quotient.mk _ _ = ideal.quotient.mk _ _, congr' with i, exact zero_mul (a i), end, } diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity.lean index 0bab3118e50aa..453fd90d7e9b9 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity.lean @@ -1090,7 +1090,7 @@ variables {S} [comm_ring S] [is_domain S] {μ : S} {n : ℕ+} (hμ : is_primitiv (R) [comm_ring R] [algebra R S] /-- The `monoid_hom` that takes an automorphism to the power of μ that μ gets mapped to under it. -/ -@[simps {attrs := []}] noncomputable def aut_to_pow : (S ≃ₐ[R] S) →* (zmod n)ˣ := +noncomputable def aut_to_pow : (S ≃ₐ[R] S) →* (zmod n)ˣ := let μ' := hμ.to_roots_of_unity in have ho : order_of μ' = n := by rw [hμ.eq_order_of, ←hμ.coe_to_roots_of_unity_coe, order_of_units, order_of_subgroup], @@ -1122,6 +1122,10 @@ monoid_hom.to_hom_units rw [←nat.cast_mul, zmod.nat_coe_eq_nat_coe_iff, ←ho, ←pow_eq_pow_iff_modeq μ', hxy] end } +-- We are not using @[simps] in aut_to_pow to avoid a timeout. +lemma coe_aut_to_pow_apply (f : S ≃ₐ[R] S) : (aut_to_pow R hμ f : zmod n) = + ((map_root_of_unity_eq_pow_self f hμ.to_roots_of_unity).some : zmod n) := rfl + @[simp] lemma aut_to_pow_spec (f : S ≃ₐ[R] S) : μ ^ (hμ.aut_to_pow R f : zmod n).val = f μ := begin From a6ca2e8fdd1bbd3c4e037cf8dd91fe8d4c863324 Mon Sep 17 00:00:00 2001 From: Vincent Beffara Date: Sat, 3 Sep 2022 22:00:07 +0000 Subject: [PATCH 149/828] feat(analysis/analytic/isolated_zeros): Principle of isolated zeros (#15908) Local version: an analytic function is either locally zero, or nonzero in a punctured neighborhood. Co-authored-by: ADedecker --- src/analysis/analytic/basic.lean | 89 +++++++++++ src/analysis/analytic/isolated_zeros.lean | 139 ++++++++++++++++++ .../calculus/formal_multilinear_series.lean | 125 ++++++++++++++++ src/analysis/normed_space/multilinear.lean | 16 ++ src/linear_algebra/multilinear/basic.lean | 17 +++ src/topology/algebra/module/multilinear.lean | 3 + src/topology/separation.lean | 10 ++ 7 files changed, 399 insertions(+) create mode 100644 src/analysis/analytic/isolated_zeros.lean diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index dded622c580d3..d141ef6a92920 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -407,12 +407,47 @@ lemma has_fpower_series_on_ball.mono has_fpower_series_on_ball f p x r' := ⟨le_trans hr hf.1, r'_pos, λ y hy, hf.has_sum (emetric.ball_subset_ball hr hy)⟩ +lemma has_fpower_series_at.congr (hf : has_fpower_series_at f p x) (hg : f =ᶠ[𝓝 x] g) : + has_fpower_series_at g p x := +begin + rcases hf with ⟨r₁, h₁⟩, + rcases emetric.mem_nhds_iff.mp hg with ⟨r₂, h₂pos, h₂⟩, + exact ⟨min r₁ r₂, (h₁.mono (lt_min h₁.r_pos h₂pos) inf_le_left).congr + (λ y hy, h₂ (emetric.ball_subset_ball inf_le_right hy))⟩ +end + protected lemma has_fpower_series_at.eventually (hf : has_fpower_series_at f p x) : ∀ᶠ r : ℝ≥0∞ in 𝓝[>] 0, has_fpower_series_on_ball f p x r := let ⟨r, hr⟩ := hf in mem_of_superset (Ioo_mem_nhds_within_Ioi (left_mem_Ico.2 hr.r_pos)) $ λ r' hr', hr.mono hr'.1 hr'.2.le +lemma has_fpower_series_on_ball.eventually_has_sum (hf : has_fpower_series_on_ball f p x r) : + ∀ᶠ y in 𝓝 0, has_sum (λn:ℕ, p n (λ(i : fin n), y)) (f (x + y)) := +by filter_upwards [emetric.ball_mem_nhds (0 : E) hf.r_pos] using λ _, hf.has_sum + +lemma has_fpower_series_at.eventually_has_sum (hf : has_fpower_series_at f p x) : + ∀ᶠ y in 𝓝 0, has_sum (λn:ℕ, p n (λ(i : fin n), y)) (f (x + y)) := +let ⟨r, hr⟩ := hf in hr.eventually_has_sum + +lemma has_fpower_series_on_ball.eventually_has_sum_sub (hf : has_fpower_series_on_ball f p x r) : + ∀ᶠ y in 𝓝 x, has_sum (λn:ℕ, p n (λ(i : fin n), y - x)) (f y) := +by filter_upwards [emetric.ball_mem_nhds x hf.r_pos] with y using hf.has_sum_sub + +lemma has_fpower_series_at.eventually_has_sum_sub (hf : has_fpower_series_at f p x) : + ∀ᶠ y in 𝓝 x, has_sum (λn:ℕ, p n (λ(i : fin n), y - x)) (f y) := +let ⟨r, hr⟩ := hf in hr.eventually_has_sum_sub + +lemma has_fpower_series_on_ball.eventually_eq_zero + (hf : has_fpower_series_on_ball f (0 : formal_multilinear_series 𝕜 E F) x r) : + ∀ᶠ z in 𝓝 x, f z = 0 := +by filter_upwards [hf.eventually_has_sum_sub] with z hz using hz.unique has_sum_zero + +lemma has_fpower_series_at.eventually_eq_zero + (hf : has_fpower_series_at f (0 : formal_multilinear_series 𝕜 E F) x) : + ∀ᶠ z in 𝓝 x, f z = 0 := +let ⟨r, hr⟩ := hf in hr.eventually_eq_zero + lemma has_fpower_series_on_ball.add (hf : has_fpower_series_on_ball f pf x r) (hg : has_fpower_series_on_ball g pg x r) : has_fpower_series_on_ball (f + g) (pf + pg) x r := @@ -840,6 +875,17 @@ theorem has_fpower_series_at.eq_formal_multilinear_series p₁ = p₂ := sub_eq_zero.mp (has_fpower_series_at.eq_zero (by simpa only [sub_self] using h₁.sub h₂)) +lemma has_fpower_series_at.eq_formal_multilinear_series_of_eventually + {p q : formal_multilinear_series 𝕜 𝕜 E} {f g : 𝕜 → E} {x : 𝕜} (hp : has_fpower_series_at f p x) + (hq : has_fpower_series_at g q x) (heq : ∀ᶠ z in 𝓝 x, f z = g z) : + p = q := +(hp.congr heq).eq_formal_multilinear_series hq + +/-- A one-dimensional formal multilinear series representing a locally zero function is zero. -/ +lemma has_fpower_series_at.eq_zero_of_eventually {p : formal_multilinear_series 𝕜 𝕜 E} {f : 𝕜 → E} + {x : 𝕜} (hp : has_fpower_series_at f p x) (hf : f =ᶠ[𝓝 x] 0) : p = 0 := +(hp.congr hf).eq_zero + /-- If a function `f : 𝕜 → E` has two power series representations at `x`, then the given radii in which convergence is guaranteed may be interchanged. This can be useful when the formal multilinear series in one representation has a particularly nice form, but the other has a larger radius. -/ @@ -1180,3 +1226,46 @@ begin end end + +section + +open formal_multilinear_series + +variables {p : formal_multilinear_series 𝕜 𝕜 E} {f : 𝕜 → E} {z₀ : 𝕜} + +/-- A function `f : 𝕜 → E` has `p` as power series expansion at a point `z₀` iff it is the sum of +`p` in a neighborhood of `z₀`. This makes some proofs easier by hiding the fact that +`has_fpower_series_at` depends on `p.radius`. -/ +lemma has_fpower_series_at_iff : has_fpower_series_at f p z₀ ↔ + ∀ᶠ z in 𝓝 0, has_sum (λ n, z ^ n • p.coeff n) (f (z₀ + z)) := +begin + refine ⟨λ ⟨r, r_le, r_pos, h⟩, eventually_of_mem (emetric.ball_mem_nhds 0 r_pos) + (λ _, by simpa using h), _⟩, + simp only [metric.eventually_nhds_iff], + rintro ⟨r, r_pos, h⟩, + refine ⟨p.radius ⊓ r.to_nnreal, by simp, _, _⟩, + { simp only [r_pos.lt, lt_inf_iff, ennreal.coe_pos, real.to_nnreal_pos, and_true], + obtain ⟨z, z_pos, le_z⟩ := normed_field.exists_norm_lt 𝕜 r_pos.lt, + have : (∥z∥₊ : ennreal) ≤ p.radius, + by { simp only [dist_zero_right] at h, + apply formal_multilinear_series.le_radius_of_tendsto, + convert tendsto_norm.comp (h le_z).summable.tendsto_at_top_zero, + funext; simp [norm_smul, mul_comm] }, + refine lt_of_lt_of_le _ this, + simp only [ennreal.coe_pos], + exact zero_lt_iff.mpr (nnnorm_ne_zero_iff.mpr (norm_pos_iff.mp z_pos)) }, + { simp only [emetric.mem_ball, lt_inf_iff, edist_lt_coe, apply_eq_pow_smul_coeff, and_imp, + dist_zero_right] at h ⊢, + refine λ y hyp hyr, h _, + simpa [nndist_eq_nnnorm, real.lt_to_nnreal_iff_coe_lt] using hyr } +end + +lemma has_fpower_series_at_iff' : has_fpower_series_at f p z₀ ↔ + ∀ᶠ z in 𝓝 z₀, has_sum (λ n, (z - z₀) ^ n • p.coeff n) (f z) := +begin + rw [← map_add_left_nhds_zero, eventually_map, has_fpower_series_at_iff], + congrm ∀ᶠ z in (𝓝 0 : filter 𝕜), has_sum (λ n, _) (f (z₀ + z)), + rw add_sub_cancel' +end + +end diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean new file mode 100644 index 0000000000000..dfc719faf9dcf --- /dev/null +++ b/src/analysis/analytic/isolated_zeros.lean @@ -0,0 +1,139 @@ +/- +Copyright (c) 2022 Vincent Beffara. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Vincent Beffara +-/ +import analysis.analytic.basic +import analysis.calculus.dslope +import analysis.calculus.fderiv_analytic +import analysis.calculus.formal_multilinear_series +import analysis.complex.basic +import topology.algebra.infinite_sum + +/-! +# Principle of isolated zeros + +This file proves the fact that the zeros of a non-constant analytic function of one variable are +isolated. It also introduces a little bit of API in the `has_fpower_series_at` namespace that +is useful in this setup. + +## Main results + +* `analytic_at.eventually_eq_zero_or_eventually_ne_zero` is the main statement that if a function is + analytic at `z₀`, then either it is identically zero in a neighborhood of `z₀`, or it does not + vanish in a punctured neighborhood of `z₀`. +-/ + +open_locale classical + +open filter function nat formal_multilinear_series emetric +open_locale topological_space big_operators + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] + {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {s : E} + {p q : formal_multilinear_series 𝕜 𝕜 E} {f g : 𝕜 → E} + {n : ℕ} {z z₀ : 𝕜} {y : fin n → 𝕜} + +namespace has_sum + +variables {a : ℕ → E} + +lemma has_sum_at_zero (a : ℕ → E) : has_sum (λ n, (0:𝕜) ^ n • a n) (a 0) := +by convert has_sum_single 0 (λ b h, _); simp [nat.pos_of_ne_zero h] <|> simp + +lemma exists_has_sum_smul_of_apply_eq_zero (hs : has_sum (λ m, z ^ m • a m) s) + (ha : ∀ k < n, a k = 0) : + ∃ t : E, z ^ n • t = s ∧ has_sum (λ m, z ^ m • a (m + n)) t := +begin + refine classical.by_cases (λ hn : n = 0, by { subst n; simpa }) (λ hn, _), + replace hn := nat.pos_of_ne_zero hn, + by_cases (z = 0), + { have : s = 0 := hs.unique (by simpa [ha 0 hn, h] using has_sum_at_zero a), + exact ⟨a n, by simp [h, hn, this], by simpa [h] using has_sum_at_zero (λ m, a (m + n))⟩ }, + { refine ⟨(z ^ n)⁻¹ • s, by field_simp [smul_smul], _⟩, + have h1 : ∑ i in finset.range n, z ^ i • a i = 0, + from finset.sum_eq_zero (λ k hk, by simp [ha k (finset.mem_range.mp hk)]), + have h2 : has_sum (λ m, z ^ (m + n) • a (m + n)) s, + by simpa [h1] using (has_sum_nat_add_iff' n).mpr hs, + convert @has_sum.const_smul E ℕ 𝕜 _ _ _ _ _ _ _ (z⁻¹ ^ n) h2, + field_simp [pow_add, smul_smul], simp only [inv_pow] } +end + +end has_sum + +namespace has_fpower_series_at + +lemma has_fpower_series_dslope_fslope (hp : has_fpower_series_at f p z₀) : + has_fpower_series_at (dslope f z₀) p.fslope z₀ := +begin + have hpd : deriv f z₀ = p.coeff 1 := hp.deriv, + have hp0 : p.coeff 0 = f z₀ := hp.coeff_zero 1, + simp only [has_fpower_series_at_iff, apply_eq_pow_smul_coeff, coeff_fslope] at hp ⊢, + refine hp.mono (λ x hx, _), + by_cases x = 0, + { convert has_sum_single 0 _; intros; simp [*] }, + { have hxx : ∀ (n : ℕ), x⁻¹ * x ^ (n + 1) = x ^ n := λ n, by field_simp [h, pow_succ'], + suffices : has_sum (λ n, x⁻¹ • x ^ (n + 1) • p.coeff (n + 1)) (x⁻¹ • (f (z₀ + x) - f z₀)), + { simpa [dslope, slope, h, smul_smul, hxx] using this }, + { simpa [hp0] using ((has_sum_nat_add_iff' 1).mpr hx).const_smul } } +end + +lemma has_fpower_series_iterate_dslope_fslope (n : ℕ) (hp : has_fpower_series_at f p z₀) : + has_fpower_series_at ((swap dslope z₀)^[n] f) (fslope^[n] p) z₀ := +begin + induction n with n ih generalizing f p, + { exact hp }, + { simpa using ih (has_fpower_series_dslope_fslope hp) } +end + +lemma iterate_dslope_fslope_ne_zero (hp : has_fpower_series_at f p z₀) (h : p ≠ 0) : + (swap dslope z₀)^[p.order] f z₀ ≠ 0 := +begin + rw [← coeff_zero (has_fpower_series_iterate_dslope_fslope p.order hp) 1], + simpa [coeff_eq_zero] using apply_order_ne_zero h +end + +lemma eq_pow_order_mul_iterate_dslope (hp : has_fpower_series_at f p z₀) : + ∀ᶠ z in 𝓝 z₀, f z = (z - z₀) ^ p.order • ((swap dslope z₀)^[p.order] f z) := +begin + have hq := has_fpower_series_at_iff'.mp (has_fpower_series_iterate_dslope_fslope p.order hp), + filter_upwards [hq, has_fpower_series_at_iff'.mp hp] with x hx1 hx2, + have : ∀ k < p.order, p.coeff k = 0, + from λ k hk, by simpa [coeff_eq_zero] using apply_eq_zero_of_lt_order hk, + obtain ⟨s, hs1, hs2⟩ := has_sum.exists_has_sum_smul_of_apply_eq_zero hx2 this, + convert hs1.symm, + simp only [coeff_iterate_fslope] at hx1, + exact hx1.unique hs2 +end + +lemma locally_ne_zero (hp : has_fpower_series_at f p z₀) (h : p ≠ 0) : + ∀ᶠ z in 𝓝[≠] z₀, f z ≠ 0 := +begin + rw [eventually_nhds_within_iff], + have h2 := (has_fpower_series_iterate_dslope_fslope p.order hp).continuous_at, + have h3 := h2.eventually_ne (iterate_dslope_fslope_ne_zero hp h), + filter_upwards [eq_pow_order_mul_iterate_dslope hp, h3] with z e1 e2 e3, + simpa [e1, e2, e3] using pow_ne_zero p.order (sub_ne_zero.mpr e3), +end + +lemma locally_zero_iff (hp : has_fpower_series_at f p z₀) : + (∀ᶠ z in 𝓝 z₀, f z = 0) ↔ p = 0 := +⟨λ hf, hp.eq_zero_of_eventually hf, λ h, eventually_eq_zero (by rwa h at hp)⟩ + +end has_fpower_series_at + +namespace analytic_at + +/-- The *principle of isolated zeros* for an analytic function, local version: if a function is +analytic at `z₀`, then either it is identically zero in a neighborhood of `z₀`, or it does not +vanish in a punctured neighborhood of `z₀`. -/ +theorem eventually_eq_zero_or_eventually_ne_zero (hf : analytic_at 𝕜 f z₀) : + (∀ᶠ z in 𝓝 z₀, f z = 0) ∨ (∀ᶠ z in 𝓝[≠] z₀, f z ≠ 0) := +begin + rcases hf with ⟨p, hp⟩, + by_cases h : p = 0, + { exact or.inl (has_fpower_series_at.eventually_eq_zero (by rwa h at hp)) }, + { exact or.inr (hp.locally_ne_zero h) } +end + +end analytic_at diff --git a/src/analysis/calculus/formal_multilinear_series.lean b/src/analysis/calculus/formal_multilinear_series.lean index d08591e5e5873..bcbd3eb378590 100644 --- a/src/analysis/calculus/formal_multilinear_series.lean +++ b/src/analysis/calculus/formal_multilinear_series.lean @@ -67,6 +67,12 @@ end module namespace formal_multilinear_series +protected lemma ext_iff {p q : formal_multilinear_series 𝕜 E F} : p = q ↔ ∀ n, p n = q n := +function.funext_iff + +protected lemma ne_iff {p q : formal_multilinear_series 𝕜 E F} : p ≠ q ↔ ∃ n, p n ≠ q n := +function.ne_iff + /-- Killing the zeroth coefficient in a formal multilinear series -/ def remove_zero (p : formal_multilinear_series 𝕜 E F) : formal_multilinear_series 𝕜 E F | 0 := 0 @@ -164,3 +170,122 @@ lemma comp_formal_multilinear_series_apply' rfl end continuous_linear_map + +namespace formal_multilinear_series + +section order + +variables [comm_ring 𝕜] {n : ℕ} + [add_comm_group E] [module 𝕜 E] [topological_space E] [topological_add_group E] + [has_continuous_const_smul 𝕜 E] + [add_comm_group F] [module 𝕜 F] [topological_space F] [topological_add_group F] + [has_continuous_const_smul 𝕜 F] + {p : formal_multilinear_series 𝕜 E F} + +/-- The index of the first non-zero coefficient in `p` (or `0` if all coefficients are zero). This + is the order of the isolated zero of an analytic function `f` at a point if `p` is the Taylor + series of `f` at that point. -/ +noncomputable def order (p : formal_multilinear_series 𝕜 E F) : ℕ := +Inf { n | p n ≠ 0 } + +@[simp] lemma order_zero : (0 : formal_multilinear_series 𝕜 E F).order = 0 := by simp [order] + +lemma ne_zero_of_order_ne_zero (hp : p.order ≠ 0) : p ≠ 0 := +λ h, by simpa [h] using hp + +lemma order_eq_find [decidable_pred (λ n, p n ≠ 0)] (hp : ∃ n, p n ≠ 0) : + p.order = nat.find hp := +by simp [order, Inf, hp] + +lemma order_eq_find' [decidable_pred (λ n, p n ≠ 0)] (hp : p ≠ 0) : + p.order = nat.find (formal_multilinear_series.ne_iff.mp hp) := +order_eq_find _ + +lemma order_eq_zero_iff (hp : p ≠ 0) : p.order = 0 ↔ p 0 ≠ 0 := +begin + classical, + have : ∃ n, p n ≠ 0 := formal_multilinear_series.ne_iff.mp hp, + simp [order_eq_find this, hp] +end + +lemma order_eq_zero_iff' : p.order = 0 ↔ p = 0 ∨ p 0 ≠ 0 := +by { by_cases h : p = 0; simp [h, order_eq_zero_iff] } + +lemma apply_order_ne_zero (hp : p ≠ 0) : p p.order ≠ 0 := +begin + classical, + let h := formal_multilinear_series.ne_iff.mp hp, + exact (order_eq_find h).symm ▸ nat.find_spec h +end + +lemma apply_order_ne_zero' (hp : p.order ≠ 0) : p p.order ≠ 0 := +apply_order_ne_zero (ne_zero_of_order_ne_zero hp) + +lemma apply_eq_zero_of_lt_order (hp : n < p.order) : p n = 0 := +begin + by_cases p = 0, + { simp [h] }, + { classical, + rw [order_eq_find' h] at hp, + simpa using nat.find_min _ hp } +end + +end order + +section coef + +variables [nontrivially_normed_field 𝕜] + [normed_add_comm_group E] [normed_space 𝕜 E] {s : E} + {p : formal_multilinear_series 𝕜 𝕜 E} {f : 𝕜 → E} + {n : ℕ} {z z₀ : 𝕜} {y : fin n → 𝕜} + +open_locale big_operators + +/-- The `n`th coefficient of `p` when seen as a power series. -/ +def coeff (p : formal_multilinear_series 𝕜 𝕜 E) (n : ℕ) : E := p n 1 + +lemma mk_pi_field_coeff_eq (p : formal_multilinear_series 𝕜 𝕜 E) (n : ℕ) : + continuous_multilinear_map.mk_pi_field 𝕜 (fin n) (p.coeff n) = p n := +(p n).mk_pi_field_apply_one_eq_self + +@[simp] lemma apply_eq_prod_smul_coeff : p n y = (∏ i, y i) • p.coeff n := +begin + convert (p n).to_multilinear_map.map_smul_univ y 1, + funext; simp only [pi.one_apply, algebra.id.smul_eq_mul, mul_one], +end + +lemma coeff_eq_zero : p.coeff n = 0 ↔ p n = 0 := +by rw [← mk_pi_field_coeff_eq p, continuous_multilinear_map.mk_pi_field_eq_zero_iff] + +@[simp] lemma apply_eq_pow_smul_coeff : p n (λ _, z) = z ^ n • p.coeff n := +by simp + +@[simp] lemma norm_apply_eq_norm_coef : ∥p n∥ = ∥coeff p n∥ := +by rw [← mk_pi_field_coeff_eq p, continuous_multilinear_map.norm_mk_pi_field] + +end coef + +section fslope + +variables [nontrivially_normed_field 𝕜] + [normed_add_comm_group E] [normed_space 𝕜 E] + {p : formal_multilinear_series 𝕜 𝕜 E} {n : ℕ} + +/-- The formal counterpart of `dslope`, corresponding to the expansion of `(f z - f 0) / z`. If `f` +has `p` as a power series, then `dslope f` has `fslope p` as a power series. -/ +noncomputable def fslope (p : formal_multilinear_series 𝕜 𝕜 E) : formal_multilinear_series 𝕜 𝕜 E := + λ n, (p (n + 1)).curry_left 1 + +@[simp] lemma coeff_fslope : p.fslope.coeff n = p.coeff (n + 1) := +begin + have : @fin.cons n (λ _, 𝕜) 1 (1 : fin n → 𝕜) = 1 := fin.cons_self_tail 1, + simp only [fslope, coeff, continuous_multilinear_map.curry_left_apply, this], +end + +@[simp] lemma coeff_iterate_fslope (k n : ℕ) : + (fslope^[k] p).coeff n = p.coeff (n + k) := +by induction k with k ih generalizing p; refl <|> simpa [ih] + +end fslope + +end formal_multilinear_series diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index cc713ff94c7d3..4415d8327d832 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -789,6 +789,22 @@ to_multilinear_map_inj f.to_multilinear_map.mk_pi_ring_apply_one_eq_self (multilinear_map.mk_continuous_norm_le _ (norm_nonneg z) _).antisymm $ by simpa using (continuous_multilinear_map.mk_pi_field 𝕜 ι z).le_op_norm (λ _, 1) +lemma mk_pi_field_eq_iff {z₁ z₂ : G} : + continuous_multilinear_map.mk_pi_field 𝕜 ι z₁ = continuous_multilinear_map.mk_pi_field 𝕜 ι z₂ ↔ + z₁ = z₂ := +begin + rw [← to_multilinear_map_inj.eq_iff], + exact multilinear_map.mk_pi_ring_eq_iff +end + +lemma mk_pi_field_zero : + continuous_multilinear_map.mk_pi_field 𝕜 ι (0 : G) = 0 := +by ext; rw [mk_pi_field_apply, smul_zero, continuous_multilinear_map.zero_apply] + +lemma mk_pi_field_eq_zero_iff (z : G) : + continuous_multilinear_map.mk_pi_field 𝕜 ι z = 0 ↔ z = 0 := +by rw [← mk_pi_field_zero, mk_pi_field_eq_iff] + variables (𝕜 ι G) /-- Continuous multilinear maps on `𝕜^n` with values in `G` are in bijection with `G`, as such a diff --git a/src/linear_algebra/multilinear/basic.lean b/src/linear_algebra/multilinear/basic.lean index 409b7078e5098..9f0501119f9af 100644 --- a/src/linear_algebra/multilinear/basic.lean +++ b/src/linear_algebra/multilinear/basic.lean @@ -878,6 +878,23 @@ begin refl end +lemma mk_pi_ring_eq_iff [fintype ι] {z₁ z₂ : M₂} : + multilinear_map.mk_pi_ring R ι z₁ = multilinear_map.mk_pi_ring R ι z₂ ↔ z₁ = z₂ := +begin + simp_rw [multilinear_map.ext_iff, mk_pi_ring_apply], + split; intro h, + { simpa using h (λ _, 1) }, + { intro x, simp [h] } +end + +lemma mk_pi_ring_zero [fintype ι] : + multilinear_map.mk_pi_ring R ι (0 : M₂) = 0 := +by ext; rw [mk_pi_ring_apply, smul_zero, multilinear_map.zero_apply] + +lemma mk_pi_ring_eq_zero_iff [fintype ι] (z : M₂) : + multilinear_map.mk_pi_ring R ι z = 0 ↔ z = 0 := +by rw [← mk_pi_ring_zero, mk_pi_ring_eq_iff] + end comm_semiring section range_add_comm_group diff --git a/src/topology/algebra/module/multilinear.lean b/src/topology/algebra/module/multilinear.lean index adf47d257ac27..8bdb55433e7a6 100644 --- a/src/topology/algebra/module/multilinear.lean +++ b/src/topology/algebra/module/multilinear.lean @@ -85,6 +85,9 @@ theorem to_multilinear_map_inj : @[ext] theorem ext {f f' : continuous_multilinear_map R M₁ M₂} (H : ∀ x, f x = f' x) : f = f' := to_multilinear_map_inj $ multilinear_map.ext H +theorem ext_iff {f f' : continuous_multilinear_map R M₁ M₂} : f = f' ↔ ∀ x, f x = f' x := +by rw [← to_multilinear_map_inj.eq_iff, multilinear_map.ext_iff]; refl + @[simp] lemma map_add (m : Πi, M₁ i) (i : ι) (x y : M₁ i) : f (update m i (x + y)) = f (update m i x) + f (update m i y) := f.map_add' m i x y diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 2ef12a338bee0..43e1623478a1c 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -605,6 +605,16 @@ have fact₁ : {f a}ᶜ ∈ 𝓝 b := compl_singleton_mem_nhds hfa.symm, have fact₂ : tendsto f (pure a) (𝓝 b) := h.comp (tendsto_id'.2 $ pure_le_nhds a), fact₂ fact₁ (eq.refl $ f a) +lemma filter.tendsto.eventually_ne [topological_space β] [t1_space β] {α : Type*} {g : α → β} + {l : filter α} {b₁ b₂ : β} (hg : tendsto g l (𝓝 b₁)) (hb : b₁ ≠ b₂) : + ∀ᶠ z in l, g z ≠ b₂ := +hg.eventually (is_open_compl_singleton.eventually_mem hb) + +lemma continuous_at.eventually_ne [topological_space β] [t1_space β] {g : α → β} + {a : α} {b : β} (hg1 : continuous_at g a) (hg2 : g a ≠ b) : + ∀ᶠ z in 𝓝 a, g z ≠ b := +hg1.tendsto.eventually_ne hg2 + /-- To prove a function to a `t1_space` is continuous at some point `a`, it suffices to prove that `f` admits *some* limit at `a`. -/ lemma continuous_at_of_tendsto_nhds [topological_space β] [t1_space β] {f : α → β} {a : α} {b : β} From f46fcef914754ad20285b06a2ed4edf93e50dbd6 Mon Sep 17 00:00:00 2001 From: Alexander Bentkamp Date: Sun, 4 Sep 2022 02:01:26 +0000 Subject: [PATCH 150/828] feat(matrix/spectrum): The determinant of a hermitian matrix is the product of its eigenvalues. (#16293) --- src/linear_algebra/matrix/spectrum.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/linear_algebra/matrix/spectrum.lean b/src/linear_algebra/matrix/spectrum.lean index 2029a46c5de7b..72e882caee369 100644 --- a/src/linear_algebra/matrix/spectrum.lean +++ b/src/linear_algebra/matrix/spectrum.lean @@ -23,6 +23,7 @@ variables {𝕜 : Type*} [is_R_or_C 𝕜] [decidable_eq 𝕜] {n : Type*} [finty variables {A : matrix n n 𝕜} open_locale matrix +open_locale big_operators namespace is_hermitian @@ -80,6 +81,13 @@ begin euclidean_space.single, pi_Lp.equiv_symm_apply'] } end +/-- The determinant of a hermitian matrix is the product of its eigenvalues. -/ +lemma det_eq_prod_eigenvalues : det A = ∏ i, hA.eigenvalues i := +begin + apply mul_left_cancel₀ (det_ne_zero_of_left_inverse (eigenvector_matrix_mul_inv hA)), + rw [←det_mul, spectral_theorem, det_mul, mul_comm, det_diagonal] +end + end is_hermitian end matrix From 443c02be8960e30f4956273aab336a5eb47b0584 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Sun, 4 Sep 2022 02:50:31 +0000 Subject: [PATCH 151/828] feat(group_theory/finiteness): Add `rank_le_of_surjective`, `rank_range_le`, and `rank_congr` (#16363) This PR adds some more API lemmas for `group.rank`. --- src/group_theory/finiteness.lean | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/src/group_theory/finiteness.lean b/src/group_theory/finiteness.lean index 1ff57337b4b2f..0e829a45dadf2 100644 --- a/src/group_theory/finiteness.lean +++ b/src/group_theory/finiteness.lean @@ -293,6 +293,28 @@ noncomputable def group.rank [h : group.fg G] := {S : finset G} (hS : subgroup.closure (S : set G) = ⊤) : group.rank G ≤ S.card := @nat.find_le _ _ (classical.dec_pred _) (group.fg_iff'.mp h) ⟨S, rfl, hS⟩ +variables {G} {G' : Type*} [group G'] + +@[to_additive] lemma group.rank_le_of_surjective [group.fg G] [group.fg G'] (f : G →* G') + (hf : function.surjective f) : group.rank G' ≤ group.rank G := +begin + classical, + obtain ⟨S, hS1, hS2⟩ := group.rank_spec G, + transitivity (S.image f).card, + { apply group.rank_le, + rw [finset.coe_image, ←monoid_hom.map_closure, hS2, subgroup.map_top_of_surjective f hf] }, + { exact finset.card_image_le.trans_eq hS1 }, +end + +@[to_additive] lemma group.rank_range_le [group.fg G] {f : G →* G'} : + group.rank f.range ≤ group.rank G := +group.rank_le_of_surjective f.range_restrict f.range_restrict_surjective + +@[to_additive] lemma group.rank_congr [group.fg G] [group.fg G'] (f : G ≃* G') : + group.rank G = group.rank G' := +le_antisymm (group.rank_le_of_surjective f.symm f.symm.surjective) + (group.rank_le_of_surjective f f.surjective) + end group section quotient_group From 1cfdf5f34e1044ecb65d10be753008baaf118edf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 4 Sep 2022 04:40:23 +0000 Subject: [PATCH 152/828] chore(linear_algebra/*): Lint (#16362) Satisfy the `fintype_finite` and `fails_quickly` linters. --- src/geometry/euclidean/basic.lean | 2 + src/geometry/euclidean/circumcenter.lean | 2 + .../affine_space/affine_subspace.lean | 6 +- .../affine_space/finite_dimensional.lean | 8 +-- src/linear_algebra/alternating.lean | 2 +- src/linear_algebra/determinant.lean | 16 +++-- src/linear_algebra/dimension.lean | 6 +- src/linear_algebra/dual.lean | 70 +++++++++---------- src/linear_algebra/finsupp_vector_space.lean | 3 +- .../free_module/finite/basic.lean | 14 ++-- .../free_module/finite/rank.lean | 2 +- src/linear_algebra/free_module/pid.lean | 4 +- src/linear_algebra/free_module/rank.lean | 13 ++-- .../matrix/finite_dimensional.lean | 7 +- src/linear_algebra/matrix/reindex.lean | 4 +- src/linear_algebra/matrix/transvection.lean | 6 +- src/linear_algebra/multilinear/basis.lean | 6 +- .../multilinear/finite_dimensional.lean | 7 +- src/linear_algebra/orientation.lean | 23 +++--- src/linear_algebra/std_basis.lean | 11 ++- src/linear_algebra/trace.lean | 7 +- 21 files changed, 118 insertions(+), 101 deletions(-) diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index d4e407a43b7b5..dc331ea86724b 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -783,6 +783,8 @@ lemma orthogonal_projection_fn_vsub_mem_direction_orthogonal {s : affine_subspac direction_mk' p s.directionᗮ ▸ vsub_mem_direction (orthogonal_projection_fn_mem_orthogonal p) (self_mem_mk' _ _) +local attribute [instance] affine_subspace.to_add_torsor + /-- The orthogonal projection of a point onto a nonempty affine subspace, whose direction is complete. The corresponding linear map (mapping a vector to the difference between the projections of two diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index 95a77aa6e6ce3..69ada84e2f8bf 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -377,6 +377,8 @@ begin (λ i, hr i (set.mem_univ _))).symm end +local attribute [instance] affine_subspace.to_add_torsor + /-- The orthogonal projection of a point `p` onto the hyperplane spanned by the simplex's points. -/ def orthogonal_projection_span {n : ℕ} (s : simplex ℝ P n) : P →ᵃ[ℝ] affine_span ℝ (set.range s.points) := diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index bad91a54e2b2c..4d444adc745bf 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -376,7 +376,9 @@ begin exact vsub_mem_direction hp hq2 } end -instance to_add_torsor (s : affine_subspace k P) [nonempty s] : add_torsor s.direction s := +/-- This is not an instance because it loops with `add_torsor.nonempty`. -/ +@[reducible] -- See note [reducible non instances] +def to_add_torsor (s : affine_subspace k P) [nonempty s] : add_torsor s.direction s := { vadd := λ a b, ⟨(a:V) +ᵥ (b:P), vadd_mem_of_mem_direction a.2 b.2⟩, zero_vadd := by simp, add_vadd := λ a b c, by { ext, apply add_vadd }, @@ -385,6 +387,8 @@ instance to_add_torsor (s : affine_subspace k P) [nonempty s] : add_torsor s.dir vsub_vadd' := λ a b, by { ext, apply add_torsor.vsub_vadd' }, vadd_vsub' := λ a b, by { ext, apply add_torsor.vadd_vsub' } } +local attribute [instance] to_add_torsor + @[simp, norm_cast] lemma coe_vsub (s : affine_subspace k P) [nonempty s] (a b : s) : ↑(a -ᵥ b) = (a:P) -ᵥ (b:P) := rfl diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index d793821c4ce5f..e9c203dd691a4 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -39,13 +39,13 @@ span_of_finite k $ h.vsub h /-- The `vector_span` of a family indexed by a `fintype` is finite-dimensional. -/ -instance finite_dimensional_vector_span_of_fintype [fintype ι] (p : ι → P) : +instance finite_dimensional_vector_span_range [_root_.finite ι] (p : ι → P) : finite_dimensional k (vector_span k (set.range p)) := finite_dimensional_vector_span_of_finite k (set.finite_range _) /-- The `vector_span` of a subset of a family indexed by a `fintype` is finite-dimensional. -/ -instance finite_dimensional_vector_span_image_of_fintype [fintype ι] (p : ι → P) +instance finite_dimensional_vector_span_image_of_finite [_root_.finite ι] (p : ι → P) (s : set ι) : finite_dimensional k (vector_span k (p '' s)) := finite_dimensional_vector_span_of_finite k (set.to_finite _) @@ -57,13 +57,13 @@ lemma finite_dimensional_direction_affine_span_of_finite {s : set P} (h : set.fi /-- The direction of the affine span of a family indexed by a `fintype` is finite-dimensional. -/ -instance finite_dimensional_direction_affine_span_of_fintype [fintype ι] (p : ι → P) : +instance finite_dimensional_direction_affine_span_range [_root_.finite ι] (p : ι → P) : finite_dimensional k (affine_span k (set.range p)).direction := finite_dimensional_direction_affine_span_of_finite k (set.finite_range _) /-- The direction of the affine span of a subset of a family indexed by a `fintype` is finite-dimensional. -/ -instance finite_dimensional_direction_affine_span_image_of_fintype [fintype ι] (p : ι → P) +instance finite_dimensional_direction_affine_span_image_of_finite [_root_.finite ι] (p : ι → P) (s : set ι) : finite_dimensional k (affine_span k (p '' s)).direction := finite_dimensional_direction_affine_span_of_finite k (set.to_finite _) diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index 62ac65c6d4027..9a9988761210b 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -956,7 +956,7 @@ section basis open alternating_map -variables {ι₁ : Type*} [fintype ι] +variables {ι₁ : Type*} [finite ι] variables {R' : Type*} {N₁ N₂ : Type*} [comm_semiring R'] [add_comm_monoid N₁] [add_comm_monoid N₂] variables [module R' N₁] [module R' N₂] diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index c4ef584883b0e..2d4b4870ca51d 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -246,11 +246,10 @@ begin simp [coe_det, H, this] } end -lemma det_zero' {ι : Type*} [fintype ι] [nonempty ι] (b : basis ι A M) : +lemma det_zero' {ι : Type*} [finite ι] [nonempty ι] (b : basis ι A M) : linear_map.det (0 : M →ₗ[A] M) = 0 := -by { haveI := classical.dec_eq ι, - rw [← det_to_matrix b, linear_equiv.map_zero, det_zero], - assumption } +by { haveI := classical.dec_eq ι, casesI nonempty_fintype ι, + rwa [← det_to_matrix b, linear_equiv.map_zero, det_zero] } /-- In a finite-dimensional vector space, the zero map has determinant `1` in dimension `0`, and `0` otherwise. We give a formula that also works in infinite dimension, where we define @@ -491,11 +490,14 @@ begin simp [alternating_map.map_perm, basis.det_self] end -@[simp] lemma alternating_map.map_basis_eq_zero_iff (f : alternating_map R M R ι) : +@[simp] lemma alternating_map.map_basis_eq_zero_iff {ι : Type*} [decidable_eq ι] [finite ι] + (e : basis ι R M) (f : alternating_map R M R ι) : f e = 0 ↔ f = 0 := -⟨λ h, by simpa [h] using f.eq_smul_basis_det e, λ h, h.symm ▸ alternating_map.zero_apply _⟩ +⟨λ h, by { casesI nonempty_fintype ι, simpa [h] using f.eq_smul_basis_det e }, + λ h, h.symm ▸ alternating_map.zero_apply _⟩ -lemma alternating_map.map_basis_ne_zero_iff (f : alternating_map R M R ι) : +lemma alternating_map.map_basis_ne_zero_iff {ι : Type*} [decidable_eq ι] [finite ι] + (e : basis ι R M) (f : alternating_map R M R ι) : f e ≠ 0 ↔ f ≠ 0 := not_congr $ f.map_basis_eq_zero_iff e diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index 551ad63201b6a..e262a9055ccba 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -988,19 +988,21 @@ begin end section fintype -variable [fintype η] variables [∀i, add_comm_group (φ i)] [∀i, module K (φ i)] open linear_map -lemma dim_pi : module.rank K (Πi, φ i) = cardinal.sum (λi, module.rank K (φ i)) := +lemma dim_pi [finite η] : module.rank K (Πi, φ i) = cardinal.sum (λi, module.rank K (φ i)) := begin + casesI nonempty_fintype η, let b := assume i, basis.of_vector_space K (φ i), let this : basis (Σ j, _) K (Π j, φ j) := pi.basis b, rw [← cardinal.lift_inj, ← this.mk_eq_dim], simp [← (b _).mk_range_eq_dim] end +variable [fintype η] + lemma dim_fun {V η : Type u} [fintype η] [add_comm_group V] [module K V] : module.rank K (η → V) = fintype.card η * module.rank K V := by rw [dim_pi, cardinal.sum_const', cardinal.mk_fintype] diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 854ecd539f682..ade8d80e9d4fc 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -171,19 +171,15 @@ end theorem to_dual_ker : b.to_dual.ker = ⊥ := ker_eq_bot'.mpr b.to_dual_inj -theorem to_dual_range [fin : fintype ι] : b.to_dual.range = ⊤ := +theorem to_dual_range [_root_.finite ι] : b.to_dual.range = ⊤ := begin - rw eq_top_iff', - intro f, + casesI nonempty_fintype ι, + refine eq_top_iff'.2 (λ f, _), rw linear_map.mem_range, - let lin_comb : ι →₀ R := finsupp.on_finset fin.elems (λ i, f.to_fun (b i)) _, - { use finsupp.total ι M R b lin_comb, - apply b.ext, - { intros i, - rw [b.to_dual_eq_repr _ i, repr_total b], - { refl } } }, - { intros a _, - apply fin.complete } + let lin_comb : ι →₀ R := finsupp.equiv_fun_on_fintype.2 (λ i, f.to_fun (b i)), + refine ⟨finsupp.total ι M R b lin_comb, b.ext $ λ i, _⟩, + rw [b.to_dual_eq_repr _ i, repr_total b], + refl, end end comm_semiring @@ -207,53 +203,52 @@ section comm_ring variables [comm_ring R] [add_comm_group M] [module R M] [decidable_eq ι] variables (b : basis ι R M) +section finite +variables [_root_.finite ι] + /-- A vector space is linearly equivalent to its dual space. -/ @[simps] -def to_dual_equiv [fintype ι] : M ≃ₗ[R] (dual R M) := +def to_dual_equiv : M ≃ₗ[R] dual R M := linear_equiv.of_bijective b.to_dual (ker_eq_bot.mp b.to_dual_ker) (range_eq_top.mp b.to_dual_range) /-- Maps a basis for `V` to a basis for the dual space. -/ -def dual_basis [fintype ι] : basis ι R (dual R M) := -b.map b.to_dual_equiv +def dual_basis : basis ι R (dual R M) := b.map b.to_dual_equiv -- We use `j = i` to match `basis.repr_self` -lemma dual_basis_apply_self [fintype ι] (i j : ι) : - b.dual_basis i (b j) = if j = i then 1 else 0 := +lemma dual_basis_apply_self (i j : ι) : b.dual_basis i (b j) = if j = i then 1 else 0 := by { convert b.to_dual_apply i j using 2, rw @eq_comm _ j i } -lemma total_dual_basis [fintype ι] (f : ι →₀ R) (i : ι) : +lemma total_dual_basis (f : ι →₀ R) (i : ι) : finsupp.total ι (dual R M) R b.dual_basis f (b i) = f i := begin + casesI nonempty_fintype ι, rw [finsupp.total_apply, finsupp.sum_fintype, linear_map.sum_apply], { simp_rw [linear_map.smul_apply, smul_eq_mul, dual_basis_apply_self, mul_boole, finset.sum_ite_eq, if_pos (finset.mem_univ i)] }, { intro, rw zero_smul }, end -lemma dual_basis_repr [fintype ι] (l : dual R M) (i : ι) : - b.dual_basis.repr l i = l (b i) := +lemma dual_basis_repr (l : dual R M) (i : ι) : b.dual_basis.repr l i = l (b i) := by rw [← total_dual_basis b, basis.total_repr b.dual_basis l] -lemma dual_basis_equiv_fun [fintype ι] (l : dual R M) (i : ι) : - b.dual_basis.equiv_fun l i = l (b i) := -by rw [basis.equiv_fun_apply, dual_basis_repr] - -lemma dual_basis_apply [fintype ι] (i : ι) (m : M) : b.dual_basis i m = b.repr m i := -b.to_dual_apply_right i m +lemma dual_basis_apply (i : ι) (m : M) : b.dual_basis i m = b.repr m i := b.to_dual_apply_right i m -@[simp] lemma coe_dual_basis [fintype ι] : - ⇑b.dual_basis = b.coord := -by { ext i x, apply dual_basis_apply } +@[simp] lemma coe_dual_basis : ⇑b.dual_basis = b.coord := by { ext i x, apply dual_basis_apply } -@[simp] lemma to_dual_to_dual [fintype ι] : - b.dual_basis.to_dual.comp b.to_dual = dual.eval R M := +@[simp] lemma to_dual_to_dual : b.dual_basis.to_dual.comp b.to_dual = dual.eval R M := begin refine b.ext (λ i, b.dual_basis.ext (λ j, _)), rw [linear_map.comp_apply, to_dual_apply_left, coe_to_dual_self, ← coe_dual_basis, dual.eval_apply, basis.repr_self, finsupp.single_apply, dual_basis_apply_self] end +end finite + +lemma dual_basis_equiv_fun [fintype ι] (l : dual R M) (i : ι) : + b.dual_basis.equiv_fun l i = l (b i) := +by rw [basis.equiv_fun_apply, dual_basis_repr] + theorem eval_ker {ι : Type*} (b : basis ι R M) : (dual.eval R M).ker = ⊥ := begin @@ -263,20 +258,20 @@ begin exact (basis.forall_coord_eq_zero_iff _).mp (λ i, hm (b.coord i)) end -lemma eval_range {ι : Type*} [fintype ι] (b : basis ι R M) : - (eval R M).range = ⊤ := +lemma eval_range {ι : Type*} [_root_.finite ι] (b : basis ι R M) : (eval R M).range = ⊤ := begin classical, + casesI nonempty_fintype ι, rw [← b.to_dual_to_dual, range_comp, b.to_dual_range, map_top, to_dual_range _], apply_instance end /-- A module with a basis is linearly equivalent to the dual of its dual space. -/ -def eval_equiv {ι : Type*} [fintype ι] (b : basis ι R M) : M ≃ₗ[R] dual R (dual R M) := +def eval_equiv {ι : Type*} [_root_.finite ι] (b : basis ι R M) : M ≃ₗ[R] dual R (dual R M) := linear_equiv.of_bijective (eval R M) (ker_eq_bot.mp b.eval_ker) (range_eq_top.mp b.eval_range) -@[simp] lemma eval_equiv_to_linear_map {ι : Type*} [fintype ι] (b : basis ι R M) : +@[simp] lemma eval_equiv_to_linear_map {ι : Type*} [_root_.finite ι] (b : basis ι R M) : (b.eval_equiv).to_linear_map = dual.eval R M := rfl section @@ -294,16 +289,17 @@ end end comm_ring /-- `simp` normal form version of `total_dual_basis` -/ -@[simp] lemma total_coord [comm_ring R] [add_comm_group M] [module R M] [fintype ι] +@[simp] lemma total_coord [comm_ring R] [add_comm_group M] [module R M] [_root_.finite ι] (b : basis ι R M) (f : ι →₀ R) (i : ι) : finsupp.total ι (dual R M) R b.coord f (b i) = f i := by { haveI := classical.dec_eq ι, rw [← coe_dual_basis, total_dual_basis] } --- TODO(jmc): generalize to rings, once `module.rank` is generalized -theorem dual_dim_eq [field K] [add_comm_group V] [module K V] [fintype ι] (b : basis ι K V) : +lemma dual_dim_eq [comm_ring K] [add_comm_group V] [module K V] [_root_.finite ι] + (b : basis ι K V) : cardinal.lift (module.rank K V) = module.rank K (dual K V) := begin classical, + casesI nonempty_fintype ι, have := linear_equiv.lift_dim_eq b.to_dual_equiv, simp only [cardinal.lift_umax] at this, rw [this, ← cardinal.lift_umax], diff --git a/src/linear_algebra/finsupp_vector_space.lean b/src/linear_algebra/finsupp_vector_space.lean index efc5670b9683f..b65a171d4ed5e 100644 --- a/src/linear_algebra/finsupp_vector_space.lean +++ b/src/linear_algebra/finsupp_vector_space.lean @@ -200,7 +200,8 @@ begin ... = _ : by rw [← cardinal.lift_inj.1 hs.mk_eq_dim, cardinal.power_def] end -lemma cardinal_lt_aleph_0_of_finite_dimensional [fintype K] [finite_dimensional K V] : #V < ℵ₀ := +lemma cardinal_lt_aleph_0_of_finite_dimensional [_root_.finite K] [finite_dimensional K V] : + #V < ℵ₀ := begin letI : is_noetherian K V := is_noetherian.iff_fg.2 infer_instance, rw cardinal_mk_eq_cardinal_mk_field_pow_dim K V, diff --git a/src/linear_algebra/free_module/finite/basic.lean b/src/linear_algebra/free_module/finite/basic.lean index bf561f6dcd798..ac6fd5cb3642d 100644 --- a/src/linear_algebra/free_module/finite/basic.lean +++ b/src/linear_algebra/free_module/finite/basic.lean @@ -59,22 +59,22 @@ begin (linear_map.to_matrix (module.free.choose_basis R M) (module.free.choose_basis R N)).symm, end -variables {R M} +variables {R} /-- A free module with a basis indexed by a `fintype` is finite. -/ -lemma _root_.module.finite.of_basis {R : Type*} {M : Type*} {ι : Type*} [comm_ring R] - [add_comm_group M] [module R M] [fintype ι] (b : basis ι R M) : module.finite R M := +lemma _root_.module.finite.of_basis {R M ι : Type*} [comm_ring R] [add_comm_group M] [module R M] + [finite ι] (b : basis ι R M) : module.finite R M := begin + casesI nonempty_fintype ι, classical, refine ⟨⟨finset.univ.image b, _⟩⟩, simp only [set.image_univ, finset.coe_univ, finset.coe_image, basis.span_eq], end -instance _root_.module.finite.matrix {ι₁ : Type*} [fintype ι₁] {ι₂ : Type*} [fintype ι₂] : +instance _root_.module.finite.matrix {ι₁ ι₂ : Type*} [finite ι₁] [finite ι₂] : module.finite R (matrix ι₁ ι₂ R) := -module.finite.of_basis $ pi.basis $ λ i, pi.basis_fun R _ - -variables (M) +by { casesI nonempty_fintype ι₁, casesI nonempty_fintype ι₂, + exact module.finite.of_basis (pi.basis $ λ i, pi.basis_fun R _) } instance _root_.module.finite.linear_map [module.finite R M] [module.finite R N] : module.finite R (M →ₗ[R] N) := diff --git a/src/linear_algebra/free_module/finite/rank.lean b/src/linear_algebra/free_module/finite/rank.lean index 81b13a1f00898..40638dfa67547 100644 --- a/src/linear_algebra/free_module/finite/rank.lean +++ b/src/linear_algebra/free_module/finite/rank.lean @@ -83,7 +83,7 @@ lemma finrank_pi_fintype {ι : Type v} [fintype ι] {M : ι → Type w} [Π (i : ι), module.finite R (M i)] : finrank R (Π i, M i) = ∑ i, finrank R (M i) := begin letI := nontrivial_of_invariant_basis_number R, - simp only [finrank, λ i, rank_eq_card_choose_basis_index R (M i), rank_pi_fintype, + simp only [finrank, λ i, rank_eq_card_choose_basis_index R (M i), rank_pi_finite, ← mk_sigma, mk_to_nat_eq_card, card_sigma], end diff --git a/src/linear_algebra/free_module/pid.lean b/src/linear_algebra/free_module/pid.lean index 7d0654dcdb2d4..597e4628030df 100644 --- a/src/linear_algebra/free_module/pid.lean +++ b/src/linear_algebra/free_module/pid.lean @@ -514,11 +514,11 @@ See also `ideal.smith_normal_form` for a version of this theorem that returns a `basis.smith_normal_form`. -/ theorem ideal.exists_smith_normal_form - [fintype ι] {S : Type*} [comm_ring S] [is_domain S] [algebra R S] + [finite ι] {S : Type*} [comm_ring S] [is_domain S] [algebra R S] (b : basis ι R S) (I : ideal S) (hI : I ≠ ⊥) : ∃ (b' : basis ι R S) (a : ι → R) (ab' : basis ι R I), ∀ i, (ab' i : S) = a i • b' i := -let ⟨bS, bI, f, a, snf⟩ := I.smith_normal_form b hI, + by casesI nonempty_fintype ι; exact let ⟨bS, bI, f, a, snf⟩ := I.smith_normal_form b hI, e : fin (fintype.card ι) ≃ ι := equiv.of_bijective f ((fintype.bijective_iff_injective_and_card f).mpr ⟨f.injective, fintype.card_fin _⟩) in have fe : ∀ i, f (e.symm i) = i := e.apply_symm_apply, diff --git a/src/linear_algebra/free_module/rank.lean b/src/linear_algebra/free_module/rank.lean index f4457a08efbbe..8527af1eb7879 100644 --- a/src/linear_algebra/free_module/rank.lean +++ b/src/linear_algebra/free_module/rank.lean @@ -65,15 +65,18 @@ begin end /-- The rank of a finite product is the sum of the ranks. -/ -@[simp] lemma rank_pi_fintype {ι : Type v} [fintype ι] {M : ι → Type w} +@[simp] lemma rank_pi_finite {ι : Type v} [finite ι] {M : ι → Type w} [Π (i : ι), add_comm_group (M i)] [Π (i : ι), module R (M i)] [Π (i : ι), module.free R (M i)] : module.rank R (Π i, M i) = cardinal.sum (λ i, module.rank R (M i)) := -by { rw [← (direct_sum.linear_equiv_fun_on_fintype _ _ M).dim_eq, rank_direct_sum] } +by { casesI nonempty_fintype ι, + rw [←(direct_sum.linear_equiv_fun_on_fintype _ _ M).dim_eq, rank_direct_sum] } /-- If `m` and `n` are `fintype`, the rank of `m × n` matrices is `(# m).lift * (# n).lift`. -/ -@[simp] lemma rank_matrix (m : Type v) (n : Type w) [fintype m] [fintype n] : +@[simp] lemma rank_matrix (m : Type v) (n : Type w) [finite m] [finite n] : module.rank R (matrix m n R) = (lift.{(max v w u) v} (# m)) * (lift.{(max v w u) w} (# n)) := begin + casesI nonempty_fintype m, + casesI nonempty_fintype n, have h := (matrix.std_basis R m n).mk_eq_dim, rw [← lift_lift.{(max v w u) (max v w)}, lift_inj] at h, simpa using h.symm, @@ -81,13 +84,13 @@ end /-- If `m` and `n` are `fintype` that lie in the same universe, the rank of `m × n` matrices is `(# n * # m).lift`. -/ -@[simp] lemma rank_matrix' (m n : Type v) [fintype m] [fintype n] : +@[simp] lemma rank_matrix' (m n : Type v) [finite m] [finite n] : module.rank R (matrix m n R) = (# m * # n).lift := by rw [rank_matrix, lift_mul, lift_umax] /-- If `m` and `n` are `fintype` that lie in the same universe as `R`, the rank of `m × n` matrices is `# m * # n`. -/ -@[simp] lemma rank_matrix'' (m n : Type u) [fintype m] [fintype n] : +@[simp] lemma rank_matrix'' (m n : Type u) [finite m] [finite n] : module.rank R (matrix m n R) = # m * # n := by simp end ring diff --git a/src/linear_algebra/matrix/finite_dimensional.lean b/src/linear_algebra/matrix/finite_dimensional.lean index aaeefb583faa5..55796c37dfd9f 100644 --- a/src/linear_algebra/matrix/finite_dimensional.lean +++ b/src/linear_algebra/matrix/finite_dimensional.lean @@ -29,17 +29,16 @@ namespace matrix section finite_dimensional -variables {m n : Type*} [fintype m] [fintype n] -variables {R : Type v} [field R] +variables {m n : Type*} {R : Type v} [field R] -instance : finite_dimensional R (matrix m n R) := +instance [finite m] [finite n] : finite_dimensional R (matrix m n R) := linear_equiv.finite_dimensional (linear_equiv.curry R m n) /-- The dimension of the space of finite dimensional matrices is the product of the number of rows and columns. -/ -@[simp] lemma finrank_matrix : +@[simp] lemma finrank_matrix [fintype m] [fintype n] : finite_dimensional.finrank R (matrix m n R) = fintype.card m * fintype.card n := by rw [@linear_equiv.finrank_eq R (matrix m n R) _ _ _ _ _ _ (linear_equiv.curry R m n).symm, finite_dimensional.finrank_fintype_fun_eq_card, fintype.card_prod] diff --git a/src/linear_algebra/matrix/reindex.lean b/src/linear_algebra/matrix/reindex.lean index 56e22e46b138b..0df5c98bf91af 100644 --- a/src/linear_algebra/matrix/reindex.lean +++ b/src/linear_algebra/matrix/reindex.lean @@ -90,10 +90,10 @@ lemma reindex_linear_equiv_mul [fintype n] [fintype n'] reindex_linear_equiv R A eₘ eₒ (M ⬝ N) := submatrix_mul_equiv M N _ _ _ -lemma mul_reindex_linear_equiv_one [fintype n] [fintype o] [decidable_eq o] (e₁ : o ≃ n) +lemma mul_reindex_linear_equiv_one [fintype n] [decidable_eq o] (e₁ : o ≃ n) (e₂ : o ≃ n') (M : matrix m n A) : M.mul (reindex_linear_equiv R A e₁ e₂ 1) = reindex_linear_equiv R A (equiv.refl m) (e₁.symm.trans e₂) M := -mul_submatrix_one _ _ _ +by { haveI := fintype.of_equiv _ e₁.symm, exact mul_submatrix_one _ _ _ } end semiring diff --git a/src/linear_algebra/matrix/transvection.lean b/src/linear_algebra/matrix/transvection.lean index 96d969dcffe5a..a50b2eb4ab470 100644 --- a/src/linear_algebra/matrix/transvection.lean +++ b/src/linear_algebra/matrix/transvection.lean @@ -81,14 +81,14 @@ def transvection (c : R) : matrix n n R := 1 + matrix.std_basis_matrix i j c by simp [transvection] section -variable [fintype n] /-- A transvection matrix is obtained from the identity by adding `c` times the `j`-th row to the `i`-th row. -/ -lemma update_row_eq_transvection (c : R) : +lemma update_row_eq_transvection [finite n] (c : R) : update_row (1 : matrix n n R) i (((1 : matrix n n R)) i + c • (1 : matrix n n R) j) = transvection i j c := begin + casesI nonempty_fintype n, ext a b, by_cases ha : i = a, by_cases hb : j = b, { simp only [update_row, transvection, ha, hb, function.update_same, std_basis_matrix.apply_same, @@ -101,6 +101,8 @@ begin pi.smul_apply, mul_zero, false_and] }, end +variables [fintype n] + lemma transvection_mul_transvection_same (h : i ≠ j) (c d : R) : transvection i j c ⬝ transvection i j d = transvection i j (c + d) := by simp [transvection, matrix.add_mul, matrix.mul_add, h, h.symm, add_smul, add_assoc, diff --git a/src/linear_algebra/multilinear/basis.lean b/src/linear_algebra/multilinear/basis.lean index ac96a198b15ec..1b1a06cac5a33 100644 --- a/src/linear_algebra/multilinear/basis.lean +++ b/src/linear_algebra/multilinear/basis.lean @@ -51,8 +51,8 @@ end are basis vectors. Unlike `basis.ext_multilinear_fin`, this only uses a single basis; a dependently-typed version would still be true, but the proof would need a dependently-typed version of `dom_dom_congr`. -/ -lemma basis.ext_multilinear [decidable_eq ι] [fintype ι] {f g : multilinear_map R (λ i : ι, M₂) M₃} +lemma basis.ext_multilinear [decidable_eq ι] [finite ι] {f g : multilinear_map R (λ i : ι, M₂) M₃} {ι₁ : Type*} (e : basis ι₁ R M₂) (h : ∀ v : ι → ι₁, f (λ i, e (v i)) = g (λ i, e (v i))) : f = g := -(dom_dom_congr_eq_iff (fintype.equiv_fin ι) f g).mp $ - basis.ext_multilinear_fin (λ i, e) (λ i, h (i ∘ _)) +by { casesI nonempty_fintype ι, exact (dom_dom_congr_eq_iff (fintype.equiv_fin ι) f g).mp + (basis.ext_multilinear_fin (λ i, e) $ λ i, h (i ∘ _)) } diff --git a/src/linear_algebra/multilinear/finite_dimensional.lean b/src/linear_algebra/multilinear/finite_dimensional.lean index 9aae60e8351c1..de26b7441d36e 100644 --- a/src/linear_algebra/multilinear/finite_dimensional.lean +++ b/src/linear_algebra/multilinear/finite_dimensional.lean @@ -21,8 +21,8 @@ there. namespace multilinear_map variables {ι R M₂ : Type*} {M₁ : ι → Type*} -variables [decidable_eq ι] -variables [fintype ι] [comm_ring R] [add_comm_group M₂] [module R M₂] +variables [decidable_eq ι] [finite ι] +variables [comm_ring R] [add_comm_group M₂] [module R M₂] variables [Π i, add_comm_group (M₁ i)] [Π i, module R (M₁ i)] variables [module.finite R M₂] [module.free R M₂] variables [∀ i, module.finite R (M₁ i)] [∀ i, module.free R (M₁ i)] @@ -36,7 +36,8 @@ begin by exactI ∀ [Π i, module R (N i)], by exactI ∀ [∀ i, module.finite R (N i)] [∀ i, module.free R (N i)], module.free R (multilinear_map R N M₂) ∧ module.finite R (multilinear_map R N M₂), - { casesI this _ (M₁ ∘ (fintype.equiv_fin ι).symm), + { casesI nonempty_fintype ι, + casesI this _ (M₁ ∘ (fintype.equiv_fin ι).symm), have e := dom_dom_congr_linear_equiv' R M₁ M₂ (fintype.equiv_fin ι), exact ⟨module.free.of_equiv e.symm, module.finite.equiv e.symm⟩, }, introsI n N _ _ _ _, diff --git a/src/linear_algebra/orientation.lean b/src/linear_algebra/orientation.lean index 3864278e92a7f..bb3a10db19672 100644 --- a/src/linear_algebra/orientation.lean +++ b/src/linear_algebra/orientation.lean @@ -77,21 +77,14 @@ variables {M N : Type*} [add_comm_group M] [add_comm_group N] [module R M] [modu namespace basis -variables {ι : Type*} [fintype ι] [decidable_eq ι] - -/-- The orientation given by a basis. -/ -protected def orientation [nontrivial R] (e : basis ι R M) : orientation R M ι := -ray_of_ne_zero R _ e.det_ne_zero - -lemma orientation_map [nontrivial R] (e : basis ι R M) - (f : M ≃ₗ[R] N) : (e.map f).orientation = orientation.map ι f e.orientation := -by simp_rw [basis.orientation, orientation.map_apply, basis.det_map'] +variables {ι : Type*} [decidable_eq ι] /-- The value of `orientation.map` when the index type has the cardinality of a basis, in terms of `f.det`. -/ -lemma map_orientation_eq_det_inv_smul (e : basis ι R M) +lemma map_orientation_eq_det_inv_smul [finite ι] (e : basis ι R M) (x : orientation R M ι) (f : M ≃ₗ[R] M) : orientation.map ι f x = (f.det)⁻¹ • x := begin + casesI nonempty_fintype ι, induction x using module.ray.ind with g hg, rw [orientation.map_apply, smul_ray_of_ne_zero, ray_eq_iff, units.smul_def, (g.comp_linear_map ↑f.symm).eq_smul_basis_det e, g.eq_smul_basis_det e, @@ -99,6 +92,16 @@ begin basis.det_self, mul_one, smul_eq_mul, mul_comm, mul_smul, linear_equiv.coe_inv_det], end +variables [fintype ι] + +/-- The orientation given by a basis. -/ +protected def orientation [nontrivial R] (e : basis ι R M) : orientation R M ι := +ray_of_ne_zero R _ e.det_ne_zero + +lemma orientation_map [nontrivial R] (e : basis ι R M) + (f : M ≃ₗ[R] N) : (e.map f).orientation = orientation.map ι f e.orientation := +by simp_rw [basis.orientation, orientation.map_apply, basis.det_map'] + /-- The orientation given by a basis derived using `units_smul`, in terms of the product of those units. -/ lemma orientation_units_smul [nontrivial R] (e : basis ι R M) (w : ι → units R) : diff --git a/src/linear_algebra/std_basis.lean b/src/linear_algebra/std_basis.lean index 8d341a0f2fd2f..90ee77f399777 100644 --- a/src/linear_algebra/std_basis.lean +++ b/src/linear_algebra/std_basis.lean @@ -113,13 +113,12 @@ begin exact le_rfl end -lemma supr_range_std_basis [fintype ι] : (⨆i:ι, range (std_basis R φ i)) = ⊤ := -have (set.univ : set ι) ⊆ ↑(finset.univ : finset ι) ∪ ∅ := by rw [finset.coe_univ, set.union_empty], +lemma supr_range_std_basis [finite ι] : (⨆ i, range (std_basis R φ i)) = ⊤ := begin - apply top_unique, - convert (infi_ker_proj_le_supr_range_std_basis R φ this), - exact infi_emptyset.symm, - exact (funext $ λi, (@supr_pos _ _ _ (λh, range (std_basis R φ i)) $ finset.mem_univ i).symm) + casesI nonempty_fintype ι, + convert top_unique (infi_emptyset.ge.trans $ infi_ker_proj_le_supr_range_std_basis R φ _), + { exact funext (λ i, (@supr_pos _ _ _ (λ h, range $ std_basis R φ i) $ finset.mem_univ i).symm) }, + { rw [finset.coe_univ, set.union_empty] } end lemma disjoint_std_basis_std_basis (I J : set ι) (h : disjoint I J) : diff --git a/src/linear_algebra/trace.lean b/src/linear_algebra/trace.lean index 9345f6947722e..1222de9f18e2f 100644 --- a/src/linear_algebra/trace.lean +++ b/src/linear_algebra/trace.lean @@ -112,14 +112,15 @@ section variables {R : Type*} [comm_ring R] {M : Type*} [add_comm_group M] [module R M] variables (N : Type*) [add_comm_group N] [module R N] -variables {ι : Type*} [fintype ι] +variables {ι : Type*} /-- The trace of a linear map correspond to the contraction pairing under the isomorphism `End(M) ≃ M* ⊗ M`-/ -lemma trace_eq_contract_of_basis (b : basis ι R M) : +lemma trace_eq_contract_of_basis [finite ι] (b : basis ι R M) : (linear_map.trace R M) ∘ₗ (dual_tensor_hom R M M) = contract_left R M := begin classical, + casesI nonempty_fintype ι, apply basis.ext (basis.tensor_product (basis.dual_basis b) b), rintros ⟨i, j⟩, simp only [function.comp_app, basis.tensor_product_apply, basis.coe_dual_basis, coe_comp], @@ -132,7 +133,7 @@ end /-- The trace of a linear map correspond to the contraction pairing under the isomorphism `End(M) ≃ M* ⊗ M`-/ -lemma trace_eq_contract_of_basis' [decidable_eq ι] (b : basis ι R M) : +lemma trace_eq_contract_of_basis' [fintype ι] [decidable_eq ι] (b : basis ι R M) : (linear_map.trace R M) = (contract_left R M) ∘ₗ (dual_tensor_hom_equiv_of_basis b).symm.to_linear_map := by simp [linear_equiv.eq_comp_to_linear_map_symm, trace_eq_contract_of_basis b] From bb0978c26626b8ce0eb39eeb1a36a21fa914b0a5 Mon Sep 17 00:00:00 2001 From: Mauricio Collares Date: Sun, 4 Sep 2022 06:25:03 +0000 Subject: [PATCH 153/828] fix(analysis/inner_product_space,geometry/euclidean): two deterministic timeout fixes (#16365) CI seems to be having some issues after #16356, which I can't reproduce with `-T100000` locally but can with `-T90000`. This PR speeds up `analysis/inner_product_space/l2_space.lean:hilbert_basis.coe_mk` (which was already investigated before in #15271) and `geometry/euclidean/oriented_angle.lean:inner_eq_norm_mul_norm_mul_cos_oangle`. --- src/analysis/inner_product_space/l2_space.lean | 2 +- src/geometry/euclidean/oriented_angle.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 1a07d1c490c3a..5257cbb58485f 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -501,7 +501,7 @@ by rw [is_hilbert_sum.linear_isometry_equiv_symm_apply_single, @[simp] protected lemma coe_mk (hsp : ⊤ ≤ (span 𝕜 (set.range v)).topological_closure) : ⇑(hilbert_basis.mk hv hsp) = v := -funext $ orthonormal.linear_isometry_equiv_symm_apply_single_one hv _ +by apply (funext $ orthonormal.linear_isometry_equiv_symm_apply_single_one hv hsp) /-- An orthonormal family of vectors whose span has trivial orthogonal complement is a Hilbert basis. -/ diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 8762d2e983886..c803acefa02d3 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -879,7 +879,7 @@ begin star_ring_end_apply, star_trivial], rw [finset.sum_insert (dec_trivial : (0 : fin 2) ∉ ({1} : finset (fin 2))), finset.sum_singleton], - field_simp [norm_ne_zero_iff.2 hx, norm_ne_zero_iff.2 hy], + field_simp only [norm_ne_zero_iff.2 hx, norm_ne_zero_iff.2 hy, ne.def, not_false_iff], ring end From e4fb6414dd4bc9779a154efaccf292d168925cf7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 4 Sep 2022 15:22:31 +0000 Subject: [PATCH 154/828] feat(order/heyting/basic): Generalize boolean algebras lemmas (#16281) Generalize lemmas from `(generalized_)boolean_algebra` to `(generalized_)coheyting_algebra`. Dualize (some of) them. Duplicate lemmas have been made aliases. Add supporting lemmas to golf the proofs. Co-authored-by: Mauricio Collares --- .../prime_spectrum/basic.lean | 4 +- .../projective_spectrum/topology.lean | 6 +- src/data/finset/basic.lean | 7 +- src/data/set/basic.lean | 13 +- src/model_theory/definability.lean | 106 ++++-------- src/order/boolean_algebra.lean | 155 ++---------------- src/order/heyting/basic.lean | 132 ++++++++++++++- src/order/hom/basic.lean | 3 +- src/order/lattice.lean | 56 +++---- .../uniform_space/compact_separated.lean | 2 +- 10 files changed, 205 insertions(+), 279 deletions(-) diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index f3ec6ef8b5dce..5e0b3631d604e 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -298,7 +298,7 @@ lemma vanishing_ideal_Union {ι : Sort*} (t : ι → set (prime_spectrum R)) : lemma zero_locus_inf (I J : ideal R) : zero_locus ((I ⊓ J : ideal R) : set R) = zero_locus I ∪ zero_locus J := -set.ext $ λ x, by simpa using x.2.inf_le +set.ext $ λ x, x.2.inf_le lemma union_zero_locus (s s' : set R) : zero_locus s ∪ zero_locus s' = zero_locus ((ideal.span s) ⊓ (ideal.span s') : ideal R) := @@ -306,7 +306,7 @@ by { rw zero_locus_inf, simp } lemma zero_locus_mul (I J : ideal R) : zero_locus ((I * J : ideal R) : set R) = zero_locus I ∪ zero_locus J := -set.ext $ λ x, by simpa using x.2.mul_le +set.ext $ λ x, x.2.mul_le lemma zero_locus_singleton_mul (f g : R) : zero_locus ({f * g} : set R) = zero_locus {f} ∪ zero_locus {g} := diff --git a/src/algebraic_geometry/projective_spectrum/topology.lean b/src/algebraic_geometry/projective_spectrum/topology.lean index a0e04d1a10277..cc3cc424344d9 100644 --- a/src/algebraic_geometry/projective_spectrum/topology.lean +++ b/src/algebraic_geometry/projective_spectrum/topology.lean @@ -252,7 +252,7 @@ by convert (gc_ideal 𝒜).u_infi; exact homogeneous_ideal.to_ideal_infi _ lemma zero_locus_inf (I J : ideal A) : zero_locus 𝒜 ((I ⊓ J : ideal A) : set A) = zero_locus 𝒜 I ∪ zero_locus 𝒜 J := -set.ext $ λ x, by simpa using x.2.1.inf_le +set.ext $ λ x, x.2.1.inf_le lemma union_zero_locus (s s' : set A) : zero_locus 𝒜 s ∪ zero_locus 𝒜 s' = zero_locus 𝒜 ((ideal.span s) ⊓ (ideal.span s'): ideal A) := @@ -260,11 +260,11 @@ by { rw zero_locus_inf, simp } lemma zero_locus_mul_ideal (I J : ideal A) : zero_locus 𝒜 ((I * J : ideal A) : set A) = zero_locus 𝒜 I ∪ zero_locus 𝒜 J := -set.ext $ λ x, by simpa using x.2.1.mul_le +set.ext $ λ x, x.2.1.mul_le lemma zero_locus_mul_homogeneous_ideal (I J : homogeneous_ideal 𝒜) : zero_locus 𝒜 ((I * J : homogeneous_ideal 𝒜) : set A) = zero_locus 𝒜 I ∪ zero_locus 𝒜 J := -set.ext $ λ x, by simpa using x.2.1.mul_le +set.ext $ λ x, x.2.1.mul_le lemma zero_locus_singleton_mul (f g : A) : zero_locus 𝒜 ({f * g} : set A) = zero_locus 𝒜 {f} ∪ zero_locus 𝒜 {g} := diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 0f7306821f6e0..253cc1aa93ecc 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -1250,14 +1250,13 @@ sdiff_le_sdiff ‹s ≤ t› ‹v ≤ u› @[simp, norm_cast] lemma coe_sdiff (s₁ s₂ : finset α) : ↑(s₁ \ s₂) = (s₁ \ s₂ : set α) := set.ext $ λ _, mem_sdiff -@[simp] theorem union_sdiff_self_eq_union : s ∪ (t \ s) = s ∪ t := sup_sdiff_self_right - -@[simp] theorem sdiff_union_self_eq_union : (s \ t) ∪ t = s ∪ t := sup_sdiff_self_left +@[simp] lemma union_sdiff_self_eq_union : s ∪ t \ s = s ∪ t := sup_sdiff_self_right _ _ +@[simp] lemma sdiff_union_self_eq_union : s \ t ∪ t = s ∪ t := sup_sdiff_self_left _ _ lemma union_sdiff_left (s t : finset α) : (s ∪ t) \ s = t \ s := sup_sdiff_left_self lemma union_sdiff_right (s t : finset α) : (s ∪ t) \ t = s \ t := sup_sdiff_right_self -lemma union_sdiff_symm : s ∪ (t \ s) = t ∪ (s \ t) := sup_sdiff_symm +lemma union_sdiff_symm : s ∪ (t \ s) = t ∪ (s \ t) := by simp [union_comm] lemma sdiff_union_inter (s t : finset α) : (s \ t) ∪ (s ∩ t) = s := sup_sdiff_inf _ _ diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index d9a60cb8af577..2d56823fba457 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -945,7 +945,7 @@ ext $ λ x, and_iff_not_or_not @[simp] theorem compl_union_self (s : set α) : sᶜ ∪ s = univ := by rw [union_comm, union_compl_self] lemma compl_subset_comm : sᶜ ⊆ t ↔ tᶜ ⊆ s := @compl_le_iff_compl_le _ s _ _ -lemma subset_compl_comm : s ⊆ tᶜ ↔ t ⊆ sᶜ := @le_compl_iff_le_compl _ t _ _ +lemma subset_compl_comm : s ⊆ tᶜ ↔ t ⊆ sᶜ := @le_compl_iff_le_compl _ _ _ t @[simp] lemma compl_subset_compl : sᶜ ⊆ tᶜ ↔ t ⊆ s := @compl_le_compl_iff_le (set α) _ _ _ @@ -1140,11 +1140,8 @@ ext $ λ x, and_congr_right $ λ hx, or_iff_right $ ne_of_mem_of_not_mem hx h lemma insert_inter_of_not_mem (h : a ∉ t) : insert a s ∩ t = s ∩ t := ext $ λ x, and_congr_left $ λ hx, or_iff_right $ ne_of_mem_of_not_mem hx h -@[simp] theorem union_diff_self {s t : set α} : s ∪ (t \ s) = s ∪ t := -sup_sdiff_self_right - -@[simp] theorem diff_union_self {s t : set α} : (s \ t) ∪ t = s ∪ t := -sup_sdiff_self_left +@[simp] lemma union_diff_self {s t : set α} : s ∪ (t \ s) = s ∪ t := sup_sdiff_self _ _ +@[simp] lemma diff_union_self {s t : set α} : (s \ t) ∪ t = s ∪ t := sdiff_sup_self _ _ @[simp] theorem diff_inter_self {a b : set α} : (b \ a) ∩ a = ∅ := inf_sdiff_self_left @@ -2137,10 +2134,10 @@ by rw [← image_univ, preimage_inl_image_inr] by rw [← image_univ, preimage_inr_image_inl] @[simp] lemma compl_range_inl : (range (sum.inl : α → α ⊕ β))ᶜ = range (sum.inr : β → α ⊕ β) := -is_compl_range_inl_range_inr.compl_eq +is_compl.compl_eq is_compl_range_inl_range_inr @[simp] lemma compl_range_inr : (range (sum.inr : β → α ⊕ β))ᶜ = range (sum.inl : α → α ⊕ β) := -is_compl_range_inl_range_inr.symm.compl_eq +is_compl.compl_eq is_compl_range_inl_range_inr.symm @[simp] theorem range_quot_mk (r : α → α → Prop) : range (quot.mk r) = univ := (surjective_quot_mk r).range_eq diff --git a/src/model_theory/definability.lean b/src/model_theory/definability.lean index 5ebe106658247..a4ddd753f5059 100644 --- a/src/model_theory/definability.lean +++ b/src/model_theory/definability.lean @@ -180,9 +180,6 @@ begin simp } end -lemma fin.coe_cast_add_zero {m : ℕ} : (fin.cast_add 0 : fin m → fin (m + 0)) = id := -funext (λ _, fin.ext rfl) - /-- This lemma is only intended as a helper for `definable.image_comp. -/ lemma definable.image_comp_sum_inl_fin (m : ℕ) {s : set ((α ⊕ fin m) → M)} (h : A.definable L s) : @@ -192,7 +189,7 @@ begin refine ⟨(bounded_formula.relabel id φ).exs, _⟩, ext x, simp only [set.mem_image, mem_set_of_eq, bounded_formula.realize_exs, - bounded_formula.realize_relabel, function.comp.right_id, fin.coe_cast_add_zero], + bounded_formula.realize_relabel, function.comp.right_id, fin.cast_add_zero, fin.cast_refl], split, { rintro ⟨y, hy, rfl⟩, exact ⟨y ∘ sum.inr, @@ -203,10 +200,11 @@ end /-- Shows that definability is closed under finite projections. -/ lemma definable.image_comp_embedding {s : set (β → M)} (h : A.definable L s) - (f : α ↪ β) [fintype β] : + (f : α ↪ β) [finite β] : A.definable L ((λ g : β → M, g ∘ f) '' s) := begin classical, + casesI nonempty_fintype β, refine (congr rfl (ext (λ x, _))).mp (((h.image_comp_equiv (equiv.set.sum_compl (range f))).image_comp_equiv (equiv.sum_congr (equiv.of_injective f f.injective) (fintype.equiv_fin _).symm)).image_comp_sum_inl_fin _), @@ -217,10 +215,12 @@ end /-- Shows that definability is closed under finite projections. -/ lemma definable.image_comp {s : set (β → M)} (h : A.definable L s) - (f : α → β) [fintype α] [fintype β] : + (f : α → β) [finite α] [finite β] : A.definable L ((λ g : β → M, g ∘ f) '' s) := begin classical, + casesI nonempty_fintype α, + casesI nonempty_fintype β, have h := (((h.image_comp_equiv (equiv.set.sum_compl (range f))).image_comp_equiv (equiv.sum_congr (_root_.equiv.refl _) (fintype.equiv_fin _).symm)).image_comp_sum_inl_fin _).preimage_comp (range_splitting f), @@ -270,90 +270,42 @@ variables (L : first_order.language.{u v}) {M : Type w} [L.Structure M] (A : set def definable_set := { s : set (α → M) // A.definable L s} namespace definable_set -variables {L} {A} {α} - -instance : has_top (L.definable_set A α) := ⟨⟨⊤, definable_univ⟩⟩ - -instance : has_bot (L.definable_set A α) := ⟨⟨⊥, definable_empty⟩⟩ - -instance : inhabited (L.definable_set A α) := ⟨⊥⟩ +variables {L A α} {s t : L.definable_set A α} {x : α → M} instance : set_like (L.definable_set A α) (α → M) := { coe := subtype.val, coe_injective' := subtype.val_injective } -@[simp] -lemma mem_top {x : α → M} : x ∈ (⊤ : L.definable_set A α) := mem_univ x - -@[simp] -lemma coe_top : ((⊤ : L.definable_set A α) : set (α → M)) = ⊤ := rfl - -@[simp] -lemma not_mem_bot {x : α → M} : ¬ x ∈ (⊥ : L.definable_set A α) := not_mem_empty x +instance : has_top (L.definable_set A α) := ⟨⟨⊤, definable_univ⟩⟩ +instance : has_bot (L.definable_set A α) := ⟨⟨⊥, definable_empty⟩⟩ +instance : has_sup (L.definable_set A α) := ⟨λ s t, ⟨s ∪ t, s.2.union t.2⟩⟩ +instance : has_inf (L.definable_set A α) := ⟨λ s t, ⟨s ∩ t, s.2.inter t.2⟩⟩ +instance : has_compl (L.definable_set A α) := ⟨λ s, ⟨sᶜ, s.2.compl⟩⟩ +instance : has_sdiff (L.definable_set A α) := ⟨λ s t, ⟨s \ t, s.2.sdiff t.2⟩⟩ -@[simp] -lemma coe_bot : ((⊥ : L.definable_set A α) : set (α → M)) = ⊥ := rfl +instance : inhabited (L.definable_set A α) := ⟨⊥⟩ -instance : lattice (L.definable_set A α) := -subtype.lattice (λ _ _, definable.union) (λ _ _, definable.inter) +lemma le_iff : s ≤ t ↔ (s : set (α → M)) ≤ (t : set (α → M)) := iff.rfl -lemma le_iff {s t : L.definable_set A α} : s ≤ t ↔ (s : set (α → M)) ≤ (t : set (α → M)) := iff.rfl +@[simp] lemma mem_top : x ∈ (⊤ : L.definable_set A α) := mem_univ x +@[simp] lemma not_mem_bot {x : α → M} : ¬ x ∈ (⊥ : L.definable_set A α) := not_mem_empty x +@[simp] lemma mem_sup : x ∈ s ⊔ t ↔ x ∈ s ∨ x ∈ t := iff.rfl +@[simp] lemma mem_inf : x ∈ s ⊓ t ↔ x ∈ s ∧ x ∈ t := iff.rfl +@[simp] lemma mem_compl : x ∈ sᶜ ↔ ¬ x ∈ s := iff.rfl +@[simp] lemma mem_sdiff : x ∈ s \ t ↔ x ∈ s ∧ ¬ x ∈ t := iff.rfl -@[simp] -lemma coe_sup {s t : L.definable_set A α} : ((s ⊔ t : L.definable_set A α) : set (α → M)) = s ∪ t := +@[simp, norm_cast] lemma coe_top : ((⊤ : L.definable_set A α) : set (α → M)) = univ := rfl +@[simp, norm_cast] lemma coe_bot : ((⊥ : L.definable_set A α) : set (α → M)) = ∅ := rfl +@[simp, norm_cast] lemma coe_sup (s t : L.definable_set A α) : (↑(s ⊔ t) : set (α → M)) = s ∪ t := rfl - -@[simp] -lemma mem_sup {s t : L.definable_set A α} {x : α → M} : x ∈ s ⊔ t ↔ x ∈ s ∨ x ∈ t := iff.rfl - -@[simp] -lemma coe_inf {s t : L.definable_set A α} : ((s ⊓ t : L.definable_set A α) : set (α → M)) = s ∩ t := +@[simp, norm_cast] lemma coe_inf (s t : L.definable_set A α) : (↑(s ⊓ t) : set (α → M)) = s ∩ t := +rfl +@[simp, norm_cast] lemma coe_compl (s : L.definable_set A α) : (↑(sᶜ) : set (α → M)) = sᶜ := rfl +@[simp, norm_cast] lemma coe_sdiff (s t : L.definable_set A α) : (↑(s \ t) : set (α → M)) = s \ t := rfl - -@[simp] -lemma mem_inf {s t : L.definable_set A α} {x : α → M} : x ∈ s ⊓ t ↔ x ∈ s ∧ x ∈ t := iff.rfl - -instance : bounded_order (L.definable_set A α) := -{ bot_le := λ s x hx, false.elim hx, - le_top := λ s x hx, mem_univ x, - .. definable_set.has_top, - .. definable_set.has_bot } - -instance : distrib_lattice (L.definable_set A α) := -{ le_sup_inf := begin - intros s t u x, - simp only [and_imp, mem_inter_eq, set_like.mem_coe, coe_sup, coe_inf, mem_union_eq, - subtype.val_eq_coe], - tauto, - end, - .. definable_set.lattice } - -/-- The complement of a definable set is also definable. -/ -@[reducible] instance : has_compl (L.definable_set A α) := -⟨λ ⟨s, hs⟩, ⟨sᶜ, hs.compl⟩⟩ - -@[simp] -lemma mem_compl {s : L.definable_set A α} {x : α → M} : x ∈ sᶜ ↔ ¬ x ∈ s := -begin - cases s with s hs, - refl, -end - -@[simp] -lemma coe_compl {s : L.definable_set A α} : ((sᶜ : L.definable_set A α) : set (α → M)) = sᶜ := -begin - ext, - simp, -end instance : boolean_algebra (L.definable_set A α) := -{ sdiff := λ s t, s ⊓ tᶜ, - sdiff_eq := λ s t, rfl, - inf_compl_le_bot := λ ⟨s, hs⟩, by simp [le_iff], - top_le_sup_compl := λ ⟨s, hs⟩, by simp [le_iff], - .. definable_set.has_compl, - .. definable_set.bounded_order, - .. definable_set.distrib_lattice } +subtype.coe_injective.boolean_algebra _ coe_sup coe_inf coe_top coe_bot coe_compl coe_sdiff end definable_set end language diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index aadd18dcdade3..ba1b3f338cae5 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -113,37 +113,17 @@ begin exact (eq_of_inf_eq_sup_eq i s).symm, end -lemma sdiff_le' : x \ y ≤ x := +-- Use `sdiff_le` +private lemma sdiff_le' : x \ y ≤ x := calc x \ y ≤ (x ⊓ y) ⊔ (x \ y) : le_sup_right ... = x : sup_inf_sdiff x y -lemma inf_sdiff_right : x ⊓ (x \ y) = x \ y := inf_of_le_right (@sdiff_le' _ x y _) -lemma inf_sdiff_left : (x \ y) ⊓ x = x \ y := by rw [inf_comm, inf_sdiff_right] +-- Use `sdiff_sup_self` +private lemma sdiff_sup_self' : y \ x ⊔ x = y ⊔ x := +calc y \ x ⊔ x = y \ x ⊔ (x ⊔ x ⊓ y) : by rw sup_inf_self + ... = (y ⊓ x) ⊔ y \ x ⊔ x : by ac_refl + ... = y ⊔ x : by rw sup_inf_sdiff -@[simp] theorem sup_sdiff_self_right : x ⊔ (y \ x) = x ⊔ y := -calc x ⊔ (y \ x) = (x ⊔ (x ⊓ y)) ⊔ (y \ x) : by rw sup_inf_self - ... = x ⊔ ((y ⊓ x) ⊔ (y \ x)) : by ac_refl - ... = x ⊔ y : by rw sup_inf_sdiff - -@[simp] theorem sup_sdiff_self_left : (y \ x) ⊔ x = y ⊔ x := -by rw [sup_comm, sup_sdiff_self_right, sup_comm] - -lemma sup_sdiff_symm : x ⊔ (y \ x) = y ⊔ (x \ y) := -by rw [sup_sdiff_self_right, sup_sdiff_self_right, sup_comm] - -lemma sup_sdiff_cancel_right (h : x ≤ y) : x ⊔ (y \ x) = y := -by conv_rhs { rw [←sup_inf_sdiff y x, inf_eq_right.2 h] } - -lemma sdiff_sup_cancel (h : y ≤ x) : x \ y ⊔ y = x := by rw [sup_comm, sup_sdiff_cancel_right h] - -lemma sup_le_of_le_sdiff_left (h : y ≤ z \ x) (hxz : x ≤ z) : x ⊔ y ≤ z := -(sup_le_sup_left h x).trans (sup_sdiff_cancel_right hxz).le - -lemma sup_le_of_le_sdiff_right (h : x ≤ z \ y) (hyz : y ≤ z) : x ⊔ y ≤ z := -(sup_le_sup_right h y).trans (sdiff_sup_cancel hyz).le - -@[simp] lemma sup_sdiff_left : x ⊔ (x \ y) = x := sup_eq_left.2 sdiff_le' -lemma sup_sdiff_right : (x \ y) ⊔ x = x := sup_eq_right.2 sdiff_le' @[simp] lemma sdiff_inf_sdiff : x \ y ⊓ (y \ x) = ⊥ := eq.symm $ @@ -155,7 +135,7 @@ eq.symm $ ... = (x ⊓ y ⊓ (x \ y)) ⊔ (x ⊓ (y \ x) ⊓ (x \ y)) : by rw [inf_sup_right, @inf_comm _ _ x y] ... = x ⊓ (y \ x) ⊓ (x \ y) : by rw [inf_inf_sdiff, bot_sup_eq] ... = x ⊓ (x \ y) ⊓ (y \ x) : by ac_refl - ... = (x \ y) ⊓ (y \ x) : by rw inf_sdiff_right + ... = (x \ y) ⊓ (y \ x) : by rw inf_of_le_right sdiff_le' lemma disjoint_sdiff_sdiff : disjoint (x \ y) (y \ x) := sdiff_inf_sdiff.le @@ -171,18 +151,18 @@ instance generalized_boolean_algebra.to_generalized_coheyting_algebra : { sdiff := (\), sdiff_le_iff := λ y x z, ⟨λ h, le_of_inf_le_sup_le (le_of_eq - (calc y ⊓ (y \ x) = y \ x : inf_sdiff_right + (calc y ⊓ (y \ x) = y \ x : inf_of_le_right sdiff_le' ... = (x ⊓ (y \ x)) ⊔ (z ⊓ (y \ x)) : by rw [inf_eq_right.2 h, inf_sdiff_self_right, bot_sup_eq] ... = (x ⊔ z) ⊓ (y \ x) : inf_sup_right.symm)) - (calc y ⊔ y \ x = y : sup_sdiff_left + (calc y ⊔ y \ x = y : sup_of_le_left sdiff_le' ... ≤ y ⊔ (x ⊔ z) : le_sup_left - ... = ((y \ x) ⊔ x) ⊔ z : by rw [←sup_assoc, ←@sup_sdiff_self_left _ x y] + ... = ((y \ x) ⊔ x) ⊔ z : by rw [←sup_assoc, ←@sdiff_sup_self' _ x y] ... = x ⊔ z ⊔ y \ x : by ac_refl), λ h, le_of_inf_le_sup_le (calc y \ x ⊓ x = ⊥ : inf_sdiff_self_left ... ≤ z ⊓ x : bot_le) - (calc y \ x ⊔ x = y ⊔ x : sup_sdiff_self_left + (calc y \ x ⊔ x = y ⊔ x : sdiff_sup_self' ... ≤ (x ⊔ z) ⊔ x : sup_le_sup_right h x ... ≤ z ⊔ x : by rw [sup_assoc, sup_comm, sup_assoc, sup_idem])⟩, ..‹generalized_boolean_algebra α›, ..generalized_boolean_algebra.to_order_bot } @@ -196,11 +176,6 @@ theorem disjoint.sdiff_eq_of_sup_eq (hi : disjoint x z) (hs : x ⊔ z = y) : y \ have h : y ⊓ x = x := inf_eq_right.2 $ le_sup_left.trans hs.le, sdiff_unique (by rw [h, hs]) (by rw [h, hi.eq_bot]) -lemma disjoint.sup_sdiff_cancel_left (h : disjoint x y) : (x ⊔ y) \ x = y := -h.sdiff_eq_of_sup_eq rfl -lemma disjoint.sup_sdiff_cancel_right (h : disjoint x y) : (x ⊔ y) \ y = x := -h.symm.sdiff_eq_of_sup_eq sup_comm - protected theorem disjoint.sdiff_unique (hd : disjoint x z) (hz : z ≤ y) (hs : y ≤ x ⊔ z) : y \ x = z := sdiff_unique @@ -250,13 +225,6 @@ lemma le_iff_eq_sup_sdiff (hz : z ≤ y) (hx : x ≤ y) : x ≤ z ↔ y = z ⊔ exact bot_le, end⟩ -lemma sup_sdiff_eq_sup (h : z ≤ x) : x ⊔ y \ z = x ⊔ y := -sup_congr_left (sdiff_le.trans le_sup_right) $ le_sup_sdiff.trans $ sup_le_sup_right h _ - --- cf. `set.union_diff_cancel'` -lemma sup_sdiff_cancel' (hx : x ≤ z) (hz : z ≤ y) : z ⊔ (y \ x) = y := -((le_iff_eq_sup_sdiff hz (hx.trans hz)).1 hx).symm - -- cf. `is_compl.sup_inf` lemma sdiff_sup : y \ (x ⊔ z) = (y \ x) ⊓ (y \ z) := sdiff_unique @@ -280,10 +248,6 @@ lemma sdiff_eq_sdiff_iff_inf_eq_inf : y \ x = y \ z ↔ y ⊓ x = y ⊓ z := (by rw [sup_inf_sdiff, h, sup_inf_sdiff]), λ h, by rw [←sdiff_inf_self_right, ←sdiff_inf_self_right z y, inf_comm, h, inf_comm]⟩ -theorem disjoint.sdiff_eq_left (h : disjoint x y) : x \ y = x := -by conv_rhs { rw [←sup_inf_sdiff x y, h.eq_bot, bot_sup_eq] } -theorem disjoint.sdiff_eq_right (h : disjoint x y) : y \ x = y := h.symm.sdiff_eq_left - theorem sdiff_eq_self_iff_disjoint : x \ y = x ↔ disjoint y x := calc x \ y = x ↔ x \ y = x \ ⊥ : by rw sdiff_bot ... ↔ x ⊓ y = x ⊓ ⊥ : sdiff_eq_sdiff_iff_inf_eq_inf @@ -300,14 +264,6 @@ begin rw [←h, inf_eq_right.mpr hx], end -lemma sdiff_sdiff_le : x \ (x \ y) ≤ y := sdiff_le_iff.2 le_sdiff_sup - -lemma sdiff_triangle (x y z : α) : x \ z ≤ x \ y ⊔ y \ z := -begin - rw [sdiff_le_iff, sup_left_comm, ←sdiff_le_iff], - exact sdiff_sdiff_le.trans (sdiff_le_iff.1 le_rfl), -end - @[simp] lemma le_sdiff_iff : x ≤ y \ x ↔ x = ⊥ := ⟨λ h, disjoint_self.1 (disjoint_sdiff_self_right.mono_right h), λ h, h.le.trans bot_le⟩ @@ -320,11 +276,6 @@ calc (x ⊓ y) ⊓ z ⊔ (y \ z) = x ⊓ (y ⊓ z) ⊔ (y \ z) : by rw inf_assoc ... = (x ⊔ (y \ z)) ⊓ y : by rw [sup_inf_right, sup_inf_sdiff] ... = (x ⊓ y) ⊔ (y \ z) : by rw [inf_sup_right, inf_sdiff_left] -@[simp] lemma inf_sdiff_sup_left : (x \ z) ⊓ (x ⊔ y) = x \ z := -by rw [inf_sup_left, inf_sdiff_left, sup_inf_self] -@[simp] lemma inf_sdiff_sup_right : (x \ z) ⊓ (y ⊔ x) = x \ z := -by rw [sup_comm, inf_sdiff_sup_left] - lemma sdiff_sdiff_right : x \ (y \ z) = (x \ y) ⊔ (x ⊓ y ⊓ z) := begin rw [sup_comm, inf_comm, ←inf_assoc, sup_inf_inf_sdiff], @@ -334,7 +285,7 @@ begin ... = (x ⊔ x ⊓ z ⊔ x \ y) ⊓ (y \ z ⊔ (x ⊓ z ⊔ x \ y)) : by ac_refl ... = x ⊓ (y \ z ⊔ x ⊓ z ⊔ x \ y) : by rw [sup_inf_self, sup_sdiff_left, ←sup_assoc] ... = x ⊓ (y \ z ⊓ (z ⊔ y) ⊔ x ⊓ (z ⊔ y) ⊔ x \ y) : - by rw [sup_inf_left, sup_sdiff_self_left, inf_sup_right, @sup_comm _ _ y] + by rw [sup_inf_left, sdiff_sup_self', inf_sup_right, @sup_comm _ _ y] ... = x ⊓ (y \ z ⊔ (x ⊓ z ⊔ x ⊓ y) ⊔ x \ y) : by rw [inf_sdiff_sup_right, @inf_sup_left _ _ x z y] ... = x ⊓ (y \ z ⊔ (x ⊓ z ⊔ (x ⊓ y ⊔ x \ y))) : by ac_refl @@ -370,24 +321,9 @@ lemma sdiff_eq_comm (hy : y ≤ x) (hz : z ≤ x) : x \ y = z ↔ x \ z = y := lemma eq_of_sdiff_eq_sdiff (hxz : x ≤ z) (hyz : y ≤ z) (h : z \ x = z \ y) : x = y := by rw [←sdiff_sdiff_eq_self hxz, h, sdiff_sdiff_eq_self hyz] -lemma sdiff_sdiff_left : (x \ y) \ z = x \ (y ⊔ z) := -begin - rw sdiff_sup, - apply sdiff_unique, - { rw [←inf_sup_left, sup_sdiff_self_right, inf_sdiff_sup_right] }, - { rw [inf_assoc, @inf_comm _ _ z, inf_assoc, inf_sdiff_self_left, inf_bot_eq, inf_bot_eq] } -end - lemma sdiff_sdiff_left' : (x \ y) \ z = (x \ y) ⊓ (x \ z) := by rw [sdiff_sdiff_left, sdiff_sup] -lemma sdiff_sdiff_comm : (x \ y) \ z = (x \ z) \ y := -by rw [sdiff_sdiff_left, sup_comm, sdiff_sdiff_left] - -@[simp] lemma sdiff_idem : x \ y \ y = x \ y := by rw [sdiff_sdiff_left, sup_idem] - -@[simp] lemma sdiff_sdiff_self : x \ y \ x = ⊥ := by rw [sdiff_sdiff_comm, sdiff_self, bot_sdiff] - lemma sdiff_sdiff_sup_sdiff : z \ (x \ y ⊔ y \ x) = z ⊓ (z \ x ⊔ y) ⊓ (z \ y ⊔ x) := calc z \ (x \ y ⊔ y \ x) = (z \ x ⊔ z ⊓ x ⊓ y) ⊓ (z \ y ⊔ z ⊓ y ⊓ x) : by rw [sdiff_sup, sdiff_sdiff_right, sdiff_sdiff_right] @@ -405,29 +341,6 @@ calc z \ (x \ y ⊔ y \ x) = ... = (z \ x) ⊓ (z \ y) ⊔ z ⊓ y ⊓ x : sup_inf_right.symm ... = z ⊓ x ⊓ y ⊔ ((z \ x) ⊓ (z \ y)) : by ac_refl -lemma sup_sdiff : (x ⊔ y) \ z = (x \ z) ⊔ (y \ z) := -sdiff_unique - (calc (x ⊔ y) ⊓ z ⊔ (x \ z ⊔ y \ z) = - (x ⊓ z ⊔ y ⊓ z) ⊔ (x \ z ⊔ y \ z) : by rw inf_sup_right - ... = x ⊓ z ⊔ x \ z ⊔ y \ z ⊔ y ⊓ z : by ac_refl - ... = x ⊔ (y ⊓ z ⊔ y \ z) : by rw [sup_inf_sdiff, sup_assoc, @sup_comm _ _ (y \ z)] - ... = x ⊔ y : by rw sup_inf_sdiff) - (calc (x ⊔ y) ⊓ z ⊓ (x \ z ⊔ y \ z) = - (x ⊓ z ⊔ y ⊓ z) ⊓ (x \ z ⊔ y \ z) : by rw inf_sup_right - ... = (x ⊓ z ⊔ y ⊓ z) ⊓ (x \ z) ⊔ ((x ⊓ z ⊔ y ⊓ z) ⊓ (y \ z)) : - by rw [@inf_sup_left _ _ (x ⊓ z ⊔ y ⊓ z)] - ... = (y ⊓ z ⊓ (x \ z)) ⊔ ((x ⊓ z ⊔ y ⊓ z) ⊓ (y \ z)) : - by rw [inf_sup_right, inf_inf_sdiff, bot_sup_eq] - ... = (x ⊓ z ⊔ y ⊓ z) ⊓ (y \ z) : by rw [inf_assoc, inf_sdiff_self_right, inf_bot_eq, bot_sup_eq] - ... = x ⊓ z ⊓ (y \ z) : by rw [inf_sup_right, inf_inf_sdiff, sup_bot_eq] - ... = ⊥ : by rw [inf_assoc, inf_sdiff_self_right, inf_bot_eq]) - -lemma sup_sdiff_right_self : (x ⊔ y) \ y = x \ y := -by rw [sup_sdiff, sdiff_self, sup_bot_eq] - -lemma sup_sdiff_left_self : (x ⊔ y) \ x = y \ x := -by rw [sup_comm, sup_sdiff_right_self] - lemma inf_sdiff : (x ⊓ y) \ z = (x \ z) ⊓ (y \ z) := sdiff_unique (calc (x ⊓ y) ⊓ z ⊔ ((x \ z) ⊓ (y \ z)) = @@ -459,10 +372,6 @@ by rw [sdiff_inf, sdiff_eq_bot_iff.2 inf_le_left, bot_sup_eq, inf_sdiff_assoc] lemma inf_sdiff_distrib_right (a b c : α) : a \ b ⊓ c = (a ⊓ c) \ (b ⊓ c) := by simp_rw [@inf_comm _ _ _ c, inf_sdiff_distrib_left] -lemma sdiff_sup_sdiff_cancel (hyx : y ≤ x) (hzy : z ≤ y) : x \ y ⊔ y \ z = x \ z := -by rw [←sup_sdiff_inf (x \ z) y, sdiff_sdiff_left, sup_eq_right.2 hzy, inf_sdiff_right_comm, - inf_eq_right.2 hyx] - lemma sup_eq_sdiff_sup_sdiff_sup_inf : x ⊔ y = (x \ y) ⊔ (y \ x) ⊔ (x ⊓ y) := eq.symm $ calc (x \ y) ⊔ (y \ x) ⊔ (x ⊓ y) = @@ -471,18 +380,6 @@ eq.symm $ ... = (x ⊔ (y \ x)) ⊓ ((x \ y) ⊔ y) : by rw [sup_sdiff_right, sup_sdiff_right] ... = x ⊔ y : by rw [sup_sdiff_self_right, sup_sdiff_self_left, inf_idem] -lemma sdiff_le_sdiff_of_sup_le_sup_left (h : z ⊔ x ≤ z ⊔ y) : x \ z ≤ y \ z := -begin - rw [←sup_sdiff_left_self, ←@sup_sdiff_left_self _ _ y], - exact sdiff_le_sdiff_right h, -end - -lemma sdiff_le_sdiff_of_sup_le_sup_right (h : x ⊔ z ≤ y ⊔ z) : x \ z ≤ y \ z := -begin - rw [←sup_sdiff_right_self, ←@sup_sdiff_right_self _ y], - exact sdiff_le_sdiff_right h, -end - lemma sup_lt_of_lt_sdiff_left (h : y < z \ x) (hxz : x ≤ z) : x ⊔ y < z := begin rw ←sup_sdiff_cancel_right hxz, @@ -556,8 +453,6 @@ lemma is_compl_compl : is_compl x xᶜ := is_compl.of_eq inf_compl_eq_bot' sup_c lemma sdiff_eq : x \ y = x ⊓ yᶜ := boolean_algebra.sdiff_eq x y lemma himp_eq : x ⇨ y = y ⊔ xᶜ := boolean_algebra.himp_eq x y -@[simp] lemma top_sdiff : ⊤ \ x = xᶜ := by rw [sdiff_eq, top_inf_eq] - @[priority 100] instance boolean_algebra.to_complemented_lattice : complemented_lattice α := ⟨λ x, ⟨xᶜ, is_compl_compl⟩⟩ @@ -579,11 +474,7 @@ instance boolean_algebra.to_biheyting_algebra : biheyting_algebra α := @[simp] lemma hnot_eq_compl : ¬x = xᶜ := rfl -theorem is_compl.eq_compl (h : is_compl x y) : x = yᶜ := -h.left_unique is_compl_compl.symm - -theorem is_compl.compl_eq (h : is_compl x y) : xᶜ = y := -(h.right_unique is_compl_compl).symm +@[simp] lemma top_sdiff : ⊤ \ x = xᶜ := top_sdiff' _ theorem eq_compl_iff_is_compl : x = yᶜ ↔ is_compl x y := ⟨λ h, by { rw h, exact is_compl_compl.symm }, is_compl.eq_compl⟩ @@ -597,11 +488,7 @@ by rw [eq_comm, compl_eq_iff_is_compl, eq_compl_iff_is_compl] theorem eq_compl_comm : x = yᶜ ↔ y = xᶜ := by rw [eq_comm, compl_eq_iff_is_compl, eq_compl_iff_is_compl] -theorem compl_unique (i : x ⊓ y = ⊥) (s : x ⊔ y = ⊤) : xᶜ = y := -(is_compl.of_eq i s).compl_eq - -@[simp] theorem compl_compl (x : α) : xᶜᶜ = x := -is_compl_compl.symm.compl_eq +@[simp] theorem compl_compl (x : α) : xᶜᶜ = x := (@is_compl_compl _ x _).symm.compl_eq theorem compl_comp_compl : compl ∘ compl = @id α := funext compl_compl @@ -628,22 +515,15 @@ is_compl_bot_top.compl_eq_iff @[simp] theorem compl_eq_bot : xᶜ = ⊥ ↔ x = ⊤ := is_compl_top_bot.compl_eq_iff -@[simp] theorem compl_inf : (x ⊓ y)ᶜ = xᶜ ⊔ yᶜ := -(is_compl_compl.inf_sup is_compl_compl).compl_eq +@[simp] theorem compl_inf : (x ⊓ y)ᶜ = xᶜ ⊔ yᶜ := hnot_inf_distrib _ _ @[simp] theorem compl_le_compl_iff_le : yᶜ ≤ xᶜ ↔ x ≤ y := ⟨assume h, by have h := compl_le_compl h; simp at h; assumption, compl_le_compl⟩ -theorem le_compl_of_le_compl (h : y ≤ xᶜ) : x ≤ yᶜ := -by simpa only [compl_compl] using compl_le_compl h - theorem compl_le_of_compl_le (h : yᶜ ≤ x) : xᶜ ≤ y := by simpa only [compl_compl] using compl_le_compl h -theorem le_compl_iff_le_compl : y ≤ xᶜ ↔ x ≤ yᶜ := -⟨le_compl_of_le_compl, le_compl_of_le_compl⟩ - theorem compl_le_iff_compl_le : xᶜ ≤ y ↔ yᶜ ≤ x := ⟨compl_le_of_compl_le, compl_le_of_compl_le⟩ @@ -671,9 +551,6 @@ by rw [←le_compl_iff_disjoint_left, compl_compl] lemma disjoint_compl_right_iff : disjoint x yᶜ ↔ x ≤ y := by rw [←le_compl_iff_disjoint_right, compl_compl] -alias disjoint_compl_left_iff ↔ _ has_le.le.disjoint_compl_left -alias disjoint_compl_right_iff ↔ _ has_le.le.disjoint_compl_right - end boolean_algebra instance Prop.boolean_algebra : boolean_algebra Prop := diff --git a/src/order/heyting/basic.lean b/src/order/heyting/basic.lean index 0b602af952915..f295ebab1a3f3 100644 --- a/src/order/heyting/basic.lean +++ b/src/order/heyting/basic.lean @@ -209,8 +209,11 @@ See also `Prop.heyting_algebra`. -/ -- `p → q → r ↔ p ∧ q → r` @[simp] lemma le_himp_iff : a ≤ b ⇨ c ↔ a ⊓ b ≤ c := generalized_heyting_algebra.le_himp_iff _ _ _ +-- `p → q → r ↔ q ∧ p → r` +lemma le_himp_iff' : a ≤ b ⇨ c ↔ b ⊓ a ≤ c := by rw [le_himp_iff, inf_comm] + -- `p → q → r ↔ q → p → r` -lemma le_himp_comm : a ≤ b ⇨ c ↔ b ≤ a ⇨ c := by rw [le_himp_iff, le_himp_iff, inf_comm] +lemma le_himp_comm : a ≤ b ⇨ c ↔ b ≤ a ⇨ c := by rw [le_himp_iff, le_himp_iff'] -- `p → q → p` lemma le_himp : a ≤ b ⇨ a := le_himp_iff.2 inf_le_left @@ -247,7 +250,7 @@ lemma himp_himp (a b c : α) : a ⇨ b ⇨ c = a ⊓ b ⇨ c := eq_of_forall_le_iff $ λ d, by simp_rw [le_himp_iff, inf_assoc] -- `(q → r) → (p → q) → q → r` -@[simp] lemma himp_le_himp_himp : b ⇨ c ≤ (a ⇨ b) ⇨ a ⇨ c := +@[simp] lemma himp_le_himp_himp_himp : b ⇨ c ≤ (a ⇨ b) ⇨ a ⇨ c := begin rw [le_himp_iff, le_himp_iff, inf_assoc, himp_inf_self, ←inf_assoc, himp_inf_self, inf_assoc], exact inf_le_left, @@ -276,6 +279,25 @@ by rw [sup_himp_distrib, himp_self, top_inf_eq] @[simp] lemma sup_himp_self_right (a b : α) : (a ⊔ b) ⇨ b = a ⇨ b := by rw [sup_himp_distrib, himp_self, inf_top_eq] +lemma codisjoint.himp_eq_right (h : codisjoint a b) : b ⇨ a = a := +by { conv_rhs { rw ←@top_himp _ _ a }, rw [←h.eq_top, sup_himp_self_left] } + +lemma codisjoint.himp_eq_left (h : codisjoint a b) : a ⇨ b = b := h.symm.himp_eq_right + +lemma codisjoint.himp_inf_cancel_right (h : codisjoint a b) : a ⇨ (a ⊓ b) = b := +by rw [himp_inf_distrib, himp_self, top_inf_eq, h.himp_eq_left] + +lemma codisjoint.himp_inf_cancel_left (h : codisjoint a b) : b ⇨ (a ⊓ b) = a := +by rw [himp_inf_distrib, himp_self, inf_top_eq, h.himp_eq_right] + +lemma le_himp_himp : a ≤ (a ⇨ b) ⇨ b := le_himp_iff.2 inf_himp_le + +lemma himp_triangle (a b c : α) : (a ⇨ b) ⊓ (b ⇨ c) ≤ a ⇨ c := +by { rw [le_himp_iff, inf_right_comm, ←le_himp_iff], exact himp_inf_le.trans le_himp_himp } + +lemma himp_inf_himp_cancel (hba : b ≤ a) (hcb : c ≤ b) : (a ⇨ b) ⊓ (b ⇨ c) = a ⇨ c := +(himp_triangle _ _ _).antisymm $ le_inf (himp_le_himp_left hcb) (himp_le_himp_right hba) + @[priority 100] -- See note [lower instance priority] instance generalized_heyting_algebra.to_distrib_lattice : distrib_lattice α := distrib_lattice.of_inf_sup_le $ λ a b c, @@ -303,7 +325,9 @@ variables [generalized_coheyting_algebra α] {a b c d : α} @[simp] lemma sdiff_le_iff : a \ b ≤ c ↔ a ≤ b ⊔ c := generalized_coheyting_algebra.sdiff_le_iff _ _ _ -lemma sdiff_le_comm : a \ b ≤ c ↔ a \ c ≤ b := by rw [sdiff_le_iff, sdiff_le_iff, sup_comm] +lemma sdiff_le_iff' : a \ b ≤ c ↔ a ≤ c ⊔ b := by rw [sdiff_le_iff, sup_comm] + +lemma sdiff_le_comm : a \ b ≤ c ↔ a \ c ≤ b := by rw [sdiff_le_iff, sdiff_le_iff'] lemma sdiff_le : a \ b ≤ a := sdiff_le_iff.2 le_sup_right @@ -317,35 +341,74 @@ lemma disjoint.disjoint_sdiff_right (h : disjoint a b) : disjoint a (b \ c) := h lemma le_sup_sdiff : a ≤ b ⊔ a \ b := sdiff_le_iff.1 le_rfl lemma le_sdiff_sup : a ≤ a \ b ⊔ b := by rw [sup_comm, ←sdiff_le_iff] +@[simp] lemma sup_sdiff_left : a ⊔ a \ b = a := sup_of_le_left sdiff_le +@[simp] lemma sup_sdiff_right : a \ b ⊔ a = a := sup_of_le_right sdiff_le +@[simp] lemma inf_sdiff_left : a \ b ⊓ a = a \ b := inf_of_le_left sdiff_le +@[simp] lemma inf_sdiff_right : a ⊓ a \ b = a \ b := inf_of_le_right sdiff_le + @[simp] lemma sup_sdiff_self (a b : α) : a ⊔ b \ a = a ⊔ b := le_antisymm (sup_le_sup_left sdiff_le _) (sup_le le_sup_left le_sup_sdiff) @[simp] lemma sdiff_sup_self (a b : α) : b \ a ⊔ a = b ⊔ a := by rw [sup_comm, sup_sdiff_self, sup_comm] +alias sdiff_sup_self ← sup_sdiff_self_left +alias sup_sdiff_self ← sup_sdiff_self_right + +lemma sup_sdiff_eq_sup (h : c ≤ a) : a ⊔ b \ c = a ⊔ b := +sup_congr_left (sdiff_le.trans le_sup_right) $ le_sup_sdiff.trans $ sup_le_sup_right h _ + +-- cf. `set.union_diff_cancel'` +lemma sup_sdiff_cancel' (hab : a ≤ b) (hbc : b ≤ c) : b ⊔ c \ a = c := +by rw [sup_sdiff_eq_sup hab, sup_of_le_right hbc] + +lemma sup_sdiff_cancel_right (h : a ≤ b) : a ⊔ b \ a = b := sup_sdiff_cancel' le_rfl h + +lemma sdiff_sup_cancel (h : b ≤ a) : a \ b ⊔ b = a := by rw [sup_comm, sup_sdiff_cancel_right h] + +lemma sup_le_of_le_sdiff_left (h : b ≤ c \ a) (hac : a ≤ c) : a ⊔ b ≤ c := +sup_le hac $ h.trans sdiff_le + +lemma sup_le_of_le_sdiff_right (h : a ≤ c \ b) (hbc : b ≤ c) : a ⊔ b ≤ c := +sup_le (h.trans sdiff_le) hbc + @[simp] lemma sdiff_eq_bot_iff : a \ b = ⊥ ↔ a ≤ b := by rw [←le_bot_iff, sdiff_le_iff, sup_bot_eq] @[simp] lemma sdiff_bot : a \ ⊥ = a := eq_of_forall_ge_iff $ λ b, by rw [sdiff_le_iff, bot_sup_eq] @[simp] lemma bot_sdiff : ⊥ \ a = ⊥ := sdiff_eq_bot_iff.2 bot_le -lemma sdiff_sdiff (a b c : α) : a \ b \ c = a \ (b ⊔ c) := -eq_of_forall_ge_iff $ λ d, by simp_rw [sdiff_le_iff, sup_assoc] - -@[simp] lemma sdiff_sdiff_le_sdiff : a \ b \ (a \ c) ≤ c \ b := +@[simp] lemma sdiff_sdiff_sdiff_le_sdiff : a \ b \ (a \ c) ≤ c \ b := begin rw [sdiff_le_iff, sdiff_le_iff, sup_left_comm, sup_sdiff_self, sup_left_comm, sdiff_sup_self, sup_left_comm], exact le_sup_left, end +lemma sdiff_sdiff (a b c : α) : a \ b \ c = a \ (b ⊔ c) := +eq_of_forall_ge_iff $ λ d, by simp_rw [sdiff_le_iff, sup_assoc] + +lemma sdiff_sdiff_left : a \ b \ c = a \ (b ⊔ c) := sdiff_sdiff _ _ _ + lemma sdiff_right_comm (a b c : α) : a \ b \ c = a \ c \ b := by simp_rw [sdiff_sdiff, sup_comm] +lemma sdiff_sdiff_comm : a \ b \ c = a \ c \ b := sdiff_right_comm _ _ _ + +@[simp] lemma sdiff_idem : a \ b \ b = a \ b := by rw [sdiff_sdiff_left, sup_idem] +@[simp] lemma sdiff_sdiff_self : a \ b \ a = ⊥ := by rw [sdiff_sdiff_comm, sdiff_self, bot_sdiff] + lemma sup_sdiff_distrib (a b c : α) : (a ⊔ b) \ c = a \ c ⊔ b \ c := eq_of_forall_ge_iff $ λ d, by simp_rw [sdiff_le_iff, sup_le_iff, sdiff_le_iff] lemma sdiff_inf_distrib (a b c : α) : a \ (b ⊓ c) = a \ b ⊔ a \ c := eq_of_forall_ge_iff $ λ d, by { rw [sup_le_iff, sdiff_le_comm, le_inf_iff], simp_rw sdiff_le_comm } +lemma sup_sdiff : (a ⊔ b) \ c = a \ c ⊔ b \ c := sup_sdiff_distrib _ _ _ + +@[simp] lemma sup_sdiff_right_self : (a ⊔ b) \ b = a \ b := +by rw [sup_sdiff, sdiff_self, sup_bot_eq] + +@[simp] lemma sup_sdiff_left_self : (a ⊔ b) \ a = b \ a := by rw [sup_comm, sup_sdiff_right_self] + lemma sdiff_le_sdiff_right (h : a ≤ b) : a \ c ≤ b \ c := sdiff_le_iff.2 $ h.trans $ le_sup_sdiff lemma sdiff_le_sdiff_left (h : a ≤ b) : c \ b ≤ c \ a := @@ -363,6 +426,36 @@ by rw [sdiff_inf, sdiff_self, bot_sup_eq] @[simp] lemma sdiff_inf_self_right (a b : α) : b \ (a ⊓ b) = b \ a := by rw [sdiff_inf, sdiff_self, sup_bot_eq] +lemma disjoint.sdiff_eq_left (h : disjoint a b) : a \ b = a := +by { conv_rhs { rw ←@sdiff_bot _ _ a }, rw [←h.eq_bot, sdiff_inf_self_left] } + +lemma disjoint.sdiff_eq_right (h : disjoint a b) : b \ a = b := h.symm.sdiff_eq_left + +lemma disjoint.sup_sdiff_cancel_left (h : disjoint a b) : (a ⊔ b) \ a = b := +by rw [sup_sdiff, sdiff_self, bot_sup_eq, h.sdiff_eq_right] + +lemma disjoint.sup_sdiff_cancel_right (h : disjoint a b) : (a ⊔ b) \ b = a := +by rw [sup_sdiff, sdiff_self, sup_bot_eq, h.sdiff_eq_left] + +lemma sdiff_sdiff_le : a \ (a \ b) ≤ b := sdiff_le_iff.2 le_sdiff_sup + +lemma sdiff_triangle (a b c : α) : a \ c ≤ a \ b ⊔ b \ c := +by { rw [sdiff_le_iff, sup_left_comm, ←sdiff_le_iff], exact sdiff_sdiff_le.trans le_sup_sdiff } + +lemma sdiff_sup_sdiff_cancel (hba : b ≤ a) (hcb : c ≤ b) : a \ b ⊔ b \ c = a \ c := +(sdiff_triangle _ _ _).antisymm' $ sup_le (sdiff_le_sdiff_left hcb) (sdiff_le_sdiff_right hba) + +lemma sdiff_le_sdiff_of_sup_le_sup_left (h : c ⊔ a ≤ c ⊔ b) : a \ c ≤ b \ c := +by { rw [←sup_sdiff_left_self, ←@sup_sdiff_left_self _ _ _ b], exact sdiff_le_sdiff_right h } + +lemma sdiff_le_sdiff_of_sup_le_sup_right (h : a ⊔ c ≤ b ⊔ c) : a \ c ≤ b \ c := +by { rw [←sup_sdiff_right_self, ←@sup_sdiff_right_self _ _ b], exact sdiff_le_sdiff_right h } + +@[simp] lemma inf_sdiff_sup_left : a \ c ⊓ (a ⊔ b) = a \ c := +inf_of_le_left $ sdiff_le.trans le_sup_left +@[simp] lemma inf_sdiff_sup_right : a \ c ⊓ (b ⊔ a) = a \ c := +inf_of_le_left $ sdiff_le.trans le_sup_right + @[priority 100] -- See note [lower instance priority] instance generalized_coheyting_algebra.to_distrib_lattice : distrib_lattice α := { le_sup_inf := λ a b c, by simp_rw [←sdiff_le_iff, le_inf_iff, sdiff_le_iff, ←le_inf_iff], @@ -415,10 +508,23 @@ by rw [le_compl_iff_disjoint_right, le_compl_iff_disjoint_left] alias le_compl_iff_disjoint_right ↔ _ disjoint.le_compl_right alias le_compl_iff_disjoint_left ↔ _ disjoint.le_compl_left +alias le_compl_comm ← le_compl_iff_le_compl +alias le_compl_comm ↔ le_compl_of_le_compl _ lemma disjoint_compl_left : disjoint aᶜ a := le_himp_iff.1 (himp_bot _).ge lemma disjoint_compl_right : disjoint a aᶜ := disjoint_compl_left.symm +lemma has_le.le.disjoint_compl_left (h : b ≤ a) : disjoint aᶜ b := disjoint_compl_left.mono_right h +lemma has_le.le.disjoint_compl_right (h : a ≤ b) : disjoint a bᶜ := disjoint_compl_right.mono_left h + +lemma is_compl.compl_eq (h : is_compl a b) : aᶜ = b := +h.1.le_compl_left.antisymm' $ disjoint.le_of_codisjoint disjoint_compl_left h.2 + +lemma is_compl.eq_compl (h : is_compl a b) : a = bᶜ := +h.1.le_compl_right.antisymm $ disjoint.le_of_codisjoint disjoint_compl_left h.2.symm + +lemma compl_unique (h₀ : a ⊓ b = ⊥) (h₁ : a ⊔ b = ⊤) : aᶜ = b := (is_compl.of_eq h₀ h₁).compl_eq + @[simp] lemma inf_compl_self (a : α) : a ⊓ aᶜ = ⊥ := disjoint_compl_right.eq_bot @[simp] lemma compl_inf_self (a : α) : aᶜ ⊓ a = ⊥ := disjoint_compl_left.eq_bot lemma inf_compl_eq_bot : a ⊓ aᶜ = ⊥ := inf_compl_self _ @@ -523,6 +629,18 @@ alias hnot_le_iff_codisjoint_left ↔ _ codisjoint.hnot_le_left lemma codisjoint_hnot_right : codisjoint a (¬a) := sdiff_le_iff.1 (top_sdiff' _).le lemma codisjoint_hnot_left : codisjoint (¬a) a := codisjoint_hnot_right.symm +lemma has_le.le.codisjoint_hnot_left (h : a ≤ b) : codisjoint (¬a) b := +codisjoint_hnot_left.mono_right h + +lemma has_le.le.codisjoint_hnot_right (h : b ≤ a) : codisjoint a (¬b) := +codisjoint_hnot_right.mono_left h + +lemma is_compl.hnot_eq (h : is_compl a b) : ¬a = b := +h.2.hnot_le_right.antisymm $ disjoint.le_of_codisjoint h.1.symm codisjoint_hnot_right + +lemma is_compl.eq_hnot (h : is_compl a b) : a = ¬b := +h.2.hnot_le_left.antisymm' $ disjoint.le_of_codisjoint h.1 codisjoint_hnot_right + @[simp] lemma sup_hnot_self (a : α) : a ⊔ ¬a = ⊤ := codisjoint_hnot_right.eq_top @[simp] lemma hnot_sup_self (a : α) : ¬a ⊔ a = ⊤ := codisjoint_hnot_left.eq_top diff --git a/src/order/hom/basic.lean b/src/order/hom/basic.lean index 074068651b763..deab8d48d7eb2 100644 --- a/src/order/hom/basic.lean +++ b/src/order/hom/basic.lean @@ -807,7 +807,8 @@ lemma order_iso.map_inf [semilattice_inf α] [semilattice_inf β] (f : α ≃o f (x ⊓ y) = f x ⊓ f y := begin refine (f.to_order_embedding.map_inf_le x y).antisymm _, - simpa [← f.symm.le_iff_le] using f.symm.to_order_embedding.map_inf_le (f x) (f y) + apply f.symm.le_iff_le.1, + simpa using f.symm.to_order_embedding.map_inf_le (f x) (f y), end lemma order_iso.map_sup [semilattice_sup α] [semilattice_sup β] (f : α ≃o β) (x y : α) : diff --git a/src/order/lattice.lean b/src/order/lattice.lean index 1022d57cd1387..66877bbc6e257 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -145,27 +145,19 @@ semilattice_sup.sup_le a b c ⟨assume h : a ⊔ b ≤ c, ⟨le_trans le_sup_left h, le_trans le_sup_right h⟩, assume ⟨h₁, h₂⟩, sup_le h₁ h₂⟩ -@[simp] theorem sup_eq_left : a ⊔ b = a ↔ b ≤ a := -le_antisymm_iff.trans $ by simp [le_refl] +@[simp] lemma sup_eq_left : a ⊔ b = a ↔ b ≤ a := le_antisymm_iff.trans $ by simp [le_rfl] +@[simp] lemma sup_eq_right : a ⊔ b = b ↔ a ≤ b := le_antisymm_iff.trans $ by simp [le_rfl] +@[simp] lemma left_eq_sup : a = a ⊔ b ↔ b ≤ a := eq_comm.trans sup_eq_left +@[simp] lemma right_eq_sup : b = a ⊔ b ↔ a ≤ b := eq_comm.trans sup_eq_right -theorem sup_of_le_left (h : b ≤ a) : a ⊔ b = a := -sup_eq_left.2 h +alias sup_eq_left ↔ _ sup_of_le_left +alias sup_eq_right ↔ le_of_sup_eq sup_of_le_right -@[simp] theorem left_eq_sup : a = a ⊔ b ↔ b ≤ a := -eq_comm.trans sup_eq_left +attribute [simp] sup_of_le_left sup_of_le_right @[simp] theorem left_lt_sup : a < a ⊔ b ↔ ¬b ≤ a := le_sup_left.lt_iff_ne.trans $ not_congr left_eq_sup -@[simp] theorem sup_eq_right : a ⊔ b = b ↔ a ≤ b := -le_antisymm_iff.trans $ by simp [le_refl] - -theorem sup_of_le_right (h : a ≤ b) : a ⊔ b = b := -sup_eq_right.2 h - -@[simp] theorem right_eq_sup : b = a ⊔ b ↔ a ≤ b := -eq_comm.trans sup_eq_right - @[simp] theorem right_lt_sup : b < a ⊔ b ↔ ¬a ≤ b := le_sup_right.lt_iff_ne.trans $ not_congr right_eq_sup @@ -189,9 +181,6 @@ sup_le_sup le_rfl h₁ theorem sup_le_sup_right (h₁ : a ≤ b) (c) : a ⊔ c ≤ b ⊔ c := sup_le_sup h₁ le_rfl -theorem le_of_sup_eq (h : a ⊔ b = b) : a ≤ b := -by { rw ← h, simp } - @[simp] theorem sup_idem : a ⊔ a = a := by apply le_antisymm; simp @@ -331,31 +320,22 @@ lt_of_le_of_lt inf_le_right h @[simp] theorem le_inf_iff : a ≤ b ⊓ c ↔ a ≤ b ∧ a ≤ c := @sup_le_iff αᵒᵈ _ _ _ _ -@[simp] theorem inf_eq_left : a ⊓ b = a ↔ a ≤ b := -le_antisymm_iff.trans $ by simp [le_refl] - -@[simp] theorem inf_lt_left : a ⊓ b < a ↔ ¬a ≤ b := @left_lt_sup αᵒᵈ _ _ _ - -theorem inf_of_le_left (h : a ≤ b) : a ⊓ b = a := -inf_eq_left.2 h +@[simp] lemma inf_eq_left : a ⊓ b = a ↔ a ≤ b := le_antisymm_iff.trans $ by simp [le_rfl] +@[simp] lemma inf_eq_right : a ⊓ b = b ↔ b ≤ a := le_antisymm_iff.trans $ by simp [le_rfl] +@[simp] lemma left_eq_inf : a = a ⊓ b ↔ a ≤ b := eq_comm.trans inf_eq_left +@[simp] lemma right_eq_inf : b = a ⊓ b ↔ b ≤ a := eq_comm.trans inf_eq_right -@[simp] theorem left_eq_inf : a = a ⊓ b ↔ a ≤ b := -eq_comm.trans inf_eq_left +alias inf_eq_left ↔ le_of_inf_eq inf_of_le_left +alias inf_eq_right ↔ _ inf_of_le_right -@[simp] theorem inf_eq_right : a ⊓ b = b ↔ b ≤ a := -le_antisymm_iff.trans $ by simp [le_refl] +attribute [simp] inf_of_le_left inf_of_le_right +@[simp] theorem inf_lt_left : a ⊓ b < a ↔ ¬a ≤ b := @left_lt_sup αᵒᵈ _ _ _ @[simp] theorem inf_lt_right : a ⊓ b < b ↔ ¬b ≤ a := @right_lt_sup αᵒᵈ _ _ _ theorem inf_lt_left_or_right (h : a ≠ b) : a ⊓ b < a ∨ a ⊓ b < b := @left_or_right_lt_sup αᵒᵈ _ _ _ h -theorem inf_of_le_right (h : b ≤ a) : a ⊓ b = b := -inf_eq_right.2 h - -@[simp] theorem right_eq_inf : b = a ⊓ b ↔ b ≤ a := -eq_comm.trans inf_eq_right - theorem inf_le_inf (h₁ : a ≤ b) (h₂ : c ≤ d) : a ⊓ c ≤ b ⊓ d := @sup_le_sup αᵒᵈ _ _ _ _ _ h₁ h₂ @@ -363,8 +343,6 @@ lemma inf_le_inf_right (a : α) {b c : α} (h : b ≤ c) : b ⊓ a ≤ c ⊓ a : lemma inf_le_inf_left (a : α) {b c : α} (h : b ≤ c) : a ⊓ b ≤ a ⊓ c := inf_le_inf le_rfl h -theorem le_of_inf_eq (h : a ⊓ b = a) : a ≤ b := inf_eq_left.1 h - @[simp] lemma inf_idem : a ⊓ a = a := @sup_idem αᵒᵈ _ _ instance inf_is_idempotent : is_idempotent α (⊓) := ⟨@inf_idem _ _⟩ @@ -526,6 +504,10 @@ begin le_sup_right.trans (Heq.symm.trans_le inf_le_left)] } end +@[simp] lemma sup_le_inf : a ⊔ b ≤ a ⊓ b ↔ a = b := +⟨λ h, le_antisymm (le_sup_left.trans $ h.trans inf_le_right) + (le_sup_right.trans $ h.trans inf_le_left), by { rintro rfl, simp }⟩ + /-! #### Distributivity laws -/ diff --git a/src/topology/uniform_space/compact_separated.lean b/src/topology/uniform_space/compact_separated.lean index e226a3b7e569e..d610588238d57 100644 --- a/src/topology/uniform_space/compact_separated.lean +++ b/src/topology/uniform_space/compact_separated.lean @@ -45,7 +45,7 @@ begin symmetry, refine le_antisymm supr_nhds_le_uniformity _, by_contra H, obtain ⟨V, hV, h⟩ : ∃ V : set (α × α), (∀ x : α, V ∈ 𝓝 (x, x)) ∧ 𝓤 α ⊓ 𝓟 Vᶜ ≠ ⊥, - { simpa [le_iff_forall_inf_principal_compl] using H }, + { simpa only [le_iff_forall_inf_principal_compl, mem_supr, not_forall, exists_prop] using H }, let F := 𝓤 α ⊓ 𝓟 Vᶜ, haveI : ne_bot F := ⟨h⟩, obtain ⟨⟨x, y⟩, hx⟩ : ∃ (p : α × α), cluster_pt p F := From a3e847c015c09648d01869b2446c940143ea502c Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Sun, 4 Sep 2022 18:07:36 +0000 Subject: [PATCH 155/828] feat(analysis/special_functions/stirling): Stirling's formula, part I (#14874) Part 1 - [x] depends on: #14881 Co-authored-by: fabian-kruse Co-authored-by: nick-kuhn <46423253+nick-kuhn@users.noreply.github.com> Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Co-authored-by: Moritz Firsching Co-authored-by: nick-kuhn <46423253+nick-kuhn@users.noreply.github.com> --- src/analysis/special_functions/stirling.lean | 205 +++++++++++++++++++ 1 file changed, 205 insertions(+) create mode 100644 src/analysis/special_functions/stirling.lean diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean new file mode 100644 index 0000000000000..5170c5e8af251 --- /dev/null +++ b/src/analysis/special_functions/stirling.lean @@ -0,0 +1,205 @@ +/- +Copyright (c) 2022. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Firsching, Fabian Kruse, Nikolas Kuhn +-/ +import analysis.p_series +import analysis.special_functions.log.deriv +import tactic.positivity + +/-! +# Stirling's formula + +This file proves Stirling's formula for the factorial. +It states that $n!$ grows asymptotically like $\sqrt{2\pi n}(\frac{n}{e})^n$. +TODO: Add Part 2 to complete the proof + +## Proof outline + +The proof follows: . + +### Part 1 +We consider the fraction sequence $a_n$ of fractions $\frac{n!}{\sqrt{2n}(\frac{n}{e})^n}$ and +prove that this sequence converges against a real, positive number $a$. For this the two main +ingredients are + - taking the logarithm of the sequence and + - use the series expansion of $\log(1 + x)$. +-/ + +open_locale topological_space big_operators +open finset filter nat real + +/-! + ### Part 1 + https://proofwiki.org/wiki/Stirling%27s_Formula#Part_1 +-/ + +/-- +Define `stirling_seq n` as $\frac{n!}{\sqrt{2n}/(\frac{n}{e})^n$. +Stirling's formula states that this sequence has limit $\sqrt(π)$. +-/ +noncomputable def stirling_seq (n : ℕ) : ℝ := +n.factorial / (sqrt (2 * n) * (n / exp 1) ^ n) + +@[simp] lemma stirling_seq_zero : stirling_seq 0 = 0 := +by rw [stirling_seq, cast_zero, mul_zero, real.sqrt_zero, zero_mul, div_zero] + +@[simp] lemma stirling_seq_one : stirling_seq 1 = exp 1 / sqrt 2 := +by rw [stirling_seq, pow_one, factorial_one, cast_one, mul_one, mul_one_div, one_div_div] + +/-- +We have the expression +`log (stirling_seq (n + 1)) = log(n + 1)! - 1 / 2 * log(2 * n) - n * log ((n + 1) / e)`. +-/ +lemma log_stirling_seq_formula (n : ℕ) : log (stirling_seq n.succ) = + log n.succ.factorial - 1 / 2 * log (2 * n.succ) - n.succ * log (n.succ / exp 1) := +begin + have h1 : (0 : ℝ) < n.succ.factorial := cast_pos.mpr n.succ.factorial_pos, + have h2 : (0 : ℝ) < (2 * n.succ) := mul_pos two_pos (cast_pos.mpr (succ_pos n)), + have h3 := real.sqrt_pos.mpr h2, + have h4 := pow_pos (div_pos (cast_pos.mpr n.succ_pos ) (exp_pos 1)) n.succ, + have h5 := mul_pos h3 h4, + rw [stirling_seq, log_div, log_mul, sqrt_eq_rpow, log_rpow, log_pow]; linarith, +end + +/-- +The sequence `log (stirling_seq (m + 1)) - log (stirling_seq (m + 2))` has the series expansion + `∑ 1 / (2 * (k + 1) + 1) * (1 / 2 * (m + 1) + 1)^(2 * (k + 1))` +-/ +lemma log_stirling_seq_diff_has_sum (m : ℕ) : + has_sum (λ k : ℕ, (1 : ℝ) / (2 * k.succ + 1) * ((1 / (2 * m.succ + 1)) ^ 2) ^ k.succ) + (log (stirling_seq m.succ) - log (stirling_seq m.succ.succ)) := +begin + change has_sum ((λ b : ℕ, 1 / (2 * (b : ℝ) + 1) * ((1 / (2 * m.succ + 1)) ^ 2) ^ b) ∘ succ) _, + refine (has_sum_nat_add_iff 1).mpr _, + convert (has_sum_log_one_add_inv $ cast_pos.mpr (succ_pos m)).mul_left ((m.succ : ℝ) + 1 / 2), + { ext k, + rw [← pow_mul, pow_add], + push_cast, + have : 2 * (k : ℝ) + 1 ≠ 0, {norm_cast, exact succ_ne_zero (2*k)}, + have : 2 * ((m : ℝ) + 1) + 1 ≠ 0, {norm_cast, exact succ_ne_zero (2*m.succ)}, + field_simp, + ring }, + { have h : ∀ (x : ℝ) (hx : x ≠ 0), 1 + x⁻¹ = (x + 1) / x, + { intros, rw [_root_.add_div, div_self hx, inv_eq_one_div], }, + simp only [log_stirling_seq_formula, log_div, log_mul, log_exp, factorial_succ, cast_mul, + cast_succ, cast_zero, range_one, sum_singleton, h] { discharger := + `[norm_cast, apply_rules [mul_ne_zero, succ_ne_zero, factorial_ne_zero, exp_ne_zero]] }, + ring }, +end + +/-- The sequence `log ∘ stirling_seq ∘ succ` is monotone decreasing -/ +lemma log_stirling_seq'_antitone : antitone (log ∘ stirling_seq ∘ succ) := +begin + have : ∀ {k : ℕ}, 0 < (1 : ℝ) / (2 * k.succ + 1) := + λ k, one_div_pos.mpr (add_pos (mul_pos two_pos (cast_pos.mpr k.succ_pos)) one_pos), + exact antitone_nat_of_succ_le (λ n, sub_nonneg.mp ((log_stirling_seq_diff_has_sum n).nonneg + (λ m, (mul_pos this (pow_pos (pow_pos this 2) m.succ)).le))), +end + +/-- +We have a bound for successive elements in the sequence `log (stirling_seq k)`. +-/ +lemma log_stirling_seq_diff_le_geo_sum (n : ℕ) : + log (stirling_seq n.succ) - log (stirling_seq n.succ.succ) ≤ + (1 / (2 * n.succ + 1)) ^ 2 / (1 - (1 / (2 * n.succ + 1)) ^ 2) := +begin + have h_nonneg : 0 ≤ ((1 / (2 * (n.succ : ℝ) + 1)) ^ 2) := sq_nonneg _, + have g : has_sum (λ k : ℕ, ((1 / (2 * (n.succ : ℝ) + 1)) ^ 2) ^ k.succ) + ((1 / (2 * n.succ + 1)) ^ 2 / (1 - (1 / (2 * n.succ + 1)) ^ 2)), + { refine (has_sum_geometric_of_lt_1 h_nonneg _).mul_left ((1 / (2 * (n.succ : ℝ) + 1)) ^ 2), + rw [one_div, inv_pow], + refine inv_lt_one (one_lt_pow ((lt_add_iff_pos_left 1).mpr + (mul_pos two_pos (cast_pos.mpr n.succ_pos))) two_ne_zero) }, + have hab : ∀ (k : ℕ), (1 / (2 * (k.succ : ℝ) + 1)) * ((1 / (2 * n.succ + 1)) ^ 2) ^ k.succ ≤ + ((1 / (2 * n.succ + 1)) ^ 2) ^ k.succ, + { refine λ k, mul_le_of_le_one_left (pow_nonneg h_nonneg k.succ) _, + rw one_div, + exact inv_le_one (le_add_of_nonneg_left (mul_pos two_pos (cast_pos.mpr k.succ_pos)).le) }, + exact has_sum_le hab (log_stirling_seq_diff_has_sum n) g, +end + +/-- +We have the bound `log (stirling_seq n) - log (stirling_seq (n+1))` ≤ 1/(4 n^2) +-/ +lemma log_stirling_seq_sub_log_stirling_seq_succ (n : ℕ) : + log (stirling_seq n.succ) - log (stirling_seq n.succ.succ) ≤ 1 / (4 * n.succ ^ 2) := +begin + have h₁ : 0 < 4 * ((n : ℝ) + 1) ^ 2 := by nlinarith [@cast_nonneg ℝ _ n], + have h₃ : 0 < (2 * ((n : ℝ) + 1) + 1) ^ 2 := by nlinarith [@cast_nonneg ℝ _ n], + have h₂ : 0 < 1 - (1 / (2 * ((n : ℝ) + 1) + 1)) ^ 2, + { rw ← mul_lt_mul_right h₃, + have H : 0 < (2 * ((n : ℝ) + 1) + 1) ^ 2 - 1 := by nlinarith [@cast_nonneg ℝ _ n], + convert H using 1; field_simp [h₃.ne'] }, + refine (log_stirling_seq_diff_le_geo_sum n).trans _, + push_cast, + rw div_le_div_iff h₂ h₁, + field_simp [h₃.ne'], + rw div_le_div_right h₃, + ring_nf, + norm_cast, + linarith, +end + +/-- For any `n`, we have `log_stirling_seq 1 - log_stirling_seq n ≤ 1/4 * ∑' 1/k^2` -/ +lemma log_stirling_seq_bounded_aux : + ∃ (c : ℝ), ∀ (n : ℕ), log (stirling_seq 1) - log (stirling_seq n.succ) ≤ c := +begin + let d := ∑' k : ℕ, (1 : ℝ) / k.succ ^ 2, + use (1 / 4 * d : ℝ), + let log_stirling_seq' : ℕ → ℝ := λ k, log (stirling_seq k.succ), + intro n, + have h₁ : ∀ k, log_stirling_seq' k - log_stirling_seq' (k + 1) ≤ 1 / 4 * (1 / k.succ ^ 2) := + by { intro k, convert log_stirling_seq_sub_log_stirling_seq_succ k using 1, field_simp, }, + have h₂ : ∑ (k : ℕ) in range n, (1 : ℝ) / (k.succ) ^ 2 ≤ d := by + { refine sum_le_tsum (range n) (λ k _, _) + ((summable_nat_add_iff 1).mpr (real.summable_one_div_nat_pow.mpr one_lt_two)), + apply le_of_lt, + rw [one_div_pos, sq_pos_iff], + exact nonzero_of_invertible (succ k), }, + calc + log (stirling_seq 1) - log (stirling_seq n.succ) = log_stirling_seq' 0 - log_stirling_seq' n : rfl + ... = ∑ k in range n, (log_stirling_seq' k - log_stirling_seq' (k + 1)) : by + rw ← sum_range_sub' log_stirling_seq' n + ... ≤ ∑ k in range n, (1/4) * (1 / k.succ^2) : sum_le_sum (λ k _, h₁ k) + ... = 1 / 4 * ∑ k in range n, 1 / k.succ ^ 2 : by rw mul_sum + ... ≤ 1 / 4 * d : (mul_le_mul_left (one_div_pos.mpr four_pos)).mpr h₂, +end + +/-- The sequence `log_stirling_seq` is bounded below for `n ≥ 1`. -/ +lemma log_stirling_seq_bounded_by_constant : ∃ c, ∀ (n : ℕ), c ≤ log (stirling_seq n.succ) := +begin + obtain ⟨d, h⟩ := log_stirling_seq_bounded_aux, + exact ⟨log (stirling_seq 1) - d, λ n, sub_le.mp (h n)⟩, +end + +/-- The sequence `stirling_seq` is positive for `n > 0` -/ +lemma stirling_seq'_pos (n : ℕ) : 0 < stirling_seq n.succ := +div_pos (cast_pos.mpr n.succ.factorial_pos) (mul_pos (real.sqrt_pos.mpr (mul_pos two_pos + (cast_pos.mpr n.succ_pos))) (pow_pos (div_pos (cast_pos.mpr n.succ_pos) (exp_pos 1)) n.succ)) + +/-- +The sequence `stirling_seq` has a positive lower bound. +-/ +lemma stirling_seq'_bounded_by_pos_constant : ∃ a, 0 < a ∧ ∀ n : ℕ, a ≤ stirling_seq n.succ := +begin + cases log_stirling_seq_bounded_by_constant with c h, + refine ⟨exp c, exp_pos _, λ n, _⟩, + rw ← le_log_iff_exp_le (stirling_seq'_pos n), + exact h n, +end + +/-- The sequence `stirling_seq ∘ succ` is monotone decreasing -/ +lemma stirling_seq'_antitone : antitone (stirling_seq ∘ succ) := +λ n m h, (log_le_log (stirling_seq'_pos m) (stirling_seq'_pos n)).mp (log_stirling_seq'_antitone h) + +/-- The limit `a` of the sequence `stirling_seq` satisfies `0 < a` -/ +lemma stirling_seq_has_pos_limit_a : + ∃ (a : ℝ), 0 < a ∧ tendsto stirling_seq at_top (𝓝 a) := +begin + obtain ⟨x, x_pos, hx⟩ := stirling_seq'_bounded_by_pos_constant, + have hx' : x ∈ lower_bounds (set.range (stirling_seq ∘ succ)) := by simpa [lower_bounds] using hx, + refine ⟨_, lt_of_lt_of_le x_pos (le_cInf (set.range_nonempty _) hx'), _⟩, + rw ←filter.tendsto_add_at_top_iff_nat 1, + exact tendsto_at_top_cinfi stirling_seq'_antitone ⟨x, hx'⟩, +end From 075b0d232669c13431aa4f4e95258c4ad525bc68 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Sun, 4 Sep 2022 19:09:59 +0000 Subject: [PATCH 156/828] doc(data/nat/{part_enat, lattice}): update doc string (#16345) --- src/data/nat/lattice.lean | 1 - src/data/nat/part_enat.lean | 1 + 2 files changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/nat/lattice.lean b/src/data/nat/lattice.lean index c21fe473aff37..cdc59f7f06f1c 100644 --- a/src/data/nat/lattice.lean +++ b/src/data/nat/lattice.lean @@ -11,7 +11,6 @@ import order.conditionally_complete_lattice In this file we * define a `conditionally_complete_linear_order_bot` structure on `ℕ`; -* define a `complete_linear_order` structure on `part_enat`; * prove a few lemmas about `supr`/`infi`/`set.Union`/`set.Inter` and natural numbers. -/ diff --git a/src/data/nat/part_enat.lean b/src/data/nat/part_enat.lean index 2c118d87571cf..9944676a33021 100644 --- a/src/data/nat/part_enat.lean +++ b/src/data/nat/part_enat.lean @@ -20,6 +20,7 @@ The following instances are defined: * `ordered_add_comm_monoid part_enat` * `canonically_ordered_add_monoid part_enat` +* `complete_linear_order part_enat` There is no additive analogue of `monoid_with_zero`; if there were then `part_enat` could be an `add_monoid_with_top`. From 86ee668979dacd95cca6b1ba4efe27b8d13d569b Mon Sep 17 00:00:00 2001 From: Mauricio Collares Date: Sun, 4 Sep 2022 22:22:14 +0000 Subject: [PATCH 157/828] fix(algebra/category/Module): speedup (#16371) This change makes `map'` go from using 95,000+ heartbeats (that is, `lean --make -T95000` fails locally) to using less than 25,000. With `set_option pp.all true`, the output of `#print map'` is the same before and after this change, but note that this is a `def` so perhaps using the `apply` tactic is not appropriate. --- src/algebra/category/Module/change_of_rings.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/src/algebra/category/Module/change_of_rings.lean b/src/algebra/category/Module/change_of_rings.lean index 747bcec4a2fb2..f09acdc2a3648 100644 --- a/src/algebra/category/Module/change_of_rings.lean +++ b/src/algebra/category/Module/change_of_rings.lean @@ -110,7 +110,8 @@ Extension of scalars is a functor where an `R`-module `M` is sent to `S ⊗ M` a `l : M1 ⟶ M2` is sent to `s ⊗ m ↦ s ⊗ l m` -/ def map' {M1 M2 : Module.{v} R} (l : M1 ⟶ M2) : (obj' f M1) ⟶ (obj' f M2) := -@linear_map.base_change R S M1 M2 _ _ ((algebra_map S _).comp f).to_algebra _ _ _ _ l +-- The "by apply" part makes this require 75% fewer heartbeats to process (#16371). +by apply (@linear_map.base_change R S M1 M2 _ _ ((algebra_map S _).comp f).to_algebra _ _ _ _ l) lemma map'_id {M : Module.{v} R} : map' f (𝟙 M) = 𝟙 _ := linear_map.ext $ λ (x : obj' f M), From 7ce2bb4a707187c85538a650c94a23701e98d56d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 5 Sep 2022 00:43:00 +0000 Subject: [PATCH 158/828] =?UTF-8?q?feat(order/filter/at=5Ftop=5Fbot):=20ad?= =?UTF-8?q?d=20lemmas=20about=20convergence=20of=20`=CE=BB=20x,=20r=20*=20?= =?UTF-8?q?f=20x`=20to=20`=C2=B1=E2=88=9E`=20(#16355)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Other API changes: * assume `m ≠ 0` instead of `1 ≤ m` in `le_self_pow`; * generalize `linear_ordered_semiring.to_no_max_order` to `ordered_semiring.to_no_max_order`. --- src/algebra/group_power/order.lean | 4 +- src/algebra/order/ring.lean | 9 +- src/data/nat/factorization/basic.lean | 2 +- src/order/filter/at_top_bot.lean | 170 +++++++++++++++++++++++--- 4 files changed, 161 insertions(+), 24 deletions(-) diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index 0580656f20723..0935f8a31ec8c 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -251,8 +251,8 @@ monotone_nat_of_le_succ $ λ n, theorem pow_le_pow (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := pow_mono ha h -theorem le_self_pow (ha : 1 ≤ a) (h : 1 ≤ m) : a ≤ a ^ m := -eq.trans_le (pow_one a).symm (pow_le_pow ha h) +theorem le_self_pow (ha : 1 ≤ a) (h : m ≠ 0) : a ≤ a ^ m := +(pow_one a).symm.trans_le (pow_le_pow ha $ pos_iff_ne_zero.mpr h) lemma strict_mono_pow (h : 1 < a) : strict_mono (λ n : ℕ, a ^ n) := have 0 < a := zero_le_one.trans_lt h, diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 2df672b7e01a7..a22119e0dd2b5 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -163,6 +163,10 @@ alias zero_lt_two ← two_pos alias zero_lt_three ← three_pos alias zero_lt_four ← four_pos +@[priority 100] -- see Note [lower instance priority] +instance ordered_semiring.to_no_max_order : no_max_order α := +⟨assume a, ⟨a + 1, lt_add_of_pos_right _ one_pos⟩⟩ + end nontrivial -- See Note [decidable namespace] @@ -726,11 +730,6 @@ le_of_not_gt (λ ha, absurd h (mul_neg_of_pos_of_neg ha hb).not_le) lemma nonpos_of_mul_nonneg_right (h : 0 ≤ a * b) (ha : a < 0) : b ≤ 0 := le_of_not_gt (λ hb, absurd h (mul_neg_of_neg_of_pos ha hb).not_le) -@[priority 100] -- see Note [lower instance priority] -instance linear_ordered_semiring.to_no_max_order {α : Type*} [linear_ordered_semiring α] : - no_max_order α := -⟨assume a, ⟨a + 1, lt_add_of_pos_right _ zero_lt_one⟩⟩ - /-- Pullback a `linear_ordered_semiring` under an injective map. See note [reducible non-instances]. -/ @[reducible] diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 8913a68033644..d9ee0db5b768a 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -720,7 +720,7 @@ begin by_cases ha1 : a = 1, { rw [ha1, mul_one], exact hp p n hp' hn }, - refine h (p^n) a ((hp'.one_lt).trans_le (le_self_pow (prime.one_lt hp').le (succ_le_iff.mpr hn))) + refine h (p^n) a ((hp'.one_lt).trans_le (le_self_pow (prime.one_lt hp').le hn.ne')) _ _ (hp _ _ hp' hn) hPa, { contrapose! hpa, simp [lt_one_iff.1 (lt_of_le_of_ne hpa ha1)] }, diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index 19d23a18ec819..ed7c0db7127a0 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -682,6 +682,8 @@ lemma tendsto_neg_at_top_at_bot : tendsto (has_neg.neg : β → β) at_top at_bo lemma tendsto_neg_at_bot_at_top : tendsto (has_neg.neg : β → β) at_bot at_top := @tendsto_neg_at_top_at_bot βᵒᵈ _ +variable {l} + @[simp] lemma tendsto_neg_at_top_iff : tendsto (λ x, -f x) l at_top ↔ tendsto f l at_bot := (order_iso.neg β).tendsto_at_bot_iff @@ -711,11 +713,7 @@ tendsto_id.at_top_mul_at_top tendsto_id /-- The monomial function `x^n` tends to `+∞` at `+∞` for any positive natural `n`. A version for positive real powers exists as `tendsto_rpow_at_top`. -/ lemma tendsto_pow_at_top {n : ℕ} (hn : n ≠ 0) : tendsto (λ x : α, x ^ n) at_top at_top := -begin - rw [← pos_iff_ne_zero, ← nat.succ_le_iff] at hn, - refine tendsto_at_top_mono' _ ((eventually_ge_at_top 1).mono $ λ x hx, _) tendsto_id, - simpa only [pow_one] using pow_le_pow hx hn -end +tendsto_at_top_mono' _ ((eventually_ge_at_top 1).mono $ λ x hx, le_self_pow hx hn) tendsto_id end ordered_semiring @@ -782,29 +780,70 @@ lemma tendsto.at_top_of_mul_const {c : α} (hc : 0 < c) (hf : tendsto (λ x, f x tendsto f l at_top := tendsto_at_top.2 $ λ b, (tendsto_at_top.1 hf (b * c)).mono $ λ x hx, le_of_mul_le_mul_right hx hc +@[simp] lemma tendsto_pow_at_top_iff {n : ℕ} : tendsto (λ x : α, x ^ n) at_top at_top ↔ n ≠ 0 := +⟨λ h hn, by simpa only [hn, pow_zero, not_tendsto_const_at_top] using h, tendsto_pow_at_top⟩ + end linear_ordered_semiring lemma nonneg_of_eventually_pow_nonneg [linear_ordered_ring α] {a : α} (h : ∀ᶠ n in at_top, 0 ≤ a ^ (n : ℕ)) : 0 ≤ a := let ⟨n, hn⟩ := (tendsto_bit1_at_top.eventually h).exists in pow_bit1_nonneg_iff.1 hn +lemma not_tendsto_pow_at_top_at_bot [linear_ordered_ring α] : + ∀ {n : ℕ}, ¬tendsto (λ x : α, x ^ n) at_top at_bot +| 0 := by simp [not_tendsto_const_at_bot] +| (n + 1) := (tendsto_pow_at_top n.succ_ne_zero).not_tendsto disjoint_at_top_at_bot + section linear_ordered_semifield variables [linear_ordered_semifield α] {l : filter β} {f : β → α} {r c : α} {n : ℕ} +/-! +### Multiplication by constant: iff lemmas +-/ + +/-- If `r` is a positive constant, then `λ x, r * f x` tends to infinity along a filter if and only +if `f` tends to infinity along the same filter. -/ +lemma tendsto_const_mul_at_top_of_pos (hr : 0 < r) : + tendsto (λ x, r * f x) l at_top ↔ tendsto f l at_top := +⟨λ h, h.at_top_of_const_mul hr, + λ h, tendsto.at_top_of_const_mul (inv_pos.2 hr) $ by simpa only [inv_mul_cancel_left₀ hr.ne']⟩ + +/-- If `r` is a positive constant, then `λ x, f x * r` tends to infinity along a filter if and only +if `f` tends to infinity along the same filter. -/ +lemma tendsto_mul_const_at_top_of_pos (hr : 0 < r) : + tendsto (λ x, f x * r) l at_top ↔ tendsto f l at_top := +by simpa only [mul_comm] using tendsto_const_mul_at_top_of_pos hr + +/-- If `f` tends to infinity along a nontrivial filter `l`, then `λ x, r * f x` tends to infinity +if and only if `0 < r. `-/ +lemma tendsto_const_mul_at_top_iff_pos [ne_bot l] (h : tendsto f l at_top) : + tendsto (λ x, r * f x) l at_top ↔ 0 < r := +begin + refine ⟨λ hrf, not_le.mp $ λ hr, _, λ hr, (tendsto_const_mul_at_top_of_pos hr).mpr h⟩, + rcases ((h.eventually_ge_at_top 0).and (hrf.eventually_gt_at_top 0)).exists with ⟨x, hx, hrx⟩, + exact (mul_nonpos_of_nonpos_of_nonneg hr hx).not_lt hrx +end + +/-- If `f` tends to infinity along a nontrivial filter `l`, then `λ x, f x * r` tends to infinity +if and only if `0 < r. `-/ +lemma tendsto_mul_const_at_top_iff_pos [ne_bot l] (h : tendsto f l at_top) : + tendsto (λ x, f x * r) l at_top ↔ 0 < r := +by simp only [mul_comm _ r, tendsto_const_mul_at_top_iff_pos h] + /-- If a function tends to infinity along a filter, then this function multiplied by a positive constant (on the left) also tends to infinity. For a version working in `ℕ` or `ℤ`, use `filter.tendsto.const_mul_at_top'` instead. -/ lemma tendsto.const_mul_at_top (hr : 0 < r) (hf : tendsto f l at_top) : tendsto (λx, r * f x) l at_top := -tendsto.at_top_of_const_mul (inv_pos.2 hr) $ by simpa only [inv_mul_cancel_left₀ hr.ne'] +(tendsto_const_mul_at_top_of_pos hr).2 hf /-- If a function tends to infinity along a filter, then this function multiplied by a positive constant (on the right) also tends to infinity. For a version working in `ℕ` or `ℤ`, use `filter.tendsto.at_top_mul_const'` instead. -/ lemma tendsto.at_top_mul_const (hr : 0 < r) (hf : tendsto f l at_top) : tendsto (λx, f x * r) l at_top := -by simpa only [mul_comm] using hf.const_mul_at_top hr +(tendsto_mul_const_at_top_of_pos hr).2 hf /-- If a function tends to infinity along a filter, then this function divided by a positive constant also tends to infinity. -/ @@ -831,31 +870,131 @@ end linear_ordered_semifield section linear_ordered_field variables [linear_ordered_field α] {l : filter β} {f : β → α} {r : α} + +/-- If `r` is a positive constant, then `λ x, r * f x` tends to negative infinity along a filter if +and only if `f` tends to negative infinity along the same filter. -/ +lemma tendsto_const_mul_at_bot_of_pos (hr : 0 < r) : + tendsto (λ x, r * f x) l at_bot ↔ tendsto f l at_bot := +by simpa only [← mul_neg, ← tendsto_neg_at_top_iff] using tendsto_const_mul_at_top_of_pos hr + +/-- If `r` is a positive constant, then `λ x, f x * r` tends to negative infinity along a filter if +and only if `f` tends to negative infinity along the same filter. -/ +lemma tendsto_mul_const_at_bot_of_pos (hr : 0 < r) : + tendsto (λ x, f x * r) l at_bot ↔ tendsto f l at_bot := +by simpa only [mul_comm] using tendsto_const_mul_at_bot_of_pos hr + +/-- If `r` is a negative constant, then `λ x, r * f x` tends to infinity along a filter if and only +if `f` tends to negative infinity along the same filter. -/ +lemma tendsto_const_mul_at_top_of_neg (hr : r < 0) : + tendsto (λ x, r * f x) l at_top ↔ tendsto f l at_bot := +by simpa only [neg_mul, tendsto_neg_at_bot_iff] using tendsto_const_mul_at_bot_of_pos (neg_pos.2 hr) + +/-- If `r` is a negative constant, then `λ x, f x * r` tends to infinity along a filter if and only +if `f` tends to negative infinity along the same filter. -/ +lemma tendsto_mul_const_at_top_of_neg (hr : r < 0) : + tendsto (λ x, f x * r) l at_top ↔ tendsto f l at_bot := +by simpa only [mul_comm] using tendsto_const_mul_at_top_of_neg hr + +/-- If `r` is a negative constant, then `λ x, r * f x` tends to negative infinity along a filter if +and only if `f` tends to infinity along the same filter. -/ +lemma tendsto_const_mul_at_bot_of_neg (hr : r < 0) : + tendsto (λ x, r * f x) l at_bot ↔ tendsto f l at_top := +by simpa only [neg_mul, tendsto_neg_at_top_iff] using tendsto_const_mul_at_top_of_pos (neg_pos.2 hr) + +/-- If `r` is a negative constant, then `λ x, f x * r` tends to negative infinity along a filter if +and only if `f` tends to infinity along the same filter. -/ +lemma tendsto_mul_const_at_bot_of_neg (hr : r < 0) : + tendsto (λ x, f x * r) l at_bot ↔ tendsto f l at_top := +by simpa only [mul_comm] using tendsto_const_mul_at_bot_of_neg hr + +/-- The function `λ x, r * f x` tends to infinity along a nontrivial filter if and only if `r > 0` +and `f` tends to infinity or `r < 0` and `f` tends to negative infinity. -/ +lemma tendsto_const_mul_at_top_iff [ne_bot l] : + tendsto (λ x, r * f x) l at_top ↔ 0 < r ∧ tendsto f l at_top ∨ r < 0 ∧ tendsto f l at_bot := +begin + rcases lt_trichotomy r 0 with hr|rfl|hr, + { simp [hr, hr.not_lt, tendsto_const_mul_at_top_of_neg] }, + { simp [not_tendsto_const_at_top] }, + { simp [hr, hr.not_lt, tendsto_const_mul_at_top_of_pos] } +end + +/-- The function `λ x, f x * r` tends to infinity along a nontrivial filter if and only if `r > 0` +and `f` tends to infinity or `r < 0` and `f` tends to negative infinity. -/ +lemma tendsto_mul_const_at_top_iff [ne_bot l] : + tendsto (λ x, f x * r) l at_top ↔ 0 < r ∧ tendsto f l at_top ∨ r < 0 ∧ tendsto f l at_bot := +by simp only [mul_comm _ r, tendsto_const_mul_at_top_iff] + +/-- The function `λ x, r * f x` tends to negative infinity along a nontrivial filter if and only if +`r > 0` and `f` tends to negative infinity or `r < 0` and `f` tends to infinity. -/ +lemma tendsto_const_mul_at_bot_iff [ne_bot l] : + tendsto (λ x, r * f x) l at_bot ↔ 0 < r ∧ tendsto f l at_bot ∨ r < 0 ∧ tendsto f l at_top := +by simp only [← tendsto_neg_at_top_iff, ← mul_neg, tendsto_const_mul_at_top_iff, neg_neg] + +/-- The function `λ x, f x * r` tends to negative infinity along a nontrivial filter if and only if +`r > 0` and `f` tends to negative infinity or `r < 0` and `f` tends to infinity. -/ +lemma tendsto_mul_const_at_bot_iff [ne_bot l] : + tendsto (λ x, f x * r) l at_bot ↔ 0 < r ∧ tendsto f l at_bot ∨ r < 0 ∧ tendsto f l at_top := +by simp only [mul_comm _ r, tendsto_const_mul_at_bot_iff] + +/-- If `f` tends to negative infinity along a nontrivial filter `l`, then `λ x, r * f x` tends to +infinity if and only if `r < 0. `-/ +lemma tendsto_const_mul_at_top_iff_neg [ne_bot l] (h : tendsto f l at_bot) : + tendsto (λ x, r * f x) l at_top ↔ r < 0 := +by simp [tendsto_const_mul_at_top_iff, h, h.not_tendsto disjoint_at_bot_at_top] + +/-- If `f` tends to negative infinity along a nontrivial filter `l`, then `λ x, f x * r` tends to +infinity if and only if `r < 0. `-/ +lemma tendsto_mul_const_at_top_iff_neg [ne_bot l] (h : tendsto f l at_bot) : + tendsto (λ x, f x * r) l at_top ↔ r < 0 := +by simp only [mul_comm _ r, tendsto_const_mul_at_top_iff_neg h] + +/-- If `f` tends to negative infinity along a nontrivial filter `l`, then `λ x, r * f x` tends to +negative infinity if and only if `0 < r. `-/ +lemma tendsto_const_mul_at_bot_iff_pos [ne_bot l] (h : tendsto f l at_bot) : + tendsto (λ x, r * f x) l at_bot ↔ 0 < r := +by simp [tendsto_const_mul_at_bot_iff, h, h.not_tendsto disjoint_at_bot_at_top] + +/-- If `f` tends to negative infinity along a nontrivial filter `l`, then `λ x, f x * r` tends to +negative infinity if and only if `0 < r. `-/ +lemma tendsto_mul_const_at_bot_iff_pos [ne_bot l] (h : tendsto f l at_bot) : + tendsto (λ x, f x * r) l at_bot ↔ 0 < r := +by simp only [mul_comm _ r, tendsto_const_mul_at_bot_iff_pos h] + +/-- If `f` tends to infinity along a nontrivial filter `l`, then `λ x, r * f x` tends to negative +infinity if and only if `r < 0. `-/ +lemma tendsto_const_mul_at_bot_iff_neg [ne_bot l] (h : tendsto f l at_top) : + tendsto (λ x, r * f x) l at_bot ↔ r < 0 := +by simp [tendsto_const_mul_at_bot_iff, h, h.not_tendsto disjoint_at_top_at_bot] + +/-- If `f` tends to infinity along a nontrivial filter `l`, then `λ x, f x * r` tends to negative +infinity if and only if `r < 0. `-/ +lemma tendsto_mul_const_at_bot_iff_neg [ne_bot l] (h : tendsto f l at_top) : + tendsto (λ x, f x * r) l at_bot ↔ r < 0 := +by simp only [mul_comm _ r, tendsto_const_mul_at_bot_iff_neg h] + /-- If a function tends to infinity along a filter, then this function multiplied by a negative constant (on the left) tends to negative infinity. -/ lemma tendsto.neg_const_mul_at_top (hr : r < 0) (hf : tendsto f l at_top) : tendsto (λ x, r * f x) l at_bot := -by simpa only [(∘), neg_mul_eq_neg_mul, neg_neg] - using tendsto_neg_at_top_at_bot.comp (hf.const_mul_at_top (neg_pos.2 hr)) +(tendsto_const_mul_at_bot_of_neg hr).2 hf /-- If a function tends to infinity along a filter, then this function multiplied by a negative constant (on the right) tends to negative infinity. -/ lemma tendsto.at_top_mul_neg_const (hr : r < 0) (hf : tendsto f l at_top) : tendsto (λ x, f x * r) l at_bot := -by simpa only [mul_comm] using hf.neg_const_mul_at_top hr +(tendsto_mul_const_at_bot_of_neg hr).2 hf /-- If a function tends to negative infinity along a filter, then this function multiplied by a positive constant (on the left) also tends to negative infinity. -/ lemma tendsto.const_mul_at_bot (hr : 0 < r) (hf : tendsto f l at_bot) : tendsto (λx, r * f x) l at_bot := -by simpa only [(∘), neg_mul_eq_mul_neg, neg_neg] - using tendsto_neg_at_top_at_bot.comp ((tendsto_neg_at_bot_at_top.comp hf).const_mul_at_top hr) +(tendsto_const_mul_at_bot_of_pos hr).2 hf /-- If a function tends to negative infinity along a filter, then this function multiplied by a positive constant (on the right) also tends to negative infinity. -/ lemma tendsto.at_bot_mul_const (hr : 0 < r) (hf : tendsto f l at_bot) : tendsto (λx, f x * r) l at_bot := -by simpa only [mul_comm] using hf.const_mul_at_bot hr +(tendsto_mul_const_at_bot_of_pos hr).2 hf /-- If a function tends to negative infinity along a filter, then this function divided by a positive constant also tends to negative infinity. -/ @@ -867,14 +1006,13 @@ by simpa only [div_eq_mul_inv] using hf.at_bot_mul_const (inv_pos.2 hr) a negative constant (on the left) tends to positive infinity. -/ lemma tendsto.neg_const_mul_at_bot (hr : r < 0) (hf : tendsto f l at_bot) : tendsto (λ x, r * f x) l at_top := -by simpa only [(∘), neg_mul_eq_neg_mul, neg_neg] - using tendsto_neg_at_bot_at_top.comp (hf.const_mul_at_bot (neg_pos.2 hr)) +(tendsto_const_mul_at_top_of_neg hr).2 hf /-- If a function tends to negative infinity along a filter, then this function multiplied by a negative constant (on the right) tends to positive infinity. -/ lemma tendsto.at_bot_mul_neg_const (hr : r < 0) (hf : tendsto f l at_bot) : tendsto (λ x, f x * r) l at_top := -by simpa only [mul_comm] using hf.neg_const_mul_at_bot hr +(tendsto_mul_const_at_top_of_neg hr).2 hf lemma tendsto_neg_const_mul_pow_at_top {c : α} {n : ℕ} (hn : n ≠ 0) (hc : c < 0) : tendsto (λ x, c * x^n) at_top at_bot := From 7fdc06de176a13a783727a44b8b96fb0e7273aa1 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Mon, 5 Sep 2022 03:19:00 +0000 Subject: [PATCH 159/828] chore(data/nat/part_enat): make `coe_inj` `@[norm_cast]` (#16334) --- src/data/nat/part_enat.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/nat/part_enat.lean b/src/data/nat/part_enat.lean index 9944676a33021..a808fd603f7b4 100644 --- a/src/data/nat/part_enat.lean +++ b/src/data/nat/part_enat.lean @@ -87,7 +87,7 @@ instance : add_monoid_with_one part_enat := lemma some_eq_coe (n : ℕ) : some n = n := rfl -@[simp] lemma coe_inj {x y : ℕ} : (x : part_enat) = y ↔ x = y := part.some_inj +@[simp, norm_cast] lemma coe_inj {x y : ℕ} : (x : part_enat) = y ↔ x = y := part.some_inj @[simp] lemma dom_coe (x : ℕ) : (x : part_enat).dom := trivial From e1d2f5bdbb8f613edd7cb11e12bcd4d55e6c16b9 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 5 Sep 2022 03:19:01 +0000 Subject: [PATCH 160/828] feat(geometry/euclidean/basic): `affine_isometry.angle_map` (#16374) Add the lemma that angles between three points are preserved by affine isometries, analogous to the one that already exists that angles between two vectors are preserved by linear isometries. --- src/geometry/euclidean/basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index dc331ea86724b..2fd08526903dd 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -407,6 +407,11 @@ begin (continuous_snd.snd.vsub continuous_snd.fst)).continuous_at end +@[simp] lemma _root_.affine_isometry.angle_map {V₂ P₂ : Type*} [inner_product_space ℝ V₂] + [metric_space P₂] [normed_add_torsor V₂ P₂] (f : P →ᵃⁱ[ℝ] P₂) (p₁ p₂ p₃ : P) : + ∠ (f p₁) (f p₂) (f p₃) = ∠ p₁ p₂ p₃ := +by simp_rw [angle, ←affine_isometry.map_vsub, linear_isometry.angle_map] + /-- The angle at a point does not depend on the order of the other two points. -/ lemma angle_comm (p1 p2 p3 : P) : ∠ p1 p2 p3 = ∠ p3 p2 p1 := From 9b11656c0355e057d739c644878cc1d62c47f1b7 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Mon, 5 Sep 2022 05:16:33 +0000 Subject: [PATCH 161/828] chore(scripts): update nolints.txt (#16384) I am happy to remove some nolints for you! --- scripts/nolints.txt | 70 --------------------------------------------- 1 file changed, 70 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 35424dd9601a4..39e90c94e7f75 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -423,82 +423,16 @@ apply_nolint subgroup.card_subgroup_dvd_card to_additive_doc -- group_theory/group_action/sub_mul_action.lean apply_nolint sub_mul_action.has_zero fails_quickly --- linear_algebra/affine_space/affine_subspace.lean -apply_nolint affine_span.nonempty fails_quickly -apply_nolint affine_subspace.to_add_torsor fails_quickly - -- linear_algebra/affine_space/basis.lean apply_nolint affine_basis.affine_independent_of_to_matrix_right_inv fintype_finite apply_nolint affine_basis.affine_span_eq_top_of_to_matrix_left_inv fintype_finite apply_nolint affine_basis.ext_elem fintype_finite --- linear_algebra/affine_space/finite_dimensional.lean -apply_nolint finite_dimensional_direction_affine_span_image_of_fintype fintype_finite -apply_nolint finite_dimensional_direction_affine_span_of_fintype fintype_finite -apply_nolint finite_dimensional_vector_span_image_of_fintype fintype_finite -apply_nolint finite_dimensional_vector_span_of_fintype fintype_finite - --- linear_algebra/alternating.lean -apply_nolint basis.ext_alternating fintype_finite - --- linear_algebra/determinant.lean -apply_nolint alternating_map.map_basis_eq_zero_iff fintype_finite -apply_nolint alternating_map.map_basis_ne_zero_iff fintype_finite -apply_nolint linear_map.det_zero' fintype_finite - --- linear_algebra/dimension.lean -apply_nolint dim_pi fintype_finite - --- linear_algebra/dual.lean -apply_nolint basis.dual_dim_eq fintype_finite -apply_nolint basis.eval_range fintype_finite -apply_nolint basis.to_dual_range fintype_finite -apply_nolint basis.total_coord fintype_finite - --- linear_algebra/finsupp_vector_space.lean -apply_nolint cardinal_lt_aleph_0_of_finite_dimensional fintype_finite - --- linear_algebra/free_module/finite/basic.lean -apply_nolint module.finite.matrix fintype_finite -apply_nolint module.finite.of_basis fintype_finite - -- linear_algebra/free_module/pid.lean -apply_nolint ideal.exists_smith_normal_form fintype_finite apply_nolint submodule.basis_of_pid_aux fintype_finite apply_nolint submodule.exists_smith_normal_form_of_le fintype_finite apply_nolint submodule.nonempty_basis_of_pid fintype_finite --- linear_algebra/free_module/rank.lean -apply_nolint module.free.rank_matrix fintype_finite -apply_nolint module.free.rank_matrix' fintype_finite -apply_nolint module.free.rank_matrix'' fintype_finite -apply_nolint module.free.rank_pi_fintype fintype_finite - --- linear_algebra/matrix/finite_dimensional.lean -apply_nolint matrix.finite_dimensional fintype_finite - --- linear_algebra/matrix/reindex.lean -apply_nolint matrix.mul_reindex_linear_equiv_one fintype_finite - --- linear_algebra/matrix/transvection.lean -apply_nolint matrix.update_row_eq_transvection fintype_finite - --- linear_algebra/multilinear/basis.lean -apply_nolint basis.ext_multilinear fintype_finite - --- linear_algebra/multilinear/finite_dimensional.lean -apply_nolint module.finite.multilinear_map fintype_finite -apply_nolint module.free.multilinear_map fintype_finite - --- linear_algebra/orientation.lean -apply_nolint basis.map_orientation_eq_det_inv_smul fintype_finite - --- linear_algebra/std_basis.lean -apply_nolint linear_map.supr_range_std_basis fintype_finite - --- linear_algebra/trace.lean -apply_nolint linear_map.trace_eq_contract_of_basis fintype_finite - -- logic/relator.lean apply_nolint relator.bi_total doc_blame apply_nolint relator.bi_unique doc_blame @@ -533,10 +467,6 @@ apply_nolint tactic.coinductive_predicate doc_blame apply_nolint tactic.interactive.coinduction doc_blame apply_nolint tactic.mono doc_blame --- model_theory/definability.lean -apply_nolint set.definable.image_comp fintype_finite -apply_nolint set.definable.image_comp_embedding fintype_finite - -- model_theory/direct_limit.lean apply_nolint first_order.language.direct_limit.exists_quotient_mk_sigma_mk_eq fintype_finite apply_nolint first_order.language.direct_limit.exists_unify_eq fintype_finite From dc8113b073382b356fa6b9b967c3eabc4ce49e03 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 5 Sep 2022 06:12:36 +0000 Subject: [PATCH 162/828] feat(linear_algebra/ray): `same_ray_map_iff` for injective linear maps (#16377) Add a stronger version of `same_ray_map_iff`, for any injective linear map rather than just a linear equivalence. --- src/linear_algebra/ray.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/src/linear_algebra/ray.lean b/src/linear_algebra/ray.lean index 8ff35ef3b4b8b..d05573e8e4a44 100644 --- a/src/linear_algebra/ray.lean +++ b/src/linear_algebra/ray.lean @@ -126,10 +126,16 @@ lemma map (f : M →ₗ[R] N) (h : same_ray R x y) : same_ray R (f x) (f y) := h.imp (λ hx, by rw [hx, map_zero]) $ or.imp (λ hy, by rw [hy, map_zero]) $ λ ⟨r₁, r₂, hr₁, hr₂, h⟩, ⟨r₁, r₂, hr₁, hr₂, by rw [←f.map_smul, ←f.map_smul, h]⟩ +/-- The images of two vectors under an injective linear map are on the same ray if and only if the +original vectors are on the same ray. -/ +lemma _root_.function.injective.same_ray_map_iff {F : Type*} [linear_map_class F R M N] {f : F} + (hf : function.injective f) : same_ray R (f x) (f y) ↔ same_ray R x y := +by simp only [same_ray, map_zero, ← hf.eq_iff, map_smul] + /-- The images of two vectors under a linear equivalence are on the same ray if and only if the original vectors are on the same ray. -/ @[simp] lemma _root_.same_ray_map_iff (e : M ≃ₗ[R] N) : same_ray R (e x) (e y) ↔ same_ray R x y := -⟨λ h, by simpa using same_ray.map e.symm.to_linear_map h, same_ray.map e.to_linear_map⟩ +function.injective.same_ray_map_iff (equiv_like.injective e) /-- If two vectors are on the same ray then both scaled by the same action are also on the same ray. -/ From 44c90247a7438134e994c1d26ba8691e205b02e7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Mon, 5 Sep 2022 08:52:17 +0000 Subject: [PATCH 163/828] feat(probability/martingale/basic): the stopped process of a submartingale is a submartingale (#16375) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Rémy Degenne --- src/probability/martingale/basic.lean | 17 +++++++++++++++++ src/probability/stopping.lean | 6 ++++++ 2 files changed, 23 insertions(+) diff --git a/src/probability/martingale/basic.lean b/src/probability/martingale/basic.lean index e5a83bcbccf49..6907199516c67 100644 --- a/src/probability/martingale/basic.lean +++ b/src/probability/martingale/basic.lean @@ -553,6 +553,23 @@ lemma submartingale_iff_expected_stopped_value_mono [is_finite_measure μ] ⟨λ hf _ _ hτ hπ hle ⟨N, hN⟩, hf.expected_stopped_value_mono hτ hπ hle hN, submartingale_of_expected_stopped_value_mono hadp hint⟩ +/-- The stopped process of a submartingale with respect to a stopping time is a submartingale. -/ +@[protected] +lemma submartingale.stopped_process [is_finite_measure μ] + {f : ℕ → Ω → ℝ} (h : submartingale f 𝒢 μ) {τ : Ω → ℕ} (hτ : is_stopping_time 𝒢 τ) : + submartingale (stopped_process f τ) 𝒢 μ := +begin + rw submartingale_iff_expected_stopped_value_mono, + { intros σ π hσ hπ hσ_le_π hπ_bdd, + simp_rw stopped_value_stopped_process, + obtain ⟨n, hπ_le_n⟩ := hπ_bdd, + exact h.expected_stopped_value_mono (hσ.min hτ) (hπ.min hτ) + (λ ω, min_le_min (hσ_le_π ω) le_rfl) (λ ω, (min_le_left _ _).trans (hπ_le_n ω)), }, + { exact adapted.stopped_process_of_nat h.adapted hτ, }, + { exact λ i, integrable_stopped_value ((is_stopping_time_const _ i).min hτ) (h.integrable) + (λ ω, min_le_left _ _), }, +end + section maximal open finset diff --git a/src/probability/stopping.lean b/src/probability/stopping.lean index 8823a96a56d0b..63bedf4533adf 100644 --- a/src/probability/stopping.lean +++ b/src/probability/stopping.lean @@ -1212,6 +1212,12 @@ Intuitively, the stopped process stops evolving once the stopping time has occur def stopped_process (u : ι → Ω → β) (τ : Ω → ι) : ι → Ω → β := λ i ω, u (min i (τ ω)) ω +lemma stopped_process_eq_stopped_value {u : ι → Ω → β} {τ : Ω → ι} : + stopped_process u τ = λ i, stopped_value u (λ ω, min i (τ ω)) := rfl + +lemma stopped_value_stopped_process {u : ι → Ω → β} {τ σ : Ω → ι} : + stopped_value (stopped_process u τ) σ = stopped_value u (λ ω, min (σ ω) (τ ω)) := rfl + lemma stopped_process_eq_of_le {u : ι → Ω → β} {τ : Ω → ι} {i : ι} {ω : Ω} (h : i ≤ τ ω) : stopped_process u τ i ω = u i ω := by simp [stopped_process, min_eq_left h] From 3a8665a831e938cbfbd06145664c066a45a2e3c2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 5 Sep 2022 11:05:45 +0000 Subject: [PATCH 164/828] feat(tactic/positivity): Handle division of integers (#16163) Extend `positivity_div` to handle division of integers. --- src/tactic/positivity.lean | 37 ++++++++++++++++++++++++++++++++----- test/positivity.lean | 6 +++--- 2 files changed, 35 insertions(+), 8 deletions(-) diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index b9c36aa5a13cb..db20574053c15 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -297,18 +297,45 @@ private lemma div_nonneg_of_nonneg_of_pos [linear_ordered_field R] {a b : R} (ha 0 ≤ a / b := div_nonneg ha hb.le +private lemma int_div_self_pos {a : ℤ} (ha : 0 < a) : 0 < a / a := +by { rw int.div_self ha.ne', exact zero_lt_one } + +private lemma int_div_nonneg_of_pos_of_nonneg {a b : ℤ} (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ a / b := +int.div_nonneg ha.le hb + +private lemma int_div_nonneg_of_nonneg_of_pos {a b : ℤ} (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ a / b := +int.div_nonneg ha hb.le + +private lemma int_div_nonneg_of_pos_of_pos {a b : ℤ} (ha : 0 < a) (hb : 0 < b) : 0 ≤ a / b := +int.div_nonneg ha.le hb.le + /-- Extension for the `positivity` tactic: division is nonnegative if both numerator and denominator are nonnegative, and strictly positive if both numerator and denominator are. -/ @[positivity] meta def positivity_div : expr → tactic strictness -| `(%%a / %%b) := do -- TODO handle eg `int.div_nonneg` +| `(@has_div.div int _ %%a %%b) := do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | positive pa, positive pb := + if a = b then -- Only attempts to prove `0 < a / a`, otherwise falls back to `0 ≤ a / b` + positive <$> mk_app ``int_div_self_pos [pa] + else + nonnegative <$> mk_app ``int_div_nonneg_of_pos_of_pos [pa, pb] + | positive pa, nonnegative pb := + nonnegative <$> mk_app ``int_div_nonneg_of_pos_of_nonneg [pa, pb] + | nonnegative pa, positive pb := + nonnegative <$> mk_app ``int_div_nonneg_of_nonneg_of_pos [pa, pb] + | nonnegative pa, nonnegative pb := nonnegative <$> mk_app ``int.div_nonneg [pa, pb] + end +| `(%%a / %%b) := do strictness_a ← core a, strictness_b ← core b, match strictness_a, strictness_b with - | (positive pa), (positive pb) := positive <$> mk_app ``div_pos [pa, pb] - | (positive pa), (nonnegative pb) := nonnegative <$> mk_app ``div_nonneg_of_pos_of_nonneg [pa, pb] - | (nonnegative pa), (positive pb) := nonnegative <$> mk_app ``div_nonneg_of_nonneg_of_pos [pa, pb] - | (nonnegative pa), (nonnegative pb) := nonnegative <$> mk_app ``div_nonneg [pa, pb] + | positive pa, positive pb := positive <$> mk_app ``div_pos [pa, pb] + | positive pa, nonnegative pb := nonnegative <$> mk_app ``div_nonneg_of_pos_of_nonneg [pa, pb] + | nonnegative pa, positive pb := nonnegative <$> mk_app ``div_nonneg_of_nonneg_of_pos [pa, pb] + | nonnegative pa, nonnegative pb := nonnegative <$> mk_app ``div_nonneg [pa, pb] end | _ := failed diff --git a/test/positivity.lean b/test/positivity.lean index 8c3baf504fe01..782184eeabf47 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -46,9 +46,9 @@ example {a : ℤ} (ha : 3 < a) : 0 < a + a := by positivity example {a b : ℚ} (ha : 3 < a) (hb : 4 ≤ b) : 0 < 3 + a * b / 7 + b + 7 + 14 := by positivity --- TODO: this fails because `div_nonneg` doesn't apply directly to `ℤ` -- it requires a linearly --- ordered field --- example {a b : ℤ} (ha : 3 < a) (hb : 4 ≤ b) : 0 < 3 + a * b / 7 + b + 7 + 14 := by positivity +example {a b : ℤ} (ha : 3 < a) (hb : 4 ≤ b) : 0 < 3 + a * b / 7 + b + 7 + 14 := by positivity + +example {a : ℤ} (ha : 0 < a) : 0 < a / a := by positivity example {a : ℕ} : 0 < a ^ 0 := by positivity From 0ab31714b8772cf968823f7aa4a9a8be66a12913 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 5 Sep 2022 12:48:21 +0000 Subject: [PATCH 165/828] feat(order/heyting/boundary): Co-Heyting boundary (#16257) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the boundary of an element in a co-Heyting algebra. This generalizes the topological boundary as an operation on `closeds α`. --- src/order/heyting/boundary.lean | 114 ++++++++++++++++++++++++++++++++ 1 file changed, 114 insertions(+) create mode 100644 src/order/heyting/boundary.lean diff --git a/src/order/heyting/boundary.lean b/src/order/heyting/boundary.lean new file mode 100644 index 0000000000000..192f0371025d8 --- /dev/null +++ b/src/order/heyting/boundary.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.boolean_algebra + +/-! +# Co-Heyting boundary + +The boundary of an element of a co-Heyting algebra is the intersection of its Heyting negation with +itself. The boundary in the co-Heyting algebra of closed sets coincides with the topological +boundary. + +## Main declarations + +* `coheyting.boundary`: Co-Heyting boundary. `coheyting.boundary a = a ⊓ ¬a` + +## Notation + +`∂ a` is notation for `coheyting.boundary a` in locale `heyting`. +-/ + +variables {α : Type*} + +namespace coheyting +variables [coheyting_algebra α] {a b : α} + +/-- The boundary of an element of a co-Heyting algebra is the intersection of its Heyting negation +with itself. Note that this is always `⊥` for a boolean algebra. -/ +def boundary (a : α) : α := a ⊓ ¬a + +localized "prefix `∂ `:120 := coheyting.boundary" in heyting + +lemma inf_hnot_self (a : α) : a ⊓ ¬a = ∂ a := rfl + +lemma boundary_le : ∂ a ≤ a := inf_le_left +lemma boundary_le_hnot : ∂ a ≤ ¬a := inf_le_right + +@[simp] lemma boundary_bot : ∂ (⊥ : α) = ⊥ := bot_inf_eq +@[simp] lemma boundary_top : ∂ (⊤ : α) = ⊥ := by rw [boundary, hnot_top, inf_bot_eq] + +lemma boundary_hnot_le (a : α) : ∂ ¬a ≤ ∂ a := inf_comm.trans_le $ inf_le_inf_right _ hnot_hnot_le + +@[simp] lemma boundary_hnot_hnot (a : α) : ∂ ¬¬a = ∂ ¬a := +by simp_rw [boundary, hnot_hnot_hnot, inf_comm] + +@[simp] lemma hnot_boundary (a : α) : ¬ ∂ a = ⊤ := by rw [boundary, hnot_inf_distrib, sup_hnot_self] + +/-- **Leibniz rule** for the co-Heyting boundary. -/ +lemma boundary_inf (a b : α) : ∂ (a ⊓ b) = ∂ a ⊓ b ⊔ a ⊓ ∂ b := +by { unfold boundary, rw [hnot_inf_distrib, inf_sup_left, inf_right_comm, ←inf_assoc] } + +lemma boundary_inf_le : ∂ (a ⊓ b) ≤ ∂ a ⊔ ∂ b := +(boundary_inf _ _).trans_le $ sup_le_sup inf_le_left inf_le_right + +lemma boundary_sup_le : ∂ (a ⊔ b) ≤ ∂ a ⊔ ∂ b := +begin + rw [boundary, inf_sup_right], + exact sup_le_sup (inf_le_inf_left _ $ hnot_anti le_sup_left) + (inf_le_inf_left _ $ hnot_anti le_sup_right), +end + +/- The intuitionistic version of `coheyting.boundary_le_boundary_sup_sup_boundary_inf_left`. Either +proof can be obtained from the other using the equivalence of Heyting algebras and intuitionistic +logic and duality between Heyting and co-Heyting algebras. It is crucial that the following proof be +intuitionistic. -/ +example (a b : Prop) : ((a ∧ b) ∨ ¬(a ∧ b)) ∧ ((a ∨ b) ∨ ¬ (a ∨ b)) → a ∨ ¬ a := +begin + rintro ⟨⟨ha, hb⟩ | hnab, (ha | hb) | hnab⟩; + try { exact or.inl ha }, + { exact or.inr (λ ha, hnab ⟨ha, hb⟩) }, + { exact or.inr (λ ha, hnab $ or.inl ha) } +end + +lemma boundary_le_boundary_sup_sup_boundary_inf_left : ∂ a ≤ ∂ (a ⊔ b) ⊔ ∂ (a ⊓ b) := +begin + simp only [boundary, sup_inf_left, sup_inf_right, sup_right_idem, le_inf_iff, sup_assoc, + @sup_comm _ _ _ a], + refine ⟨⟨⟨_, _⟩, _⟩, ⟨_, _⟩, _⟩; + try { exact le_sup_of_le_left inf_le_left }; + refine inf_le_of_right_le _, + { rw [hnot_le_iff_codisjoint_right, codisjoint_left_comm], + exact codisjoint_hnot_left }, + { refine le_sup_of_le_right _, + rw hnot_le_iff_codisjoint_right, + exact codisjoint_hnot_right.mono_right (hnot_anti inf_le_left) } +end + +lemma boundary_le_boundary_sup_sup_boundary_inf_right : ∂ b ≤ ∂ (a ⊔ b) ⊔ ∂ (a ⊓ b) := +by { rw [@sup_comm _ _ a, inf_comm], exact boundary_le_boundary_sup_sup_boundary_inf_left } + +lemma boundary_sup_sup_boundary_inf (a b : α) : ∂ (a ⊔ b) ⊔ ∂ (a ⊓ b) = ∂ a ⊔ ∂ b := +le_antisymm (sup_le boundary_sup_le boundary_inf_le) $ sup_le + boundary_le_boundary_sup_sup_boundary_inf_left boundary_le_boundary_sup_sup_boundary_inf_right + +@[simp] lemma boundary_idem (a : α) : ∂ ∂ a = ∂ a := by rw [boundary, hnot_boundary, inf_top_eq] + +lemma hnot_hnot_sup_boundary (a : α) : ¬¬a ⊔ ∂ a = a := +by { rw [boundary, sup_inf_left, hnot_sup_self, inf_top_eq, sup_eq_right], exact hnot_hnot_le } + +lemma hnot_eq_top_iff_exists_boundary : ¬a = ⊤ ↔ ∃ b, ∂ b = a := +⟨λ h, ⟨a, by rw [boundary, h, inf_top_eq]⟩, by { rintro ⟨b, rfl⟩, exact hnot_boundary _ }⟩ + +end coheyting + +open_locale heyting + +section boolean_algebra +variables [boolean_algebra α] + +@[simp] lemma coheyting.boundary_eq_bot (a : α) : ∂ a = ⊥ := inf_compl_eq_bot + +end boolean_algebra From 4b4975cf92a1ffe2ddfeff6ff91b0c46a9162bf5 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Mon, 5 Sep 2022 12:48:22 +0000 Subject: [PATCH 166/828] feat(): add a bunch of lemmas for use with Jacobi symbols (#16290) This PR introduces a number of lemmas that will be needed for proving results about the Jacobi symbol (to be introduced in a follow-up PR). (Originally, the Jacobi symbol results were included here; see the discussion below.) Discussion on [Zuilp](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Jacobi.20symbol/near/295816984). --- src/data/list/basic.lean | 17 ++++++++++ src/data/nat/factorization/basic.lean | 8 +++++ src/data/nat/parity.lean | 22 ++++++++++++ src/data/sign.lean | 7 ++++ src/data/zmod/coprime.lean | 34 +++++++++++++++++++ .../legendre_symbol/zmod_char.lean | 34 ++++++++++++++++++- src/ring_theory/int/basic.lean | 15 ++++++++ 7 files changed, 136 insertions(+), 1 deletion(-) create mode 100644 src/data/zmod/coprime.lean diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 9d8b10d2692cf..eadbb338f32c9 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -2749,6 +2749,23 @@ begin { simpa [hl] } } end +lemma pmap_append {p : ι → Prop} (f : Π (a : ι), p a → α) (l₁ l₂ : list ι) + (h : ∀ a ∈ l₁ ++ l₂, p a) : + (l₁ ++ l₂).pmap f h = l₁.pmap f (λ a ha, h a (mem_append_left l₂ ha)) ++ + l₂.pmap f (λ a ha, h a (mem_append_right l₁ ha)) := +begin + induction l₁ with _ _ ih, + { refl, }, + { dsimp only [pmap, cons_append], + rw ih, } +end + +lemma pmap_append' {α β : Type*} {p : α → Prop} (f : Π (a : α), p a → β) (l₁ l₂ : list α) + (h₁ : ∀ a ∈ l₁, p a) (h₂ : ∀ a ∈ l₂, p a) : + (l₁ ++ l₂).pmap f (λ a ha, (list.mem_append.1 ha).elim (h₁ a) (h₂ a)) = + l₁.pmap f h₁ ++ l₂.pmap f h₂ := +pmap_append f l₁ l₂ _ + /-! ### find -/ section find diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index d9ee0db5b768a..5510beae7791e 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -482,6 +482,14 @@ begin { simp [hqp, (factorization_le_iff_dvd hd0 hn0).2 hdn q] }, end +/-- If `n` is a nonzero natural number and `p ≠ 1`, then there are natural numbers `e` +and `n'` such that `n'` is not divisible by `p` and `n = p^e * n'`. -/ +lemma exists_eq_pow_mul_and_not_dvd {n : ℕ} (hn : n ≠ 0) (p : ℕ) (hp : p ≠ 1) : + ∃ e n' : ℕ, ¬ p ∣ n' ∧ n = p ^ e * n' := +let ⟨a', h₁, h₂⟩ := multiplicity.exists_eq_pow_mul_and_not_dvd + (multiplicity.finite_nat_iff.mpr ⟨hp, nat.pos_of_ne_zero hn⟩) in +⟨_, a', h₂, h₁⟩ + lemma dvd_iff_div_factorization_eq_tsub {d n : ℕ} (hd : d ≠ 0) (hdn : d ≤ n) : d ∣ n ↔ (n / d).factorization = n.factorization - d.factorization := begin diff --git a/src/data/nat/parity.lean b/src/data/nat/parity.lean index 9d899182a0ad4..b284576f9d898 100644 --- a/src/data/nat/parity.lean +++ b/src/data/nat/parity.lean @@ -227,3 +227,25 @@ variables {R : Type*} [monoid R] [has_distrib_neg R] {n : ℕ} lemma neg_one_pow_eq_one_iff_even (h : (-1 : R) ≠ 1) : (-1 : R) ^ n = 1 ↔ even n := ⟨λ h', of_not_not $ λ hn, h $ (odd.neg_one_pow $ odd_iff_not_even.mpr hn).symm.trans h', even.neg_one_pow⟩ + +/-- If `a` is even, then `n` is odd iff `n % a` is odd. -/ +lemma odd.mod_even_iff {n a : ℕ} (ha : even a) : odd (n % a) ↔ odd n := +((even_sub' $ mod_le n a).mp $ even_iff_two_dvd.mpr $ (even_iff_two_dvd.mp ha).trans $ + dvd_sub_mod n).symm + +/-- If `a` is even, then `n` is even iff `n % a` is even. -/ +lemma even.mod_even_iff {n a : ℕ} (ha : even a) : even (n % a) ↔ even n := +((even_sub $ mod_le n a).mp $ even_iff_two_dvd.mpr $ (even_iff_two_dvd.mp ha).trans $ + dvd_sub_mod n).symm + +/-- If `n` is odd and `a` is even, then `n % a` is odd. -/ +lemma odd.mod_even {n a : ℕ} (hn : odd n) (ha : even a) : odd (n % a) := +(odd.mod_even_iff ha).mpr hn + +/-- If `n` is even and `a` is even, then `n % a` is even. -/ +lemma even.mod_even {n a : ℕ} (hn : even n) (ha : even a) : even (n % a) := +(even.mod_even_iff ha).mpr hn + +/-- `2` is not a prime factor of an odd natural number. -/ +lemma odd.factors_ne_two {n p : ℕ} (hn : odd n) (hp : p ∈ n.factors) : p ≠ 2 := +by { rintro rfl, exact two_dvd_ne_zero.mpr (odd_iff.mp hn) (dvd_of_mem_factors hp) } diff --git a/src/data/sign.lean b/src/data/sign.lean index f43908af88240..092565b849f62 100644 --- a/src/data/sign.lean +++ b/src/data/sign.lean @@ -174,6 +174,13 @@ end cast map_one' := rfl, map_mul' := λ x y, by cases x; cases y; simp } +lemma range_eq {α} (f : sign_type → α) : set.range f = {f zero, f neg, f pos} := +begin + classical, + simpa only [← finset.coe_singleton, ← finset.image_singleton, + ← fintype.coe_image_univ, finset.coe_image, ← set.image_insert_eq], +end + end sign_type variables {α : Type*} diff --git a/src/data/zmod/coprime.lean b/src/data/zmod/coprime.lean new file mode 100644 index 0000000000000..e1bee15b7c3da --- /dev/null +++ b/src/data/zmod/coprime.lean @@ -0,0 +1,34 @@ +/- +Copyright (c) 2022 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import data.zmod.basic +import ring_theory.int.basic + +/-! +# Coprimality and vanishing + +We show that for prime `p`, the image of an integer `a` in `zmod p` vanishes if and only if +`a` and `p` are not coprime. +-/ + +namespace zmod + +/-- If `p` is a prime and `a` is an integer, then `a : zmod p` is zero if and only if +`gcd a p ≠ 1`. -/ +lemma eq_zero_iff_gcd_ne_one {a : ℤ} {p : ℕ} [pp : fact p.prime] : (a : zmod p) = 0 ↔ a.gcd p ≠ 1 := +by rw [ne, int.gcd_comm, int.gcd_eq_one_iff_coprime, + (nat.prime_iff_prime_int.1 pp.1).coprime_iff_not_dvd, not_not, int_coe_zmod_eq_zero_iff_dvd] + +/-- If an integer `a` and a prime `p` satisfy `gcd a p = 1`, then `a : zmod p` is nonzero. -/ +lemma ne_zero_of_gcd_eq_one {a : ℤ} {p : ℕ} (pp : p.prime) (h : a.gcd p = 1) : + (a : zmod p) ≠ 0 := +mt (@eq_zero_iff_gcd_ne_one a p ⟨pp⟩).mp (not_not.mpr h) + +/-- If an integer `a` and a prime `p` satisfy `gcd a p ≠ 1`, then `a : zmod p` is zero. -/ +lemma eq_zero_of_gcd_ne_one {a : ℤ} {p : ℕ} (pp : p.prime) (h : a.gcd p ≠ 1) : + (a : zmod p) = 0 := +(@eq_zero_iff_gcd_ne_one a p ⟨pp⟩).mpr h + +end zmod diff --git a/src/number_theory/legendre_symbol/zmod_char.lean b/src/number_theory/legendre_symbol/zmod_char.lean index 8363db2366816..a268ce915e329 100644 --- a/src/number_theory/legendre_symbol/zmod_char.lean +++ b/src/number_theory/legendre_symbol/zmod_char.lean @@ -40,6 +40,14 @@ It corresponds to the extension `ℚ(√-1)/ℚ`. -/ /-- `χ₄` takes values in `{0, 1, -1}` -/ lemma is_quadratic_χ₄ : χ₄.is_quadratic := by { intro a, dec_trivial!, } +/-- The value of `χ₄ n`, for `n : ℕ`, depends only on `n % 4`. -/ +lemma χ₄_nat_mod_four (n : ℕ) : χ₄ n = χ₄ (n % 4 : ℕ) := +by rw ← zmod.nat_cast_mod n 4 + +/-- The value of `χ₄ n`, for `n : ℤ`, depends only on `n % 4`. -/ +lemma χ₄_int_mod_four (n : ℤ) : χ₄ n = χ₄ (n % 4 : ℤ) := +by { rw ← zmod.int_cast_mod n 4, norm_cast, } + /-- An explicit description of `χ₄` on integers / naturals -/ lemma χ₄_int_eq_if_mod_four (n : ℤ) : χ₄ n = if n % 2 = 0 then 0 else if n % 4 = 1 then 1 else -1 := begin @@ -52,7 +60,7 @@ end lemma χ₄_nat_eq_if_mod_four (n : ℕ) : χ₄ n = if n % 2 = 0 then 0 else if n % 4 = 1 then 1 else -1 := by exact_mod_cast χ₄_int_eq_if_mod_four n -/-- Alternative description for odd `n : ℕ` in terms of powers of `-1` -/ +/-- Alternative description of `χ₄ n` for odd `n : ℕ` in terms of powers of `-1` -/ lemma χ₄_eq_neg_one_pow {n : ℕ} (hn : n % 2 = 1) : χ₄ n = (-1)^(n / 2) := begin rw χ₄_nat_eq_if_mod_four, @@ -67,6 +75,22 @@ begin ((nat.mod_mod_of_dvd n (by norm_num : 2 ∣ 4)).trans hn), end +/-- If `n % 4 = 1`, then `χ₄ n = 1`. -/ +lemma χ₄_nat_one_mod_four {n : ℕ} (hn : n % 4 = 1) : χ₄ n = 1 := +by { rw [χ₄_nat_mod_four, hn], refl } + +/-- If `n % 4 = 3`, then `χ₄ n = -1`. -/ +lemma χ₄_nat_three_mod_four {n : ℕ} (hn : n % 4 = 3) : χ₄ n = -1 := +by { rw [χ₄_nat_mod_four, hn], refl } + +/-- If `n % 4 = 1`, then `χ₄ n = 1`. -/ +lemma χ₄_int_one_mod_four {n : ℤ} (hn : n % 4 = 1) : χ₄ n = 1 := +by { rw [χ₄_int_mod_four, hn], refl } + +/-- If `n % 4 = 3`, then `χ₄ n = -1`. -/ +lemma χ₄_int_three_mod_four {n : ℤ} (hn : n % 4 = 3) : χ₄ n = -1 := +by { rw [χ₄_int_mod_four, hn], refl } + /-- If `n % 4 = 1`, then `(-1)^(n/2) = 1`. -/ lemma _root_.neg_one_pow_div_two_of_one_mod_four {n : ℕ} (hn : n % 4 = 1) : (-1 : ℤ) ^ (n / 2) = 1 := @@ -86,6 +110,14 @@ It corresponds to the extension `ℚ(√2)/ℚ`. -/ /-- `χ₈` takes values in `{0, 1, -1}` -/ lemma is_quadratic_χ₈ : χ₈.is_quadratic := by { intro a, dec_trivial!, } +/-- The value of `χ₈ n`, for `n : ℕ`, depends only on `n % 8`. -/ +lemma χ₈_nat_mod_eight (n : ℕ) : χ₈ n = χ₈ (n % 8 : ℕ) := +by rw ← zmod.nat_cast_mod n 8 + +/-- The value of `χ₈ n`, for `n : ℤ`, depends only on `n % 8`. -/ +lemma χ₈_int_mod_eight (n : ℤ) : χ₈ n = χ₈ (n % 8 : ℤ) := +by { rw ← zmod.int_cast_mod n 8, norm_cast, } + /-- An explicit description of `χ₈` on integers / naturals -/ lemma χ₈_int_eq_if_mod_eight (n : ℤ) : χ₈ n = if n % 2 = 0 then 0 else if n % 8 = 1 ∨ n % 8 = 7 then 1 else -1 := diff --git a/src/ring_theory/int/basic.lean b/src/ring_theory/int/basic.lean index f17203787592f..ac14779ac532b 100644 --- a/src/ring_theory/int/basic.lean +++ b/src/ring_theory/int/basic.lean @@ -176,6 +176,21 @@ end lemma coprime_iff_nat_coprime {a b : ℤ} : is_coprime a b ↔ nat.coprime a.nat_abs b.nat_abs := by rw [←gcd_eq_one_iff_coprime, nat.coprime_iff_gcd_eq_one, gcd_eq_nat_abs] +/-- If `gcd a (m * n) ≠ 1`, then `gcd a m ≠ 1` or `gcd a n ≠ 1`. -/ +lemma gcd_ne_one_iff_gcd_mul_right_ne_one {a : ℤ} {m n : ℕ} : + a.gcd (m * n) ≠ 1 ↔ a.gcd m ≠ 1 ∨ a.gcd n ≠ 1 := +by simp only [gcd_eq_one_iff_coprime, ← not_and_distrib, not_iff_not, is_coprime.mul_right_iff] + +/-- If `gcd a (m * n) = 1`, then `gcd a m = 1`. -/ +lemma gcd_eq_one_of_gcd_mul_right_eq_one_left {a : ℤ} {m n : ℕ} (h : a.gcd (m * n) = 1) : + a.gcd m = 1 := +nat.dvd_one.mp $ trans_rel_left _ (gcd_dvd_gcd_mul_right_right a m n) h + +/-- If `gcd a (m * n) = 1`, then `gcd a n = 1`. -/ +lemma gcd_eq_one_of_gcd_mul_right_eq_one_right {a : ℤ} {m n : ℕ} (h : a.gcd (m * n) = 1) : + a.gcd n = 1 := +nat.dvd_one.mp $ trans_rel_left _ (gcd_dvd_gcd_mul_left_right a n m) h + lemma sq_of_gcd_eq_one {a b c : ℤ} (h : int.gcd a b = 1) (heq : a * b = c ^ 2) : ∃ (a0 : ℤ), a = a0 ^ 2 ∨ a = - (a0 ^ 2) := begin From 06930844e3f937188af029cbb16383a65915152e Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 5 Sep 2022 12:48:23 +0000 Subject: [PATCH 167/828] refactor(linear_algebra/affine_space/finite_dimensional): change three lemmas to instances (#16331) I realised after the PR with these lemmas was approved that they are actually suitable to be instances (don't depend on any hypotheses that typeclass inference can't deduce), and making them such sometimes allows removing explicit `haveI` uses because typeclass inference can then find them automatically. --- src/linear_algebra/affine_space/finite_dimensional.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index e9c203dd691a4..7c181394b7b5e 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -266,7 +266,7 @@ begin end /-- The `vector_span` of adding a point to a finite-dimensional subspace is finite-dimensional. -/ -lemma finite_dimensional_vector_span_insert (s : affine_subspace k P) +instance finite_dimensional_vector_span_insert (s : affine_subspace k P) [finite_dimensional k s.direction] (p : P) : finite_dimensional k (vector_span k (insert p (s : set P))) := begin @@ -282,7 +282,7 @@ end /-- The direction of the affine span of adding a point to a finite-dimensional subspace is finite-dimensional. -/ -lemma finite_dimensional_direction_affine_span_insert (s : affine_subspace k P) +instance finite_dimensional_direction_affine_span_insert (s : affine_subspace k P) [finite_dimensional k s.direction] (p : P) : finite_dimensional k (affine_span k (insert p (s : set P))).direction := (direction_affine_span k (insert p (s : set P))).symm ▸ finite_dimensional_vector_span_insert s p @@ -291,7 +291,7 @@ variables (k) /-- The `vector_span` of adding a point to a set with a finite-dimensional `vector_span` is finite-dimensional. -/ -lemma finite_dimensional_vector_span_insert_set (s : set P) +instance finite_dimensional_vector_span_insert_set (s : set P) [finite_dimensional k (vector_span k s)] (p : P) : finite_dimensional k (vector_span k (insert p s)) := begin From 7df02f87d798d54ec979d57d5958aae1954b9aee Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Mon, 5 Sep 2022 12:48:24 +0000 Subject: [PATCH 168/828] =?UTF-8?q?feat(measure=5Ftheory/integral/circle?= =?UTF-8?q?=5Fintegral=5Ftransform.lean):=20name=20ch=E2=80=A6=20(#16387)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Change file name from `circle_integral_transform` to `circle_transform` to match the definitions. This is in preparation for #15356 --- .../{circle_integral_transform.lean => circle_transform.lean} | 0 1 file changed, 0 insertions(+), 0 deletions(-) rename src/measure_theory/integral/{circle_integral_transform.lean => circle_transform.lean} (100%) diff --git a/src/measure_theory/integral/circle_integral_transform.lean b/src/measure_theory/integral/circle_transform.lean similarity index 100% rename from src/measure_theory/integral/circle_integral_transform.lean rename to src/measure_theory/integral/circle_transform.lean From 3f772d42fdca637de99b4696c2fe1a31f448e934 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 5 Sep 2022 15:24:14 +0000 Subject: [PATCH 169/828] =?UTF-8?q?chore(category=5Ftheory/sites/compatibl?= =?UTF-8?q?e=5Fplus):=20speed=20up=20=CE=B9=5Fplus=5Fcomp=5Fiso=5Fhom=20(#?= =?UTF-8?q?16379)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This failed when experimenting with `-T90000`; just squeezing the `simp`s seems to help a fair bit. --- src/category_theory/sites/compatible_plus.lean | 14 ++++++++++---- 1 file changed, 10 insertions(+), 4 deletions(-) diff --git a/src/category_theory/sites/compatible_plus.lean b/src/category_theory/sites/compatible_plus.lean index 3e654c5ccbef9..bb82c2ad7f308 100644 --- a/src/category_theory/sites/compatible_plus.lean +++ b/src/category_theory/sites/compatible_plus.lean @@ -116,11 +116,17 @@ lemma ι_plus_comp_iso_hom (X) (W) : F.map (colimit.ι _ W) ≫ (J.plus_comp_iso (J.diagram_comp_iso F P X.unop).hom.app W ≫ colimit.ι _ W := begin delta diagram_comp_iso plus_comp_iso, - dsimp [is_colimit.cocone_point_unique_up_to_iso], - simp only [← category.assoc], + simp only [is_colimit.desc_cocone_morphism_hom, is_colimit.unique_up_to_iso_hom, + cocones.forget_map, iso.trans_hom, nat_iso.of_components_hom_app, functor.map_iso_hom, + ← category.assoc], erw (is_colimit_of_preserves F (colimit.is_colimit (J.diagram P (unop X)))).fac, - dsimp, - simp, + simp only [category.assoc, has_limit.iso_of_nat_iso_hom_π, iso.symm_hom, + cover.multicospan_comp_hom_inv_left, eq_to_hom_refl, category.comp_id, + limit.cone_point_unique_up_to_iso_hom_comp, functor.map_cone_π_app, + multiequalizer.multifork_π_app_left, multiequalizer.lift_ι, functor.map_comp, eq_self_iff_true, + category.assoc, iso.trans_hom, iso.cancel_iso_hom_left, nat_iso.of_components_hom_app, + colimit.cocone_ι, category.assoc, has_colimit.iso_of_nat_iso_ι_hom], + end @[simp, reassoc] From 8ac7e3bab257041b460eb00704a59c98f0cac43f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 5 Sep 2022 16:20:52 +0000 Subject: [PATCH 170/828] feat(topology/instances/nnreal): add `nnreal.pow_order_iso` (#16344) Also use it to redefine `nnreal.sqrt`. --- src/data/real/sqrt.lean | 35 ++++++++++-------------------- src/topology/instances/nnreal.lean | 7 ++++++ 2 files changed, 19 insertions(+), 23 deletions(-) diff --git a/src/data/real/sqrt.lean b/src/data/real/sqrt.lean index 594e67cf1694a..2ba270b9a1395 100644 --- a/src/data/real/sqrt.lean +++ b/src/data/real/sqrt.lean @@ -42,10 +42,7 @@ variables {x y : ℝ≥0} /-- Square root of a nonnegative real number. -/ @[pp_nodot] noncomputable def sqrt : ℝ≥0 ≃o ℝ≥0 := -order_iso.symm $ strict_mono.order_iso_of_surjective (λ x, x * x) - (λ x y h, mul_self_lt_mul_self x.2 h) $ - (continuous_id.mul continuous_id).surjective tendsto_mul_self_at_top $ - by simp [order_bot.at_bot_eq] +order_iso.symm $ pow_order_iso 2 two_ne_zero lemma sqrt_le_sqrt_iff : sqrt x ≤ sqrt y ↔ x ≤ y := sqrt.le_iff_le @@ -53,35 +50,27 @@ sqrt.le_iff_le lemma sqrt_lt_sqrt_iff : sqrt x < sqrt y ↔ x < y := sqrt.lt_iff_lt -lemma sqrt_eq_iff_sq_eq : sqrt x = y ↔ y * y = x := +lemma sqrt_eq_iff_sq_eq : sqrt x = y ↔ y ^ 2 = x := sqrt.to_equiv.apply_eq_iff_eq_symm_apply.trans eq_comm -lemma sqrt_le_iff : sqrt x ≤ y ↔ x ≤ y * y := +lemma sqrt_le_iff : sqrt x ≤ y ↔ x ≤ y ^ 2 := sqrt.to_galois_connection _ _ -lemma le_sqrt_iff : x ≤ sqrt y ↔ x * x ≤ y := +lemma le_sqrt_iff : x ≤ sqrt y ↔ x ^ 2 ≤ y := (sqrt.symm.to_galois_connection _ _).symm @[simp] lemma sqrt_eq_zero : sqrt x = 0 ↔ x = 0 := -sqrt_eq_iff_sq_eq.trans $ by rw [eq_comm, zero_mul] +sqrt_eq_iff_sq_eq.trans $ by rw [eq_comm, sq, zero_mul] @[simp] lemma sqrt_zero : sqrt 0 = 0 := sqrt_eq_zero.2 rfl - -@[simp] lemma sqrt_one : sqrt 1 = 1 := sqrt_eq_iff_sq_eq.2 $ mul_one 1 - -@[simp] lemma mul_self_sqrt (x : ℝ≥0) : sqrt x * sqrt x = x := -sqrt.symm_apply_apply x - -@[simp] lemma sqrt_mul_self (x : ℝ≥0) : sqrt (x * x) = x := sqrt.apply_symm_apply x - -@[simp] lemma sq_sqrt (x : ℝ≥0) : (sqrt x)^2 = x := -by rw [sq, mul_self_sqrt x] - -@[simp] lemma sqrt_sq (x : ℝ≥0) : sqrt (x^2) = x := -by rw [sq, sqrt_mul_self x] +@[simp] lemma sqrt_one : sqrt 1 = 1 := sqrt_eq_iff_sq_eq.2 $ one_pow _ +@[simp] lemma sq_sqrt (x : ℝ≥0) : (sqrt x)^2 = x := sqrt.symm_apply_apply x +@[simp] lemma mul_self_sqrt (x : ℝ≥0) : sqrt x * sqrt x = x := by rw [← sq, sq_sqrt] +@[simp] lemma sqrt_sq (x : ℝ≥0) : sqrt (x^2) = x := sqrt.apply_symm_apply x +@[simp] lemma sqrt_mul_self (x : ℝ≥0) : sqrt (x * x) = x := by rw [← sq, sqrt_sq x] lemma sqrt_mul (x y : ℝ≥0) : sqrt (x * y) = sqrt x * sqrt y := -by rw [sqrt_eq_iff_sq_eq, mul_mul_mul_comm, mul_self_sqrt, mul_self_sqrt] +by rw [sqrt_eq_iff_sq_eq, mul_pow, sq_sqrt, sq_sqrt] /-- `nnreal.sqrt` as a `monoid_with_zero_hom`. -/ noncomputable def sqrt_hom : ℝ≥0 →*₀ ℝ≥0 := ⟨sqrt, sqrt_zero, sqrt_one, sqrt_mul⟩ @@ -213,7 +202,7 @@ theorem sqrt_lt_sqrt (hx : 0 ≤ x) (h : x < y) : sqrt x < sqrt y := (sqrt_lt_sqrt_iff hx).2 h theorem sqrt_le_left (hy : 0 ≤ y) : sqrt x ≤ y ↔ x ≤ y ^ 2 := -by rw [sqrt, ← real.le_to_nnreal_iff_coe_le hy, nnreal.sqrt_le_iff, ← real.to_nnreal_mul hy, +by rw [sqrt, ← real.le_to_nnreal_iff_coe_le hy, nnreal.sqrt_le_iff, sq, ← real.to_nnreal_mul hy, real.to_nnreal_le_to_nnreal_iff (mul_self_nonneg y), sq] theorem sqrt_le_iff : sqrt x ≤ y ↔ 0 ≤ y ∧ x ≤ y ^ 2 := diff --git a/src/topology/instances/nnreal.lean b/src/topology/instances/nnreal.lean index be9f755cf5c69..0f8c94e0ba68a 100644 --- a/src/topology/instances/nnreal.lean +++ b/src/topology/instances/nnreal.lean @@ -219,4 +219,11 @@ begin exact tendsto_tsum_compl_at_top_zero (λ (a : α), (f a : ℝ)) end +/-- `x ↦ x ^ n` as an order isomorphism of `ℝ≥0`. -/ +def pow_order_iso (n : ℕ) (hn : n ≠ 0) : ℝ≥0 ≃o ℝ≥0 := +strict_mono.order_iso_of_surjective (λ x, x ^ n) + (λ x y h, strict_mono_on_pow hn.bot_lt (zero_le x) (zero_le y) h) $ + (continuous_id.pow _).surjective (tendsto_pow_at_top hn) $ + by simpa [order_bot.at_bot_eq, pos_iff_ne_zero] + end nnreal From d25f4d64e139dbd9afd960e4c25c8fc6982b15cb Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 5 Sep 2022 18:04:23 +0000 Subject: [PATCH 171/828] feat(linear_algebra/affine_space/affine_subspace): `vadd_mem_iff_mem_of_mem_direction` (#16383) Add a lemma similar to `vadd_mem_iff_mem_direction`, but where it's given that the vector is in `s.direction` and the `iff` is between the two points being in the subspace. --- src/linear_algebra/affine_space/affine_subspace.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index 4d444adc745bf..cdf63befdda3b 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -276,6 +276,16 @@ lemma vadd_mem_iff_mem_direction {s : affine_subspace k P} (v : V) {p : P} (hp : v +ᵥ p ∈ s ↔ v ∈ s.direction := ⟨λ h, by simpa using vsub_mem_direction h hp, λ h, vadd_mem_of_mem_direction h hp⟩ +/-- Adding a vector in the direction to a point produces a point in the subspace if and only if +the original point is in the subspace. -/ +lemma vadd_mem_iff_mem_of_mem_direction {s : affine_subspace k P} {v : V} (hv : v ∈ s.direction) + {p : P} : v +ᵥ p ∈ s ↔ p ∈ s := +begin + refine ⟨λ h, _, λ h, vadd_mem_of_mem_direction hv h⟩, + convert vadd_mem_of_mem_direction (submodule.neg_mem _ hv) h, + simp +end + /-- Given a point in an affine subspace, the set of vectors in its direction equals the set of vectors subtracting that point on the right. -/ From 037731799f20e700c618cfb495eab06825eabd35 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Mon, 5 Sep 2022 19:41:20 +0000 Subject: [PATCH 172/828] feat(algebra/group_power/basic): add `pow_mul_pow_eq_one` (#16328) --- src/algebra/group_power/basic.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/algebra/group_power/basic.lean b/src/algebra/group_power/basic.lean index 14fbb60c631cc..6448fec7ca7dc 100644 --- a/src/algebra/group_power/basic.lean +++ b/src/algebra/group_power/basic.lean @@ -142,6 +142,17 @@ by rw [pow_bit0, (commute.refl a).mul_pow] theorem pow_bit1' (a : M) (n : ℕ) : a ^ bit1 n = (a * a) ^ n * a := by rw [bit1, pow_succ', pow_bit0'] +@[to_additive] +lemma pow_mul_pow_eq_one {a b : M} (n : ℕ) (h : a * b = 1) : + a ^ n * b ^ n = 1 := +begin + induction n with n hn, + { simp }, + { calc a ^ n.succ * b ^ n.succ = a ^ n * a * (b * b ^ n) : by rw [pow_succ', pow_succ] + ... = a ^ n * (a * b) * b ^ n : by simp only [mul_assoc] + ... = 1 : by simp [h, hn] } +end + lemma dvd_pow {x y : M} (hxy : x ∣ y) : ∀ {n : ℕ} (hn : n ≠ 0), x ∣ y^n | 0 hn := (hn rfl).elim From c3762602ae9d2ae8f0e0a757f19bf5078786e02c Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Mon, 5 Sep 2022 19:41:21 +0000 Subject: [PATCH 173/828] chore(analysis/analytic/isolated_zeros): tidy exists_has_sum_smul_of_apply_eq_zero (#16389) --- src/analysis/analytic/isolated_zeros.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean index dfc719faf9dcf..92ff0d95fb527 100644 --- a/src/analysis/analytic/isolated_zeros.lean +++ b/src/analysis/analytic/isolated_zeros.lean @@ -45,9 +45,9 @@ lemma exists_has_sum_smul_of_apply_eq_zero (hs : has_sum (λ m, z ^ m • a m) s (ha : ∀ k < n, a k = 0) : ∃ t : E, z ^ n • t = s ∧ has_sum (λ m, z ^ m • a (m + n)) t := begin - refine classical.by_cases (λ hn : n = 0, by { subst n; simpa }) (λ hn, _), - replace hn := nat.pos_of_ne_zero hn, - by_cases (z = 0), + obtain rfl | hn := n.eq_zero_or_pos, + { simpa }, + by_cases h : z = 0, { have : s = 0 := hs.unique (by simpa [ha 0 hn, h] using has_sum_at_zero a), exact ⟨a n, by simp [h, hn, this], by simpa [h] using has_sum_at_zero (λ m, a (m + n))⟩ }, { refine ⟨(z ^ n)⁻¹ • s, by field_simp [smul_smul], _⟩, @@ -56,7 +56,8 @@ begin have h2 : has_sum (λ m, z ^ (m + n) • a (m + n)) s, by simpa [h1] using (has_sum_nat_add_iff' n).mpr hs, convert @has_sum.const_smul E ℕ 𝕜 _ _ _ _ _ _ _ (z⁻¹ ^ n) h2, - field_simp [pow_add, smul_smul], simp only [inv_pow] } + { field_simp [pow_add, smul_smul] }, + { simp only [inv_pow] } } end end has_sum @@ -70,7 +71,7 @@ begin have hp0 : p.coeff 0 = f z₀ := hp.coeff_zero 1, simp only [has_fpower_series_at_iff, apply_eq_pow_smul_coeff, coeff_fslope] at hp ⊢, refine hp.mono (λ x hx, _), - by_cases x = 0, + by_cases h : x = 0, { convert has_sum_single 0 _; intros; simp [*] }, { have hxx : ∀ (n : ℕ), x⁻¹ * x ^ (n + 1) = x ^ n := λ n, by field_simp [h, pow_succ'], suffices : has_sum (λ n, x⁻¹ • x ^ (n + 1) • p.coeff (n + 1)) (x⁻¹ • (f (z₀ + x) - f z₀)), From 3e067975886cf5801e597925328c335609511b1a Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 5 Sep 2022 22:37:47 +0000 Subject: [PATCH 174/828] feat(algebra/monoid_algebra/no_zero_divisors): left/right order implies no zero divisors (#16299) This PR is split off from #15983. It contains only the result about zero-divisors in the presence of a left/right-ordered appropriately cancellable semigroup. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/144837-PR-reviews/topic/.2315983.20no-zero.20divisors.20and.20right.20orderability) --- .../monoid_algebra/no_zero_divisors.lean | 155 ++++++++++++++++++ 1 file changed, 155 insertions(+) create mode 100644 src/algebra/monoid_algebra/no_zero_divisors.lean diff --git a/src/algebra/monoid_algebra/no_zero_divisors.lean b/src/algebra/monoid_algebra/no_zero_divisors.lean new file mode 100644 index 0000000000000..4a6499cd4e067 --- /dev/null +++ b/src/algebra/monoid_algebra/no_zero_divisors.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2022 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa +-/ +import algebra.monoid_algebra.support + +/-! +# Variations on non-zero divisors in `add_monoid_algebra`s + +This file studies the interaction between typeclass assumptions on two Types `R` and `A` and +whether `add_monoid_algebra R A` has non-zero zero-divisors. For some background on related +questions, see [Kaplansky's Conjectures](https://en.wikipedia.org/wiki/Kaplansky%27s_conjectures), +especially the *zero divisor conjecture*. + +_Conjecture._ +Let `K` be a field, and `G` a torsion-free group. The group ring `K[G]` does not contain +nontrivial zero divisors, that is, it is a domain. + +We formalize in this file the well-known result that if `R` is a field and `A` is a left-ordered +group, then `R[A]` contains no non-zero zero-divisors. Some of these assumptions can be trivially +weakened: below we mention what assumptions are sufficient for the proofs in this file. + +## Main results + +* `no_zero_divisors.of_left_ordered` shows that if `R` is a semiring with no non-zero + zero-divisors, `A` is a linearly ordered, add right cancel semigroup with strictly monotone + left addition, then `add_monoid_algebra R A` has no non-zero zero-divisors. +* `no_zero_divisors.of_right_ordered` shows that if `R` is a semiring with no non-zero + zero-divisors, `A` is a linearly ordered, add left cancel semigroup with strictly monotone + right addition, then `add_monoid_algebra R A` has no non-zero zero-divisors. + +The conditions on `A` imposed in `no_zero_divisors.of_left_ordered` are sometimes referred to as +`left-ordered`. +The conditions on `A` imposed in `no_zero_divisors.of_right_ordered` are sometimes referred to as +`right-ordered`. + +These conditions are sufficient, but not necessary. As mentioned above, *Kaplansky's Conjecture* +asserts that `A` being torsion-free may be enough. +-/ + +namespace add_monoid_algebra +open finsupp + +variables {R A : Type*} [semiring R] + +/-- The coefficient of a monomial in a product `f * g` that can be reached in at most one way +as a product of monomials in the supports of `f` and `g` is a product. -/ +lemma mul_apply_add_eq_mul_of_forall_ne [has_add A] {f g : add_monoid_algebra R A} {a0 b0 : A} + (h : ∀ {a b : A}, a ∈ f.support → b ∈ g.support → (a ≠ a0 ∨ b ≠ b0) → a + b ≠ a0 + b0) : + (f * g) (a0 + b0) = f a0 * g b0 := +begin + classical, + rw mul_apply, + refine (finset.sum_eq_single a0 _ _).trans _, + { exact λ b H hb, finset.sum_eq_zero (λ x H1, if_neg (h H H1 (or.inl hb))) }, + { exact λ af0, by simp [not_mem_support_iff.mp af0] }, + { refine (finset.sum_eq_single b0 (λ b bg b0, _) _).trans (if_pos rfl), + { by_cases af : a0 ∈ f.support, + { exact if_neg (h af bg (or.inr b0)) }, + { simp only [not_mem_support_iff.mp af, zero_mul, if_t_t] } }, + { exact λ bf0, by simp [not_mem_support_iff.mp bf0] } }, +end + +section left_or_right_orderability + +lemma left.exists_add_of_mem_support_single_mul [add_left_cancel_semigroup A] + {g : add_monoid_algebra R A} (a x : A) + (hx : x ∈ (single a 1 * g : add_monoid_algebra R A).support) : + ∃ b ∈ g.support, a + b = x := +by rwa [support_single_mul _ _ (λ y, by rw one_mul : ∀ y : R, 1 * y = 0 ↔ _), finset.mem_map] at hx + +lemma right.exists_add_of_mem_support_single_mul [add_right_cancel_semigroup A] + {f : add_monoid_algebra R A} (b x : A) + (hx : x ∈ (f * single b 1 : add_monoid_algebra R A).support) : + ∃ a ∈ f.support, a + b = x := +by rwa [support_mul_single _ _ (λ y, by rw mul_one : ∀ y : R, y * 1 = 0 ↔ _), finset.mem_map] at hx + +/-- If `R` is a semiring with no non-trivial zero-divisors and `A` is a left-ordered add right +cancel semigroup, then `add_monoid_algebra R A` also contains no non-zero zero-divisors. -/ +lemma no_zero_divisors.of_left_ordered [no_zero_divisors R] + [add_right_cancel_semigroup A] [linear_order A] [covariant_class A A (+) (<)] : + no_zero_divisors (add_monoid_algebra R A) := +⟨λ f g fg, begin + contrapose! fg, + let gmin : A := g.support.min' (support_nonempty_iff.mpr fg.2), + refine support_nonempty_iff.mp _, + obtain ⟨a, ha, H⟩ := right.exists_add_of_mem_support_single_mul gmin + ((f * single gmin 1 : add_monoid_algebra R A).support.min' + (by rw support_mul_single; simp [support_nonempty_iff.mpr fg.1])) (finset.min'_mem _ _), + refine ⟨a + gmin, mem_support_iff.mpr _⟩, + rw mul_apply_add_eq_mul_of_forall_ne _, + { refine mul_ne_zero _ _, + exacts [mem_support_iff.mp ha, mem_support_iff.mp (finset.min'_mem _ _)] }, + { rw H, + rintro b c bf cg (hb | hc); refine ne_of_gt _, + { refine lt_of_lt_of_le (_ : _ < b + gmin ) _, + { apply finset.min'_lt_of_mem_erase_min', + rw ← H, + apply finset.mem_erase_of_ne_of_mem, + { simpa only [ne.def, add_left_inj] }, + { rw support_mul_single _ _ (λ y, by rw mul_one : ∀ y : R, y * 1 = 0 ↔ _), + simpa only [finset.mem_map, add_right_embedding_apply, add_left_inj, exists_prop, + exists_eq_right] } }, + { haveI : covariant_class A A (+) (≤) := has_add.to_covariant_class_left A, + exact add_le_add_left (finset.min'_le _ _ cg) _ } }, + { refine lt_of_le_of_lt (_ : _ ≤ b + gmin) _, + { apply finset.min'_le, + rw support_mul_single _ _ (λ y, by rw mul_one : ∀ y : R, y * 1 = 0 ↔ _), + simp only [bf, finset.mem_map, add_right_embedding_apply, add_left_inj, exists_prop, + exists_eq_right] }, + { refine add_lt_add_left _ _, + exact finset.min'_lt_of_mem_erase_min' _ _ (finset.mem_erase.mpr ⟨hc, cg⟩) } } } +end⟩ + +/-- If `R` is a semiring with no non-trivial zero-divisors and `A` is a right-ordered add left +cancel semigroup, then `add_monoid_algebra R A` also contains no non-zero zero-divisors. -/ +lemma no_zero_divisors.of_right_ordered [no_zero_divisors R] + [add_left_cancel_semigroup A] [linear_order A] [covariant_class A A (function.swap (+)) (<)] : + no_zero_divisors (add_monoid_algebra R A) := +⟨λ f g fg, begin + contrapose! fg, + let fmin : A := f.support.min' (support_nonempty_iff.mpr fg.1), + refine support_nonempty_iff.mp _, + obtain ⟨a, ha, H⟩ := left.exists_add_of_mem_support_single_mul fmin + ((single fmin 1 * g : add_monoid_algebra R A).support.min' + (by rw support_single_mul; simp [support_nonempty_iff.mpr fg.2])) (finset.min'_mem _ _), + refine ⟨fmin + a, mem_support_iff.mpr _⟩, + rw mul_apply_add_eq_mul_of_forall_ne _, + { refine mul_ne_zero _ _, + exacts [mem_support_iff.mp (finset.min'_mem _ _), mem_support_iff.mp ha] }, + { rw H, + rintro b c bf cg (hb | hc); refine ne_of_gt _, + { refine lt_of_le_of_lt (_ : _ ≤ fmin + c) _, + { apply finset.min'_le, + rw support_single_mul _ _ (λ y, by rw one_mul : ∀ y : R, 1 * y = 0 ↔ _), + simp only [cg, finset.mem_map, add_left_embedding_apply, add_right_inj, exists_prop, + exists_eq_right] }, + { refine add_lt_add_right _ _, + exact finset.min'_lt_of_mem_erase_min' _ _ (finset.mem_erase.mpr ⟨hb, bf⟩) } }, + { refine lt_of_lt_of_le (_ : _ < fmin + c) _, + { apply finset.min'_lt_of_mem_erase_min', + rw ← H, + apply finset.mem_erase_of_ne_of_mem, + { simpa only [ne.def, add_right_inj] }, + { rw support_single_mul _ _ (λ y, by rw one_mul : ∀ y : R, 1 * y = 0 ↔ _), + simpa only [finset.mem_map, add_left_embedding_apply, add_right_inj, exists_prop, + exists_eq_right]} }, + { haveI : covariant_class A A (function.swap (+)) (≤) := has_add.to_covariant_class_right A, + exact add_le_add_right (finset.min'_le _ _ bf) _ } } } +end⟩ + +end left_or_right_orderability + +end add_monoid_algebra From 1ef341261865ddb2179ebf47ee258b5ab7ec8a72 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 6 Sep 2022 00:11:44 +0000 Subject: [PATCH 175/828] refactor(order/rel_classes): ditch `is_strict_total_order'` (#16069) Since Lean 3.46, `is_strict_total_order` and `is_strict_total_order'` are exactly the same typeclass. Hence, we remove the latter for the former. Co-authored-by: Yury G. Kudryashov --- src/data/list/lex.lean | 4 +- src/data/pi/lex.lean | 2 +- src/order/rel_classes.lean | 40 +++++-------------- src/order/rel_iso.lean | 6 +-- src/set_theory/cardinal/ordinal.lean | 2 +- .../metric_space/emetric_paracompact.lean | 2 +- src/topology/partition_of_unity.lean | 2 +- 7 files changed, 19 insertions(+), 39 deletions(-) diff --git a/src/data/list/lex.lean b/src/data/list/lex.lean index 516c7d3b7fc0c..6b8d9eb1439d7 100644 --- a/src/data/list/lex.lean +++ b/src/data/list/lex.lean @@ -93,7 +93,7 @@ instance is_asymm (r : α → α → Prop) end⟩ instance is_strict_total_order (r : α → α → Prop) - [is_strict_total_order' α r] : is_strict_total_order' (list α) (lex r) := + [is_strict_total_order α r] : is_strict_total_order (list α) (lex r) := {..is_strict_weak_order_of_is_order_connected} instance decidable_rel [decidable_eq α] (r : α → α → Prop) @@ -157,7 +157,7 @@ theorem nil_lt_cons [has_lt α] (a : α) (l : list α) : [] < a :: l := lex.nil instance [linear_order α] : linear_order (list α) := -linear_order_of_STO' (lex (<)) +linear_order_of_STO (lex (<)) --Note: this overrides an instance in core lean instance has_le' [linear_order α] : has_le (list α) := diff --git a/src/data/pi/lex.lean b/src/data/pi/lex.lean index c6ac718a57a90..ebdda5aaa60c3 100644 --- a/src/data/pi/lex.lean +++ b/src/data/pi/lex.lean @@ -83,7 +83,7 @@ partial_order_of_SO (<) /-- `Πₗ i, α i` is a linear order if the original order is well-founded. -/ noncomputable instance [linear_order ι] [is_well_order ι (<)] [∀ a, linear_order (β a)] : linear_order (lex (Π i, β i)) := -@linear_order_of_STO' (Πₗ i, β i) (<) +@linear_order_of_STO (Πₗ i, β i) (<) { to_is_trichotomous := is_trichotomous_lex _ _ is_well_founded.wf } (classical.dec_rel _) lemma lex.le_of_forall_le [linear_order ι] [is_well_order ι (<)] [Π a, linear_order (β a)] diff --git a/src/order/rel_classes.lean b/src/order/rel_classes.lean index 6b4677f24d302..07d26e85e1927 100644 --- a/src/order/rel_classes.lean +++ b/src/order/rel_classes.lean @@ -145,17 +145,11 @@ See note [reducible non-instances]. -/ (asymm h)⟩, λ ⟨h₁, h₂⟩, h₁.resolve_left (λ e, h₂ $ e ▸ or.inl rfl)⟩ } -/-- This is basically the same as `is_strict_total_order`, but that definition has a redundant -assumption `is_incomp_trans α lt`. -/ --- TODO: This is now exactly the same as `is_strict_total_order`, remove. -@[algebra] class is_strict_total_order' (α : Type u) (lt : α → α → Prop) - extends is_trichotomous α lt, is_strict_order α lt : Prop. - -/-- Construct a linear order from an `is_strict_total_order'` relation. +/-- Construct a linear order from an `is_strict_total_order` relation. See note [reducible non-instances]. -/ @[reducible] -def linear_order_of_STO' (r) [is_strict_total_order' α r] [Π x y, decidable (¬ r x y)] : +def linear_order_of_STO (r) [is_strict_total_order α r] [Π x y, decidable (¬ r x y)] : linear_order α := { le_total := λ x y, match y, trichotomous_of r x y with @@ -168,8 +162,8 @@ def linear_order_of_STO' (r) [is_strict_total_order' α r] [Π x y, decidable ( λ h, h.elim (λ h, h ▸ irrefl_of _ _) (asymm_of r)⟩, ..partial_order_of_SO r } -theorem is_strict_total_order'.swap (r) [is_strict_total_order' α r] : - is_strict_total_order' α (swap r) := +theorem is_strict_total_order.swap (r) [is_strict_total_order α r] : + is_strict_total_order α (swap r) := {..is_trichotomous.swap r, ..is_strict_order.swap r} /-! ### Order connection -/ @@ -193,25 +187,15 @@ theorem is_strict_weak_order_of_is_order_connected [is_asymm α r] ..@is_asymm.is_irrefl α r _ } @[priority 100] -- see Note [lower instance priority] -instance is_order_connected_of_is_strict_total_order' - [is_strict_total_order' α r] : is_order_connected α r := +instance is_order_connected_of_is_strict_total_order + [is_strict_total_order α r] : is_order_connected α r := ⟨λ a b c h, (trichotomous _ _).imp_right (λ o, o.elim (λ e, e ▸ h) (λ h', trans h' h))⟩ -@[priority 100] -- see Note [lower instance priority] -instance is_strict_total_order_of_is_strict_total_order' - [is_strict_total_order' α r] : is_strict_total_order α r := -{ } - -@[priority 100] -- see Note [lower instance priority] -instance is_strict_weak_order_of_is_strict_total_order' - [is_strict_total_order' α r] : is_strict_weak_order α r := -{ ..is_strict_weak_order_of_is_order_connected } - @[priority 100] -- see Note [lower instance priority] instance is_strict_weak_order_of_is_strict_total_order [is_strict_total_order α r] : is_strict_weak_order α r := -by { haveI : is_strict_total_order' α r := {}, apply_instance } +{ ..is_strict_weak_order_of_is_order_connected } /-! ### Well-order -/ @@ -281,11 +265,8 @@ theorem well_founded_lt_dual_iff (α : Type*) [has_lt α] : well_founded_lt α extends is_trichotomous α r, is_trans α r, is_well_founded α r : Prop @[priority 100] -- see Note [lower instance priority] -instance is_well_order.is_strict_total_order' {α} (r : α → α → Prop) [is_well_order α r] : - is_strict_total_order' α r := { } -@[priority 100] -- see Note [lower instance priority] instance is_well_order.is_strict_total_order {α} (r : α → α → Prop) [is_well_order α r] : - is_strict_total_order α r := by apply_instance + is_strict_total_order α r := { } @[priority 100] -- see Note [lower instance priority] instance is_well_order.is_trichotomous {α} (r : α → α → Prop) [is_well_order α r] : is_trichotomous α r := by apply_instance @@ -352,7 +333,7 @@ end well_founded_gt /-- Construct a decidable linear order from a well-founded linear order. -/ noncomputable def is_well_order.linear_order (r : α → α → Prop) [is_well_order α r] : linear_order α := -by { letI := λ x y, classical.dec (¬r x y), exact linear_order_of_STO' r } +by { letI := λ x y, classical.dec (¬r x y), exact linear_order_of_STO r } /-- Derive a `has_well_founded` instance from a `is_well_order` instance. -/ def is_well_order.to_has_well_founded [has_lt α] [hwo : is_well_order α (<)] : @@ -604,8 +585,7 @@ instance [linear_order α] : is_trichotomous α (<) := ⟨lt_trichotomy⟩ instance [linear_order α] : is_trichotomous α (>) := is_trichotomous.swap _ instance [linear_order α] : is_trichotomous α (≤) := is_total.is_trichotomous _ instance [linear_order α] : is_trichotomous α (≥) := is_total.is_trichotomous _ -instance [linear_order α] : is_strict_total_order α (<) := by apply_instance -instance [linear_order α] : is_strict_total_order' α (<) := {} +instance [linear_order α] : is_strict_total_order α (<) := {} instance [linear_order α] : is_order_connected α (<) := by apply_instance instance [linear_order α] : is_incomp_trans α (<) := by apply_instance instance [linear_order α] : is_strict_weak_order α (<) := by apply_instance diff --git a/src/order/rel_iso.lean b/src/order/rel_iso.lean index 8b326a1d484c9..aa48961a89d7f 100644 --- a/src/order/rel_iso.lean +++ b/src/order/rel_iso.lean @@ -290,8 +290,8 @@ protected theorem is_strict_order : ∀ (f : r ↪r s) [is_strict_order β s], i protected theorem is_trichotomous : ∀ (f : r ↪r s) [is_trichotomous β s], is_trichotomous α r | ⟨f, o⟩ ⟨H⟩ := ⟨λ a b, (or_congr o (or_congr f.inj'.eq_iff o)).1 (H _ _)⟩ -protected theorem is_strict_total_order' : - ∀ (f : r ↪r s) [is_strict_total_order' β s], is_strict_total_order' α r +protected theorem is_strict_total_order : + ∀ (f : r ↪r s) [is_strict_total_order β s], is_strict_total_order α r | f H := by exactI {..f.is_trichotomous, ..f.is_strict_order} protected theorem acc (f : r ↪r s) (a : α) : acc s (f a) → acc r a := @@ -305,7 +305,7 @@ protected theorem well_founded : ∀ (f : r ↪r s) (h : well_founded s), well_f | f ⟨H⟩ := ⟨λ a, f.acc _ (H _)⟩ protected theorem is_well_order : ∀ (f : r ↪r s) [is_well_order β s], is_well_order α r -| f H := by exactI {wf := f.well_founded H.wf, ..f.is_strict_total_order'} +| f H := by exactI {wf := f.well_founded H.wf, ..f.is_strict_total_order} /-- `quotient.out` as a relation embedding between the lift of a relation and the relation. -/ @[simps] noncomputable def _root_.quotient.out_rel_embedding [s : setoid α] {r : α → α → Prop} (H) : diff --git a/src/set_theory/cardinal/ordinal.lean b/src/set_theory/cardinal/ordinal.lean index b9a334642767b..676dda5722659 100644 --- a/src/set_theory/cardinal/ordinal.lean +++ b/src/set_theory/cardinal/ordinal.lean @@ -404,7 +404,7 @@ begin quotient.induction_on c $ λ α IH ol, _) h, -- consider the minimal well-order `r` on `α` (a type with cardinality `c`). rcases ord_eq α with ⟨r, wo, e⟩, resetI, - letI := linear_order_of_STO' r, + letI := linear_order_of_STO r, haveI : is_well_order α (<) := wo, -- Define an order `s` on `α × α` by writing `(a, b) < (c, d)` if `max a b < max c d`, or -- the max are equal and `a < c`, or the max are equal and `a = c` and `b < d`. diff --git a/src/topology/metric_space/emetric_paracompact.lean b/src/topology/metric_space/emetric_paracompact.lean index 5048b2434f82d..cb0a62325bc36 100644 --- a/src/topology/metric_space/emetric_paracompact.lean +++ b/src/topology/metric_space/emetric_paracompact.lean @@ -46,7 +46,7 @@ begin refine ⟨λ ι s ho hcov, _⟩, simp only [Union_eq_univ_iff] at hcov, -- choose a well founded order on `S` - letI : linear_order ι := linear_order_of_STO' well_ordering_rel, + letI : linear_order ι := linear_order_of_STO well_ordering_rel, have wf : well_founded ((<) : ι → ι → Prop) := @is_well_founded.wf ι well_ordering_rel _, -- Let `ind x` be the minimal index `s : S` such that `x ∈ s`. set ind : α → ι := λ x, wf.min {i : ι | x ∈ s i} (hcov x), diff --git a/src/topology/partition_of_unity.lean b/src/topology/partition_of_unity.lean index 7f7e1ace7b685..638019acbeeed 100644 --- a/src/topology/partition_of_unity.lean +++ b/src/topology/partition_of_unity.lean @@ -361,7 +361,7 @@ begin exact λ i hi, f.support_to_pou_fun_subset i hi }, have B : mul_support (λ i, 1 - f i x) ⊆ s, { rw [hs, mul_support_one_sub], exact λ i, id }, - letI : linear_order ι := linear_order_of_STO' well_ordering_rel, + letI : linear_order ι := linear_order_of_STO well_ordering_rel, rw [finsum_eq_sum_of_support_subset _ A, finprod_eq_prod_of_mul_support_subset _ B, finset.prod_one_sub_ordered, sub_sub_cancel], refine finset.sum_congr rfl (λ i hi, _), From 58903102dc40d8309a4ae29953b637d0699bc396 Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Tue, 6 Sep 2022 02:53:11 +0000 Subject: [PATCH 176/828] feat(src/data/*): lemmas about emptiness for finset/multiset/vector coercions (#16113) Adds basic lemmas about whether the result of `to_list` is empty for various data structures. Co-authored-by: Yury G. Kudryashov --- src/algebra/big_operators/multiset.lean | 4 +-- src/algebra/periodic.lean | 2 +- src/data/finset/basic.lean | 21 ++++++++++++---- src/data/multiset/basic.lean | 25 +++++++++++-------- src/data/vector/basic.lean | 8 ++++++ src/data/vector/mem.lean | 2 +- src/field_theory/splitting_field.lean | 2 +- .../admissible_absolute_value.lean | 2 +- 8 files changed, 45 insertions(+), 21 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index 0bf24c4984bf9..db8d3bddff2e1 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -56,13 +56,13 @@ lemma prod_cons (a : α) (s) : prod (a ::ₘ s) = a * prod s := foldr_cons _ _ _ @[simp, to_additive] lemma prod_erase [decidable_eq α] (h : a ∈ s) : a * (s.erase a).prod = s.prod := -by rw [← s.coe_to_list, coe_erase, coe_prod, coe_prod, list.prod_erase ((s.mem_to_list a).2 h)] +by rw [← s.coe_to_list, coe_erase, coe_prod, coe_prod, list.prod_erase (mem_to_list.2 h)] @[simp, to_additive] lemma prod_map_erase [decidable_eq ι] {a : ι} (h : a ∈ m) : f a * ((m.erase a).map f).prod = (m.map f).prod := by rw [← m.coe_to_list, coe_erase, coe_map, coe_map, coe_prod, coe_prod, - list.prod_map_erase f ((m.mem_to_list a).2 h)] + list.prod_map_erase f (mem_to_list.2 h)] @[simp, to_additive] lemma prod_singleton (a : α) : prod {a} = a := diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index f7146dd723aaa..1ea42572a7c38 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -85,7 +85,7 @@ end lemma _root_.multiset.periodic_prod [has_add α] [comm_monoid β] (s : multiset (α → β)) (hs : ∀ f ∈ s, periodic f c) : periodic s.prod c := -s.prod_to_list ▸ s.to_list.periodic_prod $ λ f hf, hs f $ (multiset.mem_to_list f s).mp hf +s.prod_to_list ▸ s.to_list.periodic_prod $ λ f hf, hs f $ multiset.mem_to_list.mp hf @[to_additive] lemma _root_.finset.periodic_prod [has_add α] [comm_monoid β] diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 253cc1aa93ecc..19e7b8d41f0f1 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -2463,18 +2463,29 @@ finset.val_inj.1 (multiset.dedup_map_dedup_eq _ _).symm section to_list /-- Produce a list of the elements in the finite set using choice. -/ -@[reducible] noncomputable def to_list (s : finset α) : list α := s.1.to_list +noncomputable def to_list (s : finset α) : list α := s.1.to_list lemma nodup_to_list (s : finset α) : s.to_list.nodup := by { rw [to_list, ←multiset.coe_nodup, multiset.coe_to_list], exact s.nodup } -@[simp] lemma mem_to_list {a : α} (s : finset α) : a ∈ s.to_list ↔ a ∈ s := -by { rw [to_list, ←multiset.mem_coe, multiset.coe_to_list], exact iff.rfl } +@[simp] lemma mem_to_list {a : α} {s : finset α} : a ∈ s.to_list ↔ a ∈ s := mem_to_list -@[simp] lemma to_list_empty : (∅ : finset α).to_list = [] := by simp [to_list] +@[simp] lemma to_list_eq_nil {s : finset α} : s.to_list = [] ↔ s = ∅ := +to_list_eq_nil.trans val_eq_zero + +@[simp] lemma empty_to_list {s : finset α} : s.to_list.empty ↔ s = ∅ := +list.empty_iff_eq_nil.trans to_list_eq_nil + +@[simp] lemma to_list_empty : (∅ : finset α).to_list = [] := to_list_eq_nil.mpr rfl + +lemma nonempty.to_list_ne_nil {s : finset α} (hs : s.nonempty) : s.to_list ≠ [] := +mt to_list_eq_nil.mp hs.ne_empty + +lemma nonempty.not_empty_to_list {s : finset α} (hs : s.nonempty) : ¬s.to_list.empty := +mt empty_to_list.mp hs.ne_empty @[simp, norm_cast] -lemma coe_to_list (s : finset α) : (s.to_list : multiset α) = s.val := by { classical, ext, simp } +lemma coe_to_list (s : finset α) : (s.to_list : multiset α) = s.val := s.val.coe_to_list @[simp] lemma to_list_to_finset [decidable_eq α] (s : finset α) : s.to_list.to_finset = s := by { ext, simp } diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 54ef913df33d1..2fe50738cd830 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -59,6 +59,9 @@ instance inhabited_multiset : inhabited (multiset α) := ⟨0⟩ @[simp] theorem coe_eq_zero (l : list α) : (l : multiset α) = 0 ↔ l = [] := iff.trans coe_eq_coe perm_nil +lemma coe_eq_zero_iff_empty (l : list α) : (l : multiset α) = 0 ↔ l.empty := +iff.trans (coe_eq_zero l) (empty_iff_eq_nil).symm + /-! ### `multiset.cons` -/ /-- `cons a s` is the multiset which contains `s` plus one more @@ -296,19 +299,21 @@ end subset section to_list /-- Produces a list of the elements in the multiset using choice. -/ -@[reducible] noncomputable def to_list {α : Type*} (s : multiset α) := -classical.some (quotient.exists_rep s) - -@[simp] lemma to_list_zero {α : Type*} : (multiset.to_list 0 : list α) = [] := -(multiset.coe_eq_zero _).1 (classical.some_spec (quotient.exists_rep multiset.zero)) +noncomputable def to_list (s : multiset α) := s.out' @[simp, norm_cast] -lemma coe_to_list {α : Type*} (s : multiset α) : (s.to_list : multiset α) = s := -classical.some_spec (quotient.exists_rep _) +lemma coe_to_list (s : multiset α) : (s.to_list : multiset α) = s := s.out_eq' -@[simp] -lemma mem_to_list {α : Type*} (a : α) (s : multiset α) : a ∈ s.to_list ↔ a ∈ s := -by rw [←multiset.mem_coe, multiset.coe_to_list] +@[simp] lemma to_list_eq_nil {s : multiset α} : s.to_list = [] ↔ s = 0 := +by rw [← coe_eq_zero, coe_to_list] + +@[simp] lemma empty_to_list {s : multiset α} : s.to_list.empty ↔ s = 0 := +empty_iff_eq_nil.trans to_list_eq_nil + +@[simp] lemma to_list_zero : (multiset.to_list 0 : list α) = [] := to_list_eq_nil.mpr rfl + +@[simp] lemma mem_to_list {a : α} {s : multiset α} : a ∈ s.to_list ↔ a ∈ s := +by rw [← mem_coe, coe_to_list] end to_list diff --git a/src/data/vector/basic.lean b/src/data/vector/basic.lean index fbf668184a87f..d958b566e2f0c 100644 --- a/src/data/vector/basic.lean +++ b/src/data/vector/basic.lean @@ -143,6 +143,8 @@ by simp only [←cons_head_tail, eq_iff_true_of_subsingleton] tail (of_fn f) = of_fn (λ i, f i.succ) := (of_fn_nth _).symm.trans $ by { congr, funext i, cases i, simp, } +@[simp] theorem to_list_empty (v : vector α 0) : v.to_list = [] := list.length_eq_zero.mp v.2 + /-- The list that makes up a `vector` made up of a single element, retrieved via `to_list`, is equal to the list of that single element. -/ @[simp] lemma to_list_singleton (v : vector α 1) : v.to_list = [v.head] := @@ -152,6 +154,12 @@ begin and_self, singleton_tail] end +@[simp] lemma empty_to_list_eq_ff (v : vector α (n + 1)) : v.to_list.empty = ff := +match v with | ⟨a :: as, _⟩ := rfl end + +lemma not_empty_to_list (v : vector α (n + 1)) : ¬ v.to_list.empty := +by simp only [empty_to_list_eq_ff, coe_sort_ff, not_false_iff] + /-- Mapping under `id` does not change a vector. -/ @[simp] lemma map_id {n : ℕ} (v : vector α n) : vector.map id v = v := vector.eq _ _ (by simp only [list.map_id, vector.to_list_map]) diff --git a/src/data/vector/mem.lean b/src/data/vector/mem.lean index 458960d34cc63..7db02e14c7bbe 100644 --- a/src/data/vector/mem.lean +++ b/src/data/vector/mem.lean @@ -27,7 +27,7 @@ by simp only [list.mem_iff_nth_le, fin.exists_iff, vector.nth_eq_nth_le]; lemma not_mem_nil : a ∉ (vector.nil : vector α 0).to_list := id -@[simp] lemma not_mem_zero (v : vector α 0) : a ∉ v.to_list := +lemma not_mem_zero (v : vector α 0) : a ∉ v.to_list := (vector.eq_nil v).symm ▸ (not_mem_nil a) lemma mem_cons_iff (v : vector α n) : diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field.lean index b19a3e5de8d0d..3b0ffd6ea1466 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field.lean @@ -293,7 +293,7 @@ else or.inr $ λ p hp hdp, begin { refine (hp.2.1 $ is_unit_of_dvd_unit hd _).elim, exact is_unit_C.2 ((leading_coeff_ne_zero.2 hf0).is_unit.map i) }, { obtain ⟨q, hq, hd⟩ := hp.dvd_prod_iff.1 hd, - obtain ⟨a, ha, rfl⟩ := multiset.mem_map.1 ((multiset.mem_to_list _ _).1 hq), + obtain ⟨a, ha, rfl⟩ := multiset.mem_map.1 (multiset.mem_to_list.1 hq), rw degree_eq_degree_of_associated ((hp.dvd_prime_iff_associated $ prime_X_sub_C a).1 hd), exact degree_X_sub_C a }, end diff --git a/src/number_theory/class_number/admissible_absolute_value.lean b/src/number_theory/class_number/admissible_absolute_value.lean index 9ba9b4345a66b..5f7474ce78b71 100644 --- a/src/number_theory/class_number/admissible_absolute_value.lean +++ b/src/number_theory/class_number/admissible_absolute_value.lean @@ -96,7 +96,7 @@ begin { intros i j h, ext, exact list.nodup_iff_nth_le_inj.mp (finset.nodup_to_list _) _ _ _ _ h }, have : ∀ i h, (finset.univ.filter (λ x, t x = s)).to_list.nth_le i h ∈ finset.univ.filter (λ x, t x = s), - { intros i h, exact (finset.mem_to_list _).mp (list.nth_le_mem _ _ _) }, + { intros i h, exact finset.mem_to_list.mp (list.nth_le_mem _ _ _) }, obtain ⟨_, h₀⟩ := finset.mem_filter.mp (this i₀ _), obtain ⟨_, h₁⟩ := finset.mem_filter.mp (this i₁ _), exact h₀.trans h₁.symm }, From f0c55c2ad8e3e026ff45e6667756f2c80ac0e21a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 6 Sep 2022 05:41:24 +0000 Subject: [PATCH 177/828] feat(measure_theory/measure/measure_space): add variants of `ae_restrict_Union_eq` (#16370) Co-authored-by: Yury G. Kudryashov --- src/measure_theory/measure/measure_space.lean | 23 +++++++++++++++++++ src/order/filter/basic.lean | 2 +- 2 files changed, 24 insertions(+), 1 deletion(-) diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index de48daf41be0b..3d455c8c0a2ee 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -1983,6 +1983,29 @@ le_antisymm (ae_sum_eq (λ i, μ.restrict (s i)) ▸ ae_mono restrict_Union_le) (μ.restrict (s ∪ t)).ae = (μ.restrict s).ae ⊔ (μ.restrict t).ae := by simp [union_eq_Union, supr_bool_eq] +lemma ae_restrict_bUnion_eq (s : ι → set α) {t : set ι} (ht : t.countable) : + (μ.restrict (⋃ i ∈ t, s i)).ae = ⨆ i ∈ t, (μ.restrict (s i)).ae := +begin + haveI := ht.to_subtype, + rw [bUnion_eq_Union, ae_restrict_Union_eq, ← supr_subtype''], +end + +lemma ae_restrict_bUnion_finset_eq (s : ι → set α) (t : finset ι) : + (μ.restrict (⋃ i ∈ t, s i)).ae = ⨆ i ∈ t, (μ.restrict (s i)).ae := +ae_restrict_bUnion_eq s t.countable_to_set + +lemma ae_eq_restrict_Union_iff [countable ι] (s : ι → set α) (f g : α → δ) : + f =ᵐ[μ.restrict (⋃ i, s i)] g ↔ ∀ i, f =ᵐ[μ.restrict (s i)] g := +by simp_rw [eventually_eq, ae_restrict_Union_eq, eventually_supr] + +lemma ae_eq_restrict_bUnion_iff (s : ι → set α) {t : set ι} (ht : t.countable) (f g : α → δ) : + f =ᵐ[μ.restrict (⋃ i ∈ t, s i)] g ↔ ∀ i ∈ t, f =ᵐ[μ.restrict (s i)] g := +by simp_rw [ae_restrict_bUnion_eq s ht, eventually_eq, eventually_supr] + +lemma ae_eq_restrict_bUnion_finset_iff (s : ι → set α) (t : finset ι) (f g : α → δ) : + f =ᵐ[μ.restrict (⋃ i ∈ t, s i)] g ↔ ∀ i ∈ t, f =ᵐ[μ.restrict (s i)] g := +ae_eq_restrict_bUnion_iff s t.countable_to_set f g + lemma ae_restrict_interval_oc_eq [linear_order α] (a b : α) : (μ.restrict (Ι a b)).ae = (μ.restrict (Ioc a b)).ae ⊔ (μ.restrict (Ioc b a)).ae := by simp only [interval_oc_eq_union, ae_restrict_union_eq] diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 195385af2f5dc..f33c040554aba 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1029,7 +1029,7 @@ lemma eventually_Sup {p : α → Prop} {fs : set (filter α)} : iff.rfl @[simp] -lemma eventually_supr {p : α → Prop} {fs : β → filter α} : +lemma eventually_supr {p : α → Prop} {fs : ι → filter α} : (∀ᶠ x in (⨆ b, fs b), p x) ↔ (∀ b, ∀ᶠ x in fs b, p x) := mem_supr From e11bafa5284544728bd3b189942e930e0d4701de Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 6 Sep 2022 07:23:28 +0000 Subject: [PATCH 178/828] =?UTF-8?q?feat(category=5Ftheory/limits/shapes/pr?= =?UTF-8?q?oducts):=20if=20each=20`f=20b=20=E2=9F=B6=20g=20b`=20is=20mono,?= =?UTF-8?q?=20then=20`=E2=88=8F=20f=20=E2=9F=B6=20=E2=88=8F=20g`=20is=20mo?= =?UTF-8?q?no=20(#16180)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and its dual version: if each `f b ⟶ g b` is epi, then `∐ f ⟶ ∐ g` is epi --- src/category_theory/limits/has_limits.lean | 15 +++++++++++++++ src/category_theory/limits/shapes/products.lean | 10 ++++++++++ 2 files changed, 25 insertions(+) diff --git a/src/category_theory/limits/has_limits.lean b/src/category_theory/limits/has_limits.lean index b409eb045bf50..67bfe897f66f6 100644 --- a/src/category_theory/limits/has_limits.lean +++ b/src/category_theory/limits/has_limits.lean @@ -506,6 +506,14 @@ instance : is_right_adjoint (lim : (J ⥤ C) ⥤ C) := ⟨_, const_lim_adj⟩ end lim_functor +instance lim_map_mono' {F G : J ⥤ C} [has_limits_of_shape J C] (α : F ⟶ G) + [mono α] : mono (lim_map α) := +(lim : (J ⥤ C) ⥤ C).map_mono α + +instance lim_map_mono {F G : J ⥤ C} [has_limit F] [has_limit G] (α : F ⟶ G) + [∀ j, mono (α.app j)] : mono (lim_map α) := +⟨λ Z u v h, limit.hom_ext $ λ j, (cancel_mono (α.app j)).1 $ by simpa using h =≫ limit.π _ j⟩ + /-- We can transport limits of shape `J` along an equivalence `J ≌ J'`. -/ @@ -1006,6 +1014,13 @@ instance : is_left_adjoint (colim : (J ⥤ C) ⥤ C) := ⟨_, colim_const_adj⟩ end colim_functor +instance colim_map_epi' {F G : J ⥤ C} [has_colimits_of_shape J C] (α : F ⟶ G) [epi α] : + epi (colim_map α) := (colim : (J ⥤ C) ⥤ C).map_epi α + +instance colim_map_epi {F G : J ⥤ C} [has_colimit F] [has_colimit G] (α : F ⟶ G) + [∀ j, epi (α.app j)] : epi (colim_map α) := +⟨λ Z u v h, colimit.hom_ext $ λ j, (cancel_epi (α.app j)).1 $ by simpa using colimit.ι _ j ≫= h⟩ + /-- We can transport colimits of shape `J` along an equivalence `J ≌ J'`. -/ diff --git a/src/category_theory/limits/shapes/products.lean b/src/category_theory/limits/shapes/products.lean index 4bf49c165731f..ed5d4f3caed4b 100644 --- a/src/category_theory/limits/shapes/products.lean +++ b/src/category_theory/limits/shapes/products.lean @@ -142,6 +142,11 @@ from a family of morphisms between the factors. abbreviation pi.map {f g : β → C} [has_product f] [has_product g] (p : Π b, f b ⟶ g b) : ∏ f ⟶ ∏ g := lim_map (discrete.nat_trans (λ X, p X.as)) + +instance pi.map_mono {f g : β → C} [has_product f] [has_product g] + (p : Π b, f b ⟶ g b) [Π i, mono (p i)] : mono $ pi.map p := +@@limits.lim_map_mono _ _ _ _ _ (by { dsimp, apply_instance }) + /-- Construct an isomorphism between categorical products (indexed by the same type) from a family of isomorphisms between the factors. @@ -156,6 +161,11 @@ from a family of morphisms between the factors. abbreviation sigma.map {f g : β → C} [has_coproduct f] [has_coproduct g] (p : Π b, f b ⟶ g b) : ∐ f ⟶ ∐ g := colim_map (discrete.nat_trans (λ X, p X.as)) + +instance sigma.map_epi {f g : β → C} [has_coproduct f] [has_coproduct g] + (p : Π b, f b ⟶ g b) [Π i, epi (p i)] : epi $ sigma.map p := +@@limits.colim_map_epi _ _ _ _ _ (by { dsimp, apply_instance }) + /-- Construct an isomorphism between categorical coproducts (indexed by the same type) from a family of isomorphisms between the factors. From fbebf1a05777c44916fac1512b9624772113468f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 6 Sep 2022 07:23:29 +0000 Subject: [PATCH 179/828] chore(algebra/order/ring): golf a few proofs (#16398) --- src/algebra/order/ring.lean | 46 +++++++++++++++---------------------- 1 file changed, 18 insertions(+), 28 deletions(-) diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index a22119e0dd2b5..85afcf7c3aa7f 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -1523,10 +1523,8 @@ lemma mul_coe {b : α} (hb : b ≠ 0) : ∀{a : with_top α}, a * b = a.bind (λ begin cases a; cases b; simp only [none_eq_top, some_eq_coe], { simp [← coe_mul] }, - { suffices : ⊤ * (b : with_top α) = ⊤ ↔ b ≠ 0, by simpa, - by_cases hb : b = 0; simp [hb] }, - { suffices : (a : with_top α) * ⊤ = ⊤ ↔ a ≠ 0, by simpa, - by_cases ha : a = 0; simp [ha] }, + { by_cases hb : b = 0; simp [hb] }, + { by_cases ha : a = 0; simp [ha] }, { simp [← coe_mul] } end @@ -1554,12 +1552,12 @@ instance [mul_zero_one_class α] [nontrivial α] : mul_zero_one_class (with_top one := 1, zero := 0, one_mul := λ a, match a with - | none := show ((1:α) : with_top α) * ⊤ = ⊤, by simp [-with_top.coe_one] - | (some a) := show ((1:α) : with_top α) * a = a, by simp [coe_mul.symm, -with_top.coe_one] + | ⊤ := mul_top (mt coe_eq_coe.1 one_ne_zero) + | (a : α) := by rw [← coe_one, ← coe_mul, one_mul] end, mul_one := λ a, match a with - | none := show ⊤ * ((1:α) : with_top α) = ⊤, by simp [-with_top.coe_one] - | (some a) := show ↑a * ((1:α) : with_top α) = a, by simp [coe_mul.symm, -with_top.coe_one] + | ⊤ := top_mul (mt coe_eq_coe.1 one_ne_zero) + | (a : α) := by rw [← coe_one, ← coe_mul, mul_one] end, .. with_top.mul_zero_class } @@ -1591,16 +1589,13 @@ instance [semigroup_with_zero α] [no_zero_divisors α] : semigroup_with_zero (w { mul := (*), zero := 0, mul_assoc := λ a b c, begin - cases a, - { by_cases hb : b = 0; by_cases hc : c = 0; - simp [*, none_eq_top] }, - cases b, - { by_cases ha : a = 0; by_cases hc : c = 0; - simp [*, none_eq_top, some_eq_coe] }, - cases c, - { by_cases ha : a = 0; by_cases hb : b = 0; - simp [*, none_eq_top, some_eq_coe] }, - simp [some_eq_coe, coe_mul.symm, mul_assoc] + rcases eq_or_ne a 0 with rfl|ha, { simp only [zero_mul] }, + rcases eq_or_ne b 0 with rfl|hb, { simp only [zero_mul, mul_zero] }, + rcases eq_or_ne c 0 with rfl|hc, { simp only [mul_zero] }, + induction a using with_top.rec_top_coe, { simp [hb, hc] }, + induction b using with_top.rec_top_coe, { simp [ha, hc] }, + induction c using with_top.rec_top_coe, { simp [ha, hb] }, + simp only [← coe_mul, mul_assoc] end, .. with_top.mul_zero_class } @@ -1611,22 +1606,17 @@ instance [comm_monoid_with_zero α] [no_zero_divisors α] [nontrivial α] : comm_monoid_with_zero (with_top α) := { mul := (*), zero := 0, - mul_comm := λ a b, begin - by_cases ha : a = 0, { simp [ha] }, - by_cases hb : b = 0, { simp [hb] }, - simp [ha, hb, mul_def, option.bind_comm a b, mul_comm] - end, + mul_comm := λ a b, + by simp only [or_comm, mul_def, option.bind_comm a b, mul_comm], .. with_top.monoid_with_zero } variables [canonically_ordered_comm_semiring α] private lemma distrib' (a b c : with_top α) : (a + b) * c = a * c + b * c := begin - cases c, - { show (a + b) * ⊤ = a * ⊤ + b * ⊤, - by_cases ha : a = 0; simp [ha] }, - { show (a + b) * c = a * c + b * c, - by_cases hc : c = 0, { simp [hc] }, + induction c using with_top.rec_top_coe, + { by_cases ha : a = 0; simp [ha] }, + { by_cases hc : c = 0, { simp [hc] }, simp [mul_coe hc], cases a; cases b, repeat { refl <|> exact congr_arg some (add_mul _ _ _) } } end From a67ec23dd8dc08195d77b6df2cd21f9c64989131 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 6 Sep 2022 10:21:12 +0000 Subject: [PATCH 180/828] feat(category_theory/sites/sheaf): category of sheaf is preadditive when target is preadditive (#16324) `Sheaf J A` is preadditive when `A` is --- src/category_theory/sites/sheaf.lean | 55 ++++++++++++++++++++++++++++ 1 file changed, 55 insertions(+) diff --git a/src/category_theory/sites/sheaf.lean b/src/category_theory/sites/sheaf.lean index 0da437d8201d8..e9ad13949f38a 100644 --- a/src/category_theory/sites/sheaf.lean +++ b/src/category_theory/sites/sheaf.lean @@ -7,6 +7,7 @@ Authors: Kevin Buzzard, Bhavik Mehta import category_theory.limits.preserves.shapes.equalizers import category_theory.limits.preserves.shapes.products import category_theory.limits.yoneda +import category_theory.preadditive.functor_category import category_theory.sites.sheaf_of_types /-! @@ -324,6 +325,60 @@ begin exact ⟨⟨t⟩, λ a, h.2 a (by tidy)⟩ end +section preadditive + +open preadditive + +variables [preadditive A] {P Q : Sheaf J A} + +instance Sheaf_hom_has_zsmul : has_smul ℤ (P ⟶ Q) := +{ smul := λ n f, Sheaf.hom.mk + { app := λ U, n • f.1.app U, + naturality' := λ U V i, begin + induction n using int.induction_on with n ih n ih, + { simp only [zero_smul, comp_zero, zero_comp], }, + { simpa only [add_zsmul, one_zsmul, comp_add, nat_trans.naturality, add_comp, add_left_inj] }, + { simpa only [sub_smul, one_zsmul, comp_sub, nat_trans.naturality, sub_comp, sub_left_inj] + using ih, } + end } } + +instance : has_sub (P ⟶ Q) := +{ sub := λ f g, Sheaf.hom.mk $ f.1 - g.1 } + +instance : has_neg (P ⟶ Q) := +{ neg := λ f, Sheaf.hom.mk $ -f.1 } + +instance Sheaf_hom_has_nsmul : has_smul ℕ (P ⟶ Q) := +{ smul := λ n f, Sheaf.hom.mk + { app := λ U, n • f.1.app U, + naturality' := λ U V i, begin + induction n with n ih, + { simp only [zero_smul, comp_zero, zero_comp], }, + { simp only [nat.succ_eq_add_one, add_smul, ih, one_nsmul, comp_add, nat_trans.naturality, + add_comp], }, + end } } + +instance : has_zero (P ⟶ Q) := +{ zero := Sheaf.hom.mk 0 } + +instance : has_add (P ⟶ Q) := +{ add := λ f g, Sheaf.hom.mk $ f.1 + g.1 } + +@[simp] lemma Sheaf.hom.add_app (f g : P ⟶ Q) (U) : + (f + g).1.app U = f.1.app U + g.1.app U := rfl + +instance : add_comm_group (P ⟶ Q) := +function.injective.add_comm_group (λ (f : Sheaf.hom P Q), f.1) + (λ _ _ h, Sheaf.hom.ext _ _ h) rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) + (λ _ _, by { dsimp at *, ext, simpa [*] }) (λ _ _, by { dsimp at *, ext, simpa [*] }) + +instance : preadditive (Sheaf J A) := +{ hom_group := λ P Q, infer_instance, + add_comp' := λ P Q R f f' g, by { ext, simp, }, + comp_add' := λ P Q R f g g', by { ext, simp, } } + +end preadditive + end category_theory namespace category_theory From b3b5475d55ac7963f75d8fd90a9ec06bab6139a1 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 6 Sep 2022 13:43:25 +0000 Subject: [PATCH 181/828] feat(data/finsupp/ne_locus): add a not-member lemma (#16245) This PR adds a convenience lemma characterizing non-membership in the `ne_locus`. `simp` proves this lemma, but given the position of the negations in `mem_ne_locus`, this form is also convenient. --- src/data/finsupp/ne_locus.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/data/finsupp/ne_locus.lean b/src/data/finsupp/ne_locus.lean index 5504a917ebca3..2dfb2a4d6d19b 100644 --- a/src/data/finsupp/ne_locus.lean +++ b/src/data/finsupp/ne_locus.lean @@ -36,6 +36,9 @@ def ne_locus (f g : α →₀ N) : finset α := by simpa only [ne_locus, finset.mem_filter, finset.mem_union, mem_support_iff, and_iff_right_iff_imp] using ne.ne_or_ne _ +lemma not_mem_ne_locus {f g : α →₀ N} {a : α} : a ∉ f.ne_locus g ↔ f a = g a := +mem_ne_locus.not.trans not_ne_iff + @[simp] lemma coe_ne_locus : ↑(f.ne_locus g) = {x | f x ≠ g x} := by { ext, exact mem_ne_locus } From 3f409bd9df181d26dd223170da7b6830ece18442 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 6 Sep 2022 14:33:29 +0000 Subject: [PATCH 182/828] chore(category_theory/*): Lint (#16366) Satisfy the `fintype_finite` linter. --- .../filtered_colimit_commutes_finite_limit.lean | 6 +++--- src/category_theory/limits/shapes/biproducts.lean | 8 ++++++-- src/category_theory/limits/shapes/finite_limits.lean | 8 ++++---- .../limits/shapes/finite_products.lean | 12 +++++++----- src/category_theory/preadditive/biproducts.lean | 3 ++- 5 files changed, 22 insertions(+), 15 deletions(-) diff --git a/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean b/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean index 5ac3bd400749c..4a68a0a2bda04 100644 --- a/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean +++ b/src/category_theory/limits/filtered_colimit_commutes_finite_limit.lean @@ -45,7 +45,7 @@ section Injectivity doesn't need that we have finitely many morphisms in `J`, only that there are finitely many objects. -/ -variables [fintype J] +variables [finite J] /-- This follows this proof from @@ -55,7 +55,7 @@ lemma colimit_limit_to_limit_colimit_injective : function.injective (colimit_limit_to_limit_colimit F) := begin classical, - + casesI nonempty_fintype J, -- Suppose we have two terms `x y` in the colimit (over `K`) of the limits (over `J`), -- and that these have the same image under `colimit_limit_to_limit_colimit F`. intros x y h, @@ -81,7 +81,7 @@ begin -- We now use that `K` is filtered, picking some point to the right of all these -- morphisms `f j` and `g j`. - let O : finset K := (finset.univ).image k ∪ {kx, ky}, + let O : finset K := finset.univ.image k ∪ {kx, ky}, have kxO : kx ∈ O := finset.mem_union.mpr (or.inr (by simp)), have kyO : ky ∈ O := finset.mem_union.mpr (or.inr (by simp)), have kjO : ∀ j, k j ∈ O := λ j, finset.mem_union.mpr (or.inl (by simp)), diff --git a/src/category_theory/limits/shapes/biproducts.lean b/src/category_theory/limits/shapes/biproducts.lean index 6e3f996250607..e103165ce5f4b 100644 --- a/src/category_theory/limits/shapes/biproducts.lean +++ b/src/category_theory/limits/shapes/biproducts.lean @@ -1536,7 +1536,9 @@ is_bilimit_of_is_limit _ $ /-- In a preadditive category, if the product over `f : J → C` exists, then the biproduct over `f` exists. -/ -lemma has_biproduct.of_has_product (f : J → C) [has_product f] : has_biproduct f := +lemma has_biproduct.of_has_product {J : Type} [finite J] (f : J → C) [has_product f] : + has_biproduct f := +by casesI nonempty_fintype J; exact has_biproduct.mk { bicone := _, is_bilimit := bicone_is_bilimit_of_limit_cone_of_is_limit (limit.is_limit _) } @@ -1559,7 +1561,9 @@ is_bilimit_of_is_colimit _ $ /-- In a preadditive category, if the coproduct over `f : J → C` exists, then the biproduct over `f` exists. -/ -lemma has_biproduct.of_has_coproduct (f : J → C) [has_coproduct f] : has_biproduct f := +lemma has_biproduct.of_has_coproduct {J : Type} [finite J] (f : J → C) [has_coproduct f] : + has_biproduct f := +by casesI nonempty_fintype J; exact has_biproduct.mk { bicone := _, is_bilimit := bicone_is_bilimit_of_colimit_cocone_of_is_colimit (colimit.is_colimit _) } diff --git a/src/category_theory/limits/shapes/finite_limits.lean b/src/category_theory/limits/shapes/finite_limits.lean index 56b4bc4ccb4c0..c2059ea5ba5b4 100644 --- a/src/category_theory/limits/shapes/finite_limits.lean +++ b/src/category_theory/limits/shapes/finite_limits.lean @@ -201,9 +201,9 @@ class has_finite_wide_pullbacks : Prop := (out (J : Type) [fintype J] : has_limits_of_shape (wide_pullback_shape J) C) instance has_limits_of_shape_wide_pullback_shape - (J : Type) [fintype J] [has_finite_wide_pullbacks C] : + (J : Type) [finite J] [has_finite_wide_pullbacks C] : has_limits_of_shape (wide_pullback_shape J) C := -by { haveI := @has_finite_wide_pullbacks.out C _ _ J, apply_instance } +by { casesI nonempty_fintype J, haveI := @has_finite_wide_pullbacks.out C _ _ J, apply_instance } /-- `has_finite_wide_pushouts` represents a choice of wide pushout @@ -213,9 +213,9 @@ class has_finite_wide_pushouts : Prop := (out (J : Type) [fintype J] : has_colimits_of_shape (wide_pushout_shape J) C) instance has_colimits_of_shape_wide_pushout_shape - (J : Type) [fintype J] [has_finite_wide_pushouts C] : + (J : Type) [finite J] [has_finite_wide_pushouts C] : has_colimits_of_shape (wide_pushout_shape J) C := -by { haveI := @has_finite_wide_pushouts.out C _ _ J, apply_instance } +by { casesI nonempty_fintype J, haveI := @has_finite_wide_pushouts.out C _ _ J, apply_instance } /-- Finite wide pullbacks are finite limits, so if `C` has all finite limits, diff --git a/src/category_theory/limits/shapes/finite_products.lean b/src/category_theory/limits/shapes/finite_products.lean index 56e23ddd3b4bf..dc772ea2e7d8d 100644 --- a/src/category_theory/limits/shapes/finite_products.lean +++ b/src/category_theory/limits/shapes/finite_products.lean @@ -33,9 +33,9 @@ class has_finite_products : Prop := (out (J : Type) [fintype J] : has_limits_of_shape (discrete J) C) instance has_limits_of_shape_discrete - (J : Type) [fintype J] [has_finite_products C] : + (J : Type) [finite J] [has_finite_products C] : has_limits_of_shape (discrete J) C := -by { haveI := @has_finite_products.out C _ _ J, apply_instance } +by { casesI nonempty_fintype J, haveI := @has_finite_products.out C _ _ J, apply_instance } /-- If `C` has finite limits then it has finite products. -/ @[priority 10] @@ -43,8 +43,9 @@ instance has_finite_products_of_has_finite_limits [has_finite_limits C] : has_finite_products C := ⟨λ J 𝒥, by { resetI, apply_instance }⟩ -instance has_fintype_products [has_finite_products C] (ι : Type w) [fintype ι] : +instance has_fintype_products [has_finite_products C] (ι : Type w) [finite ι] : has_limits_of_shape (discrete ι) C := +by casesI nonempty_fintype ι; exact has_limits_of_shape_of_equivalence (discrete.equivalence (equiv.ulift.{0}.trans @@ -69,9 +70,9 @@ class has_finite_coproducts : Prop := attribute [class] has_finite_coproducts instance has_colimits_of_shape_discrete - (J : Type) [fintype J] [has_finite_coproducts C] : + (J : Type) [finite J] [has_finite_coproducts C] : has_colimits_of_shape (discrete J) C := -by { haveI := @has_finite_coproducts.out C _ _ J, apply_instance } +by { casesI nonempty_fintype J, haveI := @has_finite_coproducts.out C _ _ J, apply_instance } /-- If `C` has finite colimits then it has finite coproducts. -/ @[priority 10] @@ -81,6 +82,7 @@ instance has_finite_coproducts_of_has_finite_colimits [has_finite_colimits C] : instance has_fintype_coproducts [has_finite_coproducts C] (ι : Type w) [fintype ι] : has_colimits_of_shape (discrete ι) C := +by casesI nonempty_fintype ι; exact has_colimits_of_shape_of_equivalence (discrete.equivalence (equiv.ulift.{0}.trans diff --git a/src/category_theory/preadditive/biproducts.lean b/src/category_theory/preadditive/biproducts.lean index d397c0813732a..8f8c976b4e853 100644 --- a/src/category_theory/preadditive/biproducts.lean +++ b/src/category_theory/preadditive/biproducts.lean @@ -270,11 +270,12 @@ end variables [preadditive.{v} C] lemma biproduct.column_nonzero_of_iso' - {σ τ : Type} [fintype τ] + {σ τ : Type} [finite τ] {S : σ → C} [has_biproduct S] {T : τ → C} [has_biproduct T] (s : σ) (f : ⨁ S ⟶ ⨁ T) [is_iso f] : (∀ t : τ, biproduct.ι S s ≫ f ≫ biproduct.π T t = 0) → 𝟙 (S s) = 0 := begin + casesI nonempty_fintype τ, intro z, set x := biproduct.ι S s ≫ f ≫ inv f ≫ biproduct.π S s, have h₁ : x = 𝟙 (S s), by simp [x], From 3c227fc0bf65ba9eb3048eb7845f8ce2a6084d45 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Tue, 6 Sep 2022 14:33:30 +0000 Subject: [PATCH 183/828] feat(data/nat/parity): add lemmas nat.bit{0|1}_mod_bit0 (#16396) This PR adds a pair of lemmas that complement [nat.bit0_div_bit0](https://leanprover-community.github.io/mathlib_docs/data/nat/parity.html#nat.bit0_div_bit0) and [nat.bit1_div_bit0](https://leanprover-community.github.io/mathlib_docs/data/nat/parity.html#nat.bit1_div_bit0), namely ```lean @[simp] lemma bit0_mod_bit0 : bit0 n % bit0 m = bit0 (n % m) := ... @[simp] lemma bit1_mod_bit0 : bit1 n % bit0 m = bit1 (n % m) := ... ``` They will be helpful in an upcoming PR that introduces a `norm_num` extension for computing Jacobi symbols. [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/bit0.2Fbit1.20and.20mod/near/297228893) --- src/data/nat/basic.lean | 4 ++-- src/data/nat/parity.lean | 13 +++++++++++++ 2 files changed, 15 insertions(+), 2 deletions(-) diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index de379321d77c9..a5cbafbb0391e 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -1626,9 +1626,9 @@ by { convert bit1_lt_bit0_iff, refl, } | ff := bit0_le_bit1_iff | tt := bit1_le_bit1 -@[simp] lemma bit0_mod_two : bit0 n % 2 = 0 := by { rw nat.mod_two_of_bodd, simp } +lemma bit0_mod_two : bit0 n % 2 = 0 := by { rw nat.mod_two_of_bodd, simp } -@[simp] lemma bit1_mod_two : bit1 n % 2 = 1 := by { rw nat.mod_two_of_bodd, simp } +lemma bit1_mod_two : bit1 n % 2 = 1 := by { rw nat.mod_two_of_bodd, simp } lemma pos_of_bit0_pos {n : ℕ} (h : 0 < bit0 n) : 0 < n := by { cases n, cases h, apply succ_pos, } diff --git a/src/data/nat/parity.lean b/src/data/nat/parity.lean index b284576f9d898..9851b36c3f407 100644 --- a/src/data/nat/parity.lean +++ b/src/data/nat/parity.lean @@ -210,6 +210,19 @@ by rw [bit0_eq_two_mul m, ←nat.div_div_eq_div_mul, bit0_div_two] @[simp] lemma bit1_div_bit0 : bit1 n / bit0 m = n / m := by rw [bit0_eq_two_mul, ←nat.div_div_eq_div_mul, bit1_div_two] +@[simp] lemma bit0_mod_bit0 : bit0 n % bit0 m = bit0 (n % m) := +by rw [bit0_eq_two_mul n, bit0_eq_two_mul m, bit0_eq_two_mul (n % m), nat.mul_mod_mul_left] + +@[simp] lemma bit1_mod_bit0 : bit1 n % bit0 m = bit1 (n % m) := +begin + have h₁ := congr_arg bit1 (nat.div_add_mod n m), + -- `∀ m n : ℕ, bit0 m * n = bit0 (m * n)` seems to be missing... + rw [bit1_add, bit0_eq_two_mul, ← mul_assoc, ← bit0_eq_two_mul] at h₁, + have h₂ := nat.div_add_mod (bit1 n) (bit0 m), + rw [bit1_div_bit0] at h₂, + exact add_left_cancel (h₂.trans h₁.symm), +end + -- Here are examples of how `parity_simps` can be used with `nat`. example (m n : ℕ) (h : even m) : ¬ even (n + 3) ↔ even (m^2 + m + n) := From b799b43dc0e7071921134fa07209837eabadf7ba Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 6 Sep 2022 17:23:42 +0000 Subject: [PATCH 184/828] feat(data/fin/basic): eq_succ_of_ne_zero, coe_cast_pred, succ_pred_above_succ (#16294) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces three basic lemmas about `fin`. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- src/data/fin/basic.lean | 41 +++++++++++++++++++++++++++++++++++++++++ 1 file changed, 41 insertions(+) diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index de7bea16471e7..bd02a922c3d66 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -270,6 +270,13 @@ begin { right, exact ⟨⟨j, nat.lt_of_succ_lt_succ h⟩, rfl⟩, } end +lemma eq_succ_of_ne_zero {n : ℕ} {i : fin (n + 1)} (hi : i ≠ 0) : ∃ j : fin n, i = j.succ := +begin + cases eq_zero_or_eq_succ i, + { exfalso, exact hi h, }, + { exact h, }, +end + /-- The greatest value of `fin (n+1)` -/ def last (n : ℕ) : fin (n+1) := ⟨_, n.lt_succ_self⟩ @@ -1569,6 +1576,14 @@ begin simp [cast_pred, pred_above, this] end +lemma coe_cast_pred {n : ℕ} (a : fin (n + 2)) (hx : a < fin.last _) : + (a.cast_pred : ℕ) = a := +begin + rcases a with ⟨a, ha⟩, + rw cast_pred_mk, + exacts [rfl, hx], +end + lemma pred_above_below (p : fin (n + 1)) (i : fin (n + 2)) (h : i ≤ p.cast_succ) : p.pred_above i = i.cast_pred := begin @@ -1646,6 +1661,32 @@ begin { rwa [cast_succ_pred_eq_pred_cast_succ , fin.pred_le_pred_iff] } } end +/-- `succ` commutes with `pred_above`. -/ +@[simp] +lemma succ_pred_above_succ (a : fin n) (b : fin (n+1)) : + a.succ.pred_above b.succ = (a.pred_above b).succ := +begin + obtain h₁ | h₂ := lt_or_le a.cast_succ b, + { rw [fin.pred_above_above _ _ h₁, fin.succ_pred, + fin.pred_above_above, fin.pred_succ], + simpa only [fin.lt_iff_coe_lt_coe, fin.coe_cast_succ, + fin.coe_succ, add_lt_add_iff_right] using h₁, }, + { cases n, + { exfalso, + exact not_lt_zero' a.is_lt, }, + { rw [fin.pred_above_below a b h₂, fin.pred_above_below a.succ b.succ + (by simpa only [le_iff_coe_le_coe, coe_succ, coe_cast_succ, + add_le_add_iff_right] using h₂)], + ext, + have h₀ : (b : ℕ) < n+1, + { simp only [le_iff_coe_le_coe, coe_cast_succ] at h₂, + simpa only [lt_succ_iff] using h₂.trans a.is_le, }, + have h₁ : (b.succ : ℕ) < n+2, + { rw ← nat.succ_lt_succ_iff at h₀, + simpa only [coe_succ] using h₀, }, + simp only [coe_cast_pred b h₀, coe_cast_pred b.succ h₁, coe_succ], }, }, +end + @[simp] theorem cast_pred_cast_succ (i : fin (n + 1)) : cast_pred i.cast_succ = i := by simp [cast_pred, pred_above, le_last] From a0d1a84dfb13f8971b1fd55bd19352db8e75937f Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Tue, 6 Sep 2022 20:23:05 +0000 Subject: [PATCH 185/828] feat(topology/sheaves/*): sheafification preserves zero morphisms (#16381) --- src/category_theory/sites/plus.lean | 14 ++++++++++++++ src/category_theory/sites/sheafification.lean | 5 +++++ 2 files changed, 19 insertions(+) diff --git a/src/category_theory/sites/plus.lean b/src/category_theory/sites/plus.lean index dc555f51452fd..29e1ccd916a4d 100644 --- a/src/category_theory/sites/plus.lean +++ b/src/category_theory/sites/plus.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Adam Topaz -/ import category_theory.sites.sheaf +import category_theory.preadditive.additive_functor /-! @@ -74,6 +75,11 @@ begin erw category.comp_id end +@[simp] +lemma diagram_nat_trans_zero [preadditive D] (X : C) (P Q : Cᵒᵖ ⥤ D) : + J.diagram_nat_trans (0 : P ⟶ Q) X = 0 := +by { ext j x, dsimp, rw [zero_comp, multiequalizer.lift_ι, comp_zero] } + @[simp] lemma diagram_nat_trans_comp {P Q R : Cᵒᵖ ⥤ D} (η : P ⟶ Q) (γ : Q ⟶ R) (X : C) : J.diagram_nat_trans (η ≫ γ) X = J.diagram_nat_trans η X ≫ J.diagram_nat_trans γ X := @@ -160,6 +166,10 @@ begin simp, end +@[simp] +lemma plus_map_zero [preadditive D] (P Q : Cᵒᵖ ⥤ D) : J.plus_map (0 : P ⟶ Q) = 0 := +by { ext, erw [comp_zero, colimit.ι_map, J.diagram_nat_trans_zero, zero_comp] } + @[simp] lemma plus_map_comp {P Q R : Cᵒᵖ ⥤ D} (η : P ⟶ Q) (γ : Q ⟶ R) : J.plus_map (η ≫ γ) = J.plus_map η ≫ J.plus_map γ := @@ -337,4 +347,8 @@ begin rw [← category.assoc, ← J.to_plus_naturality, category.assoc, J.to_plus_plus_lift], end +instance plus_functor_preserves_zero_morphisms [preadditive D] : + (plus_functor J D).preserves_zero_morphisms := +{ map_zero' := λ F G, by { ext, dsimp, rw [J.plus_map_zero, nat_trans.app_zero] } } + end category_theory.grothendieck_topology diff --git a/src/category_theory/sites/sheafification.lean b/src/category_theory/sites/sheafification.lean index 792eb964d8788..3fa8f802d41f3 100644 --- a/src/category_theory/sites/sheafification.lean +++ b/src/category_theory/sites/sheafification.lean @@ -606,6 +606,11 @@ def presheaf_to_Sheaf : (Cᵒᵖ ⥤ D) ⥤ Sheaf J D := map_id' := λ P, Sheaf.hom.ext _ _ $ J.sheafify_map_id _, map_comp' := λ P Q R f g, Sheaf.hom.ext _ _ $ J.sheafify_map_comp _ _ } +instance presheaf_to_Sheaf_preserves_zero_morphisms [preadditive D] : + (presheaf_to_Sheaf J D).preserves_zero_morphisms := +{ map_zero' := λ F G, by { ext, erw [colimit.ι_map, comp_zero, J.plus_map_zero, + J.diagram_nat_trans_zero, zero_comp] } } + /-- The sheafification functor is left adjoint to the forgetful functor. -/ @[simps unit_app counit_app_val] def sheafification_adjunction : presheaf_to_Sheaf J D ⊣ Sheaf_to_presheaf J D := From 18cf2ca32c43a5a9f1e8a9f07222d8ede7daf672 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Wed, 7 Sep 2022 03:48:54 +0000 Subject: [PATCH 186/828] chore(scripts): update nolints.txt (#16409) I am happy to remove some nolints for you! --- scripts/nolints.txt | 15 --------------- 1 file changed, 15 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 39e90c94e7f75..639382f051ec5 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -10,25 +10,10 @@ apply_nolint mul_hom.ext_iff to_additive_doc apply_nolint one_hom.comp_assoc to_additive_doc -- category_theory/limits/filtered_colimit_commutes_finite_limit.lean -apply_nolint category_theory.limits.colimit_limit_to_limit_colimit_injective fintype_finite apply_nolint category_theory.limits.colimit_limit_to_limit_colimit_is_iso fails_quickly --- category_theory/limits/shapes/biproducts.lean -apply_nolint category_theory.limits.has_biproduct.of_has_coproduct fintype_finite -apply_nolint category_theory.limits.has_biproduct.of_has_product fintype_finite - --- category_theory/limits/shapes/finite_limits.lean -apply_nolint category_theory.limits.has_colimits_of_shape_wide_pushout_shape fintype_finite -apply_nolint category_theory.limits.has_limits_of_shape_wide_pullback_shape fintype_finite - -- category_theory/limits/shapes/finite_products.lean -apply_nolint category_theory.limits.has_colimits_of_shape_discrete fintype_finite apply_nolint category_theory.limits.has_fintype_coproducts fintype_finite -apply_nolint category_theory.limits.has_fintype_products fintype_finite -apply_nolint category_theory.limits.has_limits_of_shape_discrete fintype_finite - --- category_theory/preadditive/biproducts.lean -apply_nolint category_theory.biproduct.column_nonzero_of_iso' fintype_finite -- computability/partrec.lean apply_nolint computable doc_blame From 914d42168720303f019f79b3da2916cb0e6e831c Mon Sep 17 00:00:00 2001 From: Alexander Bentkamp Date: Wed, 7 Sep 2022 06:39:14 +0000 Subject: [PATCH 187/828] feat(linear_algebra/matrix): A positive definite matrix has a positive determinant (#16376) --- src/linear_algebra/matrix/pos_def.lean | 20 +++++++++++++ src/linear_algebra/matrix/spectrum.lean | 39 +++++++++++++++++++++++++ 2 files changed, 59 insertions(+) diff --git a/src/linear_algebra/matrix/pos_def.lean b/src/linear_algebra/matrix/pos_def.lean index c584f1d8e917c..80149bb427397 100644 --- a/src/linear_algebra/matrix/pos_def.lean +++ b/src/linear_algebra/matrix/pos_def.lean @@ -56,6 +56,26 @@ begin apply hM.2 x hx, end +namespace pos_def + +variables {M : matrix n n ℝ} (hM : M.pos_def) +include hM + +lemma det_pos [decidable_eq n] : 0 < det M := +begin + rw hM.is_hermitian.det_eq_prod_eigenvalues, + apply finset.prod_pos, + intros i _, + rw hM.is_hermitian.eigenvalues_eq, + apply hM.2 _ (λ h, _), + have h_det : (hM.is_hermitian.eigenvector_matrix)ᵀ.det = 0, + from matrix.det_eq_zero_of_row_eq_zero i (λ j, congr_fun h j), + simpa only [h_det, not_is_unit_zero] using + is_unit_det_of_invertible hM.is_hermitian.eigenvector_matrixᵀ, +end + +end pos_def + end matrix namespace quadratic_form diff --git a/src/linear_algebra/matrix/spectrum.lean b/src/linear_algebra/matrix/spectrum.lean index 72e882caee369..1482feb242201 100644 --- a/src/linear_algebra/matrix/spectrum.lean +++ b/src/linear_algebra/matrix/spectrum.lean @@ -55,6 +55,33 @@ lemma eigenvector_matrix_mul_inv : hA.eigenvector_matrix ⬝ hA.eigenvector_matrix_inv = 1 := by apply basis.to_matrix_mul_to_matrix_flip +noncomputable instance : invertible hA.eigenvector_matrix_inv := +invertible_of_left_inverse _ _ hA.eigenvector_matrix_mul_inv + +noncomputable instance : invertible hA.eigenvector_matrix := +invertible_of_right_inverse _ _ hA.eigenvector_matrix_mul_inv + +lemma eigenvector_matrix_apply (i j : n) : hA.eigenvector_matrix i j = hA.eigenvector_basis j i := +by simp only [eigenvector_matrix, basis.to_matrix_apply, orthonormal_basis.coe_to_basis, + pi.basis_fun_repr] + +lemma eigenvector_matrix_inv_apply (i j : n) : + hA.eigenvector_matrix_inv i j = star (hA.eigenvector_basis i j) := +begin + rw [eigenvector_matrix_inv, basis.to_matrix_apply, orthonormal_basis.coe_to_basis_repr_apply, + pi.basis_fun_apply, linear_map.coe_std_basis, orthonormal_basis.repr_apply_apply], + change inner (hA.eigenvector_basis i) (euclidean_space.single j 1) = _, + rw [euclidean_space.inner_single_right], + simp only [one_mul, conj_transpose_apply, is_R_or_C.star_def], +end + +lemma conj_transpose_eigenvector_matrix_inv : hA.eigenvector_matrix_invᴴ = hA.eigenvector_matrix := +by { ext i j, + rw [conj_transpose_apply, eigenvector_matrix_inv_apply, eigenvector_matrix_apply, star_star] } + +lemma conj_transpose_eigenvector_matrix : hA.eigenvector_matrixᴴ = hA.eigenvector_matrix_inv := +by rw [← conj_transpose_eigenvector_matrix_inv, conj_transpose_conj_transpose] + /-- *Diagonalization theorem*, *spectral theorem* for matrices; A hermitian matrix can be diagonalized by a change of basis. @@ -81,6 +108,18 @@ begin euclidean_space.single, pi_Lp.equiv_symm_apply'] } end +lemma eigenvalues_eq (i : n) : + hA.eigenvalues i = + is_R_or_C.re ((star (hA.eigenvector_matrixᵀ i) ⬝ᵥ (A.mul_vec (hA.eigenvector_matrixᵀ i)))) := +begin + have := hA.spectral_theorem, + rw [←matrix.mul_inv_eq_iff_eq_mul_of_invertible] at this, + have := congr_arg is_R_or_C.re (congr_fun (congr_fun this i) i), + rw [diagonal_apply_eq, is_R_or_C.of_real_re, inv_eq_left_inv hA.eigenvector_matrix_mul_inv, + ← conj_transpose_eigenvector_matrix, mul_mul_apply] at this, + exact this.symm, +end + /-- The determinant of a hermitian matrix is the product of its eigenvalues. -/ lemma det_eq_prod_eigenvalues : det A = ∏ i, hA.eigenvalues i := begin From 9362a5b56c1f50a147e462e8acd89b12fed87383 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 7 Sep 2022 08:06:16 +0000 Subject: [PATCH 188/828] docs(data/polynomial/ring_division, field_theory/splitting_field): update docs (#16406) --- src/data/polynomial/ring_division.lean | 12 ++++++++++++ src/field_theory/splitting_field.lean | 3 --- 2 files changed, 12 insertions(+), 3 deletions(-) diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index 11dd463a3012b..2dfa0d0812559 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -14,6 +14,18 @@ import algebra.polynomial.big_operators This file starts looking like the ring theory of $ R[X] $ +## Main definitions + +* `polynomial.roots p`: The multiset containing all the roots of `p`, including their + multiplicities. +* `polynomial.root_set p E`: The set of distinct roots of `p` in an algebra `E`. + +## Main statements + +* `polynomial.C_leading_coeff_mul_prod_multiset_X_sub_C`: If a polynomial has as many roots as its + degree, it can be written as the product of its leading coefficient with `∏ (X - a)` where `a` + ranges through its roots. + -/ noncomputable theory diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field.lean index 3b0ffd6ea1466..b9c950db65b53 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field.lean @@ -25,9 +25,6 @@ if it is the smallest field extension of `K` such that `f` splits. ## Main statements -* `polynomial.C_leading_coeff_mul_prod_multiset_X_sub_C`: If a polynomial has as many roots as its - degree, it can be written as the product of its leading coefficient with `∏ (X - a)` where `a` - ranges through its roots. * `lift_of_splits`: If `K` and `L` are field extensions of a field `F` and for some finite subset `S` of `K`, the minimal polynomial of every `x ∈ K` splits as a polynomial with coefficients in `L`, then `algebra.adjoin F S` embeds into `L`. From 63014d264d0ba8f404438240d2bac09126838200 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 7 Sep 2022 09:54:37 +0000 Subject: [PATCH 189/828] feat(order/interval): Order intervals (#14807) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `nonempty_interval`, the type of nonempty intervals in an order, namely pairs where the second element is greater than the first, and `interval α` as `with_bot (nonempty_interval α)`. --- src/data/option/basic.lean | 6 +- src/data/set/lattice.lean | 12 + src/order/bounded_order.lean | 15 +- src/order/interval.lean | 427 +++++++++++++++++++++++++++++++++++ 4 files changed, 457 insertions(+), 3 deletions(-) create mode 100644 src/order/interval.lean diff --git a/src/data/option/basic.lean b/src/data/option/basic.lean index a64fd190b0bc1..f8587890de76a 100644 --- a/src/data/option/basic.lean +++ b/src/data/option/basic.lean @@ -32,7 +32,7 @@ along with a term `a : α` if the value is `true`. -/ namespace option -variables {α : Type*} {β : Type*} {γ : Type*} +variables {α β γ δ : Type*} lemma coe_def : (coe : α → option α) = some := rfl @@ -189,6 +189,10 @@ by { cases x; simp only [map_none', map_some', h, mem_def] } option.map h (option.map g x) = option.map (h ∘ g) x := by { cases x; simp only [map_none', map_some'] } +lemma map_comm {f₁ : α → β} {f₂ : α → γ} {g₁ : β → δ} {g₂ : γ → δ} (h : g₁ ∘ f₁ = g₂ ∘ f₂) (a : α) : + (option.map f₁ a).map g₁ = (option.map f₂ a).map g₂ := +by rw [map_map, h, ←map_map] + lemma comp_map (h : β → γ) (g : α → β) (x : option α) : option.map (h ∘ g) x = option.map h (option.map g x) := (map_map _ _ _).symm diff --git a/src/data/set/lattice.lean b/src/data/set/lattice.lean index 04d2d57d8a31c..6c38cce8a7b6e 100644 --- a/src/data/set/lattice.lean +++ b/src/data/set/lattice.lean @@ -294,6 +294,18 @@ explicit for this purpose. -/ lemma Inter_subset_of_subset {s : ι → set α} {t : set α} (i : ι) (h : s i ⊆ t) : (⋂ i, s i) ⊆ t := @infi_le_of_le (set α) _ _ _ _ i h +/-- This rather trivial consequence of `subset_Union₂` is convenient with `apply`, and has `i` and +`j` explicit for this purpose. -/ +lemma subset_Union₂_of_subset {s : set α} {t : Π i, κ i → set α} (i : ι) (j : κ i) (h : s ⊆ t i j) : + s ⊆ ⋃ i j, t i j := +@le_supr₂_of_le (set α) _ _ _ _ _ i j h + +/-- This rather trivial consequence of `Inter₂_subset` is convenient with `apply`, and has `i` and +`j` explicit for this purpose. -/ +lemma Inter₂_subset_of_subset {s : Π i, κ i → set α} {t : set α} (i : ι) (j : κ i) (h : s i j ⊆ t) : + (⋂ i j, s i j) ⊆ t := +@infi₂_le_of_le (set α) _ _ _ _ _ i j h + lemma Union_mono {s t : ι → set α} (h : ∀ i, s i ⊆ t i) : (⋃ i, s i) ⊆ ⋃ i, t i := @supr_mono (set α) _ _ s t h diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index 95e91ab0e5085..9708d43353b22 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -35,13 +35,13 @@ instances for `Prop` and `fun`. Typical examples include `Prop` and `set α`. -/ -open order_dual +open function order_dual set_option old_structure_cmd true universes u v -variables {α : Type u} {β : Type v} +variables {α : Type u} {β : Type v} {γ δ : Type*} /-! ### Top, bottom element -/ @@ -548,6 +548,9 @@ meta instance {α : Type} [reflected _ α] [has_reflect α] : has_reflect (with_ instance : inhabited (with_bot α) := ⟨⊥⟩ +lemma coe_injective : injective (coe : α → with_bot α) := option.some_injective _ +@[norm_cast] lemma coe_inj : (a : with_bot α) = b ↔ a = b := option.some_inj + lemma none_eq_bot : (none : with_bot α) = (⊥ : with_bot α) := rfl lemma some_eq_coe (a : α) : (some a : with_bot α) = (↑a : with_bot α) := rfl @@ -580,6 +583,10 @@ def map (f : α → β) : with_bot α → with_bot β := option.map f @[simp] lemma map_bot (f : α → β) : map f ⊥ = ⊥ := rfl @[simp] lemma map_coe (f : α → β) (a : α) : map f a = f a := rfl +lemma map_comm {f₁ : α → β} {f₂ : α → γ} {g₁ : β → δ} {g₂ : γ → δ} (h : g₁ ∘ f₁ = g₂ ∘ f₂) (a : α) : + map g₁ (map f₁ a) = map g₂ (map f₂ a) := +option.map_comm h _ + lemma ne_bot_iff_exists {x : with_bot α} : x ≠ ⊥ ↔ ∃ (a : α), ↑a = x := option.ne_none_iff_exists /-- Deconstruct a `x : with_bot α` to the underlying value in `α`, given a proof that `x ≠ ⊥`. -/ @@ -917,6 +924,10 @@ def map (f : α → β) : with_top α → with_top β := option.map f @[simp] lemma map_top (f : α → β) : map f ⊤ = ⊤ := rfl @[simp] lemma map_coe (f : α → β) (a : α) : map f a = f a := rfl +lemma map_comm {f₁ : α → β} {f₂ : α → γ} {g₁ : β → δ} {g₂ : γ → δ} (h : g₁ ∘ f₁ = g₂ ∘ f₂) (a : α) : + map g₁ (map f₁ a) = map g₂ (map f₂ a) := +option.map_comm h _ + lemma map_to_dual (f : αᵒᵈ → βᵒᵈ) (a : with_bot α) : map f (with_bot.to_dual a) = a.map (to_dual ∘ f) := rfl lemma map_of_dual (f : α → β) (a : with_bot αᵒᵈ) : diff --git a/src/order/interval.lean b/src/order/interval.lean new file mode 100644 index 0000000000000..76defaf048fa9 --- /dev/null +++ b/src/order/interval.lean @@ -0,0 +1,427 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.set.intervals.basic +import data.set.lattice +import data.set_like.basic + +/-! +# Order intervals + +This file defines (nonempty) closed intervals in an order (see `set.Icc`). This is a prototype for +interval arithmetic. + +## Main declarations + +* `nonempty_interval`: Nonempty intervals. Pairs where the second element is greater than the first. +* `interval`: Intervals. Either `∅` or a nonempty interval. +-/ + +open function order_dual set + +variables {α β γ : Type*} {ι : Sort*} {κ : ι → Sort*} + +/-- The nonempty closed intervals in an order. + +We define intervals by the pair of endpoints `fst`, `snd`. To convert intervals to the set of +elements between these endpoints, use the coercion `nonempty_interval α → set α`. -/ +@[ext] structure nonempty_interval (α : Type*) [has_le α] extends α × α := +(fst_le_snd : fst ≤ snd) + +namespace nonempty_interval +section has_le +variables [has_le α] {s t : nonempty_interval α} + +/-- The injection that induces the order on intervals. -/ +def to_dual_prod : nonempty_interval α → αᵒᵈ × α := to_prod + +@[simp] lemma to_dual_prod_apply (s : nonempty_interval α) : + s.to_dual_prod = (to_dual s.fst, s.snd) := prod.mk.eta.symm + +lemma to_dual_prod_injective : injective (to_dual_prod : nonempty_interval α → αᵒᵈ × α) := +λ s t, (ext_iff _ _).2 + +instance [is_empty α] : is_empty (nonempty_interval α) := ⟨λ s, is_empty_elim s.fst⟩ +instance [subsingleton α] : subsingleton (nonempty_interval α) := +to_dual_prod_injective.subsingleton + +instance : has_le (nonempty_interval α) := ⟨λ s t, t.fst ≤ s.fst ∧ s.snd ≤ t.snd⟩ + +lemma le_def : s ≤ t ↔ t.fst ≤ s.fst ∧ s.snd ≤ t.snd := iff.rfl + +/-- `to_dual_prod` as an order embedding. -/ +@[simps] def to_dual_prod_hom : nonempty_interval α ↪o αᵒᵈ × α := +{ to_fun := to_dual_prod, + inj' := to_dual_prod_injective, + map_rel_iff' := λ _ _, iff.rfl } + +/-- Turn an interval into an interval in the dual order. -/ +def dual : nonempty_interval α ≃ nonempty_interval αᵒᵈ := +{ to_fun := λ s, ⟨s.to_prod.swap, s.fst_le_snd⟩, + inv_fun := λ s, ⟨s.to_prod.swap, s.fst_le_snd⟩, + left_inv := λ s, ext _ _ $ prod.swap_swap _, + right_inv := λ s, ext _ _ $ prod.swap_swap _ } + +@[simp] lemma fst_dual (s : nonempty_interval α) : s.dual.fst = to_dual s.snd := rfl +@[simp] lemma snd_dual (s : nonempty_interval α) : s.dual.snd = to_dual s.fst := rfl + +end has_le + +section preorder +variables [preorder α] [preorder β] [preorder γ] + +instance : preorder (nonempty_interval α) := preorder.lift to_dual_prod + +/-- `{a}` as an interval. -/ +@[simps] def pure (a : α) : nonempty_interval α := ⟨⟨a, a⟩, le_rfl⟩ + +lemma pure_injective : injective (pure : α → nonempty_interval α) := +λ s t, congr_arg $ prod.fst ∘ to_prod + +@[simp] lemma dual_pure (a : α) : (pure a).dual = pure (to_dual a) := rfl + +instance [inhabited α] : inhabited (nonempty_interval α) := ⟨pure default⟩ +instance : ∀ [nonempty α], nonempty (nonempty_interval α) := nonempty.map pure +instance [nontrivial α] : nontrivial (nonempty_interval α) := pure_injective.nontrivial + +/-- Pushforward of nonempty intervals. -/ +@[simps] def map (f : α →o β) (a : nonempty_interval α) : nonempty_interval β := +⟨a.to_prod.map f f, f.mono a.fst_le_snd⟩ + +@[simp] lemma map_pure (f : α →o β) (a : α) : (pure a).map f = pure (f a) := rfl +@[simp] lemma map_map (g : β →o γ) (f : α →o β) (a : nonempty_interval α) : + (a.map f).map g = a.map (g.comp f) := rfl + +@[simp] lemma dual_map (f : α →o β) (a : nonempty_interval α) : + (a.map f).dual = a.dual.map f.dual := rfl + +variables [bounded_order α] + +instance : order_top (nonempty_interval α) := +{ top := ⟨⟨⊥, ⊤⟩, bot_le⟩, + le_top := λ a, ⟨bot_le, le_top⟩ } + +@[simp] lemma dual_top : (⊤ : nonempty_interval α).dual = ⊤ := rfl + +end preorder + +section partial_order +variables [partial_order α] {s t : nonempty_interval α} {x : α × α} {a : α} + +instance : partial_order (nonempty_interval α) := partial_order.lift _ to_dual_prod_injective + +/-- Consider a nonempty interval `[a, b]` as the set `[a, b]`. -/ +def coe_hom : nonempty_interval α ↪o set α := +order_embedding.of_map_le_iff (λ s, Icc s.fst s.snd) (λ s t, Icc_subset_Icc_iff s.fst_le_snd) + +instance : set_like (nonempty_interval α) α := +{ coe := λ s, Icc s.fst s.snd, + coe_injective' := coe_hom.injective } + +@[simp] lemma mem_mk {hx : x.1 ≤ x.2} : a ∈ mk x hx ↔ x.1 ≤ a ∧ a ≤ x.2 := iff.rfl +lemma mem_def : a ∈ s ↔ s.fst ≤ a ∧ a ≤ s.snd := iff.rfl + +@[simp, norm_cast] lemma coe_subset_coe : (s : set α) ⊆ t ↔ s ≤ t := (@coe_hom α _).le_iff_le +@[simp, norm_cast] lemma coe_ssubset_coe : (s : set α) ⊂ t ↔ s < t := (@coe_hom α _).lt_iff_lt + +@[simp] lemma coe_nonempty (s : nonempty_interval α) : (s : set α).nonempty := +nonempty_Icc.2 s.fst_le_snd + +@[simp] lemma coe_coe_hom : (coe_hom : nonempty_interval α → set α) = coe := rfl + +@[simp, norm_cast] lemma coe_pure (s : α) : (pure s : set α) = {s} := Icc_self _ + +@[simp, norm_cast] +lemma coe_top [bounded_order α] : ((⊤ : nonempty_interval α) : set α) = univ := Icc_bot_top + +@[simp, norm_cast] +lemma coe_dual (s : nonempty_interval α) : (s.dual : set αᵒᵈ) = of_dual ⁻¹' s := dual_Icc + +end partial_order + +section lattice +variables [lattice α] + +instance : has_sup (nonempty_interval α) := +⟨λ s t, ⟨⟨s.fst ⊓ t.fst, s.snd ⊔ t.snd⟩, inf_le_left.trans $ s.fst_le_snd.trans le_sup_left⟩⟩ + +instance : semilattice_sup (nonempty_interval α) := +to_dual_prod_injective.semilattice_sup _ $ λ _ _, rfl + +@[simp] lemma fst_sup (s t : nonempty_interval α) : (s ⊔ t).fst = s.fst ⊓ t.fst := rfl +@[simp] lemma snd_sup (s t : nonempty_interval α) : (s ⊔ t).snd = s.snd ⊔ t.snd := rfl + +end lattice +end nonempty_interval + +/-- The closed intervals in an order. + +We represent intervals either as `⊥` or a nonempty interval given by its endpoints `fst`, `snd`. +To convert intervals to the set of elements between these endpoints, use the coercion +`interval α → set α`. -/ +@[derive [inhabited, has_le, order_bot]] +def interval (α : Type*) [has_le α] := with_bot (nonempty_interval α) + +namespace interval +section has_le +variables [has_le α] {s t : interval α} + +instance : has_coe_t (nonempty_interval α) (interval α) := with_bot.has_coe_t +instance : can_lift (interval α) (nonempty_interval α) := with_bot.can_lift + +lemma coe_injective : injective (coe : nonempty_interval α → interval α) := with_bot.coe_injective +@[simp, norm_cast] lemma coe_inj {s t : nonempty_interval α} : (s : interval α) = t ↔ s = t := +with_bot.coe_inj + +@[protected] lemma «forall» {p : interval α → Prop} : + (∀ s, p s) ↔ p ⊥ ∧ ∀ s : nonempty_interval α, p s := option.forall +@[protected] lemma «exists» {p : interval α → Prop} : + (∃ s, p s) ↔ p ⊥ ∨ ∃ s : nonempty_interval α, p s := option.exists + +instance [is_empty α] : unique (interval α) := option.unique + +/-- Turn an interval into an interval in the dual order. -/ +def dual : interval α ≃ interval αᵒᵈ := nonempty_interval.dual.option_congr + +end has_le + +section preorder +variables [preorder α] [preorder β] [preorder γ] + +instance : preorder (interval α) := with_bot.preorder + +/-- `{a}` as an interval. -/ +def pure (a : α) : interval α := nonempty_interval.pure a + +lemma pure_injective : injective (pure : α → interval α) := +coe_injective.comp nonempty_interval.pure_injective + +@[simp] lemma dual_pure (a : α) : (pure a).dual = pure (to_dual a) := rfl +@[simp] lemma dual_bot : (⊥ : interval α).dual = ⊥ := rfl + +instance [nonempty α] : nontrivial (interval α) := option.nontrivial + +/-- Pushforward of intervals. -/ +def map (f : α →o β) : interval α → interval β := with_bot.map (nonempty_interval.map f) + +@[simp] lemma map_pure (f : α →o β) (a : α) : (pure a).map f = pure (f a) := rfl +@[simp] lemma map_map (g : β →o γ) (f : α →o β) (s : interval α) : + (s.map f).map g = s.map (g.comp f) := option.map_map _ _ _ + +@[simp] lemma dual_map (f : α →o β) (s : interval α) : (s.map f).dual = s.dual.map f.dual := +by { cases s, { refl }, { exact with_bot.map_comm rfl _ } } + +variables [bounded_order α] + +instance : bounded_order (interval α) := with_bot.bounded_order + +@[simp] lemma dual_top : (⊤ : interval α).dual = ⊤ := rfl + +end preorder + +section partial_order +variables [partial_order α] {s t : interval α} + +instance : partial_order (interval α) := with_bot.partial_order + +/-- Consider a interval `[a, b]` as the set `[a, b]`. -/ +def coe_hom : interval α ↪o set α := +order_embedding.of_map_le_iff (λ s, match s with + | ⊥ := ∅ + | some s := s + end) (λ s t, match s, t with + | ⊥, t := iff_of_true bot_le bot_le + | some s, ⊥ := iff_of_false (λ h, s.coe_nonempty.ne_empty $ le_bot_iff.1 h) + (with_bot.not_coe_le_bot _) + | some s, some t := (@nonempty_interval.coe_hom α _).le_iff_le.trans with_bot.some_le_some.symm + end) + +instance : set_like (interval α) α := +{ coe := coe_hom, + coe_injective' := coe_hom.injective } + +@[simp, norm_cast] lemma coe_subset_coe : (s : set α) ⊆ t ↔ s ≤ t := (@coe_hom α _).le_iff_le +@[simp, norm_cast] lemma coe_ssubset_coe : (s : set α) ⊂ t ↔ s < t := (@coe_hom α _).lt_iff_lt +@[simp, norm_cast] lemma coe_pure (a : α) : (pure a : set α) = {a} := Icc_self _ +@[simp, norm_cast] lemma coe_coe (s : nonempty_interval α) : ((s : interval α) : set α) = s := rfl +@[simp, norm_cast] lemma coe_bot : ((⊥ : interval α) : set α) = ∅ := rfl +@[simp, norm_cast] lemma coe_top [bounded_order α] : ((⊤ : interval α) : set α) = univ := +Icc_bot_top +@[simp, norm_cast] lemma coe_dual (s : interval α) : (s.dual : set αᵒᵈ) = of_dual ⁻¹' s := +by { cases s, { refl }, exact s.coe_dual } + +end partial_order + +section lattice +variables [lattice α] + +instance : semilattice_sup (interval α) := with_bot.semilattice_sup + +variables [@decidable_rel α (≤)] + +instance : lattice (interval α) := +{ inf := λ s t, match s, t with + | ⊥, t := ⊥ + | s, ⊥ := ⊥ + | some s, some t := if h : s.fst ≤ t.snd ∧ t.fst ≤ s.snd then some + ⟨⟨s.fst ⊔ t.fst, s.snd ⊓ t.snd⟩, sup_le (le_inf s.fst_le_snd h.1) $ le_inf h.2 t.fst_le_snd⟩ + else ⊥ + end, + inf_le_left := λ s t, match s, t with + | ⊥, ⊥ := bot_le + | ⊥, some t := bot_le + | some s, ⊥ := bot_le + | some s, some t := begin + change dite _ _ _ ≤ _, + split_ifs, + { exact with_bot.some_le_some.2 ⟨le_sup_left, inf_le_left⟩ }, + { exact bot_le } + end + end, + inf_le_right := λ s t, match s, t with + | ⊥, ⊥ := bot_le + | ⊥, some t := bot_le + | some s, ⊥ := bot_le + | some s, some t := begin + change dite _ _ _ ≤ _, + split_ifs, + { exact with_bot.some_le_some.2 ⟨le_sup_right, inf_le_right⟩ }, + { exact bot_le } + end + end, + le_inf := λ s t c, match s, t, c with + | ⊥, t, c := λ _ _, bot_le + | some s, t, c := λ hb hc, begin + lift t to nonempty_interval α using ne_bot_of_le_ne_bot with_bot.coe_ne_bot hb, + lift c to nonempty_interval α using ne_bot_of_le_ne_bot with_bot.coe_ne_bot hc, + change _ ≤ dite _ _ _, + simp only [with_bot.some_eq_coe, with_bot.coe_le_coe] at ⊢ hb hc, + rw [dif_pos, with_bot.coe_le_coe], + exact ⟨sup_le hb.1 hc.1, le_inf hb.2 hc.2⟩, + exact ⟨hb.1.trans $ s.fst_le_snd.trans hc.2, hc.1.trans $ s.fst_le_snd.trans hb.2⟩, + end + end, + ..interval.semilattice_sup } + +@[simp, norm_cast] lemma coe_inf (s t : interval α) : (↑(s ⊓ t) : set α) = s ∩ t := +begin + cases s, + { rw [with_bot.none_eq_bot, bot_inf_eq], + exact (empty_inter _).symm }, + cases t, + { rw [with_bot.none_eq_bot, inf_bot_eq], + exact (inter_empty _).symm }, + refine (_ : coe (dite _ _ _) = _).trans Icc_inter_Icc.symm, + split_ifs, + { refl }, + { exact (Icc_eq_empty $ λ H, + h ⟨le_sup_left.trans $ H.trans inf_le_right, le_sup_right.trans $ H.trans inf_le_left⟩).symm } +end + +@[simp, norm_cast] +lemma disjoint_coe (s t : interval α) : disjoint (s : set α) t ↔ disjoint s t := +by { rw [disjoint, disjoint, le_eq_subset, ←coe_subset_coe, coe_inf], refl } + +end lattice +end interval + +namespace nonempty_interval +section preorder +variables [preorder α] + +@[simp, norm_cast] lemma coe_pure_interval (s : α) : (pure s : interval α) = interval.pure s := rfl + +@[simp, norm_cast] +lemma coe_top_interval [bounded_order α] : ((⊤ : nonempty_interval α) : interval α) = ⊤ := rfl + +end preorder + +@[simp, norm_cast] +lemma mem_coe_interval [partial_order α] {s : nonempty_interval α} {x : α} : + x ∈ (s : interval α) ↔ x ∈ s := iff.rfl + +@[simp, norm_cast] lemma coe_sup_interval [lattice α] (s t : nonempty_interval α) : + (↑(s ⊔ t) : interval α) = s ⊔ t := rfl + +end nonempty_interval + +namespace interval +section complete_lattice +variables [complete_lattice α] + +noncomputable instance [@decidable_rel α (≤)] : complete_lattice (interval α) := +by classical; exact { Sup := λ S, if h : S ⊆ {⊥} then ⊥ else some + ⟨⟨⨅ (s : nonempty_interval α) (h : ↑s ∈ S), s.fst, + ⨆ (s : nonempty_interval α) (h : ↑s ∈ S), s.snd⟩, begin + obtain ⟨s, hs, ha⟩ := not_subset.1 h, + lift s to nonempty_interval α using ha, + exact infi₂_le_of_le s hs (le_supr₂_of_le s hs s.fst_le_snd) + end⟩, + le_Sup := λ s s ha, begin + split_ifs, + { exact (h ha).le }, + cases s, + { exact bot_le }, + { exact with_bot.some_le_some.2 ⟨infi₂_le _ ha, le_supr₂_of_le _ ha le_rfl⟩ } + end, + Sup_le := λ s s ha, begin + split_ifs, + { exact bot_le }, + obtain ⟨b, hs, hb⟩ := not_subset.1 h, + lift s to nonempty_interval α using ne_bot_of_le_ne_bot hb (ha _ hs), + exact with_bot.coe_le_coe.2 ⟨le_infi₂ $ λ c hc, (with_bot.coe_le_coe.1 $ ha _ hc).1, + supr₂_le $ λ c hc, (with_bot.coe_le_coe.1 $ ha _ hc).2⟩, + end, + Inf := λ S, if h : ⊥ ∉ S ∧ ∀ ⦃s : nonempty_interval α⦄, ↑s ∈ S → ∀ ⦃t : nonempty_interval α⦄, + ↑t ∈ S → s.fst ≤ t.snd then some + ⟨⟨⨆ (s : nonempty_interval α) (h : ↑s ∈ S), s.fst, + ⨅ (s : nonempty_interval α) (h : ↑s ∈ S), s.snd⟩, + supr₂_le $ λ s hs, le_infi₂ $ h.2 hs⟩ else ⊥, + Inf_le := λ s s ha, begin + split_ifs, + { lift s to nonempty_interval α using ne_of_mem_of_not_mem ha h.1, + exact with_bot.coe_le_coe.2 ⟨le_supr₂ s ha, infi₂_le s ha⟩ }, + { exact bot_le } + end, + le_Inf := λ S s ha, begin + cases s, + { exact bot_le }, + split_ifs, + { exact with_bot.some_le_some.2 ⟨supr₂_le $ λ t hb, (with_bot.coe_le_coe.1 $ ha _ hb).1, + le_infi₂ $ λ t hb, (with_bot.coe_le_coe.1 $ ha _ hb).2⟩ }, + rw [not_and_distrib, not_not] at h, + cases h, + { exact ha _ h }, + cases h (λ t hb c hc, (with_bot.coe_le_coe.1 $ ha _ hb).1.trans $ s.fst_le_snd.trans + (with_bot.coe_le_coe.1 $ ha _ hc).2), + end, + ..interval.lattice, ..interval.bounded_order } + +@[simp, norm_cast] lemma coe_Inf (S : set (interval α)) : ↑(Inf S) = ⋂ s ∈ S, (s : set α) := +begin + change coe (dite _ _ _) = _, + split_ifs, + { ext, + simp [with_bot.some_eq_coe, interval.forall, h.1, ←forall_and_distrib, + ←nonempty_interval.mem_def] }, + simp_rw [not_and_distrib, not_not] at h, + cases h, + { refine (eq_empty_of_subset_empty _).symm, + exact Inter₂_subset_of_subset _ h subset.rfl }, + { refine (not_nonempty_iff_eq_empty.1 _).symm, + rintro ⟨x, hx⟩, + rw mem_Inter₂ at hx, + exact h (λ s ha t hb, (hx _ ha).1.trans (hx _ hb).2) } +end + +@[simp, norm_cast] lemma coe_infi (f : ι → interval α) : ↑(⨅ i, f i) = ⋂ i, (f i : set α) := +by simp [infi] + +@[simp, norm_cast] lemma coe_infi₂ (f : Π i, κ i → interval α) : + ↑(⨅ i j, f i j) = ⋂ i j, (f i j : set α) := +by simp_rw [coe_infi] + +end complete_lattice +end interval From 41626f7d6f841571268fc92b28bd879386c99541 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 7 Sep 2022 13:23:09 +0000 Subject: [PATCH 190/828] feat(data/polynomial/degree/definitions): add `degree_X_sub_C_le` (#16404) --- src/data/polynomial/algebra_map.lean | 2 +- src/data/polynomial/degree/definitions.lean | 7 ++++--- 2 files changed, 5 insertions(+), 4 deletions(-) diff --git a/src/data/polynomial/algebra_map.lean b/src/data/polynomial/algebra_map.lean index a94b9adb93fda..296dde611bffa 100644 --- a/src/data/polynomial/algebra_map.lean +++ b/src/data/polynomial/algebra_map.lean @@ -375,7 +375,7 @@ begin have bound := calc (p * (X - C r)).nat_degree ≤ p.nat_degree + (X - C r).nat_degree : nat_degree_mul_le - ... ≤ p.nat_degree + 1 : add_le_add_left nat_degree_X_sub_C_le _ + ... ≤ p.nat_degree + 1 : add_le_add_left (nat_degree_X_sub_C_le _) _ ... < p.nat_degree + 2 : lt_add_one _, rw sum_over_range' _ _ (p.nat_degree + 2) bound, swap, diff --git a/src/data/polynomial/degree/definitions.lean b/src/data/polynomial/degree/definitions.lean index 164821fa1cea1..87273d4ba7f5f 100644 --- a/src/data/polynomial/degree/definitions.lean +++ b/src/data/polynomial/degree/definitions.lean @@ -1047,10 +1047,11 @@ calc degree (p - q) = degree (erase (nat_degree q) p + -erase (nat_degree q) q) : degree_neg (erase (nat_degree q) q) ▸ degree_add_le _ _ ... < degree p : max_lt_iff.2 ⟨hd' ▸ degree_erase_lt hp0, hd.symm ▸ degree_erase_lt hq0⟩ +lemma degree_X_sub_C_le (r : R) : (X - C r).degree ≤ 1 := +(degree_sub_le _ _).trans (max_le degree_X_le (degree_C_le.trans zero_le_one)) -lemma nat_degree_X_sub_C_le {r : R} : (X - C r).nat_degree ≤ 1 := -nat_degree_le_iff_degree_le.2 $ le_trans (degree_sub_le _ _) $ max_le degree_X_le $ -le_trans degree_C_le $ with_bot.coe_le_coe.2 zero_le_one +lemma nat_degree_X_sub_C_le (r : R) : (X - C r).nat_degree ≤ 1 := +nat_degree_le_iff_degree_le.2 $ degree_X_sub_C_le r lemma degree_sub_eq_left_of_degree_lt (h : degree q < degree p) : degree (p - q) = degree p := by { rw ← degree_neg q at h, rw [sub_eq_add_neg, degree_add_eq_left_of_degree_lt h] } From 6211945146a1b095ddd65addcf5a6f97862b4a81 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Wed, 7 Sep 2022 15:10:37 +0000 Subject: [PATCH 191/828] feat(analysis/complex/upper_half_plane): Functions bounded at infinity (#15009) The defines the notion of functions on the upper half plane being bounded at infinity and zero at infinity. This is required for #13250. --- .../functions_bounded_at_infty.lean | 99 +++++++++++++++++++ .../filter/zero_and_bounded_at_filter.lean | 78 +++++++++++++++ 2 files changed, 177 insertions(+) create mode 100644 src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean create mode 100644 src/order/filter/zero_and_bounded_at_filter.lean diff --git a/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean new file mode 100644 index 0000000000000..d8732563824f1 --- /dev/null +++ b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean @@ -0,0 +1,99 @@ +/- +Copyright (c) 2022 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck, David Loeffler +-/ + +import algebra.module.submodule.basic +import analysis.complex.upper_half_plane.basic +import order.filter.zero_and_bounded_at_filter + +/-! +# Bounded at infinity + +For complex valued functions on the upper half plane, this file defines the filter `at_im_infty` +required for defining when functions are bounded at infinity and zero at infinity. +Both of which are relevant for defining modular forms. + +-/ + +open complex filter + +open_locale topological_space upper_half_plane + +noncomputable theory + +namespace upper_half_plane + +/-- Filter for approaching `i∞`. -/ +def at_im_infty := filter.at_top.comap upper_half_plane.im + +lemma at_im_infty_basis : (at_im_infty).has_basis (λ _, true) (λ (i : ℝ), im ⁻¹' set.Ici i) := +filter.has_basis.comap upper_half_plane.im filter.at_top_basis + +lemma at_im_infty_mem (S : set ℍ) : S ∈ at_im_infty ↔ (∃ A : ℝ, ∀ z : ℍ, A ≤ im z → z ∈ S) := +begin + simp only [at_im_infty, filter.mem_comap', filter.mem_at_top_sets, ge_iff_le, set.mem_set_of_eq, + upper_half_plane.coe_im], + refine ⟨λ ⟨a, h⟩, ⟨a, (λ z hz, h (im z) hz rfl)⟩, _⟩, + rintro ⟨A, h⟩, + refine ⟨A, λ b hb x hx, h x _⟩, + rwa hx, +end + +/-- A function ` f : ℍ → α` is bounded at infinity if it is bounded along `at_im_infty`. -/ +def is_bounded_at_im_infty {α : Type*} [has_norm α] [has_one (ℍ → α)] (f : ℍ → α) : Prop := +bounded_at_filter at_im_infty f + +/-- A function ` f : ℍ → α` is zero at infinity it is zero along `at_im_infty`. -/ +def is_zero_at_im_infty {α : Type*} [has_zero α] [topological_space α] (f : ℍ → α) : Prop := +zero_at_filter at_im_infty f + +lemma zero_form_is_bounded_at_im_infty {α : Type*} [normed_field α] : + is_bounded_at_im_infty (0 : ℍ → α) := +zero_is_bounded_at_filter at_im_infty + +/-- Module of functions that are zero at infinity. -/ +def zero_at_im_infty_submodule (α : Type*) [normed_field α] : submodule α (ℍ → α) := +zero_at_filter_submodule at_im_infty + +/-- ubalgebra of functions that are bounded at infinity. -/ +def bounded_at_im_infty_subalgebra (α : Type*) [normed_field α] : subalgebra α (ℍ → α) := +bounded_filter_subalgebra at_im_infty + +lemma prod_of_bounded_is_bounded {f g : ℍ → ℂ} (hf : is_bounded_at_im_infty f) + (hg : is_bounded_at_im_infty g) : is_bounded_at_im_infty (f * g) := +by simpa only [pi.one_apply, mul_one, norm_eq_abs, complex.abs_mul] using hf.mul hg + +@[simp] lemma bounded_mem (f : ℍ → ℂ) : + is_bounded_at_im_infty f ↔ ∃ (M A : ℝ), ∀ z : ℍ, A ≤ im z → abs (f z) ≤ M := +begin + simp [is_bounded_at_im_infty, bounded_at_filter, asymptotics.is_O_iff, filter.eventually, + at_im_infty_mem], +end + +lemma zero_at_im_infty (f : ℍ → ℂ) : + is_zero_at_im_infty f ↔ ∀ ε : ℝ, 0 < ε → ∃ A : ℝ, ∀ z : ℍ, A ≤ im z → abs (f z) ≤ ε := +begin + rw [is_zero_at_im_infty, zero_at_filter, tendsto_iff_forall_eventually_mem], + split, + { simp_rw [filter.eventually, at_im_infty_mem], + intros h ε hε, + simpa using (h (metric.closed_ball (0 : ℂ) ε) (metric.closed_ball_mem_nhds (0 : ℂ) hε))}, + { simp_rw metric.mem_nhds_iff, + intros h s hs, + simp_rw [filter.eventually, at_im_infty_mem], + obtain ⟨ε, h1, h2⟩ := hs, + have h11 : 0 < (ε/2), by {linarith,}, + obtain ⟨A, hA⟩ := (h (ε/2) h11), + use A, + intros z hz, + have hzs : f z ∈ s, + { apply h2, + simp only [mem_ball_zero_iff, norm_eq_abs], + apply lt_of_le_of_lt (hA z hz), + linarith }, + apply hzs,} +end + +end upper_half_plane diff --git a/src/order/filter/zero_and_bounded_at_filter.lean b/src/order/filter/zero_and_bounded_at_filter.lean new file mode 100644 index 0000000000000..3fec25c7b891b --- /dev/null +++ b/src/order/filter/zero_and_bounded_at_filter.lean @@ -0,0 +1,78 @@ +/- +Copyright (c) 2022 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck, David Loeffler +-/ + +import order.filter +import algebra.module.submodule.basic +import topology.algebra.monoid +import analysis.asymptotics.asymptotics + +/-! +# Zero and Bounded at filter + +Given a filter `l` we define the notion of a function being `zero_at_filter` as well as being +`bounded_at_filter`. Alongside this we construct the `submodule`, `add_submonoid` of functions +that are `zero_at_filter`. Similarly, we construct the `submodule` and `subalgebra` of functions +that are `bounded_at_filter`. + +-/ + +namespace filter + +variables {α β : Type*} + +open_locale topological_space + +/-- A function `f : α → β` is `zero_at_filter` if in the limit it is zero. -/ +def zero_at_filter [has_zero β] [topological_space β] (l : filter α) (f : α → β) : Prop := +filter.tendsto f l (𝓝 0) + +lemma zero_is_zero_at_filter [has_zero β] [topological_space β] (l : filter α) : zero_at_filter l + (0 : α → β) := tendsto_const_nhds + +/-- The submodule of funtions that are `zero_at_filter`. -/ +def zero_at_filter_submodule [topological_space β] [semiring β] + [has_continuous_add β] [has_continuous_mul β] (l : filter α) : submodule β (α → β) := +{ carrier := zero_at_filter l, + zero_mem' := zero_is_zero_at_filter l, + add_mem' := by { intros a b ha hb, simpa using ha.add hb, }, + smul_mem' := by { intros c f hf, simpa using hf.const_mul c }, } + +/-- The additive submonoid of funtions that are `zero_at_filter`. -/ +def zero_at_filter_add_submonoid [topological_space β] + [add_zero_class β] [has_continuous_add β] (l : filter α) : add_submonoid (α → β) := +{ carrier := zero_at_filter l, + add_mem' := by { intros a b ha hb, simpa using ha.add hb }, + zero_mem' := zero_is_zero_at_filter l, } + +/-- A function `f: α → β` is `bounded_at_filter` if `f =O[l] 1`. -/ +def bounded_at_filter [has_norm β] [has_one (α → β)] (l : filter α) (f : α → β) : Prop := +asymptotics.is_O l f (1 : α → β) + +lemma zero_at_filter_is_bounded_at_filter [normed_field β] (l : filter α) (f : α → β) + (hf : zero_at_filter l f) : bounded_at_filter l f := +asymptotics.is_O_of_div_tendsto_nhds (by simp) _ (by { convert hf, ext1, simp, }) + +lemma zero_is_bounded_at_filter [normed_field β] (l : filter α) : + bounded_at_filter l (0 : α → β) := +(zero_at_filter_is_bounded_at_filter l _) (zero_is_zero_at_filter l) + +/-- The submodule of funtions that are `bounded_at_filter`. -/ +def bounded_filter_submodule [normed_field β] (l : filter α) : submodule β (α → β) := +{ carrier := bounded_at_filter l, + zero_mem' := zero_is_bounded_at_filter l, + add_mem' := by { intros f g hf hg, simpa using hf.add hg, }, + smul_mem' := by { intros c f hf, simpa using hf.const_mul_left c }, } + +/-- The subalgebra of funtions that are `bounded_at_filter`. -/ +def bounded_filter_subalgebra [normed_field β] (l : filter α) : + subalgebra β (α → β) := +begin + refine submodule.to_subalgebra (bounded_filter_submodule l) _ (λ f g hf hg, _), + { simpa using asymptotics.is_O_const_mul_self (1 : β) (1 : α → β) l }, + { simpa only [pi.one_apply, mul_one, norm_mul] using hf.mul hg, }, +end + +end filter From 13c6c43774cb0f94122155856b8129b671b90833 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Wed, 7 Sep 2022 15:10:39 +0000 Subject: [PATCH 192/828] feat(group_theory/subgroup/basic): Centralizer of closure is intersection of centralizers (#16394) This PR adds some more API for `subgroup.centralizer` and proves that the centralizer of a closure is the intersection of the centralizers. --- src/group_theory/subgroup/basic.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 606d4c97c01e1..e48c5119b27b0 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1787,6 +1787,8 @@ def centralizer : subgroup G := inv_mem' := λ g, set.inv_mem_centralizer, .. submonoid.centralizer ↑H } +variables {H} + @[to_additive] lemma mem_centralizer_iff {g : G} : g ∈ H.centralizer ↔ ∀ h ∈ H, h * g = g * h := iff.rfl @@ -1797,6 +1799,12 @@ by simp only [mem_centralizer_iff, mul_inv_eq_iff_eq_mul, one_mul] @[to_additive] lemma centralizer_top : centralizer ⊤ = center G := set_like.ext' (set.centralizer_univ G) +@[to_additive] lemma le_centralizer_iff : H ≤ K.centralizer ↔ K ≤ H.centralizer := +⟨λ h x hx y hy, (h hy x hx).symm, λ h x hx y hy, (h hy x hx).symm⟩ + +@[to_additive] lemma centralizer_le (h : H ≤ K) : centralizer K ≤ centralizer H := +submonoid.centralizer_le h + @[to_additive] instance subgroup.centralizer.characteristic [hH : H.characteristic] : H.centralizer.characteristic := begin @@ -2699,6 +2707,12 @@ by rw [eq_bot_iff, zpowers_le, mem_bot] subgroup.zpowers (1 : G) = ⊥ := subgroup.zpowers_eq_bot.mpr rfl +@[to_additive] lemma centralizer_closure (S : set G) : + (closure S).centralizer = ⨅ g ∈ S, (zpowers g).centralizer := +le_antisymm (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ zpowers_le.2 $ subset_closure hg) + $ le_centralizer_iff.1 $ (closure_le _).2 + $ λ g, set_like.mem_coe.2 ∘ zpowers_le.1 ∘ le_centralizer_iff.1 ∘ infi_le_of_le g ∘ infi_le _ + end subgroup namespace monoid_hom From 56e4f0d51f97220717e05f284387c17ad4a25b0d Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Wed, 7 Sep 2022 15:10:40 +0000 Subject: [PATCH 193/828] feat(analysis/seminorm): `restrict_scalars` for seminorms (#16401) --- src/analysis/seminorm.lean | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index d5963af7f647b..fd30d228d9231 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -647,6 +647,29 @@ end end module end convex + +section restrict_scalars + +variables (𝕜) {𝕜' : Type*} [normed_field 𝕜] [semi_normed_ring 𝕜'] [normed_algebra 𝕜 𝕜'] + [norm_one_class 𝕜'] [add_comm_group E] [module 𝕜' E] [has_smul 𝕜 E] [is_scalar_tower 𝕜 𝕜' E] + +/-- Reinterpret a seminorm over a field `𝕜'` as a seminorm over a smaller field `𝕜`. This will +typically be used with `is_R_or_C 𝕜'` and `𝕜 = ℝ`. -/ +protected def restrict_scalars (p : seminorm 𝕜' E) : + seminorm 𝕜 E := +{ smul' := λ a x, by rw [← smul_one_smul 𝕜' a x, p.smul', norm_smul, norm_one, mul_one], + ..p } + +@[simp] lemma coe_restrict_scalars (p : seminorm 𝕜' E) : + (p.restrict_scalars 𝕜 : E → ℝ) = p := +rfl + +@[simp] lemma restrict_scalars_ball (p : seminorm 𝕜' E) : + (p.restrict_scalars 𝕜).ball = p.ball := +rfl + +end restrict_scalars + end seminorm /-! ### The norm as a seminorm -/ From eae1ec8198020c8dd7490a338e670c075e6e88bd Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 7 Sep 2022 23:23:50 +0000 Subject: [PATCH 194/828] feat(analysis/normed/group/add_torsor): `normed_add_torsor` instance for `affine_subspace` (#16339) Add an instance that a nonempty affine subspace of a `normed_add_torsor` is itself a `normed_add_torsor` (building on the existing instance that says such a subspace of an `add_torsor` is itself an `add_torsor`). Note that this instance uses `@[nolint fails_quickly]`, because of a typeclass loop with the `add_torsor.nonempty` instance. I don't know the right way to avoid that typeclass loop, but I don't think it's meaningfully a new issue, since exactly the same nolint appears in `nolints.txt` for `affine_subspace.to_add_torsor`. --- src/analysis/normed/group/add_torsor.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/analysis/normed/group/add_torsor.lean b/src/analysis/normed/group/add_torsor.lean index 77b484b3d4dc4..fcd36ce4130b0 100644 --- a/src/analysis/normed/group/add_torsor.lean +++ b/src/analysis/normed/group/add_torsor.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Yury Kudryashov -/ import analysis.normed.group.basic +import linear_algebra.affine_space.affine_subspace import linear_algebra.affine_space.midpoint /-! @@ -37,6 +38,13 @@ variables {α V P W Q : Type*} [seminormed_add_comm_group V] [pseudo_metric_spac instance seminormed_add_comm_group.to_normed_add_torsor : normed_add_torsor V V := { dist_eq_norm' := dist_eq_norm } +/-- A nonempty affine subspace of a `normed_add_torsor` is itself a `normed_add_torsor`. -/ +@[nolint fails_quickly] -- Because of the add_torsor.nonempty instance. +instance affine_subspace.to_normed_add_torsor {R : Type*} [ring R] [module R V] + (s : affine_subspace R P) [nonempty s] : normed_add_torsor s.direction s := +{ dist_eq_norm' := λ x y, normed_add_torsor.dist_eq_norm' ↑x ↑y, + ..affine_subspace.to_add_torsor s } + include V section From 019df1c0abf070fdf123415eccd79c7a8cfa3941 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 7 Sep 2022 23:23:51 +0000 Subject: [PATCH 195/828] feat(order/max): equivalence of no_top_order and no_max_order for linear orders (#16341) --- src/order/max.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/order/max.lean b/src/order/max.lean index 465f1fd7e2781..0cd96ba775dba 100644 --- a/src/order/max.lean +++ b/src/order/max.lean @@ -79,6 +79,22 @@ instance no_min_order.to_no_bot_order (α : Type*) [preorder α] [no_min_order instance no_max_order.to_no_top_order (α : Type*) [preorder α] [no_max_order α] : no_top_order α := ⟨λ a, (exists_gt a).imp $ λ _, not_le_of_lt⟩ +lemma no_bot_order.to_no_min_order (α : Type*) [linear_order α] [no_bot_order α] : no_min_order α := +{ exists_lt := by { convert λ a : α, exists_not_ge a, simp_rw not_le, } } + +lemma no_top_order.to_no_max_order (α : Type*) [linear_order α] [no_top_order α] : no_max_order α := +{ exists_gt := by { convert λ a : α, exists_not_le a, simp_rw not_le, } } + +lemma no_bot_order_iff_no_min_order (α : Type*) [linear_order α] : + no_bot_order α ↔ no_min_order α := +⟨λ h, by { haveI := h, exact no_bot_order.to_no_min_order α }, + λ h, by { haveI := h, exact no_min_order.to_no_bot_order α }⟩ + +lemma no_top_order_iff_no_max_order (α : Type*) [linear_order α] : + no_top_order α ↔ no_max_order α := +⟨λ h, by { haveI := h, exact no_top_order.to_no_max_order α }, + λ h, by { haveI := h, exact no_max_order.to_no_top_order α }⟩ + theorem no_min_order.not_acc [has_lt α] [no_min_order α] (a : α) : ¬ acc (<) a := λ h, acc.rec_on h $ λ x _, (exists_lt x).rec_on From ed4882b179a5c6a6b607b7bed42c2cf632a5d733 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 7 Sep 2022 23:23:52 +0000 Subject: [PATCH 196/828] chore(data/finset/basic): generalize `finset.nonempty_mk_coe` to `nonempty_mk` (#16399) --- src/data/finset/basic.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 19e7b8d41f0f1..a12082e020a13 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -530,9 +530,8 @@ by simp only [mem_cons, or_imp_distrib, forall_and_distrib, forall_eq] @[simp] lemma nonempty_cons (h : a ∉ s) : (cons a s h).nonempty := ⟨a, mem_cons.2 $ or.inl rfl⟩ -@[simp] lemma nonempty_mk_coe : ∀ {l : list α} {hl}, (⟨↑l, hl⟩ : finset α).nonempty ↔ l ≠ [] -| [] hl := by simp -| (a :: l) hl := by simp [← multiset.cons_coe] +@[simp] lemma nonempty_mk {m : multiset α} {hm} : (⟨m, hm⟩ : finset α).nonempty ↔ m ≠ 0 := +by induction m using multiset.induction_on; simp @[simp] lemma coe_cons {a s h} : (@cons α a s h : set α) = insert a s := by { ext, simp } From 5ebb7d87f1e381ce15d4775f4c2918efa77a4035 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 8 Sep 2022 02:05:20 +0000 Subject: [PATCH 197/828] feat(analysis/convex/gauge): unit ball of the gauge seminorm (#16413) --- src/analysis/convex/gauge.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index 36df033059094..6235a77f2f557 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -389,6 +389,13 @@ lemma gauge_seminorm_lt_one_of_open (hs : is_open s) {x : E} (hx : x ∈ s) : gauge_seminorm hs₀ hs₁ hs₂ x < 1 := gauge_lt_one_of_mem_of_open hs₁ hs₂.zero_mem hs hx +lemma gauge_seminorm_ball_one (hs : is_open s) : + (gauge_seminorm hs₀ hs₁ hs₂).ball 0 1 = s := +begin + rw seminorm.ball_zero_eq, + exact gauge_lt_one_eq_self_of_open hs₁ hs₂.zero_mem hs +end + end is_R_or_C /-- Any seminorm arises as the gauge of its unit ball. -/ From fd5d3068a6af711ac20ff2486aa5c2315ba52b9b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 8 Sep 2022 06:45:07 +0000 Subject: [PATCH 198/828] chore(measure_theory/function/conditional_expectation): change the definition of `condexp` (#16325) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Change the definition of `condexp` slightly to have `μ[f|m] = 0` when `f` is not integrable, instead of `μ[f|m] =ᵐ[μ] 0`. --- .../conditional_expectation/basic.lean | 78 +++++++++++-------- .../conditional_expectation/indicator.lean | 8 +- .../conditional_expectation/real.lean | 2 +- src/probability/conditional_expectation.lean | 2 +- 4 files changed, 48 insertions(+), 42 deletions(-) diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 797a53a9cfd41..097d685355147 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -1827,9 +1827,12 @@ lemma condexp_L1_eq (hf : integrable f μ) : condexp_L1 hm μ f = condexp_L1_clm hm μ (hf.to_L1 f) := set_to_fun_eq (dominated_fin_meas_additive_condexp_ind F' hm μ) hf -lemma condexp_L1_zero : condexp_L1 hm μ (0 : α → F') = 0 := +@[simp] lemma condexp_L1_zero : condexp_L1 hm μ (0 : α → F') = 0 := set_to_fun_zero _ +@[simp] lemma condexp_L1_measure_zero (hm : m ≤ m0) : condexp_L1 hm (0 : measure α) f = 0 := +set_to_fun_measure_zero _ rfl + lemma ae_strongly_measurable'_condexp_L1 {f : α → F'} : ae_strongly_measurable' m (condexp_L1 hm μ f) μ := begin @@ -1904,17 +1907,20 @@ open_locale classical variables {𝕜} {m m0 : measurable_space α} {μ : measure α} {f g : α → F'} {s : set α} -/-- Conditional expectation of a function. Its value is 0 if the function is not integrable, if -the σ-algebra is not a sub-σ-algebra or if the measure is not σ-finite on that σ-algebra. -/ +/-- Conditional expectation of a function. It is defined as 0 if any one of the following conditions +is true: +- `m` is not a sub-σ-algebra of `m0`, +- `μ` is not σ-finite with respect to `m`, +- `f` is not integrable. -/ @[irreducible] def condexp (m : measurable_space α) {m0 : measurable_space α} (μ : measure α) (f : α → F') : α → F' := if hm : m ≤ m0 - then if hμ : sigma_finite (μ.trim hm) - then if (strongly_measurable[m] f ∧ integrable f μ) + then if h : sigma_finite (μ.trim hm) ∧ integrable f μ + then if strongly_measurable[m] f then f - else (@ae_strongly_measurable'_condexp_L1 _ _ _ _ _ m m0 μ hm hμ _).mk - (@condexp_L1 _ _ _ _ _ _ _ hm μ hμ f) + else (@ae_strongly_measurable'_condexp_L1 _ _ _ _ _ m m0 μ hm h.1 _).mk + (@condexp_L1 _ _ _ _ _ _ _ hm μ h.1 f) else 0 else 0 @@ -1926,19 +1932,26 @@ lemma condexp_of_not_le (hm_not : ¬ m ≤ m0) : μ[f|m] = 0 := by rw [condexp, lemma condexp_of_not_sigma_finite (hm : m ≤ m0) (hμm_not : ¬ sigma_finite (μ.trim hm)) : μ[f|m] = 0 := -by rw [condexp, dif_pos hm, dif_neg hμm_not] +by { rw [condexp, dif_pos hm, dif_neg], push_neg, exact λ h, absurd h hμm_not, } lemma condexp_of_sigma_finite (hm : m ≤ m0) [hμm : sigma_finite (μ.trim hm)] : μ[f|m] = - if (strongly_measurable[m] f ∧ integrable f μ) - then f else ae_strongly_measurable'_condexp_L1.mk (condexp_L1 hm μ f) := -by rw [condexp, dif_pos hm, dif_pos hμm] + if integrable f μ + then if strongly_measurable[m] f + then f else ae_strongly_measurable'_condexp_L1.mk (condexp_L1 hm μ f) + else 0 := +begin + rw [condexp, dif_pos hm], + simp only [hμm, ne.def, true_and], + by_cases hf : integrable f μ, + { rw [dif_pos hf, if_pos hf], }, + { rw [dif_neg hf, if_neg hf], }, +end lemma condexp_of_strongly_measurable (hm : m ≤ m0) [hμm : sigma_finite (μ.trim hm)] {f : α → F'} (hf : strongly_measurable[m] f) (hfi : integrable f μ) : μ[f|m] = f := -by { rw [condexp_of_sigma_finite hm, - if_pos (⟨hf, hfi⟩ : strongly_measurable[m] f ∧ integrable f μ)], apply_instance, } +by { rw [condexp_of_sigma_finite hm, if_pos hfi, if_pos hf], apply_instance, } lemma condexp_const (hm : m ≤ m0) (c : F') [is_finite_measure μ] : μ[(λ x : α, c)|m] = λ _, c := condexp_of_strongly_measurable hm (@strongly_measurable_const _ _ m _ _) (integrable_const c) @@ -1947,15 +1960,16 @@ lemma condexp_ae_eq_condexp_L1 (hm : m ≤ m0) [hμm : sigma_finite (μ.trim hm) (f : α → F') : μ[f|m] =ᵐ[μ] condexp_L1 hm μ f := begin rw condexp_of_sigma_finite hm, - by_cases hfm : strongly_measurable[m] f, - { by_cases hfi : integrable f μ, - { rw if_pos (⟨hfm, hfi⟩ : strongly_measurable[m] f ∧ integrable f μ), + by_cases hfi : integrable f μ, + { rw if_pos hfi, + by_cases hfm : strongly_measurable[m] f, + { rw if_pos hfm, exact (condexp_L1_of_ae_strongly_measurable' (strongly_measurable.ae_strongly_measurable' hfm) hfi).symm, }, - { simp only [hfi, if_false, and_false], + { rw if_neg hfm, exact (ae_strongly_measurable'.ae_eq_mk ae_strongly_measurable'_condexp_L1).symm, }, }, - simp only [hfm, if_false, false_and], - exact (ae_strongly_measurable'.ae_eq_mk ae_strongly_measurable'_condexp_L1).symm, + rw [if_neg hfi, condexp_L1_undef hfi], + exact (coe_fn_zero _ _ _).symm, end lemma condexp_ae_eq_condexp_L1_clm (hm : m ≤ m0) [sigma_finite (μ.trim hm)] (hf : integrable f μ) : @@ -1965,15 +1979,14 @@ begin rw condexp_L1_eq hf, end -lemma condexp_undef (hf : ¬ integrable f μ) : μ[f|m] =ᵐ[μ] 0 := +lemma condexp_undef (hf : ¬ integrable f μ) : μ[f|m] = 0 := begin by_cases hm : m ≤ m0, swap, { rw condexp_of_not_le hm, }, by_cases hμm : sigma_finite (μ.trim hm), swap, { rw condexp_of_not_sigma_finite hm hμm, }, haveI : sigma_finite (μ.trim hm) := hμm, - refine (condexp_ae_eq_condexp_L1 hm f).trans (eventually_eq.trans _ (coe_fn_zero _ 1 _)), - rw condexp_L1_undef hf, + rw [condexp_of_sigma_finite, if_neg hf], end @[simp] lemma condexp_zero : μ[(0 : α → F')|m] = 0 := @@ -1996,13 +2009,10 @@ begin haveI : sigma_finite (μ.trim hm) := hμm, rw condexp_of_sigma_finite hm, swap, { apply_instance, }, - by_cases hfm : strongly_measurable[m] f, - { by_cases hfi : integrable f μ, - { rwa if_pos (⟨hfm, hfi⟩ : strongly_measurable[m] f ∧ integrable f μ), }, - { simp only [hfi, if_false, and_false], - exact ae_strongly_measurable'.strongly_measurable_mk _, }, }, - simp only [hfm, if_false, false_and], - exact ae_strongly_measurable'.strongly_measurable_mk _, + split_ifs with hfi hfm, + { exact hfm, }, + { exact ae_strongly_measurable'.strongly_measurable_mk _, }, + { exact strongly_measurable_zero, }, end lemma condexp_congr_ae (h : f =ᵐ[μ] g) : μ[f | m] =ᵐ[μ] μ[g | m] := @@ -2130,6 +2140,8 @@ begin by_cases hμm₁ : sigma_finite (μ.trim (hm₁₂.trans hm₂)), swap, { simp_rw condexp_of_not_sigma_finite (hm₁₂.trans hm₂) hμm₁, }, haveI : sigma_finite (μ.trim (hm₁₂.trans hm₂)) := hμm₁, + by_cases hf : integrable f μ, + swap, { simp_rw [condexp_undef hf, condexp_zero], }, refine ae_eq_of_forall_set_integral_eq_of_sigma_finite' (hm₁₂.trans hm₂) (λ s hs hμs, integrable_condexp.integrable_on) (λ s hs hμs, integrable_condexp.integrable_on) _ (strongly_measurable.ae_strongly_measurable' strongly_measurable_condexp) @@ -2137,9 +2149,7 @@ begin intros s hs hμs, rw set_integral_condexp (hm₁₂.trans hm₂) integrable_condexp hs, swap, { apply_instance, }, - by_cases hf : integrable f μ, - { rw [set_integral_condexp (hm₁₂.trans hm₂) hf hs, set_integral_condexp hm₂ hf (hm₁₂ s hs)], }, - { simp_rw integral_congr_ae (ae_restrict_of_ae (condexp_undef hf)), }, + rw [set_integral_condexp (hm₁₂.trans hm₂) hf hs, set_integral_condexp hm₂ hf (hm₁₂ s hs)], end lemma condexp_mono {E} [normed_lattice_add_comm_group E] [complete_space E] [normed_space ℝ E] @@ -2162,7 +2172,7 @@ begin by_cases hfint : integrable f μ, { rw (condexp_zero.symm : (0 : α → E) = μ[0 | m]), exact condexp_mono (integrable_zero _ _ _) hfint hf }, - { exact eventually_eq.le (condexp_undef hfint).symm } + { rw condexp_undef hfint, } end lemma condexp_nonpos {E} [normed_lattice_add_comm_group E] [complete_space E] [normed_space ℝ E] @@ -2172,7 +2182,7 @@ begin by_cases hfint : integrable f μ, { rw (condexp_zero.symm : (0 : α → E) = μ[0 | m]), exact condexp_mono hfint (integrable_zero _ _ _) hf }, - { exact eventually_eq.le (condexp_undef hfint) } + { rw condexp_undef hfint, } end /-- **Lebesgue dominated convergence theorem**: sufficient conditions under which almost diff --git a/src/measure_theory/function/conditional_expectation/indicator.lean b/src/measure_theory/function/conditional_expectation/indicator.lean index fdba24c63553a..75145fe2c871c 100644 --- a/src/measure_theory/function/conditional_expectation/indicator.lean +++ b/src/measure_theory/function/conditional_expectation/indicator.lean @@ -44,7 +44,7 @@ begin { rw ← restrict_trim hm _ hs, exact restrict.sigma_finite _ s, }, by_cases hf_int : integrable f μ, - swap, { exact ae_restrict_of_ae (condexp_undef hf_int), }, + swap, { rw condexp_undef hf_int, }, refine ae_eq_of_forall_set_integral_eq_of_sigma_finite' hm _ _ _ _ _, { exact λ t ht hμt, integrable_condexp.integrable_on.integrable_on, }, { exact λ t ht hμt, (integrable_zero _ _ _).integrable_on, }, @@ -159,11 +159,7 @@ begin have hs_m₂ : measurable_set[m₂] s, { rwa [← set.inter_univ s, ← hs set.univ, set.inter_univ], }, by_cases hf_int : integrable f μ, - swap, - { filter_upwards [@condexp_undef _ _ _ _ _ m _ μ _ hf_int, - @condexp_undef _ _ _ _ _ m₂ _ μ _ hf_int] with x hxm hxm₂, - rw pi.zero_apply at hxm hxm₂, - rw [set.indicator_apply_eq_zero.2 (λ _, hxm), set.indicator_apply_eq_zero.2 (λ _, hxm₂)] }, + swap, { simp_rw condexp_undef hf_int, }, refine ((condexp_indicator hf_int hs_m).symm.trans _).trans (condexp_indicator hf_int hs_m₂), refine ae_eq_of_forall_set_integral_eq_of_sigma_finite' hm₂ (λ s hs hμs, integrable_condexp.integrable_on) diff --git a/src/measure_theory/function/conditional_expectation/real.lean b/src/measure_theory/function/conditional_expectation/real.lean index 6360daf577c23..7a0cdca071406 100644 --- a/src/measure_theory/function/conditional_expectation/real.lean +++ b/src/measure_theory/function/conditional_expectation/real.lean @@ -56,7 +56,7 @@ lemma snorm_one_condexp_le_snorm (f : α → ℝ) : snorm (μ[f | m]) 1 μ ≤ snorm f 1 μ := begin by_cases hf : integrable f μ, - swap, { rw [snorm_congr_ae (condexp_undef hf), snorm_zero], exact zero_le _ }, + swap, { rw [condexp_undef hf, snorm_zero], exact zero_le _ }, by_cases hm : m ≤ m0, swap, { rw [condexp_of_not_le hm, snorm_zero], exact zero_le _ }, by_cases hsig : sigma_finite (μ.trim hm), diff --git a/src/probability/conditional_expectation.lean b/src/probability/conditional_expectation.lean index ec903d56ad21d..b0d2ad7857919 100644 --- a/src/probability/conditional_expectation.lean +++ b/src/probability/conditional_expectation.lean @@ -38,7 +38,7 @@ lemma condexp_indep_eq μ[f | m₂] =ᵐ[μ] λ x, μ[f] := begin by_cases hfint : integrable f μ, - swap, { exact (integral_undef hfint).symm ▸ condexp_undef hfint }, + swap, { rw [condexp_undef hfint, integral_undef hfint], refl, }, have hfint₁ := hfint.trim hle₁ hf, refine (ae_eq_condexp_of_forall_set_integral_eq hle₂ hfint (λ s _ hs, integrable_on_const.2 (or.inr hs)) (λ s hms hs, _) From 572e54ab9935efa809ebc38435f6f5500db546e2 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 8 Sep 2022 08:00:09 +0000 Subject: [PATCH 199/828] fix(algebra/big_operators/multiset): rename `prod_le_sum_prod` to `prod_le_prod_map` (#16420) --- src/algebra/big_operators/multiset.lean | 2 +- src/topology/algebra/polynomial.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index db8d3bddff2e1..181cde814b5c6 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -321,7 +321,7 @@ lemma prod_map_le_prod (f : α → α) (h : ∀ x, x ∈ s → f x ≤ x) : (s.m prod_le_prod_of_rel_le $ rel_map_left.2 $ rel_refl_of_refl_on h @[to_additive] -lemma prod_le_sum_prod (f : α → α) (h : ∀ x, x ∈ s → x ≤ f x) : s.prod ≤ (s.map f).prod := +lemma prod_le_prod_map (f : α → α) (h : ∀ x, x ∈ s → x ≤ f x) : s.prod ≤ (s.map f).prod := @prod_map_le_prod αᵒᵈ _ _ f h @[to_additive card_nsmul_le_sum] diff --git a/src/topology/algebra/polynomial.lean b/src/topology/algebra/polynomial.lean index 054f74a6854c4..cc3aaef487808 100644 --- a/src/topology/algebra/polynomial.lean +++ b/src/topology/algebra/polynomial.lean @@ -156,7 +156,7 @@ begin rw multiset.map_map, suffices : ∀ r ∈ multiset.map (norm_hom ∘ prod) (powerset_len (card (map f p).roots - i) (map f p).roots), r ≤ B^(p.nat_degree - i), - { convert sum_le_sum_sum _ this, + { convert sum_le_sum_map _ this, simp only [hi, hcd, multiset.map_const, card_map, card_powerset_len, nat.choose_symm, sum_repeat, nsmul_eq_mul, mul_comm], }, { intros r hr, From 9b6eef6ed5a5e723489891f19391d131b67bd563 Mon Sep 17 00:00:00 2001 From: Chris Birkbeck Date: Thu, 8 Sep 2022 10:32:12 +0000 Subject: [PATCH 200/828] feat(number_theory/modular_forms/congruence_subgroups): Add definition of congruence subgroups. (#15159) This contains the definition of a congruence subgroup of `SL(2,Z)` and defines Gamma1, Gamma0 and full level subgroups. It also contains some basic results about them. --- src/data/matrix/basic.lean | 12 + src/group_theory/subgroup/pointwise.lean | 22 ++ src/linear_algebra/special_linear_group.lean | 23 ++ .../modular_forms/congruence_subgroups.lean | 233 ++++++++++++++++++ 4 files changed, 290 insertions(+) create mode 100644 src/number_theory/modular_forms/congruence_subgroups.lean diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index 08d0d9237e2dc..d44e796938c24 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -572,6 +572,18 @@ lemma sum_apply [add_comm_monoid α] (i : m) (j : n) (∑ c in s, g c) i j = ∑ c in s, g c i j := (congr_fun (s.sum_apply i g) j).trans (s.sum_apply j _) +lemma two_mul_expl {R : Type*} [comm_ring R] (A B : matrix (fin 2) (fin 2) R) : + (A * B) 0 0 = A 0 0 * B 0 0 + A 0 1 * B 1 0 ∧ + (A * B) 0 1 = A 0 0 * B 0 1 + A 0 1 * B 1 1 ∧ + (A * B) 1 0 = A 1 0 * B 0 0 + A 1 1 * B 1 0 ∧ + (A * B) 1 1 = A 1 0 * B 0 1 + A 1 1 * B 1 1 := +begin + split, work_on_goal 2 {split}, work_on_goal 3 {split}, + all_goals {simp only [matrix.mul_eq_mul], + rw [matrix.mul_apply, finset.sum_fin_eq_sum_range, finset.sum_range_succ, finset.sum_range_succ], + simp}, +end + section add_comm_monoid variables [add_comm_monoid α] [has_mul α] diff --git a/src/group_theory/subgroup/pointwise.lean b/src/group_theory/subgroup/pointwise.lean index 91fe83041c42f..ee899b4054625 100644 --- a/src/group_theory/subgroup/pointwise.lean +++ b/src/group_theory/subgroup/pointwise.lean @@ -5,6 +5,7 @@ Authors: Eric Wieser -/ import group_theory.subgroup.basic import group_theory.submonoid.pointwise +import group_theory.group_action.conj_act /-! # Pointwise instances on `subgroup` and `add_subgroup`s @@ -130,6 +131,27 @@ begin exact H.mul_mem hh hh', end +lemma normal.conj_act {G : Type*} [group G] {H : subgroup G} (hH : H.normal ) (g : conj_act G) : + g • H = H := +begin + ext, + split, + { intro h, + have := hH.conj_mem (g⁻¹ • x) _ (conj_act.of_conj_act g), + rw subgroup.mem_pointwise_smul_iff_inv_smul_mem at h, + dsimp at *, + rw conj_act.smul_def at *, + simp only [conj_act.of_conj_act_inv, conj_act.of_conj_act_to_conj_act, inv_inv] at *, + convert this, + simp only [←mul_assoc, mul_right_inv, one_mul, mul_inv_cancel_right], + rw subgroup.mem_pointwise_smul_iff_inv_smul_mem at h, + exact h}, + { intro h, + rw [subgroup.mem_pointwise_smul_iff_inv_smul_mem, conj_act.smul_def], + apply hH.conj_mem, + exact h} +end + end group section group_with_zero diff --git a/src/linear_algebra/special_linear_group.lean b/src/linear_algebra/special_linear_group.lean index f085db85e748c..f7293f4a62192 100644 --- a/src/linear_algebra/special_linear_group.lean +++ b/src/linear_algebra/special_linear_group.lean @@ -210,6 +210,29 @@ subtype.ext $ (@ring_hom.map_matrix n _ _ _ _ _ _ (int.cast_ring_hom R)).map_neg end has_neg +section special_cases + +lemma SL2_inv_expl_det (A : SL(2,R)) : det ![![A.1 1 1, -A.1 0 1], ![-A.1 1 0 , A.1 0 0]] = 1 := +begin + rw [matrix.det_fin_two, mul_comm], + simp only [subtype.val_eq_coe, cons_val_zero, cons_val_one, head_cons, mul_neg, neg_mul, neg_neg], + have := A.2, + rw matrix.det_fin_two at this, + convert this, +end + +lemma SL2_inv_expl (A : SL(2, R)) : A⁻¹ = ⟨![![A.1 1 1, -A.1 0 1], ![-A.1 1 0 , A.1 0 0]], + SL2_inv_expl_det A⟩ := +begin + ext, + have := matrix.adjugate_fin_two A.1, + simp only [subtype.val_eq_coe] at this, + rw [coe_inv, this], + refl, +end + +end special_cases + -- this section should be last to ensure we do not use it in lemmas section coe_fn_instance diff --git a/src/number_theory/modular_forms/congruence_subgroups.lean b/src/number_theory/modular_forms/congruence_subgroups.lean new file mode 100644 index 0000000000000..5dfd42d0e3b01 --- /dev/null +++ b/src/number_theory/modular_forms/congruence_subgroups.lean @@ -0,0 +1,233 @@ +/- +Copyright (c) 2022 Chris Birkbeck. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Birkbeck +-/ +import linear_algebra.special_linear_group +import data.zmod.basic +import group_theory.subgroup.pointwise +import group_theory.group_action.conj_act +/-! +# Congruence subgroups + +This defines congruence subgroups of `SL(2,ℤ)` such as `Γ(N)`, `Γ₀(N)` and `Γ₁(N)` for `N` a +natural number. + +It also contains basic results about congruence subgroups. + +-/ + +local notation `SL(` n `, ` R `)`:= matrix.special_linear_group (fin n) R + +local attribute [-instance] matrix.special_linear_group.has_coe_to_fun + +local prefix `↑ₘ`:1024 := @coe _ (matrix (fin 2) (fin 2) _) _ + +open matrix.special_linear_group matrix + +variable (N : ℕ) + +local notation `SLMOD(`N`)` := @matrix.special_linear_group.map (fin 2) _ _ _ _ _ _ + (int.cast_ring_hom (zmod N)) + +@[simp] +lemma SL_reduction_mod_hom_val (N : ℕ) (γ : SL(2, ℤ)) : ∀ (i j : fin 2), + ((SLMOD(N) γ) : (matrix (fin 2) (fin 2) (zmod N))) i j = + (((↑ₘγ i j) : ℤ) : zmod N) := λ i j, rfl + +/--The full level `N` congruence subgroup of `SL(2,ℤ)` of matrices that reduce to the identity +modulo `N`.-/ +def Gamma (N : ℕ) : subgroup SL(2, ℤ) := (SLMOD(N)).ker + +lemma Gamma_mem' (N : ℕ) (γ : SL(2, ℤ)) : γ ∈ Gamma N ↔ SLMOD(N) γ = 1 := iff.rfl + +@[simp] +lemma Gamma_mem (N : ℕ) (γ : SL(2, ℤ)) : γ ∈ Gamma N ↔ (((↑ₘγ 0 0) : ℤ) : zmod N) = 1 ∧ + (((↑ₘγ 0 1) : ℤ) : zmod N) = 0 ∧ (((↑ₘγ 1 0) : ℤ) : zmod N) = 0 ∧ + (((↑ₘγ 1 1) : ℤ) : zmod N) = 1 := +begin + rw Gamma_mem', + split, + { intro h, + simp [←(SL_reduction_mod_hom_val N γ), h] }, + { intro h, + ext, + rw SL_reduction_mod_hom_val N γ, + fin_cases i; fin_cases j, + all_goals {simp_rw h, refl} } +end + +lemma Gamma_normal (N : ℕ) : subgroup.normal (Gamma N) := (SLMOD(N)).normal_ker + +lemma Gamma_one_top : Gamma 1 = ⊤ := +begin + ext, + simp, +end + +lemma Gamma_zero_bot : Gamma 0 = ⊥ := +begin + ext, + simp only [Gamma_mem, coe_coe, coe_matrix_coe, int.coe_cast_ring_hom, map_apply, int.cast_id, + subgroup.mem_bot], + split, + { intro h, + ext, + fin_cases i; fin_cases j, + any_goals {simp [h]} }, + { intro h, + simp [h] } +end + +/--The congruence subgroup of `SL(2,ℤ)` of matrices whose lower left-hand entry reduces to zero +modulo `N`. -/ +def Gamma0 (N : ℕ) : subgroup SL(2, ℤ) := +{ carrier := { g : SL(2, ℤ) | ((↑ₘg 1 0 : ℤ) : zmod N) = 0 }, + one_mem' := by { simp }, + mul_mem':= by {intros a b ha hb, + simp only [ set.mem_set_of_eq], + have h := ((matrix.two_mul_expl a.1 b.1).2.2.1), + simp only [coe_coe, coe_matrix_coe, coe_mul, int.coe_cast_ring_hom, map_apply, + set.mem_set_of_eq, subtype.val_eq_coe, mul_eq_mul] at *, + rw h, + simp [ha, hb] }, + inv_mem':= by {intros a ha, + simp only [ set.mem_set_of_eq, subtype.val_eq_coe], + rw (SL2_inv_expl a), + simp only [subtype.val_eq_coe, cons_val_zero, cons_val_one, head_cons, coe_coe, coe_matrix_coe, + coe_mk, int.coe_cast_ring_hom, map_apply, int.cast_neg, neg_eq_zero, set.mem_set_of_eq] at *, + exact ha } } + +@[simp] +lemma Gamma0_mem (N : ℕ) (A: SL(2, ℤ)) : A ∈ Gamma0 N ↔ (((↑ₘA) 1 0 : ℤ) : zmod N) = 0 := iff.rfl + +lemma Gamma0_det (N : ℕ) (A : Gamma0 N) : (A.1.1.det : zmod N) = 1 := +by {simp [A.1.property]} + +/--The group homomorphism from `Gamma0` to `zmod N` given by mapping a matrix to its lower +right-hand entry. -/ +def Gamma_0_map (N : ℕ): Gamma0 N →* zmod N := +{ to_fun := λ g, ((↑ₘg 1 1 : ℤ) : zmod N), + map_one' := by { simp, }, + map_mul' := by {intros A B, + have := (two_mul_expl A.1.1 B.1.1).2.2.2, + simp only [coe_coe, subgroup.coe_mul, coe_matrix_coe, coe_mul, int.coe_cast_ring_hom, map_apply, + subtype.val_eq_coe, mul_eq_mul] at *, + rw this, + have ha := A.property, + simp only [int.cast_add, int.cast_mul, add_left_eq_self, subtype.val_eq_coe, Gamma0_mem, coe_coe, + coe_matrix_coe, int.coe_cast_ring_hom, map_apply] at *, + rw ha, + simp,} } + +/--The congruence subgroup `Gamma1` (as a subgroup of `Gamma0`) of matrices whose bottom +row is congruent to `(0,1)` modulo `N`.-/ +def Gamma1' (N : ℕ) : subgroup (Gamma0 N) := (Gamma_0_map N).ker + +@[simp] +lemma Gamma1_mem' (N : ℕ) (γ : Gamma0 N) : γ ∈ Gamma1' N ↔ (Gamma_0_map N) γ = 1 := iff.rfl + +lemma Gamma1_to_Gamma0_mem (N : ℕ) (A : Gamma0 N) : A ∈ Gamma1' N ↔ + ((↑ₘA 0 0 : ℤ) : zmod N) = 1 ∧ ((↑ₘA 1 1 : ℤ) : zmod N) = 1 ∧ ((↑ₘA 1 0 : ℤ) : zmod N) = 0 := +begin + split, + { intro ha, + have hA := A.property, + rw Gamma0_mem at hA, + have adet := Gamma0_det N A, + rw matrix.det_fin_two at adet, + simp only [Gamma_0_map, coe_coe, coe_matrix_coe, int.coe_cast_ring_hom, map_apply, Gamma1_mem', + monoid_hom.coe_mk, subtype.val_eq_coe, int.cast_sub, int.cast_mul] at *, + rw [hA, ha] at adet, + simp only [mul_one, mul_zero, sub_zero] at adet, + simp only [adet, hA, ha, eq_self_iff_true, and_self]}, + { intro ha, + simp only [Gamma1_mem', Gamma_0_map, monoid_hom.coe_mk, coe_coe, coe_matrix_coe, + int.coe_cast_ring_hom, map_apply], + exact ha.2.1,} +end + +/--The congruence subgroup `Gamma1` of `SL(2,ℤ)` consisting of matrices whose bottom +row is congruent to `(0,1)` modulo `N`. -/ +def Gamma1 (N : ℕ) : subgroup SL(2, ℤ) := subgroup.map +(((Gamma0 N).subtype).comp (Gamma1' N).subtype) ⊤ + +@[simp] +lemma Gamma1_mem (N : ℕ) (A : SL(2, ℤ)) : A ∈ Gamma1 N ↔ + ((↑ₘA 0 0 : ℤ) : zmod N) = 1 ∧ ((↑ₘA 1 1 : ℤ) : zmod N) = 1 ∧ ((↑ₘA 1 0 : ℤ) : zmod N) = 0 := +begin + split, + { intro ha, + simp_rw [Gamma1, subgroup.mem_map] at ha, + simp at ha, + obtain ⟨⟨x, hx⟩, hxx⟩ := ha, + rw Gamma1_to_Gamma0_mem at hx, + rw ←hxx, + convert hx }, + { intro ha, + simp_rw [Gamma1, subgroup.mem_map], + have hA : A ∈ (Gamma0 N), by {simp [ha.right.right, Gamma0_mem, subtype.val_eq_coe],}, + have HA : (⟨A , hA⟩ : Gamma0 N) ∈ Gamma1' N, + by {simp only [Gamma1_to_Gamma0_mem, subgroup.coe_mk, coe_coe, coe_matrix_coe, + int.coe_cast_ring_hom, map_apply], + exact ha,}, + refine ⟨(⟨(⟨A , hA⟩ : Gamma0 N), HA ⟩ : (( Gamma1' N ) : subgroup (Gamma0 N))), _⟩, + simp } +end + +lemma Gamma1_in_Gamma0 (N : ℕ) : Gamma1 N ≤ Gamma0 N := +begin + intros x HA, + simp only [Gamma0_mem, Gamma1_mem, coe_coe, coe_matrix_coe, int.coe_cast_ring_hom, + map_apply] at *, + exact HA.2.2, +end + +section congruence_subgroup + +/--A congruence subgroup is a subgroup of `SL(2,ℤ)` which contains some `Gamma N` for some +`(N : ℕ+)`. -/ +def is_congruence_subgroup (Γ : subgroup SL(2, ℤ)) : Prop := ∃ (N : ℕ+), Gamma N ≤ Γ + +lemma is_congruence_subgroup_trans (H K : subgroup SL(2, ℤ)) (h: H ≤ K) + (h2 : is_congruence_subgroup H) : is_congruence_subgroup K := +begin + obtain ⟨N , hN⟩ := h2, + refine ⟨N, le_trans hN h⟩, +end + +lemma Gamma_is_cong_sub (N : ℕ+) : is_congruence_subgroup (Gamma N) := +⟨N, by {simp only [le_refl]}⟩ + +lemma Gamma1_is_congruence (N : ℕ+) : is_congruence_subgroup (Gamma1 N) := +begin + refine ⟨N, _⟩, + intros A hA, + simp only [Gamma1_mem, Gamma_mem] at *, + simp only [hA, eq_self_iff_true, and_self], +end + +lemma Gamma0_is_congruence (N : ℕ+) : is_congruence_subgroup (Gamma0 N) := +is_congruence_subgroup_trans _ _ (Gamma1_in_Gamma0 N) (Gamma1_is_congruence N) + +end congruence_subgroup + +section conjugation + +open_locale pointwise + +lemma Gamma_cong_eq_self (N : ℕ) (g : conj_act SL(2, ℤ)) : g • (Gamma N) = Gamma N := +begin + apply subgroup.normal.conj_act (Gamma_normal N), +end + +lemma conj_cong_is_cong (g : conj_act SL(2, ℤ)) (Γ : subgroup SL(2, ℤ)) + (h : is_congruence_subgroup Γ) : is_congruence_subgroup (g • Γ) := +begin + obtain ⟨N, HN⟩ := h, + refine ⟨N, _⟩, + rw [←Gamma_cong_eq_self N g, subgroup.pointwise_smul_le_pointwise_smul_iff], + exact HN, +end + +end conjugation From 4c0aa6e7d93638e9c9608b8fca27bf698c319765 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Thu, 8 Sep 2022 10:32:13 +0000 Subject: [PATCH 201/828] feat(analysis/schwartz): Definition of the Schwartz space (#15850) This PR adds the definition of the Schwartz space and by abstract results also of tempered distributions. We prove basic algebraic and topological properties of the Schwartz space. Co-authored-by: sgouezel --- docs/overview.yaml | 3 + docs/undergrad.yaml | 2 +- src/analysis/schwartz_space.lean | 383 +++++++++++++++++++++++++++++++ 3 files changed, 387 insertions(+), 1 deletion(-) create mode 100644 src/analysis/schwartz_space.lean diff --git a/docs/overview.yaml b/docs/overview.yaml index 448a2d9240055..5d93e9ff13425 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -343,6 +343,9 @@ Analysis: Phragmen-Lindelöf principle: 'phragmen_lindelof.horizontal_strip' fundamental theorem of algebra: 'complex.is_alg_closed' + Distribution theory: + Schwartz space: 'schwartz_map' + Probability Theory: Definitions in probability theory: probability measure: 'measure_theory.is_probability_measure' diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index c80477a9d3728..bc923730f01af 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -592,7 +592,7 @@ Distribution calculus: convergence of sequences of distributions: '' support of a distribution: '' Spaces $\mathcal{S}(\R^d)$: - Schwartz space of rapidly decreasing functions: '' + Schwartz space of rapidly decreasing functions: 'schwartz_map' stability by derivation: '' stability by multiplication by a slowly growing smooth function: '' gaussian functions: '' diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean new file mode 100644 index 0000000000000..8e573be7fa024 --- /dev/null +++ b/src/analysis/schwartz_space.lean @@ -0,0 +1,383 @@ +/- +Copyright (c) 2022 Moritz Doll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Doll +-/ + +import analysis.calculus.cont_diff +import analysis.complex.basic +import analysis.locally_convex.with_seminorms +import topology.algebra.uniform_filter_basis +import tactic.positivity + +/-! +# Schwartz space + +This file defines the Schwartz space. Usually, the Schwartz space is defined as the set of smooth +functions $f : ℝ^n → ℂ$ such that there exists $C_{αβ} > 0$ with $$|x^α ∂^β f(x)| < C_{αβ}$$ for +all $x ∈ ℝ^n$ and for all multiindices $α, β$. +In mathlib, we use a slightly different approach and define define the Schwartz space as all +smooth functions `f : E → F`, where `E` and `F` are real normed vector spaces such that for all +natural numbers `k` and `n` we have uniform bounds `∥x∥^k * ∥iterated_fderiv ℝ n f x∥ < C`. +This approach completely avoids using partial derivatives as well as polynomials. +We construct the topology on the Schwartz space by a family of seminorms, which are the best +constants in the above estimates, which is by abstract theory from +`seminorm_family.module_filter_basis` and `seminorm_family.to_locally_convex_space` turns the +Schwartz space into a locally convex topological vector space. + +## Main definitions + +* `schwartz_map`: The Schwartz space is the space of smooth functions such that all derivatives +decay faster than any power of `∥x∥`. +* `schwartz_map.seminorm`: The family of seminorms as described above + +## Main statements + +* `schwartz_map.uniform_add_group` and `schwartz_map.locally_convex`: The Schwartz space is a +locally convex topological vector space. + +## Implementation details + +The implementation of the seminorms is taken almost literally from `continuous_linear_map.op_norm`. + +## Notation + +* `𝓢(E, F)`: The Schwartz space `schwartz_map E F` localized in `schwartz_space` + +## Tags + +Schwartz space, tempered distributions +-/ + +noncomputable theory + +variables {𝕜 𝕜' E F : Type*} + +variables [normed_add_comm_group E] [normed_space ℝ E] +variables [normed_add_comm_group F] [normed_space ℝ F] + +variables (E F) + +/-- A function is a Schwartz function if it is smooth and all derivatives decay faster than + any power of `∥x∥`. -/ +structure schwartz_map := + (to_fun : E → F) + (smooth' : cont_diff ℝ ⊤ to_fun) + (decay' : ∀ (k n : ℕ), ∃ (C : ℝ), ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n to_fun x∥ ≤ C) + +localized "notation `𝓢(` E `, ` F `)` := schwartz_map E F" in schwartz_space + +variables {E F} + +namespace schwartz_map + +instance : has_coe 𝓢(E, F) (E → F) := ⟨to_fun⟩ + +instance fun_like : fun_like 𝓢(E, F) E (λ _, F) := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ +instance : has_coe_to_fun 𝓢(E, F) (λ _, E → F) := ⟨λ p, p.to_fun⟩ + +/-- All derivatives of a Schwartz function are rapidly decaying. -/ +lemma decay (f : 𝓢(E, F)) (k n : ℕ) : ∃ (C : ℝ) (hC : 0 < C), + ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ C := +begin + rcases f.decay' k n with ⟨C, hC⟩, + exact ⟨max C 1, by positivity, λ x, (hC x).trans (le_max_left _ _)⟩, +end + +/-- Every Schwartz function is smooth. -/ +lemma smooth (f : 𝓢(E, F)) (n : ℕ∞) : cont_diff ℝ n f := f.smooth'.of_le le_top + +@[ext] lemma ext {f g : 𝓢(E, F)} (h : ∀ x, (f : E → F) x = g x) : f = g := fun_like.ext f g h + +section aux + +lemma bounds_nonempty (k n : ℕ) (f : 𝓢(E, F)) : + ∃ (c : ℝ), c ∈ {c : ℝ | 0 ≤ c ∧ ∀ (x : E), ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ c} := +let ⟨M, hMp, hMb⟩ := f.decay k n in ⟨M, le_of_lt hMp, hMb⟩ + +lemma bounds_bdd_below (k n : ℕ) (f : 𝓢(E, F)) : + bdd_below {c | 0 ≤ c ∧ ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ c} := +⟨0, λ _ ⟨hn, _⟩, hn⟩ + +lemma decay_add_le_aux (k n : ℕ) (f g : 𝓢(E, F)) (x : E) : + ∥x∥^k * ∥iterated_fderiv ℝ n (f+g) x∥ ≤ + ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ + + ∥x∥^k * ∥iterated_fderiv ℝ n g x∥ := +begin + rw ←mul_add, + refine mul_le_mul_of_nonneg_left _ (by positivity), + convert norm_add_le _ _, + exact iterated_fderiv_add_apply (f.smooth _) (g.smooth _), +end + +lemma decay_neg_aux (k n : ℕ) (f : 𝓢(E, F)) (x : E) : + ∥x∥ ^ k * ∥iterated_fderiv ℝ n (-f) x∥ = ∥x∥ ^ k * ∥iterated_fderiv ℝ n f x∥ := +begin + nth_rewrite 3 ←norm_neg, + congr, + exact iterated_fderiv_neg_apply, +end + +variables [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] + +lemma decay_smul_aux (k n : ℕ) (f : 𝓢(E, F)) (c : 𝕜) (x : E) : + ∥x∥ ^ k * ∥iterated_fderiv ℝ n (c • f) x∥ = + ∥c∥ * ∥x∥ ^ k * ∥iterated_fderiv ℝ n f x∥ := +by rw [mul_comm (∥c∥), mul_assoc, iterated_fderiv_const_smul_apply (f.smooth _), norm_smul] + +end aux + +section seminorm_aux + +/-- Helper definition for the seminorms of the Schwartz space. -/ +@[protected] +def seminorm_aux (k n : ℕ) (f : 𝓢(E, F)) : ℝ := +Inf {c | 0 ≤ c ∧ ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ c} + +lemma seminorm_aux_nonneg (k n : ℕ) (f : 𝓢(E, F)) : 0 ≤ f.seminorm_aux k n := +le_cInf (bounds_nonempty k n f) (λ _ ⟨hx, _⟩, hx) + +lemma le_seminorm_aux (k n : ℕ) (f : 𝓢(E, F)) (x : E) : + ∥x∥ ^ k * ∥iterated_fderiv ℝ n ⇑f x∥ ≤ f.seminorm_aux k n := +le_cInf (bounds_nonempty k n f) (λ y ⟨_, h⟩, h x) + +/-- If one controls the norm of every `A x`, then one controls the norm of `A`. -/ +lemma seminorm_aux_le_bound (k n : ℕ) (f : 𝓢(E, F)) {M : ℝ} (hMp: 0 ≤ M) + (hM : ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ M) : + f.seminorm_aux k n ≤ M := +cInf_le (bounds_bdd_below k n f) ⟨hMp, hM⟩ + +end seminorm_aux + +/-! ### Algebraic properties -/ + +section smul + +variables [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] + [normed_field 𝕜'] [normed_space 𝕜' F] [smul_comm_class ℝ 𝕜' F] + +instance : has_smul 𝕜 𝓢(E, F) := +⟨λ c f, { to_fun := c • f, + smooth' := (f.smooth _).const_smul c, + decay' := λ k n, begin + refine ⟨f.seminorm_aux k n * (∥c∥+1), λ x, _⟩, + have hc : 0 ≤ ∥c∥ := by positivity, + refine le_trans _ ((mul_le_mul_of_nonneg_right (f.le_seminorm_aux k n x) hc).trans _), + { apply eq.le, + rw [mul_comm _ (∥c∥), ← mul_assoc], + exact decay_smul_aux k n f c x }, + { apply mul_le_mul_of_nonneg_left _ (f.seminorm_aux_nonneg k n), + linarith } + end}⟩ + +@[simp] lemma smul_apply {f : 𝓢(E, F)} {c : 𝕜} {x : E} : (c • f) x = c • (f x) := rfl + +instance +[has_smul 𝕜 𝕜'] [is_scalar_tower 𝕜 𝕜' F] : is_scalar_tower 𝕜 𝕜' 𝓢(E, F) := +⟨λ a b f, ext $ λ x, smul_assoc a b (f x)⟩ + +instance [smul_comm_class 𝕜 𝕜' F] : smul_comm_class 𝕜 𝕜' 𝓢(E, F) := +⟨λ a b f, ext $ λ x, smul_comm a b (f x)⟩ + +lemma seminorm_aux_smul_le (k n : ℕ) (c : 𝕜) (f : 𝓢(E, F)) : + (c • f).seminorm_aux k n ≤ ∥c∥ * f.seminorm_aux k n := +begin + refine (c • f).seminorm_aux_le_bound k n (mul_nonneg (norm_nonneg _) (seminorm_aux_nonneg _ _ _)) + (λ x, (decay_smul_aux k n f c x).le.trans _), + rw mul_assoc, + exact mul_le_mul_of_nonneg_left (f.le_seminorm_aux k n x) (norm_nonneg _), +end + +instance has_nsmul : has_smul ℕ 𝓢(E, F) := +⟨λ c f, { to_fun := c • f, + smooth' := (f.smooth _).const_smul c, + decay' := begin + have : c • (f : E → F) = (c : ℝ) • f, + { ext x, simp only [pi.smul_apply, ← nsmul_eq_smul_cast] }, + simp only [this], + exact ((c : ℝ) • f).decay', + end}⟩ + +instance has_zsmul : has_smul ℤ 𝓢(E, F) := +⟨λ c f, { to_fun := c • f, + smooth' := (f.smooth _).const_smul c, + decay' := begin + have : c • (f : E → F) = (c : ℝ) • f, + { ext x, simp only [pi.smul_apply, ← zsmul_eq_smul_cast] }, + simp only [this], + exact ((c : ℝ) • f).decay', + end}⟩ + +end smul + +section zero + +instance : has_zero 𝓢(E, F) := +⟨{ to_fun := λ _, 0, + smooth' := cont_diff_const, + decay' := λ _ _, ⟨1, λ _, by simp⟩ }⟩ + +instance : inhabited 𝓢(E, F) := ⟨0⟩ + +lemma coe_zero : ↑(0 : 𝓢(E, F)) = (0 : E → F) := rfl + +@[simp] lemma coe_fn_zero : coe_fn (0 : 𝓢(E, F)) = (0 : E → F) := rfl + +@[simp] lemma zero_apply {x : E} : (0 : 𝓢(E, F)) x = 0 := rfl + +lemma seminorm_aux_zero (k n : ℕ) : + (0 : 𝓢(E, F)).seminorm_aux k n = 0 := +le_antisymm (seminorm_aux_le_bound k n _ rfl.le (λ _, by simp [pi.zero_def])) + (seminorm_aux_nonneg _ _ _) + +end zero + +section neg + +instance : has_neg 𝓢(E, F) := +⟨λ f, ⟨-f, (f.smooth _).neg, λ k n, + ⟨f.seminorm_aux k n, λ x, (decay_neg_aux k n f x).le.trans (f.le_seminorm_aux k n x)⟩⟩⟩ + +end neg + +section add + +instance : has_add 𝓢(E, F) := +⟨λ f g, ⟨f + g, (f.smooth _).add (g.smooth _), λ k n, + ⟨f.seminorm_aux k n + g.seminorm_aux k n, λ x, (decay_add_le_aux k n f g x).trans + (add_le_add (f.le_seminorm_aux k n x) (g.le_seminorm_aux k n x))⟩⟩⟩ + +@[simp] lemma add_apply {f g : 𝓢(E, F)} {x : E} : (f + g) x = f x + g x := rfl + +lemma seminorm_aux_add_le (k n : ℕ) (f g : 𝓢(E, F)) : + (f + g).seminorm_aux k n ≤ f.seminorm_aux k n + g.seminorm_aux k n := +(f + g).seminorm_aux_le_bound k n + (add_nonneg (seminorm_aux_nonneg _ _ _) (seminorm_aux_nonneg _ _ _)) $ + λ x, (decay_add_le_aux k n f g x).trans $ + add_le_add (f.le_seminorm_aux k n x) (g.le_seminorm_aux k n x) + +end add + +section sub + +instance : has_sub 𝓢(E, F) := +⟨λ f g, ⟨f - g, (f.smooth _).sub (g.smooth _), + begin + intros k n, + refine ⟨f.seminorm_aux k n + g.seminorm_aux k n, λ x, _⟩, + refine le_trans _ (add_le_add (f.le_seminorm_aux k n x) (g.le_seminorm_aux k n x)), + rw sub_eq_add_neg, + rw ←decay_neg_aux k n g x, + convert decay_add_le_aux k n f (-g) x, + -- exact fails with deterministic timeout + end⟩ ⟩ + +@[simp] lemma sub_apply {f g : 𝓢(E, F)} {x : E} : (f - g) x = f x - g x := rfl + +end sub + +section add_comm_group + +instance : add_comm_group 𝓢(E, F) := +fun_like.coe_injective.add_comm_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) + (λ _ _, rfl) + +variables (E F) + +/-- Coercion as an additive homomorphism. -/ +def coe_hom : 𝓢(E, F) →+ (E → F) := +{ to_fun := λ f, f, map_zero' := coe_zero, map_add' := λ _ _, rfl } + +variables {E F} + +lemma coe_coe_hom : (coe_hom E F : 𝓢(E, F) → (E → F)) = coe_fn := rfl + +lemma coe_hom_injective : function.injective (coe_hom E F) := +by { rw coe_coe_hom, exact fun_like.coe_injective } + +end add_comm_group + +section module + +variables [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] + +instance : module 𝕜 𝓢(E, F) := +coe_hom_injective.module 𝕜 (coe_hom E F) (λ _ _, rfl) + +end module + +section seminorms + +/-! ### Seminorms on Schwartz space-/ + +variables [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] +variable (𝕜) + +/-- The seminorms of the Schwartz space given by the best constants in the definition of +`𝓢(E, F)`. -/ +@[protected] +def seminorm (k n : ℕ) : seminorm 𝕜 𝓢(E, F) := seminorm.of_smul_le (seminorm_aux k n) + (seminorm_aux_zero k n) (seminorm_aux_add_le k n) (seminorm_aux_smul_le k n) + +/-- If one controls the seminorm for every `x`, then one controls the seminorm. -/ +lemma seminorm_le_bound (k n : ℕ) (f : 𝓢(E, F)) {M : ℝ} (hMp: 0 ≤ M) + (hM : ∀ x, ∥x∥^k * ∥iterated_fderiv ℝ n f x∥ ≤ M) : seminorm 𝕜 k n f ≤ M := +f.seminorm_aux_le_bound k n hMp hM + +/-- The seminorm controls the Schwartz estimate for any fixed `x`. -/ +lemma le_seminorm (k n : ℕ) (f : 𝓢(E, F)) (x : E) : + ∥x∥ ^ k * ∥iterated_fderiv ℝ n f x∥ ≤ seminorm 𝕜 k n f := +f.le_seminorm_aux k n x + +end seminorms + +section topology + +/-! ### The topology on the Schwartz space-/ + +variables [normed_field 𝕜] [normed_space 𝕜 F] [smul_comm_class ℝ 𝕜 F] +variables (𝕜 E F) + +/-- The family of Schwartz seminorms. -/ +def _root_.schwartz_seminorm_family : seminorm_family 𝕜 𝓢(E, F) (ℕ × ℕ) := +λ n, seminorm 𝕜 n.1 n.2 + +instance : topological_space 𝓢(E, F) := +(schwartz_seminorm_family ℝ E F).module_filter_basis.topology' + +lemma _root_.schwartz_with_seminorms : with_seminorms (schwartz_seminorm_family 𝕜 E F) := +begin + have A : with_seminorms (schwartz_seminorm_family ℝ E F) := ⟨rfl⟩, + rw seminorm_family.with_seminorms_iff_nhds_eq_infi at ⊢ A, + rw A, + refl +end + +variables {𝕜 E F} + +instance : has_continuous_smul 𝕜 𝓢(E, F) := +begin + rw seminorm_family.with_seminorms_eq (schwartz_with_seminorms 𝕜 E F), + exact (schwartz_seminorm_family 𝕜 E F).module_filter_basis.has_continuous_smul, +end + +instance : topological_add_group 𝓢(E, F) := +(schwartz_seminorm_family ℝ E F).module_filter_basis.to_add_group_filter_basis + .is_topological_add_group + +instance : uniform_space 𝓢(E, F) := +(schwartz_seminorm_family ℝ E F).module_filter_basis.to_add_group_filter_basis.uniform_space + +instance : uniform_add_group 𝓢(E, F) := +(schwartz_seminorm_family ℝ E F).module_filter_basis.to_add_group_filter_basis.uniform_add_group + +instance : locally_convex_space ℝ 𝓢(E, F) := +seminorm_family.to_locally_convex_space (schwartz_with_seminorms ℝ E F) + +end topology + +end schwartz_map From a0735864ba72769da4b378673d3dbe2453924fde Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Thu, 8 Sep 2022 10:32:14 +0000 Subject: [PATCH 202/828] fix(*): fix padding around notations (#16333) I noticed a bunch of issues around padding in mathport, so I went through and fixed all the padding issues I could find in mathlib. Co-authored-by: Yury G. Kudryashov --- archive/examples/prop_encodable.lean | 2 +- archive/sensitivity.lean | 4 ++-- src/algebra/big_operators/intervals.lean | 2 +- .../presheafed_space/gluing.lean | 8 ++++---- .../projective_spectrum/scheme.lean | 10 +++++----- src/algebraic_topology/simplicial_object.lean | 4 ++-- src/analysis/inner_product_space/l2_space.lean | 2 +- src/data/bracket.lean | 2 +- src/data/bundle.lean | 2 +- src/data/finset/fold.lean | 2 +- src/data/list/basic.lean | 4 ++-- src/data/list/perm.lean | 4 ++-- src/data/multiset/fold.lean | 2 +- src/data/nat/factorization/basic.lean | 4 ++-- src/data/rat/nnrat.lean | 2 +- src/data/rbtree/insert.lean | 2 +- src/data/real/nnreal.lean | 2 +- src/geometry/manifold/derivation_bundle.lean | 2 +- src/geometry/manifold/diffeomorph.lean | 4 ++-- src/group_theory/eckmann_hilton.lean | 2 +- src/linear_algebra/matrix/determinant.lean | 2 +- src/linear_algebra/span.lean | 2 +- src/linear_algebra/special_linear_group.lean | 2 +- src/linear_algebra/tensor_power.lean | 2 +- src/logic/function/basic.lean | 2 +- src/logic/relator.lean | 2 +- .../integral/divergence_theorem.lean | 8 ++++---- src/measure_theory/integral/torus_integral.lean | 16 ++++++++-------- src/order/basic.lean | 4 ++-- src/ring_theory/derivation.lean | 2 +- src/ring_theory/witt_vector/isocrystal.lean | 14 +++++++------- src/ring_theory/witt_vector/witt_polynomial.lean | 2 +- src/set_theory/cardinal/basic.lean | 2 +- src/set_theory/zfc/basic.lean | 2 +- src/tactic/omega/find_ees.lean | 10 +++++----- src/tactic/omega/int/preterm.lean | 6 +++--- src/tactic/omega/misc.lean | 2 +- src/tactic/omega/nat/preterm.lean | 2 +- src/tactic/reserved_notation.lean | 2 +- src/tactic/rewrite.lean | 2 +- .../uniform_space/abstract_completion.lean | 2 +- src/topology/uniform_space/basic.lean | 2 +- .../uniform_convergence_topology.lean | 6 +++--- test/noncomm_ring.lean | 2 +- test/simps.lean | 2 +- 45 files changed, 83 insertions(+), 83 deletions(-) diff --git a/archive/examples/prop_encodable.lean b/archive/examples/prop_encodable.lean index 64da2837b4f0b..4890ca2837d89 100644 --- a/archive/examples/prop_encodable.lean +++ b/archive/examples/prop_encodable.lean @@ -61,7 +61,7 @@ namespace prop_form private def constructors (α : Type*) := α ⊕ unit ⊕ unit ⊕ unit -local notation `cvar` a := sum.inl a +local notation `cvar ` a := sum.inl a local notation `cnot` := sum.inr (sum.inl unit.star) local notation `cand` := sum.inr (sum.inr (sum.inr unit.star)) local notation `cor` := sum.inr (sum.inr (sum.inl unit.star)) diff --git a/archive/sensitivity.lean b/archive/sensitivity.lean index 9c0342613cc0e..f1960b94e62ee 100644 --- a/archive/sensitivity.lean +++ b/archive/sensitivity.lean @@ -332,7 +332,7 @@ In this section, in order to enforce that `n` is positive, we write it as `m + 1` for some natural number `m`. -/ /-! `dim X` will denote the dimension of a subspace `X` as a cardinal. -/ -notation `dim` X:70 := module.rank ℝ ↥X +notation `dim ` X:70 := module.rank ℝ ↥X /-! `fdim X` will denote the (finite) dimension of a subspace `X` as a natural number. -/ notation `fdim` := finrank ℝ @@ -341,7 +341,7 @@ notation `Span` := submodule.span ℝ /-! `Card X` will denote the cardinal of a subset of a finite type, as a natural number. -/ -notation `Card` X:70 := X.to_finset.card +notation `Card ` X:70 := X.to_finset.card /-! In the following, `⊓` and `⊔` will denote intersection and sums of ℝ-subspaces, equipped with their subspace structures. The notations come from the general diff --git a/src/algebra/big_operators/intervals.lean b/src/algebra/big_operators/intervals.lean index 8a500766ee19a..baf2ae60a8faf 100644 --- a/src/algebra/big_operators/intervals.lean +++ b/src/algebra/big_operators/intervals.lean @@ -234,7 +234,7 @@ section module variables {R M : Type*} [ring R] [add_comm_group M] [module R M] (f : ℕ → R) (g : ℕ → M) {m n : ℕ} open finset -- The partial sum of `g`, starting from zero -local notation `G` n:80 := ∑ i in range n, g i +local notation `G ` n:80 := ∑ i in range n, g i /-- **Summation by parts**, also known as **Abel's lemma** or an **Abel transformation** -/ theorem sum_Ico_by_parts (hmn : m < n) : diff --git a/src/algebraic_geometry/presheafed_space/gluing.lean b/src/algebraic_geometry/presheafed_space/gluing.lean index 94ef63d65c728..144ed3c721d9e 100644 --- a/src/algebraic_geometry/presheafed_space/gluing.lean +++ b/src/algebraic_geometry/presheafed_space/gluing.lean @@ -98,11 +98,11 @@ namespace glue_data variables {C} (D : glue_data C) local notation `𝖣` := D.to_glue_data -local notation `π₁` i `,` j `,` k := @pullback.fst _ _ _ _ _ (D.f i j) (D.f i k) _ -local notation `π₂` i `,` j `,` k := @pullback.snd _ _ _ _ _ (D.f i j) (D.f i k) _ -local notation `π₁⁻¹` i `,` j `,` k := +local notation `π₁ `i`, `j`, `k := @pullback.fst _ _ _ _ _ (D.f i j) (D.f i k) _ +local notation `π₂ `i`, `j`, `k := @pullback.snd _ _ _ _ _ (D.f i j) (D.f i k) _ +local notation `π₁⁻¹ `i`, `j`, `k := (PresheafedSpace.is_open_immersion.pullback_fst_of_right (D.f i j) (D.f i k)).inv_app -local notation `π₂⁻¹` i `,` j `,` k := +local notation `π₂⁻¹ `i`, `j`, `k := (PresheafedSpace.is_open_immersion.pullback_snd_of_left (D.f i j) (D.f i k)).inv_app /-- The glue data of topological spaces associated to a family of glue data of PresheafedSpaces. -/ diff --git a/src/algebraic_geometry/projective_spectrum/scheme.lean b/src/algebraic_geometry/projective_spectrum/scheme.lean index d99ce60d766e8..06965cb6ab393 100644 --- a/src/algebraic_geometry/projective_spectrum/scheme.lean +++ b/src/algebraic_geometry/projective_spectrum/scheme.lean @@ -79,13 +79,13 @@ local notation `Proj| ` U := Proj .restrict (opens.open_embedding (U : opens Pro local notation `Proj.T| ` U := (Proj .restrict (opens.open_embedding (U : opens Proj.T))).to_SheafedSpace.to_PresheafedSpace.1 -- the underlying topological space of `Proj` restricted to some open set -local notation `pbo` x := projective_spectrum.basic_open 𝒜 x +local notation `pbo ` x := projective_spectrum.basic_open 𝒜 x -- basic open sets in `Proj` -local notation `sbo` f := prime_spectrum.basic_open f +local notation `sbo ` f := prime_spectrum.basic_open f -- basic open sets in `Spec` -local notation `Spec` ring := Spec.LocallyRingedSpace_obj (CommRing.of ring) +local notation `Spec ` ring := Spec.LocallyRingedSpace_obj (CommRing.of ring) -- `Spec` as a locally ringed space -local notation `Spec.T` ring := +local notation `Spec.T ` ring := (Spec.LocallyRingedSpace_obj (CommRing.of ring)).to_SheafedSpace.to_PresheafedSpace.1 -- the underlying topological space of `Spec` @@ -121,7 +121,7 @@ def degree_zero_part {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) : subring (away f) end -local notation `A⁰_` f_deg := degree_zero_part f_deg +local notation `A⁰_ ` f_deg := degree_zero_part f_deg section diff --git a/src/algebraic_topology/simplicial_object.lean b/src/algebraic_topology/simplicial_object.lean index 3bbdcf4a7e6ad..b6fc69fe884cc 100644 --- a/src/algebraic_topology/simplicial_object.lean +++ b/src/algebraic_topology/simplicial_object.lean @@ -36,7 +36,7 @@ def simplicial_object := simplex_categoryᵒᵖ ⥤ C namespace simplicial_object -localized "notation (name := simplicial_object.at) X `_[`:1000 n `]` := +localized "notation (name := simplicial_object.at) X ` _[`:1000 n `]` := (X : category_theory.simplicial_object hole!).obj (opposite.op (simplex_category.mk n))" in simplicial @@ -265,7 +265,7 @@ def cosimplicial_object := simplex_category ⥤ C namespace cosimplicial_object -localized "notation (name := cosimplicial_object.at) X `_[`:1000 n `]` := +localized "notation (name := cosimplicial_object.at) X ` _[`:1000 n `]` := (X : category_theory.cosimplicial_object hole!).obj (simplex_category.mk n)" in simplicial instance {J : Type v} [small_category J] [has_limits_of_shape J C] : diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 5257cbb58485f..4d7510f4504f2 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -89,7 +89,7 @@ variables {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [inner_product_space 𝕜 variables {G : ι → Type*} [Π i, inner_product_space 𝕜 (G i)] local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y -notation `ℓ²(` ι `,` 𝕜 `)` := lp (λ i : ι, 𝕜) 2 +notation `ℓ²(`ι`, `𝕜`)` := lp (λ i : ι, 𝕜) 2 /-! ### Inner product space structure on `lp G 2` -/ diff --git a/src/data/bracket.lean b/src/data/bracket.lean index e5cb5e55f2ff4..4372992dcfd76 100644 --- a/src/data/bracket.lean +++ b/src/data/bracket.lean @@ -34,4 +34,4 @@ these are the Unicode "square with quill" brackets rather than the usual square `K` of a group. -/ class has_bracket (L M : Type*) := (bracket : L → M → M) -notation `⁅`x`,` y`⁆` := has_bracket.bracket x y +notation `⁅`x`, `y`⁆` := has_bracket.bracket x y diff --git a/src/data/bundle.lean b/src/data/bundle.lean index 946d03f62ac00..30eb9875bdb20 100644 --- a/src/data/bundle.lean +++ b/src/data/bundle.lean @@ -73,7 +73,7 @@ instance {x : B} : has_coe_t (E x) (total_space E) := ⟨total_space_mk x⟩ lemma to_total_space_coe {x : B} (v : E x) : (v : total_space E) = total_space_mk x v := rfl -- notation for the direct sum of two bundles over the same base -notation E₁ `×ᵇ`:100 E₂ := λ x, E₁ x × E₂ x +notation E₁ ` ×ᵇ `:100 E₂ := λ x, E₁ x × E₂ x /-- `bundle.trivial B F` is the trivial bundle over `B` of fiber `F`. -/ def trivial (B : Type*) (F : Type*) : B → Type* := function.const B F diff --git a/src/data/finset/fold.lean b/src/data/finset/fold.lean index a6db4aa26f38e..31120d6e4bc95 100644 --- a/src/data/finset/fold.lean +++ b/src/data/finset/fold.lean @@ -18,7 +18,7 @@ variables {α β γ : Type*} /-! ### fold -/ section fold variables (op : β → β → β) [hc : is_commutative β op] [ha : is_associative β op] -local notation (name := op) a * b := op a b +local notation (name := op) a ` * ` b := op a b include hc ha /-- `fold op b f s` folds the commutative associative operation `op` over the diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index eadbb338f32c9..1dca88a4290fc 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -2459,8 +2459,8 @@ end foldl_eq_foldlr' section variables {op : α → α → α} [ha : is_associative α op] [hc : is_commutative α op] -local notation (name := op) a * b := op a b -local notation (name := foldl) l <*> a := foldl op a l +local notation (name := op) a ` * ` b := op a b +local notation (name := foldl) l ` <*> ` a := foldl op a l include ha diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index cacfad8a95820..73c8662d851c8 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -457,8 +457,8 @@ end section variables {op : α → α → α} [is_associative α op] [is_commutative α op] -local notation (name := op) a * b := op a b -local notation (name := foldl) l <*> a := foldl op a l +local notation (name := op) a ` * ` b := op a b +local notation (name := foldl) l ` <*> ` a := foldl op a l lemma perm.fold_op_eq {l₁ l₂ : list α} {a : α} (h : l₁ ~ l₂) : l₁ <*> a = l₂ <*> a := h.foldl_eq (right_comm _ is_commutative.comm is_associative.assoc) _ diff --git a/src/data/multiset/fold.lean b/src/data/multiset/fold.lean index ff8c7fc44d233..94e79d0543047 100644 --- a/src/data/multiset/fold.lean +++ b/src/data/multiset/fold.lean @@ -16,7 +16,7 @@ variables {α β : Type*} /-! ### fold -/ section fold variables (op : α → α → α) [hc : is_commutative α op] [ha : is_associative α op] -local notation (name := op) a * b := op a b +local notation (name := op) a ` * ` b := op a b include hc ha /-- `fold op b s` folds a commutative associative operation `op` over diff --git a/src/data/nat/factorization/basic.lean b/src/data/nat/factorization/basic.lean index 5510beae7791e..ca5f59e2c4e80 100644 --- a/src/data/nat/factorization/basic.lean +++ b/src/data/nat/factorization/basic.lean @@ -291,8 +291,8 @@ the $p$-adic order/valuation of a number, and `proj` and `compl` are for the pro complementary projection. The term `n.factorization p` is the $p$-adic order itself. For example, `ord_proj[2] n` is the even part of `n` and `ord_compl[2] n` is the odd part. -/ -notation `ord_proj[` p `]` n:max := p ^ (nat.factorization n p) -notation `ord_compl[` p `]` n:max := n / ord_proj[p] n +notation `ord_proj[` p `] ` n:max := p ^ (nat.factorization n p) +notation `ord_compl[` p `] ` n:max := n / ord_proj[p] n @[simp] lemma ord_proj_of_not_prime (n p : ℕ) (hp : ¬ p.prime) : ord_proj[p] n = 1 := by simp [factorization_eq_zero_of_non_prime n hp] diff --git a/src/data/rat/nnrat.lean b/src/data/rat/nnrat.lean index 8b14701c39ca4..29e05d2d872e2 100644 --- a/src/data/rat/nnrat.lean +++ b/src/data/rat/nnrat.lean @@ -31,7 +31,7 @@ open_locale big_operators densely_ordered, archimedean, inhabited]] def nnrat := {q : ℚ // 0 ≤ q} -localized "notation ` ℚ≥0 ` := nnrat" in nnrat +localized "notation (name := nnrat) `ℚ≥0` := nnrat" in nnrat namespace nnrat variables {α : Type*} {p q : ℚ≥0} diff --git a/src/data/rbtree/insert.lean b/src/data/rbtree/insert.lean index 22bbd2934e4a9..955270781fd37 100644 --- a/src/data/rbtree/insert.lean +++ b/src/data/rbtree/insert.lean @@ -188,7 +188,7 @@ parameters {α : Type u} (lt : α → α → Prop) local attribute [simp] mem balance1_node balance2_node -local infix (name := mem) `∈` := mem lt +local infix (name := mem) ` ∈ ` := mem lt lemma mem_balance1_node_of_mem_left {x s} (v) (t : rbnode α) : x ∈ s → x ∈ balance1_node s v t := begin diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 12a5b2942c8f6..7af447cae0bee 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -59,7 +59,7 @@ open_locale classical big_operators linear_ordered_semiring, ordered_comm_semiring, canonically_ordered_comm_semiring, has_sub, has_ordered_sub, has_div, inhabited]] def nnreal := {r : ℝ // 0 ≤ r} -localized "notation (name := nnreal) ` ℝ≥0 ` := nnreal" in nnreal +localized "notation (name := nnreal) `ℝ≥0` := nnreal" in nnreal namespace nnreal diff --git a/src/geometry/manifold/derivation_bundle.lean b/src/geometry/manifold/derivation_bundle.lean index 0bc52955e9860..f463781aa276b 100644 --- a/src/geometry/manifold/derivation_bundle.lean +++ b/src/geometry/manifold/derivation_bundle.lean @@ -35,7 +35,7 @@ instance smooth_functions_tower : is_scalar_tower 𝕜 C^∞⟮I, M; 𝕜⟯ C^ which is defined as `f • r = f(x) * r`. -/ @[nolint unused_arguments] def pointed_smooth_map (x : M) := C^n⟮I, M; 𝕜⟯ -localized "notation (name := pointed_smooth_map) `C^` n `⟮` I `,` M `;` 𝕜 `⟯⟨` x `⟩` := +localized "notation (name := pointed_smooth_map) `C^` n `⟮` I `, ` M `; ` 𝕜 `⟯⟨` x `⟩` := pointed_smooth_map 𝕜 I M n x" in derivation variables {𝕜 M} diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index a9968eaaf5ffc..cead8e7582a07 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -77,9 +77,9 @@ structure diffeomorph extends M ≃ M' := end defs -localized "notation (name := diffeomorph) M ` ≃ₘ^` n:1000 `⟮`:50 I `,` J `⟯ ` N := +localized "notation (name := diffeomorph) M ` ≃ₘ^` n:1000 `⟮`:50 I `, ` J `⟯ ` N := diffeomorph I J M N n" in manifold -localized "notation (name := diffeomorph.top) M ` ≃ₘ⟮` I `,` J `⟯ ` N := +localized "notation (name := diffeomorph.top) M ` ≃ₘ⟮` I `, ` J `⟯ ` N := diffeomorph I J M N ⊤" in manifold localized "notation (name := diffeomorph.self) E ` ≃ₘ^` n:1000 `[`:50 𝕜 `] ` E' := diffeomorph (model_with_corners_self 𝕜 E) (model_with_corners_self 𝕜 E') E E' n" in manifold diff --git a/src/group_theory/eckmann_hilton.lean b/src/group_theory/eckmann_hilton.lean index 9ac8e44c83bdc..3cbc0aadc8db5 100644 --- a/src/group_theory/eckmann_hilton.lean +++ b/src/group_theory/eckmann_hilton.lean @@ -27,7 +27,7 @@ universe u namespace eckmann_hilton variables {X : Type u} -local notation a `<`m`>` b := m a b +local notation a ` <`m`> ` b := m a b /-- `is_unital m e` expresses that `e : X` is a left and right unit for the binary operation `m : X → X → X`. -/ diff --git a/src/linear_algebra/matrix/determinant.lean b/src/linear_algebra/matrix/determinant.lean index fd0e426dc7da5..8c398e048380f 100644 --- a/src/linear_algebra/matrix/determinant.lean +++ b/src/linear_algebra/matrix/determinant.lean @@ -47,7 +47,7 @@ open_locale matrix big_operators variables {m n : Type*} [decidable_eq n] [fintype n] [decidable_eq m] [fintype m] variables {R : Type v} [comm_ring R] -local notation `ε` σ:max := ((sign σ : ℤ ) : R) +local notation `ε ` σ:max := ((sign σ : ℤ) : R) /-- `det` is an `alternating_map` in the rows of the matrix. -/ diff --git a/src/linear_algebra/span.lean b/src/linear_algebra/span.lean index 068fdef67622b..0ead5fe1ea603 100644 --- a/src/linear_algebra/span.lean +++ b/src/linear_algebra/span.lean @@ -202,7 +202,7 @@ by rw [submodule.span_union, p.span_eq] /- Note that the character `∙` U+2219 used below is different from the scalar multiplication character `•` U+2022 and the matrix multiplication character `⬝` U+2B1D. -/ -notation R`∙`:1000 x := span R (@singleton _ _ set.has_singleton x) +notation R` ∙ `:1000 x := span R (@singleton _ _ set.has_singleton x) lemma span_eq_supr_of_singleton_spans (s : set M) : span R s = ⨆ x ∈ s, R ∙ x := by simp only [←span_Union, set.bUnion_of_singleton s] diff --git a/src/linear_algebra/special_linear_group.lean b/src/linear_algebra/special_linear_group.lean index f7293f4a62192..d371eab1cda05 100644 --- a/src/linear_algebra/special_linear_group.lean +++ b/src/linear_algebra/special_linear_group.lean @@ -63,7 +63,7 @@ def special_linear_group := { A : matrix n n R // A.det = 1 } end localized "notation (name := special_linear_group.fin) - `SL(` n `,` R `)`:= matrix.special_linear_group (fin n) R" in matrix_groups + `SL(`n`, `R`)`:= matrix.special_linear_group (fin n) R" in matrix_groups namespace special_linear_group diff --git a/src/linear_algebra/tensor_power.lean b/src/linear_algebra/tensor_power.lean index 360f289e3870e..abb55292695bc 100644 --- a/src/linear_algebra/tensor_power.lean +++ b/src/linear_algebra/tensor_power.lean @@ -66,7 +66,7 @@ def mul_equiv {n m : ℕ} : (⨂[R]^n M) ⊗[R] (⨂[R]^m M) ≃ₗ[R] ⨂[R]^(n instance ghas_mul : graded_monoid.ghas_mul (λ i, ⨂[R]^i M) := { mul := λ i j a b, mul_equiv (a ⊗ₜ b) } -local infix `ₜ*`:70 := @graded_monoid.ghas_mul.mul ℕ (λ i, ⨂[R]^i M) _ _ _ _ +local infix ` ₜ* `:70 := @graded_monoid.ghas_mul.mul ℕ (λ i, ⨂[R]^i M) _ _ _ _ lemma ghas_mul_def {i j} (a : ⨂[R]^i M) (b : ⨂[R]^j M) : a ₜ* b = mul_equiv (a ⊗ₜ b) := rfl diff --git a/src/logic/function/basic.lean b/src/logic/function/basic.lean index 4a8e328352e8e..427d6f658d8db 100644 --- a/src/logic/function/basic.lean +++ b/src/logic/function/basic.lean @@ -622,7 +622,7 @@ def bicompr (f : γ → δ) (g : α → β → γ) (a b) := f (g a b) -- Suggested local notation: -local notation f `∘₂` g := bicompr f g +local notation f ` ∘₂ ` g := bicompr f g lemma uncurry_bicompr (f : α → β → γ) (g : γ → δ) : uncurry (g ∘₂ f) = (g ∘ uncurry f) := rfl diff --git a/src/logic/relator.lean b/src/logic/relator.lean index 30aa8a5230e81..c9ea9b0afedac 100644 --- a/src/logic/relator.lean +++ b/src/logic/relator.lean @@ -26,7 +26,7 @@ variables (R : α → β → Prop) (S : γ → δ → Prop) def lift_fun (f : α → γ) (g : β → δ) : Prop := ∀⦃a b⦄, R a b → S (f a) (g b) -infixr ⇒ := lift_fun +infixr ` ⇒ ` := lift_fun end diff --git a/src/measure_theory/integral/divergence_theorem.lean b/src/measure_theory/integral/divergence_theorem.lean index 35585f5fc6dcb..8210300518012 100644 --- a/src/measure_theory/integral/divergence_theorem.lean +++ b/src/measure_theory/integral/divergence_theorem.lean @@ -61,7 +61,7 @@ variables {n : ℕ} local notation `ℝⁿ` := fin n → ℝ local notation `ℝⁿ⁺¹` := fin (n + 1) → ℝ local notation `Eⁿ⁺¹` := fin (n + 1) → E -local notation `e` i := pi.single i 1 +local notation `e ` i := pi.single i 1 section @@ -230,9 +230,9 @@ end variables (a b : ℝⁿ⁺¹) -local notation `face` i := set.Icc (a ∘ fin.succ_above i) (b ∘ fin.succ_above i) -local notation `front_face` i:2000 := fin.insert_nth i (b i) -local notation `back_face` i:2000 := fin.insert_nth i (a i) +local notation `face ` i := set.Icc (a ∘ fin.succ_above i) (b ∘ fin.succ_above i) +local notation `front_face ` i:2000 := fin.insert_nth i (b i) +local notation `back_face ` i:2000 := fin.insert_nth i (a i) /-- **Divergence theorem** for Bochner integral. If `f : ℝⁿ⁺¹ → Eⁿ⁺¹` is continuous on a rectangular box `[a, b] : set ℝⁿ⁺¹`, `a ≤ b`, is differentiable on its interior with derivative diff --git a/src/measure_theory/integral/torus_integral.lean b/src/measure_theory/integral/torus_integral.lean index 01fdf2694ed15..d13836f451806 100644 --- a/src/measure_theory/integral/torus_integral.lean +++ b/src/measure_theory/integral/torus_integral.lean @@ -60,14 +60,14 @@ noncomputable theory open complex set measure_theory function filter topological_space open_locale real big_operators -local notation `ℝ⁰`:= fin 0 → ℝ -local notation `ℂ⁰`:= fin 0 → ℂ -local notation `ℝ¹`:= fin 1 → ℝ -local notation `ℂ¹`:= fin 1 → ℂ -local notation `ℝⁿ`:= fin n → ℝ -local notation `ℂⁿ`:= fin n → ℂ -local notation `ℝⁿ⁺¹`:= fin (n + 1) → ℝ -local notation `ℂⁿ⁺¹`:= fin (n + 1) → ℂ +local notation `ℝ⁰` := fin 0 → ℝ +local notation `ℂ⁰` := fin 0 → ℂ +local notation `ℝ¹` := fin 1 → ℝ +local notation `ℂ¹` := fin 1 → ℂ +local notation `ℝⁿ` := fin n → ℝ +local notation `ℂⁿ` := fin n → ℂ +local notation `ℝⁿ⁺¹` := fin (n + 1) → ℝ +local notation `ℂⁿ⁺¹` := fin (n + 1) → ℂ /-! ### `torus_map`, a generalization of a torus diff --git a/src/order/basic.lean b/src/order/basic.lean index 1e4188049bb58..cb0ffdeaba3da 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -555,8 +555,8 @@ end min_max_rec /-- Typeclass for the `⊓` (`\glb`) notation -/ @[notation_class] class has_inf (α : Type u) := (inf : α → α → α) -infix ⊔ := has_sup.sup -infix ⊓ := has_inf.inf +infix ` ⊔ ` := has_sup.sup +infix ` ⊓ ` := has_inf.inf /-! ### Lifts of order instances -/ diff --git a/src/ring_theory/derivation.lean b/src/ring_theory/derivation.lean index 3f249a07dd6f6..903b95af667f5 100644 --- a/src/ring_theory/derivation.lean +++ b/src/ring_theory/derivation.lean @@ -563,7 +563,7 @@ Note that the slash is `\textfractionsolidus`. @[derive [add_comm_group, module R, module S, module (S ⊗[R] S)]] def kaehler_differential : Type* := (kaehler_differential.ideal R S).cotangent -notation `Ω[ `:100 S ` ⁄ `:0 R ` ]`:0 := kaehler_differential R S +notation `Ω[`:100 S `⁄`:0 R `]`:0 := kaehler_differential R S instance : nonempty Ω[S⁄R] := ⟨0⟩ diff --git a/src/ring_theory/witt_vector/isocrystal.lean b/src/ring_theory/witt_vector/isocrystal.lean index 853f192542547..74f49caeefa05 100644 --- a/src/ring_theory/witt_vector/isocrystal.lean +++ b/src/ring_theory/witt_vector/isocrystal.lean @@ -59,7 +59,7 @@ namespace witt_vector variables (p : ℕ) [fact p.prime] variables (k : Type*) [comm_ring k] localized "notation (name := witt_vector.fraction_ring) - `K(` p`,` k`)` := fraction_ring (witt_vector p k)" in isocrystal + `K(`p`, `k`)` := fraction_ring (witt_vector p k)" in isocrystal section perfect_ring variables [is_domain k] [char_p k p] [perfect_ring k p] @@ -74,7 +74,7 @@ is_fraction_ring.field_equiv_of_ring_equiv (frobenius_equiv p k) def fraction_ring.frobenius_ring_hom : K(p, k) →+* K(p, k) := fraction_ring.frobenius p k localized "notation (name := witt_vector.frobenius_ring_hom) - `φ(` p`,` k`)` := witt_vector.fraction_ring.frobenius_ring_hom p k" in isocrystal + `φ(`p`, `k`)` := witt_vector.fraction_ring.frobenius_ring_hom p k" in isocrystal instance inv_pair₁ : ring_hom_inv_pair (φ(p, k)) _ := ring_hom_inv_pair.of_ring_equiv (fraction_ring.frobenius p k) @@ -83,9 +83,9 @@ instance inv_pair₂ : ring_hom_inv_pair ((fraction_ring.frobenius p k).symm : K(p, k) →+* K(p, k)) _ := ring_hom_inv_pair.of_ring_equiv (fraction_ring.frobenius p k).symm -localized "notation (name := frobenius_ring_hom.linear_map) M ` →ᶠˡ[`:50 p `,` k `] ` M₂ := +localized "notation (name := frobenius_ring_hom.linear_map) M ` →ᶠˡ[`:50 p `, ` k `] ` M₂ := linear_map (witt_vector.fraction_ring.frobenius_ring_hom p k) M M₂" in isocrystal -localized "notation (name := frobenius_ring_hom.linear_equiv) M ` ≃ᶠˡ[`:50 p `,` k `] ` M₂ := +localized "notation (name := frobenius_ring_hom.linear_equiv) M ` ≃ᶠˡ[`:50 p `, ` k `] ` M₂ := linear_equiv (witt_vector.fraction_ring.frobenius_ring_hom p k) M M₂" in isocrystal /-! ### Isocrystals -/ @@ -108,7 +108,7 @@ Project the Frobenius automorphism from an isocrystal. Denoted by `Φ(p, k)` whe def isocrystal.frobenius : V ≃ᶠˡ[p, k] V := @isocrystal.frob p _ k _ _ _ _ _ _ _ variables (V) -localized "notation `Φ(` p`,` k`)` := witt_vector.isocrystal.frobenius p k" in isocrystal +localized "notation `Φ(`p`, `k`)` := witt_vector.isocrystal.frobenius p k" in isocrystal /-- A homomorphism between isocrystals respects the Frobenius map. -/ @[nolint has_nonempty_instance] @@ -121,9 +121,9 @@ structure isocrystal_equiv extends V ≃ₗ[K(p, k)] V₂ := ( frob_equivariant : ∀ x : V, Φ(p, k) (to_linear_equiv x) = to_linear_equiv (Φ(p, k) x) ) localized "notation (name := isocrystal_hom) - M ` →ᶠⁱ[`:50 p `,` k `] ` M₂ := witt_vector.isocrystal_hom p k M M₂" in isocrystal + M ` →ᶠⁱ[`:50 p `, ` k `] ` M₂ := witt_vector.isocrystal_hom p k M M₂" in isocrystal localized "notation (name := isocrystal_equiv) - M ` ≃ᶠⁱ[`:50 p `,` k `] ` M₂ := witt_vector.isocrystal_equiv p k M M₂" in isocrystal + M ` ≃ᶠⁱ[`:50 p `, ` k `] ` M₂ := witt_vector.isocrystal_equiv p k M M₂" in isocrystal end perfect_ring diff --git a/src/ring_theory/witt_vector/witt_polynomial.lean b/src/ring_theory/witt_vector/witt_polynomial.lean index 8a4a5d7e6fc46..5d20e8308e910 100644 --- a/src/ring_theory/witt_vector/witt_polynomial.lean +++ b/src/ring_theory/witt_vector/witt_polynomial.lean @@ -90,7 +90,7 @@ This allows us to simply write `W n` or `W_ ℤ n`. -/ -- Notation with ring of coefficients explicit localized "notation (name := witt_polynomial) `W_` := witt_polynomial p" in witt -- Notation with ring of coefficients implicit -localized "notation (name := witt_polynomial.infer) `W` := witt_polynomial p hole!" in witt +localized "notation (name := witt_polynomial.infer) `W` := witt_polynomial p hole!" in witt open_locale witt open mv_polynomial diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index 2798b655518f7..e94aa2580d581 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -103,7 +103,7 @@ namespace cardinal /-- The cardinal number of a type -/ def mk : Type u → cardinal := quotient.mk -localized "notation (name := cardinal.mk) `#` := cardinal.mk" in cardinal +localized "prefix (name := cardinal.mk) `#` := cardinal.mk" in cardinal instance can_lift_cardinal_Type : can_lift cardinal.{u} (Type u) := ⟨mk, λ c, true, λ c _, quot.induction_on c $ λ α, ⟨α, rfl⟩⟩ diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index 3447accd70398..df5bbe6c3c911 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -931,7 +931,7 @@ mem_univ.2 $ or.elim (classical.em $ ∃ x, ∀ y, A y ↔ y = x) /-- Function value -/ def fval (F A : Class.{u}) : Class.{u} := iota (λ y, to_Set (λ x, F (Set.pair x y)) A) -infixl `′`:100 := fval +infixl ` ′ `:100 := fval theorem fval_ex (F A : Class.{u}) : F ′ A ∈ univ.{u} := iota_ex _ diff --git a/src/tactic/omega/find_ees.lean b/src/tactic/omega/find_ees.lean index 9529649b46846..83968c777d898 100644 --- a/src/tactic/omega/find_ees.lean +++ b/src/tactic/omega/find_ees.lean @@ -79,7 +79,7 @@ do x ← ((t1 >>= return ∘ some) <|> return none), | (some a) := t3 a end -local notation t1 `!>>=` t2 `;` t3 := ee_commit t1 t2 t3 +local notation t1 ` !>>= ` t2 `; ` t3 := ee_commit t1 t2 t3 private meta def of_tactic {α : Type} : tactic α → eqelim α := state_t.lift @@ -126,8 +126,8 @@ do (i,n) ← find_min_coeff_core t.snd, meta def elim_eq : eqelim unit := do t ← head_eq, i ← get_gcd t, - factor i t !>>= (set_eqs [] >> add_ee (ee.nondiv i)) ; -λ s, find_min_coeff s !>>= add_ee ee.drop ; + factor i t !>>= (set_eqs [] >> add_ee (ee.nondiv i)); +λ s, find_min_coeff s !>>= add_ee ee.drop; λ ⟨i, n, u⟩, if i = 1 then do eqs ← get_eqs, @@ -147,11 +147,11 @@ else let v : term := coeffs_reduce n u.fst u.snd in /-- Find and return the sequence of steps for eliminating all equality constraints in the current state. -/ meta def elim_eqs : eqelim (list ee) := -elim_eq !>>= get_ees ; λ _, elim_eqs +elim_eq !>>= get_ees; λ _, elim_eqs /-- Given a linear constrain clause, return a list of steps for eliminating its equality constraints. -/ meta def find_ees : clause → tactic (list ee) -| (eqs,les) := run eqs les elim_eqs +| (eqs, les) := run eqs les elim_eqs end omega diff --git a/src/tactic/omega/int/preterm.lean b/src/tactic/omega/int/preterm.lean index dafc9fc939d60..5d793e5325c99 100644 --- a/src/tactic/omega/int/preterm.lean +++ b/src/tactic/omega/int/preterm.lean @@ -32,9 +32,9 @@ inductive preterm : Type | var : int → nat → preterm | add : preterm → preterm → preterm -localized "notation (name := preterm.cst) `&` k := omega.int.preterm.cst k" in omega.int -localized "infix (name := preterm.var) ` ** ` : 300 := omega.int.preterm.var" in omega.int -localized "notation (name := preterm.add) t `+*` s := omega.int.preterm.add t s" in omega.int +localized "notation (name := preterm.cst) `&` k := omega.int.preterm.cst k" in omega.int +localized "infix (name := preterm.var) ` ** `:300 := omega.int.preterm.var" in omega.int +localized "notation (name := preterm.add) t ` +* ` s := omega.int.preterm.add t s" in omega.int namespace preterm diff --git a/src/tactic/omega/misc.lean b/src/tactic/omega/misc.lean index adb61ec932185..2931b7282e1eb 100644 --- a/src/tactic/omega/misc.lean +++ b/src/tactic/omega/misc.lean @@ -31,7 +31,7 @@ lemma pred_mono_2' {c : Prop → Prop → Prop} {a1 a2 b1 b2 : Prop} : def update (m : nat) (a : α) (v : nat → α) : nat → α | n := if n = m then a else v n -localized "notation (name := omega.update) v ` ⟨` m ` ↦ ` a `⟩` := omega.update m a v" in omega +localized "notation (name := omega.update) v ` ⟨`m` ↦ `a`⟩` := omega.update m a v" in omega lemma update_eq (m : nat) (a : α) (v : nat → α) : (v ⟨m ↦ a⟩) m = a := by simp only [update, if_pos rfl] diff --git a/src/tactic/omega/nat/preterm.lean b/src/tactic/omega/nat/preterm.lean index a0ef3c187fda2..b55af2eb58577 100644 --- a/src/tactic/omega/nat/preterm.lean +++ b/src/tactic/omega/nat/preterm.lean @@ -37,7 +37,7 @@ inductive preterm : Type | sub : preterm → preterm → preterm localized "notation (name := preterm.cst) `&` k := omega.nat.preterm.cst k" in omega.nat -localized "infix (name := preterm.var) ` ** ` : 300 := omega.nat.preterm.var" in omega.nat +localized "infix (name := preterm.var) ` ** `:300 := omega.nat.preterm.var" in omega.nat localized "notation (name := preterm.add) t ` +* ` s := omega.nat.preterm.add t s" in omega.nat localized "notation (name := preterm.sub) t ` -* ` s := omega.nat.preterm.sub t s" in omega.nat diff --git a/src/tactic/reserved_notation.lean b/src/tactic/reserved_notation.lean index 0f18d64545518..d34358ece4121 100644 --- a/src/tactic/reserved_notation.lean +++ b/src/tactic/reserved_notation.lean @@ -60,4 +60,4 @@ reserve infixl ` ⊔ `:68 reserve infix ` ≃ₗ `:25 -- used in `data/matrix/notation.lean` -reserve notation `!![` +reserve notation `!![` diff --git a/src/tactic/rewrite.lean b/src/tactic/rewrite.lean index 6fee7e9216cb9..8398ceae981d8 100644 --- a/src/tactic/rewrite.lean +++ b/src/tactic/rewrite.lean @@ -204,7 +204,7 @@ It works for any function `f` for which an `is_associative f` instance can be fo ``` example {α : Type*} (f : α → α → α) [is_associative α f] (a b c d x : α) : - let infix `~` := f in + let infix ` ~ ` := f in b ~ c = x → (a ~ b ~ c ~ d) = (a ~ x ~ d) := begin intro h, diff --git a/src/topology/uniform_space/abstract_completion.lean b/src/topology/uniform_space/abstract_completion.lean index 6c5960557923b..75c54a385929b 100644 --- a/src/topology/uniform_space/abstract_completion.lean +++ b/src/topology/uniform_space/abstract_completion.lean @@ -299,7 +299,7 @@ variables {γ : Type*} [uniform_space γ] (pkg'' : abstract_completion γ) local notation `hatγ` := pkg''.space local notation `ι''` := pkg''.coe -local notation f `∘₂` g := bicompr f g +local notation f ` ∘₂ ` g := bicompr f g /-- Lift two variable maps to completions. -/ protected def map₂ (f : α → β → γ) : hatα → hatβ → hatγ := diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index c63b90d52d037..fe33e6b76513c 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -1473,7 +1473,7 @@ open uniform_space function variables {δ' : Type*} [uniform_space α] [uniform_space β] [uniform_space γ] [uniform_space δ] [uniform_space δ'] -local notation f `∘₂` g := function.bicompr f g +local notation f ` ∘₂ ` g := function.bicompr f g /-- Uniform continuity for functions of two variables. -/ def uniform_continuous₂ (f : α → β → γ) := uniform_continuous (uncurry f) diff --git a/src/topology/uniform_space/uniform_convergence_topology.lean b/src/topology/uniform_space/uniform_convergence_topology.lean index 5915f96223940..92fdd99b70693 100644 --- a/src/topology/uniform_space/uniform_convergence_topology.lean +++ b/src/topology/uniform_space/uniform_convergence_topology.lean @@ -222,7 +222,7 @@ protected def uniform_space : uniform_space (α → β) := uniform_space.of_core (uniform_convergence.uniform_core α β) local attribute [instance] uniform_convergence.uniform_space -local notation `𝒰(` α `,` β `,` u `)` := @uniform_convergence.uniform_space α β u +local notation `𝒰(`α`, `β`, `u`)` := @uniform_convergence.uniform_space α β u /-- By definition, the uniformity of `α → β` endowed with the structure of uniform convergence on `α` admits the family `{(f, g) | ∀ x, (f x, g x) ∈ V}` for `V ∈ 𝓤 β` as a filter basis. -/ @@ -472,7 +472,7 @@ namespace uniform_convergence_on variables (α β : Type*) {γ ι : Type*} [uniform_space β] (𝔖 : set (set α)) variables {F : ι → α → β} {f : α → β} {s s' : set α} {x : α} {p : filter ι} {g : ι → α} -local notation `𝒰(` α `,` β `,` u `)` := @uniform_convergence.uniform_space α β u +local notation `𝒰(`α`, `β`, `u`)` := @uniform_convergence.uniform_space α β u /-- Uniform structure of `𝔖`-convergence, i.e uniform convergence on the elements of `𝔖`. It is defined as the infimum, for `S ∈ 𝔖`, of the pullback of `𝒰 S β` by `S.restrict`, the @@ -482,7 +482,7 @@ protected def uniform_space : uniform_space (α → β) := ⨅ (s : set α) (hs : s ∈ 𝔖), uniform_space.comap s.restrict (𝒰(s, β, _)) -local notation `𝒱(` α `,` β `,` 𝔖 `,` u `)` := @uniform_convergence_on.uniform_space α β u 𝔖 +local notation `𝒱(`α`, `β`, `𝔖`, `u`)` := @uniform_convergence_on.uniform_space α β u 𝔖 /-- Topology of `𝔖`-convergence, i.e uniform convergence on the elements of `𝔖`. -/ protected def topological_space : topological_space (α → β) := diff --git a/test/noncomm_ring.lean b/test/noncomm_ring.lean index 47e487c84a3e4..8fcb1ce3d0545 100644 --- a/test/noncomm_ring.lean +++ b/test/noncomm_ring.lean @@ -1,6 +1,6 @@ import tactic.noncomm_ring -local notation (name := commutator) `⁅`a`,` b`⁆` := a * b - b * a +local notation (name := commutator) `⁅`a`, `b`⁆` := a * b - b * a local infix (name := op) ` ⚬ `:70 := λ a b, a * b + b * a variables {R : Type*} [ring R] diff --git a/test/simps.lean b/test/simps.lean index 6ab8a76833210..74a558af4b53f 100644 --- a/test/simps.lean +++ b/test/simps.lean @@ -421,7 +421,7 @@ example {α β} [semigroup α] [semigroup β] (x y : α × β) : (x * y).1 = x.1 structure Semigroup := (G : Type*) (op : G → G → G) - (infix (name := op) * := op) + (infix (name := op) ` * ` := op) (op_assoc : ∀ (x y z : G), (x * y) * z = x * (y * z)) namespace Group From 370dabe700255712a96284510866629e56e9fb41 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 8 Sep 2022 13:17:09 +0000 Subject: [PATCH 203/828] feat(algebra/big_operators/multiset): add `prod_map_le_prod_map` (#16418) --- src/algebra/big_operators/multiset.lean | 5 +++++ src/algebra/big_operators/order.lean | 8 +------- 2 files changed, 6 insertions(+), 7 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index 181cde814b5c6..c5f6340148181 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -316,6 +316,11 @@ begin exact mul_le_mul' rh rt } end +@[to_additive] +lemma prod_map_le_prod_map {s : multiset ι} (f : ι → α) (g : ι → α) (h : ∀ i, i ∈ s → f i ≤ g i) : + (s.map f).prod ≤ (s.map g).prod := +prod_le_prod_of_rel_le $ rel_map.2 $ rel_refl_of_refl_on h + @[to_additive] lemma prod_map_le_prod (f : α → α) (h : ∀ x, x ∈ s → f x ≤ x) : (s.map f).prod ≤ s.prod := prod_le_prod_of_rel_le $ rel_map_left.2 $ rel_refl_of_refl_on h diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 8c7030a2ecf03..68bb20bb7cc68 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -106,13 +106,7 @@ equal to the corresponding factor `g i` of another finite product, then `∏ i in s, f i ≤ ∏ i in s, g i`. -/ @[to_additive sum_le_sum] lemma prod_le_prod'' (h : ∀ i ∈ s, f i ≤ g i) : ∏ i in s, f i ≤ ∏ i in s, g i := -begin - classical, - induction s using finset.induction_on with i s hi ihs h, - { refl }, - { simp only [prod_insert hi], - exact mul_le_mul' (h _ (mem_insert_self _ _)) (ihs $ λ j hj, h j (mem_insert_of_mem hj)) } -end +multiset.prod_map_le_prod_map f g h /-- In an ordered additive commutative monoid, if each summand `f i` of one finite sum is less than or equal to the corresponding summand `g i` of another finite sum, then From d75314b3ad91c89df7fe401e8a2dc6a855c3c441 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 8 Sep 2022 13:17:12 +0000 Subject: [PATCH 204/828] fix(data/nat/interval): do not dedup when implementing `finset.Icc` etc (#16423) This means that `finset.Iic n` no longer has quadratic complexity. `#eval (finset.Iic 200000).card` is now almost instant rather than taking a very long time. --- src/data/nat/interval.lean | 40 +++++++++++++++----------------------- 1 file changed, 16 insertions(+), 24 deletions(-) diff --git a/src/data/nat/interval.lean b/src/data/nat/interval.lean index 67abcfce8f957..96181cf439d5d 100644 --- a/src/data/nat/interval.lean +++ b/src/data/nat/interval.lean @@ -20,26 +20,26 @@ and subsequently be moved upstream to `data.finset.locally_finite`. open finset nat instance : locally_finite_order ℕ := -{ finset_Icc := λ a b, (list.range' a (b + 1 - a)).to_finset, - finset_Ico := λ a b, (list.range' a (b - a)).to_finset, - finset_Ioc := λ a b, (list.range' (a + 1) (b - a)).to_finset, - finset_Ioo := λ a b, (list.range' (a + 1) (b - a - 1)).to_finset, +{ finset_Icc := λ a b, ⟨list.range' a (b + 1 - a), list.nodup_range' _ _⟩, + finset_Ico := λ a b, ⟨list.range' a (b - a), list.nodup_range' _ _⟩, + finset_Ioc := λ a b, ⟨list.range' (a + 1) (b - a), list.nodup_range' _ _⟩, + finset_Ioo := λ a b, ⟨list.range' (a + 1) (b - a - 1), list.nodup_range' _ _⟩, finset_mem_Icc := λ a b x, begin - rw [list.mem_to_finset, list.mem_range'], + rw [finset.mem_mk, multiset.mem_coe, list.mem_range'], cases le_or_lt a b, { rw [add_tsub_cancel_of_le (nat.lt_succ_of_le h).le, nat.lt_succ_iff] }, { rw [tsub_eq_zero_iff_le.2 (succ_le_of_lt h), add_zero], exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2)) } end, finset_mem_Ico := λ a b x, begin - rw [list.mem_to_finset, list.mem_range'], + rw [finset.mem_mk, multiset.mem_coe, list.mem_range'], cases le_or_lt a b, { rw [add_tsub_cancel_of_le h] }, { rw [tsub_eq_zero_iff_le.2 h.le, add_zero], exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.trans hx.2.le)) } end, finset_mem_Ioc := λ a b x, begin - rw [list.mem_to_finset, list.mem_range'], + rw [finset.mem_mk, multiset.mem_coe, list.mem_range'], cases le_or_lt a b, { rw [←succ_sub_succ, add_tsub_cancel_of_le (succ_le_succ h), nat.lt_succ_iff, nat.succ_le_iff] }, @@ -47,7 +47,7 @@ instance : locally_finite_order ℕ := exact iff_of_false (λ hx, hx.2.not_le hx.1) (λ hx, h.not_le (hx.1.le.trans hx.2)) } end, finset_mem_Ioo := λ a b x, begin - rw [list.mem_to_finset, list.mem_range', ← tsub_add_eq_tsub_tsub], + rw [finset.mem_mk, multiset.mem_coe, list.mem_range', ← tsub_add_eq_tsub_tsub], cases le_or_lt (a + 1) b, { rw [add_tsub_cancel_of_le h, nat.succ_le_iff] }, { rw [tsub_eq_zero_iff_le.2 h.le, add_zero], @@ -58,10 +58,10 @@ variables (a b c : ℕ) namespace nat -lemma Icc_eq_range' : Icc a b = (list.range' a (b + 1 - a)).to_finset := rfl -lemma Ico_eq_range' : Ico a b = (list.range' a (b - a)).to_finset := rfl -lemma Ioc_eq_range' : Ioc a b = (list.range' (a + 1) (b - a)).to_finset := rfl -lemma Ioo_eq_range' : Ioo a b = (list.range' (a + 1) (b - a - 1)).to_finset := rfl +lemma Icc_eq_range' : Icc a b = ⟨list.range' a (b + 1 - a), list.nodup_range' _ _⟩ := rfl +lemma Ico_eq_range' : Ico a b = ⟨list.range' a (b - a), list.nodup_range' _ _⟩ := rfl +lemma Ioc_eq_range' : Ioc a b = ⟨list.range' (a + 1) (b - a), list.nodup_range' _ _⟩ := rfl +lemma Ioo_eq_range' : Ioo a b = ⟨list.range' (a + 1) (b - a - 1), list.nodup_range' _ _⟩ := rfl lemma Iio_eq_range : Iio = range := by { ext b x, rw [mem_Iio, mem_range] } @@ -69,18 +69,10 @@ lemma Iio_eq_range : Iio = range := by { ext b x, rw [mem_Iio, mem_range] } lemma _root_.finset.range_eq_Ico : range = Ico 0 := Ico_zero_eq_range.symm -@[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := -by rw [Icc_eq_range', list.card_to_finset, (list.nodup_range' _ _).dedup, list.length_range'] - -@[simp] lemma card_Ico : (Ico a b).card = b - a := -by rw [Ico_eq_range', list.card_to_finset, (list.nodup_range' _ _).dedup, list.length_range'] - -@[simp] lemma card_Ioc : (Ioc a b).card = b - a := -by rw [Ioc_eq_range', list.card_to_finset, (list.nodup_range' _ _).dedup, list.length_range'] - -@[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := -by rw [Ioo_eq_range', list.card_to_finset, (list.nodup_range' _ _).dedup, list.length_range'] - +@[simp] lemma card_Icc : (Icc a b).card = b + 1 - a := list.length_range' _ _ +@[simp] lemma card_Ico : (Ico a b).card = b - a := list.length_range' _ _ +@[simp] lemma card_Ioc : (Ioc a b).card = b - a := list.length_range' _ _ +@[simp] lemma card_Ioo : (Ioo a b).card = b - a - 1 := list.length_range' _ _ @[simp] lemma card_Iic : (Iic b).card = b + 1 := by rw [Iic_eq_Icc, card_Icc, bot_eq_zero, tsub_zero] From aff61c0f687c5d366edb248791f26fa69a4adfe2 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 8 Sep 2022 13:17:13 +0000 Subject: [PATCH 205/828] chore(algebra/big_operators/basic): golf (#16427) --- src/algebra/big_operators/basic.lean | 7 ++----- 1 file changed, 2 insertions(+), 5 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index a77d695ab33be..13446fadb00c1 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -1129,9 +1129,7 @@ begin end @[simp, to_additive] lemma prod_const (b : β) : (∏ x in s, b) = b ^ s.card := -by haveI := classical.dec_eq α; exact -finset.induction_on s (by simp) (λ a s has ih, -by rw [prod_insert has, card_insert_of_not_mem has, pow_succ, ih]) +(congr_arg _ $ s.val.map_const b).trans $ multiset.prod_repeat b s.card @[to_additive] lemma pow_eq_prod_const (b : β) : ∀ n, b ^ n = ∏ k in range n, b := by simp @@ -1139,8 +1137,7 @@ lemma pow_eq_prod_const (b : β) : ∀ n, b ^ n = ∏ k in range n, b := by simp @[to_additive] lemma prod_pow (s : finset α) (n : ℕ) (f : α → β) : ∏ x in s, f x ^ n = (∏ x in s, f x) ^ n := -by haveI := classical.dec_eq α; exact -finset.induction_on s (by simp) (by simp [mul_pow] {contextual := tt}) +multiset.prod_map_pow @[to_additive] lemma prod_flip {n : ℕ} (f : ℕ → β) : From d645d0e8fd7c243aa54c94d5bfa4cdf595dfc496 Mon Sep 17 00:00:00 2001 From: Kevin Wilson Date: Thu, 8 Sep 2022 16:05:58 +0000 Subject: [PATCH 206/828] feat(analysis/calculus/uniform_limits_deriv): Swap limits and derivatives (#14090) This commit proves that the derivative of the pointwise limit of a series of functions is the limit of the derivatives when the derivatives converge uniformly in some closed ball. This and #13500 are two fundamental theorems for bootstrapping the theory of Dirichlet series. This theory underlies my ongoing attempts to prove the density of squarefree integers. --- .../calculus/uniform_limits_deriv.lean | 544 ++++++++++++++++++ src/analysis/normed/group/basic.lean | 36 ++ src/order/filter/curry.lean | 89 +++ src/topology/metric_space/basic.lean | 11 + 4 files changed, 680 insertions(+) create mode 100644 src/analysis/calculus/uniform_limits_deriv.lean create mode 100644 src/order/filter/curry.lean diff --git a/src/analysis/calculus/uniform_limits_deriv.lean b/src/analysis/calculus/uniform_limits_deriv.lean new file mode 100644 index 0000000000000..4416629d8836d --- /dev/null +++ b/src/analysis/calculus/uniform_limits_deriv.lean @@ -0,0 +1,544 @@ +/- +Copyright (c) 2022 Kevin H. Wilson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin H. Wilson +-/ +import analysis.calculus.mean_value +import analysis.normed_space.is_R_or_C +import order.filter.curry + +/-! +# Swapping limits and derivatives via uniform convergence + +The purpose of this file is to prove that the derivative of the pointwise limit of a sequence of +functions is the pointwise limit of the functions' derivatives when the derivatives converge +_uniformly_. The formal statement appears as `has_fderiv_at_of_tendsto_locally_uniformly_at`. + +## Main statements + +* `uniform_cauchy_seq_on_filter_of_tendsto_uniformly_on_filter_fderiv`: If + 1. `f : ℕ → E → G` is a sequence of functions which have derivatives + `f' : ℕ → E → (E →L[𝕜] G)` on a neighborhood of `x`, + 2. the functions `f` converge at `x`, and + 3. the derivatives `f'` converge uniformly on a neighborhood of `x`, + then the `f` converge _uniformly_ on a neighborhood of `x` +* `has_fderiv_at_of_tendsto_uniformly_on_filter` : Suppose (1), (2), and (3) above are true. Let + `g` (resp. `g'`) be the limiting function of the `f` (resp. `g'`). Then `f'` is the derivative of + `g` on a neighborhood of `x` +* `has_fderiv_at_of_tendsto_uniformly_on`: An often-easier-to-use version of the above theorem when + *all* the derivatives exist and functions converge on a common open set and the derivatives + converge uniformly there. + +Each of the above statements also has variations that support `deriv` instead of `fderiv`. + +## Implementation notes + +Our technique for proving the main result is the famous "`ε / 3` proof." In words, you can find it +explained, for instance, at [this StackExchange post](https://math.stackexchange.com/questions/214218/uniform-convergence-of-derivatives-tao-14-2-7). +The subtlety is that we want to prove that the difference quotients of the `g` converge to the `g'`. +That is, we want to prove something like: + +``` +∀ ε > 0, ∃ δ > 0, ∀ y ∈ B_δ(x), |y - x|⁻¹ * |(g y - g x) - g' x (y - x)| < ε. +``` + +To do so, we will need to introduce a pair of quantifers + +```lean +∀ ε > 0, ∃ N, ∀ n ≥ N, ∃ δ > 0, ∀ y ∈ B_δ(x), |y - x|⁻¹ * |(g y - g x) - g' x (y - x)| < ε. +``` + +So how do we write this in terms of filters? Well, the initial definition of the derivative is + +```lean +tendsto (|y - x|⁻¹ * |(g y - g x) - g' x (y - x)|) (𝓝 x) (𝓝 0) +``` + +There are two ways we might introduce `n`. We could do: + +```lean +∀ᶠ (n : ℕ) in at_top, tendsto (|y - x|⁻¹ * |(g y - g x) - g' x (y - x)|) (𝓝 x) (𝓝 0) +``` + +but this is equivalent to the quantifier order `∃ N, ∀ n ≥ N, ∀ ε > 0, ∃ δ > 0, ∀ y ∈ B_δ(x)`, +which _implies_ our desired `∀ ∃ ∀ ∃ ∀` but is _not_ equivalent to it. On the other hand, we might +try + +```lean +tendsto (|y - x|⁻¹ * |(g y - g x) - g' x (y - x)|) (at_top ×ᶠ 𝓝 x) (𝓝 0) +``` + +but this is equivalent to the quantifer order `∀ ε > 0, ∃ N, ∃ δ > 0, ∀ n ≥ N, ∀ y ∈ B_δ(x)`, which +again _implies_ our desired `∀ ∃ ∀ ∃ ∀` but is not equivalent to it. + +So to get the quantifier order we want, we need to introduce a new filter construction, which we +call a "curried filter" + +```lean +tendsto (|y - x|⁻¹ * |(g y - g x) - g' x (y - x)|) (at_top.curry (𝓝 x)) (𝓝 0) +``` + +Then the above implications are `filter.tendsto.curry` and +`filter.tendsto.mono_left filter.curry_le_prod`. We will use both of these deductions as part of +our proof. + +We note that if you loosen the assumptions of the main theorem then the proof becomes quite a bit +easier. In particular, if you assume there is a common neighborhood `s` where all of the three +assumptions of `has_fderiv_at_of_tendsto_uniformly_on_filter` hold and that the `f'` are +continuous, then you can avoid the mean value theorem and much of the work around curried filters. + +## Tags + +uniform convergence, limits of derivatives +-/ + +open filter +open_locale uniformity filter topological_space + +section limits_of_derivatives + +variables {ι : Type*} {l : filter ι} [ne_bot l] + {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] + {𝕜 : Type*} [is_R_or_C 𝕜] [normed_space 𝕜 E] + {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] + {f : ι → E → G} {g : E → G} {f' : ι → (E → (E →L[𝕜] G))} {g' : E → (E →L[𝕜] G)} + {x : E} + +/-- If a sequence of functions real or complex functions are eventually differentiable on a +neighborhood of `x`, they converge pointwise _at_ `x`, and their derivatives +converge uniformly in a neighborhood of `x`, then the functions form a uniform Cauchy sequence +in a neighborhood of `x`. -/ +lemma uniform_cauchy_seq_on_filter_of_tendsto_uniformly_on_filter_fderiv + (hf' : uniform_cauchy_seq_on_filter f' l (𝓝 x)) + (hf : ∀ᶠ (n : ι × E) in (l ×ᶠ 𝓝 x), has_fderiv_at (f n.1) (f' n.1 n.2) n.2) + (hfg : tendsto (λ n, f n x) l (𝓝 (g x))) : + uniform_cauchy_seq_on_filter f l (𝓝 x) := +begin + rw normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero at + hf' ⊢, + + suffices : tendsto_uniformly_on_filter + (λ (n : ι × ι) (z : E), f n.1 z - f n.2 z - (f n.1 x - f n.2 x)) 0 (l ×ᶠ l) (𝓝 x) ∧ + tendsto_uniformly_on_filter (λ (n : ι × ι) (z : E), f n.1 x - f n.2 x) 0 (l ×ᶠ l) (𝓝 x), + { have := this.1.add this.2, + rw add_zero at this, + exact this.congr (by simp), }, + split, + { -- This inequality follows from the mean value theorem. To apply it, we will need to shrink our + -- neighborhood to small enough ball + rw metric.tendsto_uniformly_on_filter_iff at hf' ⊢, + intros ε hε, + have := (tendsto_swap4_prod.eventually (hf.prod_mk hf)).diag_of_prod_right, + obtain ⟨a, b, c, d, e⟩ := eventually_prod_iff.1 ((hf' ε hε).and this), + obtain ⟨R, hR, hR'⟩ := metric.nhds_basis_ball.eventually_iff.mp d, + let r := min 1 R, + have hr : 0 < r, { simp [hR], }, + have hr' : ∀ ⦃y : E⦄, y ∈ metric.ball x r → c y, + { exact (λ y hy, hR' (lt_of_lt_of_le (metric.mem_ball.mp hy) (min_le_right _ _))), }, + have hxy : ∀ (y : E), y ∈ metric.ball x r → ∥y - x∥ < 1, + { intros y hy, + rw [metric.mem_ball, dist_eq_norm] at hy, + exact lt_of_lt_of_le hy (min_le_left _ _), }, + have hxyε : ∀ (y : E), y ∈ metric.ball x r → ε * ∥y - x∥ < ε, + { intros y hy, + exact (mul_lt_iff_lt_one_right hε.lt).mpr (hxy y hy), }, + + -- With a small ball in hand, apply the mean value theorem + refine eventually_prod_iff.mpr ⟨_, b, (λ e : E, metric.ball x r e), + eventually_mem_set.mpr (metric.nhds_basis_ball.mem_of_mem hr), (λ n hn y hy, _)⟩, + simp only [pi.zero_apply, dist_zero_left] at e ⊢, + refine lt_of_le_of_lt _ (hxyε y hy), + exact convex.norm_image_sub_le_of_norm_has_fderiv_within_le + (λ y hy, ((e hn (hr' hy)).2.1.sub (e hn (hr' hy)).2.2).has_fderiv_within_at) + (λ y hy, (e hn (hr' hy)).1.le) + (convex_ball x r) (metric.mem_ball_self hr) hy, }, + { -- This is just `hfg` run through `eventually_prod_iff` + refine metric.tendsto_uniformly_on_filter_iff.mpr (λ ε hε, _), + obtain ⟨t, ht, ht'⟩ := (metric.cauchy_iff.mp hfg.cauchy_map).2 ε hε, + exact eventually_prod_iff.mpr + ⟨ (λ (n : ι × ι), (f n.1 x ∈ t) ∧ (f n.2 x ∈ t)), + eventually_prod_iff.mpr ⟨_, ht, _, ht, (λ n hn n' hn', ⟨hn, hn'⟩)⟩, + (λ y, true), + (by simp), + (λ n hn y hy, by simpa [norm_sub_rev, dist_eq_norm] using ht' _ hn.1 _ hn.2)⟩, }, +end + +/-- A variant of the second fundamental theorem of calculus (FTC-2): If a sequence of functions +between real or complex normed spaces are differentiable on a ball centered at `x`, they +converge pointwise _at_ `x`, and their derivatives converge uniformly on the ball, then the +functions form a uniform Cauchy sequence on the ball. + +NOTE: The fact that we work on a ball is typically all that is necessary to work with power series +and Dirichlet series (our primary use case). However, this can be generalized by replacing the ball +with any connected, bounded, open set and replacing uniform convergence with local uniform +convergence. +-/ +lemma uniform_cauchy_seq_on_ball_of_tendsto_uniformly_on_ball_fderiv + {r : ℝ} (hr : 0 < r) + (hf' : uniform_cauchy_seq_on f' l (metric.ball x r)) + (hf : ∀ n : ι, ∀ y : E, y ∈ metric.ball x r → has_fderiv_at (f n) (f' n y) y) + (hfg : tendsto (λ n, f n x) l (𝓝 (g x))) : + uniform_cauchy_seq_on f l (metric.ball x r) := +begin + rw normed_add_comm_group.uniform_cauchy_seq_on_iff_tendsto_uniformly_on_zero at hf' ⊢, + + suffices : tendsto_uniformly_on + (λ (n : ι × ι) (z : E), f n.1 z - f n.2 z - (f n.1 x - f n.2 x)) 0 + (l ×ᶠ l) (metric.ball x r) ∧ + tendsto_uniformly_on (λ (n : ι × ι) (z : E), f n.1 x - f n.2 x) 0 + (l ×ᶠ l) (metric.ball x r), + { have := this.1.add this.2, + rw add_zero at this, + refine this.congr _, + apply eventually_of_forall, + intros n z hz, + simp, }, + split, + { -- This inequality follows from the mean value theorem + rw metric.tendsto_uniformly_on_iff at hf' ⊢, + intros ε hε, + obtain ⟨q, hqpos, hq⟩ : ∃ q : ℝ, 0 < q ∧ q * r < ε, + { simp_rw mul_comm, + exact exists_pos_mul_lt hε.lt r, }, + apply (hf' q hqpos.gt).mono, + intros n hn y hy, + simp_rw [dist_eq_norm, pi.zero_apply, zero_sub, norm_neg] at hn ⊢, + have mvt := convex.norm_image_sub_le_of_norm_has_fderiv_within_le + (λ z hz, ((hf n.1 z hz).sub (hf n.2 z hz)).has_fderiv_within_at) + (λ z hz, (hn z hz).le) (convex_ball x r) (metric.mem_ball_self hr) hy, + refine lt_of_le_of_lt mvt _, + have : q * ∥y - x∥ < q * r, + { exact mul_lt_mul' rfl.le (by simpa only [dist_eq_norm] using metric.mem_ball.mp hy) + (norm_nonneg _) hqpos, }, + exact this.trans hq, }, + { -- This is just `hfg` run through `eventually_prod_iff` + refine metric.tendsto_uniformly_on_iff.mpr (λ ε hε, _), + obtain ⟨t, ht, ht'⟩ := (metric.cauchy_iff.mp hfg.cauchy_map).2 ε hε, + rw eventually_prod_iff, + refine ⟨(λ n, f n x ∈ t), ht, (λ n, f n x ∈ t), ht, _⟩, + intros n hn n' hn' z hz, + rw [dist_eq_norm, pi.zero_apply, zero_sub, norm_neg, ←dist_eq_norm], + exact (ht' _ hn _ hn'), }, +end + +/-- If `f_n → g` pointwise and the derivatives `(f_n)' → h` _uniformly_ converge, then +in fact for a fixed `y`, the difference quotients `∥z - y∥⁻¹ • (f_n z - f_n y)` converge +_uniformly_ to `∥z - y∥⁻¹ • (g z - g y)` -/ +lemma difference_quotients_converge_uniformly + (hf' : tendsto_uniformly_on_filter f' g' l (𝓝 x)) + (hf : ∀ᶠ (n : ι × E) in (l ×ᶠ 𝓝 x), has_fderiv_at (f n.1) (f' n.1 n.2) n.2) + (hfg : ∀ᶠ (y : E) in 𝓝 x, tendsto (λ n, f n y) l (𝓝 (g y))) : + tendsto_uniformly_on_filter + (λ n : ι, λ y : E, (∥y - x∥⁻¹ : 𝕜) • (f n y - f n x)) + (λ y : E, (∥y - x∥⁻¹ : 𝕜) • (g y - g x)) + l (𝓝 x) := +begin + refine uniform_cauchy_seq_on_filter.tendsto_uniformly_on_filter_of_tendsto _ + ((hfg.and (eventually_const.mpr hfg.self_of_nhds)).mono (λ y hy, (hy.1.sub hy.2).const_smul _)), + rw normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero, + rw metric.tendsto_uniformly_on_filter_iff, + + have hfg' := hf'.uniform_cauchy_seq_on_filter, + rw normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero at + hfg', + rw metric.tendsto_uniformly_on_filter_iff at hfg', + intros ε hε, + obtain ⟨q, hqpos, hqε⟩ := exists_pos_rat_lt hε, + specialize hfg' (q : ℝ) (by simp [hqpos]), + + have := (tendsto_swap4_prod.eventually (hf.prod_mk hf)).diag_of_prod_right, + obtain ⟨a, b, c, d, e⟩ := eventually_prod_iff.1 (hfg'.and this), + obtain ⟨r, hr, hr'⟩ := metric.nhds_basis_ball.eventually_iff.mp d, + + rw eventually_prod_iff, + refine ⟨_, b, (λ e : E, metric.ball x r e), + eventually_mem_set.mpr (metric.nhds_basis_ball.mem_of_mem hr), (λ n hn y hy, _)⟩, + simp only [pi.zero_apply, dist_zero_left], + rw [← smul_sub, norm_smul, norm_inv, is_R_or_C.norm_coe_norm], + refine lt_of_le_of_lt _ hqε, + by_cases hyz' : x = y, { simp [hyz', hqpos.le], }, + have hyz : 0 < ∥y - x∥, + {rw norm_pos_iff, intros hy', exact hyz' (eq_of_sub_eq_zero hy').symm, }, + rw [inv_mul_le_iff hyz, mul_comm, sub_sub_sub_comm], + simp only [pi.zero_apply, dist_zero_left] at e, + refine convex.norm_image_sub_le_of_norm_has_fderiv_within_le + (λ y hy, ((e hn (hr' hy)).2.1.sub (e hn (hr' hy)).2.2).has_fderiv_within_at) + (λ y hy, (e hn (hr' hy)).1.le) + (convex_ball x r) (metric.mem_ball_self hr) hy, +end + +/-- `(d/dx) lim_{n → ∞} f n x = lim_{n → ∞} f' n x` when the `f' n` converge +_uniformly_ to their limit at `x`. + +In words the assumptions mean the following: + * `hf'`: The `f'` converge "uniformly at" `x` to `g'`. This does not mean that the `f' n` even + converge away from `x`! + * `hf`: For all `(y, n)` with `y` sufficiently close to `x` and `n` sufficiently large, `f' n` is + the derivative of `f n` + * `hfg`: The `f n` converge pointwise to `g` on a neighborhood of `x` -/ +lemma has_fderiv_at_of_tendsto_uniformly_on_filter + (hf' : tendsto_uniformly_on_filter f' g' l (𝓝 x)) + (hf : ∀ᶠ (n : ι × E) in (l ×ᶠ 𝓝 x), has_fderiv_at (f n.1) (f' n.1 n.2) n.2) + (hfg : ∀ᶠ y in 𝓝 x, tendsto (λ n, f n y) l (𝓝 (g y))) : + has_fderiv_at g (g' x) x := +begin + -- The proof strategy follows several steps: + -- 1. The quantifiers in the definition of the derivative are + -- `∀ ε > 0, ∃δ > 0, ∀y ∈ B_δ(x)`. We will introduce a quantifier in the middle: + -- `∀ ε > 0, ∃N, ∀n ≥ N, ∃δ > 0, ∀y ∈ B_δ(x)` which will allow us to introduce the `f(') n` + -- 2. The order of the quantifiers `hfg` are opposite to what we need. We will be able to swap + -- the quantifiers using the uniform convergence assumption + rw has_fderiv_at_iff_tendsto, + + -- Introduce extra quantifier via curried filters + suffices : tendsto + (λ (y : ι × E), ∥y.2 - x∥⁻¹ * ∥g y.2 - g x - (g' x) (y.2 - x)∥) (l.curry (𝓝 x)) (𝓝 0), + { rw metric.tendsto_nhds at this ⊢, + intros ε hε, + specialize this ε hε, + rw eventually_curry_iff at this, + simp only at this, + exact (eventually_const.mp this).mono (by simp only [imp_self, forall_const]), }, + + -- With the new quantifier in hand, we can perform the famous `ε/3` proof. Specifically, + -- we will break up the limit (the difference functions minus the derivative go to 0) into 3: + -- * The difference functions of the `f n` converge *uniformly* to the difference functions + -- of the `g n` + -- * The `f' n` are the derivatives of the `f n` + -- * The `f' n` converge to `g'` at `x` + conv + { congr, funext, + rw [←norm_norm, ←norm_inv,←@is_R_or_C.norm_of_real 𝕜 _ _, + is_R_or_C.of_real_inv, ←norm_smul], }, + rw ←tendsto_zero_iff_norm_tendsto_zero, + have : (λ a : ι × E, (∥a.2 - x∥⁻¹ : 𝕜) • (g a.2 - g x - (g' x) (a.2 - x))) = + (λ a : ι × E, (∥a.2 - x∥⁻¹ : 𝕜) • (g a.2 - g x - (f a.1 a.2 - f a.1 x))) + + (λ a : ι × E, (∥a.2 - x∥⁻¹ : 𝕜) • ((f a.1 a.2 - f a.1 x) - + ((f' a.1 x) a.2 - (f' a.1 x) x))) + + (λ a : ι × E, (∥a.2 - x∥⁻¹ : 𝕜) • ((f' a.1 x - g' x) (a.2 - x))), + { ext, simp only [pi.add_apply], rw [←smul_add, ←smul_add], congr, + simp only [map_sub, sub_add_sub_cancel, continuous_linear_map.coe_sub', pi.sub_apply], }, + simp_rw this, + have : 𝓝 (0 : G) = 𝓝 (0 + 0 + 0), simp only [add_zero], + rw this, + refine tendsto.add (tendsto.add _ _) _, + simp only, + { have := difference_quotients_converge_uniformly hf' hf hfg, + rw metric.tendsto_uniformly_on_filter_iff at this, + rw metric.tendsto_nhds, + intros ε hε, + apply ((this ε hε).filter_mono curry_le_prod).mono, + intros n hn, + rw dist_eq_norm at hn ⊢, + rw ← smul_sub at hn, + rwa sub_zero, }, + { -- (Almost) the definition of the derivatives + rw metric.tendsto_nhds, + intros ε hε, + rw eventually_curry_iff, + refine hf.curry.mono (λ n hn, _), + have := hn.self_of_nhds, + rw [has_fderiv_at_iff_tendsto, metric.tendsto_nhds] at this, + refine (this ε hε).mono (λ y hy, _), + rw dist_eq_norm at hy ⊢, + simp only [sub_zero, map_sub, norm_mul, norm_inv, norm_norm] at hy ⊢, + rw [norm_smul, norm_inv, is_R_or_C.norm_coe_norm], + exact hy, }, + { -- hfg' after specializing to `x` and applying the definition of the operator norm + refine tendsto.mono_left _ curry_le_prod, + have h1: tendsto (λ n : ι × E, g' n.2 - f' n.1 n.2) (l ×ᶠ 𝓝 x) (𝓝 0), + { rw metric.tendsto_uniformly_on_filter_iff at hf', + exact metric.tendsto_nhds.mpr (λ ε hε, by simpa using hf' ε hε), }, + have h2: tendsto (λ n : ι, g' x - f' n x) l (𝓝 0), + { rw metric.tendsto_nhds at h1 ⊢, + exact (λ ε hε, (h1 ε hε).curry.mono (λ n hn, hn.self_of_nhds)), }, + have := (tendsto_fst.comp (h2.prod_map tendsto_id)), + refine squeeze_zero_norm _ (tendsto_zero_iff_norm_tendsto_zero.mp this), + intros n, + simp_rw [norm_smul, norm_inv, is_R_or_C.norm_coe_norm], + by_cases hx : x = n.2, { simp [hx], }, + have hnx : 0 < ∥n.2 - x∥, + { rw norm_pos_iff, intros hx', exact hx (eq_of_sub_eq_zero hx').symm, }, + rw [inv_mul_le_iff hnx, mul_comm], + simp only [function.comp_app, prod_map], + rw norm_sub_rev, + exact (f' n.1 x - g' x).le_op_norm (n.2 - x), }, +end + +/-- `(d/dx) lim_{n → ∞} f n x = lim_{n → ∞} f' n x` when the `f' n` converge +_uniformly_ to their limit on an open set containing `x`. -/ +lemma has_fderiv_at_of_tendsto_uniformly_on + {s : set E} (hs : is_open s) + (hf' : tendsto_uniformly_on f' g' l s) + (hf : ∀ (n : ι), ∀ (x : E), x ∈ s → has_fderiv_at (f n) (f' n x) x) + (hfg : ∀ (x : E), x ∈ s → tendsto (λ n, f n x) l (𝓝 (g x))) : + ∀ (x : E), x ∈ s → has_fderiv_at g (g' x) x := +begin + intros x hx, + have hf : ∀ᶠ (n : ι × E) in (l ×ᶠ 𝓝 x), has_fderiv_at (f n.1) (f' n.1 n.2) n.2, + { exact eventually_prod_iff.mpr ⟨(λ y, true), (by simp), (λ y, y ∈ s), + eventually_mem_set.mpr (mem_nhds_iff.mpr ⟨s, rfl.subset, hs, hx⟩), + (λ n hn y hy, hf n y hy)⟩, }, + + have hfg : ∀ᶠ y in 𝓝 x, tendsto (λ n, f n y) l (𝓝 (g y)), + { exact eventually_iff.mpr (mem_nhds_iff.mpr ⟨s, set.subset_def.mpr hfg, hs, hx⟩), }, + + have hfg' := hf'.tendsto_uniformly_on_filter.mono_right (calc + 𝓝 x = 𝓝[s] x : ((hs.nhds_within_eq hx).symm) + ... ≤ 𝓟 s : (by simp only [nhds_within, inf_le_right])), + + exact has_fderiv_at_of_tendsto_uniformly_on_filter hfg' hf hfg, +end + +/-- `(d/dx) lim_{n → ∞} f n x = lim_{n → ∞} f' n x` when the `f' n` converge +_uniformly_ to their limit. -/ +lemma has_fderiv_at_of_tendsto_uniformly + (hf' : tendsto_uniformly f' g' l) + (hf : ∀ (n : ι), ∀ (x : E), has_fderiv_at (f n) (f' n x) x) + (hfg : ∀ (x : E), tendsto (λ n, f n x) l (𝓝 (g x))) : + ∀ (x : E), has_fderiv_at g (g' x) x := +begin + intros x, + have hf : ∀ (n : ι), ∀ (x : E), x ∈ set.univ → has_fderiv_at (f n) (f' n x) x, { simp [hf], }, + have hfg : ∀ (x : E), x ∈ set.univ → tendsto (λ n, f n x) l (𝓝 (g x)), { simp [hfg], }, + have hf' : tendsto_uniformly_on f' g' l set.univ, { rwa tendsto_uniformly_on_univ, }, + refine has_fderiv_at_of_tendsto_uniformly_on is_open_univ hf' hf hfg x (set.mem_univ x), +end + +end limits_of_derivatives + +section deriv + +/-! ### `deriv` versions of above theorems + +In this section, we provide `deriv` equivalents of the `fderiv` lemmas in the previous section. +The protected function `promote_deriv` provides the translation between derivatives and Fréchet +derivatives +-/ + +variables {ι : Type*} {l : filter ι} + {𝕜 : Type*} [is_R_or_C 𝕜] + {G : Type*} [normed_add_comm_group G] [normed_space 𝕜 G] + {f : ι → 𝕜 → G} {g : 𝕜 → G} {f' : ι → 𝕜 → G} {g' : 𝕜 → G} + {x : 𝕜} + +/-- If our derivatives converge uniformly, then the Fréchet derivatives converge uniformly -/ +lemma uniform_cauchy_seq_on_filter.one_smul_right {l' : filter 𝕜} + (hf' : uniform_cauchy_seq_on_filter f' l l') : + uniform_cauchy_seq_on_filter (λ n, λ z, (1 : 𝕜 →L[𝕜] 𝕜).smul_right (f' n z)) l l' := +begin + -- The tricky part of this proof is that operator norms are written in terms of `≤` whereas + -- metrics are written in terms of `<`. So we need to shrink `ε` utilizing the archimedean + -- property of `ℝ` + + rw [normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero, + metric.tendsto_uniformly_on_filter_iff] at hf' ⊢, + intros ε hε, + obtain ⟨q, hq, hq'⟩ := exists_between hε.lt, + apply (hf' q hq).mono, + intros n hn, + refine lt_of_le_of_lt _ hq', + simp only [dist_eq_norm, pi.zero_apply, zero_sub, norm_neg] at hn ⊢, + refine continuous_linear_map.op_norm_le_bound _ hq.le _, + intros z, + simp only [continuous_linear_map.coe_sub', pi.sub_apply, continuous_linear_map.smul_right_apply, + continuous_linear_map.one_apply], + rw [←smul_sub, norm_smul, mul_comm], + exact mul_le_mul hn.le rfl.le (norm_nonneg _) hq.le, +end + +variables [ne_bot l] + +lemma uniform_cauchy_seq_on_filter_of_tendsto_uniformly_on_filter_deriv + (hf' : uniform_cauchy_seq_on_filter f' l (𝓝 x)) + (hf : ∀ᶠ (n : ι × 𝕜) in (l ×ᶠ 𝓝 x), has_deriv_at (f n.1) (f' n.1 n.2) n.2) + (hfg : tendsto (λ n, f n x) l (𝓝 (g x))) : + uniform_cauchy_seq_on_filter f l (𝓝 x) := +begin + simp_rw has_deriv_at_iff_has_fderiv_at at hf, + exact uniform_cauchy_seq_on_filter_of_tendsto_uniformly_on_filter_fderiv + hf'.one_smul_right hf hfg, +end + +lemma uniform_cauchy_seq_on_ball_of_tendsto_uniformly_on_ball_deriv + {r : ℝ} (hr : 0 < r) + (hf' : uniform_cauchy_seq_on f' l (metric.ball x r)) + (hf : ∀ n : ι, ∀ y : 𝕜, y ∈ metric.ball x r → has_deriv_at (f n) (f' n y) y) + (hfg : tendsto (λ n, f n x) l (𝓝 (g x))) : + uniform_cauchy_seq_on f l (metric.ball x r) := +begin + simp_rw has_deriv_at_iff_has_fderiv_at at hf, + rw uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter at hf', + have hf' : uniform_cauchy_seq_on (λ n, λ z, (1 : 𝕜 →L[𝕜] 𝕜).smul_right (f' n z)) l + (metric.ball x r), + { rw uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter, + exact hf'.one_smul_right, }, + exact uniform_cauchy_seq_on_ball_of_tendsto_uniformly_on_ball_fderiv hr hf' hf hfg, +end + +lemma has_deriv_at_of_tendsto_uniformly_on_filter + (hf' : tendsto_uniformly_on_filter f' g' l (𝓝 x)) + (hf : ∀ᶠ (n : ι × 𝕜) in (l ×ᶠ 𝓝 x), has_deriv_at (f n.1) (f' n.1 n.2) n.2) + (hfg : ∀ᶠ y in 𝓝 x, tendsto (λ n, f n y) l (𝓝 (g y))) : + has_deriv_at g (g' x) x := +begin + -- The first part of the proof rewrites `hf` and the goal to be functions so that Lean + -- can recognize them when we apply `has_fderiv_at_of_tendsto_uniformly_on_filter` + let F' := (λ n, λ z, (1 : 𝕜 →L[𝕜] 𝕜).smul_right (f' n z)), + let G' := λ z, (1 : 𝕜 →L[𝕜] 𝕜).smul_right (g' z), + simp_rw has_deriv_at_iff_has_fderiv_at at hf ⊢, + + -- Now we need to rewrite hf' in terms of continuous_linear_maps. The tricky part is that + -- operator norms are written in terms of `≤` whereas metrics are written in terms of `<`. So we + -- need to shrink `ε` utilizing the archimedean property of `ℝ` + have hf' : tendsto_uniformly_on_filter F' G' l (𝓝 x), + { rw metric.tendsto_uniformly_on_filter_iff at hf' ⊢, + intros ε hε, + obtain ⟨q, hq, hq'⟩ := exists_between hε.lt, + apply (hf' q hq).mono, + intros n hn, + refine lt_of_le_of_lt _ hq', + simp only [F', G', dist_eq_norm] at hn ⊢, + refine continuous_linear_map.op_norm_le_bound _ hq.le _, + intros z, + simp only [continuous_linear_map.coe_sub', pi.sub_apply, continuous_linear_map.smul_right_apply, + continuous_linear_map.one_apply], + rw [←smul_sub, norm_smul, mul_comm], + exact mul_le_mul hn.le rfl.le (norm_nonneg _) hq.le, }, + exact has_fderiv_at_of_tendsto_uniformly_on_filter hf' hf hfg, +end + +lemma has_deriv_at_of_tendsto_uniformly_on + {s : set 𝕜} (hs : is_open s) + (hf' : tendsto_uniformly_on f' g' l s) + (hf : ∀ (n : ι), ∀ (x : 𝕜), x ∈ s → has_deriv_at (f n) (f' n x) x) + (hfg : ∀ (x : 𝕜), x ∈ s → tendsto (λ n, f n x) l (𝓝 (g x))) : + ∀ (x : 𝕜), x ∈ s → has_deriv_at g (g' x) x := +begin + intros x hx, + have hsx : s ∈ 𝓝 x, { exact mem_nhds_iff.mpr ⟨s, rfl.subset, hs, hx⟩, }, + rw tendsto_uniformly_on_iff_tendsto_uniformly_on_filter at hf', + have hf' := hf'.mono_right (le_principal_iff.mpr hsx), + have hfg : ∀ᶠ y in 𝓝 x, tendsto (λ n, f n y) l (𝓝 (g y)), + { exact eventually_iff_exists_mem.mpr ⟨s, hsx, hfg⟩, }, + have hf : ∀ᶠ (n : ι × 𝕜) in (l ×ᶠ 𝓝 x), has_deriv_at (f n.1) (f' n.1 n.2) n.2, + { rw eventually_prod_iff, + refine ⟨(λ y, true), by simp, (λ y, y ∈ s), _, (λ n hn y hy, hf n y hy)⟩, + exact eventually_mem_set.mpr hsx, }, + exact has_deriv_at_of_tendsto_uniformly_on_filter hf' hf hfg, +end + +lemma has_deriv_at_of_tendsto_uniformly + (hf' : tendsto_uniformly f' g' l) + (hf : ∀ (n : ι), ∀ (x : 𝕜), has_deriv_at (f n) (f' n x) x) + (hfg : ∀ (x : 𝕜), tendsto (λ n, f n x) l (𝓝 (g x))) : + ∀ (x : 𝕜), has_deriv_at g (g' x) x := +begin + intros x, + have hf : ∀ (n : ι), ∀ (x : 𝕜), x ∈ set.univ → has_deriv_at (f n) (f' n x) x, { simp [hf], }, + have hfg : ∀ (x : 𝕜), x ∈ set.univ → tendsto (λ n, f n x) l (𝓝 (g x)), { simp [hfg], }, + have hf' : tendsto_uniformly_on f' g' l set.univ, { rwa tendsto_uniformly_on_univ, }, + exact has_deriv_at_of_tendsto_uniformly_on is_open_univ hf' hf hfg x (set.mem_univ x), +end + +end deriv diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 37e31726e714b..89520e23fda3d 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -479,6 +479,42 @@ begin simp [dist_eq_norm] end +section tendsto_uniformly +/-- The results in this section do not require `E'` to any particular structure -/ +variables {E' : Type*} {f : ι → E' → G} {s : set E'} {l : filter ι} + +lemma normed_add_comm_group.tendsto_uniformly_on_zero : + tendsto_uniformly_on f 0 l s ↔ ∀ ε > 0, ∀ᶠ (N : ι) in l, ∀ x : E', x ∈ s → ∥f N x∥ < ε := +by simp_rw [tendsto_uniformly_on_iff, pi.zero_apply, dist_zero_left] + +lemma normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero + {l' : filter E'} : uniform_cauchy_seq_on_filter f l l' ↔ + tendsto_uniformly_on_filter (λ n : ι × ι, λ z : E', f n.fst z - f n.snd z) 0 (l.prod l) l' := +begin + split, + { intros hf u hu, + obtain ⟨ε, hε, H⟩ := uniformity_basis_dist.mem_uniformity_iff.mp hu, + refine (hf {p : G × G | dist p.fst p.snd < ε} $ dist_mem_uniformity hε).mono (λ x hx, + H 0 (f x.fst.fst x.snd - f x.fst.snd x.snd) _), + simpa [dist_eq_norm, norm_sub_rev] using hx, }, + { intros hf u hu, + obtain ⟨ε, hε, H⟩ := uniformity_basis_dist.mem_uniformity_iff.mp hu, + refine (hf {p : G × G | dist p.fst p.snd < ε} $ dist_mem_uniformity hε).mono (λ x hx, + H (f x.fst.fst x.snd) (f x.fst.snd x.snd) _), + simpa [dist_eq_norm, norm_sub_rev] using hx, }, +end + +lemma normed_add_comm_group.uniform_cauchy_seq_on_iff_tendsto_uniformly_on_zero : + uniform_cauchy_seq_on f l s ↔ + tendsto_uniformly_on (λ n : ι × ι, λ z : E', f n.fst z - f n.snd z) 0 (l.prod l) s := +begin + rw tendsto_uniformly_on_iff_tendsto_uniformly_on_filter, + rw uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter, + exact normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero, +end + +end tendsto_uniformly + open finset /-- A homomorphism `f` of seminormed groups is Lipschitz, if there exists a constant `C` such that diff --git a/src/order/filter/curry.lean b/src/order/filter/curry.lean new file mode 100644 index 0000000000000..e5aeb997aca00 --- /dev/null +++ b/src/order/filter/curry.lean @@ -0,0 +1,89 @@ +/- +Copyright (c) 2022 Kevin H. Wilson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kevin H. Wilson +-/ +import order.filter.prod + +/-! +# Curried Filters + +This file provides an operation (`filter.curry`) on filters which provides the equivalence +`∀ᶠ a in l, ∀ᶠ b in l', p (a, b) ↔ ∀ᶠ c in (l.curry l'), p c` (see `filter.eventually_curry_iff`). + +To understand when this operation might arise, it is helpful to think of `∀ᶠ` as a combination of +the quantifiers `∃ ∀`. For instance, `∀ᶠ n in at_top, p n ↔ ∃ N, ∀ n ≥ N, p n`. A curried filter +yields the quantifier order `∃ ∀ ∃ ∀`. For instance, +`∀ᶠ n in at_top.curry at_top, p n ↔ ∃ M, ∀ m ≥ M, ∃ N, ∀ n ≥ N, p (m, n)`. + +This is different from a product filter, which instead yields a quantifier order `∃ ∃ ∀ ∀`. For +instance, `∀ᶠ n in at_top ×ᶠ at_top, p n ↔ ∃ M, ∃ N, ∀ m ≥ M, ∀ n ≥ N, p (m, n)`. This makes it +clear that if something eventually occurs on the product filter, it eventually occurs on the curried +filter (see `filter.curry_le_prod` and `filter.eventually.curry`), but the converse is not true. + +Another way to think about the curried versus the product filter is that tending to some limit on +the product filter is a version of uniform convergence (see `tendsto_prod_filter_iff`) whereas +tending to some limit on a curried filter is just iterated limits (see `tendsto.curry`). + +## Main definitions + +* `filter.curry`: A binary operation on filters which represents iterated limits + +## Main statements + +* `filter.eventually_curry_iff`: An alternative definition of a curried filter +* `filter.curry_le_prod`: Something that is eventually true on the a product filter is eventually + true on the curried filter + +## Tags + +uniform convergence, curried filters, product filters +-/ + +namespace filter + +variables {α β γ : Type*} + +/-- This filter is characterized by `filter.eventually_curry_iff`: +`(∀ᶠ (x : α × β) in f.curry g, p x) ↔ ∀ᶠ (x : α) in f, ∀ᶠ (y : β) in g, p (x, y)`. Useful +in adding quantifiers to the middle of `tendsto`s. See +`has_fderiv_at_of_tendsto_uniformly_on_filter`. -/ +def curry (f : filter α) (g : filter β) : filter (α × β) := +{ sets := { s | ∀ᶠ (a : α) in f, ∀ᶠ (b : β) in g, (a, b) ∈ s }, + univ_sets := (by simp only [set.mem_set_of_eq, set.mem_univ, eventually_true]), + sets_of_superset := begin + intros x y hx hxy, + simp only [set.mem_set_of_eq] at hx ⊢, + exact hx.mono (λ a ha, ha.mono(λ b hb, set.mem_of_subset_of_mem hxy hb)), + end, + inter_sets := begin + intros x y hx hy, + simp only [set.mem_set_of_eq, set.mem_inter_eq] at hx hy ⊢, + exact (hx.and hy).mono (λ a ha, (ha.1.and ha.2).mono (λ b hb, hb)), + end, } + +lemma eventually_curry_iff {f : filter α} {g : filter β} {p : α × β → Prop} : + (∀ᶠ (x : α × β) in f.curry g, p x) ↔ ∀ᶠ (x : α) in f, ∀ᶠ (y : β) in g, p (x, y) := +iff.rfl + +lemma curry_le_prod {f : filter α} {g : filter β} : + f.curry g ≤ f.prod g := +begin + intros u hu, + rw ←eventually_mem_set at hu ⊢, + rw eventually_curry_iff, + exact hu.curry, +end + +lemma tendsto.curry {f : α → β → γ} {la : filter α} {lb : filter β} {lc : filter γ} : + (∀ᶠ a in la, tendsto (λ b : β, f a b) lb lc) → tendsto ↿f (la.curry lb) lc := +begin + intros h, + rw tendsto_def, + simp only [curry, filter.mem_mk, set.mem_set_of_eq, set.mem_preimage], + simp_rw tendsto_def at h, + refine (λ s hs, h.mono (λ a ha, eventually_iff.mpr _)), + simpa [function.has_uncurry.uncurry, set.preimage] using ha s hs, +end + +end filter diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 3f567456db3aa..8dcbdf059ec13 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -800,6 +800,17 @@ begin exact hs _ (dist_mem_uniformity ε_pos), end +/-- Expressing uniform convergence using `dist` -/ +lemma tendsto_uniformly_on_filter_iff {ι : Type*} + {F : ι → β → α} {f : β → α} {p : filter ι} {p' : filter β} : + tendsto_uniformly_on_filter F f p p' ↔ + ∀ ε > 0, ∀ᶠ (n : ι × β) in (p ×ᶠ p'), dist (f n.snd) (F n.fst n.snd) < ε := +begin + refine ⟨λ H ε hε, H _ (dist_mem_uniformity hε), λ H u hu, _⟩, + rcases mem_uniformity_dist.1 hu with ⟨ε, εpos, hε⟩, + refine (H ε εpos).mono (λ n hn, hε hn), +end + /-- Expressing locally uniform convergence on a set using `dist`. -/ lemma tendsto_locally_uniformly_on_iff {ι : Type*} [topological_space β] {F : ι → β → α} {f : β → α} {p : filter ι} {s : set β} : From cd7f0626a0b04be1dda223a26123313514a55fb4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 8 Sep 2022 17:38:33 +0000 Subject: [PATCH 207/828] =?UTF-8?q?feat(combinatorics/simple=5Fgraph/trian?= =?UTF-8?q?gle/basic):=20add=20=CE=B5-triangle-free=20far=20graph=20predic?= =?UTF-8?q?ate=20(#12988)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the property of being triangle-free far. This comes up in the triangle counting and triangle removal lemmas. Co-authored-by: Bhavik Mehta Co-authored-by: Kyle Miller --- src/combinatorics/simple_graph/basic.lean | 61 +++++++++++++- src/combinatorics/simple_graph/clique.lean | 4 +- .../simple_graph/triangle/basic.lean | 79 +++++++++++++++++++ 3 files changed, 140 insertions(+), 4 deletions(-) create mode 100644 src/combinatorics/simple_graph/triangle/basic.lean diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index e51ba9debd13d..f3c745efcdc02 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -133,9 +133,8 @@ def complete_bipartite_graph (V W : Type*) : simple_graph (V ⊕ W) := end } namespace simple_graph - -variables {V : Type u} {W : Type v} {X : Type w} (G : simple_graph V) (G' : simple_graph W) - {a b c u v w : V} {e : sym2 V} +variables {𝕜 : Type*} {V : Type u} {W : Type v} {X : Type w} (G : simple_graph V) + (G' : simple_graph W) {a b c u v w : V} {e : sym2 V} @[simp] protected lemma irrefl {v : V} : ¬G.adj v v := G.loopless v @@ -282,6 +281,9 @@ def edge_set : set (sym2 V) := sym2.from_rel G.symm @[simp] lemma mem_edge_set : ⟦(v, w)⟧ ∈ G.edge_set ↔ G.adj v w := iff.rfl +lemma edge_set_mono {G G' : simple_graph V} (h : G ≤ G') : G.edge_set ⊆ G'.edge_set := +λ e, sym2.ind (λ v w, @h v w) e + /-- Two vertices are adjacent iff there is an edge between them. The condition `v ≠ w` ensures they are different endpoints of the edge, @@ -463,6 +465,14 @@ set.to_finset G.edge_set e ∈ G.edge_finset ↔ e ∈ G.edge_set := set.mem_to_finset +@[simp, norm_cast] lemma coe_edge_finset [fintype G.edge_set] : + (G.edge_finset : set (sym2 V)) = G.edge_set := +set.coe_to_finset _ + +lemma edge_finset_mono {G G' : simple_graph V} [fintype G.edge_set] [fintype G'.edge_set] : + G ≤ G' → G.edge_finset ⊆ G'.edge_finset := +by { simp_rw [←coe_subset, coe_edge_finset], exact edge_set_mono } + lemma edge_finset_card [fintype G.edge_set] : G.edge_finset.card = fintype.card G.edge_set := set.to_finset_card _ @@ -638,6 +648,51 @@ lemma delete_edges_eq_inter_edge_set (s : set (sym2 V)) : G.delete_edges s = G.delete_edges (s ∩ G.edge_set) := by { ext, simp [imp_false] { contextual := tt } } +lemma delete_edges_sdiff_eq_of_le {H : simple_graph V} (h : H ≤ G) : + G.delete_edges (G.edge_set \ H.edge_set) = H := +by { ext v w, split; simp [@h v w] { contextual := tt } } + +lemma edge_set_delete_edges (s : set (sym2 V)) : + (G.delete_edges s).edge_set = G.edge_set \ s := +by { ext e, refine sym2.ind _ e, simp } + +lemma edge_finset_delete_edges [fintype V] [decidable_eq V] [decidable_rel G.adj] + (s : finset (sym2 V)) [decidable_rel (G.delete_edges s).adj] : + (G.delete_edges s).edge_finset = G.edge_finset \ s := +by { ext e, simp [edge_set_delete_edges] } + +section delete_far +variables (G) [ordered_ring 𝕜] [fintype V] [decidable_eq V] [decidable_rel G.adj] + {p : simple_graph V → Prop} {r r₁ r₂ : 𝕜} + +/-- A graph is `r`-*delete-far* from a property `p` if we must delete at least `r` edges from it to +get a graph with the property `p`. -/ +def delete_far (p : simple_graph V → Prop) (r : 𝕜) : Prop := +∀ ⦃s⦄, s ⊆ G.edge_finset → p (G.delete_edges s) → r ≤ s.card + +open_locale classical + +variables {G} + +lemma delete_far_iff : + G.delete_far p r ↔ ∀ ⦃H⦄, H ≤ G → p H → r ≤ G.edge_finset.card - H.edge_finset.card := +begin + refine ⟨λ h H hHG hH, _, λ h s hs hG, _⟩, + { have := h (sdiff_subset G.edge_finset H.edge_finset), + simp only [delete_edges_sdiff_eq_of_le _ hHG, edge_finset_mono hHG, card_sdiff, + card_le_of_subset, coe_sdiff, coe_edge_finset, nat.cast_sub] at this, + exact this hH }, + { simpa [card_sdiff hs, edge_finset_delete_edges, -set.to_finset_card, nat.cast_sub, + card_le_of_subset hs] using h (G.delete_edges_le s) hG } +end + +alias delete_far_iff ↔ delete_far.le_card_sub_card _ + +lemma delete_far.mono (h : G.delete_far p r₂) (hr : r₁ ≤ r₂) : G.delete_far p r₁ := +λ s hs hG, hr.trans $ h hs hG + +end delete_far + /-! ## Map and comap -/ /-- Given an injective function, there is an covariant induced map on graphs by pushing forward diff --git a/src/combinatorics/simple_graph/clique.lean b/src/combinatorics/simple_graph/clique.lean index d16e1a82bb090..09741ee7e4ebf 100644 --- a/src/combinatorics/simple_graph/clique.lean +++ b/src/combinatorics/simple_graph/clique.lean @@ -135,7 +135,9 @@ variables {m n : ℕ} /-- `G.clique_free n` means that `G` has no `n`-cliques. -/ def clique_free (n : ℕ) : Prop := ∀ t, ¬ G.is_n_clique n t -variables {G H} +variables {G H} {s : finset α} + +lemma is_n_clique.not_clique_free (hG : G.is_n_clique n s) : ¬ G.clique_free n := λ h, h _ hG lemma not_clique_free_of_top_embedding {n : ℕ} (f : (⊤ : simple_graph (fin n)) ↪g G) : ¬ G.clique_free n := diff --git a/src/combinatorics/simple_graph/triangle/basic.lean b/src/combinatorics/simple_graph/triangle/basic.lean new file mode 100644 index 0000000000000..ed1a2512fedaf --- /dev/null +++ b/src/combinatorics/simple_graph/triangle/basic.lean @@ -0,0 +1,79 @@ +/- +Copyright (c) 2022 Yaël Dillies, Bhavik Mehta. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Bhavik Mehta +-/ +import combinatorics.simple_graph.clique + +/-! +# Triangles in graphs + +A *triangle* in a simple graph is a `3`-clique, namely a set of three vertices that are +pairwise adjacent. + +This module defines and proves properties about triangles in simple graphs. + +## Main declarations + +* `simple_graph.far_from_triangle_free`: Predicate for a graph to have enough triangles that, to + remove all of them, one must one must remove a lot of edges. This is the crux of the Triangle + Removal lemma. + +## TODO + +* Generalise `far_from_triangle_free` to other graphs, to state and prove the Graph Removal Lemma. +* Find a better name for `far_from_triangle_free`. Added 4/26/2022. Remove this TODO if it gets old. +-/ + +open finset fintype nat +open_locale classical + +namespace simple_graph +variables {α 𝕜 : Type*} [fintype α] [linear_ordered_field 𝕜] {G H : simple_graph α} {ε δ : 𝕜} + {n : ℕ} {s : finset α} + +/-- A simple graph is *`ε`-triangle-free far* if one must remove at least `ε * (card α)^2` edges to +make it triangle-free. -/ +def far_from_triangle_free (G : simple_graph α) (ε : 𝕜) : Prop := +G.delete_far (λ H, H.clique_free 3) $ ε * (card α^2 : ℕ) + +lemma far_from_triangle_free_iff : + G.far_from_triangle_free ε ↔ + ∀ ⦃H⦄, H ≤ G → H.clique_free 3 → ε * (card α^2 : ℕ) ≤ G.edge_finset.card - H.edge_finset.card := +delete_far_iff + +alias far_from_triangle_free_iff ↔ far_from_triangle_free.le_card_sub_card _ + +lemma far_from_triangle_free.mono (hε : G.far_from_triangle_free ε) (h : δ ≤ ε) : + G.far_from_triangle_free δ := +hε.mono $ mul_le_mul_of_nonneg_right h $ cast_nonneg _ + +lemma far_from_triangle_free.clique_finset_nonempty' (hH : H ≤ G) (hG : G.far_from_triangle_free ε) + (hcard : (G.edge_finset.card - H.edge_finset.card : 𝕜) < ε * (card α ^ 2 : ℕ)) : + (H.clique_finset 3).nonempty := +nonempty_of_ne_empty $ H.clique_finset_eq_empty_iff.not.2 $ λ hH', + (hG.le_card_sub_card hH hH').not_lt hcard + +variables [nonempty α] + +lemma far_from_triangle_free.nonpos (h₀ : G.far_from_triangle_free ε) (h₁ : G.clique_free 3) : + ε ≤ 0 := +begin + have := h₀ (empty_subset _), + rw [coe_empty, finset.card_empty, cast_zero, delete_edges_empty_eq] at this, + exact nonpos_of_mul_nonpos_left (this h₁) (cast_pos.2 $ sq_pos_of_pos fintype.card_pos), +end + +lemma clique_free.not_far_from_triangle_free (hG : G.clique_free 3) (hε : 0 < ε) : + ¬ G.far_from_triangle_free ε := +λ h, (h.nonpos hG).not_lt hε + +lemma far_from_triangle_free.not_clique_free (hG : G.far_from_triangle_free ε) (hε : 0 < ε) : + ¬ G.clique_free 3 := +λ h, (hG.nonpos h).not_lt hε + +lemma far_from_triangle_free.clique_finset_nonempty (hG : G.far_from_triangle_free ε) (hε : 0 < ε) : + (G.clique_finset 3).nonempty := +nonempty_of_ne_empty $ G.clique_finset_eq_empty_iff.not.2 $ hG.not_clique_free hε + +end simple_graph From 00ab77614e085c9ef49479babba1a7d826d3232e Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 8 Sep 2022 20:32:26 +0000 Subject: [PATCH 208/828] feat(ring_theory/dedekind_domain/integer_unit): define S-integers and S-units (#15646) --- .../dedekind_domain/S_integer.lean | 102 ++++++++++++++++++ src/ring_theory/subring/basic.lean | 7 ++ .../valuation/valuation_subring.lean | 12 ++- 3 files changed, 119 insertions(+), 2 deletions(-) create mode 100644 src/ring_theory/dedekind_domain/S_integer.lean diff --git a/src/ring_theory/dedekind_domain/S_integer.lean b/src/ring_theory/dedekind_domain/S_integer.lean new file mode 100644 index 0000000000000..ec1944af6e824 --- /dev/null +++ b/src/ring_theory/dedekind_domain/S_integer.lean @@ -0,0 +1,102 @@ +/- +Copyright (c) 2022 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata +-/ + +import ring_theory.dedekind_domain.adic_valuation + +/-! +# `S`-integers and `S`-units of fraction fields of Dedekind domains + +Let `K` be the field of fractions of a Dedekind domain `R`, and let `S` be a set of prime ideals in +the height one spectrum of `R`. An `S`-integer of `K` is defined to have `v`-adic valuation at most +one for all primes ideals `v` away from `S`, whereas an `S`-unit of `Kˣ` is defined to have `v`-adic +valuation exactly one for all prime ideals `v` away from `S`. + +This file defines the subalgebra of `S`-integers of `K` and the subgroup of `S`-units of `Kˣ`, where +`K` can be specialised to the case of a number field or a function field separately. + +## Main definitions + + * `set.integer`: `S`-integers. + * `set.unit`: `S`-units. + * TODO: localised notation for `S`-integers. + +## Main statements + + * `set.unit_equiv_units_integer`: `S`-units are units of `S`-integers. + * TODO: proof that `S`-units is the kernel of a map to a product. + * TODO: proof that `∅`-integers is the usual ring of integers. + * TODO: finite generation of `S`-units and Dirichlet's `S`-unit theorem. + +## References + + * [D Marcus, *Number Fields*][marcus1977number] + * [J W S Cassels, A Frölich, *Algebraic Number Theory*][cassels1967algebraic] + * [J Neukirch, *Algebraic Number Theory*][Neukirch1992] + +## Tags + +S integer, S-integer, S unit, S-unit +-/ + +namespace set + +noncomputable theory + +open is_dedekind_domain + +open_locale non_zero_divisors + +universes u v + +variables {R : Type u} [comm_ring R] [is_domain R] [is_dedekind_domain R] + (S : set $ height_one_spectrum R) (K : Type v) [field K] [algebra R K] [is_fraction_ring R K] + +/-! ## `S`-integers -/ + +/-- The `R`-subalgebra of `S`-integers of `K`. -/ +@[simps] def integer : subalgebra R K := +{ algebra_map_mem' := λ x v _, v.valuation_le_one x, + .. (⨅ v ∉ S, (v : height_one_spectrum R).valuation.valuation_subring.to_subring).copy + {x : K | ∀ v ∉ S, (v : height_one_spectrum R).valuation x ≤ 1} $ set.ext $ λ _, + by simpa only [set_like.mem_coe, subring.mem_infi] } + +lemma integer_eq : + (S.integer K).to_subring + = ⨅ v ∉ S, (v : height_one_spectrum R).valuation.valuation_subring.to_subring := +set_like.ext' $ by simpa only [integer, subring.copy_eq] + +lemma integer_valuation_le_one (x : S.integer K) {v : height_one_spectrum R} (hv : v ∉ S) : + v.valuation (x : K) ≤ 1 := +x.property v hv + +/-! ## `S`-units -/ + +/-- The subgroup of `S`-units of `Kˣ`. -/ +@[simps] def unit : subgroup Kˣ := +(⨅ v ∉ S, (v : height_one_spectrum R).valuation.valuation_subring.unit_group).copy + {x : Kˣ | ∀ v ∉ S, (v : height_one_spectrum R).valuation (x : K) = 1} $ set.ext $ λ _, + by simpa only [set_like.mem_coe, subgroup.mem_infi, valuation.mem_unit_group_iff] + +lemma unit_eq : + S.unit K = ⨅ v ∉ S, (v : height_one_spectrum R).valuation.valuation_subring.unit_group := +subgroup.copy_eq _ _ _ + +lemma unit_valuation_eq_one (x : S.unit K) {v : height_one_spectrum R} (hv : v ∉ S) : + v.valuation (x : K) = 1 := +x.property v hv + +/-- The group of `S`-units is the group of units of the ring of `S`-integers. -/ +@[simps] def unit_equiv_units_integer : S.unit K ≃* (S.integer K)ˣ := +{ to_fun := λ x, ⟨⟨x, λ v hv, (x.property v hv).le⟩, ⟨↑x⁻¹, λ v hv, ((x⁻¹).property v hv).le⟩, + subtype.ext x.val.val_inv, subtype.ext x.val.inv_val⟩, + inv_fun := λ x, ⟨units.mk0 x $ λ hx, x.ne_zero ((subring.coe_eq_zero_iff _).mp hx), + λ v hv, eq_one_of_one_le_mul_left (x.val.property v hv) (x.inv.property v hv) $ eq.ge $ + by { rw [← map_mul], convert v.valuation.map_one, exact subtype.mk_eq_mk.mp x.val_inv }⟩, + left_inv := λ _, by { ext, refl }, + right_inv := λ _, by { ext, refl }, + map_mul' := λ _ _, by { ext, refl } } + +end set diff --git a/src/ring_theory/subring/basic.lean b/src/ring_theory/subring/basic.lean index 58fb30a99bd6c..a02054040b424 100644 --- a/src/ring_theory/subring/basic.lean +++ b/src/ring_theory/subring/basic.lean @@ -560,6 +560,13 @@ instance : has_Inf (subring R) := lemma mem_Inf {S : set (subring R)} {x : R} : x ∈ Inf S ↔ ∀ p ∈ S, x ∈ p := set.mem_Inter₂ +@[simp, norm_cast] lemma coe_infi {ι : Sort*} {S : ι → subring R} : + (↑(⨅ i, S i) : set R) = ⋂ i, S i := +by simp only [infi, coe_Inf, set.bInter_range] + +lemma mem_infi {ι : Sort*} {S : ι → subring R} {x : R} : (x ∈ ⨅ i, S i) ↔ ∀ i, x ∈ S i := +by simp only [infi, mem_Inf, set.forall_range_iff] + @[simp] lemma Inf_to_submonoid (s : set (subring R)) : (Inf s).to_submonoid = ⨅ t ∈ s, subring.to_submonoid t := mk'_to_submonoid _ _ diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index dd880990215ec..3e05886ec1cea 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -371,7 +371,7 @@ section unit_group def unit_group : subgroup Kˣ := (A.valuation.to_monoid_with_zero_hom.to_monoid_hom.comp (units.coe_hom K)).ker -lemma mem_unit_group_iff (x : Kˣ) : x ∈ A.unit_group ↔ A.valuation x = 1 := iff.rfl +@[simp] lemma mem_unit_group_iff (x : Kˣ) : x ∈ A.unit_group ↔ A.valuation x = 1 := iff.rfl /-- For a valuation subring `A`, `A.unit_group` agrees with the units of `A`. -/ def unit_group_mul_equiv : A.unit_group ≃* Aˣ := @@ -393,7 +393,6 @@ lemma coe_unit_group_mul_equiv_apply (a : A.unit_group) : lemma coe_unit_group_mul_equiv_symm_apply (a : Aˣ) : (A.unit_group_mul_equiv.symm a : K) = a := rfl - lemma unit_group_le_unit_group {A B : valuation_subring K} : A.unit_group ≤ B.unit_group ↔ A ≤ B := begin @@ -718,3 +717,12 @@ set.subset_set_smul_iff end pointwise_actions end valuation_subring + +namespace valuation + +variables {Γ : Type*} [linear_ordered_comm_group_with_zero Γ] (v : valuation K Γ) (x : Kˣ) + +@[simp] lemma mem_unit_group_iff : x ∈ v.valuation_subring.unit_group ↔ v x = 1 := +(valuation.is_equiv_iff_val_eq_one _ _).mp (valuation.is_equiv_valuation_valuation_subring _).symm + +end valuation From e0bdbbe1294964d621838c1623211c130eaf6a46 Mon Sep 17 00:00:00 2001 From: jakelev Date: Thu, 8 Sep 2022 20:32:28 +0000 Subject: [PATCH 209/828] feat(combinatorics/young_diagram): add transposes, rows and columns of Young diagrams (#16120) Add transposes, rows (and row lengths), columns (and column lengths) of Young diagrams. --- src/combinatorics/young_diagram.lean | 151 +++++++++++++++++++++++++++ 1 file changed, 151 insertions(+) diff --git a/src/combinatorics/young_diagram.lean b/src/combinatorics/young_diagram.lean index 7886783d1f48a..2d80a2b00291e 100644 --- a/src/combinatorics/young_diagram.lean +++ b/src/combinatorics/young_diagram.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jake Levinson -/ import order.upper_lower +import data.finset.preimage /-! # Young diagrams @@ -27,6 +28,8 @@ to say that `(i, j)` (in matrix coordinates) is in the Young diagram `μ`. - `young_diagram.card` : the number of cells in a Young diagram (its *cardinality*) - `young_diagram.distrib_lattice` : a distributive lattice instance for Young diagrams ordered by containment, with `(⊥ : young_diagram)` the empty diagram. +- `young_diagram.row` and `young_diagram.row_len`: rows of a Young diagram and their lengths +- `young_diagram.col` and `young_diagram.col_len`: columns of a Young diagram and their lengths ## Notation @@ -63,6 +66,12 @@ instance : set_like young_diagram (ℕ × ℕ) := @[simp] lemma mem_cells {μ : young_diagram} (c : ℕ × ℕ) : c ∈ μ.cells ↔ c ∈ μ := iff.rfl +@[simp] lemma mem_mk (c : ℕ × ℕ) (cells) (is_lower_set) : + c ∈ young_diagram.mk cells is_lower_set ↔ c ∈ cells := iff.rfl + +instance decidable_mem (μ : young_diagram) : decidable_pred (∈ μ) := +show decidable_pred (∈ μ.cells), by apply_instance + /-- In "English notation", a Young diagram is drawn so that (i1, j1) ≤ (i2, j2) means (i1, j1) is weakly up-and-left of (i2, j2). -/ lemma up_left_mem (μ : young_diagram) {i1 i2 j1 j2 : ℕ} @@ -123,4 +132,146 @@ end distrib_lattice /-- Cardinality of a Young diagram -/ @[reducible] protected def card (μ : young_diagram) : ℕ := μ.cells.card +section transpose + +/-- The `transpose` of a Young diagram is obtained by swapping i's with j's. -/ +def transpose (μ : young_diagram) : young_diagram := +{ cells := (equiv.prod_comm _ _).finset_congr μ.cells, + is_lower_set := λ _ _ h, begin + simp only [finset.mem_coe, equiv.finset_congr_apply, finset.mem_map_equiv], + intro hcell, + apply μ.is_lower_set _ hcell, + simp [h], + end } + +@[simp] lemma mem_transpose {μ : young_diagram} {c : ℕ × ℕ} : c ∈ μ.transpose ↔ c.swap ∈ μ := +by simp [transpose] + +@[simp] lemma transpose_transpose (μ : young_diagram) : μ.transpose.transpose = μ := +by { ext, simp } + +lemma transpose_eq_iff_eq_transpose {μ ν : young_diagram} : + μ.transpose = ν ↔ μ = ν.transpose := +by { split; { rintro rfl, simp } } + +@[simp] lemma transpose_eq_iff {μ ν : young_diagram} : + μ.transpose = ν.transpose ↔ μ = ν := +by { rw transpose_eq_iff_eq_transpose, simp } + +-- This is effectively both directions of `transpose_le_iff` below. +protected lemma le_of_transpose_le {μ ν : young_diagram} (h_le : μ.transpose ≤ ν) : + μ ≤ ν.transpose := +λ c hc, by { simp only [mem_transpose], apply h_le, simpa } + +@[simp] lemma transpose_le_iff {μ ν : young_diagram} : μ.transpose ≤ ν.transpose ↔ μ ≤ ν := +⟨ λ h, by { convert young_diagram.le_of_transpose_le h, simp }, + λ h, by { convert @young_diagram.le_of_transpose_le _ _ _, simpa } ⟩ + +@[mono] +protected lemma transpose_mono {μ ν : young_diagram} (h_le : μ ≤ ν) : μ.transpose ≤ ν.transpose := +transpose_le_iff.mpr h_le + +/-- Transposing Young diagrams is an `order_iso`. -/ +@[simps] def transpose_order_iso : young_diagram ≃o young_diagram := +⟨⟨transpose, transpose, λ _, by simp, λ _, by simp⟩, by simp⟩ + +end transpose + +section rows +/-! ### Rows and row lengths of Young diagrams. + +This section defines `μ.row` and `μ.row_len`, with the following API: + 1. `(i, j) ∈ μ ↔ j < μ.row_len i` + 2. `μ.row i = {i} ×ˢ (finset.range (μ.row_len i))` + 3. `μ.row_len i = (μ.row i).card` + 4. `∀ {i1 i2}, i1 ≤ i2 → μ.row_len i2 ≤ μ.row_len i1` + +Note: #3 is not convenient for defining `μ.row_len`; instead, `μ.row_len` is defined +as the smallest `j` such that `(i, j) ∉ μ`. -/ + +/-- The `i`-th row of a Young diagram consists of the cells whose first coordinate is `i`. -/ +def row (μ : young_diagram) (i : ℕ) : finset (ℕ × ℕ) := μ.cells.filter (λ c, c.fst = i) + +lemma mem_row_iff {μ : young_diagram} {i : ℕ} {c : ℕ × ℕ} : c ∈ μ.row i ↔ c ∈ μ ∧ c.fst = i := +by simp [row] + +lemma mk_mem_row_iff {μ : young_diagram} {i j : ℕ} : (i, j) ∈ μ.row i ↔ (i, j) ∈ μ := +by simp [row] + +protected lemma exists_not_mem_row (μ : young_diagram) (i : ℕ) : ∃ j, (i, j) ∉ μ := +begin + obtain ⟨j, hj⟩ := infinite.exists_not_mem_finset + ((μ.cells).preimage (prod.mk i) (λ _ _ _ _ h, by {cases h, refl})), + rw finset.mem_preimage at hj, + exact ⟨j, hj⟩, +end + +/-- Length of a row of a Young diagram -/ +def row_len (μ : young_diagram) (i : ℕ) : ℕ := nat.find $ μ.exists_not_mem_row i + +lemma mem_iff_lt_row_len {μ : young_diagram} {i j : ℕ} : (i, j) ∈ μ ↔ j < μ.row_len i := +by { rw [row_len, nat.lt_find_iff], push_neg, + exact ⟨λ h _ hmj, μ.up_left_mem (by refl) hmj h, λ h, h _ (by refl)⟩ } + +lemma row_eq_prod {μ : young_diagram} {i : ℕ} : μ.row i = {i} ×ˢ finset.range (μ.row_len i) := +by { ext ⟨a, b⟩, + simp only [finset.mem_product, finset.mem_singleton, finset.mem_range, + mem_row_iff, mem_iff_lt_row_len, and_comm, and.congr_right_iff], + rintro rfl, refl } + +lemma row_len_eq_card (μ : young_diagram) {i : ℕ} : μ.row_len i = (μ.row i).card := +by simp [row_eq_prod] + +@[mono] +lemma row_len_anti (μ : young_diagram) (i1 i2 : ℕ) (hi : i1 ≤ i2) : μ.row_len i2 ≤ μ.row_len i1 := +by { by_contra' h_lt, rw ← lt_self_iff_false (μ.row_len i1), + rw ← mem_iff_lt_row_len at h_lt ⊢, + exact μ.up_left_mem hi (by refl) h_lt } + +end rows + +section columns +/-! ### Columns and column lengths of Young diagrams. + +This section has an identical API to the rows section. -/ + +/-- The `j`-th column of a Young diagram consists of the cells whose second coordinate is `j`. -/ +def col (μ : young_diagram) (j : ℕ) : finset (ℕ × ℕ) := μ.cells.filter (λ c, c.snd = j) + +lemma mem_col_iff {μ : young_diagram} {j : ℕ} {c : ℕ × ℕ} : c ∈ μ.col j ↔ c ∈ μ ∧ c.snd = j := +by simp [col] + +lemma mk_mem_col_iff {μ : young_diagram} {i j : ℕ} : (i, j) ∈ μ.col j ↔ (i, j) ∈ μ := +by simp [col] + +protected lemma exists_not_mem_col (μ : young_diagram) (j : ℕ) : ∃ i, (i, j) ∉ μ.cells := +by { convert μ.transpose.exists_not_mem_row j, simp } + +/-- Length of a column of a Young diagram -/ +def col_len (μ : young_diagram) (j : ℕ) : ℕ := nat.find $ μ.exists_not_mem_col j + +@[simp] lemma col_len_transpose (μ : young_diagram) (j : ℕ) : μ.transpose.col_len j = μ.row_len j := +by simp [row_len, col_len] + +@[simp] lemma row_len_transpose (μ : young_diagram) (i : ℕ) : μ.transpose.row_len i = μ.col_len i := +by simp [row_len, col_len] + +lemma mem_iff_lt_col_len {μ : young_diagram} {i j : ℕ} : (i, j) ∈ μ ↔ i < μ.col_len j := +by { rw [← row_len_transpose, ← mem_iff_lt_row_len], simp } + +lemma col_eq_prod {μ : young_diagram} {j : ℕ} : μ.col j = (finset.range (μ.col_len j)) ×ˢ {j} := +by { ext ⟨a, b⟩, + simp only [finset.mem_product, finset.mem_singleton, finset.mem_range, + mem_col_iff, mem_iff_lt_col_len, and_comm, and.congr_right_iff], + rintro rfl, refl } + +lemma col_len_eq_card (μ : young_diagram) {j : ℕ} : μ.col_len j = (μ.col j).card := +by simp [col_eq_prod] + +@[mono] +lemma col_len_anti (μ : young_diagram) (j1 j2 : ℕ) (hj : j1 ≤ j2) : μ.col_len j2 ≤ μ.col_len j1 := +by { convert μ.transpose.row_len_anti j1 j2 hj; simp } + +end columns + end young_diagram From 662786a896a103da7528bf3e807f7016a356b422 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 8 Sep 2022 20:32:29 +0000 Subject: [PATCH 210/828] refactor(data/{finite,fintype}): redefine `infinite` in terms of `finite`, move to `data.finite.defs` (#16390) --- src/data/finite/basic.lean | 8 ------ src/data/finite/defs.lean | 30 +++++++++++++++++----- src/data/fintype/basic.lean | 31 +++++------------------ src/data/fintype/card_embedding.lean | 5 ++-- src/data/set/finite.lean | 5 ++-- src/field_theory/is_alg_closed/basic.lean | 2 +- 6 files changed, 37 insertions(+), 44 deletions(-) diff --git a/src/data/finite/basic.lean b/src/data/finite/basic.lean index cd984fca51e83..060e39a1a81f0 100644 --- a/src/data/finite/basic.lean +++ b/src/data/finite/basic.lean @@ -18,7 +18,6 @@ defined. former lemma takes `fintype α` as an explicit argument while the latter takes it as an instance argument. * `fintype.of_finite` noncomputably creates a `fintype` instance from a `finite` instance. -* `finite_or_infinite` is that every type is either `finite` or `infinite`. ## Implementation notes @@ -42,13 +41,6 @@ open_locale classical variables {α β γ : Type*} -lemma finite_or_infinite (α : Type*) : - finite α ∨ infinite α := -begin - rw ← not_finite_iff_infinite, - apply em -end - namespace finite @[priority 100] -- see Note [lower instance priority] diff --git a/src/data/finite/defs.lean b/src/data/finite/defs.lean index 3ca11998c3577..ef16b491da196 100644 --- a/src/data/finite/defs.lean +++ b/src/data/finite/defs.lean @@ -9,7 +9,8 @@ import logic.equiv.basic # Definition of the `finite` typeclass This file defines a typeclass `finite` saying that `α : Sort*` is finite. A type is `finite` if it -is equivalent to `fin n` for some `n`. +is equivalent to `fin n` for some `n`. We also define `infinite α` as a typeclass equivalent to +`¬finite α`. The `finite` predicate has no computational relevance and, being `Prop`-valued, gets to enjoy proof irrelevance -- it represents the mere fact that the type is finite. While the `fintype` class also @@ -26,6 +27,7 @@ instead. ## Main definitions * `finite α` denotes that `α` is a finite type. +* `infinite α` denotes that `α` is an infinite type. ## Implementation notes @@ -66,11 +68,27 @@ lemma equiv.finite_iff (f : α ≃ β) : finite α ↔ finite β := lemma function.bijective.finite_iff {f : α → β} (h : bijective f) : finite α ↔ finite β := (equiv.of_bijective f h).finite_iff -namespace finite +lemma finite.of_bijective [finite α] {f : α → β} (h : bijective f) : finite β := h.finite_iff.mp ‹_› -lemma of_bijective [finite α] {f : α → β} (h : bijective f) : finite β := h.finite_iff.mp ‹_› +instance [finite α] : finite (plift α) := finite.of_equiv α equiv.plift.symm +instance {α : Type v} [finite α] : finite (ulift.{u} α) := finite.of_equiv α equiv.ulift.symm -instance [finite α] : finite (plift α) := of_equiv α equiv.plift.symm -instance {α : Type v} [finite α] : finite (ulift.{u} α) := of_equiv α equiv.ulift.symm +/-- A type is said to be infinite if it is not finite. Note that `infinite α` is equivalent to +`is_empty (fintype α)` or `is_empty (finite α)`. -/ +class infinite (α : Sort*) : Prop := +(not_finite : ¬finite α) -end finite +@[simp] lemma not_finite_iff_infinite : ¬finite α ↔ infinite α := ⟨infinite.mk, λ h, h.1⟩ + +@[simp] lemma not_infinite_iff_finite : ¬infinite α ↔ finite α := +not_finite_iff_infinite.not_right.symm + +lemma finite_or_infinite (α : Sort*) : finite α ∨ infinite α := +or_iff_not_imp_left.2 $ not_finite_iff_infinite.1 + +lemma not_finite (α : Sort*) [infinite α] [finite α] : false := @infinite.not_finite α ‹_› ‹_› + +protected lemma finite.false [infinite α] (h : finite α) : false := not_finite α +protected lemma infinite.false [finite α] (h : infinite α) : false := not_finite α + +alias not_infinite_iff_finite ↔ finite.of_not_infinite finite.not_infinite diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 0c2591208b25c..fe89a2a923be8 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -1918,34 +1918,15 @@ well_founded_of_trans_of_irrefl _ end finite -/-- A type is said to be infinite if it has no fintype instance. - Note that `infinite α` is equivalent to `is_empty (fintype α)`. -/ -class infinite (α : Type*) : Prop := -(not_fintype : fintype α → false) - -lemma not_finite (α : Type*) [infinite α] [finite α] : false := -by { casesI nonempty_fintype α, exact infinite.not_fintype ‹_› } - -protected lemma finite.false [infinite α] (h : finite α) : false := not_finite α - @[nolint fintype_finite] protected lemma fintype.false [infinite α] (h : fintype α) : false := not_finite α -protected lemma infinite.false [finite α] (h : infinite α) : false := not_finite α @[simp] lemma is_empty_fintype {α : Type*} : is_empty (fintype α) ↔ infinite α := -⟨λ ⟨x⟩, ⟨x⟩, λ ⟨x⟩, ⟨x⟩⟩ - -lemma not_finite_iff_infinite : ¬ finite α ↔ infinite α := -by rw [← is_empty_fintype, finite_iff_nonempty_fintype, not_nonempty_iff] - -lemma not_infinite_iff_finite : ¬ infinite α ↔ finite α := not_finite_iff_infinite.not_right.symm - -alias not_finite_iff_infinite ↔ infinite.of_not_finite infinite.not_finite -alias not_infinite_iff_finite ↔ finite.of_not_infinite finite.not_infinite +⟨λ ⟨h⟩, ⟨λ h', (@nonempty_fintype α h').elim h⟩, λ ⟨h⟩, ⟨λ h', h h'.finite⟩⟩ /-- A non-infinite type is a fintype. -/ noncomputable def fintype_of_not_infinite {α : Type*} (h : ¬ infinite α) : fintype α := -nonempty.some $ by rwa [← not_is_empty_iff, is_empty_fintype] +@fintype.of_finite _ (not_infinite_iff_finite.mp h) section open_locale classical @@ -1975,6 +1956,8 @@ lemma finset.exists_maximal {α : Type*} [preorder α] (s : finset α) (h : s.no namespace infinite +lemma of_not_fintype (h : fintype α → false) : infinite α := is_empty_fintype.mp ⟨h⟩ + lemma exists_not_mem_finset [infinite α] (s : finset α) : ∃ x, x ∉ s := not_forall.1 $ λ h, fintype.false ⟨s, h⟩ @@ -1988,15 +1971,15 @@ protected lemma nonempty (α : Type*) [infinite α] : nonempty α := by apply_instance lemma of_injective [infinite β] (f : β → α) (hf : injective f) : infinite α := -⟨λ I, by exactI (fintype.of_injective f hf).false⟩ +⟨λ I, by exactI (finite.of_injective f hf).false⟩ lemma of_surjective [infinite β] (f : α → β) (hf : surjective f) : infinite α := -⟨λ I, by { classical, exactI (fintype.of_surjective f hf).false }⟩ +⟨λ I, by exactI (finite.of_surjective f hf).false⟩ end infinite instance : infinite ℕ := -⟨λ ⟨s, hs⟩, finset.not_mem_range_self $ s.subset_range_sup_succ (hs _)⟩ +infinite.of_not_fintype $ λ ⟨s, hs⟩, finset.not_mem_range_self $ s.subset_range_sup_succ (hs _) instance : infinite ℤ := infinite.of_injective int.of_nat (λ _ _, int.of_nat.inj) diff --git a/src/data/fintype/card_embedding.lean b/src/data/fintype/card_embedding.lean index 6bd0a76a67166..16296b33ea36e 100644 --- a/src/data/fintype/card_embedding.lean +++ b/src/data/fintype/card_embedding.lean @@ -43,8 +43,9 @@ end /- The cardinality of embeddings from an infinite type to a finite type is zero. This is a re-statement of the pigeonhole principle. -/ -@[simp] lemma card_embedding_eq_of_infinite {α β} [infinite α] [fintype β] [fintype (α ↪ β)] : +@[simp] lemma card_embedding_eq_of_infinite {α β : Type*} [infinite α] [fintype β] + [fintype (α ↪ β)] : ‖α ↪ β‖ = 0 := -card_eq_zero_iff.mpr function.embedding.is_empty +card_eq_zero end fintype diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 516e835d046b2..dd47aab43790c 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -853,10 +853,9 @@ theorem infinite_univ [h : infinite α] : (@univ α).infinite := infinite_univ_iff.2 h theorem infinite_coe_iff {s : set α} : infinite s ↔ s.infinite := -⟨λ ⟨h₁⟩ h₂, h₁ h₂.fintype, λ h₁, ⟨λ h₂, h₁ ⟨h₂⟩⟩⟩ +not_finite_iff_infinite.symm.trans finite_coe_iff.not -theorem infinite.to_subtype {s : set α} (h : s.infinite) : infinite s := -infinite_coe_iff.2 h +alias infinite_coe_iff ↔ _ infinite.to_subtype /-- Embedding of `ℕ` into an infinite set. -/ noncomputable def infinite.nat_embedding (s : set α) (h : s.infinite) : ℕ ↪ s := diff --git a/src/field_theory/is_alg_closed/basic.lean b/src/field_theory/is_alg_closed/basic.lean index 6f9df61bb0b2a..5f7495206061a 100644 --- a/src/field_theory/is_alg_closed/basic.lean +++ b/src/field_theory/is_alg_closed/basic.lean @@ -348,7 +348,7 @@ perfect_ring.of_surjective k p $ λ x, is_alg_closed.exists_pow_nat_eq _ $ ne_ze @[priority 500] instance {K : Type*} [field K] [is_alg_closed K] : infinite K := begin - apply infinite.mk, + apply infinite.of_not_fintype, introsI hfin, set n := fintype.card K with hn, set f := (X : K[X]) ^ (n + 1) - 1 with hf, From d2d6652653ae5a80c20544121f244b0fa2891756 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 8 Sep 2022 20:32:30 +0000 Subject: [PATCH 211/828] feat(number_theory/legendre_symbol/jacobi_symbol): add new file with the definition and properties of the Jacobi symbol (#16395) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR consists of a new file, `number_theory.legendre_symbol.jacobi_symbol`, that contains a definition of the Jacobi symbol (and sets up notation `[a | b]ⱼ` for it) and then proves a number of results, e.g., multiplicativity in both arguments, expressions for the cases `a = -1`, `a = 2`, `a = -2`, and several versions of quadratic reciprocity. It also proves that the symbol depends on a only through its residue class `mod b` and that it depends on `b` only through its residue class `mod 4*a` (quadratic reciprocity is needed for the latter). The code has already been quite thoroughly reviewed while it was part of #16290 (which then was reduced to the auxiliary results needed for the Jacobi symbol code). Discussion on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Jacobi.20symbol/near/295816984) --- .../legendre_symbol/jacobi_symbol.lean | 385 ++++++++++++++++++ 1 file changed, 385 insertions(+) create mode 100644 src/number_theory/legendre_symbol/jacobi_symbol.lean diff --git a/src/number_theory/legendre_symbol/jacobi_symbol.lean b/src/number_theory/legendre_symbol/jacobi_symbol.lean new file mode 100644 index 0000000000000..f1f5e9ac3c0c6 --- /dev/null +++ b/src/number_theory/legendre_symbol/jacobi_symbol.lean @@ -0,0 +1,385 @@ +/- +Copyright (c) 2022 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import number_theory.legendre_symbol.quadratic_reciprocity +import data.zmod.coprime + +/-! +# The Jacobi Symbol + +We define the Jacobi symbol and prove its main properties. + +## Main definitions + +We define the Jacobi symbol, `jacobi_sym a b`, for integers `a` and natural numbers `b` +as the product over the prime factors `p` of `b` of the Legendre symbols `zmod.legendre_sym p a`. +This agrees with the mathematical definition when `b` is odd. + +The prime factors are obtained via `nat.factors`. Since `nat.factors 0 = []`, +this implies in particular that `jacobi_sym a 0 = 1` for all `a`. + +## Main statements + +We prove the main properties of the Jacobi symbol, including the following. + +* Multiplicativity in both arguments (`jacobi_sym_mul_left`, `jacobi_sym_mul_right`) + +* The value of the symbol is `1` or `-1` when the arguments are coprime + (`jacobi_sym_eq_one_or_neg_one`) + +* The symbol vanishes if and only if `b ≠ 0` and the arguments are not coprime + (`jacobi_sym_eq_zero_iff`) + +* If the symbol has the value `-1`, then `a : zmod b` is not a square + (`nonsquare_of_jacobi_sym_eq_neg_one`); the converse holds when `b = p` is a prime + (`nonsquare_iff_jacobi_sym_eq_neg_one`); in particular, in this case `a` is a + square mod `p` when the symbol has the value `1` (`is_square_of_jacobi_sym_eq_one`). + +* Quadratic reciprocity (`jacobi_sym_quadratic_reciprocity`, + `jacobi_sym_quadratic_reciprocity_one_mod_four`, + `jacobi_sym_quadratic_reciprocity_three_mod_four`) + +* The supplementary laws for `a = -1`, `a = 2`, `a = -2` (`jacobi_sym_neg_one`, `jacobi_sym_two`, + `jacobi_sym_neg_two`) + +* The symbol depends on `a` only via its residue class mod `b` (`jacobi_sym_mod_left`) + and on `b` only via its residue class mod `4*a` (`jacobi_sym_mod_right`) + +## Notations + +We define the notation `J(a | b)` for `jacobi_sym a b`, localized to `number_theory_symbols`. + +## Tags +Jacobi symbol, quadratic reciprocity +-/ + +section jacobi + +/-! +### Definition of the Jacobi symbol + +We define the Jacobi symbol $\Bigl\frac{a}{b}\Bigr$ for integers `a` and natural numbers `b` as the +product of the Legendre symbols $\Bigl\frac{a}{p}\Bigr$, where `p` runs through the prime divisors +(with multiplicity) of `b`, as provided by `b.factors`. This agrees with the Jacobi symbol +when `b` is odd and gives less meaningful values when it is not (e.g., the symbol is `1` +when `b = 0`). This is called `jacobi_sym a b`. + +We define localized notation (locale `number_theory_symbols`) `J(a | b)` for the Jacobi +symbol `jacobi_sym a b`. (Unfortunately, there is no subscript "J" in unicode.) +-/ + +namespace zmod +open nat + +/-- The Jacobi symbol of `a` and `b` -/ +-- Since we need the fact that the factors are prime, we use `list.pmap`. +def jacobi_sym (a : ℤ) (b : ℕ) : ℤ := +(b.factors.pmap (λ p pp, @legendre_sym p ⟨pp⟩ a) (λ p pf, prime_of_mem_factors pf)).prod + +-- Notation for the Jacobi symbol. +localized "notation `J(` a ` | ` b `)` := jacobi_sym a b" in number_theory_symbols + +/-! +### Properties of the Jacobi symbol +-/ + +/-- The symbol `J(a | 0)` has the value `1`. -/ +@[simp] lemma jacobi_sym_zero_right (a : ℤ) : J(a | 0) = 1 := +by simp only [jacobi_sym, factors_zero, list.prod_nil, list.pmap] + +/-- The symbol `J(a | 1)` has the value `1`. -/ +@[simp] lemma jacobi_sym_one_right (a : ℤ) : J(a | 1) = 1 := +by simp only [jacobi_sym, factors_one, list.prod_nil, list.pmap] + +/-- The Legendre symbol `legendre_sym p a` with an integer `a` and a prime number `p` +is the same as the Jacobi symbol `J(a | p)`. -/ +lemma legendre_sym.to_jacobi_sym {p : ℕ} [fp : fact p.prime] {a : ℤ} : + legendre_sym p a = J(a | p) := +by simp only [jacobi_sym, factors_prime fp.1, list.prod_cons, list.prod_nil, mul_one, list.pmap] + +/-- The Jacobi symbol is multiplicative in its second argument. -/ +lemma jacobi_sym_mul_right' (a : ℤ) {b₁ b₂ : ℕ} (hb₁ : b₁ ≠ 0) (hb₂ : b₂ ≠ 0) : + J(a | b₁ * b₂) = J(a | b₁) * J(a | b₂) := +begin + rw [jacobi_sym, ((perm_factors_mul hb₁ hb₂).pmap _).prod_eq, list.pmap_append, list.prod_append], + exacts [rfl, λ p hp, (list.mem_append.mp hp).elim prime_of_mem_factors prime_of_mem_factors], +end + +/-- The Jacobi symbol is multiplicative in its second argument. -/ +lemma jacobi_sym_mul_right (a : ℤ) (b₁ b₂ : ℕ) [ne_zero b₁] [ne_zero b₂] : + J(a | b₁ * b₂) = J(a | b₁) * J(a | b₂) := +jacobi_sym_mul_right' a (ne_zero.ne b₁) (ne_zero.ne b₂) + +/-- The Jacobi symbol takes only the values `0`, `1` and `-1`. -/ +lemma jacobi_sym_trichotomy (a : ℤ) (b : ℕ) : J(a | b) = 0 ∨ J(a | b) = 1 ∨ J(a | b) = -1 := +((@sign_type.cast_hom ℤ _ _).to_monoid_hom.mrange.copy {0, 1, -1} $ + by {rw set.pair_comm, exact (sign_type.range_eq sign_type.cast_hom).symm}).list_prod_mem +begin + intros _ ha', + rcases list.mem_pmap.mp ha' with ⟨p, hp, rfl⟩, + haveI : fact p.prime := ⟨prime_of_mem_factors hp⟩, + exact quadratic_char_is_quadratic (zmod p) a, +end + +/-- The symbol `J(1 | b)` has the value `1`. -/ +@[simp] lemma jacobi_sym_one_left (b : ℕ) : J(1 | b) = 1 := +list.prod_eq_one (λ z hz, let ⟨p, hp, he⟩ := list.mem_pmap.1 hz in by rw [← he, legendre_sym_one]) + +/-- The Jacobi symbol is multiplicative in its first argument. -/ +lemma jacobi_sym_mul_left (a₁ a₂ : ℤ) (b : ℕ) : J(a₁ * a₂ | b) = J(a₁ | b) * J(a₂ | b) := +by { simp_rw [jacobi_sym, list.pmap_eq_map_attach, legendre_sym_mul], exact list.prod_map_mul } + +/-- The symbol `J(a | b)` vanishes iff `a` and `b` are not coprime (assuming `b ≠ 0`). -/ +lemma jacobi_sym_eq_zero_iff_not_coprime {a : ℤ} {b : ℕ} [ne_zero b] : + J(a | b) = 0 ↔ a.gcd b ≠ 1 := +list.prod_eq_zero_iff.trans begin + rw [list.mem_pmap, int.gcd_eq_nat_abs, ne, prime.not_coprime_iff_dvd], + simp_rw [legendre_sym_eq_zero_iff, int_coe_zmod_eq_zero_iff_dvd, mem_factors (ne_zero.ne b), + ← int.coe_nat_dvd_left, int.coe_nat_dvd, exists_prop, and_assoc, and_comm], +end + +/-- The symbol `J(a | b)` is nonzero when `a` and `b` are coprime. -/ +lemma jacobi_sym_ne_zero {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a | b) ≠ 0 := +begin + casesI eq_zero_or_ne_zero b with hb, + { rw [hb, jacobi_sym_zero_right], + exact one_ne_zero }, + { contrapose! h, exact jacobi_sym_eq_zero_iff_not_coprime.1 h }, +end + +/-- The symbol `J(a | b)` vanishes if and only if `b ≠ 0` and `a` and `b` are not coprime. -/ +lemma jacobi_sym_eq_zero_iff {a : ℤ} {b : ℕ} : J(a | b) = 0 ↔ b ≠ 0 ∧ a.gcd b ≠ 1 := +⟨λ h, begin + casesI eq_or_ne b 0 with hb hb, + { rw [hb, jacobi_sym_zero_right] at h, cases h }, + exact ⟨hb, mt jacobi_sym_ne_zero $ not_not.2 h⟩, +end, λ ⟨hb, h⟩, by { rw ← ne_zero_iff at hb, exactI jacobi_sym_eq_zero_iff_not_coprime.2 h }⟩ + +/-- The symbol `J(0 | b)` vanishes when `b > 1`. -/ +lemma jacobi_sym_zero_left {b : ℕ} (hb : 1 < b) : J(0 | b) = 0 := +(@jacobi_sym_eq_zero_iff_not_coprime 0 b ⟨ne_zero_of_lt hb⟩).mpr $ + by { rw [int.gcd_zero_left, int.nat_abs_of_nat], exact hb.ne' } + +/-- The symbol `J(a | b)` takes the value `1` or `-1` if `a` and `b` are coprime. -/ +lemma jacobi_sym_eq_one_or_neg_one {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : + J(a | b) = 1 ∨ J(a | b) = -1 := +(jacobi_sym_trichotomy a b).resolve_left $ jacobi_sym_ne_zero h + +/-- We have that `J(a^e | b) = J(a | b)^e`. -/ +lemma jacobi_sym_pow_left (a : ℤ) (e b : ℕ) : J(a ^ e | b) = J(a | b) ^ e := +nat.rec_on e (by rw [pow_zero, pow_zero, jacobi_sym_one_left]) $ + λ _ ih, by rw [pow_succ, pow_succ, jacobi_sym_mul_left, ih] + +/-- We have that `J(a | b^e) = J(a | b)^e`. -/ +lemma jacobi_sym_pow_right (a : ℤ) (b e : ℕ) : J(a | b ^ e) = J(a | b) ^ e := +begin + induction e with e ih, + { rw [pow_zero, pow_zero, jacobi_sym_one_right], }, + { casesI eq_zero_or_ne_zero b with hb, + { rw [hb, zero_pow (succ_pos e), jacobi_sym_zero_right, one_pow], }, + { rw [pow_succ, pow_succ, jacobi_sym_mul_right, ih], } } +end + +/-- The square of `J(a | b)` is `1` when `a` and `b` are coprime. -/ +lemma jacobi_sym_sq_one {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a | b) ^ 2 = 1 := +by cases jacobi_sym_eq_one_or_neg_one h with h₁ h₁; rw h₁; refl + +/-- The symbol `J(a^2 | b)` is `1` when `a` and `b` are coprime. -/ +lemma jacobi_sym_sq_one' {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a ^ 2 | b) = 1 := +by rw [jacobi_sym_pow_left, jacobi_sym_sq_one h] + +/-- The symbol `J(a | b)` depends only on `a` mod `b`. -/ +lemma jacobi_sym_mod_left (a : ℤ) (b : ℕ) : J(a | b) = J(a % b | b) := +congr_arg list.prod $ list.pmap_congr _ begin + rintro p hp _ _, + conv_rhs { rw [legendre_sym_mod, int.mod_mod_of_dvd _ + (int.coe_nat_dvd.2 $ dvd_of_mem_factors hp), ← legendre_sym_mod] }, +end + +/-- The symbol `J(a | b)` depends only on `a` mod `b`. -/ +lemma jacobi_sym_mod_left' {a₁ a₂ : ℤ} {b : ℕ} (h : a₁ % b = a₂ % b) : J(a₁ | b) = J(a₂ | b) := +by rw [jacobi_sym_mod_left, h, ← jacobi_sym_mod_left] + +/-- If `J(a | b)` is `-1`, then `a` is not a square modulo `b`. -/ +lemma nonsquare_of_jacobi_sym_eq_neg_one {a : ℤ} {b : ℕ} (h : J(a | b) = -1) : + ¬ is_square (a : zmod b) := +λ ⟨r, ha⟩, begin + rw [← r.coe_val_min_abs, ← int.cast_mul, int_coe_eq_int_coe_iff', ← sq] at ha, + apply (by norm_num : ¬ (0 : ℤ) ≤ -1), + rw [← h, jacobi_sym_mod_left, ha, ← jacobi_sym_mod_left, jacobi_sym_pow_left], + apply sq_nonneg, +end + +/-- If `p` is prime, then `J(a | p)` is `-1` iff `a` is not a square modulo `p`. -/ +lemma nonsquare_iff_jacobi_sym_eq_neg_one {a : ℤ} {p : ℕ} [fact p.prime] : + J(a | p) = -1 ↔ ¬ is_square (a : zmod p) := +by { rw [← legendre_sym.to_jacobi_sym], exact legendre_sym_eq_neg_one_iff p } + +/-- If `p` is prime and `J(a | p) = 1`, then `a` is q square mod `p`. -/ +lemma is_square_of_jacobi_sym_eq_one {a : ℤ} {p : ℕ} [fact p.prime] (h : J(a | p) = 1) : + is_square (a : zmod p) := +not_not.mp $ mt nonsquare_iff_jacobi_sym_eq_neg_one.mpr $ + λ hf, one_ne_zero $ neg_eq_self_iff.mp $ hf.symm.trans h + +/-! +### Values at `-1`, `2` and `-2` +-/ + +/-- If `χ` is a multiplicative function such that `J(a | p) = χ p` for all odd primes `p`, +then `J(a | b)` equals `χ b` for all odd natural numbers `b`. -/ +lemma jacobi_sym_value (a : ℤ) {R : Type*} [comm_semiring R] (χ : R →* ℤ) + (hp : ∀ (p : ℕ) (pp : p.prime) (h2 : p ≠ 2), @legendre_sym p ⟨pp⟩ a = χ p) {b : ℕ} (hb : odd b) : + J(a | b) = χ b := +begin + conv_rhs { rw [← prod_factors hb.pos.ne', cast_list_prod, χ.map_list_prod] }, + rw [jacobi_sym, list.map_map, ← list.pmap_eq_map nat.prime _ _ (λ _, prime_of_mem_factors)], + congr' 1, apply list.pmap_congr, + exact λ p h pp _, hp p pp (hb.factors_ne_two h), +end + +/-- If `b` is odd, then `J(-1 | b)` is given by `χ₄ b`. -/ +lemma jacobi_sym_neg_one {b : ℕ} (hb : odd b) : J(-1 | b) = χ₄ b := +jacobi_sym_value (-1) χ₄ (λ p pp, @legendre_sym_neg_one p ⟨pp⟩) hb + +/-- If `b` is odd, then `J(-a | b) = χ₄ b * J(a | b)`. -/ +lemma jacobi_sym_neg (a : ℤ) {b : ℕ} (hb : odd b) : J(-a | b) = χ₄ b * J(a | b) := +by rw [neg_eq_neg_one_mul, jacobi_sym_mul_left, jacobi_sym_neg_one hb] + +/-- If `b` is odd, then `J(2 | b)` is given by `χ₈ b`. -/ +lemma jacobi_sym_two {b : ℕ} (hb : odd b) : J(2 | b) = χ₈ b := +jacobi_sym_value 2 χ₈ (λ p pp, @legendre_sym_two p ⟨pp⟩) hb + +/-- If `b` is odd, then `J(-2 | b)` is given by `χ₈' b`. -/ +lemma jacobi_sym_neg_two {b : ℕ} (hb : odd b) : J(-2 | b) = χ₈' b := +jacobi_sym_value (-2) χ₈' (λ p pp, @legendre_sym_neg_two p ⟨pp⟩) hb + + +/-! +### Quadratic Reciprocity +-/ + +/-- The bi-multiplicative map giving the sign in the Law of Quadratic Reciprocity -/ +def qr_sign (m n : ℕ) : ℤ := J(χ₄ m | n) + +/-- We can express `qr_sign m n` as a power of `-1` when `m` and `n` are odd. -/ +lemma qr_sign_neg_one_pow {m n : ℕ} (hm : odd m) (hn : odd n) : + qr_sign m n = (-1) ^ ((m / 2) * (n / 2)) := +begin + rw [qr_sign, pow_mul, ← χ₄_eq_neg_one_pow (odd_iff.mp hm)], + cases odd_mod_four_iff.mp (odd_iff.mp hm) with h h, + { rw [χ₄_nat_one_mod_four h, jacobi_sym_one_left, one_pow], }, + { rw [χ₄_nat_three_mod_four h, ← χ₄_eq_neg_one_pow (odd_iff.mp hn), jacobi_sym_neg_one hn], } +end + +/-- When `m` and `n` are odd, then the square of `qr_sign m n` is `1`. -/ +lemma qr_sign_sq_eq_one {m n : ℕ} (hm : odd m) (hn : odd n) : (qr_sign m n) ^ 2 = 1 := +by rw [qr_sign_neg_one_pow hm hn, ← pow_mul, mul_comm, pow_mul, neg_one_sq, one_pow] + +/-- `qr_sign` is multiplicative in the first argument. -/ +lemma qr_sign_mul_left (m₁ m₂ n : ℕ) : qr_sign (m₁ * m₂) n = qr_sign m₁ n * qr_sign m₂ n := +by simp_rw [qr_sign, nat.cast_mul, map_mul, jacobi_sym_mul_left] + +/-- `qr_sign` is multiplicative in the second argument. -/ +lemma qr_sign_mul_right (m n₁ n₂ : ℕ) [ne_zero n₁] [ne_zero n₂] : + qr_sign m (n₁ * n₂) = qr_sign m n₁ * qr_sign m n₂ := +jacobi_sym_mul_right (χ₄ m) n₁ n₂ + +/-- `qr_sign` is symmetric when both arguments are odd. -/ +lemma qr_sign_symm {m n : ℕ} (hm : odd m) (hn : odd n) : qr_sign m n = qr_sign n m := +by rw [qr_sign_neg_one_pow hm hn, qr_sign_neg_one_pow hn hm, mul_comm (m / 2)] + +/-- We can move `qr_sign m n` from one side of an equality to the other when `m` and `n` are odd. -/ +lemma qr_sign_eq_iff_eq {m n : ℕ} (hm : odd m) (hn : odd n) (x y : ℤ) : + qr_sign m n * x = y ↔ x = qr_sign m n * y := +by refine ⟨λ h', let h := h'.symm in _, λ h, _⟩; + rw [h, ← mul_assoc, ← pow_two, qr_sign_sq_eq_one hm hn, one_mul] + +/-- The Law of Quadratic Reciprocity for the Jacobi symbol, version with `qr_sign` -/ +lemma jacobi_sym_quadratic_reciprocity' {a b : ℕ} (ha : odd a) (hb : odd b) : + J(a | b) = qr_sign b a * J(b | a) := +begin + -- define the right hand side for fixed `a` as a `ℕ →* ℤ` + let rhs : ℕ → ℕ →* ℤ := λ a, + { to_fun := λ x, qr_sign x a * J(x | a), + map_one' := by { convert ← mul_one _, symmetry, all_goals { apply jacobi_sym_one_left } }, + map_mul' := λ x y, by rw [qr_sign_mul_left, nat.cast_mul, jacobi_sym_mul_left, + mul_mul_mul_comm] }, + have rhs_apply : ∀ (a b : ℕ), rhs a b = qr_sign b a * J(b | a) := λ a b, rfl, + refine jacobi_sym_value a (rhs a) (λ p pp hp, eq.symm _) hb, + have hpo := pp.eq_two_or_odd'.resolve_left hp, + rw [@legendre_sym.to_jacobi_sym p ⟨pp⟩, rhs_apply, nat.cast_id, + qr_sign_eq_iff_eq hpo ha, qr_sign_symm hpo ha], + refine jacobi_sym_value p (rhs p) (λ q pq hq, _) ha, + have hqo := pq.eq_two_or_odd'.resolve_left hq, + rw [rhs_apply, nat.cast_id, ← @legendre_sym.to_jacobi_sym p ⟨pp⟩, qr_sign_symm hqo hpo, + qr_sign_neg_one_pow hpo hqo, @quadratic_reciprocity' p q ⟨pp⟩ ⟨pq⟩ hp hq], +end + +/-- The Law of Quadratic Reciprocity for the Jacobi symbol -/ +lemma jacobi_sym_quadratic_reciprocity {a b : ℕ} (ha : odd a) (hb : odd b) : + J(a | b) = (-1) ^ ((a / 2) * (b / 2)) * J(b | a) := +by rw [← qr_sign_neg_one_pow ha hb, qr_sign_symm ha hb, jacobi_sym_quadratic_reciprocity' ha hb] + +/-- The Law of Quadratic Reciprocity for the Jacobi symbol: if `a` and `b` are natural numbers +with `a % 4 = 1` and `b` odd, then `J(a | b) = J(b | a)`. -/ +theorem jacobi_sym_quadratic_reciprocity_one_mod_four {a b : ℕ} (ha : a % 4 = 1) (hb : odd b) : + J(a | b) = J(b | a) := +by rw [jacobi_sym_quadratic_reciprocity (odd_iff.mpr (odd_of_mod_four_eq_one ha)) hb, + pow_mul, neg_one_pow_div_two_of_one_mod_four ha, one_pow, one_mul] + +/-- The Law of Quadratic Reciprocity for the Jacobi symbol: if `a` and `b` are natural numbers +with `a` odd and `b % 4 = 1`, then `J(a | b) = J(b | a)`. -/ +theorem jacobi_sym_quadratic_reciprocity_one_mod_four' {a b : ℕ} (ha : odd a) (hb : b % 4 = 1) : + J(a | b) = J(b | a) := +(jacobi_sym_quadratic_reciprocity_one_mod_four hb ha).symm + +/-- The Law of Quadratic Reciprocityfor the Jacobi symbol: if `a` and `b` are natural numbers +both congruent to `3` mod `4`, then `J(a | b) = -J(b | a)`. -/ +theorem jacobi_sym_quadratic_reciprocity_three_mod_four + {a b : ℕ} (ha : a % 4 = 3) (hb : b % 4 = 3) : + J(a | b) = - J(b | a) := +let nop := @neg_one_pow_div_two_of_three_mod_four in begin + rw [jacobi_sym_quadratic_reciprocity, pow_mul, nop ha, nop hb, neg_one_mul]; + rwa [odd_iff, odd_of_mod_four_eq_three], +end + +/-- The Jacobi symbol `J(a | b)` depends only on `b` mod `4*a` (version for `a : ℕ`). -/ +lemma jacobi_sym_mod_right' (a : ℕ) {b : ℕ} (hb : odd b) : J(a | b) = J(a | b % (4 * a)) := +begin + rcases eq_or_ne a 0 with rfl | ha₀, + { rw [mul_zero, mod_zero], }, + have hb' : odd (b % (4 * a)) := hb.mod_even (even.mul_right (by norm_num) _), + rcases exists_eq_pow_mul_and_not_dvd ha₀ 2 (by norm_num) with ⟨e, a', ha₁', ha₂⟩, + have ha₁ := odd_iff.mpr (two_dvd_ne_zero.mp ha₁'), + nth_rewrite 1 [ha₂], nth_rewrite 0 [ha₂], + rw [nat.cast_mul, jacobi_sym_mul_left, jacobi_sym_mul_left, + jacobi_sym_quadratic_reciprocity' ha₁ hb, jacobi_sym_quadratic_reciprocity' ha₁ hb', + nat.cast_pow, jacobi_sym_pow_left, jacobi_sym_pow_left, + (show ((2 : ℕ) : ℤ) = 2, by refl), jacobi_sym_two hb, jacobi_sym_two hb'], + congr' 1, swap, congr' 1, + { simp_rw [qr_sign], + rw [χ₄_nat_mod_four, χ₄_nat_mod_four (b % (4 * a)), mod_mod_of_dvd b (dvd_mul_right 4 a) ] }, + { rw [jacobi_sym_mod_left ↑(b % _), jacobi_sym_mod_left b, int.coe_nat_mod, + int.mod_mod_of_dvd b], + simp only [ha₂, nat.cast_mul, ← mul_assoc], + exact dvd_mul_left a' _, }, + cases e, { refl }, + { rw [χ₈_nat_mod_eight, χ₈_nat_mod_eight (b % (4 * a)), mod_mod_of_dvd b], + use 2 ^ e * a', rw [ha₂, pow_succ], ring, } +end + +/-- The Jacobi symbol `J(a | b)` depends only on `b` mod `4*a`. -/ +lemma jacobi_sym_mod_right (a : ℤ) {b : ℕ} (hb : odd b) : J(a | b) = J(a | b % (4 * a.nat_abs)) := +begin + cases int.nat_abs_eq a with ha ha; nth_rewrite 1 [ha]; nth_rewrite 0 [ha], + { exact jacobi_sym_mod_right' a.nat_abs hb, }, + { have hb' : odd (b % (4 * a.nat_abs)) := hb.mod_even (even.mul_right (by norm_num) _), + rw [jacobi_sym_neg _ hb, jacobi_sym_neg _ hb', jacobi_sym_mod_right' _ hb, χ₄_nat_mod_four, + χ₄_nat_mod_four (b % (4 * _)), mod_mod_of_dvd b (dvd_mul_right 4 _)], } +end + +end zmod + +end jacobi From ba346d8b82f7e35270418ac2ff8c59c4db3b5af5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 8 Sep 2022 20:32:32 +0000 Subject: [PATCH 212/828] feat(order/upper_lower): `Iic` is not bounded below (#16425) Nonempty lower sets are not bounded below in a `no_min_order`, and dual. --- src/order/upper_lower.lean | 30 ++++++++++++++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/src/order/upper_lower.lean b/src/order/upper_lower.lean index abb12e02c4aed..e89c2245b90c0 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower.lean @@ -185,6 +185,36 @@ lemma is_lower_set.not_bot_mem (hs : is_lower_set s) : ⊥ ∉ s ↔ s = ∅ := hs.bot_mem.not.trans not_nonempty_iff_eq_empty end order_bot + +section no_max_order +variables [no_max_order α] (a) + +lemma is_upper_set.not_bdd_above (hs : is_upper_set s) : s.nonempty → ¬ bdd_above s := +begin + rintro ⟨a, ha⟩ ⟨b, hb⟩, + obtain ⟨c, hc⟩ := exists_gt b, + exact hc.not_le (hb $ hs ((hb ha).trans hc.le) ha), +end + +lemma not_bdd_above_Ici : ¬ bdd_above (Ici a) := (is_upper_set_Ici _).not_bdd_above nonempty_Ici +lemma not_bdd_above_Ioi : ¬ bdd_above (Ioi a) := (is_upper_set_Ioi _).not_bdd_above nonempty_Ioi + +end no_max_order + +section no_min_order +variables [no_min_order α] (a) + +lemma is_lower_set.not_bdd_below (hs : is_lower_set s) : s.nonempty → ¬ bdd_below s := +begin + rintro ⟨a, ha⟩ ⟨b, hb⟩, + obtain ⟨c, hc⟩ := exists_lt b, + exact hc.not_le (hb $ hs (hc.le.trans $ hb ha) ha), +end + +lemma not_bdd_below_Iic : ¬ bdd_below (Iic a) := (is_lower_set_Iic _).not_bdd_below nonempty_Iic +lemma not_bdd_below_Iio : ¬ bdd_below (Iio a) := (is_lower_set_Iio _).not_bdd_below nonempty_Iio + +end no_min_order end preorder /-! ### Bundled upper/lower sets -/ From 75cc1ae96453f283a9cd2f42152cb5e724dac53a Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 8 Sep 2022 23:06:22 +0000 Subject: [PATCH 213/828] feat(analysis/normed/group/basic): add `norm_multiset_sum_le` (#16419) --- src/analysis/normed/group/basic.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 89520e23fda3d..069242539e902 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -235,6 +235,10 @@ lemma ne_zero_of_norm_ne_zero {g : E} : ∥g∥ ≠ 0 → g ≠ 0 := mt $ by { r @[nontriviality] lemma norm_of_subsingleton [subsingleton E] (x : E) : ∥x∥ = 0 := by rw [subsingleton.elim x 0, norm_zero] +lemma norm_multiset_sum_le (m : multiset E) : + ∥m.sum∥ ≤ (m.map (λ x, ∥x∥)).sum := +m.le_sum_of_subadditive norm norm_zero norm_add_le + lemma norm_sum_le (s : finset ι) (f : ι → E) : ∥∑ i in s, f i∥ ≤ ∑ i in s, ∥f i∥ := s.le_sum_of_subadditive norm norm_zero norm_add_le f @@ -708,6 +712,10 @@ by simp only [sub_eq_add_neg, edist_add_left, edist_neg_neg] @[simp] lemma edist_sub_right (g₁ g₂ h : E) : edist (g₁ - h) (g₂ - h) = edist g₁ g₂ := by simpa only [sub_eq_add_neg] using edist_add_right _ _ _ +lemma nnnorm_multiset_sum_le (m : multiset E) : + ∥m.sum∥₊ ≤ (m.map (λ x, ∥x∥₊)).sum := +m.le_sum_of_subadditive nnnorm nnnorm_zero nnnorm_add_le + lemma nnnorm_sum_le (s : finset ι) (f : ι → E) : ∥∑ a in s, f a∥₊ ≤ ∑ a in s, ∥f a∥₊ := s.le_sum_of_subadditive nnnorm nnnorm_zero nnnorm_add_le f From 3a0b839b9f8d6c2f88b043b039db4bded3eb9572 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 9 Sep 2022 12:28:35 +0000 Subject: [PATCH 214/828] feat(analysis/{normed/group}/seminorm): Hom classes for seminorms (#16227) Introduce `add_group_seminorm_class`, `group_seminorm_class`, `seminorm_class`, the hom classes of `add_group_seminorm`, `group_seminorm`, `seminorm`. --- .../seminorm_lattice_not_distrib.lean | 13 +- src/algebra/order/absolute_value.lean | 21 ++- src/algebra/order/hom/basic.lean | 60 +++++++ src/analysis/convex/gauge.lean | 14 +- .../locally_convex/with_seminorms.lean | 11 +- src/analysis/normed/group/seminorm.lean | 158 ++++++++++-------- src/analysis/seminorm.lean | 134 +++++++-------- 7 files changed, 244 insertions(+), 167 deletions(-) create mode 100644 src/algebra/order/hom/basic.lean diff --git a/counterexamples/seminorm_lattice_not_distrib.lean b/counterexamples/seminorm_lattice_not_distrib.lean index ca4c5798c29fb..7963e3e7a07ef 100644 --- a/counterexamples/seminorm_lattice_not_distrib.lean +++ b/counterexamples/seminorm_lattice_not_distrib.lean @@ -17,13 +17,10 @@ This proves the lattice `seminorm ℝ (ℝ × ℝ)` is not distributive. * https://en.wikipedia.org/wiki/Seminorm#Examples -/ -namespace seminorm_not_distrib +open seminorm open_locale nnreal -private lemma bdd_below_range_add {𝕜 E : Type*} [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] - (x : E) (p q : seminorm 𝕜 E) : - bdd_below (set.range (λ (u : E), p u + q (x - u))) := -by { use 0, rintro _ ⟨x, rfl⟩, exact add_nonneg (p.nonneg _) (q.nonneg _) } +namespace seminorm_not_distrib @[simps] noncomputable def p : seminorm ℝ (ℝ×ℝ) := (norm_seminorm ℝ ℝ).comp (linear_map.fst _ _ _) ⊔ (norm_seminorm ℝ ℝ).comp (linear_map.snd _ _ _) @@ -37,7 +34,7 @@ by { use 0, rintro _ ⟨x, rfl⟩, exact add_nonneg (p.nonneg _) (q.nonneg _) } lemma eq_one : (p ⊔ (q1 ⊓ q2)) (1, 1) = 1 := begin suffices : (⨅ x : ℝ × ℝ, q1 x + q2 (1 - x)) ≤ 1, by simpa, - apply cinfi_le_of_le (bdd_below_range_add _ _ _) ((0, 1) : ℝ×ℝ), dsimp [q1, q2], + apply cinfi_le_of_le bdd_below_range_add ((0, 1) : ℝ×ℝ), dsimp [q1, q2], simp only [abs_zero, smul_zero, sub_self, add_zero, zero_le_one], end @@ -56,7 +53,7 @@ begin ... ≤ 4 * |1 - x.snd| : (mul_le_mul_left zero_lt_four).mpr (le_abs_self _) ... = q2 ((1, 1) - x) : rfl ... ≤ (p ⊔ q2) ((1, 1) - x) : le_sup_right - ... ≤ (p ⊔ q1) x + (p ⊔ q2) ((1, 1) - x) : le_add_of_nonneg_left ((p ⊔ q1).nonneg _) }, + ... ≤ (p ⊔ q1) x + (p ⊔ q2) ((1, 1) - x) : le_add_of_nonneg_left (map_nonneg _ _) }, { calc 4/3 = 2/3 + (1 - 1/3) : by norm_num ... ≤ x.snd + (1 - x.fst) : add_le_add (le_of_lt h2) (sub_le_sub_left h1 _) ... ≤ |x.snd| + |1 - x.fst| : add_le_add (le_abs_self _) (le_abs_self _) @@ -67,7 +64,7 @@ begin ... ≤ 4 * |x.fst| : (mul_le_mul_left zero_lt_four).mpr (le_abs_self _) ... = q1 x : rfl ... ≤ (p ⊔ q1) x : le_sup_right - ... ≤ (p ⊔ q1) x + (p ⊔ q2) ((1, 1) - x) : le_add_of_nonneg_right ((p ⊔ q2).nonneg _) } + ... ≤ (p ⊔ q1) x + (p ⊔ q2) ((1, 1) - x) : le_add_of_nonneg_right (map_nonneg _ _) } end end seminorm_not_distrib diff --git a/src/algebra/order/absolute_value.lean b/src/algebra/order/absolute_value.lean index def9cda1ba8af..e5036eb1c1d9c 100644 --- a/src/algebra/order/absolute_value.lean +++ b/src/algebra/order/absolute_value.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Anne Baanen -/ import algebra.order.field +import algebra.order.hom.basic /-! # Absolute values @@ -39,12 +40,26 @@ section semiring variables {R S : Type*} [semiring R] [ordered_semiring S] (abv : absolute_value R S) -instance mul_hom_class : mul_hom_class (absolute_value R S) R S := +instance zero_hom_class : zero_hom_class (absolute_value R S) R S := { coe := λ f, f.to_fun, coe_injective' := λ f g h, by { obtain ⟨⟨_, _⟩, _⟩ := f, obtain ⟨⟨_, _⟩, _⟩ := g, congr' }, - map_mul := λ f, f.map_mul' } + map_zero := λ f, (f.eq_zero' _).2 rfl } + +instance mul_hom_class : mul_hom_class (absolute_value R S) R S := +{ map_mul := λ f, f.map_mul' + ..absolute_value.zero_hom_class } + +instance nonneg_hom_class : nonneg_hom_class (absolute_value R S) R S := +{ map_nonneg := λ f, f.nonneg', + ..absolute_value.zero_hom_class } + +instance subadditive_hom_class : subadditive_hom_class (absolute_value R S) R S := +{ map_add_le_add := λ f, f.add_le', + ..absolute_value.zero_hom_class } -instance : has_coe_to_fun (absolute_value R S) (λ f, R → S) := ⟨λ f, f.to_fun⟩ +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun` +directly. -/ +instance : has_coe_to_fun (absolute_value R S) (λ f, R → S) := fun_like.has_coe_to_fun @[simp] lemma coe_to_mul_hom : ⇑abv.to_mul_hom = abv := rfl diff --git a/src/algebra/order/hom/basic.lean b/src/algebra/order/hom/basic.lean new file mode 100644 index 0000000000000..b36d1b7a08760 --- /dev/null +++ b/src/algebra/order/hom/basic.lean @@ -0,0 +1,60 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import algebra.hom.group +import algebra.order.with_zero +import order.hom.basic + +/-! +# Algebraic order homomorphism classes + +This file defines hom classes for common properties at the intersection of order theory and algebra. + +## Typeclasses + +* `nonneg_hom_class`: Homs are nonnegative: `∀ f a, 0 ≤ f a` +* `subadditive_hom_class`: Homs are subadditive: `∀ f a b, f (a + b) ≤ f a + f b` +* `submultiplicative_hom_class`: Homs are submultiplicative: `∀ f a b, f (a * b) ≤ f a * f b` +* `mul_le_add_hom_class`: `∀ f a b, f (a * b) ≤ f a + f b` +-/ + +set_option old_structure_cmd true + +open function + +variables {F α β γ δ : Type*} + +/-- `nonneg_hom_class F α β` states that `F` is a type of nonnegative morphisms. -/ +class nonneg_hom_class (F : Type*) (α β : out_param $ Type*) [has_zero β] [has_le β] + extends fun_like F α (λ _, β) := +(map_nonneg (f : F) : ∀ a, 0 ≤ f a) + +/-- `subadditive_hom_class F α β` states that `F` is a type of subadditive morphisms. -/ +class subadditive_hom_class (F : Type*) (α β : out_param $ Type*) [has_add α] [has_add β] [has_le β] + extends fun_like F α (λ _, β) := +(map_add_le_add (f : F) : ∀ a b, f (a + b) ≤ f a + f b) + +/-- `submultiplicative_hom_class F α β` states that `F` is a type of submultiplicative morphisms. -/ +@[to_additive subadditive_hom_class] +class submultiplicative_hom_class (F : Type*) (α β : out_param $ Type*) [has_mul α] [has_mul β] + [has_le β] extends fun_like F α (λ _, β) := +(map_mul_le_mul (f : F) : ∀ a b, f (a * b) ≤ f a * f b) + +/-- `map_add_le_class F α β` states that `F` is a type of subadditive morphisms. -/ +@[to_additive subadditive_hom_class] +class mul_le_add_hom_class (F : Type*) (α β : out_param $ Type*) [has_mul α] [has_add β] [has_le β] + extends fun_like F α (λ _, β) := +(map_mul_le_add (f : F) : ∀ a b, f (a * b) ≤ f a + f b) + +export nonneg_hom_class (map_nonneg) +export subadditive_hom_class (map_add_le_add) +export submultiplicative_hom_class (map_mul_le_mul) +export mul_le_add_hom_class (map_mul_le_add) + +attribute [simp] map_nonneg + +@[to_additive] lemma le_map_add_map_div [group α] [add_comm_semigroup β] [has_le β] + [mul_le_add_hom_class F α β] (f : F) (a b : α) : f a ≤ f b + f (a / b) := +by simpa only [add_comm, div_mul_cancel'] using map_mul_le_add f (a / b) b diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index 6235a77f2f557..2b1855b092fd1 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -405,21 +405,21 @@ begin obtain hp | hp := {r : ℝ | 0 < r ∧ x ∈ r • p.ball 0 1}.eq_empty_or_nonempty, { rw [gauge, hp, real.Inf_empty], by_contra, - have hpx : 0 < p x := (p.nonneg x).lt_of_ne h, + have hpx : 0 < p x := (map_nonneg _ _).lt_of_ne h, have hpx₂ : 0 < 2 * p x := mul_pos zero_lt_two hpx, refine hp.subset ⟨hpx₂, (2 * p x)⁻¹ • x, _, smul_inv_smul₀ hpx₂.ne' _⟩, - rw [p.mem_ball_zero, p.smul, real.norm_eq_abs, abs_of_pos (inv_pos.2 hpx₂), inv_mul_lt_iff hpx₂, - mul_one], + rw [p.mem_ball_zero, map_smul_eq_mul, real.norm_eq_abs, abs_of_pos (inv_pos.2 hpx₂), + inv_mul_lt_iff hpx₂, mul_one], exact lt_mul_of_one_lt_left hpx one_lt_two }, refine is_glb.cInf_eq ⟨λ r, _, λ r hr, le_of_forall_pos_le_add $ λ ε hε, _⟩ hp, { rintro ⟨hr, y, hy, rfl⟩, rw p.mem_ball_zero at hy, - rw [p.smul, real.norm_eq_abs, abs_of_pos hr], + rw [map_smul_eq_mul, real.norm_eq_abs, abs_of_pos hr], exact mul_le_of_le_one_right hr.le hy.le }, - { have hpε : 0 < p x + ε := add_pos_of_nonneg_of_pos (p.nonneg _) hε, + { have hpε : 0 < p x + ε := add_pos_of_nonneg_of_pos (map_nonneg _ _) hε, refine hr ⟨hpε, (p x + ε)⁻¹ • x, _, smul_inv_smul₀ hpε.ne' _⟩, - rw [p.mem_ball_zero, p.smul, real.norm_eq_abs, abs_of_pos (inv_pos.2 hpε), inv_mul_lt_iff hpε, - mul_one], + rw [p.mem_ball_zero, map_smul_eq_mul, real.norm_eq_abs, abs_of_pos (inv_pos.2 hpε), + inv_mul_lt_iff hpε, mul_one], exact lt_add_of_pos_right _ hε } end diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 3f1cba49ad904..6afbaf2ddc9d2 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -128,12 +128,12 @@ lemma basis_sets_smul_right (v : E) (U : set E) begin rcases p.basis_sets_iff.mp hU with ⟨s, r, hr, hU⟩, rw [hU, filter.eventually_iff], - simp_rw [(s.sup p).mem_ball_zero, (s.sup p).smul], + simp_rw [(s.sup p).mem_ball_zero, map_smul_eq_mul], by_cases h : 0 < (s.sup p) v, { simp_rw (lt_div_iff h).symm, rw ←_root_.ball_zero_eq, exact metric.ball_mem_nhds 0 (div_pos hr h) }, - simp_rw [le_antisymm (not_lt.mp h) ((s.sup p).nonneg v), mul_zero, hr], + simp_rw [le_antisymm (not_lt.mp h) (map_nonneg _ v), mul_zero, hr], exact is_open.mem_nhds is_open_univ (mem_univ 0), end @@ -228,11 +228,8 @@ lemma is_bounded_sup {p : ι → seminorm 𝕜 E} {q : ι' → seminorm 𝕜 F} ∃ (C : ℝ≥0) (s : finset ι), 0 < C ∧ (s'.sup q).comp f ≤ C • (s.sup p) := begin classical, - by_cases hs' : ¬s'.nonempty, - { refine ⟨1, ∅, zero_lt_one, _⟩, - rw [finset.not_nonempty_iff_eq_empty.mp hs', finset.sup_empty, seminorm.bot_eq_zero, zero_comp], - exact seminorm.nonneg _ }, - rw not_not at hs', + obtain rfl | hs' := s'.eq_empty_or_nonempty, + { exact ⟨1, ∅, zero_lt_one, by simp [seminorm.bot_eq_zero]⟩ }, choose fₛ fC hf using hf, use [s'.card • s'.sup fC, finset.bUnion s' fₛ], split, diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index 3442f5374e264..67890137b6ace 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 María Inés de Frutos-Fernández, Yaël Dillies. All rights Released under Apache 2.0 license as described in the file LICENSE. Authors: María Inés de Frutos-Fernández, Yaël Dillies -/ +import algebra.order.hom.basic import data.real.nnreal /-! @@ -51,35 +52,75 @@ structure group_seminorm (G : Type*) [group G] := attribute [nolint doc_blame] add_group_seminorm.to_zero_hom +/-- `add_group_seminorm_class F α` states that `F` is a type of seminorms on the additive group `α`. + +You should extend this class when you extend `add_group_seminorm`. -/ +class add_group_seminorm_class (F : Type*) (α : out_param $ Type*) [add_group α] + extends subadditive_hom_class F α ℝ := +(map_zero (f : F) : f 0 = 0) +(map_neg_eq_map (f : F) (a : α) : f (-a) = f a) + +/-- `group_seminorm_class F α` states that `F` is a type of seminorms on the group `α`. + +You should extend this class when you extend `group_seminorm`. -/ +@[to_additive] +class group_seminorm_class (F : Type*) (α : out_param $ Type*) [group α] + extends mul_le_add_hom_class F α ℝ := +(map_one_eq_zero (f : F) : f 1 = 0) +(map_inv_eq_map (f : F) (a : α) : f a⁻¹ = f a) + +attribute [to_additive] group_seminorm_class.to_mul_le_add_hom_class + +export add_group_seminorm_class (map_neg_eq_map) +export group_seminorm_class (map_one_eq_zero map_inv_eq_map) + +attribute [simp, to_additive map_zero] map_one_eq_zero +attribute [simp] map_neg_eq_map +attribute [simp, to_additive] map_inv_eq_map + +@[priority 100] -- See note [lower instance priority] +instance add_group_seminorm_class.to_zero_hom_class [add_group E] [add_group_seminorm_class F E] : + zero_hom_class F E ℝ := +{ ..‹add_group_seminorm_class F E› } + +section group +variables [group E] [group_seminorm_class F E] (f : F) (x y : E) +include E + +@[to_additive] lemma map_div_le_add : f (x / y) ≤ f x + f y := +by { rw [div_eq_mul_inv, ←map_inv_eq_map f y], exact map_mul_le_add _ _ _ } + +@[to_additive] lemma map_div_rev : f (x / y) = f (y / x) := by rw [←inv_div, map_inv_eq_map] + +@[to_additive] lemma le_map_add_map_div' : f x ≤ f y + f (y / x) := +by simpa only [add_comm, map_div_rev, div_mul_cancel'] using map_mul_le_add f (x / y) y + +end group + +@[to_additive, priority 100] -- See note [lower instance priority] +instance group_seminorm_class.to_nonneg_hom_class [group E] [group_seminorm_class F E] : + nonneg_hom_class F E ℝ := +{ map_nonneg := λ f a, nonneg_of_mul_nonneg_right + (by { rw [two_mul, ←map_one_eq_zero f, ←div_self' a], exact map_div_le_add _ _ _ }) two_pos, + ..‹group_seminorm_class F E› } + namespace group_seminorm section group -variables [group E] [group F] [group G] +variables [group E] [group F] [group G] {p q : group_seminorm E} -@[to_additive] instance fun_like : fun_like (group_seminorm E) E (λ _, ℝ) := +@[to_additive] instance group_seminorm_class : group_seminorm_class (group_seminorm E) E := { coe := λ f, f.to_fun, - coe_injective' := λ f g h, by cases f; cases g; congr' } + coe_injective' := λ f g h, by cases f; cases g; congr', + map_one_eq_zero := λ f, f.map_one', + map_mul_le_add := λ f, f.mul_le', + map_inv_eq_map := λ f, f.inv' } /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ @[to_additive "Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. "] instance : has_coe_to_fun (group_seminorm E) (λ _, E → ℝ) := ⟨to_fun⟩ -@[ext, to_additive] lemma ext {p q : group_seminorm E} : (∀ x, (p : E → ℝ) x = q x) → p = q := -fun_like.ext p q - -variables (p q : group_seminorm E) (x y : E) (r : ℝ) - -@[simp, to_additive] protected lemma map_one : p 1 = 0 := p.map_one' -@[to_additive] protected lemma mul_le : p (x * y) ≤ p x + p y := p.mul_le' _ _ -@[simp, to_additive] protected lemma inv : p x⁻¹ = p x := p.inv' _ - -@[to_additive] protected lemma div_le : p (x / y) ≤ p x + p y := -by { rw [div_eq_mul_inv, ←p.inv y], exact p.mul_le _ _ } - -@[to_additive] protected lemma nonneg : 0 ≤ p x := -nonneg_of_mul_nonneg_right (by { rw [two_mul, ←p.map_one, ←div_self' x], apply p.div_le }) two_pos - -@[to_additive] lemma div_rev : p (x / y) = p (y / x) := by rw [←inv_div, p.inv] +@[ext, to_additive] lemma ext : (∀ x, p x = q x) → p = q := fun_like.ext p q @[to_additive] instance : partial_order (group_seminorm E) := partial_order.lift _ fun_like.coe_injective @@ -87,8 +128,6 @@ partial_order.lift _ fun_like.coe_injective @[to_additive] lemma le_def : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl @[to_additive] lemma lt_def : p < q ↔ (p : E → ℝ) < q := iff.rfl -variables {p q} - @[simp, to_additive] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl @[simp, to_additive] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl @@ -108,10 +147,10 @@ variables (p q) (f : F →* E) @[to_additive] instance : has_add (group_seminorm E) := ⟨λ p q, { to_fun := λ x, p x + q x, - map_one' := by rw [p.map_one, q.map_one, zero_add], - mul_le' := λ _ _, (add_le_add (p.mul_le _ _) $ q.mul_le _ _).trans_eq $ + map_one' := by rw [map_one_eq_zero p, map_one_eq_zero q, zero_add], + mul_le' := λ _ _, (add_le_add (map_mul_le_add p _ _) $ map_mul_le_add q _ _).trans_eq $ add_add_add_comm _ _ _ _, - inv' := λ x, by rw [p.inv, q.inv] }⟩ + inv' := λ x, by rw [map_inv_eq_map p, map_inv_eq_map q] }⟩ @[simp, to_additive] lemma coe_add : ⇑(p + q) = p + q := rfl @[simp, to_additive] lemma add_apply (x : E) : (p + q) x = p x + q x := rfl @@ -121,11 +160,12 @@ variables (p q) (f : F →* E) @[to_additive] noncomputable instance : has_sup (group_seminorm E) := ⟨λ p q, { to_fun := p ⊔ q, - map_one' := by rw [pi.sup_apply, ←p.map_one, sup_eq_left, p.map_one, q.map_one], + map_one' := + by rw [pi.sup_apply, ←map_one_eq_zero p, sup_eq_left, map_one_eq_zero p, map_one_eq_zero q], mul_le' := λ x y, sup_le - ((p.mul_le x y).trans $ add_le_add le_sup_left le_sup_left) - ((q.mul_le x y).trans $ add_le_add le_sup_right le_sup_right), - inv' := λ x, by rw [pi.sup_apply, pi.sup_apply, p.inv, q.inv] }⟩ + ((map_mul_le_add p x y).trans $ add_le_add le_sup_left le_sup_left) + ((map_mul_le_add q x y).trans $ add_le_add le_sup_right le_sup_right), + inv' := λ x, by rw [pi.sup_apply, pi.sup_apply, map_inv_eq_map p, map_inv_eq_map q] }⟩ @[simp, to_additive] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl @[simp, to_additive] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl @@ -138,14 +178,14 @@ fun_like.coe_injective.semilattice_sup _ coe_sup additive group seminorm."] def comp (p : group_seminorm E) (f : F →* E) : group_seminorm F := { to_fun := λ x, p (f x), - map_one' := by rw [f.map_one, p.map_one], - mul_le' := λ _ _, (congr_arg p $ f.map_mul _ _).trans_le $ p.mul_le _ _, - inv' := λ x, by rw [map_inv, p.inv] } + map_one' := by rw [f.map_one, map_one_eq_zero p], + mul_le' := λ _ _, (congr_arg p $ f.map_mul _ _).trans_le $ map_mul_le_add p _ _, + inv' := λ x, by rw [map_inv, map_inv_eq_map p] } @[simp, to_additive] lemma coe_comp : ⇑(p.comp f) = p ∘ f := rfl @[simp, to_additive] lemma comp_apply (x : F) : (p.comp f) x = p (f x) := rfl @[simp, to_additive] lemma comp_id : p.comp (monoid_hom.id _) = p := ext $ λ _, rfl -@[simp, to_additive] lemma comp_zero : p.comp (1 : F →* E) = 0 := ext $ λ _, p.map_one +@[simp, to_additive] lemma comp_zero : p.comp (1 : F →* E) = 0 := ext $ λ _, map_one_eq_zero p @[simp, to_additive] lemma zero_comp : (0 : group_seminorm E).comp f = 0 := ext $ λ _, rfl @[to_additive] lemma comp_assoc (g : F →* E) (f : G →* F) : p.comp (g.comp f) = (p.comp g).comp f := @@ -162,75 +202,61 @@ end group section comm_group variables [comm_group E] [comm_group F] (p q : group_seminorm E) (x y : E) -/-- The direct path from `1` to `y` is shorter than the path with `x` "inserted" in between. -/ -@[to_additive "The direct path from `0` to `y` is shorter than the path with `x` \"inserted\" in -between."] -lemma le_insert : p y ≤ p x + p (x / y) := (p.div_le _ _).trans_eq' $ by rw div_div_cancel - -/-- The direct path from `1` to `x` is shorter than the path with `y` "inserted" in between. -/ -@[to_additive "The direct path from `0` to `x` is shorter than the path with `y` \"inserted\" in -between."] -lemma le_insert' : p x ≤ p y + p (x / y) := by { rw div_rev, exact le_insert _ _ _ } - @[to_additive] lemma comp_mul_le (f g : F →* E) : p.comp (f * g) ≤ p.comp f + p.comp g := -λ _, p.mul_le _ _ +λ _, map_mul_le_add p _ _ @[to_additive] private lemma mul_bdd_below_range_add {p q : group_seminorm E} {x : E} : bdd_below (range $ λ y, p y + q (x / y)) := -⟨0, by { rintro _ ⟨x, rfl⟩, exact add_nonneg (p.nonneg _) (q.nonneg _) }⟩ +⟨0, by { rintro _ ⟨x, rfl⟩, exact add_nonneg (map_nonneg p _) (map_nonneg q _) }⟩ @[to_additive] noncomputable instance : has_inf (group_seminorm E) := ⟨λ p q, { to_fun := λ x, ⨅ y, p y + q (x / y), map_one' := cinfi_eq_of_forall_ge_of_forall_gt_exists_lt - (λ x, add_nonneg (p.nonneg _) (q.nonneg _)) - (λ r hr, ⟨1, by rwa [div_one, p.map_one, q.map_one, add_zero]⟩), + (λ x, add_nonneg (map_nonneg p _) (map_nonneg q _)) + (λ r hr, ⟨1, by rwa [div_one, map_one_eq_zero p, map_one_eq_zero q, add_zero]⟩), mul_le' := λ x y, le_cinfi_add_cinfi $ λ u v, begin refine cinfi_le_of_le mul_bdd_below_range_add (u * v) _, rw [mul_div_mul_comm, add_add_add_comm], - exact add_le_add (p.mul_le _ _) (q.mul_le _ _), + exact add_le_add (map_mul_le_add p _ _) (map_mul_le_add q _ _), end, - inv' := λ x, (inv_surjective.infi_comp _).symm.trans $ by simp_rw [p.inv, ←inv_div', q.inv] }⟩ + inv' := λ x, (inv_surjective.infi_comp _).symm.trans $ + by simp_rw [map_inv_eq_map p, ←inv_div', map_inv_eq_map q] }⟩ @[simp, to_additive] lemma inf_apply : (p ⊓ q) x = ⨅ y, p y + q (x / y) := rfl @[to_additive] noncomputable instance : lattice (group_seminorm E) := { inf := (⊓), inf_le_left := λ p q x, cinfi_le_of_le mul_bdd_below_range_add x $ - by rw [div_self', q.map_one, add_zero], + by rw [div_self', map_one_eq_zero q, add_zero], inf_le_right := λ p q x, cinfi_le_of_le mul_bdd_below_range_add (1 : E) $ - by simp only [div_one, p.map_one, zero_add], - le_inf := λ a b c hb hc x, le_cinfi $ λ u, (a.le_insert' _ _).trans $ add_le_add (hb _) (hc _), + by simp only [div_one, map_one_eq_zero p, zero_add], + le_inf := λ a b c hb hc x, le_cinfi $ λ u, (le_map_add_map_div a _ _).trans $ + add_le_add (hb _) (hc _), ..group_seminorm.semilattice_sup } end comm_group end group_seminorm -namespace add_group_seminorm -variables [add_group E] (p : add_group_seminorm E) (x y : E) (r : ℝ) - -instance zero_hom_class : zero_hom_class (add_group_seminorm E) E ℝ := -{ coe := λ f, f.to_fun, - coe_injective' := λ f g h, by cases f; cases g; congr', - map_zero := λ f, f.map_zero' } - /- TODO: All the following ought to be automated using `to_additive`. The problem is that it doesn't see that `has_smul R ℝ` should be fixed because `ℝ` is fixed. -/ -variables [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] +namespace add_group_seminorm +variables [add_group E] [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] + (p : add_group_seminorm E) /-- Any action on `ℝ` which factors through `ℝ≥0` applies to an `add_group_seminorm`. -/ instance : has_smul R (add_group_seminorm E) := ⟨λ r p, { to_fun := λ x, r • p x, map_zero' := by simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul, - p.map_zero, mul_zero], + map_zero, mul_zero], add_le' := λ _ _, begin simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul], - exact (mul_le_mul_of_nonneg_left (p.add_le _ _) (nnreal.coe_nonneg _)).trans_eq + exact (mul_le_mul_of_nonneg_left (map_add_le_add _ _ _) $ nnreal.coe_nonneg _).trans_eq (mul_add _ _ _), end, - neg' := λ x, by rw p.neg }⟩ + neg' := λ x, by rw map_neg_eq_map }⟩ @[simp] lemma coe_smul (r : R) (p : add_group_seminorm E) : ⇑(r • p) = r • p := rfl @[simp] lemma smul_apply (r : R) (p : add_group_seminorm E) (x : E) : (r • p) x = r • p x := rfl @@ -256,13 +282,13 @@ variables [group E] [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ ⟨λ r p, { to_fun := λ x, r • p x, map_one' := by simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul, - p.map_one, mul_zero], + map_one_eq_zero p, mul_zero], mul_le' := λ _ _, begin simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul], - exact (mul_le_mul_of_nonneg_left (p.mul_le _ _) $ nnreal.coe_nonneg _).trans_eq + exact (mul_le_mul_of_nonneg_left (map_mul_le_add p _ _) $ nnreal.coe_nonneg _).trans_eq (mul_add _ _ _), end, - inv' := λ x, by rw p.inv }⟩ + inv' := λ x, by rw map_inv_eq_map p }⟩ @[to_additive add_group_seminorm.is_scalar_tower] instance [has_smul R' ℝ] [has_smul R' ℝ≥0] [is_scalar_tower R' ℝ≥0 ℝ] [has_smul R R'] diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index fd30d228d9231..651205d77c805 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -48,6 +48,18 @@ structure seminorm (𝕜 : Type*) (E : Type*) [semi_normed_ring 𝕜] [add_group attribute [nolint doc_blame] seminorm.to_add_group_seminorm +/-- `seminorm_class F 𝕜 E` states that `F` is a type of seminorms on the `𝕜`-module E. + +You should extend this class when you extend `seminorm`. -/ +class seminorm_class (F : Type*) (𝕜 E : out_param $ Type*) [semi_normed_ring 𝕜] [add_group E] + [has_smul 𝕜 E] extends add_group_seminorm_class F E := +(map_smul_eq_mul (f : F) (a : 𝕜) (x : E) : f (a • x) = ∥a∥ * f x) + +export seminorm_class (map_smul_eq_mul) + +-- `𝕜` is an `out_param`, so this is a false positive. +attribute [nolint dangerous_instance] seminorm_class.to_add_group_seminorm_class + section of /-- Alternative constructor for a `seminorm` on an `add_comm_group E` that is a module over a @@ -92,13 +104,16 @@ variables [add_group E] section has_smul variables [has_smul 𝕜 E] -instance zero_hom_class : zero_hom_class (seminorm 𝕜 E) E ℝ := +instance seminorm_class : seminorm_class (seminorm 𝕜 E) 𝕜 E := { coe := λ f, f.to_fun, coe_injective' := λ f g h, by cases f; cases g; congr', - map_zero := λ f, f.map_zero' } + map_zero := λ f, f.map_zero', + map_add_le_add := λ f, f.add_le', + map_neg_eq_map := λ f, f.neg', + map_smul_eq_mul := λ f, f.smul' } /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ -instance : has_coe_to_fun (seminorm 𝕜 E) (λ _, E → ℝ) := ⟨λ p, p.to_fun⟩ +instance : has_coe_to_fun (seminorm 𝕜 E) (λ _, E → ℝ) := fun_like.has_coe_to_fun @[ext] lemma ext {p q : seminorm 𝕜 E} (h : ∀ x, (p : E → ℝ) x = q x) : p = q := fun_like.ext p q h @@ -114,11 +129,6 @@ instance : inhabited (seminorm 𝕜 E) := ⟨0⟩ variables (p : seminorm 𝕜 E) (c : 𝕜) (x y : E) (r : ℝ) -protected lemma nonneg : 0 ≤ p x := p.to_add_group_seminorm.nonneg _ -protected lemma map_zero : p 0 = 0 := p.map_zero' -protected lemma smul : p (c • x) = ∥c∥ * p x := p.smul' _ _ -protected lemma add_le : p (x + y) ≤ p x + p y := p.add_le' _ _ - /-- Any action on `ℝ` which factors through `ℝ≥0` applies to a seminorm. -/ instance [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] : has_smul R (seminorm 𝕜 E) := @@ -126,7 +136,7 @@ instance [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] : { to_fun := λ x, r • p x, smul' := λ _ _, begin simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul], - rw [p.smul, mul_left_comm], + rw [map_smul_eq_mul, mul_left_comm], end, ..(r • p.to_add_group_seminorm) }} @@ -145,7 +155,7 @@ lemma coe_smul [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 instance : has_add (seminorm 𝕜 E) := { add := λ p q, { to_fun := λ x, p x + q x, - smul' := λ a x, by simp only [p.smul, q.smul, mul_add], + smul' := λ a x, by simp only [map_smul_eq_mul, map_smul_eq_mul, mul_add], ..(p.to_add_group_seminorm + q.to_add_group_seminorm) }} lemma coe_add (p q : seminorm 𝕜 E) : ⇑(p + q) = p + q := rfl @@ -187,7 +197,7 @@ instance [semiring R] [module R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ noncomputable instance : has_sup (seminorm 𝕜 E) := { sup := λ p q, { to_fun := p ⊔ q, - smul' := λ x v, (congr_arg2 max (p.smul x v) (q.smul x v)).trans $ + smul' := λ x v, (congr_arg2 max (map_smul_eq_mul p x v) (map_smul_eq_mul q x v)).trans $ (mul_max_of_nonneg _ _ $ norm_nonneg x).symm, ..(p.to_add_group_seminorm ⊔ q.to_add_group_seminorm) } } @@ -223,7 +233,7 @@ variables [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] /-- Composition of a seminorm with a linear map is a seminorm. -/ def comp (p : seminorm 𝕜 F) (f : E →ₗ[𝕜] F) : seminorm 𝕜 E := { to_fun := λ x, p (f x), - smul' := λ _ _, (congr_arg p (f.map_smul _ _)).trans (p.smul _ _), + smul' := λ _ _, (congr_arg p (f.map_smul _ _)).trans (map_smul_eq_mul p _ _), ..(p.to_add_group_seminorm.comp f.to_add_monoid_hom) } lemma coe_comp (p : seminorm 𝕜 F) (f : E →ₗ[𝕜] F) : ⇑(p.comp f) = p ∘ f := rfl @@ -247,7 +257,7 @@ lemma add_comp (p q : seminorm 𝕜 F) (f : E →ₗ[𝕜] F) : (p + q).comp f = ext $ λ _, rfl lemma comp_add_le (p : seminorm 𝕜 F) (f g : E →ₗ[𝕜] F) : p.comp (f + g) ≤ p.comp f + p.comp g := -λ _, p.add_le _ _ +λ _, map_add_le_add p _ _ lemma smul_comp (p : seminorm 𝕜 F) (f : E →ₗ[𝕜] F) (c : R) : (c • p).comp f = c • (p.comp f) := ext $ λ _, rfl @@ -256,36 +266,10 @@ lemma comp_mono {p : seminorm 𝕜 F} {q : seminorm 𝕜 F} (f : E →ₗ[𝕜] p.comp f ≤ q.comp f := λ _, hp _ /-- The composition as an `add_monoid_hom`. -/ -@[simps] def pullback (f : E →ₗ[𝕜] F) : add_monoid_hom (seminorm 𝕜 F) (seminorm 𝕜 E) := +@[simps] def pullback (f : E →ₗ[𝕜] F) : seminorm 𝕜 F →+ seminorm 𝕜 E := ⟨λ p, p.comp f, zero_comp f, λ p q, add_comp p q f⟩ -section -variables (p : seminorm 𝕜 E) - -@[simp] -protected lemma neg (x : E) : p (-x) = p x := -by rw [←neg_one_smul 𝕜, seminorm.smul, norm_neg, ←seminorm.smul, one_smul] - -protected lemma sub_le (x y : E) : p (x - y) ≤ p x + p y := -calc - p (x - y) - = p (x + -y) : by rw sub_eq_add_neg - ... ≤ p x + p (-y) : p.add_le x (-y) - ... = p x + p y : by rw p.neg - -lemma sub_rev (x y : E) : p (x - y) = p (y - x) := by rw [←neg_sub, p.neg] - -/-- The direct path from 0 to y is shorter than the path with x "inserted" in between. -/ -lemma le_insert (x y : E) : p y ≤ p x + p (x - y) := -calc p y = p (x - (x - y)) : by rw sub_sub_cancel -... ≤ p x + p (x - y) : p.sub_le _ _ - -/-- The direct path from 0 to x is shorter than the path with y "inserted" in between. -/ -lemma le_insert' (x y : E) : p x ≤ p y + p (x - y) := by { rw sub_rev, exact le_insert _ _ _ } - -end - -instance : order_bot (seminorm 𝕜 E) := ⟨0, seminorm.nonneg⟩ +instance : order_bot (seminorm 𝕜 E) := ⟨0, map_nonneg⟩ @[simp] lemma coe_bot : ⇑(⊥ : seminorm 𝕜 E) = 0 := rfl @@ -297,11 +281,11 @@ begin simp_rw [le_def, pi.le_def, coe_smul], intros x, simp_rw [pi.smul_apply, nnreal.smul_def, smul_eq_mul], - exact mul_le_mul hab (hpq x) (p.nonneg x) (nnreal.coe_nonneg b), + exact mul_le_mul hab (hpq x) (map_nonneg p x) (nnreal.coe_nonneg b), end lemma finset_sup_apply (p : ι → seminorm 𝕜 E) (s : finset ι) (x : E) : - s.sup p x = ↑(s.sup (λ i, ⟨p i x, (p i).nonneg x⟩) : ℝ≥0) := + s.sup p x = ↑(s.sup (λ i, ⟨p i x, map_nonneg (p i) x⟩) : ℝ≥0) := begin induction s using finset.cons_induction_on with a s ha ih, { rw [finset.sup_empty, finset.sup_empty, coe_bot, _root_.bot_eq_zero, pi.zero_apply, @@ -344,20 +328,20 @@ variables [semi_normed_comm_ring 𝕜] [add_comm_group E] [add_comm_group F] [mo lemma comp_smul (p : seminorm 𝕜 F) (f : E →ₗ[𝕜] F) (c : 𝕜) : p.comp (c • f) = ∥c∥₊ • p.comp f := -ext $ λ _, by rw [comp_apply, smul_apply, linear_map.smul_apply, p.smul, nnreal.smul_def, +ext $ λ _, by rw [comp_apply, smul_apply, linear_map.smul_apply, map_smul_eq_mul, nnreal.smul_def, coe_nnnorm, smul_eq_mul, comp_apply] lemma comp_smul_apply (p : seminorm 𝕜 F) (f : E →ₗ[𝕜] F) (c : 𝕜) (x : E) : - p.comp (c • f) x = ∥c∥ * p (f x) := p.smul _ _ + p.comp (c • f) x = ∥c∥ * p (f x) := map_smul_eq_mul p _ _ end semi_normed_comm_ring section normed_field -variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] +variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {p q : seminorm 𝕜 E} {x : E} -private lemma bdd_below_range_add (x : E) (p q : seminorm 𝕜 E) : - bdd_below (range (λ (u : E), p u + q (x - u))) := -by { use 0, rintro _ ⟨x, rfl⟩, exact add_nonneg (p.nonneg _) (q.nonneg _) } +/-- Auxiliary lemma to show that the infimum of seminorms is well-defined. -/ +lemma bdd_below_range_add : bdd_below (range $ λ u, p u + q (x - u)) := +⟨0, by { rintro _ ⟨x, rfl⟩, exact add_nonneg (map_nonneg p _) (map_nonneg q _) }⟩ noncomputable instance : has_inf (seminorm 𝕜 E) := { inf := λ p q, @@ -368,9 +352,10 @@ noncomputable instance : has_inf (seminorm 𝕜 E) := obtain rfl | ha := eq_or_ne a 0, { rw [norm_zero, zero_mul, zero_smul], refine cinfi_eq_of_forall_ge_of_forall_gt_exists_lt - (λ i, add_nonneg (p.nonneg _) (q.nonneg _)) + (λ i, add_nonneg (map_nonneg p _) (map_nonneg q _)) (λ x hx, ⟨0, by rwa [map_zero, sub_zero, map_zero, add_zero]⟩) }, - simp_rw [real.mul_infi_of_nonneg (norm_nonneg a), mul_add, ←p.smul, ←q.smul, smul_sub], + simp_rw [real.mul_infi_of_nonneg (norm_nonneg a), mul_add, ←map_smul_eq_mul p, + ←map_smul_eq_mul q, smul_sub], refine function.surjective.infi_congr ((•) a⁻¹ : E → E) (λ u, ⟨a • u, inv_smul_smul₀ ha u⟩) (λ u, _), rw smul_inv_smul₀ ha @@ -381,16 +366,12 @@ noncomputable instance : has_inf (seminorm 𝕜 E) := noncomputable instance : lattice (seminorm 𝕜 E) := { inf := (⊓), - inf_le_left := λ p q x, begin - apply cinfi_le_of_le (bdd_below_range_add _ _ _) x, - simp only [sub_self, map_zero, add_zero], - end, - inf_le_right := λ p q x, begin - apply cinfi_le_of_le (bdd_below_range_add _ _ _) (0:E), - simp only [sub_self, map_zero, zero_add, sub_zero], - end, + inf_le_left := λ p q x, cinfi_le_of_le bdd_below_range_add x $ + by simp only [sub_self, map_zero, add_zero], + inf_le_right := λ p q x, cinfi_le_of_le bdd_below_range_add 0 $ + by simp only [sub_self, map_zero, zero_add, sub_zero], le_inf := λ a b c hab hac x, - le_cinfi $ λ u, le_trans (a.le_insert' _ _) (add_le_add (hab _) (hac _)), + le_cinfi $ λ u, (le_map_add_map_sub a _ _).trans $ add_le_add (hab _) (hac _), ..seminorm.semilattice_sup } lemma smul_inf [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] @@ -461,7 +442,7 @@ lemma ball_add_ball_subset (p : seminorm 𝕜 E) (r₁ r₂ : ℝ) (x₁ x₂ : begin rintros x ⟨y₁, y₂, hy₁, hy₂, rfl⟩, rw [mem_ball, add_sub_add_comm], - exact (p.add_le _ _).trans_lt (add_lt_add hy₁ hy₂), + exact (map_add_le_add p _ _).trans_lt (add_lt_add hy₁ hy₂), end end has_smul @@ -486,7 +467,7 @@ begin ext x, simp only [mem_ball, sub_zero, mem_preimage, mem_ball_zero_iff], rw real.norm_of_nonneg, - exact p.nonneg _, + exact map_nonneg p _, end @[simp] lemma ball_bot {r : ℝ} (x : E) (hr : 0 < r) : @@ -497,8 +478,8 @@ ball_zero' x hr lemma balanced_ball_zero (r : ℝ) : balanced 𝕜 (ball p 0 r) := begin rintro a ha x ⟨y, hy, hx⟩, - rw [mem_ball_zero, ←hx, p.smul], - calc _ ≤ p y : mul_le_of_le_one_left (p.nonneg _) ha + rw [mem_ball_zero, ←hx, map_smul_eq_mul], + calc _ ≤ p y : mul_le_of_le_one_left (map_nonneg p _) ha ... < r : by rwa mem_ball_zero at hy, end @@ -524,15 +505,16 @@ begin intros x hx, rw set.mem_smul at hx, rcases hx with ⟨a, y, ha, hy, hx⟩, - rw [←hx, mem_ball_zero, seminorm.smul], - exact mul_lt_mul'' (mem_ball_zero_iff.mp ha) (p.mem_ball_zero.mp hy) (norm_nonneg a) (p.nonneg y), + rw [←hx, mem_ball_zero, map_smul_eq_mul], + exact mul_lt_mul'' (mem_ball_zero_iff.mp ha) (p.mem_ball_zero.mp hy) (norm_nonneg a) + (map_nonneg p y), end @[simp] lemma ball_eq_emptyset (p : seminorm 𝕜 E) {x : E} {r : ℝ} (hr : r ≤ 0) : p.ball x r = ∅ := begin ext, rw [seminorm.mem_ball, set.mem_empty_eq, iff_false, not_lt], - exact hr.trans (p.nonneg _), + exact hr.trans (map_nonneg p _), end end module @@ -550,11 +532,11 @@ begin rw [set.mem_smul_set, seminorm.mem_ball_zero], split; intro h, { rcases h with ⟨y, hy, h⟩, - rw [←h, seminorm.smul], + rw [←h, map_smul_eq_mul], rw seminorm.mem_ball_zero at hy, exact (mul_lt_mul_left hk).mpr hy }, refine ⟨k⁻¹ • x, _, _⟩, - { rw [seminorm.mem_ball_zero, seminorm.smul, norm_inv, ←(mul_lt_mul_left hk), + { rw [seminorm.mem_ball_zero, map_smul_eq_mul, norm_inv, ←(mul_lt_mul_left hk), ←mul_assoc, ←(div_eq_mul_inv ∥k∥ ∥k∥), div_self (ne_of_gt hk), one_mul], exact h}, rw [←smul_assoc, smul_eq_mul, ←div_eq_mul_inv, div_self (norm_pos_iff.mp hk), one_smul], @@ -582,11 +564,11 @@ protected lemma absorbent_ball_zero (hr : 0 < r) : absorbent 𝕜 (ball p (0 : E begin rw absorbent_iff_nonneg_lt, rintro x, - have hxr : 0 ≤ p x/r := div_nonneg (p.nonneg _) hr.le, + have hxr : 0 ≤ p x/r := div_nonneg (map_nonneg p _) hr.le, refine ⟨p x/r, hxr, λ a ha, _⟩, have ha₀ : 0 < ∥a∥ := hxr.trans_lt ha, refine ⟨a⁻¹ • x, _, smul_inv_smul₀ (norm_pos_iff.1 ha₀) x⟩, - rwa [mem_ball_zero, p.smul, norm_inv, inv_mul_lt_iff ha₀, ←div_lt_iff hr], + rwa [mem_ball_zero, map_smul_eq_mul, norm_inv, inv_mul_lt_iff ha₀, ←div_lt_iff hr], end /-- Seminorm-balls containing the origin are absorbent. -/ @@ -594,7 +576,7 @@ protected lemma absorbent_ball (hpr : p x < r) : absorbent 𝕜 (ball p x r) := begin refine (p.absorbent_ball_zero $ sub_pos.2 hpr).subset (λ y hy, _), rw p.mem_ball_zero at hy, - exact p.mem_ball.2 ((p.sub_le _ _).trans_lt $ add_lt_of_lt_sub_right hy), + exact p.mem_ball.2 ((map_sub_le_add p _ _).trans_lt $ add_lt_of_lt_sub_right hy), end lemma symmetric_ball_zero (r : ℝ) (hx : x ∈ ball p 0 r) : -x ∈ ball p 0 r := @@ -603,13 +585,13 @@ balanced_ball_zero p r (-1) (by rw [norm_neg, norm_one]) ⟨x, hx, by rw [neg_sm @[simp] lemma neg_ball (p : seminorm 𝕜 E) (r : ℝ) (x : E) : -ball p x r = ball p (-x) r := -by { ext, rw [mem_neg, mem_ball, mem_ball, ←neg_add', sub_neg_eq_add, p.neg], } +by { ext, rw [mem_neg, mem_ball, mem_ball, ←neg_add', sub_neg_eq_add, map_neg_eq_map] } @[simp] lemma smul_ball_preimage (p : seminorm 𝕜 E) (y : E) (r : ℝ) (a : 𝕜) (ha : a ≠ 0) : ((•) a) ⁻¹' p.ball y r = p.ball (a⁻¹ • y) (r / ∥a∥) := set.ext $ λ _, by rw [mem_preimage, mem_ball, mem_ball, - lt_div_iff (norm_pos_iff.mpr ha), mul_comm, ←p.smul, smul_sub, smul_inv_smul₀ ha] + lt_div_iff (norm_pos_iff.mpr ha), mul_comm, ←map_smul_eq_mul p, smul_sub, smul_inv_smul₀ ha] end normed_field @@ -623,9 +605,9 @@ variables [has_smul ℝ E] [is_scalar_tower ℝ 𝕜 E] (p : seminorm 𝕜 E) protected lemma convex_on : convex_on ℝ univ p := begin refine ⟨convex_univ, λ x _ y _ a b ha hb hab, _⟩, - calc p (a • x + b • y) ≤ p (a • x) + p (b • y) : p.add_le _ _ + calc p (a • x + b • y) ≤ p (a • x) + p (b • y) : map_add_le_add p _ _ ... = ∥a • (1 : 𝕜)∥ * p x + ∥b • (1 : 𝕜)∥ * p y - : by rw [←p.smul, ←p.smul, smul_one_smul, smul_one_smul] + : by rw [←map_smul_eq_mul p, ←map_smul_eq_mul p, smul_one_smul, smul_one_smul] ... = a * p x + b * p y : by rw [norm_smul, norm_smul, norm_one, mul_one, mul_one, real.norm_of_nonneg ha, real.norm_of_nonneg hb], From c3d2b3b86a45d3c8a5457c4dacde1e723e61ab5f Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Fri, 9 Sep 2022 15:00:27 +0000 Subject: [PATCH 215/828] feat(analysis/inner_product_space/basic): completion of an `inner_product_space` is an `inner_product_space` (#16407) --- src/analysis/inner_product_space/basic.lean | 60 +++++++++++++++++++++ src/topology/algebra/group_completion.lean | 7 +++ 2 files changed, 67 insertions(+) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index cf780211e6a01..3e0ed29ece738 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -6,6 +6,7 @@ Authors: Zhouhang Zhou, Sébastien Gouëzel, Frédéric Dupuis import algebra.direct_sum.module import analysis.complex.basic import analysis.convex.uniform +import analysis.normed_space.completion import analysis.normed_space.bounded_linear_maps import analysis.normed_space.banach import linear_algebra.bilinear_form @@ -2324,3 +2325,62 @@ lemma submodule.orthogonal_family_self : | ff ff := absurd rfl end orthogonal + +namespace uniform_space.completion + +open uniform_space function + +instance {𝕜' E' : Type*} [topological_space 𝕜'] [uniform_space E'] [has_inner 𝕜' E'] : + has_inner 𝕜' (completion E') := +{ inner := curry $ (dense_inducing_coe.prod dense_inducing_coe).extend (uncurry inner) } + +@[simp] lemma inner_coe (a b : E) : + inner (a : completion E) (b : completion E) = (inner a b : 𝕜) := +(dense_inducing_coe.prod dense_inducing_coe).extend_eq + (continuous_inner : continuous (uncurry inner : E × E → 𝕜)) (a, b) + +protected lemma continuous_inner : + continuous (uncurry inner : completion E × completion E → 𝕜) := +begin + let inner' : E →+ E →+ 𝕜 := + { to_fun := λ x, (innerₛₗ x).to_add_monoid_hom, + map_zero' := by ext x; exact inner_zero_left, + map_add' := λ x y, by ext z; exact inner_add_left }, + have : continuous (λ p : E × E, inner' p.1 p.2) := continuous_inner, + rw [completion.has_inner, uncurry_curry _], + change continuous (((dense_inducing_to_compl E).prod (dense_inducing_to_compl E)).extend + (λ p : E × E, inner' p.1 p.2)), + exact (dense_inducing_to_compl E).extend_Z_bilin (dense_inducing_to_compl E) this, +end + +protected lemma continuous.inner {α : Type*} [topological_space α] + {f g : α → completion E} (hf : continuous f) (hg : continuous g) : + continuous (λ x : α, inner (f x) (g x) : α → 𝕜) := +uniform_space.completion.continuous_inner.comp (hf.prod_mk hg : _) + +instance : inner_product_space 𝕜 (completion E) := +{ norm_sq_eq_inner := λ x, completion.induction_on x + (is_closed_eq + (continuous_norm.pow 2) + (continuous_re.comp (continuous.inner continuous_id' continuous_id'))) + (λ a, by simp only [norm_coe, inner_coe, inner_self_eq_norm_sq]), + conj_sym := λ x y, completion.induction_on₂ x y + (is_closed_eq + (continuous_conj.comp (continuous.inner continuous_snd continuous_fst)) + (continuous.inner continuous_fst continuous_snd)) + (λ a b, by simp only [inner_coe, inner_conj_sym]), + add_left := λ x y z, completion.induction_on₃ x y z + (is_closed_eq + (continuous.inner (continuous_fst.add (continuous_fst.comp continuous_snd)) + (continuous_snd.comp continuous_snd)) + ((continuous.inner continuous_fst (continuous_snd.comp continuous_snd)).add + (continuous.inner (continuous_fst.comp continuous_snd) + (continuous_snd.comp continuous_snd)))) + (λ a b c, by simp only [← coe_add, inner_coe, inner_add_left]), + smul_left := λ x y c, completion.induction_on₂ x y + (is_closed_eq + (continuous.inner (continuous_fst.const_smul c) continuous_snd) + ((continuous_mul_left _).comp (continuous.inner continuous_fst continuous_snd))) + (λ a b, by simp only [← coe_smul c a, inner_coe, inner_smul_left]) } + +end uniform_space.completion diff --git a/src/topology/algebra/group_completion.lean b/src/topology/algebra/group_completion.lean index 92074999a9132..afe26321bb8f4 100644 --- a/src/topology/algebra/group_completion.lean +++ b/src/topology/algebra/group_completion.lean @@ -149,6 +149,13 @@ instance {M} [monoid M] [distrib_mul_action M α] [has_uniform_continuous_const_ lemma continuous_to_compl : continuous (to_compl : α → completion α) := continuous_coe α +variable (α) + +lemma dense_inducing_to_compl : dense_inducing (to_compl : α → completion α) := +dense_inducing_coe + +variable {α} + end uniform_add_group section uniform_add_comm_group From 22a9066dbd977d5b1122d23975b0bb9e6cd9f6c6 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Fri, 9 Sep 2022 16:10:08 +0000 Subject: [PATCH 216/828] feat(topology/algebra/module/weak_dual): `weak_space` is functorial (#16441) --- src/topology/algebra/module/weak_dual.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/topology/algebra/module/weak_dual.lean b/src/topology/algebra/module/weak_dual.lean index 122106d3c27f9..627aa29f3062c 100644 --- a/src/topology/algebra/module/weak_dual.lean +++ b/src/topology/algebra/module/weak_dual.lean @@ -248,6 +248,22 @@ def weak_space (𝕜 E) [comm_semiring 𝕜] [topological_space 𝕜] [has_conti [has_continuous_const_smul 𝕜 𝕜] [add_comm_monoid E] [module 𝕜 E] [topological_space E] := weak_bilin (top_dual_pairing 𝕜 E).flip +namespace weak_space + +variables {𝕜 E F} [add_comm_monoid F] [module 𝕜 F] [topological_space F] + +/-- A continuous linear map from `E` to `F` is still continuous when `E` and `F` are equipped with +their weak topologies. -/ +def map (f : E →L[𝕜] F) : + weak_space 𝕜 E →L[𝕜] weak_space 𝕜 F := +{ cont := weak_bilin.continuous_of_continuous_eval _ (λ l, weak_bilin.eval_continuous _ (l ∘L f)), + ..f } + +lemma map_apply (f : E →L[𝕜] F) (x : E) : weak_space.map f x = f x := rfl +@[simp] lemma coe_map (f : E →L[𝕜] F) : (weak_space.map f : E → F) = f := rfl + +end weak_space + theorem tendsto_iff_forall_eval_tendsto_top_dual_pairing {l : filter α} {f : α → weak_dual 𝕜 E} {x : weak_dual 𝕜 E} : tendsto f l (𝓝 x) ↔ From 48e3d6a4ebeb3208f50ebc3f6dfc292d92433c89 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Fri, 9 Sep 2022 23:42:10 +0000 Subject: [PATCH 217/828] feat(linear_algebra/linear_pmap): add_action and mul_action (#16068) This PR defines the instances `has_vadd`, `add_action` and `mul_action` for `linear_pmap`. Co-authored-by: Moritz Doll Co-authored-by: Yury G. Kudryashov --- src/linear_algebra/linear_pmap.lean | 37 ++++++++++++++++++++++++++--- 1 file changed, 34 insertions(+), 3 deletions(-) diff --git a/src/linear_algebra/linear_pmap.lean b/src/linear_algebra/linear_pmap.lean index 34491f4490cae..0151edba7d051 100644 --- a/src/linear_algebra/linear_pmap.lean +++ b/src/linear_algebra/linear_pmap.lean @@ -73,6 +73,8 @@ lemma ext_iff {f g : E →ₗ.[R] F} : ∀ ⦃x : f.domain⦄ ⦃y : g.domain⦄ (h : (x:E) = y), f x = g y := ⟨λ EQ, EQ ▸ ⟨rfl, λ x y h, by { congr, exact_mod_cast h }⟩, λ ⟨deq, feq⟩, ext deq feq⟩ +lemma ext' {s : submodule R E} {f g : s →ₗ[R] F} (h : f = g) : mk s f = mk s g := h ▸ rfl + lemma map_add (f : E →ₗ.[R] F) (x y : f.domain) : f (x + y) = f x + f y := f.to_fun.map_add x y @@ -322,19 +324,48 @@ instance : has_smul M (E →ₗ.[R] F) := { domain := f.domain, to_fun := a • f.to_fun }⟩ +@[simp] lemma smul_domain (a : M) (f : E →ₗ.[R] F) : (a • f).domain = f.domain := rfl + lemma smul_apply (a : M) (f : E →ₗ.[R] F) (x : ((a • f).domain)) : (a • f) x = a • f x := rfl @[simp] lemma coe_smul (a : M) (f : E →ₗ.[R] F) : ⇑(a • f) = a • f := rfl instance [smul_comm_class M N F] : smul_comm_class M N (E →ₗ.[R] F) := -⟨λ a b f, ext rfl $ λ x y hxy, by simp_rw [smul_apply, subtype.eq hxy, smul_comm]⟩ +⟨λ a b f, ext' $ smul_comm a b f.to_fun⟩ instance [has_smul M N] [is_scalar_tower M N F] : is_scalar_tower M N (E →ₗ.[R] F) := -⟨λ a b f, ext rfl $ λ x y hxy, by simp_rw [smul_apply, subtype.eq hxy, smul_assoc]⟩ +⟨λ a b f, ext' $ smul_assoc a b f.to_fun⟩ + +instance : mul_action M (E →ₗ.[R] F) := +{ smul := (•), + one_smul := λ ⟨s, f⟩, ext' $ one_smul M f, + mul_smul := λ a b f, ext' $ mul_smul a b f.to_fun } end smul +section vadd + +instance : has_vadd (E →ₗ[R] F) (E →ₗ.[R] F) := +⟨λ f g, + { domain := g.domain, + to_fun := f.comp g.domain.subtype + g.to_fun }⟩ + +@[simp] lemma vadd_domain (f : E →ₗ[R] F) (g : E →ₗ.[R] F) : (f +ᵥ g).domain = g.domain := rfl + +lemma vadd_apply (f : E →ₗ[R] F) (g : E →ₗ.[R] F) (x : (f +ᵥ g).domain) : + (f +ᵥ g) x = f x + g x := rfl + +@[simp] lemma coe_vadd (f : E →ₗ[R] F) (g : E →ₗ.[R] F) : + ⇑(f +ᵥ g) = f.comp g.domain.subtype + g := rfl + +instance : add_action (E →ₗ[R] F) (E →ₗ.[R] F) := +{ vadd := (+ᵥ), + zero_vadd := λ ⟨s, f⟩, ext' $ zero_add _, + add_vadd := λ f₁ f₂ ⟨s, g⟩, ext' $ linear_map.ext $ λ x, add_assoc _ _ _ } + +end vadd + section variables {K : Type*} [division_ring K] [module K E] [module K F] @@ -505,7 +536,7 @@ variables {M : Type*} [monoid M] [distrib_mul_action M F] [smul_comm_class R M F /-- The graph of `z • f` as a pushforward. -/ lemma smul_graph (f : E →ₗ.[R] F) (z : M) : - (z • f).graph = + (z • f).graph = f.graph.map ((linear_map.id : E →ₗ[R] E).prod_map (z • (linear_map.id : F →ₗ[R] F))) := begin ext x, cases x, From f7d03fec62a92ccff816658c135ded0b83c7c60c Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 10 Sep 2022 01:42:57 +0000 Subject: [PATCH 218/828] feat(counterexamples/sorgenfrey_line): new file (#14820) Define the Sorgenfrey line and prove that it is a normal topological space but its product with itself is not a normal topological space. --- counterexamples/sorgenfrey_line.lean | 293 ++++++++++++++++++++++++++ src/order/filter/bases.lean | 14 ++ src/topology/algebra/order/basic.lean | 4 + src/topology/connected.lean | 7 +- src/topology/paracompact.lean | 2 +- 5 files changed, 315 insertions(+), 5 deletions(-) create mode 100644 counterexamples/sorgenfrey_line.lean diff --git a/counterexamples/sorgenfrey_line.lean b/counterexamples/sorgenfrey_line.lean new file mode 100644 index 0000000000000..549f2eab337cd --- /dev/null +++ b/counterexamples/sorgenfrey_line.lean @@ -0,0 +1,293 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import topology.instances.irrational +import topology.algebra.order.archimedean +import topology.paracompact +import topology.metric_space.metrizable +import topology.metric_space.emetric_paracompact +import data.set.intervals.monotone + +/-! +# Sorgenfrey line + +In this file we define `sorgenfrey_line` (notation: `ℝₗ`) to be the Sorgenfrey line. It is the real +line with the topology space structure generated by half-open intervals `set.Ico a b`. + +We prove that this line is a normal space but its product with itself is not a normal space. In +particular, this implies that the topology on `ℝₗ` is neither metrizable, nor second countable. + +## Notations + +- `ℝₗ`: Sorgenfrey line. + +## TODO + +Prove that the Sorgenfrey line is a paracompact space. + +-/ + +open set filter topological_space +open_locale topological_space filter +noncomputable theory + +/-- The Sorgenfrey line. It is the real line with the topology space structure generated by +half-open intervals `set.Ico a b`. -/ +@[derive [conditionally_complete_linear_order, linear_ordered_field, archimedean]] +def sorgenfrey_line : Type := ℝ + +localized "notation (name := sorgenfrey_line) `ℝₗ` := sorgenfrey_line" in sorgenfrey_line + +namespace sorgenfrey_line + +/-- Ring homomorphism between the Sorgenfrey line and the standard real line. -/ +def to_real : ℝₗ ≃+* ℝ := ring_equiv.refl ℝ + +instance : topological_space ℝₗ := +topological_space.generate_from {s : set ℝₗ | ∃ a b : ℝₗ, Ico a b = s} + +lemma is_open_Ico (a b : ℝₗ) : is_open (Ico a b) := +topological_space.generate_open.basic _ ⟨a, b, rfl⟩ + +lemma is_open_Ici (a : ℝₗ) : is_open (Ici a) := +Union_Ico_right a ▸ is_open_Union (is_open_Ico a) + +lemma nhds_basis_Ico (a : ℝₗ) : (𝓝 a).has_basis (λ b, a < b) (λ b, Ico a b) := +begin + rw topological_space.nhds_generate_from, + haveI : nonempty {x // x ≤ a} := set.nonempty_Iic_subtype, + have : (⨅ (x : {i // i ≤ a}), 𝓟 (Ici ↑x)) = 𝓟 (Ici a), + { refine (is_least.is_glb _).infi_eq, + exact ⟨⟨⟨a, le_rfl⟩, rfl⟩, forall_range_iff.2 $ + λ b, principal_mono.2 $ Ici_subset_Ici.2 b.2⟩, }, + simp only [mem_set_of_eq, infi_and, infi_exists, @infi_comm _ (_ ∈ _), + @infi_comm _ (set ℝₗ), infi_infi_eq_right], + simp_rw [@infi_comm _ ℝₗ (_ ≤ _), infi_subtype', ← Ici_inter_Iio, ← inf_principal, ← inf_infi, + ← infi_inf, this, infi_subtype], + suffices : (⨅ x ∈ Ioi a, 𝓟 (Iio x)).has_basis ((<) a) Iio, from this.principal_inf _, + refine has_basis_binfi_principal _ nonempty_Ioi, + exact directed_on_iff_directed.2 (directed_of_inf $ λ x y hxy, Iio_subset_Iio hxy), +end + +lemma nhds_basis_Ico_rat (a : ℝₗ) : + (𝓝 a).has_countable_basis (λ r : ℚ, a < r) (λ r, Ico a r) := +begin + refine ⟨(nhds_basis_Ico a).to_has_basis (λ b hb, _) (λ r hr, ⟨_, hr, subset.rfl⟩), + set.to_countable _⟩, + rcases exists_rat_btwn hb with ⟨r, har, hrb⟩, + exact ⟨r, har, Ico_subset_Ico_right hrb.le⟩ +end + +lemma nhds_basis_Ico_inv_pnat (a : ℝₗ) : + (𝓝 a).has_basis (λ n : ℕ+, true) (λ n, Ico a (a + n⁻¹)) := +begin + refine (nhds_basis_Ico a).to_has_basis (λ b hb, _) + (λ n hn, ⟨_, lt_add_of_pos_right _ (inv_pos.2 $ nat.cast_pos.2 n.pos), subset.rfl⟩), + rcases exists_nat_one_div_lt (sub_pos.2 hb) with ⟨k, hk⟩, + rw [one_div] at hk, + rw [← nat.cast_add_one] at hk, + exact ⟨k.succ_pnat, trivial, Ico_subset_Ico_right (le_sub_iff_add_le'.1 hk.le)⟩ +end + +lemma nhds_countable_basis_Ico_inv_pnat (a : ℝₗ) : + (𝓝 a).has_countable_basis (λ n : ℕ+, true) (λ n, Ico a (a + n⁻¹)) := +⟨nhds_basis_Ico_inv_pnat a, set.to_countable _⟩ + +lemma nhds_antitone_basis_Ico_inv_pnat (a : ℝₗ) : + (𝓝 a).has_antitone_basis (λ n : ℕ+, Ico a (a + n⁻¹)) := +⟨nhds_basis_Ico_inv_pnat a, monotone_const.Ico $ + antitone.const_add (λ k l hkl, inv_le_inv_of_le (nat.cast_pos.2 k.pos) (nat.mono_cast hkl)) _⟩ + +lemma is_open_iff {s : set ℝₗ} : is_open s ↔ ∀ x ∈ s, ∃ y > x, Ico x y ⊆ s := +is_open_iff_mem_nhds.trans $ forall₂_congr $ λ x hx, (nhds_basis_Ico x).mem_iff + +lemma is_closed_iff {s : set ℝₗ} : is_closed s ↔ ∀ x ∉ s, ∃ y > x, disjoint (Ico x y) s := +by simp only [← is_open_compl_iff, is_open_iff, mem_compl_iff, subset_compl_iff_disjoint_right] + +lemma exists_Ico_disjoint_closed {a : ℝₗ} {s : set ℝₗ} (hs : is_closed s) (ha : a ∉ s) : + ∃ b > a, disjoint (Ico a b) s := +is_closed_iff.1 hs a ha + +@[simp] lemma map_to_real_nhds (a : ℝₗ) : map to_real (𝓝 a) = 𝓝[≥] (to_real a) := +begin + refine ((nhds_basis_Ico a).map _).eq_of_same_basis _, + simpa only [to_real.image_eq_preimage] using nhds_within_Ici_basis_Ico (to_real a) +end + +lemma nhds_eq_map (a : ℝₗ) : 𝓝 a = map to_real.symm (𝓝[≥] a.to_real) := +by simp_rw [← map_to_real_nhds, map_map, (∘), to_real.symm_apply_apply, map_id'] + +lemma nhds_eq_comap (a : ℝₗ) : 𝓝 a = comap to_real (𝓝[≥] a.to_real) := +by rw [← map_to_real_nhds, comap_map to_real.injective] + +@[continuity] lemma continuous_to_real : continuous to_real := +continuous_iff_continuous_at.2 $ λ x, + by { rw [continuous_at, tendsto, map_to_real_nhds], exact inf_le_left } + +instance : order_closed_topology ℝₗ := +⟨is_closed_le_prod.preimage (continuous_to_real.prod_map continuous_to_real)⟩ + +instance : has_continuous_add ℝₗ := +begin + refine ⟨continuous_iff_continuous_at.2 _⟩, + rintro ⟨x, y⟩, + simp only [continuous_at, nhds_prod_eq, nhds_eq_map, nhds_eq_comap (x + y), prod_map_map_eq, + tendsto_comap_iff, tendsto_map'_iff, (∘), ← nhds_within_prod_eq], + exact (continuous_add.tendsto _).inf (maps_to.tendsto $ λ x hx, add_le_add hx.1 hx.2) +end + +lemma is_clopen_Ici (a : ℝₗ) : is_clopen (Ici a) := ⟨is_open_Ici a, is_closed_Ici⟩ + +lemma is_clopen_Iio (a : ℝₗ) : is_clopen (Iio a) := +by simpa only [compl_Ici] using (is_clopen_Ici a).compl + +lemma is_clopen_Ico (a b : ℝₗ) : is_clopen (Ico a b) := +(is_clopen_Ici a).inter (is_clopen_Iio b) + +instance : totally_disconnected_space ℝₗ := +⟨λ s hs' hs x hx y hy, le_antisymm (hs.subset_clopen (is_clopen_Ici x) ⟨x, hx, le_rfl⟩ hy) + (hs.subset_clopen (is_clopen_Ici y) ⟨y, hy, le_rfl⟩ hx)⟩ + +instance : first_countable_topology ℝₗ := ⟨λ x, (nhds_basis_Ico_rat x).is_countably_generated⟩ + +/-- Sorgenfrey line is a normal topological space. -/ +instance : normal_space ℝₗ := +begin + /- Let `s` and `t` be disjoint closed sets. For each `x ∈ s` we choose `X x` such that + `set.Ico x (X x)` is disjoint with `t`. Similarly, for each `y ∈ t` we choose `Y y` such that + `set.Ico y (Y y)` is disjoint with `s`. Then `⋃ x ∈ s, Ico x (X x)` and `⋃ y ∈ t, Ico y (Y y)` are + disjoint open sets that include `s` and `t`. -/ + refine ⟨λ s t hs ht hd, _⟩, + choose! X hX hXd using λ x (hx : x ∈ s), exists_Ico_disjoint_closed ht (disjoint_left.1 hd hx), + choose! Y hY hYd using λ y (hy : y ∈ t), exists_Ico_disjoint_closed hs (disjoint_right.1 hd hy), + refine ⟨⋃ x ∈ s, Ico x (X x), ⋃ y ∈ t, Ico y (Y y), is_open_bUnion $ λ _ _, is_open_Ico _ _, + is_open_bUnion $ λ _ _, is_open_Ico _ _, _, _, _⟩, + { exact λ x hx, mem_Union₂.2 ⟨x, hx, left_mem_Ico.2 $ hX x hx⟩ }, + { exact λ y hy, mem_Union₂.2 ⟨y, hy, left_mem_Ico.2 $ hY y hy⟩ }, + { simp only [disjoint_Union_left, disjoint_Union_right, Ico_disjoint_Ico], + intros y hy x hx, + clear hs ht hd hX hY, + cases le_total x y with hle hle, + { calc min (X x) (Y y) ≤ X x : min_le_left _ _ + ... ≤ y : not_lt.1 (λ hyx, hXd x hx ⟨⟨hle, hyx⟩, hy⟩) + ... ≤ max x y : le_max_right _ _ }, + { calc min (X x) (Y y) ≤ Y y : min_le_right _ _ + ... ≤ x : not_lt.1 $ λ hxy, hYd y hy ⟨⟨hle, hxy⟩, hx⟩ + ... ≤ max x y : le_max_left _ _ } } +end + +lemma dense_range_coe_rat : dense_range (coe : ℚ → ℝₗ) := +begin + refine dense_iff_inter_open.2 _, + rintro U Uo ⟨x, hx⟩, + rcases is_open_iff.1 Uo _ hx with ⟨y, hxy, hU⟩, + rcases exists_rat_btwn hxy with ⟨z, hxz, hzy⟩, + exact ⟨z, hU ⟨hxz.le, hzy⟩, mem_range_self _⟩ +end + +instance : separable_space ℝₗ := ⟨⟨_, countable_range _, dense_range_coe_rat⟩⟩ + +lemma is_closed_antidiagonal (c : ℝₗ) : is_closed {x : ℝₗ × ℝₗ | x.1 + x.2 = c} := +is_closed_singleton.preimage continuous_add + +lemma is_clopen_Ici_prod (x : ℝₗ × ℝₗ) : is_clopen (Ici x) := +(Ici_prod_eq x).symm ▸ (is_clopen_Ici _).prod (is_clopen_Ici _) + +/-- Any subset of an antidiagonal `{(x, y) : ℝₗ × ℝₗ| x + y = c}` is a closed set. -/ +lemma is_closed_of_subset_antidiagonal {s : set (ℝₗ × ℝₗ)} {c : ℝₗ} + (hs : ∀ x : ℝₗ × ℝₗ, x ∈ s → x.1 + x.2 = c) : is_closed s := +begin + rw [← closure_subset_iff_is_closed], + rintro ⟨x, y⟩ H, + obtain rfl : x + y = c, + { change (x, y) ∈ {p : ℝₗ × ℝₗ | p.1 + p.2 = c}, + exact closure_minimal (hs : s ⊆ {x | x.1 + x.2 = c}) (is_closed_antidiagonal c) H }, + rcases mem_closure_iff.1 H (Ici (x, y)) (is_clopen_Ici_prod _).1 le_rfl + with ⟨⟨x', y'⟩, ⟨hx : x ≤ x', hy : y ≤ y'⟩, H⟩, + convert H, + { refine hx.antisymm _, + rwa [← add_le_add_iff_right, hs _ H, add_le_add_iff_left] }, + { refine hy.antisymm _, + rwa [← add_le_add_iff_left, hs _ H, add_le_add_iff_right] } +end + +lemma nhds_prod_antitone_basis_inv_pnat (x y : ℝₗ) : + (𝓝 (x, y)).has_antitone_basis (λ n : ℕ+, Ico x (x + n⁻¹) ×ˢ Ico y (y + n⁻¹)) := +begin + rw [nhds_prod_eq], + exact (nhds_antitone_basis_Ico_inv_pnat x).prod (nhds_antitone_basis_Ico_inv_pnat y) +end + +/-- The product of the Sorgenfrey line and itself is not a normal topological space. -/ +lemma not_normal_space_prod : ¬normal_space (ℝₗ × ℝₗ) := +begin + have h₀ : ∀ {n : ℕ+}, (0 : ℝ) < n⁻¹, from λ n, inv_pos.2 (nat.cast_pos.2 n.pos), + have h₀' : ∀ {n : ℕ+} {x : ℝ}, x < x + n⁻¹, from λ n x, lt_add_of_pos_right _ h₀, + introI, + /- Let `S` be the set of points `(x, y)` on the line `x + y = 0` such that `x` is rational. + Let `T` be the set of points `(x, y)` on the line `x + y = 0` such that `x` is irrational. + These sets are closed, see `sorgenfrey_line.is_closed_of_subset_antidiagonal`, and disjoint. -/ + set S := {x : ℝₗ × ℝₗ | x.1 + x.2 = 0 ∧ ∃ r : ℚ, ↑r = x.1}, + set T := {x : ℝₗ × ℝₗ | x.1 + x.2 = 0 ∧ irrational x.1.to_real}, + have hSc : is_closed S, from is_closed_of_subset_antidiagonal (λ x hx, hx.1), + have hTc : is_closed T, from is_closed_of_subset_antidiagonal (λ x hx, hx.1), + have hd : disjoint S T, + { rintro ⟨x, y⟩ ⟨⟨-, r, rfl : _ = x⟩, -, hr⟩, + exact r.not_irrational hr }, + /- Consider disjoint open sets `U ⊇ S` and `V ⊇ T`. -/ + rcases normal_separation hSc hTc hd with ⟨U, V, Uo, Vo, SU, TV, UV⟩, + /- For each point `(x, -x) ∈ T`, choose a neighborhood + `Ico x (x + k⁻¹) ×ˢ Ico (-x) (-x + k⁻¹) ⊆ V`. -/ + have : ∀ x : ℝₗ, irrational x.to_real → + ∃ k : ℕ+, Ico x (x + k⁻¹) ×ˢ Ico (-x) (-x + k⁻¹) ⊆ V, + { intros x hx, + have hV : V ∈ 𝓝 (x, -x), from Vo.mem_nhds (@TV (x, -x) ⟨add_neg_self x, hx⟩), + exact (nhds_prod_antitone_basis_inv_pnat _ _).mem_iff.1 hV }, + choose! k hkV, + /- Since the set of irrational numbers is a dense Gδ set in the usual topology of `ℝ`, there + exists `N > 0` such that the set `C N = {x : ℝ | irrational x ∧ k x = N}` is dense in a nonempty + interval. In other words, the closure of this set has a nonempty interior. -/ + set C : ℕ+ → set ℝ := λ n, closure {x | irrational x ∧ k (to_real.symm x) = n}, + have H : {x : ℝ | irrational x} ⊆ ⋃ n, C n, + from λ x hx, mem_Union.2 ⟨_, subset_closure ⟨hx, rfl⟩⟩, + have Hd : dense (⋃ n, interior (C n)) := + is_Gδ_irrational.dense_Union_interior_of_closed dense_irrational (λ _, is_closed_closure) H, + obtain ⟨N, hN⟩ : ∃ n : ℕ+, (interior $ C n).nonempty, from nonempty_Union.mp Hd.nonempty, + /- Choose a rational number `r` in the interior of the closure of `C N`, then choose `n ≥ N > 0` + such that `Ico r (r + n⁻¹) × Ico (-r) (-r + n⁻¹) ⊆ U`. -/ + rcases rat.dense_range_cast.exists_mem_open is_open_interior hN with ⟨r, hr⟩, + have hrU : ((r, -r) : ℝₗ × ℝₗ) ∈ U, from @SU (r, -r) ⟨add_neg_self _, r, rfl⟩, + obtain ⟨n, hnN, hn⟩ : ∃ n (hnN : N ≤ n), Ico (r : ℝₗ) (r + n⁻¹) ×ˢ Ico (-r : ℝₗ) (-r + n⁻¹) ⊆ U, + from ((nhds_prod_antitone_basis_inv_pnat _ _).has_basis_ge N).mem_iff.1 (Uo.mem_nhds hrU), + /- Finally, choose `x ∈ Ioo (r : ℝ) (r + n⁻¹) ∩ C N`. Then `(x, -r)` belongs both to `U` and `V`, + so they are not disjoint. This contradiction completes the proof. -/ + obtain ⟨x, hxn, hx_irr, rfl⟩ : + ∃ x : ℝ, x ∈ Ioo (r : ℝ) (r + n⁻¹) ∧ irrational x ∧ k (to_real.symm x) = N, + { have : (r : ℝ) ∈ closure (Ioo (r : ℝ) (r + n⁻¹)), + { rw [closure_Ioo h₀'.ne, left_mem_Icc], exact h₀'.le }, + rcases mem_closure_iff_nhds.1 this _ (mem_interior_iff_mem_nhds.1 hr) with ⟨x', hx', hx'ε⟩, + exact mem_closure_iff.1 hx' _ is_open_Ioo hx'ε }, + refine @UV (to_real.symm x, -r) ⟨hn ⟨_, _⟩, hkV (to_real.symm x) hx_irr ⟨_, _⟩⟩, + { exact Ioo_subset_Ico_self hxn }, + { exact left_mem_Ico.2 h₀' }, + { exact left_mem_Ico.2 h₀' }, + { refine (nhds_antitone_basis_Ico_inv_pnat (-x)).2 hnN ⟨neg_le_neg hxn.1.le, _⟩, + simp only [add_neg_lt_iff_le_add', lt_neg_add_iff_add_lt], + exact hxn.2 } +end + +/-- Topology on the Sorgenfrey line is not metrizable. -/ +lemma not_metrizable_space : ¬metrizable_space ℝₗ := +begin + introI, + letI := metrizable_space_metric ℝₗ, + exact not_normal_space_prod infer_instance +end + +/-- Topology on the Sorgenfrey line is not second countable. -/ +lemma not_second_countable_topology : ¬second_countable_topology ℝₗ := +by { introI, exact not_metrizable_space (metrizable_space_of_t3_second_countable _) } + +end sorgenfrey_line diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index 99daf52276edb..5c311780c311a 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -543,6 +543,10 @@ lemma has_basis.inf_principal (hl : l.has_basis p s) (s' : set α) : ⟨λ t, by simp only [mem_inf_principal, hl.mem_iff, subset_def, mem_set_of_eq, mem_inter_iff, and_imp]⟩ +lemma has_basis.principal_inf (hl : l.has_basis p s) (s' : set α) : + (𝓟 s' ⊓ l).has_basis p (λ i, s' ∩ s i) := +by simpa only [inf_comm, inter_comm] using hl.inf_principal s' + lemma has_basis.inf_basis_ne_bot_iff (hl : l.has_basis p s) (hl' : l'.has_basis p' s') : ne_bot (l ⊓ l') ↔ ∀ ⦃i⦄ (hi : p i) ⦃i'⦄ (hi' : p' i'), (s i ∩ s' i').nonempty := (hl.inf' hl').ne_bot_iff.trans $ by simp [@forall_swap _ ι'] @@ -886,10 +890,20 @@ countable_binfi_eq_infi_seq' Bcbl 𝓟 principal_univ section is_countably_generated +protected lemma has_antitone_basis.mem_iff [preorder ι] {l : filter α} {s : ι → set α} + (hs : l.has_antitone_basis s) {t : set α} : t ∈ l ↔ ∃ i, s i ⊆ t := +hs.to_has_basis.mem_iff.trans $ by simp only [exists_prop, true_and] + protected lemma has_antitone_basis.mem [preorder ι] {l : filter α} {s : ι → set α} (hs : l.has_antitone_basis s) (i : ι) : s i ∈ l := hs.to_has_basis.mem_of_mem trivial +lemma has_antitone_basis.has_basis_ge [preorder ι] [is_directed ι (≤)] {l : filter α} + {s : ι → set α} (hs : l.has_antitone_basis s) (i : ι) : + l.has_basis (λ j, i ≤ j) s := +hs.1.to_has_basis (λ j _, (exists_ge_ge i j).imp $ λ k hk, ⟨hk.1, hs.2 hk.2⟩) + (λ j hj, ⟨j, trivial, subset.rfl⟩) + /-- If `f` is countably generated and `f.has_basis p s`, then `f` admits a decreasing basis enumerated by natural numbers such that all sets have the form `s i`. More precisely, there is a sequence `i n` such that `p (i n)` for all `n` and `s (i n)` is a decreasing sequence of sets which diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index 4ae424f91eb66..70e0be39a3c5e 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -1519,6 +1519,10 @@ lemma mem_nhds_within_Ici_iff_exists_Ico_subset [no_max_order α] {a : α} {s : s ∈ 𝓝[≥] a ↔ ∃u ∈ Ioi a, Ico a u ⊆ s := let ⟨u', hu'⟩ := exists_gt a in mem_nhds_within_Ici_iff_exists_Ico_subset' hu' +lemma nhds_within_Ici_basis_Ico [no_max_order α] (a : α) : + (𝓝[≥] a).has_basis (λ u, a < u) (Ico a) := +⟨λ s, mem_nhds_within_Ici_iff_exists_Ico_subset⟩ + /-- A set is a neighborhood of `a` within `[a, +∞)` if and only if it contains an interval `[a, u]` with `a < u`. -/ lemma mem_nhds_within_Ici_iff_exists_Icc_subset [no_max_order α] [densely_ordered α] diff --git a/src/topology/connected.lean b/src/topology/connected.lean index bc34f21e5a63d..02c407cd88282 100644 --- a/src/topology/connected.lean +++ b/src/topology/connected.lean @@ -683,10 +683,9 @@ set.disjoint_left.2 (λ a h1 h2, h theorem is_closed_connected_component {x : α} : is_closed (connected_component x) := -closure_eq_iff_is_closed.1 $ subset.antisymm - (is_connected_connected_component.closure.subset_connected_component - (subset_closure mem_connected_component)) - subset_closure +closure_subset_iff_is_closed.1 $ + is_connected_connected_component.closure.subset_connected_component $ + subset_closure mem_connected_component lemma continuous.image_connected_component_subset [topological_space β] {f : α → β} (h : continuous f) (a : α) : f '' connected_component a ⊆ connected_component (f a) := diff --git a/src/topology/paracompact.lean b/src/topology/paracompact.lean index f045a6edb3fc5..cc09cecc8b346 100644 --- a/src/topology/paracompact.lean +++ b/src/topology/paracompact.lean @@ -34,7 +34,7 @@ We also prove the following facts. the instance graph. * Every `emetric_space` is a paracompact space, see instance `emetric_space.paracompact_space` in - `topology/metric_space/emetric_space`. + `topology/metric_space/emetric_paracompact`. ## TODO From 7fd4a0bf38496c1cc5f350dd7bcadf36ce59b685 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sat, 10 Sep 2022 10:02:38 +0000 Subject: [PATCH 219/828] =?UTF-8?q?feat(probability/stopping):=20generaliz?= =?UTF-8?q?e=20`mem=5F=E2=84=92p=5Fstopped=5Fvalue`=20(#16369)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also rename many lemmas which had names `*_of_encodable` but took `countable` arguments since a recent PR which replaced all `encodable` by `countable` but did not touch the names. Co-authored-by: Rémy Degenne Co-authored-by: RemyDegenne --- src/probability/martingale/basic.lean | 69 +++++----- src/probability/stopping.lean | 178 +++++++++++++++++--------- 2 files changed, 154 insertions(+), 93 deletions(-) diff --git a/src/probability/martingale/basic.lean b/src/probability/martingale/basic.lean index 6907199516c67..b396fc11fe517 100644 --- a/src/probability/martingale/basic.lean +++ b/src/probability/martingale/basic.lean @@ -91,7 +91,7 @@ integrable_condexp.congr (hf.condexp_ae_eq (le_refl i)) lemma set_integral_eq [sigma_finite_filtration μ ℱ] (hf : martingale f ℱ μ) {i j : ι} (hij : i ≤ j) {s : set Ω} (hs : measurable_set[ℱ i] s) : - ∫ x in s, f i x ∂μ = ∫ x in s, f j x ∂μ := + ∫ ω in s, f i ω ∂μ = ∫ ω in s, f j ω ∂μ := begin rw ← @set_integral_condexp _ _ _ _ _ (ℱ i) m0 _ _ _ (ℱ.le i) _ (hf.integrable j) hs, refine set_integral_congr_ae (ℱ.le i s hs) _, @@ -155,7 +155,7 @@ hf.2.1 i j hij lemma set_integral_le [sigma_finite_filtration μ ℱ] {f : ι → Ω → ℝ} (hf : supermartingale f ℱ μ) {i j : ι} (hij : i ≤ j) {s : set Ω} (hs : measurable_set[ℱ i] s) : - ∫ x in s, f j x ∂μ ≤ ∫ x in s, f i x ∂μ := + ∫ ω in s, f j ω ∂μ ≤ ∫ ω in s, f i ω ∂μ := begin rw ← set_integral_condexp (ℱ.le i) (hf.integrable j) hs, refine set_integral_mono_ae integrable_condexp.integrable_on (hf.integrable i).integrable_on _, @@ -229,7 +229,7 @@ end /-- The converse of this lemma is `measure_theory.submartingale_of_set_integral_le`. -/ lemma set_integral_le [sigma_finite_filtration μ ℱ] {f : ι → Ω → ℝ} (hf : submartingale f ℱ μ) {i j : ι} (hij : i ≤ j) {s : set Ω} (hs : measurable_set[ℱ i] s) : - ∫ x in s, f i x ∂μ ≤ ∫ x in s, f j x ∂μ := + ∫ ω in s, f i ω ∂μ ≤ ∫ ω in s, f j ω ∂μ := begin rw [← neg_le_neg_iff, ← integral_neg, ← integral_neg], exact supermartingale.set_integral_le hf.neg hij hs, @@ -268,7 +268,7 @@ section submartingale lemma submartingale_of_set_integral_le [is_finite_measure μ] {f : ι → Ω → ℝ} (hadp : adapted ℱ f) (hint : ∀ i, integrable (f i) μ) (hf : ∀ i j : ι, i ≤ j → ∀ s : set Ω, measurable_set[ℱ i] s → - ∫ x in s, f i x ∂μ ≤ ∫ x in s, f j x ∂μ) : + ∫ ω in s, f i ω ∂μ ≤ ∫ ω in s, f j ω ∂μ) : submartingale f ℱ μ := begin refine ⟨hadp, λ i j hij, _, hint⟩, @@ -389,7 +389,7 @@ variables {𝒢 : filtration ℕ m0} lemma submartingale_of_set_integral_le_succ [is_finite_measure μ] {f : ℕ → Ω → ℝ} (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) - (hf : ∀ i, ∀ s : set Ω, measurable_set[𝒢 i] s → ∫ x in s, f i x ∂μ ≤ ∫ x in s, f (i + 1) x ∂μ) : + (hf : ∀ i, ∀ s : set Ω, measurable_set[𝒢 i] s → ∫ ω in s, f i ω ∂μ ≤ ∫ ω in s, f (i + 1) ω ∂μ) : submartingale f 𝒢 μ := begin refine submartingale_of_set_integral_le hadp hint (λ i j hij s hs, _), @@ -400,7 +400,7 @@ end lemma supermartingale_of_set_integral_succ_le [is_finite_measure μ] {f : ℕ → Ω → ℝ} (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) - (hf : ∀ i, ∀ s : set Ω, measurable_set[𝒢 i] s → ∫ x in s, f (i + 1) x ∂μ ≤ ∫ x in s, f i x ∂μ) : + (hf : ∀ i, ∀ s : set Ω, measurable_set[𝒢 i] s → ∫ ω in s, f (i + 1) ω ∂μ ≤ ∫ ω in s, f i ω ∂μ) : supermartingale f 𝒢 μ := begin rw ← neg_neg f, @@ -410,7 +410,7 @@ end lemma martingale_of_set_integral_eq_succ [is_finite_measure μ] {f : ℕ → Ω → ℝ} (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) - (hf : ∀ i, ∀ s : set Ω, measurable_set[𝒢 i] s → ∫ x in s, f i x ∂μ = ∫ x in s, f (i + 1) x ∂μ) : + (hf : ∀ i, ∀ s : set Ω, measurable_set[𝒢 i] s → ∫ ω in s, f i ω ∂μ = ∫ ω in s, f (i + 1) ω ∂μ) : martingale f 𝒢 μ := martingale_iff.2 ⟨supermartingale_of_set_integral_succ_le hadp hint $ λ i s hs, (hf i s hs).ge, @@ -422,7 +422,7 @@ lemma submartingale_nat [is_finite_measure μ] submartingale f 𝒢 μ := begin refine submartingale_of_set_integral_le_succ hadp hint (λ i s hs, _), - have : ∫ x in s, f (i + 1) x ∂μ = ∫ x in s, μ[f (i + 1)|𝒢 i] x ∂μ := + have : ∫ ω in s, f (i + 1) ω ∂μ = ∫ ω in s, μ[f (i + 1)|𝒢 i] ω ∂μ := (set_integral_condexp (𝒢.le i) (hint _) hs).symm, rw this, exact set_integral_mono_ae (hint i).integrable_on integrable_condexp.integrable_on (hf i), @@ -482,10 +482,11 @@ end namespace submartingale +@[protected] lemma integrable_stopped_value [has_le E] {f : ℕ → Ω → E} (hf : submartingale f 𝒢 μ) {τ : Ω → ℕ} - (hτ : is_stopping_time 𝒢 τ) {N : ℕ} (hbdd : ∀ x, τ x ≤ N) : + (hτ : is_stopping_time 𝒢 τ) {N : ℕ} (hbdd : ∀ ω, τ ω ≤ N) : integrable (stopped_value f τ) μ := -integrable_stopped_value hτ hf.integrable hbdd +integrable_stopped_value ℕ hτ hf.integrable hbdd -- We may generalize the below lemma to functions taking value in a `normed_lattice_add_comm_group`. -- Similarly, generalize `(super/)submartingale.set_integral_le`. @@ -496,7 +497,7 @@ This is the forward direction of the optional stopping theorem. -/ lemma expected_stopped_value_mono [sigma_finite_filtration μ 𝒢] {f : ℕ → Ω → ℝ} (hf : submartingale f 𝒢 μ) {τ π : Ω → ℕ} (hτ : is_stopping_time 𝒢 τ) (hπ : is_stopping_time 𝒢 π) (hle : τ ≤ π) - {N : ℕ} (hbdd : ∀ x, π x ≤ N) : + {N : ℕ} (hbdd : ∀ ω, π ω ≤ N) : μ[stopped_value f τ] ≤ μ[stopped_value f π] := begin rw [← sub_nonneg, ← integral_sub', stopped_value_sub_eq_sum' hle hbdd], @@ -517,7 +518,7 @@ begin exact integrable.indicator (integrable.sub (hf.integrable _) (hf.integrable _)) (𝒢.le _ _ (this _)) }, { exact hf.integrable_stopped_value hπ hbdd }, - { exact hf.integrable_stopped_value hτ (λ x, le_trans (hle x) (hbdd x)) } + { exact hf.integrable_stopped_value hτ (λ ω, le_trans (hle ω) (hbdd ω)) } end end submartingale @@ -527,7 +528,7 @@ is a submartingale if for all bounded stopping times `τ` and `π` such that `τ stopped value of `f` at `τ` has expectation smaller than its stopped value at `π`. -/ lemma submartingale_of_expected_stopped_value_mono [is_finite_measure μ] {f : ℕ → Ω → ℝ} (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) - (hf : ∀ τ π : Ω → ℕ, is_stopping_time 𝒢 τ → is_stopping_time 𝒢 π → τ ≤ π → (∃ N, ∀ x, π x ≤ N) → + (hf : ∀ τ π : Ω → ℕ, is_stopping_time 𝒢 τ → is_stopping_time 𝒢 π → τ ≤ π → (∃ N, ∀ ω, π ω ≤ N) → μ[stopped_value f τ] ≤ μ[stopped_value f π]) : submartingale f 𝒢 μ := begin @@ -566,7 +567,7 @@ begin exact h.expected_stopped_value_mono (hσ.min hτ) (hπ.min hτ) (λ ω, min_le_min (hσ_le_π ω) le_rfl) (λ ω, (min_le_left _ _).trans (hπ_le_n ω)), }, { exact adapted.stopped_process_of_nat h.adapted hτ, }, - { exact λ i, integrable_stopped_value ((is_stopping_time_const _ i).min hτ) (h.integrable) + { exact λ i, h.integrable_stopped_value ((is_stopping_time_const _ i).min hτ) (λ ω, min_le_left _ _), }, end @@ -576,14 +577,14 @@ open finset lemma smul_le_stopped_value_hitting [is_finite_measure μ] {f : ℕ → Ω → ℝ} (hsub : submartingale f 𝒢 μ) {ε : ℝ≥0} (n : ℕ) : - ε • μ {x | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k x)} ≤ - ennreal.of_real (∫ x in {x | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k x)}, - stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) x ∂μ) := + ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} ≤ + ennreal.of_real (∫ ω in {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)}, + stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω ∂μ) := begin have hn : set.Icc 0 n = {k | k ≤ n}, { ext x, simp }, - have : ∀ x, ((ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k x)) → - (ε : ℝ) ≤ stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) x, + have : ∀ ω, ((ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) → + (ε : ℝ) ≤ stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω, { intros x hx, simp_rw [le_sup'_iff, mem_range, nat.lt_succ_iff] at hx, refine stopped_value_hitting_mem _, @@ -592,8 +593,8 @@ begin have h := set_integral_ge_of_const_le (measurable_set_le measurable_const (finset.measurable_range_sup'' (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n)))) (measure_ne_top _ _) this - (integrable.integrable_on (integrable_stopped_value (hitting_is_stopping_time - hsub.adapted measurable_set_Ici) hsub.integrable hitting_le)), + (integrable.integrable_on (hsub.integrable_stopped_value + (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)), rw [ennreal.le_of_real_iff_to_real_le, ennreal.to_real_smul], { exact h }, { exact ennreal.mul_ne_top (by simp) (measure_ne_top _ _) }, @@ -601,19 +602,19 @@ begin end /-- **Doob's maximal inequality**: Given a non-negative submartingale `f`, for all `ε : ℝ≥0`, -we have `ε • μ {ε ≤ f* n} ≤ ∫ x in {ε ≤ f* n}, f n` where `f* n x = max_{k ≤ n}, f k x`. +we have `ε • μ {ε ≤ f* n} ≤ ∫ ω in {ε ≤ f* n}, f n` where `f* n ω = max_{k ≤ n}, f k ω`. In some literature, the Doob's maximal inequality refers to what we call Doob's Lp inequality (which is a corollary of this lemma and will be proved in an upcomming PR). -/ lemma maximal_ineq [is_finite_measure μ] {f : ℕ → Ω → ℝ} (hsub : submartingale f 𝒢 μ) (hnonneg : 0 ≤ f) {ε : ℝ≥0} (n : ℕ) : - ε • μ {x | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k x)} ≤ - ennreal.of_real (∫ x in {x | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k x)}, - f n x ∂μ) := + ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} ≤ + ennreal.of_real (∫ ω in {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)}, + f n ω ∂μ) := begin - suffices : ε • μ {x | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k x)} + - ennreal.of_real (∫ x in {x | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k x)) < ε}, - f n x ∂μ) ≤ ennreal.of_real (μ[f n]), + suffices : ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} + + ennreal.of_real (∫ ω in {ω | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) < ε}, + f n ω ∂μ) ≤ ennreal.of_real (μ[f n]), { have hadd : ennreal.of_real (∫ ω, f n ω ∂μ) = ennreal.of_real (∫ ω in {ω | ↑ε ≤ ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω))}, f n ω ∂μ) + @@ -642,8 +643,8 @@ begin begin refine add_le_add (smul_le_stopped_value_hitting hsub _) (ennreal.of_real_le_of_real (set_integral_mono_on (hsub.integrable n).integrable_on - (integrable.integrable_on (integrable_stopped_value - (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hsub.integrable hitting_le)) + (integrable.integrable_on (hsub.integrable_stopped_value + (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)) (measurable_set_lt (finset.measurable_range_sup'' (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n))) measurable_const) _)), intros ω hω, @@ -668,10 +669,10 @@ begin exact (not_le.2 hω₂) hω₁ }, { exact (measurable_set_lt (finset.measurable_range_sup'' (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n))) measurable_const) }, - { exact (integrable.integrable_on (integrable_stopped_value - (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hsub.integrable hitting_le)) }, - { exact (integrable.integrable_on (integrable_stopped_value - (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hsub.integrable hitting_le)) }, + { exact (integrable.integrable_on (hsub.integrable_stopped_value + (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)) }, + { exact (integrable.integrable_on (hsub.integrable_stopped_value + (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)) }, exacts [integral_nonneg (λ x, hnonneg _ _), integral_nonneg (λ x, hnonneg _ _)], end ... ≤ ennreal.of_real (μ[f n]) : diff --git a/src/probability/stopping.lean b/src/probability/stopping.lean index 63bedf4533adf..f279b54015d6c 100644 --- a/src/probability/stopping.lean +++ b/src/probability/stopping.lean @@ -461,7 +461,7 @@ def is_stopping_time [preorder ι] (f : filtration ι m) (τ : Ω → ι) := ∀ i : ι, measurable_set[f i] $ {ω | τ ω ≤ i} lemma is_stopping_time_const [preorder ι] (f : filtration ι m) (i : ι) : - is_stopping_time f (λ x, i) := + is_stopping_time f (λ ω, i) := λ j, by simp only [measurable_set.const] section measurable_set @@ -496,7 +496,7 @@ namespace is_stopping_time variables [partial_order ι] {τ : Ω → ι} {f : filtration ι m} -protected lemma measurable_set_eq_of_countable +protected lemma measurable_set_eq_of_countable_range (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : measurable_set[f i] {ω | τ ω = i} := begin @@ -521,38 +521,39 @@ begin exact @measurable_set.empty _ (f i), }, end -protected lemma measurable_set_eq_of_encodable [countable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_eq_of_countable [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[f i] {ω | τ ω = i} := -hτ.measurable_set_eq_of_countable (set.to_countable _) i +hτ.measurable_set_eq_of_countable_range (set.to_countable _) i -protected lemma measurable_set_lt_of_countable +protected lemma measurable_set_lt_of_countable_range (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : measurable_set[f i] {ω | τ ω < i} := begin have : {ω | τ ω < i} = {ω | τ ω ≤ i} \ {ω | τ ω = i}, { ext1 ω, simp [lt_iff_le_and_ne], }, rw this, - exact (hτ.measurable_set_le i).diff (hτ.measurable_set_eq_of_countable h_countable i), + exact (hτ.measurable_set_le i).diff (hτ.measurable_set_eq_of_countable_range h_countable i), end -protected lemma measurable_set_lt_of_encodable [countable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_lt_of_countable [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[f i] {ω | τ ω < i} := -hτ.measurable_set_lt_of_countable (set.to_countable _) i +hτ.measurable_set_lt_of_countable_range (set.to_countable _) i -protected lemma measurable_set_ge_of_countable {ι} [linear_order ι] {τ : Ω → ι} {f : filtration ι m} +protected lemma measurable_set_ge_of_countable_range {ι} [linear_order ι] {τ : Ω → ι} + {f : filtration ι m} (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : measurable_set[f i] {ω | i ≤ τ ω} := begin have : {ω | i ≤ τ ω} = {ω | τ ω < i}ᶜ, { ext1 ω, simp only [set.mem_set_of_eq, set.mem_compl_eq, not_lt], }, rw this, - exact (hτ.measurable_set_lt_of_countable h_countable i).compl, + exact (hτ.measurable_set_lt_of_countable_range h_countable i).compl, end -protected lemma measurable_set_ge_of_encodable {ι} [linear_order ι] {τ : Ω → ι} {f : filtration ι m} +protected lemma measurable_set_ge_of_countable {ι} [linear_order ι] {τ : Ω → ι} {f : filtration ι m} [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[f i] {ω | i ≤ τ ω} := -hτ.measurable_set_ge_of_countable (set.to_countable _) i +hτ.measurable_set_ge_of_countable_range (set.to_countable _) i end is_stopping_time @@ -780,7 +781,7 @@ begin exact le_trans (hle _) hle' }, end -lemma measurable_space_le_of_encodable [countable ι] (hτ : is_stopping_time f τ) : +lemma measurable_space_le_of_countable [countable ι] (hτ : is_stopping_time f τ) : hτ.measurable_space ≤ m := begin intros s hs, @@ -872,6 +873,10 @@ lemma measurable_space_le_of_le_const (hτ : is_stopping_time f τ) {i : ι} (h hτ.measurable_space ≤ f i := (measurable_space_mono hτ _ hτ_le).trans (measurable_space_const _ _).le +lemma measurable_space_le_of_le (hτ : is_stopping_time f τ) {n : ι} (hτ_le : ∀ ω, τ ω ≤ n) : + hτ.measurable_space ≤ m := +(hτ.measurable_space_le_of_le_const hτ_le).trans (f.le n) + lemma le_measurable_space_of_const_le (hτ : is_stopping_time f τ) {i : ι} (hτ_le : ∀ ω, i ≤ τ ω) : f i ≤ hτ.measurable_space := (measurable_space_const _ _).symm.le.trans (measurable_space_mono _ hτ hτ_le) @@ -890,6 +895,17 @@ begin { apply_instance, }, end +instance sigma_finite_stopping_time_of_le {ι} [semilattice_sup ι] [order_bot ι] + {μ : measure Ω} {f : filtration ι m} {τ : Ω → ι} + [sigma_finite_filtration μ f] (hτ : is_stopping_time f τ) {n : ι} (hτ_le : ∀ ω, τ ω ≤ n) : + sigma_finite (μ.trim (hτ.measurable_space_le_of_le hτ_le)) := +begin + refine sigma_finite_trim_mono (hτ.measurable_space_le_of_le hτ_le) _, + { exact f ⊥, }, + { exact hτ.le_measurable_space_of_const_le (λ _, bot_le), }, + { apply_instance, }, +end + section linear_order variables [linear_order ι] {f : filtration ι m} {τ π : Ω → ι} @@ -948,19 +964,19 @@ end section countable -protected lemma measurable_set_eq_of_countable' +protected lemma measurable_set_eq_of_countable_range' (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : measurable_set[hτ.measurable_space] {ω | τ ω = i} := begin rw [← set.univ_inter {ω | τ ω = i}, measurable_set_inter_eq_iff, set.univ_inter], - exact hτ.measurable_set_eq_of_countable h_countable i, + exact hτ.measurable_set_eq_of_countable_range h_countable i, end -protected lemma measurable_set_eq_of_encodable' [countable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_eq_of_countable' [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[hτ.measurable_space] {ω | τ ω = i} := -hτ.measurable_set_eq_of_countable' (set.to_countable _) i +hτ.measurable_set_eq_of_countable_range' (set.to_countable _) i -protected lemma measurable_set_ge_of_countable' +protected lemma measurable_set_ge_of_countable_range' (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : measurable_set[hτ.measurable_space] {ω | i ≤ τ ω} := begin @@ -969,14 +985,14 @@ begin simp only [le_iff_lt_or_eq, set.mem_set_of_eq, set.mem_union_eq], rw [@eq_comm _ i, or_comm], }, rw this, - exact (hτ.measurable_set_eq_of_countable' h_countable i).union (hτ.measurable_set_gt' i), + exact (hτ.measurable_set_eq_of_countable_range' h_countable i).union (hτ.measurable_set_gt' i), end -protected lemma measurable_set_ge_of_encodable' [countable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_ge_of_countable' [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[hτ.measurable_space] {ω | i ≤ τ ω} := -hτ.measurable_set_ge_of_countable' (set.to_countable _) i +hτ.measurable_set_ge_of_countable_range' (set.to_countable _) i -protected lemma measurable_set_lt_of_countable' +protected lemma measurable_set_lt_of_countable_range' (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) (i : ι) : measurable_set[hτ.measurable_space] {ω | τ ω < i} := begin @@ -984,14 +1000,14 @@ begin { ext1 ω, simp only [lt_iff_le_and_ne, set.mem_set_of_eq, set.mem_diff], }, rw this, - exact (hτ.measurable_set_le' i).diff (hτ.measurable_set_eq_of_countable' h_countable i), + exact (hτ.measurable_set_le' i).diff (hτ.measurable_set_eq_of_countable_range' h_countable i), end -protected lemma measurable_set_lt_of_encodable' [countable ι] (hτ : is_stopping_time f τ) (i : ι) : +protected lemma measurable_set_lt_of_countable' [countable ι] (hτ : is_stopping_time f τ) (i : ι) : measurable_set[hτ.measurable_space] {ω | τ ω < i} := -hτ.measurable_set_lt_of_countable' (set.to_countable _) i +hτ.measurable_set_lt_of_countable_range' (set.to_countable _) i -protected lemma measurable_space_le_of_countable (hτ : is_stopping_time f τ) +protected lemma measurable_space_le_of_countable_range (hτ : is_stopping_time f τ) (h_countable : (set.range τ).countable) : hτ.measurable_space ≤ m := begin @@ -1101,6 +1117,17 @@ begin exact h.1, }, end +lemma measurable_set_inter_le_const_iff (hτ : is_stopping_time f τ) (s : set Ω) (i : ι) : + measurable_set[hτ.measurable_space] (s ∩ {ω | τ ω ≤ i}) + ↔ measurable_set[(hτ.min_const i).measurable_space] (s ∩ {ω | τ ω ≤ i}) := +begin + rw [is_stopping_time.measurable_set_min_iff hτ (is_stopping_time_const _ i), + is_stopping_time.measurable_space_const, is_stopping_time.measurable_set], + refine ⟨λ h, ⟨h, _⟩, λ h j, h.1 j⟩, + specialize h i, + rwa [set.inter_assoc, set.inter_self] at h, +end + lemma measurable_set_le_stopping_time [topological_space ι] [second_countable_topology ι] [order_topology ι] [measurable_space ι] [borel_space ι] (hτ : is_stopping_time f τ) (hπ : is_stopping_time f π) : @@ -1196,7 +1223,7 @@ section linear_order /-! ## Stopped value and stopped process -/ /-- Given a map `u : ι → Ω → E`, its stopped value with respect to the stopping -time `τ` is the map `x ↦ u (τ ω) x`. -/ +time `τ` is the map `x ↦ u (τ ω) ω`. -/ def stopped_value (u : ι → Ω → β) (τ : Ω → ι) : Ω → β := λ ω, u (τ ω) ω @@ -1205,8 +1232,8 @@ rfl variable [linear_order ι] -/-- Given a map `u : ι → Ω → E`, the stopped process with respect to `τ` is `u i x` if -`i ≤ τ ω`, and `u (τ ω) x` otherwise. +/-- Given a map `u : ι → Ω → E`, the stopped process with respect to `τ` is `u i ω` if +`i ≤ τ ω`, and `u (τ ω) ω` otherwise. Intuitively, the stopped process stops evolving once the stopping time has occured. -/ def stopped_process (u : ι → Ω → β) (τ : Ω → ι) : ι → Ω → β := @@ -1318,6 +1345,66 @@ end prog_measurable end linear_order +section stopped_value_of_mem_finset + +variables {μ : measure Ω} {τ σ : Ω → ι} {E : Type*} {p : ℝ≥0∞} {u : ι → Ω → E} + +lemma stopped_value_eq_of_mem_finset [add_comm_monoid E] {s : finset ι} (hbdd : ∀ ω, τ ω ∈ s) : + stopped_value u τ = ∑ i in s, set.indicator {ω | τ ω = i} (u i) := +begin + ext y, + rw [stopped_value, finset.sum_apply, finset.sum_indicator_eq_sum_filter], + suffices : finset.filter (λ i, y ∈ {ω : Ω | τ ω = i}) s = ({τ y} : finset ι), + by rw [this, finset.sum_singleton], + ext1 ω, + simp only [set.mem_set_of_eq, finset.mem_filter, finset.mem_singleton], + split; intro h, + { exact h.2.symm, }, + { refine ⟨_, h.symm⟩, rw h, exact hbdd y, }, +end + +lemma stopped_value_eq' [preorder ι] [locally_finite_order_bot ι] [add_comm_monoid E] + {N : ι} (hbdd : ∀ ω, τ ω ≤ N) : + stopped_value u τ = ∑ i in finset.Iic N, set.indicator {ω | τ ω = i} (u i) := +stopped_value_eq_of_mem_finset (λ ω, finset.mem_Iic.mpr (hbdd ω)) + +variables [partial_order ι] {ℱ : filtration ι m} [normed_add_comm_group E] + +lemma mem_ℒp_stopped_value_of_mem_finset (hτ : is_stopping_time ℱ τ) (hu : ∀ n, mem_ℒp (u n) p μ) + {s : finset ι} (hbdd : ∀ ω, τ ω ∈ s) : + mem_ℒp (stopped_value u τ) p μ := +begin + rw stopped_value_eq_of_mem_finset hbdd, + swap, apply_instance, + refine mem_ℒp_finset_sum' _ (λ i hi, mem_ℒp.indicator _ (hu i)), + refine ℱ.le i {a : Ω | τ a = i} (hτ.measurable_set_eq_of_countable_range _ i), + refine ((finset.finite_to_set s).subset (λ ω hω, _)).countable, + obtain ⟨y, rfl⟩ := hω, + exact hbdd y, +end + +lemma mem_ℒp_stopped_value [locally_finite_order_bot ι] + (hτ : is_stopping_time ℱ τ) (hu : ∀ n, mem_ℒp (u n) p μ) {N : ι} (hbdd : ∀ ω, τ ω ≤ N) : + mem_ℒp (stopped_value u τ) p μ := +mem_ℒp_stopped_value_of_mem_finset hτ hu (λ ω, finset.mem_Iic.mpr (hbdd ω)) + +lemma integrable_stopped_value_of_mem_finset (hτ : is_stopping_time ℱ τ) + (hu : ∀ n, integrable (u n) μ) {s : finset ι} (hbdd : ∀ ω, τ ω ∈ s) : + integrable (stopped_value u τ) μ := +begin + simp_rw ← mem_ℒp_one_iff_integrable at hu ⊢, + exact mem_ℒp_stopped_value_of_mem_finset hτ hu hbdd, +end + +variables (ι) + +lemma integrable_stopped_value [locally_finite_order_bot ι] + (hτ : is_stopping_time ℱ τ) (hu : ∀ n, integrable (u n) μ) {N : ι} (hbdd : ∀ ω, τ ω ≤ N) : + integrable (stopped_value u τ) μ := +integrable_stopped_value_of_mem_finset hτ hu (λ ω, finset.mem_Iic.mpr (hbdd ω)) +end stopped_value_of_mem_finset + + section nat /-! ### Filtrations indexed by `ℕ` -/ @@ -1392,16 +1479,7 @@ hu.prog_measurable_of_nat.strongly_measurable_stopped_process hτ n lemma stopped_value_eq {N : ℕ} (hbdd : ∀ ω, τ ω ≤ N) : stopped_value u τ = λ x, (∑ i in finset.range (N + 1), set.indicator {ω | τ ω = i} (u i)) x := -begin - ext y, - rw [stopped_value, finset.sum_apply, finset.sum_eq_single (τ y)], - { rw set.indicator_of_mem, - exact rfl }, - { exact λ i hi hneq, set.indicator_of_not_mem hneq.symm _ }, - { intro hy, - rw set.indicator_of_not_mem, - exact λ _, hy (finset.mem_range.2 $ lt_of_le_of_lt (hbdd _) (nat.lt_succ_self _)) } -end +stopped_value_eq_of_mem_finset (λ ω, finset.mem_range_succ_iff.mpr (hbdd ω)) lemma stopped_process_eq (n : ℕ) : stopped_process u τ n = @@ -1465,23 +1543,6 @@ lemma integrable_stopped_process (hτ : is_stopping_time f τ) integrable (stopped_process u τ n) μ := by { simp_rw ← mem_ℒp_one_iff_integrable at hu ⊢, exact mem_ℒp_stopped_process hτ hu n, } -lemma mem_ℒp_stopped_value (hτ : is_stopping_time f τ) - (hu : ∀ n, mem_ℒp (u n) p μ) {N : ℕ} (hbdd : ∀ ω, τ ω ≤ N) : - mem_ℒp (stopped_value u τ) p μ := -begin - rw stopped_value_eq hbdd, - suffices : mem_ℒp (λ x, ∑ (i : ℕ) in finset.range (N + 1), - {a : Ω | τ a = i}.indicator (u i) x) p μ, - { convert this, ext1 ω, simp only [finset.sum_apply] }, - refine mem_ℒp_finset_sum _ (λ i hi, mem_ℒp.indicator _ (hu i)), - exact f.le i {a : Ω | τ a = i} (hτ.measurable_set_eq i) -end - -lemma integrable_stopped_value (hτ : is_stopping_time f τ) - (hu : ∀ n, integrable (u n) μ) {N : ℕ} (hbdd : ∀ ω, τ ω ≤ N) : - integrable (stopped_value u τ) μ := -by { simp_rw ← mem_ℒp_one_iff_integrable at hu ⊢, exact mem_ℒp_stopped_value hτ hu hbdd, } - end normed_add_comm_group end nat @@ -1494,13 +1555,12 @@ variables [preorder ι] {𝒢 : filtration ι m} {τ η : Ω → ι} {i j : ι} /-- Given stopping times `τ` and `η` which are bounded below, `set.piecewise s τ η` is also a stopping time with respect to the same filtration. -/ lemma is_stopping_time.piecewise_of_le (hτ_st : is_stopping_time 𝒢 τ) - (hη_st : is_stopping_time 𝒢 η) (hτ : ∀ ω, i ≤ τ ω) (hη : ∀ x, i ≤ η x) + (hη_st : is_stopping_time 𝒢 η) (hτ : ∀ ω, i ≤ τ ω) (hη : ∀ ω, i ≤ η ω) (hs : measurable_set[𝒢 i] s) : is_stopping_time 𝒢 (s.piecewise τ η) := begin intro n, - have : {x | s.piecewise τ η x ≤ n} - = (s ∩ {ω | τ ω ≤ n}) ∪ (sᶜ ∩ {x | η x ≤ n}), + have : {ω | s.piecewise τ η ω ≤ n} = (s ∩ {ω | τ ω ≤ n}) ∪ (sᶜ ∩ {ω | η ω ≤ n}), { ext1 ω, simp only [set.piecewise, set.mem_inter_eq, set.mem_set_of_eq, and.congr_right_iff], by_cases hx : ω ∈ s; simp [hx], }, From ba6b1422904951dbdfa78c49f7a84ad128977e8f Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 10 Sep 2022 18:07:07 +0000 Subject: [PATCH 220/828] refactor(analysis/normed_space/spectrum): generalize Gelfand-Mazur away from `normed_division_ring` (#16448) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This generalizes the Gelfand-Mazur theorem so that it does not require the hypothesis that the algebra is a `normed_division_ring`, but instead is only a `normed_ring` in which the nonzero elements are precisely the units. This is important for applications of the Gelfand-Mazur theorem in which one may not be able to prove *a priori* that `∥a * b∥ = ∥a∥ * ∥b∥`, even though it may follow *a posteriori* from the Gelfand-Mazur isomorphism. An explicit example of this is the quotient of a complex Banach algebra by a maximal ideal, which is the use case that prompted this change. --- src/analysis/normed_space/spectrum.lean | 29 ++++++++++++++++--------- 1 file changed, 19 insertions(+), 10 deletions(-) diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 55fb086b524ed..4dfe44551d13f 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -16,6 +16,9 @@ This file contains the basic theory for the resolvent and spectrum of a Banach a ## Main definitions * `spectral_radius : ℝ≥0∞`: supremum of `∥k∥₊` for all `k ∈ spectrum 𝕜 a` +* `normed_ring.alg_equiv_complex_of_complete`: **Gelfand-Mazur theorem** For a complex + Banach division algebra, the natural `algebra_map ℂ A` is an algebra isomorphism whose inverse + is given by selecting the (unique) element of `spectrum ℂ a` ## Main statements @@ -29,9 +32,6 @@ This file contains the basic theory for the resolvent and spectrum of a Banach a * `spectrum.pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius`: Gelfand's formula for the spectral radius in Banach algebras over `ℂ`. * `spectrum.nonempty`: the spectrum of any element in a complex Banach algebra is nonempty. -* `normed_division_ring.alg_equiv_complex_of_complete`: **Gelfand-Mazur theorem** For a complex - Banach division algebra, the natural `algebra_map ℂ A` is an algebra isomorphism whose inverse - is given by selecting the (unique) element of `spectrum ℂ a` ## TODO @@ -347,23 +347,32 @@ end section gelfand_mazur_isomorphism -variables [normed_division_ring A] [normed_algebra ℂ A] +variables [normed_ring A] [normed_algebra ℂ A] (hA : ∀ {a : A}, is_unit a ↔ a ≠ 0) +include hA local notation `σ` := spectrum ℂ lemma algebra_map_eq_of_mem {a : A} {z : ℂ} (h : z ∈ σ a) : algebra_map ℂ A z = a := -by rwa [mem_iff, is_unit_iff_ne_zero, not_not, sub_eq_zero] at h +by rwa [mem_iff, hA, not_not, sub_eq_zero] at h /-- **Gelfand-Mazur theorem**: For a complex Banach division algebra, the natural `algebra_map ℂ A` is an algebra isomorphism whose inverse is given by selecting the (unique) element of -`spectrum ℂ a`. In addition, `algebra_map_isometry` guarantees this map is an isometry. -/ +`spectrum ℂ a`. In addition, `algebra_map_isometry` guarantees this map is an isometry. + +Note: because `normed_division_ring` requires the field `norm_mul' : ∀ a b, ∥a * b∥ = ∥a∥ * ∥b∥`, we +don't use this type class and instead opt for a `normed_ring` in which the nonzero elements are +precisely the units. This allows for the application of this isomorphism in broader contexts, e.g., +to the quotient of a complex Banach algebra by a maximal ideal. In the case when `A` is actually a +`normed_division_ring`, one may fill in the argument `hA` with the lemma `is_unit_iff_ne_zero`. -/ @[simps] -noncomputable def _root_.normed_division_ring.alg_equiv_complex_of_complete +noncomputable def _root_.normed_ring.alg_equiv_complex_of_complete [complete_space A] : ℂ ≃ₐ[ℂ] A := +let nt : nontrivial A := ⟨⟨1, 0, hA.mp ⟨⟨1, 1, mul_one _, mul_one _⟩, rfl⟩⟩⟩ in { to_fun := algebra_map ℂ A, - inv_fun := λ a, (spectrum.nonempty a).some, - left_inv := λ z, by simpa only [scalar_eq] using (spectrum.nonempty $ algebra_map ℂ A z).some_mem, - right_inv := λ a, algebra_map_eq_of_mem (spectrum.nonempty a).some_mem, + inv_fun := λ a, (@spectrum.nonempty _ _ _ _ nt a).some, + left_inv := λ z, by simpa only [@scalar_eq _ _ _ _ _ nt _] using + (@spectrum.nonempty _ _ _ _ nt $ algebra_map ℂ A z).some_mem, + right_inv := λ a, algebra_map_eq_of_mem @hA (@spectrum.nonempty _ _ _ _ nt a).some_mem, ..algebra.of_id ℂ A } end gelfand_mazur_isomorphism From 5e78882f839f37d086ea067a1ad44a8f0d8a51a2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 10 Sep 2022 19:00:22 +0000 Subject: [PATCH 221/828] feat(algebra/parity): Primes are not square (#16430) and a few other basic results about `is_square`. --- src/algebra/parity.lean | 46 +++++++++++++++++++++++++++++++++++------ 1 file changed, 40 insertions(+), 6 deletions(-) diff --git a/src/algebra/parity.lean b/src/algebra/parity.lean index 6c65fb9f65854..07a69fc759459 100644 --- a/src/algebra/parity.lean +++ b/src/algebra/parity.lean @@ -67,7 +67,7 @@ lemma is_square.map [mul_one_class α] [mul_one_class β] [monoid_hom_class F α by { rintro ⟨m, rfl⟩, exact ⟨f m, by simp⟩ } section monoid -variables [monoid α] +variables [monoid α] {n : ℕ} {a : α} @[to_additive even_iff_exists_two_nsmul] lemma is_square_iff_exists_sq (m : α) : is_square m ↔ ∃ c, m = c ^ 2 := @@ -81,10 +81,17 @@ attribute [to_additive even.exists_two_nsmul "Alias of the forwards direction of attribute [to_additive even_of_exists_two_nsmul "Alias of the backwards direction of `even_iff_exists_two_nsmul`."] is_square_of_exists_sq +@[to_additive even.nsmul] lemma is_square.pow (n : ℕ) : is_square a → is_square (a ^ n) := +by { rintro ⟨a, rfl⟩, exact ⟨a ^ n, (commute.refl _).mul_pow _⟩ } + +@[simp, to_additive even.nsmul'] +lemma even.is_square_pow : even n → ∀ a : α, is_square (a ^ n) := +by { rintro ⟨n, rfl⟩ a, exact ⟨a ^ n, pow_add _ _ _⟩ } + @[simp, to_additive even_two_nsmul] lemma is_square_sq (a : α) : is_square (a ^ 2) := ⟨a, pow_two _⟩ -variables [has_distrib_neg α] {n : ℕ} +variables [has_distrib_neg α] lemma even.neg_pow : even n → ∀ a : α, (-a) ^ n = a ^ n := by { rintro ⟨c, rfl⟩ a, simp_rw [←two_mul, pow_mul, neg_sq] } @@ -93,14 +100,34 @@ lemma even.neg_one_pow (h : even n) : (-1 : α) ^ n = 1 := by rw [h.neg_pow, one end monoid -/-- `0` is always a square (in a monoid with zero). -/ -lemma is_square_zero (M : Type*) [monoid_with_zero M] : is_square (0 : M) := -by { use 0, simp only [mul_zero] } - @[to_additive] lemma is_square.mul [comm_semigroup α] {a b : α} : is_square a → is_square b → is_square (a * b) := by { rintro ⟨a, rfl⟩ ⟨b, rfl⟩, exact ⟨a * b, mul_mul_mul_comm _ _ _ _⟩ } +section comm_monoid +variables [comm_monoid α] {a : α} + +lemma irreducible.not_square (ha : irreducible a) : ¬ is_square a := +by { rintro ⟨b, rfl⟩, simp only [irreducible_mul_iff, or_self] at ha, exact ha.1.not_unit ha.2 } + +lemma is_square.not_irreducible (ha : is_square a) : ¬ irreducible a := λ h, h.not_square ha + +end comm_monoid + +variables (α) + +@[simp] lemma is_square_zero [mul_zero_class α] : is_square (0 : α) := ⟨0, (mul_zero _).symm⟩ + +variables {α} + +section cancel_comm_monoid_with_zero +variables [cancel_comm_monoid_with_zero α] {a : α} + +lemma prime.not_square (ha : prime a) : ¬ is_square a := ha.irreducible.not_square +lemma is_square.not_prime (ha : is_square a) : ¬ prime a := λ h, h.not_square ha + +end cancel_comm_monoid_with_zero + section division_monoid variables [division_monoid α] {a : α} @@ -116,6 +143,9 @@ alias is_square_inv ↔ _ is_square.inv attribute [to_additive] is_square.inv +@[to_additive even.zsmul] lemma is_square.zpow (n : ℤ) : is_square a → is_square (a ^ n) := +by { rintro ⟨a, rfl⟩, exact ⟨a ^ n, (commute.refl _).mul_zpow _⟩ } + variables [has_distrib_neg α] {n : ℤ} lemma even.neg_zpow : even n → ∀ a : α, (-a) ^ n = a ^ n := @@ -133,6 +163,10 @@ lemma is_square.div [division_comm_monoid α] {a b : α} (ha : is_square a) (hb is_square (a / b) := by { rw div_eq_mul_inv, exact ha.mul hb.inv } +@[simp, to_additive even.zsmul'] +lemma even.is_square_zpow [group α] {n : ℤ} : even n → ∀ a : α, is_square (a ^ n) := +by { rintro ⟨n, rfl⟩ a, exact ⟨a ^ n, zpow_add _ _ _⟩ } + -- `odd.tsub` requires `canonically_linear_ordered_semiring`, which we don't have lemma even.tsub [canonically_linear_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] [contravariant_class α α (+) (≤)] {m n : α} (hm : even m) (hn : even n) : even (m - n) := From c8c740dfa9504301d99b65b4ec443b02679f3fd6 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Sun, 11 Sep 2022 12:05:08 +0000 Subject: [PATCH 222/828] feat(data/fintype/basic): `fintype.to_finset_prod` (#16459) --- src/data/fintype/basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index fe89a2a923be8..4b43f316e564e 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -675,6 +675,10 @@ theorem to_finset_compl [decidable_eq α] [fintype α] (s : set α) [fintype s] (sᶜ).to_finset = s.to_finsetᶜ := by { ext, simp } +lemma to_finset_prod (s : set α) (t : set β) [fintype s] [fintype t] [fintype (s ×ˢ t)] : + (s ×ˢ t).to_finset = s.to_finset ×ˢ t.to_finset := +by { ext, simp } + /- TODO Without the coercion arrow (`↥`) there is an elaboration bug; it essentially infers `fintype.{v} (set.univ.{u} : set α)` with `v` and `u` distinct. Reported in leanprover-community/lean#672 -/ From 279173913e368987ab8fc27a9f5097466d9cd8c5 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sun, 11 Sep 2022 19:22:57 +0000 Subject: [PATCH 223/828] feat(topology/uniform_space/compact_separated): drop unneeded assumptions (#16457) Lemmas in this file don't need `[separated_space _]` / `is_separated _` assumption. Also add `nhds_set` versions of some lemmas. --- src/topology/algebra/uniform_group.lean | 7 +- src/topology/nhds_set.lean | 3 + src/topology/subset_properties.lean | 3 + src/topology/uniform_space/basic.lean | 57 +++++++--- .../uniform_space/compact_separated.lean | 104 ++++++++---------- 5 files changed, 95 insertions(+), 79 deletions(-) diff --git a/src/topology/algebra/uniform_group.lean b/src/topology/algebra/uniform_group.lean index 1833669c5c04b..af3ad1cdbdeea 100644 --- a/src/topology/algebra/uniform_group.lean +++ b/src/topology/algebra/uniform_group.lean @@ -379,14 +379,14 @@ variables (G : Type*) [group G] [topological_space G] [topological_group G] Warning: in general the right and left uniformities do not coincide and so one does not obtain a `uniform_group` structure. Two important special cases where they _do_ coincide are for -commutative groups (see `topological_comm_group_is_uniform`) and for compact Hausdorff groups (see +commutative groups (see `topological_comm_group_is_uniform`) and for compact groups (see `topological_group_is_uniform_of_compact_space`). -/ @[to_additive "The right uniformity on a topological additive group (as opposed to the left uniformity). Warning: in general the right and left uniformities do not coincide and so one does not obtain a `uniform_add_group` structure. Two important special cases where they _do_ coincide are for -commutative additive groups (see `topological_add_comm_group_is_uniform`) and for compact Hausdorff +commutative additive groups (see `topological_add_comm_group_is_uniform`) and for compact additive groups (see `topological_add_comm_group_is_uniform_of_compact_space`)."] def topological_group.to_uniform_space : uniform_space G := { uniformity := comap (λp:G×G, p.2 / p.1) (𝓝 1), @@ -440,9 +440,8 @@ local attribute [instance] topological_group.to_uniform_space 𝓤 G = comap (λp:G×G, p.2 / p.1) (𝓝 (1 : G)) := rfl @[to_additive] lemma topological_group_is_uniform_of_compact_space - [compact_space G] [t2_space G] : uniform_group G := + [compact_space G] : uniform_group G := ⟨begin - haveI : separated_space G := separated_iff_t2.mpr (by apply_instance), apply compact_space.uniform_continuous_of_continuous, exact continuous_div', end⟩ diff --git a/src/topology/nhds_set.lean b/src/topology/nhds_set.lean index 907437b3152ed..04fc5c72a962b 100644 --- a/src/topology/nhds_set.lean +++ b/src/topology/nhds_set.lean @@ -35,6 +35,9 @@ Sup (nhds '' s) localized "notation (name := nhds_set) `𝓝ˢ` := nhds_set" in topological_space +lemma nhds_set_diagonal (α) [topological_space (α × α)] : 𝓝ˢ (diagonal α) = ⨆ x, 𝓝 (x, x) := +by { rw [nhds_set, ← range_diag, ← range_comp], refl } + lemma mem_nhds_set_iff_forall : s ∈ 𝓝ˢ t ↔ ∀ (x : α), x ∈ t → s ∈ 𝓝 x := by simp_rw [nhds_set, filter.mem_Sup, ball_image_iff] diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index 8dbe23d765f0b..627d55a1d3044 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -803,6 +803,9 @@ lemma is_compact_range [compact_space α] {f : α → β} (hf : continuous f) : is_compact (range f) := by rw ← image_univ; exact compact_univ.image hf +lemma is_compact_diagonal [compact_space α] : is_compact (diagonal α) := +@range_diag α ▸ is_compact_range (continuous_id.prod_mk continuous_id) + /-- If X is is_compact then pr₂ : X × Y → Y is a closed map -/ theorem is_closed_proj_of_is_compact {X : Type*} [topological_space X] [compact_space X] diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index fe33e6b76513c..22e1275bd01b6 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -5,6 +5,8 @@ Authors: Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import order.filter.small_sets import topology.subset_properties +import topology.nhds_set + /-! # Uniform spaces @@ -681,6 +683,32 @@ lemma mem_nhds_right (y : α) {s : set (α×α)} (h : s ∈ 𝓤 α) : {x : α | (x, y) ∈ s} ∈ 𝓝 y := mem_nhds_left _ (symm_le_uniformity h) +lemma is_compact.nhds_set_basis_uniformity {p : ι → Prop} {s : ι → set (α × α)} + (hU : (𝓤 α).has_basis p s) {K : set α} (hK : is_compact K) : + (𝓝ˢ K).has_basis p (λ i, ⋃ x ∈ K, ball x (s i)) := +begin + refine ⟨λ U, _⟩, + simp only [mem_nhds_set_iff_forall, (nhds_basis_uniformity' hU).mem_iff, Union₂_subset_iff], + refine ⟨λ H, _, λ ⟨i, hpi, hi⟩ x hx, ⟨i, hpi, hi x hx⟩⟩, + replace H : ∀ x ∈ K, ∃ i : {i // p i}, ball x (s i ○ s i) ⊆ U, + { intros x hx, + rcases H x hx with ⟨i, hpi, hi⟩, + rcases comp_mem_uniformity_sets (hU.mem_of_mem hpi) with ⟨t, ht_mem, ht⟩, + rcases hU.mem_iff.1 ht_mem with ⟨j, hpj, hj⟩, + exact ⟨⟨j, hpj⟩, subset.trans (ball_mono ((comp_rel_mono hj hj).trans ht) _) hi⟩ }, + haveI : nonempty {a // p a}, from nonempty_subtype.2 hU.ex_mem, + choose! I hI using H, + rcases hK.elim_nhds_subcover (λ x, ball x $ s (I x)) + (λ x hx, ball_mem_nhds _ $ hU.mem_of_mem (I x).2) with ⟨t, htK, ht⟩, + obtain ⟨i, hpi, hi⟩ : ∃ i (hpi : p i), s i ⊆ ⋂ x ∈ t, s (I x), + from hU.mem_iff.1 ((bInter_finset_mem t).2 (λ x hx, hU.mem_of_mem (I x).2)), + rw [subset_Inter₂_iff] at hi, + refine ⟨i, hpi, λ x hx, _⟩, + rcases mem_Union₂.1 (ht hx) with ⟨z, hzt : z ∈ t, hzx : x ∈ ball z (s (I z))⟩, + calc ball x (s i) ⊆ ball z (s (I z) ○ s (I z)) : λ y hy, ⟨x, hzx, hi z hzt hy⟩ + ... ⊆ U : hI z (htK z hzt), +end + lemma tendsto_right_nhds_uniformity {a : α} : tendsto (λa', (a', a)) (𝓝 a) (𝓤 α) := assume s, mem_nhds_right a @@ -764,6 +792,10 @@ end lemma supr_nhds_le_uniformity : (⨆ x : α, 𝓝 (x, x)) ≤ 𝓤 α := supr_le nhds_le_uniformity +/-- Entourages are neighborhoods of the diagonal. -/ +lemma nhds_set_diagonal_le_uniformity : 𝓝ˢ (diagonal α) ≤ 𝓤 α := +(nhds_set_diagonal α).trans_le supr_nhds_le_uniformity + /-! ### Closure and interior in uniform spaces -/ @@ -1355,25 +1387,22 @@ theorem uniformity_prod [uniform_space α] [uniform_space β] : 𝓤 (α × β) (𝓤 β).comap (λp:(α × β) × α × β, (p.1.2, p.2.2)) := rfl +lemma uniformity_prod_eq_comap_prod [uniform_space α] [uniform_space β] : + 𝓤 (α × β) = comap (λ p : (α × β) × (α × β), ((p.1.1, p.2.1), (p.1.2, p.2.2))) (𝓤 α ×ᶠ 𝓤 β) := +by rw [uniformity_prod, filter.prod, comap_inf, comap_comap, comap_comap] + lemma uniformity_prod_eq_prod [uniform_space α] [uniform_space β] : 𝓤 (α × β) = map (λ p : (α × α) × (β × β), ((p.1.1, p.2.1), (p.1.2, p.2.2))) (𝓤 α ×ᶠ 𝓤 β) := -by rw [map_swap4_eq_comap, uniformity_prod, filter.prod, comap_inf, comap_comap, comap_comap] - -lemma mem_map_iff_exists_image' {α : Type*} {β : Type*} {f : filter α} {m : α → β} {t : set β} : - t ∈ (map m f).sets ↔ (∃s∈f, m '' s ⊆ t) := -mem_map_iff_exists_image +by rw [map_swap4_eq_comap, uniformity_prod_eq_comap_prod] -lemma mem_uniformity_of_uniform_continuous_invariant [uniform_space α] {s:set (α×α)} {f : α → α → α} - (hf : uniform_continuous (λp:α×α, f p.1 p.2)) (hs : s ∈ 𝓤 α) : - ∃u∈𝓤 α, ∀a b c, (a, b) ∈ u → (f a c, f b c) ∈ s := +lemma mem_uniformity_of_uniform_continuous_invariant [uniform_space α] [uniform_space β] + {s : set (β × β)} {f : α → α → β} (hf : uniform_continuous (λ p : α × α, f p.1 p.2)) + (hs : s ∈ 𝓤 β) : + ∃ u ∈ 𝓤 α, ∀ a b c, (a, b) ∈ u → (f a c, f b c) ∈ s := begin rw [uniform_continuous, uniformity_prod_eq_prod, tendsto_map'_iff, (∘)] at hf, - rcases mem_map_iff_exists_image'.1 (hf hs) with ⟨t, ht, hts⟩, clear hf, - rcases mem_prod_iff.1 ht with ⟨u, hu, v, hv, huvt⟩, clear ht, - refine ⟨u, hu, assume a b c hab, hts $ (mem_image _ _ _).2 ⟨⟨⟨a, b⟩, ⟨c, c⟩⟩, huvt ⟨_, _⟩, _⟩⟩, - exact hab, - exact refl_mem_uniformity hv, - refl + rcases mem_prod_iff.1 (mem_map.1 $ hf hs) with ⟨u, hu, v, hv, huvt⟩, + exact ⟨u, hu, λ a b c hab, @huvt ((_, _), (_, _)) ⟨hab, refl_mem_uniformity hv⟩⟩ end lemma mem_uniform_prod [t₁ : uniform_space α] [t₂ : uniform_space β] {a : set (α × α)} diff --git a/src/topology/uniform_space/compact_separated.lean b/src/topology/uniform_space/compact_separated.lean index d610588238d57..9fddded951809 100644 --- a/src/topology/uniform_space/compact_separated.lean +++ b/src/topology/uniform_space/compact_separated.lean @@ -1,22 +1,25 @@ /- Copyright (c) 2020 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Patrick Massot +Authors: Patrick Massot, Yury Kudryashov -/ -import topology.uniform_space.separation import topology.uniform_space.uniform_convergence +import topology.separation + /-! # Compact separated uniform spaces ## Main statements -* `compact_space_uniformity`: On a separated compact uniform space, the topology determines the +* `compact_space_uniformity`: On a compact uniform space, the topology determines the uniform structure, entourages are exactly the neighborhoods of the diagonal. + * `uniform_space_of_compact_t2`: every compact T2 topological structure is induced by a uniform structure. This uniform structure is described in the previous item. -* Heine-Cantor theorem: continuous functions on compact separated uniform spaces with values in - uniform spaces are automatically uniformly continuous. There are several variations, the main one - is `compact_space.uniform_continuous_of_continuous`. + +* **Heine-Cantor** theorem: continuous functions on compact uniform spaces with values in uniform + spaces are automatically uniformly continuous. There are several variations, the main one is + `compact_space.uniform_continuous_of_continuous`. ## Implementation notes @@ -33,51 +36,40 @@ open filter uniform_space set variables {α β γ : Type*} [uniform_space α] [uniform_space β] - /-! -### Uniformity on compact separated spaces +### Uniformity on compact spaces -/ -/-- On a separated compact uniform space, the topology determines the uniform structure, entourages -are exactly the neighborhoods of the diagonal. -/ -lemma compact_space_uniformity [compact_space α] [separated_space α] : 𝓤 α = ⨆ x : α, 𝓝 (x, x) := +/-- On a compact uniform space, the topology determines the uniform structure, entourages are +exactly the neighborhoods of the diagonal. -/ +lemma nhds_set_diagonal_eq_uniformity [compact_space α] : 𝓝ˢ (diagonal α) = 𝓤 α := begin - symmetry, refine le_antisymm supr_nhds_le_uniformity _, - by_contra H, - obtain ⟨V, hV, h⟩ : ∃ V : set (α × α), (∀ x : α, V ∈ 𝓝 (x, x)) ∧ 𝓤 α ⊓ 𝓟 Vᶜ ≠ ⊥, - { simpa only [le_iff_forall_inf_principal_compl, mem_supr, not_forall, exists_prop] using H }, - let F := 𝓤 α ⊓ 𝓟 Vᶜ, - haveI : ne_bot F := ⟨h⟩, - obtain ⟨⟨x, y⟩, hx⟩ : ∃ (p : α × α), cluster_pt p F := - cluster_point_of_compact F, - have : cluster_pt (x, y) (𝓤 α) := - hx.of_inf_left, - obtain rfl : x = y := eq_of_uniformity_inf_nhds this, - have : cluster_pt (x, x) (𝓟 Vᶜ) := - hx.of_inf_right, - have : (x, x) ∉ interior V, - { have : (x, x) ∈ closure Vᶜ, by rwa mem_closure_iff_cluster_pt, - rwa closure_compl at this }, - have : (x, x) ∈ interior V, - { rw mem_interior_iff_mem_nhds, - exact hV x }, - contradiction + refine nhds_set_diagonal_le_uniformity.antisymm _, + have : (𝓤 (α × α)).has_basis (λ U, U ∈ 𝓤 α) + (λ U, (λ p : (α × α) × α × α, ((p.1.1, p.2.1), p.1.2, p.2.2)) ⁻¹' U ×ˢ U), + { rw [uniformity_prod_eq_comap_prod], + exact (𝓤 α).basis_sets.prod_self.comap _ }, + refine (is_compact_diagonal.nhds_set_basis_uniformity this).ge_iff.2 (λ U hU, _), + exact mem_of_superset hU (λ ⟨x, y⟩ hxy, mem_Union₂.2 ⟨(x, x), rfl, refl_mem_uniformity hU, hxy⟩) end -lemma unique_uniformity_of_compact_t2 [t : topological_space γ] [compact_space γ] -[t2_space γ] {u u' : uniform_space γ} -(h : u.to_topological_space = t) (h' : u'.to_topological_space = t) : u = u' := +/-- On a compact uniform space, the topology determines the uniform structure, entourages are +exactly the neighborhoods of the diagonal. -/ +lemma compact_space_uniformity [compact_space α] : 𝓤 α = ⨆ x, 𝓝 (x, x) := +nhds_set_diagonal_eq_uniformity.symm.trans (nhds_set_diagonal _) + +lemma unique_uniformity_of_compact [t : topological_space γ] [compact_space γ] + {u u' : uniform_space γ} (h : u.to_topological_space = t) (h' : u'.to_topological_space = t) : + u = u' := begin apply uniform_space_eq, change uniformity _ = uniformity _, - haveI : @compact_space γ u.to_topological_space, { rw h ; assumption }, - haveI : @compact_space γ u'.to_topological_space, { rw h' ; assumption }, - haveI : @separated_space γ u, { rwa [separated_iff_t2, h] }, - haveI : @separated_space γ u', { rwa [separated_iff_t2, h'] }, + haveI : @compact_space γ u.to_topological_space, { rwa h }, + haveI : @compact_space γ u'.to_topological_space, { rwa h' }, rw [compact_space_uniformity, compact_space_uniformity, h, h'] end -/-- The unique uniform structure inducing a given compact Hausdorff topological structure. -/ +/-- The unique uniform structure inducing a given compact topological structure. -/ def uniform_space_of_compact_t2 [topological_space γ] [compact_space γ] [t2_space γ] : uniform_space γ := { uniformity := ⨆ x, 𝓝 (x, x), @@ -188,7 +180,7 @@ def uniform_space_of_compact_t2 [topological_space γ] [compact_space γ] [t2_sp /-- Heine-Cantor: a continuous function on a compact separated uniform space is uniformly continuous. -/ -lemma compact_space.uniform_continuous_of_continuous [compact_space α] [separated_space α] +lemma compact_space.uniform_continuous_of_continuous [compact_space α] {f : α → β} (h : continuous f) : uniform_continuous f := calc map (prod.map f f) (𝓤 α) = map (prod.map f f) (⨆ x, 𝓝 (x, x)) : by rw compact_space_uniformity @@ -197,44 +189,34 @@ map (prod.map f f) (𝓤 α) = map (prod.map f f) (⨆ x, 𝓝 (x, x)) : by rw ... ≤ ⨆ y, 𝓝 (y, y) : supr_comp_le (λ y, 𝓝 (y, y)) f ... ≤ 𝓤 β : supr_nhds_le_uniformity -/-- Heine-Cantor: a continuous function on a compact separated set of a uniform space is -uniformly continuous. -/ -lemma is_compact.uniform_continuous_on_of_continuous' {s : set α} {f : α → β} - (hs : is_compact s) (hs' : is_separated s) (hf : continuous_on f s) : uniform_continuous_on f s := +/-- Heine-Cantor: a continuous function on a compact set of a uniform space is uniformly +continuous. -/ +lemma is_compact.uniform_continuous_on_of_continuous {s : set α} {f : α → β} + (hs : is_compact s) (hf : continuous_on f s) : uniform_continuous_on f s := begin rw uniform_continuous_on_iff_restrict, - rw is_separated_iff_induced at hs', rw is_compact_iff_compact_space at hs, rw continuous_on_iff_continuous_restrict at hf, resetI, exact compact_space.uniform_continuous_of_continuous hf, end -/-- Heine-Cantor: a continuous function on a compact set of a separated uniform space -is uniformly continuous. -/ -lemma is_compact.uniform_continuous_on_of_continuous [separated_space α] {s : set α} {f : α → β} - (hs : is_compact s) (hf : continuous_on f s) : uniform_continuous_on f s := -hs.uniform_continuous_on_of_continuous' (is_separated_of_separated_space s) hf - /-- A family of functions `α → β → γ` tends uniformly to its value at `x` if `α` is locally compact, -`β` is compact and separated and `f` is continuous on `U × (univ : set β)` for some separated -neighborhood `U` of `x`. -/ +`β` is compact and `f` is continuous on `U × (univ : set β)` for some neighborhood `U` of `x`. -/ lemma continuous_on.tendsto_uniformly [locally_compact_space α] [compact_space β] - [separated_space β] [uniform_space γ] {f : α → β → γ} {x : α} {U : set α} - (hxU : U ∈ 𝓝 x) (hU : is_separated U) (h : continuous_on ↿f (U ×ˢ univ)) : + [uniform_space γ] {f : α → β → γ} {x : α} {U : set α} + (hxU : U ∈ 𝓝 x) (h : continuous_on ↿f (U ×ˢ univ)) : tendsto_uniformly f (f x) (𝓝 x) := begin rcases locally_compact_space.local_compact_nhds _ _ hxU with ⟨K, hxK, hKU, hK⟩, have : uniform_continuous_on ↿f (K ×ˢ univ), - { refine is_compact.uniform_continuous_on_of_continuous' (hK.prod compact_univ) _ + from is_compact.uniform_continuous_on_of_continuous (hK.prod compact_univ) (h.mono $ prod_mono hKU subset.rfl), - exact (hU.mono hKU).prod (is_separated_of_separated_space _) }, exact this.tendsto_uniformly hxK end /-- A continuous family of functions `α → β → γ` tends uniformly to its value at `x` if `α` is -locally compact and `β` is compact and separated. -/ -lemma continuous.tendsto_uniformly [separated_space α] [locally_compact_space α] - [compact_space β] [separated_space β] [uniform_space γ] +locally compact and `β` is compact. -/ +lemma continuous.tendsto_uniformly [locally_compact_space α] [compact_space β] [uniform_space γ] (f : α → β → γ) (h : continuous ↿f) (x : α) : tendsto_uniformly f (f x) (𝓝 x) := -h.continuous_on.tendsto_uniformly univ_mem $ is_separated_of_separated_space _ +h.continuous_on.tendsto_uniformly univ_mem From d04db5859f3964af91d0081e8e8e5c5be93cdb8a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 11 Sep 2022 21:08:50 +0000 Subject: [PATCH 224/828] refactor(algebra/order/monoid): Remove redundant field from `ordered_cancel_comm_monoid` (#16445) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `mul_left_cancel : ∀ a b c, a + b = a + c → b = c` can be inferred from `le_of_mul_le_mul_left : ∀ a b c, a + b ≤ a + c → b ≤ c` and antisymmetry. --- src/algebra/order/group.lean | 4 +--- src/algebra/order/monoid.lean | 26 +++++++++++++++---------- src/algebra/order/positive/ring.lean | 4 +--- src/algebra/order/ring.lean | 16 +++------------ src/algebra/punit_instances.lean | 3 +-- src/data/dfinsupp/order.lean | 5 ----- src/data/finsupp/order.lean | 1 - src/data/multiset/basic.lean | 1 - src/data/nat/basic.lean | 7 ++----- src/data/num/lemmas.lean | 3 +-- src/number_theory/bertrand.lean | 2 +- src/order/filter/germ.lean | 2 +- src/set_theory/ordinal/natural_ops.lean | 1 - 13 files changed, 27 insertions(+), 48 deletions(-) diff --git a/src/algebra/order/group.lean b/src/algebra/order/group.lean index 2b2c30317c7cd..daec169b69b28 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group.lean @@ -54,8 +54,7 @@ instance units.ordered_comm_group [ordered_comm_monoid α] : ordered_comm_group instance ordered_comm_group.to_ordered_cancel_comm_monoid (α : Type u) [s : ordered_comm_group α] : ordered_cancel_comm_monoid α := -{ mul_left_cancel := λ a b c, (mul_right_inj a).mp, - le_of_mul_le_mul_left := λ a b c, (mul_le_mul_iff_left a).mp, +{ le_of_mul_le_mul_left := λ a b c, (mul_le_mul_iff_left a).mp, ..s } @[priority 100, to_additive] -- See note [lower instance priority] @@ -872,7 +871,6 @@ variables [linear_ordered_comm_group α] {a b c : α} instance linear_ordered_comm_group.to_linear_ordered_cancel_comm_monoid : linear_ordered_cancel_comm_monoid α := { le_of_mul_le_mul_left := λ x y z, le_of_mul_le_mul_left', - mul_left_cancel := λ x y z, mul_left_cancel, ..‹linear_ordered_comm_group α› } /-- Pullback a `linear_ordered_comm_group` under an injective map. diff --git a/src/algebra/order/monoid.lean b/src/algebra/order/monoid.lean index 3dfe4abb65071..5d592a6fcfc04 100644 --- a/src/algebra/order/monoid.lean +++ b/src/algebra/order/monoid.lean @@ -564,24 +564,27 @@ end canonically_linear_ordered_monoid /-- An ordered cancellative additive commutative monoid is an additive commutative monoid with a partial order, in which addition is cancellative and monotone. -/ -@[protect_proj, ancestor add_cancel_comm_monoid partial_order] -class ordered_cancel_add_comm_monoid (α : Type u) - extends add_cancel_comm_monoid α, partial_order α := +@[protect_proj, ancestor add_comm_monoid partial_order] +class ordered_cancel_add_comm_monoid (α : Type u) extends add_comm_monoid α, partial_order α := (add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) (le_of_add_le_add_left : ∀ a b c : α, a + b ≤ a + c → b ≤ c) /-- An ordered cancellative commutative monoid is a commutative monoid with a partial order, in which multiplication is cancellative and monotone. -/ -@[protect_proj, ancestor cancel_comm_monoid partial_order, to_additive] -class ordered_cancel_comm_monoid (α : Type u) - extends cancel_comm_monoid α, partial_order α := +@[protect_proj, ancestor comm_monoid partial_order, to_additive] +class ordered_cancel_comm_monoid (α : Type u) extends comm_monoid α, partial_order α := (mul_le_mul_left : ∀ a b : α, a ≤ b → ∀ c : α, c * a ≤ c * b) (le_of_mul_le_mul_left : ∀ a b c : α, a * b ≤ a * c → b ≤ c) section ordered_cancel_comm_monoid variables [ordered_cancel_comm_monoid α] {a b c d : α} +@[priority 200, to_additive] -- see Note [lower instance priority] +instance ordered_cancel_comm_monoid.to_contravariant_class_le_left : + contravariant_class α α (*) (≤) := +⟨ordered_cancel_comm_monoid.le_of_mul_le_mul_left⟩ + @[to_additive] lemma ordered_cancel_comm_monoid.lt_of_mul_lt_mul_left : ∀ a b c : α, a * b < a * c → b < c := λ a b c h, lt_of_le_not_le @@ -608,6 +611,12 @@ contravariant_swap_mul_lt_of_contravariant_mul_lt M instance ordered_cancel_comm_monoid.to_ordered_comm_monoid : ordered_comm_monoid α := { ..‹ordered_cancel_comm_monoid α› } +@[priority 100, to_additive] -- see Note [lower instance priority] +instance ordered_cancel_comm_monoid.to_cancel_comm_monoid : cancel_comm_monoid α := +{ mul_left_cancel := λ a b c h, + (le_of_mul_le_mul_left' h.le).antisymm $ le_of_mul_le_mul_left' h.ge, + ..‹ordered_cancel_comm_monoid α› } + /-- Pullback an `ordered_cancel_comm_monoid` under an injective map. See note [reducible non-instances]. -/ @[reducible, to_additive function.injective.ordered_cancel_add_comm_monoid @@ -619,7 +628,6 @@ def function.injective.ordered_cancel_comm_monoid {β : Type*} ordered_cancel_comm_monoid β := { le_of_mul_le_mul_left := λ a b c (bc : f (a * b) ≤ f (a * c)), (mul_le_mul_iff_left (f a)).mp (by rwa [← mul, ← mul]), - ..hf.left_cancel_semigroup f mul, ..hf.ordered_comm_monoid f one mul npow } end ordered_cancel_comm_monoid @@ -832,7 +840,7 @@ instance [ordered_comm_monoid α] [ordered_comm_monoid β] : ordered_comm_monoid instance [ordered_cancel_comm_monoid M] [ordered_cancel_comm_monoid N] : ordered_cancel_comm_monoid (M × N) := { le_of_mul_le_mul_left := λ a b c h, ⟨le_of_mul_le_mul_left' h.1, le_of_mul_le_mul_left' h.2⟩, - .. prod.cancel_comm_monoid, .. prod.ordered_comm_monoid } + .. prod.ordered_comm_monoid } @[to_additive] instance [has_le α] [has_le β] [has_mul α] [has_mul β] [has_exists_mul_of_le α] [has_exists_mul_of_le β] : has_exists_mul_of_le (α × β) := @@ -1342,12 +1350,10 @@ instance [ordered_comm_monoid α] : ordered_add_comm_monoid (additive α) := instance [ordered_cancel_add_comm_monoid α] : ordered_cancel_comm_monoid (multiplicative α) := { le_of_mul_le_mul_left := @ordered_cancel_add_comm_monoid.le_of_add_le_add_left α _, - ..multiplicative.left_cancel_semigroup, ..multiplicative.ordered_comm_monoid } instance [ordered_cancel_comm_monoid α] : ordered_cancel_add_comm_monoid (additive α) := { le_of_add_le_add_left := @ordered_cancel_comm_monoid.le_of_mul_le_mul_left α _, - ..additive.add_left_cancel_semigroup, ..additive.ordered_add_comm_monoid } instance [linear_ordered_add_comm_monoid α] : linear_ordered_comm_monoid (multiplicative α) := diff --git a/src/algebra/order/positive/ring.lean b/src/algebra/order/positive/ring.lean index 6fad786474e51..623a60ea8f757 100644 --- a/src/algebra/order/positive/ring.lean +++ b/src/algebra/order/positive/ring.lean @@ -106,9 +106,7 @@ ordered cancellative commutative monoid. We don't have a typeclass for linear or semirings, so we assume `[linear_ordered_semiring R] [is_commutative R (*)] instead. -/ instance [linear_ordered_semiring R] [is_commutative R (*)] [nontrivial R] : linear_ordered_cancel_comm_monoid {x : R // 0 < x} := -{ mul_left_cancel := λ a b c h, subtype.ext $ (strict_mono_mul_left_of_pos a.2).injective $ - by convert congr_arg subtype.val h, - le_of_mul_le_mul_left := λ a b c h, subtype.coe_le_coe.1 $ (mul_le_mul_left a.2).1 h, +{ le_of_mul_le_mul_left := λ a b c h, subtype.coe_le_coe.1 $ (mul_le_mul_left a.2).1 h, .. subtype.linear_order _, .. @positive.subtype.ordered_comm_monoid R { mul_comm := is_commutative.comm, .. ‹linear_ordered_semiring R› } _ } diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 85afcf7c3aa7f..15ed407dafdd0 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -852,13 +852,10 @@ end @[priority 100] -- see Note [lower instance priority] instance ordered_ring.to_ordered_semiring : ordered_semiring α := -{ mul_zero := mul_zero, - zero_mul := zero_mul, - add_left_cancel := @add_left_cancel α _, - le_of_add_le_add_left := @le_of_add_le_add_left α _ _ _, +{ le_of_add_le_add_left := @le_of_add_le_add_left α _ _ _, mul_lt_mul_of_pos_left := @ordered_ring.mul_lt_mul_of_pos_left α _, mul_lt_mul_of_pos_right := @ordered_ring.mul_lt_mul_of_pos_right α _, - ..‹ordered_ring α› } + ..‹ordered_ring α›, ..ring.to_semiring } -- See Note [decidable namespace] protected lemma decidable.mul_le_mul_of_nonpos_left [@decidable_rel α (≤)] @@ -1024,14 +1021,7 @@ local attribute [instance] linear_ordered_ring.decidable_le linear_ordered_ring. @[priority 100] -- see Note [lower instance priority] instance linear_ordered_ring.to_linear_ordered_semiring : linear_ordered_semiring α := -{ mul_zero := mul_zero, - zero_mul := zero_mul, - add_left_cancel := @add_left_cancel α _, - le_of_add_le_add_left := @le_of_add_le_add_left α _ _ _, - mul_lt_mul_of_pos_left := @mul_lt_mul_of_pos_left α _, - mul_lt_mul_of_pos_right := @mul_lt_mul_of_pos_right α _, - le_total := linear_ordered_ring.le_total, - ..‹linear_ordered_ring α› } +{ ..‹linear_ordered_ring α›, ..ordered_ring.to_ordered_semiring } @[priority 100] -- see Note [lower instance priority] instance linear_ordered_ring.is_domain : is_domain α := diff --git a/src/algebra/punit_instances.lean b/src/algebra/punit_instances.lean index fdbf6d1aabda4..e0d96bc9658da 100644 --- a/src/algebra/punit_instances.lean +++ b/src/algebra/punit_instances.lean @@ -77,8 +77,7 @@ by refine intros; trivial instance : linear_ordered_cancel_add_comm_monoid punit := -{ add_left_cancel := λ _ _ _ _, subsingleton.elim _ _, - le_of_add_le_add_left := λ _ _ _ _, trivial, +{ le_of_add_le_add_left := λ _ _ _ _, trivial, .. punit.canonically_ordered_add_monoid, ..punit.linear_order } instance : has_smul R punit := diff --git a/src/data/dfinsupp/order.lean b/src/data/dfinsupp/order.lean index e5f390ebc8a6d..1c37f6a388ff7 100644 --- a/src/data/dfinsupp/order.lean +++ b/src/data/dfinsupp/order.lean @@ -111,11 +111,6 @@ instance (α : ι → Type*) [Π i, ordered_cancel_add_comm_monoid (α i)] : rw [add_apply, add_apply] at H, exact le_of_add_le_add_left H, end, - add_left_cancel := λ f g h H, ext $ λ i, begin - refine add_left_cancel _, - exact f i, - rw [←add_apply, ←add_apply, H], - end, .. dfinsupp.ordered_add_comm_monoid α } instance [Π i, ordered_add_comm_monoid (α i)] [Π i, contravariant_class (α i) (α i) (+) (≤)] : diff --git a/src/data/finsupp/order.lean b/src/data/finsupp/order.lean index 607664a15a715..80ac71bacb218 100644 --- a/src/data/finsupp/order.lean +++ b/src/data/finsupp/order.lean @@ -96,7 +96,6 @@ instance [ordered_add_comm_monoid α] : ordered_add_comm_monoid (ι →₀ α) : instance [ordered_cancel_add_comm_monoid α] : ordered_cancel_add_comm_monoid (ι →₀ α) := { le_of_add_le_add_left := λ f g i h s, le_of_add_le_add_left (h s), - add_left_cancel := λ f g i h, ext $ λ s, add_left_cancel (ext_iff.1 h s), .. finsupp.ordered_add_comm_monoid } instance [ordered_add_comm_monoid α] [contravariant_class α α (+) (≤)] : diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 2fe50738cd830..e6197b00ab1c8 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -428,7 +428,6 @@ instance : ordered_cancel_add_comm_monoid (multiset α) := congr_arg coe $ append_assoc l₁ l₂ l₃, zero_add := λ s, quot.induction_on s $ λ l, rfl, add_zero := λ s, quotient.induction_on s $ λ l, congr_arg coe $ append_nil l, - add_left_cancel := λ a b c, add_left_cancel'', add_le_add_left := λ s₁ s₂, add_le_add_left, le_of_add_le_add_left := λ s₁ s₂ s₃, le_of_add_le_add_left, ..@multiset.partial_order α } diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index a5cbafbb0391e..69396d6710e40 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -54,8 +54,7 @@ instance : comm_semiring ℕ := by rw [nat.succ_eq_add_one, nat.add_comm, nat.right_distrib, nat.one_mul] } instance : linear_ordered_semiring nat := -{ add_left_cancel := @nat.add_left_cancel, - lt := nat.lt, +{ lt := nat.lt, add_le_add_left := @nat.add_le_add_left, le_of_add_le_add_left := @nat.le_of_add_le_add_left, zero_le_one := nat.le_of_lt (nat.zero_lt_succ 0), @@ -66,9 +65,7 @@ instance : linear_ordered_semiring nat := ..nat.comm_semiring, ..nat.linear_order } -- all the fields are already included in the linear_ordered_semiring instance -instance : linear_ordered_cancel_add_comm_monoid ℕ := -{ add_left_cancel := @nat.add_left_cancel, - ..nat.linear_ordered_semiring } +instance : linear_ordered_cancel_add_comm_monoid ℕ := { ..nat.linear_ordered_semiring } instance : linear_ordered_comm_monoid_with_zero ℕ := { mul_le_mul_left := λ a b h c, nat.mul_le_mul_left c h, diff --git a/src/data/num/lemmas.lean b/src/data/num/lemmas.lean index e37619a4f3844..ca40d688db3e9 100644 --- a/src/data/num/lemmas.lean +++ b/src/data/num/lemmas.lean @@ -320,8 +320,7 @@ try { intros, refl }; try { transfer }; simp [add_comm, mul_add, add_mul, mul_assoc, mul_comm, mul_left_comm] instance : ordered_cancel_add_comm_monoid num := -{ add_left_cancel := by {intros a b c, transfer_rw, apply add_left_cancel}, - lt := (<), +{ lt := (<), lt_iff_le_not_le := by {intros a b, transfer_rw, apply lt_iff_le_not_le}, le := (≤), le_refl := by transfer, diff --git a/src/number_theory/bertrand.lean b/src/number_theory/bertrand.lean index 27480499ca9b3..dd790a5440843 100644 --- a/src/number_theory/bertrand.lean +++ b/src/number_theory/bertrand.lean @@ -68,7 +68,7 @@ begin (convex_Ioi 0.5)).add ((strict_concave_on_sqrt_mul_log_Ioi.concave_on.comp_linear_map ((2 : ℝ) • linear_map.id)).subset (λ a ha, lt_of_eq_of_lt _ ((mul_lt_mul_left two_pos).mpr ha)) (convex_Ioi 0.5))).sub - ((convex_on_id (convex_Ioi 0.5)).smul (div_nonneg (log_nonneg _) _)); norm_num1 }, + ((convex_on_id (convex_Ioi (0.5 : ℝ))).smul (div_nonneg (log_nonneg _) _)); norm_num1 }, suffices : ∃ x1 x2, 0.5 < x1 ∧ x1 < x2 ∧ x2 ≤ x ∧ 0 ≤ f x1 ∧ f x2 ≤ 0, { obtain ⟨x1, x2, h1, h2, h0, h3, h4⟩ := this, exact (h.right_le_of_le_left'' h1 ((h1.trans h2).trans_le h0) h2 h0 (h4.trans h3)).trans h4 }, diff --git a/src/order/filter/germ.lean b/src/order/filter/germ.lean index 2ccd6960c3068..c2baca43deeaa 100644 --- a/src/order/filter/germ.lean +++ b/src/order/filter/germ.lean @@ -537,7 +537,7 @@ instance [ordered_cancel_comm_monoid β] : ordered_cancel_comm_monoid (germ l β H.mono $ λ x H, mul_le_mul_left' H _, le_of_mul_le_mul_left := λ f g h, induction_on₃ f g h $ λ f g h H, H.mono $ λ x, le_of_mul_le_mul_left', - .. germ.partial_order, .. germ.comm_monoid, .. germ.left_cancel_semigroup } + .. germ.partial_order, .. germ.comm_monoid } @[to_additive] instance ordered_comm_group [ordered_comm_group β] : ordered_comm_group (germ l β) := diff --git a/src/set_theory/ordinal/natural_ops.lean b/src/set_theory/ordinal/natural_ops.lean index d44b0193ea09c..31a7f0ae3ebeb 100644 --- a/src/set_theory/ordinal/natural_ops.lean +++ b/src/set_theory/ordinal/natural_ops.lean @@ -261,7 +261,6 @@ instance add_contravariant_class_le : instance : ordered_cancel_add_comm_monoid nat_ordinal := { add := (+), add_assoc := nadd_assoc, - add_left_cancel := λ a b c, add_left_cancel'', add_le_add_left := λ a b, add_le_add_left, le_of_add_le_add_left := λ a b c, le_of_add_le_add_left, zero := 0, From 87714908ff34b908c84307827a565eb2c1278a33 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sun, 11 Sep 2022 21:08:51 +0000 Subject: [PATCH 225/828] chore(probability/martingale): move the optional stopping theorem to its own file (#16464) The idea is to reserve martingale/basic.lean to really basic API about (sub/super)martingales. --- src/probability/martingale/basic.lean | 197 --------------- .../martingale/optional_stopping.lean | 230 ++++++++++++++++++ 2 files changed, 230 insertions(+), 197 deletions(-) create mode 100644 src/probability/martingale/optional_stopping.lean diff --git a/src/probability/martingale/basic.lean b/src/probability/martingale/basic.lean index b396fc11fe517..af312b7e84da4 100644 --- a/src/probability/martingale/basic.lean +++ b/src/probability/martingale/basic.lean @@ -488,205 +488,8 @@ lemma integrable_stopped_value [has_le E] {f : ℕ → Ω → E} (hf : submartin integrable (stopped_value f τ) μ := integrable_stopped_value ℕ hτ hf.integrable hbdd --- We may generalize the below lemma to functions taking value in a `normed_lattice_add_comm_group`. --- Similarly, generalize `(super/)submartingale.set_integral_le`. - -/-- Given a submartingale `f` and bounded stopping times `τ` and `π` such that `τ ≤ π`, the -expectation of `stopped_value f τ` is less than or equal to the expectation of `stopped_value f π`. -This is the forward direction of the optional stopping theorem. -/ -lemma expected_stopped_value_mono [sigma_finite_filtration μ 𝒢] - {f : ℕ → Ω → ℝ} (hf : submartingale f 𝒢 μ) {τ π : Ω → ℕ} - (hτ : is_stopping_time 𝒢 τ) (hπ : is_stopping_time 𝒢 π) (hle : τ ≤ π) - {N : ℕ} (hbdd : ∀ ω, π ω ≤ N) : - μ[stopped_value f τ] ≤ μ[stopped_value f π] := -begin - rw [← sub_nonneg, ← integral_sub', stopped_value_sub_eq_sum' hle hbdd], - { simp only [finset.sum_apply], - have : ∀ i, measurable_set[𝒢 i] {ω : Ω | τ ω ≤ i ∧ i < π ω}, - { intro i, - refine (hτ i).inter _, - convert (hπ i).compl, - ext x, - simpa }, - rw integral_finset_sum, - { refine finset.sum_nonneg (λ i hi, _), - rw [integral_indicator (𝒢.le _ _ (this _)), integral_sub', sub_nonneg], - { exact hf.set_integral_le (nat.le_succ i) (this _) }, - { exact (hf.integrable _).integrable_on }, - { exact (hf.integrable _).integrable_on } }, - intros i hi, - exact integrable.indicator (integrable.sub (hf.integrable _) (hf.integrable _)) - (𝒢.le _ _ (this _)) }, - { exact hf.integrable_stopped_value hπ hbdd }, - { exact hf.integrable_stopped_value hτ (λ ω, le_trans (hle ω) (hbdd ω)) } -end - end submartingale -/-- The converse direction of the optional stopping theorem, i.e. an adapted integrable process `f` -is a submartingale if for all bounded stopping times `τ` and `π` such that `τ ≤ π`, the -stopped value of `f` at `τ` has expectation smaller than its stopped value at `π`. -/ -lemma submartingale_of_expected_stopped_value_mono [is_finite_measure μ] - {f : ℕ → Ω → ℝ} (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) - (hf : ∀ τ π : Ω → ℕ, is_stopping_time 𝒢 τ → is_stopping_time 𝒢 π → τ ≤ π → (∃ N, ∀ ω, π ω ≤ N) → - μ[stopped_value f τ] ≤ μ[stopped_value f π]) : - submartingale f 𝒢 μ := -begin - refine submartingale_of_set_integral_le hadp hint (λ i j hij s hs, _), - classical, - specialize hf (s.piecewise (λ _, i) (λ _, j)) _ - (is_stopping_time_piecewise_const hij hs) - (is_stopping_time_const 𝒢 j) (λ x, (ite_le_sup _ _ _).trans (max_eq_right hij).le) - ⟨j, λ x, le_rfl⟩, - rwa [stopped_value_const, stopped_value_piecewise_const, - integral_piecewise (𝒢.le _ _ hs) (hint _).integrable_on (hint _).integrable_on, - ← integral_add_compl (𝒢.le _ _ hs) (hint j), add_le_add_iff_right] at hf, -end - -/-- **The optional stopping theorem** (fair game theorem): an adapted integrable process `f` -is a submartingale if and only if for all bounded stopping times `τ` and `π` such that `τ ≤ π`, the -stopped value of `f` at `τ` has expectation smaller than its stopped value at `π`. -/ -lemma submartingale_iff_expected_stopped_value_mono [is_finite_measure μ] - {f : ℕ → Ω → ℝ} (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) : - submartingale f 𝒢 μ ↔ - ∀ τ π : Ω → ℕ, is_stopping_time 𝒢 τ → is_stopping_time 𝒢 π → τ ≤ π → (∃ N, ∀ x, π x ≤ N) → - μ[stopped_value f τ] ≤ μ[stopped_value f π] := -⟨λ hf _ _ hτ hπ hle ⟨N, hN⟩, hf.expected_stopped_value_mono hτ hπ hle hN, - submartingale_of_expected_stopped_value_mono hadp hint⟩ - -/-- The stopped process of a submartingale with respect to a stopping time is a submartingale. -/ -@[protected] -lemma submartingale.stopped_process [is_finite_measure μ] - {f : ℕ → Ω → ℝ} (h : submartingale f 𝒢 μ) {τ : Ω → ℕ} (hτ : is_stopping_time 𝒢 τ) : - submartingale (stopped_process f τ) 𝒢 μ := -begin - rw submartingale_iff_expected_stopped_value_mono, - { intros σ π hσ hπ hσ_le_π hπ_bdd, - simp_rw stopped_value_stopped_process, - obtain ⟨n, hπ_le_n⟩ := hπ_bdd, - exact h.expected_stopped_value_mono (hσ.min hτ) (hπ.min hτ) - (λ ω, min_le_min (hσ_le_π ω) le_rfl) (λ ω, (min_le_left _ _).trans (hπ_le_n ω)), }, - { exact adapted.stopped_process_of_nat h.adapted hτ, }, - { exact λ i, h.integrable_stopped_value ((is_stopping_time_const _ i).min hτ) - (λ ω, min_le_left _ _), }, -end - -section maximal - -open finset - -lemma smul_le_stopped_value_hitting [is_finite_measure μ] - {f : ℕ → Ω → ℝ} (hsub : submartingale f 𝒢 μ) {ε : ℝ≥0} (n : ℕ) : - ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} ≤ - ennreal.of_real (∫ ω in {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)}, - stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω ∂μ) := -begin - have hn : set.Icc 0 n = {k | k ≤ n}, - { ext x, simp }, - have : ∀ ω, ((ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) → - (ε : ℝ) ≤ stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω, - { intros x hx, - simp_rw [le_sup'_iff, mem_range, nat.lt_succ_iff] at hx, - refine stopped_value_hitting_mem _, - simp only [set.mem_set_of_eq, exists_prop, hn], - exact let ⟨j, hj₁, hj₂⟩ := hx in ⟨j, hj₁, hj₂⟩ }, - have h := set_integral_ge_of_const_le (measurable_set_le measurable_const - (finset.measurable_range_sup'' (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n)))) - (measure_ne_top _ _) this - (integrable.integrable_on (hsub.integrable_stopped_value - (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)), - rw [ennreal.le_of_real_iff_to_real_le, ennreal.to_real_smul], - { exact h }, - { exact ennreal.mul_ne_top (by simp) (measure_ne_top _ _) }, - { exact le_trans (mul_nonneg ε.coe_nonneg ennreal.to_real_nonneg) h } -end - -/-- **Doob's maximal inequality**: Given a non-negative submartingale `f`, for all `ε : ℝ≥0`, -we have `ε • μ {ε ≤ f* n} ≤ ∫ ω in {ε ≤ f* n}, f n` where `f* n ω = max_{k ≤ n}, f k ω`. - -In some literature, the Doob's maximal inequality refers to what we call Doob's Lp inequality -(which is a corollary of this lemma and will be proved in an upcomming PR). -/ -lemma maximal_ineq [is_finite_measure μ] - {f : ℕ → Ω → ℝ} (hsub : submartingale f 𝒢 μ) (hnonneg : 0 ≤ f) {ε : ℝ≥0} (n : ℕ) : - ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} ≤ - ennreal.of_real (∫ ω in {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)}, - f n ω ∂μ) := -begin - suffices : ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} + - ennreal.of_real (∫ ω in {ω | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) < ε}, - f n ω ∂μ) ≤ ennreal.of_real (μ[f n]), - { have hadd : ennreal.of_real (∫ ω, f n ω ∂μ) = - ennreal.of_real (∫ ω in - {ω | ↑ε ≤ ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω))}, f n ω ∂μ) + - ennreal.of_real (∫ ω in - {ω | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) < ↑ε}, f n ω ∂μ), - { rw [← ennreal.of_real_add, ← integral_union], - { conv_lhs { rw ← integral_univ }, - convert rfl, - ext ω, - change (ε : ℝ) ≤ _ ∨ _ < (ε : ℝ) ↔ _, - simp only [le_or_lt, true_iff] }, - { rintro ω ⟨hω₁ : _ ≤ _, hω₂ : _ < _⟩, - exact (not_le.2 hω₂) hω₁ }, - { exact (measurable_set_lt (finset.measurable_range_sup'' - (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n))) measurable_const) }, - exacts [(hsub.integrable _).integrable_on, (hsub.integrable _).integrable_on, - integral_nonneg (hnonneg _), integral_nonneg (hnonneg _)] }, - rwa [hadd, ennreal.add_le_add_iff_right ennreal.of_real_ne_top] at this }, - calc ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} - + ennreal.of_real (∫ ω in {ω | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) < ε}, - f n ω ∂μ) - ≤ ennreal.of_real (∫ ω in {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)}, - stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω ∂μ) - + ennreal.of_real (∫ ω in {ω | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) < ε}, - stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω ∂μ) : - begin - refine add_le_add (smul_le_stopped_value_hitting hsub _) - (ennreal.of_real_le_of_real (set_integral_mono_on (hsub.integrable n).integrable_on - (integrable.integrable_on (hsub.integrable_stopped_value - (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)) - (measurable_set_lt (finset.measurable_range_sup'' - (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n))) measurable_const) _)), - intros ω hω, - rw set.mem_set_of_eq at hω, - have : hitting f {y : ℝ | ↑ε ≤ y} 0 n ω = n, - { simp only [hitting, set.mem_set_of_eq, exists_prop, pi.coe_nat, nat.cast_id, - ite_eq_right_iff, forall_exists_index, and_imp], - intros m hm hεm, - exact false.elim ((not_le.2 hω) - ((le_sup'_iff _).2 ⟨m, mem_range.2 (nat.lt_succ_of_le hm.2), hεm⟩)) }, - simp_rw [stopped_value, this], - end - ... = ennreal.of_real (∫ ω, stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω ∂μ) : - begin - rw [← ennreal.of_real_add, ← integral_union], - { conv_rhs { rw ← integral_univ }, - convert rfl, - ext ω, - change _ ↔ (ε : ℝ) ≤ _ ∨ _ < (ε : ℝ), - simp only [le_or_lt, iff_true] }, - { rintro ω ⟨hω₁ : _ ≤ _, hω₂ : _ < _⟩, - exact (not_le.2 hω₂) hω₁ }, - { exact (measurable_set_lt (finset.measurable_range_sup'' - (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n))) measurable_const) }, - { exact (integrable.integrable_on (hsub.integrable_stopped_value - (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)) }, - { exact (integrable.integrable_on (hsub.integrable_stopped_value - (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)) }, - exacts [integral_nonneg (λ x, hnonneg _ _), integral_nonneg (λ x, hnonneg _ _)], - end - ... ≤ ennreal.of_real (μ[f n]) : - begin - refine ennreal.of_real_le_of_real _, - rw ← stopped_value_const f n, - exact hsub.expected_stopped_value_mono - (hitting_is_stopping_time hsub.adapted measurable_set_Ici) - (is_stopping_time_const _ _) (λ ω, hitting_le ω) (λ ω, le_rfl : ∀ ω, n ≤ n), - end -end - -end maximal - lemma submartingale.sum_mul_sub [is_finite_measure μ] {R : ℝ} {ξ f : ℕ → Ω → ℝ} (hf : submartingale f 𝒢 μ) (hξ : adapted 𝒢 ξ) (hbdd : ∀ n ω, ξ n ω ≤ R) (hnonneg : ∀ n ω, 0 ≤ ξ n ω) : diff --git a/src/probability/martingale/optional_stopping.lean b/src/probability/martingale/optional_stopping.lean new file mode 100644 index 0000000000000..86c7f5f882d4b --- /dev/null +++ b/src/probability/martingale/optional_stopping.lean @@ -0,0 +1,230 @@ +/- +Copyright (c) 2022 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying +-/ + +import probability.martingale.basic + +/-! # Optional stopping theorem (fair game theorem) + +The optional stopping theorem states that an adapted integrable process `f` is a submartingale if +and only if for all bounded stopping times `τ` and `π` such that `τ ≤ π`, the +stopped value of `f` at `τ` has expectation smaller than its stopped value at `π`. + +This file also contains Doob's maximal inequality: given a non-negative submartingale `f`, for all +`ε : ℝ≥0`, we have `ε • μ {ε ≤ f* n} ≤ ∫ ω in {ε ≤ f* n}, f n` where `f* n ω = max_{k ≤ n}, f k ω`. + +### Main results + +* `measure_theory.submartingale_iff_expected_stopped_value_mono`: the optional stopping theorem. +* `measure_theory.submartingale.stopped_process`: the stopped process of a submartingale with + respect to a stopping time is a submartingale. +* `measure_theory.maximal_ineq`: Doob's maximal inequality. + + -/ + +open_locale nnreal ennreal measure_theory probability_theory + +namespace measure_theory + +variables {Ω : Type*} {m0 : measurable_space Ω} {μ : measure Ω} {𝒢 : filtration ℕ m0} + {f : ℕ → Ω → ℝ} {τ π : Ω → ℕ} + +-- We may generalize the below lemma to functions taking value in a `normed_lattice_add_comm_group`. +-- Similarly, generalize `(super/)submartingale.set_integral_le`. + +/-- Given a submartingale `f` and bounded stopping times `τ` and `π` such that `τ ≤ π`, the +expectation of `stopped_value f τ` is less than or equal to the expectation of `stopped_value f π`. +This is the forward direction of the optional stopping theorem. -/ +lemma submartingale.expected_stopped_value_mono [sigma_finite_filtration μ 𝒢] + (hf : submartingale f 𝒢 μ) (hτ : is_stopping_time 𝒢 τ) (hπ : is_stopping_time 𝒢 π) (hle : τ ≤ π) + {N : ℕ} (hbdd : ∀ ω, π ω ≤ N) : + μ[stopped_value f τ] ≤ μ[stopped_value f π] := +begin + rw [← sub_nonneg, ← integral_sub', stopped_value_sub_eq_sum' hle hbdd], + { simp only [finset.sum_apply], + have : ∀ i, measurable_set[𝒢 i] {ω : Ω | τ ω ≤ i ∧ i < π ω}, + { intro i, + refine (hτ i).inter _, + convert (hπ i).compl, + ext x, + simpa }, + rw integral_finset_sum, + { refine finset.sum_nonneg (λ i hi, _), + rw [integral_indicator (𝒢.le _ _ (this _)), integral_sub', sub_nonneg], + { exact hf.set_integral_le (nat.le_succ i) (this _) }, + { exact (hf.integrable _).integrable_on }, + { exact (hf.integrable _).integrable_on } }, + intros i hi, + exact integrable.indicator (integrable.sub (hf.integrable _) (hf.integrable _)) + (𝒢.le _ _ (this _)) }, + { exact hf.integrable_stopped_value hπ hbdd }, + { exact hf.integrable_stopped_value hτ (λ ω, le_trans (hle ω) (hbdd ω)) } +end + +/-- The converse direction of the optional stopping theorem, i.e. an adapted integrable process `f` +is a submartingale if for all bounded stopping times `τ` and `π` such that `τ ≤ π`, the +stopped value of `f` at `τ` has expectation smaller than its stopped value at `π`. -/ +lemma submartingale_of_expected_stopped_value_mono [is_finite_measure μ] + (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) + (hf : ∀ τ π : Ω → ℕ, is_stopping_time 𝒢 τ → is_stopping_time 𝒢 π → τ ≤ π → (∃ N, ∀ ω, π ω ≤ N) → + μ[stopped_value f τ] ≤ μ[stopped_value f π]) : + submartingale f 𝒢 μ := +begin + refine submartingale_of_set_integral_le hadp hint (λ i j hij s hs, _), + classical, + specialize hf (s.piecewise (λ _, i) (λ _, j)) _ + (is_stopping_time_piecewise_const hij hs) + (is_stopping_time_const 𝒢 j) (λ x, (ite_le_sup _ _ _).trans (max_eq_right hij).le) + ⟨j, λ x, le_rfl⟩, + rwa [stopped_value_const, stopped_value_piecewise_const, + integral_piecewise (𝒢.le _ _ hs) (hint _).integrable_on (hint _).integrable_on, + ← integral_add_compl (𝒢.le _ _ hs) (hint j), add_le_add_iff_right] at hf, +end + +/-- **The optional stopping theorem** (fair game theorem): an adapted integrable process `f` +is a submartingale if and only if for all bounded stopping times `τ` and `π` such that `τ ≤ π`, the +stopped value of `f` at `τ` has expectation smaller than its stopped value at `π`. -/ +lemma submartingale_iff_expected_stopped_value_mono [is_finite_measure μ] + (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) : + submartingale f 𝒢 μ ↔ + ∀ τ π : Ω → ℕ, is_stopping_time 𝒢 τ → is_stopping_time 𝒢 π → τ ≤ π → (∃ N, ∀ x, π x ≤ N) → + μ[stopped_value f τ] ≤ μ[stopped_value f π] := +⟨λ hf _ _ hτ hπ hle ⟨N, hN⟩, hf.expected_stopped_value_mono hτ hπ hle hN, + submartingale_of_expected_stopped_value_mono hadp hint⟩ + +/-- The stopped process of a submartingale with respect to a stopping time is a submartingale. -/ +@[protected] +lemma submartingale.stopped_process [is_finite_measure μ] + (h : submartingale f 𝒢 μ) (hτ : is_stopping_time 𝒢 τ) : + submartingale (stopped_process f τ) 𝒢 μ := +begin + rw submartingale_iff_expected_stopped_value_mono, + { intros σ π hσ hπ hσ_le_π hπ_bdd, + simp_rw stopped_value_stopped_process, + obtain ⟨n, hπ_le_n⟩ := hπ_bdd, + exact h.expected_stopped_value_mono (hσ.min hτ) (hπ.min hτ) + (λ ω, min_le_min (hσ_le_π ω) le_rfl) (λ ω, (min_le_left _ _).trans (hπ_le_n ω)), }, + { exact adapted.stopped_process_of_nat h.adapted hτ, }, + { exact λ i, h.integrable_stopped_value ((is_stopping_time_const _ i).min hτ) + (λ ω, min_le_left _ _), }, +end + +section maximal + +open finset + +lemma smul_le_stopped_value_hitting [is_finite_measure μ] + (hsub : submartingale f 𝒢 μ) {ε : ℝ≥0} (n : ℕ) : + ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} ≤ + ennreal.of_real (∫ ω in {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)}, + stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω ∂μ) := +begin + have hn : set.Icc 0 n = {k | k ≤ n}, + { ext x, simp }, + have : ∀ ω, ((ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) → + (ε : ℝ) ≤ stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω, + { intros x hx, + simp_rw [le_sup'_iff, mem_range, nat.lt_succ_iff] at hx, + refine stopped_value_hitting_mem _, + simp only [set.mem_set_of_eq, exists_prop, hn], + exact let ⟨j, hj₁, hj₂⟩ := hx in ⟨j, hj₁, hj₂⟩ }, + have h := set_integral_ge_of_const_le (measurable_set_le measurable_const + (finset.measurable_range_sup'' (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n)))) + (measure_ne_top _ _) this + (integrable.integrable_on (hsub.integrable_stopped_value + (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)), + rw [ennreal.le_of_real_iff_to_real_le, ennreal.to_real_smul], + { exact h }, + { exact ennreal.mul_ne_top (by simp) (measure_ne_top _ _) }, + { exact le_trans (mul_nonneg ε.coe_nonneg ennreal.to_real_nonneg) h } +end + +/-- **Doob's maximal inequality**: Given a non-negative submartingale `f`, for all `ε : ℝ≥0`, +we have `ε • μ {ε ≤ f* n} ≤ ∫ ω in {ε ≤ f* n}, f n` where `f* n ω = max_{k ≤ n}, f k ω`. + +In some literature, the Doob's maximal inequality refers to what we call Doob's Lp inequality +(which is a corollary of this lemma and will be proved in an upcomming PR). -/ +lemma maximal_ineq [is_finite_measure μ] + (hsub : submartingale f 𝒢 μ) (hnonneg : 0 ≤ f) {ε : ℝ≥0} (n : ℕ) : + ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} ≤ + ennreal.of_real (∫ ω in {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)}, + f n ω ∂μ) := +begin + suffices : ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} + + ennreal.of_real (∫ ω in {ω | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) < ε}, + f n ω ∂μ) ≤ ennreal.of_real (μ[f n]), + { have hadd : ennreal.of_real (∫ ω, f n ω ∂μ) = + ennreal.of_real (∫ ω in + {ω | ↑ε ≤ ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω))}, f n ω ∂μ) + + ennreal.of_real (∫ ω in + {ω | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) < ↑ε}, f n ω ∂μ), + { rw [← ennreal.of_real_add, ← integral_union], + { conv_lhs { rw ← integral_univ }, + convert rfl, + ext ω, + change (ε : ℝ) ≤ _ ∨ _ < (ε : ℝ) ↔ _, + simp only [le_or_lt, true_iff] }, + { rintro ω ⟨hω₁ : _ ≤ _, hω₂ : _ < _⟩, + exact (not_le.2 hω₂) hω₁ }, + { exact (measurable_set_lt (finset.measurable_range_sup'' + (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n))) measurable_const) }, + exacts [(hsub.integrable _).integrable_on, (hsub.integrable _).integrable_on, + integral_nonneg (hnonneg _), integral_nonneg (hnonneg _)] }, + rwa [hadd, ennreal.add_le_add_iff_right ennreal.of_real_ne_top] at this }, + calc ε • μ {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)} + + ennreal.of_real (∫ ω in {ω | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) < ε}, + f n ω ∂μ) + ≤ ennreal.of_real (∫ ω in {ω | (ε : ℝ) ≤ (range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)}, + stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω ∂μ) + + ennreal.of_real (∫ ω in {ω | ((range (n + 1)).sup' nonempty_range_succ (λ k, f k ω)) < ε}, + stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω ∂μ) : + begin + refine add_le_add (smul_le_stopped_value_hitting hsub _) + (ennreal.of_real_le_of_real (set_integral_mono_on (hsub.integrable n).integrable_on + (integrable.integrable_on (hsub.integrable_stopped_value + (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)) + (measurable_set_lt (finset.measurable_range_sup'' + (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n))) measurable_const) _)), + intros ω hω, + rw set.mem_set_of_eq at hω, + have : hitting f {y : ℝ | ↑ε ≤ y} 0 n ω = n, + { simp only [hitting, set.mem_set_of_eq, exists_prop, pi.coe_nat, nat.cast_id, + ite_eq_right_iff, forall_exists_index, and_imp], + intros m hm hεm, + exact false.elim ((not_le.2 hω) + ((le_sup'_iff _).2 ⟨m, mem_range.2 (nat.lt_succ_of_le hm.2), hεm⟩)) }, + simp_rw [stopped_value, this], + end + ... = ennreal.of_real (∫ ω, stopped_value f (hitting f {y : ℝ | ↑ε ≤ y} 0 n) ω ∂μ) : + begin + rw [← ennreal.of_real_add, ← integral_union], + { conv_rhs { rw ← integral_univ }, + convert rfl, + ext ω, + change _ ↔ (ε : ℝ) ≤ _ ∨ _ < (ε : ℝ), + simp only [le_or_lt, iff_true] }, + { rintro ω ⟨hω₁ : _ ≤ _, hω₂ : _ < _⟩, + exact (not_le.2 hω₂) hω₁ }, + { exact (measurable_set_lt (finset.measurable_range_sup'' + (λ n _, (hsub.strongly_measurable n).measurable.le (𝒢.le n))) measurable_const) }, + { exact (integrable.integrable_on (hsub.integrable_stopped_value + (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)) }, + { exact (integrable.integrable_on (hsub.integrable_stopped_value + (hitting_is_stopping_time hsub.adapted measurable_set_Ici) hitting_le)) }, + exacts [integral_nonneg (λ x, hnonneg _ _), integral_nonneg (λ x, hnonneg _ _)], + end + ... ≤ ennreal.of_real (μ[f n]) : + begin + refine ennreal.of_real_le_of_real _, + rw ← stopped_value_const f n, + exact hsub.expected_stopped_value_mono + (hitting_is_stopping_time hsub.adapted measurable_set_Ici) + (is_stopping_time_const _ _) (λ ω, hitting_le ω) (λ ω, le_rfl : ∀ ω, n ≤ n), + end +end + +end maximal + +end measure_theory From 3c0061266fe1a5b8cc24b9a19654673fea5e4a16 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Mon, 12 Sep 2022 01:34:41 +0000 Subject: [PATCH 226/828] feat(data/{seq,stream}): cons_injective2 (#15832) Also some simple API lemmas for `nth` I didn't change the style of the `stream/init` file. Co-authored-by: Yakov Pechersky --- src/data/seq/seq.lean | 10 ++++++++++ src/data/stream/init.lean | 14 +++++++++++++- 2 files changed, 23 insertions(+), 1 deletion(-) diff --git a/src/data/seq/seq.lean b/src/data/seq/seq.lean index bcb2a44a7315f..74d600162366b 100644 --- a/src/data/seq/seq.lean +++ b/src/data/seq/seq.lean @@ -59,6 +59,16 @@ def nth : seq α → ℕ → option α := subtype.val @[ext] protected lemma ext {s t : seq α} (h : ∀ n : ℕ, s.nth n = t.nth n) : s = t := subtype.eq $ funext h +lemma cons_injective2 : function.injective2 (cons : α → seq α → seq α) := +λ x y s t h, ⟨by rw [←option.some_inj, ←nth_cons_zero, h, nth_cons_zero], + seq.ext $ λ n, by simp_rw [←nth_cons_succ x s n, h, nth_cons_succ]⟩ + +lemma cons_left_injective (s : seq α) : function.injective (λ x, cons x s) := +cons_injective2.left _ + +lemma cons_right_injective (x : α) : function.injective (cons x) := +cons_injective2.right _ + /-- A sequence has terminated at position `n` if the value at position `n` equals `none`. -/ def terminated_at (s : seq α) (n : ℕ) : Prop := s.nth n = none diff --git a/src/data/stream/init.lean b/src/data/stream/init.lean index 867b2a3c97cf6..29fec36b7894c 100644 --- a/src/data/stream/init.lean +++ b/src/data/stream/init.lean @@ -25,7 +25,7 @@ instance {α} [inhabited α] : inhabited (stream α) := protected theorem eta (s : stream α) : head s :: tail s = s := funext (λ i, begin cases i; refl end) -theorem nth_zero_cons (a : α) (s : stream α) : nth (a :: s) 0 = a := rfl +@[simp] theorem nth_zero_cons (a : α) (s : stream α) : nth (a :: s) 0 = a := rfl theorem head_cons (a : α) (s : stream α) : head (a :: s) = a := rfl @@ -43,6 +43,8 @@ funext (λ i, begin unfold drop, rw nat.add_assoc end) theorem nth_succ (n : nat) (s : stream α) : nth s (succ n) = nth (tail s) n := rfl +@[simp] lemma nth_succ_cons (n : nat) (s : stream α) (x : α) : nth (x :: s) n.succ = nth s n := rfl + theorem drop_succ (n : nat) (s : stream α) : drop (succ n) s = drop n (tail s) := rfl @[simp] lemma head_drop {α} (a : stream α) (n : ℕ) : (a.drop n).head = a.nth n := @@ -51,6 +53,16 @@ by simp only [drop, head, nat.zero_add, stream.nth] @[ext] protected theorem ext {s₁ s₂ : stream α} : (∀ n, nth s₁ n = nth s₂ n) → s₁ = s₂ := assume h, funext h +lemma cons_injective2 : function.injective2 (cons : α → stream α → stream α) := +λ x y s t h, ⟨by rw [←nth_zero_cons x s, h, nth_zero_cons], + stream.ext (λ n, by rw [←nth_succ_cons n _ x, h, nth_succ_cons])⟩ + +lemma cons_injective_left (s : stream α) : function.injective (λ x, cons x s) := +cons_injective2.left _ + +lemma cons_injective_right (x : α) : function.injective (cons x) := +cons_injective2.right _ + theorem all_def (p : α → Prop) (s : stream α) : all p s = ∀ n, p (nth s n) := rfl theorem any_def (p : α → Prop) (s : stream α) : any p s = ∃ n, p (nth s n) := rfl From 745e23d56380eae0a96504bf6ad6d48864746d1b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Mon, 12 Sep 2022 09:20:02 +0000 Subject: [PATCH 227/828] feat(measure_theory/function/conditional_expectation): conditional expectation with respect to the bot sigma algebra (#16342) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove that for a probability measure, `μ[f|⊥] = λ _, ∫ x, f x ∂μ`. Co-authored-by: RemyDegenne --- .../conditional_expectation/basic.lean | 42 ++++++++++++++++++- .../function/strongly_measurable.lean | 20 +++++++++ src/measure_theory/integral/lebesgue.lean | 27 ++++++++++++ src/measure_theory/measure/measure_space.lean | 42 +++++++++++++++++-- 4 files changed, 127 insertions(+), 4 deletions(-) diff --git a/src/measure_theory/function/conditional_expectation/basic.lean b/src/measure_theory/function/conditional_expectation/basic.lean index 097d685355147..b122f9e743a4c 100644 --- a/src/measure_theory/function/conditional_expectation/basic.lean +++ b/src/measure_theory/function/conditional_expectation/basic.lean @@ -2056,7 +2056,7 @@ begin exact set_integral_condexp_L1 hf hs, end -lemma integral_condexp {hm : m ≤ m0} [hμm : sigma_finite (μ.trim hm)] +lemma integral_condexp (hm : m ≤ m0) [hμm : sigma_finite (μ.trim hm)] (hf : integrable f μ) : ∫ x, μ[f|m] x ∂μ = ∫ x, f x ∂μ := begin suffices : ∫ x in set.univ, μ[f|m] x ∂μ = ∫ x in set.univ, f x ∂μ, @@ -2080,6 +2080,46 @@ begin rw [hg_eq s hs hμs, set_integral_condexp hm hf hs], end +lemma condexp_bot' [hμ : μ.ae.ne_bot] (f : α → F') : + μ[f|⊥] = λ _, (μ set.univ).to_real⁻¹ • ∫ x, f x ∂μ := +begin + by_cases hμ_finite : is_finite_measure μ, + swap, + { have h : ¬ sigma_finite (μ.trim bot_le), + { rwa sigma_finite_trim_bot_iff, }, + rw not_is_finite_measure_iff at hμ_finite, + rw [condexp_of_not_sigma_finite bot_le h], + simp only [hμ_finite, ennreal.top_to_real, inv_zero, zero_smul], + refl, }, + haveI : is_finite_measure μ := hμ_finite, + by_cases hf : integrable f μ, + swap, { rw [integral_undef hf, smul_zero, condexp_undef hf], refl, }, + have h_meas : strongly_measurable[⊥] (μ[f|⊥]) := strongly_measurable_condexp, + obtain ⟨c, h_eq⟩ := strongly_measurable_bot_iff.mp h_meas, + rw h_eq, + have h_integral : ∫ x, μ[f|⊥] x ∂μ = ∫ x, f x ∂μ := integral_condexp bot_le hf, + simp_rw [h_eq, integral_const] at h_integral, + rw [← h_integral, ← smul_assoc, smul_eq_mul, inv_mul_cancel, one_smul], + rw [ne.def, ennreal.to_real_eq_zero_iff, auto.not_or_eq, measure.measure_univ_eq_zero, + ← ae_eq_bot, ← ne.def, ← ne_bot_iff], + exact ⟨hμ, measure_ne_top μ set.univ⟩, +end + +lemma condexp_bot_ae_eq (f : α → F') : + μ[f|⊥] =ᵐ[μ] λ _, (μ set.univ).to_real⁻¹ • ∫ x, f x ∂μ := +begin + by_cases μ.ae.ne_bot, + { refine eventually_of_forall (λ x, _), + rw condexp_bot' f, + exact h, }, + { rw [ne_bot_iff, not_not, ae_eq_bot] at h, + simp only [h, ae_zero], }, +end + +lemma condexp_bot [is_probability_measure μ] (f : α → F') : + μ[f|⊥] = λ _, ∫ x, f x ∂μ := +by { refine (condexp_bot' f).trans _, rw [measure_univ, ennreal.one_to_real, inv_one, one_smul], } + lemma condexp_add (hf : integrable f μ) (hg : integrable g μ) : μ[f + g | m] =ᵐ[μ] μ[f|m] + μ[g|m] := begin diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index cf30dc9951d47..e24a038c63a6a 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -250,6 +250,26 @@ begin { rwa div_le_one (lt_of_le_of_ne (norm_nonneg _) (ne.symm h0)), }, }, end +lemma _root_.strongly_measurable_bot_iff [nonempty β] [t2_space β] : + strongly_measurable[⊥] f ↔ ∃ c, f = λ _, c := +begin + casesI is_empty_or_nonempty α with hα hα, + { simp only [subsingleton.strongly_measurable', eq_iff_true_of_subsingleton, exists_const], }, + refine ⟨λ hf, _, λ hf_eq, _⟩, + { refine ⟨f hα.some, _⟩, + let fs := hf.approx, + have h_fs_tendsto : ∀ x, tendsto (λ n, fs n x) at_top (𝓝 (f x)) := hf.tendsto_approx, + have : ∀ n, ∃ c, ∀ x, fs n x = c := λ n, simple_func.simple_func_bot (fs n), + let cs := λ n, (this n).some, + have h_cs_eq : ∀ n, ⇑(fs n) = (λ x, cs n) := λ n, funext (this n).some_spec, + simp_rw h_cs_eq at h_fs_tendsto, + have h_tendsto : tendsto cs at_top (𝓝 (f hα.some)) := h_fs_tendsto hα.some, + ext1 x, + exact tendsto_nhds_unique (h_fs_tendsto x) h_tendsto, }, + { obtain ⟨c, rfl⟩ := hf_eq, + exact strongly_measurable_const, }, +end + end basic_properties_in_any_topological_space lemma fin_strongly_measurable_of_set_sigma_finite [topological_space β] [has_zero β] diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 885c6d67a1a82..7280d3d89bd72 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -126,6 +126,33 @@ lemma range_const_subset (α) [measurable_space α] (b : β) : (const α b).range ⊆ {b} := finset.coe_subset.1 $ by simp +lemma simple_func_bot {α} (f : @simple_func α ⊥ β) [nonempty β] : ∃ c, ∀ x, f x = c := +begin + have hf_meas := @simple_func.measurable_set_fiber α _ ⊥ f, + simp_rw measurable_space.measurable_set_bot_iff at hf_meas, + casesI is_empty_or_nonempty α, + { simp only [is_empty.forall_iff, exists_const], }, + { specialize hf_meas (f h.some), + cases hf_meas, + { exfalso, + refine set.not_mem_empty h.some _, + rw [← hf_meas, set.mem_preimage], + exact set.mem_singleton _, }, + { refine ⟨f h.some, λ x, _⟩, + have : x ∈ f ⁻¹' {f h.some}, + { rw hf_meas, exact set.mem_univ x, }, + rwa [set.mem_preimage, set.mem_singleton_iff] at this, }, }, +end + +lemma simple_func_bot' {α} [nonempty β] (f : @simple_func α ⊥ β) : + ∃ c, f = @simple_func.const α _ ⊥ c := +begin + obtain ⟨c, h_eq⟩ := simple_func_bot f, + refine ⟨c, _⟩, + ext1 x, + rw [h_eq x, simple_func.coe_const], +end + lemma measurable_set_cut (r : α → β → Prop) (f : α →ₛ β) (h : ∀b, measurable_set {a | r a b}) : measurable_set {a | r a (f a)} := begin diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 3d455c8c0a2ee..b250280907da5 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -2219,6 +2219,13 @@ include m0 /-- A measure `μ` is called finite if `μ univ < ∞`. -/ class is_finite_measure (μ : measure α) : Prop := (measure_univ_lt_top : μ univ < ∞) +lemma not_is_finite_measure_iff : ¬ is_finite_measure μ ↔ μ set.univ = ∞ := +begin + refine ⟨λ h, _, λ h, λ h', h'.measure_univ_lt_top.ne h⟩, + by_contra h', + exact h ⟨lt_top_iff_ne_top.mpr h'⟩, +end + instance restrict.is_finite_measure (μ : measure α) [hs : fact (μ s < ∞)] : is_finite_measure (μ.restrict s) := ⟨by simp [hs.elim]⟩ @@ -2753,14 +2760,35 @@ lemma sigma_finite_of_le (μ : measure α) [hs : sigma_finite μ] end measure -include m0 - /-- Every finite measure is σ-finite. -/ @[priority 100] -instance is_finite_measure.to_sigma_finite (μ : measure α) [is_finite_measure μ] : +instance is_finite_measure.to_sigma_finite {m0 : measurable_space α} (μ : measure α) + [is_finite_measure μ] : sigma_finite μ := ⟨⟨⟨λ _, univ, λ _, trivial, λ _, measure_lt_top μ _, Union_const _⟩⟩⟩ +lemma sigma_finite_bot_iff (μ : @measure α ⊥) : sigma_finite μ ↔ is_finite_measure μ := +begin + refine ⟨λ h, ⟨_⟩, λ h, by { haveI := h, apply_instance, }⟩, + haveI : sigma_finite μ := h, + let s := spanning_sets μ, + have hs_univ : (⋃ i, s i) = set.univ := Union_spanning_sets μ, + have hs_meas : ∀ i, measurable_set[⊥] (s i) := measurable_spanning_sets μ, + simp_rw measurable_space.measurable_set_bot_iff at hs_meas, + by_cases h_univ_empty : set.univ = ∅, + { rw [h_univ_empty, measure_empty], exact ennreal.zero_ne_top.lt_top, }, + obtain ⟨i, hsi⟩ : ∃ i, s i = set.univ, + { by_contra h_not_univ, + push_neg at h_not_univ, + have h_empty : ∀ i, s i = ∅, by simpa [h_not_univ] using hs_meas, + simp [h_empty] at hs_univ, + exact h_univ_empty hs_univ.symm, }, + rw ← hsi, + exact measure_spanning_sets_lt_top μ i, +end + +include m0 + instance restrict.sigma_finite (μ : measure α) [sigma_finite μ] (s : set α) : sigma_finite (μ.restrict s) := begin @@ -3314,6 +3342,14 @@ begin ... < ∞ : measure_spanning_sets_lt_top _ _, }, end +lemma sigma_finite_trim_bot_iff : sigma_finite (μ.trim bot_le) ↔ is_finite_measure μ := +begin + rw sigma_finite_bot_iff, + refine ⟨λ h, ⟨_⟩, λ h, ⟨_⟩⟩; have h_univ := h.measure_univ_lt_top, + { rwa trim_measurable_set_eq bot_le measurable_set.univ at h_univ, }, + { rwa trim_measurable_set_eq bot_le measurable_set.univ, }, +end + end trim end measure_theory From d5ef06846236f927e7aedbd62fe990b4d62206ce Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 12 Sep 2022 11:04:50 +0000 Subject: [PATCH 228/828] feat(linear_algebra/finsupp): A variant of `finsupp.total` for `fintype`s. (#15462) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/linear_algebra/finsupp.lean | 58 +++++++++++++++++++++++++++++++++ 1 file changed, 58 insertions(+) diff --git a/src/linear_algebra/finsupp.lean b/src/linear_algebra/finsupp.lean index 76e48ee87c79f..533caf150664d 100644 --- a/src/linear_algebra/finsupp.lean +++ b/src/linear_algebra/finsupp.lean @@ -823,6 +823,64 @@ end prod end finsupp +section fintype + +variables {α M : Type*} (R : Type*) [fintype α] [semiring R] [add_comm_monoid M] [module R M] +variables (S : Type*) [semiring S] [module S M] [smul_comm_class R S M] +variable (v : α → M) + +/-- `fintype.total R S v f` is the linear combination of vectors in `v` with weights in `f`. +This variant of `finsupp.total` is defined on fintype indexed vectors. + +This map is linear in `v` if `R` is commutative, and always linear in `f`. +See note [bundled maps over different rings] for why separate `R` and `S` semirings are used. +-/ +protected def fintype.total : (α → M) →ₗ[S] (α → R) →ₗ[R] M := +{ to_fun := λ v, { to_fun := λ f, ∑ i, f i • v i, + map_add' := λ f g, by { simp_rw [← finset.sum_add_distrib, ← add_smul], refl }, + map_smul' := λ r f, by { simp_rw [finset.smul_sum, smul_smul], refl } }, + map_add' := λ u v, by { ext, simp [finset.sum_add_distrib, pi.add_apply, smul_add] }, + map_smul' := λ r v, by { ext, simp [finset.smul_sum, smul_comm _ r] } } + +variables {S} + +lemma fintype.total_apply (f) : fintype.total R S v f = ∑ i, f i • v i := rfl + +@[simp] +lemma fintype.total_apply_single (i : α) (r : R) : + fintype.total R S v (pi.single i r) = r • v i := +begin + simp_rw [fintype.total_apply, pi.single_apply, ite_smul, zero_smul], + rw [finset.sum_ite_eq', if_pos (finset.mem_univ _)] +end + +variables (S) + +lemma finsupp.total_eq_fintype_total_apply (x : α → R) : + finsupp.total α M R v ((finsupp.linear_equiv_fun_on_fintype R R α).symm x) = + fintype.total R S v x := +begin + apply finset.sum_subset, + { exact finset.subset_univ _ }, + { intros x _ hx, + rw finsupp.not_mem_support_iff.mp hx, + exact zero_smul _ _ } +end + +lemma finsupp.total_eq_fintype_total : + (finsupp.total α M R v).comp (finsupp.linear_equiv_fun_on_fintype R R α).symm.to_linear_map = + fintype.total R S v := +linear_map.ext $ finsupp.total_eq_fintype_total_apply R S v + +variables {S} + +@[simp] +lemma fintype.range_total : (fintype.total R S v).range = submodule.span R (set.range v) := +by rw [← finsupp.total_eq_fintype_total, linear_map.range_comp, + linear_equiv.to_linear_map_eq_coe, linear_equiv.range, submodule.map_top, finsupp.range_total] + +end fintype + variables {R : Type*} {M : Type*} {N : Type*} variables [semiring R] [add_comm_monoid M] [module R M] [add_comm_monoid N] [module R N] From 1c4e18434eeb5546b212e830b2b39de6a83c473c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 12 Sep 2022 11:04:51 +0000 Subject: [PATCH 229/828] =?UTF-8?q?refactor(set=5Ftheory/basic):=20match?= =?UTF-8?q?=20`x=20<=20=E2=84=B5=E2=82=80`=20lemmas=20with=20`x=20?= =?UTF-8?q?=E2=89=A4=20=E2=84=B5=E2=82=80`=20lemmas=20(#15662)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Mark `cardinal.lt_aleph_0_iff_set_finite` as `@[simp]` lemma. * Add `cardinal.lt_aleph_0_iff_subtype_finite` and `cardinal.mk_le_aleph_0_iff`; drop `cardinal.encodable_iff`. * Rename `cardinal.mk_set_le_aleph_0` to `cardinal.le_aleph_0_iff_set_countable`. * Rename `cardinal.mk_subtype_le_aleph_0` to `cardinal.le_aleph_0_iff_subtype_countable`. * Make `first_order.language.countable` protected. * Use `[countable _]` instead of `[encodable _]` or `[nonempty (encodable _)]` here and there. Co-authored-by: Yury G. Kudryashov --- src/algebra/algebraic_card.lean | 2 +- src/analysis/complex/cauchy_integral.lean | 3 +- src/data/complex/cardinality.lean | 2 +- src/data/real/cardinality.lean | 2 +- src/model_theory/basic.lean | 12 +++++--- src/model_theory/encoding.lean | 13 +++----- src/model_theory/finitely_generated.lean | 10 +++--- src/model_theory/fraisse.lean | 4 +-- src/model_theory/substructures.lean | 6 ++-- src/set_theory/cardinal/basic.lean | 37 ++++++++++++----------- src/set_theory/cardinal/ordinal.lean | 2 +- 11 files changed, 47 insertions(+), 46 deletions(-) diff --git a/src/algebra/algebraic_card.lean b/src/algebra/algebraic_card.lean index f2595c32199a5..6a84a8c238ff5 100644 --- a/src/algebra/algebraic_card.lean +++ b/src/algebra/algebraic_card.lean @@ -73,7 +73,7 @@ variable [encodable R] @[simp] theorem countable_of_encodable : set.countable {x : A | is_algebraic R x} := begin - rw [←mk_set_le_aleph_0, ←lift_le], + rw [←le_aleph_0_iff_set_countable, ←lift_le], apply (cardinal_mk_lift_le_max R A).trans, simp end diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index d109234451380..02821474cb302 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -477,7 +477,8 @@ begin { refine nonempty_diff.2 (λ hsub, _), have : (Ioo l u).countable, from (hs.preimage ((add_right_injective w).comp of_real_injective)).mono hsub, - rw [← cardinal.mk_set_le_aleph_0, cardinal.mk_Ioo_real (hlu₀.1.trans hlu₀.2)] at this, + rw [← cardinal.le_aleph_0_iff_set_countable, + cardinal.mk_Ioo_real (hlu₀.1.trans hlu₀.2)] at this, exact this.not_lt cardinal.aleph_0_lt_continuum }, exact ⟨g x, (hlu_sub hx.1).1, (hlu_sub hx.1).2, hx.2⟩ end diff --git a/src/data/complex/cardinality.lean b/src/data/complex/cardinality.lean index ddd2f1d659f15..1bbabbe9cde5e 100644 --- a/src/data/complex/cardinality.lean +++ b/src/data/complex/cardinality.lean @@ -27,4 +27,4 @@ by rw [mk_univ, mk_complex] /-- The complex numbers are not countable. -/ lemma not_countable_complex : ¬ (set.univ : set ℂ).countable := -by { rw [← mk_set_le_aleph_0, not_le, mk_univ_complex], apply cantor } +by { rw [← le_aleph_0_iff_set_countable, not_le, mk_univ_complex], apply cantor } diff --git a/src/data/real/cardinality.lean b/src/data/real/cardinality.lean index 4fa0609a4fa4a..bc8e478aee486 100644 --- a/src/data/real/cardinality.lean +++ b/src/data/real/cardinality.lean @@ -164,7 +164,7 @@ by rw [mk_univ, mk_real] /-- **Non-Denumerability of the Continuum**: The reals are not countable. -/ lemma not_countable_real : ¬ (set.univ : set ℝ).countable := -by { rw [← mk_set_le_aleph_0, not_le, mk_univ_real], apply cantor } +by { rw [← le_aleph_0_iff_set_countable, not_le, mk_univ_real], apply cantor } /-- The cardinality of the interval (a, ∞). -/ lemma mk_Ioi_real (a : ℝ) : #(Ioi a) = 𝔠 := diff --git a/src/model_theory/basic.lean b/src/model_theory/basic.lean index 33dd3d5c7da4c..07a2c899efc82 100644 --- a/src/model_theory/basic.lean +++ b/src/model_theory/basic.lean @@ -35,6 +35,10 @@ structures. to the `L`-structure `N` that commutes with the interpretations of functions, and which preserves the interpretations of relations in both directions. +## TODO + +Use `[countable L.symbols]` instead of `[L.countable]`. + ## References For the Flypitch project: - [J. Han, F. van Doorn, *A formal proof of the independence of the continuum hypothesis*] @@ -125,7 +129,7 @@ rfl def card : cardinal := # L.symbols /-- A language is countable when it has countably many symbols. -/ -class countable : Prop := (card_le_aleph_0' : L.card ≤ ℵ₀) +@[protected] class countable : Prop := (card_le_aleph_0' : L.card ≤ ℵ₀) lemma card_le_aleph_0 [L.countable] : L.card ≤ ℵ₀ := countable.card_le_aleph_0' @@ -194,8 +198,8 @@ instance subsingleton_mk₂_relations {c f₁ f₂ : Type u} {r₁ r₂ : Type v nat.cases_on n ⟨λ x, pempty.elim x⟩ (λ n, nat.cases_on n h1 (λ n, nat.cases_on n h2 (λ n, ⟨λ x, pempty.elim x⟩))) -lemma encodable.countable [h : encodable L.symbols] : L.countable := -⟨cardinal.encodable_iff.1 ⟨h⟩⟩ +lemma encodable.countable [_root_.countable L.symbols] : L.countable := +⟨cardinal.mk_le_aleph_0⟩ @[simp] lemma empty_card : language.empty.card = 0 := by simp [card_eq_card_functions_add_card_relations] @@ -211,7 +215,7 @@ instance countable_empty : language.empty.countable := end⟩ lemma encodable.countable_functions [h : encodable (Σl, L.functions l)] : L.countable_functions := -⟨cardinal.encodable_iff.1 ⟨h⟩⟩ +⟨cardinal.mk_le_aleph_0⟩ @[priority 100] instance is_relational.countable_functions [L.is_relational] : L.countable_functions := diff --git a/src/model_theory/encoding.lean b/src/model_theory/encoding.lean index 41b3d2bf7835f..10fdaa84c80ba 100644 --- a/src/model_theory/encoding.lean +++ b/src/model_theory/encoding.lean @@ -117,13 +117,10 @@ begin rw [mk_nat, lift_aleph_0, mul_eq_max_of_aleph_0_le_left le_rfl, max_le_iff, csupr_le_iff' (bdd_above_range _)], { refine ⟨le_max_left _ _, λ i, card_le.trans _⟩, - rw max_le_iff, - refine ⟨le_max_left _ _, _⟩, + refine max_le (le_max_left _ _) _, rw [← add_eq_max le_rfl, mk_sum, mk_sum, mk_sum, add_comm (cardinal.lift (#α)), lift_add, - add_assoc, lift_lift, lift_lift], - refine add_le_add_right _ _, - rw [lift_le_aleph_0, ← encodable_iff], - exact ⟨infer_instance⟩ }, + add_assoc, lift_lift, lift_lift, mk_fin, lift_nat_cast], + exact add_le_add_right (nat_lt_aleph_0 _).le _ }, { rw [← one_le_iff_ne_zero], refine trans _ (le_csupr (bdd_above_range _) 1), rw [one_le_iff_ne_zero, mk_ne_zero_iff], @@ -154,13 +151,13 @@ encodable.of_left_injection list_encode (λ l, (list_decode l).head'.join) simp only [option.join, head', list.map, option.some_bind, id.def], end) -lemma card_le_aleph_0 [h1 : nonempty (encodable α)] [h2 : L.countable_functions] : +lemma card_le_aleph_0 [h1 : countable α] [h2 : L.countable_functions] : # (L.term α) ≤ ℵ₀ := begin refine (card_le.trans _), rw [max_le_iff], simp only [le_refl, mk_sum, add_le_aleph_0, lift_le_aleph_0, true_and], - exact ⟨encodable_iff.1 h1, L.card_functions_le_aleph_0⟩, + exact ⟨mk_le_aleph_0, L.card_functions_le_aleph_0⟩, end instance small [small.{u} α] : diff --git a/src/model_theory/finitely_generated.lean b/src/model_theory/finitely_generated.lean index dc9e22e225e22..19d55eac3ae5f 100644 --- a/src/model_theory/finitely_generated.lean +++ b/src/model_theory/finitely_generated.lean @@ -157,9 +157,9 @@ begin end theorem cg_iff_countable [L.countable_functions] {s : L.substructure M} : - s.cg ↔ nonempty (encodable s) := + s.cg ↔ countable s := begin - refine ⟨_, λ h, ⟨s, h, s.closure_eq⟩⟩, + refine ⟨_, λ h, ⟨s, h.to_set, s.closure_eq⟩⟩, rintro ⟨s, h, rfl⟩, exact h.substructure_closure L end @@ -224,10 +224,8 @@ begin exact h.range f, end -lemma cg_iff_countable [L.countable_functions] : - cg L M ↔ nonempty (encodable M) := -by rw [cg_def, cg_iff_countable, cardinal.encodable_iff, cardinal.encodable_iff, - top_equiv.to_equiv.cardinal_eq] +lemma cg_iff_countable [L.countable_functions] : cg L M ↔ countable M := +by rw [cg_def, cg_iff_countable, top_equiv.to_equiv.countable_iff] lemma fg.cg (h : fg L M) : cg L M := cg_def.2 (fg_def.1 h).cg diff --git a/src/model_theory/fraisse.lean b/src/model_theory/fraisse.lean index 4b8b79d05201f..2415f5397cd1a 100644 --- a/src/model_theory/fraisse.lean +++ b/src/model_theory/fraisse.lean @@ -237,8 +237,8 @@ begin age.hereditary M, age.joint_embedding M⟩, }, { rintros ⟨Kn, eqinv, cq, hfg, hp, jep⟩, obtain ⟨M, hM, rfl⟩ := exists_cg_is_age_of Kn eqinv cq hfg hp jep, - haveI := ((Structure.cg_iff_countable).1 hM).some, - refine ⟨M, to_countable _, rfl⟩, } + haveI : countable M := Structure.cg_iff_countable.1 hM, + exact ⟨M, to_countable _, rfl⟩, } end variables {K} (L) (M) diff --git a/src/model_theory/substructures.lean b/src/model_theory/substructures.lean index 7104f63c491a3..e2411eb0f75d8 100644 --- a/src/model_theory/substructures.lean +++ b/src/model_theory/substructures.lean @@ -274,10 +274,10 @@ variable (L) lemma _root_.set.countable.substructure_closure [L.countable_functions] (h : s.countable) : - nonempty (encodable (closure L s)) := + countable (closure L s) := begin - haveI : nonempty (encodable s) := h, - rw [encodable_iff, ← lift_le_aleph_0], + haveI : countable s := h.to_subtype, + rw [← mk_le_aleph_0_iff, ← lift_le_aleph_0], exact lift_card_closure_le_card_term.trans term.card_le_aleph_0, end diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index e94aa2580d581..e7b7eefa30b7c 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -979,11 +979,29 @@ lt_aleph_0_iff_finite.trans (finite_iff_nonempty_fintype _) theorem lt_aleph_0_of_finite (α : Type u) [finite α] : #α < ℵ₀ := lt_aleph_0_iff_finite.2 ‹_› -theorem lt_aleph_0_iff_set_finite {α} {S : set α} : #S < ℵ₀ ↔ S.finite := +@[simp] theorem lt_aleph_0_iff_set_finite {S : set α} : #S < ℵ₀ ↔ S.finite := lt_aleph_0_iff_finite.trans finite_coe_iff alias lt_aleph_0_iff_set_finite ↔ _ _root_.set.finite.lt_aleph_0 +@[simp] theorem lt_aleph_0_iff_subtype_finite {p : α → Prop} : + #{x // p x} < ℵ₀ ↔ {x | p x}.finite := +lt_aleph_0_iff_set_finite + +lemma mk_le_aleph_0_iff : #α ≤ ℵ₀ ↔ countable α := +by rw [countable_iff_nonempty_embedding, aleph_0, ← lift_uzero (#α), lift_mk_le'] + +@[simp] lemma mk_le_aleph_0 [countable α] : #α ≤ ℵ₀ := mk_le_aleph_0_iff.mpr ‹_› + +@[simp] lemma le_aleph_0_iff_set_countable {s : set α} : #s ≤ ℵ₀ ↔ s.countable := +by rw [mk_le_aleph_0_iff, countable_coe_iff] + +alias le_aleph_0_iff_set_countable ↔ _ _root_.set.countable.le_aleph_0 + +@[simp] lemma le_aleph_0_iff_subtype_countable {p : α → Prop} : + #{x // p x} ≤ ℵ₀ ↔ {x | p x}.countable := +le_aleph_0_iff_set_countable + instance can_lift_cardinal_nat : can_lift cardinal ℕ := ⟨ coe, λ x, x < ℵ₀, λ x hx, let ⟨n, hn⟩ := lt_aleph_0.mp hx in ⟨n, hn.symm⟩⟩ @@ -1066,13 +1084,6 @@ by rw [← not_lt, lt_aleph_0_iff_finite, not_finite_iff_infinite] @[simp] lemma aleph_0_le_mk (α : Type u) [infinite α] : ℵ₀ ≤ #α := infinite_iff.1 ‹_› -lemma encodable_iff {α : Type u} : nonempty (encodable α) ↔ #α ≤ ℵ₀ := -⟨λ ⟨h⟩, ⟨(@encodable.encode' α h).trans equiv.ulift.symm.to_embedding⟩, - λ ⟨h⟩, ⟨encodable.of_inj _ (h.trans equiv.ulift.to_embedding).injective⟩⟩ - -@[simp] lemma mk_le_aleph_0 [countable α] : #α ≤ ℵ₀ := -encodable_iff.1 $ encodable.nonempty_encodable.2 ‹_› - lemma denumerable_iff {α : Type u} : nonempty (denumerable α) ↔ #α = ℵ₀ := ⟨λ ⟨h⟩, mk_congr ((@denumerable.eqv α h).trans equiv.ulift.symm), λ h, by { cases quotient.exact h with f, exact ⟨denumerable.mk' $ f.trans equiv.ulift⟩ }⟩ @@ -1080,16 +1091,6 @@ lemma denumerable_iff {α : Type u} : nonempty (denumerable α) ↔ #α = ℵ₀ @[simp] lemma mk_denumerable (α : Type u) [denumerable α] : #α = ℵ₀ := denumerable_iff.1 ⟨‹_›⟩ -@[simp] lemma mk_set_le_aleph_0 (s : set α) : #s ≤ ℵ₀ ↔ s.countable := -begin - rw [set.countable_iff_exists_injective], split, - { rintro ⟨f'⟩, cases embedding.trans f' equiv.ulift.to_embedding with f hf, exact ⟨f, hf⟩ }, - { rintro ⟨f, hf⟩, exact ⟨embedding.trans ⟨f, hf⟩ equiv.ulift.symm.to_embedding⟩ } -end - -@[simp] lemma mk_subtype_le_aleph_0 (p : α → Prop) : #{x // p x} ≤ ℵ₀ ↔ {x | p x}.countable := -mk_set_le_aleph_0 _ - @[simp] lemma aleph_0_add_aleph_0 : ℵ₀ + ℵ₀ = ℵ₀ := mk_denumerable _ lemma aleph_0_mul_aleph_0 : ℵ₀ * ℵ₀ = ℵ₀ := mk_denumerable _ diff --git a/src/set_theory/cardinal/ordinal.lean b/src/set_theory/cardinal/ordinal.lean index 676dda5722659..7503ee6a185ad 100644 --- a/src/set_theory/cardinal/ordinal.lean +++ b/src/set_theory/cardinal/ordinal.lean @@ -284,7 +284,7 @@ lemma aleph_0_lt_aleph_one : ℵ₀ < aleph 1 := by { rw ←succ_aleph_0, apply lt_succ } lemma countable_iff_lt_aleph_one {α : Type*} (s : set α) : s.countable ↔ #s < aleph 1 := -by rw [←succ_aleph_0, lt_succ_iff, mk_set_le_aleph_0] +by rw [←succ_aleph_0, lt_succ_iff, le_aleph_0_iff_set_countable] /-- Ordinals that are cardinals are unbounded. -/ theorem ord_card_unbounded : unbounded (<) {b : ordinal | b.card.ord = b} := From e1847377ff725977bd916fc2214abf20effb4783 Mon Sep 17 00:00:00 2001 From: Wrenna Robson Date: Mon, 12 Sep 2022 11:04:52 +0000 Subject: [PATCH 230/828] feat(data/real/ennreal): Add `to_real_{min, sup, inf}` (#16233) This adds `to_real_min` to match `to_real_max`, and adds `inf` and `sup` spellings of these lemmas. Co-authored-by: Wrenna Robson <34025592+linesthatinterlace@users.noreply.github.com> --- src/data/real/ennreal.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index c518cfdd1cc10..99a303554d9a8 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -1549,6 +1549,18 @@ lemma to_real_max (hr : a ≠ ∞) (hp : b ≠ ∞) : (λ h, by simp only [h, (ennreal.to_real_le_to_real hr hp).2 h, max_eq_right]) (λ h, by simp only [h, (ennreal.to_real_le_to_real hp hr).2 h, max_eq_left]) +lemma to_real_min {a b : ℝ≥0∞} (hr : a ≠ ∞) (hp : b ≠ ∞) : + ennreal.to_real (min a b) = min (ennreal.to_real a) (ennreal.to_real b) := +(le_total a b).elim + (λ h, by simp only [h, (ennreal.to_real_le_to_real hr hp).2 h, min_eq_left]) + (λ h, by simp only [h, (ennreal.to_real_le_to_real hp hr).2 h, min_eq_right]) + +lemma to_real_sup {a b : ℝ≥0∞} + : a ≠ ∞ → b ≠ ∞ → (a ⊔ b).to_real = a.to_real ⊔ b.to_real := to_real_max + +lemma to_real_inf {a b : ℝ≥0∞} + : a ≠ ∞ → b ≠ ∞ → (a ⊓ b).to_real = a.to_real ⊓ b.to_real := to_real_min + lemma to_nnreal_pos_iff : 0 < a.to_nnreal ↔ (0 < a ∧ a < ∞) := by { induction a using with_top.rec_top_coe; simp } From 28b23c9435ea2e2db1e1173d231b4cad83ccb29e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 12 Sep 2022 13:12:22 +0000 Subject: [PATCH 231/828] feat(analysis/normed/group/seminorm): Group norms (#16237) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define (additive) group norms, which are group seminorms `f` such that `f x = 0 → x = 0` (resp. `f x = 0 → x = 1`), along with their hom classes. --- src/analysis/normed/group/seminorm.lean | 180 ++++++++++++++++++++++-- 1 file changed, 171 insertions(+), 9 deletions(-) diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index 67890137b6ace..792d03dc20ec5 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -9,8 +9,8 @@ import data.real.nnreal /-! # Group seminorms -This file defines seminorms in a group. A group seminorm is a function to the reals which is -positive-semidefinite and subadditive. +This file defines norms and seminorms in a group. A group seminorm is a function to the reals which +is positive-semidefinite and subadditive. A norm further only maps zero to zero. ## Main declarations @@ -18,6 +18,8 @@ positive-semidefinite and subadditive. takes nonnegative values, is subadditive and such that `f (-x) = f x` for all `x`. * `group_seminorm`: A function `f` from a group `G` to the reals that sends one to zero, takes nonnegative values, is submultiplicative and such that `f x⁻¹ = f x` for all `x`. +* `add_group_norm`: A seminorm `f` such that `f x = 0 → x = 0` for all `x`. +* `group_norm`: A seminorm `f` such that `f x = 0 → x = 1` for all `x`. ## References @@ -25,7 +27,7 @@ positive-semidefinite and subadditive. ## Tags -seminorm +norm, seminorm -/ set_option old_structure_cmd true @@ -50,7 +52,20 @@ structure group_seminorm (G : Type*) [group G] := (mul_le' : ∀ x y, to_fun (x * y) ≤ to_fun x + to_fun y) (inv' : ∀ x, to_fun x⁻¹ = to_fun x) -attribute [nolint doc_blame] add_group_seminorm.to_zero_hom +/-- A norm on an additive group `G` is a function `f : G → ℝ` that preserves zero, is subadditive +and such that `f (-x) = f x` and `f x = 0 → x = 0` for all `x`. -/ +@[protect_proj] +structure add_group_norm (G : Type*) [add_group G] extends add_group_seminorm G := +(eq_zero_of_map_eq_zero' : ∀ x, to_fun x = 0 → x = 0) + +/-- A seminorm on a group `G` is a function `f : G → ℝ` that sends one to zero, is submultiplicative +and such that `f x⁻¹ = f x` and `f x = 0 → x = 1` for all `x`. -/ +@[protect_proj, to_additive] +structure group_norm (G : Type*) [group G] extends group_seminorm G := +(eq_one_of_map_eq_zero' : ∀ x, to_fun x = 0 → x = 1) + +attribute [nolint doc_blame] add_group_seminorm.to_zero_hom add_group_norm.to_add_group_seminorm + group_norm.to_group_seminorm /-- `add_group_seminorm_class F α` states that `F` is a type of seminorms on the additive group `α`. @@ -69,21 +84,39 @@ class group_seminorm_class (F : Type*) (α : out_param $ Type*) [group α] (map_one_eq_zero (f : F) : f 1 = 0) (map_inv_eq_map (f : F) (a : α) : f a⁻¹ = f a) -attribute [to_additive] group_seminorm_class.to_mul_le_add_hom_class +/-- `add_group_norm_class F α` states that `F` is a type of norms on the additive group `α`. + +You should extend this class when you extend `add_group_norm`. -/ +class add_group_norm_class (F : Type*) (α : out_param $ Type*) [add_group α] + extends add_group_seminorm_class F α := +(eq_zero_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 0) + +/-- `group_norm_class F α` states that `F` is a type of norms on the group `α`. + +You should extend this class when you extend `group_norm`. -/ +@[to_additive] +class group_norm_class (F : Type*) (α : out_param $ Type*) [group α] + extends group_seminorm_class F α := +(eq_one_of_map_eq_zero (f : F) {a : α} : f a = 0 → a = 1) export add_group_seminorm_class (map_neg_eq_map) -export group_seminorm_class (map_one_eq_zero map_inv_eq_map) + group_seminorm_class (map_one_eq_zero map_inv_eq_map) + add_group_norm_class (eq_zero_of_map_eq_zero) + group_norm_class (eq_one_of_map_eq_zero) attribute [simp, to_additive map_zero] map_one_eq_zero attribute [simp] map_neg_eq_map attribute [simp, to_additive] map_inv_eq_map +attribute [to_additive] group_seminorm_class.to_mul_le_add_hom_class +attribute [to_additive] group_norm.to_group_seminorm +attribute [to_additive] group_norm_class.to_group_seminorm_class @[priority 100] -- See note [lower instance priority] instance add_group_seminorm_class.to_zero_hom_class [add_group E] [add_group_seminorm_class F E] : zero_hom_class F E ℝ := { ..‹add_group_seminorm_class F E› } -section group +section group_seminorm_class variables [group E] [group_seminorm_class F E] (f : F) (x y : E) include E @@ -95,7 +128,7 @@ by { rw [div_eq_mul_inv, ←map_inv_eq_map f y], exact map_mul_le_add _ _ _ } @[to_additive] lemma le_map_add_map_div' : f x ≤ f y + f (y / x) := by simpa only [add_comm, map_div_rev, div_mul_cancel'] using map_mul_le_add f (x / y) y -end group +end group_seminorm_class @[to_additive, priority 100] -- See note [lower instance priority] instance group_seminorm_class.to_nonneg_hom_class [group E] [group_seminorm_class F E] : @@ -104,6 +137,22 @@ instance group_seminorm_class.to_nonneg_hom_class [group E] [group_seminorm_clas (by { rw [two_mul, ←map_one_eq_zero f, ←div_self' a], exact map_div_le_add _ _ _ }) two_pos, ..‹group_seminorm_class F E› } +section group_norm_class +variables [group E] [group_norm_class F E] (f : F) {x : E} +include E + +@[to_additive] lemma map_pos_of_ne_one (hx : x ≠ 1) : 0 < f x := +(map_nonneg _ _).lt_of_ne $ λ h, hx $ eq_one_of_map_eq_zero _ h.symm + +@[simp, to_additive] lemma map_eq_zero_iff_eq_one : f x = 0 ↔ x = 1 := +⟨eq_one_of_map_eq_zero _, by { rintro rfl, exact map_one_eq_zero _ }⟩ + +@[to_additive] lemma map_ne_zero_iff_ne_one : f x ≠ 0 ↔ x ≠ 1 := (map_eq_zero_iff_eq_one _).not + +end group_norm_class + +/-! ### Seminorms -/ + namespace group_seminorm section group variables [group E] [group F] [group G] {p q : group_seminorm E} @@ -205,7 +254,7 @@ variables [comm_group E] [comm_group F] (p q : group_seminorm E) (x y : E) @[to_additive] lemma comp_mul_le (f g : F →* E) : p.comp (f * g) ≤ p.comp f + p.comp g := λ _, map_mul_le_add p _ _ -@[to_additive] private lemma mul_bdd_below_range_add {p q : group_seminorm E} {x : E} : +@[to_additive] lemma mul_bdd_below_range_add {p q : group_seminorm E} {x : E} : bdd_below (range $ λ y, p y + q (x / y)) := ⟨0, by { rintro _ ⟨x, rfl⟩, exact add_nonneg (map_nonneg p _) (map_nonneg q _) }⟩ @@ -245,6 +294,20 @@ namespace add_group_seminorm variables [add_group E] [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] (p : add_group_seminorm E) +instance [decidable_eq E] : has_one (add_group_seminorm E) := +⟨{ to_fun := λ x, if x = 0 then 0 else 1, + map_zero' := if_pos rfl, + add_le' := λ x y, begin + by_cases hx : x = 0, + { rw [if_pos hx, hx, zero_add, zero_add] }, + { rw if_neg hx, + refine le_add_of_le_of_nonneg _ _; split_ifs; norm_num } + end, + neg' := λ x, by simp_rw neg_eq_zero }⟩ + +@[simp] lemma apply_one [decidable_eq E] (x : E) : + (1 : add_group_seminorm E) x = if x = 0 then 0 else 1 := rfl + /-- Any action on `ℝ` which factors through `ℝ≥0` applies to an `add_group_seminorm`. -/ instance : has_smul R (add_group_seminorm E) := ⟨λ r p, @@ -277,6 +340,21 @@ end add_group_seminorm namespace group_seminorm variables [group E] [has_smul R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] +@[to_additive add_group_seminorm.has_one] +instance [decidable_eq E] : has_one (group_seminorm E) := +⟨{ to_fun := λ x, if x = 1 then 0 else 1, + map_one' := if_pos rfl, + mul_le' := λ x y, begin + by_cases hx : x = 1, + { rw [if_pos hx, hx, one_mul, zero_add] }, + { rw if_neg hx, + refine le_add_of_le_of_nonneg _ _; split_ifs; norm_num } + end, + inv' := λ x, by simp_rw inv_eq_one }⟩ + +@[simp, to_additive add_group_seminorm.apply_one] lemma apply_one [decidable_eq E] (x : E) : + (1 : group_seminorm E) x = if x = 1 then 0 else 1 := rfl + /-- Any action on `ℝ` which factors through `ℝ≥0` applies to an `add_group_seminorm`. -/ @[to_additive add_group_seminorm.has_smul] instance : has_smul R (group_seminorm E) := ⟨λ r p, @@ -309,3 +387,87 @@ from λ x y, by simpa only [←smul_eq_mul, ←nnreal.smul_def, smul_one_smul ext $ λ x, real.smul_max _ _ end group_seminorm + +/-! ### Norms -/ + +namespace group_norm +section group +variables [group E] [group F] [group G] {p q : group_norm E} + +@[to_additive] instance group_norm_class : group_norm_class (group_norm E) E := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_one_eq_zero := λ f, f.map_one', + map_mul_le_add := λ f, f.mul_le', + map_inv_eq_map := λ f, f.inv', + eq_one_of_map_eq_zero := λ f, f.eq_one_of_map_eq_zero' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun` +directly. -/ +@[to_additive "Helper instance for when there's too many metavariables to apply +`fun_like.has_coe_to_fun` directly. "] +instance : has_coe_to_fun (group_norm E) (λ _, E → ℝ) := fun_like.has_coe_to_fun + +@[ext, to_additive] lemma ext : (∀ x, p x = q x) → p = q := fun_like.ext p q + +@[to_additive] instance : partial_order (group_norm E) := +partial_order.lift _ fun_like.coe_injective + +@[to_additive] lemma le_def : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl +@[to_additive] lemma lt_def : p < q ↔ (p : E → ℝ) < q := iff.rfl + +@[simp, to_additive] lemma coe_le_coe : (p : E → ℝ) ≤ q ↔ p ≤ q := iff.rfl +@[simp, to_additive] lemma coe_lt_coe : (p : E → ℝ) < q ↔ p < q := iff.rfl + +variables (p q) (f : F →* E) + +@[to_additive] instance : has_add (group_norm E) := +⟨λ p q, { eq_one_of_map_eq_zero' := λ x hx, of_not_not $ λ h, + hx.not_gt $ add_pos (map_pos_of_ne_one p h) (map_pos_of_ne_one q h), + ..p.to_group_seminorm + q.to_group_seminorm }⟩ + +@[simp, to_additive] lemma coe_add : ⇑(p + q) = p + q := rfl +@[simp, to_additive] lemma add_apply (x : E) : (p + q) x = p x + q x := rfl + +-- TODO: define `has_Sup` +@[to_additive] noncomputable instance : has_sup (group_norm E) := +⟨λ p q, + { eq_one_of_map_eq_zero' := λ x hx, of_not_not $ λ h, hx.not_gt $ + lt_sup_iff.2 $ or.inl $ map_pos_of_ne_one p h, + ..p.to_group_seminorm ⊔ q.to_group_seminorm }⟩ + +@[simp, to_additive] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl +@[simp, to_additive] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl + +@[to_additive] noncomputable instance : semilattice_sup (group_norm E) := +fun_like.coe_injective.semilattice_sup _ coe_sup + +end group +end group_norm + +namespace add_group_norm +variables [add_group E] [decidable_eq E] + +instance : has_one (add_group_norm E) := +⟨{ eq_zero_of_map_eq_zero' := λ x, zero_ne_one.ite_eq_left_iff.1, + ..(1 : add_group_seminorm E) }⟩ + +@[simp] lemma apply_one (x : E) : (1 : add_group_norm E) x = if x = 0 then 0 else 1 := rfl + +instance : inhabited (add_group_norm E) := ⟨1⟩ + +end add_group_norm + +namespace group_norm +variables [group E] [decidable_eq E] + +@[to_additive add_group_norm.has_one] instance : has_one (group_norm E) := +⟨{ eq_one_of_map_eq_zero' := λ x, zero_ne_one.ite_eq_left_iff.1, + ..(1 : group_seminorm E) }⟩ + +@[simp, to_additive add_group_norm.apply_one] +lemma apply_one (x : E) : (1 : group_norm E) x = if x = 1 then 0 else 1 := rfl + +@[to_additive] instance : inhabited (group_norm E) := ⟨1⟩ + +end group_norm From 86be023b16b433dcd61a816870b46848f1c95701 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Mon, 12 Sep 2022 13:12:23 +0000 Subject: [PATCH 232/828] feat(data/polynomial/algebra_map): add `aeval_X_left_apply` (#16433) --- src/data/polynomial/algebra_map.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/data/polynomial/algebra_map.lean b/src/data/polynomial/algebra_map.lean index 296dde611bffa..fc08bdd5bb11f 100644 --- a/src/data/polynomial/algebra_map.lean +++ b/src/data/polynomial/algebra_map.lean @@ -215,6 +215,9 @@ alg_hom_ext $ by simp only [aeval_X, alg_hom.comp_apply] @[simp] theorem aeval_X_left : aeval (X : R[X]) = alg_hom.id R R[X] := alg_hom_ext $ aeval_X X +theorem aeval_X_left_apply (p : R[X]) : aeval X p = p := +alg_hom.congr_fun (@aeval_X_left R _) p + theorem eval_unique (φ : R[X] →ₐ[R] A) (p) : φ p = eval₂ (algebra_map R A) (φ X) p := by rw [← aeval_def, aeval_alg_hom, aeval_X_left, alg_hom.comp_id] From 60f72157418bc571ab9bca5f6460c995c6128f76 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 12 Sep 2022 13:12:24 +0000 Subject: [PATCH 233/828] chore(data/real/cau_seq_completion): golf using existing lemmas (#16469) This extracts `cau_seq.mul_equiv_mul` and `cau_seq.sub_equiv_sub` into standalone lemmas, and uses the existing lemmas for `add` and `neg` rather than reproving the result from scratch. --- src/data/real/cau_seq.lean | 23 +++++++++++------------ src/data/real/cau_seq_completion.lean | 18 ++++-------------- 2 files changed, 15 insertions(+), 26 deletions(-) diff --git a/src/data/real/cau_seq.lean b/src/data/real/cau_seq.lean index 205e834413fa5..6d5d9ec430435 100644 --- a/src/data/real/cau_seq.lean +++ b/src/data/real/cau_seq.lean @@ -321,20 +321,14 @@ instance equiv : setoid (cau_seq β abv) := lemma add_equiv_add {f1 f2 g1 g2 : cau_seq β abv} (hf : f1 ≈ f2) (hg : g1 ≈ g2) : f1 + g1 ≈ f2 + g2 := -begin - change lim_zero ((f1 + g1) - _), - convert add_lim_zero hf hg using 1, - simp only [sub_eq_add_neg, add_assoc], - rw add_comm (-f2), simp only [add_assoc], - congr' 2, simp -end +by simpa only [←add_sub_add_comm] using add_lim_zero hf hg lemma neg_equiv_neg {f g : cau_seq β abv} (hf : f ≈ g) : -f ≈ -g := -begin - show lim_zero (-f - -g), - rw ←neg_sub', - exact neg_lim_zero hf, -end +by simpa only [neg_sub'] using neg_lim_zero hf + +lemma sub_equiv_sub {f1 f2 g1 g2 : cau_seq β abv} (hf : f1 ≈ f2) (hg : g1 ≈ g2) : + f1 - g1 ≈ f2 - g2 := +by simpa only [sub_eq_add_neg] using add_equiv_add hf (neg_equiv_neg hg) theorem equiv_def₃ {f g : cau_seq β abv} (h : f ≈ g) {ε : α} (ε0 : 0 < ε) : ∃ i, ∀ j ≥ i, ∀ k ≥ j, abv (f k - g j) < ε := @@ -417,6 +411,11 @@ variables {β : Type*} [comm_ring β] {abv : β → α} [is_absolute_value abv] lemma mul_equiv_zero' (g : cau_seq _ abv) {f : cau_seq _ abv} (hf : f ≈ 0) : f * g ≈ 0 := by rw mul_comm; apply mul_equiv_zero _ hf +lemma mul_equiv_mul {f1 f2 g1 g2 : cau_seq β abv} (hf : f1 ≈ f2) (hg : g1 ≈ g2) : + f1 * g1 ≈ f2 * g2 := +by simpa only [mul_sub, mul_comm, sub_add_sub_cancel] + using add_lim_zero (mul_lim_zero_right g1 hf) (mul_lim_zero_right f2 hg) + end comm_ring section is_domain diff --git a/src/data/real/cau_seq_completion.lean b/src/data/real/cau_seq_completion.lean index 30c208f6672bf..d0f1dfc984bfc 100644 --- a/src/data/real/cau_seq_completion.lean +++ b/src/data/real/cau_seq_completion.lean @@ -44,32 +44,22 @@ by have : mk f = 0 ↔ lim_zero (f - 0) := quotient.eq; rwa sub_zero at this instance : has_add Cauchy := -⟨λ x y, quotient.lift_on₂ x y (λ f g, mk (f + g)) $ - λ f₁ g₁ f₂ g₂ hf hg, quotient.sound $ - by simpa [(≈), setoid.r, sub_eq_add_neg, add_comm, add_left_comm, add_assoc] - using add_lim_zero hf hg⟩ +⟨quotient.map₂ (+) $ λ f₁ g₁ hf f₂ g₂ hg, add_equiv_add hf hg⟩ @[simp] theorem mk_add (f g : cau_seq β abv) : mk f + mk g = mk (f + g) := rfl instance : has_neg Cauchy := -⟨λ x, quotient.lift_on x (λ f, mk (-f)) $ - λ f₁ f₂ hf, quotient.sound $ - by simpa [neg_sub', (≈), setoid.r] using neg_lim_zero hf⟩ +⟨quotient.map has_neg.neg $ λ f₁ f₂ hf, neg_equiv_neg hf⟩ @[simp] theorem mk_neg (f : cau_seq β abv) : -mk f = mk (-f) := rfl instance : has_mul Cauchy := -⟨λ x y, quotient.lift_on₂ x y (λ f g, mk (f * g)) $ - λ f₁ g₁ f₂ g₂ hf hg, quotient.sound $ - by simpa [(≈), setoid.r, mul_add, mul_comm, add_assoc, sub_eq_add_neg] using - add_lim_zero (mul_lim_zero_right g₁ hf) (mul_lim_zero_right f₂ hg)⟩ +⟨quotient.map₂ (*) $ λ f₁ g₁ hf f₂ g₂ hg, mul_equiv_mul hf hg⟩ @[simp] theorem mk_mul (f g : cau_seq β abv) : mk f * mk g = mk (f * g) := rfl instance : has_sub Cauchy := -⟨λ x y, quotient.lift_on₂ x y (λ f g, mk (f - g)) $ - λ f₁ g₁ f₂ g₂ hf hg, quotient.sound $ show ((f₁ - g₁) - (f₂ - g₂)).lim_zero, - by simpa [sub_eq_add_neg, add_assoc, add_comm, add_left_comm] using sub_lim_zero hf hg⟩ +⟨quotient.map₂ has_sub.sub $ λ f₁ g₁ hf f₂ g₂ hg, sub_equiv_sub hf hg⟩ @[simp] theorem mk_sub (f g : cau_seq β abv) : mk f - mk g = mk (f - g) := rfl From 38963b536190e8a52f2a124631d4f985ce15f69e Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Mon, 12 Sep 2022 16:02:12 +0000 Subject: [PATCH 234/828] feat(algebra/algebra/basic): add two simp lemmas (#16096) Two simp lemmas that are useful for dealing with the group of automorphisms of an algebra. I also had to fix two proofs from `field_theory/krull_topology`. --- src/algebra/algebra/basic.lean | 2 +- src/field_theory/krull_topology.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index 198f5cf791ba0..cb7c209b1a450 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -1166,7 +1166,7 @@ by { ext, refl } end of_linear_equiv -@[simps mul one {attrs := []}] instance aut : group (A₁ ≃ₐ[R] A₁) := +@[simps mul one inv {attrs := []}] instance aut : group (A₁ ≃ₐ[R] A₁) := { mul := λ ϕ ψ, ψ.trans ϕ, mul_assoc := λ ϕ ψ χ, rfl, one := refl, diff --git a/src/field_theory/krull_topology.lean b/src/field_theory/krull_topology.lean index 2dbbace10155d..ba2447e065658 100644 --- a/src/field_theory/krull_topology.lean +++ b/src/field_theory/krull_topology.lean @@ -276,7 +276,7 @@ begin haveI := intermediate_field.adjoin.finite_dimensional (h_int x), refine ⟨left_coset σ E.fixing_subgroup, ⟨E.fixing_subgroup_is_open.left_coset σ, E.fixing_subgroup_is_closed.left_coset σ⟩, - ⟨1, E.fixing_subgroup.one_mem', by simp⟩, _⟩, + ⟨1, E.fixing_subgroup.one_mem', mul_one σ⟩, _⟩, simp only [mem_left_coset_iff, set_like.mem_coe, intermediate_field.mem_fixing_subgroup_iff, not_forall], exact ⟨x, intermediate_field.mem_adjoin_simple_self K x, hx⟩, From f058af302cbea6525c9ccb3d991c37f1fd7db446 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 12 Sep 2022 18:27:04 +0000 Subject: [PATCH 235/828] feat(analysis/inner_product_space/pi_L2): change of basis matrix between orthonormal bases is unitary (#16474) Formalized as part of the Sphere Eversion project. --- src/analysis/inner_product_space/pi_L2.lean | 56 ++++++++++++++++++++- src/linear_algebra/unitary_group.lean | 26 +++++++++- 2 files changed, 79 insertions(+), 3 deletions(-) diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 1ac02569e0933..1a460c5986670 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Sébastien Gouëzel, Heather Macbeth -/ import analysis.inner_product_space.projection -import linear_algebra.finite_dimensional import analysis.normed_space.pi_Lp +import linear_algebra.finite_dimensional +import linear_algebra.unitary_group /-! # `L²` inner product space structure on finite products of inner product spaces @@ -526,6 +527,59 @@ by simp [complex.isometry_of_orthonormal, (dec_trivial : (finset.univ : finset ( open finite_dimensional +/-! ### Matrix representation of an orthonormal basis with respect to another -/ + +section to_matrix +variables [decidable_eq ι] + +section +variables (a b : orthonormal_basis ι 𝕜 E) + +/-- The change-of-basis matrix between two orthonormal bases `a`, `b` is a unitary matrix. -/ +lemma orthonormal_basis.to_matrix_orthonormal_basis_mem_unitary : + a.to_basis.to_matrix b ∈ matrix.unitary_group ι 𝕜 := +begin + rw matrix.mem_unitary_group_iff', + ext i j, + convert a.repr.inner_map_map (b i) (b j), + rw orthonormal_iff_ite.mp b.orthonormal i j, + refl, +end + +/-- The determinant of the change-of-basis matrix between two orthonormal bases `a`, `b` has +unit length. -/ +@[simp] lemma orthonormal_basis.det_to_matrix_orthonormal_basis : + ∥a.to_basis.det b∥ = 1 := +begin + have : (norm_sq (a.to_basis.det b) : 𝕜) = 1, + { simpa [is_R_or_C.mul_conj] + using (matrix.det_of_mem_unitary (a.to_matrix_orthonormal_basis_mem_unitary b)).2 }, + norm_cast at this, + rwa [← sqrt_norm_sq_eq_norm, sqrt_eq_one], +end + +end + +section real +variables (a b : orthonormal_basis ι ℝ F) + +/-- The change-of-basis matrix between two orthonormal bases `a`, `b` is an orthogonal matrix. -/ +lemma orthonormal_basis.to_matrix_orthonormal_basis_mem_orthogonal : + a.to_basis.to_matrix b ∈ matrix.orthogonal_group ι ℝ := +a.to_matrix_orthonormal_basis_mem_unitary b + +/-- The determinant of the change-of-basis matrix between two orthonormal bases `a`, `b` is ±1. -/ +lemma orthonormal_basis.det_to_matrix_orthonormal_basis_real : + a.to_basis.det b = 1 ∨ a.to_basis.det b = -1 := +begin + rw ← sq_eq_one_iff, + simpa [unitary, sq] using matrix.det_of_mem_unitary (a.to_matrix_orthonormal_basis_mem_unitary b) +end + +end real + +end to_matrix + /-! ### Existence of orthonormal basis, etc. -/ section finite_dimensional diff --git a/src/linear_algebra/unitary_group.lean b/src/linear_algebra/unitary_group.lean index c699d2b41f89c..e27a84a586643 100644 --- a/src/linear_algebra/unitary_group.lean +++ b/src/linear_algebra/unitary_group.lean @@ -59,7 +59,22 @@ lemma mem_unitary_group_iff {A : matrix n n α} : A ∈ matrix.unitary_group n α ↔ A * star A = 1 := begin refine ⟨and.right, λ hA, ⟨_, hA⟩⟩, - simpa only [matrix.mul_eq_mul, matrix.mul_eq_one_comm] using hA + simpa only [mul_eq_mul, mul_eq_one_comm] using hA +end + +lemma mem_unitary_group_iff' {A : matrix n n α} : + A ∈ matrix.unitary_group n α ↔ star A * A = 1 := +begin + refine ⟨and.left, λ hA, ⟨hA, _⟩⟩, + rwa [mul_eq_mul, mul_eq_one_comm] at hA, +end + +lemma det_of_mem_unitary {A : matrix n n α} (hA : A ∈ matrix.unitary_group n α) : + A.det ∈ unitary α := +begin + split, + { simpa [star, det_transpose] using congr_arg det hA.1 }, + { simpa [star, det_transpose] using congr_arg det hA.2 }, end namespace unitary_group @@ -163,7 +178,14 @@ lemma mem_orthogonal_group_iff {A : matrix n n β} : A ∈ matrix.orthogonal_group n β ↔ A * star A = 1 := begin refine ⟨and.right, λ hA, ⟨_, hA⟩⟩, - simpa only [matrix.mul_eq_mul, matrix.mul_eq_one_comm] using hA + simpa only [mul_eq_mul, mul_eq_one_comm] using hA +end + +lemma mem_orthogonal_group_iff' {A : matrix n n β} : + A ∈ matrix.orthogonal_group n β ↔ star A * A = 1 := +begin + refine ⟨and.left, λ hA, ⟨hA, _⟩⟩, + rwa [mul_eq_mul, mul_eq_one_comm] at hA, end end orthogonal_group From 1b3556cc838f818f5679cad86f5b7d7f89e6dadf Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Mon, 12 Sep 2022 18:27:05 +0000 Subject: [PATCH 236/828] chore(linear_algebra): rename dual_pair (#16482) This only renames `dual_pair` to `module.dual_bases`, avoiding to put a very generic name into the root name space. This `dual_pair` was used only in the archive, I introduced it a very long time ago to streamline the proof of the sensitivity conjecture. --- archive/sensitivity.lean | 26 ++++++++++++++------------ src/linear_algebra/dual.lean | 32 +++++++++++++++++--------------- 2 files changed, 31 insertions(+), 27 deletions(-) diff --git a/archive/sensitivity.lean b/archive/sensitivity.lean index f1960b94e62ee..5b96d7ddbffa0 100644 --- a/archive/sensitivity.lean +++ b/archive/sensitivity.lean @@ -43,7 +43,7 @@ open_locale big_operators notation `√` := real.sqrt -open function bool linear_map fintype finite_dimensional dual_pair +open function bool linear_map fintype finite_dimensional module.dual_bases /-! ### The hypercube @@ -217,9 +217,11 @@ begin rwa show p = π q, by { ext, simp [q, fin.succ_ne_zero, π] } } } end +open module + /-- `e` and `ε` are dual families of vectors. It implies that `e` is indeed a basis and `ε` computes coefficients of decompositions of vectors on that basis. -/ -def dual_pair_e_ε (n : ℕ) : dual_pair (@e n) (@ε n) := +def dual_bases_e_ε (n : ℕ) : dual_bases (@e n) (@ε n) := { eval := duality, total := @epsilon_total _ } @@ -228,11 +230,11 @@ since this cardinal is finite, as a natural number in `finrank_V` -/ lemma dim_V : module.rank ℝ (V n) = 2^n := have module.rank ℝ (V n) = (2^n : ℕ), - by { rw [dim_eq_card_basis (dual_pair_e_ε _).basis, Q.card]; apply_instance }, + by { rw [dim_eq_card_basis (dual_bases_e_ε _).basis, Q.card]; apply_instance }, by assumption_mod_cast instance : finite_dimensional ℝ (V n) := -finite_dimensional.of_fintype_basis (dual_pair_e_ε _).basis +finite_dimensional.of_fintype_basis (dual_bases_e_ε _).basis lemma finrank_V : finrank ℝ (V n) = 2^n := have _ := @dim_V n, @@ -367,12 +369,12 @@ begin apply dim_V }, have dimW : dim W = card H, { have li : linear_independent ℝ (H.restrict e), - { convert (dual_pair_e_ε _).basis.linear_independent.comp _ subtype.val_injective, - rw (dual_pair_e_ε _).coe_basis }, + { convert (dual_bases_e_ε _).basis.linear_independent.comp _ subtype.val_injective, + rw (dual_bases_e_ε _).coe_basis }, have hdW := dim_span li, rw set.range_restrict at hdW, convert hdW, - rw [← (dual_pair_e_ε _).coe_basis, cardinal.mk_image_eq (dual_pair_e_ε _).basis.injective, + rw [← (dual_bases_e_ε _).coe_basis, cardinal.mk_image_eq (dual_bases_e_ε _).basis.injective, cardinal.mk_fintype] }, rw ← finrank_eq_dim ℝ at ⊢ dim_le dim_add dimW, rw [← finrank_eq_dim ℝ, ← finrank_eq_dim ℝ] at dim_add, @@ -387,28 +389,28 @@ theorem huang_degree_theorem (H : set (Q (m + 1))) (hH : Card H ≥ 2^m + 1) : ∃ q, q ∈ H ∧ √(m + 1) ≤ Card (H ∩ q.adjacent) := begin rcases exists_eigenvalue H hH with ⟨y, ⟨⟨y_mem_H, y_mem_g⟩, y_ne⟩⟩, - have coeffs_support : ((dual_pair_e_ε (m+1)).coeffs y).support ⊆ H.to_finset, + have coeffs_support : ((dual_bases_e_ε (m+1)).coeffs y).support ⊆ H.to_finset, { intros p p_in, rw finsupp.mem_support_iff at p_in, rw set.mem_to_finset, - exact (dual_pair_e_ε _).mem_of_mem_span y_mem_H p p_in }, + exact (dual_bases_e_ε _).mem_of_mem_span y_mem_H p p_in }, obtain ⟨q, H_max⟩ : ∃ q : Q (m+1), ∀ q' : Q (m+1), |(ε q' : _) y| ≤ |ε q y|, from finite.exists_max _, have H_q_pos : 0 < |ε q y|, { contrapose! y_ne, exact epsilon_total (λ p, abs_nonpos_iff.mp (le_trans (H_max p) y_ne)) }, - refine ⟨q, (dual_pair_e_ε _).mem_of_mem_span y_mem_H q (abs_pos.mp H_q_pos), _⟩, + refine ⟨q, (dual_bases_e_ε _).mem_of_mem_span y_mem_H q (abs_pos.mp H_q_pos), _⟩, let s := √(m+1), suffices : s * |ε q y| ≤ ↑(_) * |ε q y|, from (mul_le_mul_right H_q_pos).mp ‹_›, - let coeffs := (dual_pair_e_ε (m+1)).coeffs, + let coeffs := (dual_bases_e_ε (m+1)).coeffs, calc s * |ε q y| = |ε q (s • y)| : by rw [map_smul, smul_eq_mul, abs_mul, abs_of_nonneg (real.sqrt_nonneg _)] ... = |ε q (f (m+1) y)| : by rw [← f_image_g y (by simpa using y_mem_g)] - ... = |ε q (f (m+1) (lc _ (coeffs y)))| : by rw (dual_pair_e_ε _).lc_coeffs y + ... = |ε q (f (m+1) (lc _ (coeffs y)))| : by rw (dual_bases_e_ε _).lc_coeffs y ... = |(coeffs y).sum (λ (i : Q (m + 1)) (a : ℝ), a • ((ε q) ∘ (f (m + 1)) ∘ λ (i : Q (m + 1)), e i) i)| : by erw [(f $ m + 1).map_finsupp_total, (ε q).map_finsupp_total, finsupp.total_apply] diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index ade8d80e9d4fc..48f3ddbfe7342 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -18,7 +18,7 @@ The dual space of an R-module M is the R-module of linear maps `M → R`. * `dual R M` defines the dual space of M over R. * Given a basis for an `R`-module `M`, `basis.to_dual` produces a map from `M` to `dual R M`. -* Given families of vectors `e` and `ε`, `dual_pair e ε` states that these families have the +* Given families of vectors `e` and `ε`, `module.dual_bases e ε` states that these families have the characteristic properties of a basis and a dual. * `dual_annihilator W` is the submodule of `dual R M` where every element annihilates `W`. @@ -26,8 +26,8 @@ The dual space of an R-module M is the R-module of linear maps `M → R`. * `to_dual_equiv` : the linear equivalence between the dual module and primal module, given a finite basis. -* `dual_pair.basis` and `dual_pair.eq_dual`: if `e` and `ε` form a dual pair, `e` is a basis and - `ε` is its dual basis. +* `module.dual_bases.basis` and `module.dual_bases.eq_dual`: if `e` and `ε` form a dual pair, `e` + is a basis and `ε` is its dual basis. * `quot_equiv_annihilator`: the quotient by a subspace is isomorphic to its dual annihilator. ## Notation @@ -342,7 +342,7 @@ variables {K V} end module -section dual_pair +section dual_bases open module @@ -351,14 +351,14 @@ variables [comm_semiring R] [add_comm_monoid M] [module R M] [decidable_eq ι] /-- `e` and `ε` have characteristic properties of a basis and its dual -/ @[nolint has_nonempty_instance] -structure dual_pair (e : ι → M) (ε : ι → (dual R M)) := +structure module.dual_bases (e : ι → M) (ε : ι → (dual R M)) := (eval : ∀ i j : ι, ε i (e j) = if i = j then 1 else 0) (total : ∀ {m : M}, (∀ i, ε i m = 0) → m = 0) [finite : ∀ m : M, fintype {i | ε i m ≠ 0}] -end dual_pair +end dual_bases -namespace dual_pair +namespace module.dual_bases open module module.dual linear_map function @@ -367,12 +367,12 @@ variables [comm_ring R] [add_comm_group M] [module R M] variables {e : ι → M} {ε : ι → dual R M} /-- The coefficients of `v` on the basis `e` -/ -def coeffs [decidable_eq ι] (h : dual_pair e ε) (m : M) : ι →₀ R := +def coeffs [decidable_eq ι] (h : dual_bases e ε) (m : M) : ι →₀ R := { to_fun := λ i, ε i m, support := by { haveI := h.finite m, exact {i : ι | ε i m ≠ 0}.to_finset }, mem_support_to_fun := by {intro i, rw set.mem_to_finset, exact iff.rfl } } -@[simp] lemma coeffs_apply [decidable_eq ι] (h : dual_pair e ε) (m : M) (i : ι) : +@[simp] lemma coeffs_apply [decidable_eq ι] (h : dual_bases e ε) (m : M) (i : ι) : h.coeffs m i = ε i m := rfl /-- linear combinations of elements of `e`. @@ -381,10 +381,12 @@ def lc {ι} (e : ι → M) (l : ι →₀ R) : M := l.sum (λ (i : ι) (a : R), lemma lc_def (e : ι → M) (l : ι →₀ R) : lc e l = finsupp.total _ _ _ e l := rfl -variables [decidable_eq ι] (h : dual_pair e ε) +open module + +variables [decidable_eq ι] (h : dual_bases e ε) include h -lemma dual_lc (l : ι →₀ R) (i : ι) : ε i (dual_pair.lc e l) = l i := +lemma dual_lc (l : ι →₀ R) (i : ι) : ε i (dual_bases.lc e l) = l i := begin erw linear_map.map_sum, simp only [h.eval, map_smul, smul_eq_mul], @@ -397,19 +399,19 @@ begin end @[simp] -lemma coeffs_lc (l : ι →₀ R) : h.coeffs (dual_pair.lc e l) = l := +lemma coeffs_lc (l : ι →₀ R) : h.coeffs (dual_bases.lc e l) = l := by { ext i, rw [h.coeffs_apply, h.dual_lc] } /-- For any m : M n, \sum_{p ∈ Q n} (ε p m) • e p = m -/ @[simp] -lemma lc_coeffs (m : M) : dual_pair.lc e (h.coeffs m) = m := +lemma lc_coeffs (m : M) : dual_bases.lc e (h.coeffs m) = m := begin refine eq_of_sub_eq_zero (h.total _), intros i, simp [-sub_eq_add_neg, linear_map.map_sub, h.dual_lc, sub_eq_zero] end -/-- `(h : dual_pair e ε).basis` shows the family of vectors `e` forms a basis. -/ +/-- `(h : dual_bases e ε).basis` shows the family of vectors `e` forms a basis. -/ @[simps] def basis : basis ι R M := basis.of_repr @@ -438,7 +440,7 @@ lemma coe_dual_basis [fintype ι] : ⇑h.basis.dual_basis = ε := funext (λ i, h.basis.ext (λ j, by rw [h.basis.dual_basis_apply_self, h.coe_basis, h.eval, if_congr eq_comm rfl rfl])) -end dual_pair +end module.dual_bases namespace submodule From a280c1e533419a68a0556591a8b8fed8dd35be0d Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Mon, 12 Sep 2022 19:53:46 +0000 Subject: [PATCH 237/828] feat(data/nat/{choose/multinomial,factorial/big_operations}): add multinomial coefficient definition with basic lemma's (#16170) Chose a definition based on finset and a function for multiplicity. This allows the definition to be used in various context, including in a multiset context, where multiset.count yields the function describing multiplicity. Supporting lemma's concerning factorials are introduced in a new file, since importing big_operators in the main factorial file yields a circular import. Co-authored-by: Kyle Miller --- src/data/nat/choose/multinomial.lean | 119 ++++++++++++++++++++++ src/data/nat/factorial/big_operators.lean | 38 +++++++ 2 files changed, 157 insertions(+) create mode 100644 src/data/nat/choose/multinomial.lean create mode 100644 src/data/nat/factorial/big_operators.lean diff --git a/src/data/nat/choose/multinomial.lean b/src/data/nat/choose/multinomial.lean new file mode 100644 index 0000000000000..9ed8012077508 --- /dev/null +++ b/src/data/nat/choose/multinomial.lean @@ -0,0 +1,119 @@ +/- +Copyright (c) 2022 Pim Otte. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kyle Miller, Pim Otte +-/ + +import algebra.big_operators.fin +import algebra.big_operators.order +import data.nat.choose.basic +import data.nat.factorial.big_operators +import data.fin.vec_notation + +import tactic.linarith + +/-! +# Multinomial + +This file defines the multinomial coefficient and several small lemma's for manipulating it. + +## Main declarations + +- `nat.multinomial`: the multinomial coefficient + +-/ + +open_locale big_operators nat + +namespace nat + +variables {α : Type*} (s : finset α) (f : α → ℕ) +variables {a b : α} + +/-- The multinomial coefficient. Gives the number of strings consisting of symbols +from `s`, where `c ∈ s` appears with multiplicity `f c`. + +Defined as `(∑ i in s, f i)! / ∏ i in s, (f i)!`. +-/ +def multinomial : ℕ := (∑ i in s, f i)! / ∏ i in s, (f i)! + +lemma multinomial_pos : 0 < multinomial s f := nat.div_pos + (le_of_dvd (factorial_pos _) (prod_factorial_dvd_factorial_sum s f)) (prod_factorial_pos s f) + +lemma multinomial_spec : (∏ i in s, (f i)!) * multinomial s f = (∑ i in s, f i)! := +nat.mul_div_cancel' (prod_factorial_dvd_factorial_sum s f) + +@[simp] lemma multinomial_nil : multinomial ∅ f = 1 := rfl + +@[simp] lemma multinomial_singleton : multinomial {a} f = 1 := +by simp [multinomial, nat.div_self (factorial_pos (f a))] + +@[simp] lemma multinomial_insert_one [decidable_eq α] (h : a ∉ s) (h₁ : f a = 1) : + multinomial (insert a s) f = (s.sum f).succ * multinomial s f := +begin + simp only [multinomial, one_mul, factorial], + rw [finset.sum_insert h, finset.prod_insert h, h₁, add_comm, ←succ_eq_add_one, factorial_succ], + simp only [factorial_one, one_mul, function.comp_app, factorial], + rw nat.mul_div_assoc _ (prod_factorial_dvd_factorial_sum _ _), +end + +lemma multinomial_insert [decidable_eq α] (h : a ∉ s) : + multinomial (insert a s) f = (f a + s.sum f).choose (f a) * multinomial s f := +begin + rw choose_eq_factorial_div_factorial (le.intro rfl), + simp only [multinomial, nat.add_sub_cancel_left, finset.sum_insert h, finset.prod_insert h, + function.comp_app], + rw [div_mul_div_comm ((f a).factorial_mul_factorial_dvd_factorial_add (s.sum f)) + (prod_factorial_dvd_factorial_sum _ _), mul_comm (f a)! (s.sum f)!, mul_assoc, + mul_comm _ (s.sum f)!, nat.mul_div_mul _ _ (factorial_pos _)], +end + +/-! ### Connection to binomial coefficients -/ + +lemma binomial_eq [decidable_eq α] (h : a ≠ b) : + multinomial {a, b} f = (f a + f b)! / ((f a)! * (f b)!) := +by simp [multinomial, finset.sum_pair h, finset.prod_pair h] + +lemma binomial_eq_choose [decidable_eq α] (h : a ≠ b) : + multinomial {a, b} f = (f a + f b).choose (f a) := +by simp [binomial_eq _ h, choose_eq_factorial_div_factorial (nat.le_add_right _ _)] + +lemma binomial_spec [decidable_eq α] (hab : a ≠ b) : + (f a)! * (f b)! * multinomial {a, b} f = (f a + f b)! := +by simpa [finset.sum_pair hab, finset.prod_pair hab] using multinomial_spec {a, b} f + +@[simp] lemma binomial_one [decidable_eq α] (h : a ≠ b) (h₁ : f a = 1) : + multinomial {a, b} f = (f b).succ := +by simp [multinomial_insert_one {b} f (finset.not_mem_singleton.mpr h) h₁] + +lemma binomial_succ_succ [decidable_eq α] (h : a ≠ b) : + multinomial {a, b} (function.update (function.update f a (f a).succ) b (f b).succ) = + multinomial {a, b} (function.update f a (f a).succ) + + multinomial {a, b} (function.update f b (f b).succ) := +begin + simp only [binomial_eq_choose, function.update_apply, function.update_noteq, + succ_add, add_succ, choose_succ_succ, h, ne.def, not_false_iff, function.update_same], + rw if_neg h.symm, + ring, +end + +lemma succ_mul_binomial [decidable_eq α] (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) := +begin + rw [binomial_eq_choose _ h, binomial_eq_choose _ h, mul_comm (f a).succ, + function.update_same, function.update_noteq (ne_comm.mp h)], + convert succ_mul_choose_eq (f a + f b) (f a), + exact succ_add (f a) (f b), +end + +/-! ### Simple cases -/ + +lemma multinomial_univ_two (a b : ℕ) : multinomial finset.univ ![a, b] = (a + b)! / (a! * b!) := +by simp [multinomial, fin.sum_univ_two, fin.prod_univ_two] + +lemma multinomial_univ_three (a b c : ℕ) : multinomial finset.univ ![a, b, c] = + (a + b + c)! / (a! * b! * c!) := +by simp [multinomial, fin.sum_univ_three, fin.prod_univ_three] + +end nat diff --git a/src/data/nat/factorial/big_operators.lean b/src/data/nat/factorial/big_operators.lean new file mode 100644 index 0000000000000..0c095d9f8f6b4 --- /dev/null +++ b/src/data/nat/factorial/big_operators.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2022 Pim Otte. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kyle Miller, Pim Otte +-/ +import data.nat.factorial.basic +import algebra.big_operators.order + +/-! +# Factorial with big operators + +This file contains some lemmas on factorials in combination with big operators. + +While in terms of semantics they could be in the `basic.lean` file, importing +`algebra.big_operators.basic` leads to a cyclic import. + +-/ + +open_locale nat big_operators + +namespace nat + +variables {α : Type*} (s : finset α) (f : α → ℕ) + +lemma prod_factorial_pos : 0 < ∏ i in s, (f i)! := +finset.prod_pos (λ i _, factorial_pos (f i)) + +lemma prod_factorial_dvd_factorial_sum : (∏ i in s, (f i)!) ∣ (∑ i in s, f i)! := +begin + classical, + induction s using finset.induction with a' s' has ih, + { simp only [finset.sum_empty, finset.prod_empty, factorial], }, + { simp only [finset.prod_insert has, finset.sum_insert has], + refine dvd_trans (mul_dvd_mul_left ((f a')!) ih) _, + apply nat.factorial_mul_factorial_dvd_factorial_add, }, +end + +end nat From 36e0130108febccf38bc1f361f450b52b65aece0 Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Mon, 12 Sep 2022 19:53:47 +0000 Subject: [PATCH 238/828] feat(ring_theory/dedekind_domain/ideal): construct map between the sets of prime factors of ideal `I` and `J` induced by an isomorphism between `R/I` and `S/J` (#16455) These lemmas generalize results that were originally in #15000. Co-authored-by: Anne Baanen --- src/ring_theory/dedekind_domain/ideal.lean | 135 +++++++++++++++++---- 1 file changed, 113 insertions(+), 22 deletions(-) diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index bdc93ddfcbcf2..78f6e62740ad2 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -9,6 +9,7 @@ import order.hom.basic import ring_theory.dedekind_domain.basic import ring_theory.fractional_ideal import ring_theory.principal_ideal_domain +import ring_theory.chain_of_divisors /-! # Dedekind domains and ideals @@ -925,13 +926,12 @@ begin map_comap_of_surjective J^.quotient.mk quotient.mk_surjective, map_map], end -variables [is_domain R] [is_dedekind_domain R] +variables [is_domain R] [is_dedekind_domain R] (f : R ⧸ I ≃+* A ⧸ J) /-- The bijection between ideals of `R` dividing `I` and the ideals of `A` dividing `J` induced by an isomorphism `f : R/I ≅ A/J`. -/ @[simps] -def ideal_factors_equiv_of_quot_equiv (f : R ⧸ I ≃+* A ⧸ J) : - {p : ideal R | p ∣ I} ≃o {p : ideal A | p ∣ J} := +def ideal_factors_equiv_of_quot_equiv : {p : ideal R | p ∣ I} ≃o {p : ideal A | p ∣ J} := order_iso.of_hom_inv (ideal_factors_fun_of_quot_hom (show function.surjective (f : R ⧸I →+* A ⧸ J), from f.surjective)) @@ -946,6 +946,71 @@ order_iso.of_hom_inv ← ring_equiv.to_ring_hom_eq_coe, ← ring_equiv.to_ring_hom_trans, ring_equiv.self_trans_symm, ring_equiv.to_ring_hom_refl]) +lemma ideal_factors_equiv_of_quot_equiv_symm : + (ideal_factors_equiv_of_quot_equiv f).symm = ideal_factors_equiv_of_quot_equiv f.symm := rfl + +lemma ideal_factors_equiv_of_quot_equiv_is_dvd_iso {L M : ideal R} (hL : L ∣ I) (hM : M ∣ I) : + (ideal_factors_equiv_of_quot_equiv f ⟨L, hL⟩ : ideal A) ∣ + ideal_factors_equiv_of_quot_equiv f ⟨M, hM⟩ ↔ L ∣ M := +begin + suffices : ideal_factors_equiv_of_quot_equiv f ⟨M, hM⟩ ≤ + ideal_factors_equiv_of_quot_equiv f ⟨L, hL⟩ ↔ (⟨M, hM⟩ : {p : ideal R | p ∣ I}) ≤ ⟨L, hL⟩, + { rw [dvd_iff_le, dvd_iff_le, subtype.coe_le_coe, this, subtype.mk_le_mk] }, + exact (ideal_factors_equiv_of_quot_equiv f).le_iff_le, +end + +open unique_factorization_monoid + +variables [decidable_eq (ideal R)] [decidable_eq (ideal A)] + +lemma ideal_factors_equiv_of_quot_equiv_mem_normalized_factors_of_mem_normalized_factors + (hJ : J ≠ ⊥) {L : ideal R} (hL : L ∈ normalized_factors I) : + ↑(ideal_factors_equiv_of_quot_equiv f + ⟨L, dvd_of_mem_normalized_factors hL⟩) ∈ normalized_factors J := +begin + by_cases hI : I = ⊥, + { exfalso, + rw [hI, bot_eq_zero, normalized_factors_zero, ← multiset.empty_eq_zero] at hL, + exact hL, }, + { apply mem_normalized_factors_factor_dvd_iso_of_mem_normalized_factors hI hJ hL _, + rintros ⟨l, hl⟩ ⟨l', hl'⟩, + rw [subtype.coe_mk, subtype.coe_mk], + apply ideal_factors_equiv_of_quot_equiv_is_dvd_iso f } +end + +/-- The bijection between the sets of normalized factors of I and J induced by a ring + isomorphism `f : R/I ≅ A/J`. -/ +@[simps apply] +def normalized_factors_equiv_of_quot_equiv (hI : I ≠ ⊥) (hJ : J ≠ ⊥) : + {L : ideal R | L ∈ normalized_factors I } ≃ {M : ideal A | M ∈ normalized_factors J } := +{ to_fun := λ j, ⟨ideal_factors_equiv_of_quot_equiv f ⟨↑j, dvd_of_mem_normalized_factors j.prop⟩, + ideal_factors_equiv_of_quot_equiv_mem_normalized_factors_of_mem_normalized_factors f hJ j.prop⟩, + inv_fun := λ j, ⟨(ideal_factors_equiv_of_quot_equiv f).symm + ⟨↑j, dvd_of_mem_normalized_factors j.prop⟩, by { rw ideal_factors_equiv_of_quot_equiv_symm, + exact ideal_factors_equiv_of_quot_equiv_mem_normalized_factors_of_mem_normalized_factors + f.symm hI j.prop} ⟩, + left_inv := λ ⟨j, hj⟩, by simp, + right_inv := λ ⟨j, hj⟩, by simp } + +@[simp] +lemma normalized_factors_equiv_of_quot_equiv_symm (hI : I ≠ ⊥) (hJ : J ≠ ⊥) : + (normalized_factors_equiv_of_quot_equiv f hI hJ).symm = + normalized_factors_equiv_of_quot_equiv f.symm hJ hI := +rfl + +variable [decidable_rel ((∣) : ideal R → ideal R → Prop)] +variable [decidable_rel ((∣) : ideal A → ideal A → Prop)] + +/-- The map `normalized_factors_equiv_of_quot_equiv` preserves multiplicities. -/ +lemma normalized_factors_equiv_of_quot_equiv_multiplicity_eq_multiplicity (hI : I ≠ ⊥) (hJ : J ≠ ⊥) + (L : ideal R) (hL : L ∈ normalized_factors I) : + multiplicity ↑(normalized_factors_equiv_of_quot_equiv f hI hJ ⟨L, hL⟩) J = multiplicity L I := +begin + rw [normalized_factors_equiv_of_quot_equiv, equiv.coe_fn_mk, subtype.coe_mk], + exact multiplicity_factor_dvd_iso_eq_multiplicity_of_mem_normalized_factor hI hJ hL + (λ ⟨l, hl⟩ ⟨l', hl'⟩, ideal_factors_equiv_of_quot_equiv_is_dvd_iso f hl hl'), +end + end section chinese_remainder @@ -1148,13 +1213,33 @@ begin exact (prime_of_normalized_factor a ha).ne_zero (span_singleton_eq_bot.mp h) }, end +lemma multiplicity_eq_multiplicity_span [decidable_rel ((∣) : R → R → Prop)] + [decidable_rel ((∣) : ideal R → ideal R → Prop)] {a b : R} : + multiplicity (ideal.span {a}) (ideal.span ({b} : set R)) = multiplicity a b := +begin + by_cases h : finite a b, + { rw ← part_enat.coe_get (finite_iff_dom.mp h), + refine (multiplicity.unique + (show (ideal.span {a})^(((multiplicity a b).get h)) ∣ (ideal.span {b}), from _) _).symm ; + rw [ideal.span_singleton_pow, span_singleton_dvd_span_singleton_iff_dvd], + exact pow_multiplicity_dvd h , + { exact multiplicity.is_greatest ((part_enat.lt_coe_iff _ _).mpr (exists.intro + (finite_iff_dom.mp h) (nat.lt_succ_self _))) } }, + { suffices : ¬ (finite (ideal.span ({a} : set R)) (ideal.span ({b} : set R))), + { rw [finite_iff_dom, part_enat.not_dom_iff_eq_top] at h this, + rw [h, this] }, + refine not_finite_iff_forall.mpr (λ n, by {rw [ideal.span_singleton_pow, + span_singleton_dvd_span_singleton_iff_dvd], exact not_finite_iff_forall.mp h n }) } +end + +variables [decidable_eq R] [decidable_eq (ideal R)] [normalization_monoid R] + /-- The bijection between the (normalized) prime factors of `r` and the (normalized) prime factors of `span {r}` -/ @[simps] -noncomputable def normalized_factors_equiv_span_normalized_factors [normalization_monoid R] - [decidable_eq R] [decidable_eq (ideal R)] {r : R} (hr : r ≠ 0) : +noncomputable def normalized_factors_equiv_span_normalized_factors {r : R} (hr : r ≠ 0) : {d : R | d ∈ normalized_factors r} ≃ - {I : ideal R | I ∈ normalized_factors (ideal.span ({r} : set R))} := + {I : ideal R | I ∈ normalized_factors (ideal.span ({r} : set R))} := equiv.of_bijective (λ d, ⟨ideal.span {↑d}, singleton_span_mem_normalized_factors_of_mem_normalized_factors d.prop⟩) begin @@ -1176,23 +1261,29 @@ begin dvd_iff_le.mp (dvd_of_mem_normalized_factors hi))) (mem_span_singleton.mpr (dvd_refl r))) } } end -lemma multiplicity_eq_multiplicity_span [decidable_rel ((∣) : R → R → Prop)] - [decidable_rel ((∣) : ideal R → ideal R → Prop)] {a b : R} : - multiplicity (ideal.span {a}) (ideal.span ({b} : set R)) = multiplicity a b := +variables [decidable_rel ((∣) : R → R → Prop)] [decidable_rel ((∣) : ideal R → ideal R → Prop)] + +/-- The bijection `normalized_factors_equiv_span_normalized_factors` between the set of prime + factors of `r` and the set of prime factors of the ideal `⟨r⟩` preserves multiplicities. -/ +lemma multiplicity_normalized_factors_equiv_span_normalized_factors_eq_multiplicity {r d: R} + (hr : r ≠ 0) (hd : d ∈ normalized_factors r) : + multiplicity d r = + multiplicity (normalized_factors_equiv_span_normalized_factors hr ⟨d, hd⟩ : ideal R) + (ideal.span {r}) := +by simp only [normalized_factors_equiv_span_normalized_factors, multiplicity_eq_multiplicity_span, + subtype.coe_mk, equiv.of_bijective_apply] + +/-- The bijection `normalized_factors_equiv_span_normalized_factors.symm` between the set of prime + factors of the ideal `⟨r⟩` and the set of prime factors of `r` preserves multiplicities. -/ +lemma multiplicity_normalized_factors_equiv_span_normalized_factors_symm_eq_multiplicity + {r : R} (hr : r ≠ 0) (I : {I : ideal R | I ∈ normalized_factors (ideal.span ({r} : set R))}) : + multiplicity ((normalized_factors_equiv_span_normalized_factors hr).symm I : R) r = + multiplicity (I : ideal R) (ideal.span {r}) := begin - by_cases h : finite a b, - { rw ← part_enat.coe_get (finite_iff_dom.mp h), - refine (multiplicity.unique - (show (ideal.span {a})^(((multiplicity a b).get h)) ∣ (ideal.span {b}), from _) _).symm ; - rw [ideal.span_singleton_pow, span_singleton_dvd_span_singleton_iff_dvd], - exact pow_multiplicity_dvd h , - { exact multiplicity.is_greatest ((part_enat.lt_coe_iff _ _).mpr (exists.intro - (finite_iff_dom.mp h) (nat.lt_succ_self _))) } }, - { suffices : ¬ (finite (ideal.span ({a} : set R)) (ideal.span ({b} : set R))), - { rw [finite_iff_dom, part_enat.not_dom_iff_eq_top] at h this, - rw [h, this] }, - refine not_finite_iff_forall.mpr (λ n, by {rw [ideal.span_singleton_pow, - span_singleton_dvd_span_singleton_iff_dvd], exact not_finite_iff_forall.mp h n }) } + obtain ⟨x, hx⟩ := (normalized_factors_equiv_span_normalized_factors hr).surjective I, + obtain ⟨a, ha⟩ := x, + rw [hx.symm, equiv.symm_apply_apply, subtype.coe_mk, + multiplicity_normalized_factors_equiv_span_normalized_factors_eq_multiplicity hr ha, hx], end end PID From 06ae0493a9b4e22f58a7b4cf5c8846cd0e683bd3 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 13 Sep 2022 03:04:28 +0000 Subject: [PATCH 239/828] feat(data/complex/module): define `real_part` and `imaginary_part` (#16438) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In a generic `star_module` one always has a decomposition of any element into a `self_adjoint_part` and a `skew_self_adjoint` part, but in a star module over `ℂ` we can instead decompose any element into a `real_part` and an `imaginary_part`, both of which are self-adjoint. Here we define these as `ℝ`-linear maps from the star module into the type of its self-adjoint elements and describe the basic relationships between these maps. The decomposition into real and imaginary parts is often useful for reducing arguments about elements of a star module over `ℂ` to the case when the element is self-adjoint. --- src/data/complex/module.lean | 77 ++++++++++++++++++++++++++++++++++++ 1 file changed, 77 insertions(+) diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index c9e8450b855eb..b9246960022e5 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -7,6 +7,7 @@ import algebra.order.smul import data.complex.basic import data.fin.vec_notation import field_theory.tower +import algebra.char_p.invertible /-! # Complex number as a vector space over `ℝ` @@ -31,6 +32,13 @@ part, the embedding of `ℝ` in `ℂ`, and the complex conjugate): It also provides a universal property of the complex numbers `complex.lift`, which constructs a `ℂ →ₐ[ℝ] A` into any `ℝ`-algebra `A` given a square root of `-1`. +In addition, this file provides a decomposition into `real_part` and `imaginary_part` for any +element of a `star_module` over `ℂ`. + +## Notation + +* `ℜ` and `ℑ` for the `real_part` and `imaginary_part`, respectively, in the locale + `complex_star_module`. -/ namespace complex @@ -300,3 +308,72 @@ alg_hom_ext $ (lift_aux_apply_I _ _).trans conj_I.symm end lift end complex + +section real_imaginary_part + +open complex + +variables {A : Type*} [add_comm_group A] [module ℂ A] [star_add_monoid A] [star_module ℂ A] + +/-- Create a `self_adjoint` element from a `skew_adjoint` element by multiplying by the scalar +`-complex.I`. -/ +@[simps] def skew_adjoint.neg_I_smul : skew_adjoint A →ₗ[ℝ] self_adjoint A := +{ to_fun := λ a, ⟨-I • a, by simp only [self_adjoint.mem_iff, neg_smul, star_neg, star_smul, + star_def, conj_I, skew_adjoint.star_coe_eq, neg_smul_neg]⟩, + map_add' := λ a b, by { ext, simp only [add_subgroup.coe_add, smul_add, add_mem_class.mk_add_mk]}, + map_smul' := λ a b, by { ext, simp only [neg_smul, skew_adjoint.coe_smul, add_subgroup.coe_mk, + ring_hom.id_apply, self_adjoint.coe_smul, smul_neg, neg_inj], rw smul_comm, } } + +lemma skew_adjoint.I_smul_neg_I (a : skew_adjoint A) : + I • (skew_adjoint.neg_I_smul a : A) = a := +by simp only [smul_smul, skew_adjoint.neg_I_smul_apply_coe, neg_smul, smul_neg, I_mul_I, one_smul, + neg_neg] + +/-- The real part `ℜ a` of an element `a` of a star module over `ℂ`, as a linear map. This is just +`self_adjoint_part ℝ`, but we provide it as a separate definition in order to link it with lemmas +concerning the `imaginary_part`, which doesn't exist in star modules over other rings. -/ +noncomputable def real_part : A →ₗ[ℝ] self_adjoint A := self_adjoint_part ℝ + +/-- The imaginary part `ℑ a` of an element `a` of a star module over `ℂ`, as a linear map into the +self adjoint elements. In a general star module, we have a decomposition into the `self_adjoint` +and `skew_adjoint` parts, but in a star module over `ℂ` we have +`real_part_add_I_smul_imaginary_part`, which allows us to decompose into a linear combination of +`self_adjoint`s. -/ +noncomputable +def imaginary_part : A →ₗ[ℝ] self_adjoint A := skew_adjoint.neg_I_smul.comp (skew_adjoint_part ℝ) + +localized "notation `ℜ` := real_part" in complex_star_module +localized "notation `ℑ` := imaginary_part" in complex_star_module + +@[simp] lemma real_part_apply_coe (a : A) : + (ℜ a : A) = (2 : ℝ)⁻¹ • (a + star a) := +by { unfold real_part, simp only [self_adjoint_part_apply_coe, inv_of_eq_inv]} + +@[simp] lemma imaginary_part_apply_coe (a : A) : + (ℑ a : A) = -I • (2 : ℝ)⁻¹ • (a - star a) := +begin + unfold imaginary_part, + simp only [linear_map.coe_comp, skew_adjoint.neg_I_smul_apply_coe, skew_adjoint_part_apply_coe, + inv_of_eq_inv], +end + +/-- The standard decomposition of `ℜ a + complex.I • ℑ a = a` of an element of a star module over +`ℂ` into a linear combination of self adjoint elements. -/ +lemma real_part_add_I_smul_imaginary_part (a : A) : (ℜ a + I • ℑ a : A) = a := +by simpa only [smul_smul, real_part_apply_coe, imaginary_part_apply_coe, neg_smul, I_mul_I, + one_smul, neg_sub, add_add_sub_cancel, smul_sub, smul_add, neg_sub_neg, inv_of_eq_inv] + using inv_of_two_smul_add_inv_of_two_smul ℝ a + +@[simp] lemma real_part_I_smul (a : A) : ℜ (I • a) = - ℑ a := +by { ext, simp [smul_comm I, smul_sub, sub_eq_add_neg, add_comm] } + +@[simp] lemma imaginary_part_I_smul (a : A) : ℑ (I • a) = ℜ a := +by { ext, simp [smul_comm I, smul_smul I] } + +lemma real_part_smul (z : ℂ) (a : A) : ℜ (z • a) = z.re • ℜ a - z.im • ℑ a := +by { nth_rewrite 0 ←re_add_im z, simp [-re_add_im, add_smul, ←smul_smul, sub_eq_add_neg] } + +lemma imaginary_part_smul (z : ℂ) (a : A) : ℑ (z • a) = z.re • ℑ a + z.im • ℜ a := +by { nth_rewrite 0 ←re_add_im z, simp [-re_add_im, add_smul, ←smul_smul] } + +end real_imaginary_part From 5bd4f627d7e2a4b52d24bd9a17162d27cf80d08c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 13 Sep 2022 04:22:54 +0000 Subject: [PATCH 240/828] chore(ring_theory/local_properties): remove coercions from `ring_hom` to `monoid_hom` (#16453) Now that `submonoid.map` takes `monoid_hom_class`, these aren't necessary. Also golfs a pair of proofs. --- src/algebra/algebra/basic.lean | 4 +-- src/number_theory/bernoulli_polynomials.lean | 14 +++----- src/ring_theory/local_properties.lean | 35 ++++++++----------- src/ring_theory/localization/basic.lean | 4 +-- .../localization/inv_submonoid.lean | 8 ++--- .../localization_localization.lean | 2 +- 6 files changed, 28 insertions(+), 39 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index cb7c209b1a450..61c3425366a65 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -401,8 +401,8 @@ lemma algebra_map_of_subring_apply {R : Type*} [comm_ring R] (S : subring R) (x /-- Explicit characterization of the submonoid map in the case of an algebra. `S` is made explicit to help with type inference -/ def algebra_map_submonoid (S : Type*) [semiring S] [algebra R S] - (M : submonoid R) : (submonoid S) := -submonoid.map (algebra_map R S : R →* S) M + (M : submonoid R) : submonoid S := +M.map (algebra_map R S) lemma mem_algebra_map_submonoid_of_mem {S : Type*} [semiring S] [algebra R S] {M : submonoid R} (x : M) : (algebra_map R S x) ∈ algebra_map_submonoid S M := diff --git a/src/number_theory/bernoulli_polynomials.lean b/src/number_theory/bernoulli_polynomials.lean index 163162901b42d..73b8263a76085 100644 --- a/src/number_theory/bernoulli_polynomials.lean +++ b/src/number_theory/bernoulli_polynomials.lean @@ -213,14 +213,10 @@ begin simp only [ring_hom.map_sub, tsub_self, constant_coeff_one, constant_coeff_exp, coeff_zero_eq_constant_coeff, mul_zero, sub_self, add_zero], -- Let's multiply both sides by (n+1)! (OK because it's a unit) - set u : units ℚ := ⟨(n+1)!, (n+1)!⁻¹, - mul_inv_cancel (by exact_mod_cast factorial_ne_zero (n+1)), - inv_mul_cancel (by exact_mod_cast factorial_ne_zero (n+1))⟩ with hu, - rw ←units.mul_right_inj (units.map (algebra_map ℚ A).to_monoid_hom u), - -- now tidy up unit mess and generally do trivial rearrangements - -- to make RHS (n+1)*t^n - rw [units.coe_map, mul_left_comm, ring_hom.to_monoid_hom_eq_coe, - ring_hom.coe_monoid_hom, ←ring_hom.map_mul, hu, units.coe_mk], + have hnp1 : is_unit ((n+1)! : ℚ) := is_unit.mk0 _ (by exact_mod_cast factorial_ne_zero (n+1)), + rw ←(hnp1.map (algebra_map ℚ A)).mul_right_inj, + -- do trivial rearrangements to make RHS (n+1)*t^n + rw [mul_left_comm, ←ring_hom.map_mul], change _ = t^n * algebra_map ℚ A (((n+1)*n! : ℕ)*(1/n!)), rw [cast_mul, mul_assoc, mul_one_div_cancel (show (n! : ℚ) ≠ 0, from cast_ne_zero.2 (factorial_ne_zero n)), mul_one, mul_comm (t^n), @@ -235,7 +231,7 @@ begin -- deal with coefficients of e^X-1 simp only [nat.cast_choose ℚ (mem_range_le hi), coeff_mk, if_neg (mem_range_sub_ne_zero hi), one_div, alg_hom.map_smul, power_series.coeff_one, - units.coe_mk, coeff_exp, sub_zero, linear_map.map_sub, algebra.smul_mul_assoc, algebra.smul_def, + coeff_exp, sub_zero, linear_map.map_sub, algebra.smul_mul_assoc, algebra.smul_def, mul_right_comm _ ((aeval t) _), ←mul_assoc, ← ring_hom.map_mul, succ_eq_add_one, ← polynomial.C_eq_algebra_map, polynomial.aeval_mul, polynomial.aeval_C], -- finally cancel the Bernoulli polynomial and the algebra_map diff --git a/src/ring_theory/local_properties.lean b/src/ring_theory/local_properties.lean index f4909bd2138b2..cb3e04017c5db 100644 --- a/src/ring_theory/local_properties.lean +++ b/src/ring_theory/local_properties.lean @@ -403,16 +403,16 @@ span of `finset_integer_multiple _ s` over `R`. -/ lemma is_localization.smul_mem_finset_integer_multiple_span [algebra R S] [algebra R S'] [is_scalar_tower R S S'] - [is_localization (M.map (algebra_map R S : R →* S)) S'] (x : S) + [is_localization (M.map (algebra_map R S)) S'] (x : S) (s : finset S') (hx : algebra_map S S' x ∈ submodule.span R (s : set S')) : ∃ m : M, m • x ∈ submodule.span R - (is_localization.finset_integer_multiple (M.map (algebra_map R S : R →* S)) s : set S) := + (is_localization.finset_integer_multiple (M.map (algebra_map R S)) s : set S) := begin let g : S →ₐ[R] S' := alg_hom.mk' (algebra_map S S') (λ c x, by simp [algebra.algebra_map_eq_smul_one]), -- We first obtain the `y' ∈ M` such that `s' = y' • s` is falls in the image of `S` in `S'`. - let y := is_localization.common_denom_of_finset (M.map (algebra_map R S : R →* S)) s, + let y := is_localization.common_denom_of_finset (M.map (algebra_map R S)) s, have hx₁ : (y : S) • ↑s = g '' _ := (is_localization.finset_integer_multiple_image _ s).symm, obtain ⟨y', hy', e : algebra_map R S y' = y⟩ := y.prop, have : algebra_map R S y' • (s : set S') = y' • s := @@ -428,10 +428,10 @@ begin -- Thus `a • (y' • x) = a • x' ∈ span s'` in `S` for some `a ∈ M`. obtain ⟨x', hx', hx'' : algebra_map _ _ _ = _⟩ := hx, obtain ⟨⟨_, a, ha₁, rfl⟩, ha₂⟩ := (is_localization.eq_iff_exists - (M.map (algebra_map R S : R →* S)) S').mp hx'', + (M.map (algebra_map R S)) S').mp hx'', use (⟨a, ha₁⟩ : M) * (⟨y', hy'⟩ : M), convert (submodule.span R (is_localization.finset_integer_multiple - (submonoid.map (algebra_map R S : R →* S) M) s : set S)).smul_mem a hx' using 1, + (submonoid.map (algebra_map R S) M) s : set S)).smul_mem a hx' using 1, convert ha₂.symm, { rw [mul_comm (y' • x), subtype.coe_mk, submonoid.smul_def, submonoid.coe_mul, ← smul_smul], exact algebra.smul_def _ _ }, @@ -488,7 +488,7 @@ begin resetI, letI := f.to_algebra, letI := λ (r : s), (localization.away_map f r).to_algebra, - haveI : ∀ r : s, is_localization ((submonoid.powers (r : R)).map (algebra_map R S : R →* S)) + haveI : ∀ r : s, is_localization ((submonoid.powers (r : R)).map (algebra_map R S)) (localization.away (f r)), { intro r, rw submonoid.map_powers, exact localization.is_localization }, haveI : ∀ r : s, is_scalar_tower R (localization.away (r : R)) (localization.away (f r)) := @@ -519,12 +519,9 @@ begin obtain ⟨⟨_, n₂, rfl⟩, hn₂⟩ := is_localization.smul_mem_finset_integer_multiple_span (submonoid.powers (r : R)) (localization.away (f r)) _ (s₁ r) hn₁, rw [submonoid.smul_def, ← algebra.smul_def, smul_smul, subtype.coe_mk, ← pow_add] at hn₂, + simp_rw submonoid.map_powers at hn₂, use n₂ + n₁, - refine le_supr (λ (x : s), submodule.span R (sf x : set S)) r _, - change _ ∈ submodule.span R - ((is_localization.finset_integer_multiple _ (s₁ r) : finset S) : set S), - convert hn₂, - rw submonoid.map_powers, refl, + exact le_supr (λ (x : s), submodule.span R (sf x : set S)) r hn₂, end end finite @@ -609,13 +606,13 @@ adjoin of `finset_integer_multiple _ s` over `R`. -/ lemma is_localization.lift_mem_adjoin_finset_integer_multiple [algebra R S] [algebra R S'] [is_scalar_tower R S S'] - [is_localization (M.map (algebra_map R S : R →* S)) S'] (x : S) + [is_localization (M.map (algebra_map R S)) S'] (x : S) (s : finset S') (hx : algebra_map S S' x ∈ algebra.adjoin R (s : set S')) : ∃ m : M, m • x ∈ algebra.adjoin R - (is_localization.finset_integer_multiple (M.map (algebra_map R S : R →* S)) s : set S) := + (is_localization.finset_integer_multiple (M.map (algebra_map R S)) s : set S) := begin obtain ⟨⟨_, a, ha, rfl⟩, e⟩ := is_localization.exists_smul_mem_of_mem_adjoin - (M.map (algebra_map R S : R →* S)) x s (algebra.adjoin R _) algebra.subset_adjoin _ hx, + (M.map (algebra_map R S)) x s (algebra.adjoin R _) algebra.subset_adjoin _ hx, { exact ⟨⟨a, ha⟩, by simpa [submonoid.smul_def] using e⟩ }, { rintros _ ⟨a, ha, rfl⟩, exact subalgebra.algebra_map_mem _ a } end @@ -628,7 +625,7 @@ begin resetI, letI := f.to_algebra, letI := λ (r : s), (localization.away_map f r).to_algebra, - haveI : ∀ r : s, is_localization ((submonoid.powers (r : R)).map (algebra_map R S : R →* S)) + haveI : ∀ r : s, is_localization ((submonoid.powers (r : R)).map (algebra_map R S)) (localization.away (f r)), { intro r, rw submonoid.map_powers, exact localization.is_localization }, haveI : ∀ r : s, is_scalar_tower R (localization.away (r : R)) (localization.away (f r)) := @@ -653,13 +650,9 @@ begin obtain ⟨⟨_, n₂, rfl⟩, hn₂⟩ := is_localization.lift_mem_adjoin_finset_integer_multiple (submonoid.powers (r : R)) _ (s₁ r) hn₁, rw [submonoid.smul_def, ← algebra.smul_def, smul_smul, subtype.coe_mk, ← pow_add] at hn₂, + simp_rw submonoid.map_powers at hn₂, use n₂ + n₁, - refine le_supr (λ (x : s), algebra.adjoin R (sf x : set S)) r _, - change _ ∈ algebra.adjoin R - ((is_localization.finset_integer_multiple _ (s₁ r) : finset S) : set S), - convert hn₂, - rw submonoid.map_powers, - refl, + exact le_supr (λ (x : s), algebra.adjoin R (sf x : set S)) r hn₂ end end finite_type diff --git a/src/ring_theory/localization/basic.lean b/src/ring_theory/localization/basic.lean index 67da6bb0ee506..1283dd6e3a0f7 100644 --- a/src/ring_theory/localization/basic.lean +++ b/src/ring_theory/localization/basic.lean @@ -710,7 +710,7 @@ variables (M S) include M lemma non_zero_divisors_le_comap [is_localization M S] : - non_zero_divisors R ≤ (non_zero_divisors S).comap (algebra_map R S) := + non_zero_divisors R ≤ (non_zero_divisors S).comap (algebra_map R S) := begin rintros a ha b (e : b * algebra_map R S a = 0), obtain ⟨x, s, rfl⟩ := mk'_surjective M b, @@ -723,7 +723,7 @@ begin end lemma map_non_zero_divisors_le [is_localization M S] : - (non_zero_divisors R).map (algebra_map R S).to_monoid_hom ≤ non_zero_divisors S := + (non_zero_divisors R).map (algebra_map R S) ≤ non_zero_divisors S := submonoid.map_le_iff_le_comap.mpr (non_zero_divisors_le_comap M S) end is_localization diff --git a/src/ring_theory/localization/inv_submonoid.lean b/src/ring_theory/localization/inv_submonoid.lean index 4c1ba1919ad43..bbab92af6e54b 100644 --- a/src/ring_theory/localization/inv_submonoid.lean +++ b/src/ring_theory/localization/inv_submonoid.lean @@ -38,17 +38,17 @@ section inv_submonoid variables (M S) /-- The submonoid of `S = M⁻¹R` consisting of `{ 1 / x | x ∈ M }`. -/ -def inv_submonoid : submonoid S := (M.map (algebra_map R S : R →* S)).left_inv +def inv_submonoid : submonoid S := (M.map (algebra_map R S)).left_inv variable [is_localization M S] -lemma submonoid_map_le_is_unit : M.map (algebra_map R S : R →* S) ≤ is_unit.submonoid S := +lemma submonoid_map_le_is_unit : M.map (algebra_map R S) ≤ is_unit.submonoid S := by { rintros _ ⟨a, ha, rfl⟩, exact is_localization.map_units S ⟨_, ha⟩ } /-- There is an equivalence of monoids between the image of `M` and `inv_submonoid`. -/ noncomputable -abbreviation equiv_inv_submonoid : M.map (algebra_map R S : R →* S) ≃* inv_submonoid M S := -((M.map (algebra_map R S : R →* S)).left_inv_equiv (submonoid_map_le_is_unit M S)).symm +abbreviation equiv_inv_submonoid : M.map (algebra_map R S) ≃* inv_submonoid M S := +((M.map (algebra_map R S)).left_inv_equiv (submonoid_map_le_is_unit M S)).symm /-- There is a canonical map from `M` to `inv_submonoid` sending `x` to `1 / x`. -/ noncomputable diff --git a/src/ring_theory/localization/localization_localization.lean b/src/ring_theory/localization/localization_localization.lean index 0bb2ef93f76c6..a21991101fca7 100644 --- a/src/ring_theory/localization/localization_localization.lean +++ b/src/ring_theory/localization/localization_localization.lean @@ -203,7 +203,7 @@ localization_algebra_of_submonoid_le _ _ x.prime_compl (non_zero_divisors R) lemma is_localization_of_submonoid_le (M N : submonoid R) (h : M ≤ N) [is_localization M S] [is_localization N T] [algebra S T] [is_scalar_tower R S T] : - is_localization (N.map (algebra_map R S).to_monoid_hom) T := + is_localization (N.map (algebra_map R S)) T := { map_units := begin rintro ⟨_, ⟨y, hy, rfl⟩⟩, convert is_localization.map_units T ⟨y, hy⟩, From 4ff102129218b2a4896abbaa4f755dd5e2af30a7 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 13 Sep 2022 06:29:29 +0000 Subject: [PATCH 241/828] feat(analysis/normed_space/units): maximal ideals in complete normed rings are closed (#16303) --- src/analysis/normed_space/units.lean | 35 ++++++++++++++++++++++++++++ src/topology/algebra/ring.lean | 6 ++--- 2 files changed, 38 insertions(+), 3 deletions(-) diff --git a/src/analysis/normed_space/units.lean b/src/analysis/normed_space/units.lean index 81dbba4a8e261..1dd2104580a84 100644 --- a/src/analysis/normed_space/units.lean +++ b/src/analysis/normed_space/units.lean @@ -81,6 +81,19 @@ is_open.mem_nhds units.is_open x.is_unit end units +namespace nonunits + +/-- The `nonunits` in a complete normed ring are contained in the complement of the ball of radius +`1` centered at `1 : R`. -/ +lemma subset_compl_ball : nonunits R ⊆ (metric.ball (1 : R) 1)ᶜ := +set.subset_compl_comm.mp $ λ x hx, by simpa [sub_sub_self, units.coe_one_sub] using + (units.one_sub (1 - x) (by rwa [metric.mem_ball, dist_eq_norm, norm_sub_rev] at hx)).is_unit + +/- The `nonunits` in a complete normed ring are a closed set -/ +protected lemma is_closed : is_closed (nonunits R) := units.is_open.is_closed_compl + +end nonunits + namespace normed_ring open_locale classical big_operators open asymptotics filter metric finset ring @@ -288,3 +301,25 @@ lemma open_embedding_coe : open_embedding (coe : Rˣ → R) := open_embedding_of_continuous_injective_open continuous_coe ext is_open_map_coe end units + +namespace ideal + +/-- An ideal which contains an element within `1` of `1 : R` is the unit ideal. -/ +lemma eq_top_of_norm_lt_one (I : ideal R) {x : R} (hxI : x ∈ I) (hx : ∥1 - x∥ < 1) : I = ⊤ := +let u := units.one_sub (1 - x) hx in (I.eq_top_iff_one.mpr $ + by simpa only [show u.inv * x = 1, by simp] using I.mul_mem_left u.inv hxI) + +/-- The `ideal.closure` of a proper ideal in a complete normed ring is proper. -/ +lemma closure_ne_top (I : ideal R) (hI : I ≠ ⊤) : I.closure ≠ ⊤ := +have h : _ := closure_minimal (coe_subset_nonunits hI) nonunits.is_closed, + by simpa only [I.closure.eq_top_iff_one, ne.def] using mt (@h 1) one_not_mem_nonunits + +/-- The `ideal.closure` of a maximal ideal in a complete normed ring is the ideal itself. -/ +lemma is_maximal.closure_eq {I : ideal R} (hI : I.is_maximal) : I.closure = I := +(hI.eq_of_le (I.closure_ne_top hI.ne_top) subset_closure).symm + +/-- Maximal ideals in complete normed rings are closed. -/ +instance is_maximal.is_closed {I : ideal R} [hI : I.is_maximal] : is_closed (I : set R) := +is_closed_of_closure_subset $ eq.subset $ congr_arg (coe : ideal R → set R) hI.closure_eq + +end ideal diff --git a/src/topology/algebra/ring.lean b/src/topology/algebra/ring.lean index 46d628aabc25e..889c8398c11a8 100644 --- a/src/topology/algebra/ring.lean +++ b/src/topology/algebra/ring.lean @@ -285,8 +285,8 @@ def subring.comm_ring_topological_closure [t2_space α] (s : subring α) end topological_semiring -section topological_comm_ring -variables {α : Type*} [topological_space α] [comm_ring α] [topological_ring α] +section topological_ring +variables {α : Type*} [topological_space α] [ring α] [topological_ring α] /-- The closure of an ideal in a topological ring as an ideal. -/ def ideal.closure (S : ideal α) : ideal α := @@ -296,7 +296,7 @@ def ideal.closure (S : ideal α) : ideal α := @[simp] lemma ideal.coe_closure (S : ideal α) : (S.closure : set α) = closure S := rfl -end topological_comm_ring +end topological_ring section topological_ring variables {α : Type*} [topological_space α] [comm_ring α] (N : ideal α) From c64fb26a29e6987c51f43cb930de1997e06ea199 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 13 Sep 2022 09:02:48 +0000 Subject: [PATCH 242/828] feat(topology/algebra/order/intermediate_value): intervals are connected (#16473) `topology.algebra.order.intermediate_value` has a series of lemmas that different kinds of intervals are preconnected. Add a corresponding series of lemmas that intervals are connected (with appropriate extra conditions on the order or the endpoints as needed). --- .../algebra/order/intermediate_value.lean | 22 +++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/src/topology/algebra/order/intermediate_value.lean b/src/topology/algebra/order/intermediate_value.lean index 83e6f74b73028..ea6373adce4de 100644 --- a/src/topology/algebra/order/intermediate_value.lean +++ b/src/topology/algebra/order/intermediate_value.lean @@ -426,6 +426,28 @@ lemma is_preconnected_Ioo : is_preconnected (Ioo a b) := ord_connected_Ioo.is_pr lemma is_preconnected_Ioc : is_preconnected (Ioc a b) := ord_connected_Ioc.is_preconnected lemma is_preconnected_Ico : is_preconnected (Ico a b) := ord_connected_Ico.is_preconnected +lemma is_connected_Ici : is_connected (Ici a) := ⟨nonempty_Ici, is_preconnected_Ici⟩ + +lemma is_connected_Iic : is_connected (Iic a) := ⟨nonempty_Iic, is_preconnected_Iic⟩ + +lemma is_connected_Ioi [no_max_order α] : is_connected (Ioi a) := +⟨nonempty_Ioi, is_preconnected_Ioi⟩ + +lemma is_connected_Iio [no_min_order α] : is_connected (Iio a) := +⟨nonempty_Iio, is_preconnected_Iio⟩ + +lemma is_connected_Icc (h : a ≤ b) : is_connected (Icc a b) := +⟨nonempty_Icc.2 h, is_preconnected_Icc⟩ + +lemma is_connected_Ioo (h : a < b) : is_connected (Ioo a b) := +⟨nonempty_Ioo.2 h, is_preconnected_Ioo⟩ + +lemma is_connected_Ioc (h : a < b) : is_connected (Ioc a b) := +⟨nonempty_Ioc.2 h, is_preconnected_Ioc⟩ + +lemma is_connected_Ico (h : a < b) : is_connected (Ico a b) := +⟨nonempty_Ico.2 h, is_preconnected_Ico⟩ + @[priority 100] instance ordered_connected_space : preconnected_space α := ⟨ord_connected_univ.is_preconnected⟩ From a945b3769cb82bc238ee004b4327201a6864e7e0 Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Tue, 13 Sep 2022 10:49:18 +0000 Subject: [PATCH 243/828] feat(data/sym/basic): add `fill` `filter_ne` and `sigma_ext` (#16316) Add definitions and lemmas to support a proof for the multinomial theorem. Co-authored-by: Junyan Xu Co-authored-by: Kyle Miller --- src/data/sym/basic.lean | 51 +++++++++++++++++++++++++++++++++++++++++ 1 file changed, 51 insertions(+) diff --git a/src/data/sym/basic.lean b/src/data/sym/basic.lean index 3ea65790af6e9..89bb6b59db930 100644 --- a/src/data/sym/basic.lean +++ b/src/data/sym/basic.lean @@ -314,6 +314,57 @@ multiset.mem_attach _ _ := coe_injective $ multiset.attach_cons _ _ +/-- Change the length of a `sym` using an equality. +The simp-normal form is for the `cast` to be pushed outward. -/ +protected def cast {n m : ℕ} (h : n = m) : sym α n ≃ sym α m := +{ to_fun := λ s, ⟨s.val, s.2.trans h⟩, + inv_fun := λ s, ⟨s.val, s.2.trans h.symm⟩, + left_inv := λ s, subtype.ext rfl, + right_inv := λ s, subtype.ext rfl } + +@[simp] lemma cast_rfl : sym.cast rfl s = s := subtype.ext rfl + +@[simp] lemma cast_cast {n n' n'' : ℕ} (s : sym α n) (h : n = n') (h' : n' = n'') : + sym.cast h' (sym.cast h s) = sym.cast (h.trans h') s := rfl + +@[simp] lemma coe_cast {n m : ℕ} (h : n = m) (s : sym α n) : + (sym.cast h s : multiset α) = s := rfl + +/-- Append a pair of `sym` terms. -/ +def append {n n' : ℕ} (s : sym α n) (s' : sym α n') : sym α (n + n') := +⟨s.1 + s'.1, by simp_rw [← s.2, ← s'.2, map_add]⟩ + +@[simp] lemma append_inj_right {n n' : ℕ} (s : sym α n) (t t' : sym α n') : + s.append t = s.append t' ↔ t = t' := +subtype.ext_iff.trans $ (add_right_inj _).trans subtype.ext_iff.symm + +@[simp] lemma append_inj_left {n n' : ℕ} (s s' : sym α n) (t : sym α n') : + s.append t = s'.append t ↔ s = s' := +subtype.ext_iff.trans $ (add_left_inj _).trans subtype.ext_iff.symm + +lemma append_comm {n n' : ℕ} (s : sym α n) (s' : sym α n') : + s.append s' = sym.cast (add_comm _ _) (s'.append s) := +by { ext, simp [append, add_comm], } + +/-- Fill a term `m : sym α (n - i)` with `i` copies of `a` to obtain a term of `sym α n`. +This is a convenience wrapper for `m.append (repeat a i)` that adjusts the term using `sym.cast`. -/ +def fill (a : α) (i : fin (n + 1)) (m : sym α (n - i)) : sym α n := +sym.cast (nat.sub_add_cancel i.is_le) (m.append (repeat a i)) + +/-- Remove every `a` from a given `sym α n`. +Yields the number of copies `i` and a term of `sym α (n - i)`. -/ +def filter_ne [decidable_eq α] (a : α) {n : ℕ} (m : sym α n) : Σ i : fin (n + 1), sym α (n - i) := +⟨⟨m.1.count a, (multiset.count_le_card _ _).trans_lt $ by rw [m.2, nat.lt_succ_iff]⟩, + m.1.filter ((≠) a), eq_tsub_of_add_eq $ eq.trans begin + rw [← multiset.countp_eq_card_filter, add_comm], + exact (multiset.card_eq_countp_add_countp _ _).symm, + end m.2⟩ + +lemma sigma_sub_ext (m₁ m₂ : Σ i : fin (n + 1), sym α (n - i)) + (h : (m₁.2 : multiset α) = m₂.2) : m₁ = m₂ := +sigma.subtype_ext (fin.ext $ by rw [← nat.sub_sub_self m₁.1.is_le, ← nat.sub_sub_self m₂.1.is_le, + ← m₁.2.2, ← m₂.2.2, subtype.val_eq_coe, subtype.val_eq_coe, h]) h + end sym section equiv From 3c93c656c2eaf4677080d5251c2a1e263ea9695a Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 13 Sep 2022 13:37:32 +0000 Subject: [PATCH 244/828] feat(analysis/inner_product_space/l2_space): the family of subspaces spanned by finitely many elements of a Hilbert basis has dense supremum (#15821) --- src/analysis/inner_product_space/l2_space.lean | 10 ++++++++++ src/linear_algebra/span.lean | 3 +++ 2 files changed, 13 insertions(+) diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 4d7510f4504f2..518ff11a4b9a9 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -485,6 +485,16 @@ end (b.to_orthonormal_basis : ι → E) = b := orthonormal_basis.coe_mk _ _ +lemma finite_spans_dense (b : hilbert_basis ι 𝕜 E) : + (⨆ J : finset ι, span 𝕜 (J.image b : set E)).topological_closure = ⊤ := +eq_top_iff.mpr $ b.dense_span.ge.trans +begin + simp_rw [← submodule.span_Union], + exact topological_closure_mono (span_mono $ set.range_subset_iff.mpr $ + λ i, set.mem_Union_of_mem {i} $ finset.mem_coe.mpr $ finset.mem_image_of_mem _ $ + finset.mem_singleton_self i) +end + variables {v : ι → E} (hv : orthonormal 𝕜 v) include hv cplt diff --git a/src/linear_algebra/span.lean b/src/linear_algebra/span.lean index 0ead5fe1ea603..463c5e021dec6 100644 --- a/src/linear_algebra/span.lean +++ b/src/linear_algebra/span.lean @@ -51,6 +51,9 @@ lemma span_le {p} : span R s ≤ p ↔ s ⊆ p := lemma span_mono (h : s ⊆ t) : span R s ≤ span R t := span_le.2 $ subset.trans h subset_span +lemma span_monotone : monotone (span R : set M → submodule R M) := +λ _ _, span_mono + lemma span_eq_of_le (h₁ : s ⊆ p) (h₂ : p ≤ span R s) : span R s = p := le_antisymm (span_le.2 h₁) h₂ From 01a8a9f83f089ff6c6d393ea32b7d61703b67bac Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Tue, 13 Sep 2022 13:37:33 +0000 Subject: [PATCH 245/828] feat(convex/specific_functions): specific case of Jensen's inequality for powers (#16186) Proves a specific case of Jensen's inequality: a powers of a sum divided by the cardinality of the `finset` is less than or equal to the sum of the powers. --- src/analysis/convex/specific_functions.lean | 26 ++++++++++++++++++++- 1 file changed, 25 insertions(+), 1 deletion(-) diff --git a/src/analysis/convex/specific_functions.lean b/src/analysis/convex/specific_functions.lean index 1c2e1b8b27ae6..f8ad2ed96ed8b 100644 --- a/src/analysis/convex/specific_functions.lean +++ b/src/analysis/convex/specific_functions.lean @@ -31,7 +31,7 @@ For `p : ℝ`, prove that `λ x, x ^ p` is concave when `0 ≤ p ≤ 1` and stri -/ open real set -open_locale big_operators +open_locale big_operators nnreal /-- `exp` is strictly convex on the whole real line. -/ lemma strict_convex_on_exp : strict_convex_on ℝ univ exp := @@ -82,6 +82,30 @@ begin nat.sub_pos_of_lt hn) (nat.cast_pos.2 $ zero_lt_two.trans_le hn), end +/-- Specific case of Jensen's inequality for sums of powers -/ +lemma real.pow_sum_div_card_le_sum_pow {α : Type*} {s : finset α} {f : α → ℝ} (n : ℕ) + (hf : ∀ a ∈ s, 0 ≤ f a) : (∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) := +begin + by_cases hs0 : s = ∅, + { simp_rw [hs0, finset.sum_empty, zero_pow' _ (nat.succ_ne_zero n), zero_div] }, + { have hs : s.card ≠ 0 := hs0 ∘ finset.card_eq_zero.1, + have hs' : (s.card : ℝ) ≠ 0 := (nat.cast_ne_zero.2 hs), + have hs'' : 0 < (s.card : ℝ) := nat.cast_pos.2 (nat.pos_of_ne_zero hs), + suffices : (∑ x in s, f x / s.card) ^ (n + 1) ≤ ∑ x in s, (f x ^ (n + 1) / s.card), + by rwa [← finset.sum_div, ← finset.sum_div, div_pow, pow_succ' (s.card : ℝ), + ← div_div, div_le_iff hs'', div_mul, div_self hs', div_one] at this, + have := @convex_on.map_sum_le ℝ ℝ ℝ α _ _ _ _ _ _ (set.Ici 0) (λ x, x ^ (n + 1)) s + (λ _, 1 / s.card) (coe ∘ f) (convex_on_pow (n + 1)) _ _ (λ i hi, set.mem_Ici.2 (hf i hi)), + { simpa only [inv_mul_eq_div, one_div, algebra.id.smul_eq_mul] using this }, + { simp only [one_div, inv_nonneg, nat.cast_nonneg, implies_true_iff] }, + { simpa only [one_div, finset.sum_const, nsmul_eq_mul] using mul_inv_cancel hs' }} +end + +lemma nnreal.pow_sum_div_card_le_sum_pow {α : Type*} (s : finset α) (f : α → ℝ≥0) (n : ℕ) : + (∑ x in s, f x) ^ (n + 1) / s.card ^ n ≤ ∑ x in s, (f x) ^ (n + 1) := +by simpa only [← nnreal.coe_le_coe, nnreal.coe_sum, nonneg.coe_div, nnreal.coe_pow] using + @real.pow_sum_div_card_le_sum_pow α s (coe ∘ f) n (λ _ _, nnreal.coe_nonneg _) + lemma finset.prod_nonneg_of_card_nonpos_even {α β : Type*} [linear_ordered_comm_ring β] {f : α → β} [decidable_pred (λ x, f x ≤ 0)] From c0ad3bf92a6d4926d4eb9dcb31ed75621bf0b80a Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 13 Sep 2022 13:37:34 +0000 Subject: [PATCH 246/828] feat(topology/algebra/uniform_group): the quotient of a first countable complete topological group by a normal subgroup is itself complete (#16368) --- docs/references.bib | 11 +++ src/order/filter/bases.lean | 5 + src/topology/algebra/group.lean | 41 ++++++++ src/topology/algebra/uniform_group.lean | 119 ++++++++++++++++++++++++ 4 files changed, 176 insertions(+) diff --git a/docs/references.bib b/docs/references.bib index a425febd4dfc2..03eb047dc6dbc 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -251,6 +251,17 @@ @Book{ bourbaki1966 mrnumber = {0205210} } +@Book{ bourbaki1966b, + author = {Bourbaki, Nicolas}, + title = {Elements of mathematics. {G}eneral topology. {P}art 2}, + publisher = {Hermann, Paris; Addison-Wesley Publishing Co., Reading, + Mass.-London-Don Mills, Ont.}, + year = {1966}, + pages = {iv+363}, + mrclass = {54-02 (00A05 54-01)}, + mrnumber = {979295}, +} + @Book{ bourbaki1968, author = {Bourbaki, Nicolas}, title = {Lie groups and {L}ie algebras. {C}hapters 4--6}, diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index 5c311780c311a..6f87ab1839f77 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -954,6 +954,11 @@ begin ⟨hs.to_has_basis.inf ht.to_has_basis, set.to_countable _⟩ end +instance map.is_countably_generated (l : filter α) [l.is_countably_generated] (f : α → β) : + (map f l).is_countably_generated := +let ⟨x, hxl⟩ := l.exists_antitone_basis in +has_countable_basis.is_countably_generated ⟨hxl.map.to_has_basis, to_countable _⟩ + instance comap.is_countably_generated (l : filter β) [l.is_countably_generated] (f : α → β) : (comap f l).is_countably_generated := let ⟨x, hxl⟩ := l.exists_antitone_basis in diff --git a/src/topology/algebra/group.lean b/src/topology/algebra/group.lean index f004a03a18af3..b95b052dc9dc1 100644 --- a/src/topology/algebra/group.lean +++ b/src/topology/algebra/group.lean @@ -749,6 +749,47 @@ instance topological_group_quotient [N.normal] : topological_group (G ⧸ N) := end, continuous_inv := by convert (@continuous_inv G _ _ _).quotient_map' _ } +/-- Neighborhoods in the quotient are precisely the map of neighborhoods in the prequotient. -/ +@[to_additive "Neighborhoods in the quotient are precisely the map of neighborhoods in +the prequotient."] +lemma quotient_group.nhds_eq (x : G) : 𝓝 (x : G ⧸ N) = map coe (𝓝 x) := +le_antisymm ((quotient_group.is_open_map_coe N).nhds_le x) continuous_quot_mk.continuous_at + +variables (G) [first_countable_topology G] + +/-- Any first countable topological group has an antitone neighborhood basis `u : ℕ → set G` for +which `(u (n + 1)) ^ 2 ⊆ u n`. The existence of such a neighborhood basis is a key tool for +`quotient_group.complete_space` -/ +@[to_additive "Any first countable topological additive group has an antitone neighborhood basis +`u : ℕ → set G` for which `u (n + 1) + u (n + 1) ⊆ u n`. The existence of such a neighborhood basis +is a key tool for `quotient_add_group.complete_space`"] +lemma topological_group.exists_antitone_basis_nhds_one : + ∃ (u : ℕ → set G), (𝓝 1).has_antitone_basis u ∧ (∀ n, u (n + 1) * u (n + 1) ⊆ u n) := +begin + rcases (𝓝 (1 : G)).exists_antitone_basis with ⟨u, hu, u_anti⟩, + have := ((hu.prod_nhds hu).tendsto_iff hu).mp + (by simpa only [mul_one] using continuous_mul.tendsto ((1, 1) : G × G)), + simp only [and_self, mem_prod, and_imp, prod.forall, exists_true_left, prod.exists, + forall_true_left] at this, + have event_mul : ∀ n : ℕ, ∀ᶠ m in at_top, u m * u m ⊆ u n, + { intros n, + rcases this n with ⟨j, k, h⟩, + refine at_top_basis.eventually_iff.mpr ⟨max j k, true.intro, λ m hm, _⟩, + rintro - ⟨a, b, ha, hb, rfl⟩, + exact h a b (u_anti ((le_max_left _ _).trans hm) ha) (u_anti ((le_max_right _ _).trans hm) hb)}, + obtain ⟨φ, -, hφ, φ_anti_basis⟩ := has_antitone_basis.subbasis_with_rel ⟨hu, u_anti⟩ event_mul, + exact ⟨u ∘ φ, φ_anti_basis, λ n, hφ n.lt_succ_self⟩, +end + +include n + +/-- In a first countable topological group `G` with normal subgroup `N`, `1 : G ⧸ N` has a +countable neighborhood basis. -/ +@[to_additive "In a first countable topological additive group `G` with normal additive subgroup +`N`, `0 : G ⧸ N` has a countable neighborhood basis."] +instance quotient_group.nhds_one_is_countably_generated : (𝓝 (1 : G ⧸ N)).is_countably_generated := +(quotient_group.nhds_eq N 1).symm ▸ map.is_countably_generated _ _ + end quotient_topological_group /-- A typeclass saying that `λ p : G × G, p.1 - p.2` is a continuous function. This property diff --git a/src/topology/algebra/uniform_group.lean b/src/topology/algebra/uniform_group.lean index af3ad1cdbdeea..df29665f1b193 100644 --- a/src/topology/algebra/uniform_group.lean +++ b/src/topology/algebra/uniform_group.lean @@ -27,6 +27,10 @@ group naturally induces a uniform structure. to construct a canonical uniformity for a topological add group. * extension of ℤ-bilinear maps to complete groups (useful for ring completions) + +* `quotient_group.complete_space` and `quotient_add_group.complete_space` guarantee that quotients + of first countable topological groups by normal subgroups are themselves complete. In particular, + the quotient of a Banach space by a subspace is complete. -/ noncomputable theory @@ -745,3 +749,118 @@ begin apply h ; tauto } } end end dense_inducing + +section complete_quotient + +universe u +open topological_space classical + +/-- The quotient `G ⧸ N` of a complete first countable topological group `G` by a normal subgroup +is itself complete. [N. Bourbaki, *General Topology*, IX.3.1 Proposition 4][bourbaki1966b] + +Because a topological group is not equipped with a `uniform_space` instance by default, we must +explicitly provide it in order to consider completeness. See `quotient_group.complete_space` for a +version in which `G` is already equipped with a uniform structure. -/ +@[to_additive "The quotient `G ⧸ N` of a complete first countable topological additive group +`G` by a normal additive subgroup is itself complete. Consequently, quotients of Banach spaces by +subspaces are complete. [N. Bourbaki, *General Topology*, IX.3.1 Proposition 4][bourbaki1966b] + +Because an additive topological group is not equipped with a `uniform_space` instance by default, +we must explicitly provide it in order to consider completeness. See +`quotient_add_group.complete_space` for a version in which `G` is already equipped with a uniform +structure."] +instance quotient_group.complete_space' (G : Type u) [group G] [topological_space G] + [topological_group G] [first_countable_topology G] (N : subgroup G) [N.normal] + [@complete_space G (topological_group.to_uniform_space G)] : + @complete_space (G ⧸ N) (topological_group.to_uniform_space (G ⧸ N)) := +begin + /- Since `G ⧸ N` is a topological group it is a uniform space, and since `G` is first countable + the uniformities of both `G` and `G ⧸ N` are countably generated. Moreover, we may choose a + sequential antitone neighborhood basis `u` for `𝓝 (1 : G)` so that `(u (n + 1)) ^ 2 ⊆ u n`, and + this descends to an antitone neighborhood basis `v` for `𝓝 (1 : G ⧸ N)`. Since `𝓤 (G ⧸ N)` is + countably generated, it suffices to show any Cauchy sequence `x` converges. -/ + letI : uniform_space (G ⧸ N) := topological_group.to_uniform_space (G ⧸ N), + letI : uniform_space G := topological_group.to_uniform_space G, + haveI : (𝓤 (G ⧸ N)).is_countably_generated := comap.is_countably_generated _ _, + obtain ⟨u, hu, u_mul⟩ := topological_group.exists_antitone_basis_nhds_one G, + obtain ⟨hv, v_anti⟩ := @has_antitone_basis.map _ _ _ _ _ _ (coe : G → G ⧸ N) hu, + rw [←quotient_group.nhds_eq N 1, quotient_group.coe_one] at hv, + refine uniform_space.complete_of_cauchy_seq_tendsto (λ x hx, _), + /- Given `n : ℕ`, for sufficiently large `a b : ℕ`, given any lift of `x b`, we can find a lift + of `x a` such that the quotient of the lifts lies in `u n`. -/ + have key₀ : ∀ i j : ℕ, ∃ M : ℕ, + j < M ∧ ∀ a b : ℕ, M ≤ a → M ≤ b → ∀ g : G, x b = g → ∃ g' : G, g / g' ∈ u i ∧ x a = g', + { have h𝓤GN : (𝓤 (G ⧸ N)).has_basis (λ _, true) (λ i, {x | x.snd / x.fst ∈ coe '' u i}), + { simpa [uniformity_eq_comap_nhds_one'] using hv.comap _ }, + simp only [h𝓤GN.cauchy_seq_iff, ge_iff_le, mem_set_of_eq, forall_true_left, mem_image] at hx, + intros i j, + rcases hx i with ⟨M, hM⟩, + refine ⟨max j M + 1, (le_max_left _ _).trans_lt (lt_add_one _), λ a b ha hb g hg, _⟩, + obtain ⟨y, y_mem, hy⟩ := hM a (((le_max_right j _).trans (lt_add_one _).le).trans ha) b + (((le_max_right j _).trans (lt_add_one _).le).trans hb), + refine ⟨y⁻¹ * g, + by simpa only [div_eq_mul_inv, mul_inv_rev, inv_inv, mul_inv_cancel_left] using y_mem, _⟩, + rw [quotient_group.coe_mul, quotient_group.coe_inv, hy, hg, inv_div, div_mul_cancel'], }, + /- Inductively construct a subsequence `φ : ℕ → ℕ` using `key₀` so that if `a b : ℕ` exceed + `φ (n + 1)`, then we may find lifts whose quotients lie within `u n`. -/ + set φ : ℕ → ℕ := λ n, nat.rec_on n (some $ key₀ 0 0) (λ k yk, some $ key₀ (k + 1) yk), + have hφ : ∀ n : ℕ, φ n < φ (n + 1) ∧ ∀ a b : ℕ, φ (n + 1) ≤ a → φ (n + 1) ≤ b → + (∀ g : G, x b = g → ∃ g' : G, g / g' ∈ u (n + 1) ∧ x a = g'), + from λ n, some_spec (key₀ (n + 1) (φ n)), + /- Inductively construct a sequence `x' n : G` of lifts of `x (φ (n + 1))` such that quotients of + successive terms lie in `x' n / x' (n + 1) ∈ u (n + 1)`. We actually need the proofs that each + term is a lift to construct the next term, so we use a Σ-type. -/ + set x' : Π n, psigma (λ g : G, x (φ (n + 1)) = g) := + λ n, nat.rec_on n + ⟨some (quotient_group.mk_surjective (x (φ 1))), + (some_spec (quotient_group.mk_surjective (x (φ 1)))).symm⟩ + (λ k hk, ⟨some $ (hφ k).2 _ _ (hφ (k + 1)).1.le le_rfl hk.fst hk.snd, + (some_spec $ (hφ k).2 _ _ (hφ (k + 1)).1.le le_rfl hk.fst hk.snd).2⟩), + have hx' : ∀ n : ℕ, (x' n).fst / (x' (n + 1)).fst ∈ u (n + 1) := + λ n, (some_spec $ (hφ n).2 _ _ (hφ (n + 1)).1.le le_rfl (x' n).fst (x' n).snd).1, + /- The sequence `x'` is Cauchy. This is where we exploit the condition on `u`. The key idea + is to show by decreasing induction that `x' m / x' n ∈ u m` if `m ≤ n`. -/ + have x'_cauchy : cauchy_seq (λ n, (x' n).fst), + { have h𝓤G : (𝓤 G).has_basis (λ _, true) (λ i, {x | x.snd / x.fst ∈ u i}), + { simpa [uniformity_eq_comap_nhds_one'] using hu.to_has_basis.comap _ }, + simp only [h𝓤G.cauchy_seq_iff', ge_iff_le, mem_set_of_eq, forall_true_left], + exact λ m, ⟨m, λ n hmn, nat.decreasing_induction' + (λ k hkn hkm hk, u_mul k ⟨_, _, hx' k, hk, div_mul_div_cancel' _ _ _⟩) + hmn (by simpa only [div_self'] using mem_of_mem_nhds (hu.mem _))⟩ }, + /- Since `G` is complete, `x'` converges to some `x₀`, and so the image of this sequence under + the quotient map converges to `↑x₀`. The image of `x'` is a convergent subsequence of `x`, and + since `x` is Cauchy, this implies it converges. -/ + rcases cauchy_seq_tendsto_of_complete x'_cauchy with ⟨x₀, hx₀⟩, + refine ⟨↑x₀, tendsto_nhds_of_cauchy_seq_of_subseq hx + (strict_mono_nat_of_lt_succ $ λ n, (hφ (n + 1)).1).tendsto_at_top _⟩, + convert ((continuous_coinduced_rng : continuous (coe : G → G ⧸ N)).tendsto x₀).comp hx₀, + exact funext (λ n, (x' n).snd), +end + +/-- The quotient `G ⧸ N` of a complete first countable uniform group `G` by a normal subgroup +is itself complete. In constrast to `quotient_group.complete_space'`, in this version `G` is +already equipped with a uniform structure. +[N. Bourbaki, *General Topology*, IX.3.1 Proposition 4][bourbaki1966b] + +Even though `G` is equipped with a uniform structure, the quotient `G ⧸ N` does not inherit a +uniform structure, so it is still provided manually via `topological_group.to_uniform_space`. +In the most common use cases, this coincides (definitionally) with the uniform structure on the +quotient obtained via other means. -/ +@[to_additive "The quotient `G ⧸ N` of a complete first countable uniform additive group +`G` by a normal additive subgroup is itself complete. Consequently, quotients of Banach spaces by +subspaces are complete. In constrast to `quotient_add_group.complete_space'`, in this version +`G` is already equipped with a uniform structure. +[N. Bourbaki, *General Topology*, IX.3.1 Proposition 4][bourbaki1966b] + +Even though `G` is equipped with a uniform structure, the quotient `G ⧸ N` does not inherit a +uniform structure, so it is still provided manually via `topological_add_group.to_uniform_space`. +In the most common use case ─ quotients of normed additive commutative groups by subgroups ─ +significant care was taken so that the uniform structure inherent in that setting coincides +(definitionally) with the uniform structure provided here."] +instance quotient_group.complete_space (G : Type u) [group G] [us : uniform_space G] + [uniform_group G] [first_countable_topology G] (N : subgroup G) [N.normal] + [hG : complete_space G] : @complete_space (G ⧸ N) (topological_group.to_uniform_space (G ⧸ N)) := +by { unfreezingI { rw ←@uniform_group.to_uniform_space_eq _ us _ _ at hG }, apply_instance } + + +end complete_quotient From d268bb7a32e0007f86b495109c71baa8e491f660 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Tue, 13 Sep 2022 13:37:35 +0000 Subject: [PATCH 247/828] feat(data/polynomial/eval): add polynomial.map_bit0/1 (#16481) --- src/data/polynomial/eval.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/src/data/polynomial/eval.lean b/src/data/polynomial/eval.lean index 5972293e4330e..76c84b381ed03 100644 --- a/src/data/polynomial/eval.lean +++ b/src/data/polynomial/eval.lean @@ -538,7 +538,7 @@ end @[simp] protected lemma map_mul : (p * q).map f = p.map f * q.map f := by { rw [map, eval₂_mul_noncomm], exact λ k, (commute_X _).symm } -@[simp] lemma map_smul (r : R) : (r • p).map f = f r • p.map f := +@[simp] protected lemma map_smul (r : R) : (r • p).map f = f r • p.map f := by rw [map, eval₂_smul, ring_hom.comp_apply, C_mul'] /-- `polynomial.map` as a `ring_hom`. -/ @@ -561,6 +561,12 @@ def map_ring_hom (f : R →+* S) : R[X] →+* S[X] := @[simp] protected theorem map_nat_cast (n : ℕ) : (n : R[X]).map f = n := map_nat_cast (map_ring_hom f) n +@[simp] protected lemma map_bit0 : (bit0 p).map f = bit0 (p.map f) := +map_bit0 (map_ring_hom f) p + +@[simp] protected lemma map_bit1 : (bit1 p).map f = bit1 (p.map f) := +map_bit1 (map_ring_hom f) p + @[simp] lemma coeff_map (n : ℕ) : coeff (p.map f) n = f (coeff p n) := begin From ed9be880b0938405dcc2b829611ae8b938c60f6e Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 13 Sep 2022 19:19:23 +0000 Subject: [PATCH 248/828] docs(category_theory): mention notation for identity homs (#16410) --- src/category_theory/category/basic.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/category_theory/category/basic.lean b/src/category_theory/category/basic.lean index 75db02fc3617b..60477f4fe162f 100644 --- a/src/category_theory/category/basic.lean +++ b/src/category_theory/category/basic.lean @@ -14,8 +14,9 @@ Defines a category, as a type class parametrised by the type of objects. ## Notations Introduces notations -* `X ⟶ Y` for the morphism spaces, -* `f ≫ g` for composition in the 'arrows' convention. +* `X ⟶ Y` for the morphism spaces (type as `\hom`), +* `𝟙 X` for the identity morphism on `X` (type as `\b1`), +* `f ≫ g` for composition in the 'arrows' convention (type as `\gg`). Users may like to add `f ⊚ g` for composition in the standard convention, using ```lean From 3d36908752dc7abb4647d088fef03d1806a7e56d Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 13 Sep 2022 21:26:46 +0000 Subject: [PATCH 249/828] fix(order/bounded_order): remove classical axiom from Prop.distrib_lattice (#16497) Used all three axioms before, now only `propext`. --- src/order/bounded_order.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index 9708d43353b22..629e1d1335d0d 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -355,8 +355,7 @@ instance Prop.distrib_lattice : distrib_lattice Prop := inf_le_left := @and.left, inf_le_right := @and.right, le_inf := λ a b c Hab Hac Ha, and.intro (Hab Ha) (Hac Ha), - le_sup_inf := λ a b c H, or_iff_not_imp_left.2 $ - λ Ha, ⟨H.1.resolve_left Ha, H.2.resolve_left Ha⟩, + le_sup_inf := λ a b c, or_and_distrib_left.2, ..Prop.partial_order } /-- Propositions form a bounded order. -/ From 2ff9204589246c7e2fed82eaa9da7bd9169ea5e1 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Wed, 14 Sep 2022 00:27:32 +0000 Subject: [PATCH 250/828] feat(analysis/inner_product_space/projection): remove useless `complete_space E` assumption (#15682) The current proof decomposes v as its projection on K and its projection on the orthogonal complement of K, thus requiring `E` to be complete. Instead, we decompose it as `v = p v + (v - p v)`, where `p` is the projection on K. --- .../inner_product_space/projection.lean | 31 +++++++++++-------- 1 file changed, 18 insertions(+), 13 deletions(-) diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 73d8df8dd8ca3..65d4ad647837b 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -43,7 +43,7 @@ open_locale big_operators topological_space variables {𝕜 E F : Type*} [is_R_or_C 𝕜] variables [inner_product_space 𝕜 E] [inner_product_space ℝ F] -local notation `⟪`x`, `y`⟫` := @inner 𝕜 E _ x y +local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y local notation `absR` := has_abs.abs /-! ### Orthogonal projection in inner product spaces -/ @@ -915,7 +915,7 @@ begin let p2 := orthogonal_projection Sᗮ, have x_decomp : x = p1 x + p2 x := eq_sum_orthogonal_projection_self_orthogonal_complement S x, - have x_orth : ⟪ p1 x, p2 x ⟫ = 0 := + have x_orth : ⟪ (p1 x : E), p2 x ⟫ = 0 := submodule.inner_right_of_mem_orthogonal (set_like.coe_mem (p1 x)) (set_like.coe_mem (p2 x)), nth_rewrite 0 [x_decomp], simp only [sq, norm_add_sq_eq_norm_sq_add_norm_sq_of_inner_eq_zero ((p1 x) : E) (p2 x) x_orth, @@ -932,22 +932,27 @@ lemma id_eq_sum_orthogonal_projection_self_orthogonal_complement + Kᗮ.subtypeL.comp (orthogonal_projection Kᗮ) := by { ext w, exact eq_sum_orthogonal_projection_self_orthogonal_complement K w } +@[simp] lemma inner_orthogonal_projection_eq_of_mem_right [complete_space K] (u : K) (v : E) : + ⟪orthogonal_projection K v, u⟫ = ⟪v, u⟫ := +calc ⟪orthogonal_projection K v, u⟫ + = ⟪(orthogonal_projection K v : E), u⟫ : K.coe_inner _ _ +... = ⟪(orthogonal_projection K v : E), u⟫ + ⟪v - orthogonal_projection K v, u⟫ : + by rw [orthogonal_projection_inner_eq_zero _ _ (submodule.coe_mem _), add_zero] +... = ⟪v, u⟫ : + by rw [← inner_add_left, add_sub_cancel'_right] + +@[simp] lemma inner_orthogonal_projection_eq_of_mem_left [complete_space K] (u : K) (v : E) : + ⟪u, orthogonal_projection K v⟫ = ⟪(u : E), v⟫ := +by rw [← inner_conj_sym, ← inner_conj_sym (u : E), inner_orthogonal_projection_eq_of_mem_right] + /-- The orthogonal projection is self-adjoint. -/ -lemma inner_orthogonal_projection_left_eq_right [complete_space E] +lemma inner_orthogonal_projection_left_eq_right [complete_space K] (u v : E) : ⟪↑(orthogonal_projection K u), v⟫ = ⟪u, orthogonal_projection K v⟫ := -begin - nth_rewrite 0 eq_sum_orthogonal_projection_self_orthogonal_complement K v, - nth_rewrite 1 eq_sum_orthogonal_projection_self_orthogonal_complement K u, - rw [inner_add_left, inner_add_right, - submodule.inner_right_of_mem_orthogonal (submodule.coe_mem (orthogonal_projection K u)) - (submodule.coe_mem (orthogonal_projection Kᗮ v)), - submodule.inner_left_of_mem_orthogonal (submodule.coe_mem (orthogonal_projection K v)) - (submodule.coe_mem (orthogonal_projection Kᗮ u))], -end +by rw [← inner_orthogonal_projection_eq_of_mem_left, inner_orthogonal_projection_eq_of_mem_right] /-- The orthogonal projection is symmetric. -/ -lemma orthogonal_projection_is_symmetric [complete_space E] +lemma orthogonal_projection_is_symmetric [complete_space K] : (K.subtypeL ∘L orthogonal_projection K : E →ₗ[𝕜] E).is_symmetric := inner_orthogonal_projection_left_eq_right K From 8b8cd99cfd2c6ef5caffe15a358f56609e494e8e Mon Sep 17 00:00:00 2001 From: mcdoll Date: Wed, 14 Sep 2022 05:48:04 +0000 Subject: [PATCH 251/828] feat(analysis/calculus): Taylor's theorem (#15087) Proves Taylor's theorem for the mean value form of the remainder. There are four different versions: the general case for real-valued functions, the Lagrange form, the Cauchy form, and a version for vector valued functions. Co-authored-by: Moritz Doll --- docs/100.yaml | 4 + docs/overview.yaml | 1 + docs/undergrad.yaml | 2 +- src/algebra/group_with_zero/basic.lean | 3 + src/analysis/calculus/tangent_cone.lean | 4 + src/analysis/calculus/taylor.lean | 408 ++++++++++++++++++++++++ src/data/polynomial/module.lean | 4 + src/topology/algebra/monoid.lean | 31 +- 8 files changed, 452 insertions(+), 5 deletions(-) create mode 100644 src/analysis/calculus/taylor.lean diff --git a/docs/100.yaml b/docs/100.yaml index 2c5bb5c929690..0d2e10244a6a8 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -127,6 +127,10 @@ author : Anatole Dedecker, Yury Kudryashov 35: title : Taylor’s Theorem + decls : + - taylor_mean_remainder_lagrange + - taylor_mean_remainder_cauchy + author : Moritz Doll 36: title : Brouwer Fixed Point Theorem 37: diff --git a/docs/overview.yaml b/docs/overview.yaml index 5d93e9ff13425..9404beff926e3 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -282,6 +282,7 @@ Analysis: derivative of the inverse of a function: 'has_fderiv_at.of_local_left_inverse' Rolle's theorem: 'exists_deriv_eq_zero' mean value theorem: 'exists_ratio_deriv_eq_ratio_slope' + Taylor's theorem: 'analysis/calculus/taylor.html' $C^k$ function: 'cont_diff' Leibniz formula: 'fderiv_mul' local extrema: 'is_local_min.fderiv_eq_zero' diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index bc923730f01af..e37577deec32e 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -323,7 +323,7 @@ Single Variable Real Analysis: Taylor-like theorems: Taylor's theorem with little-o remainder: "https://en.wikipedia.org/wiki/Taylor%27s_theorem#Taylor's_theorem_in_one_real_variable" Taylor's theorem with integral form for remainder: 'https://en.wikipedia.org/wiki/Taylor%27s_theorem#Explicit_formulas_for_the_remainder' - Taylor's theorem with Lagrange form for remainder: 'https://en.wikipedia.org/wiki/Taylor%27s_theorem#Explicit_formulas_for_the_remainder' + Taylor's theorem with Lagrange form for remainder: 'taylor_mean_remainder_lagrange' Taylor series expansions: 'https://en.wikipedia.org/wiki/Taylor_series' Elementary functions (trigonometric, rational, $\exp$, $\log$, etc): polynomial functions: 'polynomial.eval' diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index fbaa1bb4b5609..228f125178567 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -905,6 +905,9 @@ lemma div_div_cancel' (ha : a ≠ 0) : a / (a / b) = b := ha.is_unit.div_div_can lemma div_helper (b : G₀) (h : a ≠ 0) : 1 / (a * b) * a = 1 / b := by rw [div_mul_eq_mul_div, one_mul, div_mul_right _ h] +lemma div_mul_eq_mul_div₀ (a b c : G₀) : (a / c) * b = a * b / c := +by simp_rw [div_eq_mul_inv, mul_assoc, mul_comm c⁻¹] + end comm_group_with_zero namespace semiconj_by diff --git a/src/analysis/calculus/tangent_cone.lean b/src/analysis/calculus/tangent_cone.lean index 824d445ab946d..cfdc15a72bf7e 100644 --- a/src/analysis/calculus/tangent_cone.lean +++ b/src/analysis/calculus/tangent_cone.lean @@ -427,6 +427,10 @@ is_open_Ioo.unique_diff_on lemma unique_diff_on_Icc_zero_one : unique_diff_on ℝ (Icc (0:ℝ) 1) := unique_diff_on_Icc zero_lt_one +lemma unique_diff_within_at_Ioo {a b t : ℝ} (ht : t ∈ set.Ioo a b) : + unique_diff_within_at ℝ (set.Ioo a b) t := +is_open.unique_diff_within_at is_open_Ioo ht + lemma unique_diff_within_at_Ioi (a : ℝ) : unique_diff_within_at ℝ (Ioi a) a := unique_diff_within_at_convex (convex_Ioi a) (by simp) (by simp) diff --git a/src/analysis/calculus/taylor.lean b/src/analysis/calculus/taylor.lean new file mode 100644 index 0000000000000..34c213296961f --- /dev/null +++ b/src/analysis/calculus/taylor.lean @@ -0,0 +1,408 @@ +/- +Copyright (c) 2022 Moritz Doll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Doll +-/ + +import analysis.calculus.iterated_deriv +import analysis.calculus.mean_value +import measure_theory.integral.interval_integral +import data.polynomial.basic +import data.polynomial.module + +/-! +# Taylor's theorem + +This file defines the Taylor polynomial of a real function `f : ℝ → E`, +where `E` is a normed vector space over `ℝ` and proves Taylor's theorem, +which states that if `f` is sufficiently smooth, then +`f` can be approximated by the Taylor polynomial up to an explicit error term. + +## Main definitions + +* `taylor_coeff_within`: the Taylor coefficient using `deriv_within` +* `taylor_within`: the Taylor polynomial using `deriv_within` + +## Main statements + +* `taylor_mean_remainder`: Taylor's theorem with the general form of the remainder term +* `taylor_mean_remainder_lagrange`: Taylor's theorem with the Lagrange remainder +* `taylor_mean_remainder_cauchy`: Taylor's theorem with the Cauchy remainder +* `exists_taylor_mean_remainder_bound`: Taylor's theorem for vector valued functions with a +polynomial bound on the remainder + +## TODO + +* the Peano form of the remainder +* the integral form of the remainder +* Generalization to higher dimensions + +## Tags + +Taylor polynomial, Taylor's theorem +-/ + + +open_locale big_operators interval topological_space nat +open set + +variables {𝕜 E F : Type*} +variables [normed_add_comm_group E] [normed_space ℝ E] + +/-- The `k`th coefficient of the Taylor polynomial. -/ +noncomputable +def taylor_coeff_within (f : ℝ → E) (k : ℕ) (s : set ℝ) (x₀ : ℝ) : E := +(k! : ℝ)⁻¹ • (iterated_deriv_within k f s x₀) + +/-- The Taylor polynomial with derivatives inside of a set `s`. + +The Taylor polynomial is given by +$$∑_{k=0}^n \\frac{(x - x₀)^k}{k!} f^{(k)}(x₀),$$ +where $f^{(k)}(x₀)$ denotes the iterated derivative in the set `s`. -/ +noncomputable +def taylor_within (f : ℝ → E) (n : ℕ) (s : set ℝ) (x₀ : ℝ) : polynomial_module ℝ E := +(finset.range (n+1)).sum (λ k, + polynomial_module.comp (polynomial.X - polynomial.C x₀) + (polynomial_module.single ℝ k (taylor_coeff_within f k s x₀))) + +/-- The Taylor polynomial with derivatives inside of a set `s` considered as a function `ℝ → E`-/ +noncomputable +def taylor_within_eval (f : ℝ → E) (n : ℕ) (s : set ℝ) (x₀ x : ℝ) : E := +polynomial_module.eval x (taylor_within f n s x₀) + +lemma taylor_within_succ (f : ℝ → E) (n : ℕ) (s : set ℝ) (x₀ : ℝ) : + taylor_within f (n+1) s x₀ = taylor_within f n s x₀ + + polynomial_module.comp (polynomial.X - polynomial.C x₀) + (polynomial_module.single ℝ (n+1) (taylor_coeff_within f (n+1) s x₀)) := +begin + dunfold taylor_within, + rw finset.sum_range_succ, +end + +@[simp] lemma taylor_within_eval_succ (f : ℝ → E) (n : ℕ) (s : set ℝ) (x₀ x : ℝ) : + taylor_within_eval f (n+1) s x₀ x = taylor_within_eval f n s x₀ x + + (((n + 1 : ℝ) * n!)⁻¹ * (x - x₀)^(n+1)) • iterated_deriv_within (n + 1) f s x₀ := +begin + simp_rw [taylor_within_eval, taylor_within_succ, linear_map.map_add, polynomial_module.comp_eval], + congr, + simp only [polynomial.eval_sub, polynomial.eval_X, polynomial.eval_C, + polynomial_module.eval_single, mul_inv_rev], + dunfold taylor_coeff_within, + rw [←mul_smul, mul_comm, nat.factorial_succ, nat.cast_mul, nat.cast_add, nat.cast_one, + mul_inv_rev], +end + +/-- The Taylor polynomial of order zero evaluates to `f x`. -/ +@[simp] lemma taylor_within_zero_eval (f : ℝ → E) (s : set ℝ) (x₀ x : ℝ) : + taylor_within_eval f 0 s x₀ x = f x₀ := +begin + dunfold taylor_within_eval, + dunfold taylor_within, + dunfold taylor_coeff_within, + simp, +end + +/-- Evaluating the Taylor polynomial at `x = x₀` yields `f x`. -/ +@[simp] lemma taylor_within_eval_self (f : ℝ → E) (n : ℕ) (s : set ℝ) (x₀ : ℝ) : + taylor_within_eval f n s x₀ x₀ = f x₀ := +begin + induction n with k hk, + { exact taylor_within_zero_eval _ _ _ _}, + simp [hk] +end + +lemma taylor_within_apply (f : ℝ → E) (n : ℕ) (s : set ℝ) (x₀ x : ℝ) : + taylor_within_eval f n s x₀ x = ∑ k in finset.range (n+1), + ((k! : ℝ)⁻¹ * (x - x₀)^k) • iterated_deriv_within k f s x₀ := +begin + induction n with k hk, + { simp }, + rw [taylor_within_eval_succ, finset.sum_range_succ, hk], + simp, +end + +/-- If `f` is `n` times continuous differentiable on a set `s`, then the Taylor polynomial + `taylor_within_eval f n s x₀ x` is continuous in `x₀`. -/ +lemma continuous_on_taylor_within_eval {f : ℝ → E} {x : ℝ} {n : ℕ} {s : set ℝ} + (hs : unique_diff_on ℝ s) (hf : cont_diff_on ℝ n f s) : + continuous_on (λ t, taylor_within_eval f n s t x) s := +begin + simp_rw taylor_within_apply, + refine continuous_on_finset_sum (finset.range (n+1)) (λ i hi, _), + refine (continuous_on_const.mul ((continuous_on_const.sub continuous_on_id).pow _)).smul _, + rw cont_diff_on_iff_continuous_on_differentiable_on_deriv hs at hf, + cases hf, + specialize hf_left i, + simp only [finset.mem_range] at hi, + refine (hf_left _), + simp only [with_top.coe_le_coe], + exact nat.lt_succ_iff.mp hi, +end + +/-- Helper lemma for calculating the derivative of the monomial that appears in Taylor expansions.-/ +lemma monomial_has_deriv_aux (t x : ℝ) (n : ℕ) : + has_deriv_at (λ y, (x - y)^(n+1)) (-(n+1) * (x - t)^n) t := +begin + simp_rw sub_eq_neg_add, + rw [←neg_one_mul, mul_comm (-1 : ℝ), mul_assoc, mul_comm (-1 : ℝ), ←mul_assoc], + convert @has_deriv_at.pow _ _ _ _ _ (n+1) ((has_deriv_at_id t).neg.add_const x), + simp only [nat.cast_add, nat.cast_one], +end + +lemma has_deriv_within_at_taylor_coeff_within {f : ℝ → E} {x y : ℝ} {k : ℕ} {s s' : set ℝ} + (hs'_unique : unique_diff_within_at ℝ s' y) + (hs' : s' ∈ 𝓝[s] y) (hy : y ∈ s') (h : s' ⊆ s) + (hf' : differentiable_on ℝ (iterated_deriv_within (k+1) f s) s') : + has_deriv_within_at (λ t, + (((k+1 : ℝ) * k!)⁻¹ * (x - t)^(k+1)) • iterated_deriv_within (k+1) f s t) + ((((k+1 : ℝ) * k!)⁻¹ * (x - y)^(k+1)) • iterated_deriv_within (k+2) f s y - + ((k! : ℝ)⁻¹ * (x - y)^k) • iterated_deriv_within (k+1) f s y) s' y := +begin + have hf'' : has_deriv_within_at (λ t, iterated_deriv_within (k+1) f s t) + (iterated_deriv_within (k+2) f s y) s' y := + begin + convert (hf' y hy).has_deriv_within_at, + rw iterated_deriv_within_succ (hs'_unique.mono h), + refine (deriv_within_subset h hs'_unique _).symm, + exact (hf' y hy).antimono h hs', + end, + have : has_deriv_within_at (λ t, (((k+1 : ℝ) * k!)⁻¹ * (x - t)^(k+1))) + (-((k! : ℝ)⁻¹ * (x - y)^k)) s' y := + begin + -- Commuting the factors: + have : (-((k! : ℝ)⁻¹ * (x - y)^k)) = + (((k+1 : ℝ) * k!)⁻¹ * (-(k+1) *(x - y)^k)) := + by { field_simp [nat.cast_add_one_ne_zero k, nat.factorial_ne_zero k], ring_nf }, + rw this, + exact (monomial_has_deriv_aux y x _).has_deriv_within_at.const_mul _, + end, + convert this.smul hf'', + field_simp [nat.cast_add_one_ne_zero k, nat.factorial_ne_zero k], + rw [neg_div, neg_smul, sub_eq_add_neg], +end + +/-- Calculate the derivative of the Taylor polynomial with respect to `x₀`. + +Version for arbitrary sets -/ +lemma has_deriv_within_at_taylor_within_eval {f : ℝ → E} {x y : ℝ} {n : ℕ} {s s' : set ℝ} + (hs'_unique : unique_diff_within_at ℝ s' y) (hs_unique : unique_diff_on ℝ s) + (hs' : s' ∈ 𝓝[s] y) (hy : y ∈ s') (h : s' ⊆ s) + (hf : cont_diff_on ℝ n f s) + (hf' : differentiable_on ℝ (iterated_deriv_within n f s) s') : + has_deriv_within_at (λ t, taylor_within_eval f n s t x) + (((n! : ℝ)⁻¹ * (x - y)^n) • (iterated_deriv_within (n+1) f s y)) s' y := +begin + induction n with k hk, + { simp only [taylor_within_zero_eval, nat.factorial_zero, nat.cast_one, inv_one, pow_zero, + mul_one, zero_add, one_smul], + simp only [iterated_deriv_within_zero] at hf', + rw iterated_deriv_within_one hs_unique (h hy), + refine has_deriv_within_at.mono _ h, + refine differentiable_within_at.has_deriv_within_at _, + exact (hf' y hy).antimono h hs' }, + simp_rw [nat.add_succ, taylor_within_eval_succ], + simp only [add_zero, nat.factorial_succ, nat.cast_mul, nat.cast_add, nat.cast_one], + have hdiff : differentiable_on ℝ (iterated_deriv_within k f s) s' := + begin + have coe_lt_succ : (k : with_top ℕ) < k.succ := + by { rw [with_top.coe_lt_coe], exact lt_add_one k }, + refine differentiable_on.mono _ h, + exact hf.differentiable_on_iterated_deriv_within coe_lt_succ hs_unique, + end, + specialize hk (cont_diff_on.of_succ hf) hdiff, + convert hk.add (has_deriv_within_at_taylor_coeff_within hs'_unique hs' hy h hf'), + exact (add_sub_cancel'_right _ _).symm, +end + +/-- Calculate the derivative of the Taylor polynomial with respect to `x₀`. + +Version for open intervals -/ +lemma taylor_within_eval_has_deriv_at_Ioo {f : ℝ → E} {a b t : ℝ} (x : ℝ) {n : ℕ} + (hx : a < b) (ht : t ∈ Ioo a b) + (hf : cont_diff_on ℝ n f (Icc a b)) + (hf' : differentiable_on ℝ (iterated_deriv_within n f (Icc a b)) (Ioo a b)) : + has_deriv_at (λ y, taylor_within_eval f n (Icc a b) y x) + (((n! : ℝ)⁻¹ * (x - t)^n) • (iterated_deriv_within (n+1) f (Icc a b) t)) t := +begin + have h_nhds := is_open.mem_nhds is_open_Ioo ht, + exact (has_deriv_within_at_taylor_within_eval (unique_diff_within_at_Ioo ht) + (unique_diff_on_Icc hx) (nhds_within_le_nhds h_nhds) ht Ioo_subset_Icc_self hf hf') + .has_deriv_at h_nhds, +end + +/-- Calculate the derivative of the Taylor polynomial with respect to `x₀`. + +Version for closed intervals -/ +lemma has_deriv_within_taylor_within_eval_at_Icc {f : ℝ → E} {a b t : ℝ} (x : ℝ) {n : ℕ} + (hx : a < b) (ht : t ∈ Icc a b) (hf : cont_diff_on ℝ n f (Icc a b)) + (hf' : differentiable_on ℝ (iterated_deriv_within n f (Icc a b)) (Icc a b)) : + has_deriv_within_at (λ y, taylor_within_eval f n (Icc a b) y x) + (((n! : ℝ)⁻¹ * (x - t)^n) • (iterated_deriv_within (n+1) f (Icc a b) t)) (Icc a b) t := +has_deriv_within_at_taylor_within_eval (unique_diff_on_Icc hx t ht) (unique_diff_on_Icc hx) + self_mem_nhds_within ht rfl.subset hf hf' + +/-! ### Taylor's theorem with mean value type remainder estimate -/ + +/-- **Taylor's theorem** with the general mean value form of the remainder. + +We assume that `f` is `n+1`-times continuously differentiable in the closed set `Icc x₀ x` and +`n+1`-times differentiable on the open set `Ioo x₀ x`, and `g` is a differentiable function on +`Ioo x₀ x` and continuous on `Icc x₀ x`. Then there exists a `x' ∈ Ioo x₀ x` such that +$$f(x) - (P_n f)(x₀, x) = \\frac{(x - x')^n}{n!} \\frac{g(x) - g(x₀)}{g' x'},$$ +where $P_n f$ denotes the Taylor polynomial of degree $n$. -/ +lemma taylor_mean_remainder {f : ℝ → ℝ} {g g' : ℝ → ℝ} {x x₀ : ℝ} {n : ℕ} (hx : x₀ < x) + (hf : cont_diff_on ℝ n f (Icc x₀ x)) + (hf' : differentiable_on ℝ (iterated_deriv_within n f (Icc x₀ x)) (Ioo x₀ x)) + (gcont : continuous_on g (Icc x₀ x)) + (gdiff : ∀ (x_1 : ℝ), x_1 ∈ Ioo x₀ x → has_deriv_at g (g' x_1) x_1) + (g'_ne : ∀ (x_1 : ℝ), x_1 ∈ Ioo x₀ x → g' x_1 ≠ 0) : + ∃ (x' : ℝ) (hx' : x' ∈ Ioo x₀ x), f x - taylor_within_eval f n (Icc x₀ x) x₀ x = + ((x - x')^n /n! * (g x - g x₀) / g' x') • + (iterated_deriv_within (n+1) f (Icc x₀ x) x') + := +begin + -- We apply the mean value theorem + rcases exists_ratio_has_deriv_at_eq_ratio_slope (λ t, taylor_within_eval f n (Icc x₀ x) t x) + (λ t, ((n! : ℝ)⁻¹ * (x - t)^n) • (iterated_deriv_within (n+1) f (Icc x₀ x) t)) hx + (continuous_on_taylor_within_eval (unique_diff_on_Icc hx) hf) + (λ _ hy, taylor_within_eval_has_deriv_at_Ioo x hx hy hf hf') + g g' gcont gdiff with ⟨y, hy, h⟩, + use [y, hy], + -- The rest is simplifications and trivial calculations + simp only [taylor_within_eval_self] at h, + rw [mul_comm, ←div_left_inj' (g'_ne y hy), mul_div_cancel _ (g'_ne y hy)] at h, + rw ←h, + field_simp [g'_ne y hy, nat.factorial_ne_zero n], + ring, +end + +/-- **Taylor's theorem** with the Lagrange form of the remainder. + +We assume that `f` is `n+1`-times continuously differentiable in the closed set `Icc x₀ x` and +`n+1`-times differentiable on the open set `Ioo x₀ x`. Then there exists a `x' ∈ Ioo x₀ x` such that +$$f(x) - (P_n f)(x₀, x) = \\frac{f^{(n+1)}(x') (x - x₀)^{n+1}}{(n+1)!},$$ +where $P_n f$ denotes the Taylor polynomial of degree $n$ and $f^{(n+1)}$ is the $n+1$-th iterated +derivative. -/ +lemma taylor_mean_remainder_lagrange {f : ℝ → ℝ} {x x₀ : ℝ} {n : ℕ} (hx : x₀ < x) + (hf : cont_diff_on ℝ n f (Icc x₀ x)) + (hf' : differentiable_on ℝ (iterated_deriv_within n f (Icc x₀ x)) (Ioo x₀ x)) : + ∃ (x' : ℝ) (hx' : x' ∈ Ioo x₀ x), f x - taylor_within_eval f n (Icc x₀ x) x₀ x = + (iterated_deriv_within (n+1) f (Icc x₀ x) x') * (x - x₀)^(n+1) /(n+1)! := +begin + have gcont : continuous_on (λ (t : ℝ), (x - t) ^ (n + 1)) (Icc x₀ x) := + by { refine continuous.continuous_on _, continuity }, + have xy_ne : ∀ (y : ℝ), y ∈ Ioo x₀ x → (x - y)^n ≠ 0 := + begin + intros y hy, + refine pow_ne_zero _ _, + rw [mem_Ioo] at hy, + rw sub_ne_zero, + exact hy.2.ne.symm, + end, + have hg' : ∀ (y : ℝ), y ∈ Ioo x₀ x → -(↑n + 1) * (x - y) ^ n ≠ 0 := + λ y hy, mul_ne_zero (neg_ne_zero.mpr (nat.cast_add_one_ne_zero n)) (xy_ne y hy), + -- We apply the general theorem with g(t) = (x - t)^(n+1) + rcases taylor_mean_remainder hx hf hf' gcont (λ y _, monomial_has_deriv_aux y x _) hg' + with ⟨y, hy, h⟩, + use [y, hy], + simp only [sub_self, zero_pow', ne.def, nat.succ_ne_zero, not_false_iff, zero_sub, mul_neg] at h, + rw [h, neg_div, ←div_neg, neg_mul, neg_neg], + field_simp [nat.cast_add_one_ne_zero n, nat.factorial_ne_zero n, xy_ne y hy], + ring, +end + +/-- **Taylor's theorem** with the Cauchy form of the remainder. + +We assume that `f` is `n+1`-times continuously differentiable on the closed set `Icc x₀ x` and +`n+1`-times differentiable on the open set `Ioo x₀ x`. Then there exists a `x' ∈ Ioo x₀ x` such that +$$f(x) - (P_n f)(x₀, x) = \\frac{f^{(n+1)}(x') (x - x')^n (x-x₀)}{n!},$$ +where $P_n f$ denotes the Taylor polynomial of degree $n$ and $f^{(n+1)}$ is the $n+1$-th iterated +derivative. -/ +lemma taylor_mean_remainder_cauchy {f : ℝ → ℝ} {x x₀ : ℝ} {n : ℕ} (hx : x₀ < x) + (hf : cont_diff_on ℝ n f (Icc x₀ x)) + (hf' : differentiable_on ℝ (iterated_deriv_within n f (Icc x₀ x)) (Ioo x₀ x)) : + ∃ (x' : ℝ) (hx' : x' ∈ Ioo x₀ x), f x - taylor_within_eval f n (Icc x₀ x) x₀ x = + (iterated_deriv_within (n+1) f (Icc x₀ x) x') * (x - x')^n /n! * (x - x₀) := +begin + have gcont : continuous_on id (Icc x₀ x) := continuous.continuous_on (by continuity), + have gdiff : (∀ (x_1 : ℝ), x_1 ∈ Ioo x₀ x → has_deriv_at id + ((λ (t : ℝ), (1 : ℝ)) x_1) x_1) := λ _ _, has_deriv_at_id _, + -- We apply the general theorem with g = id + rcases taylor_mean_remainder hx hf hf' gcont gdiff (λ _ _, by simp) with ⟨y, hy, h⟩, + use [y, hy], + rw h, + field_simp [nat.factorial_ne_zero n], + ring, +end + +/-- **Taylor's theorem** with a polynomial bound on the remainder + +We assume that `f` is `n+1`-times continuously differentiable on the closed set `Icc a b`. +The difference of `f` and its `n`-th Taylor polynomial can be estimated by +`C * (x - a)^(n+1) / n!` where `C` is a bound for the `n+1`-th iterated derivative of `f`. -/ +lemma taylor_mean_remainder_bound {f : ℝ → E} {a b C x : ℝ} {n : ℕ} + (hab : a ≤ b) (hf : cont_diff_on ℝ (n+1) f (Icc a b)) (hx : x ∈ Icc a b) + (hC : ∀ y ∈ Icc a b, ∥iterated_deriv_within (n + 1) f (Icc a b) y∥ ≤ C) : + ∥f x - taylor_within_eval f n (Icc a b) a x∥ ≤ C * (x - a)^(n+1) / n! := +begin + rcases eq_or_lt_of_le hab with rfl|h, + { rw [Icc_self, mem_singleton_iff] at hx, + simp [hx] }, + -- The nth iterated derivative is differentiable + have hf' : differentiable_on ℝ (iterated_deriv_within n f (Icc a b)) (Icc a b) := + hf.differentiable_on_iterated_deriv_within (with_top.coe_lt_coe.mpr n.lt_succ_self) + (unique_diff_on_Icc h), + -- natural numbers are non-negative + have fac_nonneg : 0 ≤ (n! : ℝ) := n!.cast_nonneg, + -- We can uniformly bound the derivative of the Taylor polynomial + have h' : ∀ (y : ℝ) (hy : y ∈ Ico a x), + ∥((n! : ℝ)⁻¹ * (x - y) ^ n) • iterated_deriv_within (n + 1) f (Icc a b) y∥ + ≤ (n! : ℝ)⁻¹ * |(x - a)|^n * C, + { intros y hy, + rw [norm_smul, real.norm_eq_abs], + -- Estimate the iterated derivative by `C` + refine mul_le_mul _ (hC y ⟨hy.1, hy.2.le.trans hx.2⟩) (by positivity) (by positivity), + -- The rest is a trivial calculation + rw [abs_mul, abs_pow, abs_inv, nat.abs_cast], + mono*, + any_goals { positivity }, + { exact hy.1 }, + { linarith [hx.1, hy.2] } }, + -- Apply the mean value theorem for vector valued functions: + have A : ∀ t ∈ Icc a x, has_deriv_within_at (λ y, taylor_within_eval f n (Icc a b) y x) + (((↑n!)⁻¹ * (x - t) ^ n) • iterated_deriv_within (n + 1) f (Icc a b) t) (Icc a x) t, + { assume t ht, + have I : Icc a x ⊆ Icc a b := Icc_subset_Icc_right hx.2, + exact (has_deriv_within_taylor_within_eval_at_Icc x h (I ht) hf.of_succ hf').mono I }, + have := norm_image_sub_le_of_norm_deriv_le_segment' A h' x (right_mem_Icc.2 hx.1), + simp only [taylor_within_eval_self] at this, + refine le_trans this (le_of_eq _), + -- The rest is a trivial calculation + rw [abs_of_nonneg (sub_nonneg.mpr hx.1)], + ring_exp, +end + + +/-- **Taylor's theorem** with a polynomial bound on the remainder + +We assume that `f` is `n+1`-times continuously differentiable on the closed set `Icc a b`. +There exists a constant `C` such that for all `x ∈ Icc a b` The difference of `f` and its `n`-th +Taylor polynomial can be estimated by `C * (x - a)^(n+1) / n!` where `C` is a bound for the `n+1`-th +iterated derivative of `f`. -/ +lemma exists_taylor_mean_remainder_bound {f : ℝ → E} {a b : ℝ} {n : ℕ} + (hab : a ≤ b) (hf : cont_diff_on ℝ (n+1) f (Icc a b)) : + ∃ C, ∀ x ∈ Icc a b, ∥f x - taylor_within_eval f n (Icc a b) a x∥ ≤ C * (x - a)^(n+1) := +begin + rcases eq_or_lt_of_le hab with rfl|h, + { refine ⟨0, λ x hx, _⟩, + have : a = x, by simpa [← le_antisymm_iff] using hx, + simp [← this] }, + -- We estimate by the supremum of the norm of the iterated derivative + let g : ℝ → ℝ := λ y, ∥iterated_deriv_within (n + 1) f (Icc a b) y∥, + use [has_Sup.Sup (g '' Icc a b) / n!], + intros x hx, + rw div_mul_eq_mul_div₀, + refine taylor_mean_remainder_bound hab hf hx (λ y, _), + exact (hf.continuous_on_iterated_deriv_within rfl.le $ unique_diff_on_Icc h) + .norm.le_Sup_image_Icc, +end diff --git a/src/data/polynomial/module.lean b/src/data/polynomial/module.lean index a38b84fdf0c30..cb7062752b9df 100644 --- a/src/data/polynomial/module.lean +++ b/src/data/polynomial/module.lean @@ -231,6 +231,10 @@ def eval (r : R) : polynomial_module R M →ₗ[R] M := lemma eval_single (r : R) (i : ℕ) (m : M) : eval r (single R i m) = r ^ i • m := finsupp.sum_single_index (smul_zero _) +@[simp] +lemma eval_lsingle (r : R) (i : ℕ) (m : M) : eval r (lsingle R i m) = r ^ i • m := +eval_single r i m + lemma eval_smul (p : R[X]) (q : polynomial_module R M) (r : R) : eval r (p • q) = p.eval r • eval r q := begin diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index f7123d4ceeb80..6cc5cf80fe184 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -358,11 +358,24 @@ lemma tendsto_list_prod {f : ι → α → M} {x : filter α} {a : ι → M} : @[to_additive] lemma continuous_list_prod {f : ι → X → M} (l : list ι) - (h : ∀i∈l, continuous (f i)) : - continuous (λa, (l.map (λi, f i a)).prod) := + (h : ∀ i ∈ l, continuous (f i)) : + continuous (λ a, (l.map (λ i, f i a)).prod) := continuous_iff_continuous_at.2 $ assume x, tendsto_list_prod l $ assume c hc, continuous_iff_continuous_at.1 (h c hc) x +@[to_additive] +lemma continuous_on_list_prod {f : ι → X → M} (l : list ι) {t : set X} + (h : ∀ i ∈ l, continuous_on (f i) t) : + continuous_on (λ a, (l.map (λ i, f i a)).prod) t := +begin + intros x hx, + rw continuous_within_at_iff_continuous_at_restrict _ hx, + refine tendsto_list_prod _ (λ i hi, _), + specialize h i hi x hx, + rw continuous_within_at_iff_continuous_at_restrict _ hx at h, + exact h, +end + @[continuity, to_additive] lemma continuous_pow : ∀ n : ℕ, continuous (λ a : M, a ^ n) | 0 := by simpa using continuous_const @@ -493,14 +506,24 @@ tendsto_multiset_prod _ @[continuity, to_additive] lemma continuous_multiset_prod {f : ι → X → M} (s : multiset ι) : - (∀i ∈ s, continuous (f i)) → continuous (λ a, (s.map (λ i, f i a)).prod) := + (∀ i ∈ s, continuous (f i)) → continuous (λ a, (s.map (λ i, f i a)).prod) := by { rcases s with ⟨l⟩, simpa using continuous_list_prod l } +@[to_additive] +lemma continuous_on_multiset_prod {f : ι → X → M} (s : multiset ι) {t : set X} : + (∀i ∈ s, continuous_on (f i) t) → continuous_on (λ a, (s.map (λ i, f i a)).prod) t := +by { rcases s with ⟨l⟩, simpa using continuous_on_list_prod l } + @[continuity, to_additive] lemma continuous_finset_prod {f : ι → X → M} (s : finset ι) : - (∀ i ∈ s, continuous (f i)) → continuous (λa, ∏ i in s, f i a) := + (∀ i ∈ s, continuous (f i)) → continuous (λ a, ∏ i in s, f i a) := continuous_multiset_prod _ +@[to_additive] +lemma continuous_on_finset_prod {f : ι → X → M} (s : finset ι) {t : set X} : + (∀ i ∈ s, continuous_on (f i) t) → continuous_on (λ a, ∏ i in s, f i a) t := +continuous_on_multiset_prod _ + @[to_additive] lemma eventually_eq_prod {X M : Type*} [comm_monoid M] {s : finset ι} {l : filter X} {f g : ι → X → M} (hs : ∀ i ∈ s, f i =ᶠ[l] g i) : ∏ i in s, f i =ᶠ[l] ∏ i in s, g i := From f9f5d51e123317f6b2a0b66ca7ae70ba1bfe2379 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 14 Sep 2022 08:28:08 +0000 Subject: [PATCH 252/828] feat(algebra/order/group): some more lemmas about `min`, `max`, and `abs` (#16485) These are trivially true, but require a bit of annoying casework, making them handy to package into lemmas. --- src/algebra/order/group.lean | 22 +++++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/src/algebra/order/group.lean b/src/algebra/order/group.lean index daec169b69b28..081f861d78bae 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group.lean @@ -1172,11 +1172,31 @@ begin rintro (rfl|rfl); simp only [abs_neg, abs_of_nonneg hb] end -lemma abs_le_max_abs_abs (hab : a ≤ b) (hbc : b ≤ c) : |b| ≤ max (|a|) (|c|) := +lemma abs_le_max_abs_abs (hab : a ≤ b) (hbc : b ≤ c) : |b| ≤ max (|a|) (|c|) := abs_le'.2 ⟨by simp [hbc.trans (le_abs_self c)], by simp [(neg_le_neg_iff.mpr hab).trans (neg_le_abs_self a)]⟩ +lemma min_abs_abs_le_abs_max : min (|a|) (|b|) ≤ |max a b| := +(le_total a b).elim + (λ h, (min_le_right _ _).trans_eq $ congr_arg _ (max_eq_right h).symm) + (λ h, (min_le_left _ _).trans_eq $ congr_arg _ (max_eq_left h).symm) + +lemma min_abs_abs_le_abs_min : min (|a|) (|b|) ≤ |min a b| := +(le_total a b).elim + (λ h, (min_le_left _ _).trans_eq $ congr_arg _ (min_eq_left h).symm) + (λ h, (min_le_right _ _).trans_eq $ congr_arg _ (min_eq_right h).symm) + +lemma abs_max_le_max_abs_abs : |max a b| ≤ max (|a|) (|b|) := +(le_total a b).elim + (λ h, (congr_arg _ $ max_eq_right h).trans_le $ le_max_right _ _) + (λ h, (congr_arg _ $ max_eq_left h).trans_le $ le_max_left _ _) + +lemma abs_min_le_max_abs_abs : |min a b| ≤ max (|a|) (|b|) := +(le_total a b).elim + (λ h, (congr_arg _ $ min_eq_left h).trans_le $ le_max_left _ _) + (λ h, (congr_arg _ $ min_eq_right h).trans_le $ le_max_right _ _) + lemma eq_of_abs_sub_eq_zero {a b : α} (h : |a - b| = 0) : a = b := sub_eq_zero.1 $ abs_eq_zero.1 h From a3bb2050cb912e56a99b84efcb267c347c378d28 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 14 Sep 2022 08:28:10 +0000 Subject: [PATCH 253/828] feat(topology/separation): generalize&rename a lemma (#16503) * rename `tot_sep_of_zero_dim` to `totally_separated_space_of_t1_of_basis_clopen`; * generalize it from a `t2_space` to a `t1_space`. --- src/data/set/basic.lean | 3 +-- src/topology/separation.lean | 22 ++++++++++------------ 2 files changed, 11 insertions(+), 14 deletions(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 2d56823fba457..e357532fd41ec 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -159,8 +159,7 @@ end set_coe lemma subtype.mem {α : Type*} {s : set α} (p : s) : (p : α) ∈ s := p.prop /-- Duplicate of `eq.subset'`, which currently has elaboration problems. -/ -lemma eq.subset {α} {s t : set α} : s = t → s ⊆ t := -by { rintro rfl x hx, exact hx } +lemma eq.subset {α} {s t : set α} : s = t → s ⊆ t := eq.subset' namespace set diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 43e1623478a1c..7f254e8977f1b 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -58,7 +58,8 @@ This file defines the predicate `separated`, and common separation axioms * `is_compact.is_closed`: All compact sets are closed. * `locally_compact_of_compact_nhds`: If every point has a compact neighbourhood, then the space is locally compact. -* `tot_sep_of_zero_dim`: If `α` has a clopen basis, it is a `totally_separated_space`. +* `totally_separated_space_of_t1_of_basis_clopen`: If `α` has a clopen basis, then + it is a `totally_separated_space`. * `loc_compact_t2_tot_disc_iff_tot_sep`: A locally compact T₂ space is totally disconnected iff it is totally separated. @@ -1544,22 +1545,19 @@ end section profinite -variables [t2_space α] - -/-- A Hausdorff space with a clopen basis is totally separated. -/ -lemma tot_sep_of_zero_dim (h : is_topological_basis {s : set α | is_clopen s}) : +/-- A T1 space with a clopen basis is totally separated. -/ +lemma totally_separated_space_of_t1_of_basis_clopen [t1_space α] + (h : is_topological_basis {s : set α | is_clopen s}) : totally_separated_space α := begin constructor, rintros x - y - hxy, - obtain ⟨u, v, hu, hv, xu, yv, disj⟩ := t2_separation hxy, - obtain ⟨w, hw : is_clopen w, xw, wu⟩ := (is_topological_basis.mem_nhds_iff h).1 - (is_open.mem_nhds hu xu), - refine ⟨w, wᶜ, hw.1, hw.compl.1, xw, λ h, disj ⟨wu h, yv⟩, _, disjoint_compl_right⟩, - rw set.union_compl_self, + rcases h.mem_nhds_iff.mp (is_open_ne.mem_nhds hxy) with ⟨U, hU, hxU, hyU⟩, + exact ⟨U, Uᶜ, hU.is_open, hU.compl.is_open, hxU, λ h, hyU h rfl, + (union_compl_self U).superset, disjoint_compl_right⟩ end -variables [compact_space α] +variables [t2_space α] [compact_space α] /-- A compact Hausdorff space is totally disconnected if and only if it is totally separated, this is also true for locally compact spaces. -/ @@ -1678,7 +1676,7 @@ theorem loc_compact_t2_tot_disc_iff_tot_sep : begin split, { introI h, - exact tot_sep_of_zero_dim loc_compact_Haus_tot_disc_of_zero_dim, }, + exact totally_separated_space_of_t1_of_basis_clopen loc_compact_Haus_tot_disc_of_zero_dim }, apply totally_separated_space.totally_disconnected_space, end From b118db5c468cb6db526f21a4d176f73e6f2f7465 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 14 Sep 2022 11:13:02 +0000 Subject: [PATCH 254/828] feat(ring_theory/roots_of_unity): add roots_of_unity.norm_one (#16426) This is the proof that the norm of the image of a root of unity by a ring homomorphism is always equal to one. This is the counterpart of the result proved in #15143 that an algebraic integer whose conjugates are all of absolute value one is a root of unity. From flt-regular Co-authored-by: Alex J. Best [alex.j.best@gmail.com](mailto:alex.j.best@gmail.com) --- src/analysis/normed/field/basic.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 9c66d6aa5ea05..68fea94d1ad80 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -481,6 +481,23 @@ begin simp, end +lemma norm_one_of_pow_eq_one {x : α} {k : ℕ+} (h : x ^ (k : ℕ) = 1) : + ∥x∥ = 1 := +begin + rw ( _ : ∥x∥ = 1 ↔ ∥x∥₊ = 1), + apply (@pow_left_inj nnreal _ _ _ ↑k zero_le' zero_le' (pnat.pos k)).mp, + { rw [← nnnorm_pow, one_pow, h, nnnorm_one], }, + { exact subtype.mk_eq_mk.symm, }, +end + +lemma norm_map_one_of_pow_eq_one [comm_monoid β] (φ : β →* α) {x : β} {k : ℕ+} + (h : x ^ (k : ℕ) = 1) : + ∥φ x∥ = 1 := +begin + have : (φ x) ^ (k : ℕ) = 1 := by rw [← monoid_hom.map_pow, h, monoid_hom.map_one], + exact norm_one_of_pow_eq_one this, +end + end normed_division_ring /-- A normed field is a field with a norm satisfying ∥x y∥ = ∥x∥ ∥y∥. -/ From e1452ef24e117a92df20b6d45f80b53cabe5a177 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Wed, 14 Sep 2022 11:13:07 +0000 Subject: [PATCH 255/828] chore(number_theory/legendre_symbol/*): move zmod.{legendre|jacobi}_sym* to other namespaces (#16461) This PR moves `legendre_sym` and `jacobi_sym` (and `qr_sign`) to the root namespace. It also moves definitions and lemmas named `legendre_sym_...` to the `legendre_sym` namespace and analogous for `jacobi_sym_...` and `qr_sign_...`. We change names like `jacobi_sym_two` to `jacobi_sym.at_two`. This mostly affects the files `number_theory.legendre_symbol.quadratic_reciprocity` and `number_theory.legendre_symbol.jacobi_symbol`. Note that names like `exists_sq_eq_neg_one_iff` have been left in `zmod`. We prefix `quadratic_reciprocity*` by `legendre_sym` to disambiguate with the versions for the Jacobi symbol. (The links in `docs/100.yaml` and `docs/overview.yaml` are updated accordingly.) We also move the results in `number_theory.legendre_symbol.gauss_eisenstein_lemmas` from the `legendre_symbol` namespace to the `zmod` namespace. See this [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Jacobi.20symbol/near/297792676) and the comments at [#16395](https://github.com/leanprover-community/mathlib/pull/16395) for discussion. --- docs/100.yaml | 8 +- docs/overview.yaml | 2 +- .../gauss_eisenstein_lemmas.lean | 16 +- .../legendre_symbol/jacobi_symbol.lean | 246 ++++++++++-------- .../quadratic_reciprocity.lean | 138 ++++++---- 5 files changed, 237 insertions(+), 173 deletions(-) diff --git a/docs/100.yaml b/docs/100.yaml index 0d2e10244a6a8..0f742a3fdc935 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -17,11 +17,13 @@ 5: title : Prime Number Theorem 6: - title : Godel’s Incompleteness Theorem + title : Gödel’s Incompleteness Theorem 7: title : Law of Quadratic Reciprocity - decl : zmod.quadratic_reciprocity - author : Chris Hughes + decls : + - legendre_sym.quadratic_reciprocity + - jacobi_sym.quadratic_reciprocity + author : Chris Hughes (first) and Michael Stoll (second) 8: title : The Impossibility of Trisecting the Angle and Doubling the Cube 9: diff --git a/docs/overview.yaml b/docs/overview.yaml index 9404beff926e3..b59896809a824 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -127,7 +127,7 @@ General algebra: Number theory: sum of two squares: 'nat.prime.sq_add_sq' sum of four squares: 'nat.sum_four_squares' - quadratic reciprocity: 'zmod.quadratic_reciprocity' + quadratic reciprocity: 'legendre_sym.quadratic_reciprocity' "solutions to Pell's equation": 'pell.pell' "Matiyasevič's theorem": 'pell.matiyasevic' arithmetic functions: 'nat.arithmetic_function' diff --git a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean index 3fed5d4fece90..ce813915fe53b 100644 --- a/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean +++ b/src/number_theory/legendre_symbol/gauss_eisenstein_lemmas.lean @@ -9,9 +9,8 @@ import number_theory.legendre_symbol.quadratic_reciprocity # Lemmas of Gauss and Eisenstein This file contains code for the proof of the Lemmas of Gauss and Eisenstein -on the Legendre symbol. The main results are `gauss_lemma_aux` and -`eisenstein_lemma_aux`; they are used in `quadratic_reciprocity.lean` -to prove `gauss_lemma` and `eisenstein_lemma`, respectively. +on the Legendre symbol. The main results are `zmod.gauss_lemma_aux` and +`zmod.eisenstein_lemma_aux`. -/ open function finset nat finite_field zmod @@ -23,7 +22,6 @@ section wilson variables (p : ℕ) [fact p.prime] --- One can probably deduce the following from `finite_field.prod_univ_units_id_eq_neg_one` /-- **Wilson's Lemma**: the product of `1`, ..., `p-1` is `-1` modulo `p`. -/ @[simp] lemma wilsons_lemma : ((p - 1)! : zmod p) = -1 := begin @@ -65,7 +63,7 @@ end zmod section gauss_eisenstein -namespace legendre_symbol +namespace zmod /-- The image of the map sending a non zero natural number `x ≤ p / 2` to the absolute value of the element of interger in the interval `(-p/2, p/2]` congruent to `a * x` mod p is the set @@ -161,8 +159,8 @@ begin haveI hp' : fact (p % 2 = 1) := ⟨nat.prime.mod_two_eq_one_iff_ne_two.mpr hp⟩, have : (legendre_sym p a : zmod p) = (((-1)^((Ico 1 (p / 2).succ).filter (λ x : ℕ, p / 2 < (a * x : zmod p).val)).card : ℤ) : zmod p) := - by { rw [legendre_sym_eq_pow, legendre_symbol.gauss_lemma_aux p ha0]; simp }, - cases legendre_sym_eq_one_or_neg_one p ha0; + by { rw [legendre_sym.eq_pow, gauss_lemma_aux p ha0]; simp }, + cases legendre_sym.eq_one_or_neg_one p ha0; cases neg_one_pow_eq_or ℤ ((Ico 1 (p / 2).succ).filter (λ x : ℕ, p / 2 < (a * x : zmod p).val)).card; simp [*, ne_neg_self p one_ne_zero, (ne_neg_self p one_ne_zero).symm] at * @@ -299,9 +297,9 @@ begin have ha0' : ((a : ℤ) : zmod p) ≠ 0 := by { norm_cast, exact ha0 }, rw [neg_one_pow_eq_pow_mod_two, gauss_lemma hp ha0', neg_one_pow_eq_pow_mod_two, (by norm_cast : ((a : ℤ) : zmod p) = (a : zmod p)), - show _ = _, from legendre_symbol.eisenstein_lemma_aux p ha1 ha0] + show _ = _, from eisenstein_lemma_aux p ha1 ha0] end -end legendre_symbol +end zmod end gauss_eisenstein diff --git a/src/number_theory/legendre_symbol/jacobi_symbol.lean b/src/number_theory/legendre_symbol/jacobi_symbol.lean index f1f5e9ac3c0c6..15ef4760319c1 100644 --- a/src/number_theory/legendre_symbol/jacobi_symbol.lean +++ b/src/number_theory/legendre_symbol/jacobi_symbol.lean @@ -14,7 +14,7 @@ We define the Jacobi symbol and prove its main properties. ## Main definitions We define the Jacobi symbol, `jacobi_sym a b`, for integers `a` and natural numbers `b` -as the product over the prime factors `p` of `b` of the Legendre symbols `zmod.legendre_sym p a`. +as the product over the prime factors `p` of `b` of the Legendre symbols `legendre_sym p a`. This agrees with the mathematical definition when `b` is odd. The prime factors are obtained via `nat.factors`. Since `nat.factors 0 = []`, @@ -24,28 +24,28 @@ this implies in particular that `jacobi_sym a 0 = 1` for all `a`. We prove the main properties of the Jacobi symbol, including the following. -* Multiplicativity in both arguments (`jacobi_sym_mul_left`, `jacobi_sym_mul_right`) +* Multiplicativity in both arguments (`jacobi_sym.mul_left`, `jacobi_sym.mul_right`) * The value of the symbol is `1` or `-1` when the arguments are coprime - (`jacobi_sym_eq_one_or_neg_one`) + (`jacobi_sym.eq_one_or_neg_one`) * The symbol vanishes if and only if `b ≠ 0` and the arguments are not coprime - (`jacobi_sym_eq_zero_iff`) + (`jacobi_sym.eq_zero_iff`) * If the symbol has the value `-1`, then `a : zmod b` is not a square - (`nonsquare_of_jacobi_sym_eq_neg_one`); the converse holds when `b = p` is a prime - (`nonsquare_iff_jacobi_sym_eq_neg_one`); in particular, in this case `a` is a - square mod `p` when the symbol has the value `1` (`is_square_of_jacobi_sym_eq_one`). + (`zmod.nonsquare_of_jacobi_sym_eq_neg_one`); the converse holds when `b = p` is a prime + (`zmod.nonsquare_iff_jacobi_sym_eq_neg_one`); in particular, in this case `a` is a + square mod `p` when the symbol has the value `1` (`zmod.is_square_of_jacobi_sym_eq_one`). -* Quadratic reciprocity (`jacobi_sym_quadratic_reciprocity`, - `jacobi_sym_quadratic_reciprocity_one_mod_four`, - `jacobi_sym_quadratic_reciprocity_three_mod_four`) +* Quadratic reciprocity (`jacobi_sym.quadratic_reciprocity`, + `jacobi_sym.quadratic_reciprocity_one_mod_four`, + `jacobi_sym.quadratic_reciprocity_three_mod_four`) -* The supplementary laws for `a = -1`, `a = 2`, `a = -2` (`jacobi_sym_neg_one`, `jacobi_sym_two`, - `jacobi_sym_neg_two`) +* The supplementary laws for `a = -1`, `a = 2`, `a = -2` (`jacobi_sym.at_neg_one`, + `jacobi_sym.at_two`, `jacobi_sym.at_neg_two`) -* The symbol depends on `a` only via its residue class mod `b` (`jacobi_sym_mod_left`) - and on `b` only via its residue class mod `4*a` (`jacobi_sym_mod_right`) +* The symbol depends on `a` only via its residue class mod `b` (`jacobi_sym.mod_left`) + and on `b` only via its residue class mod `4*a` (`jacobi_sym.mod_right`) ## Notations @@ -60,18 +60,17 @@ section jacobi /-! ### Definition of the Jacobi symbol -We define the Jacobi symbol $\Bigl\frac{a}{b}\Bigr$ for integers `a` and natural numbers `b` as the -product of the Legendre symbols $\Bigl\frac{a}{p}\Bigr$, where `p` runs through the prime divisors -(with multiplicity) of `b`, as provided by `b.factors`. This agrees with the Jacobi symbol -when `b` is odd and gives less meaningful values when it is not (e.g., the symbol is `1` -when `b = 0`). This is called `jacobi_sym a b`. +We define the Jacobi symbol $\Bigl(\frac{a}{b}\Bigr)$ for integers `a` and natural numbers `b` +as the product of the Legendre symbols $(\Bigl\frac{a}{p}\Bigr)$, where `p` runs through the +prime divisors (with multiplicity) of `b`, as provided by `b.factors`. This agrees with the +Jacobi symbol when `b` is odd and gives less meaningful values when it is not (e.g., the symbol +is `1` when `b = 0`). This is called `jacobi_sym a b`. We define localized notation (locale `number_theory_symbols`) `J(a | b)` for the Jacobi symbol `jacobi_sym a b`. (Unfortunately, there is no subscript "J" in unicode.) -/ -namespace zmod -open nat +open nat zmod /-- The Jacobi symbol of `a` and `b` -/ -- Since we need the fact that the factors are prime, we use `list.pmap`. @@ -84,23 +83,24 @@ localized "notation `J(` a ` | ` b `)` := jacobi_sym a b" in number_theory_symbo /-! ### Properties of the Jacobi symbol -/ +namespace jacobi_sym /-- The symbol `J(a | 0)` has the value `1`. -/ -@[simp] lemma jacobi_sym_zero_right (a : ℤ) : J(a | 0) = 1 := +@[simp] lemma zero_right (a : ℤ) : J(a | 0) = 1 := by simp only [jacobi_sym, factors_zero, list.prod_nil, list.pmap] /-- The symbol `J(a | 1)` has the value `1`. -/ -@[simp] lemma jacobi_sym_one_right (a : ℤ) : J(a | 1) = 1 := +@[simp] lemma one_right (a : ℤ) : J(a | 1) = 1 := by simp only [jacobi_sym, factors_one, list.prod_nil, list.pmap] /-- The Legendre symbol `legendre_sym p a` with an integer `a` and a prime number `p` is the same as the Jacobi symbol `J(a | p)`. -/ -lemma legendre_sym.to_jacobi_sym {p : ℕ} [fp : fact p.prime] {a : ℤ} : +lemma _root_.legendre_sym.to_jacobi_sym (p : ℕ) [fp : fact p.prime] (a : ℤ) : legendre_sym p a = J(a | p) := by simp only [jacobi_sym, factors_prime fp.1, list.prod_cons, list.prod_nil, mul_one, list.pmap] /-- The Jacobi symbol is multiplicative in its second argument. -/ -lemma jacobi_sym_mul_right' (a : ℤ) {b₁ b₂ : ℕ} (hb₁ : b₁ ≠ 0) (hb₂ : b₂ ≠ 0) : +lemma mul_right' (a : ℤ) {b₁ b₂ : ℕ} (hb₁ : b₁ ≠ 0) (hb₂ : b₂ ≠ 0) : J(a | b₁ * b₂) = J(a | b₁) * J(a | b₂) := begin rw [jacobi_sym, ((perm_factors_mul hb₁ hb₂).pmap _).prod_eq, list.pmap_append, list.prod_append], @@ -108,12 +108,12 @@ begin end /-- The Jacobi symbol is multiplicative in its second argument. -/ -lemma jacobi_sym_mul_right (a : ℤ) (b₁ b₂ : ℕ) [ne_zero b₁] [ne_zero b₂] : +lemma mul_right (a : ℤ) (b₁ b₂ : ℕ) [ne_zero b₁] [ne_zero b₂] : J(a | b₁ * b₂) = J(a | b₁) * J(a | b₂) := -jacobi_sym_mul_right' a (ne_zero.ne b₁) (ne_zero.ne b₂) +mul_right' a (ne_zero.ne b₁) (ne_zero.ne b₂) /-- The Jacobi symbol takes only the values `0`, `1` and `-1`. -/ -lemma jacobi_sym_trichotomy (a : ℤ) (b : ℕ) : J(a | b) = 0 ∨ J(a | b) = 1 ∨ J(a | b) = -1 := +lemma trichotomy (a : ℤ) (b : ℕ) : J(a | b) = 0 ∨ J(a | b) = 1 ∨ J(a | b) = -1 := ((@sign_type.cast_hom ℤ _ _).to_monoid_hom.mrange.copy {0, 1, -1} $ by {rw set.pair_comm, exact (sign_type.range_eq sign_type.cast_hom).symm}).list_prod_mem begin @@ -124,83 +124,89 @@ begin end /-- The symbol `J(1 | b)` has the value `1`. -/ -@[simp] lemma jacobi_sym_one_left (b : ℕ) : J(1 | b) = 1 := -list.prod_eq_one (λ z hz, let ⟨p, hp, he⟩ := list.mem_pmap.1 hz in by rw [← he, legendre_sym_one]) +@[simp] lemma one_left (b : ℕ) : J(1 | b) = 1 := +list.prod_eq_one (λ z hz, + let ⟨p, hp, he⟩ := list.mem_pmap.1 hz in by rw [← he, legendre_sym.at_one]) /-- The Jacobi symbol is multiplicative in its first argument. -/ -lemma jacobi_sym_mul_left (a₁ a₂ : ℤ) (b : ℕ) : J(a₁ * a₂ | b) = J(a₁ | b) * J(a₂ | b) := -by { simp_rw [jacobi_sym, list.pmap_eq_map_attach, legendre_sym_mul], exact list.prod_map_mul } +lemma mul_left (a₁ a₂ : ℤ) (b : ℕ) : J(a₁ * a₂ | b) = J(a₁ | b) * J(a₂ | b) := +by { simp_rw [jacobi_sym, list.pmap_eq_map_attach, legendre_sym.mul], exact list.prod_map_mul } /-- The symbol `J(a | b)` vanishes iff `a` and `b` are not coprime (assuming `b ≠ 0`). -/ -lemma jacobi_sym_eq_zero_iff_not_coprime {a : ℤ} {b : ℕ} [ne_zero b] : - J(a | b) = 0 ↔ a.gcd b ≠ 1 := +lemma eq_zero_iff_not_coprime {a : ℤ} {b : ℕ} [ne_zero b] : J(a | b) = 0 ↔ a.gcd b ≠ 1 := list.prod_eq_zero_iff.trans begin rw [list.mem_pmap, int.gcd_eq_nat_abs, ne, prime.not_coprime_iff_dvd], - simp_rw [legendre_sym_eq_zero_iff, int_coe_zmod_eq_zero_iff_dvd, mem_factors (ne_zero.ne b), + simp_rw [legendre_sym.eq_zero_iff, int_coe_zmod_eq_zero_iff_dvd, mem_factors (ne_zero.ne b), ← int.coe_nat_dvd_left, int.coe_nat_dvd, exists_prop, and_assoc, and_comm], end /-- The symbol `J(a | b)` is nonzero when `a` and `b` are coprime. -/ -lemma jacobi_sym_ne_zero {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a | b) ≠ 0 := +protected +lemma ne_zero {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a | b) ≠ 0 := begin casesI eq_zero_or_ne_zero b with hb, - { rw [hb, jacobi_sym_zero_right], + { rw [hb, zero_right], exact one_ne_zero }, - { contrapose! h, exact jacobi_sym_eq_zero_iff_not_coprime.1 h }, + { contrapose! h, exact eq_zero_iff_not_coprime.1 h }, end /-- The symbol `J(a | b)` vanishes if and only if `b ≠ 0` and `a` and `b` are not coprime. -/ -lemma jacobi_sym_eq_zero_iff {a : ℤ} {b : ℕ} : J(a | b) = 0 ↔ b ≠ 0 ∧ a.gcd b ≠ 1 := +lemma eq_zero_iff {a : ℤ} {b : ℕ} : J(a | b) = 0 ↔ b ≠ 0 ∧ a.gcd b ≠ 1 := ⟨λ h, begin casesI eq_or_ne b 0 with hb hb, - { rw [hb, jacobi_sym_zero_right] at h, cases h }, - exact ⟨hb, mt jacobi_sym_ne_zero $ not_not.2 h⟩, -end, λ ⟨hb, h⟩, by { rw ← ne_zero_iff at hb, exactI jacobi_sym_eq_zero_iff_not_coprime.2 h }⟩ + { rw [hb, zero_right] at h, cases h }, + exact ⟨hb, mt jacobi_sym.ne_zero $ not_not.2 h⟩, +end, λ ⟨hb, h⟩, by { rw ← ne_zero_iff at hb, exactI eq_zero_iff_not_coprime.2 h }⟩ /-- The symbol `J(0 | b)` vanishes when `b > 1`. -/ -lemma jacobi_sym_zero_left {b : ℕ} (hb : 1 < b) : J(0 | b) = 0 := -(@jacobi_sym_eq_zero_iff_not_coprime 0 b ⟨ne_zero_of_lt hb⟩).mpr $ +lemma zero_left {b : ℕ} (hb : 1 < b) : J(0 | b) = 0 := +(@eq_zero_iff_not_coprime 0 b ⟨ne_zero_of_lt hb⟩).mpr $ by { rw [int.gcd_zero_left, int.nat_abs_of_nat], exact hb.ne' } /-- The symbol `J(a | b)` takes the value `1` or `-1` if `a` and `b` are coprime. -/ -lemma jacobi_sym_eq_one_or_neg_one {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : - J(a | b) = 1 ∨ J(a | b) = -1 := -(jacobi_sym_trichotomy a b).resolve_left $ jacobi_sym_ne_zero h +lemma eq_one_or_neg_one {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a | b) = 1 ∨ J(a | b) = -1 := +(trichotomy a b).resolve_left $ jacobi_sym.ne_zero h /-- We have that `J(a^e | b) = J(a | b)^e`. -/ -lemma jacobi_sym_pow_left (a : ℤ) (e b : ℕ) : J(a ^ e | b) = J(a | b) ^ e := -nat.rec_on e (by rw [pow_zero, pow_zero, jacobi_sym_one_left]) $ - λ _ ih, by rw [pow_succ, pow_succ, jacobi_sym_mul_left, ih] +lemma pow_left (a : ℤ) (e b : ℕ) : J(a ^ e | b) = J(a | b) ^ e := +nat.rec_on e (by rw [pow_zero, pow_zero, one_left]) $ + λ _ ih, by rw [pow_succ, pow_succ, mul_left, ih] /-- We have that `J(a | b^e) = J(a | b)^e`. -/ -lemma jacobi_sym_pow_right (a : ℤ) (b e : ℕ) : J(a | b ^ e) = J(a | b) ^ e := +lemma pow_right (a : ℤ) (b e : ℕ) : J(a | b ^ e) = J(a | b) ^ e := begin induction e with e ih, - { rw [pow_zero, pow_zero, jacobi_sym_one_right], }, + { rw [pow_zero, pow_zero, one_right], }, { casesI eq_zero_or_ne_zero b with hb, - { rw [hb, zero_pow (succ_pos e), jacobi_sym_zero_right, one_pow], }, - { rw [pow_succ, pow_succ, jacobi_sym_mul_right, ih], } } + { rw [hb, zero_pow (succ_pos e), zero_right, one_pow], }, + { rw [pow_succ, pow_succ, mul_right, ih], } } end /-- The square of `J(a | b)` is `1` when `a` and `b` are coprime. -/ -lemma jacobi_sym_sq_one {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a | b) ^ 2 = 1 := -by cases jacobi_sym_eq_one_or_neg_one h with h₁ h₁; rw h₁; refl +lemma sq_one {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a | b) ^ 2 = 1 := +by cases eq_one_or_neg_one h with h₁ h₁; rw h₁; refl /-- The symbol `J(a^2 | b)` is `1` when `a` and `b` are coprime. -/ -lemma jacobi_sym_sq_one' {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a ^ 2 | b) = 1 := -by rw [jacobi_sym_pow_left, jacobi_sym_sq_one h] +lemma sq_one' {a : ℤ} {b : ℕ} (h : a.gcd b = 1) : J(a ^ 2 | b) = 1 := +by rw [pow_left, sq_one h] /-- The symbol `J(a | b)` depends only on `a` mod `b`. -/ -lemma jacobi_sym_mod_left (a : ℤ) (b : ℕ) : J(a | b) = J(a % b | b) := +lemma mod_left (a : ℤ) (b : ℕ) : J(a | b) = J(a % b | b) := congr_arg list.prod $ list.pmap_congr _ begin rintro p hp _ _, - conv_rhs { rw [legendre_sym_mod, int.mod_mod_of_dvd _ - (int.coe_nat_dvd.2 $ dvd_of_mem_factors hp), ← legendre_sym_mod] }, + conv_rhs { rw [legendre_sym.mod, int.mod_mod_of_dvd _ + (int.coe_nat_dvd.2 $ dvd_of_mem_factors hp), ← legendre_sym.mod] }, end /-- The symbol `J(a | b)` depends only on `a` mod `b`. -/ -lemma jacobi_sym_mod_left' {a₁ a₂ : ℤ} {b : ℕ} (h : a₁ % b = a₂ % b) : J(a₁ | b) = J(a₂ | b) := -by rw [jacobi_sym_mod_left, h, ← jacobi_sym_mod_left] +lemma mod_left' {a₁ a₂ : ℤ} {b : ℕ} (h : a₁ % b = a₂ % b) : J(a₁ | b) = J(a₂ | b) := +by rw [mod_left, h, ← mod_left] + +end jacobi_sym + +namespace zmod + +open jacobi_sym /-- If `J(a | b)` is `-1`, then `a` is not a square modulo `b`. -/ lemma nonsquare_of_jacobi_sym_eq_neg_one {a : ℤ} {b : ℕ} (h : J(a | b) = -1) : @@ -208,14 +214,14 @@ lemma nonsquare_of_jacobi_sym_eq_neg_one {a : ℤ} {b : ℕ} (h : J(a | b) = -1) λ ⟨r, ha⟩, begin rw [← r.coe_val_min_abs, ← int.cast_mul, int_coe_eq_int_coe_iff', ← sq] at ha, apply (by norm_num : ¬ (0 : ℤ) ≤ -1), - rw [← h, jacobi_sym_mod_left, ha, ← jacobi_sym_mod_left, jacobi_sym_pow_left], + rw [← h, mod_left, ha, ← mod_left, pow_left], apply sq_nonneg, end /-- If `p` is prime, then `J(a | p)` is `-1` iff `a` is not a square modulo `p`. -/ lemma nonsquare_iff_jacobi_sym_eq_neg_one {a : ℤ} {p : ℕ} [fact p.prime] : J(a | p) = -1 ↔ ¬ is_square (a : zmod p) := -by { rw [← legendre_sym.to_jacobi_sym], exact legendre_sym_eq_neg_one_iff p } +by { rw [← legendre_sym.to_jacobi_sym], exact legendre_sym.eq_neg_one_iff p } /-- If `p` is prime and `J(a | p) = 1`, then `a` is q square mod `p`. -/ lemma is_square_of_jacobi_sym_eq_one {a : ℤ} {p : ℕ} [fact p.prime] (h : J(a | p) = 1) : @@ -223,13 +229,17 @@ lemma is_square_of_jacobi_sym_eq_one {a : ℤ} {p : ℕ} [fact p.prime] (h : J(a not_not.mp $ mt nonsquare_iff_jacobi_sym_eq_neg_one.mpr $ λ hf, one_ne_zero $ neg_eq_self_iff.mp $ hf.symm.trans h +end zmod + /-! ### Values at `-1`, `2` and `-2` -/ +namespace jacobi_sym + /-- If `χ` is a multiplicative function such that `J(a | p) = χ p` for all odd primes `p`, then `J(a | b)` equals `χ b` for all odd natural numbers `b`. -/ -lemma jacobi_sym_value (a : ℤ) {R : Type*} [comm_semiring R] (χ : R →* ℤ) +lemma value_at (a : ℤ) {R : Type*} [comm_semiring R] (χ : R →* ℤ) (hp : ∀ (p : ℕ) (pp : p.prime) (h2 : p ≠ 2), @legendre_sym p ⟨pp⟩ a = χ p) {b : ℕ} (hb : odd b) : J(a | b) = χ b := begin @@ -240,21 +250,23 @@ begin end /-- If `b` is odd, then `J(-1 | b)` is given by `χ₄ b`. -/ -lemma jacobi_sym_neg_one {b : ℕ} (hb : odd b) : J(-1 | b) = χ₄ b := -jacobi_sym_value (-1) χ₄ (λ p pp, @legendre_sym_neg_one p ⟨pp⟩) hb +lemma at_neg_one {b : ℕ} (hb : odd b) : J(-1 | b) = χ₄ b := +value_at (-1) χ₄ (λ p pp, @legendre_sym.at_neg_one p ⟨pp⟩) hb /-- If `b` is odd, then `J(-a | b) = χ₄ b * J(a | b)`. -/ -lemma jacobi_sym_neg (a : ℤ) {b : ℕ} (hb : odd b) : J(-a | b) = χ₄ b * J(a | b) := -by rw [neg_eq_neg_one_mul, jacobi_sym_mul_left, jacobi_sym_neg_one hb] +protected +lemma neg (a : ℤ) {b : ℕ} (hb : odd b) : J(-a | b) = χ₄ b * J(a | b) := +by rw [neg_eq_neg_one_mul, mul_left, at_neg_one hb] /-- If `b` is odd, then `J(2 | b)` is given by `χ₈ b`. -/ -lemma jacobi_sym_two {b : ℕ} (hb : odd b) : J(2 | b) = χ₈ b := -jacobi_sym_value 2 χ₈ (λ p pp, @legendre_sym_two p ⟨pp⟩) hb +lemma at_two {b : ℕ} (hb : odd b) : J(2 | b) = χ₈ b := +value_at 2 χ₈ (λ p pp, @legendre_sym.at_two p ⟨pp⟩) hb /-- If `b` is odd, then `J(-2 | b)` is given by `χ₈' b`. -/ -lemma jacobi_sym_neg_two {b : ℕ} (hb : odd b) : J(-2 | b) = χ₈' b := -jacobi_sym_value (-2) χ₈' (λ p pp, @legendre_sym_neg_two p ⟨pp⟩) hb +lemma at_neg_two {b : ℕ} (hb : odd b) : J(-2 | b) = χ₈' b := +value_at (-2) χ₈' (λ p pp, @legendre_sym.at_neg_two p ⟨pp⟩) hb +end jacobi_sym /-! ### Quadratic Reciprocity @@ -263,90 +275,96 @@ jacobi_sym_value (-2) χ₈' (λ p pp, @legendre_sym_neg_two p ⟨pp⟩) hb /-- The bi-multiplicative map giving the sign in the Law of Quadratic Reciprocity -/ def qr_sign (m n : ℕ) : ℤ := J(χ₄ m | n) +namespace qr_sign + /-- We can express `qr_sign m n` as a power of `-1` when `m` and `n` are odd. -/ -lemma qr_sign_neg_one_pow {m n : ℕ} (hm : odd m) (hn : odd n) : +lemma neg_one_pow {m n : ℕ} (hm : odd m) (hn : odd n) : qr_sign m n = (-1) ^ ((m / 2) * (n / 2)) := begin rw [qr_sign, pow_mul, ← χ₄_eq_neg_one_pow (odd_iff.mp hm)], cases odd_mod_four_iff.mp (odd_iff.mp hm) with h h, - { rw [χ₄_nat_one_mod_four h, jacobi_sym_one_left, one_pow], }, - { rw [χ₄_nat_three_mod_four h, ← χ₄_eq_neg_one_pow (odd_iff.mp hn), jacobi_sym_neg_one hn], } + { rw [χ₄_nat_one_mod_four h, jacobi_sym.one_left, one_pow], }, + { rw [χ₄_nat_three_mod_four h, ← χ₄_eq_neg_one_pow (odd_iff.mp hn), jacobi_sym.at_neg_one hn], } end /-- When `m` and `n` are odd, then the square of `qr_sign m n` is `1`. -/ -lemma qr_sign_sq_eq_one {m n : ℕ} (hm : odd m) (hn : odd n) : (qr_sign m n) ^ 2 = 1 := -by rw [qr_sign_neg_one_pow hm hn, ← pow_mul, mul_comm, pow_mul, neg_one_sq, one_pow] +lemma sq_eq_one {m n : ℕ} (hm : odd m) (hn : odd n) : (qr_sign m n) ^ 2 = 1 := +by rw [neg_one_pow hm hn, ← pow_mul, mul_comm, pow_mul, neg_one_sq, one_pow] /-- `qr_sign` is multiplicative in the first argument. -/ -lemma qr_sign_mul_left (m₁ m₂ n : ℕ) : qr_sign (m₁ * m₂) n = qr_sign m₁ n * qr_sign m₂ n := -by simp_rw [qr_sign, nat.cast_mul, map_mul, jacobi_sym_mul_left] +lemma mul_left (m₁ m₂ n : ℕ) : qr_sign (m₁ * m₂) n = qr_sign m₁ n * qr_sign m₂ n := +by simp_rw [qr_sign, nat.cast_mul, map_mul, jacobi_sym.mul_left] /-- `qr_sign` is multiplicative in the second argument. -/ -lemma qr_sign_mul_right (m n₁ n₂ : ℕ) [ne_zero n₁] [ne_zero n₂] : +lemma mul_right (m n₁ n₂ : ℕ) [ne_zero n₁] [ne_zero n₂] : qr_sign m (n₁ * n₂) = qr_sign m n₁ * qr_sign m n₂ := -jacobi_sym_mul_right (χ₄ m) n₁ n₂ +jacobi_sym.mul_right (χ₄ m) n₁ n₂ /-- `qr_sign` is symmetric when both arguments are odd. -/ -lemma qr_sign_symm {m n : ℕ} (hm : odd m) (hn : odd n) : qr_sign m n = qr_sign n m := -by rw [qr_sign_neg_one_pow hm hn, qr_sign_neg_one_pow hn hm, mul_comm (m / 2)] +protected +lemma symm {m n : ℕ} (hm : odd m) (hn : odd n) : qr_sign m n = qr_sign n m := +by rw [neg_one_pow hm hn, neg_one_pow hn hm, mul_comm (m / 2)] /-- We can move `qr_sign m n` from one side of an equality to the other when `m` and `n` are odd. -/ -lemma qr_sign_eq_iff_eq {m n : ℕ} (hm : odd m) (hn : odd n) (x y : ℤ) : +lemma eq_iff_eq {m n : ℕ} (hm : odd m) (hn : odd n) (x y : ℤ) : qr_sign m n * x = y ↔ x = qr_sign m n * y := by refine ⟨λ h', let h := h'.symm in _, λ h, _⟩; - rw [h, ← mul_assoc, ← pow_two, qr_sign_sq_eq_one hm hn, one_mul] + rw [h, ← mul_assoc, ← pow_two, sq_eq_one hm hn, one_mul] + +end qr_sign + +namespace jacobi_sym /-- The Law of Quadratic Reciprocity for the Jacobi symbol, version with `qr_sign` -/ -lemma jacobi_sym_quadratic_reciprocity' {a b : ℕ} (ha : odd a) (hb : odd b) : +lemma quadratic_reciprocity' {a b : ℕ} (ha : odd a) (hb : odd b) : J(a | b) = qr_sign b a * J(b | a) := begin -- define the right hand side for fixed `a` as a `ℕ →* ℤ` let rhs : ℕ → ℕ →* ℤ := λ a, { to_fun := λ x, qr_sign x a * J(x | a), - map_one' := by { convert ← mul_one _, symmetry, all_goals { apply jacobi_sym_one_left } }, - map_mul' := λ x y, by rw [qr_sign_mul_left, nat.cast_mul, jacobi_sym_mul_left, + map_one' := by { convert ← mul_one _, symmetry, all_goals { apply one_left } }, + map_mul' := λ x y, by rw [qr_sign.mul_left, nat.cast_mul, mul_left, mul_mul_mul_comm] }, have rhs_apply : ∀ (a b : ℕ), rhs a b = qr_sign b a * J(b | a) := λ a b, rfl, - refine jacobi_sym_value a (rhs a) (λ p pp hp, eq.symm _) hb, + refine value_at a (rhs a) (λ p pp hp, eq.symm _) hb, have hpo := pp.eq_two_or_odd'.resolve_left hp, rw [@legendre_sym.to_jacobi_sym p ⟨pp⟩, rhs_apply, nat.cast_id, - qr_sign_eq_iff_eq hpo ha, qr_sign_symm hpo ha], - refine jacobi_sym_value p (rhs p) (λ q pq hq, _) ha, + qr_sign.eq_iff_eq hpo ha, qr_sign.symm hpo ha], + refine value_at p (rhs p) (λ q pq hq, _) ha, have hqo := pq.eq_two_or_odd'.resolve_left hq, - rw [rhs_apply, nat.cast_id, ← @legendre_sym.to_jacobi_sym p ⟨pp⟩, qr_sign_symm hqo hpo, - qr_sign_neg_one_pow hpo hqo, @quadratic_reciprocity' p q ⟨pp⟩ ⟨pq⟩ hp hq], + rw [rhs_apply, nat.cast_id, ← @legendre_sym.to_jacobi_sym p ⟨pp⟩, qr_sign.symm hqo hpo, + qr_sign.neg_one_pow hpo hqo, @legendre_sym.quadratic_reciprocity' p q ⟨pp⟩ ⟨pq⟩ hp hq], end /-- The Law of Quadratic Reciprocity for the Jacobi symbol -/ -lemma jacobi_sym_quadratic_reciprocity {a b : ℕ} (ha : odd a) (hb : odd b) : +lemma quadratic_reciprocity {a b : ℕ} (ha : odd a) (hb : odd b) : J(a | b) = (-1) ^ ((a / 2) * (b / 2)) * J(b | a) := -by rw [← qr_sign_neg_one_pow ha hb, qr_sign_symm ha hb, jacobi_sym_quadratic_reciprocity' ha hb] +by rw [← qr_sign.neg_one_pow ha hb, qr_sign.symm ha hb, quadratic_reciprocity' ha hb] /-- The Law of Quadratic Reciprocity for the Jacobi symbol: if `a` and `b` are natural numbers with `a % 4 = 1` and `b` odd, then `J(a | b) = J(b | a)`. -/ -theorem jacobi_sym_quadratic_reciprocity_one_mod_four {a b : ℕ} (ha : a % 4 = 1) (hb : odd b) : +theorem quadratic_reciprocity_one_mod_four {a b : ℕ} (ha : a % 4 = 1) (hb : odd b) : J(a | b) = J(b | a) := -by rw [jacobi_sym_quadratic_reciprocity (odd_iff.mpr (odd_of_mod_four_eq_one ha)) hb, +by rw [quadratic_reciprocity (odd_iff.mpr (odd_of_mod_four_eq_one ha)) hb, pow_mul, neg_one_pow_div_two_of_one_mod_four ha, one_pow, one_mul] /-- The Law of Quadratic Reciprocity for the Jacobi symbol: if `a` and `b` are natural numbers with `a` odd and `b % 4 = 1`, then `J(a | b) = J(b | a)`. -/ -theorem jacobi_sym_quadratic_reciprocity_one_mod_four' {a b : ℕ} (ha : odd a) (hb : b % 4 = 1) : +theorem quadratic_reciprocity_one_mod_four' {a b : ℕ} (ha : odd a) (hb : b % 4 = 1) : J(a | b) = J(b | a) := -(jacobi_sym_quadratic_reciprocity_one_mod_four hb ha).symm +(quadratic_reciprocity_one_mod_four hb ha).symm /-- The Law of Quadratic Reciprocityfor the Jacobi symbol: if `a` and `b` are natural numbers both congruent to `3` mod `4`, then `J(a | b) = -J(b | a)`. -/ -theorem jacobi_sym_quadratic_reciprocity_three_mod_four - {a b : ℕ} (ha : a % 4 = 3) (hb : b % 4 = 3) : +theorem quadratic_reciprocity_three_mod_four {a b : ℕ} (ha : a % 4 = 3) (hb : b % 4 = 3) : J(a | b) = - J(b | a) := let nop := @neg_one_pow_div_two_of_three_mod_four in begin - rw [jacobi_sym_quadratic_reciprocity, pow_mul, nop ha, nop hb, neg_one_mul]; + rw [quadratic_reciprocity, pow_mul, nop ha, nop hb, neg_one_mul]; rwa [odd_iff, odd_of_mod_four_eq_three], end /-- The Jacobi symbol `J(a | b)` depends only on `b` mod `4*a` (version for `a : ℕ`). -/ -lemma jacobi_sym_mod_right' (a : ℕ) {b : ℕ} (hb : odd b) : J(a | b) = J(a | b % (4 * a)) := +lemma mod_right' (a : ℕ) {b : ℕ} (hb : odd b) : J(a | b) = J(a | b % (4 * a)) := begin rcases eq_or_ne a 0 with rfl | ha₀, { rw [mul_zero, mod_zero], }, @@ -354,15 +372,13 @@ begin rcases exists_eq_pow_mul_and_not_dvd ha₀ 2 (by norm_num) with ⟨e, a', ha₁', ha₂⟩, have ha₁ := odd_iff.mpr (two_dvd_ne_zero.mp ha₁'), nth_rewrite 1 [ha₂], nth_rewrite 0 [ha₂], - rw [nat.cast_mul, jacobi_sym_mul_left, jacobi_sym_mul_left, - jacobi_sym_quadratic_reciprocity' ha₁ hb, jacobi_sym_quadratic_reciprocity' ha₁ hb', - nat.cast_pow, jacobi_sym_pow_left, jacobi_sym_pow_left, - (show ((2 : ℕ) : ℤ) = 2, by refl), jacobi_sym_two hb, jacobi_sym_two hb'], + rw [nat.cast_mul, mul_left, mul_left, quadratic_reciprocity' ha₁ hb, + quadratic_reciprocity' ha₁ hb', nat.cast_pow, pow_left, pow_left, + nat.cast_two, at_two hb, at_two hb'], congr' 1, swap, congr' 1, { simp_rw [qr_sign], rw [χ₄_nat_mod_four, χ₄_nat_mod_four (b % (4 * a)), mod_mod_of_dvd b (dvd_mul_right 4 a) ] }, - { rw [jacobi_sym_mod_left ↑(b % _), jacobi_sym_mod_left b, int.coe_nat_mod, - int.mod_mod_of_dvd b], + { rw [mod_left ↑(b % _), mod_left b, int.coe_nat_mod, int.mod_mod_of_dvd b], simp only [ha₂, nat.cast_mul, ← mul_assoc], exact dvd_mul_left a' _, }, cases e, { refl }, @@ -371,15 +387,15 @@ begin end /-- The Jacobi symbol `J(a | b)` depends only on `b` mod `4*a`. -/ -lemma jacobi_sym_mod_right (a : ℤ) {b : ℕ} (hb : odd b) : J(a | b) = J(a | b % (4 * a.nat_abs)) := +lemma mod_right (a : ℤ) {b : ℕ} (hb : odd b) : J(a | b) = J(a | b % (4 * a.nat_abs)) := begin cases int.nat_abs_eq a with ha ha; nth_rewrite 1 [ha]; nth_rewrite 0 [ha], - { exact jacobi_sym_mod_right' a.nat_abs hb, }, + { exact mod_right' a.nat_abs hb, }, { have hb' : odd (b % (4 * a.nat_abs)) := hb.mod_even (even.mul_right (by norm_num) _), - rw [jacobi_sym_neg _ hb, jacobi_sym_neg _ hb', jacobi_sym_mod_right' _ hb, χ₄_nat_mod_four, + rw [jacobi_sym.neg _ hb, jacobi_sym.neg _ hb', mod_right' _ hb, χ₄_nat_mod_four, χ₄_nat_mod_four (b % (4 * _)), mod_mod_of_dvd b (dvd_mul_right 4 _)], } end -end zmod +end jacobi_sym end jacobi diff --git a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean index 696f9243aa41a..93215318f6248 100644 --- a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean +++ b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean @@ -10,23 +10,26 @@ import number_theory.legendre_symbol.quadratic_char This file contains results about quadratic residues modulo a prime number. -We define the Legendre symbol `(a / p)` as `zmod.legendre_sym p a`. -Note the order of arguments! The advantage of this form is that then `zmod.legendre_sym p` +We define the Legendre symbol $\Bigl(\frac{a}{p}\Bigr)$ as `legendre_sym p a`. +Note the order of arguments! The advantage of this form is that then `legendre_sym p` is a multiplicative map. +The Legendre symbol is used to define the Jacobi symbol, `jacobi_sym a b`, for integers `a` +and (odd) natural numbers `b`, which extends the Legendre symbol. + ## Main results -We prove the law of quadratic reciprocity, see `zmod.quadratic_reciprocity` and -`zmod.quadratic_reciprocity'`, as well as the +We prove the law of quadratic reciprocity, see `legendre_sym.quadratic_reciprocity` and +`legendre_sym.quadratic_reciprocity'`, as well as the interpretations in terms of existence of square roots depending on the congruence mod 4, `zmod.exists_sq_eq_prime_iff_of_mod_four_eq_one`, and `zmod.exists_sq_eq_prime_iff_of_mod_four_eq_three`. We also prove the supplementary laws that give conditions for when `-1` or `2` (or `-2`) is a square modulo a prime `p`: -`zmod.legende_sym_neg_one` and `zmod.exists_sq_eq_neg_one_iff` for `-1`, -`zmod.legendre_sym_two` and `zmod.exists_sq_eq_two_iff` for `2`, -`zmod.legendre_sym_neg_two` and `zmod.exists_sq_eq_neg_two_iff` for `-2`. +`legendre_sym.at_neg_one` and `zmod.exists_sq_eq_neg_one_iff` for `-1`, +`legendre_sym.at_two` and `zmod.exists_sq_eq_two_iff` for `2`, +`legendre_sym.at_neg_two` and `zmod.exists_sq_eq_neg_two_iff` for `-2`. ## Implementation notes @@ -41,10 +44,10 @@ quadratic residue, quadratic nonresidue, Legendre symbol, quadratic reciprocity open nat -namespace zmod - section euler +namespace zmod + variables (p : ℕ) [fact p.prime] /-- Euler's Criterion: A unit `x` of `zmod p` is a square if and only if `x ^ (p / 2) = 1`. -/ @@ -82,6 +85,8 @@ begin exact pow_card_sub_one_eq_one ha end +end zmod + end euler section legendre @@ -90,6 +95,8 @@ section legendre ### Definition of the Legendre symbol and basic properties -/ +open zmod + variables (p : ℕ) [fact p.prime] /-- The Legendre symbol of `a : ℤ` and a prime `p`, `legendre_sym p a`, @@ -104,8 +111,10 @@ that `legendre_sym p` is a multiplicative function `ℤ → ℤ`. -/ def legendre_sym (a : ℤ) : ℤ := quadratic_char (zmod p) a +namespace legendre_sym + /-- We have the congruence `legendre_sym p a ≡ a ^ (p / 2) mod p`. -/ -lemma legendre_sym_eq_pow (a : ℤ) : (legendre_sym p a : zmod p) = a ^ (p / 2) := +lemma eq_pow (a : ℤ) : (legendre_sym p a : zmod p) = a ^ (p / 2) := begin cases eq_or_ne (ring_char (zmod p)) 2 with hc hc, { by_cases ha : (a : zmod p) = 0, @@ -122,83 +131,93 @@ begin end /-- If `p ∤ a`, then `legendre_sym p a` is `1` or `-1`. -/ -lemma legendre_sym_eq_one_or_neg_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : +lemma eq_one_or_neg_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : legendre_sym p a = 1 ∨ legendre_sym p a = -1 := quadratic_char_dichotomy ha -lemma legendre_sym_eq_neg_one_iff_not_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : +lemma eq_neg_one_iff_not_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : legendre_sym p a = -1 ↔ ¬ legendre_sym p a = 1 := quadratic_char_eq_neg_one_iff_not_one ha /-- The Legendre symbol of `p` and `a` is zero iff `p ∣ a`. -/ -lemma legendre_sym_eq_zero_iff (a : ℤ) : legendre_sym p a = 0 ↔ (a : zmod p) = 0 := +lemma eq_zero_iff (a : ℤ) : legendre_sym p a = 0 ↔ (a : zmod p) = 0 := quadratic_char_eq_zero_iff -@[simp] lemma legendre_sym_zero : legendre_sym p 0 = 0 := +@[simp] lemma at_zero : legendre_sym p 0 = 0 := by rw [legendre_sym, int.cast_zero, mul_char.map_zero] -@[simp] lemma legendre_sym_one : legendre_sym p 1 = 1 := +@[simp] lemma at_one : legendre_sym p 1 = 1 := by rw [legendre_sym, int.cast_one, mul_char.map_one] /-- The Legendre symbol is multiplicative in `a` for `p` fixed. -/ -lemma legendre_sym_mul (a b : ℤ) : legendre_sym p (a * b) = legendre_sym p a * legendre_sym p b := +protected +lemma mul (a b : ℤ) : legendre_sym p (a * b) = legendre_sym p a * legendre_sym p b := by simp only [legendre_sym, int.cast_mul, map_mul] /-- The Legendre symbol is a homomorphism of monoids with zero. -/ -@[simps] def legendre_sym_hom : ℤ →*₀ ℤ := +@[simps] def hom : ℤ →*₀ ℤ := { to_fun := legendre_sym p, - map_zero' := legendre_sym_zero p, - map_one' := legendre_sym_one p, - map_mul' := legendre_sym_mul p } + map_zero' := at_zero p, + map_one' := at_one p, + map_mul' := legendre_sym.mul p } /-- The square of the symbol is 1 if `p ∤ a`. -/ -theorem legendre_sym_sq_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : (legendre_sym p a) ^ 2 = 1 := +theorem sq_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : (legendre_sym p a) ^ 2 = 1 := quadratic_char_sq_one ha /-- The Legendre symbol of `a^2` at `p` is 1 if `p ∤ a`. -/ -theorem legendre_sym_sq_one' {a : ℤ} (ha : (a : zmod p) ≠ 0) : legendre_sym p (a ^ 2) = 1 := +theorem sq_one' {a : ℤ} (ha : (a : zmod p) ≠ 0) : legendre_sym p (a ^ 2) = 1 := by exact_mod_cast quadratic_char_sq_one' ha /-- The Legendre symbol depends only on `a` mod `p`. -/ -theorem legendre_sym_mod (a : ℤ) : legendre_sym p a = legendre_sym p (a % p) := +protected +theorem mod (a : ℤ) : legendre_sym p a = legendre_sym p (a % p) := by simp only [legendre_sym, int_cast_mod] /-- When `p ∤ a`, then `legendre_sym p a = 1` iff `a` is a square mod `p`. -/ -lemma legendre_sym_eq_one_iff {a : ℤ} (ha0 : (a : zmod p) ≠ 0) : +lemma eq_one_iff {a : ℤ} (ha0 : (a : zmod p) ≠ 0) : legendre_sym p a = 1 ↔ is_square (a : zmod p) := quadratic_char_one_iff_is_square ha0 -lemma legendre_sym_eq_one_iff' {a : ℕ} (ha0 : (a : zmod p) ≠ 0) : +lemma eq_one_iff' {a : ℕ} (ha0 : (a : zmod p) ≠ 0) : legendre_sym p a = 1 ↔ is_square (a : zmod p) := -by {rw legendre_sym_eq_one_iff, norm_cast, exact_mod_cast ha0} +by {rw eq_one_iff, norm_cast, exact_mod_cast ha0} /-- `legendre_sym p a = -1` iff `a` is a nonsquare mod `p`. -/ -lemma legendre_sym_eq_neg_one_iff {a : ℤ} : legendre_sym p a = -1 ↔ ¬ is_square (a : zmod p) := +lemma eq_neg_one_iff {a : ℤ} : legendre_sym p a = -1 ↔ ¬ is_square (a : zmod p) := quadratic_char_neg_one_iff_not_is_square -lemma legendre_sym_eq_neg_one_iff' {a : ℕ} : legendre_sym p a = -1 ↔ ¬ is_square (a : zmod p) := -by {rw legendre_sym_eq_neg_one_iff, norm_cast} +lemma eq_neg_one_iff' {a : ℕ} : legendre_sym p a = -1 ↔ ¬ is_square (a : zmod p) := +by {rw eq_neg_one_iff, norm_cast} /-- The number of square roots of `a` modulo `p` is determined by the Legendre symbol. -/ -lemma legendre_sym_card_sqrts (hp : p ≠ 2) (a : ℤ) : +lemma card_sqrts (hp : p ≠ 2) (a : ℤ) : ↑{x : zmod p | x^2 = a}.to_finset.card = legendre_sym p a + 1 := quadratic_char_card_sqrts ((ring_char_zmod_n p).substr hp) a +end legendre_sym + end legendre section values /-! ### The value of the Legendre symbol at `-1` + +See `jacobi_sym.at_neg_one` for the corresponding statement for the Jacobi symbol. -/ variables {p : ℕ} [fact p.prime] +open zmod + /-- `legendre_sym p (-1)` is given by `χ₄ p`. -/ -lemma legendre_sym_neg_one (hp : p ≠ 2) : legendre_sym p (-1) = χ₄ p := +lemma legendre_sym.at_neg_one (hp : p ≠ 2) : legendre_sym p (-1) = χ₄ p := by simp only [legendre_sym, card p, quadratic_char_neg_one ((ring_char_zmod_n p).substr hp), int.cast_neg, int.cast_one] +namespace zmod + /-- `-1` is a square in `zmod p` iff `p` is not congruent to `3` mod `4`. -/ lemma exists_sq_eq_neg_one_iff : is_square (-1 : zmod p) ↔ p % 4 ≠ 3 := by rw [finite_field.is_square_neg_one_iff, card p] @@ -218,18 +237,37 @@ lemma mod_four_ne_three_of_sq_eq_neg_sq {x y : zmod p} (hx : x ≠ 0) (hxy : x ^ p % 4 ≠ 3 := mod_four_ne_three_of_sq_eq_neg_sq' hx (eq_neg_iff_eq_neg.1 hxy) +end zmod + /-! ### The value of the Legendre symbol at `2` and `-2` + +See `jacobi_sym.at_two` and `jacobi_sym.at_neg_two` for the corresponding statements +for the Jacobi symbol. -/ +namespace legendre_sym + variables (hp : p ≠ 2) include hp /-- `legendre_sym p 2` is given by `χ₈ p`. -/ -lemma legendre_sym_two : legendre_sym p 2 = χ₈ p := +lemma at_two : legendre_sym p 2 = χ₈ p := by simp only [legendre_sym, card p, quadratic_char_two ((ring_char_zmod_n p).substr hp), int.cast_bit0, int.cast_one] +/-- `legendre_sym p (-2)` is given by `χ₈' p`. -/ +lemma at_neg_two : legendre_sym p (-2) = χ₈' p := +by simp only [legendre_sym, card p, quadratic_char_neg_two ((ring_char_zmod_n p).substr hp), + int.cast_bit0, int.cast_one, int.cast_neg] + +end legendre_sym + +namespace zmod + +variables (hp : p ≠ 2) +include hp + /-- `2` is a square modulo an odd prime `p` iff `p` is congruent to `1` or `7` mod `8`. -/ lemma exists_sq_eq_two_iff : is_square (2 : zmod p) ↔ p % 8 = 1 ∨ p % 8 = 7 := begin @@ -242,11 +280,6 @@ begin dec_trivial!, end -/-- `legendre_sym p (-2)` is given by `χ₈' p`. -/ -lemma legendre_sym_neg_two : legendre_sym p (-2) = χ₈' p := -by simp only [legendre_sym, card p, quadratic_char_neg_two ((ring_char_zmod_n p).substr hp), - int.cast_bit0, int.cast_one, int.cast_neg] - /-- `-2` is a square modulo an odd prime `p` iff `p` is congruent to `1` or `3` mod `8`. -/ lemma exists_sq_eq_neg_two_iff : is_square (-2 : zmod p) ↔ p % 8 = 1 ∨ p % 8 = 3 := begin @@ -259,16 +292,25 @@ begin dec_trivial!, end +end zmod + end values section reciprocity /-! ### The Law of Quadratic Reciprocity + +See `jacobi_sym.quadratic_reciprocity` and variants for a version of Quadratic Reciprocity +for the Jacobi symbol. -/ variables {p q : ℕ} [fact p.prime] [fact q.prime] +namespace legendre_sym + +open zmod + /-- The Law of Quadratic Reciprocity: if `p` and `q` are distinct odd primes, then `(q / p) * (p / q) = (-1)^((p-1)(q-1)/4)`. -/ theorem quadratic_reciprocity (hp : p ≠ 2) (hq : q ≠ 2) (hpq : p ≠ q) : @@ -294,17 +336,18 @@ theorem quadratic_reciprocity' (hp : p ≠ 2) (hq : q ≠ 2) : begin cases eq_or_ne p q with h h, { substI p, - rw [(legendre_sym_eq_zero_iff q q).mpr (by exact_mod_cast nat_cast_self q), mul_zero] }, + rw [(eq_zero_iff q q).mpr (by exact_mod_cast nat_cast_self q), mul_zero] }, { have qr := congr_arg (* legendre_sym p q) (quadratic_reciprocity hp hq h), have : ((q : ℤ) : zmod p) ≠ 0 := by exact_mod_cast prime_ne_zero p q h, - simpa only [mul_assoc, ← pow_two, legendre_sym_sq_one p this, mul_one] using qr } + simpa only [mul_assoc, ← pow_two, sq_one p this, mul_one] using qr } end /-- The Law of Quadratic Reciprocity: if `p` and `q` are odd primes and `p % 4 = 1`, then `(q / p) = (p / q)`. -/ theorem quadratic_reciprocity_one_mod_four (hp : p % 4 = 1) (hq : q ≠ 2) : legendre_sym q p = legendre_sym p q := -by rw [quadratic_reciprocity' (prime.mod_two_eq_one_iff_ne_two.mp (odd_of_mod_four_eq_one hp)) hq, +by rw [quadratic_reciprocity' (prime.mod_two_eq_one_iff_ne_two.mp + (odd_of_mod_four_eq_one hp)) hq, pow_mul, neg_one_pow_div_two_of_one_mod_four hp, one_pow, one_mul] /-- The Law of Quadratic Reciprocity: if `p` and `q` are primes that are both congruent @@ -316,6 +359,12 @@ let nop := @neg_one_pow_div_two_of_three_mod_four in begin rwa [← prime.mod_two_eq_one_iff_ne_two, odd_of_mod_four_eq_three], end +end legendre_sym + +namespace zmod + +open legendre_sym + /-- If `p` and `q` are odd primes and `p % 4 = 1`, then `q` is a square mod `p` iff `p` is a square mod `q`. -/ lemma exists_sq_eq_prime_iff_of_mod_four_eq_one (hp1 : p % 4 = 1) (hq1 : q ≠ 2) : @@ -323,8 +372,7 @@ lemma exists_sq_eq_prime_iff_of_mod_four_eq_one (hp1 : p % 4 = 1) (hq1 : q ≠ 2 begin cases eq_or_ne p q with h h, { substI p }, - { rw [← legendre_sym_eq_one_iff' p (prime_ne_zero p q h), - ← legendre_sym_eq_one_iff' q (prime_ne_zero q p h.symm), + { rw [← eq_one_iff' p (prime_ne_zero p q h), ← eq_one_iff' q (prime_ne_zero q p h.symm), quadratic_reciprocity_one_mod_four hp1 hq1], } end @@ -333,9 +381,9 @@ a square mod `p` iff `p` is a nonsquare mod `q`. -/ lemma exists_sq_eq_prime_iff_of_mod_four_eq_three (hp3 : p % 4 = 3) (hq3 : q % 4 = 3) (hpq : p ≠ q) : is_square (q : zmod p) ↔ ¬ is_square (p : zmod q) := -by rw [← legendre_sym_eq_one_iff' p (prime_ne_zero p q hpq), ← legendre_sym_eq_neg_one_iff' q, +by rw [← eq_one_iff' p (prime_ne_zero p q hpq), ← eq_neg_one_iff' q, quadratic_reciprocity_three_mod_four hp3 hq3, neg_inj] -end reciprocity - end zmod + +end reciprocity From 11878a45b2e160498dfbf8488959290576411640 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Wed, 14 Sep 2022 12:38:51 +0000 Subject: [PATCH 256/828] =?UTF-8?q?feat(topology/algebra/uniform=5Fconverg?= =?UTF-8?q?ence):=20maps=20to=20a=20uniform=20group=20form=20a=20uniform?= =?UTF-8?q?=20group=20when=20equipped=20with=20the=20uniform=20structure?= =?UTF-8?q?=20of=20`=F0=9D=94=96`-convergence=20(#14693)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/topology/algebra/uniform_convergence.lean | 81 +++++++++++++++++++ 1 file changed, 81 insertions(+) create mode 100644 src/topology/algebra/uniform_convergence.lean diff --git a/src/topology/algebra/uniform_convergence.lean b/src/topology/algebra/uniform_convergence.lean new file mode 100644 index 0000000000000..461c976b2f718 --- /dev/null +++ b/src/topology/algebra/uniform_convergence.lean @@ -0,0 +1,81 @@ +/- +Copyright (c) 2022 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anatole Dedecker +-/ +import topology.uniform_space.uniform_convergence_topology +import topology.algebra.uniform_group + +/-! +# Algebraic facts about the topology of uniform convergence + +This file contains algebraic compatibility results about the uniform structure of uniform +convergence / `𝔖`-convergence. They will mostly be useful for defining strong topologies on the +space of continuous linear maps between two topological vector spaces. + +## Main statements + +* `uniform_convergence.uniform_group` : if `G` is a uniform group, then the uniform structure of + uniform convergence makes `α → G` a uniform group +* `uniform_convergence_on.uniform_group` : if `G` is a uniform group, then the uniform structure of + `𝔖`-convergence, for any `𝔖 : set (set α)`, makes `α → G` a uniform group + +## TODO + +* Let `E` be a TVS, `𝔖 : set (set α)` and `H` a submodule of `α → E`. If the image of any `S ∈ 𝔖` + by any `u ∈ H` is bounded (in the sense of `bornology.is_vonN_bounded`), then `H`, equipped with + the topology of `𝔖`-convergence, is a TVS. + +## References + +* [N. Bourbaki, *General Topology, Chapter X*][bourbaki1966] + +## Tags + +uniform convergence, strong dual + +-/ + +section group + +variables {α G : Type*} [group G] [uniform_space G] [uniform_group G] {𝔖 : set $ set α} + +local attribute [-instance] Pi.uniform_space + +/-- If `G` is a uniform group, then the uniform structure of uniform convergence makes `α → G` +a uniform group as well. -/ +@[to_additive "If `G` is a uniform additive group, then the uniform structure of uniform +convergence makes `α → G` a uniform additive group as well."] +protected lemma uniform_convergence.uniform_group : + @uniform_group (α → G) (uniform_convergence.uniform_space α G) _ := +begin + -- Since `(/) : G × G → G` is uniformly continuous, + -- `uniform_convergence.postcomp_uniform_continuous` tells us that + -- `((/) ∘ —) : (α → G × G) → (α → G)` is uniformly continuous too. By precomposing with + -- `uniform_convergence.uniform_equiv_prod_arrow`, this gives that + -- `(/) : (α → G) × (α → G) → (α → G)` is also uniformly continuous + letI : uniform_space (α → G) := uniform_convergence.uniform_space α G, + letI : uniform_space (α → G × G) := uniform_convergence.uniform_space α (G × G), + exact ⟨(uniform_convergence.postcomp_uniform_continuous uniform_continuous_div).comp + uniform_convergence.uniform_equiv_prod_arrow.symm.uniform_continuous⟩ +end + +/-- Let `𝔖 : set (set α)`. If `G` is a uniform group, then the uniform structure of +`𝔖`-convergence makes `α → G` a uniform group as well. -/ +@[to_additive "Let `𝔖 : set (set α)`. If `G` is a uniform additive group, then the uniform +structure of `𝔖`-convergence makes `α → G` a uniform additive group as well. "] +protected lemma uniform_convergence_on.uniform_group : + @uniform_group (α → G) (uniform_convergence_on.uniform_space α G 𝔖) _ := +begin + -- Since `(/) : G × G → G` is uniformly continuous, + -- `uniform_convergence_on.postcomp_uniform_continuous` tells us that + -- `((/) ∘ —) : (α → G × G) → (α → G)` is uniformly continuous too. By precomposing with + -- `uniform_convergence_on.uniform_equiv_prod_arrow`, this gives that + -- `(/) : (α → G) × (α → G) → (α → G)` is also uniformly continuous + letI : uniform_space (α → G) := uniform_convergence_on.uniform_space α G 𝔖, + letI : uniform_space (α → G × G) := uniform_convergence_on.uniform_space α (G × G) 𝔖, + exact ⟨(uniform_convergence_on.postcomp_uniform_continuous uniform_continuous_div).comp + uniform_convergence_on.uniform_equiv_prod_arrow.symm.uniform_continuous⟩ +end + +end group From db0c0a79afe25e59baaf233e4817f37f077b3c53 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Wed, 14 Sep 2022 13:30:25 +0000 Subject: [PATCH 257/828] refactor (equiv.perm.basic): base of_subtype on extend_domain (#16484) There were two distinct ways to extend a permutation from a subtype to the ambient type. As suggested by Johann Commelin in https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/permutations.20.3A.20of_subtype.20vs.20extend_domain/near/298180355 this PR unifies them by basing .of_subtype on the more general .extend_domain. Some proofs have been redone, some other simplified. For symmetry, I also added cycle_type_of_subtype (it just calls cycle_type_extend_domain. Co-authored-by: Antoine Chambert-Loir --- src/group_theory/perm/basic.lean | 42 ++++++++------------------ src/group_theory/perm/cycle/basic.lean | 18 +++++++++-- src/group_theory/perm/cycle/type.lean | 3 ++ src/group_theory/perm/sign.lean | 11 +++---- src/group_theory/perm/support.lean | 24 +++++++++++---- 5 files changed, 54 insertions(+), 44 deletions(-) diff --git a/src/group_theory/perm/basic.lean b/src/group_theory/perm/basic.lean index 0e0e1c801dbe3..02f99f3bce424 100644 --- a/src/group_theory/perm/basic.lean +++ b/src/group_theory/perm/basic.lean @@ -234,39 +234,24 @@ equiv.ext $ λ ⟨_, _⟩, rfl /-- The inclusion map of permutations on a subtype of `α` into permutations of `α`, fixing the other points. -/ def of_subtype {p : α → Prop} [decidable_pred p] : perm (subtype p) →* perm α := -{ to_fun := λ f, - ⟨λ x, if h : p x then f ⟨x, h⟩ else x, λ x, if h : p x then f⁻¹ ⟨x, h⟩ else x, - λ x, have h : ∀ h : p x, p (f ⟨x, h⟩), from λ h, (f ⟨x, h⟩).2, - by { simp only [], split_ifs at *; - simp only [perm.inv_apply_self, subtype.coe_eta, subtype.coe_mk, not_true, *] at * }, - λ x, have h : ∀ h : p x, p (f⁻¹ ⟨x, h⟩), from λ h, (f⁻¹ ⟨x, h⟩).2, - by { simp only [], split_ifs at *; - simp only [perm.apply_inv_self, subtype.coe_eta, subtype.coe_mk, not_true, *] at * }⟩, - map_one' := begin ext, dsimp, split_ifs; refl, end, - map_mul' := λ f g, equiv.ext $ λ x, begin - by_cases h : p x, - { have h₁ : p (f (g ⟨x, h⟩)), from (f (g ⟨x, h⟩)).2, - have h₂ : p (g ⟨x, h⟩), from (g ⟨x, h⟩).2, - simp only [h, h₂, coe_fn_mk, perm.mul_apply, dif_pos, subtype.coe_eta] }, - { simp only [h, coe_fn_mk, perm.mul_apply, dif_neg, not_false_iff] } -end } +{ to_fun := λ f, extend_domain f (equiv.refl (subtype p)), + map_one' := equiv.perm.extend_domain_one _, + map_mul' := λ f g, (equiv.perm.extend_domain_mul _ f g).symm, } lemma of_subtype_subtype_perm {f : perm α} {p : α → Prop} [decidable_pred p] (h₁ : ∀ x, p x ↔ p (f x)) (h₂ : ∀ x, f x ≠ x → p x) : of_subtype (subtype_perm f h₁) = f := equiv.ext $ λ x, begin - rw [of_subtype, subtype_perm], by_cases hx : p x, - { simp only [hx, coe_fn_mk, dif_pos, monoid_hom.coe_mk, subtype.coe_mk]}, - { haveI := classical.prop_decidable, - simp only [hx, not_not.mp (mt (h₂ x) hx), coe_fn_mk, dif_neg, not_false_iff, - monoid_hom.coe_mk] } + { exact (subtype_perm f h₁).extend_domain_apply_subtype _ hx, }, + { rw [of_subtype, monoid_hom.coe_mk, equiv.perm.extend_domain_apply_not_subtype], + { exact not_not.mp (λ h, hx (h₂ x (ne.symm h))), }, + { exact hx, }, } end lemma of_subtype_apply_of_mem {p : α → Prop} [decidable_pred p] (f : perm (subtype p)) {x : α} (hx : p x) : - of_subtype f x = f ⟨x, hx⟩ := -dif_pos hx + of_subtype f x = f ⟨x, hx⟩ := extend_domain_apply_subtype f _ hx @[simp] lemma of_subtype_apply_coe {p : α → Prop} [decidable_pred p] (f : perm (subtype p)) (x : subtype p) : @@ -275,20 +260,19 @@ subtype.cases_on x $ λ _, of_subtype_apply_of_mem f lemma of_subtype_apply_of_not_mem {p : α → Prop} [decidable_pred p] (f : perm (subtype p)) {x : α} (hx : ¬ p x) : - of_subtype f x = x := -dif_neg hx + of_subtype f x = x := extend_domain_apply_not_subtype f (equiv.refl (subtype p)) hx lemma mem_iff_of_subtype_apply_mem {p : α → Prop} [decidable_pred p] (f : perm (subtype p)) (x : α) : p x ↔ p ((of_subtype f : α → α) x) := -if h : p x then by simpa only [of_subtype, h, coe_fn_mk, dif_pos, true_iff, monoid_hom.coe_mk] - using (f ⟨x, h⟩).2 +if h : p x then +by simpa only [h, true_iff, monoid_hom.coe_mk, of_subtype_apply_of_mem f h] using (f ⟨x, h⟩).2 else by simp [h, of_subtype_apply_of_not_mem f h] @[simp] lemma subtype_perm_of_subtype {p : α → Prop} [decidable_pred p] (f : perm (subtype p)) : subtype_perm (of_subtype f) (mem_iff_of_subtype_apply_mem f) = f := -equiv.ext $ λ ⟨x, hx⟩, by { dsimp [subtype_perm, of_subtype], - simp only [show p x, from hx, dif_pos, subtype.coe_eta] } +equiv.ext $ λ ⟨x, hx⟩, + subtype.coe_injective (of_subtype_apply_of_mem f hx) @[simp] lemma default_perm {n : Type*} : (default : perm n) = 1 := rfl diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index 1e70027498abd..ad15315bd79e6 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -579,7 +579,13 @@ def cycle_of [fintype α] (f : perm α) (x : α) : perm α := of_subtype (@subtype_perm _ f (same_cycle f x) (λ _, same_cycle_apply.symm)) lemma cycle_of_apply [fintype α] (f : perm α) (x y : α) : - cycle_of f x y = if same_cycle f x y then f y else y := rfl + cycle_of f x y = if same_cycle f x y then f y else y := +begin + dsimp only [cycle_of], + split_ifs, + { apply of_subtype_apply_of_mem, exact h, }, + { apply of_subtype_apply_of_not_mem, exact h }, +end lemma cycle_of_inv [fintype α] (f : perm α) (x : α) : (cycle_of f x)⁻¹ = cycle_of f⁻¹ x := @@ -602,10 +608,16 @@ end zpow_neg_succ_of_nat, ← inv_pow, cycle_of_pow_apply_self] lemma same_cycle.cycle_of_apply [fintype α] {f : perm α} {x y : α} (h : same_cycle f x y) : - cycle_of f x y = f y := dif_pos h + cycle_of f x y = f y := +begin + apply of_subtype_apply_of_mem, exact h, +end lemma cycle_of_apply_of_not_same_cycle [fintype α] {f : perm α} {x y : α} (h : ¬same_cycle f x y) : - cycle_of f x y = y := dif_neg h + cycle_of f x y = y := +begin + apply of_subtype_apply_of_not_mem, exact h, +end lemma same_cycle.cycle_of_eq [fintype α] {f : perm α} {x y : α} (h : same_cycle f x y) : cycle_of f x = cycle_of f y := diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index b60349105f90d..b00aac036a120 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -303,6 +303,9 @@ begin rw [hd.cycle_type, ← extend_domain_mul, (hd.extend_domain f).cycle_type, hσ, hτ] } end +lemma cycle_type_of_subtype {p : α → Prop} [decidable_pred p] {g : perm (subtype p)}: + cycle_type (g.of_subtype) = cycle_type g := cycle_type_extend_domain (equiv.refl (subtype p)) + lemma mem_cycle_type_iff {n : ℕ} {σ : perm α} : n ∈ cycle_type σ ↔ ∃ c τ : perm α, σ = c * τ ∧ disjoint c τ ∧ is_cycle c ∧ c.support.card = n := begin diff --git a/src/group_theory/perm/sign.lean b/src/group_theory/perm/sign.lean index 704ea1e95cc82..f558b168c5a76 100644 --- a/src/group_theory/perm/sign.lean +++ b/src/group_theory/perm/sign.lean @@ -564,11 +564,6 @@ have hl'₂ : (l.1.map of_subtype).prod = f, by { conv { congr, rw ← l.2.1, skip, rw ← hl'₂ }, rw [sign_prod_list_swap l.2.2, sign_prod_list_swap hl', list.length_map] } -@[simp] lemma sign_of_subtype {p : α → Prop} [decidable_pred p] - (f : perm (subtype p)) : sign (of_subtype f) = sign f := -have ∀ x, of_subtype f x ≠ x → p x, from λ x, not_imp_comm.1 (of_subtype_apply_of_not_mem f), -by conv {to_rhs, rw [← subtype_perm_of_subtype f, sign_subtype_perm _ _ this]} - lemma sign_eq_sign_of_equiv [decidable_eq β] [fintype β] (f : perm α) (g : perm β) (e : α ≃ β) (h : ∀ x, e (f x) = g (e x)) : sign f = sign g := have hg : g = (e.symm.trans f).trans e, from equiv.ext $ by simp [h], @@ -685,7 +680,11 @@ by simp [subtype_congr] @[simp] lemma sign_extend_domain (e : perm α) {p : β → Prop} [decidable_pred p] (f : α ≃ subtype p) : equiv.perm.sign (e.extend_domain f) = equiv.perm.sign e := -by simp [equiv.perm.extend_domain] +by simp only [equiv.perm.extend_domain, sign_subtype_congr, sign_perm_congr, sign_refl, mul_one] + +@[simp] lemma sign_of_subtype {p : α → Prop} [decidable_pred p] + (f : equiv.perm (subtype p)) : equiv.perm.sign (f.of_subtype) = equiv.perm.sign f := +sign_extend_domain f (equiv.refl (subtype p)) end congr diff --git a/src/group_theory/perm/support.lean b/src/group_theory/perm/support.lean index 19830aa6d1a6a..05851de4298be 100644 --- a/src/group_theory/perm/support.lean +++ b/src/group_theory/perm/support.lean @@ -172,16 +172,28 @@ variable [decidable_eq α] /-- `f.is_swap` indicates that the permutation `f` is a transposition of two elements. -/ def is_swap (f : perm α) : Prop := ∃ x y, x ≠ y ∧ f = swap x y +@[simp] lemma of_subtype_swap_eq {p : α → Prop} [decidable_pred p] + (x y : subtype p) : + (equiv.swap x y).of_subtype = equiv.swap ↑x ↑y := +equiv.ext $ λ z, begin + by_cases hz : p z, + { rw [swap_apply_def, of_subtype_apply_of_mem _ hz], + split_ifs with hzx hzy, + { simp_rw [hzx, subtype.coe_eta, swap_apply_left], }, + { simp_rw [hzy, subtype.coe_eta, swap_apply_right], }, + { rw swap_apply_of_ne_of_ne, refl, + intro h, apply hzx, rw ← h, refl, + intro h, apply hzy, rw ← h, refl, } }, + { rw [of_subtype_apply_of_not_mem _ hz, swap_apply_of_ne_of_ne], + intro h, apply hz, rw h, exact subtype.prop x, + intro h, apply hz, rw h, exact subtype.prop y, } +end + lemma is_swap.of_subtype_is_swap {p : α → Prop} [decidable_pred p] {f : perm (subtype p)} (h : f.is_swap) : (of_subtype f).is_swap := let ⟨⟨x, hx⟩, ⟨y, hy⟩, hxy⟩ := h in ⟨x, y, by { simp only [ne.def] at hxy, exact hxy.1 }, - equiv.ext $ λ z, begin - rw [hxy.2, of_subtype], - simp only [swap_apply_def, coe_fn_mk, swap_inv, subtype.mk_eq_mk, monoid_hom.coe_mk], - split_ifs; - rw subtype.coe_mk <|> cc, - end⟩ + by { simp only [hxy.2, of_subtype_swap_eq], refl, }⟩ lemma ne_and_ne_of_swap_mul_apply_ne_self {f : perm α} {x y : α} (hy : (swap x (f x) * f) y ≠ y) : f y ≠ y ∧ y ≠ x := From 71b8c46ae3c48efb2565dad23f8293d968f0c303 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 14 Sep 2022 16:17:49 +0000 Subject: [PATCH 258/828] feat(ring_theory/multiplicity): generalize `multiplicity` (#16330) Divisibility is defined for any `semigroup`, so I also reduce the assumption of `multiplicity` to any `monoid`. At least it can be used for the elements which commute with every element, such as `power_series.X`. - [x] depends on: #16328 --- src/ring_theory/multiplicity.lean | 106 ++++++++++++++++++------------ 1 file changed, 65 insertions(+), 41 deletions(-) diff --git a/src/ring_theory/multiplicity.lean b/src/ring_theory/multiplicity.lean index 19eeea7a630ef..e48057423a0f9 100644 --- a/src/ring_theory/multiplicity.lean +++ b/src/ring_theory/multiplicity.lean @@ -29,13 +29,13 @@ open_locale big_operators /-- `multiplicity a b` returns the largest natural number `n` such that `a ^ n ∣ b`, as an `part_enat` or natural with infinity. If `∀ n, a ^ n ∣ b`, then it returns `⊤`-/ -def multiplicity [comm_monoid α] [decidable_rel ((∣) : α → α → Prop)] (a b : α) : part_enat := +def multiplicity [monoid α] [decidable_rel ((∣) : α → α → Prop)] (a b : α) : part_enat := part_enat.find $ λ n, ¬a ^ (n + 1) ∣ b namespace multiplicity -section comm_monoid -variables [comm_monoid α] +section monoid +variables [monoid α] /-- `multiplicity.finite a b` indicates that the multiplicity of `a` in `b` is finite. -/ @[reducible] def finite (a b : α) : Prop := ∃ n : ℕ, ¬a ^ (n + 1) ∣ b @@ -45,6 +45,9 @@ lemma finite_iff_dom [decidable_rel ((∣) : α → α → Prop)] {a b : α} : lemma finite_def {a b : α} : finite a b ↔ ∃ n : ℕ, ¬a ^ (n + 1) ∣ b := iff.rfl +lemma not_dvd_one_of_finite_one_right {a : α} : finite a 1 → ¬a ∣ 1 := +λ ⟨n, hn⟩ ⟨d, hd⟩, hn ⟨d ^ (n + 1), (pow_mul_pow_eq_one (n + 1) hd.symm).symm⟩ + @[norm_cast] theorem int.coe_nat_multiplicity (a b : ℕ) : multiplicity (a : ℤ) (b : ℤ) = multiplicity a b := @@ -61,14 +64,10 @@ lemma not_finite_iff_forall {a b : α} : (¬ finite a b) ↔ ∀ n : ℕ, a ^ n by simp [finite, multiplicity, not_not]; tauto⟩ lemma not_unit_of_finite {a b : α} (h : finite a b) : ¬is_unit a := -let ⟨n, hn⟩ := h in mt (is_unit_iff_forall_dvd.1 ∘ is_unit.pow (n + 1)) $ -λ h, hn (h b) - -lemma finite_of_finite_mul_left {a b c : α} : finite a (b * c) → finite a c := -λ ⟨n, hn⟩, ⟨n, λ h, hn (h.trans (by simp [mul_pow]))⟩ +let ⟨n, hn⟩ := h in hn ∘ is_unit.dvd ∘ is_unit.pow (n + 1) lemma finite_of_finite_mul_right {a b c : α} : finite a (b * c) → finite a b := -by rw mul_comm; exact finite_of_finite_mul_left +λ ⟨n, hn⟩, ⟨n, λ h, hn (h.trans (dvd_mul_right _ _))⟩ variable [decidable_rel ((∣) : α → α → Prop)] @@ -133,29 +132,19 @@ by { simp only [not_not], exact ⟨λ h n, nat.cases_on n (by { rw pow_zero, exact one_dvd _}) (λ n, h _), λ h n, h _⟩ } @[simp] lemma is_unit_left {a : α} (b : α) (ha : is_unit a) : multiplicity a b = ⊤ := -eq_top_iff.2 (λ _, is_unit_iff_forall_dvd.1 (ha.pow _) _) - -lemma is_unit_right {a b : α} (ha : ¬is_unit a) (hb : is_unit b) : - multiplicity a b = 0 := -eq_coe_iff.2 ⟨show a ^ 0 ∣ b, by simp only [pow_zero, one_dvd], - by { rw pow_one, exact λ h, mt (is_unit_of_dvd_unit h) ha hb }⟩ +eq_top_iff.2 (λ _, is_unit.dvd (ha.pow _)) @[simp] lemma one_left (b : α) : multiplicity 1 b = ⊤ := is_unit_left b is_unit_one -lemma one_right {a : α} (ha : ¬is_unit a) : multiplicity a 1 = 0 := is_unit_right ha is_unit_one - @[simp] lemma get_one_right {a : α} (ha : finite a 1) : get (multiplicity a 1) ha = 0 := begin rw [part_enat.get_eq_iff_eq_coe, eq_coe_iff, pow_zero], - simpa [is_unit_iff_dvd_one.symm] using not_unit_of_finite ha, + simp [not_dvd_one_of_finite_one_right ha], end @[simp] lemma unit_left (a : α) (u : αˣ) : multiplicity (u : α) a = ⊤ := is_unit_left a u.is_unit -lemma unit_right {a : α} (ha : ¬is_unit a) (u : αˣ) : multiplicity a u = 0 := -is_unit_right ha u.is_unit - lemma multiplicity_eq_zero_of_not_dvd {a b : α} (ha : ¬a ∣ b) : multiplicity a b = 0 := by { rw [← nat.cast_zero, eq_coe_iff], simpa } @@ -192,14 +181,11 @@ lemma multiplicity_le_multiplicity_iff {a b c d : α} : multiplicity a b ≤ mul by rw [eq_top_iff_not_finite.2 hab, eq_top_iff_not_finite.2 (not_finite_iff_forall.2 this)]⟩ -lemma multiplicity_le_multiplicity_of_dvd_left {a b c : α} (hdvd : a ∣ b) : - multiplicity b c ≤ multiplicity a c := -multiplicity_le_multiplicity_iff.2 $ λ n h, (pow_dvd_pow_of_dvd hdvd n).trans h - -lemma eq_of_associated_left {a b c : α} (h : associated a b) : - multiplicity b c = multiplicity a c := -le_antisymm (multiplicity_le_multiplicity_of_dvd_left h.dvd) - (multiplicity_le_multiplicity_of_dvd_left h.symm.dvd) +lemma multiplicity_eq_multiplicity_iff {a b c d : α} : multiplicity a b = multiplicity c d ↔ + (∀ n : ℕ, a ^ n ∣ b ↔ c ^ n ∣ d) := +⟨λ h n, ⟨multiplicity_le_multiplicity_iff.mp h.le n, multiplicity_le_multiplicity_iff.mp h.ge n⟩, + λ h, le_antisymm (multiplicity_le_multiplicity_iff.mpr (λ n, (h n).mp)) + (multiplicity_le_multiplicity_iff.mpr (λ n, (h n).mpr))⟩ lemma multiplicity_le_multiplicity_of_dvd_right {a b c : α} (h : b ∣ c) : multiplicity a b ≤ multiplicity a c := @@ -240,11 +226,44 @@ end alias dvd_iff_multiplicity_pos ↔ _ _root_.has_dvd.dvd.multiplicity_pos +end monoid + +section comm_monoid +variables [comm_monoid α] + +lemma finite_of_finite_mul_left {a b c : α} : finite a (b * c) → finite a c := +by rw mul_comm; exact finite_of_finite_mul_right + +variable [decidable_rel ((∣) : α → α → Prop)] + +lemma is_unit_right {a b : α} (ha : ¬is_unit a) (hb : is_unit b) : + multiplicity a b = 0 := +eq_coe_iff.2 ⟨show a ^ 0 ∣ b, by simp only [pow_zero, one_dvd], + by { rw pow_one, exact λ h, mt (is_unit_of_dvd_unit h) ha hb }⟩ + +lemma one_right {a : α} (ha : ¬is_unit a) : multiplicity a 1 = 0 := is_unit_right ha is_unit_one + +lemma unit_right {a : α} (ha : ¬is_unit a) (u : αˣ) : multiplicity a u = 0 := +is_unit_right ha u.is_unit + +open_locale classical + +lemma multiplicity_le_multiplicity_of_dvd_left {a b c : α} (hdvd : a ∣ b) : + multiplicity b c ≤ multiplicity a c := +multiplicity_le_multiplicity_iff.2 $ λ n h, (pow_dvd_pow_of_dvd hdvd n).trans h + +lemma eq_of_associated_left {a b c : α} (h : associated a b) : + multiplicity b c = multiplicity a c := +le_antisymm (multiplicity_le_multiplicity_of_dvd_left h.dvd) + (multiplicity_le_multiplicity_of_dvd_left h.symm.dvd) + +alias dvd_iff_multiplicity_pos ↔ _ _root_.has_dvd.dvd.multiplicity_pos + end comm_monoid -section comm_monoid_with_zero +section monoid_with_zero -variable [comm_monoid_with_zero α] +variable [monoid_with_zero α] lemma ne_zero_of_finite {a b : α} (h : finite a b) : b ≠ 0 := let ⟨n, hn⟩ := h in λ hb, by simpa [hb] using hn @@ -260,6 +279,14 @@ begin rwa zero_dvd_iff, end +end monoid_with_zero + +section comm_monoid_with_zero + +variable [comm_monoid_with_zero α] + +variable [decidable_rel ((∣) : α → α → Prop)] + lemma multiplicity_mk_eq_multiplicity [decidable_rel ((∣) : associates α → associates α → Prop)] {a b : α} : multiplicity (associates.mk a) (associates.mk b) = multiplicity a b := begin @@ -274,15 +301,15 @@ begin { suffices : ¬ (finite (associates.mk a) (associates.mk b)), { rw [finite_iff_dom, part_enat.not_dom_iff_eq_top] at h this, rw [h, this] }, - refine not_finite_iff_forall.mpr (λ n, by {rw [← associates.mk_pow, associates.mk_dvd_mk], + refine not_finite_iff_forall.mpr (λ n, by { rw [← associates.mk_pow, associates.mk_dvd_mk], exact not_finite_iff_forall.mp h n }) } end end comm_monoid_with_zero -section comm_semiring +section semiring -variables [comm_semiring α] [decidable_rel ((∣) : α → α → Prop)] +variables [semiring α] [decidable_rel ((∣) : α → α → Prop)] lemma min_le_multiplicity_add {p a b : α} : min (multiplicity p a) (multiplicity p b) ≤ multiplicity p (a + b) := @@ -292,13 +319,11 @@ lemma min_le_multiplicity_add {p a b : α} : (λ h, by rw [min_eq_right h, multiplicity_le_multiplicity_iff]; exact λ n hn, dvd_add (multiplicity_le_multiplicity_iff.1 h n hn) hn) -end comm_semiring - -section comm_ring +end semiring -variables [comm_ring α] [decidable_rel ((∣) : α → α → Prop)] +section ring -open_locale classical +variables [ring α] [decidable_rel ((∣) : α → α → Prop)] @[simp] protected lemma neg (a b : α) : multiplicity a (-b) = multiplicity a b := part.ext' (by simp only [multiplicity, part_enat.find, dvd_neg]) @@ -340,7 +365,7 @@ begin { rw [multiplicity_add_of_gt hab, min_eq_right], exact le_of_lt hab}, end -end comm_ring +end ring section cancel_comm_monoid_with_zero @@ -476,7 +501,6 @@ lemma multiplicity_pow_self_of_prime {p : α} (hp : prime p) (n : ℕ) : multiplicity p (p ^ n) = n := multiplicity_pow_self hp.ne_zero hp.not_unit n - end cancel_comm_monoid_with_zero section valuation From 9532aeba3ad43fb4dfc0ed02eea0b7a13d890813 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 14 Sep 2022 16:17:51 +0000 Subject: [PATCH 259/828] chore(data/mv_polynomial/{basic, monad}): move lemmas `aeval_X_left` and `aeval_X_left_apply` (#16391) --- src/data/mv_polynomial/basic.lean | 6 ++++++ src/data/mv_polynomial/monad.lean | 6 ------ 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/src/data/mv_polynomial/basic.lean b/src/data/mv_polynomial/basic.lean index 2ad8f7d923525..01173043ffae5 100644 --- a/src/data/mv_polynomial/basic.lean +++ b/src/data/mv_polynomial/basic.lean @@ -1110,6 +1110,12 @@ theorem aeval_unique (φ : mv_polynomial σ R →ₐ[R] S₁) : φ = aeval (φ ∘ X) := by { ext i, simp } +lemma aeval_X_left : aeval X = alg_hom.id R (mv_polynomial σ R) := +(aeval_unique (alg_hom.id R _)).symm + +lemma aeval_X_left_apply (p : mv_polynomial σ R) : aeval X p = p := +alg_hom.congr_fun aeval_X_left p + lemma comp_aeval {B : Type*} [comm_semiring B] [algebra R B] (φ : S₁ →ₐ[R] B) : φ.comp (aeval f) = aeval (λ i, φ (f i)) := diff --git a/src/data/mv_polynomial/monad.lean b/src/data/mv_polynomial/monad.lean index b681972fda45b..a1b138f35320f 100644 --- a/src/data/mv_polynomial/monad.lean +++ b/src/data/mv_polynomial/monad.lean @@ -133,12 +133,6 @@ eval₂_hom_X' f X i lemma bind₁_X_left : bind₁ (X : σ → mv_polynomial σ R) = alg_hom.id R _ := by { ext1 i, simp } -lemma aeval_X_left : aeval (X : σ → mv_polynomial σ R) = alg_hom.id R _ := -by rw [aeval_eq_bind₁, bind₁_X_left] - -lemma aeval_X_left_apply (φ : mv_polynomial σ R) : aeval X φ = φ := -by rw [aeval_eq_bind₁, bind₁_X_left, alg_hom.id_apply] - variable (f : σ → mv_polynomial τ R) @[simp] From 94b2c2cba7d64557ac30df4df8a25e9bdfa50911 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 14 Sep 2022 16:17:52 +0000 Subject: [PATCH 260/828] chore(category_theory/sites/limits): generalise universe level (#16408) --- src/category_theory/sites/limits.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/category_theory/sites/limits.lean b/src/category_theory/sites/limits.lean index af9b4b93f99ca..50ff8db9774df 100644 --- a/src/category_theory/sites/limits.lean +++ b/src/category_theory/sites/limits.lean @@ -33,10 +33,10 @@ open opposite section limits -universes w v u +universes w v u z variables {C : Type (max v u)} [category.{v} C] {J : grothendieck_topology C} variables {D : Type w} [category.{max v u} D] -variables {K : Type (max v u)} [small_category K] +variables {K : Type z} [small_category K] noncomputable theory From 5cbe3c4ff5b9d121c9a047f02946f5c49d4f8e0c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Wed, 14 Sep 2022 16:17:53 +0000 Subject: [PATCH 261/828] chore(data/finset/powerset): generalize powerset_len_nonempty (#16429) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `powerset_len_nonempty` holds for `n ≤ s.card`. --- src/data/finset/powerset.lean | 11 +++++++---- 1 file changed, 7 insertions(+), 4 deletions(-) diff --git a/src/data/finset/powerset.lean b/src/data/finset/powerset.lean index 8e4f0e13f3896..fc7401ccb03fc 100644 --- a/src/data/finset/powerset.lean +++ b/src/data/finset/powerset.lean @@ -195,15 +195,17 @@ begin simp [card_insert_of_not_mem this, nat.succ_inj'] end -lemma powerset_len_nonempty {n : ℕ} {s : finset α} (h : n < s.card) : +lemma powerset_len_nonempty {n : ℕ} {s : finset α} (h : n ≤ s.card) : (powerset_len n s).nonempty := begin classical, induction s using finset.induction_on with x s hx IH generalizing n, - { simpa using h }, + { rw [card_empty, le_zero_iff] at h, + rw [h, powerset_len_zero], + exact finset.singleton_nonempty _, }, { cases n, { simp }, - { rw [card_insert_of_not_mem hx, nat.succ_lt_succ_iff] at h, + { rw [card_insert_of_not_mem hx, nat.succ_le_succ_iff] at h, rw powerset_len_succ_insert hx, refine nonempty.mono _ ((IH h).image (insert x)), convert (subset_union_right _ _) } } @@ -247,7 +249,8 @@ begin { refine ⟨insert x t, _, mem_insert_self _ _⟩, rw [←insert_erase hx, powerset_len_succ_insert (not_mem_erase _ _)], exact mem_union_right _ (mem_image_of_mem _ ht) }, - { rwa [card_erase_of_mem hx, lt_tsub_iff_right] } } } + { rw [card_erase_of_mem hx], + exact nat.le_pred_of_lt hn, } } } end @[simp] From d8bc1823dd0ad41f05fbc453012d2d4fdd48ca16 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 14 Sep 2022 16:17:54 +0000 Subject: [PATCH 262/828] chore(order/filter/basic): golf (#16472) Reorder lemmas, golf --- src/order/filter/basic.lean | 16 ++++++---------- 1 file changed, 6 insertions(+), 10 deletions(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index f33c040554aba..1249ab1e3031c 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -2035,18 +2035,14 @@ lemma disjoint_map {m : α → β} (hm : injective m) {f₁ f₂ : filter α} : disjoint (map m f₁) (map m f₂) ↔ disjoint f₁ f₂ := by simp only [disjoint_iff, ← map_inf hm, map_eq_bot_iff] -lemma map_eq_comap_of_inverse {f : filter α} {m : α → β} {n : β → α} - (h₁ : m ∘ n = id) (h₂ : n ∘ m = id) : map m f = comap n f := -le_antisymm - (λ b ⟨a, ha, (h : preimage n a ⊆ b)⟩, f.sets_of_superset ha $ - calc a = preimage (n ∘ m) a : by simp only [h₂, preimage_id, eq_self_iff_true] - ... ⊆ preimage m b : preimage_mono h) - (λ b (hb : preimage m b ∈ f), - ⟨preimage m b, hb, show preimage (m ∘ n) b ⊆ b, by simp only [h₁]; apply subset.refl⟩) - lemma map_equiv_symm (e : α ≃ β) (f : filter β) : map e.symm f = comap e f := -map_eq_comap_of_inverse e.symm_comp_self e.self_comp_symm +map_injective e.injective $ by rw [map_map, e.self_comp_symm, map_id, + map_comap_of_surjective e.surjective] + +lemma map_eq_comap_of_inverse {f : filter α} {m : α → β} {n : β → α} + (h₁ : m ∘ n = id) (h₂ : n ∘ m = id) : map m f = comap n f := +map_equiv_symm ⟨n, m, congr_fun h₁, congr_fun h₂⟩ f lemma comap_equiv_symm (e : α ≃ β) (f : filter α) : comap e.symm f = map e f := From 0602c59878ff3d5f71dea69c2d32ccf2e93e5398 Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Wed, 14 Sep 2022 16:17:56 +0000 Subject: [PATCH 263/828] refactor(model_theory/*): Use `countable` in model theory (#16496) Uses `countable` instead of custom classes in the model theory library. Co-authored-by: Yury G. Kudryashov --- src/model_theory/basic.lean | 37 +++++-------------- src/model_theory/encoding.lean | 9 +++-- src/model_theory/finitely_generated.lean | 4 +-- src/model_theory/fraisse.lean | 45 ++++++++++++------------ src/model_theory/substructures.lean | 8 ++--- 5 files changed, 40 insertions(+), 63 deletions(-) diff --git a/src/model_theory/basic.lean b/src/model_theory/basic.lean index 07a2c899efc82..047ac9598fa70 100644 --- a/src/model_theory/basic.lean +++ b/src/model_theory/basic.lean @@ -128,11 +128,6 @@ rfl /-- The cardinality of a language is the cardinality of its type of symbols. -/ def card : cardinal := # L.symbols -/-- A language is countable when it has countably many symbols. -/ -@[protected] class countable : Prop := (card_le_aleph_0' : L.card ≤ ℵ₀) - -lemma card_le_aleph_0 [L.countable] : L.card ≤ ℵ₀ := countable.card_le_aleph_0' - /-- A language is relational when it has no function symbols. -/ class is_relational : Prop := (empty_functions : ∀ n, is_empty (L.functions n)) @@ -141,12 +136,6 @@ class is_relational : Prop := class is_algebraic : Prop := (empty_relations : ∀ n, is_empty (L.relations n)) -/-- A language is countable when it has countably many symbols. -/ -class countable_functions : Prop := (card_functions_le_aleph_0' : # (Σ l, L.functions l) ≤ ℵ₀) - -lemma card_functions_le_aleph_0 [L.countable_functions] : #(Σ l, L.functions l) ≤ ℵ₀ := -countable_functions.card_functions_le_aleph_0' - variables {L} {L' : language.{u' v'}} lemma card_eq_card_functions_add_card_relations : @@ -198,28 +187,18 @@ instance subsingleton_mk₂_relations {c f₁ f₂ : Type u} {r₁ r₂ : Type v nat.cases_on n ⟨λ x, pempty.elim x⟩ (λ n, nat.cases_on n h1 (λ n, nat.cases_on n h2 (λ n, ⟨λ x, pempty.elim x⟩))) -lemma encodable.countable [_root_.countable L.symbols] : L.countable := -⟨cardinal.mk_le_aleph_0⟩ - @[simp] lemma empty_card : language.empty.card = 0 := by simp [card_eq_card_functions_add_card_relations] -instance countable_empty : language.empty.countable := -⟨by simp⟩ - -@[priority 100] instance countable.countable_functions [L.countable] : L.countable_functions := -⟨begin - refine lift_le_aleph_0.1 (trans _ L.card_le_aleph_0), - rw [card, symbols, mk_sum], - exact le_self_add -end⟩ - -lemma encodable.countable_functions [h : encodable (Σl, L.functions l)] : L.countable_functions := -⟨cardinal.mk_le_aleph_0⟩ +instance is_empty_empty : is_empty language.empty.symbols := +begin + simp only [language.symbols, is_empty_sum, is_empty_sigma], + exact ⟨λ _, infer_instance, λ _, infer_instance⟩, +end -@[priority 100] instance is_relational.countable_functions [L.is_relational] : - L.countable_functions := -encodable.countable_functions +instance countable.countable_functions [h : countable L.symbols] : + countable (Σl, L.functions l) := +@function.injective.countable _ _ h _ sum.inl_injective @[simp] lemma card_functions_sum (i : ℕ) : #((L.sum L').functions i) = (#(L.functions i)).lift + cardinal.lift.{u} (#(L'.functions i)) := diff --git a/src/model_theory/encoding.lean b/src/model_theory/encoding.lean index 10fdaa84c80ba..133e6badd613a 100644 --- a/src/model_theory/encoding.lean +++ b/src/model_theory/encoding.lean @@ -151,13 +151,12 @@ encodable.of_left_injection list_encode (λ l, (list_decode l).head'.join) simp only [option.join, head', list.map, option.some_bind, id.def], end) -lemma card_le_aleph_0 [h1 : countable α] [h2 : L.countable_functions] : - # (L.term α) ≤ ℵ₀ := +instance [h1 : countable α] [h2 : countable (Σl, L.functions l)] : + countable (L.term α) := begin - refine (card_le.trans _), - rw [max_le_iff], + refine mk_le_aleph_0_iff.1 (card_le.trans (max_le_iff.2 _)), simp only [le_refl, mk_sum, add_le_aleph_0, lift_le_aleph_0, true_and], - exact ⟨mk_le_aleph_0, L.card_functions_le_aleph_0⟩, + exact ⟨cardinal.mk_le_aleph_0, cardinal.mk_le_aleph_0⟩, end instance small [small.{u} α] : diff --git a/src/model_theory/finitely_generated.lean b/src/model_theory/finitely_generated.lean index 19d55eac3ae5f..e178102185101 100644 --- a/src/model_theory/finitely_generated.lean +++ b/src/model_theory/finitely_generated.lean @@ -156,7 +156,7 @@ begin exact hom.map_le_range h' end -theorem cg_iff_countable [L.countable_functions] {s : L.substructure M} : +theorem cg_iff_countable [countable (Σl, L.functions l)] {s : L.substructure M} : s.cg ↔ countable s := begin refine ⟨_, λ h, ⟨s, h.to_set, s.closure_eq⟩⟩, @@ -224,7 +224,7 @@ begin exact h.range f, end -lemma cg_iff_countable [L.countable_functions] : cg L M ↔ countable M := +lemma cg_iff_countable [countable (Σl, L.functions l)] : cg L M ↔ countable M := by rw [cg_def, cg_iff_countable, top_equiv.to_equiv.countable_iff] lemma fg.cg (h : fg L M) : cg L M := diff --git a/src/model_theory/fraisse.lean b/src/model_theory/fraisse.lean index 2415f5397cd1a..d819e48cea930 100644 --- a/src/model_theory/fraisse.lean +++ b/src/model_theory/fraisse.lean @@ -146,21 +146,20 @@ lemma age.joint_embedding : joint_embedding (L.age M) := /-- The age of a countable structure is essentially countable (has countably many isomorphism classes). -/ -lemma age.countable_quotient (h : (univ : set M).countable) : - (quotient.mk '' (L.age M)).countable := +lemma age.countable_quotient [h : countable M] : + (quotient.mk '' L.age M).countable := begin - refine eq.mp (congr rfl (set.ext _)) ((countable_set_of_finite_subset h).image - (λ s, ⟦⟨closure L s, infer_instance⟩⟧)), - rw forall_quotient_iff, - intro N, - simp only [subset_univ, and_true, mem_image, mem_set_of_eq, quotient.eq], + classical, + refine (congr_arg _ (set.ext $ forall_quotient_iff.2 $ λ N, _)).mp + (countable_range $ λ s : finset M, ⟦⟨closure L (s : set M), infer_instance⟩⟧), + simp only [mem_image, mem_range, mem_set_of_eq, quotient.eq], split, - { rintro ⟨s, hs1, hs2⟩, - use bundled.of ↥(closure L s), - exact ⟨⟨(fg_iff_Structure_fg _).1 (fg_closure hs1), ⟨subtype _⟩⟩, hs2⟩ }, + { rintro ⟨s, hs⟩, + use bundled.of ↥(closure L (s : set M)), + exact ⟨⟨(fg_iff_Structure_fg _).1 (fg_closure s.finite_to_set), ⟨subtype _⟩⟩, hs⟩ }, { rintro ⟨P, ⟨⟨s, hs⟩, ⟨PM⟩⟩, hP2⟩, - refine ⟨PM '' s, set.finite.image PM s.finite_to_set, setoid.trans _ hP2⟩, - rw [← embedding.coe_to_hom, closure_image PM.to_hom, hs, ← hom.range_eq_map], + refine ⟨s.image PM, setoid.trans _ hP2⟩, + rw [← embedding.coe_to_hom, finset.coe_image, closure_image PM.to_hom, hs, ← hom.range_eq_map], exact ⟨PM.equiv_range.symm⟩ } end @@ -221,8 +220,8 @@ begin { exact (hFP _ n).some } end -theorem exists_countable_is_age_of_iff [L.countable_functions] : - (∃ (M : bundled.{w} L.Structure), (univ : set M).countable ∧ L.age M = K) ↔ +theorem exists_countable_is_age_of_iff [countable (Σl, L.functions l)] : + (∃ (M : bundled.{w} L.Structure), countable M ∧ L.age M = K) ↔ K.nonempty ∧ (∀ (M N : bundled.{w} L.Structure), nonempty (M ≃[L] N) → (M ∈ K ↔ N ∈ K)) ∧ (quotient.mk '' K).countable ∧ @@ -233,12 +232,11 @@ begin split, { rintros ⟨M, h1, h2, rfl⟩, resetI, - refine ⟨age.nonempty M, age.is_equiv_invariant L M, age.countable_quotient M h1, λ N hN, hN.1, + refine ⟨age.nonempty M, age.is_equiv_invariant L M, age.countable_quotient M, λ N hN, hN.1, age.hereditary M, age.joint_embedding M⟩, }, { rintros ⟨Kn, eqinv, cq, hfg, hp, jep⟩, obtain ⟨M, hM, rfl⟩ := exists_cg_is_age_of Kn eqinv cq hfg hp jep, - haveI : countable M := Structure.cg_iff_countable.1 hM, - exact ⟨M, to_countable _, rfl⟩, } + exact ⟨M, Structure.cg_iff_countable.1 hM, rfl⟩ } end variables {K} (L) (M) @@ -253,9 +251,9 @@ variables {L} (K) /-- A structure `M` is a Fraïssé limit for a class `K` if it is countably generated, ultrahomogeneous, and has age `K`. -/ -structure is_fraisse_limit [countable_functions L] : Prop := +@[protect_proj] structure is_fraisse_limit [countable (Σl, L.functions l)] + [countable M] : Prop := (ultrahomogeneous : is_ultrahomogeneous L M) -(countable : (univ : set M).countable) (age : L.age M = K) variables {L} {M} @@ -279,17 +277,18 @@ begin set.coe_inclusion, embedding.equiv_range_apply, hgn], end -lemma is_ultrahomogeneous.age_is_fraisse (hc : (univ : set M).countable) +lemma is_ultrahomogeneous.age_is_fraisse [countable M] (h : L.is_ultrahomogeneous M) : is_fraisse (L.age M) := -⟨age.nonempty M, λ _ hN, hN.1, age.is_equiv_invariant L M, age.countable_quotient M hc, +⟨age.nonempty M, λ _ hN, hN.1, age.is_equiv_invariant L M, age.countable_quotient M, age.hereditary M, age.joint_embedding M, h.amalgamation_age⟩ namespace is_fraisse_limit /-- If a class has a Fraïssé limit, it must be Fraïssé. -/ -theorem is_fraisse [countable_functions L] (h : is_fraisse_limit K M) : is_fraisse K := -(congr rfl h.age).mp (h.ultrahomogeneous.age_is_fraisse h.countable) +theorem is_fraisse [countable (Σl, L.functions l)] [countable M] (h : is_fraisse_limit K M) : + is_fraisse K := +(congr rfl h.age).mp h.ultrahomogeneous.age_is_fraisse end is_fraisse_limit diff --git a/src/model_theory/substructures.lean b/src/model_theory/substructures.lean index e2411eb0f75d8..8228ffe11dfbf 100644 --- a/src/model_theory/substructures.lean +++ b/src/model_theory/substructures.lean @@ -272,13 +272,13 @@ end variable (L) -lemma _root_.set.countable.substructure_closure - [L.countable_functions] (h : s.countable) : - countable (closure L s) := +lemma _root_.set.countable.substructure_closure [countable (Σl, L.functions l)] + (h : s.countable) : + countable.{w + 1} (closure L s) := begin haveI : countable s := h.to_subtype, rw [← mk_le_aleph_0_iff, ← lift_le_aleph_0], - exact lift_card_closure_le_card_term.trans term.card_le_aleph_0, + exact lift_card_closure_le_card_term.trans mk_le_aleph_0 end variables {L} (S) From 660bde43158a8f5f7bcb605bd710ac3e5c55e81d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 14 Sep 2022 18:51:34 +0000 Subject: [PATCH 264/828] =?UTF-8?q?refactor(algebra/order/monoid=5Flemma?= =?UTF-8?q?=5Fzero=5Flt):=20Use=20`{x=20:=20=CE=B1=20//=200=20=E2=89=A4=20?= =?UTF-8?q?=CE=B1}`=20(#16447)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `pos_mul_mono`/`mul_pos_mono` and `pos_mul_reflect_lt`/`mul_pos_reflect_lt` were stated as covariance/contravariance of `{x : α // 0 < α}` over `α` even though the extension to `{x : α // 0 ≤ α}` also holds. This meant that many lemmas were relying on antisymmetry only to treat the cases `a = 0` and `0 < a` separately, which made them need `partial_order` and depend on `classical.choice`. This prevented #16172 from actually removing the `decidable` lemma in `algebra.order.ring` after #16189 got merged. Further, #16189 did not actually get rid of the temporary `zero_lt` namespace, so name conflicts that arise were not fixed. Finally, `le_mul_of_one_le_left` and its seven friends were reproved inline about five time each. Hence in this PR we * restate `pos_mul_mono`, `mul_pos_mono`, `pos_mul_reflect_lt`, `mul_pos_reflect_lt` using `{x : α // 0 ≤ α}` * provide the (classical) equivalence with the previous definitions in a `partial_order`. * generalise lemmas from `partial_order` to `preorder`. * delete all lemmas in the `preorder` namespace as they are now fully generalized. This in essence reverts (parts of) #12961, #13296 and #13299. * replace most of the various `has_le`/`has_lt` assumptions by a blank `preorder` one, hence simplifying the file sectioning * remove the `zero_lt` namespace. * rename lemmas to fix name conflicts. * delete a few lemmas that were left in `algebra.order.ring`. * golf proofs involving `le_mul_of_one_le_left` and its seven friends. This is why the PR has a -450 lines diff. --- src/algebra/order/monoid_lemmas_zero_lt.lean | 1247 ++++++----------- src/algebra/order/ring.lean | 51 +- src/analysis/normed_space/basic.lean | 2 +- src/analysis/special_functions/bernstein.lean | 8 +- src/measure_theory/integral/set_to_l1.lean | 23 +- 5 files changed, 427 insertions(+), 904 deletions(-) diff --git a/src/algebra/order/monoid_lemmas_zero_lt.lean b/src/algebra/order/monoid_lemmas_zero_lt.lean index 83f6fcdb9471e..f0d03dda5612d 100644 --- a/src/algebra/order/monoid_lemmas_zero_lt.lean +++ b/src/algebra/order/monoid_lemmas_zero_lt.lean @@ -13,8 +13,6 @@ Let `α` be a type with `<` and `0`. We use the type `{x : α // 0 < x}` of pos to prove results about monotonicity of multiplication. We also introduce the local notation `α>0` for the subtype `{x : α // 0 < x}`: -* the notation `α>0` to stands for `{x : α // 0 < x}`. - If the type `α` also has a multiplication, then we combine this with (`contravariant_`) `covariant_class`es to assume that multiplication by positive elements is (strictly) monotone on a `mul_zero_class`, `monoid_with_zero`,... @@ -41,118 +39,100 @@ More specifically, we use extensively the following typeclasses: * * `contravariant_class α>0 α (λ x y, y * x) (<)`, abbreviated `mul_pos_reflect_lt α`, expressing that multiplication by positive elements on the right is strictly reverse monotone. -## Formalization comments - -We use `α>0 = {x : α // 0 < x}` with a strict inequality since in most cases what happens with `0` -is clear. This creates a few bumps in the first couple of proofs, where we have to split cases on -whether an element is `0` or not, but goes smoothly after that. A further advantage is that we -only introduce notation for the positive elements and we do not need also the non-negative ones. +## Notation -Some lemmas for `partial_order` also have a variant for `preorder`, where the `preorder` version has -stronger hypotheses. In this case we put the `preorder` lemma in the `preorder` namespace. +The following is local notation in this file: +* `α≥0`: `{x : α // 0 ≤ x}` +* `α>0`: `{x : α // 0 < x}` -/ -/- I am changing the file `algebra/order/monoid_lemmas` incrementally, with the idea of -reproducing almost all of the proofs in `algebra/order/ring` with weaker assumptions. -/ +variable (α : Type*) -universe u -variable {α : Type u} - -/- Notation for positive elements +/- Notations for nonnegative and positive elements https:// leanprover.zulipchat.com/#narrow/stream/113488-general/topic/notation.20for.20positive.20elements -/ +local notation `α≥0` := {x : α // 0 ≤ x} local notation `α>0` := {x : α // 0 < x} -namespace zero_lt +section abbreviations +variables [has_mul α] [has_zero α] [preorder α] + +/-- `pos_mul_mono α` is an abbreviation for `covariant_class α≥0 α (λ x y, x * y) (≤)`, +expressing that multiplication by nonnegative elements on the left is monotone. -/ +abbreviation pos_mul_mono : Prop := covariant_class α≥0 α (λ x y, x * y) (≤) -section abbreviations_strict_mono -variables (X : Type u) [has_mul X] [has_zero X] [has_lt X] +/-- `mul_pos_mono α` is an abbreviation for `covariant_class α≥0 α (λ x y, y * x) (≤)`, +expressing that multiplication by nonnegative elements on the right is monotone. -/ +abbreviation mul_pos_mono : Prop := covariant_class α≥0 α (λ x y, y * x) (≤) -/-- `zero_lt.pos_mul_strict_mono α` is an abbreviation for -`covariant_class α>0 α (λ x y, x * y) (<)`, +/-- `pos_mul_strict_mono α` is an abbreviation for `covariant_class α>0 α (λ x y, x * y) (<)`, expressing that multiplication by positive elements on the left is strictly monotone. -/ -abbreviation pos_mul_strict_mono : Prop := -covariant_class {x : X // 0 < x} X (λ x y, x * y) (<) +abbreviation pos_mul_strict_mono : Prop := covariant_class α>0 α (λ x y, x * y) (<) -/-- `zero_lt.mul_pos_strict_mono α` is an abbreviation for -`covariant_class α>0 α (λ x y, y * x) (<)`, +/-- `mul_pos_strict_mono α` is an abbreviation for `covariant_class α>0 α (λ x y, y * x) (<)`, expressing that multiplication by positive elements on the right is strictly monotone. -/ -abbreviation mul_pos_strict_mono : Prop := -covariant_class {x : X // 0 < x} X (λ x y, y * x) (<) - -/-- `zero_lt.pos_mul_reflect_lt α` is an abbreviation for -`contravariant_class α>0 α (λ x y, x * y) (<)`, -expressing that multiplication by positive elements on the left is strictly reverse monotone. -/ -abbreviation pos_mul_reflect_lt : Prop := -contravariant_class {x : X // 0 < x} X (λ x y, x * y) (<) - -/-- `zero_lt.mul_pos_reflect_lt α` is an abbreviation for -`contravariant_class α>0 α (λ x y, y * x) (<)`, -expressing that multiplication by positive elements on the right is strictly reverse monotone. -/ -abbreviation mul_pos_reflect_lt : Prop := -contravariant_class {x : X // 0 < x} X (λ x y, y * x) (<) - -end abbreviations_strict_mono - -section abbreviations_mono -variables (X : Type*) [has_mul X] [has_zero X] [has_lt X] [has_le X] - -/-- `zero_lt.pos_mul_mono α` is an abbreviation for -`covariant_class α>0 α (λ x y, x * y) (≤)`, -expressing that multiplication by positive elements on the left is monotone. -/ -abbreviation pos_mul_mono : Prop := -covariant_class {x : X // 0 < x} X (λ x y, x * y) (≤) - -/-- `zero_lt.mul_pos_mono α` is an abbreviation for -`covariant_class α>0 α (λ x y, y * x) (≤)`, -expressing that multiplication by positive elements on the right is monotone. -/ -abbreviation mul_pos_mono : Prop := -covariant_class {x : X // 0 < x} X (λ x y, y * x) (≤) - -/-- `zero_lt.pos_mul_mono_rev α` is an abbreviation for -`contravariant_class α>0 α (λ x y, x * y) (≤)`, +abbreviation mul_pos_strict_mono : Prop := covariant_class α>0 α (λ x y, y * x) (<) + +/-- `pos_mul_reflect_lt α` is an abbreviation for `contravariant_class α≥0 α (λ x y, x * y) (<)`, +expressing that multiplication by nonnegative elements on the left is strictly reverse monotone. -/ +abbreviation pos_mul_reflect_lt : Prop := contravariant_class α≥0 α (λ x y, x * y) (<) + +/-- `mul_pos_reflect_lt α` is an abbreviation for `contravariant_class α≥0 α (λ x y, y * x) (<)`, +expressing that multiplication by nonnegative elements on the right is strictly reverse monotone. -/ +abbreviation mul_pos_reflect_lt : Prop := contravariant_class α≥0 α (λ x y, y * x) (<) + +/-- `pos_mul_mono_rev α` is an abbreviation for `contravariant_class α>0 α (λ x y, x * y) (≤)`, expressing that multiplication by positive elements on the left is reverse monotone. -/ -abbreviation pos_mul_mono_rev : Prop := -contravariant_class {x : X // 0 < x} X (λ x y, x * y) (≤) +abbreviation pos_mul_mono_rev : Prop := contravariant_class α>0 α (λ x y, x * y) (≤) -/-- `zero_lt.mul_pos_mono_rev α` is an abbreviation for -`contravariant_class α>0 α (λ x y, y * x) (≤)`, +/-- `mul_pos_mono_rev α` is an abbreviation for `contravariant_class α>0 α (λ x y, y * x) (≤)`, expressing that multiplication by positive elements on the right is reverse monotone. -/ -abbreviation mul_pos_mono_rev : Prop := -contravariant_class {x : X // 0 < x} X (λ x y, y * x) (≤) +abbreviation mul_pos_mono_rev : Prop := contravariant_class α>0 α (λ x y, y * x) (≤) -end abbreviations_mono +end abbreviations -variables {a b c d : α} +variables {α} {a b c d : α} section has_mul_zero variables [has_mul α] [has_zero α] -section has_lt -variables [has_lt α] +section preorder +variables [preorder α] + +instance pos_mul_mono.to_covariant_class_pos_mul_le [pos_mul_mono α] : + covariant_class α>0 α (λ x y, x * y) (≤) := +⟨λ a b c bc, @covariant_class.elim α≥0 α (λ x y, x * y) (≤) _ ⟨_, a.2.le⟩ _ _ bc⟩ + +instance mul_pos_mono.to_covariant_class_pos_mul_le [mul_pos_mono α] : + covariant_class α>0 α (λ x y, y * x) (≤) := +⟨λ a b c bc, @covariant_class.elim α≥0 α (λ x y, y * x) (≤) _ ⟨_, a.2.le⟩ _ _ bc⟩ + +instance pos_mul_reflect_lt.to_contravariant_class_pos_mul_lt [pos_mul_reflect_lt α] : + contravariant_class α>0 α (λ x y, x * y) (<) := +⟨λ a b c bc, @contravariant_class.elim α≥0 α (λ x y, x * y) (<) _ ⟨_, a.2.le⟩ _ _ bc⟩ + +instance mul_pos_reflect_lt.to_contravariant_class_pos_mul_lt [mul_pos_reflect_lt α] : + contravariant_class α>0 α (λ x y, y * x) (<) := +⟨λ a b c bc, @contravariant_class.elim α≥0 α (λ x y, y * x) (<) _ ⟨_, a.2.le⟩ _ _ bc⟩ + +lemma mul_le_mul_of_nonneg_left [pos_mul_mono α] (h : b ≤ c) (ha : 0 ≤ a) : a * b ≤ a * c := +@covariant_class.elim α≥0 α (λ x y, x * y) (≤) _ ⟨a, ha⟩ _ _ h -lemma mul_lt_mul_left' [pos_mul_strict_mono α] - (bc : b < c) (a0 : 0 < a) : - a * b < a * c := -@covariant_class.elim α>0 α (λ x y, x * y) (<) _ ⟨a, a0⟩ _ _ bc +lemma mul_le_mul_of_nonneg_right [mul_pos_mono α] (h : b ≤ c) (ha : 0 ≤ a) : b * a ≤ c * a := +@covariant_class.elim α≥0 α (λ x y, y * x) (≤) _ ⟨a, ha⟩ _ _ h -lemma mul_lt_mul_right' [mul_pos_strict_mono α] - (bc : b < c) (a0 : 0 < a) : - b * a < c * a := -@covariant_class.elim α>0 α (λ x y, y * x) (<) _ ⟨a, a0⟩ _ _ bc +lemma mul_lt_mul_of_pos_left [pos_mul_strict_mono α] (bc : b < c) (ha : 0 < a) : a * b < a * c := +@covariant_class.elim α>0 α (λ x y, x * y) (<) _ ⟨a, ha⟩ _ _ bc --- proven with `a0 : 0 ≤ a` as `lt_of_mul_lt_mul_left` -lemma lt_of_mul_lt_mul_left' [pos_mul_reflect_lt α] - (bc : a * b < a * c) (a0 : 0 < a) : - b < c := -@contravariant_class.elim α>0 α (λ x y, x * y) (<) _ ⟨a, a0⟩ _ _ bc +lemma mul_lt_mul_of_pos_right [mul_pos_strict_mono α] (bc : b < c) (ha : 0 < a) : b * a < c * a := +@covariant_class.elim α>0 α (λ x y, y * x) (<) _ ⟨a, ha⟩ _ _ bc --- proven with `a0 : 0 ≤ a` as `lt_of_mul_lt_mul_right` -lemma lt_of_mul_lt_mul_right' [mul_pos_reflect_lt α] - (bc : b * a < c * a) (a0 : 0 < a) : - b < c := -@contravariant_class.elim α>0 α (λ x y, y * x) (<) _ ⟨a, a0⟩ _ _ bc +lemma lt_of_mul_lt_mul_left [pos_mul_reflect_lt α] (h : a * b < a * c) (ha : 0 ≤ a) : b < c := +@contravariant_class.elim α≥0 α (λ x y, x * y) (<) _ ⟨a, ha⟩ _ _ h + +lemma lt_of_mul_lt_mul_right [mul_pos_reflect_lt α] (h : b * a < c * a) (ha : 0 ≤ a) : b < c := +@contravariant_class.elim α≥0 α (λ x y, y * x) (<) _ ⟨a, ha⟩ _ _ h @[simp] lemma mul_lt_mul_left [pos_mul_strict_mono α] [pos_mul_reflect_lt α] @@ -166,193 +146,113 @@ lemma mul_lt_mul_right [mul_pos_strict_mono α] [mul_pos_reflect_lt α] b * a < c * a ↔ b < c := @rel_iff_cov α>0 α (λ x y, y * x) (<) _ _ ⟨a, a0⟩ _ _ -end has_lt - -section has_lt_le -variables [has_lt α] [has_le α] - --- proven with `a0 : 0 ≤ a` as `mul_le_mul_left` -lemma mul_le_mul_left_of_le_of_pos [pos_mul_mono α] - (bc : b ≤ c) (a0 : 0 < a) : - a * b ≤ a * c := -@covariant_class.elim α>0 α (λ x y, x * y) (≤) _ ⟨a, a0⟩ _ _ bc - --- proven with `a0 : 0 ≤ a` as `mul_le_mul_right` -lemma mul_le_mul_right_of_le_of_pos [mul_pos_mono α] - (bc : b ≤ c) (a0 : 0 < a) : - b * a ≤ c * a := -@covariant_class.elim α>0 α (λ x y, y * x) (≤) _ ⟨a, a0⟩ _ _ bc - -lemma le_of_mul_le_mul_left' [pos_mul_mono_rev α] +lemma le_of_mul_le_mul_of_pos_left [pos_mul_mono_rev α] (bc : a * b ≤ a * c) (a0 : 0 < a) : b ≤ c := @contravariant_class.elim α>0 α (λ x y, x * y) (≤) _ ⟨a, a0⟩ _ _ bc -lemma le_of_mul_le_mul_right' [mul_pos_mono_rev α] +lemma le_of_mul_le_mul_of_pos_right [mul_pos_mono_rev α] (bc : b * a ≤ c * a) (a0 : 0 < a) : b ≤ c := @contravariant_class.elim α>0 α (λ x y, y * x) (≤) _ ⟨a, a0⟩ _ _ bc -@[simp] -lemma mul_le_mul_left [pos_mul_mono α] [pos_mul_mono_rev α] - (a0 : 0 < a) : +alias le_of_mul_le_mul_of_pos_left ← le_of_mul_le_mul_left +alias le_of_mul_le_mul_of_pos_right ← le_of_mul_le_mul_right + +@[simp] lemma mul_le_mul_left [pos_mul_mono α] [pos_mul_mono_rev α] (ha : 0 < a) : a * b ≤ a * c ↔ b ≤ c := -@rel_iff_cov α>0 α (λ x y, x * y) (≤) _ _ ⟨a, a0⟩ _ _ +@rel_iff_cov α>0 α (λ x y, x * y) (≤) _ _ ⟨a, ha⟩ _ _ -@[simp] -lemma mul_le_mul_right [mul_pos_mono α] [mul_pos_mono_rev α] - (a0 : 0 < a) : +@[simp] lemma mul_le_mul_right [mul_pos_mono α] [mul_pos_mono_rev α] (ha : 0 < a) : b * a ≤ c * a ↔ b ≤ c := -@rel_iff_cov α>0 α (λ x y, y * x) (≤) _ _ ⟨a, a0⟩ _ _ +@rel_iff_cov α>0 α (λ x y, y * x) (≤) _ _ ⟨a, ha⟩ _ _ -end has_lt_le - -section preorder -variables [preorder α] +lemma mul_lt_mul_of_pos_of_nonneg [pos_mul_strict_mono α] [mul_pos_mono α] + (h₁ : a ≤ b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 ≤ d) : a * c < b * d := +(mul_lt_mul_of_pos_left h₂ a0).trans_le (mul_le_mul_of_nonneg_right h₁ d0) --- proven with `a0 : 0 ≤ a` `d0 : 0 ≤ d` as `mul_le_mul_of_le_of_le` -lemma preorder.mul_le_mul_of_le_of_le [pos_mul_mono α] [mul_pos_mono α] - (h₁ : a ≤ b) (h₂ : c ≤ d) (a0 : 0 < a) (d0 : 0 < d) : a * c ≤ b * d := -(mul_le_mul_left_of_le_of_pos h₂ a0).trans (mul_le_mul_right_of_le_of_pos h₁ d0) +lemma mul_lt_mul_of_le_of_le' [pos_mul_strict_mono α] [mul_pos_mono α] + (h₁ : a ≤ b) (h₂ : c < d) (b0 : 0 < b) (c0 : 0 ≤ c) : a * c < b * d := +(mul_le_mul_of_nonneg_right h₁ c0).trans_lt (mul_lt_mul_of_pos_left h₂ b0) --- proven with `b0 : 0 ≤ b` `c0 : 0 ≤ c` as `mul_le_mul_of_le_of_le'` -lemma preorder.mul_le_mul_of_le_of_le' [pos_mul_mono α] [mul_pos_mono α] - (h₁ : a ≤ b) (h₂ : c ≤ d) (b0 : 0 < b) (c0 : 0 < c) : a * c ≤ b * d := -(mul_le_mul_right_of_le_of_pos h₁ c0).trans (mul_le_mul_left_of_le_of_pos h₂ b0) +lemma mul_lt_mul_of_nonneg_of_pos [pos_mul_mono α] [mul_pos_strict_mono α] + (h₁ : a < b) (h₂ : c ≤ d) (a0 : 0 ≤ a) (d0 : 0 < d) : a * c < b * d := +(mul_le_mul_of_nonneg_left h₂ a0).trans_lt (mul_lt_mul_of_pos_right h₁ d0) -lemma mul_lt_mul_of_le_of_lt [pos_mul_strict_mono α] [mul_pos_mono α] - (h₁ : a ≤ b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := -(mul_lt_mul_left' h₂ a0).trans_le (mul_le_mul_right_of_le_of_pos h₁ d0) +lemma mul_lt_mul_of_le_of_lt' [pos_mul_mono α] [mul_pos_strict_mono α] + (h₁ : a < b) (h₂ : c ≤ d) (b0 : 0 ≤ b) (c0 : 0 < c) : a * c < b * d := +(mul_lt_mul_of_pos_right h₁ c0).trans_le (mul_le_mul_of_nonneg_left h₂ b0) -lemma mul_lt_mul_of_le_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] - (h₁ : a ≤ b) (h₂ : c < d) (b0 : 0 < b) (c0 : 0 < c) : a * c < b * d := -(mul_le_mul_right_of_le_of_pos h₁ c0).trans_lt (mul_lt_mul_left' h₂ b0) - -lemma mul_lt_mul_of_lt_of_le [pos_mul_mono α] [mul_pos_strict_mono α] - (h₁ : a < b) (h₂ : c ≤ d) (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := -(mul_le_mul_left_of_le_of_pos h₂ a0).trans_lt (mul_lt_mul_right' h₁ d0) - -lemma mul_lt_mul_of_lt_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] - (h₁ : a < b) (h₂ : c ≤ d) (b0 : 0 < b) (c0 : 0 < c) : a * c < b * d := -(mul_lt_mul_right' h₁ c0).trans_le (mul_le_mul_left_of_le_of_pos h₂ b0) - -lemma mul_lt_mul_of_lt_of_lt [pos_mul_strict_mono α] [mul_pos_strict_mono α] +lemma mul_lt_mul_of_pos_of_pos [pos_mul_strict_mono α] [mul_pos_strict_mono α] (h₁ : a < b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := -(mul_lt_mul_left' h₂ a0).trans (mul_lt_mul_right' h₁ d0) +(mul_lt_mul_of_pos_left h₂ a0).trans (mul_lt_mul_of_pos_right h₁ d0) lemma mul_lt_mul_of_lt_of_lt' [pos_mul_strict_mono α] [mul_pos_strict_mono α] (h₁ : a < b) (h₂ : c < d) (b0 : 0 < b) (c0 : 0 < c) : a * c < b * d := -(mul_lt_mul_right' h₁ c0).trans (mul_lt_mul_left' h₂ b0) +(mul_lt_mul_of_pos_right h₁ c0).trans (mul_lt_mul_of_pos_left h₂ b0) --- proven with `a0 : 0 ≤ a` as `mul_le_of_mul_le_left` -lemma preorder.mul_le_of_mul_le_left [pos_mul_mono α] - (h : a * b ≤ c) (hle : d ≤ b) (a0 : 0 < a) : - a * d ≤ c := -(mul_le_mul_left_of_le_of_pos hle a0).trans h - -lemma mul_lt_of_mul_lt_left [pos_mul_mono α] - (h : a * b < c) (hle : d ≤ b) (a0 : 0 < a) : +lemma mul_lt_of_mul_lt_of_nonneg_left [pos_mul_mono α] (h : a * b < c) (hdb : d ≤ b) (ha : 0 ≤ a) : a * d < c := -(mul_le_mul_left_of_le_of_pos hle a0).trans_lt h - --- proven with `b0 : 0 ≤ b` as `le_mul_of_le_mul_left` -lemma preorder.le_mul_of_le_mul_left [pos_mul_mono α] - (h : a ≤ b * c) (hle : c ≤ d) (b0 : 0 < b) : - a ≤ b * d := -h.trans (mul_le_mul_left_of_le_of_pos hle b0) +(mul_le_mul_of_nonneg_left hdb ha).trans_lt h -lemma lt_mul_of_lt_mul_left [pos_mul_mono α] - (h : a < b * c) (hle : c ≤ d) (b0 : 0 < b) : +lemma lt_mul_of_lt_mul_of_nonneg_left [pos_mul_mono α] (h : a < b * c) (hcd : c ≤ d) (hb : 0 ≤ b) : a < b * d := -h.trans_le (mul_le_mul_left_of_le_of_pos hle b0) +h.trans_le $ mul_le_mul_of_nonneg_left hcd hb --- proven with `b0 : 0 ≤ b` as `mul_le_of_mul_le_right` -lemma preorder.mul_le_of_mul_le_right [mul_pos_mono α] - (h : a * b ≤ c) (hle : d ≤ a) (b0 : 0 < b) : - d * b ≤ c := -(mul_le_mul_right_of_le_of_pos hle b0).trans h - -lemma mul_lt_of_mul_lt_right [mul_pos_mono α] - (h : a * b < c) (hle : d ≤ a) (b0 : 0 < b) : +lemma mul_lt_of_mul_lt_of_nonneg_right [mul_pos_mono α] (h : a * b < c) (hda : d ≤ a) (hb : 0 ≤ b) : d * b < c := -(mul_le_mul_right_of_le_of_pos hle b0).trans_lt h - --- proven with `c0 : 0 ≤ c` as `le_mul_of_le_mul_right` -lemma preorder.le_mul_of_le_mul_right [mul_pos_mono α] - (h : a ≤ b * c) (hle : b ≤ d) (c0 : 0 < c) : - a ≤ d * c := -h.trans (mul_le_mul_right_of_le_of_pos hle c0) +(mul_le_mul_of_nonneg_right hda hb).trans_lt h -lemma lt_mul_of_lt_mul_right [mul_pos_mono α] - (h : a < b * c) (hle : b ≤ d) (c0 : 0 < c) : +lemma lt_mul_of_lt_mul_of_nonneg_right [mul_pos_mono α] (h : a < b * c) (hbd : b ≤ d) (hc : 0 ≤ c) : a < d * c := -h.trans_le (mul_le_mul_right_of_le_of_pos hle c0) +h.trans_le $ mul_le_mul_of_nonneg_right hbd hc end preorder -section partial_order -variables [partial_order α] - -@[priority 100] -- see Note [lower instance priority] -instance pos_mul_strict_mono.to_pos_mul_mono [pos_mul_strict_mono α] : pos_mul_mono α := -⟨λ x a b h, h.eq_or_lt.elim (λ h', h' ▸ le_rfl) (λ h', (mul_lt_mul_left' h' x.prop).le)⟩ - -@[priority 100] -- see Note [lower instance priority] -instance mul_pos_strict_mono.to_mul_pos_mono [mul_pos_strict_mono α] : mul_pos_mono α := -⟨λ x a b h, h.eq_or_lt.elim (λ h', h' ▸ le_rfl) (λ h', (mul_lt_mul_right' h' x.prop).le)⟩ - -@[priority 100] -- see Note [lower instance priority] -instance pos_mul_mono_rev.to_pos_mul_reflect_lt [pos_mul_mono_rev α] : pos_mul_reflect_lt α := -⟨λ x a b h, lt_of_le_of_ne (le_of_mul_le_mul_left' h.le x.prop) (λ h', by simpa [h'] using h)⟩ - -@[priority 100] -- see Note [lower instance priority] -instance mul_pos_mono_rev.to_mul_pos_reflect_lt [mul_pos_mono_rev α] : mul_pos_reflect_lt α := -⟨λ x a b h, lt_of_le_of_ne (le_of_mul_le_mul_right' h.le x.prop) (λ h', by simpa [h'] using h)⟩ - -end partial_order - section linear_order variables [linear_order α] @[priority 100] -- see Note [lower instance priority] instance pos_mul_strict_mono.to_pos_mul_mono_rev [pos_mul_strict_mono α] : pos_mul_mono_rev α := -⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt (mul_lt_mul_left' h' x.prop)⟩ +⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt $ mul_lt_mul_of_pos_left h' x.prop⟩ @[priority 100] -- see Note [lower instance priority] instance mul_pos_strict_mono.to_mul_pos_mono_rev [mul_pos_strict_mono α] : mul_pos_mono_rev α := -⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt (mul_lt_mul_right' h' x.prop)⟩ +⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt $ mul_lt_mul_of_pos_right h' x.prop⟩ lemma pos_mul_mono_rev.to_pos_mul_strict_mono [pos_mul_mono_rev α] : pos_mul_strict_mono α := -⟨λ x a b h, lt_of_not_le $ λ h', h.not_le (le_of_mul_le_mul_left' h' x.prop)⟩ +⟨λ x a b h, lt_of_not_le $ λ h', h.not_le $ le_of_mul_le_mul_of_pos_left h' x.prop⟩ lemma mul_pos_mono_rev.to_mul_pos_strict_mono [mul_pos_mono_rev α] : mul_pos_strict_mono α := -⟨λ x a b h, lt_of_not_le $ λ h', h.not_le (le_of_mul_le_mul_right' h' x.prop)⟩ +⟨λ x a b h, lt_of_not_le $ λ h', h.not_le $ le_of_mul_le_mul_of_pos_right h' x.prop⟩ lemma pos_mul_strict_mono_iff_pos_mul_mono_rev : pos_mul_strict_mono α ↔ pos_mul_mono_rev α := -⟨@zero_lt.pos_mul_strict_mono.to_pos_mul_mono_rev _ _ _ _, - @pos_mul_mono_rev.to_pos_mul_strict_mono _ _ _ _⟩ +⟨@pos_mul_strict_mono.to_pos_mul_mono_rev _ _ _ _, @pos_mul_mono_rev.to_pos_mul_strict_mono _ _ _ _⟩ lemma mul_pos_strict_mono_iff_mul_pos_mono_rev : mul_pos_strict_mono α ↔ mul_pos_mono_rev α := -⟨@zero_lt.mul_pos_strict_mono.to_mul_pos_mono_rev _ _ _ _, - @mul_pos_mono_rev.to_mul_pos_strict_mono _ _ _ _⟩ +⟨@mul_pos_strict_mono.to_mul_pos_mono_rev _ _ _ _, @mul_pos_mono_rev.to_mul_pos_strict_mono _ _ _ _⟩ -lemma pos_mul_reflect_lt.to_pos_mul_mono [pos_mul_reflect_lt α] : pos_mul_mono α := -⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt (lt_of_mul_lt_mul_left' h' x.prop)⟩ +lemma pos_mul_reflect_lt.to_pos_mul_mono {α : Type*} [mul_zero_class α] [linear_order α] + [pos_mul_reflect_lt α] : pos_mul_mono α := +⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt $ lt_of_mul_lt_mul_left h' x.prop⟩ -lemma mul_pos_reflect_lt.to_mul_pos_mono [mul_pos_reflect_lt α] : mul_pos_mono α := -⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt (lt_of_mul_lt_mul_right' h' x.prop)⟩ +lemma mul_pos_reflect_lt.to_mul_pos_mono {α : Type*} [mul_zero_class α] [linear_order α] + [mul_pos_reflect_lt α] : mul_pos_mono α := +⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt $ lt_of_mul_lt_mul_right h' x.prop⟩ lemma pos_mul_mono.to_pos_mul_reflect_lt [pos_mul_mono α] : pos_mul_reflect_lt α := -⟨λ x a b h, lt_of_not_le $ λ h', h.not_le (mul_le_mul_left_of_le_of_pos h' x.prop)⟩ +⟨λ x a b h, lt_of_not_le $ λ h', h.not_le $ mul_le_mul_of_nonneg_left h' x.prop⟩ lemma mul_pos_mono.to_mul_pos_reflect_lt [mul_pos_mono α] : mul_pos_reflect_lt α := -⟨λ x a b h, lt_of_not_le $ λ h', h.not_le (mul_le_mul_right_of_le_of_pos h' x.prop)⟩ +⟨λ x a b h, lt_of_not_le $ λ h', h.not_le $ mul_le_mul_of_nonneg_right h' x.prop⟩ -lemma pos_mul_mono_iff_pos_mul_reflect_lt : pos_mul_mono α ↔ pos_mul_reflect_lt α := -⟨@pos_mul_mono.to_pos_mul_reflect_lt _ _ _ _, @pos_mul_reflect_lt.to_pos_mul_mono _ _ _ _⟩ +lemma pos_mul_mono_iff_pos_mul_reflect_lt {α : Type*} [mul_zero_class α] [linear_order α] : + pos_mul_mono α ↔ pos_mul_reflect_lt α := +⟨@pos_mul_mono.to_pos_mul_reflect_lt _ _ _ _, @pos_mul_reflect_lt.to_pos_mul_mono _ _ _⟩ -lemma mul_pos_mono_iff_mul_pos_reflect_lt : mul_pos_mono α ↔ mul_pos_reflect_lt α := -⟨@mul_pos_mono.to_mul_pos_reflect_lt _ _ _ _, @mul_pos_reflect_lt.to_mul_pos_mono _ _ _ _⟩ +lemma mul_pos_mono_iff_mul_pos_reflect_lt {α : Type*} [mul_zero_class α] [linear_order α] : + mul_pos_mono α ↔ mul_pos_reflect_lt α := +⟨@mul_pos_mono.to_mul_pos_reflect_lt _ _ _ _, @mul_pos_reflect_lt.to_mul_pos_mono _ _ _⟩ end linear_order @@ -365,108 +265,49 @@ section preorder variables [preorder α] /-- Assumes left covariance. -/ -lemma left.mul_pos [pos_mul_strict_mono α] - (ha : 0 < a) (hb : 0 < b) : - 0 < a * b := -have h : a * 0 < a * b, from mul_lt_mul_left' hb ha, -by rwa [mul_zero] at h +lemma left.mul_pos [pos_mul_strict_mono α] (ha : 0 < a) (hb : 0 < b) : 0 < a * b := +by simpa only [mul_zero] using mul_lt_mul_of_pos_left hb ha alias left.mul_pos ← _root_.mul_pos -lemma mul_neg_of_pos_of_neg [pos_mul_strict_mono α] - (ha : 0 < a) (hb : b < 0) : - a * b < 0 := -have h : a * b < a * 0, from mul_lt_mul_left' hb ha, -by rwa [mul_zero] at h +lemma mul_neg_of_pos_of_neg [pos_mul_strict_mono α] (ha : 0 < a) (hb : b < 0) : a * b < 0 := +by simpa only [mul_zero] using mul_lt_mul_of_pos_left hb ha @[simp] lemma zero_lt_mul_left [pos_mul_strict_mono α] [pos_mul_reflect_lt α] (h : 0 < c) : 0 < c * b ↔ 0 < b := by { convert mul_lt_mul_left h, simp } /-- Assumes right covariance. -/ -lemma right.mul_pos [mul_pos_strict_mono α] - (ha : 0 < a) (hb : 0 < b) : - 0 < a * b := -have h : 0 * b < a * b, from mul_lt_mul_right' ha hb, -by rwa [zero_mul] at h - -lemma mul_neg_of_neg_of_pos [mul_pos_strict_mono α] - (ha : a < 0) (hb : 0 < b) : - a * b < 0 := -have h : a * b < 0 * b, from mul_lt_mul_right' ha hb, -by rwa [zero_mul] at h +lemma right.mul_pos [mul_pos_strict_mono α] (ha : 0 < a) (hb : 0 < b) : 0 < a * b := +by simpa only [zero_mul] using mul_lt_mul_of_pos_right ha hb + +lemma mul_neg_of_neg_of_pos [mul_pos_strict_mono α] (ha : a < 0) (hb : 0 < b) : a * b < 0 := +by simpa only [zero_mul] using mul_lt_mul_of_pos_right ha hb @[simp] lemma zero_lt_mul_right [mul_pos_strict_mono α] [mul_pos_reflect_lt α] (h : 0 < c) : 0 < b * c ↔ 0 < b := by { convert mul_lt_mul_right h, simp } -end preorder - -section partial_order -variables [partial_order α] - -lemma mul_le_mul_of_nonneg_left [pos_mul_mono α] - (bc : b ≤ c) (a0 : 0 ≤ a) : - a * b ≤ a * c := -a0.lt_or_eq.elim (mul_le_mul_left_of_le_of_pos bc) (λ h, by simp only [← h, zero_mul]) - -lemma mul_le_mul_of_nonneg_right [mul_pos_mono α] - (bc : b ≤ c) (a0 : 0 ≤ a) : - b * a ≤ c * a := -a0.lt_or_eq.elim (mul_le_mul_right_of_le_of_pos bc) (λ h, by simp only [← h, mul_zero]) - /-- Assumes left covariance. -/ -lemma left.mul_nonneg [pos_mul_mono α] - (ha : 0 ≤ a) (hb : 0 ≤ b) : - 0 ≤ a * b := -have h : a * 0 ≤ a * b, from mul_le_mul_of_nonneg_left hb ha, -by rwa [mul_zero] at h +lemma left.mul_nonneg [pos_mul_mono α] (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := +by simpa only [mul_zero] using mul_le_mul_of_nonneg_left hb ha alias left.mul_nonneg ← mul_nonneg -lemma mul_nonpos_of_nonneg_of_nonpos [pos_mul_mono α] - (ha : 0 ≤ a) (hb : b ≤ 0) : - a * b ≤ 0 := -have h : a * b ≤ a * 0, from mul_le_mul_of_nonneg_left hb ha, -by rwa [mul_zero] at h +lemma mul_nonpos_of_nonneg_of_nonpos [pos_mul_mono α] (ha : 0 ≤ a) (hb : b ≤ 0) : a * b ≤ 0 := +by simpa only [mul_zero] using mul_le_mul_of_nonneg_left hb ha /-- Assumes right covariance. -/ -lemma right.mul_nonneg [mul_pos_mono α] - (ha : 0 ≤ a) (hb : 0 ≤ b) : - 0 ≤ a * b := -have h : 0 * b ≤ a * b, from mul_le_mul_of_nonneg_right ha hb, -by rwa [zero_mul] at h - -lemma mul_nonpos_of_nonpos_of_nonneg [mul_pos_mono α] - (ha : a ≤ 0) (hb : 0 ≤ b) : - a * b ≤ 0 := -have h : a * b ≤ 0 * b, from mul_le_mul_of_nonneg_right ha hb, -by rwa [zero_mul] at h - -lemma lt_of_mul_lt_mul_left [pos_mul_reflect_lt α] - (bc : a * b < a * c) (a0 : 0 ≤ a) : - b < c := -begin - by_cases a₀ : a = 0, - { exact (lt_irrefl (0 : α) (by simpa only [a₀, zero_mul] using bc)).elim }, - { exact lt_of_mul_lt_mul_left' bc ((ne.symm a₀).le_iff_lt.mp a0) } -end +lemma right.mul_nonneg [mul_pos_mono α] (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := +by simpa only [zero_mul] using mul_le_mul_of_nonneg_right ha hb -lemma pos_of_mul_pos_right [pos_mul_reflect_lt α] (h : 0 < a * b) (ha : 0 ≤ a) : - 0 < b := -lt_of_mul_lt_mul_left ((mul_zero a).symm ▸ h : a * 0 < a * b) ha +lemma mul_nonpos_of_nonpos_of_nonneg [mul_pos_mono α] (ha : a ≤ 0) (hb : 0 ≤ b) : a * b ≤ 0 := +by simpa only [zero_mul] using mul_le_mul_of_nonneg_right ha hb -lemma lt_of_mul_lt_mul_right [mul_pos_reflect_lt α] - (ab : a * c < b * c) (c0 : 0 ≤ c) : - a < b := -begin - by_cases c₀ : c = 0, - { exact (lt_irrefl (0 : α) (by simpa only [c₀, mul_zero] using ab)).elim }, - { exact lt_of_mul_lt_mul_right' ab ((ne.symm c₀).le_iff_lt.mp c0) } -end +lemma pos_of_mul_pos_right [pos_mul_reflect_lt α] (h : 0 < a * b) (ha : 0 ≤ a) : 0 < b := +lt_of_mul_lt_mul_left ((mul_zero a).symm ▸ h : a * 0 < a * b) ha -lemma pos_of_mul_pos_left [mul_pos_reflect_lt α] (h : 0 < a * b) (hb : 0 ≤ b) : - 0 < a := +lemma pos_of_mul_pos_left [mul_pos_reflect_lt α] (h : 0 < a * b) (hb : 0 ≤ b) : 0 < a := lt_of_mul_lt_mul_right ((zero_mul b).symm ▸ h : 0 * b < a * b) hb lemma pos_iff_pos_of_mul_pos [pos_mul_reflect_lt α] [mul_pos_reflect_lt α] (hab : 0 < a * b) : @@ -481,37 +322,84 @@ lemma mul_le_mul [pos_mul_mono α] [mul_pos_mono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (c0 : 0 ≤ c) (b0 : 0 ≤ b) : a * c ≤ b * d := (mul_le_mul_of_nonneg_right h₁ c0).trans $ mul_le_mul_of_nonneg_left h₂ b0 -lemma mul_le_of_mul_le_left [pos_mul_mono α] - (h : a * b ≤ c) (hle : d ≤ b) (a0 : 0 ≤ a) : +lemma mul_le_of_mul_le_of_nonneg_left [pos_mul_mono α] (h : a * b ≤ c) (hle : d ≤ b) (a0 : 0 ≤ a) : a * d ≤ c := (mul_le_mul_of_nonneg_left hle a0).trans h -lemma le_mul_of_le_mul_left [pos_mul_mono α] - (h : a ≤ b * c) (hle : c ≤ d) (b0 : 0 ≤ b) : +lemma le_mul_of_le_mul_of_nonneg_left [pos_mul_mono α] (h : a ≤ b * c) (hle : c ≤ d) (b0 : 0 ≤ b) : a ≤ b * d := h.trans (mul_le_mul_of_nonneg_left hle b0) -lemma mul_le_of_mul_le_right [mul_pos_mono α] - (h : a * b ≤ c) (hle : d ≤ a) (b0 : 0 ≤ b) : +lemma mul_le_of_mul_le_of_nonneg_right [mul_pos_mono α] (h : a * b ≤ c) (hle : d ≤ a) (b0 : 0 ≤ b) : d * b ≤ c := (mul_le_mul_of_nonneg_right hle b0).trans h -lemma le_mul_of_le_mul_right [mul_pos_mono α] - (h : a ≤ b * c) (hle : b ≤ d) (c0 : 0 ≤ c) : +lemma le_mul_of_le_mul_of_nonneg_right [mul_pos_mono α] (h : a ≤ b * c) (hle : b ≤ d) (c0 : 0 ≤ c) : a ≤ d * c := h.trans (mul_le_mul_of_nonneg_right hle c0) -lemma mul_left_cancel_iff_of_pos [pos_mul_mono_rev α] - (a0 : 0 < a) : - a * b = a * c ↔ b = c := -⟨λ h, (le_of_mul_le_mul_left' h.le a0).antisymm (le_of_mul_le_mul_left' h.ge a0), congr_arg _⟩ +end preorder -lemma mul_right_cancel_iff_of_pos [mul_pos_mono_rev α] - (b0 : 0 < b) : - a * b = c * b ↔ a = c := -⟨λ h, (le_of_mul_le_mul_right' h.le b0).antisymm (le_of_mul_le_mul_right' h.ge b0), congr_arg _⟩ +section partial_order +variables [partial_order α] + +lemma pos_mul_mono_iff_covariant_pos : pos_mul_mono α ↔ covariant_class α>0 α (λ x y, x * y) (≤) := +⟨@pos_mul_mono.to_covariant_class_pos_mul_le _ _ _ _, λ h, ⟨λ a b c h, begin + obtain ha | ha := a.prop.eq_or_gt, + { simp only [ha, zero_mul] }, + { exactI @covariant_class.elim α>0 α (λ x y, x * y) (≤) _ ⟨_, ha⟩ _ _ h } + end⟩⟩ + +lemma mul_pos_mono_iff_covariant_pos : mul_pos_mono α ↔ covariant_class α>0 α (λ x y, y * x) (≤) := +⟨@mul_pos_mono.to_covariant_class_pos_mul_le _ _ _ _, λ h, ⟨λ a b c h, begin + obtain ha | ha := a.prop.eq_or_gt, + { simp only [ha, mul_zero] }, + { exactI @covariant_class.elim α>0 α (λ x y, y * x) (≤) _ ⟨_, ha⟩ _ _ h } + end⟩⟩ + +lemma pos_mul_reflect_lt_iff_contravariant_pos : + pos_mul_reflect_lt α ↔ contravariant_class α>0 α (λ x y, x * y) (<) := +⟨@pos_mul_reflect_lt.to_contravariant_class_pos_mul_lt _ _ _ _, λ h, ⟨λ a b c h, begin + obtain ha | ha := a.prop.eq_or_gt, + { simpa [ha] using h }, + { exactI (@contravariant_class.elim α>0 α (λ x y, x * y) (<) _ ⟨_, ha⟩ _ _ h) } + end⟩⟩ + +lemma mul_pos_reflect_lt_iff_contravariant_pos : + mul_pos_reflect_lt α ↔ contravariant_class α>0 α (λ x y, y * x) (<) := +⟨@mul_pos_reflect_lt.to_contravariant_class_pos_mul_lt _ _ _ _, λ h, ⟨λ a b c h, begin + obtain ha | ha := a.prop.eq_or_gt, + { simpa [ha] using h }, + { exactI (@contravariant_class.elim α>0 α (λ x y, y * x) (<) _ ⟨_, ha⟩ _ _ h) } + end⟩⟩ + +@[priority 100] -- see Note [lower instance priority] +instance pos_mul_strict_mono.to_pos_mul_mono [pos_mul_strict_mono α] : pos_mul_mono α := +pos_mul_mono_iff_covariant_pos.2 $ ⟨λ a, strict_mono.monotone $ @covariant_class.elim _ _ _ _ _ _⟩ -lemma mul_eq_mul_iff_eq_and_eq [pos_mul_strict_mono α] [mul_pos_strict_mono α] +@[priority 100] -- see Note [lower instance priority] +instance mul_pos_strict_mono.to_mul_pos_mono [mul_pos_strict_mono α] : mul_pos_mono α := +mul_pos_mono_iff_covariant_pos.2 $ ⟨λ a, strict_mono.monotone $ @covariant_class.elim _ _ _ _ _ _⟩ + +@[priority 100] -- see Note [lower instance priority] +instance pos_mul_mono_rev.to_pos_mul_reflect_lt [pos_mul_mono_rev α] : pos_mul_reflect_lt α := +pos_mul_reflect_lt_iff_contravariant_pos.2 + ⟨λ a b c h, (le_of_mul_le_mul_of_pos_left h.le a.2).lt_of_ne $ by { rintro rfl, simpa using h }⟩ + +@[priority 100] -- see Note [lower instance priority] +instance mul_pos_mono_rev.to_mul_pos_reflect_lt [mul_pos_mono_rev α] : mul_pos_reflect_lt α := +mul_pos_reflect_lt_iff_contravariant_pos.2 + ⟨λ a b c h, (le_of_mul_le_mul_of_pos_right h.le a.2).lt_of_ne $ by { rintro rfl, simpa using h }⟩ + +lemma mul_left_cancel_iff_of_pos [pos_mul_mono_rev α] (a0 : 0 < a) : a * b = a * c ↔ b = c := +⟨λ h, (le_of_mul_le_mul_of_pos_left h.le a0).antisymm $ le_of_mul_le_mul_of_pos_left h.ge a0, + congr_arg _⟩ + +lemma mul_right_cancel_iff_of_pos [mul_pos_mono_rev α] (b0 : 0 < b) : a * b = c * b ↔ a = c := +⟨λ h, (le_of_mul_le_mul_of_pos_right h.le b0).antisymm $ le_of_mul_le_mul_of_pos_right h.ge b0, + congr_arg _⟩ + +lemma mul_eq_mul_iff_eq_and_eq_of_pos [pos_mul_strict_mono α] [mul_pos_strict_mono α] [pos_mul_mono_rev α] [mul_pos_mono_rev α] (hac : a ≤ b) (hbd : c ≤ d) (a0 : 0 < a) (d0 : 0 < d) : a * c = b * d ↔ a = b ∧ c = d := @@ -521,10 +409,10 @@ begin { exact ⟨rfl, (mul_left_cancel_iff_of_pos a0).mp h⟩ }, rcases eq_or_lt_of_le hbd with rfl | hbd, { exact ⟨(mul_right_cancel_iff_of_pos d0).mp h, rfl⟩ }, - exact ((mul_lt_mul_of_lt_of_lt hac hbd a0 d0).ne h).elim, + exact ((mul_lt_mul_of_pos_of_pos hac hbd a0 d0).ne h).elim, end -lemma mul_eq_mul_iff_eq_and_eq' [pos_mul_strict_mono α] [mul_pos_strict_mono α] +lemma mul_eq_mul_iff_eq_and_eq_of_pos' [pos_mul_strict_mono α] [mul_pos_strict_mono α] [pos_mul_mono_rev α] [mul_pos_mono_rev α] (hac : a ≤ b) (hbd : c ≤ d) (b0 : 0 < b) (c0 : 0 < c) : a * c = b * d ↔ a = b ∧ c = d := @@ -659,588 +547,263 @@ lemma mul_lt_iff_lt_one_left a * b < b ↔ a < 1 := iff.trans (by rw [one_mul]) (mul_lt_mul_right b0) -/-! Lemmas of the form `b ≤ c → a ≤ 1 → b * a ≤ c`. -/ - --- proven with `b0 : 0 ≤ b` as `mul_le_of_le_of_le_one` -lemma preorder.mul_le_of_le_of_le_one [pos_mul_mono α] - (bc : b ≤ c) (ha : a ≤ 1) (b0 : 0 < b) : b * a ≤ c := -calc b * a ≤ b * 1 : mul_le_mul_left_of_le_of_pos ha b0 - ... = b : mul_one b - ... ≤ c : bc - -lemma mul_lt_of_le_of_lt_one [pos_mul_strict_mono α] - (bc : b ≤ c) (ha : a < 1) (b0 : 0 < b) : b * a < c := -calc b * a < b * 1 : mul_lt_mul_left' ha b0 - ... = b : mul_one b - ... ≤ c : bc - -lemma mul_lt_of_lt_of_le_one [pos_mul_mono α] - (bc : b < c) (ha : a ≤ 1) (b0 : 0 < b) : b * a < c := -calc b * a ≤ b * 1 : mul_le_mul_left_of_le_of_pos ha b0 - ... = b : mul_one b - ... < c : bc - -lemma mul_lt_of_lt_of_lt_one [pos_mul_strict_mono α] - (bc : b < c) (ha : a < 1) (b0 : 0 < b) : b * a < c := -calc b * a < b * 1 : mul_lt_mul_left' ha b0 - ... = b : mul_one b - ... < c : bc - --- proven with `a0 : 0 ≤ a` as `left.mul_le_one_of_le_of_le` -/-- Assumes left covariance. -/ -lemma preorder.left.mul_le_one_of_le_of_le' [pos_mul_mono α] - (ha : a ≤ 1) (hb : b ≤ 1) (a0 : 0 < a) : a * b ≤ 1 := -preorder.mul_le_of_le_of_le_one ha hb a0 - -/-- Assumes left covariance. -/ -lemma left.mul_lt_one_of_le_of_lt [pos_mul_strict_mono α] - (ha : a ≤ 1) (hb : b < 1) (a0 : 0 < a) : a * b < 1 := -mul_lt_of_le_of_lt_one ha hb a0 - -/-- Assumes left covariance. -/ -lemma left.mul_lt_one_of_lt_of_le [pos_mul_mono α] - (ha : a < 1) (hb : b ≤ 1) (a0 : 0 < a) : a * b < 1 := -mul_lt_of_lt_of_le_one ha hb a0 - -/-- Assumes left covariance. -/ -lemma left.mul_lt_one_of_lt_of_lt [pos_mul_strict_mono α] - (ha : a < 1) (hb : b < 1) (a0 : 0 < a) : a * b < 1 := -mul_lt_of_lt_of_lt_one ha hb a0 - --- proven with `a0 : 0 ≤ a` and `c0 : 0 ≤ c` as `mul_le_of_le_of_le_one'` -lemma preorder.mul_le_of_le_of_le_one' [pos_mul_mono α] [mul_pos_mono α] - (bc : b ≤ c) (ha : a ≤ 1) (a0 : 0 < a) (c0 : 0 < c) : b * a ≤ c := -calc b * a ≤ c * a : mul_le_mul_right_of_le_of_pos bc a0 - ... ≤ c * 1 : mul_le_mul_left_of_le_of_pos ha c0 - ... = c : mul_one c - --- proven with `c0 : 0 ≤ c` as `mul_lt_of_lt_of_le_one'` -lemma preorder.mul_lt_of_lt_of_le_one' [pos_mul_mono α] [mul_pos_strict_mono α] - (bc : b < c) (ha : a ≤ 1) (a0 : 0 < a) (c0 : 0 < c) : b * a < c := -calc b * a < c * a : mul_lt_mul_right' bc a0 - ... ≤ c * 1 : mul_le_mul_left_of_le_of_pos ha c0 - ... = c : mul_one c - --- proven with `a0 : 0 ≤ a` as `mul_lt_of_le_of_lt_one'` -lemma preorder.mul_lt_of_le_of_lt_one' [pos_mul_strict_mono α] [mul_pos_mono α] - (bc : b ≤ c) (ha : a < 1) (a0 : 0 < a) (c0 : 0 < c) : b * a < c := -calc b * a ≤ c * a : mul_le_mul_right_of_le_of_pos bc a0 - ... < c * 1 : mul_lt_mul_left' ha c0 - ... = c : mul_one c - -lemma mul_lt_of_lt_of_lt_one' [pos_mul_strict_mono α] [mul_pos_strict_mono α] - (bc : b < c) (ha : a < 1) (a0 : 0 < a) (c0 : 0 < c) : b * a < c := -calc b * a < c * a : mul_lt_mul_right' bc a0 - ... < c * 1 : mul_lt_mul_left' ha c0 - ... = c : mul_one c - -/-! Lemmas of the form `b ≤ c → 1 ≤ a → b ≤ c * a`. -/ - --- proven with `c0 : 0 ≤ c` as `le_mul_of_le_of_one_le` -lemma preorder.le_mul_of_le_of_one_le [pos_mul_mono α] - (bc : b ≤ c) (ha : 1 ≤ a) (c0 : 0 < c) : b ≤ c * a := -calc b ≤ c : bc - ... = c * 1 : (mul_one c).symm - ... ≤ c * a : mul_le_mul_left_of_le_of_pos ha c0 - -lemma lt_mul_of_le_of_one_lt [pos_mul_strict_mono α] - (bc : b ≤ c) (ha : 1 < a) (c0 : 0 < c) : b < c * a := -calc b ≤ c : bc - ... = c * 1 : (mul_one c).symm - ... < c * a : mul_lt_mul_left' ha c0 - -lemma lt_mul_of_lt_of_one_le [pos_mul_mono α] - (bc : b < c) (ha : 1 ≤ a) (c0 : 0 < c) : b < c * a := -calc b < c : bc - ... = c * 1 : (mul_one c).symm - ... ≤ c * a : mul_le_mul_left_of_le_of_pos ha c0 - -lemma lt_mul_of_lt_of_one_lt [pos_mul_strict_mono α] - (bc : b < c) (ha : 1 < a) (c0 : 0 < c) : b < c * a := -calc b < c : bc - ... = c * 1 : (mul_one _).symm - ... < c * a : mul_lt_mul_left' ha c0 - --- proven with `a0 : 0 ≤ a` as `left.one_le_mul_of_le_of_le` -/-- Assumes left covariance. -/ -lemma preorder.left.one_le_mul_of_le_of_le [pos_mul_mono α] - (ha : 1 ≤ a) (hb : 1 ≤ b) (a0 : 0 < a) : 1 ≤ a * b := -preorder.le_mul_of_le_of_one_le ha hb a0 - -/-- Assumes left covariance. -/ -lemma left.one_lt_mul_of_le_of_lt [pos_mul_strict_mono α] - (ha : 1 ≤ a) (hb : 1 < b) (a0 : 0 < a) : 1 < a * b := -lt_mul_of_le_of_one_lt ha hb a0 - -/-- Assumes left covariance. -/ -lemma left.one_lt_mul_of_lt_of_le [pos_mul_mono α] - (ha : 1 < a) (hb : 1 ≤ b) (a0 : 0 < a) : 1 < a * b := -lt_mul_of_lt_of_one_le ha hb a0 +/-! Lemmas of the form `1 ≤ b → a ≤ a * b`. -/ -/-- Assumes left covariance. -/ -lemma left.one_lt_mul_of_lt_of_lt [pos_mul_strict_mono α] - (ha : 1 < a) (hb : 1 < b) (a0 : 0 < a) : 1 < a * b := -lt_mul_of_lt_of_one_lt ha hb a0 - --- proven with `a0 : 0 ≤ a` and `b0 : 0 ≤ b` as `le_mul_of_le_of_one_le'` -lemma preorder.le_mul_of_le_of_one_le' [pos_mul_mono α] [mul_pos_mono α] - (bc : b ≤ c) (ha : 1 ≤ a) (a0 : 0 < a) (b0 : 0 < b) : b ≤ c * a := -calc b = b * 1 : (mul_one b).symm - ... ≤ b * a : mul_le_mul_left_of_le_of_pos ha b0 - ... ≤ c * a : mul_le_mul_right_of_le_of_pos bc a0 - --- proven with `a0 : 0 ≤ a` as `lt_mul_of_le_of_one_lt'` -lemma preorder.lt_mul_of_le_of_one_lt' [pos_mul_strict_mono α] [mul_pos_mono α] - (bc : b ≤ c) (ha : 1 < a) (a0 : 0 < a) (b0 : 0 < b) : b < c * a := -calc b = b * 1 : (mul_one b).symm - ... < b * a : mul_lt_mul_left' ha b0 - ... ≤ c * a : mul_le_mul_right_of_le_of_pos bc a0 - --- proven with `b0 : 0 ≤ b` as `lt_mul_of_lt_of_one_le'` -lemma preorder.lt_mul_of_lt_of_one_le' [pos_mul_mono α] [mul_pos_strict_mono α] - (bc : b < c) (ha : 1 ≤ a) (a0 : 0 < a) (b0 : 0 < b) : b < c * a := -calc b = b * 1 : (mul_one b).symm - ... ≤ b * a : mul_le_mul_left_of_le_of_pos ha b0 - ... < c * a : mul_lt_mul_right' bc a0 - -lemma lt_mul_of_lt_of_one_lt' [pos_mul_strict_mono α] [mul_pos_strict_mono α] - (bc : b < c) (ha : 1 < a) (a0 : 0 < a) (b0 : 0 < b) : b < c * a := -calc b = b * 1 : (mul_one b).symm - ... < b * a : mul_lt_mul_left' ha b0 - ... < c * a : mul_lt_mul_right' bc a0 +lemma mul_le_of_le_one_left [mul_pos_mono α] (hb : 0 ≤ b) (h : a ≤ 1) : a * b ≤ b := +by simpa only [one_mul] using mul_le_mul_of_nonneg_right h hb -/-! Lemmas of the form `a ≤ 1 → b ≤ c → a * b ≤ c`. -/ +lemma le_mul_of_one_le_left [mul_pos_mono α] (hb : 0 ≤ b) (h : 1 ≤ a) : b ≤ a * b := +by simpa only [one_mul] using mul_le_mul_of_nonneg_right h hb --- proven with `b0 : 0 ≤ b` as `mul_le_of_le_one_of_le` -lemma preorder.mul_le_of_le_one_of_le [mul_pos_mono α] - (ha : a ≤ 1) (bc : b ≤ c) (b0 : 0 < b) : a * b ≤ c := -calc a * b ≤ 1 * b : mul_le_mul_right_of_le_of_pos ha b0 - ... = b : one_mul b - ... ≤ c : bc - -lemma mul_lt_of_lt_one_of_le [mul_pos_strict_mono α] - (ha : a < 1) (bc : b ≤ c) (b0 : 0 < b) : a * b < c := -calc a * b < 1 * b : mul_lt_mul_right' ha b0 - ... = b : one_mul b - ... ≤ c : bc - -lemma mul_lt_of_le_one_of_lt [mul_pos_mono α] - (ha : a ≤ 1) (hb : b < c) (b0 : 0 < b) : a * b < c := -calc a * b ≤ 1 * b : mul_le_mul_right_of_le_of_pos ha b0 - ... = b : one_mul b - ... < c : hb - -lemma mul_lt_of_lt_one_of_lt [mul_pos_strict_mono α] - (ha : a < 1) (bc : b < c) (b0 : 0 < b) : a * b < c := -calc a * b < 1 * b : mul_lt_mul_right' ha b0 - ... = b : one_mul b - ... < c : bc - --- proven with `b0 : 0 ≤ b` as `right.mul_le_one_of_le_of_le` -/-- Assumes right covariance. -/ -lemma preorder.right.mul_le_one_of_le_of_le [mul_pos_mono α] - (ha : a ≤ 1) (hb : b ≤ 1) (b0 : 0 < b) : a * b ≤ 1 := -preorder.mul_le_of_le_one_of_le ha hb b0 +lemma mul_le_of_le_one_right [pos_mul_mono α] (ha : 0 ≤ a) (h : b ≤ 1) : a * b ≤ a := +by simpa only [mul_one] using mul_le_mul_of_nonneg_left h ha -/-- Assumes right covariance. -/ -lemma right.mul_lt_one_of_lt_of_le [mul_pos_strict_mono α] - (ha : a < 1) (hb : b ≤ 1) (b0 : 0 < b) : a * b < 1 := -mul_lt_of_lt_one_of_le ha hb b0 +lemma le_mul_of_one_le_right [pos_mul_mono α] (ha : 0 ≤ a) (h : 1 ≤ b) : a ≤ a * b := +by simpa only [mul_one] using mul_le_mul_of_nonneg_left h ha -/-- Assumes right covariance. -/ -lemma right.mul_lt_one_of_le_of_lt [mul_pos_mono α] - (ha : a ≤ 1) (hb : b < 1) (b0 : 0 < b) : a * b < 1 := -mul_lt_of_le_one_of_lt ha hb b0 +lemma mul_lt_of_lt_one_left [mul_pos_strict_mono α] (hb : 0 < b) (h : a < 1) : a * b < b := +by simpa only [one_mul] using mul_lt_mul_of_pos_right h hb -/-- Assumes right covariance. -/ -lemma right.mul_lt_one_of_lt_of_lt [mul_pos_strict_mono α] - (ha : a < 1) (hb : b < 1) (b0 : 0 < b) : a * b < 1 := -mul_lt_of_lt_one_of_lt ha hb b0 - --- proven with `a0 : 0 ≤ a` and `c0 : 0 ≤ c` as `mul_le_of_le_one_of_le'` -lemma preorder.mul_le_of_le_one_of_le' [pos_mul_mono α] [mul_pos_mono α] - (ha : a ≤ 1) (bc : b ≤ c) (a0 : 0 < a) (c0 : 0 < c) : a * b ≤ c := -calc a * b ≤ a * c : mul_le_mul_left_of_le_of_pos bc a0 - ... ≤ 1 * c : mul_le_mul_right_of_le_of_pos ha c0 - ... = c : one_mul c - --- proven with `a0 : 0 ≤ a` as `mul_lt_of_lt_one_of_le'` -lemma preorder.mul_lt_of_lt_one_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] - (ha : a < 1) (bc : b ≤ c) (a0 : 0 < a) (c0 : 0 < c) : a * b < c := -calc a * b ≤ a * c : mul_le_mul_left_of_le_of_pos bc a0 - ... < 1 * c : mul_lt_mul_right' ha c0 - ... = c : one_mul c - --- proven with `c0 : 0 ≤ c` as `mul_lt_of_le_one_of_lt'` -lemma preorder.mul_lt_of_le_one_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] - (ha : a ≤ 1) (bc : b < c) (a0 : 0 < a) (c0 : 0 < c) : a * b < c := -calc a * b < a * c : mul_lt_mul_left' bc a0 - ... ≤ 1 * c : mul_le_mul_right_of_le_of_pos ha c0 - ... = c : one_mul c - -lemma mul_lt_of_lt_one_of_lt' [pos_mul_strict_mono α] [mul_pos_strict_mono α] - (ha : a < 1) (bc : b < c) (a0 : 0 < a) (c0 : 0 < c) : a * b < c := -calc a * b < a * c : mul_lt_mul_left' bc a0 - ... < 1 * c : mul_lt_mul_right' ha c0 - ... = c : one_mul c +lemma lt_mul_of_one_lt_left [mul_pos_strict_mono α] (hb : 0 < b) (h : 1 < a) : b < a * b := +by simpa only [one_mul] using mul_lt_mul_of_pos_right h hb -/-! Lemmas of the form `1 ≤ a → b ≤ c → b ≤ a * c`. -/ +lemma mul_lt_of_lt_one_right [pos_mul_strict_mono α] (ha : 0 < a) (h : b < 1) : a * b < a := +by simpa only [mul_one] using mul_lt_mul_of_pos_left h ha --- proven with `c0 : 0 ≤ c` as `le_mul_of_one_le_of_le` -lemma preorder.le_mul_of_one_le_of_le [mul_pos_mono α] - (ha : 1 ≤ a) (bc : b ≤ c) (c0 : 0 < c) : b ≤ a * c := -calc b ≤ c : bc - ... = 1 * c : (one_mul c).symm - ... ≤ a * c : mul_le_mul_right_of_le_of_pos ha c0 - -lemma lt_mul_of_one_lt_of_le [mul_pos_strict_mono α] - (ha : 1 < a) (bc : b ≤ c) (c0 : 0 < c) : b < a * c := -calc b ≤ c : bc - ... = 1 * c : (one_mul c).symm - ... < a * c : mul_lt_mul_right' ha c0 - -lemma lt_mul_of_one_le_of_lt [mul_pos_mono α] - (ha : 1 ≤ a) (bc : b < c) (c0 : 0 < c) : b < a * c := -calc b < c : bc - ... = 1 * c : (one_mul c).symm - ... ≤ a * c : mul_le_mul_right_of_le_of_pos ha c0 - -lemma lt_mul_of_one_lt_of_lt [mul_pos_strict_mono α] - (ha : 1 < a) (bc : b < c) (c0 : 0 < c) : b < a * c := -calc b < c : bc - ... = 1 * c : (one_mul c).symm - ... < a * c : mul_lt_mul_right' ha c0 - --- proven with `b0 : 0 ≤ b` as `right.one_le_mul_of_le_of_le` -/-- Assumes right covariance. -/ -lemma preorder.right.one_le_mul_of_le_of_le [mul_pos_mono α] - (ha : 1 ≤ a) (hb : 1 ≤ b) (b0 : 0 < b) : 1 ≤ a * b := -preorder.le_mul_of_one_le_of_le ha hb b0 +lemma lt_mul_of_one_lt_right [pos_mul_strict_mono α] (ha : 0 < a) (h : 1 < b) : a < a * b := +by simpa only [mul_one] using mul_lt_mul_of_pos_left h ha -/-- Assumes right covariance. -/ -lemma right.one_lt_mul_of_lt_of_le [mul_pos_strict_mono α] - (ha : 1 < a) (hb : 1 ≤ b) (b0 : 0 < b) : 1 < a * b := -lt_mul_of_one_lt_of_le ha hb b0 - -/-- Assumes right covariance. -/ -lemma right.one_lt_mul_of_le_of_lt [mul_pos_mono α] - (ha : 1 ≤ a) (hb : 1 < b) (b0 : 0 < b) : 1 < a * b := -lt_mul_of_one_le_of_lt ha hb b0 - -/-- Assumes right covariance. -/ -lemma right.one_lt_mul_of_lt_of_lt [mul_pos_strict_mono α] - (ha : 1 < a) (hb : 1 < b) (b0 : 0 < b) : 1 < a * b := -lt_mul_of_one_lt_of_lt ha hb b0 - --- proven with `a0 : 0 ≤ a` and `b0 : 0 ≤ b` as `le_mul_of_one_le_of_le'` -lemma preorder.le_mul_of_one_le_of_le' [pos_mul_mono α] [mul_pos_mono α] - (ha : 1 ≤ a) (bc : b ≤ c) (a0 : 0 < a) (b0 : 0 < b) : b ≤ a * c := -calc b = 1 * b : (one_mul b).symm - ... ≤ a * b : mul_le_mul_right_of_le_of_pos ha b0 - ... ≤ a * c : mul_le_mul_left_of_le_of_pos bc a0 - --- proven with `a0 : 0 ≤ a` as `lt_mul_of_one_lt_of_le'` -lemma preorder.lt_mul_of_one_lt_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] - (ha : 1 < a) (bc : b ≤ c) (a0 : 0 < a) (b0 : 0 < b) : b < a * c := -calc b = 1 * b : (one_mul b).symm - ... < a * b : mul_lt_mul_right' ha b0 - ... ≤ a * c : mul_le_mul_left_of_le_of_pos bc a0 - --- proven with `b0 : 0 ≤ b` as `lt_mul_of_one_le_of_lt'` -lemma preorder.lt_mul_of_one_le_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] - (ha : 1 ≤ a) (bc : b < c) (a0 : 0 < a) (b0 : 0 < b) : b < a * c := -calc b = 1 * b : (one_mul b).symm - ... ≤ a * b : mul_le_mul_right_of_le_of_pos ha b0 - ... < a * c : mul_lt_mul_left' bc a0 - -lemma lt_mul_of_one_lt_of_lt' [pos_mul_strict_mono α] [mul_pos_strict_mono α] - (ha : 1 < a) (bc : b < c) (a0 : 0 < a) (b0 : 0 < b) : b < a * c := -calc b = 1 * b : (one_mul b).symm - ... < a * b : mul_lt_mul_right' ha b0 - ... < a * c : mul_lt_mul_left' bc a0 - --- proven with `a0 : 0 ≤ a` as `mul_le_of_le_one_right` -lemma preorder.mul_le_of_le_one_right [pos_mul_mono α] (h : b ≤ 1) (a0 : 0 < a) : - a * b ≤ a := -preorder.mul_le_of_le_of_le_one le_rfl h a0 - --- proven with `a0 : 0 ≤ a` as `le_mul_of_one_le_right` -lemma preorder.le_mul_of_one_le_right [pos_mul_mono α] (h : 1 ≤ b) (a0 : 0 < a) : - a ≤ a * b := -preorder.le_mul_of_le_of_one_le le_rfl h a0 - --- proven with `b0 : 0 ≤ b` as `mul_le_of_le_one_left` -lemma preorder.mul_le_of_le_one_left [mul_pos_mono α] (h : a ≤ 1) (b0 : 0 < b) : - a * b ≤ b := -preorder.mul_le_of_le_one_of_le h le_rfl b0 - --- proven with `b0 : 0 ≤ b` as `le_mul_of_one_le_left` -lemma preorder.le_mul_of_one_le_left [mul_pos_mono α] (h : 1 ≤ a) (b0 : 0 < b) : - b ≤ a * b := -preorder.le_mul_of_one_le_of_le h le_rfl b0 - -lemma mul_lt_of_lt_one_right [pos_mul_strict_mono α] (a0 : 0 < a) (h : b < 1) : - a * b < a := -mul_lt_of_le_of_lt_one le_rfl h a0 - -lemma lt_mul_of_one_lt_right [pos_mul_strict_mono α] (a0 : 0 < a) (h : 1 < b) : - a < a * b := -lt_mul_of_le_of_one_lt le_rfl h a0 - -lemma mul_lt_of_lt_one_left [mul_pos_strict_mono α] (b0 : 0 < b) (h : a < 1) : - a * b < b := -mul_lt_of_lt_one_of_le h le_rfl b0 - -lemma lt_mul_of_one_lt_left [mul_pos_strict_mono α] (b0 : 0 < b) (h : 1 < a) : - b < a * b := -lt_mul_of_one_lt_of_le h le_rfl b0 - --- proven with `a0 : 0 ≤ a` as `le_of_mul_le_of_one_le_left` -lemma preorder.le_of_mul_le_of_one_le_left [pos_mul_mono α] - (h : a * b ≤ c) (hle : 1 ≤ b) (a0 : 0 < a) : - a ≤ c := -(preorder.le_mul_of_one_le_right hle a0).trans h - -lemma lt_of_mul_lt_of_one_le_left [pos_mul_mono α] - (h : a * b < c) (hle : 1 ≤ b) (a0 : 0 < a) : - a < c := -(preorder.le_mul_of_one_le_right hle a0).trans_lt h - --- proven with `b0 : 0 ≤ b` as `le_of_le_mul_of_le_one_left` -lemma preorder.le_of_le_mul_of_le_one_left [pos_mul_mono α] - (h : a ≤ b * c) (hle : c ≤ 1) (b0 : 0 < b) : - a ≤ b := -h.trans (preorder.mul_le_of_le_one_right hle b0) - -lemma lt_of_lt_mul_of_le_one_left [pos_mul_mono α] - (h : a < b * c) (hle : c ≤ 1) (b0 : 0 < b) : - a < b := -h.trans_le (preorder.mul_le_of_le_one_right hle b0) - --- proven with `b0 : 0 ≤ b` as `le_of_mul_le_of_one_le_right` -lemma preorder.le_of_mul_le_of_one_le_right [mul_pos_mono α] - (h : a * b ≤ c) (hle : 1 ≤ a) (b0 : 0 < b) : - b ≤ c := -(preorder.le_mul_of_one_le_left hle b0).trans h - -lemma lt_of_mul_lt_of_one_le_right [mul_pos_mono α] - (h : a * b < c) (hle : 1 ≤ a) (b0 : 0 < b) : - b < c := -(preorder.le_mul_of_one_le_left hle b0).trans_lt h - --- proven with `c0 : 0 ≤ b` as `le_of_le_mul_of_le_one_right` -lemma preorder.le_of_le_mul_of_le_one_right [mul_pos_mono α] - (h : a ≤ b * c) (hle : b ≤ 1) (c0 : 0 < c) : - a ≤ c := -h.trans (preorder.mul_le_of_le_one_left hle c0) - -lemma lt_of_lt_mul_of_le_one_right [mul_pos_mono α] - (h : a < b * c) (hle : b ≤ 1) (c0 : 0 < c) : - a < c := -h.trans_le (preorder.mul_le_of_le_one_left hle c0) - -end preorder - -section linear_order -variables [linear_order α] - --- proven with `a0 : 0 ≤ a` as `exists_square_le` -lemma exists_square_le' [pos_mul_strict_mono α] - (a0 : 0 < a) : ∃ (b : α), b * b ≤ a := -begin - by_cases h : a < 1, - { use a, - have : a*a < a*1, - exact mul_lt_mul_left' h a0, - rw mul_one at this, - exact le_of_lt this }, - { use 1, - push_neg at h, - rwa mul_one } -end - -end linear_order +/-! Lemmas of the form `b ≤ c → a ≤ 1 → b * a ≤ c`. -/ +/- Yaël: What's the point of these lemmas? They just chain an existing lemma with an assumption in +all possible ways, thereby artificially inflating the API and making the truly relevant lemmas hard +to find -/ -end mul_one_class +lemma mul_le_of_le_of_le_one_of_nonneg [pos_mul_mono α] (h : b ≤ c) (ha : a ≤ 1) (hb : 0 ≤ b) : + b * a ≤ c := +(mul_le_of_le_one_right hb ha).trans h -section mul_zero_one_class -variables [mul_zero_one_class α] +lemma mul_lt_of_le_of_lt_one_of_pos [pos_mul_strict_mono α] (bc : b ≤ c) (ha : a < 1) (b0 : 0 < b) : + b * a < c := +(mul_lt_of_lt_one_right b0 ha).trans_le bc -section partial_order -variables [partial_order α] +lemma mul_lt_of_lt_of_le_one_of_nonneg [pos_mul_mono α] (h : b < c) (ha : a ≤ 1) (hb : 0 ≤ b) : + b * a < c := +(mul_le_of_le_one_right hb ha).trans_lt h -/-! Lemmas of the form `b ≤ c → a ≤ 1 → b * a ≤ c`. -/ +/-- Assumes left covariance. -/ +lemma left.mul_le_one_of_le_of_le [pos_mul_mono α] (ha : a ≤ 1) (hb : b ≤ 1) (a0 : 0 ≤ a) : + a * b ≤ 1 := +mul_le_of_le_of_le_one_of_nonneg ha hb a0 -lemma mul_le_of_le_of_le_one [pos_mul_mono α] - (bc : b ≤ c) (ha : a ≤ 1) (b0 : 0 ≤ b) : b * a ≤ c := -b0.lt_or_eq.elim (preorder.mul_le_of_le_of_le_one bc ha) - (λ h, by rw [← h, zero_mul]; exact b0.trans bc) +/-- Assumes left covariance. -/ +lemma left.mul_lt_of_le_of_lt_one_of_pos [pos_mul_strict_mono α] + (ha : a ≤ 1) (hb : b < 1) (a0 : 0 < a) : a * b < 1 := +mul_lt_of_le_of_lt_one_of_pos ha hb a0 /-- Assumes left covariance. -/ -lemma left.mul_le_one_of_le_of_le [pos_mul_mono α] - (ha : a ≤ 1) (hb : b ≤ 1) (a0 : 0 ≤ a) : a * b ≤ 1 := -mul_le_of_le_of_le_one ha hb a0 +lemma left.mul_lt_of_lt_of_le_one_of_nonneg [pos_mul_mono α] + (ha : a < 1) (hb : b ≤ 1) (a0 : 0 ≤ a) : a * b < 1 := +mul_lt_of_lt_of_le_one_of_nonneg ha hb a0 -lemma mul_nonneg_le_one_le [pos_mul_mono α] [mul_pos_mono α] - (c0 : 0 ≤ c) (bc : b ≤ c) (a0 : 0 ≤ a) (ha : a ≤ 1) : b * a ≤ c := -calc b * a ≤ c * a : mul_le_mul_of_nonneg_right bc a0 - ... ≤ c * 1 : mul_le_mul_of_nonneg_left ha c0 - ... = c : mul_one c +lemma mul_le_of_le_of_le_one' [pos_mul_mono α] [mul_pos_mono α] + (bc : b ≤ c) (ha : a ≤ 1) (a0 : 0 ≤ a) (c0 : 0 ≤ c) : b * a ≤ c := +(mul_le_mul_of_nonneg_right bc a0).trans $ mul_le_of_le_one_right c0 ha lemma mul_lt_of_lt_of_le_one' [pos_mul_mono α] [mul_pos_strict_mono α] (bc : b < c) (ha : a ≤ 1) (a0 : 0 < a) (c0 : 0 ≤ c) : b * a < c := -calc b * a < c * a : mul_lt_mul_right' bc a0 - ... ≤ c * 1 : mul_le_mul_of_nonneg_left ha c0 - ... = c : mul_one c +(mul_lt_mul_of_pos_right bc a0).trans_le $ mul_le_of_le_one_right c0 ha lemma mul_lt_of_le_of_lt_one' [pos_mul_strict_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : a < 1) (a0 : 0 ≤ a) (c0 : 0 < c) : b * a < c := -calc b * a ≤ c * a : mul_le_mul_of_nonneg_right bc a0 - ... < c * 1 : mul_lt_mul_left' ha c0 - ... = c : mul_one c +(mul_le_mul_of_nonneg_right bc a0).trans_lt $ mul_lt_of_lt_one_right c0 ha + +lemma mul_lt_of_lt_of_lt_one_of_pos [pos_mul_mono α] [mul_pos_strict_mono α] + (bc : b < c) (ha : a ≤ 1) (a0 : 0 < a) (c0 : 0 ≤ c) : b * a < c := +(mul_lt_mul_of_pos_right bc a0).trans_le $ mul_le_of_le_one_right c0 ha /-! Lemmas of the form `b ≤ c → 1 ≤ a → b ≤ c * a`. -/ -lemma le_mul_of_le_of_one_le [pos_mul_mono α] - (bc : b ≤ c) (ha : 1 ≤ a) (c0 : 0 ≤ c) : b ≤ c * a := -c0.lt_or_eq.elim (preorder.le_mul_of_le_of_one_le bc ha) - (λ h, by rw [← h, zero_mul] at *; exact bc) +lemma le_mul_of_le_of_one_le_of_nonneg [pos_mul_mono α] (h : b ≤ c) (ha : 1 ≤ a) (hc : 0 ≤ c) : + b ≤ c * a := +h.trans $ le_mul_of_one_le_right hc ha + +lemma lt_mul_of_le_of_one_lt_of_pos [pos_mul_strict_mono α] (bc : b ≤ c) (ha : 1 < a) (c0 : 0 < c) : + b < c * a := +bc.trans_lt $ lt_mul_of_one_lt_right c0 ha + +lemma lt_mul_of_lt_of_one_le_of_nonneg [pos_mul_mono α] (h : b < c) (ha : 1 ≤ a) (hc : 0 ≤ c) : + b < c * a := +h.trans_le $ le_mul_of_one_le_right hc ha + +/-- Assumes left covariance. -/ +lemma left.one_le_mul_of_le_of_le [pos_mul_mono α] (ha : 1 ≤ a) (hb : 1 ≤ b) (a0 : 0 ≤ a) : + 1 ≤ a * b := +le_mul_of_le_of_one_le_of_nonneg ha hb a0 + +/-- Assumes left covariance. -/ +lemma left.one_lt_mul_of_le_of_lt_of_pos [pos_mul_strict_mono α] + (ha : 1 ≤ a) (hb : 1 < b) (a0 : 0 < a) : 1 < a * b := +lt_mul_of_le_of_one_lt_of_pos ha hb a0 /-- Assumes left covariance. -/ -lemma left.one_le_mul_of_le_of_le [pos_mul_mono α] - (ha : 1 ≤ a) (hb : 1 ≤ b) (a0 : 0 ≤ a) : 1 ≤ a * b := -le_mul_of_le_of_one_le ha hb a0 +lemma left.lt_mul_of_lt_of_one_le_of_nonneg [pos_mul_mono α] + (ha : 1 < a) (hb : 1 ≤ b) (a0 : 0 ≤ a) : 1 < a * b := +lt_mul_of_lt_of_one_le_of_nonneg ha hb a0 lemma le_mul_of_le_of_one_le' [pos_mul_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : 1 ≤ a) (a0 : 0 ≤ a) (b0 : 0 ≤ b) : b ≤ c * a := -calc b = b * 1 : (mul_one b).symm - ... ≤ b * a : mul_le_mul_of_nonneg_left ha b0 - ... ≤ c * a : mul_le_mul_of_nonneg_right bc a0 +(le_mul_of_one_le_right b0 ha).trans $ mul_le_mul_of_nonneg_right bc a0 lemma lt_mul_of_le_of_one_lt' [pos_mul_strict_mono α] [mul_pos_mono α] (bc : b ≤ c) (ha : 1 < a) (a0 : 0 ≤ a) (b0 : 0 < b) : b < c * a := -calc b = b * 1 : (mul_one b).symm - ... < b * a : mul_lt_mul_left' ha b0 - ... ≤ c * a : mul_le_mul_of_nonneg_right bc a0 +(lt_mul_of_one_lt_right b0 ha).trans_le $ mul_le_mul_of_nonneg_right bc a0 lemma lt_mul_of_lt_of_one_le' [pos_mul_mono α] [mul_pos_strict_mono α] (bc : b < c) (ha : 1 ≤ a) (a0 : 0 < a) (b0 : 0 ≤ b) : b < c * a := -calc b = b * 1 : (mul_one b).symm - ... ≤ b * a : mul_le_mul_of_nonneg_left ha b0 - ... < c * a : mul_lt_mul_right' bc a0 +(le_mul_of_one_le_right b0 ha).trans_lt $ mul_lt_mul_of_pos_right bc a0 + +lemma lt_mul_of_lt_of_one_lt_of_pos [pos_mul_strict_mono α] [mul_pos_strict_mono α] + (bc : b < c) (ha : 1 < a) (a0 : 0 < a) (b0 : 0 < b) : b < c * a := +(lt_mul_of_one_lt_right b0 ha).trans $ mul_lt_mul_of_pos_right bc a0 /-! Lemmas of the form `a ≤ 1 → b ≤ c → a * b ≤ c`. -/ -lemma mul_le_of_le_one_of_le [mul_pos_mono α] - (ha : a ≤ 1) (bc : b ≤ c) (b0 : 0 ≤ b) : a * b ≤ c := -b0.lt_or_eq.elim (preorder.mul_le_of_le_one_of_le ha bc) - (λ h, by rw [← h, mul_zero] at *; exact bc) +lemma mul_le_of_le_one_of_le_of_nonneg [mul_pos_mono α] (ha : a ≤ 1) (h : b ≤ c) (hb : 0 ≤ b) : + a * b ≤ c := +(mul_le_of_le_one_left hb ha).trans h + +lemma mul_lt_of_lt_one_of_le_of_pos [mul_pos_strict_mono α] (ha : a < 1) (h : b ≤ c) (hb : 0 < b) : + a * b < c := +(mul_lt_of_lt_one_left hb ha).trans_le h + +lemma mul_lt_of_le_one_of_lt_of_nonneg [mul_pos_mono α] (ha : a ≤ 1) (h : b < c) (hb : 0 ≤ b) : + a * b < c := +(mul_le_of_le_one_left hb ha).trans_lt h + +/-- Assumes right covariance. -/ +lemma right.mul_lt_one_of_lt_of_le_of_pos [mul_pos_strict_mono α] + (ha : a < 1) (hb : b ≤ 1) (b0 : 0 < b) : a * b < 1 := +mul_lt_of_lt_one_of_le_of_pos ha hb b0 + +/-- Assumes right covariance. -/ +lemma right.mul_lt_one_of_le_of_lt_of_nonneg [mul_pos_mono α] + (ha : a ≤ 1) (hb : b < 1) (b0 : 0 ≤ b) : a * b < 1 := +mul_lt_of_le_one_of_lt_of_nonneg ha hb b0 + +lemma mul_lt_of_lt_one_of_lt_of_pos [pos_mul_strict_mono α] [mul_pos_strict_mono α] + (ha : a < 1) (bc : b < c) (a0 : 0 < a) (c0 : 0 < c) : a * b < c := +(mul_lt_mul_of_pos_left bc a0).trans $ mul_lt_of_lt_one_left c0 ha /-- Assumes right covariance. -/ lemma right.mul_le_one_of_le_of_le [mul_pos_mono α] - (ha : a ≤ 1) (hb : b ≤ 1) (b0 : 0 < b) : a * b ≤ 1 := -preorder.mul_le_of_le_one_of_le ha hb b0 + (ha : a ≤ 1) (hb : b ≤ 1) (b0 : 0 ≤ b) : a * b ≤ 1 := +mul_le_of_le_one_of_le_of_nonneg ha hb b0 lemma mul_le_of_le_one_of_le' [pos_mul_mono α] [mul_pos_mono α] (ha : a ≤ 1) (bc : b ≤ c) (a0 : 0 ≤ a) (c0 : 0 ≤ c) : a * b ≤ c := -calc a * b ≤ a * c : mul_le_mul_of_nonneg_left bc a0 - ... ≤ 1 * c : mul_le_mul_of_nonneg_right ha c0 - ... = c : one_mul c +(mul_le_mul_of_nonneg_left bc a0).trans $ mul_le_of_le_one_left c0 ha lemma mul_lt_of_lt_one_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] (ha : a < 1) (bc : b ≤ c) (a0 : 0 ≤ a) (c0 : 0 < c) : a * b < c := -calc a * b ≤ a * c : mul_le_mul_of_nonneg_left bc a0 - ... < 1 * c : mul_lt_mul_right' ha c0 - ... = c : one_mul c +(mul_le_mul_of_nonneg_left bc a0).trans_lt $ mul_lt_of_lt_one_left c0 ha lemma mul_lt_of_le_one_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] (ha : a ≤ 1) (bc : b < c) (a0 : 0 < a) (c0 : 0 ≤ c) : a * b < c := -calc a * b < a * c : mul_lt_mul_left' bc a0 - ... ≤ 1 * c : mul_le_mul_of_nonneg_right ha c0 - ... = c : one_mul c +(mul_lt_mul_of_pos_left bc a0).trans_le $ mul_le_of_le_one_left c0 ha /-! Lemmas of the form `1 ≤ a → b ≤ c → b ≤ a * c`. -/ -lemma le_mul_of_one_le_of_le [mul_pos_mono α] - (ha : 1 ≤ a) (bc : b ≤ c) (c0 : 0 ≤ c) : b ≤ a * c := -c0.lt_or_eq.elim (preorder.le_mul_of_one_le_of_le ha bc) - (λ h, by rw [← h, mul_zero] at *; exact bc) +lemma lt_mul_of_one_lt_of_le_of_pos [mul_pos_strict_mono α] (ha : 1 < a) (h : b ≤ c) (hc : 0 < c) : + b < a * c := +h.trans_lt $ lt_mul_of_one_lt_left hc ha + +lemma lt_mul_of_one_le_of_lt_of_nonneg [mul_pos_mono α] (ha : 1 ≤ a) (h : b < c) (hc : 0 ≤ c) : + b < a * c := +h.trans_le $ le_mul_of_one_le_left hc ha + +lemma lt_mul_of_one_lt_of_lt_of_pos [mul_pos_strict_mono α] (ha : 1 < a) (h : b < c) (hc : 0 < c) : + b < a * c := +h.trans $ lt_mul_of_one_lt_left hc ha + +/-- Assumes right covariance. -/ +lemma right.one_lt_mul_of_lt_of_le_of_pos [mul_pos_strict_mono α] + (ha : 1 < a) (hb : 1 ≤ b) (b0 : 0 < b) : 1 < a * b := +lt_mul_of_one_lt_of_le_of_pos ha hb b0 /-- Assumes right covariance. -/ -lemma right.one_le_mul_of_le_of_le [mul_pos_mono α] - (ha : 1 ≤ a) (hb : 1 ≤ b) (b0 : 0 ≤ b) : 1 ≤ a * b := -le_mul_of_one_le_of_le ha hb b0 - -lemma le_mul_of_one_le_of_le' [pos_mul_mono α] [mul_pos_mono α] - (ha : 1 ≤ a) (bc : b ≤ c) (a0 : 0 ≤ a) (b0 : 0 ≤ b) : b ≤ a * c := -calc b = 1 * b : (one_mul b).symm - ... ≤ a * b : mul_le_mul_of_nonneg_right ha b0 - ... ≤ a * c : mul_le_mul_of_nonneg_left bc a0 - -lemma lt_mul_of_one_lt_of_le' [pos_mul_mono α] [mul_pos_strict_mono α] - (ha : 1 < a) (bc : b ≤ c) (a0 : 0 ≤ a) (b0 : 0 < b) : b < a * c := -calc b = 1 * b : (one_mul b).symm - ... < a * b : mul_lt_mul_right' ha b0 - ... ≤ a * c : mul_le_mul_of_nonneg_left bc a0 - -lemma lt_mul_of_one_le_of_lt' [pos_mul_strict_mono α] [mul_pos_mono α] - (ha : 1 ≤ a) (bc : b < c) (a0 : 0 < a) (b0 : 0 ≤ b) : b < a * c := -calc b = 1 * b : (one_mul b).symm - ... ≤ a * b : mul_le_mul_of_nonneg_right ha b0 - ... < a * c : mul_lt_mul_left' bc a0 - -lemma mul_le_of_le_one_right [pos_mul_mono α] (a0 : 0 ≤ a) (h : b ≤ 1) : - a * b ≤ a := -mul_le_of_le_of_le_one le_rfl h a0 - -lemma le_mul_of_one_le_right [pos_mul_mono α] (a0 : 0 ≤ a) (h : 1 ≤ b) : - a ≤ a * b := -le_mul_of_le_of_one_le le_rfl h a0 - -lemma mul_le_of_le_one_left [mul_pos_mono α] (b0 : 0 ≤ b) (h : a ≤ 1) : - a * b ≤ b := -mul_le_of_le_one_of_le h le_rfl b0 - -lemma le_mul_of_one_le_left [mul_pos_mono α] (b0 : 0 ≤ b) (h : 1 ≤ a) : - b ≤ a * b := -le_mul_of_one_le_of_le h le_rfl b0 - -lemma le_of_mul_le_of_one_le_of_nonneg_left [pos_mul_mono α] - (h : a * b ≤ c) (hle : 1 ≤ b) (a0 : 0 ≤ a) : +lemma right.one_lt_mul_of_le_of_lt_of_nonneg [mul_pos_mono α] + (ha : 1 ≤ a) (hb : 1 < b) (b0 : 0 ≤ b) : 1 < a * b := +lt_mul_of_one_le_of_lt_of_nonneg ha hb b0 + +/-- Assumes right covariance. -/ +lemma right.one_lt_mul_of_lt_of_lt [mul_pos_strict_mono α] + (ha : 1 < a) (hb : 1 < b) (b0 : 0 < b) : 1 < a * b := +lt_mul_of_one_lt_of_lt_of_pos ha hb b0 + +lemma lt_mul_of_one_lt_of_lt_of_nonneg [mul_pos_mono α] (ha : 1 ≤ a) (h : b < c) (hc : 0 ≤ c) : + b < a * c := +h.trans_le $ le_mul_of_one_le_left hc ha + +lemma lt_of_mul_lt_of_one_le_of_nonneg_left [pos_mul_mono α] (h : a * b < c) (hle : 1 ≤ b) + (ha : 0 ≤ a) : + a < c := +(le_mul_of_one_le_right ha hle).trans_lt h + +lemma lt_of_lt_mul_of_le_one_of_nonneg_left [pos_mul_mono α] (h : a < b * c) (hc : c ≤ 1) + (hb : 0 ≤ b) : + a < b := +h.trans_le $ mul_le_of_le_one_right hb hc + +lemma lt_of_lt_mul_of_le_one_of_nonneg_right [mul_pos_mono α] (h : a < b * c) (hb : b ≤ 1) + (hc : 0 ≤ c) : + a < c := +h.trans_le $ mul_le_of_le_one_left hc hb + +lemma le_mul_of_one_le_of_le_of_nonneg [mul_pos_mono α] (ha : 1 ≤ a) (bc : b ≤ c) (c0 : 0 ≤ c) : + b ≤ a * c := +bc.trans $ le_mul_of_one_le_left c0 ha + +/-- Assumes right covariance. -/ +lemma right.one_le_mul_of_le_of_le [mul_pos_mono α] (ha : 1 ≤ a) (hb : 1 ≤ b) (b0 : 0 ≤ b) : + 1 ≤ a * b := +le_mul_of_one_le_of_le_of_nonneg ha hb b0 + +lemma le_of_mul_le_of_one_le_of_nonneg_left [pos_mul_mono α] (h : a * b ≤ c) (hb : 1 ≤ b) + (ha : 0 ≤ a) : a ≤ c := -a0.lt_or_eq.elim (preorder.le_of_mul_le_of_one_le_left h hle) - (λ ha, by simpa only [← ha, zero_mul] using h) +(le_mul_of_one_le_right ha hb).trans h -lemma le_of_le_mul_of_le_one_of_nonneg_left [pos_mul_mono α] - (h : a ≤ b * c) (hle : c ≤ 1) (b0 : 0 ≤ b) : +lemma le_of_le_mul_of_le_one_of_nonneg_left [pos_mul_mono α] (h : a ≤ b * c) (hc : c ≤ 1) + (hb : 0 ≤ b) : a ≤ b := -b0.lt_or_eq.elim (preorder.le_of_le_mul_of_le_one_left h hle) - (λ hb, by simpa only [← hb, zero_mul] using h) +h.trans $ mul_le_of_le_one_right hb hc -lemma le_of_mul_le_of_one_le_nonneg_right [mul_pos_mono α] - (h : a * b ≤ c) (hle : 1 ≤ a) (b0 : 0 ≤ b) : +lemma le_of_mul_le_of_one_le_nonneg_right [mul_pos_mono α] (h : a * b ≤ c) (ha : 1 ≤ a) + (hb : 0 ≤ b) : b ≤ c := -b0.lt_or_eq.elim (preorder.le_of_mul_le_of_one_le_right h hle) - (λ ha, by simpa only [← ha, mul_zero] using h) +(le_mul_of_one_le_left hb ha).trans h -lemma le_of_le_mul_of_le_one_right [mul_pos_mono α] - (h : a ≤ b * c) (hle : b ≤ 1) (c0 : 0 ≤ c) : +lemma le_of_le_mul_of_le_one_of_nonneg_right [mul_pos_mono α] (h : a ≤ b * c) (hb : b ≤ 1) + (hc : 0 ≤ c) : a ≤ c := -c0.lt_or_eq.elim (preorder.le_of_le_mul_of_le_one_right h hle) - (λ ha, by simpa only [← ha, mul_zero] using h) +h.trans $ mul_le_of_le_one_left hc hb -end partial_order +end preorder section linear_order variables [linear_order α] -lemma exists_square_le [pos_mul_strict_mono α] - (a0 : 0 ≤ a) : ∃ (b : α), b * b ≤ a := -a0.lt_or_eq.elim exists_square_le' (λ h, by rw [← h]; exact ⟨0, by simp⟩) +-- Yaël: What's the point of this lemma? If we have `0 * 0 = 0`, then we can just take `b = 0`. +-- proven with `a0 : 0 ≤ a` as `exists_square_le` +lemma exists_square_le' [pos_mul_strict_mono α] (a0 : 0 < a) : ∃ (b : α), b * b ≤ a := +begin + obtain ha | ha := lt_or_le a 1, + { exact ⟨a, (mul_lt_of_lt_one_right a0 ha).le⟩ }, + { exact ⟨1, by rwa mul_one⟩ } +end end linear_order - -end mul_zero_one_class +end mul_one_class section cancel_monoid_with_zero @@ -1250,41 +813,37 @@ section partial_order variables [partial_order α] lemma pos_mul_mono.to_pos_mul_strict_mono [pos_mul_mono α] : pos_mul_strict_mono α := -⟨λ x a b h, (mul_le_mul_left_of_le_of_pos h.le x.2).lt_of_ne (h.ne ∘ mul_left_cancel₀ x.2.ne')⟩ +⟨λ x a b h, (mul_le_mul_of_nonneg_left h.le x.2.le).lt_of_ne (h.ne ∘ mul_left_cancel₀ x.2.ne')⟩ lemma pos_mul_mono_iff_pos_mul_strict_mono : pos_mul_mono α ↔ pos_mul_strict_mono α := -⟨@pos_mul_mono.to_pos_mul_strict_mono α _ _, @zero_lt.pos_mul_strict_mono.to_pos_mul_mono α _ _ _⟩ +⟨@pos_mul_mono.to_pos_mul_strict_mono α _ _, @pos_mul_strict_mono.to_pos_mul_mono α _ _⟩ lemma mul_pos_mono.to_mul_pos_strict_mono [mul_pos_mono α] : mul_pos_strict_mono α := -⟨λ x a b h, (mul_le_mul_right_of_le_of_pos h.le x.2).lt_of_ne (h.ne ∘ mul_right_cancel₀ x.2.ne')⟩ +⟨λ x a b h, (mul_le_mul_of_nonneg_right h.le x.2.le).lt_of_ne (h.ne ∘ mul_right_cancel₀ x.2.ne')⟩ lemma mul_pos_mono_iff_mul_pos_strict_mono : mul_pos_mono α ↔ mul_pos_strict_mono α := -⟨@mul_pos_mono.to_mul_pos_strict_mono α _ _, @zero_lt.mul_pos_strict_mono.to_mul_pos_mono α _ _ _⟩ +⟨@mul_pos_mono.to_mul_pos_strict_mono α _ _, @mul_pos_strict_mono.to_mul_pos_mono α _ _⟩ lemma pos_mul_reflect_lt.to_pos_mul_mono_rev [pos_mul_reflect_lt α] : pos_mul_mono_rev α := ⟨λ x a b h, h.eq_or_lt.elim (le_of_eq ∘ mul_left_cancel₀ x.2.ne.symm) - (λ h', (lt_of_mul_lt_mul_left' h' x.2).le)⟩ + (λ h', (lt_of_mul_lt_mul_left h' x.2.le).le)⟩ lemma pos_mul_mono_rev_iff_pos_mul_reflect_lt : pos_mul_mono_rev α ↔ pos_mul_reflect_lt α := -⟨@zero_lt.pos_mul_mono_rev.to_pos_mul_reflect_lt α _ _ _, - @pos_mul_reflect_lt.to_pos_mul_mono_rev α _ _⟩ +⟨@pos_mul_mono_rev.to_pos_mul_reflect_lt α _ _, @pos_mul_reflect_lt.to_pos_mul_mono_rev α _ _⟩ lemma mul_pos_reflect_lt.to_mul_pos_mono_rev [mul_pos_reflect_lt α] : mul_pos_mono_rev α := ⟨λ x a b h, h.eq_or_lt.elim (le_of_eq ∘ mul_right_cancel₀ x.2.ne.symm) - (λ h', (lt_of_mul_lt_mul_right' h' x.2).le)⟩ + (λ h', (lt_of_mul_lt_mul_right h' x.2.le).le)⟩ lemma mul_pos_mono_rev_iff_mul_pos_reflect_lt : mul_pos_mono_rev α ↔ mul_pos_reflect_lt α := -⟨@zero_lt.mul_pos_mono_rev.to_mul_pos_reflect_lt α _ _ _, - @mul_pos_reflect_lt.to_mul_pos_mono_rev α _ _⟩ +⟨@mul_pos_mono_rev.to_mul_pos_reflect_lt α _ _, @mul_pos_reflect_lt.to_mul_pos_mono_rev α _ _⟩ end partial_order end cancel_monoid_with_zero section comm_semigroup_has_zero -variables [comm_semigroup α] [has_zero α] - -variables [has_lt α] +variables [comm_semigroup α] [has_zero α] [preorder α] lemma pos_mul_strict_mono_iff_mul_pos_strict_mono : pos_mul_strict_mono α ↔ mul_pos_strict_mono α := @@ -1294,8 +853,6 @@ lemma pos_mul_reflect_lt_iff_mul_pos_reflect_lt : pos_mul_reflect_lt α ↔ mul_pos_reflect_lt α := by simp ! only [mul_comm] -variables [has_le α] - lemma pos_mul_mono_iff_mul_pos_mono : pos_mul_mono α ↔ mul_pos_mono α := by simp ! only [mul_comm] @@ -1305,7 +862,3 @@ lemma pos_mul_mono_rev_iff_mul_pos_mono_rev : by simp ! only [mul_comm] end comm_semigroup_has_zero - -end zero_lt - -export zero_lt diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 15ed407dafdd0..21da4334cc8e9 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -113,24 +113,18 @@ class ordered_semiring (α : Type u) extends semiring α, ordered_cancel_add_com section ordered_semiring variables [ordered_semiring α] {a b c d : α} -lemma mul_lt_mul_of_pos_left (h₁ : a < b) (h₂ : 0 < c) : c * a < c * b := -ordered_semiring.mul_lt_mul_of_pos_left a b c h₁ h₂ +@[priority 200] -- see Note [lower instance priority] +instance ordered_semiring.pos_mul_strict_mono : pos_mul_strict_mono α := +⟨λ x a b h, ordered_semiring.mul_lt_mul_of_pos_left _ _ _ h x.prop⟩ -lemma mul_lt_mul_of_pos_right (h₁ : a < b) (h₂ : 0 < c) : a * c < b * c := -ordered_semiring.mul_lt_mul_of_pos_right a b c h₁ h₂ +@[priority 200] -- see Note [lower instance priority] +instance ordered_semiring.mul_pos_strict_mono : mul_pos_strict_mono α := +⟨λ x a b h, ordered_semiring.mul_lt_mul_of_pos_right _ _ _ h x.prop⟩ @[priority 100] -- see Note [lower instance priority] instance ordered_semiring.zero_le_one_class : zero_le_one_class α := { ..‹ordered_semiring α› } -@[priority 200] -- see Note [lower instance priority] -instance ordered_semiring.pos_mul_strict_mono : zero_lt.pos_mul_strict_mono α := -⟨λ x a b h, mul_lt_mul_of_pos_left h x.prop⟩ - -@[priority 200] -- see Note [lower instance priority] -instance ordered_semiring.mul_pos_strict_mono : zero_lt.mul_pos_strict_mono α := -⟨λ x a b h, mul_lt_mul_of_pos_right h x.prop⟩ - section nontrivial variables [nontrivial α] @@ -623,12 +617,6 @@ local attribute [instance] linear_ordered_semiring.decidable_le -- with only a `linear_ordered_semiring` typeclass argument. lemma zero_lt_one' : 0 < (1 : α) := zero_lt_one -lemma le_of_mul_le_mul_left (h : c * a ≤ c * b) (hc : 0 < c) : a ≤ b := -(strict_mono_mul_left_of_pos hc).le_iff_le.1 h - -lemma le_of_mul_le_mul_right (h : a * c ≤ b * c) (hc : 0 < c) : a ≤ b := -(strict_mono_mul_right_of_pos hc).le_iff_le.1 h - lemma nonneg_and_nonneg_or_nonpos_and_nonpos_of_mul_nnonneg (hab : 0 ≤ a * b) : (0 ≤ a ∧ 0 ≤ b) ∨ (a ≤ 0 ∧ b ≤ 0) := begin @@ -1410,13 +1398,11 @@ begin end @[priority 200] -- see Note [lower instance priority] -instance canonically_ordered_comm_semiring.pos_mul_mono : - zero_lt.pos_mul_mono α := +instance canonically_ordered_comm_semiring.to_pos_mul_mono : pos_mul_mono α := ⟨λ x a b h, by { obtain ⟨d, rfl⟩ := exists_add_of_le h, simp_rw [left_distrib, le_self_add], }⟩ @[priority 200] -- see Note [lower instance priority] -instance canonically_ordered_comm_semiring.mul_pos_mono : - zero_lt.mul_pos_mono α := +instance canonically_ordered_comm_semiring.to_mul_pos_mono : mul_pos_mono α := ⟨λ x a b h, by { obtain ⟨d, rfl⟩ := exists_add_of_le h, simp_rw [right_distrib, le_self_add], }⟩ /-- A version of `zero_lt_one : 0 < 1` for a `canonically_ordered_comm_semiring`. -/ @@ -1706,20 +1692,17 @@ with_top.comm_monoid_with_zero instance [canonically_ordered_comm_semiring α] [nontrivial α] : comm_semiring (with_bot α) := with_top.comm_semiring -instance [canonically_ordered_comm_semiring α] [nontrivial α] : - zero_lt.pos_mul_mono (with_bot α) := -⟨ begin +instance [canonically_ordered_comm_semiring α] [nontrivial α] : pos_mul_mono (with_bot α) := +pos_mul_mono_iff_covariant_pos.2 ⟨begin rintros ⟨x, x0⟩ a b h, simp only [subtype.coe_mk], - induction x using with_bot.rec_bot_coe, - { have := bot_lt_coe (0 : α), rw [coe_zero] at this, exact absurd x0.le this.not_le, }, - { induction a using with_bot.rec_bot_coe, { simp_rw [mul_bot x0.ne.symm, bot_le], }, - induction b using with_bot.rec_bot_coe, { exact absurd h (bot_lt_coe a).not_le, }, - { simp only [← coe_mul, coe_le_coe] at *, - exact mul_le_mul_left' h x, }, }, + lift x to α using x0.ne_bot, + induction a using with_bot.rec_bot_coe, { simp_rw [mul_bot x0.ne.symm, bot_le] }, + induction b using with_bot.rec_bot_coe, { exact absurd h (bot_lt_coe a).not_le }, + simp only [← coe_mul, coe_le_coe] at *, + exact mul_le_mul_left' h x, end ⟩ -instance [canonically_ordered_comm_semiring α] [nontrivial α] : - zero_lt.mul_pos_mono (with_bot α) := -zero_lt.pos_mul_mono_iff_mul_pos_mono.mp zero_lt.pos_mul_mono +instance [canonically_ordered_comm_semiring α] [nontrivial α] : mul_pos_mono (with_bot α) := +pos_mul_mono_iff_mul_pos_mono.mp infer_instance end with_bot diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 5f0dc2f80afe5..462b1fb9034ba 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -173,7 +173,7 @@ begin obtain ⟨k, rfl⟩ := add_subgroup.mem_zmultiples_iff.mp hx, rw [mem_preimage, mem_ball_zero_iff, add_subgroup.coe_mk, mem_singleton_iff, subtype.ext_iff, add_subgroup.coe_mk, add_subgroup.coe_zero, norm_zsmul ℚ k e, - int.norm_cast_rat, int.norm_eq_abs, ← int.cast_abs, zero_lt.mul_lt_iff_lt_one_left + int.norm_cast_rat, int.norm_eq_abs, ← int.cast_abs, mul_lt_iff_lt_one_left (norm_pos_iff.mpr he), ← @int.cast_one ℝ _, int.cast_lt, int.abs_lt_one_iff, smul_eq_zero, or_iff_left he], }, end diff --git a/src/analysis/special_functions/bernstein.lean b/src/analysis/special_functions/bernstein.lean index 9965481e75340..f7532b07c6960 100644 --- a/src/analysis/special_functions/bernstein.lean +++ b/src/analysis/special_functions/bernstein.lean @@ -299,11 +299,7 @@ begin ... = (2 * ∥f∥) * δ^(-2 : ℤ) * x * (1-x) / n : by { rw variance npos, ring, } ... ≤ (2 * ∥f∥) * δ^(-2 : ℤ) / n - : (div_le_div_right npos).mpr - begin - apply mul_nonneg_le_one_le w₂, - apply mul_nonneg_le_one_le w₂ le_rfl, - all_goals { unit_interval, }, - end + : (div_le_div_right npos).mpr $ + by refine mul_le_of_le_of_le_one' (mul_le_of_le_one_right w₂ _) _ _ w₂; unit_interval ... < ε/2 : nh, } end diff --git a/src/measure_theory/integral/set_to_l1.lean b/src/measure_theory/integral/set_to_l1.lean index a0548ef4a96e1..30d3fb6be1626 100644 --- a/src/measure_theory/integral/set_to_l1.lean +++ b/src/measure_theory/integral/set_to_l1.lean @@ -599,15 +599,9 @@ lemma norm_set_to_simple_func_le_sum_mul_norm (T : set α → F →L[ℝ] F') {C ∥f.set_to_simple_func T∥ ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ∥x∥ := calc ∥f.set_to_simple_func T∥ ≤ ∑ x in f.range, ∥T (f ⁻¹' {x})∥ * ∥x∥ : norm_set_to_simple_func_le_sum_op_norm T f -... ≤ ∑ x in f.range, C * (μ (f ⁻¹' {x})).to_real * ∥x∥ : - begin - refine finset.sum_le_sum (λ b hb, _), - by_cases hb : ∥b∥ = 0, - { rw hb, simp, }, - refine (zero_lt.mul_le_mul_right _).mpr _, swap, -- remove swap or inline the second subcase - { exact hT_norm _ (simple_func.measurable_set_fiber _ _), }, - { exact lt_of_le_of_ne (norm_nonneg _) (ne.symm hb), }, - end +... ≤ ∑ x in f.range, C * (μ (f ⁻¹' {x})).to_real * ∥x∥ + : sum_le_sum $ λ b hb, mul_le_mul_of_nonneg_right + (hT_norm _ $ simple_func.measurable_set_fiber _ _) $ norm_nonneg _ ... ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ∥x∥ : by simp_rw [mul_sum, ← mul_assoc] lemma norm_set_to_simple_func_le_sum_mul_norm_of_integrable (T : set α → E →L[ℝ] F') {C : ℝ} @@ -619,13 +613,10 @@ calc ∥f.set_to_simple_func T∥ ... ≤ ∑ x in f.range, C * (μ (f ⁻¹' {x})).to_real * ∥x∥ : begin refine finset.sum_le_sum (λ b hb, _), - by_cases hb : ∥b∥ = 0, - { rw hb, simp, }, - refine (zero_lt.mul_le_mul_right _).mpr _, swap, -- remove swap or inline the second subcase - { refine hT_norm _ (simple_func.measurable_set_fiber _ _) - (simple_func.measure_preimage_lt_top_of_integrable _ hf _), - rwa norm_eq_zero at hb, }, - { exact lt_of_le_of_ne (norm_nonneg _) (ne.symm hb), }, + obtain rfl | hb := eq_or_ne b 0, + { simp }, + exact mul_le_mul_of_nonneg_right (hT_norm _ (simple_func.measurable_set_fiber _ _) $ + simple_func.measure_preimage_lt_top_of_integrable _ hf hb) (norm_nonneg _), end ... ≤ C * ∑ x in f.range, (μ (f ⁻¹' {x})).to_real * ∥x∥ : by simp_rw [mul_sum, ← mul_assoc] From 1afd6c75e76e46221da42cf0ece4ad2111b3984a Mon Sep 17 00:00:00 2001 From: mcdoll Date: Wed, 14 Sep 2022 18:51:35 +0000 Subject: [PATCH 265/828] docs(analysis/calculus/taylor): Latex fixes (#16508) Fixes backslash error and corrects one docstring. --- src/analysis/calculus/taylor.lean | 13 ++++++------- 1 file changed, 6 insertions(+), 7 deletions(-) diff --git a/src/analysis/calculus/taylor.lean b/src/analysis/calculus/taylor.lean index 34c213296961f..6afa45503665c 100644 --- a/src/analysis/calculus/taylor.lean +++ b/src/analysis/calculus/taylor.lean @@ -57,7 +57,7 @@ def taylor_coeff_within (f : ℝ → E) (k : ℕ) (s : set ℝ) (x₀ : ℝ) : E /-- The Taylor polynomial with derivatives inside of a set `s`. The Taylor polynomial is given by -$$∑_{k=0}^n \\frac{(x - x₀)^k}{k!} f^{(k)}(x₀),$$ +$$∑_{k=0}^n \frac{(x - x₀)^k}{k!} f^{(k)}(x₀),$$ where $f^{(k)}(x₀)$ denotes the iterated derivative in the set `s`. -/ noncomputable def taylor_within (f : ℝ → E) (n : ℕ) (s : set ℝ) (x₀ : ℝ) : polynomial_module ℝ E := @@ -248,7 +248,7 @@ has_deriv_within_at_taylor_within_eval (unique_diff_on_Icc hx t ht) (unique_diff We assume that `f` is `n+1`-times continuously differentiable in the closed set `Icc x₀ x` and `n+1`-times differentiable on the open set `Ioo x₀ x`, and `g` is a differentiable function on `Ioo x₀ x` and continuous on `Icc x₀ x`. Then there exists a `x' ∈ Ioo x₀ x` such that -$$f(x) - (P_n f)(x₀, x) = \\frac{(x - x')^n}{n!} \\frac{g(x) - g(x₀)}{g' x'},$$ +$$f(x) - (P_n f)(x₀, x) = \frac{(x - x')^n}{n!} \frac{g(x) - g(x₀)}{g' x'},$$ where $P_n f$ denotes the Taylor polynomial of degree $n$. -/ lemma taylor_mean_remainder {f : ℝ → ℝ} {g g' : ℝ → ℝ} {x x₀ : ℝ} {n : ℕ} (hx : x₀ < x) (hf : cont_diff_on ℝ n f (Icc x₀ x)) @@ -280,7 +280,7 @@ end We assume that `f` is `n+1`-times continuously differentiable in the closed set `Icc x₀ x` and `n+1`-times differentiable on the open set `Ioo x₀ x`. Then there exists a `x' ∈ Ioo x₀ x` such that -$$f(x) - (P_n f)(x₀, x) = \\frac{f^{(n+1)}(x') (x - x₀)^{n+1}}{(n+1)!},$$ +$$f(x) - (P_n f)(x₀, x) = \frac{f^{(n+1)}(x') (x - x₀)^{n+1}}{(n+1)!},$$ where $P_n f$ denotes the Taylor polynomial of degree $n$ and $f^{(n+1)}$ is the $n+1$-th iterated derivative. -/ lemma taylor_mean_remainder_lagrange {f : ℝ → ℝ} {x x₀ : ℝ} {n : ℕ} (hx : x₀ < x) @@ -315,7 +315,7 @@ end We assume that `f` is `n+1`-times continuously differentiable on the closed set `Icc x₀ x` and `n+1`-times differentiable on the open set `Ioo x₀ x`. Then there exists a `x' ∈ Ioo x₀ x` such that -$$f(x) - (P_n f)(x₀, x) = \\frac{f^{(n+1)}(x') (x - x')^n (x-x₀)}{n!},$$ +$$f(x) - (P_n f)(x₀, x) = \frac{f^{(n+1)}(x') (x - x')^n (x-x₀)}{n!},$$ where $P_n f$ denotes the Taylor polynomial of degree $n$ and $f^{(n+1)}$ is the $n+1$-th iterated derivative. -/ lemma taylor_mean_remainder_cauchy {f : ℝ → ℝ} {x x₀ : ℝ} {n : ℕ} (hx : x₀ < x) @@ -386,9 +386,8 @@ end /-- **Taylor's theorem** with a polynomial bound on the remainder We assume that `f` is `n+1`-times continuously differentiable on the closed set `Icc a b`. -There exists a constant `C` such that for all `x ∈ Icc a b` The difference of `f` and its `n`-th -Taylor polynomial can be estimated by `C * (x - a)^(n+1) / n!` where `C` is a bound for the `n+1`-th -iterated derivative of `f`. -/ +There exists a constant `C` such that for all `x ∈ Icc a b` the difference of `f` and its `n`-th +Taylor polynomial can be estimated by `C * (x - a)^(n+1)`. -/ lemma exists_taylor_mean_remainder_bound {f : ℝ → E} {a b : ℝ} {n : ℕ} (hab : a ≤ b) (hf : cont_diff_on ℝ (n+1) f (Icc a b)) : ∃ C, ∀ x ∈ Icc a b, ∥f x - taylor_within_eval f n (Icc a b) a x∥ ≤ C * (x - a)^(n+1) := From e702f02117d89e3757e9cab0c9c8a2f7c01f6ddb Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 14 Sep 2022 21:29:45 +0000 Subject: [PATCH 266/828] feat(topology/algebra/module/character_space, analysis/normed_space/star/spectrum): define the Gelfand transform as an `alg_hom` (#16451) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This defines the `gelfand_transform` as an algebra homomorphism from a `𝕜`-algebra `A` into the continuous functions from the `character_space 𝕜 A` into the base field `𝕜`, where the map is given by evaluation. For this definition it is only supposed that `𝕜` is a topological ring and `A` is a topological space. When the algebra `A` is a C⋆-algebra over `ℂ`, algebra homomorphisms of `A` into `ℂ` (hence also terms of `character_space ℂ A`) are automatically star-preserving. Therefore, in this setting the Gelfand transform may be upgraded to a `star_alg_hom`. However, we do not implement that here because, with more work, one may show that this is actually an equivalence. - [x] depends on: #16438 --- src/analysis/normed_space/star/spectrum.lean | 39 +++++++++++++++++++ .../algebra/module/character_space.lean | 25 ++++++++++++ 2 files changed, 64 insertions(+) diff --git a/src/analysis/normed_space/star/spectrum.lean b/src/analysis/normed_space/star/spectrum.lean index 69938599375c8..863448cb51d90 100644 --- a/src/analysis/normed_space/star/spectrum.lean +++ b/src/analysis/normed_space/star/spectrum.lean @@ -15,6 +15,8 @@ In this file, we establish various propreties related to the spectrum of element local postfix `⋆`:std.prec.max_plus := star +section + open_locale topological_space ennreal open filter ennreal spectrum cstar_ring @@ -151,3 +153,40 @@ noncomputable instance : continuous_linear_map_class F ℂ A B := .. alg_hom_class.linear_map_class } end star_alg_hom + +end + +namespace weak_dual + +open continuous_map complex +open_locale complex_star_module + +variables {F A : Type*} [normed_ring A] [normed_algebra ℂ A] [nontrivial A] [complete_space A] + [star_ring A] [cstar_ring A] [star_module ℂ A] [hF : alg_hom_class F ℂ A ℂ] + +include hF + +/-- This instance is provided instead of `star_alg_hom_class` to avoid type class inference loops. +See note [lower instance priority] -/ +@[priority 100] +noncomputable instance : star_hom_class F A ℂ := +{ coe := λ φ, φ, + coe_injective' := fun_like.coe_injective', + map_star := λ φ a, + begin + suffices hsa : ∀ s : self_adjoint A, (φ s)⋆ = φ s, + { rw ←real_part_add_I_smul_imaginary_part a, + simp only [map_add, map_smul, star_add, star_smul, hsa, self_adjoint.star_coe_eq] }, + { intros s, + have := alg_hom.apply_mem_spectrum φ (s : A), + rw self_adjoint.coe_re_map_spectrum s at this, + rcases this with ⟨⟨_, _⟩, _, heq⟩, + rw [←heq, is_R_or_C.star_def, is_R_or_C.conj_of_real] } + end } + +/-- This is not an instance to avoid type class inference loops. See +`weak_dual.complex.star_hom_class`. -/ +noncomputable def _root_.alg_hom_class.star_alg_hom_class : star_alg_hom_class F ℂ A ℂ := +{ .. hF, .. weak_dual.complex.star_hom_class } + +end weak_dual diff --git a/src/topology/algebra/module/character_space.lean b/src/topology/algebra/module/character_space.lean index dfce8973e4cbc..df3da5dab22e6 100644 --- a/src/topology/algebra/module/character_space.lean +++ b/src/topology/algebra/module/character_space.lean @@ -6,6 +6,7 @@ Authors: Frédéric Dupuis import topology.algebra.module.weak_dual import algebra.algebra.spectrum +import topology.continuous_function.algebra /-! # Character space of a topological algebra @@ -166,4 +167,28 @@ end ring end character_space +section gelfand_transform + +open continuous_map + +variables (𝕜 A) [comm_ring 𝕜] [no_zero_divisors 𝕜] [topological_space 𝕜] + [topological_ring 𝕜] [topological_space A] [semiring A] [algebra 𝕜 A] + +/-- The **Gelfand transform** is an algebra homomorphism (over `𝕜`) from a topological `𝕜`-algebra +`A` into the `𝕜`-algebra of continuous `𝕜`-valued functions on the `character_space 𝕜 A`. +The character space itself consists of all algebra homomorphisms from `A` to `𝕜`. -/ +@[simps] def gelfand_transform : A →ₐ[𝕜] C(character_space 𝕜 A, 𝕜) := +{ to_fun := λ a, + { to_fun := λ φ, φ a, + continuous_to_fun := (eval_continuous a).comp continuous_induced_dom }, + map_one' := by {ext, simp only [coe_mk, coe_one, pi.one_apply, map_one a] }, + map_mul' := λ a b, by {ext, simp only [map_mul, coe_mk, coe_mul, pi.mul_apply] }, + map_zero' := by {ext, simp only [map_zero, coe_mk, coe_mul, coe_zero, pi.zero_apply], }, + map_add' := λ a b, by {ext, simp only [map_add, coe_mk, coe_add, pi.add_apply] }, + commutes' := λ k, by {ext, simp only [alg_hom_class.commutes, algebra.id.map_eq_id, + ring_hom.id_apply, coe_mk, algebra_map_apply, algebra.id.smul_eq_mul, mul_one] } } + +end gelfand_transform + + end weak_dual From fd2fb7b36e6c90d8fbc42d7c48fe6d5ed23aa010 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Thu, 15 Sep 2022 00:06:23 +0000 Subject: [PATCH 267/828] feat(topology/unit_interval): add lemma add_pos (#16507) Add `lemma add_pos` stating that for every element of the unit interval and for every positive real number, their sum (as real number) is positive It is needed for #16029 defining H-spaces --- src/topology/unit_interval.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/topology/unit_interval.lean b/src/topology/unit_interval.lean index 0618eb9204244..098dfaf050564 100644 --- a/src/topology/unit_interval.lean +++ b/src/topology/unit_interval.lean @@ -101,6 +101,8 @@ lemma nonneg (x : I) : 0 ≤ (x : ℝ) := x.2.1 lemma one_minus_nonneg (x : I) : 0 ≤ 1 - (x : ℝ) := by simpa using x.2.2 lemma le_one (x : I) : (x : ℝ) ≤ 1 := x.2.2 lemma one_minus_le_one (x : I) : 1 - (x : ℝ) ≤ 1 := by simpa using x.2.1 +lemma add_pos {t : I} {x : ℝ} (hx : 0 < x) : 0 < (x + t : ℝ) := +add_pos_of_pos_of_nonneg hx $ nonneg _ /-- like `unit_interval.nonneg`, but with the inequality in `I`. -/ lemma nonneg' {t : I} : 0 ≤ t := t.2.1 From a4895002a025a68474efc68d389e060d440b9c80 Mon Sep 17 00:00:00 2001 From: Pim Otte Date: Thu, 15 Sep 2022 01:42:16 +0000 Subject: [PATCH 268/828] feat(data/sym/basic): add `fill_mem`, `append_mem` and supporting lemmas (#16486) Add lemmas for membership of fill and append, with supporting coercion and casting lemmas. --- src/data/sym/basic.lean | 34 +++++++++++++++++++++++----------- 1 file changed, 23 insertions(+), 11 deletions(-) diff --git a/src/data/sym/basic.lean b/src/data/sym/basic.lean index 89bb6b59db930..91bd2e9a887b1 100644 --- a/src/data/sym/basic.lean +++ b/src/data/sym/basic.lean @@ -53,7 +53,7 @@ local attribute [instance] vector.perm.is_setoid namespace sym -variables {α β : Type*} {n : ℕ} {s : sym α n} {a b : α} +variables {α β : Type*} {n n' m : ℕ} {s : sym α n} {a b : α} lemma coe_injective : injective (coe : sym α n → multiset α) := subtype.coe_injective @@ -117,10 +117,12 @@ s.1.decidable_mem _ @[simp] lemma mem_mk (a : α) (s : multiset α) (h : s.card = n) : a ∈ mk s h ↔ a ∈ s := iff.rfl -@[simp] lemma mem_cons {a b : α} {s : sym α n} : a ∈ b ::ₛ s ↔ a = b ∨ a ∈ s := +@[simp] lemma mem_cons : a ∈ b ::ₛ s ↔ a = b ∨ a ∈ s := multiset.mem_cons -lemma mem_cons_of_mem {a b : α} {s : sym α n} (h : a ∈ s) : a ∈ b ::ₛ s := +@[simp] lemma mem_coe : a ∈ (s : multiset α) ↔ a ∈ s := iff.rfl + +lemma mem_cons_of_mem (h : a ∈ s) : a ∈ b ::ₛ s := multiset.mem_cons_of_mem h @[simp] lemma mem_cons_self (a : α) (s : sym α n) : a ∈ a ::ₛ s := @@ -324,36 +326,46 @@ protected def cast {n m : ℕ} (h : n = m) : sym α n ≃ sym α m := @[simp] lemma cast_rfl : sym.cast rfl s = s := subtype.ext rfl -@[simp] lemma cast_cast {n n' n'' : ℕ} (s : sym α n) (h : n = n') (h' : n' = n'') : +@[simp] lemma cast_cast {n'' : ℕ} (h : n = n') (h' : n' = n'') : sym.cast h' (sym.cast h s) = sym.cast (h.trans h') s := rfl -@[simp] lemma coe_cast {n m : ℕ} (h : n = m) (s : sym α n) : - (sym.cast h s : multiset α) = s := rfl +@[simp] lemma coe_cast (h : n = m) : (sym.cast h s : multiset α) = s := rfl + +@[simp] lemma mem_cast (h : n = m) : a ∈ sym.cast h s ↔ a ∈ s := iff.rfl /-- Append a pair of `sym` terms. -/ -def append {n n' : ℕ} (s : sym α n) (s' : sym α n') : sym α (n + n') := +def append (s : sym α n) (s' : sym α n') : sym α (n + n') := ⟨s.1 + s'.1, by simp_rw [← s.2, ← s'.2, map_add]⟩ -@[simp] lemma append_inj_right {n n' : ℕ} (s : sym α n) (t t' : sym α n') : +@[simp] lemma append_inj_right (s : sym α n) {t t' : sym α n'} : s.append t = s.append t' ↔ t = t' := subtype.ext_iff.trans $ (add_right_inj _).trans subtype.ext_iff.symm -@[simp] lemma append_inj_left {n n' : ℕ} (s s' : sym α n) (t : sym α n') : +@[simp] lemma append_inj_left {s s' : sym α n} (t : sym α n') : s.append t = s'.append t ↔ s = s' := subtype.ext_iff.trans $ (add_left_inj _).trans subtype.ext_iff.symm -lemma append_comm {n n' : ℕ} (s : sym α n) (s' : sym α n') : +lemma append_comm (s : sym α n') (s' : sym α n') : s.append s' = sym.cast (add_comm _ _) (s'.append s) := by { ext, simp [append, add_comm], } +@[simp, norm_cast] lemma coe_append (s : sym α n') (s' : sym α n') : + (s.append s' : multiset α) = s + s' := rfl + +lemma mem_append_iff {s' : sym α m} : a ∈ s.append s' ↔ a ∈ s ∨ a ∈ s' := multiset.mem_add + /-- Fill a term `m : sym α (n - i)` with `i` copies of `a` to obtain a term of `sym α n`. This is a convenience wrapper for `m.append (repeat a i)` that adjusts the term using `sym.cast`. -/ def fill (a : α) (i : fin (n + 1)) (m : sym α (n - i)) : sym α n := sym.cast (nat.sub_add_cancel i.is_le) (m.append (repeat a i)) +lemma mem_fill_iff (a b : α) (i : fin (n + 1)) (s : sym α (n - i)) : + a ∈ sym.fill b i s ↔ ((i : ℕ) ≠ 0 ∧ a = b) ∨ a ∈ s := +by rw [fill, mem_cast, mem_append_iff, or_comm, mem_repeat] + /-- Remove every `a` from a given `sym α n`. Yields the number of copies `i` and a term of `sym α (n - i)`. -/ -def filter_ne [decidable_eq α] (a : α) {n : ℕ} (m : sym α n) : Σ i : fin (n + 1), sym α (n - i) := +def filter_ne [decidable_eq α] (a : α) (m : sym α n) : Σ i : fin (n + 1), sym α (n - i) := ⟨⟨m.1.count a, (multiset.count_le_card _ _).trans_lt $ by rw [m.2, nat.lt_succ_iff]⟩, m.1.filter ((≠) a), eq_tsub_of_add_eq $ eq.trans begin rw [← multiset.countp_eq_card_filter, add_comm], From 53e5dd300e5d752bda3516889049f7c35f3d4427 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 15 Sep 2022 04:20:53 +0000 Subject: [PATCH 269/828] chore(category/limits/opposites): instances (#16511) Upgrade some lemmas to instances, per [zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/limits.20in.20opposite.20categories). Co-authored-by: Scott Morrison --- src/algebraic_geometry/AffineScheme.lean | 1 - src/algebraic_geometry/pullbacks.lean | 2 -- src/category_theory/abelian/generator.lean | 1 - src/category_theory/abelian/opposite.lean | 1 - src/category_theory/generator.lean | 2 -- src/category_theory/limits/opposites.lean | 25 +++++++++++----------- 6 files changed, 12 insertions(+), 20 deletions(-) diff --git a/src/algebraic_geometry/AffineScheme.lean b/src/algebraic_geometry/AffineScheme.lean index c80a5ebebafd5..114556cb2195b 100644 --- a/src/algebraic_geometry/AffineScheme.lean +++ b/src/algebraic_geometry/AffineScheme.lean @@ -110,7 +110,6 @@ end instance : has_colimits AffineScheme.{u} := begin haveI := adjunction.has_limits_of_equivalence.{u} Γ.{u}, - haveI : has_colimits AffineScheme.{u} ᵒᵖᵒᵖ := has_colimits_op_of_has_limits, exactI adjunction.has_colimits_of_equivalence.{u} (op_op_equivalence AffineScheme.{u}).inverse end diff --git a/src/algebraic_geometry/pullbacks.lean b/src/algebraic_geometry/pullbacks.lean index 8b2e24f67d4d1..d3fbcd6d6fb5a 100644 --- a/src/algebraic_geometry/pullbacks.lean +++ b/src/algebraic_geometry/pullbacks.lean @@ -514,8 +514,6 @@ end lemma has_pullback_of_cover : has_pullback f g := ⟨⟨⟨_, glued_is_limit 𝒰 f g⟩⟩⟩ -instance : has_limits CommRingᵒᵖ := has_limits_op_of_has_colimits - instance affine_has_pullback {A B C : CommRing} (f : Spec.obj (opposite.op A) ⟶ Spec.obj (opposite.op C)) (g : Spec.obj (opposite.op B) ⟶ Spec.obj (opposite.op C)) : has_pullback f g := diff --git a/src/category_theory/abelian/generator.lean b/src/category_theory/abelian/generator.lean index e5ff33f877a57..fbde8c6e412c8 100644 --- a/src/category_theory/abelian/generator.lean +++ b/src/category_theory/abelian/generator.lean @@ -50,7 +50,6 @@ end theorem has_projective_separator [has_colimits C] [enough_projectives C] (G : C) (hG : is_coseparator G) : ∃ G : C, projective G ∧ is_separator G := begin - haveI : has_limits Cᵒᵖ := has_limits_op_of_has_colimits, obtain ⟨T, hT₁, hT₂⟩ := has_injective_coseparator (op G) ((is_separator_op_iff _).2 hG), exactI ⟨unop T, infer_instance, (is_separator_unop_iff _).2 hT₂⟩ end diff --git a/src/category_theory/abelian/opposite.lean b/src/category_theory/abelian/opposite.lean index 5f7725253b888..c89fbb827fada 100644 --- a/src/category_theory/abelian/opposite.lean +++ b/src/category_theory/abelian/opposite.lean @@ -23,7 +23,6 @@ variables (C : Type*) [category C] [abelian C] local attribute [instance] has_finite_limits_of_has_equalizers_and_finite_products has_finite_colimits_of_has_coequalizers_and_finite_coproducts - has_finite_limits_opposite has_finite_colimits_opposite has_finite_products_opposite instance : abelian Cᵒᵖ := { normal_mono_of_mono := λ X Y f m, by exactI diff --git a/src/category_theory/generator.lean b/src/category_theory/generator.lean index 23fa91a45b432..997fc72f54a32 100644 --- a/src/category_theory/generator.lean +++ b/src/category_theory/generator.lean @@ -146,7 +146,6 @@ lemma is_detecting.is_separating [has_equalizers C] {𝒢 : set C} (h𝒢 : is_d by exactI eq_of_epi_equalizer section -local attribute [instance] has_equalizers_opposite lemma is_codetecting.is_coseparating [has_coequalizers C] {𝒢 : set C} : is_codetecting 𝒢 → is_coseparating 𝒢 := @@ -293,7 +292,6 @@ end lemma has_terminal_of_is_separating [well_powered Cᵒᵖ] [has_colimits C] {𝒢 : set C} [small.{v₁} 𝒢] (h𝒢 : is_separating 𝒢) : has_terminal C := begin - haveI : has_limits Cᵒᵖ := has_limits_op_of_has_colimits, haveI : small.{v₁} 𝒢.op := small_of_injective (set.op_equiv_self 𝒢).injective, haveI : has_initial Cᵒᵖ := has_initial_of_is_coseparating ((is_coseparating_op_iff _).2 h𝒢), exact has_terminal_of_has_initial_op diff --git a/src/category_theory/limits/opposites.lean b/src/category_theory/limits/opposites.lean index d80da9d4b511f..0b0e88529db86 100644 --- a/src/category_theory/limits/opposites.lean +++ b/src/category_theory/limits/opposites.lean @@ -228,8 +228,7 @@ local attribute [instance] has_limits_of_shape_op_of_has_colimits_of_shape /-- If `C` has colimits, we can construct limits for `Cᵒᵖ`. -/ -lemma has_limits_op_of_has_colimits [has_colimits C] : has_limits Cᵒᵖ := ⟨infer_instance⟩ - +instance has_limits_op_of_has_colimits [has_colimits C] : has_limits Cᵒᵖ := ⟨infer_instance⟩ /-- If `F.left_op : Jᵒᵖ ⥤ C` has a limit, we can construct a colimit for `F : J ⥤ Cᵒᵖ`. @@ -251,13 +250,13 @@ local attribute [instance] has_colimits_of_shape_op_of_has_limits_of_shape /-- If `C` has limits, we can construct colimits for `Cᵒᵖ`. -/ -lemma has_colimits_op_of_has_limits [has_limits C] : has_colimits Cᵒᵖ := ⟨infer_instance⟩ +instance has_colimits_op_of_has_limits [has_limits C] : has_colimits Cᵒᵖ := ⟨infer_instance⟩ variables (X : Type v₁) /-- If `C` has products indexed by `X`, then `Cᵒᵖ` has coproducts indexed by `X`. -/ -lemma has_coproducts_opposite [has_products_of_shape X C] : +instance has_coproducts_opposite [has_products_of_shape X C] : has_coproducts_of_shape X Cᵒᵖ := begin haveI : has_limits_of_shape (discrete X)ᵒᵖ C := @@ -268,7 +267,7 @@ end /-- If `C` has coproducts indexed by `X`, then `Cᵒᵖ` has products indexed by `X`. -/ -lemma has_products_opposite [has_coproducts_of_shape X C] : +instance has_products_opposite [has_coproducts_of_shape X C] : has_products_of_shape X Cᵒᵖ := begin haveI : has_colimits_of_shape (discrete X)ᵒᵖ C := @@ -276,7 +275,7 @@ begin apply_instance end -lemma has_finite_coproducts_opposite [has_finite_products C] : has_finite_coproducts Cᵒᵖ := +instance has_finite_coproducts_opposite [has_finite_products C] : has_finite_coproducts Cᵒᵖ := { out := λ J 𝒟, begin resetI, haveI : has_limits_of_shape (discrete J)ᵒᵖ C := @@ -284,7 +283,7 @@ lemma has_finite_coproducts_opposite [has_finite_products C] : has_finite_coprod apply_instance, end } -lemma has_finite_products_opposite [has_finite_coproducts C] : has_finite_products Cᵒᵖ := +instance has_finite_products_opposite [has_finite_coproducts C] : has_finite_products Cᵒᵖ := { out := λ J 𝒟, begin resetI, haveI : has_colimits_of_shape (discrete J)ᵒᵖ C := @@ -292,36 +291,36 @@ lemma has_finite_products_opposite [has_finite_coproducts C] : has_finite_produc apply_instance, end } -lemma has_equalizers_opposite [has_coequalizers C] : has_equalizers Cᵒᵖ := +instance has_equalizers_opposite [has_coequalizers C] : has_equalizers Cᵒᵖ := begin haveI : has_colimits_of_shape walking_parallel_pairᵒᵖ C := has_colimits_of_shape_of_equivalence walking_parallel_pair_op_equiv, apply_instance end -lemma has_coequalizers_opposite [has_equalizers C] : has_coequalizers Cᵒᵖ := +instance has_coequalizers_opposite [has_equalizers C] : has_coequalizers Cᵒᵖ := begin haveI : has_limits_of_shape walking_parallel_pairᵒᵖ C := has_limits_of_shape_of_equivalence walking_parallel_pair_op_equiv, apply_instance end -lemma has_finite_colimits_opposite [has_finite_limits C] : +instance has_finite_colimits_opposite [has_finite_limits C] : has_finite_colimits Cᵒᵖ := { out := λ J 𝒟 𝒥, by { resetI, apply_instance, }, } -lemma has_finite_limits_opposite [has_finite_colimits C] : +instance has_finite_limits_opposite [has_finite_colimits C] : has_finite_limits Cᵒᵖ := { out := λ J 𝒟 𝒥, by { resetI, apply_instance, }, } -lemma has_pullbacks_opposite [has_pushouts C] : has_pullbacks Cᵒᵖ := +instance has_pullbacks_opposite [has_pushouts C] : has_pullbacks Cᵒᵖ := begin haveI : has_colimits_of_shape walking_cospanᵒᵖ C := has_colimits_of_shape_of_equivalence walking_cospan_op_equiv.symm, apply has_limits_of_shape_op_of_has_colimits_of_shape, end -lemma has_pushouts_opposite [has_pullbacks C] : has_pushouts Cᵒᵖ := +instance has_pushouts_opposite [has_pullbacks C] : has_pushouts Cᵒᵖ := begin haveI : has_limits_of_shape walking_spanᵒᵖ C := has_limits_of_shape_of_equivalence walking_span_op_equiv.symm, From a0ae2165d015aef83609c5d48f1e58320697df82 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 15 Sep 2022 06:43:26 +0000 Subject: [PATCH 270/828] feat(field_theory/splitting_field): generalize `splits` for `comm_ring K` (#16405) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The definition of `splits` has been slightly changed for this change. Hence 2 lemmas `splits_iff` and `splits.def` are added for the `field K` case. I guess it may be helpful if we are able to talk about something like `splits i (p : ℤ[X])`? --- src/field_theory/abel_ruffini.lean | 2 +- src/field_theory/normal.lean | 4 +- src/field_theory/splitting_field.lean | 219 ++++++++++++++++---------- 3 files changed, 137 insertions(+), 88 deletions(-) diff --git a/src/field_theory/abel_ruffini.lean b/src/field_theory/abel_ruffini.lean index aa9ee0ae382e9..2ecab32a062ae 100644 --- a/src/field_theory/abel_ruffini.lean +++ b/src/field_theory/abel_ruffini.lean @@ -133,7 +133,7 @@ begin have hn''' : (X ^ n - 1 : F[X]) ≠ 0 := λ h, one_ne_zero ((leading_coeff_X_pow_sub_one hn').symm.trans (congr_arg leading_coeff h)), have mem_range : ∀ {c}, c ^ n = 1 → ∃ d, algebra_map F (X ^ n - C a).splitting_field d = c := - λ c hc, ring_hom.mem_range.mp (minpoly.mem_range_of_degree_eq_one F c (or.resolve_left h hn''' + λ c hc, ring_hom.mem_range.mp (minpoly.mem_range_of_degree_eq_one F c (h.def.resolve_left hn''' (minpoly.irreducible ((splitting_field.normal (X ^ n - C a)).is_integral c)) (minpoly.dvd F c (by rwa [map_id, alg_hom.map_sub, sub_eq_zero, aeval_X_pow, aeval_one])))), apply is_solvable_of_comm, diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index 5e104283fa2ec..db49716d7fbc2 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -96,7 +96,7 @@ lemma alg_hom.normal_bijective [h : normal F E] (ϕ : E →ₐ[F] K) : function. ⟨ϕ.to_ring_hom.injective, λ x, by { letI : algebra E K := ϕ.to_ring_hom.to_algebra, obtain ⟨h1, h2⟩ := h.out (algebra_map K E x), - cases minpoly.mem_range_of_degree_eq_one E x (or.resolve_left h2 (minpoly.ne_zero h1) + cases minpoly.mem_range_of_degree_eq_one E x (h2.def.resolve_left (minpoly.ne_zero h1) (minpoly.irreducible (is_integral_of_is_scalar_tower x ((is_integral_algebra_map_iff (algebra_map K E).injective).mp h1))) (minpoly.dvd E x ((algebra_map K E).injective (by @@ -246,7 +246,7 @@ def alg_hom.restrict_normal_aux [h : normal F E] : rintros x ⟨y, ⟨z, hy⟩, hx⟩, rw [←hx, ←hy], apply minpoly.mem_range_of_degree_eq_one E, - exact or.resolve_left (h.splits z) (minpoly.ne_zero (h.is_integral z)) + exact or.resolve_left (h.splits z).def (minpoly.ne_zero (h.is_integral z)) (minpoly.irreducible $ is_integral_of_is_scalar_tower _ $ is_integral_alg_hom ϕ $ is_integral_alg_hom _ $ h.is_integral z) (minpoly.dvd E _ $ by rw [aeval_map, aeval_alg_hom, aeval_alg_hom, alg_hom.comp_apply, diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field.lean index b9c950db65b53..f4522aa6a7dcc 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field.lean @@ -17,8 +17,9 @@ if it is the smallest field extension of `K` such that `f` splits. ## Main definitions -* `polynomial.splits i f`: A predicate on a field homomorphism `i : K → L` and a polynomial `f` - saying that `f` is zero or all of its irreducible factors over `L` have degree `1`. +* `polynomial.splits i f`: A predicate on a homomorphism `i : K →+* L` from a commutative ring to a + field and a polynomial `f` saying that `f.map i` is zero or all of its irreducible factors over + `L` have degree `1`. * `polynomial.splitting_field f`: A fixed splitting field of the polynomial `f`. * `polynomial.is_splitting_field`: A predicate on a field to be a splitting field of a polynomial `f`. @@ -44,29 +45,32 @@ variables {F : Type u} {K : Type v} {L : Type w} namespace polynomial -variables [field K] [field L] [field F] open polynomial section splits +section comm_ring +variables [comm_ring K] [field L] [field F] variables (i : K →+* L) /-- A polynomial `splits` iff it is zero or all of its irreducible factors have `degree` 1. -/ def splits (f : K[X]) : Prop := -f = 0 ∨ ∀ {g : L[X]}, irreducible g → g ∣ f.map i → degree g = 1 +f.map i = 0 ∨ ∀ {g : L[X]}, irreducible g → g ∣ f.map i → degree g = 1 -@[simp] lemma splits_zero : splits i (0 : K[X]) := or.inl rfl +@[simp] lemma splits_zero : splits i (0 : K[X]) := or.inl (polynomial.map_zero i) -@[simp] lemma splits_C (a : K) : splits i (C a) := -if ha : a = 0 then ha.symm ▸ (@C_0 K _).symm ▸ splits_zero i -else -have hia : i a ≠ 0, from mt ((injective_iff_map_eq_zero i).1 i.injective _) ha, -or.inr $ λ g hg ⟨p, hp⟩, absurd hg.1 (not_not.2 (is_unit_iff_degree_eq_zero.2 $ - by have := congr_arg degree hp; - simp [degree_C hia, @eq_comm (with_bot ℕ) 0, - nat.with_bot.add_eq_zero_iff] at this; clear _fun_match; tauto)) +lemma splits_of_map_eq_C {f : K[X]} {a : L} (h : f.map i = C a) : splits i f := +if ha : a = 0 then or.inl (h.trans (ha.symm ▸ C_0)) +else or.inr $ λ g hg ⟨p, hp⟩, absurd hg.1 $ not_not.2 $ is_unit_iff_degree_eq_zero.2 $ +begin + have := congr_arg degree hp, + rw [h, degree_C ha, degree_mul, @eq_comm (with_bot ℕ) 0, nat.with_bot.add_eq_zero_iff] at this, + exact this.1, +end -lemma splits_of_degree_eq_one {f : K[X]} (hf : degree f = 1) : splits i f := +@[simp] lemma splits_C (a : K) : splits i (C a) := splits_of_map_eq_C i (map_C i) + +lemma splits_of_map_degree_eq_one {f : K[X]} (hf : degree (f.map i) = 1) : splits i f := or.inr $ λ g hg ⟨p, hp⟩, by have := congr_arg degree hp; simp [nat.with_bot.add_eq_one_iff, hf, @eq_comm (with_bot ℕ) 1, @@ -74,18 +78,16 @@ or.inr $ λ g hg ⟨p, hp⟩, clear _fun_match; tauto lemma splits_of_degree_le_one {f : K[X]} (hf : degree f ≤ 1) : splits i f := -begin - cases h : degree f with n, - { rw [degree_eq_bot.1 h]; exact splits_zero i }, - { cases n with n, - { rw [eq_C_of_degree_le_zero (trans_rel_right (≤) h le_rfl)]; - exact splits_C _ _ }, - { have hn : n = 0, - { rw h at hf, - cases n, { refl }, { exact absurd hf dec_trivial } }, - exact splits_of_degree_eq_one _ (by rw [h, hn]; refl) } } +if hif : degree (f.map i) ≤ 0 then splits_of_map_eq_C i (degree_le_zero_iff.mp hif) +else begin + push_neg at hif, + rw [← order.succ_le_iff, ← with_bot.coe_zero, with_bot.succ_coe, nat.succ_eq_succ] at hif, + exact splits_of_map_degree_eq_one i (le_antisymm ((degree_map_le i _).trans hf) hif), end +lemma splits_of_degree_eq_one {f : K[X]} (hf : degree f = 1) : splits i f := +splits_of_degree_le_one i hf.le + lemma splits_of_nat_degree_le_one {f : K[X]} (hf : nat_degree f ≤ 1) : splits i f := splits_of_degree_le_one i (degree_le_of_nat_degree_le hf) @@ -93,31 +95,19 @@ lemma splits_of_nat_degree_eq_one {f : K[X]} (hf : nat_degree f = 1) : splits i splits_of_nat_degree_le_one i (le_of_eq hf) lemma splits_mul {f g : K[X]} (hf : splits i f) (hg : splits i g) : splits i (f * g) := -if h : f * g = 0 then by simp [h] +if h : (f * g).map i = 0 then or.inl h else or.inr $ λ p hp hpf, ((principal_ideal_ring.irreducible_iff_prime.1 hp).2.2 _ _ (show p ∣ map i f * map i g, by convert hpf; rw polynomial.map_mul)).elim (hf.resolve_left (λ hf, by simpa [hf] using h) hp) (hg.resolve_left (λ hg, by simpa [hg] using h) hp) -lemma splits_of_splits_mul {f g : K[X]} (hfg : f * g ≠ 0) (h : splits i (f * g)) : +lemma splits_of_splits_mul' {f g : K[X]} (hfg : (f * g).map i ≠ 0) (h : splits i (f * g)) : splits i f ∧ splits i g := ⟨or.inr $ λ g hgi hg, or.resolve_left h hfg hgi (by rw polynomial.map_mul; exact hg.trans (dvd_mul_right _ _)), or.inr $ λ g hgi hg, or.resolve_left h hfg hgi (by rw polynomial.map_mul; exact hg.trans (dvd_mul_left _ _))⟩ -lemma splits_of_splits_of_dvd {f g : K[X]} (hf0 : f ≠ 0) (hf : splits i f) (hgf : g ∣ f) : - splits i g := -by { obtain ⟨f, rfl⟩ := hgf, exact (splits_of_splits_mul i hf0 hf).1 } - -lemma splits_of_splits_gcd_left {f g : K[X]} (hf0 : f ≠ 0) (hf : splits i f) : - splits i (euclidean_domain.gcd f g) := -polynomial.splits_of_splits_of_dvd i hf0 hf (euclidean_domain.gcd_dvd_left f g) - -lemma splits_of_splits_gcd_right {f g : K[X]} (hg0 : g ≠ 0) (hg : splits i g) : - splits i (euclidean_domain.gcd f g) := -polynomial.splits_of_splits_of_dvd i hg0 hg (euclidean_domain.gcd_dvd_right f g) - lemma splits_map_iff (j : L →+* F) {f : K[X]} : splits j (f.map i) ↔ splits (j.comp i) f := by simp [splits, polynomial.map_map] @@ -125,22 +115,14 @@ by simp [splits, polynomial.map_map] theorem splits_one : splits i 1 := splits_C i 1 -theorem splits_of_is_unit {u : K[X]} (hu : is_unit u) : u.splits i := -splits_of_splits_of_dvd i one_ne_zero (splits_one _) $ is_unit_iff_dvd_one.1 hu +theorem splits_of_is_unit [is_domain K] {u : K[X]} (hu : is_unit u) : u.splits i := +(is_unit_iff.mp hu).some_spec.2 ▸ splits_C _ _ theorem splits_X_sub_C {x : K} : (X - C x).splits i := -splits_of_degree_eq_one _ $ degree_X_sub_C x +splits_of_degree_le_one _ $ degree_X_sub_C_le _ theorem splits_X : X.splits i := -splits_of_degree_eq_one _ $ degree_X - -theorem splits_id_iff_splits {f : K[X]} : - (f.map i).splits (ring_hom.id L) ↔ f.splits i := -by rw [splits_map_iff, ring_hom.id_comp] - -theorem splits_mul_iff {f g : K[X]} (hf : f ≠ 0) (hg : g ≠ 0) : - (f * g).splits i ↔ f.splits i ∧ g.splits i := -⟨splits_of_splits_mul i (mul_ne_zero hf hg), λ ⟨hfs, hgs⟩, splits_mul i hfs hgs⟩ +splits_of_degree_le_one _ degree_X_le theorem splits_prod {ι : Type u} {s : ι → K[X]} {t : finset ι} : (∀ j ∈ t, (s j).splits i) → (∏ x in t, s x).splits i := @@ -158,6 +140,91 @@ end lemma splits_X_pow (n : ℕ) : (X ^ n).splits i := splits_pow i (splits_X i) n +theorem splits_id_iff_splits {f : K[X]} : + (f.map i).splits (ring_hom.id L) ↔ f.splits i := +by rw [splits_map_iff, ring_hom.id_comp] + +lemma exists_root_of_splits' {f : K[X]} (hs : splits i f) (hf0 : degree (f.map i) ≠ 0) : + ∃ x, eval₂ i x f = 0 := +if hf0' : f.map i = 0 then by simp [eval₂_eq_eval_map, hf0'] +else + let ⟨g, hg⟩ := wf_dvd_monoid.exists_irreducible_factor + (show ¬ is_unit (f.map i), from mt is_unit_iff_degree_eq_zero.1 hf0) hf0' in + let ⟨x, hx⟩ := exists_root_of_degree_eq_one (hs.resolve_left hf0' hg.1 hg.2) in + let ⟨i, hi⟩ := hg.2 in + ⟨x, by rw [← eval_map, hi, eval_mul, show _ = _, from hx, zero_mul]⟩ + +lemma roots_ne_zero_of_splits' {f : K[X]} (hs : splits i f) (hf0 : nat_degree (f.map i) ≠ 0) : + (f.map i).roots ≠ 0 := +let ⟨x, hx⟩ := exists_root_of_splits' i hs (λ h, hf0 $ nat_degree_eq_of_degree_eq_some h) in +λ h, by { rw ← eval_map at hx, + cases h.subst ((mem_roots _).2 hx), exact ne_zero_of_nat_degree_gt (nat.pos_of_ne_zero hf0) } + +/-- Pick a root of a polynomial that splits. See `root_of_splits` for polynomials over a field +which has simpler assumptions. -/ +def root_of_splits' {f : K[X]} (hf : f.splits i) (hfd : (f.map i).degree ≠ 0) : L := +classical.some $ exists_root_of_splits' i hf hfd + +theorem map_root_of_splits' {f : K[X]} (hf : f.splits i) (hfd) : + f.eval₂ i (root_of_splits' i hf hfd) = 0 := +classical.some_spec $ exists_root_of_splits' i hf hfd + +lemma nat_degree_eq_card_roots' {p : K[X]} {i : K →+* L} + (hsplit : splits i p) : (p.map i).nat_degree = (p.map i).roots.card := +begin + by_cases hp : p.map i = 0, + { rw [hp, nat_degree_zero, roots_zero, multiset.card_zero] }, + obtain ⟨q, he, hd, hr⟩ := exists_prod_multiset_X_sub_C_mul (p.map i), + rw [← splits_id_iff_splits, ← he] at hsplit, + rw ← he at hp, + have hq : q ≠ 0 := λ h, hp (by rw [h, mul_zero]), + rw [← hd, add_right_eq_self], + by_contra, + have h' : (map (ring_hom.id L) q).nat_degree ≠ 0, { simp [h], }, + have := roots_ne_zero_of_splits' (ring_hom.id L) (splits_of_splits_mul' _ _ hsplit).2 h', + { rw map_id at this, exact this hr }, + { rw [map_id], exact mul_ne_zero monic_prod_multiset_X_sub_C.ne_zero hq }, +end + +lemma degree_eq_card_roots' {p : K[X]} {i : K →+* L} (p_ne_zero : p.map i ≠ 0) + (hsplit : splits i p) : (p.map i).degree = (p.map i).roots.card := +by rw [degree_eq_nat_degree p_ne_zero, nat_degree_eq_card_roots' hsplit] + +end comm_ring + +variables [field K] [field L] [field F] +variables (i : K →+* L) + +/-- This lemma is for polynomials over a field. -/ +lemma 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] + +/-- This lemma is for polynomials over a field. -/ +lemma splits.def {i : K →+* L} {f : K[X]} (h : splits i f) : + f = 0 ∨ ∀ {g : L[X]}, irreducible g → g ∣ f.map i → degree g = 1 := +(splits_iff i f).mp h + +lemma splits_of_splits_mul {f g : K[X]} (hfg : f * g ≠ 0) (h : splits i (f * g)) : + splits i f ∧ splits i g := +splits_of_splits_mul' i (map_ne_zero hfg) h + +lemma splits_of_splits_of_dvd {f g : K[X]} (hf0 : f ≠ 0) (hf : splits i f) (hgf : g ∣ f) : + splits i g := +by { obtain ⟨f, rfl⟩ := hgf, exact (splits_of_splits_mul i hf0 hf).1 } + +lemma splits_of_splits_gcd_left {f g : K[X]} (hf0 : f ≠ 0) (hf : splits i f) : + splits i (euclidean_domain.gcd f g) := +polynomial.splits_of_splits_of_dvd i hf0 hf (euclidean_domain.gcd_dvd_left f g) + +lemma splits_of_splits_gcd_right {f g : K[X]} (hg0 : g ≠ 0) (hg : splits i g) : + splits i (euclidean_domain.gcd f g) := +polynomial.splits_of_splits_of_dvd i hg0 hg (euclidean_domain.gcd_dvd_right f g) + +theorem splits_mul_iff {f g : K[X]} (hf : f ≠ 0) (hg : g ≠ 0) : + (f * g).splits i ↔ f.splits i ∧ g.splits i := +⟨splits_of_splits_mul i (mul_ne_zero hf hg), λ ⟨hfs, hgs⟩, splits_mul i hfs hgs⟩ + theorem splits_prod_iff {ι : Type u} {s : ι → K[X]} {t : finset ι} : (∀ j ∈ t, s j ≠ 0) → ((∏ x in t, s x).splits i ↔ ∀ j ∈ t, (s j).splits i) := begin @@ -166,57 +233,39 @@ begin rw [finset.prod_insert hat, splits_mul_iff i ht.1 (finset.prod_ne_zero_iff.2 ht.2), ih ht.2] end -lemma degree_eq_one_of_irreducible_of_splits {p : L[X]} - (hp : irreducible p) (hp_splits : splits (ring_hom.id L) p) : +lemma degree_eq_one_of_irreducible_of_splits {p : K[X]} + (hp : irreducible p) (hp_splits : splits (ring_hom.id K) p) : p.degree = 1 := begin - by_cases h_nz : p = 0, - { exfalso, simp * at *, }, rcases hp_splits, - { contradiction }, + { exfalso, simp * at *, }, { apply hp_splits hp, simp } end lemma exists_root_of_splits {f : K[X]} (hs : splits i f) (hf0 : degree f ≠ 0) : ∃ x, eval₂ i x f = 0 := -if hf0 : f = 0 then by simp [hf0] -else - let ⟨g, hg⟩ := wf_dvd_monoid.exists_irreducible_factor - (show ¬ is_unit (f.map i), from mt is_unit_iff_degree_eq_zero.1 (by rwa degree_map)) - (map_ne_zero hf0) in - let ⟨x, hx⟩ := exists_root_of_degree_eq_one (hs.resolve_left hf0 hg.1 hg.2) in - let ⟨i, hi⟩ := hg.2 in - ⟨x, by rw [← eval_map, hi, eval_mul, show _ = _, from hx, zero_mul]⟩ +exists_root_of_splits' i hs ((f.degree_map i).symm ▸ hf0) lemma roots_ne_zero_of_splits {f : K[X]} (hs : splits i f) (hf0 : nat_degree f ≠ 0) : (f.map i).roots ≠ 0 := -let ⟨x, hx⟩ := exists_root_of_splits i hs (λ h, hf0 $ nat_degree_eq_of_degree_eq_some h) in -λ h, by { rw ← eval_map at hx, - cases h.subst ((mem_roots _).2 hx), exact map_ne_zero (λ h, (h.subst hf0) rfl) } +roots_ne_zero_of_splits' i hs (ne_of_eq_of_ne (nat_degree_map i) hf0) -/-- Pick a root of a polynomial that splits. -/ +/-- Pick a root of a polynomial that splits. This version is for polynomials over a field and has +simpler assumptions. -/ def root_of_splits {f : K[X]} (hf : f.splits i) (hfd : f.degree ≠ 0) : L := -classical.some $ exists_root_of_splits i hf hfd +root_of_splits' i hf ((f.degree_map i).symm ▸ hfd) + +/-- `root_of_splits'` is definitionally equal to `root_of_splits`. -/ +lemma root_of_splits'_eq_root_of_splits {f : K[X]} (hf : f.splits i) (hfd) : + root_of_splits' i hf hfd = root_of_splits i hf (f.degree_map i ▸ hfd) := rfl theorem map_root_of_splits {f : K[X]} (hf : f.splits i) (hfd) : f.eval₂ i (root_of_splits i hf hfd) = 0 := -classical.some_spec $ exists_root_of_splits i hf hfd +map_root_of_splits' i hf (ne_of_eq_of_ne (degree_map f i) hfd) lemma nat_degree_eq_card_roots {p : K[X]} {i : K →+* L} (hsplit : splits i p) : p.nat_degree = (p.map i).roots.card := -begin - by_cases hp : p = 0, - { rw [hp, nat_degree_zero, polynomial.map_zero, roots_zero, multiset.card_zero] }, - obtain ⟨q, he, hd, hr⟩ := exists_prod_multiset_X_sub_C_mul (p.map i), - rw [← splits_id_iff_splits, ← he] at hsplit, - have hpm : p.map i ≠ 0 := map_ne_zero hp, rw ← he at hpm, - have hq : q ≠ 0 := λ h, hpm (by rw [h, mul_zero]), - rw [← nat_degree_map i, ← hd, add_right_eq_self], - by_contra, - have := roots_ne_zero_of_splits (ring_hom.id L) (splits_of_splits_mul _ _ hsplit).2 h, - { rw map_id at this, exact this hr }, - { exact mul_ne_zero monic_prod_multiset_X_sub_C.ne_zero hq }, -end +(nat_degree_map i).symm.trans $ nat_degree_eq_card_roots' hsplit lemma degree_eq_card_roots {p : K[X]} {i : K →+* L} (p_ne_zero : p ≠ 0) (hsplit : splits i p) : p.degree = (p.map i).roots.card := @@ -282,7 +331,7 @@ open unique_factorization_monoid associates lemma splits_of_exists_multiset {f : K[X]} {s : multiset L} (hs : f.map i = C (i f.leading_coeff) * (s.map (λ a : L, X - C a)).prod) : splits i f := -if hf0 : f = 0 then or.inl hf0 +if hf0 : f = 0 then hf0.symm ▸ splits_zero i else or.inr $ λ p hp hdp, begin rw irreducible_iff_prime at hp, rw [hs, ← multiset.prod_to_list] at hdp, @@ -295,13 +344,13 @@ else or.inr $ λ p hp hdp, begin exact degree_X_sub_C a }, end -lemma splits_of_splits_id {f : K[X]} : splits (ring_hom.id _) f → splits i f := +lemma splits_of_splits_id {f : K[X]} : splits (ring_hom.id K) f → splits i f := unique_factorization_monoid.induction_on_prime f (λ _, splits_zero _) (λ _ hu _, splits_of_degree_le_one _ ((is_unit_iff_degree_eq_zero.1 hu).symm ▸ dec_trivial)) (λ a p ha0 hp ih hfi, splits_mul _ (splits_of_degree_eq_one _ - ((splits_of_splits_mul _ (mul_ne_zero hp.1 ha0) hfi).1.resolve_left + ((splits_of_splits_mul _ (mul_ne_zero hp.1 ha0) hfi).1.def.resolve_left hp.1 hp.irreducible (by rw map_id))) (ih (splits_of_splits_mul _ (mul_ne_zero hp.1 ha0) hfi).2)) From 7afe23e08102bfcc6d58cd33cc721c9f7b8c3942 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 15 Sep 2022 09:41:40 +0000 Subject: [PATCH 271/828] feat(linear_algebra/affine_space/affine_subspace): more `mem_map` lemmas (#16378) Add two more lemmas about membership of an affine subspace given by `affine_subspace.map`. --- src/linear_algebra/affine_space/affine_subspace.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index cdf63befdda3b..19d7d019d5e01 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -1252,6 +1252,13 @@ def map (s : affine_subspace k P₁) : affine_subspace k P₂ := @[simp] lemma mem_map {f : P₁ →ᵃ[k] P₂} {x : P₂} {s : affine_subspace k P₁} : x ∈ s.map f ↔ ∃ y ∈ s, f y = x := mem_image_iff_bex +lemma mem_map_of_mem {x : P₁} {s : affine_subspace k P₁} (h : x ∈ s) : f x ∈ s.map f := +set.mem_image_of_mem _ h + +lemma mem_map_iff_mem_of_injective {f : P₁ →ᵃ[k] P₂} {x : P₁} {s : affine_subspace k P₁} + (hf : function.injective f) : f x ∈ s.map f ↔ x ∈ s := +hf.mem_set_image + @[simp] lemma map_bot : (⊥ : affine_subspace k P₁).map f = ⊥ := coe_injective $ image_empty f From 400aa37580b65418f114c0314f822f1a1943d914 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 15 Sep 2022 09:41:42 +0000 Subject: [PATCH 272/828] feat(data/polynomial/eval): add `map_{eq/ne}_zero_iff` (#16440) Co-authored-by: Junyan Xu --- src/data/polynomial/eval.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/data/polynomial/eval.lean b/src/data/polynomial/eval.lean index 76c84b381ed03..9e677fd6dc6c5 100644 --- a/src/data/polynomial/eval.lean +++ b/src/data/polynomial/eval.lean @@ -619,6 +619,12 @@ nat_degree_le_nat_degree (degree_map_le f p) variables {f} +protected lemma map_eq_zero_iff (hf : function.injective f) : p.map f = 0 ↔ p = 0 := +map_eq_zero_iff (map_ring_hom f) (map_injective f hf) + +protected lemma map_ne_zero_iff (hf : function.injective f) : p.map f ≠ 0 ↔ p ≠ 0 := +(polynomial.map_eq_zero_iff hf).not + lemma map_monic_eq_zero_iff (hp : p.monic) : p.map f = 0 ↔ ∀ x, f x = 0 := ⟨ λ hfp x, calc f x = f x * f p.leading_coeff : by simp only [mul_one, hp.leading_coeff, f.map_one] ... = f x * (p.map f).coeff p.nat_degree : congr_arg _ (coeff_map _ _).symm From 0f287f82063da0fbaa21341d8fdb7122d4db1210 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 15 Sep 2022 09:41:43 +0000 Subject: [PATCH 273/828] chore(topology/separation): move/add `simp` lemmas (#16470) * add `inseparable_iff_eq` and `inseparable_eq_eq`; * add `specializes_eq_eq`, move `@[simp]` from `specializes_iff_eq`; * golf `pure_le_nhds_iff` and `nhds_le_nhds_iff`. --- src/topology/separation.lean | 28 ++++++++++++++++------------ 1 file changed, 16 insertions(+), 12 deletions(-) diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 7f254e8977f1b..6123bd318017a 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -178,9 +178,15 @@ t0_space_iff_inseparable α lemma nhds_injective [t0_space α] : injective (𝓝 : α → filter α) := (t0_space_iff_nhds_injective α).1 ‹_› +lemma inseparable_iff_eq [t0_space α] {x y : α} : inseparable x y ↔ x = y := +nhds_injective.eq_iff + @[simp] lemma nhds_eq_nhds_iff [t0_space α] {a b : α} : 𝓝 a = 𝓝 b ↔ a = b := nhds_injective.eq_iff +@[simp] lemma inseparable_eq_eq [t0_space α] : inseparable = @eq α := +funext₂ $ λ x y, propext inseparable_iff_eq + lemma t0_space_iff_exists_is_open_xor_mem (α : Type u) [topological_space α] : t0_space α ↔ ∀ x y, x ≠ y → ∃ U:set α, is_open U ∧ (xor (x ∈ U) (y ∈ U)) := by simp only [t0_space_iff_not_inseparable, xor_iff_not_iff, not_forall, exists_prop, @@ -438,9 +444,18 @@ t1_space_iff_disjoint_nhds_pure.mp ‹_› h lemma specializes.eq [t1_space α] {x y : α} (h : x ⤳ y) : x = y := t1_space_iff_specializes_imp_eq.1 ‹_› h -@[simp] lemma specializes_iff_eq [t1_space α] {x y : α} : x ⤳ y ↔ x = y := +lemma specializes_iff_eq [t1_space α] {x y : α} : x ⤳ y ↔ x = y := ⟨specializes.eq, λ h, h ▸ specializes_rfl⟩ +@[simp] lemma specializes_eq_eq [t1_space α] : (⤳) = @eq α := +funext₂ $ λ x y, propext specializes_iff_eq + +@[simp] lemma pure_le_nhds_iff [t1_space α] {a b : α} : pure a ≤ 𝓝 b ↔ a = b := +specializes_iff_pure.symm.trans specializes_iff_eq + +@[simp] lemma nhds_le_nhds_iff [t1_space α] {a b : α} : 𝓝 a ≤ 𝓝 b ↔ a = b := +specializes_iff_eq + instance {α : Type*} : t1_space (cofinite_topology α) := t1_space_iff_continuous_cofinite_of.mpr continuous_id @@ -533,17 +548,6 @@ begin exact ⟨i, hi, λ h, hsub h rfl⟩ end -@[simp] lemma pure_le_nhds_iff [t1_space α] {a b : α} : pure a ≤ 𝓝 b ↔ a = b := -begin - refine ⟨λ h, _, λ h, h ▸ pure_le_nhds a⟩, - by_contra hab, - simpa only [mem_pure, mem_compl_iff, mem_singleton, not_true] using - h (compl_singleton_mem_nhds $ ne.symm hab) -end - -@[simp] lemma nhds_le_nhds_iff [t1_space α] {a b : α} : 𝓝 a ≤ 𝓝 b ↔ a = b := -⟨λ h, pure_le_nhds_iff.mp $ (pure_le_nhds a).trans h, λ h, h ▸ le_rfl⟩ - @[simp] lemma compl_singleton_mem_nhds_set_iff [t1_space α] {x : α} {s : set α} : {x}ᶜ ∈ 𝓝ˢ s ↔ x ∉ s := by rwa [is_open_compl_singleton.mem_nhds_set, subset_compl_singleton_iff] From 2178a7f76c7a435ab9f02c4589ac4714726d3c7a Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Thu, 15 Sep 2022 11:57:44 +0000 Subject: [PATCH 274/828] feat(ring_theory/{ideal/basic, adjoin_root): some lemmas from #15000 (#16450) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR contains some lemmas that were originally in #15000. The main result that is proven here is `quotient_equiv_quotient_minpoly_map`, which says that if `α` has minimal polynomial `f` over `R` and `I` is an ideal of `R`, then rings `R[α] / I[α]` and `(R/I)[X] / (f mod I)` are isomorphic as R-algebras. Co-authored-by: Anne Baanen --- src/algebra/algebra/basic.lean | 24 +++++++-- src/ring_theory/adjoin_root.lean | 77 ++++++++++++++++++++++++++- src/ring_theory/ideal/operations.lean | 28 +++++++++- 3 files changed, 124 insertions(+), 5 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index 61c3425366a65..75def7eace672 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -907,6 +907,7 @@ rfl @[simp] lemma to_ring_equiv_eq_coe : e.to_ring_equiv = e := rfl @[simp, norm_cast] lemma coe_ring_equiv : ((e : A₁ ≃+* A₂) : A₁ → A₂) = e := rfl + lemma coe_ring_equiv' : (e.to_ring_equiv : A₁ → A₂) = e := rfl lemma coe_ring_equiv_injective : function.injective (coe : (A₁ ≃ₐ[R] A₂) → (A₁ ≃+* A₂)) := @@ -996,8 +997,12 @@ symm_bijective.injective $ ext $ λ x, rfl { to_fun := f', inv_fun := f, ..(⟨f, f', h₁, h₂, h₃, h₄, h₅⟩ : A₁ ≃ₐ[R] A₂).symm } := rfl -@[simp] -theorem refl_symm : (alg_equiv.refl : A₁ ≃ₐ[R] A₁).symm = alg_equiv.refl := rfl +@[simp] theorem refl_symm : (alg_equiv.refl : A₁ ≃ₐ[R] A₁).symm = alg_equiv.refl := rfl + +--this should be a simp lemma but causes a lint timeout +lemma to_ring_equiv_symm (f : A₁ ≃ₐ[R] A₁) : (f : A₁ ≃+* A₁).symm = f.symm := rfl + +@[simp] lemma symm_to_ring_equiv : (e.symm : A₂ ≃+* A₁) = (e : A₁ ≃+* A₂).symm := rfl /-- Algebra equivalences are transitive. -/ @[trans] @@ -1166,7 +1171,20 @@ by { ext, refl } end of_linear_equiv -@[simps mul one inv {attrs := []}] instance aut : group (A₁ ≃ₐ[R] A₁) := +section of_ring_equiv + +/-- Promotes a linear ring_equiv to an alg_equiv. -/ +@[simps] +def of_ring_equiv {f : A₁ ≃+* A₂} + (hf : ∀ x, f (algebra_map R A₁ x) = algebra_map R A₂ x) : A₁ ≃ₐ[R] A₂ := +{ to_fun := f, + inv_fun := f.symm, + commutes' := hf, + .. f } + +end of_ring_equiv + +@[simps mul one {attrs := []}] instance aut : group (A₁ ≃ₐ[R] A₁) := { mul := λ ϕ ψ, ψ.trans ϕ, mul_assoc := λ ϕ ψ χ, rfl, one := refl, diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index 3156cdeb9947a..a7da2bea5ea4e 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Chris Hughes -/ +import algebra.algebra.basic import data.polynomial.field_division import linear_algebra.finite_dimensional import ring_theory.adjoin.basic @@ -608,7 +609,7 @@ by simp only [polynomial.quot_quot_equiv_comm, quotient_equiv_symm_mk, polynomial_quotient_equiv_quotient_polynomial_symm_mk] /-- The natural isomorphism `R[α]/I[α] ≅ (R/I)[X]/(f mod I)` for `α` a root of `f : polynomial R` - and `I : ideal R`-/ + and `I : ideal R`.-/ def quot_adjoin_root_equiv_quot_polynomial_quot : (adjoin_root f) ⧸ (I.map (of f)) ≃+* polynomial (R ⧸ I) ⧸ (span ({f.map (I^.quotient.mk)} : set (polynomial (R ⧸ I)))) := (quot_map_of_equiv_quot_map_C_map_span_mk I f).trans @@ -641,6 +642,80 @@ by rw [quot_adjoin_root_equiv_quot_polynomial_quot, ring_equiv.symm_trans_apply, quot_map_C_map_span_mk_equiv_quot_map_C_quot_map_span_mk_symm_quot_quot_mk, quot_map_of_equiv_quot_map_C_map_span_mk_symm_mk] +/-- Promote `adjoin_root.quot_adjoin_root_equiv_quot_polynomial_quot` to an alg_equiv. -/ +@[simps apply symm_apply] +noncomputable def quot_equiv_quot_map (f : R[X]) (I : ideal R) : + ((adjoin_root f) ⧸ (ideal.map (of f) I)) ≃ₐ[R] + ((R ⧸ I) [X]) ⧸ (ideal.span ({polynomial.map I^.quotient.mk f} : set ((R ⧸ I) [X]))) := +alg_equiv.of_ring_equiv (show ∀ x, (quot_adjoin_root_equiv_quot_polynomial_quot I f) + (algebra_map R _ x) = algebra_map R _ x, from λ x, begin + have : algebra_map R ((adjoin_root f) ⧸ (ideal.map (of f) I)) x = ideal.quotient.mk + (ideal.map (adjoin_root.of f) I) ((mk f) (C x)) := rfl, + simpa only [this, quot_adjoin_root_equiv_quot_polynomial_quot_mk_of, map_C] + end) + +@[simp] +lemma quot_equiv_quot_map_apply_mk (f g : polynomial R) (I : ideal R) : + adjoin_root.quot_equiv_quot_map f I (ideal.quotient.mk _ (adjoin_root.mk f g)) = + ideal.quotient.mk _ (g.map I^.quotient.mk) := +by rw [adjoin_root.quot_equiv_quot_map_apply, + adjoin_root.quot_adjoin_root_equiv_quot_polynomial_quot_mk_of] + +@[simp] +lemma quot_equiv_quot_map_symm_apply_mk (f g : polynomial R) (I : ideal R) : + (adjoin_root.quot_equiv_quot_map f I).symm (ideal.quotient.mk _ (map (ideal.quotient.mk I) g)) = + ideal.quotient.mk _ (adjoin_root.mk f g) := +by rw [adjoin_root.quot_equiv_quot_map_symm_apply, + adjoin_root.quot_adjoin_root_equiv_quot_polynomial_quot_symm_mk_mk] + end end adjoin_root + +namespace power_basis + +open adjoin_root alg_equiv + +variables [comm_ring R] [is_domain R] [comm_ring S] [is_domain S] [algebra R S] + +/-- Let `α` have minimal polynomial `f` over `R` and `I` be an ideal of `R`, +then `R[α] / (I) = (R[x] / (f)) / pS = (R/p)[x] / (f mod p)`. -/ +@[simps apply symm_apply] +noncomputable def quotient_equiv_quotient_minpoly_map (pb : power_basis R S) + (I : ideal R) : + (S ⧸ I.map (algebra_map R S)) ≃ₐ[R] (polynomial (R ⧸ I)) ⧸ + (ideal.span ({(minpoly R pb.gen).map I^.quotient.mk} : set (polynomial (R ⧸ I)))) := +(of_ring_equiv + (show ∀ x, (ideal.quotient_equiv _ (ideal.map (adjoin_root.of (minpoly R pb.gen)) I) + (adjoin_root.equiv' (minpoly R pb.gen) pb + (by rw [adjoin_root.aeval_eq, adjoin_root.mk_self]) + (minpoly.aeval _ _)).symm.to_ring_equiv + (by rw [ideal.map_map, alg_equiv.to_ring_equiv_eq_coe, ← alg_equiv.coe_ring_hom_commutes, + ← adjoin_root.algebra_map_eq, alg_hom.comp_algebra_map])) + (algebra_map R (S ⧸ I.map (algebra_map R S)) x) = algebra_map R _ x, from + (λ x, by rw [← ideal.quotient.mk_algebra_map, ideal.quotient_equiv_apply, + ring_hom.to_fun_eq_coe, ideal.quotient_map_mk, alg_equiv.to_ring_equiv_eq_coe, + ring_equiv.coe_to_ring_hom, alg_equiv.coe_ring_equiv, alg_equiv.commutes, + quotient.mk_algebra_map]))).trans (adjoin_root.quot_equiv_quot_map _ _) + +@[simp] +lemma quotient_equiv_quotient_minpoly_map_apply_mk (pb : power_basis R S) (I : ideal R) + (g : polynomial R) : pb.quotient_equiv_quotient_minpoly_map I + (ideal.quotient.mk _ (aeval pb.gen g)) = ideal.quotient.mk _ (g.map I^.quotient.mk) := +by rw [power_basis.quotient_equiv_quotient_minpoly_map, alg_equiv.trans_apply, + alg_equiv.of_ring_equiv_apply, quotient_equiv_mk, alg_equiv.coe_ring_equiv', + adjoin_root.equiv'_symm_apply, power_basis.lift_aeval, + adjoin_root.aeval_eq, adjoin_root.quot_equiv_quot_map_apply_mk] + +@[simp] +lemma quotient_equiv_quotient_minpoly_map_symm_apply_mk (pb : power_basis R S) (I : ideal R) + (g : polynomial R) : (pb.quotient_equiv_quotient_minpoly_map I).symm + (ideal.quotient.mk _ (g.map I^.quotient.mk)) = (ideal.quotient.mk _ (aeval pb.gen g)) := +begin simp only [quotient_equiv_quotient_minpoly_map, to_ring_equiv_eq_coe, symm_trans_apply, + quot_equiv_quot_map_symm_apply_mk, of_ring_equiv_symm_apply, quotient_equiv_symm_mk, + to_ring_equiv_symm, ring_equiv.symm_symm, adjoin_root.equiv'_apply, coe_ring_equiv, + lift_hom_mk, symm_to_ring_equiv], + +end + +end power_basis diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index b2ce5384f0c12..edb84dddcba19 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -2109,7 +2109,11 @@ end ring_hom namespace double_quot open ideal -variables {R : Type u} [comm_ring R] (I J : ideal R) +variable {R : Type u} + +section + +variables [comm_ring R] (I J : ideal R) /-- The obvious ring hom `R/I → R/(I ⊔ J)` -/ def quot_left_to_quot_sup : R ⧸ I →+* R ⧸ (I ⊔ J) := @@ -2178,4 +2182,26 @@ lemma quot_quot_equiv_comm_symm : (quot_quot_equiv_comm I J).symm = quot_quot_equiv_comm J I := rfl +end + +section algebra + +@[simp] +lemma quot_quot_equiv_comm_mk_mk [comm_ring R] (I J : ideal R) (x : R) : + quot_quot_equiv_comm I J (ideal.quotient.mk _ (ideal.quotient.mk _ x)) = + algebra_map R _ x := rfl + +variables [comm_semiring R] {A : Type v} [comm_ring A] [algebra R A] (I J : ideal A) + +@[simp] +lemma quot_quot_equiv_quot_sup_quot_quot_algebra_map (x : R) : + double_quot.quot_quot_equiv_quot_sup I J (algebra_map R _ x) = algebra_map _ _ x := +rfl + +@[simp] +lemma quot_quot_equiv_comm_algebra_map (x : R) : + quot_quot_equiv_comm I J (algebra_map R _ x) = algebra_map _ _ x := rfl + +end algebra + end double_quot From 54882236dea768346062be46a12c966861c49c7c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Thu, 15 Sep 2022 11:57:46 +0000 Subject: [PATCH 275/828] feat(probabiliy/martingale): centering lemma (#16517) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We provide the decomposition of an adapted and integrable stochastic process into a predictable part and a martingale part. Co-authored-by: Rémy Degenne --- src/probability/martingale/centering.lean | 139 ++++++++++++++++++++++ src/probability/stopping.lean | 5 + 2 files changed, 144 insertions(+) create mode 100644 src/probability/martingale/centering.lean diff --git a/src/probability/martingale/centering.lean b/src/probability/martingale/centering.lean new file mode 100644 index 0000000000000..9e244d8b878b1 --- /dev/null +++ b/src/probability/martingale/centering.lean @@ -0,0 +1,139 @@ +/- +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 probability.martingale.basic + +/-! +# Centering lemma for stochastic processes + +Any `ℕ`-indexed stochastic process which is adapted and integrable can be written as the sum of a +martingale and a predictable process. This result is also known as **Doob's decomposition theorem**. +From a process `f`, a filtration `ℱ` and a measure `μ`, we define two processes +`martingale_part ℱ μ f` and `predictable_part ℱ μ f`. + +## Main definitions + +* `measure_theory.predictable_part ℱ μ f`: a predictable process such that + `f = predictable_part ℱ μ f + martingale_part ℱ μ f` +* `measure_theory.martingale_part ℱ μ f`: a martingale such that + `f = predictable_part ℱ μ f + martingale_part ℱ μ f` + +## Main statements + +* `measure_theory.adapted_predictable_part`: `(λ n, predictable_part ℱ μ f (n+1))` is adapted. That + is, `predictable_part` is predictable. +* `measure_theory.martingale_martingale_part`: `martingale_part ℱ μ f` is a martingale. + +-/ + + +open topological_space filter +open_locale nnreal ennreal measure_theory probability_theory big_operators + +namespace measure_theory + +variables {Ω E : Type*} {m0 : measurable_space Ω} {μ : measure Ω} + [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + {f : ℕ → Ω → E} {ℱ : filtration ℕ m0} {n : ℕ} + +/-- Any `ℕ`-indexed stochastic process can be written as the sum of a martingale and a predictable +process. This is the predictable process. See `martingale_part` for the martingale. -/ +noncomputable +def predictable_part {m0 : measurable_space Ω} + (ℱ : filtration ℕ m0) (μ : measure Ω) (f : ℕ → Ω → E) : ℕ → Ω → E := +λ n, ∑ i in finset.range n, μ[f (i+1) - f i | ℱ i] + +@[simp] lemma predictable_part_zero : predictable_part ℱ μ f 0 = 0 := +by simp_rw [predictable_part, finset.range_zero, finset.sum_empty] + +lemma adapted_predictable_part : adapted ℱ (λ n, predictable_part ℱ μ f (n+1)) := +λ n, finset.strongly_measurable_sum' _ + (λ i hin, strongly_measurable_condexp.mono (ℱ.mono (finset.mem_range_succ_iff.mp hin))) + +lemma adapted_predictable_part' : adapted ℱ (λ n, predictable_part ℱ μ f n) := +λ n, finset.strongly_measurable_sum' _ + (λ i hin, strongly_measurable_condexp.mono (ℱ.mono (finset.mem_range_le hin))) + +/-- Any `ℕ`-indexed stochastic process can be written as the sum of a martingale and a predictable +process. This is the martingale. See `predictable_part` for the predictable process. -/ +noncomputable +def martingale_part {m0 : measurable_space Ω} + (ℱ : filtration ℕ m0) (μ : measure Ω) (f : ℕ → Ω → E) : ℕ → Ω → E := +λ n, f n - predictable_part ℱ μ f n + +lemma martingale_part_add_predictable_part (f : ℕ → Ω → E) : + martingale_part ℱ μ f + predictable_part ℱ μ f = f := +sub_add_cancel _ _ + +lemma martingale_part_eq_sum : + martingale_part ℱ μ f = + λ n, f 0 + ∑ i in finset.range n, (f (i+1) - f i - μ[f (i+1) - f i | ℱ i]) := +begin + rw [martingale_part, predictable_part], + ext1 n, + rw [finset.eq_sum_range_sub f n, ← add_sub, ← finset.sum_sub_distrib], +end + +lemma adapted_martingale_part (hf : adapted ℱ f) : + adapted ℱ (martingale_part ℱ μ f) := +adapted.sub hf adapted_predictable_part' + +lemma integrable_martingale_part (hf_int : ∀ n, integrable (f n) μ) (n : ℕ) : + integrable (martingale_part ℱ μ f n) μ := +begin + rw martingale_part_eq_sum, + exact (hf_int 0).add + (integrable_finset_sum' _ (λ i hi, ((hf_int _).sub (hf_int _)).sub integrable_condexp)), +end + +lemma martingale_martingale_part (hf : adapted ℱ f) (hf_int : ∀ n, integrable (f n) μ) + [sigma_finite_filtration μ ℱ] : + martingale (martingale_part ℱ μ f) ℱ μ := +begin + refine ⟨adapted_martingale_part hf, λ i j hij, _⟩, + -- ⊢ μ[martingale_part ℱ μ f j | ℱ i] =ᵐ[μ] martingale_part ℱ μ f i + have h_eq_sum : μ[martingale_part ℱ μ f j | ℱ i] + =ᵐ[μ] f 0 + ∑ k in finset.range j, (μ[f (k+1) - f k | ℱ i] - μ[μ[f (k+1) - f k | ℱ k] | ℱ i]), + { rw martingale_part_eq_sum, + refine (condexp_add (hf_int 0) _).trans _, + { exact integrable_finset_sum' _ + (λ i hij, ((hf_int _).sub (hf_int _)).sub integrable_condexp), }, + refine (eventually_eq.add eventually_eq.rfl (condexp_finset_sum (λ i hij, _))).trans _, + { exact ((hf_int _).sub (hf_int _)).sub integrable_condexp, }, + refine eventually_eq.add _ _, + { rw condexp_of_strongly_measurable (ℱ.le _) _ (hf_int 0), + { apply_instance, }, + { exact (hf 0).mono (ℱ.mono (zero_le i)), }, }, + { exact eventually_eq_sum (λ k hkj, condexp_sub ((hf_int _).sub (hf_int _)) + integrable_condexp), }, }, + refine h_eq_sum.trans _, + have h_ge : ∀ k, i ≤ k → μ[f (k + 1) - f k|ℱ i] - μ[μ[f (k + 1) - f k|ℱ k]|ℱ i] =ᵐ[μ] 0, + { intros k hk, + have : μ[μ[f (k + 1) - f k|ℱ k]|ℱ i] =ᵐ[μ] μ[f (k + 1) - f k|ℱ i], + { exact condexp_condexp_of_le (ℱ.mono hk) (ℱ.le k), }, + filter_upwards [this] with x hx, + rw [pi.sub_apply, pi.zero_apply, hx, sub_self], }, + have h_lt : ∀ k, k < i → μ[f (k + 1) - f k|ℱ i] - μ[μ[f (k + 1) - f k|ℱ k]|ℱ i] + =ᵐ[μ] f (k + 1) - f k - μ[f (k + 1) - f k|ℱ k], + { refine λ k hk, eventually_eq.sub _ _, + { rw condexp_of_strongly_measurable, + { exact ((hf (k+1)).mono (ℱ.mono (nat.succ_le_of_lt hk))).sub + ((hf k).mono (ℱ.mono hk.le)), }, + { exact (hf_int _).sub (hf_int _), }, }, + { rw condexp_of_strongly_measurable, + { exact strongly_measurable_condexp.mono (ℱ.mono hk.le), }, + { exact integrable_condexp, }, }, }, + rw martingale_part_eq_sum, + refine eventually_eq.add eventually_eq.rfl _, + rw [← finset.sum_range_add_sum_Ico _ hij, + ← add_zero (∑ i in finset.range i, (f (i + 1) - f i - μ[f (i + 1) - f i | ℱ i]))], + refine (eventually_eq_sum (λ k hk, h_lt k (finset.mem_range.mp hk))).add _, + refine (eventually_eq_sum (λ k hk, h_ge k (finset.mem_Ico.mp hk).1)).trans _, + simp only [finset.sum_const_zero, pi.zero_apply], + refl, +end + +end measure_theory diff --git a/src/probability/stopping.lean b/src/probability/stopping.lean index f279b54015d6c..93263f9bc8810 100644 --- a/src/probability/stopping.lean +++ b/src/probability/stopping.lean @@ -243,6 +243,11 @@ namespace adapted adapted f (u * v) := λ i, (hu i).mul (hv i) +@[protected, to_additive] lemma div [has_div β] [has_continuous_div β] + (hu : adapted f u) (hv : adapted f v) : + adapted f (u / v) := +λ i, (hu i).div (hv i) + @[protected, to_additive] lemma inv [group β] [topological_group β] (hu : adapted f u) : adapted f u⁻¹ := λ i, (hu i).inv From 3077b72c1c3d42be2f8e34b1ace7d49074f61222 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 15 Sep 2022 15:17:12 +0000 Subject: [PATCH 276/828] feat(analysis/normed/group/quotient): transfer norm structures on quotients of groups to quotients of modules by submodules and of rings by ideals (#16446) This takes the existing norm structures on quotients of additive groups and transfers it along the definitional equality to quotients of modules by submodules and quotients of rings by ideals. In addition, this puts the extra norm structures on these objects where appropriate including `complete_space`, `normed_space`, `semi_normed_comm_ring`, `normed_comm_ring` and `normed_algebra`. - [x] depends on: #16368 --- src/analysis/normed/group/quotient.lean | 107 ++++++++++++++++++++++++ 1 file changed, 107 insertions(+) diff --git a/src/analysis/normed/group/quotient.lean b/src/analysis/normed/group/quotient.lean index d93fc7cb894a7..2d01162f39f48 100644 --- a/src/analysis/normed/group/quotient.lean +++ b/src/analysis/normed/group/quotient.lean @@ -20,6 +20,13 @@ to a normed group hom defined on `M ⧸ S`. This file also introduces a predicate `is_quotient` characterizing normed group homs that are isomorphic to the canonical projection onto a normed group quotient. +In addition, this file also provides normed structures for quotients of modules by submodules, and +of (commutative) rings by ideals. The `seminormed_add_comm_group` and `normed_add_comm_group` +instances described above are transferred directly, but we also define instances of `normed_space`, +`semi_normed_comm_ring`, `normed_comm_ring` and `normed_algebra` under appropriate type class +assumptions on the original space. Moreover, while `quotient_add_group.complete_space` works +out-of-the-box for quotients of `normed_add_comm_group`s by `add_subgroup`s, we need to transfer +this instance in `submodule.quotient.complete_space` so that it applies to these other quotients. ## Main definitions @@ -528,3 +535,103 @@ begin end end normed_add_group_hom + +/-! +### Submodules and ideals + +In what follows, the norm structures created above for quotients of (semi)`normed_add_comm_group`s +by `add_subgroup`s are transferred via definitional equality to quotients of modules by submodules, +and of rings by ideals, thereby preserving the definitional equality for the topological group and +uniform structures worked for above. Completeness is also transferred via this definitional +equality. + +In addition, instances are constructed for `normed_space`, `semi_normed_comm_ring`, +`normed_comm_ring` and `normed_algebra` under the appropriate hypotheses. Currently, we do not +have quotients of rings by two-sided ideals, hence the commutativity hypotheses are required. + -/ + +section submodule + +variables {R : Type*} [ring R] [module R M] (S : submodule R M) + +instance submodule.quotient.seminormed_add_comm_group : + seminormed_add_comm_group (M ⧸ S) := +add_subgroup.seminormed_add_comm_group_quotient S.to_add_subgroup + +instance submodule.quotient.normed_add_comm_group [hS : is_closed (S : set M)] : + normed_add_comm_group (M ⧸ S) := +@add_subgroup.normed_add_comm_group_quotient _ _ S.to_add_subgroup hS + +instance submodule.quotient.complete_space [complete_space M] : complete_space (M ⧸ S) := +quotient_add_group.complete_space M S.to_add_subgroup + +/-- For any `x : M ⧸ S` and any `0 < ε`, there is `m : M` such that `submodule.quotient.mk m = x` +and `∥m∥ < ∥x∥ + ε`. -/ +lemma submodule.quotient.norm_mk_lt {S : submodule R M} (x : M ⧸ S) {ε : ℝ} (hε : 0 < ε) : + ∃ m : M, submodule.quotient.mk m = x ∧ ∥m∥ < ∥x∥ + ε := +norm_mk_lt x hε + +lemma submodule.quotient.norm_mk_le (m : M) : + ∥(submodule.quotient.mk m : M ⧸ S)∥ ≤ ∥m∥ := +quotient_norm_mk_le S.to_add_subgroup m + +instance submodule.quotient.normed_space (𝕜 : Type*) [normed_field 𝕜] [normed_space 𝕜 M] + [has_smul 𝕜 R] [is_scalar_tower 𝕜 R M] : normed_space 𝕜 (M ⧸ S) := +{ norm_smul_le := λ k x, le_of_forall_pos_le_add $ λ ε hε, + begin + have := (nhds_basis_ball.tendsto_iff nhds_basis_ball).mp + ((@real.uniform_continuous_const_mul (∥k∥)).continuous.tendsto (∥x∥)) ε hε, + simp only [mem_ball, exists_prop, dist, abs_sub_lt_iff] at this, + rcases this with ⟨δ, hδ, h⟩, + obtain ⟨a, rfl, ha⟩ := submodule.quotient.norm_mk_lt x hδ, + specialize h (∥a∥) (⟨by linarith, by linarith [submodule.quotient.norm_mk_le S a]⟩), + calc _ ≤ ∥k∥ * ∥a∥ : (quotient_norm_mk_le S.to_add_subgroup (k • a)).trans_eq (norm_smul k a) + ... ≤ _ : (sub_lt_iff_lt_add'.mp h.1).le + end, + .. submodule.quotient.module' S, } + +end submodule + +section ideal + +variables {R : Type*} [semi_normed_comm_ring R] (I : ideal R) + +lemma ideal.quotient.norm_mk_lt {I : ideal R} (x : R ⧸ I) {ε : ℝ} (hε : 0 < ε) : + ∃ r : R, ideal.quotient.mk I r = x ∧ ∥r∥ < ∥x∥ + ε := +norm_mk_lt x hε + +lemma ideal.quotient.norm_mk_le (r : R) : + ∥ideal.quotient.mk I r∥ ≤ ∥r∥ := +quotient_norm_mk_le I.to_add_subgroup r + +instance ideal.quotient.semi_normed_comm_ring : semi_normed_comm_ring (R ⧸ I) := +{ mul_comm := mul_comm, + norm_mul := λ x y, le_of_forall_pos_le_add $ λ ε hε, + begin + have := ((nhds_basis_ball.prod_nhds nhds_basis_ball).tendsto_iff nhds_basis_ball).mp + (real.continuous_mul.tendsto (∥x∥, ∥y∥)) ε hε, + simp only [set.mem_prod, mem_ball, and_imp, prod.forall, exists_prop, prod.exists] at this, + rcases this with ⟨ε₁, ε₂, ⟨h₁, h₂⟩, h⟩, + obtain ⟨⟨a, rfl, ha⟩, ⟨b, rfl, hb⟩⟩ := + ⟨ideal.quotient.norm_mk_lt x h₁, ideal.quotient.norm_mk_lt y h₂⟩, + simp only [dist, abs_sub_lt_iff] at h, + specialize h (∥a∥) (∥b∥) (⟨by linarith, by linarith [ideal.quotient.norm_mk_le I a]⟩) + (⟨by linarith, by linarith [ideal.quotient.norm_mk_le I b]⟩), + calc _ ≤ ∥a∥ * ∥b∥ : (ideal.quotient.norm_mk_le I (a * b)).trans (norm_mul_le a b) + ... ≤ _ : (sub_lt_iff_lt_add'.mp h.1).le + end, + .. submodule.quotient.seminormed_add_comm_group I } + +instance ideal.quotient.normed_comm_ring [is_closed (I : set R)] : + normed_comm_ring (R ⧸ I) := +{ .. ideal.quotient.semi_normed_comm_ring I, + .. submodule.quotient.normed_add_comm_group I } + +variables (𝕜 : Type*) [normed_field 𝕜] + +instance ideal.quotient.normed_algebra [normed_algebra 𝕜 R] : + normed_algebra 𝕜 (R ⧸ I) := +{ .. submodule.quotient.normed_space I 𝕜, + .. ideal.quotient.algebra 𝕜 } + +end ideal From 97befb833b5fd4e01328c14faef052f6d89d31a7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Thu, 15 Sep 2022 21:54:34 +0000 Subject: [PATCH 277/828] feat(analysis/normed/ring/seminorm): add ring_seminorm, ring_norm (#14115) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We define structures `ring_seminorm` and `ring_norm`. These definitions are useful when one needs to consider multiple (semi)norms on a given ring. Co-authored-by: Yaël Dillies --- docs/references.bib | 9 ++ src/analysis/normed/group/basic.lean | 10 ++ src/analysis/normed/group/seminorm.lean | 4 + src/analysis/normed/ring/seminorm.lean | 187 ++++++++++++++++++++++++ 4 files changed, 210 insertions(+) create mode 100644 src/analysis/normed/ring/seminorm.lean diff --git a/docs/references.bib b/docs/references.bib index 03eb047dc6dbc..601d8fff59768 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -240,6 +240,15 @@ @Book{ borceux-vol3 publisher = {Cambridge University Press} } +@Book{ bosch-guntzer-remmert, + title = {Non-Archimedean Analysis : A Systematic Approach to Rigid Analytic Geometry}, + author = {S. Bosch and U. G{\"{u}}ntzer and R. Remmert}, + series = {Grundlehren der mathematischen Wissenschaften}, + volume = {261}, + year = {1984}, + publisher = {Springer-Verlag Berlin } +} + @Book{ bourbaki1966, author = {Bourbaki, Nicolas}, title = {Elements of mathematics. {G}eneral topology. {P}art 1}, diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 069242539e902..67af706a829f8 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -1200,6 +1200,16 @@ by rw [← nnreal.coe_eq_zero, coe_nnnorm, norm_eq_zero] lemma nnnorm_ne_zero_iff {g : E} : ∥g∥₊ ≠ 0 ↔ g ≠ 0 := not_congr nnnorm_eq_zero +variables (E) + +/-- The norm of a normed group as an additive group norm. -/ +def norm_add_group_norm : add_group_norm E := ⟨norm, norm_zero, norm_add_le, norm_neg, + λ {g : E}, norm_eq_zero.mp⟩ + +@[simp] lemma coe_norm_add_group_norm : ⇑(norm_add_group_norm E) = norm := rfl + +variables {E} + /-- An injective group homomorphism from an `add_comm_group` to a `normed_add_comm_group` induces a `normed_add_comm_group` structure on the domain. diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index 792d03dc20ec5..abf47e5688a05 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -169,6 +169,8 @@ variables [group E] [group F] [group G] {p q : group_seminorm E} `fun_like.has_coe_to_fun`. "] instance : has_coe_to_fun (group_seminorm E) (λ _, E → ℝ) := ⟨to_fun⟩ +@[simp, to_additive] lemma to_fun_eq_coe : p.to_fun = p := rfl + @[ext, to_additive] lemma ext : (∀ x, p x = q x) → p = q := fun_like.ext p q @[to_additive] instance : partial_order (group_seminorm E) := @@ -408,6 +410,8 @@ directly. -/ `fun_like.has_coe_to_fun` directly. "] instance : has_coe_to_fun (group_norm E) (λ _, E → ℝ) := fun_like.has_coe_to_fun +@[simp, to_additive] lemma to_fun_eq_coe : p.to_fun = p := rfl + @[ext, to_additive] lemma ext : (∀ x, p x = q x) → p = q := fun_like.ext p q @[to_additive] instance : partial_order (group_norm E) := diff --git a/src/analysis/normed/ring/seminorm.lean b/src/analysis/normed/ring/seminorm.lean new file mode 100644 index 0000000000000..dfc994e58364d --- /dev/null +++ b/src/analysis/normed/ring/seminorm.lean @@ -0,0 +1,187 @@ +/- +Copyright (c) 2022 María Inés de Frutos-Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: María Inés de Frutos-Fernández +-/ +import analysis.seminorm + +/-! +# Seminorms and norms on rings + +This file defines seminorms and norms on rings. These definitions are useful when one needs to +consider multiple (semi)norms on a given ring. + +## Main declarations + +For a ring `R`: +* `ring_seminorm`: A seminorm on a ring `R` is a function `f : R → ℝ` that preserves zero, takes + nonnegative values, is subadditive and submultiplicative and such that `f (-x) = f x` for all + `x ∈ R`. +* `ring_norm`: A seminorm `f` is a norm if `f x = 0` if and only if `x = 0`. + +## References + +* [S. Bosch, U. Güntzer, R. Remmert, *Non-Archimedean Analysis*][bosch-guntzer-remmert] + +## Tags +ring_seminorm, ring_norm +-/ + +set_option old_structure_cmd true + +open_locale nnreal + +variables {R S : Type*} (x y : R) (r : ℝ) + +/-- A seminorm on a ring `R` is a function `f : R → ℝ` that preserves zero, takes nonnegative + values, is subadditive and submultiplicative and such that `f (-x) = f x` for all `x ∈ R`. -/ +structure ring_seminorm (R : Type*) [non_unital_ring R] + extends add_group_seminorm R := +(mul_le' : ∀ x y : R, to_fun (x * y) ≤ to_fun x * to_fun y) + +/-- A function `f : R → ℝ` is a norm on a (nonunital) ring if it is a seminorm and `f x = 0` + implies `x = 0`. -/ +structure ring_norm (R : Type*) [non_unital_ring R] extends add_group_norm R, ring_seminorm R + +attribute [nolint doc_blame] ring_seminorm.to_add_group_seminorm ring_norm.to_add_group_norm + ring_norm.to_ring_seminorm + +/-- `ring_seminorm_class F α` states that `F` is a type of seminorms on the ring `α`. + +You should extend this class when you extend `ring_seminorm`. -/ +class ring_seminorm_class (F : Type*) (α : out_param $ Type*) [non_unital_ring α] + extends add_group_seminorm_class F α, submultiplicative_hom_class F α ℝ + +/-- `ring_norm_class F α` states that `F` is a type of norms on the ring `α`. + +You should extend this class when you extend `ring_norm`. -/ +class ring_norm_class (F : Type*) (α : out_param $ Type*) [non_unital_ring α] + extends ring_seminorm_class F α, add_group_norm_class F α + +namespace ring_seminorm + +section non_unital_ring + +variables [non_unital_ring R] + +instance ring_seminorm_class : ring_seminorm_class (ring_seminorm R) R := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_zero := λ f, f.map_zero', + map_add_le_add := λ f, f.add_le', + map_mul_le_mul := λ f, f.mul_le', + map_neg_eq_map := λ f, f.neg' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ +instance : has_coe_to_fun (ring_seminorm R) (λ _, R → ℝ) := fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe (p : ring_seminorm R) : p.to_fun = p := rfl + +@[ext] lemma ext {p q : ring_seminorm R} : (∀ x, p x = q x) → p = q := fun_like.ext p q + +instance : has_zero (ring_seminorm R) := +⟨{ mul_le' := λ _ _, (zero_mul _).ge, + ..add_group_seminorm.has_zero.zero }⟩ + +lemma eq_zero_iff {p : ring_seminorm R} : p = 0 ↔ ∀ x, p x = 0 := fun_like.ext_iff +lemma ne_zero_iff {p : ring_seminorm R} : p ≠ 0 ↔ ∃ x, p x ≠ 0 := by simp [eq_zero_iff] + +instance : inhabited (ring_seminorm R) := ⟨0⟩ + +/-- The trivial seminorm on a ring `R` is the `ring_seminorm` taking value `0` at `0` and `1` at +every other element. -/ +instance [decidable_eq R] : has_one (ring_seminorm R) := +⟨{ mul_le' := λ x y, begin + by_cases h : x * y = 0, + { refine (if_pos h).trans_le (mul_nonneg _ _); + { change _ ≤ ite _ _ _, + split_ifs, + exacts [le_rfl, zero_le_one] } }, + { change ite _ _ _ ≤ ite _ _ _ * ite _ _ _, + simp only [if_false, h, left_ne_zero_of_mul h, right_ne_zero_of_mul h, mul_one] } + end, + ..(1 : add_group_seminorm R) }⟩ + +@[simp] lemma apply_one [decidable_eq R] (x : R) : + (1 : ring_seminorm R) x = if x = 0 then 0 else 1 := rfl + +end non_unital_ring + +section ring + +variables [ring R] (p : ring_seminorm R) + +lemma seminorm_one_eq_one_iff_ne_zero (hp : p 1 ≤ 1) : + p 1 = 1 ↔ p ≠ 0 := +begin + refine ⟨λ h, ne_zero_iff.mpr ⟨1, by {rw h, exact one_ne_zero}⟩, λ h, _⟩, + obtain hp0 | hp0 := (map_nonneg p (1 : R)).eq_or_gt, + { cases h (ext $ λ x, (map_nonneg _ _).antisymm' _), + simpa only [hp0, mul_one, mul_zero] using map_mul_le_mul p x 1}, + { refine hp.antisymm ((le_mul_iff_one_le_left hp0).1 _), + simpa only [one_mul] using map_mul_le_mul p (1 : R) _ } +end + +end ring + +end ring_seminorm + +/-- The norm of a `non_unital_semi_normed_ring` as a `ring_seminorm`. -/ +def norm_ring_seminorm (R : Type*) [non_unital_semi_normed_ring R] : + ring_seminorm R := +{ to_fun := norm, + mul_le' := norm_mul_le, + ..(norm_add_group_seminorm R) } + +namespace ring_norm + +variable [non_unital_ring R] + +instance ring_norm_class : ring_norm_class (ring_norm R) R := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_zero := λ f, f.map_zero', + map_add_le_add := λ f, f.add_le', + map_mul_le_mul := λ f, f.mul_le', + map_neg_eq_map := λ f, f.neg', + eq_zero_of_map_eq_zero := λ f, f.eq_zero_of_map_eq_zero' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ +instance : has_coe_to_fun (ring_norm R) (λ _, R → ℝ) := ⟨λ p, p.to_fun⟩ + +@[simp] lemma to_fun_eq_coe (p : ring_norm R) : p.to_fun = p := rfl + +@[ext] lemma ext {p q : ring_norm R} : (∀ x, p x = q x) → p = q := fun_like.ext p q + +variable (R) + +/-- The trivial norm on a ring `R` is the `ring_norm` taking value `0` at `0` and `1` at every + other element. -/ +instance [decidable_eq R] : has_one (ring_norm R) := +⟨{ ..(1 : ring_seminorm R), ..(1 : add_group_norm R) }⟩ + +@[simp] lemma apply_one [decidable_eq R] (x : R) : (1 : ring_norm R) x = if x = 0 then 0 else 1 := +rfl + +instance [decidable_eq R] : inhabited (ring_norm R) := ⟨1⟩ + +end ring_norm + +/-- A nonzero ring seminorm on a field `K` is a ring norm. -/ +def ring_seminorm.to_ring_norm {K : Type*} [field K] (f : ring_seminorm K) (hnt : f ≠ 0) : + ring_norm K := +{ eq_zero_of_map_eq_zero' := λ x hx, + begin + obtain ⟨c, hc⟩ := ring_seminorm.ne_zero_iff.mp hnt, + by_contradiction hn0, + have hc0 : f c = 0, + { rw [← mul_one c, ← mul_inv_cancel hn0, ← mul_assoc, mul_comm c, mul_assoc], + exact le_antisymm (le_trans (map_mul_le_mul f _ _) + (by rw [← ring_seminorm.to_fun_eq_coe, hx, zero_mul])) (map_nonneg f _) }, + exact hc hc0, + end, + ..f } + +/-- The norm of a normed_ring as a ring_norm. -/ +@[simps] def norm_ring_norm (R : Type*) [non_unital_normed_ring R] : ring_norm R := +{ ..norm_add_group_norm R, ..norm_ring_seminorm R } From f1397dc14b8f1ae04900b3beb0c4f2aa8ee18f40 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 16 Sep 2022 05:29:21 +0000 Subject: [PATCH 278/828] =?UTF-8?q?feat(analysis/normed=5Fspace/star/gelfa?= =?UTF-8?q?nd=5Fduality):=20Show=20the=20Gelfand=20transform=20is=20a=20bi?= =?UTF-8?q?jective=20isometry=20for=20C=E2=8B=86-algebras=20over=20?= =?UTF-8?q?=E2=84=82=20(#16488)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - [x] depends on: #16451 - [x] depends on: #16438 - [x] depends on: #16368 - [x] depends on: #16303 - [x] depends on: #16446 - [x] depends on: #16448 --- src/analysis/normed_space/spectrum.lean | 25 +++ .../normed_space/star/gelfand_duality.lean | 169 ++++++++++++++++++ src/ring_theory/ideal/basic.lean | 5 + src/topology/continuous_function/bounded.lean | 4 + src/topology/continuous_function/compact.lean | 3 + 5 files changed, 206 insertions(+) create mode 100644 src/analysis/normed_space/star/gelfand_duality.lean diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 4dfe44551d13f..1529e7b246779 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -8,6 +8,7 @@ import analysis.special_functions.pow import analysis.special_functions.exponential import analysis.complex.liouville import analysis.analytic.radius_liminf +import topology.algebra.module.character_space /-! # The spectrum of elements in a complete normed algebra @@ -453,3 +454,27 @@ continuous_linear_map.op_norm_eq_of_bounds zero_le_one end nontrivially_normed_field end alg_hom + +namespace weak_dual + +namespace character_space + +variables [normed_field 𝕜] [normed_ring A] [complete_space A] [norm_one_class A] +variables [normed_algebra 𝕜 A] + +/-- The equivalence between characters and algebra homomorphisms into the base field. -/ +def equiv_alg_hom : (character_space 𝕜 A) ≃ (A →ₐ[𝕜] 𝕜) := +{ to_fun := to_alg_hom, + inv_fun := λ f, + { val := f.to_continuous_linear_map, + property := by { rw eq_set_map_one_map_mul, exact ⟨map_one f, map_mul f⟩ } }, + left_inv := λ f, subtype.ext $ continuous_linear_map.ext $ λ x, rfl, + right_inv := λ f, alg_hom.ext $ λ x, rfl } + +@[simp] lemma equiv_alg_hom_coe (f : character_space 𝕜 A) : ⇑(equiv_alg_hom f) = f := rfl + +@[simp] lemma equiv_alg_hom_symm_coe (f : A →ₐ[𝕜] 𝕜) : ⇑(equiv_alg_hom.symm f) = f := rfl + +end character_space + +end weak_dual diff --git a/src/analysis/normed_space/star/gelfand_duality.lean b/src/analysis/normed_space/star/gelfand_duality.lean new file mode 100644 index 0000000000000..9c79754e4b49d --- /dev/null +++ b/src/analysis/normed_space/star/gelfand_duality.lean @@ -0,0 +1,169 @@ +/- +Copyright (c) 2022 Jireh Loreaux. All rights reserved. +Reeased under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import analysis.normed_space.star.spectrum +import analysis.normed.group.quotient +import analysis.normed_space.algebra +import topology.continuous_function.units +import topology.continuous_function.compact +import topology.algebra.algebra +import topology.continuous_function.stone_weierstrass + +/-! +# Gelfand Duality + +The `gelfand_transform` is an algebra homomorphism from a topological `𝕜`-algebra `A` to +`C(character_space 𝕜 A, 𝕜)`. In the case where `A` is a commutative complex Banach algebra, then +the Gelfand transform is actually spectrum-preserving (`spectrum.gelfand_transform_eq`). Moreover, +when `A` is a commutative C⋆-algebra over `ℂ`, then the Gelfand transform is a surjective isometry, +and even an equivalence between C⋆-algebras. + +## Main definitions + +* `ideal.to_character_space` : constructs an element of the character space from a maximal ideal in + a commutative complex Banach algebra + +## Main statements + +* `spectrum.gelfand_transform_eq` : the Gelfand transform is spectrum-preserving when the algebra is + a commutative complex Banach algebra. +* `gelfand_transform_isometry` : the Gelfand transform is an isometry when the algebra is a + commutative (unital) C⋆-algebra over `ℂ`. +* `gelfand_transform_bijective` : the Gelfand transform is bijective when the algebra is a + commutative (unital) C⋆-algebra over `ℂ`. + +## TODO + +* After `star_alg_equiv` is defined, realize `gelfand_transform` as a `star_alg_equiv`. +* Prove that if `A` is the unital C⋆-algebra over `ℂ` generated by a fixed normal element `x` in + a larger C⋆-algebra `B`, then `character_space ℂ A` is homeomorphic to `spectrum ℂ x`. +* From the previous result, construct the **continuous functional calculus**. +* Show that if `X` is a compact Hausdorff space, then `X` is (canonically) homeomorphic to + `character_space ℂ C(X, ℂ)`. +* Conclude using the previous fact that the functors `C(⬝, ℂ)` and `character_space ℂ ⬝` along with + the canonical homeomorphisms described above constitute a natural contravariant equivalence of + the categories of compact Hausdorff spaces (with continuous maps) and commutative unital + C⋆-algebras (with unital ⋆-algebra homomoprhisms); this is known as **Gelfand duality**. + +## Tags + +Gelfand transform, character space, C⋆-algebra +-/ + +open weak_dual +open_locale nnreal + +section complex_banach_algebra +open ideal + +variables {A : Type*} [normed_comm_ring A] [normed_algebra ℂ A] [complete_space A] + [norm_one_class A] (I : ideal A) [ideal.is_maximal I] + +/-- Every maximal ideal in a commutative complex Banach algebra gives rise to a character on that +algebra. In particular, the character, which may be identified as an algebra homomorphism due to +`weak_dual.character_space.equiv_alg_hom`, is given by the composition of the quotient map and +the Gelfand-Mazur isomorphism `normed_ring.alg_equiv_complex_of_complete`. -/ +noncomputable def ideal.to_character_space : character_space ℂ A := +character_space.equiv_alg_hom.symm $ ((@normed_ring.alg_equiv_complex_of_complete (A ⧸ I) _ _ + (by { letI := quotient.field I, exact @is_unit_iff_ne_zero (A ⧸ I) _ }) _).symm : + A ⧸ I →ₐ[ℂ] ℂ).comp + (quotient.mkₐ ℂ I) + +lemma ideal.to_character_space_apply_eq_zero_of_mem {a : A} (ha : a ∈ I) : + I.to_character_space a = 0 := +begin + unfold ideal.to_character_space, + simpa only [character_space.equiv_alg_hom_symm_coe, alg_hom.coe_comp, + alg_equiv.coe_alg_hom, quotient.mkₐ_eq_mk, function.comp_app, quotient.eq_zero_iff_mem.mpr ha, + spectrum.zero_eq, normed_ring.alg_equiv_complex_of_complete_symm_apply] + using set.eq_of_mem_singleton (set.singleton_nonempty (0 : ℂ)).some_mem, +end + +/-- If `a : A` is not a unit, then some character takes the value zero at `a`. This is equivlaent +to `gelfand_transform ℂ A a` takes the value zero at some character. -/ +lemma weak_dual.character_space.exists_apply_eq_zero {a : A} (ha : ¬ is_unit a) : + ∃ f : character_space ℂ A, f a = 0 := +begin + unfreezingI { obtain ⟨M, hM, haM⟩ := (span {a}).exists_le_maximal (span_singleton_ne_top ha) }, + exact ⟨M.to_character_space, M.to_character_space_apply_eq_zero_of_mem + (haM (mem_span_singleton.mpr ⟨1, (mul_one a).symm⟩))⟩, +end + +/-- The Gelfand transform is spectrum-preserving. -/ +lemma spectrum.gelfand_transform_eq (a : A) : spectrum ℂ (gelfand_transform ℂ A a) = spectrum ℂ a := +begin + refine set.subset.antisymm (alg_hom.spectrum_apply_subset (gelfand_transform ℂ A) a) (λ z hz, _), + obtain ⟨f, hf⟩ := weak_dual.character_space.exists_apply_eq_zero hz, + simp only [map_sub, sub_eq_zero, alg_hom_class.commutes, algebra.id.map_eq_id, ring_hom.id_apply] + at hf, + exact (continuous_map.spectrum_eq_range (gelfand_transform ℂ A a)).symm ▸ ⟨f, hf.symm⟩, +end + +instance : nonempty (character_space ℂ A) := +begin + haveI := norm_one_class.nontrivial A, + exact ⟨classical.some $ + weak_dual.character_space.exists_apply_eq_zero (zero_mem_nonunits.mpr zero_ne_one)⟩, +end + +end complex_banach_algebra + +section complex_cstar_algebra + +variables (A : Type*) [normed_comm_ring A] [normed_algebra ℂ A] [complete_space A] +variables [star_ring A] [cstar_ring A] [star_module ℂ A] [nontrivial A] + +/-- The Gelfand transform is an isometry when the algebra is a C⋆-algebra over `ℂ`. -/ +lemma gelfand_transform_isometry : isometry (gelfand_transform ℂ A) := +begin + refine add_monoid_hom_class.isometry_of_norm (gelfand_transform ℂ A) (λ a, _), + have gt_map_star : gelfand_transform ℂ A (star a) = star (gelfand_transform ℂ A a), + from continuous_map.ext (λ φ, map_star φ a), + /- By `spectrum.gelfand_transform_eq`, the spectra of `star a * a` and its + `gelfand_transform` coincide. Therefore, so do their spectral radii, and since they are + self-adjoint, so also do their norms. Applying the C⋆-property of the norm and taking square + roots shows that the norm is preserved. -/ + have : spectral_radius ℂ (gelfand_transform ℂ A (star a * a)) = spectral_radius ℂ (star a * a), + { unfold spectral_radius, rw spectrum.gelfand_transform_eq, }, + simp only [map_mul, gt_map_star, (is_self_adjoint.star_mul_self _).spectral_radius_eq_nnnorm, + ennreal.coe_eq_coe, cstar_ring.nnnorm_star_mul_self, ←sq] at this, + simpa only [function.comp_app, nnreal.sqrt_sq] + using congr_arg ((coe : ℝ≥0 → ℝ) ∘ ⇑nnreal.sqrt) this, +end + +/-- The Gelfand transform is bijective when the algebra is a C⋆-algebra over `ℂ`. -/ +lemma gelfand_transform_bijective : function.bijective (gelfand_transform ℂ A) := +begin + refine ⟨(gelfand_transform_isometry A).injective, _⟩, + suffices : (gelfand_transform ℂ A).range = ⊤, + { exact λ x, this.symm ▸ (gelfand_transform ℂ A).mem_range.mp (this.symm ▸ algebra.mem_top) }, + /- Because the `gelfand_transform ℂ A` is an isometry, it has closed range, and so by the + Stone-Weierstrass theorem, it suffices to show that the image of the Gelfand transform separates + points in `C(character_space ℂ A, ℂ)` and is closed under `star`. -/ + have h : (gelfand_transform ℂ A).range.topological_closure = (gelfand_transform ℂ A).range, + from le_antisymm (subalgebra.topological_closure_minimal _ le_rfl + (gelfand_transform_isometry A).closed_embedding.closed_range) + (subalgebra.subalgebra_topological_closure _), + refine h ▸ continuous_map.subalgebra_is_R_or_C_topological_closure_eq_top_of_separates_points + _ (λ _ _, _) (λ f hf, _), + /- Separating points just means that elements of the `character_space` which agree at all points + of `A` are the same functional, which is just extensionality. -/ + { contrapose!, + exact λ h, subtype.ext (continuous_linear_map.ext $ + λ a, h (gelfand_transform ℂ A a) ⟨gelfand_transform ℂ A a, ⟨a, rfl⟩, rfl⟩), }, + /- If `f = gelfand_transform ℂ A a`, then `star f` is also in the range of `gelfand_transform ℂ A` + using the argument `star a`. The key lemma below may be hard to spot; it's `map_star` coming from + `weak_dual.star_hom_class`, which is a nontrivial result. -/ + { obtain ⟨f, ⟨a, rfl⟩, rfl⟩ := subalgebra.mem_map.mp hf, + refine ⟨star a, continuous_map.ext $ λ ψ, _⟩, + simpa only [gelfand_transform_apply_apply, map_star, ring_hom.coe_monoid_hom, + alg_equiv.coe_alg_hom, ring_hom.to_monoid_hom_eq_coe, alg_equiv.to_alg_hom_eq_coe, + ring_hom.to_fun_eq_coe, continuous_map.coe_mk, is_R_or_C.conj_ae_coe, + alg_hom.coe_to_ring_hom, monoid_hom.to_fun_eq_coe, ring_hom.comp_left_continuous_apply, + monoid_hom.comp_left_continuous_apply, continuous_map.comp_apply, + alg_hom.to_ring_hom_eq_coe, alg_hom.comp_left_continuous_apply] }, +end + +end complex_cstar_algebra diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index 15a1f9efd1d43..31cc795eb9cf9 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -127,6 +127,11 @@ lemma span_eq_bot {s : set α} : span s = ⊥ ↔ ∀ x ∈ s, (x:α) = 0 := sub @[simp] lemma span_singleton_eq_bot {x} : span ({x} : set α) = ⊥ ↔ x = 0 := submodule.span_singleton_eq_bot +lemma span_singleton_ne_top {α : Type*} [comm_semiring α] {x : α} (hx : ¬ is_unit x) : + ideal.span ({x} : set α) ≠ ⊤ := +(ideal.ne_top_iff_one _).mpr $ λ h1, let ⟨y, hy⟩ := ideal.mem_span_singleton'.mp h1 in + hx ⟨⟨x, y, mul_comm y x ▸ hy, hy⟩, rfl⟩ + @[simp] lemma span_zero : span (0 : set α) = ⊥ := by rw [←set.singleton_zero, span_singleton_eq_bot] @[simp] lemma span_one : span (1 : set α) = ⊤ := by rw [←set.singleton_one, span_singleton_one] diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index f5328761a9e29..4462c792638b4 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -799,6 +799,10 @@ lemma bdd_above_range_norm_comp : bdd_above $ set.range $ norm ∘ f := lemma norm_eq_supr_norm : ∥f∥ = ⨆ x : α, ∥f x∥ := by simp_rw [norm_def, dist_eq_supr, coe_zero, pi.zero_apply, dist_zero_right] +/-- If `∥(1 : β)∥ = 1`, then `∥(1 : α →ᵇ β)∥ = 1` if `α` is nonempty. -/ +instance [nonempty α] [has_one β] [norm_one_class β] : norm_one_class (α →ᵇ β) := +{ norm_one := by simp only [norm_eq_supr_norm, coe_one, pi.one_apply, norm_one, csupr_const] } + /-- The pointwise opposite of a bounded continuous function is again bounded continuous. -/ instance : has_neg (α →ᵇ β) := ⟨λf, of_normed_add_comm_group (-f) f.continuous.neg ∥f∥ $ λ x, diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index b415647de2ea7..0ec59004b2c0f 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -159,6 +159,9 @@ instance : normed_add_comm_group C(α, E) := rw [← norm_mk_of_compact, ← dist_mk_of_compact, dist_eq_norm, mk_of_compact_sub], dist := dist, norm := norm, .. continuous_map.metric_space _ _, .. continuous_map.add_comm_group } +instance [nonempty α] [has_one E] [norm_one_class E] : norm_one_class C(α, E) := +{ norm_one := by simp only [←norm_mk_of_compact, mk_of_compact_one, norm_one] } + section variables (f : C(α, E)) -- The corresponding lemmas for `bounded_continuous_function` are stated with `{f}`, From 69be6fe368f5617a607ebb6b356a7a71419cede3 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Fri, 16 Sep 2022 07:49:46 +0000 Subject: [PATCH 279/828] fix(tactic/monotonicity/interactive): fix the order of calling `unify_with_instance` (#16480) [Zulip](https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/.60apply_instance.60.20fails.20to.20infer.20instance.3F) --- src/tactic/monotonicity/interactive.lean | 2 +- src/tactic/monotonicity/lemmas.lean | 2 +- test/monotonicity.lean | 18 ++++++++++++++++++ 3 files changed, 20 insertions(+), 2 deletions(-) diff --git a/src/tactic/monotonicity/interactive.lean b/src/tactic/monotonicity/interactive.lean index b7bf3eaf8adec..fd9ab6624ceb2 100644 --- a/src/tactic/monotonicity/interactive.lean +++ b/src/tactic/monotonicity/interactive.lean @@ -99,7 +99,7 @@ return () private meta def match_rule_head (p : expr) : list expr → expr → expr → tactic expr | vs e t := -(unify t p >> mmap' unify_with_instance vs >> instantiate_mvars e) +(unify t p >> mmap' unify_with_instance vs.reverse >> instantiate_mvars e) <|> do (expr.pi _ _ d b) ← return t | failed, v ← mk_meta_var d, diff --git a/src/tactic/monotonicity/lemmas.lean b/src/tactic/monotonicity/lemmas.lean index e803747a64979..396611f573581 100644 --- a/src/tactic/monotonicity/lemmas.lean +++ b/src/tactic/monotonicity/lemmas.lean @@ -82,7 +82,7 @@ attribute [mono] upper_bounds_mono_set lower_bounds_mono_set attribute [mono] add_le_add mul_le_mul neg_le_neg mul_lt_mul_of_pos_left mul_lt_mul_of_pos_right imp_imp_imp le_implies_le_of_le_of_le - sub_le_sub tsub_le_tsub tsub_le_tsub_right abs_le_abs sup_le_sup + tsub_le_tsub abs_le_abs sup_le_sup inf_le_inf attribute [mono left] add_lt_add_of_le_of_lt mul_lt_mul' attribute [mono right] add_lt_add_of_lt_of_le mul_lt_mul diff --git a/test/monotonicity.lean b/test/monotonicity.lean index afde3d080330e..d2bb7cd321199 100644 --- a/test/monotonicity.lean +++ b/test/monotonicity.lean @@ -38,6 +38,14 @@ begin { ac_mono }, end +example (x y z k : ℕ) + (h : 3 ≤ (4 : ℕ)) + (h' : z ≤ y) +: (k + 3 + x) - y ≤ (k + 4 + x) - z := +begin + mono, norm_num +end + example (x y z k : ℤ) (h : 3 ≤ (4 : ℤ)) (h' : z ≤ y) @@ -46,6 +54,16 @@ begin mono, norm_num end +example (x y z a b : ℕ) + (h : a ≤ (b : ℕ)) + (h' : z ≤ y) +: (1 + a + x) - y ≤ (1 + b + x) - z := +begin + transitivity (1 + a + x - z), + { mono, }, + { mono, mono, mono }, +end + example (x y z a b : ℤ) (h : a ≤ (b : ℤ)) (h' : z ≤ y) From a5b382d82275e20e2915ef73f679a47224e93409 Mon Sep 17 00:00:00 2001 From: Matej Penciak Date: Fri, 16 Sep 2022 10:27:28 +0000 Subject: [PATCH 280/828] feat(linear_algebra/symplectic_group): add definition of symplectic group (#15513) This PR defines the symplectic group over arbitrary ring (with characteristic not 2) It also moves the definition of the `symplectic.J` into the new `linear_algebra/symplectic_group` file, and adds a lemma `from_blocks_neg` to `data/matrix/block`. Co-authored-by: Moritz Doll Co-authored-by: Moritz Doll --- src/algebra/lie/classical.lean | 6 +- src/data/matrix/block.lean | 7 + src/linear_algebra/symplectic_group.lean | 209 +++++++++++++++++++++++ 3 files changed, 218 insertions(+), 4 deletions(-) create mode 100644 src/linear_algebra/symplectic_group.lean diff --git a/src/algebra/lie/classical.lean b/src/algebra/lie/classical.lean index 14d909dd9b5cd..62b341b45b446 100644 --- a/src/algebra/lie/classical.lean +++ b/src/algebra/lie/classical.lean @@ -9,6 +9,7 @@ import data.matrix.dmatrix import algebra.lie.abelian import linear_algebra.matrix.trace import algebra.lie.skew_adjoint +import linear_algebra.symplectic_group /-! # Classical Lie algebras @@ -119,13 +120,10 @@ end special_linear namespace symplectic -/-- The matrix defining the canonical skew-symmetric bilinear form. -/ -def J : matrix (l ⊕ l) (l ⊕ l) R := matrix.from_blocks 0 (-1) 1 0 - /-- The symplectic Lie algebra: skew-adjoint matrices with respect to the canonical skew-symmetric bilinear form. -/ def sp [fintype l] : lie_subalgebra R (matrix (l ⊕ l) (l ⊕ l) R) := - skew_adjoint_matrices_lie_subalgebra (J l R) + skew_adjoint_matrices_lie_subalgebra (matrix.J l R) end symplectic diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index 1d2004bf7b3a2..2cf32751d5ad0 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -175,6 +175,13 @@ begin ext i j, rcases i; rcases j; simp [from_blocks], end +lemma from_blocks_neg [has_neg R] + (A : matrix n l R) (B : matrix n m R) (C : matrix o l R) (D : matrix o m R) : + - (from_blocks A B C D) = from_blocks (-A) (-B) (-C) (-D) := +begin + ext i j, cases i; cases j; simp [from_blocks], +end + lemma from_blocks_add [has_add α] (A : matrix n l α) (B : matrix n m α) (C : matrix o l α) (D : matrix o m α) (A' : matrix n l α) (B' : matrix n m α) (C' : matrix o l α) (D' : matrix o m α) : diff --git a/src/linear_algebra/symplectic_group.lean b/src/linear_algebra/symplectic_group.lean new file mode 100644 index 0000000000000..a9fab15586eb1 --- /dev/null +++ b/src/linear_algebra/symplectic_group.lean @@ -0,0 +1,209 @@ +/- +Copyright (c) 2022 Matej Penciak. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Matej Penciak, Moritz Doll, Fabien Clery +-/ + +import data.real.basic +import linear_algebra.matrix.nonsingular_inverse + +/-! +# The Symplectic Group + +This file defines the symplectic group and proves elementary properties. + +## Main Definitions + +`matrix.J`: the canonical `2n × 2n` skew-symmetric matrix +`symplectic_group`: the group of symplectic matrices + +## TODO +* Every symplectic matrix has determinant 1. +* For `n = 1` the symplectic group coincides with the special linear group. +-/ + +open_locale matrix + +variables {l R : Type*} + +namespace matrix + +variables (l) [decidable_eq l] (R) [comm_ring R] + +section J_matrix_lemmas + +/-- The matrix defining the canonical skew-symmetric bilinear form. -/ +def J : matrix (l ⊕ l) (l ⊕ l) R := matrix.from_blocks 0 (-1) 1 0 + +@[simp] lemma J_transpose : (J l R)ᵀ = - (J l R) := +begin + rw [J, from_blocks_transpose, ←neg_one_smul R (from_blocks _ _ _ _), from_blocks_smul, + matrix.transpose_zero, matrix.transpose_one, transpose_neg], + simp [from_blocks], +end + +variables [fintype l] + +lemma J_squared : (J l R) ⬝ (J l R) = -1 := +begin + rw [J, from_blocks_multiply], + simp only [matrix.zero_mul, matrix.neg_mul, zero_add, neg_zero', matrix.one_mul, add_zero], + rw [← neg_zero, ← matrix.from_blocks_neg, ← from_blocks_one], +end + +lemma J_inv : (J l R)⁻¹ = -(J l R) := +begin + refine matrix.inv_eq_right_inv _, + rw [matrix.mul_neg, J_squared], + exact neg_neg 1, +end + +lemma J_det_mul_J_det : (det (J l R)) * (det (J l R)) = 1 := +begin + rw [←det_mul, J_squared], + rw [←one_smul R (-1 : matrix _ _ R)], + rw [smul_neg, ←neg_smul, det_smul], + simp only [fintype.card_sum, det_one, mul_one], + apply even.neg_one_pow, + exact even_add_self _ +end + +lemma is_unit_det_J : is_unit (det (J l R)) := +is_unit_iff_exists_inv.mpr ⟨det (J l R), J_det_mul_J_det _ _⟩ + +end J_matrix_lemmas + +variable [fintype l] + +/-- The group of symplectic matrices over a ring `R`. -/ +def symplectic_group : submonoid (matrix (l ⊕ l) (l ⊕ l) R) := +{ carrier := { A | A ⬝ (J l R) ⬝ Aᵀ = J l R}, + mul_mem' := + begin + intros a b ha hb, + simp only [mul_eq_mul, set.mem_set_of_eq, transpose_mul] at *, + rw [←matrix.mul_assoc, a.mul_assoc, a.mul_assoc, hb], + exact ha, + end, + one_mem' := by simp } + +end matrix + +namespace symplectic_group + +variables {l} {R} [decidable_eq l] [fintype l] [comm_ring R] + +open matrix + +lemma mem_iff {A : matrix (l ⊕ l) (l ⊕ l) R} : + A ∈ symplectic_group l R ↔ A ⬝ (J l R) ⬝ Aᵀ = J l R := +by simp [symplectic_group] + +instance coe_matrix : has_coe (symplectic_group l R) (matrix (l ⊕ l) (l ⊕ l) R) +:= by apply_instance + +section symplectic_J + +variables (l) (R) + +lemma J_mem : (J l R) ∈ symplectic_group l R := +begin + rw [mem_iff, J, from_blocks_multiply, from_blocks_transpose, from_blocks_multiply], + simp, +end + +/-- The canonical skew-symmetric matrix as an element in the symplectic group. -/ +def sym_J : symplectic_group l R := ⟨J l R, J_mem l R⟩ + +variables {l} {R} + +@[simp] lemma coe_J : ↑(sym_J l R) = J l R := rfl + +end symplectic_J + +variables {R} {A : matrix (l ⊕ l) (l ⊕ l) R} + +lemma neg_mem (h : A ∈ symplectic_group l R) : -A ∈ symplectic_group l R := +begin + rw mem_iff at h ⊢, + simp [h], +end + +lemma symplectic_det (hA : A ∈ symplectic_group l R) : is_unit $ det A := +begin + rw is_unit_iff_exists_inv, + use A.det, + refine (is_unit_det_J l R).mul_left_cancel _, + rw [mul_one], + rw mem_iff at hA, + apply_fun det at hA, + simp only [det_mul, det_transpose] at hA, + rw [mul_comm A.det, mul_assoc] at hA, + exact hA, +end + +lemma transpose_mem (hA : A ∈ symplectic_group l R) : + Aᵀ ∈ symplectic_group l R := +begin + rw mem_iff at ⊢ hA, + rw transpose_transpose, + have huA := symplectic_det hA, + have huAT : is_unit (Aᵀ).det := + begin + rw matrix.det_transpose, + exact huA, + end, + calc Aᵀ ⬝ J l R ⬝ A + = - Aᵀ ⬝ (J l R)⁻¹ ⬝ A : by {rw J_inv, simp} + ... = - Aᵀ ⬝ (A ⬝ J l R ⬝ Aᵀ)⁻¹ ⬝ A : by rw hA + ... = - (Aᵀ ⬝ (Aᵀ⁻¹ ⬝ (J l R)⁻¹)) ⬝ A⁻¹ ⬝ A : by simp only [matrix.mul_inv_rev, + matrix.mul_assoc, matrix.neg_mul] + ... = - (J l R)⁻¹ : by rw [mul_nonsing_inv_cancel_left _ _ huAT, + nonsing_inv_mul_cancel_right _ _ huA] + ... = (J l R) : by simp [J_inv] +end + +@[simp] lemma transpose_mem_iff : Aᵀ ∈ symplectic_group l R ↔ A ∈ symplectic_group l R := +⟨λ hA, by simpa using transpose_mem hA , transpose_mem⟩ + +lemma mem_iff' : A ∈ symplectic_group l R ↔ Aᵀ ⬝ (J l R) ⬝ A = J l R := +by rw [←transpose_mem_iff, mem_iff, transpose_transpose] + +instance : has_inv (symplectic_group l R) := +{ inv := λ A, ⟨- (J l R) ⬝ (A : matrix (l ⊕ l) (l ⊕ l) R)ᵀ ⬝ (J l R), + mul_mem (mul_mem (neg_mem $ J_mem _ _) $ transpose_mem A.2) $ J_mem _ _⟩ } + +lemma coe_inv (A : symplectic_group l R) : + (↑(A⁻¹) : matrix _ _ _) = - J l R ⬝ (↑A)ᵀ ⬝ J l R := rfl + +lemma inv_left_mul_aux (hA : A ∈ symplectic_group l R) : + -(J l R ⬝ Aᵀ ⬝ J l R ⬝ A) = 1 := +calc -(J l R ⬝ Aᵀ ⬝ J l R ⬝ A) + = - J l R ⬝ (Aᵀ ⬝ J l R ⬝ A) : by simp only [matrix.mul_assoc, matrix.neg_mul] +... = - J l R ⬝ J l R : by {rw mem_iff' at hA, rw hA} +... = (-1 : R) • (J l R ⬝ J l R) : by simp only [matrix.neg_mul, neg_smul, one_smul] +... = (-1 : R) • -1 : by rw J_squared +... = 1 : by simp only [neg_smul_neg, one_smul] + +lemma coe_inv' (A : symplectic_group l R) : (↑(A⁻¹) : matrix (l ⊕ l) (l ⊕ l) R) = A⁻¹ := +begin + refine (coe_inv A).trans (inv_eq_left_inv _).symm, + simp [inv_left_mul_aux, coe_inv], +end + +lemma inv_eq_symplectic_inv (A : matrix (l ⊕ l) (l ⊕ l) R) (hA : A ∈ symplectic_group l R) : + A⁻¹ = - (J l R) ⬝ Aᵀ ⬝ (J l R) := +inv_eq_left_inv (by simp only [matrix.neg_mul, inv_left_mul_aux hA]) + +instance : group (symplectic_group l R) := +{ mul_left_inv := λ A, + begin + apply subtype.ext, + simp only [submonoid.coe_one, submonoid.coe_mul, matrix.neg_mul, coe_inv], + rw [matrix.mul_eq_mul, matrix.neg_mul], + exact inv_left_mul_aux A.2, + end, + .. symplectic_group.has_inv, + .. submonoid.to_monoid _ } + +end symplectic_group From 104f141e8ea8e4447092df913e191f01ab0c51e3 Mon Sep 17 00:00:00 2001 From: bottine Date: Fri, 16 Sep 2022 10:27:29 +0000 Subject: [PATCH 281/828] feat(combinatorics/simple_graph/prod): add `locally_finite` instance for the box product of `locally_finite` graphs, plus two related lemmas (#16431) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Chris Birkbeck Co-authored-by: mcdoll Co-authored-by: Mario Carneiro Co-authored-by: negiizhao Co-authored-by: Eric Wieser Co-authored-by: Kevin Wilson Co-authored-by: Yaël Dillies Co-authored-by: David Kurniadi Angdinata Co-authored-by: jakelev Co-authored-by: Yury G. Kudryashov Co-authored-by: Michael Stoll Co-authored-by: Anatole Dedecker Co-authored-by: Rémy Degenne Co-authored-by: Jireh Loreaux Co-authored-by: Kyle Miller --- src/combinatorics/simple_graph/prod.lean | 34 ++++++++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/src/combinatorics/simple_graph/prod.lean b/src/combinatorics/simple_graph/prod.lean index 8c1eaa9c9c0f1..2a1dd54310f7a 100644 --- a/src/combinatorics/simple_graph/prod.lean +++ b/src/combinatorics/simple_graph/prod.lean @@ -50,6 +50,14 @@ by rw [box_prod_adj, and_iff_left rfl, or_iff_left (λ h : H.adj b b ∧ _, h.1. @[simp] lemma box_prod_adj_right : (G □ H).adj (a, b₁) (a, b₂) ↔ H.adj b₁ b₂ := by rw [box_prod_adj, and_iff_left rfl, or_iff_right (λ h : G.adj a a ∧ _, h.1.ne rfl)] +lemma box_prod_neighbor_set (x : α × β) : + (G □ H).neighbor_set x = ((G.neighbor_set x.1) ×ˢ {x.2}) ∪ ({x.1} ×ˢ (H.neighbor_set x.2)) := +begin + ext ⟨a',b'⟩, + simp only [mem_neighbor_set, set.mem_union, box_prod_adj, set.mem_prod, set.mem_singleton_iff], + simp only [eq_comm, and_comm], +end + variables (G H I) /-- The box product is commutative up to isomorphism. `equiv.prod_comm` as a graph isomorphism. -/ @@ -165,4 +173,30 @@ by { haveI := (nonempty_prod.1 h.nonempty).1, haveI := (nonempty_prod.1 h.nonemp @[simp] lemma box_prod_connected : (G □ H).connected ↔ G.connected ∧ H.connected := ⟨λ h, ⟨h.of_box_prod_left, h.of_box_prod_right⟩, λ h, h.1.box_prod h.2⟩ +instance [decidable_eq α] [decidable_eq β] (x : α × β) + [fintype (G.neighbor_set x.1)] [fintype (H.neighbor_set x.2)] : + fintype ((G □ H).neighbor_set x) := +begin + rw box_prod_neighbor_set, + apply_instance, +end + +lemma box_prod_degree (x : α × β) + [fintype (G.neighbor_set x.1)] [fintype (H.neighbor_set x.2)] + [fintype ((G □ H).neighbor_set x)] : + (G □ H).degree x = G.degree x.1 + H.degree x.2 := +begin + classical, + simp_rw [← card_neighbor_set_eq_degree, box_prod_neighbor_set, + ← set.to_finset_card, set.to_finset_union], + convert finset.card_disjoint_union _; + simp only [set.to_finset_prod, finset.card_product, set.to_finset_card, + set.card_singleton, mul_one, one_mul], + { rintro ⟨_,_⟩ q, + simp only [finset.inf_eq_inter, finset.mem_inter, finset.mem_product, set.mem_to_finset, + mem_neighbor_set, set.mem_singleton_iff] at q, + obtain ⟨⟨q, rfl⟩, ⟨rfl, _⟩⟩ := q, + exact (q.ne rfl).elim, }, +end + end simple_graph From fc35e1e9e523f853ddb45971382ddcf31a1f6d48 Mon Sep 17 00:00:00 2001 From: Moritz Firsching Date: Fri, 16 Sep 2022 12:10:57 +0000 Subject: [PATCH 282/828] feat(archive/100-theorems-list): Stirling, part 2 (#14875) feat(archive/100-theorems-list): Stirling Part 2 - [x] depends on: #14874 [part 1] Co-authored-by: fabian-kruse Co-authored-by: nick-kuhn <46423253+nick-kuhn@users.noreply.github.com> Co-authored-by: Moritz Firsching Co-authored-by: nick-kuhn <46423253+nick-kuhn@users.noreply.github.com> --- docs/100.yaml | 3 + src/analysis/special_functions/stirling.lean | 104 ++++++++++++++++++- 2 files changed, 102 insertions(+), 5 deletions(-) diff --git a/docs/100.yaml b/docs/100.yaml index 0f742a3fdc935..50be57d8c9453 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -368,6 +368,9 @@ author : Chris Hughes 90: title : Stirling’s Formula + decls : + - stirling.tendsto_stirling_seq_sqrt_pi + author : mathlib (Moritz Firsching, Fabian Kruse, Nikolas Kuhn, Heather Macbeth) 91: title : The Triangle Inequality decl : norm_add_le diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index 5170c5e8af251..fa59c375ca816 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -6,29 +6,36 @@ Authors: Moritz Firsching, Fabian Kruse, Nikolas Kuhn import analysis.p_series import analysis.special_functions.log.deriv import tactic.positivity +import data.real.pi.wallis /-! # Stirling's formula This file proves Stirling's formula for the factorial. It states that $n!$ grows asymptotically like $\sqrt{2\pi n}(\frac{n}{e})^n$. -TODO: Add Part 2 to complete the proof ## Proof outline The proof follows: . +We proceed in two parts. + ### Part 1 We consider the fraction sequence $a_n$ of fractions $\frac{n!}{\sqrt{2n}(\frac{n}{e})^n}$ and prove that this sequence converges against a real, positive number $a$. For this the two main ingredients are - taking the logarithm of the sequence and - use the series expansion of $\log(1 + x)$. + +### Part 2 +We use the fact that the series defined in part 1 converges againt a real number $a$ and prove that +$a = \sqrt{\pi}$. Here the main ingredient is the convergence of the Wallis product. -/ -open_locale topological_space big_operators +open_locale topological_space real big_operators nat open finset filter nat real +namespace stirling /-! ### Part 1 https://proofwiki.org/wiki/Stirling%27s_Formula#Part_1 @@ -39,7 +46,7 @@ Define `stirling_seq n` as $\frac{n!}{\sqrt{2n}/(\frac{n}{e})^n$. Stirling's formula states that this sequence has limit $\sqrt(π)$. -/ noncomputable def stirling_seq (n : ℕ) : ℝ := -n.factorial / (sqrt (2 * n) * (n / exp 1) ^ n) +n! / (sqrt (2 * n) * (n / exp 1) ^ n) @[simp] lemma stirling_seq_zero : stirling_seq 0 = 0 := by rw [stirling_seq, cast_zero, mul_zero, real.sqrt_zero, zero_mul, div_zero] @@ -52,9 +59,9 @@ We have the expression `log (stirling_seq (n + 1)) = log(n + 1)! - 1 / 2 * log(2 * n) - n * log ((n + 1) / e)`. -/ lemma log_stirling_seq_formula (n : ℕ) : log (stirling_seq n.succ) = - log n.succ.factorial - 1 / 2 * log (2 * n.succ) - n.succ * log (n.succ / exp 1) := + log n.succ!- 1 / 2 * log (2 * n.succ) - n.succ * log (n.succ / exp 1) := begin - have h1 : (0 : ℝ) < n.succ.factorial := cast_pos.mpr n.succ.factorial_pos, + have h1 : (0 : ℝ) < n.succ!:= cast_pos.mpr n.succ.factorial_pos, have h2 : (0 : ℝ) < (2 * n.succ) := mul_pos two_pos (cast_pos.mpr (succ_pos n)), have h3 := real.sqrt_pos.mpr h2, have h4 := pow_pos (div_pos (cast_pos.mpr n.succ_pos ) (exp_pos 1)) n.succ, @@ -203,3 +210,90 @@ begin rw ←filter.tendsto_add_at_top_iff_nat 1, exact tendsto_at_top_cinfi stirling_seq'_antitone ⟨x, hx'⟩, end + +/-! + ### Part 2 + https://proofwiki.org/wiki/Stirling%27s_Formula#Part_2 +-/ + +/-- For `n : ℕ`, define `w n` as `2^(4*n) * n!^4 / ((2*n)!^2 * (2*n + 1))` -/ +noncomputable def w (n : ℕ) : ℝ := +(2 ^ (4 * n) * n! ^ 4) / ((2 * n)!^ 2 * (2 * n + 1)) + +/-- The sequence `w n` converges to `π/2` -/ +lemma tendsto_w_at_top: tendsto (λ (n : ℕ), w n) at_top (𝓝 (π/2)) := +begin + convert tendsto_prod_pi_div_two, + funext n, + induction n with n ih, + { rw [w, prod_range_zero, cast_zero, mul_zero, pow_zero, one_mul, mul_zero, factorial_zero, + cast_one, one_pow, one_pow, one_mul, mul_zero, zero_add, div_one] }, + rw [w, prod_range_succ, ←ih, w, _root_.div_mul_div_comm, _root_.div_mul_div_comm], + refine (div_eq_div_iff (ne_of_gt _) (ne_of_gt _)).mpr _, + { exact mul_pos (pow_pos (cast_pos.mpr (2 * n.succ).factorial_pos) 2) + (add_pos (mul_pos two_pos (cast_pos.mpr n.succ_pos)) one_pos) }, + { have h : 0 < 2 * (n : ℝ) + 1 := + add_pos_of_nonneg_of_pos (mul_nonneg zero_le_two n.cast_nonneg) one_pos, + exact mul_pos (mul_pos (pow_pos (cast_pos.mpr (2 * n).factorial_pos) 2) h) + (mul_pos h (add_pos_of_nonneg_of_pos (mul_nonneg zero_le_two n.cast_nonneg) three_pos)) }, + { simp_rw [nat.mul_succ, factorial_succ, succ_eq_add_one, pow_succ], + push_cast, + ring_nf }, +end + +/-- The sequence `n / (2 * n + 1)` tends to `1/2` -/ +lemma tendsto_self_div_two_mul_self_add_one : + tendsto (λ (n : ℕ), (n : ℝ) / (2 * n + 1)) at_top (𝓝 (1 / 2)) := +begin + conv { congr, skip, skip, rw [one_div, ←add_zero (2 : ℝ)] }, + refine (((tendsto_const_div_at_top_nhds_0_nat 1).const_add (2 : ℝ)).inv₀ + ((add_zero (2 : ℝ)).symm ▸ two_ne_zero)).congr' (eventually_at_top.mpr ⟨1, λ n hn, _⟩), + rw [add_div' (1 : ℝ) (2 : ℝ) (n : ℝ) (cast_ne_zero.mpr (one_le_iff_ne_zero.mp hn)), inv_div], +end + + +/-- For any `n ≠ 0`, we have the identity +`(stirling_seq n)^4/(stirling_seq (2*n))^2 * (n / (2 * n + 1)) = w n`. -/ +lemma stirling_seq_pow_four_div_stirling_seq_pow_two_eq (n : ℕ) (hn : n ≠ 0) : + ((stirling_seq n) ^ 4 / (stirling_seq (2 * n)) ^ 2) * (n / (2 * n + 1)) = w n := +begin + rw [bit0_eq_two_mul, stirling_seq, pow_mul, stirling_seq, w], + simp_rw [div_pow, mul_pow], + rw [sq_sqrt (mul_nonneg two_pos.le n.cast_nonneg), + sq_sqrt (mul_nonneg two_pos.le (2 * n).cast_nonneg)], + have : (n : ℝ) ≠ 0, from cast_ne_zero.mpr hn, + have : (exp 1) ≠ 0, from exp_ne_zero 1, + have : ((2 * n)!: ℝ) ≠ 0, from cast_ne_zero.mpr (factorial_ne_zero (2 * n)), + have : 2 * (n : ℝ) + 1 ≠ 0, by {norm_cast, exact succ_ne_zero (2*n)}, + field_simp, + simp only [mul_pow, mul_comm 2 n, mul_comm 4 n, pow_mul], + ring, +end + +/-- +Suppose the sequence `stirling_seq` (defined above) has the limit `a ≠ 0`. +Then the sequence `w` has limit `a^2/2` +-/ +lemma second_wallis_limit (a : ℝ) (hane : a ≠ 0) (ha : tendsto stirling_seq at_top (𝓝 a)) : + tendsto w at_top (𝓝 (a ^ 2 / 2)):= +begin + refine tendsto.congr' (eventually_at_top.mpr ⟨1, λ n hn, + stirling_seq_pow_four_div_stirling_seq_pow_two_eq n (one_le_iff_ne_zero.mp hn)⟩) _, + have h : a ^ 2 / 2 = (a ^ 4 / a ^ 2) * (1 / 2), + { rw [mul_one_div, ←mul_one_div (a ^ 4) (a ^ 2), one_div, ←pow_sub_of_lt a], + norm_num }, + rw h, + exact ((ha.pow 4).div ((ha.comp (tendsto_id.const_mul_at_top' two_pos)).pow 2) + (pow_ne_zero 2 hane)).mul tendsto_self_div_two_mul_self_add_one, +end + +/-- **Stirling's Formula** -/ +theorem tendsto_stirling_seq_sqrt_pi : tendsto (λ (n : ℕ), stirling_seq n) at_top (𝓝 (sqrt π)) := +begin + obtain ⟨a, hapos, halimit⟩ := stirling_seq_has_pos_limit_a, + have hπ : π / 2 = a ^ 2 / 2 := tendsto_nhds_unique tendsto_w_at_top + (second_wallis_limit a (ne_of_gt hapos) halimit), + rwa [(div_left_inj' (show (2 : ℝ) ≠ 0, from two_ne_zero)).mp hπ, sqrt_sq hapos.le], +end + +end stirling From d46774d43797f5d1f507a63a6e904f7a533ae74a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 16 Sep 2022 12:10:58 +0000 Subject: [PATCH 283/828] feat(linear_algebra/clifford_algebra/basic): lemmas about triple products of repeated vectors (#16153) --- .../clifford_algebra/basic.lean | 22 +++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/src/linear_algebra/clifford_algebra/basic.lean b/src/linear_algebra/clifford_algebra/basic.lean index 725d054f47365..f6828941f3658 100644 --- a/src/linear_algebra/clifford_algebra/basic.lean +++ b/src/linear_algebra/clifford_algebra/basic.lean @@ -206,6 +206,16 @@ calc ι Q a * ι Q b + ι Q b * ι Q a by rw [←ring_hom.map_sub, ←ring_hom.map_sub] ... = algebra_map R _ (quadratic_form.polar Q a b) : rfl +lemma ι_mul_comm (a b : M) : + ι Q a * ι Q b = algebra_map R _ (quadratic_form.polar Q a b) - ι Q b * ι Q a := +eq_sub_of_add_eq (ι_mul_ι_add_swap a b) + +/-- $aba$ is a vector. -/ +lemma ι_mul_ι_mul_ι (a b : M) : + ι Q a * ι Q b * ι Q a = ι Q (quadratic_form.polar Q a b • a - Q a • b) := +by rw [ι_mul_comm, sub_mul, mul_assoc, ι_sq_scalar, ←algebra.smul_def, ←algebra.commutes, + ←algebra.smul_def, ←map_smul, ←map_smul, ←map_sub] + @[simp] lemma ι_range_map_lift (f : M →ₗ[R] A) (cond : ∀ m, f m * f m = algebra_map _ _ (Q m)) : (ι Q).range.map (lift Q ⟨f, cond⟩).to_linear_map = f.range := @@ -317,6 +327,18 @@ begin exactI is_unit_of_invertible (ι Q m), end +/-- $aba^{-1}$ is a vector. -/ +lemma ι_mul_ι_mul_inv_of_ι (a b : M) [invertible (ι Q a)] [invertible (Q a)] : + ι Q a * ι Q b * ⅟(ι Q a) = ι Q ((⅟(Q a) * quadratic_form.polar Q a b) • a - b) := +by rw [inv_of_ι, map_smul, mul_smul_comm, ι_mul_ι_mul_ι, ←map_smul, smul_sub, smul_smul, smul_smul, + inv_of_mul_self, one_smul] + +/-- $a^{-1}ba$ is a vector. -/ +lemma inv_of_ι_mul_ι_mul_ι (a b : M) [invertible (ι Q a)] [invertible (Q a)] : + ⅟(ι Q a) * ι Q b * ι Q a = ι Q ((⅟(Q a) * quadratic_form.polar Q a b) • a - b) := +by rw [inv_of_ι, map_smul, smul_mul_assoc, smul_mul_assoc, ι_mul_ι_mul_ι, ←map_smul, smul_sub, + smul_smul, smul_smul, inv_of_mul_self, one_smul] + end clifford_algebra namespace tensor_algebra From 894c8c22e19744ad3f8f6ea0bc07073079dcb985 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Fri, 16 Sep 2022 13:03:05 +0000 Subject: [PATCH 284/828] refactor(number_theory/padics/padic_val): make prime implicit (#15221) Make the variable denoting the prime in theorems related to `padic_val_nat`, `padic_val_int`, `padic_val_rat`, and `padic_norm` implicit, since they seem to be redundant and inferable in all use cases, but the variable in their definitions remain explicit. --- src/number_theory/padics/padic_integers.lean | 79 +++--- src/number_theory/padics/padic_norm.lean | 84 +++--- src/number_theory/padics/padic_numbers.lean | 87 +++--- src/number_theory/padics/padic_val.lean | 250 +++++++++--------- .../polynomial/cyclotomic/eval.lean | 4 +- 5 files changed, 246 insertions(+), 258 deletions(-) diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index 7543b9e3de9c3..3a151d38e601a 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -56,7 +56,9 @@ def padic_int (p : ℕ) [fact p.prime] := {x : ℚ_[p] // ∥x∥ ≤ 1} notation `ℤ_[`p`]` := padic_int p namespace padic_int + /-! ### Ring structure and coercion to `ℚ_[p]` -/ + variables {p : ℕ} [fact p.prime] instance : has_coe ℤ_[p] ℚ_[p] := ⟨subtype.val⟩ @@ -156,8 +158,8 @@ def of_int_seq (seq : ℕ → ℤ) (h : is_cau_seq (padic_norm p) (λ n, seq n)) end padic_int namespace padic_int -/-! -### Instances + +/-! ### Instances We now show that `ℤ_[p]` is a * complete metric space @@ -202,7 +204,9 @@ function.injective.is_domain (subring p).subtype subtype.coe_injective end padic_int namespace padic_int + /-! ### Norm -/ + variables {p : ℕ} [fact p.prime] lemma norm_le_one (z : ℤ_[p]) : ∥z∥ ≤ 1 := z.2 @@ -260,8 +264,8 @@ end padic_int namespace padic_int -variables (p : ℕ) [hp_prime : fact p.prime] -include hp_prime +variables (p : ℕ) [hp : fact p.prime] +include hp lemma exists_pow_neg_lt {ε : ℝ} (hε : 0 < ε) : ∃ (k : ℕ), ↑p ^ -((k : ℕ) : ℤ) < ε := @@ -274,8 +278,8 @@ begin norm_cast, apply le_of_lt, convert nat.lt_pow_self _ _ using 1, - exact hp_prime.1.one_lt }, - { exact_mod_cast hp_prime.1.pos } + exact hp.1.one_lt }, + { exact_mod_cast hp.1.pos } end lemma exists_pow_neg_lt_rat {ε : ℚ} (hε : 0 < ε) : @@ -289,13 +293,13 @@ end variable {p} -lemma norm_int_lt_one_iff_dvd (k : ℤ) : ∥(k : ℤ_[p])∥ < 1 ↔ ↑p ∣ k := +lemma norm_int_lt_one_iff_dvd (k : ℤ) : ∥(k : ℤ_[p])∥ < 1 ↔ (p : ℤ) ∣ k := suffices ∥(k : ℚ_[p])∥ < 1 ↔ ↑p ∣ k, by rwa norm_int_cast_eq_padic_norm, padic_norm_e.norm_int_lt_one_iff_dvd k -lemma norm_int_le_pow_iff_dvd {k : ℤ} {n : ℕ} : ∥(k : ℤ_[p])∥ ≤ ((↑p)^(-n : ℤ)) ↔ ↑p^n ∣ k := -suffices ∥(k : ℚ_[p])∥ ≤ ((↑p)^(-n : ℤ)) ↔ ↑(p^n) ∣ k, by simpa [norm_int_cast_eq_padic_norm], -padic_norm_e.norm_int_le_pow_iff_dvd _ _ +lemma norm_int_le_pow_iff_dvd {k : ℤ} {n : ℕ} : ∥(k : ℤ_[p])∥ ≤ p ^ (-n : ℤ) ↔ (p^n : ℤ) ∣ k := +suffices ∥(k : ℚ_[p])∥ ≤ ((↑p)^(-n : ℤ)) ↔ ↑(p^n) ∣ k, +by simpa [norm_int_cast_eq_padic_norm], padic_norm_e.norm_int_le_pow_iff_dvd _ _ /-! ### Valuation on `ℤ_[p]` -/ @@ -323,31 +327,32 @@ lemma valuation_nonneg (x : ℤ_[p]) : 0 ≤ x.valuation := begin by_cases hx : x = 0, { simp [hx] }, - have h : (1 : ℝ) < p := by exact_mod_cast hp_prime.1.one_lt, + have h : (1 : ℝ) < p := by exact_mod_cast hp.1.one_lt, rw [← neg_nonpos, ← (zpow_strict_mono h).le_iff_le], show (p : ℝ) ^ -valuation x ≤ p ^ 0, rw [← norm_eq_pow_val hx], - simpa using x.property, + simpa using x.property end @[simp] lemma valuation_p_pow_mul (n : ℕ) (c : ℤ_[p]) (hc : c ≠ 0) : (↑p ^ n * c).valuation = n + c.valuation := begin - have : ∥↑p ^ n * c∥ = ∥(p ^ n : ℤ_[p])∥ * ∥c∥, + have : ∥(↑p ^ n * c)∥ = ∥(p ^ n : ℤ_[p])∥ * ∥c∥, { exact norm_mul _ _ }, - have aux : ↑p ^ n * c ≠ 0, + have aux : (↑p ^ n * c) ≠ 0, { contrapose! hc, rw mul_eq_zero at hc, cases hc, - { refine (hp_prime.1.ne_zero _).elim, + { refine (hp.1.ne_zero _).elim, exact_mod_cast (pow_eq_zero hc) }, { exact hc } }, rwa [norm_eq_pow_val aux, norm_p_pow, norm_eq_pow_val hc, ← zpow_add₀, ← neg_add, zpow_inj, neg_inj] at this, - { exact_mod_cast hp_prime.1.pos }, - { exact_mod_cast hp_prime.1.ne_one }, - { exact_mod_cast hp_prime.1.ne_zero }, + { exact_mod_cast hp.1.pos }, + { exact_mod_cast hp.1.ne_one }, + { exact_mod_cast hp.1.ne_zero } end section units + /-! ### Units of `ℤ_[p]` -/ local attribute [reducible] padic_int @@ -360,7 +365,7 @@ lemma mul_inv : ∀ {z : ℤ_[p]}, ∥z∥ = 1 → z * z.inv = 1 rw [norm_eq_padic_norm] at h, rw dif_pos h, apply subtype.ext_iff_val.2, - simp [mul_inv_cancel hk], + simp [mul_inv_cancel hk] end lemma inv_mul {z : ℤ_[p]} (hz : ∥z∥ = 1) : z.inv * z = 1 := @@ -400,7 +405,7 @@ See `unit_coeff_spec`. -/ def unit_coeff {x : ℤ_[p]} (hx : x ≠ 0) : ℤ_[p]ˣ := let u : ℚ_[p] := x*p^(-x.valuation) in have hu : ∥u∥ = 1, -by simp [hx, nat.zpow_ne_zero_of_pos (by exact_mod_cast hp_prime.1.pos) x.valuation, +by simp [hx, nat.zpow_ne_zero_of_pos (by exact_mod_cast hp.1.pos) x.valuation, norm_eq_pow_val, zpow_neg, inv_mul_cancel], mk_units hu @@ -415,14 +420,15 @@ begin have repr : (x : ℚ_[p]) = (unit_coeff hx) * p ^ x.valuation, { rw [unit_coeff_coe, mul_assoc, ← zpow_add₀], { simp }, - { exact_mod_cast hp_prime.1.ne_zero } }, + { exact_mod_cast hp.1.ne_zero } }, convert repr using 2, - rw [← zpow_coe_nat, int.nat_abs_of_nonneg (valuation_nonneg x)], + rw [← zpow_coe_nat, int.nat_abs_of_nonneg (valuation_nonneg x)] end end units section norm_le_iff + /-! ### Various characterizations of open unit balls -/ lemma norm_le_pow_iff_le_valuation (x : ℤ_[p]) (hx : x ≠ 0) (n : ℕ) : @@ -432,11 +438,11 @@ begin lift x.valuation to ℕ using x.valuation_nonneg with k hk, simp only [int.coe_nat_le, zpow_neg, zpow_coe_nat], have aux : ∀ n : ℕ, 0 < (p ^ n : ℝ), - { apply pow_pos, exact_mod_cast hp_prime.1.pos }, + { apply pow_pos, exact_mod_cast hp.1.pos }, rw [inv_le_inv (aux _) (aux _)], - have : p ^ n ≤ p ^ k ↔ n ≤ k := (strict_mono_pow hp_prime.1.one_lt).le_iff_le, + have : p ^ n ≤ p ^ k ↔ n ≤ k := (strict_mono_pow hp.1.one_lt).le_iff_le, rw [← this], - norm_cast, + norm_cast end lemma mem_span_pow_iff_le_valuation (x : ℤ_[p]) (hx : x ≠ 0) (n : ℕ) : @@ -446,14 +452,14 @@ begin split, { rintro ⟨c, rfl⟩, suffices : c ≠ 0, - { rw [valuation_p_pow_mul _ _ this, le_add_iff_nonneg_right], apply valuation_nonneg, }, - contrapose! hx, rw [hx, mul_zero], }, + { rw [valuation_p_pow_mul _ _ this, le_add_iff_nonneg_right], apply valuation_nonneg }, + contrapose! hx, rw [hx, mul_zero] }, { rw [unit_coeff_spec hx] { occs := occurrences.pos [2] }, lift x.valuation to ℕ using x.valuation_nonneg with k hk, simp only [int.nat_abs_of_nat, units.is_unit, is_unit.dvd_mul_left, int.coe_nat_le], intro H, obtain ⟨k, rfl⟩ := nat.exists_eq_add_of_le H, - simp only [pow_add, dvd_mul_right], } + simp only [pow_add, dvd_mul_right] } end lemma norm_le_pow_iff_mem_span_pow (x : ℤ_[p]) (n : ℕ) : @@ -463,7 +469,7 @@ begin { subst hx, simp only [norm_zero, zpow_neg, zpow_coe_nat, inv_nonneg, iff_true, submodule.zero_mem], exact_mod_cast nat.zero_le _ }, - rw [norm_le_pow_iff_le_valuation x hx, mem_span_pow_iff_le_valuation x hx], + rw [norm_le_pow_iff_le_valuation x hx, mem_span_pow_iff_le_valuation x hx] end lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℤ_[p]) (n : ℤ) : @@ -481,7 +487,7 @@ begin have := norm_le_pow_iff_mem_span_pow x 1, rw [ideal.mem_span_singleton, pow_one] at this, rw [← this, norm_le_pow_iff_norm_lt_pow_add_one], - simp only [zpow_zero, int.coe_nat_zero, int.coe_nat_succ, add_left_neg, zero_add], + simp only [zpow_zero, int.coe_nat_zero, int.coe_nat_succ, add_left_neg, zero_add] end @[simp] lemma pow_p_dvd_int_iff (n : ℕ) (a : ℤ) : (p ^ n : ℤ_[p]) ∣ a ↔ ↑p ^ n ∣ a := @@ -490,13 +496,14 @@ by rw [← norm_int_le_pow_iff_dvd, norm_le_pow_iff_mem_span_pow, ideal.mem_span end norm_le_iff section dvr + /-! ### Discrete valuation ring -/ instance : local_ring ℤ_[p] := local_ring.of_nonunits_add $ by simp only [mem_nonunits]; exact λ x y, norm_lt_one_add lemma p_nonnunit : (p : ℤ_[p]) ∈ nonunits ℤ_[p] := -have (p : ℝ)⁻¹ < 1, from inv_lt_one $ by exact_mod_cast hp_prime.1.one_lt, +have (p : ℝ)⁻¹ < 1, from inv_lt_one $ by exact_mod_cast hp.1.one_lt, by simp [this] lemma maximal_ideal_eq_span_p : maximal_ideal ℤ_[p] = ideal.span {p} := @@ -505,7 +512,7 @@ begin { intros x hx, rw ideal.mem_span_singleton, simp only [local_ring.mem_maximal_ideal, mem_nonunits] at hx, - rwa ← norm_lt_one_iff_dvd, }, + rwa ← norm_lt_one_iff_dvd }, { rw [ideal.span_le, set.singleton_subset_iff], exact p_nonnunit } end @@ -513,7 +520,7 @@ lemma prime_p : prime (p : ℤ_[p]) := begin rw [← ideal.span_singleton_prime, ← maximal_ideal_eq_span_p], { apply_instance }, - { exact_mod_cast hp_prime.1.ne_zero } + { exact_mod_cast hp.1.ne_zero } end lemma irreducible_p : irreducible (p : ℤ_[p]) := @@ -539,14 +546,14 @@ instance : is_adic_complete (maximal_ideal ℤ_[p]) ℤ_[p] := { intros ε hε, obtain ⟨m, hm⟩ := exists_pow_neg_lt p hε, refine ⟨m, λ n hn, lt_of_le_of_lt _ hm⟩, rw [← neg_sub, norm_neg], exact hx hn }, { refine ⟨x'.lim, λ n, _⟩, - have : (0:ℝ) < p ^ (-n : ℤ), { apply zpow_pos_of_pos, exact_mod_cast hp_prime.1.pos }, + have : (0:ℝ) < p ^ (-n : ℤ), { apply zpow_pos_of_pos, exact_mod_cast hp.1.pos }, obtain ⟨i, hi⟩ := equiv_def₃ (equiv_lim x') this, by_cases hin : i ≤ n, - { exact (hi i le_rfl n hin).le, }, + { exact (hi i le_rfl n hin).le }, { push_neg at hin, specialize hi i le_rfl i le_rfl, specialize hx hin.le, have := nonarchimedean (x n - x i) (x i - x'.lim), rw [sub_add_sub_cancel] at this, - refine this.trans (max_le_iff.mpr ⟨hx, hi.le⟩), } }, + refine this.trans (max_le_iff.mpr ⟨hx, hi.le⟩) } } end } end dvr diff --git a/src/number_theory/padics/padic_norm.lean b/src/number_theory/padics/padic_norm.lean index 818567dbd0e01..e44df9e92c04e 100644 --- a/src/number_theory/padics/padic_norm.lean +++ b/src/number_theory/padics/padic_norm.lean @@ -47,7 +47,7 @@ if q = 0 then 0 else (↑p : ℚ) ^ (-(padic_val_rat p q)) namespace padic_norm open padic_val_rat -variables (p : ℕ) +variables {p : ℕ} /-- Unfolds the definition of the p-adic norm of `q` when `q ≠ 0`. -/ @[simp] protected lemma eq_zpow_of_nonzero {q : ℚ} (hq : q ≠ 0) : @@ -75,7 +75,7 @@ The p-adic norm of `p` is `1/p` if `p > 1`. See also `padic_norm.padic_norm_p_of_prime` for a version that assumes `p` is prime. -/ -lemma padic_norm_p {p : ℕ} (hp : 1 < p) : padic_norm p p = 1 / p := +lemma padic_norm_p (hp : 1 < p) : padic_norm p p = 1 / p := by simp [padic_norm, (pos_of_gt hp).ne', padic_val_nat.self hp] /-- @@ -83,11 +83,11 @@ The p-adic norm of `p` is `1/p` if `p` is prime. See also `padic_norm.padic_norm_p` for a version that assumes `1 < p`. -/ -@[simp] lemma padic_norm_p_of_prime (p : ℕ) [fact p.prime] : padic_norm p p = 1 / p := +@[simp] lemma padic_norm_p_of_prime [fact p.prime] : padic_norm p p = 1 / p := padic_norm_p $ nat.prime.one_lt (fact.out _) /-- The p-adic norm of `q` is `1` if `q` is prime and not equal to `p`. -/ -lemma padic_norm_of_prime_of_ne {p q : ℕ} [p_prime : fact p.prime] [q_prime : fact q.prime] +lemma padic_norm_of_prime_of_ne {q : ℕ} [p_prime : fact p.prime] [q_prime : fact q.prime] (neq : p ≠ q) : padic_norm p q = 1 := begin have p : padic_val_rat p q = 0, @@ -100,7 +100,7 @@ The p-adic norm of `p` is less than 1 if `1 < p`. See also `padic_norm.padic_norm_p_lt_one_of_prime` for a version assuming `prime p`. -/ -lemma padic_norm_p_lt_one {p : ℕ} (hp : 1 < p) : padic_norm p p < 1 := +lemma padic_norm_p_lt_one (hp : 1 < p) : padic_norm p p < 1 := begin rw [padic_norm_p hp, div_lt_iff, one_mul], { exact_mod_cast hp }, @@ -112,7 +112,7 @@ The p-adic norm of `p` is less than 1 if `p` is prime. See also `padic_norm.padic_norm_p_lt_one` for a version assuming `1 < p`. -/ -lemma padic_norm_p_lt_one_of_prime (p : ℕ) [fact p.prime] : padic_norm p p < 1 := +lemma padic_norm_p_lt_one_of_prime [fact p.prime] : padic_norm p p < 1 := padic_norm_p_lt_one $ nat.prime.one_lt (fact.out _) /-- `padic_norm p q` takes discrete values `p ^ -z` for `z : ℤ`. -/ @@ -130,7 +130,7 @@ include hp /-- If `q ≠ 0`, then `padic_norm p q ≠ 0`. -/ protected lemma nonzero {q : ℚ} (hq : q ≠ 0) : padic_norm p q ≠ 0 := begin - rw padic_norm.eq_zpow_of_nonzero p hq, + rw padic_norm.eq_zpow_of_nonzero hq, apply zpow_ne_zero_of_ne_zero, exact_mod_cast ne_of_gt hp.1.pos end @@ -153,13 +153,13 @@ else if hr : r = 0 then by simp [hr] else have q*r ≠ 0, from mul_ne_zero hq hr, - have (↑p : ℚ) ≠ 0, by simp [hp.1.ne_zero], + have (p : ℚ) ≠ 0, by simp [hp.1.ne_zero], by simp [padic_norm, *, padic_val_rat.mul, zpow_add₀ this, mul_comm] /-- The p-adic norm respects division. -/ @[simp] protected theorem div (q r : ℚ) : padic_norm p (q / r) = padic_norm p q / padic_norm p r := if hr : r = 0 then by simp [hr] else -eq_div_of_mul_eq (padic_norm.nonzero _ hr) (by rw [←padic_norm.mul, div_mul_cancel _ hr]) +eq_div_of_mul_eq (padic_norm.nonzero hr) (by rw [←padic_norm.mul, div_mul_cancel _ hr]) /-- The p-adic norm of an integer is at most 1. -/ protected theorem of_int (z : ℤ) : padic_norm p ↑z ≤ 1 := @@ -168,7 +168,7 @@ begin unfold padic_norm, rw [if_neg _], { refine zpow_le_one_of_nonpos _ _, - { exact_mod_cast le_of_lt hp.1.one_lt, }, + { exact_mod_cast le_of_lt hp.1.one_lt }, { rw [padic_val_rat.of_int, neg_nonpos], norm_cast, simp }}, exact_mod_cast hz, @@ -176,8 +176,8 @@ end private lemma nonarchimedean_aux {q r : ℚ} (h : padic_val_rat p q ≤ padic_val_rat p r) : padic_norm p (q + r) ≤ max (padic_norm p q) (padic_norm p r) := -have hnqp : padic_norm p q ≥ 0, from padic_norm.nonneg _ _, -have hnrp : padic_norm p r ≥ 0, from padic_norm.nonneg _ _, +have hnqp : padic_norm p q ≥ 0, from padic_norm.nonneg _, +have hnrp : padic_norm p r ≥ 0, from padic_norm.nonneg _, if hq : q = 0 then by simp [hq, max_eq_right hnrp, le_max_right] else if hr : r = 0 then @@ -206,8 +206,8 @@ the norm of `q`. protected theorem nonarchimedean {q r : ℚ} : padic_norm p (q + r) ≤ max (padic_norm p q) (padic_norm p r) := begin - wlog hle := le_total (padic_val_rat p q) (padic_val_rat p r) using [q r], - exact nonarchimedean_aux p hle + wlog hle := le_total (padic_val_rat p q) (padic_val_rat p r) using [q r], + exact nonarchimedean_aux hle end /-- @@ -215,16 +215,16 @@ The p-adic norm respects the triangle inequality: the norm of `p + q` is at most plus the norm of `q`. -/ theorem triangle_ineq (q r : ℚ) : padic_norm p (q + r) ≤ padic_norm p q + padic_norm p r := -calc padic_norm p (q + r) ≤ max (padic_norm p q) (padic_norm p r) : padic_norm.nonarchimedean p +calc padic_norm p (q + r) ≤ max (padic_norm p q) (padic_norm p r) : padic_norm.nonarchimedean ... ≤ padic_norm p q + padic_norm p r : - max_le_add_of_nonneg (padic_norm.nonneg p _) (padic_norm.nonneg p _) + max_le_add_of_nonneg (padic_norm.nonneg _) (padic_norm.nonneg _) /-- The p-adic norm of a difference is at most the max of each component. Restates the archimedean property of the p-adic norm. -/ protected theorem sub {q r : ℚ} : padic_norm p (q - r) ≤ max (padic_norm p q) (padic_norm p r) := -by rw [sub_eq_add_neg, ←padic_norm.neg p r]; apply padic_norm.nonarchimedean +by rw [sub_eq_add_neg, ←padic_norm.neg r]; apply padic_norm.nonarchimedean /-- If the p-adic norms of `q` and `r` are different, then the norm of `q + r` is equal to the max of @@ -237,7 +237,7 @@ begin have hlt : padic_norm p r < padic_norm p q, from lt_of_le_of_ne hle hne.symm, have : padic_norm p q ≤ max (padic_norm p (q + r)) (padic_norm p r), from calc padic_norm p q = padic_norm p (q + r - r) : by congr; ring - ... ≤ max (padic_norm p (q + r)) (padic_norm p (-r)) : padic_norm.nonarchimedean p + ... ≤ max (padic_norm p (q + r)) (padic_norm p (-r)) : padic_norm.nonarchimedean ... = max (padic_norm p (q + r)) (padic_norm p r) : by simp, have hnge : padic_norm p r ≤ padic_norm p (q + r), { apply le_of_not_gt, @@ -247,9 +247,8 @@ begin assumption }, have : padic_norm p q ≤ padic_norm p (q + r), by rwa [max_eq_left hnge] at this, apply _root_.le_antisymm, - { apply padic_norm.nonarchimedean p }, - { rw max_eq_left_of_lt hlt, - assumption } + { apply padic_norm.nonarchimedean }, + { rwa max_eq_left_of_lt hlt } end /-- @@ -257,20 +256,12 @@ The p-adic norm is an absolute value: positive-definite and multiplicative, sati inequality. -/ instance : is_absolute_value (padic_norm p) := -{ abv_nonneg := padic_norm.nonneg p, - abv_eq_zero := - begin - intros, - constructor; intro, - { apply zero_of_padic_norm_eq_zero p, assumption }, - { simp [*] } - end, - abv_add := padic_norm.triangle_ineq p, - abv_mul := padic_norm.mul p } - -variable {p} - -lemma dvd_iff_norm_le {n : ℕ} {z : ℤ} : ↑(p^n) ∣ z ↔ padic_norm p z ≤ ↑p ^ (-n : ℤ) := +{ abv_nonneg := padic_norm.nonneg, + abv_eq_zero := λ _, ⟨zero_of_padic_norm_eq_zero, λ hx, by simpa only [hx]⟩, + abv_add := padic_norm.triangle_ineq, + abv_mul := padic_norm.mul } + +lemma dvd_iff_norm_le {n : ℕ} {z : ℤ} : ↑(p ^ n) ∣ z ↔ padic_norm p z ≤ p ^ (-n : ℤ) := begin unfold padic_norm, split_ifs with hz, { norm_cast at hz, @@ -314,8 +305,7 @@ begin simp only [padic_norm.of_int, true_and], end -lemma of_nat (m : ℕ) : padic_norm p m ≤ 1 := -padic_norm.of_int p (m : ℤ) +lemma of_nat (m : ℕ) : padic_norm p m ≤ 1 := padic_norm.of_int (m : ℤ) /-- The `p`-adic norm of a natural `m` is one iff `p` doesn't divide `m`. -/ lemma nat_eq_one_iff (m : ℕ) : padic_norm p m = 1 ↔ ¬ p ∣ m := @@ -327,47 +317,43 @@ by simp only [←int.coe_nat_dvd, ←int_lt_one_iff, int.cast_coe_nat] open_locale big_operators lemma sum_lt {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} : - s.nonempty → (∀ i ∈ s, padic_norm p (F i) < t) → - padic_norm p (∑ i in s, F i) < t := + s.nonempty → (∀ i ∈ s, padic_norm p (F i) < t) → padic_norm p (∑ i in s, F i) < t := begin classical, refine s.induction_on (by { rintro ⟨-, ⟨⟩⟩, }) _, rintro a S haS IH - ht, by_cases hs : S.nonempty, { rw finset.sum_insert haS, - exact lt_of_le_of_lt (padic_norm.nonarchimedean p) (max_lt + exact lt_of_le_of_lt padic_norm.nonarchimedean (max_lt (ht a (finset.mem_insert_self a S)) (IH hs (λ b hb, ht b (finset.mem_insert_of_mem hb)))), }, { simp * at *, }, end lemma sum_le {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} : - s.nonempty → (∀ i ∈ s, padic_norm p (F i) ≤ t) → - padic_norm p (∑ i in s, F i) ≤ t := + s.nonempty → (∀ i ∈ s, padic_norm p (F i) ≤ t) → padic_norm p (∑ i in s, F i) ≤ t := begin classical, refine s.induction_on (by { rintro ⟨-, ⟨⟩⟩, }) _, rintro a S haS IH - ht, by_cases hs : S.nonempty, { rw finset.sum_insert haS, - exact (padic_norm.nonarchimedean p).trans (max_le + exact padic_norm.nonarchimedean.trans (max_le (ht a (finset.mem_insert_self a S)) (IH hs (λ b hb, ht b (finset.mem_insert_of_mem hb)))), }, { simp * at *, }, end -lemma sum_lt' {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} - (hF : ∀ i ∈ s, padic_norm p (F i) < t) (ht : 0 < t) : - padic_norm p (∑ i in s, F i) < t := +lemma sum_lt' {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} (hF : ∀ i ∈ s, padic_norm p (F i) < t) + (ht : 0 < t) : padic_norm p (∑ i in s, F i) < t := begin obtain rfl | hs := finset.eq_empty_or_nonempty s, { simp [ht], }, { exact sum_lt hs hF, }, end -lemma sum_le' {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} - (hF : ∀ i ∈ s, padic_norm p (F i) ≤ t) (ht : 0 ≤ t) : - padic_norm p (∑ i in s, F i) ≤ t := +lemma sum_le' {α : Type*} {F : α → ℚ} {t : ℚ} {s : finset α} (hF : ∀ i ∈ s, padic_norm p (F i) ≤ t) + (ht : 0 ≤ t) : padic_norm p (∑ i in s, F i) ≤ t := begin obtain rfl | hs := finset.eq_empty_or_nonempty s, { simp [ht], }, diff --git a/src/number_theory/padics/padic_numbers.lean b/src/number_theory/padics/padic_numbers.lean index 83f1fca5f7df0..830a6142985b7 100644 --- a/src/number_theory/padics/padic_numbers.lean +++ b/src/number_theory/padics/padic_numbers.lean @@ -88,10 +88,10 @@ let ⟨ε, hε, N1, hN1⟩ := this, from lt_max_iff.2 (or.inl this), begin by_contradiction hne, - rw ←padic_norm.neg p (f m) at hne, - have hnam := add_eq_max_of_ne p hne, + rw ← padic_norm.neg (f m) at hne, + have hnam := add_eq_max_of_ne hne, rw [padic_norm.neg, max_comm] at hnam, - rw [←hnam, sub_eq_add_neg, add_comm] at this, + rw [← hnam, sub_eq_add_neg, add_comm] at this, apply _root_.lt_irrefl _ this end ⟩ @@ -215,7 +215,7 @@ begin use (stationary_point hf), intros n hn, rw stationary_point_spec hf le_rfl hn, - simpa [H] using hε, + simpa [H] using hε end lemma val_eq_iff_norm_eq {f g : padic_seq p} (hf : ¬ f ≈ 0) (hg : ¬ g ≈ 0) : @@ -223,7 +223,7 @@ lemma val_eq_iff_norm_eq {f g : padic_seq p} (hf : ¬ f ≈ 0) (hg : ¬ g ≈ 0) begin rw [norm_eq_pow_val hf, norm_eq_pow_val hg, ← neg_inj, zpow_inj], { exact_mod_cast (fact.out p.prime).pos }, - { exact_mod_cast (fact.out p.prime).ne_one }, + { exact_mod_cast (fact.out p.prime).ne_one } end end valuation @@ -295,7 +295,7 @@ else lemma norm_values_discrete (a : padic_seq p) (ha : ¬ a ≈ 0) : (∃ (z : ℤ), a.norm = ↑p ^ (-z)) := let ⟨k, hk, hk'⟩ := norm_eq_norm_app_of_nonzero ha in -by simpa [hk] using padic_norm.values_discrete p hk' +by simpa [hk] using padic_norm.values_discrete hk' lemma norm_one : norm (1 : padic_seq p) = 1 := have h1 : ¬ (1 : padic_seq p) ≈ 0, from one_not_equiv_zero _, @@ -314,8 +314,8 @@ begin have hN' := hN _ hi, padic_index_simp [N, hf, hg] at hN' h hlt, have hpne : padic_norm p (f i) ≠ padic_norm p (-(g i)), - by rwa [ ←padic_norm.neg p (g i)] at h, - let hpnem := add_eq_max_of_ne p hpne, + by rwa [← padic_norm.neg (g i)] at h, + let hpnem := add_eq_max_of_ne hpne, have hpeq : padic_norm p ((f - g) i) = max (padic_norm p (f i)) (padic_norm p (g i)), { rwa padic_norm.neg at hpnem }, rw [hpeq, max_eq_left_of_lt hlt] at hN', @@ -423,7 +423,7 @@ else begin unfold norm at ⊢ hfgne, split_ifs at ⊢ hfgne, padic_index_simp [hfg, hf, hg] at ⊢ hfgne, - exact padic_norm.add_eq_max_of_ne p hfgne + exact padic_norm.add_eq_max_of_ne hfgne end end embedding @@ -438,7 +438,6 @@ namespace padic section completion variables {p : ℕ} [fact p.prime] -/-- The discrete field structure on `ℚ_p` is inherited from the Cauchy completion construction. -/ instance : field (ℚ_[p]) := Cauchy.field instance : inhabited ℚ_[p] := ⟨0⟩ @@ -458,17 +457,13 @@ instance : add_comm_group ℚ_[p] := by apply_instance /-- Builds the equivalence class of a Cauchy sequence of rationals. -/ def mk : padic_seq p → ℚ_[p] := quotient.mk -end completion -section completion -variables (p : ℕ) [fact p.prime] +variables (p) lemma mk_eq {f g : padic_seq p} : mk f = mk g ↔ f ≈ g := quotient.eq lemma const_equiv {q r : ℚ} : const (padic_norm p) q ≈ const (padic_norm p) r ↔ q = r := -⟨ λ heq : lim_zero (const (padic_norm p) (q - r)), - eq_of_sub_eq_zero $ const_lim_zero.1 heq, - λ heq, by rw heq; apply setoid.refl _ ⟩ +⟨ λ heq, eq_of_sub_eq_zero $ const_lim_zero.1 heq, λ heq, by rw heq; apply setoid.refl _ ⟩ @[norm_cast] lemma coe_inj {q r : ℚ} : (↑q : ℚ_[p]) = ↑r ↔ q = r := ⟨(const_equiv p).1 ∘ quotient.eq.1, λ h, by rw h⟩ @@ -510,8 +505,8 @@ begin cases em (N ≤ stationary_point hne) with hgen hngen, { apply hN _ hgen _ hi }, { have := stationary_point_spec hne le_rfl (le_of_not_le hngen), - rw ←this, - exact hN _ le_rfl _ hi }, + rw ← this, + exact hN _ le_rfl _ hi } end protected lemma nonneg (q : ℚ_[p]) : 0 ≤ padic_norm_e q := @@ -618,7 +613,7 @@ is a sequence of rationals with the same limit point as `f`. -/ def lim_seq : ℕ → ℚ := λ n, classical.some (rat_dense' (f n) (div_nat_pos n)) lemma exi_rat_seq_conv {ε : ℚ} (hε : 0 < ε) : - ∃ N, ∀ i ≥ N, padic_norm_e (f i - ((lim_seq f) i : ℚ_[p])) < ε := + ∃ N, ∀ i ≥ N, padic_norm_e (f i - (lim_seq f i : ℚ_[p])) < ε := begin refine (exists_nat_gt (1/ε)).imp (λ N hN i hi, _), have h := classical.some_spec (rat_dense' (f i) (div_nat_pos i)), @@ -638,7 +633,7 @@ begin existsi max N N2, intros j hj, suffices : - padic_norm_e ((↑(lim_seq f j) - f (max N N2)) + (f (max N N2) - lim_seq f (max N N2))) < ε, + padic_norm_e ((lim_seq f j - f (max N N2)) + (f (max N N2) - lim_seq f (max N N2))) < ε, { ring_nf at this ⊢, rw [← padic_norm_e.eq_padic_norm'], exact_mod_cast this }, @@ -697,7 +692,7 @@ instance : has_dist ℚ_[p] := ⟨λ x y, padic_norm_e (x - y)⟩ instance : metric_space ℚ_[p] := { dist_self := by simp [dist], dist := dist, - dist_comm := λ x y, by unfold dist; rw ←padic_norm_e.neg (x - y); simp, + dist_comm := λ x y, by unfold dist; rw ← padic_norm_e.neg (x - y); simp, dist_triangle := begin intros, unfold dist, @@ -820,14 +815,14 @@ theorem norm_rat_le_one : ∀ {q : ℚ} (hq : ¬ p ∣ q.denom), ∥(q : ℚ_[p] from mt rat.zero_iff_num_zero.1 hnz, rw [padic_norm_e.eq_padic_norm], norm_cast, - rw [padic_norm.eq_zpow_of_nonzero p hnz', padic_val_rat, neg_sub, + rw [padic_norm.eq_zpow_of_nonzero hnz', padic_val_rat, neg_sub, padic_val_nat.eq_zero_of_not_dvd hq], norm_cast, rw [zero_sub, zpow_neg, zpow_coe_nat], apply inv_le_one, { norm_cast, apply one_le_pow, - exact hp.1.pos, }, + exact hp.1.pos } end theorem norm_int_le_one (z : ℤ) : ∥(z : ℚ_[p])∥ ≤ 1 := @@ -855,7 +850,7 @@ begin convert zpow_zero _, rw [neg_eq_zero, padic_val_rat.of_int], norm_cast, - apply padic_val_int.eq_zero_of_not_dvd h, } }, + apply padic_val_int.eq_zero_of_not_dvd h } }, { rintro ⟨x, rfl⟩, push_cast, rw padic_norm_e.mul, @@ -864,15 +859,15 @@ begin ... < 1 : _, { rw [mul_one, padic_norm_e.norm_p], apply inv_lt_one, - exact_mod_cast hp.1.one_lt }, }, + exact_mod_cast hp.1.one_lt } } end -lemma norm_int_le_pow_iff_dvd (k : ℤ) (n : ℕ) : ∥(k : ℚ_[p])∥ ≤ ((↑p)^(-n : ℤ)) ↔ ↑(p^n) ∣ k := +lemma norm_int_le_pow_iff_dvd (k : ℤ) (n : ℕ) : ∥(k : ℚ_[p])∥ ≤ (↑p)^(-n : ℤ) ↔ ↑(p^n) ∣ k := begin have : (p : ℝ) ^ (-n : ℤ) = ↑((p ^ (-n : ℤ) : ℚ)), {simp}, rw [show (k : ℚ_[p]) = ((k : ℚ) : ℚ_[p]), by norm_cast, eq_padic_norm, this], norm_cast, - rw padic_norm.dvd_iff_norm_le, + rw ← padic_norm.dvd_iff_norm_le end lemma eq_of_norm_add_lt_right {p : ℕ} {hp : fact p.prime} {z1 z2 : ℚ_[p]} @@ -889,8 +884,8 @@ end normed_space end padic_norm_e namespace padic -variables {p : ℕ} [hp_prime : fact p.prime] -include hp_prime +variables {p : ℕ} [hp : fact p.prime] +include hp set_option eqn_compiler.zeta true instance complete : cau_seq.is_complete ℚ_[p] norm := @@ -935,9 +930,7 @@ begin exact this.imp (λ N hN n hn, hε (hN n hn)) end -/-! -### Valuation on `ℚ_[p]` --/ +/-! ### Valuation on `ℚ_[p]` -/ /-- `padic.valuation` lifts the p-adic valuation on rationals to `ℚ_[p]`. @@ -950,7 +943,7 @@ begin simp [hf, hg, padic_seq.valuation] }, { have hg : ¬ g ≈ 0, from (λ hg, hf (setoid.trans h hg)), rw padic_seq.val_eq_iff_norm_eq hf hg, - exact padic_seq.norm_equiv h }, + exact padic_seq.norm_equiv h } end) @[simp] lemma valuation_zero : valuation (0 : ℚ_[p]) = 0 := @@ -962,7 +955,7 @@ begin have h : ¬ cau_seq.const (padic_norm p) 1 ≈ 0, { assume H, erw const_equiv p at H, exact one_ne_zero H }, rw dif_neg h, - simp, + simp end lemma norm_eq_pow_val {x : ℚ_[p]} : x ≠ 0 → ∥x∥ = p^(-x.valuation) := @@ -976,7 +969,7 @@ begin { apply cau_seq.not_lim_zero_of_not_congr_zero, contrapose! hf, apply quotient.sound, - simpa using hf, } + simpa using hf } end @[simp] lemma valuation_p : valuation (p : ℚ_[p]) = 1 := @@ -999,7 +992,7 @@ begin { have h_norm : ∥x + y∥ ≤ (max ∥x∥ ∥y∥) := padic_norm_e.nonarchimedean x y, have hp_one : (1 : ℝ) < p, { rw [← nat.cast_one, nat.cast_lt], - exact nat.prime.one_lt hp_prime.elim, }, + exact nat.prime.one_lt hp.elim }, rw [norm_eq_pow_val hx, norm_eq_pow_val hy, norm_eq_pow_val hxy] at h_norm, exact min_le_of_zpow_le_max hp_one h_norm }} end @@ -1010,13 +1003,13 @@ begin have h_norm : ∥x * y∥ = ∥x∥ * ∥y∥ := norm_mul x y, have hp_ne_one : (p : ℝ) ≠ 1, { rw [← nat.cast_one, ne.def, nat.cast_inj], - exact nat.prime.ne_one hp_prime.elim, }, + exact nat.prime.ne_one hp.elim }, have hp_pos : (0 : ℝ) < p, { rw [← nat.cast_zero, nat.cast_lt], - exact nat.prime.pos hp_prime.elim }, + exact nat.prime.pos hp.elim }, rw [norm_eq_pow_val hx, norm_eq_pow_val hy, norm_eq_pow_val (mul_ne_zero hx hy), ← zpow_add₀ (ne_of_gt hp_pos), zpow_inj hp_pos hp_ne_one, ← neg_add, neg_inj] at h_norm, - exact h_norm, + exact h_norm end /-- The additive p-adic valuation on `ℚ_p`, with values in `with_top ℤ`. -/ @@ -1027,7 +1020,7 @@ def add_valuation_def : ℚ_[p] → (with_top ℤ) := by simp only [add_valuation_def, if_pos (eq.refl _)] @[simp] lemma add_valuation.map_one : add_valuation_def (1 : ℚ_[p]) = 0 := -by simp only [add_valuation_def, if_neg (one_ne_zero), valuation_one, +by simp only [add_valuation_def, if_neg one_ne_zero, valuation_one, with_top.coe_zero] lemma add_valuation.map_mul (x y : ℚ_[p]) : @@ -1048,11 +1041,11 @@ begin simp only [add_valuation_def], by_cases hxy : x + y = 0, { rw [hxy, if_pos (eq.refl _)], - exact le_top, }, + exact le_top }, { by_cases hx : x = 0, { simp only [hx, if_pos (eq.refl _), min_eq_right, le_top, zero_add, le_refl] }, { by_cases hy : y = 0, - { simp only [hy, if_pos (eq.refl _), min_eq_left, le_top, add_zero, le_refl], }, + { simp only [hy, if_pos (eq.refl _), min_eq_left, le_top, add_zero, le_refl] }, { rw [if_neg hx, if_neg hy, if_neg hxy, ← with_top.coe_min, with_top.coe_le_coe], exact valuation_map_add hxy }}} end @@ -1067,17 +1060,19 @@ add_valuation.of add_valuation_def add_valuation.map_zero add_valuation.map_one by simp only [add_valuation, add_valuation.of_apply, add_valuation_def, if_neg hx] section norm_le_iff + /-! ### Various characterizations of open unit balls -/ + lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℚ_[p]) (n : ℤ) : ∥x∥ ≤ p ^ n ↔ ∥x∥ < p ^ (n + 1) := begin have aux : ∀ n : ℤ, 0 < (p ^ n : ℝ), - { apply nat.zpow_pos_of_pos, exact hp_prime.1.pos }, - by_cases hx0 : x = 0, { simp [hx0, norm_zero, aux, le_of_lt (aux _)], }, + { apply nat.zpow_pos_of_pos, exact hp.1.pos }, + by_cases hx0 : x = 0, { simp [hx0, norm_zero, aux, le_of_lt (aux _)] }, rw norm_eq_pow_val hx0, - have h1p : 1 < (p : ℝ), { exact_mod_cast hp_prime.1.one_lt }, + have h1p : 1 < (p : ℝ), { exact_mod_cast hp.1.one_lt }, have H := zpow_strict_mono h1p, - rw [H.le_iff_le, H.lt_iff_lt, int.lt_add_one_iff], + rw [H.le_iff_le, H.lt_iff_lt, int.lt_add_one_iff] end lemma norm_lt_pow_iff_norm_le_pow_sub_one (x : ℚ_[p]) (n : ℤ) : diff --git a/src/number_theory/padics/padic_val.lean b/src/number_theory/padics/padic_val.lean index e89f21ed2e3c2..b32672aca4be5 100644 --- a/src/number_theory/padics/padic_val.lean +++ b/src/number_theory/padics/padic_val.lean @@ -69,25 +69,20 @@ by simp [padic_val_nat] /-- `padic_val_nat p 1` is 0 for any `p`. -/ @[simp] protected lemma one : padic_val_nat p 1 = 0 := -by unfold padic_val_nat; split_ifs; simp * +by { unfold padic_val_nat, split_ifs, simp } /-- For `p ≠ 0, p ≠ 1, `padic_val_rat p p` is 1. -/ @[simp] lemma self (hp : 1 < p) : padic_val_nat p p = 1 := begin have neq_one : (¬ p = 1) ↔ true, - { exact iff_of_true ((ne_of_lt hp).symm) trivial, }, + { exact iff_of_true ((ne_of_lt hp).symm) trivial }, have eq_zero_false : (p = 0) ↔ false, { exact iff_false_intro ((ne_of_lt (trans zero_lt_one hp)).symm) }, - simp [padic_val_nat, neq_one, eq_zero_false], + simp [padic_val_nat, neq_one, eq_zero_false] end lemma eq_zero_of_not_dvd {n : ℕ} (h : ¬ p ∣ n) : padic_val_nat p n = 0 := -begin - rw padic_val_nat, - split_ifs, - { simp [multiplicity_eq_zero_of_not_dvd h], }, - refl, -end +by { rw padic_val_nat, split_ifs; simp [multiplicity_eq_zero_of_not_dvd h] } end padic_val_nat @@ -107,8 +102,8 @@ lemma of_ne_one_ne_zero {z : ℤ} (hp : p ≠ 1) (hz : z ≠ 0) : padic_val_int (multiplicity (p : ℤ) z).get (by {apply multiplicity.finite_int_iff.2, simp [hp, hz]}) := begin rw [padic_val_int, padic_val_nat, dif_pos (and.intro hp (int.nat_abs_pos_of_ne_zero hz))], - simp_rw multiplicity.int.nat_abs p z, - refl, + simp only [multiplicity.int.nat_abs p z], + refl end /-- `padic_val_int p 0` is 0 for any `p`. -/ @@ -130,10 +125,7 @@ by simp [padic_val_nat.self hp] lemma eq_zero_of_not_dvd {z : ℤ} (h : ¬ (p : ℤ) ∣ z) : padic_val_int p z = 0 := begin rw [padic_val_int, padic_val_nat], - split_ifs, - { simp_rw multiplicity.int.nat_abs, - simp [multiplicity_eq_zero_of_not_dvd h], }, - refl, + split_ifs; simp [multiplicity.int.nat_abs, multiplicity_eq_zero_of_not_dvd h], end end padic_val_int @@ -147,7 +139,9 @@ def padic_val_rat (p : ℕ) (q : ℚ) : ℤ := padic_val_int p q.num - padic_val_nat p q.denom namespace padic_val_rat + open multiplicity + variables {p : ℕ} /-- `padic_val_rat p q` is symmetric in `q`. -/ @@ -155,37 +149,33 @@ variables {p : ℕ} by simp [padic_val_rat, padic_val_int] /-- `padic_val_rat p 0` is 0 for any `p`. -/ -@[simp] -protected lemma zero (m : nat) : padic_val_rat m 0 = 0 := by simp [padic_val_rat, padic_val_int] +@[simp] protected lemma zero : padic_val_rat p 0 = 0 := by simp [padic_val_rat] /-- `padic_val_rat p 1` is 0 for any `p`. -/ -@[simp] protected lemma one : padic_val_rat p 1 = 0 := by simp [padic_val_rat, padic_val_int] +@[simp] protected lemma one : padic_val_rat p 1 = 0 := by simp [padic_val_rat] /-- The p-adic value of an integer `z ≠ 0` is its p-adic_value as a rational -/ -@[simp] lemma of_int {z : ℤ} : padic_val_rat p (z : ℚ) = padic_val_int p z := -by simp [padic_val_rat] +@[simp] lemma of_int {z : ℤ} : padic_val_rat p z = padic_val_int p z := by simp [padic_val_rat] /-- The p-adic value of an integer `z ≠ 0` is the multiplicity of `p` in `z`. -/ -lemma of_int_multiplicity (z : ℤ) (hp : p ≠ 1) (hz : z ≠ 0) : +lemma of_int_multiplicity {z : ℤ} (hp : p ≠ 1) (hz : z ≠ 0) : padic_val_rat p (z : ℚ) = (multiplicity (p : ℤ) z).get (finite_int_iff.2 ⟨hp, hz⟩) := by rw [of_int, padic_val_int.of_ne_one_ne_zero hp hz] -lemma multiplicity_sub_multiplicity {q : ℚ} (hp : p ≠ 1) (hq : q ≠ 0) : - padic_val_rat p q = +lemma multiplicity_sub_multiplicity {q : ℚ} (hp : p ≠ 1) (hq : q ≠ 0) : padic_val_rat p q = (multiplicity (p : ℤ) q.num).get (finite_int_iff.2 ⟨hp, rat.num_ne_zero_of_ne_zero hq⟩) - (multiplicity p q.denom).get - (by { rw [←finite_iff_dom, finite_nat_iff, and_iff_right hp], exact q.pos }) := + (by { rw [← finite_iff_dom, finite_nat_iff], exact ⟨hp, q.pos⟩ }) := begin rw [padic_val_rat, padic_val_int.of_ne_one_ne_zero hp, padic_val_nat, dif_pos], { refl }, { exact ⟨hp, q.pos⟩ }, - { exact rat.num_ne_zero_of_ne_zero hq }, + { exact rat.num_ne_zero_of_ne_zero hq } end /-- The p-adic value of an integer `z ≠ 0` is its p-adic_value as a rational -/ -@[simp] lemma of_nat {n : ℕ} : padic_val_rat p (n : ℚ) = padic_val_nat p n := -by simp [padic_val_rat, padic_val_int] +@[simp] lemma of_nat {n : ℕ} : padic_val_rat p n = padic_val_nat p n := by simp [padic_val_rat] /-- For `p ≠ 0, p ≠ 1, `padic_val_rat p p` is 1. -/ lemma self (hp : 1 < p) : padic_val_rat p p = 1 := by simp [of_nat, hp] @@ -194,46 +184,48 @@ end padic_val_rat section padic_val_nat -lemma zero_le_padic_val_rat_of_nat (p n : ℕ) : 0 ≤ padic_val_rat p n := by simp +variables {p : ℕ} + +lemma zero_le_padic_val_rat_of_nat (n : ℕ) : 0 ≤ padic_val_rat p n := by simp --- /-- `padic_val_rat` coincides with `padic_val_nat`. -/ -@[norm_cast] lemma padic_val_rat_of_nat (p n : ℕ) : +/-- `padic_val_rat` coincides with `padic_val_nat`. -/ +@[norm_cast] lemma padic_val_rat_of_nat (n : ℕ) : ↑(padic_val_nat p n) = padic_val_rat p n := -by simp [padic_val_rat, padic_val_int] +by simp /-- A simplification of `padic_val_nat` when one input is prime, by analogy with `padic_val_rat_def`. -/ -lemma padic_val_nat_def {p : ℕ} [hp : fact p.prime] {n : ℕ} (hn : 0 < n) : +lemma padic_val_nat_def [hp : fact p.prime] {n : ℕ} (hn : 0 < n) : padic_val_nat p n = (multiplicity p n).get (multiplicity.finite_nat_iff.2 ⟨nat.prime.ne_one hp.1, hn⟩) := begin simp [padic_val_nat], split_ifs, - { refl, }, + { refl }, { exfalso, - apply h ⟨(hp.out).ne_one, hn⟩, } + exact h ⟨hp.out.ne_one, hn⟩ } end -lemma padic_val_nat_def' {n p : ℕ} (hp : p ≠ 1) (hn : 0 < n) : +lemma padic_val_nat_def' {n : ℕ} (hp : p ≠ 1) (hn : 0 < n) : ↑(padic_val_nat p n) = multiplicity p n := by simp [padic_val_nat, hp, hn] -@[simp] lemma padic_val_nat_self (p : ℕ) [fact p.prime] : padic_val_nat p p = 1 := +@[simp] lemma padic_val_nat_self [fact p.prime] : padic_val_nat p p = 1 := by simp [padic_val_nat_def (fact.out p.prime).pos] -lemma one_le_padic_val_nat_of_dvd - {n p : nat} [prime : fact p.prime] (n_pos : 0 < n) (div : p ∣ n) : +lemma one_le_padic_val_nat_of_dvd {n : ℕ} [hp : fact p.prime] (hn : 0 < n) (div : p ∣ n) : 1 ≤ padic_val_nat p n := begin - rw @padic_val_nat_def _ prime _ n_pos, + rw padic_val_nat_def hn, let one_le_mul : _ ≤ multiplicity p n := - @multiplicity.le_multiplicity_of_pow_dvd _ _ _ p n 1 (begin norm_num, exact div end), + multiplicity.le_multiplicity_of_pow_dvd (by { rw [pow_one], exact div }), simp only [nat.cast_one] at one_le_mul, rcases one_le_mul with ⟨_, q⟩, dsimp at q, solve_by_elim, + exact hp end lemma dvd_iff_padic_val_nat_ne_zero {p n : ℕ} [fact p.prime] (hn0 : n ≠ 0) : @@ -244,36 +236,39 @@ lemma dvd_iff_padic_val_nat_ne_zero {p n : ℕ} [fact p.prime] (hn0 : n ≠ 0) : end padic_val_nat namespace padic_val_rat + open multiplicity -variables (p : ℕ) [p_prime : fact p.prime] -include p_prime + +variables {p : ℕ} [hp : fact p.prime] + +include hp /-- The multiplicity of `p : ℕ` in `a : ℤ` is finite exactly when `a ≠ 0`. -/ -lemma finite_int_prime_iff {p : ℕ} [p_prime : fact p.prime] {a : ℤ} : finite (p : ℤ) a ↔ a ≠ 0 := -by simp [finite_int_iff, ne.symm (ne_of_lt (p_prime.1.one_lt))] +lemma finite_int_prime_iff {a : ℤ} : finite (p : ℤ) a ↔ a ≠ 0 := +by simp [finite_int_iff, ne.symm (ne_of_lt hp.1.one_lt)] /-- A rewrite lemma for `padic_val_rat p q` when `q` is expressed in terms of `rat.mk`. -/ -protected lemma defn {q : ℚ} {n d : ℤ} (hqz : q ≠ 0) (qdf : q = n /. d) : - padic_val_rat p q = (multiplicity (p : ℤ) n).get (finite_int_iff.2 - ⟨ne.symm $ ne_of_lt p_prime.1.one_lt, λ hn, by simp * at *⟩) - - (multiplicity (p : ℤ) d).get (finite_int_iff.2 ⟨ne.symm $ ne_of_lt p_prime.1.one_lt, - λ hd, by simp * at *⟩) := +protected lemma defn (p : ℕ) [hp : fact p.prime] {q : ℚ} {n d : ℤ} (hqz : q ≠ 0) + (qdf : q = n /. d) : padic_val_rat p q + = (multiplicity (p : ℤ) n).get + (finite_int_iff.2 ⟨ne.symm $ ne_of_lt hp.1.one_lt, λ hn, by simp * at *⟩) + - (multiplicity (p : ℤ) d).get + (finite_int_iff.2 ⟨ne.symm $ ne_of_lt hp.1.one_lt, λ hd, by simp * at *⟩) := have hd : d ≠ 0, from rat.mk_denom_ne_zero_of_ne_zero hqz qdf, let ⟨c, hc1, hc2⟩ := rat.num_denom_mk hd qdf in begin rw [padic_val_rat.multiplicity_sub_multiplicity]; - simp [hc1, hc2, multiplicity.mul' (nat.prime_iff_prime_int.1 p_prime.1), - (ne.symm (ne_of_lt p_prime.1.one_lt)), hqz, pos_iff_ne_zero], - simp_rw [int.coe_nat_multiplicity p q.denom], + simp [hc1, hc2, multiplicity.mul' (nat.prime_iff_prime_int.1 hp.1), + ne.symm (ne_of_lt hp.1.one_lt), hqz, pos_iff_ne_zero, int.coe_nat_multiplicity p q.denom] end /-- A rewrite lemma for `padic_val_rat p (q * r)` with conditions `q ≠ 0`, `r ≠ 0`. -/ protected lemma mul {q r : ℚ} (hq : q ≠ 0) (hr : r ≠ 0) : padic_val_rat p (q * r) = padic_val_rat p q + padic_val_rat p r := -have q*r = (q.num * r.num) /. (↑q.denom * ↑r.denom), by rw_mod_cast rat.mul_num_denom, +have q*r = (q.num * r.num) /. (q.denom * r.denom), by rw_mod_cast rat.mul_num_denom, have hq' : q.num /. q.denom ≠ 0, by rw rat.num_denom; exact hq, have hr' : r.num /. r.denom ≠ 0, by rw rat.num_denom; exact hr, -have hp' : _root_.prime (p : ℤ), from nat.prime_iff_prime_int.1 p_prime.1, +have hp' : _root_.prime (p : ℤ), from nat.prime_iff_prime_int.1 hp.1, begin rw [padic_val_rat.defn p (mul_ne_zero hq hr) this], conv_rhs { rw [←(@rat.num_denom q), padic_val_rat.defn p hq', @@ -284,8 +279,7 @@ end /-- A rewrite lemma for `padic_val_rat p (q^k)` with condition `q ≠ 0`. -/ protected lemma pow {q : ℚ} (hq : q ≠ 0) {k : ℕ} : padic_val_rat p (q ^ k) = k * padic_val_rat p q := -by induction k; simp [*, padic_val_rat.mul _ hq (pow_ne_zero _ hq), - pow_succ, add_mul, add_comm] +by induction k; simp [*, padic_val_rat.mul hq (pow_ne_zero _ hq), pow_succ, add_mul, add_comm] /-- A rewrite lemma for `padic_val_rat p (q⁻¹)` with condition `q ≠ 0`. @@ -294,16 +288,19 @@ protected lemma inv (q : ℚ) : padic_val_rat p (q⁻¹) = -padic_val_rat p q := begin by_cases hq : q = 0, - { simp [hq], }, - { rw [eq_neg_iff_add_eq_zero, ← padic_val_rat.mul p (inv_ne_zero hq) hq, - inv_mul_cancel hq, padic_val_rat.one] }, + { simp [hq] }, + { rw [eq_neg_iff_add_eq_zero, ← padic_val_rat.mul (inv_ne_zero hq) hq, + inv_mul_cancel hq, padic_val_rat.one], + exact hp }, end /-- A rewrite lemma for `padic_val_rat p (q / r)` with conditions `q ≠ 0`, `r ≠ 0`. -/ protected lemma div {q r : ℚ} (hq : q ≠ 0) (hr : r ≠ 0) : padic_val_rat p (q / r) = padic_val_rat p q - padic_val_rat p r := -by rw [div_eq_mul_inv, padic_val_rat.mul p hq (inv_ne_zero hr), - padic_val_rat.inv p r, sub_eq_add_neg] +begin + rw [div_eq_mul_inv, padic_val_rat.mul hq (inv_ne_zero hr), padic_val_rat.inv r, sub_eq_add_neg], + all_goals { exact hp } +end /-- A condition for `padic_val_rat p (n₁ / d₁) ≤ padic_val_rat p (n₂ / d₂), @@ -325,8 +322,8 @@ have hf2 : finite (p : ℤ) (n₂ * d₁), ← add_sub_assoc, le_sub_iff_add_le], norm_cast, - rw [← multiplicity.mul' (nat.prime_iff_prime_int.1 p_prime.1) hf1, add_comm, - ← multiplicity.mul' (nat.prime_iff_prime_int.1 p_prime.1) hf2, + rw [← multiplicity.mul' (nat.prime_iff_prime_int.1 hp.1) hf1, add_comm, + ← multiplicity.mul' (nat.prime_iff_prime_int.1 hp.1) hf2, part_enat.get_le_get, multiplicity_le_multiplicity_iff] } /-- @@ -349,18 +346,18 @@ have hqrd : q.num * ↑(r.denom) + ↑(q.denom) * r.num ≠ 0, from rat.mk_num_ne_zero_of_ne_zero hqr hqreq, begin conv_lhs { rw ←(@rat.num_denom q) }, - rw [hqreq, padic_val_rat_le_padic_val_rat_iff p hqn hqrd hqd (mul_ne_zero hqd hrd), + rw [hqreq, padic_val_rat_le_padic_val_rat_iff hqn hqrd hqd (mul_ne_zero hqd hrd), ← multiplicity_le_multiplicity_iff, mul_left_comm, - multiplicity.mul (nat.prime_iff_prime_int.1 p_prime.1), add_mul], + multiplicity.mul (nat.prime_iff_prime_int.1 hp.1), add_mul], rw [←(@rat.num_denom q), ←(@rat.num_denom r), - padic_val_rat_le_padic_val_rat_iff p hqn hrn hqd hrd, ← multiplicity_le_multiplicity_iff] at h, + padic_val_rat_le_padic_val_rat_iff hqn hrn hqd hrd, ← multiplicity_le_multiplicity_iff] at h, calc _ ≤ min (multiplicity ↑p (q.num * ↑(r.denom) * ↑(q.denom))) (multiplicity ↑p (↑(q.denom) * r.num * ↑(q.denom))) : (le_min - (by rw [@multiplicity.mul _ _ _ _ (_ * _) _ (nat.prime_iff_prime_int.1 p_prime.1), add_comm]) + (by rw [@multiplicity.mul _ _ _ _ (_ * _) _ (nat.prime_iff_prime_int.1 hp.1), add_comm]) (by rw [mul_assoc, @multiplicity.mul _ _ _ _ (q.denom : ℤ) - (_ * _) (nat.prime_iff_prime_int.1 p_prime.1)]; - exact add_le_add_left h _)) - ... ≤ _ : min_le_multiplicity_add + (_ * _) (nat.prime_iff_prime_int.1 hp.1)]; exact add_le_add_left h _)) + ... ≤ _ : min_le_multiplicity_add, + all_goals { exact hp } end /-- @@ -369,8 +366,8 @@ The minimum of the valuations of `q` and `r` is less than or equal to the valuat theorem min_le_padic_val_rat_add {q r : ℚ} (hqr : q + r ≠ 0) : min (padic_val_rat p q) (padic_val_rat p r) ≤ padic_val_rat p (q + r) := (le_total (padic_val_rat p q) (padic_val_rat p r)).elim - (λ h, by rw [min_eq_left h]; exact le_padic_val_rat_add_of_le _ hqr h) - (λ h, by rw [min_eq_right h, add_comm]; exact le_padic_val_rat_add_of_le _ + (λ h, by rw [min_eq_left h]; exact le_padic_val_rat_add_of_le hqr h) + (λ h, by rw [min_eq_right h, add_comm]; exact le_padic_val_rat_add_of_le (by rwa add_comm) h) open_locale big_operators @@ -387,81 +384,82 @@ begin by_cases h : ∑ (x : ℕ) in finset.range d, F x = 0, { rw [h, zero_add], exact hF d (lt_add_one _) }, - { refine lt_of_lt_of_le _ (min_le_padic_val_rat_add p hn0), + { refine lt_of_lt_of_le _ (min_le_padic_val_rat_add hn0), { refine lt_min (hd (λ i hi, _) h) (hF d (lt_add_one _)), - exact hF _ (lt_trans hi (lt_add_one _)) }, } } + exact hF _ (lt_trans hi (lt_add_one _)) } } } end end padic_val_rat namespace padic_val_nat -/-- A rewrite lemma for `padic_val_nat p (q * r)` with conditions `q ≠ 0`, `r ≠ 0`. -/ -protected lemma mul (p : ℕ) [p_prime : fact p.prime] {q r : ℕ} (hq : q ≠ 0) (hr : r ≠ 0) : - padic_val_nat p (q * r) = padic_val_nat p q + padic_val_nat p r := +variables {p a b : ℕ} [hp : fact p.prime] + +include hp + +/-- A rewrite lemma for `padic_val_nat p (a * b)` with conditions `a ≠ 0`, `b ≠ 0`. -/ +protected lemma mul (ha : a ≠ 0) (hb : b ≠ 0) : + padic_val_nat p (a * b) = padic_val_nat p a + padic_val_nat p b := begin apply int.coe_nat_inj, simp only [padic_val_rat_of_nat, nat.cast_mul], rw padic_val_rat.mul, norm_cast, - exact cast_ne_zero.mpr hq, - exact cast_ne_zero.mpr hr, + exact cast_ne_zero.mpr ha, + exact cast_ne_zero.mpr hb end -protected lemma div_of_dvd (p : ℕ) [hp : fact p.prime] {a b : ℕ} (h : b ∣ a) : +protected lemma div_of_dvd (h : b ∣ a) : padic_val_nat p (a / b) = padic_val_nat p a - padic_val_nat p b := begin rcases eq_or_ne a 0 with rfl | ha, { simp }, obtain ⟨k, rfl⟩ := h, obtain ⟨hb, hk⟩ := mul_ne_zero_iff.mp ha, - rw [mul_comm, k.mul_div_cancel hb.bot_lt, padic_val_nat.mul p hk hb, nat.add_sub_cancel] + rw [mul_comm, k.mul_div_cancel hb.bot_lt, padic_val_nat.mul hk hb, nat.add_sub_cancel], + exact hp end /-- Dividing out by a prime factor reduces the padic_val_nat by 1. -/ -protected lemma div {p : ℕ} [p_prime : fact p.prime] {b : ℕ} (dvd : p ∣ b) : - (padic_val_nat p (b / p)) = (padic_val_nat p b) - 1 := +protected lemma div (dvd : p ∣ b) : padic_val_nat p (b / p) = (padic_val_nat p b) - 1 := begin - convert padic_val_nat.div_of_dvd p dvd, - rw padic_val_nat_self p + convert padic_val_nat.div_of_dvd dvd, + rw padic_val_nat_self, + exact hp end /-- A version of `padic_val_rat.pow` for `padic_val_nat` -/ -protected lemma pow (p q n : ℕ) [fact p.prime] (hq : q ≠ 0) : - padic_val_nat p (q ^ n) = n * padic_val_nat p q := -begin - apply @nat.cast_injective ℤ, - push_cast, - exact padic_val_rat.pow _ (cast_ne_zero.mpr hq), -end +protected lemma pow (n : ℕ) (ha : a ≠ 0) : + padic_val_nat p (a ^ n) = n * padic_val_nat p a := +by simpa only [← @nat.cast_inj ℤ] with push_cast using padic_val_rat.pow (cast_ne_zero.mpr ha) -@[simp] protected lemma prime_pow (p n : ℕ) [fact p.prime] : padic_val_nat p (p ^ n) = n := -by rw [padic_val_nat.pow p _ _ (fact.out p.prime).ne_zero, padic_val_nat_self p, mul_one] +@[simp] protected lemma prime_pow (n : ℕ) : padic_val_nat p (p ^ n) = n := +by rwa [padic_val_nat.pow _ (fact.out p.prime).ne_zero, padic_val_nat_self, mul_one] -protected lemma div_pow {p : ℕ} [p_prime : fact p.prime] {b k : ℕ} (dvd : p ^ k ∣ b) : - (padic_val_nat p (b / p ^ k)) = (padic_val_nat p b) - k := +protected lemma div_pow (dvd : p ^ a ∣ b) : padic_val_nat p (b / p ^ a) = (padic_val_nat p b) - a := begin - convert padic_val_nat.div_of_dvd p dvd, - rw padic_val_nat.prime_pow + convert padic_val_nat.div_of_dvd dvd, + rw padic_val_nat.prime_pow, + exact hp end end padic_val_nat section padic_val_nat -lemma dvd_of_one_le_padic_val_nat {n p : nat} (hp : 1 ≤ padic_val_nat p n) : +lemma dvd_of_one_le_padic_val_nat {p n : ℕ} (hp : 1 ≤ padic_val_nat p n) : p ∣ n := begin by_contra h, rw padic_val_nat.eq_zero_of_not_dvd h at hp, - exact lt_irrefl 0 (lt_of_lt_of_le zero_lt_one hp), + exact lt_irrefl 0 (lt_of_lt_of_le zero_lt_one hp) end lemma pow_padic_val_nat_dvd {p n : ℕ} : p ^ (padic_val_nat p n) ∣ n := begin rcases n.eq_zero_or_pos with rfl | hn, { simp }, rcases eq_or_ne p 1 with rfl | hp, { simp }, - rw [multiplicity.pow_dvd_iff_le_multiplicity, padic_val_nat_def']; assumption, + rw [multiplicity.pow_dvd_iff_le_multiplicity, padic_val_nat_def']; assumption end lemma pow_succ_padic_val_nat_not_dvd {p n : ℕ} [hp : fact (nat.prime p)] (hn : 0 < n) : @@ -476,7 +474,7 @@ begin { apply_instance } end -lemma padic_val_nat_dvd_iff (p : ℕ) [hp :fact p.prime] (n : ℕ) (a : ℕ) : +lemma padic_val_nat_dvd_iff {p : ℕ} (n : ℕ) [hp : fact p.prime] (a : ℕ) : p^n ∣ a ↔ a = 0 ∨ n ≤ padic_val_nat p a := begin split, @@ -488,17 +486,17 @@ begin exact λ hn, or.inl h } }, { rintro (rfl|h), { exact dvd_zero (p ^ n) }, - { exact dvd_trans (pow_dvd_pow p h) pow_padic_val_nat_dvd } }, + { exact dvd_trans (pow_dvd_pow p h) pow_padic_val_nat_dvd } } end -lemma padic_val_nat_primes {p q : ℕ} [p_prime : fact p.prime] [q_prime : fact q.prime] +lemma padic_val_nat_primes {p q : ℕ} [hp : fact p.prime] [hq : fact q.prime] (neq : p ≠ q) : padic_val_nat p q = 0 := @padic_val_nat.eq_zero_of_not_dvd p q $ -(not_congr (iff.symm (prime_dvd_prime_iff_eq p_prime.1 q_prime.1))).mp neq + (not_congr (iff.symm (prime_dvd_prime_iff_eq hp.1 hq.1))).mp neq -protected lemma padic_val_nat.div' {p : ℕ} [p_prime : fact p.prime] : +protected lemma padic_val_nat.div' {p : ℕ} [hp : fact p.prime] : ∀ {m : ℕ} (cpm : coprime p m) {b : ℕ} (dvd : m ∣ b), padic_val_nat p (b / m) = padic_val_nat p b -| 0 := λ cpm b dvd, by { rw zero_dvd_iff at dvd, rw [dvd, nat.zero_div], } +| 0 := λ cpm b dvd, by { rw zero_dvd_iff at dvd, rw [dvd, nat.zero_div] } | (n + 1) := λ cpm b dvd, begin @@ -509,9 +507,9 @@ protected lemma padic_val_nat.div' {p : ℕ} [p_prime : fact p.prime] : { suffices : ¬ p ∣ (n+1), { rw [padic_val_nat.eq_zero_of_not_dvd this, zero_add] }, contrapose! cpm, - exact p_prime.1.dvd_iff_not_coprime.mp cpm }, + exact hp.1.dvd_iff_not_coprime.mp cpm }, { exact nat.succ_ne_zero _ }, - { exact hc } }, + { exact hc } } end open_locale big_operators @@ -526,7 +524,7 @@ begin exact ⟨(pow_dvd_pow p $ by linarith).trans pow_padic_val_nat_dvd, hn⟩ end -lemma range_pow_padic_val_nat_subset_divisors' {n : ℕ} (p : ℕ) [h : fact p.prime] : +lemma range_pow_padic_val_nat_subset_divisors' {p n : ℕ} [hp : fact p.prime] : (finset.range (padic_val_nat p n)).image (λ t, p ^ (t + 1)) ⊆ (n.divisors \ {1}) := begin rcases eq_or_ne n 0 with rfl | hn, @@ -537,42 +535,44 @@ begin rw [finset.mem_sdiff, nat.mem_divisors], refine ⟨⟨(pow_dvd_pow p $ by linarith).trans pow_padic_val_nat_dvd, hn⟩, _⟩, rw [finset.mem_singleton], - nth_rewrite 1 ←one_pow (k + 1), - exact (nat.pow_lt_pow_of_lt_left h.1.one_lt $ nat.succ_pos k).ne', + nth_rewrite 1 ← one_pow (k + 1), + exact (nat.pow_lt_pow_of_lt_left hp.1.one_lt $ nat.succ_pos k).ne' end end padic_val_nat section padic_val_int -variables (p : ℕ) [p_prime : fact p.prime] -lemma padic_val_int_dvd_iff (p : ℕ) [fact p.prime] (n : ℕ) (a : ℤ) : - ↑p^n ∣ a ↔ a = 0 ∨ n ≤ padic_val_int p a := -by rw [padic_val_int, ←int.nat_abs_eq_zero, ←padic_val_nat_dvd_iff, ←int.coe_nat_dvd_left, - int.coe_nat_pow] +variables {p : ℕ} [hp : fact p.prime] + +include hp + +lemma padic_val_int_dvd_iff (n : ℕ) (a : ℤ) : (p : ℤ) ^ n ∣ a ↔ a = 0 ∨ n ≤ padic_val_int p a := +by rw [padic_val_int, ← int.nat_abs_eq_zero, ← padic_val_nat_dvd_iff, ← int.coe_nat_dvd_left, + int.coe_nat_pow] -lemma padic_val_int_dvd (p : ℕ) [fact p.prime] (a : ℤ) : ↑p^(padic_val_int p a) ∣ a := +lemma padic_val_int_dvd (a : ℤ) : (p : ℤ) ^ padic_val_int p a ∣ a := begin rw padic_val_int_dvd_iff, - exact or.inr le_rfl, + exact or.inr le_rfl end -lemma padic_val_int_self (p : ℕ) [pp : fact p.prime] : padic_val_int p p = 1 := -padic_val_int.self pp.out.one_lt +lemma padic_val_int_self : padic_val_int p p = 1 := padic_val_int.self hp.out.one_lt -lemma padic_val_int.mul (p : ℕ) [fact p.prime] {a b : ℤ} (ha : a ≠ 0) (hb : b ≠ 0) : +lemma padic_val_int.mul {a b : ℤ} (ha : a ≠ 0) (hb : b ≠ 0) : padic_val_int p (a*b) = padic_val_int p a + padic_val_int p b := begin simp_rw padic_val_int, rw [int.nat_abs_mul, padic_val_nat.mul]; - rwa int.nat_abs_ne_zero, + rwa int.nat_abs_ne_zero end -lemma padic_val_int_mul_eq_succ (p : ℕ) [pp : fact p.prime] (a : ℤ) (ha : a ≠ 0) : +lemma padic_val_int_mul_eq_succ (a : ℤ) (ha : a ≠ 0) : padic_val_int p (a * p) = (padic_val_int p a) + 1 := begin - rw padic_val_int.mul p ha (int.coe_nat_ne_zero.mpr (pp.out).ne_zero), + rw padic_val_int.mul ha (int.coe_nat_ne_zero.mpr hp.out.ne_zero), simp only [eq_self_iff_true, padic_val_int.of_nat, padic_val_nat_self], + exact hp end end padic_val_int diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 63d26e79b4296..79c5b3ed040c4 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -162,8 +162,8 @@ begin have := prod_cyclotomic_eq_geom_sum hn' ℤ, apply_fun eval 1 at this, - rw [eval_geom_sum, one_geom_sum, eval_prod, eq_comm, - ←finset.prod_sdiff $ range_pow_padic_val_nat_subset_divisors' p, finset.prod_image] at this, + rw [eval_geom_sum, one_geom_sum, eval_prod, eq_comm, ←finset.prod_sdiff $ + @range_pow_padic_val_nat_subset_divisors' p _ _, finset.prod_image] at this, simp_rw [eval_one_cyclotomic_prime_pow, finset.prod_const, finset.card_range, mul_comm] at this, rw [←finset.prod_sdiff $ show {n} ⊆ _, from _] at this, any_goals {apply_instance}, From 8922aaa52081d9a96582af697934e0f7099951a7 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Fri, 16 Sep 2022 14:24:33 +0000 Subject: [PATCH 285/828] feat(number_theory/legendre_symbol/norm_num): add file with `norm_num` extension for the Jacobi symbol (#16519) This PR adds the file `number_theory.legendre_symbol.norm_num`, which provides an extension to the `norm_num` tactic, so that it can compute the value of a Jacobi (or Legendre) symbol when the entries are concrete numbers. See this [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Jacobi.20symbol/near/298768851) for discussion. --- .../legendre_symbol/jacobi_symbol.lean | 2 +- .../legendre_symbol/norm_num.lean | 440 ++++++++++++++++++ test/norm_num_ext.lean | 23 + 3 files changed, 464 insertions(+), 1 deletion(-) create mode 100644 src/number_theory/legendre_symbol/norm_num.lean diff --git a/src/number_theory/legendre_symbol/jacobi_symbol.lean b/src/number_theory/legendre_symbol/jacobi_symbol.lean index 15ef4760319c1..9694294930a3f 100644 --- a/src/number_theory/legendre_symbol/jacobi_symbol.lean +++ b/src/number_theory/legendre_symbol/jacobi_symbol.lean @@ -61,7 +61,7 @@ section jacobi ### Definition of the Jacobi symbol We define the Jacobi symbol $\Bigl(\frac{a}{b}\Bigr)$ for integers `a` and natural numbers `b` -as the product of the Legendre symbols $(\Bigl\frac{a}{p}\Bigr)$, where `p` runs through the +as the product of the Legendre symbols $\Bigl(\frac{a}{p}\Bigr)$, where `p` runs through the prime divisors (with multiplicity) of `b`, as provided by `b.factors`. This agrees with the Jacobi symbol when `b` is odd and gives less meaningful values when it is not (e.g., the symbol is `1` when `b = 0`). This is called `jacobi_sym a b`. diff --git a/src/number_theory/legendre_symbol/norm_num.lean b/src/number_theory/legendre_symbol/norm_num.lean new file mode 100644 index 0000000000000..4d77458655448 --- /dev/null +++ b/src/number_theory/legendre_symbol/norm_num.lean @@ -0,0 +1,440 @@ +/- +Copyright (c) 2022 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import number_theory.legendre_symbol.jacobi_symbol + +/-! +# A `norm_num` extension for Jacobi and Legendre symbols + +We extend the `tactic.interactive.norm_num` tactic so that it can be used to provably compute +the value of the Jacobi symbol `J(a | b)` or the Legendre symbol `legendre_sym p a` when +the arguments are numerals. + +## Implementation notes + +We use the Law of Quadratic Reciprocity for the Jacobi symbol to compute the value of `J(a | b)` +efficiently, roughly comparable in effort with the euclidean algorithm for the computation +of the gcd of `a` and `b`. More precisely, the computation is done in the following steps. + +* Use `J(a | 0) = 1` (an artifact of the definition) and `J(a | 1) = 1` to deal + with corner cases. + +* Use `J(a | b) = J(a % b | b)` to reduce to the case that `a` is a natural number. + We define a version of the Jacobi symbol restricted to natural numbers for use in + the following steps; see `norm_num.jacobi_sym_nat`. (But we'll continue to write `J(a | b)` + in this description.) + +* Remove powers of two from `b`. This is done via `J(2a | 2b) = 0` and + `J(2a+1 | 2b) = J(2a+1 | b)` (another artifact of the definition). + +* Now `0 ≤ a < b` and `b` is odd. If `b = 1`, then the value is `1`. + If `a = 0` (and `b > 1`), then the value is `0`. Otherwise, we remove powers of two from `a` + via `J(4a | b) = J(a | b)` and `J(2a | b) = ±J(a | b)`, where the sign is determined + by the residue class of `b` mod 8, to reduce to `a` odd. + +* Once `a` is odd, we use Quadratic Reciprocity (QR) in the form + `J(a | b) = ±J(b % a | a)`, where the sign is determined by the residue classes + of `a` and `b` mod 4. We are then back in the previous case. + +We provide customized versions of these results for the various reduction steps, +where we encode the residue classes mod 2, mod 4, or mod 8 by using terms like +`bit1 (bit0 a)`. In this way, the only divisions we have to compute and prove +are the ones occurring in the use of QR above. +-/ + +section lemmas + +namespace norm_num + +/-- The Jacobi symbol restricted to natural numbers in both arguments. -/ +def jacobi_sym_nat (a b : ℕ) : ℤ := jacobi_sym a b + +/-! +### API Lemmas + +We repeat part of the API for `jacobi_sym` with `norm_num.jacobi_sym_nat` and without implicit +arguments, in a form that is suitable for constructing proofs in `norm_num`. +-/ + +/-- Base cases: `b = 0`, `b = 1`, `a = 0`, `a = 1`. -/ +lemma jacobi_sym_nat.zero_right (a : ℕ) : jacobi_sym_nat a 0 = 1 := +by rwa [jacobi_sym_nat, jacobi_sym.zero_right] + +lemma jacobi_sym_nat.one_right (a : ℕ) : jacobi_sym_nat a 1 = 1 := +by rwa [jacobi_sym_nat, jacobi_sym.one_right] + +lemma jacobi_sym_nat.zero_left_even (b : ℕ) (hb : b ≠ 0) : jacobi_sym_nat 0 (bit0 b) = 0 := +by rw [jacobi_sym_nat, nat.cast_zero, jacobi_sym.zero_left (nat.one_lt_bit0 hb)] + +lemma jacobi_sym_nat.zero_left_odd (b : ℕ) (hb : b ≠ 0) : jacobi_sym_nat 0 (bit1 b) = 0 := +by rw [jacobi_sym_nat, nat.cast_zero, jacobi_sym.zero_left (nat.one_lt_bit1 hb)] + +lemma jacobi_sym_nat.one_left_even (b : ℕ) : jacobi_sym_nat 1 (bit0 b) = 1 := +by rw [jacobi_sym_nat, nat.cast_one, jacobi_sym.one_left] + +lemma jacobi_sym_nat.one_left_odd (b : ℕ) : jacobi_sym_nat 1 (bit1 b) = 1 := +by rw [jacobi_sym_nat, nat.cast_one, jacobi_sym.one_left] + +/-- Turn a Legendre symbol into a Jacobi symbol. -/ +lemma legendre_sym.to_jacobi_sym (p : ℕ) (pp : fact (p.prime)) (a r : ℤ) (hr : jacobi_sym a p = r) : + legendre_sym p a = r := +by rwa [@legendre_sym.to_jacobi_sym p pp a] + +/-- The value depends only on the residue class of `a` mod `b`. -/ +lemma jacobi_sym.mod_left (a : ℤ) (b ab' : ℕ) (ab r b' : ℤ) (hb' : (b : ℤ) = b') + (hab : a % b' = ab) (h : (ab' : ℤ) = ab) (hr : jacobi_sym_nat ab' b = r) : + jacobi_sym a b = r := +by rw [← hr, jacobi_sym_nat, jacobi_sym.mod_left, hb', hab, ← h] + +lemma jacobi_sym_nat.mod_left (a b ab : ℕ) (r : ℤ) (hab : a % b = ab) + (hr : jacobi_sym_nat ab b = r) : + jacobi_sym_nat a b = r := +by { rw [← hr, jacobi_sym_nat, jacobi_sym_nat, _root_.jacobi_sym.mod_left a b, ← hab], refl, } + +/-- The symbol vanishes when both entries are even (and `b ≠ 0`). -/ +lemma jacobi_sym_nat.even_even (a b : ℕ) (hb₀ : b ≠ 0) : + jacobi_sym_nat (bit0 a) (bit0 b) = 0 := +begin + refine jacobi_sym.eq_zero_iff.mpr ⟨nat.bit0_ne_zero hb₀, λ hf, _⟩, + have h : 2 ∣ (bit0 a).gcd (bit0 b) := nat.dvd_gcd two_dvd_bit0 two_dvd_bit0, + change 2 ∣ (bit0 a : ℤ).gcd (bit0 b) at h, + rw [← nat.cast_bit0, ← nat.cast_bit0, hf, ← even_iff_two_dvd] at h, + exact nat.not_even_one h, +end + +/-- When `a` is odd and `b` is even, we can replace `b` by `b / 2`. -/ +lemma jacobi_sym_nat.odd_even (a b : ℕ) (r : ℤ) (hr : jacobi_sym_nat (bit1 a) b = r) : + jacobi_sym_nat (bit1 a) (bit0 b) = r := +begin + have ha : legendre_sym 2 (bit1 a) = 1 := + by simp only [legendre_sym, quadratic_char_apply, quadratic_char_fun_one, int.cast_bit1, + char_two.bit1_eq_one, pi.one_apply], + cases eq_or_ne b 0 with hb hb, + { rw [← hr, hb, jacobi_sym_nat.zero_right], }, + { haveI : ne_zero b := ⟨hb⟩, -- for `jacobi_sym.mul_right` + rwa [bit0_eq_two_mul b, jacobi_sym_nat, jacobi_sym.mul_right, + ← _root_.legendre_sym.to_jacobi_sym, nat.cast_bit1, ha, one_mul], } +end + +/-- If `a` is divisible by `4` and `b` is odd, then we can remove the factor `4` from `a`. -/ +lemma jacobi_sym_nat.double_even (a b : ℕ) (r : ℤ) (hr : jacobi_sym_nat a (bit1 b) = r) : + jacobi_sym_nat (bit0 (bit0 a)) (bit1 b) = r := +begin + have : ((2 : ℕ) : ℤ).gcd ((bit1 b) : ℕ) = 1, + { rw [int.coe_nat_gcd, nat.bit1_eq_succ_bit0, bit0_eq_two_mul b, nat.succ_eq_add_one, + nat.gcd_mul_left_add_right, nat.gcd_one_right], }, + rwa [bit0_eq_two_mul a, bit0_eq_two_mul (2 * a), ← mul_assoc, ← pow_two, jacobi_sym_nat, + nat.cast_mul, nat.cast_pow, jacobi_sym.mul_left, jacobi_sym.sq_one' this, one_mul], +end + +/-- If `a` is even and `b` is odd, then we can remove a factor `2` from `a`, +but we may have to change the sign, depending on `b % 8`. +We give one version for each of the four odd residue classes mod `8`. -/ +lemma jacobi_sym_nat.even_odd₁ (a b : ℕ) (r : ℤ) + (hr : jacobi_sym_nat a (bit1 (bit0 (bit0 b))) = r) : + jacobi_sym_nat (bit0 a) (bit1 (bit0 (bit0 b))) = r := +begin + have hb : (bit1 (bit0 (bit0 b))) % 8 = 1, + { rw [nat.bit1_mod_bit0, nat.bit0_mod_bit0, nat.bit0_mod_two], }, + rw [jacobi_sym_nat, bit0_eq_two_mul a, nat.cast_mul, jacobi_sym.mul_left, + nat.cast_two, jacobi_sym.at_two (odd_bit1 _), zmod.χ₈_nat_mod_eight, hb], + norm_num, + exact hr, +end + +lemma jacobi_sym_nat.even_odd₇ (a b : ℕ) (r : ℤ) + (hr : jacobi_sym_nat a (bit1 (bit1 (bit1 b))) = r) : + jacobi_sym_nat (bit0 a) (bit1 (bit1 (bit1 b))) = r := +begin + have hb : (bit1 (bit1 (bit1 b))) % 8 = 7, + { rw [nat.bit1_mod_bit0, nat.bit1_mod_bit0, nat.bit1_mod_two], }, + rw [jacobi_sym_nat, bit0_eq_two_mul a, nat.cast_mul, jacobi_sym.mul_left, + nat.cast_two, jacobi_sym.at_two (odd_bit1 _), zmod.χ₈_nat_mod_eight, hb], + norm_num, + exact hr, +end + +lemma jacobi_sym_nat.even_odd₃ (a b : ℕ) (r : ℤ) + (hr : jacobi_sym_nat a (bit1 (bit1 (bit0 b))) = r) : + jacobi_sym_nat (bit0 a) (bit1 (bit1 (bit0 b))) = -r := +begin + have hb : (bit1 (bit1 (bit0 b))) % 8 = 3, + { rw [nat.bit1_mod_bit0, nat.bit1_mod_bit0, nat.bit0_mod_two], }, + rw [jacobi_sym_nat, bit0_eq_two_mul a, nat.cast_mul, jacobi_sym.mul_left, + nat.cast_two, jacobi_sym.at_two (odd_bit1 _), zmod.χ₈_nat_mod_eight, hb], + norm_num, + exact hr, +end + +lemma jacobi_sym_nat.even_odd₅ (a b : ℕ) (r : ℤ) + (hr : jacobi_sym_nat a (bit1 (bit0 (bit1 b))) = r) : + jacobi_sym_nat (bit0 a) (bit1 (bit0 (bit1 b))) = -r := +begin + have hb : (bit1 (bit0 (bit1 b))) % 8 = 5, + { rw [nat.bit1_mod_bit0, nat.bit0_mod_bit0, nat.bit1_mod_two], }, + rw [jacobi_sym_nat, bit0_eq_two_mul a, nat.cast_mul, jacobi_sym.mul_left, + nat.cast_two, jacobi_sym.at_two (odd_bit1 _), zmod.χ₈_nat_mod_eight, hb], + norm_num, + exact hr, +end + +/-- Use quadratic reciproity to reduce to smaller `b`. -/ +lemma jacobi_sym_nat.qr₁ (a b : ℕ) (r : ℤ) (hr : jacobi_sym_nat (bit1 b) (bit1 (bit0 a)) = r) : + jacobi_sym_nat (bit1 (bit0 a)) (bit1 b) = r := +begin + have ha : (bit1 (bit0 a)) % 4 = 1, + { rw [nat.bit1_mod_bit0, nat.bit0_mod_two], }, + have hb := nat.bit1_mod_two, + rwa [jacobi_sym_nat, jacobi_sym.quadratic_reciprocity_one_mod_four ha (nat.odd_iff.mpr hb)], +end + +lemma jacobi_sym_nat.qr₁_mod (a b ab : ℕ) (r : ℤ) (hab : (bit1 b) % (bit1 (bit0 a)) = ab) + (hr : jacobi_sym_nat ab (bit1 (bit0 a)) = r) : + jacobi_sym_nat (bit1 (bit0 a)) (bit1 b) = r := +jacobi_sym_nat.qr₁ _ _ _ $ jacobi_sym_nat.mod_left _ _ ab r hab hr + +lemma jacobi_sym_nat.qr₁' (a b : ℕ) (r : ℤ) (hr : jacobi_sym_nat (bit1 (bit0 b)) (bit1 a) = r) : + jacobi_sym_nat (bit1 a) (bit1 (bit0 b)) = r := +begin + have hb : (bit1 (bit0 b)) % 4 = 1, + { rw [nat.bit1_mod_bit0, nat.bit0_mod_two], }, + have ha := nat.bit1_mod_two, + rwa [jacobi_sym_nat, ← jacobi_sym.quadratic_reciprocity_one_mod_four hb (nat.odd_iff.mpr ha)] +end + +lemma jacobi_sym_nat.qr₁'_mod (a b ab : ℕ) (r : ℤ) (hab : (bit1 (bit0 b)) % (bit1 a) = ab) + (hr : jacobi_sym_nat ab (bit1 a) = r) : + jacobi_sym_nat (bit1 a) (bit1 (bit0 b)) = r := +jacobi_sym_nat.qr₁' _ _ _ $ jacobi_sym_nat.mod_left _ _ ab r hab hr + +lemma jacobi_sym_nat.qr₃ (a b : ℕ) (r : ℤ) + (hr : jacobi_sym_nat (bit1 (bit1 b)) (bit1 (bit1 a)) = r) : + jacobi_sym_nat (bit1 (bit1 a)) (bit1 (bit1 b)) = -r := +begin + have hb : (bit1 (bit1 b)) % 4 = 3, + { rw [nat.bit1_mod_bit0, nat.bit1_mod_two], }, + have ha : (bit1 (bit1 a)) % 4 = 3, + { rw [nat.bit1_mod_bit0, nat.bit1_mod_two], }, + rwa [jacobi_sym_nat, jacobi_sym.quadratic_reciprocity_three_mod_four ha hb, neg_inj] +end + +lemma jacobi_sym_nat.qr₃_mod (a b ab : ℕ) (r : ℤ) (hab : (bit1 (bit1 b)) % (bit1 (bit1 a)) = ab) + (hr : jacobi_sym_nat ab (bit1 (bit1 a)) = r) : + jacobi_sym_nat (bit1 (bit1 a)) (bit1 (bit1 b)) = -r := +jacobi_sym_nat.qr₃ _ _ _ $ jacobi_sym_nat.mod_left _ _ ab r hab hr + +end norm_num + +end lemmas + +section evaluation + +-- The following is to prevent strange error messages from occurring. +instance : ring ℚ := division_ring.to_ring ℚ + +/-! +### Certified evaluation of the Jacobi symbol + +The following functions recursively evaluate a Jacobi symbol and construct the +corresponding proof term. +-/ + +namespace norm_num +open tactic + +/-- This evaluates `r := jacobi_sym_nat a b` recursively using quadratic reciprocity +and produces a proof term for the equality, assuming that `a < b` and `b` is odd. -/ +meta def prove_jacobi_sym_odd : instance_cache → instance_cache → expr → expr → + tactic (instance_cache × instance_cache × expr × expr) +| zc nc ea eb := do + match match_numeral eb with + | match_numeral_result.one := -- `b = 1`, result is `1` + pure (zc, nc, `(1 : ℤ), `(jacobi_sym_nat.one_right).mk_app [ea]) + | match_numeral_result.bit1 eb₁ := do -- `b > 1` (recall that `b` is odd) + match match_numeral ea with + | match_numeral_result.zero := do -- `a = 0`, result is `0` + b ← eb₁.to_nat, + (nc, phb₀) ← prove_ne nc eb₁ `(0 : ℕ) b 0, -- proof of `b ≠ 0` + pure (zc, nc, `(0 : ℤ), `(jacobi_sym_nat.zero_left_odd).mk_app [eb₁, phb₀]) + | match_numeral_result.one := do -- `a = 1`, result is `1` + pure (zc, nc, `(1 : ℤ), `(jacobi_sym_nat.one_left_odd).mk_app [eb₁]) + | match_numeral_result.bit0 ea₁ := do -- `a` is even; check if divisible by `4` + match match_numeral ea₁ with + | match_numeral_result.bit0 ea₂ := do + (zc, nc, er, p) ← prove_jacobi_sym_odd zc nc ea₂ eb, -- compute `jacobi_sym_nat (a / 4) b` + pure (zc, nc, er, `(jacobi_sym_nat.double_even).mk_app [ea₂, eb₁, er, p]) + | _ := do -- reduce to `a / 2`; need to consider `b % 8` + (zc, nc, er, p) ← prove_jacobi_sym_odd zc nc ea₁ eb, -- compute `jacobi_sym_nat (a / 2) b` + match match_numeral eb₁ with + -- | match_numeral_result.zero := -- `b = 1`, not reached + | match_numeral_result.one := do -- `b = 3` + r ← er.to_int, + (zc, er') ← zc.of_int (- r), + pure (zc, nc, er', `(jacobi_sym_nat.even_odd₃).mk_app [ea₁, `(0 : ℕ), er, p]) + | match_numeral_result.bit0 eb₂ := do -- `b % 4 = 1` + match match_numeral eb₂ with + -- | match_numeral_result.zero := -- not reached + | match_numeral_result.one := do -- `b = 5` + r ← er.to_int, + (zc, er') ← zc.of_int (- r), + pure (zc, nc, er', `(jacobi_sym_nat.even_odd₅).mk_app [ea₁, `(0 : ℕ), er, p]) + | match_numeral_result.bit0 eb₃ := do -- `b % 8 = 1` + pure (zc, nc, er, `(jacobi_sym_nat.even_odd₁).mk_app [ea₁, eb₃, er, p]) + | match_numeral_result.bit1 eb₃ := do -- `b % 8 = 5` + r ← er.to_int, + (zc, er') ← zc.of_int (- r), + pure (zc, nc, er', `(jacobi_sym_nat.even_odd₅).mk_app [ea₁, eb₃, er, p]) + | _ := failed + end + | match_numeral_result.bit1 eb₂ := do -- `b % 4 = 3` + match match_numeral eb₂ with + -- | match_numeral_result.zero := -- not reached + | match_numeral_result.one := do -- `b = 7` + pure (zc, nc, er, `(jacobi_sym_nat.even_odd₇).mk_app [ea₁, `(0 : ℕ), er, p]) + | match_numeral_result.bit0 eb₃ := do -- `b % 8 = 3` + r ← er.to_int, + (zc, er') ← zc.of_int (- r), + pure (zc, nc, er', `(jacobi_sym_nat.even_odd₃).mk_app [ea₁, eb₃, er, p]) + | match_numeral_result.bit1 eb₃ := do -- `b % 8 = 7` + pure (zc, nc, er, `(jacobi_sym_nat.even_odd₇).mk_app [ea₁, eb₃, er, p]) + | _ := failed + end + | _ := failed + end + end + | match_numeral_result.bit1 ea₁ := do -- `a` is odd + -- use Quadratic Reciprocity; look at `a` and `b` mod `4` + (nc, bma, phab) ← prove_div_mod nc eb ea tt, -- compute `b % a` + (zc, nc, er, p) ← prove_jacobi_sym_odd zc nc bma ea, -- compute `jacobi_sym_nat (b % a) a` + match match_numeral ea₁ with + -- | match_numeral_result.zero := -- `a = 1`, not reached + | match_numeral_result.one := do -- `a = 3`; need to consider `b` + match match_numeral eb₁ with + -- | match_numeral_result.zero := -- `b = 1`, not reached + -- | match_numeral_result.one := -- `b = 3`, not reached, since `a < b` + | match_numeral_result.bit0 eb₂ := do -- `b % 4 = 1` + pure (zc, nc, er, `(jacobi_sym_nat.qr₁'_mod).mk_app [ea₁, eb₂, bma, er, phab, p]) + | match_numeral_result.bit1 eb₂ := do -- `b % 4 = 3` + r ← er.to_int, + (zc, er') ← zc.of_int (- r), + pure (zc, nc, er', `(jacobi_sym_nat.qr₃_mod).mk_app [`(0 : ℕ), eb₂, bma, er, phab, p]) + | _ := failed + end + | match_numeral_result.bit0 ea₂ := do -- `a % 4 = 1` + pure (zc, nc, er, `(jacobi_sym_nat.qr₁_mod).mk_app [ea₂, eb₁, bma, er, phab, p]) + | match_numeral_result.bit1 ea₂ := do -- `a % 4 = 3`; need to consider `b` + match match_numeral eb₁ with + -- | match_numeral_result.zero := do -- `b = 1`, not reached + -- | match_numeral_result.one := do -- `b = 3`, not reached, since `a < b` + | match_numeral_result.bit0 eb₂ := do -- `b % 4 = 1` + pure (zc, nc, er, `(jacobi_sym_nat.qr₁'_mod).mk_app [ea₁, eb₂, bma, er, phab, p]) + | match_numeral_result.bit1 eb₂ := do -- `b % 4 = 3` + r ← er.to_int, + (zc, er') ← zc.of_int (- r), + pure (zc, nc, er', `(jacobi_sym_nat.qr₃_mod).mk_app [ea₂, eb₂, bma, er, phab, p]) + | _ := failed + end + | _ := failed + end + | _ := failed + end + | _ := failed + end + +/-- This evaluates `r := jacobi_sym_nat a b` and produces a proof term for the equality +by removing powers of `2` from `b` and then calling `prove_jacobi_sym_odd`. -/ +meta def prove_jacobi_sym_nat : instance_cache → instance_cache → expr → expr → + tactic (instance_cache × instance_cache × expr × expr) +| zc nc ea eb := do + match match_numeral eb with + | match_numeral_result.zero := -- `b = 0`, result is `1` + pure (zc, nc, `(1 : ℤ), `(jacobi_sym_nat.zero_right).mk_app [ea]) + | match_numeral_result.one := -- `b = 1`, result is `1` + pure (zc, nc, `(1 : ℤ), `(jacobi_sym_nat.one_right).mk_app [ea]) + | match_numeral_result.bit0 eb₁ := -- `b` is even and nonzero + match match_numeral ea with + | match_numeral_result.zero := do -- `a = 0`, result is `0` + b ← eb₁.to_nat, + (nc, phb₀) ← prove_ne nc eb₁ `(0 : ℕ) b 0, -- proof of `b ≠ 0` + pure (zc, nc, `(0 : ℤ), `(jacobi_sym_nat.zero_left_even).mk_app [eb₁, phb₀]) + | match_numeral_result.one := do -- `a = 1`, result is `1` + pure (zc, nc, `(1 : ℤ), `(jacobi_sym_nat.one_left_even).mk_app [eb₁]) + | match_numeral_result.bit0 ea₁ := do -- `a` is even, result is `0` + b ← eb₁.to_nat, + (nc, phb₀) ← prove_ne nc eb₁ `(0 : ℕ) b 0, -- proof of `b ≠ 0` + let er : expr := `(0 : ℤ), + pure (zc, nc, er, `(jacobi_sym_nat.even_even).mk_app [ea₁, eb₁, phb₀]) + | match_numeral_result.bit1 ea₁ := do -- `a` is odd, reduce to `b / 2` + (zc, nc, er, p) ← prove_jacobi_sym_nat zc nc ea eb₁, + pure (zc, nc, er, `(jacobi_sym_nat.odd_even).mk_app [ea₁, eb₁, er, p]) + | _ := failed + end + | match_numeral_result.bit1 eb₁ := do -- `b` is odd + a ← ea.to_nat, + b ← eb.to_nat, + if b ≤ a then do -- reduce to `jacobi_sym_nat (a % b) b` + (nc, amb, phab) ← prove_div_mod nc ea eb tt, -- compute `a % b` + (zc, nc, er, p) ← prove_jacobi_sym_odd zc nc amb eb, -- compute `jacobi_sym_nat (a % b) b` + pure (zc, nc, er, `(jacobi_sym_nat.mod_left).mk_app [ea, eb, amb, er, phab, p]) + else + prove_jacobi_sym_odd zc nc ea eb + | _ := failed + end + +/-- This evaluates `r := jacobi_sym a b` and produces a proof term for the equality. +This is done by reducing to `r := jacobi_sym_nat (a % b) b`. -/ +meta def prove_jacobi_sym : instance_cache → instance_cache → expr → expr + → tactic (instance_cache × instance_cache × expr × expr) +| zc nc ea eb := do + match match_numeral eb with -- deal with simple cases right away + | match_numeral_result.zero := pure (zc, nc, `(1 : ℤ), `(jacobi_sym.zero_right).mk_app [ea]) + | match_numeral_result.one := pure (zc, nc, `(1 : ℤ), `(jacobi_sym.one_right).mk_app [ea]) + | _ := do -- Now `1 < b`. Compute `jacobi_sym_nat (a % b) b` instead. + b ← eb.to_nat, + (zc, eb') ← zc.of_int (b : ℤ), + -- Get the proof that `(b : ℤ) = b'` (where `eb'` is the numeral representing `b'`). + -- This is important to avoid inefficient matching between the two. + (zc, nc, eb₁, pb') ← prove_nat_uncast zc nc eb', + (zc, amb, phab) ← prove_div_mod zc ea eb' tt, -- compute `a % b` + (zc, nc, amb', phab') ← prove_nat_uncast zc nc amb, -- `a % b` as a natural number + (zc, nc, er, p) ← prove_jacobi_sym_nat zc nc amb' eb₁, -- compute `jacobi_sym_nat (a % b) b` + pure (zc, nc, er, + `(jacobi_sym.mod_left).mk_app [ea, eb₁, amb', amb, er, eb', pb', phab, phab', p]) + end + +end norm_num + +end evaluation + +section tactic + +/-! +### The `norm_num` plug-in +-/ + +namespace tactic +namespace norm_num + +/-- This is the `norm_num` plug-in that evaluates Jacobi and Legendre symbols. -/ +@[norm_num] meta def eval_jacobi_sym : expr → tactic (expr × expr) +| `(jacobi_sym %%ea %%eb) := do -- Jacobi symbol + zc ← mk_instance_cache `(ℤ), + nc ← mk_instance_cache `(ℕ), + (prod.snd ∘ prod.snd) <$> norm_num.prove_jacobi_sym zc nc ea eb +| `(norm_num.jacobi_sym_nat %%ea %%eb) := do -- Jacobi symbol on natural numbers + zc ← mk_instance_cache `(ℤ), + nc ← mk_instance_cache `(ℕ), + (prod.snd ∘ prod.snd) <$> norm_num.prove_jacobi_sym_nat zc nc ea eb +| `(@legendre_sym %%ep %%inst %%ea) := do -- Legendre symbol + zc ← mk_instance_cache `(ℤ), + nc ← mk_instance_cache `(ℕ), + (zc, nc, er, pf) ← norm_num.prove_jacobi_sym zc nc ea ep, + pure (er, `(norm_num.legendre_sym.to_jacobi_sym).mk_app [ep, inst, ea, er, pf]) +| _ := failed + +end norm_num +end tactic + +end tactic diff --git a/test/norm_num_ext.lean b/test/norm_num_ext.lean index ef6d5a3d196fe..371b5db362d01 100644 --- a/test/norm_num_ext.lean +++ b/test/norm_num_ext.lean @@ -10,6 +10,7 @@ import data.nat.fib import data.nat.prime import data.nat.sqrt_norm_num import analysis.special_functions.pow +import number_theory.legendre_symbol.norm_num /-! # Tests for `norm_num` extensions @@ -301,3 +302,25 @@ example : ∏ i in {1, 4, 9, 16}, nat.sqrt i = 24 := by norm_num example : ∑ i : fin 2, ∑ j : fin 2, ![![0, 1], ![2, 3]] i j = 6 := by norm_num end big_operators + +section jacobi + +-- Jacobi and Legendre symbols + +open_locale number_theory_symbols + +example : J(123 | 335) = -1 := by norm_num +example : J(-2345 | 6789) = -1 := by norm_num +example : J(-1 | 1655801) = 1 := by norm_num +example : J(-102334155 | 165580141) = -1 := by norm_num + +example : J(58378362899022564339483801989973056405585914719065 | + 53974350278769849773003214636618718468638750007307) = -1 := by norm_num + +example : J(3 + 4 | 3 * 5) = -1 := by norm_num +example : J(J(-1 | 7) | 11) = -1 := by norm_num + +instance prime_1000003 : fact (nat.prime 1000003) := ⟨by norm_num⟩ +example : legendre_sym 1000003 7 = -1 := by norm_num + +end jacobi From 7a9e6ca4026238cd69531b05a919c7d6479375c8 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 16 Sep 2022 15:31:11 +0000 Subject: [PATCH 286/828] doc(category/abelian/ext): update module doc (#16512) Update future work described in module doc, to reflect previous contributions to mathlib. Co-authored-by: Scott Morrison --- src/category_theory/abelian/ext.lean | 9 +++------ 1 file changed, 3 insertions(+), 6 deletions(-) diff --git a/src/category_theory/abelian/ext.lean b/src/category_theory/abelian/ext.lean index e3a47ba8fcfcd..ee8dac92b68e7 100644 --- a/src/category_theory/abelian/ext.lean +++ b/src/category_theory/abelian/ext.lean @@ -14,7 +14,7 @@ import category_theory.abelian.projective # Ext We define `Ext R C n : Cᵒᵖ ⥤ C ⥤ Module R` for any `R`-linear abelian category `C` -by deriving in the first argument of the bifunctor `(X, Y) ↦ Module.of R (unop X ⟶ Y)`. +by (left) deriving in the first argument of the bifunctor `(X, Y) ↦ Module.of R (unop X ⟶ Y)`. ## Implementation @@ -22,11 +22,8 @@ It's not actually necessary here to assume `C` is abelian, but the hypotheses, involving both `C` and `Cᵒᵖ`, are quite lengthy, and in practice the abelian case is hopefully enough. -PROJECT we don't yet have injective resolutions and right derived functors -(although this is only a copy-and-paste dualisation) -so we can't even state the alternative definition -in terms of right-deriving in the first argument, -let alone start the harder project of showing they agree. +PROJECT: State the alternative definition in terms of +right deriving in the second argument, and show these agree. -/ noncomputable theory From e763499b99499a5afadd339f73895d27de691fd0 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Fri, 16 Sep 2022 19:44:02 +0000 Subject: [PATCH 287/828] chore(algebra/order/monoid_lemmas_zero_lt): remove redundant assumptions (#16524) The third part of #16449 --- src/algebra/order/monoid_lemmas_zero_lt.lean | 18 +++++++----------- 1 file changed, 7 insertions(+), 11 deletions(-) diff --git a/src/algebra/order/monoid_lemmas_zero_lt.lean b/src/algebra/order/monoid_lemmas_zero_lt.lean index f0d03dda5612d..78c077cb399e3 100644 --- a/src/algebra/order/monoid_lemmas_zero_lt.lean +++ b/src/algebra/order/monoid_lemmas_zero_lt.lean @@ -232,12 +232,10 @@ lemma pos_mul_strict_mono_iff_pos_mul_mono_rev : pos_mul_strict_mono α ↔ pos_ lemma mul_pos_strict_mono_iff_mul_pos_mono_rev : mul_pos_strict_mono α ↔ mul_pos_mono_rev α := ⟨@mul_pos_strict_mono.to_mul_pos_mono_rev _ _ _ _, @mul_pos_mono_rev.to_mul_pos_strict_mono _ _ _ _⟩ -lemma pos_mul_reflect_lt.to_pos_mul_mono {α : Type*} [mul_zero_class α] [linear_order α] - [pos_mul_reflect_lt α] : pos_mul_mono α := +lemma pos_mul_reflect_lt.to_pos_mul_mono [pos_mul_reflect_lt α] : pos_mul_mono α := ⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt $ lt_of_mul_lt_mul_left h' x.prop⟩ -lemma mul_pos_reflect_lt.to_mul_pos_mono {α : Type*} [mul_zero_class α] [linear_order α] - [mul_pos_reflect_lt α] : mul_pos_mono α := +lemma mul_pos_reflect_lt.to_mul_pos_mono [mul_pos_reflect_lt α] : mul_pos_mono α := ⟨λ x a b h, le_of_not_lt $ λ h', h.not_lt $ lt_of_mul_lt_mul_right h' x.prop⟩ lemma pos_mul_mono.to_pos_mul_reflect_lt [pos_mul_mono α] : pos_mul_reflect_lt α := @@ -246,13 +244,11 @@ lemma pos_mul_mono.to_pos_mul_reflect_lt [pos_mul_mono α] : pos_mul_reflect_lt lemma mul_pos_mono.to_mul_pos_reflect_lt [mul_pos_mono α] : mul_pos_reflect_lt α := ⟨λ x a b h, lt_of_not_le $ λ h', h.not_le $ mul_le_mul_of_nonneg_right h' x.prop⟩ -lemma pos_mul_mono_iff_pos_mul_reflect_lt {α : Type*} [mul_zero_class α] [linear_order α] : - pos_mul_mono α ↔ pos_mul_reflect_lt α := -⟨@pos_mul_mono.to_pos_mul_reflect_lt _ _ _ _, @pos_mul_reflect_lt.to_pos_mul_mono _ _ _⟩ +lemma pos_mul_mono_iff_pos_mul_reflect_lt : pos_mul_mono α ↔ pos_mul_reflect_lt α := +⟨@pos_mul_mono.to_pos_mul_reflect_lt _ _ _ _, @pos_mul_reflect_lt.to_pos_mul_mono _ _ _ _⟩ -lemma mul_pos_mono_iff_mul_pos_reflect_lt {α : Type*} [mul_zero_class α] [linear_order α] : - mul_pos_mono α ↔ mul_pos_reflect_lt α := -⟨@mul_pos_mono.to_mul_pos_reflect_lt _ _ _ _, @mul_pos_reflect_lt.to_mul_pos_mono _ _ _⟩ +lemma mul_pos_mono_iff_mul_pos_reflect_lt : mul_pos_mono α ↔ mul_pos_reflect_lt α := +⟨@mul_pos_mono.to_mul_pos_reflect_lt _ _ _ _, @mul_pos_reflect_lt.to_mul_pos_mono _ _ _ _⟩ end linear_order @@ -268,7 +264,7 @@ variables [preorder α] lemma left.mul_pos [pos_mul_strict_mono α] (ha : 0 < a) (hb : 0 < b) : 0 < a * b := by simpa only [mul_zero] using mul_lt_mul_of_pos_left hb ha -alias left.mul_pos ← _root_.mul_pos +alias left.mul_pos ← mul_pos lemma mul_neg_of_pos_of_neg [pos_mul_strict_mono α] (ha : 0 < a) (hb : b < 0) : a * b < 0 := by simpa only [mul_zero] using mul_lt_mul_of_pos_left hb ha From 1e225257dc98379df3360baeea6cb678b65334fe Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 16 Sep 2022 22:29:26 +0000 Subject: [PATCH 288/828] feat(number_theory): Dedekind-Kummer theorem (#15000) It's PR #15000! The Kummer-Dedekind theorem on the splitting of ideals in monogenic ring extensions. [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/144837-PR-reviews/topic/.2315000.3A.20the.20Kummer-Dedekind.20theorem) Co-Authored-By: Paul Lezeau Co-authored-by: Paul-Lez <72892199+Paul-Lez@users.noreply.github.com> Co-authored-by: Johan Commelin Co-authored-by: Anne Baanen --- src/data/multiset/basic.lean | 10 ++ src/number_theory/kummer_dedekind.lean | 168 +++++++++++++++++++++++++ src/ring_theory/ideal/operations.lean | 6 + 3 files changed, 184 insertions(+) create mode 100644 src/number_theory/kummer_dedekind.lean diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index e6197b00ab1c8..515aed2fffe64 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -1104,6 +1104,9 @@ lemma attach_cons (a : α) (m : multiset α) : quotient.induction_on m $ assume l, congr_arg coe $ congr_arg (list.cons _) $ by rw [list.map_pmap]; exact list.pmap_congr _ (λ _ _ _ _, subtype.eq rfl) +@[simp] +lemma attach_map_coe (m : multiset α) : multiset.map (coe : _ → α) m.attach = m := m.attach_map_val + section decidable_pi_exists variables {m : multiset α} @@ -1866,6 +1869,13 @@ begin contradiction } end +@[simp] +lemma attach_count_eq_count_coe (m : multiset α) (a) : m.attach.count a = m.count (a : α) := +calc m.attach.count a + = (m.attach.map (coe : _ → α)).count (a : α) : + (multiset.count_map_eq_count' _ _ subtype.coe_injective _).symm +... = m.count (a : α) : congr_arg _ m.attach_map_coe + lemma filter_eq' (s : multiset α) (b : α) : s.filter (= b) = repeat b (count b s) := begin ext a, diff --git a/src/number_theory/kummer_dedekind.lean b/src/number_theory/kummer_dedekind.lean new file mode 100644 index 0000000000000..789eb779ba264 --- /dev/null +++ b/src/number_theory/kummer_dedekind.lean @@ -0,0 +1,168 @@ +/- +Copyright (c) 2021 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Paul Lezeau +-/ + +import ring_theory.adjoin_root +import ring_theory.dedekind_domain.ideal +import ring_theory.algebra_tower + +/-! +# Kummer-Dedekind theorem + +This file proves the monogenic version of the Kummer-Dedekind theorem on the splitting of prime +ideals in an extension of the ring of integers. This states that if `I` is a prime ideal of +Dedekind domain `R` and `S = R[α]` for some `α` that is integral over `R` with minimal polynomial +`f`, then the prime factorisations of `I * S` and `f mod I` have the same shape, i.e. they have the +same number of prime factors, and each prime factors of `I * S` can be paired with a prime factor +of `f mod I` in a way that ensures multiplicities match (in fact, this pairing can be made explicit +with a formula). + +## Main definitions + + * `normalized_factors_map_equiv_normalized_factors_min_poly_mk` : The bijection in the + Kummer-Dedekind theorem. This is the pairing between the prime factors of `I * S` and the prime + factors of `f mod I`. + +## Main results + + * `normalized_factors_ideal_map_eq_normalized_factors_min_poly_mk_map` : The Kummer-Dedekind + theorem. + * `ideal.irreducible_map_of_irreducible_minpoly` : `I.map (algebra_map R S)` is irreducible if + `(map I^.quotient.mk (minpoly R pb.gen))` is irreducible, where `pb` is a power basis of `S` + over `R`. + +## TODO + + * Define the conductor ideal and prove the Kummer-Dedekind theorem in full generality. + + * Prove the converse of `ideal.irreducible_map_of_irreducible_minpoly`. + + * Prove that `normalized_factors_map_equiv_normalized_factors_min_poly_mk` can be expressed as + `normalized_factors_map_equiv_normalized_factors_min_poly_mk g = ⟨I, G(α)⟩` for `g` a prime + factor of `f mod I` and `G` a lift of `g` to `R[X]`. + +## References + + * [J. Neukirch, *Algebraic Number Theory*][Neukirch1992] + +## Tags + +kummer, dedekind, kummer dedekind, dedekind-kummer, dedekind kummer +-/ + +namespace kummer_dedekind + +open_locale big_operators polynomial classical + +open ideal polynomial double_quot unique_factorization_monoid + +variables {R : Type*} [comm_ring R] +variables {S : Type*} [comm_ring S] [is_domain S] [is_dedekind_domain S] [algebra R S] +variables (pb : power_basis R S) {I : ideal R} + +local attribute [instance] ideal.quotient.field + +variables [is_domain R] + +/-- The first half of the **Kummer-Dedekind Theorem** in the monogenic case, stating that the prime + factors of `I*S` are in bijection with those of the minimal polynomial of the generator of `S` + over `R`, taken `mod I`.-/ +noncomputable def normalized_factors_map_equiv_normalized_factors_min_poly_mk (hI : is_maximal I) + (hI' : I ≠ ⊥) : {J : ideal S | J ∈ normalized_factors (I.map (algebra_map R S) )} ≃ + {d : (R ⧸ I)[X] | d ∈ normalized_factors (map I^.quotient.mk (minpoly R pb.gen)) } := +((normalized_factors_equiv_of_quot_equiv ↑(pb.quotient_equiv_quotient_minpoly_map I) + --show that `I * S` ≠ ⊥ + (show I.map (algebra_map R S) ≠ ⊥, + by rwa [ne.def, map_eq_bot_iff_of_injective pb.basis.algebra_map_injective, ← ne.def]) + --show that the ideal spanned by `(minpoly R pb.gen) mod I` is non-zero + (by {by_contra, exact (show (map I^.quotient.mk (minpoly R pb.gen) ≠ 0), from + polynomial.map_monic_ne_zero (minpoly.monic pb.is_integral_gen)) + (span_singleton_eq_bot.mp h) } )).trans + (normalized_factors_equiv_span_normalized_factors + (show (map I^.quotient.mk (minpoly R pb.gen)) ≠ 0, from + polynomial.map_monic_ne_zero (minpoly.monic pb.is_integral_gen))).symm) + +/-- The second half of the **Kummer-Dedekind Theorem** in the monogenic case, stating that the + bijection `factors_equiv'` defined in the first half preserves multiplicities. -/ +theorem multiplicity_factors_map_eq_multiplicity (hI : is_maximal I) (hI' : I ≠ ⊥) {J : ideal S} + (hJ : J ∈ normalized_factors (I.map (algebra_map R S))) : + multiplicity J (I.map (algebra_map R S)) = + multiplicity ↑(normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI' ⟨J, hJ⟩) + (map I^.quotient.mk (minpoly R pb.gen)) := +by rw [normalized_factors_map_equiv_normalized_factors_min_poly_mk, equiv.coe_trans, + function.comp_app, + multiplicity_normalized_factors_equiv_span_normalized_factors_symm_eq_multiplicity, + normalized_factors_equiv_of_quot_equiv_multiplicity_eq_multiplicity] + +/-- The **Kummer-Dedekind Theorem**. -/ +theorem normalized_factors_ideal_map_eq_normalized_factors_min_poly_mk_map (hI : is_maximal I) + (hI' : I ≠ ⊥) : normalized_factors (I.map (algebra_map R S)) = multiset.map + (λ f, ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm f : ideal S)) + (normalized_factors (polynomial.map I^.quotient.mk (minpoly R pb.gen))).attach := +begin + ext J, + -- WLOG, assume J is a normalized factor + by_cases hJ : J ∈ normalized_factors (I.map (algebra_map R S)), swap, + { rw [multiset.count_eq_zero.mpr hJ, eq_comm, multiset.count_eq_zero, multiset.mem_map], + simp only [multiset.mem_attach, true_and, not_exists], + rintros J' rfl, + exact hJ + ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm J').prop }, + + -- Then we just have to compare the multiplicities, which we already proved are equal. + have := multiplicity_factors_map_eq_multiplicity pb hI hI' hJ, + rw [multiplicity_eq_count_normalized_factors, multiplicity_eq_count_normalized_factors, + unique_factorization_monoid.normalize_normalized_factor _ hJ, + unique_factorization_monoid.normalize_normalized_factor, + part_enat.coe_inj] + at this, + refine this.trans _, + -- Get rid of the `map` by applying the equiv to both sides. + generalize hJ' : (normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI') + ⟨J, hJ⟩ = J', + have : ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm + J' : ideal S) = J, + { rw [← hJ', equiv.symm_apply_apply _ _, subtype.coe_mk] }, + subst this, + -- Get rid of the `attach` by applying the subtype `coe` to both sides. + rw [multiset.count_map_eq_count' (λ f, + ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm f + : ideal S)), + multiset.attach_count_eq_count_coe], + { exact subtype.coe_injective.comp (equiv.injective _) }, + { exact (normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI' _).prop }, + { exact irreducible_of_normalized_factor _ + (normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI' _).prop }, + { exact polynomial.map_monic_ne_zero (minpoly.monic pb.is_integral_gen) }, + { exact irreducible_of_normalized_factor _ hJ }, + { rwa [← bot_eq_zero, ne.def, map_eq_bot_iff_of_injective pb.basis.algebra_map_injective] }, +end + +theorem ideal.irreducible_map_of_irreducible_minpoly (hI : is_maximal I) (hI' : I ≠ ⊥) + (hf : irreducible (map I^.quotient.mk (minpoly R pb.gen))) : + irreducible (I.map (algebra_map R S)) := +begin + have mem_norm_factors : normalize (map I^.quotient.mk (minpoly R pb.gen)) ∈ normalized_factors + (map I^.quotient.mk (minpoly R pb.gen)) := by simp [normalized_factors_irreducible hf], + suffices : ∃ x, normalized_factors (I.map (algebra_map R S)) = {x}, + { obtain ⟨x, hx⟩ := this, + have h := normalized_factors_prod (show I.map (algebra_map R S) ≠ 0, by + rwa [← bot_eq_zero, ne.def, map_eq_bot_iff_of_injective pb.basis.algebra_map_injective]), + rw [associated_iff_eq, hx, multiset.prod_singleton] at h, + rw ← h, + exact irreducible_of_normalized_factor x + (show x ∈ normalized_factors (I.map (algebra_map R S)), by simp [hx]) }, + rw normalized_factors_ideal_map_eq_normalized_factors_min_poly_mk_map pb hI hI', + use ((normalized_factors_map_equiv_normalized_factors_min_poly_mk pb hI hI').symm + ⟨normalize (map I^.quotient.mk (minpoly R pb.gen)), mem_norm_factors⟩ : ideal S), + rw multiset.map_eq_singleton, + use ⟨normalize (map I^.quotient.mk (minpoly R pb.gen)), mem_norm_factors⟩, + refine ⟨_, rfl⟩, + apply multiset.map_injective subtype.coe_injective, + rw [multiset.attach_map_coe, multiset.map_singleton, subtype.coe_mk], + exact normalized_factors_irreducible hf +end + +end kummer_dedekind diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index edb84dddcba19..bc0a4c1cce383 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -1530,6 +1530,7 @@ by { rw [set_like.ext'_iff, ker_eq, set.ext_iff], exact injective_iff_map_eq_zer lemma ker_eq_bot_iff_eq_zero : ker f = ⊥ ↔ ∀ x, f x = 0 → x = 0 := by { rw [← injective_iff_map_eq_zero f, injective_iff_ker_eq_bot] } + omit rc @[simp] lemma ker_coe_equiv (f : R ≃+* S) : @@ -1677,6 +1678,11 @@ begin abel }, exact (H.mem_or_mem this).imp (λ h, ha ▸ mem_map_of_mem f h) (λ h, hb ▸ mem_map_of_mem f h) } end + +lemma map_eq_bot_iff_of_injective {I : ideal R} {f : F} (hf : function.injective f) : + I.map f = ⊥ ↔ I = ⊥ := +by rw [map_eq_bot_iff_le_ker, (ring_hom.injective_iff_ker_eq_bot f).mp hf, le_bot_iff] + omit rc theorem map_is_prime_of_equiv {F' : Type*} [ring_equiv_class F' R S] From 86a14674c6add6a846e0cdea75b0815b1dec3720 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Sat, 17 Sep 2022 06:13:12 +0000 Subject: [PATCH 289/828] docs(data/list/defs): fix `list.dedup` docstring (#16535) --- src/data/list/defs.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/list/defs.lean b/src/data/list/defs.lean index 641e6ad01e818..76c77a323cee4 100644 --- a/src/data/list/defs.lean +++ b/src/data/list/defs.lean @@ -644,7 +644,7 @@ def nodup : list α → Prop := pairwise (≠) instance nodup_decidable [decidable_eq α] : ∀ l : list α, decidable (nodup l) := list.decidable_pairwise -/-- `dedup l` removes duplicates from `l` (taking only the first occurrence). +/-- `dedup l` removes duplicates from `l` (taking only the last occurrence). Defined as `pw_filter (≠)`. dedup [1, 0, 2, 2, 1] = [0, 2, 1] -/ From a4f6c41edc28e44bc032ec0f3829af48e2b54bff Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 17 Sep 2022 09:00:18 +0000 Subject: [PATCH 290/828] move(algebra/field_power): Put lemmas in the correct place (#16465) `algebra.field_power` is only made of lemmas that belong elsewhere. Move them and delete the file. Golf proofs while we're at it and make `min_le_of_zpow_le_max` an iff. --- src/algebra/field/basic.lean | 16 +- src/algebra/field_power.lean | 153 -------------------- src/algebra/group_power/order.lean | 28 ++-- src/algebra/order/archimedean.lean | 1 - src/algebra/order/field.lean | 103 ++++++++++++- src/algebra/order/positive/field.lean | 2 +- src/algebra/parity.lean | 7 +- src/algebra/star/basic.lean | 7 +- src/analysis/hofer.lean | 9 +- src/data/int/log.lean | 3 +- src/data/rat/cast.lean | 2 + src/number_theory/frobenius_number.lean | 2 + src/number_theory/padics/padic_numbers.lean | 4 +- src/number_theory/padics/padic_val.lean | 4 +- src/topology/metric_space/pi_nat.lean | 2 +- 15 files changed, 142 insertions(+), 201 deletions(-) delete mode 100644 src/algebra/field_power.lean diff --git a/src/algebra/field/basic.lean b/src/algebra/field/basic.lean index 4315534eb93b8..3ba5aa0e83b4e 100644 --- a/src/algebra/field/basic.lean +++ b/src/algebra/field/basic.lean @@ -94,6 +94,18 @@ class division_ring (K : Type u) extends ring K, div_inv_monoid K, nontrivial K, (qsmul : ℚ → K → K := qsmul_rec rat_cast) (qsmul_eq_mul' : ∀ (a : ℚ) (x : K), qsmul a x = rat_cast a * x . try_refl_tac) +@[priority 100] -- see Note [lower instance priority] +instance division_ring.to_division_semiring [division_ring α] : division_semiring α := +{ ..‹division_ring α›, ..(infer_instance : semiring α) } + +section division_ring +variables [division_ring α] + +@[simp] lemma zpow_bit1_neg (x : α) (n : ℤ) : (-x) ^ bit1 n = - x ^ bit1 n := +by rw [zpow_bit1', zpow_bit1', neg_mul_neg, neg_mul_eq_mul_neg] + +end division_ring + /-- A `semifield` is a `comm_semiring` with multiplicative inverses for nonzero elements. -/ @[protect_proj, ancestor comm_semiring division_semiring comm_group_with_zero] class semifield (α : Type*) extends comm_semiring α, division_semiring α, comm_group_with_zero α @@ -111,10 +123,6 @@ See also Note [forgetful inheritance]. @[protect_proj, ancestor comm_ring div_inv_monoid nontrivial] class field (K : Type u) extends comm_ring K, division_ring K -@[priority 100] -- see Note [lower instance priority] -instance division_ring.to_division_semiring [division_ring α] : division_semiring α := -{ ..‹division_ring α›, ..(infer_instance : semiring α) } - section division_semiring variables [division_semiring α] {a b c : α} diff --git a/src/algebra/field_power.lean b/src/algebra/field_power.lean deleted file mode 100644 index 9ccaed4ef65ae..0000000000000 --- a/src/algebra/field_power.lean +++ /dev/null @@ -1,153 +0,0 @@ -/- -Copyright (c) 2018 Robert Y. Lewis. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Robert Y. Lewis --/ -import algebra.group_with_zero.power -import algebra.ring.equiv -import tactic.linarith - -/-! -# Integer power operation on fields and division rings - -This file collects basic facts about the operation of raising an element of a `division_ring` to an -integer power. More specialised results are provided in the case of a linearly ordered field. --/ - -open function int - -variables {α β : Type*} - -section division_ring -variables [division_ring α] [division_ring β] - -@[simp] lemma zpow_bit1_neg (x : α) (n : ℤ) : (-x) ^ bit1 n = - x ^ bit1 n := -by rw [zpow_bit1', zpow_bit1', neg_mul_neg, neg_mul_eq_mul_neg] - -@[simp, norm_cast] lemma rat.cast_zpow [char_zero α] (q : ℚ) (n : ℤ) : ((q ^ n : ℚ) : α) = q ^ n := -map_zpow₀ (rat.cast_hom α) q n - -end division_ring - -section linear_ordered_semifield -variables [linear_ordered_semifield α] {a b x : α} {m n : ℤ} - -lemma zpow_nonneg (ha : 0 ≤ a) : ∀ (z : ℤ), 0 ≤ a ^ z -| (n : ℕ) := by { rw zpow_coe_nat, exact pow_nonneg ha _ } -| -[1+n] := by { rw zpow_neg_succ_of_nat, exact inv_nonneg.2 (pow_nonneg ha _) } - -lemma zpow_pos_of_pos (ha : 0 < a) : ∀ (z : ℤ), 0 < a ^ z -| (n : ℕ) := by { rw zpow_coe_nat, exact pow_pos ha _ } -| -[1+n] := by { rw zpow_neg_succ_of_nat, exact inv_pos.2 (pow_pos ha _) } - -lemma zpow_le_of_le (ha : 1 ≤ a) (h : m ≤ n) : a ^ m ≤ a ^ n := -begin - have ha₀ : 0 < a, from one_pos.trans_le ha, - lift n - m to ℕ using sub_nonneg.2 h with k hk, - calc a ^ m = a ^ m * 1 : (mul_one _).symm - ... ≤ a ^ m * a ^ k : mul_le_mul_of_nonneg_left (one_le_pow_of_one_le ha _) (zpow_nonneg ha₀.le _) - ... = a ^ n : by rw [← zpow_coe_nat, ← zpow_add₀ ha₀.ne', hk, add_sub_cancel'_right] -end - -lemma pow_le_max_of_min_le (hx : 1 ≤ x) {a b c : ℤ} (h : min a b ≤ c) : - x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b)) := -begin - wlog hle : a ≤ b, - have hnle : -b ≤ -a, from neg_le_neg hle, - have hfle : x ^ (-b) ≤ x ^ (-a), from zpow_le_of_le hx hnle, - have : x ^ (-c) ≤ x ^ (-a), - { apply zpow_le_of_le hx, - simpa only [min_eq_left hle, neg_le_neg_iff] using h }, - simpa only [max_eq_left hfle] -end - -lemma zpow_le_one_of_nonpos (ha : 1 ≤ a) (hn : n ≤ 0) : a ^ n ≤ 1 := -(zpow_le_of_le ha hn).trans_eq $ zpow_zero _ - -lemma one_le_zpow_of_nonneg (ha : 1 ≤ a) (hn : 0 ≤ n) : 1 ≤ a ^ n := -(zpow_zero _).symm.trans_le $ zpow_le_of_le ha hn - -protected lemma nat.zpow_pos_of_pos {a : ℕ} (h : 0 < a) (n : ℤ) : 0 < (a : α)^n := -by { apply zpow_pos_of_pos, exact_mod_cast h } - -lemma nat.zpow_ne_zero_of_pos {a : ℕ} (h : 0 < a) (n : ℤ) : (a : α)^n ≠ 0 := -(nat.zpow_pos_of_pos h n).ne' - -lemma one_lt_zpow (ha : 1 < a) : ∀ n : ℤ, 0 < n → 1 < a ^ n -| (n : ℕ) h := (zpow_coe_nat _ _).symm.subst (one_lt_pow ha $ int.coe_nat_ne_zero.mp h.ne') -| -[1+ n] h := ((int.neg_succ_not_pos _).mp h).elim - -lemma zpow_strict_mono (hx : 1 < x) : strict_mono ((^) x : ℤ → α) := -strict_mono_int_of_lt_succ $ λ n, -have xpos : 0 < x, from zero_lt_one.trans hx, -calc x ^ n < x ^ n * x : lt_mul_of_one_lt_right (zpow_pos_of_pos xpos _) hx -... = x ^ (n + 1) : (zpow_add_one₀ xpos.ne' _).symm - -lemma zpow_strict_anti (h₀ : 0 < x) (h₁ : x < 1) : strict_anti ((^) x : ℤ → α) := -strict_anti_int_of_succ_lt $ λ n, -calc x ^ (n + 1) = x ^ n * x : zpow_add_one₀ h₀.ne' _ -... < x ^ n * 1 : (mul_lt_mul_left $ zpow_pos_of_pos h₀ _).2 h₁ -... = x ^ n : mul_one _ - -@[simp] lemma zpow_lt_iff_lt (hx : 1 < x) : x ^ m < x ^ n ↔ m < n := (zpow_strict_mono hx).lt_iff_lt -@[simp] lemma zpow_le_iff_le (hx : 1 < x) : x ^ m ≤ x ^ n ↔ m ≤ n := (zpow_strict_mono hx).le_iff_le - -lemma min_le_of_zpow_le_max (hx : 1 < x) {a b c : ℤ} - (h_max : x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b))) : min a b ≤ c := -begin - rw min_le_iff, - refine or.imp (λ h, _) (λ h, _) (le_max_iff.mp h_max); - rwa [zpow_le_iff_le hx, neg_le_neg_iff] at h -end - -@[simp] lemma pos_div_pow_pos (ha : 0 < a) (hb : 0 < b) (k : ℕ) : 0 < a/b^k := -div_pos ha (pow_pos hb k) - -@[simp] lemma div_pow_le (ha : 0 < a) (hb : 1 ≤ b) (k : ℕ) : a/b^k ≤ a := -(div_le_iff $ pow_pos (lt_of_lt_of_le zero_lt_one hb) k).mpr -(calc a = a * 1 : (mul_one a).symm - ... ≤ a*b^k : (mul_le_mul_left ha).mpr $ one_le_pow_of_one_le hb _) - -lemma zpow_injective (h₀ : 0 < x) (h₁ : x ≠ 1) : injective ((^) x : ℤ → α) := -begin - intros m n h, - rcases h₁.lt_or_lt with H|H, - { apply (zpow_strict_mono (one_lt_inv h₀ H)).injective, - show x⁻¹ ^ m = x⁻¹ ^ n, - rw [← zpow_neg_one, ← zpow_mul, ← zpow_mul, mul_comm _ m, mul_comm _ n, zpow_mul, zpow_mul, - h], }, - { exact (zpow_strict_mono H).injective h, }, -end - -@[simp] lemma zpow_inj (h₀ : 0 < x) (h₁ : x ≠ 1) : x ^ m = x ^ n ↔ m = n := -(zpow_injective h₀ h₁).eq_iff - -end linear_ordered_semifield - -section linear_ordered_field -variables [linear_ordered_field α] {a x : α} {m n : ℤ} - -lemma zpow_bit0_nonneg (a : α) (n : ℤ) : 0 ≤ a ^ bit0 n := -(mul_self_nonneg _).trans_eq $ (zpow_bit0 _ _).symm - -lemma zpow_two_nonneg (a : α) : 0 ≤ a ^ (2 : ℤ) := zpow_bit0_nonneg _ _ - -lemma zpow_bit0_pos (h : a ≠ 0) (n : ℤ) : 0 < a ^ bit0 n := -(zpow_bit0_nonneg a n).lt_of_ne (zpow_ne_zero _ h).symm - -lemma zpow_two_pos_of_ne_zero (h : a ≠ 0) : 0 < a ^ (2 : ℤ) := zpow_bit0_pos h _ - -@[simp] theorem zpow_bit1_neg_iff : a ^ bit1 n < 0 ↔ a < 0 := -⟨λ h, not_le.1 $ λ h', not_le.2 h $ zpow_nonneg h' _, - λ h, by rw [bit1, zpow_add_one₀ h.ne]; exact mul_neg_of_pos_of_neg (zpow_bit0_pos h.ne _) h⟩ - -@[simp] theorem zpow_bit1_nonneg_iff : 0 ≤ a ^ bit1 n ↔ 0 ≤ a := -le_iff_le_iff_lt_iff_lt.2 zpow_bit1_neg_iff - -@[simp] theorem zpow_bit1_nonpos_iff : a ^ bit1 n ≤ 0 ↔ a ≤ 0 := -by rw [le_iff_lt_or_eq, le_iff_lt_or_eq, zpow_bit1_neg_iff, zpow_eq_zero_iff (bit1_ne_zero n)] - -@[simp] theorem zpow_bit1_pos_iff : 0 < a ^ bit1 n ↔ 0 < a := -lt_iff_lt_of_le_iff_le zpow_bit1_nonpos_iff - -end linear_ordered_field diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index 0935f8a31ec8c..58d0cbda3e30a 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -222,23 +222,6 @@ begin by { rw [pow_succ _ n], exact mul_le_mul_of_nonneg_left (ih (nat.succ_ne_zero k)) h2 } end -theorem pow_lt_pow_of_lt_left (Hxy : x < y) (Hxpos : 0 ≤ x) (Hnpos : 0 < n) : - x ^ n < y ^ n := -begin - cases lt_or_eq_of_le Hxpos, - { rw ← tsub_add_cancel_of_le (nat.succ_le_of_lt Hnpos), - induction (n - 1), { simpa only [pow_one] }, - rw [pow_add, pow_add, nat.succ_eq_add_one, pow_one, pow_one], - apply mul_lt_mul ih (le_of_lt Hxy) h (le_of_lt (pow_pos (lt_trans h Hxy) _)) }, - { rw [←h, zero_pow Hnpos], apply pow_pos (by rwa ←h at Hxy : 0 < y),} -end - -lemma pow_lt_one (h₀ : 0 ≤ a) (h₁ : a < 1) {n : ℕ} (hn : n ≠ 0) : a ^ n < 1 := -(one_pow n).subst (pow_lt_pow_of_lt_left h₁ h₀ (nat.pos_of_ne_zero hn)) - -theorem strict_mono_on_pow (hn : 0 < n) : strict_mono_on (λ x : R, x ^ n) (set.Ici 0) := -λ x hx y hy h, pow_lt_pow_of_lt_left h hx hn - theorem one_le_pow_of_one_le (H : 1 ≤ a) : ∀ (n : ℕ), 1 ≤ a ^ n | 0 := by rw [pow_zero] | (n+1) := by { rw pow_succ, simpa only [mul_one] using mul_le_mul H (one_le_pow_of_one_le n) @@ -283,6 +266,14 @@ lemma pow_lt_pow_of_lt_one (h : 0 < a) (ha : a < 1) {i j : ℕ} (hij : i < j) : | (k+1) := by { rw [pow_succ, pow_succ], exact mul_le_mul hab (pow_le_pow_of_le_left _) (pow_nonneg ha _) (le_trans ha hab) } +lemma pow_lt_pow_of_lt_left (h : x < y) (hx : 0 ≤ x) : ∀ {n : ℕ}, 0 < n → x ^ n < y ^ n +| 0 hn := hn.false.elim +| (n + 1) _ := by simpa only [pow_succ'] using + mul_lt_mul' (pow_le_pow_of_le_left hx h.le _) h hx (pow_pos (hx.trans_lt h) _) + +lemma pow_lt_one (h₀ : 0 ≤ a) (h₁ : a < 1) {n : ℕ} (hn : n ≠ 0) : a ^ n < 1 := +(one_pow n).subst (pow_lt_pow_of_lt_left h₁ h₀ (nat.pos_of_ne_zero hn)) + lemma one_lt_pow (ha : 1 < a) {n : ℕ} (hn : n ≠ 0) : 1 < a ^ n := pow_zero a ▸ pow_lt_pow ha (pos_iff_ne_zero.2 hn) @@ -290,6 +281,9 @@ lemma pow_le_one : ∀ (n : ℕ) (h₀ : 0 ≤ a) (h₁ : a ≤ 1), a ^ n ≤ 1 | 0 h₀ h₁ := (pow_zero a).le | (n + 1) h₀ h₁ := (pow_succ' a n).le.trans (mul_le_one (pow_le_one n h₀ h₁) h₀ h₁) +lemma strict_mono_on_pow (hn : 0 < n) : strict_mono_on (λ x : R, x ^ n) (set.Ici 0) := +λ x hx y hy h, pow_lt_pow_of_lt_left h hx hn + lemma sq_pos_of_pos (ha : 0 < a) : 0 < a ^ 2 := by { rw sq, exact mul_pos ha ha } end ordered_semiring diff --git a/src/algebra/order/archimedean.lean b/src/algebra/order/archimedean.lean index 39a16c706e9e3..b04d316a1169d 100644 --- a/src/algebra/order/archimedean.lean +++ b/src/algebra/order/archimedean.lean @@ -3,7 +3,6 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import algebra.field_power import data.int.least_greatest import data.rat.floor diff --git a/src/algebra/order/field.lean b/src/algebra/order/field.lean index 897665de9ccf1..b04e1c3153d83 100644 --- a/src/algebra/order/field.lean +++ b/src/algebra/order/field.lean @@ -22,7 +22,7 @@ A linear ordered (semi)field is a (semi)field equipped with a linear order such set_option old_structure_cmd true -open order_dual +open function order_dual variables {α β : Type*} @@ -66,7 +66,7 @@ def injective.linear_ordered_field [linear_ordered_field α] [has_zero β] [has_ end function section linear_ordered_semifield -variables [linear_ordered_semifield α] {a b c d e : α} +variables [linear_ordered_semifield α] {a b c d e : α} {m n : ℤ} /-- `equiv.mul_left₀` as an order_iso. -/ @[simps {simp_rhs := tt}] @@ -124,6 +124,14 @@ by { rw div_eq_mul_inv, exact mul_nonpos_of_nonpos_of_nonneg ha (inv_nonneg.2 hb lemma div_nonpos_of_nonneg_of_nonpos (ha : 0 ≤ a) (hb : b ≤ 0) : a / b ≤ 0 := by { rw div_eq_mul_inv, exact mul_nonpos_of_nonneg_of_nonpos ha (inv_nonpos.2 hb) } +lemma zpow_nonneg (ha : 0 ≤ a) : ∀ n : ℤ, 0 ≤ a ^ n +| (n : ℕ) := by { rw zpow_coe_nat, exact pow_nonneg ha _ } +| -[1+n] := by { rw zpow_neg_succ_of_nat, exact inv_nonneg.2 (pow_nonneg ha _) } + +lemma zpow_pos_of_pos (ha : 0 < a) : ∀ n : ℤ, 0 < a ^ n +| (n : ℕ) := by { rw zpow_coe_nat, exact pow_pos ha _ } +| -[1+n] := by { rw zpow_neg_succ_of_nat, exact inv_pos.2 (pow_pos ha _) } + /-! ### Relating one division with another term. -/ @@ -417,6 +425,72 @@ by rwa [lt_one_div (@zero_lt_one α _ _) h1, one_div_one] lemma one_le_one_div (h1 : 0 < a) (h2 : a ≤ 1) : 1 ≤ 1 / a := by rwa [le_one_div (@zero_lt_one α _ _) h1, one_div_one] +/-! ### Integer powers -/ + +lemma zpow_le_of_le (ha : 1 ≤ a) (h : m ≤ n) : a ^ m ≤ a ^ n := +begin + have ha₀ : 0 < a, from one_pos.trans_le ha, + lift n - m to ℕ using sub_nonneg.2 h with k hk, + calc a ^ m = a ^ m * 1 : (mul_one _).symm + ... ≤ a ^ m * a ^ k : mul_le_mul_of_nonneg_left (one_le_pow_of_one_le ha _) (zpow_nonneg ha₀.le _) + ... = a ^ n : by rw [← zpow_coe_nat, ← zpow_add₀ ha₀.ne', hk, add_sub_cancel'_right] +end + +lemma zpow_le_one_of_nonpos (ha : 1 ≤ a) (hn : n ≤ 0) : a ^ n ≤ 1 := +(zpow_le_of_le ha hn).trans_eq $ zpow_zero _ + +lemma one_le_zpow_of_nonneg (ha : 1 ≤ a) (hn : 0 ≤ n) : 1 ≤ a ^ n := +(zpow_zero _).symm.trans_le $ zpow_le_of_le ha hn + +protected lemma nat.zpow_pos_of_pos {a : ℕ} (h : 0 < a) (n : ℤ) : 0 < (a : α)^n := +by { apply zpow_pos_of_pos, exact_mod_cast h } + +lemma nat.zpow_ne_zero_of_pos {a : ℕ} (h : 0 < a) (n : ℤ) : (a : α)^n ≠ 0 := +(nat.zpow_pos_of_pos h n).ne' + +lemma one_lt_zpow (ha : 1 < a) : ∀ n : ℤ, 0 < n → 1 < a ^ n +| (n : ℕ) h := (zpow_coe_nat _ _).symm.subst (one_lt_pow ha $ int.coe_nat_ne_zero.mp h.ne') +| -[1+ n] h := ((int.neg_succ_not_pos _).mp h).elim + +lemma zpow_strict_mono (hx : 1 < a) : strict_mono ((^) a : ℤ → α) := +strict_mono_int_of_lt_succ $ λ n, +have xpos : 0 < a, from zero_lt_one.trans hx, +calc a ^ n < a ^ n * a : lt_mul_of_one_lt_right (zpow_pos_of_pos xpos _) hx +... = a ^ (n + 1) : (zpow_add_one₀ xpos.ne' _).symm + +lemma zpow_strict_anti (h₀ : 0 < a) (h₁ : a < 1) : strict_anti ((^) a : ℤ → α) := +strict_anti_int_of_succ_lt $ λ n, +calc a ^ (n + 1) = a ^ n * a : zpow_add_one₀ h₀.ne' _ +... < a ^ n * 1 : (mul_lt_mul_left $ zpow_pos_of_pos h₀ _).2 h₁ +... = a ^ n : mul_one _ + +@[simp] lemma zpow_lt_iff_lt (hx : 1 < a) : a ^ m < a ^ n ↔ m < n := (zpow_strict_mono hx).lt_iff_lt +@[simp] lemma zpow_le_iff_le (hx : 1 < a) : a ^ m ≤ a ^ n ↔ m ≤ n := (zpow_strict_mono hx).le_iff_le + +@[simp] lemma div_pow_le (ha : 0 ≤ a) (hb : 1 ≤ b) (k : ℕ) : a/b^k ≤ a := +div_le_self ha $ one_le_pow_of_one_le hb _ + +lemma zpow_injective (h₀ : 0 < a) (h₁ : a ≠ 1) : injective ((^) a : ℤ → α) := +begin + rcases h₁.lt_or_lt with H|H, + { exact (zpow_strict_anti h₀ H).injective }, + { exact (zpow_strict_mono H).injective } +end + +@[simp] lemma zpow_inj (h₀ : 0 < a) (h₁ : a ≠ 1) : a ^ m = a ^ n ↔ m = n := +(zpow_injective h₀ h₁).eq_iff + +lemma zpow_le_max_of_min_le {x : α} (hx : 1 ≤ x) {a b c : ℤ} (h : min a b ≤ c) : + x ^ -c ≤ max (x ^ -a) (x ^ -b) := +begin + have : antitone (λ n : ℤ, x ^ -n) := λ m n h, zpow_le_of_le hx (neg_le_neg h), + exact (this h).trans_eq this.map_min, +end + +lemma zpow_le_max_iff_min_le {x : α} (hx : 1 < x) {a b c : ℤ} : + x ^ -c ≤ max (x ^ -a) (x ^ -b) ↔ min a b ≤ c := +by simp_rw [le_max_iff, min_le_iff, zpow_le_iff_le hx, neg_le_neg_iff] + /-! ### Results about halving. @@ -564,7 +638,7 @@ by simpa [mul_comm] using hs.mul_left ha end linear_ordered_semifield section -variables [linear_ordered_field α] {a b c d : α} +variables [linear_ordered_field α] {a b c d : α} {n : ℤ} /-! ### Lemmas about pos, nonneg, nonpos, neg -/ @@ -589,6 +663,29 @@ div_neg_iff.2 $ or.inr ⟨ha, hb⟩ lemma div_neg_of_pos_of_neg (ha : 0 < a) (hb : b < 0) : a / b < 0 := div_neg_iff.2 $ or.inl ⟨ha, hb⟩ +lemma zpow_bit0_nonneg (a : α) (n : ℤ) : 0 ≤ a ^ bit0 n := +(mul_self_nonneg _).trans_eq $ (zpow_bit0 _ _).symm + +lemma zpow_two_nonneg (a : α) : 0 ≤ a ^ (2 : ℤ) := zpow_bit0_nonneg _ _ + +lemma zpow_bit0_pos (h : a ≠ 0) (n : ℤ) : 0 < a ^ bit0 n := +(zpow_bit0_nonneg a n).lt_of_ne (zpow_ne_zero _ h).symm + +lemma zpow_two_pos_of_ne_zero (h : a ≠ 0) : 0 < a ^ (2 : ℤ) := zpow_bit0_pos h _ + +@[simp] lemma zpow_bit1_neg_iff : a ^ bit1 n < 0 ↔ a < 0 := +⟨λ h, not_le.1 $ λ h', not_le.2 h $ zpow_nonneg h' _, + λ h, by rw [bit1, zpow_add_one₀ h.ne]; exact mul_neg_of_pos_of_neg (zpow_bit0_pos h.ne _) h⟩ + +@[simp] lemma zpow_bit1_nonneg_iff : 0 ≤ a ^ bit1 n ↔ 0 ≤ a := +le_iff_le_iff_lt_iff_lt.2 zpow_bit1_neg_iff + +@[simp] lemma zpow_bit1_nonpos_iff : a ^ bit1 n ≤ 0 ↔ a ≤ 0 := +by rw [le_iff_lt_or_eq, le_iff_lt_or_eq, zpow_bit1_neg_iff, zpow_eq_zero_iff (int.bit1_ne_zero n)] + +@[simp] lemma zpow_bit1_pos_iff : 0 < a ^ bit1 n ↔ 0 < a := +lt_iff_lt_of_le_iff_le zpow_bit1_nonpos_iff + /-! ### Relating one division with another term -/ lemma div_le_iff_of_neg (hc : c < 0) : b / c ≤ a ↔ a * c ≤ b := diff --git a/src/algebra/order/positive/field.lean b/src/algebra/order/positive/field.lean index 1cdc637185be3..ea8e3f6f66f01 100644 --- a/src/algebra/order/positive/field.lean +++ b/src/algebra/order/positive/field.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import algebra.order.field import algebra.order.positive.ring -import algebra.field_power /-! # Algebraic structures on the set of positive numbers diff --git a/src/algebra/parity.lean b/src/algebra/parity.lean index 07a69fc759459..8492a7488ed21 100644 --- a/src/algebra/parity.lean +++ b/src/algebra/parity.lean @@ -3,12 +3,7 @@ Copyright (c) 2022 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ - -import algebra.ring.basic -import algebra.algebra.basic -import algebra.group_power.basic -import algebra.field_power -import algebra.opposites +import algebra.associated /-! # Squares, even and odd elements diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index 07ab3af09d291..c3550fa6430b9 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -3,13 +3,10 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import tactic.apply_fun -import algebra.field.opposite -import algebra.field_power import algebra.ring.aut -import group_theory.group_action.units -import group_theory.group_action.opposite import algebra.ring.comp_typeclasses +import data.rat.cast +import group_theory.group_action.opposite /-! # Star monoids, rings, and modules diff --git a/src/analysis/hofer.lean b/src/analysis/hofer.lean index f428bb649b913..8212d269731c0 100644 --- a/src/analysis/hofer.lean +++ b/src/analysis/hofer.lean @@ -23,6 +23,10 @@ open filter finset local notation `d` := dist +@[simp] lemma pos_div_pow_pos {α : Type*} [linear_ordered_semifield α] {a b : α} (ha : 0 < a) + (hb : 0 < b) (k : ℕ) : 0 < a/b^k := +div_pos ha (pow_pos hb k) + lemma hofer {X: Type*} [metric_space X] [complete_space X] (x : X) (ε : ℝ) (ε_pos : 0 < ε) {ϕ : X → ℝ} (cont : continuous ϕ) (nonneg : ∀ y, 0 ≤ ϕ y) : @@ -35,14 +39,13 @@ begin have reformulation : ∀ x' (k : ℕ), ε * ϕ x ≤ ε / 2 ^ k * ϕ x' ↔ 2^k * ϕ x ≤ ϕ x', { intros x' k, rw [div_mul_eq_mul_div, le_div_iff, mul_assoc, mul_le_mul_left ε_pos, mul_comm], - exact pow_pos (by norm_num) k, }, + positivity }, -- Now let's specialize to `ε/2^k` replace H : ∀ k : ℕ, ∀ x', d x' x ≤ 2 * ε ∧ 2^k * ϕ x ≤ ϕ x' → ∃ y, d x' y ≤ ε/2^k ∧ 2 * ϕ x' < ϕ y, { intros k x', push_neg at H, - simpa [reformulation] using - H (ε/2^k) (by simp [ε_pos, zero_lt_two]) x' (by simp [ε_pos, zero_lt_two, one_le_two]) }, + simpa [reformulation] using H (ε/2^k) (by simp [ε_pos]) x' (by simp [ε_pos.le, one_le_two]) }, clear reformulation, haveI : nonempty X := ⟨x⟩, choose! F hF using H, -- Use the axiom of choice diff --git a/src/data/int/log.lean b/src/data/int/log.lean index 62658ff083540..4e27e0e5675d2 100644 --- a/src/data/int/log.lean +++ b/src/data/int/log.lean @@ -3,9 +3,8 @@ Copyright (c) 2022 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import data.nat.log import algebra.order.floor -import algebra.field_power +import data.nat.log /-! # Integer logarithms in a field with respect to a natural base diff --git a/src/data/rat/cast.lean b/src/data/rat/cast.lean index b37b323ebb688..780ae00f4e71f 100644 --- a/src/data/rat/cast.lean +++ b/src/data/rat/cast.lean @@ -204,6 +204,8 @@ variable {α} @[simp, norm_cast] theorem cast_inv (n) : ((n⁻¹ : ℚ) : α) = n⁻¹ := map_inv₀ (cast_hom α) _ @[simp, norm_cast] theorem cast_div (m n) : ((m / n : ℚ) : α) = m / n := map_div₀ (cast_hom α) _ _ +@[simp, norm_cast] theorem cast_zpow (q : ℚ) (n : ℤ) : ((q ^ n : ℚ) : α) = q ^ n := +map_zpow₀ (cast_hom α) q n @[norm_cast] theorem cast_mk (a b : ℤ) : ((a /. b) : α) = a / b := by simp only [mk_eq_div, cast_div, cast_coe_int] diff --git a/src/number_theory/frobenius_number.lean b/src/number_theory/frobenius_number.lean index 5d6ce548e546f..d740e474a73c8 100644 --- a/src/number_theory/frobenius_number.lean +++ b/src/number_theory/frobenius_number.lean @@ -6,6 +6,8 @@ Authors: Alex Zhao import data.nat.modeq import group_theory.submonoid.basic import group_theory.submonoid.membership +import tactic.ring +import tactic.zify /-! # Frobenius Number in Two Variables diff --git a/src/number_theory/padics/padic_numbers.lean b/src/number_theory/padics/padic_numbers.lean index 830a6142985b7..c11594bed6308 100644 --- a/src/number_theory/padics/padic_numbers.lean +++ b/src/number_theory/padics/padic_numbers.lean @@ -993,8 +993,8 @@ begin have hp_one : (1 : ℝ) < p, { rw [← nat.cast_one, nat.cast_lt], exact nat.prime.one_lt hp.elim }, - rw [norm_eq_pow_val hx, norm_eq_pow_val hy, norm_eq_pow_val hxy] at h_norm, - exact min_le_of_zpow_le_max hp_one h_norm }} + rwa [norm_eq_pow_val hx, norm_eq_pow_val hy, norm_eq_pow_val hxy, + zpow_le_max_iff_min_le hp_one] at h_norm } } end @[simp] lemma valuation_map_mul {x y : ℚ_[p]} (hx : x ≠ 0) (hy : y ≠ 0) : diff --git a/src/number_theory/padics/padic_val.lean b/src/number_theory/padics/padic_val.lean index b32672aca4be5..55bd9cb0ea5ff 100644 --- a/src/number_theory/padics/padic_val.lean +++ b/src/number_theory/padics/padic_val.lean @@ -4,11 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Y. Lewis -/ import algebra.order.absolute_value -import algebra.field_power +import number_theory.divisors import ring_theory.int.basic -import tactic.basic import tactic.ring_exp -import number_theory.divisors /-! # p-adic Valuation diff --git a/src/topology/metric_space/pi_nat.lean b/src/topology/metric_space/pi_nat.lean index 31070b61bb700..0a7287c0e7b4f 100644 --- a/src/topology/metric_space/pi_nat.lean +++ b/src/topology/metric_space/pi_nat.lean @@ -393,7 +393,7 @@ begin { simp only [le_infi_iff, le_principal_iff], assume n, refine mem_infi_of_mem ((1/2)^n) _, - refine mem_infi_of_mem (by norm_num) _, + refine mem_infi_of_mem (by positivity) _, simp only [mem_principal, set_of_subset_set_of, prod.forall], assume x y hxy, exact apply_eq_of_dist_lt hxy le_rfl } From 0ff46b9cf6c2e285614b28a64534923de2c78b31 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 17 Sep 2022 14:47:01 +0000 Subject: [PATCH 291/828] doc(order/filter/basic): fix notations in doc of filter.comap (#16530) The notations in the doc of ```filter.comap``` did not match the notations used in the lemma and thus were a bit confusing (at least for me). --- src/order/filter/basic.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 1249ab1e3031c..60e4c891ce7bd 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1546,12 +1546,12 @@ end map section comap -/-- The inverse map of a filter. A set `s` belongs to `filter.comap f l` if either of the following +/-- The inverse map of a filter. A set `s` belongs to `filter.comap m f` if either of the following equivalent conditions hold. -1. There exists a set `t ∈ l` such that `f ⁻¹' t ⊆ s`. This is used as a definition. -2. The set `{y | ∀ x, f x = y → x ∈ s}` belongs to `l`, see `filter.mem_comap'`. -3. The set `(f '' sᶜ)ᶜ` belongs to `l`, see `filter.mem_comap_iff_compl` and +1. There exists a set `t ∈ f` such that `m ⁻¹' t ⊆ s`. This is used as a definition. +2. The set `{y | ∀ x, m x = y → x ∈ s}` belongs to `f`, see `filter.mem_comap'`. +3. The set `(m '' sᶜ)ᶜ` belongs to `f`, see `filter.mem_comap_iff_compl` and `filter.compl_mem_comap`. -/ def comap (m : α → β) (f : filter β) : filter α := { sets := { s | ∃ t ∈ f, m ⁻¹' t ⊆ s }, From f7e477fdd84a5c43675436be1f6c8d747875e4e6 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 17 Sep 2022 16:29:31 +0000 Subject: [PATCH 292/828] =?UTF-8?q?feat(topology/separation):=20define=20`?= =?UTF-8?q?T=E2=82=85`=20spaces=20(#16533)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I'm going to prove that any `order_topology` is a `t5_space` in a follow-up PR. This will imply that any set in an `order_topology` is a normal space. --- counterexamples/sorgenfrey_line.lean | 43 +++++++++++++------------ src/topology/maps.lean | 5 +++ src/topology/nhds_set.lean | 3 ++ src/topology/separation.lean | 48 +++++++++++++++++++++++++--- 4 files changed, 73 insertions(+), 26 deletions(-) diff --git a/counterexamples/sorgenfrey_line.lean b/counterexamples/sorgenfrey_line.lean index 549f2eab337cd..cac0747a22745 100644 --- a/counterexamples/sorgenfrey_line.lean +++ b/counterexamples/sorgenfrey_line.lean @@ -16,8 +16,9 @@ import data.set.intervals.monotone In this file we define `sorgenfrey_line` (notation: `ℝₗ`) to be the Sorgenfrey line. It is the real line with the topology space structure generated by half-open intervals `set.Ico a b`. -We prove that this line is a normal space but its product with itself is not a normal space. In -particular, this implies that the topology on `ℝₗ` is neither metrizable, nor second countable. +We prove that this line is a completely normal Hausdorff space but its product with itself is not a +normal space. In particular, this implies that the topology on `ℝₗ` is neither metrizable, nor +second countable. ## Notations @@ -152,30 +153,30 @@ instance : totally_disconnected_space ℝₗ := instance : first_countable_topology ℝₗ := ⟨λ x, (nhds_basis_Ico_rat x).is_countably_generated⟩ -/-- Sorgenfrey line is a normal topological space. -/ -instance : normal_space ℝₗ := +/-- Sorgenfrey line is a completely normal Hausdorff topological space. -/ +instance : t5_space ℝₗ := begin /- Let `s` and `t` be disjoint closed sets. For each `x ∈ s` we choose `X x` such that `set.Ico x (X x)` is disjoint with `t`. Similarly, for each `y ∈ t` we choose `Y y` such that `set.Ico y (Y y)` is disjoint with `s`. Then `⋃ x ∈ s, Ico x (X x)` and `⋃ y ∈ t, Ico y (Y y)` are disjoint open sets that include `s` and `t`. -/ - refine ⟨λ s t hs ht hd, _⟩, - choose! X hX hXd using λ x (hx : x ∈ s), exists_Ico_disjoint_closed ht (disjoint_left.1 hd hx), - choose! Y hY hYd using λ y (hy : y ∈ t), exists_Ico_disjoint_closed hs (disjoint_right.1 hd hy), - refine ⟨⋃ x ∈ s, Ico x (X x), ⋃ y ∈ t, Ico y (Y y), is_open_bUnion $ λ _ _, is_open_Ico _ _, - is_open_bUnion $ λ _ _, is_open_Ico _ _, _, _, _⟩, - { exact λ x hx, mem_Union₂.2 ⟨x, hx, left_mem_Ico.2 $ hX x hx⟩ }, - { exact λ y hy, mem_Union₂.2 ⟨y, hy, left_mem_Ico.2 $ hY y hy⟩ }, - { simp only [disjoint_Union_left, disjoint_Union_right, Ico_disjoint_Ico], - intros y hy x hx, - clear hs ht hd hX hY, - cases le_total x y with hle hle, - { calc min (X x) (Y y) ≤ X x : min_le_left _ _ - ... ≤ y : not_lt.1 (λ hyx, hXd x hx ⟨⟨hle, hyx⟩, hy⟩) - ... ≤ max x y : le_max_right _ _ }, - { calc min (X x) (Y y) ≤ Y y : min_le_right _ _ - ... ≤ x : not_lt.1 $ λ hxy, hYd y hy ⟨⟨hle, hxy⟩, hx⟩ - ... ≤ max x y : le_max_left _ _ } } + refine ⟨λ s t hd₁ hd₂, _⟩, + choose! X hX hXd + using λ x (hx : x ∈ s), exists_Ico_disjoint_closed is_closed_closure (disjoint_left.1 hd₂ hx), + choose! Y hY hYd + using λ y (hy : y ∈ t), exists_Ico_disjoint_closed is_closed_closure (disjoint_right.1 hd₁ hy), + refine disjoint_of_disjoint_of_mem _ + (bUnion_mem_nhds_set $ λ x hx, (is_open_Ico x (X x)).mem_nhds $ left_mem_Ico.2 (hX x hx)) + (bUnion_mem_nhds_set $ λ y hy, (is_open_Ico y (Y y)).mem_nhds $ left_mem_Ico.2 (hY y hy)), + simp only [disjoint_Union_left, disjoint_Union_right, Ico_disjoint_Ico], + intros y hy x hx, + cases le_total x y with hle hle, + { calc min (X x) (Y y) ≤ X x : min_le_left _ _ + ... ≤ y : not_lt.1 (λ hyx, hXd x hx ⟨⟨hle, hyx⟩, subset_closure hy⟩) + ... ≤ max x y : le_max_right _ _ }, + { calc min (X x) (Y y) ≤ Y y : min_le_right _ _ + ... ≤ x : not_lt.1 $ λ hxy, hYd y hy ⟨⟨hle, hxy⟩, subset_closure hx⟩ + ... ≤ max x y : le_max_left _ _ } end lemma dense_range_coe_rat : dense_range (coe : ℚ → ℝₗ) := diff --git a/src/topology/maps.lean b/src/topology/maps.lean index 53dac03657330..45c4b0b3318ef 100644 --- a/src/topology/maps.lean +++ b/src/topology/maps.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 topology.order +import topology.nhds_set /-! # Specific classes of maps between topological spaces @@ -77,6 +78,10 @@ lemma inducing.nhds_eq_comap {f : α → β} (hf : inducing f) : ∀ (a : α), 𝓝 a = comap f (𝓝 $ f a) := inducing_iff_nhds.1 hf +lemma inducing.nhds_set_eq_comap {f : α → β} (hf : inducing f) (s : set α) : + 𝓝ˢ s = comap f (𝓝ˢ (f '' s)) := +by simp only [nhds_set, Sup_image, comap_supr, hf.nhds_eq_comap, supr_image] + lemma inducing.map_nhds_eq {f : α → β} (hf : inducing f) (a : α) : (𝓝 a).map f = 𝓝[range f] (f a) := hf.induced.symm ▸ map_nhds_induced_eq a diff --git a/src/topology/nhds_set.lean b/src/topology/nhds_set.lean index 04fc5c72a962b..c50a4a94bb142 100644 --- a/src/topology/nhds_set.lean +++ b/src/topology/nhds_set.lean @@ -41,6 +41,9 @@ by { rw [nhds_set, ← range_diag, ← range_comp], refl } lemma mem_nhds_set_iff_forall : s ∈ 𝓝ˢ t ↔ ∀ (x : α), x ∈ t → s ∈ 𝓝 x := by simp_rw [nhds_set, filter.mem_Sup, ball_image_iff] +lemma bUnion_mem_nhds_set {t : α → set α} (h : ∀ x ∈ s, t x ∈ 𝓝 x) : (⋃ x ∈ s, t x) ∈ 𝓝ˢ s := +mem_nhds_set_iff_forall.2 $ λ x hx, mem_of_superset (h x hx) (subset_Union₂ x hx) + lemma subset_interior_iff_mem_nhds_set : s ⊆ interior t ↔ t ∈ 𝓝ˢ s := by simp_rw [mem_nhds_set_iff_forall, subset_interior_iff_nhds] diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 6123bd318017a..e47e5bb58dddc 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -31,6 +31,7 @@ This file defines the predicate `separated`, and common separation axioms * `normal_space`: A T₄ space (sometimes referred to as normal, but authors vary on whether this includes T₂; `mathlib` does), is one where given two disjoint closed sets, we can find two open sets that separate them. In `mathlib`, T₄ implies T₃. +* `t5_space`: A T₅ space, also known as a *completely normal Hausdorff space* ## Main results @@ -105,6 +106,11 @@ def separated : set α → set α → Prop := λ (s t : set α), ∃ U V : (set α), (is_open U) ∧ is_open V ∧ (s ⊆ U) ∧ (t ⊆ V) ∧ disjoint U V +lemma separated_iff_disjoint {s t : set α} : + separated s t ↔ disjoint (𝓝ˢ s) (𝓝ˢ t) := +by simp only [(has_basis_nhds_set s).disjoint_iff (has_basis_nhds_set t), separated, exists_prop, + ← exists_and_distrib_left, and.assoc, and.comm, and.left_comm] + namespace separated open separated @@ -142,11 +148,7 @@ lemma mono {s₁ s₂ t₁ t₂ : set α} (h : separated s₂ t₂) (hs : s₁ let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h in ⟨U, V, hU, hV, hs.trans hsU, ht.trans htV, hd⟩ lemma union_left {a b c : set α} : separated a c → separated b c → separated (a ∪ b) c := -λ ⟨U, V, oU, oV, aU, bV, UV⟩ ⟨W, X, oW, oX, aW, bX, WX⟩, - ⟨U ∪ W, V ∩ X, is_open.union oU oW, is_open.inter oV oX, - union_subset_union aU aW, subset_inter bV bX, set.disjoint_union_left.mpr - ⟨disjoint_of_subset_right (inter_subset_left _ _) UV, - disjoint_of_subset_right (inter_subset_right _ _) WX⟩⟩ +by simpa only [separated_iff_disjoint, nhds_set_union, disjoint_sup_left] using and.intro lemma union_right {a b c : set α} (ab : separated a b) (ac : separated a c) : separated a (b ∪ c) := @@ -1484,6 +1486,42 @@ end end normality +section completely_normal + +/-- A topological space `α` is a *completely normal Hausdorff space* if each subspace `s : set α` is +a normal Hausdorff space. Equivalently, `α` is a `T₁` space and for any two sets `s`, `t` such that +`closure s` is disjoint with `t` and `s` is disjoint with `closure t`, there exist disjoint +neighbourhoods of `s` and `t`. -/ +class t5_space (α : Type u) [topological_space α] extends t1_space α : Prop := +(completely_normal : ∀ ⦃s t : set α⦄, disjoint (closure s) t → disjoint s (closure t) → + disjoint (𝓝ˢ s) (𝓝ˢ t)) + +export t5_space (completely_normal) + +lemma embedding.t5_space [topological_space β] [t5_space β] {e : α → β} (he : embedding e) : + t5_space α := +begin + haveI := he.t1_space, + refine ⟨λ s t hd₁ hd₂, _⟩, + simp only [he.to_inducing.nhds_set_eq_comap], + refine disjoint_comap (completely_normal _ _), + { rwa [← subset_compl_iff_disjoint_left, image_subset_iff, preimage_compl, + ← he.closure_eq_preimage_closure_image, subset_compl_iff_disjoint_left] }, + { rwa [← subset_compl_iff_disjoint_right, image_subset_iff, preimage_compl, + ← he.closure_eq_preimage_closure_image, subset_compl_iff_disjoint_right] } +end + +/-- A subspace of a `T₅` space is a `T₅` space. -/ +instance [t5_space α] {p : α → Prop} : t5_space {x // p x} := embedding_subtype_coe.t5_space + +/-- A `T₅` space is a `T₄` space. -/ +@[priority 100] -- see Note [lower instance priority] +instance t5_space.to_normal_space [t5_space α] : normal_space α := +⟨λ s t hs ht hd, separated_iff_disjoint.2 $ + completely_normal (by rwa [hs.closure_eq]) (by rwa [ht.closure_eq])⟩ + +end completely_normal + /-- In a compact t2 space, the connected component of a point equals the intersection of all its clopen neighbourhoods. -/ lemma connected_component_eq_Inter_clopen [t2_space α] [compact_space α] (x : α) : From 57bbd9ec2ebd75995038b0ea9aa7d448a5c569c5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 17 Sep 2022 18:12:09 +0000 Subject: [PATCH 293/828] chore(order/symm_diff): Generalize to co-Heyting algebras (#16282) Generalize the `symm_diff` material from `(generalized_)boolean_algebra` to `(generalized_)coheyting_algebra`. --- src/data/dfinsupp/interval.lean | 8 +- src/data/finset/basic.lean | 7 + src/data/set/basic.lean | 22 ++- src/measure_theory/decomposition/jordan.lean | 11 +- src/order/symm_diff.lean | 158 ++++++++++++------- 5 files changed, 135 insertions(+), 71 deletions(-) diff --git a/src/data/dfinsupp/interval.lean b/src/data/dfinsupp/interval.lean index e231a1c5849d6..2516db9bbc19a 100644 --- a/src/data/dfinsupp/interval.lean +++ b/src/data/dfinsupp/interval.lean @@ -92,10 +92,10 @@ def range_Icc (f g : Π₀ i, α i) : Π₀ i, finset (α i) := { to_fun := λ i, Icc (f i) (g i), support' := f.support'.bind $ λ fs, g.support'.map $ λ gs, ⟨fs + gs, λ i, or_iff_not_imp_left.2 $ λ h, begin - have hf : f i = 0 := - (fs.prop i).resolve_left (multiset.not_mem_mono (multiset.le_add_right _ _).subset h), - have hg : g i = 0 := - (gs.prop i).resolve_left (multiset.not_mem_mono (multiset.le_add_left _ _).subset h), + have hf : f i = 0 := (fs.prop i).resolve_left + (multiset.not_mem_mono (multiset.le.subset $ multiset.le_add_right _ _) h), + have hg : g i = 0 := (gs.prop i).resolve_left + (multiset.not_mem_mono (multiset.le.subset $ multiset.le_add_left _ _) h), rw [hf, hg], exact Icc_self _, end⟩ } diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index a12082e020a13..cc369ee8984ad 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -1339,6 +1339,13 @@ sup_eq_sdiff_sup_sdiff_sup_inf lemma erase_eq_empty_iff (s : finset α) (a : α) : s.erase a = ∅ ↔ s = ∅ ∨ s = {a} := by rw [←sdiff_singleton_eq_erase, sdiff_eq_empty_iff_subset, subset_singleton_iff] +/-! ### Symmetric difference -/ + +lemma mem_symm_diff : a ∈ s ∆ t ↔ a ∈ s ∧ a ∉ t ∨ a ∈ t ∧ a ∉ s := +by simp_rw [symm_diff, sup_eq_union, mem_union, mem_sdiff] + +@[simp, norm_cast] lemma coe_symm_diff : (↑(s ∆ t) : set α) = s ∆ t := set.ext $ λ _, mem_symm_diff + end decidable_eq /-! ### attach -/ diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index e357532fd41ec..4e3af611b712c 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ -import order.boolean_algebra +import order.symm_diff /-! # Basic properties of sets @@ -77,7 +77,7 @@ universes u v w x namespace set -variable {α : Type*} +variables {α : Type*} {s t : set α} instance : has_le (set α) := ⟨λ s t, ∀ ⦃x⦄, x ∈ s → x ∈ t⟩ instance : has_subset (set α) := ⟨(≤)⟩ @@ -104,6 +104,12 @@ instance : has_inter (set α) := ⟨(⊓)⟩ @[simp] lemma le_eq_subset : ((≤) : set α → set α → Prop) = (⊆) := rfl @[simp] lemma lt_eq_ssubset : ((<) : set α → set α → Prop) = (⊂) := rfl +lemma le_iff_subset : s ≤ t ↔ s ⊆ t := iff.rfl +lemma lt_iff_ssubset : s ≤ t ↔ s ⊆ t := iff.rfl + +alias le_iff_subset ↔ _root_.has_le.le.subset _root_.has_subset.subset.le +alias lt_iff_ssubset ↔ _root_.has_lt.lt.ssubset _root_.has_ssubset.ssubset.lt + /-- Coercion from a set to the corresponding subtype. -/ instance {α : Type u} : has_coe_to_sort (set α) (Type u) := ⟨λ s, {x // x ∈ s}⟩ @@ -1178,6 +1184,18 @@ lemma union_eq_diff_union_diff_union_inter (s t : set α) : s ∪ t = (s \ t) ∪ (t \ s) ∪ (s ∩ t) := sup_eq_sdiff_sup_sdiff_sup_inf +/-! ### Symmetric difference -/ + +lemma mem_symm_diff : a ∈ s ∆ t ↔ a ∈ s ∧ a ∉ t ∨ a ∈ t ∧ a ∉ s := iff.rfl + +lemma symm_diff_subset_union : s ∆ t ⊆ s ∪ t := @symm_diff_le_sup (set α) _ _ _ + +lemma inter_symm_diff_distrib_left (s t u : set α) : s ∩ t ∆ u = (s ∩ t) ∆ (s ∩ u) := +inf_symm_diff_distrib_left _ _ _ + +lemma inter_symm_diff_distrib_right (s t u : set α) : s ∆ t ∩ u = (s ∩ u) ∆ (t ∩ u) := +inf_symm_diff_distrib_right _ _ _ + /-! ### Powerset -/ /-- `𝒫 s = set.powerset s` is the set of all subsets of `s`. -/ diff --git a/src/measure_theory/decomposition/jordan.lean b/src/measure_theory/decomposition/jordan.lean index e1cad89cc5703..e8d762e73b939 100644 --- a/src/measure_theory/decomposition/jordan.lean +++ b/src/measure_theory/decomposition/jordan.lean @@ -169,7 +169,7 @@ end jordan_decomposition namespace signed_measure -open measure vector_measure jordan_decomposition classical +open classical jordan_decomposition measure set vector_measure variables {s : signed_measure α} {μ ν : measure α} [is_finite_measure μ] [is_finite_measure ν] @@ -293,12 +293,9 @@ lemma of_inter_eq_of_symm_diff_eq_zero_positive begin have hwuv : s ((w ∩ u) ∆ (w ∩ v)) = 0, { refine subset_positive_null_set (hu.union hv) ((hw.inter hu).symm_diff (hw.inter hv)) - (hu.symm_diff hv) (restrict_le_restrict_union _ _ hu hsu hv hsv) hs _ _, - { exact symm_diff_le_sup u v }, - { rintro x (⟨⟨hxw, hxu⟩, hx⟩ | ⟨⟨hxw, hxv⟩, hx⟩); - rw [set.mem_inter_eq, not_and] at hx, - { exact or.inl ⟨hxu, hx hxw⟩ }, - { exact or.inr ⟨hxv, hx hxw⟩ } } }, + (hu.symm_diff hv) (restrict_le_restrict_union _ _ hu hsu hv hsv) hs symm_diff_subset_union _, + rw ←inter_symm_diff_distrib_left, + exact inter_subset_right _ _ }, obtain ⟨huv, hvu⟩ := of_diff_eq_zero_of_symm_diff_eq_zero_positive (hw.inter hu) (hw.inter hv) (restrict_le_restrict_subset _ _ hu hsu (w.inter_subset_right u)) diff --git a/src/order/symm_diff.lean b/src/order/symm_diff.lean index d9ca1c0007027..dd9529d82c243 100644 --- a/src/order/symm_diff.lean +++ b/src/order/symm_diff.lean @@ -44,23 +44,23 @@ boolean ring, generalized boolean algebra, boolean algebra, symmetric difference open function +variables {α : Type*} + /-- The symmetric difference operator on a type with `⊔` and `\` is `(A \ B) ⊔ (B \ A)`. -/ -def symm_diff {α : Type*} [has_sup α] [has_sdiff α] (A B : α) : α := (A \ B) ⊔ (B \ A) +def symm_diff [has_sup α] [has_sdiff α] (a b : α) : α := a \ b ⊔ b \ a /- This notation might conflict with the Laplacian once we have it. Feel free to put it in locale `order` or `symm_diff` if that happens. -/ infix ` ∆ `:100 := symm_diff -lemma symm_diff_def {α : Type*} [has_sup α] [has_sdiff α] (A B : α) : - A ∆ B = (A \ B) ⊔ (B \ A) := -rfl +lemma symm_diff_def [has_sup α] [has_sdiff α] (a b : α) : a ∆ b = (a \ b) ⊔ (b \ a) := rfl lemma symm_diff_eq_xor (p q : Prop) : p ∆ q = xor p q := rfl @[simp] lemma bool.symm_diff_eq_bxor : ∀ p q : bool, p ∆ q = bxor p q := dec_trivial -section generalized_boolean_algebra -variables {α : Type*} [generalized_boolean_algebra α] (a b c d : α) +section generalized_coheyting_algebra +variables [generalized_coheyting_algebra α] (a b c d : α) lemma symm_diff_comm : a ∆ b = b ∆ a := by simp only [(∆), sup_comm] @@ -70,8 +70,86 @@ instance symm_diff_is_comm : is_commutative α (∆) := ⟨symm_diff_comm⟩ @[simp] lemma symm_diff_bot : a ∆ ⊥ = a := by rw [(∆), sdiff_bot, bot_sdiff, sup_bot_eq] @[simp] lemma bot_symm_diff : ⊥ ∆ a = a := by rw [symm_diff_comm, symm_diff_bot] -lemma symm_diff_eq_sup_sdiff_inf : a ∆ b = (a ⊔ b) \ (a ⊓ b) := -by simp [sup_sdiff, sdiff_inf, sup_comm, (∆)] +lemma symm_diff_of_le {a b : α} (h : a ≤ b) : a ∆ b = b \ a := +by rw [symm_diff, sdiff_eq_bot_iff.2 h, bot_sup_eq] + +lemma symm_diff_of_ge {a b : α} (h : b ≤ a) : a ∆ b = a \ b := +by rw [symm_diff, sdiff_eq_bot_iff.2 h, sup_bot_eq] + +lemma symm_diff_le {a b c : α} (ha : a ≤ b ⊔ c) (hb : b ≤ a ⊔ c) : a ∆ b ≤ c := +sup_le (sdiff_le_iff.2 ha) $ sdiff_le_iff.2 hb + +lemma symm_diff_le_iff {a b c : α} : a ∆ b ≤ c ↔ a ≤ b ⊔ c ∧ b ≤ a ⊔ c := +by simp_rw [symm_diff, sup_le_iff, sdiff_le_iff] + +@[simp] lemma symm_diff_le_sup {a b : α} : a ∆ b ≤ a ⊔ b := sup_le_sup sdiff_le sdiff_le + +lemma symm_diff_eq_sup_sdiff_inf : a ∆ b = (a ⊔ b) \ (a ⊓ b) := by simp [sup_sdiff, symm_diff] + +lemma disjoint.symm_diff_eq_sup {a b : α} (h : disjoint a b) : a ∆ b = a ⊔ b := +by rw [(∆), h.sdiff_eq_left, h.sdiff_eq_right] + +lemma symm_diff_sdiff : (a ∆ b) \ c = a \ (b ⊔ c) ⊔ b \ (a ⊔ c) := +by rw [symm_diff, sup_sdiff_distrib, sdiff_sdiff_left, sdiff_sdiff_left] + +@[simp] lemma symm_diff_sdiff_inf : a ∆ b \ (a ⊓ b) = a ∆ b := +by { rw symm_diff_sdiff, simp [symm_diff] } + +@[simp] lemma symm_diff_sdiff_eq_sup : a ∆ (b \ a) = a ⊔ b := +begin + rw [symm_diff, sdiff_idem], + exact le_antisymm (sup_le_sup sdiff_le sdiff_le) + (sup_le le_sdiff_sup $ le_sdiff_sup.trans $ sup_le le_sup_right le_sdiff_sup), +end + +@[simp] lemma sdiff_symm_diff_eq_sup : (a \ b) ∆ b = a ⊔ b := +by rw [symm_diff_comm, symm_diff_sdiff_eq_sup, sup_comm] + +@[simp] lemma symm_diff_sup_inf : a ∆ b ⊔ a ⊓ b = a ⊔ b := +begin + refine le_antisymm (sup_le symm_diff_le_sup inf_le_sup) _, + rw [sup_inf_left, symm_diff], + refine sup_le (le_inf le_sup_right _) (le_inf _ le_sup_right), + { rw sup_right_comm, + exact le_sup_of_le_left le_sdiff_sup }, + { rw sup_assoc, + exact le_sup_of_le_right le_sdiff_sup } +end + +@[simp] lemma inf_sup_symm_diff : a ⊓ b ⊔ a ∆ b = a ⊔ b := by rw [sup_comm, symm_diff_sup_inf] + +@[simp] lemma symm_diff_symm_diff_inf : a ∆ b ∆ (a ⊓ b) = a ⊔ b := +by rw [←symm_diff_sdiff_inf a, sdiff_symm_diff_eq_sup, symm_diff_sup_inf] + +@[simp] lemma inf_symm_diff_symm_diff : (a ⊓ b) ∆ (a ∆ b) = a ⊔ b := +by rw [symm_diff_comm, symm_diff_symm_diff_inf] + +lemma symm_diff_triangle : a ∆ c ≤ a ∆ b ⊔ b ∆ c := +begin + refine (sup_le_sup (sdiff_triangle a b c) $ sdiff_triangle _ b _).trans_eq _, + rw [@sup_comm _ _ (c \ b), sup_sup_sup_comm, symm_diff, symm_diff], +end + +end generalized_coheyting_algebra + +section coheyting_algebra +variables [coheyting_algebra α] (a : α) + +@[simp] lemma symm_diff_top' : a ∆ ⊤ = ¬a := by simp [symm_diff] +@[simp] lemma top_symm_diff' : ⊤ ∆ a = ¬a := by simp [symm_diff] + +@[simp] lemma hnot_symm_diff_self : (¬a) ∆ a = ⊤ := +by { rw [eq_top_iff, symm_diff, hnot_sdiff, sup_sdiff_self], exact codisjoint_hnot_left } + +@[simp] lemma symm_diff_hnot_self : a ∆ ¬a = ⊤ := by rw [symm_diff_comm, hnot_symm_diff_self] + +lemma is_compl.symm_diff_eq_top {a b : α} (h : is_compl a b) : a ∆ b = ⊤ := +by rw [h.eq_hnot, hnot_symm_diff_self] + +end coheyting_algebra + +section generalized_boolean_algebra +variables [generalized_boolean_algebra α] (a b c d : α) @[simp] lemma sup_sdiff_symm_diff : (a ⊔ b) \ (a ∆ b) = a ⊓ b := sdiff_eq_symm inf_le_sup (by rw symm_diff_eq_sup_sdiff_inf) @@ -82,8 +160,6 @@ begin exact disjoint_sdiff_self_left, end -lemma symm_diff_le_sup : a ∆ b ≤ a ⊔ b := by { rw symm_diff_eq_sup_sdiff_inf, exact sdiff_le } - lemma inf_symm_diff_distrib_left : a ⊓ (b ∆ c) = (a ⊓ b) ∆ (a ⊓ c) := by rw [symm_diff_eq_sup_sdiff_inf, inf_sdiff_distrib_left, inf_sup_left, inf_inf_distrib_left, symm_diff_eq_sup_sdiff_inf] @@ -97,9 +173,6 @@ by simp only [(∆), sdiff_sdiff_sup_sdiff'] lemma sdiff_symm_diff' : c \ (a ∆ b) = (c ⊓ a ⊓ b) ⊔ (c \ (a ⊔ b)) := by rw [sdiff_symm_diff, sdiff_sup, sup_comm] -lemma symm_diff_sdiff : (a ∆ b) \ c = (a \ (b ⊔ c)) ⊔ (b \ (a ⊔ c)) := -by rw [symm_diff_def, sup_sdiff, sdiff_sdiff_left, sdiff_sdiff_left] - @[simp] lemma symm_diff_sdiff_left : (a ∆ b) \ a = b \ a := by rw [symm_diff_def, sup_sdiff, sdiff_idem, sdiff_sdiff_self, bot_sup_eq] @@ -108,31 +181,11 @@ by rw [symm_diff_comm, symm_diff_sdiff_left] @[simp] lemma sdiff_symm_diff_self : a \ (a ∆ b) = a ⊓ b := by simp [sdiff_symm_diff] -lemma symm_diff_eq_iff_sdiff_eq {a b c : α} (ha : a ≤ c) : - a ∆ b = c ↔ c \ a = b := -begin - split; intro h, - { have hba : disjoint (a ⊓ b) c := begin - rw [←h, disjoint.comm], - exact disjoint_symm_diff_inf _ _, - end, - have hca : _ := congr_arg (\ a) h, - rw [symm_diff_sdiff_left] at hca, - rw [←hca, sdiff_eq_self_iff_disjoint], - exact hba.of_disjoint_inf_of_le ha }, - { have hd : disjoint a b := by { rw ←h, exact disjoint_sdiff_self_right }, - rw [symm_diff_def, hd.sdiff_eq_left, hd.sdiff_eq_right, ←h, sup_sdiff_cancel_right ha] } -end - -lemma disjoint.symm_diff_eq_sup {a b : α} (h : disjoint a b) : a ∆ b = a ⊔ b := -by rw [(∆), h.sdiff_eq_left, h.sdiff_eq_right] - lemma symm_diff_eq_sup : a ∆ b = a ⊔ b ↔ disjoint a b := begin - split; intro h, - { rw [symm_diff_eq_sup_sdiff_inf, sdiff_eq_self_iff_disjoint] at h, - exact h.of_disjoint_inf_of_le le_sup_left, }, - { exact h.symm_diff_eq_sup, }, + refine ⟨λ h, _, disjoint.symm_diff_eq_sup⟩, + rw [symm_diff_eq_sup_sdiff_inf, sdiff_eq_self_iff_disjoint] at h, + exact h.of_disjoint_inf_of_le le_sup_left, end @[simp] lemma le_symm_diff_iff_left : a ≤ a ∆ b ↔ disjoint a b := @@ -163,19 +216,6 @@ calc a ∆ (b ∆ c) = (a \ (b ∆ c)) ⊔ ((b ∆ c) \ a) : symm_diff_def _ _ ... = (a \ (b ⊔ c)) ⊔ (b \ (a ⊔ c)) ⊔ (c \ (a ⊔ b)) ⊔ (a ⊓ b ⊓ c) : by ac_refl -@[simp] lemma symm_diff_symm_diff_inf : a ∆ b ∆ (a ⊓ b) = a ⊔ b := -by rw [symm_diff_eq_iff_sdiff_eq (symm_diff_le_sup _ _), sup_sdiff_symm_diff] - -@[simp] lemma inf_symm_diff_symm_diff : (a ⊓ b) ∆ (a ∆ b) = a ⊔ b := -by rw [symm_diff_comm, symm_diff_symm_diff_inf] - -lemma symm_diff_triangle : a ∆ c ≤ a ∆ b ⊔ b ∆ c := -begin - refine (sup_le_sup (sdiff_triangle a b c) $ sdiff_triangle _ b _).trans_eq _, - rw [@sup_comm _ _ (c \ b), sup_sup_sup_comm], - refl, -end - lemma symm_diff_assoc : a ∆ b ∆ c = a ∆ (b ∆ c) := by rw [symm_diff_symm_diff_left, symm_diff_symm_diff_right] @@ -228,15 +268,21 @@ protected lemma disjoint.symm_diff_right (ha : disjoint a b) (hb : disjoint a c) disjoint a (b ∆ c) := (ha.symm.symm_diff_left hb.symm).symm +lemma symm_diff_eq_iff_sdiff_eq (ha : a ≤ c) : a ∆ b = c ↔ c \ a = b := +begin + rw ←symm_diff_of_le ha, + exact ((symm_diff_right_involutive a).to_perm _).apply_eq_iff_eq_symm_apply.trans eq_comm, +end + end generalized_boolean_algebra section boolean_algebra -variables {α : Type*} [boolean_algebra α] (a b c : α) +variables [boolean_algebra α] (a b c : α) lemma symm_diff_eq : a ∆ b = (a ⊓ bᶜ) ⊔ (b ⊓ aᶜ) := by simp only [(∆), sdiff_eq] -@[simp] lemma symm_diff_top : a ∆ ⊤ = aᶜ := by simp [symm_diff_eq] -@[simp] lemma top_symm_diff : ⊤ ∆ a = aᶜ := by rw [symm_diff_comm, symm_diff_top] +lemma symm_diff_top : a ∆ ⊤ = aᶜ := symm_diff_top' _ +lemma top_symm_diff : ⊤ ∆ a = aᶜ := top_symm_diff' _ lemma compl_symm_diff : (a ∆ b)ᶜ = (a ⊓ b) ⊔ (aᶜ ⊓ bᶜ) := by simp only [←top_sdiff, sdiff_symm_diff, top_inf_eq] @@ -244,12 +290,8 @@ by simp only [←top_sdiff, sdiff_symm_diff, top_inf_eq] lemma symm_diff_eq_top_iff : a ∆ b = ⊤ ↔ is_compl a b := by rw [symm_diff_eq_iff_sdiff_eq le_top, top_sdiff, compl_eq_iff_is_compl] -lemma is_compl.symm_diff_eq_top (h : is_compl a b) : a ∆ b = ⊤ := (symm_diff_eq_top_iff a b).2 h - -@[simp] lemma compl_symm_diff_self : aᶜ ∆ a = ⊤ := -by simp only [symm_diff_eq, compl_compl, inf_idem, compl_sup_eq_top] - -@[simp] lemma symm_diff_compl_self : a ∆ aᶜ = ⊤ := by rw [symm_diff_comm, compl_symm_diff_self] +@[simp] lemma compl_symm_diff_self : aᶜ ∆ a = ⊤ := hnot_symm_diff_self _ +@[simp] lemma symm_diff_compl_self : a ∆ aᶜ = ⊤ := symm_diff_hnot_self _ lemma symm_diff_symm_diff_right' : a ∆ (b ∆ c) = (a ⊓ b ⊓ c) ⊔ (a ⊓ bᶜ ⊓ cᶜ) ⊔ (aᶜ ⊓ b ⊓ cᶜ) ⊔ (aᶜ ⊓ bᶜ ⊓ c) := From f8dfe8102eea8138413a4d328fc68cec320b8d07 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 17 Sep 2022 18:12:10 +0000 Subject: [PATCH 294/828] chore(algebra/order/nonneg): add a `distrib_lattice` instance (#16493) --- src/algebra/order/nonneg.lean | 6 ++++++ src/order/lattice_intervals.lean | 4 ++++ 2 files changed, 10 insertions(+) diff --git a/src/algebra/order/nonneg.lean b/src/algebra/order/nonneg.lean index c45acf1280d9a..e0d7c627300e7 100644 --- a/src/algebra/order/nonneg.lean +++ b/src/algebra/order/nonneg.lean @@ -50,9 +50,15 @@ lemma bot_eq [preorder α] {a : α} : (⊥ : {x : α // a ≤ x}) = ⟨a, le_rfl instance no_max_order [partial_order α] [no_max_order α] {a : α} : no_max_order {x : α // a ≤ x} := set.Ici.no_max_order +instance semilattice_sup [semilattice_sup α] {a : α} : semilattice_sup {x : α // a ≤ x} := +set.Ici.semilattice_sup + instance semilattice_inf [semilattice_inf α] {a : α} : semilattice_inf {x : α // a ≤ x} := set.Ici.semilattice_inf +instance distrib_lattice [distrib_lattice α] {a : α} : distrib_lattice {x : α // a ≤ x} := +set.Ici.distrib_lattice + instance densely_ordered [preorder α] [densely_ordered α] {a : α} : densely_ordered {x : α // a ≤ x} := show densely_ordered (Ici a), from set.densely_ordered diff --git a/src/order/lattice_intervals.lean b/src/order/lattice_intervals.lean index 47257e7c0cc23..48e4b07455219 100644 --- a/src/order/lattice_intervals.lean +++ b/src/order/lattice_intervals.lean @@ -110,6 +110,10 @@ instance [lattice α] {a : α} : lattice (Ici a) := { .. Ici.semilattice_inf, .. Ici.semilattice_sup } +instance [distrib_lattice α] {a : α} : distrib_lattice (Ici a) := +{ le_sup_inf := λ a b c, le_sup_inf, + .. Ici.lattice } + instance [preorder α] {a : α} : order_bot (Ici a) := { bot := ⟨a, le_refl a⟩, bot_le := λ x, x.prop } From ba584a4c8bd41618ce6fab1f3c26678709485a2e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 17 Sep 2022 23:24:12 +0000 Subject: [PATCH 295/828] chore(topology/constructions): deduplicate (#16505) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * rename `mem_closure_of_continuous` to `set.maps_to.closure_left`, move near `map_mem_closure`; * merge `map_mem_closure2` and `mem_closure_of_continuous2` into `map_mem_closure₂`; * fix&golf some proofs broken by these changes. --- src/analysis/convex/topology.lean | 9 ++-- src/topology/algebra/group.lean | 5 +-- src/topology/algebra/monoid.lean | 8 +--- src/topology/basic.lean | 9 +++- src/topology/constructions.lean | 63 ++++++++++------------------ src/topology/instances/ennreal.lean | 2 +- src/topology/metric_space/basic.lean | 3 +- 7 files changed, 39 insertions(+), 60 deletions(-) diff --git a/src/analysis/convex/topology.lean b/src/analysis/convex/topology.lean index f648923a4fed2..1efbd7ec8b029 100644 --- a/src/analysis/convex/topology.lean +++ b/src/analysis/convex/topology.lean @@ -199,11 +199,10 @@ convex_iff_open_segment_subset.mpr $ λ x hx y hy, protected lemma convex.closure {s : set E} (hs : convex 𝕜 s) : convex 𝕜 (closure s) := λ x hx y hy a b ha hb hab, let f : E → E → E := λ x' y', a • x' + b • y' in -have hf : continuous (λ p : E × E, f p.1 p.2), from - (continuous_fst.const_smul _).add (continuous_snd.const_smul _), -show f x y ∈ closure s, from - mem_closure_of_continuous2 hf hx hy (λ x' hx' y' hy', subset_closure - (hs hx' hy' ha hb hab)) +have hf : continuous (function.uncurry f), + from (continuous_fst.const_smul _).add (continuous_snd.const_smul _), +show f x y ∈ closure s, + from map_mem_closure₂ hf hx hy (λ x' hx' y' hy', hs hx' hy' ha hb hab) end has_continuous_const_smul diff --git a/src/topology/algebra/group.lean b/src/topology/algebra/group.lean index b95b052dc9dc1..ef95bcf08edb2 100644 --- a/src/topology/algebra/group.lean +++ b/src/topology/algebra/group.lean @@ -557,9 +557,8 @@ lemma subgroup.is_normal_topological_closure {G : Type*} [topological_space G] [ (subgroup.topological_closure N).normal := { conj_mem := λ n hn g, begin - apply mem_closure_of_continuous (topological_group.continuous_conj g) hn, - intros m hm, - exact subset_closure (subgroup.normal.conj_mem infer_instance m hm g), + apply map_mem_closure (topological_group.continuous_conj g) hn, + exact λ m hm, subgroup.normal.conj_mem infer_instance m hm g end } @[to_additive] lemma mul_mem_connected_component_one {G : Type*} [topological_space G] diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 6cc5cf80fe184..5886701a9a683 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -251,12 +251,8 @@ variables [topological_space M] [monoid M] [has_continuous_mul M] @[to_additive] lemma submonoid.top_closure_mul_self_subset (s : submonoid M) : closure (s : set M) * closure s ⊆ closure s := -calc -closure (s : set M) * closure s - = (λ p : M × M, p.1 * p.2) '' closure (s ×ˢ s) : by simp [closure_prod_eq] -... ⊆ closure ((λ p : M × M, p.1 * p.2) '' s ×ˢ s) : - image_closure_subset_closure_image continuous_mul -... = closure s : by simp [s.coe_mul_self_eq] +image2_subset_iff.2 $ λ x hx y hy, map_mem_closure₂ continuous_mul hx hy $ + λ a ha b hb, s.mul_mem ha hb @[to_additive] lemma submonoid.top_closure_mul_self_eq (s : submonoid M) : diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 4f5fb87d0eb77..190b273fb2a6a 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -1409,8 +1409,13 @@ lemma closure_subset_preimage_closure_image {f : α → β} {s : set α} (h : co by { rw ← set.image_subset_iff, exact image_closure_subset_closure_image h } lemma map_mem_closure {s : set α} {t : set β} {f : α → β} {a : α} - (hf : continuous f) (ha : a ∈ closure s) (ht : ∀a∈s, f a ∈ t) : f a ∈ closure t := -set.maps_to.closure ht hf ha + (hf : continuous f) (ha : a ∈ closure s) (ht : maps_to f s t) : f a ∈ closure t := +ht.closure hf ha + +/-- If a continuous map `f` maps `s` to a closed set `t`, then it maps `closure s` to `t`. -/ +lemma set.maps_to.closure_left {s : set α} {t : set β} {f : α → β} (h : maps_to f s t) + (hc : continuous f) (ht : is_closed t) : maps_to f (closure s) t := +ht.closure_eq ▸ h.closure hc /-! ### Function with dense range diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index 5f71b93167519..1a069cdb859cb 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -34,7 +34,7 @@ product, sum, disjoint union, subspace, quotient space noncomputable theory -open topological_space set filter +open topological_space set filter function open_locale classical topological_space filter universes u v @@ -388,15 +388,15 @@ lemma continuous_swap : continuous (prod.swap : α × β → β × α) := continuous_snd.prod_mk continuous_fst lemma continuous_uncurry_left {f : α → β → γ} (a : α) - (h : continuous (function.uncurry f)) : continuous (f a) := -show continuous (function.uncurry f ∘ (λ b, (a, b))), from h.comp (by continuity) + (h : continuous (uncurry f)) : continuous (f a) := +show continuous (uncurry f ∘ (λ b, (a, b))), from h.comp (by continuity) lemma continuous_uncurry_right {f : α → β → γ} (b : β) - (h : continuous (function.uncurry f)) : continuous (λ a, f a b) := -show continuous (function.uncurry f ∘ (λ a, (a, b))), from h.comp (by continuity) + (h : continuous (uncurry f)) : continuous (λ a, f a b) := +show continuous (uncurry f ∘ (λ a, (a, b))), from h.comp (by continuity) lemma continuous_curry {g : α × β → γ} (a : α) - (h : continuous g) : continuous (function.curry g a) := + (h : continuous g) : continuous (curry g a) := show continuous (g ∘ (λ b, (a, b))), from h.comp (by continuity) lemma is_open.prod {s : set α} {t : set β} (hs : is_open s) (ht : is_open t) : @@ -409,7 +409,7 @@ by rw [filter.prod, prod.topological_space, nhds_inf, nhds_induced, nhds_induced /-- If a function `f x y` is such that `y ↦ f x y` is continuous for all `x`, and `x` lives in a discrete space, then `f` is continuous. -/ lemma continuous_uncurry_of_discrete_topology [discrete_topology α] - {f : α → β → γ} (hf : ∀ a, continuous (f a)) : continuous (function.uncurry f) := + {f : α → β → γ} (hf : ∀ a, continuous (f a)) : continuous (uncurry f) := begin apply continuous_iff_continuous_at.2, rintros ⟨a, x⟩, @@ -544,10 +544,10 @@ lemma prod_induced_induced {α γ : Type*} (f : α → β) (g : γ → δ) : by simp_rw [prod.topological_space, induced_inf, induced_compose] lemma continuous_uncurry_of_discrete_topology_left [discrete_topology α] - {f : α → β → γ} (h : ∀ a, continuous (f a)) : continuous (function.uncurry f) := + {f : α → β → γ} (h : ∀ a, continuous (f a)) : continuous (uncurry f) := continuous_iff_continuous_at.2 $ λ ⟨a, b⟩, by simp only [continuous_at, nhds_prod_eq, nhds_discrete α, pure_prod, tendsto_map'_iff, (∘), - function.uncurry, (h a).tendsto] + uncurry, (h a).tendsto] /-- Given a neighborhood `s` of `(x, x)`, then `(x, x)` has a square open neighborhood that is a subset of `s`. -/ @@ -638,13 +638,13 @@ by simp [frontier_prod_eq] frontier ((univ : set α) ×ˢ s) = univ ×ˢ frontier s := by simp [frontier_prod_eq] -lemma map_mem_closure2 {s : set α} {t : set β} {u : set γ} {f : α → β → γ} {a : α} {b : β} - (hf : continuous (λp:α×β, f p.1 p.2)) (ha : a ∈ closure s) (hb : b ∈ closure t) - (hu : ∀a b, a ∈ s → b ∈ t → f a b ∈ u) : +lemma map_mem_closure₂ {f : α → β → γ} {a : α} {b : β} {s : set α} {t : set β} {u : set γ} + (hf : continuous (uncurry f)) (ha : a ∈ closure s) (hb : b ∈ closure t) + (h : ∀ (a ∈ s) (b ∈ t), f a b ∈ u) : f a b ∈ closure u := -have (a, b) ∈ closure (s ×ˢ t), by rw [closure_prod_eq]; from ⟨ha, hb⟩, -show (λp:α×β, f p.1 p.2) (a, b) ∈ closure u, from - map_mem_closure hf this $ assume ⟨a, b⟩ ⟨ha, hb⟩, hu a b ha hb +have H₁ : (a, b) ∈ closure (s ×ˢ t), by simpa only [closure_prod_eq] using mk_mem_prod ha hb, +have H₂ : maps_to (uncurry f) (s ×ˢ t) u, from forall_prod_set.2 h, +H₂.closure hf H₁ lemma is_closed.prod {s₁ : set α} {s₂ : set β} (h₁ : is_closed s₁) (h₂ : is_closed s₂) : is_closed (s₁ ×ˢ s₂) := @@ -1013,23 +1013,23 @@ tendsto_pi_nhds lemma filter.tendsto.update [∀i, topological_space (π i)] [decidable_eq ι] {l : filter α} {f : α → Π i, π i} {x : Π i, π i} (hf : tendsto f l (𝓝 x)) (i : ι) {g : α → π i} {xi : π i} (hg : tendsto g l (𝓝 xi)) : - tendsto (λ a, function.update (f a) i (g a)) l (𝓝 $ function.update x i xi) := + tendsto (λ a, update (f a) i (g a)) l (𝓝 $ update x i xi) := tendsto_pi_nhds.2 $ λ j, by { rcases em (j = i) with rfl|hj; simp [*, hf.apply] } lemma continuous_at.update [∀i, topological_space (π i)] [topological_space α] [decidable_eq ι] {f : α → Π i, π i} {a : α} (hf : continuous_at f a) (i : ι) {g : α → π i} (hg : continuous_at g a) : - continuous_at (λ a, function.update (f a) i (g a)) a := + continuous_at (λ a, update (f a) i (g a)) a := hf.update i hg lemma continuous.update [∀i, topological_space (π i)] [topological_space α] [decidable_eq ι] {f : α → Π i, π i} (hf : continuous f) (i : ι) {g : α → π i} (hg : continuous g) : - continuous (λ a, function.update (f a) i (g a)) := + continuous (λ a, update (f a) i (g a)) := continuous_iff_continuous_at.2 $ λ x, hf.continuous_at.update i hg.continuous_at /-- `function.update f i x` is continuous in `(f, x)`. -/ @[continuity] lemma continuous_update [∀i, topological_space (π i)] [decidable_eq ι] (i : ι) : - continuous (λ f : (Π j, π j) × π i, function.update f.1 i f.2) := + continuous (λ f : (Π j, π j) × π i, update f.1 i f.2) := continuous_fst.update i continuous_snd lemma filter.tendsto.fin_insert_nth {n} {π : fin (n + 1) → Type*} [Π i, topological_space (π i)] @@ -1095,7 +1095,7 @@ lemma pi_eq_generate_from [∀a, topological_space (π a)] : le_antisymm (le_generate_from $ assume g ⟨s, i, hi, eq⟩, eq.symm ▸ is_open_set_pi (finset.finite_to_set _) hi) (le_infi $ assume a s ⟨t, ht, s_eq⟩, generate_open.basic _ $ - ⟨function.update (λa, univ) a t, {a}, by simpa using ht, s_eq ▸ by ext f; simp [set.pi]⟩) + ⟨update (λa, univ) a t, {a}, by simpa using ht, s_eq ▸ by ext f; simp [set.pi]⟩) lemma pi_generate_from_eq {g : Πa, set (set (π a))} : @Pi.topological_space ι π (λa, generate_from (g a)) = @@ -1110,7 +1110,7 @@ begin apply is_open_bInter (finset.finite_to_set _), assume a ha, show ((generate_from G).coinduced (λf:Πa, π a, f a)).is_open (t a), refine le_generate_from _ _ (hi a ha), - exact assume s hs, generate_open.basic _ ⟨function.update (λa, univ) a s, {a}, by simp [hs]⟩ } + exact assume s hs, generate_open.basic _ ⟨update (λa, univ) a s, {a}, by simp [hs]⟩ } end lemma pi_generate_from_eq_finite {g : Πa, set (set (π a))} [finite ι] (hg : ∀a, ⋃₀ g a = univ) : @@ -1156,7 +1156,7 @@ instance Pi.discrete_topology : discrete_topology (Π i, π i) := singletons_open_iff_discrete.mp (λ x, begin rw show {x} = ⋂ i, {y : Π i, π i | y i = x i}, - { ext, simp only [function.funext_iff, set.mem_singleton_iff, set.mem_Inter, set.mem_set_of_eq] }, + { ext, simp only [funext_iff, set.mem_singleton_iff, set.mem_Inter, set.mem_set_of_eq] }, exact is_open_Inter (λ i, (continuous_apply i).is_open_preimage {x i} (is_open_discrete {x i})) end) @@ -1300,22 +1300,3 @@ continuous_induced_dom continuous_induced_rng.2 continuous_id end ulift - -lemma mem_closure_of_continuous [topological_space α] [topological_space β] - {f : α → β} {a : α} {s : set α} {t : set β} - (hf : continuous f) (ha : a ∈ closure s) (h : maps_to f s (closure t)) : - f a ∈ closure t := -calc f a ∈ f '' closure s : mem_image_of_mem _ ha - ... ⊆ closure (f '' s) : image_closure_subset_closure_image hf - ... ⊆ closure t : closure_minimal h.image_subset is_closed_closure - -lemma mem_closure_of_continuous2 [topological_space α] [topological_space β] [topological_space γ] - {f : α → β → γ} {a : α} {b : β} {s : set α} {t : set β} {u : set γ} - (hf : continuous (λp:α×β, f p.1 p.2)) (ha : a ∈ closure s) (hb : b ∈ closure t) - (h : ∀a∈s, ∀b∈t, f a b ∈ closure u) : - f a b ∈ closure u := -have (a,b) ∈ closure (s ×ˢ t), - by simp [closure_prod_eq, ha, hb], -show f (a, b).1 (a, b).2 ∈ closure u, - from @mem_closure_of_continuous (α×β) _ _ _ (λp:α×β, f p.1 p.2) (a,b) _ u hf this $ - assume ⟨p₁, p₂⟩ ⟨h₁, h₂⟩, h p₁ h₁ p₂ h₂ diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 78c2d96d4ff96..f4ea915ed9763 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -1361,7 +1361,7 @@ is_closed_le (continuous_id.edist continuous_const) continuous_const begin refine le_antisymm (diam_le $ λ x hx y hy, _) (diam_mono subset_closure), have : edist x y ∈ closure (Iic (diam s)), - from map_mem_closure2 (@continuous_edist α _) hx hy (λ _ _, edist_le_diam_of_mem), + from map_mem_closure₂ continuous_edist hx hy (λ x hx y hy, edist_le_diam_of_mem hx hy), rwa closure_Iic at this end diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 8dcbdf059ec13..197a82f8ebdf3 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -2122,8 +2122,7 @@ end lemma bounded_closure_of_bounded (h : bounded s) : bounded (closure s) := let ⟨C, h⟩ := h in -⟨C, λ a ha b hb, (is_closed_le' C).closure_subset $ map_mem_closure2 continuous_dist ha hb -$ ball_mem_comm.mp h⟩ +⟨C, λ a ha b hb, (is_closed_le' C).closure_subset $ map_mem_closure₂ continuous_dist ha hb h⟩ alias bounded_closure_of_bounded ← bounded.closure From 1dc531a0dfefe94b3d9a1eda4f941234ccf2e67a Mon Sep 17 00:00:00 2001 From: mcdoll Date: Sun, 18 Sep 2022 07:54:17 +0000 Subject: [PATCH 296/828] feat(measure_theory/integral): finiteness of lower Lebesgue integral (#16514) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR proves that if a function is integrable and non-negative, then it's lower Lebesgue integral is finite. Co-authored-by: Sébastien Gouëzel Co-authored-by: Remy Degenne Co-authored-by: Rémy Degenne --- src/measure_theory/function/l1_space.lean | 7 +------ .../integral/integrable_on.lean | 10 +++++++++ src/measure_theory/integral/lebesgue.lean | 21 +++++++++++++++++++ 3 files changed, 32 insertions(+), 6 deletions(-) diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index e3b03c584789b..552e0764bfc5a 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -109,12 +109,7 @@ by simp only [has_finite_integral_iff_norm, edist_dist, dist_zero_right] lemma has_finite_integral_iff_of_real {f : α → ℝ} (h : 0 ≤ᵐ[μ] f) : has_finite_integral f μ ↔ ∫⁻ a, ennreal.of_real (f a) ∂μ < ∞ := -have lintegral_eq : ∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ = ∫⁻ a, ennreal.of_real (f a) ∂μ := -begin - refine lintegral_congr_ae (h.mono $ λ a h, _), - rwa [real.norm_eq_abs, abs_of_nonneg] -end, -by rw [has_finite_integral_iff_norm, lintegral_eq] +by rw [has_finite_integral, lintegral_nnnorm_eq_of_ae_nonneg h] lemma has_finite_integral_iff_of_nnreal {f : α → ℝ≥0} : has_finite_integral (λ x, (f x : ℝ)) μ ↔ ∫⁻ a, f a ∂μ < ∞ := diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index 95a616ed92867..e97a455cab5b3 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -246,6 +246,16 @@ begin exact ((Lp.mem_ℒp _).restrict s).mem_ℒp_of_exponent_le hp, end +lemma integrable.lintegral_lt_top {f : α → ℝ} (hf : integrable f μ) : + ∫⁻ x, ennreal.of_real (f x) ∂μ < ∞ := +calc ∫⁻ x, ennreal.of_real (f x) ∂μ + ≤ ∫⁻ x, ↑∥f x∥₊ ∂μ : lintegral_to_real_le_lintegral_nnnorm f +... < ∞ : hf.2 + +lemma integrable_on.set_lintegral_lt_top {f : α → ℝ} {s : set α} (hf : integrable_on f s μ) : + ∫⁻ x in s, ennreal.of_real (f x) ∂μ < ∞ := +integrable.lintegral_lt_top hf + /-- We say that a function `f` is *integrable at filter* `l` if it is integrable on some set `s ∈ l`. Equivalently, it is eventually integrable on `s` in `l.small_sets`. -/ def integrable_at_filter (f : α → E) (l : filter α) (μ : measure α . volume_tac) := diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 7280d3d89bd72..3348bf5a9124c 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1227,6 +1227,27 @@ lemma set_lintegral_congr_fun {f g : α → ℝ≥0∞} {s : set α} (hs : measu ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ := by { rw lintegral_congr_ae, rw eventually_eq, rwa ae_restrict_iff' hs, } +lemma lintegral_to_real_le_lintegral_nnnorm (f : α → ℝ) : + ∫⁻ x, ennreal.of_real (f x) ∂μ ≤ ∫⁻ x, ∥f x∥₊ ∂μ := +begin + simp_rw ← of_real_norm_eq_coe_nnnorm, + refine lintegral_mono (λ x, ennreal.of_real_le_of_real _), + rw real.norm_eq_abs, + exact le_abs_self (f x), +end + +lemma lintegral_nnnorm_eq_of_ae_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ᵐ[μ] f) : + ∫⁻ x, ∥f x∥₊ ∂μ = ∫⁻ x, ennreal.of_real (f x) ∂μ := +begin + apply lintegral_congr_ae, + filter_upwards [h_nonneg] with x hx, + rw [real.nnnorm_of_nonneg hx, ennreal.of_real_eq_coe_nnreal hx], +end + +lemma lintegral_nnnorm_eq_of_nonneg {f : α → ℝ} (h_nonneg : 0 ≤ f) : + ∫⁻ x, ∥f x∥₊ ∂μ = ∫⁻ x, ennreal.of_real (f x) ∂μ := +lintegral_nnnorm_eq_of_ae_nonneg (filter.eventually_of_forall h_nonneg) + /-- Monotone convergence theorem -- sometimes called Beppo-Levi convergence. See `lintegral_supr_directed` for a more general form. -/ From 7e74869c61390ac9f421cb6fd60d78ffcb5b6517 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sun, 18 Sep 2022 10:17:08 +0000 Subject: [PATCH 297/828] doc(number_theory/legendre_symbol/jacobi_symbol): remove now spurious reference to 'subscript J' (#16541) This PR just fixes the docstring for the notation. The reference to "no subscript J in unicode" doesn't make sense anymore after the notation was changed to `J(a | b)`. --- src/number_theory/legendre_symbol/jacobi_symbol.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/number_theory/legendre_symbol/jacobi_symbol.lean b/src/number_theory/legendre_symbol/jacobi_symbol.lean index 9694294930a3f..3f0528362dc0d 100644 --- a/src/number_theory/legendre_symbol/jacobi_symbol.lean +++ b/src/number_theory/legendre_symbol/jacobi_symbol.lean @@ -67,7 +67,7 @@ Jacobi symbol when `b` is odd and gives less meaningful values when it is not (e is `1` when `b = 0`). This is called `jacobi_sym a b`. We define localized notation (locale `number_theory_symbols`) `J(a | b)` for the Jacobi -symbol `jacobi_sym a b`. (Unfortunately, there is no subscript "J" in unicode.) +symbol `jacobi_sym a b`. -/ open nat zmod From 827a83eb48a93ea15859686de67e38ee8e0bc2bc Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Sun, 18 Sep 2022 11:24:05 +0000 Subject: [PATCH 298/828] feat(archive/birthday): improve birthday problem (#16528) This adds a measure-theoretic interpretation of the birthday problem; not yet the one with independent uniform random variables, but some form of a step in that direction. It also removes a misleading statement from the YAML to make it clearer that `card_embedding_eq` is not all the work done on the Birthday problem on mathlib. cc @digama0 Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- .../93_birthday_problem.lean | 64 ++++++++++++++++++- docs/100.yaml | 1 - src/data/set/finite.lean | 4 ++ src/measure_theory/measure/measure_space.lean | 4 ++ src/probability/cond_count.lean | 12 +++- 5 files changed, 80 insertions(+), 5 deletions(-) diff --git a/archive/100-theorems-list/93_birthday_problem.lean b/archive/100-theorems-list/93_birthday_problem.lean index 8f11c73701025..921d534c83eab 100644 --- a/archive/100-theorems-list/93_birthday_problem.lean +++ b/archive/100-theorems-list/93_birthday_problem.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Rodriguez -/ import data.fintype.card_embedding -import tactic.norm_num +import probability.cond_count +import probability.notation /-! # Birthday Problem @@ -16,12 +17,69 @@ in terms of injective functions. The general result about `fintype.card (α ↪ uses is `fintype.card_embedding_eq`. -/ -local notation `‖` x `‖` := fintype.card x +local notation (name := finset.card) `|` x `|` := finset.card x +local notation (name := fintype.card) `‖` x `‖` := fintype.card x -/-- **Birthday Problem** -/ +/-- **Birthday Problem**: set cardinality interpretation. -/ theorem birthday : 2 * ‖fin 23 ↪ fin 365‖ < ‖fin 23 → fin 365‖ ∧ 2 * ‖fin 22 ↪ fin 365‖ > ‖fin 22 → fin 365‖ := begin simp only [nat.desc_factorial, fintype.card_fin, fintype.card_embedding_eq, fintype.card_fun], norm_num end + +section measure_theory + +open measure_theory probability_theory +open_locale probability_theory ennreal + +variables {n m : ℕ} + +/- In order for Lean to understand that we can take probabilities in `fin 23 → fin 365`, we must +tell Lean that there is a `measurable_space` structure on the space. Note that this instance +is only for `fin m` - Lean automatically figures out that the function space `fin n → fin m` +is _also_ measurable, by using `measurable_space.pi` -/ + +instance : measurable_space (fin m) := ⊤ + +/- We then endow the space with a canonical measure, which is called ℙ. +We define this to be the conditional counting measure. -/ +noncomputable instance : measure_space (fin n → fin m) := ⟨cond_count set.univ⟩ + +-- Singletons are measurable; therefore, as `fin n → fin m` is finite, all sets are measurable. +instance : measurable_singleton_class (fin n → fin m) := +⟨λ f, begin + convert measurable_set.pi set.finite_univ.countable + (show ∀ i, i ∈ set.univ → measurable_set ({f i} : set (fin m)), from λ _ _, trivial), + ext g, + simp only [function.funext_iff, set.mem_singleton_iff, set.mem_univ_pi], +end⟩ + +/- The canonical measure on `fin n → fin m` is a probability measure (except on an empty space). -/ +example : is_probability_measure (ℙ : measure (fin n → fin (m + 1))) := +cond_count_is_probability_measure set.finite_univ set.univ_nonempty + +lemma fin_fin.measure_apply {s : set $ fin n → fin m} : + ℙ s = (|s.to_finite.to_finset|) / ‖fin n → fin m‖ := +by erw [cond_count_univ, measure.count_apply_finite] + +/-- **Birthday Problem**: first probabilistic interpretation. -/ +theorem birthday_measure : ℙ {f : fin 23 → fin 365 | function.injective f} < 1 / 2 := +begin + -- most of this proof is essentially converting it to the same form as `birthday`. + rw [fin_fin.measure_apply], + generalize_proofs hfin, + have : |hfin.to_finset| = 42200819302092359872395663074908957253749760700776448000000, + { transitivity ‖fin 23 ↪ fin 365‖, + { simp_rw [←fintype.card_coe, set.finite.coe_sort_to_finset, set.coe_set_of], + exact fintype.card_congr (equiv.subtype_injective_equiv_embedding _ _) }, + { simp only [fintype.card_embedding_eq, fintype.card_fin, nat.desc_factorial], + norm_num } }, + rw [this, ennreal.lt_div_iff_mul_lt, mul_comm, mul_div, ennreal.div_lt_iff], + rotate, iterate 2 { right, norm_num }, iterate 2 { left, norm_num }, + norm_cast, + simp only [fintype.card_pi, fintype.card_fin], + norm_num +end + +end measure_theory diff --git a/docs/100.yaml b/docs/100.yaml index 50be57d8c9453..dde7c7548708f 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -379,7 +379,6 @@ title : Pick’s Theorem 93: title : The Birthday Problem - decl : fintype.card_embedding_eq links : mathlib archive: https://github.com/leanprover-community/mathlib/blob/master/archive/100-theorems-list/93_birthday_problem.lean author : Eric Rodriguez diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index dd47aab43790c..ddc4def3d29cd 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -141,6 +141,10 @@ by rw [← finset.coe_sort_coe _, h.coe_to_finset] @[simp] lemma finite_empty_to_finset (h : (∅ : set α).finite) : h.to_finset = ∅ := by rw [← finset.coe_inj, h.coe_to_finset, finset.coe_empty] +@[simp] lemma finite_univ_to_finset [fintype α] (h : (set.univ : set α).finite) : + h.to_finset = finset.univ := +finset.ext $ by simp + @[simp] lemma finite.to_finset_inj {s t : set α} {hs : s.finite} {ht : t.finite} : hs.to_finset = ht.to_finset ↔ s = t := by simp only [←finset.coe_inj, finite.coe_to_finset] diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index b250280907da5..74128d22e0b4f 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -2337,6 +2337,10 @@ begin exact ne_of_lt (measure_lt_top _ _) end +instance [finite α] [measurable_space α] : is_finite_measure (measure.count : measure α) := +⟨by { casesI nonempty_fintype α, + simpa [measure.count_apply, tsum_fintype] using (ennreal.nat_ne_top _).lt_top }⟩ + end is_finite_measure section is_probability_measure diff --git a/src/probability/cond_count.lean b/src/probability/cond_count.lean index 2a708cef40cb8..86f1aef08120e 100644 --- a/src/probability/cond_count.lean +++ b/src/probability/cond_count.lean @@ -62,6 +62,16 @@ begin simpa [cond_count, cond, measure.count_apply_infinite hs'] using h, end +lemma cond_count_univ [fintype Ω] {s : set Ω} : + cond_count set.univ s = measure.count s / fintype.card Ω := +begin + rw [cond_count, cond_apply _ measurable_set.univ, ←ennreal.div_eq_inv_mul, set.univ_inter], + congr', + rw [←finset.coe_univ, measure.count_apply, finset.univ.tsum_subtype' (λ _, (1 : ennreal))], + { simp [finset.card_univ] }, + { exact (@finset.coe_univ Ω _).symm ▸ measurable_set.univ } +end + variables [measurable_singleton_class Ω] lemma cond_count_is_probability_measure {s : set Ω} (hs : s.finite) (hs' : s.nonempty) : @@ -126,7 +136,7 @@ lemma cond_count_eq_zero_iff (hs : s.finite) : by simp [cond_count, cond_apply _ hs.measurable_set, measure.count_apply_eq_top, set.not_infinite.2 hs, measure.count_apply_finite _ (hs.inter_of_left _)] -lemma cond_count_univ (hs : s.finite) (hs' : s.nonempty) : +lemma cond_count_of_univ (hs : s.finite) (hs' : s.nonempty) : cond_count s set.univ = 1 := cond_count_eq_one_of hs hs' s.subset_univ From 8c9342fd8949ec8421ed6649685d5233a19e0c08 Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Sun, 18 Sep 2022 13:45:12 +0000 Subject: [PATCH 299/828] =?UTF-8?q?feat(probability/martingale/borel=5Fcan?= =?UTF-8?q?telli):=20L=C3=A9vy's=20generalized=20Borel-Cantelli=20(#16358)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR proves the one sided martingale bound and uses it to prove the Lévy's generalized Borel-Cantelli lemma. With the generalized Borel-Cantelli, the still missing second Borel-Cantelli lemma follows by choosing an appropriate filtration. Co-authored-by: RemyDegenne --- .../conditional_expectation/real.lean | 102 ++++- src/measure_theory/integral/bochner.lean | 61 +++ src/measure_theory/integral/set_integral.lean | 24 ++ src/probability/martingale/basic.lean | 10 + .../martingale/borel_cantelli.lean | 397 ++++++++++++++++++ src/probability/martingale/centering.lean | 27 +- src/topology/algebra/order/liminf_limsup.lean | 106 +++++ 7 files changed, 724 insertions(+), 3 deletions(-) create mode 100644 src/probability/martingale/borel_cantelli.lean diff --git a/src/measure_theory/function/conditional_expectation/real.lean b/src/measure_theory/function/conditional_expectation/real.lean index 7a0cdca071406..e9f7021abe50c 100644 --- a/src/measure_theory/function/conditional_expectation/real.lean +++ b/src/measure_theory/function/conditional_expectation/real.lean @@ -50,8 +50,8 @@ begin (signed_measure.measurable_rn_deriv _ _).strongly_measurable }, end -/-- TODO: this should be generalized and proved using Jensen's inequality -for the conditional expectation (not in mathlib yet) .-/ +-- TODO: the following couple of lemmas should be generalized and proved using Jensen's inequality +-- for the conditional expectation (not in mathlib yet) . lemma snorm_one_condexp_le_snorm (f : α → ℝ) : snorm (μ[f | m]) 1 μ ≤ snorm f 1 μ := begin @@ -90,6 +90,104 @@ begin end end +lemma integral_abs_condexp_le (f : α → ℝ) : + ∫ x, |μ[f | m] x| ∂μ ≤ ∫ x, |f x| ∂μ := +begin + by_cases hm : m ≤ m0, + swap, + { simp_rw [condexp_of_not_le hm, pi.zero_apply, abs_zero, integral_zero], + exact integral_nonneg (λ x, abs_nonneg _) }, + by_cases hfint : integrable f μ, + swap, + { simp only [condexp_undef hfint, pi.zero_apply, abs_zero, integral_const, + algebra.id.smul_eq_mul, mul_zero], + exact integral_nonneg (λ x, abs_nonneg _) }, + rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae], + { rw ennreal.to_real_le_to_real; + simp_rw [← real.norm_eq_abs, of_real_norm_eq_coe_nnnorm], + { rw [← snorm_one_eq_lintegral_nnnorm, ← snorm_one_eq_lintegral_nnnorm], + exact snorm_one_condexp_le_snorm _ }, + { exact ne_of_lt integrable_condexp.2 }, + { exact ne_of_lt hfint.2 } }, + { exact eventually_of_forall (λ x, abs_nonneg _) }, + { simp_rw ← real.norm_eq_abs, + exact hfint.1.norm }, + { exact eventually_of_forall (λ x, abs_nonneg _) }, + { simp_rw ← real.norm_eq_abs, + exact (strongly_measurable_condexp.mono hm).ae_strongly_measurable.norm } +end + +lemma set_integral_abs_condexp_le {s : set α} (hs : measurable_set[m] s) (f : α → ℝ) : + ∫ x in s, |μ[f | m] x| ∂μ ≤ ∫ x in s, |f x| ∂μ := +begin + by_cases hnm : m ≤ m0, + swap, + { simp_rw [condexp_of_not_le hnm, pi.zero_apply, abs_zero, integral_zero], + exact integral_nonneg (λ x, abs_nonneg _) }, + by_cases hfint : integrable f μ, + swap, + { simp only [condexp_undef hfint, pi.zero_apply, abs_zero, integral_const, + algebra.id.smul_eq_mul, mul_zero], + exact integral_nonneg (λ x, abs_nonneg _) }, + have : ∫ x in s, |μ[f | m] x| ∂μ = ∫ x, |μ[s.indicator f | m] x| ∂μ, + { rw ← integral_indicator, + swap, { exact hnm _ hs }, + refine integral_congr_ae _, + have : (λ x, |μ[s.indicator f | m] x|) =ᵐ[μ] λ x, |s.indicator (μ[f | m]) x| := + eventually_eq.fun_comp (condexp_indicator hfint hs) _, + refine eventually_eq.trans (eventually_of_forall $ λ x, _) this.symm, + rw [← real.norm_eq_abs, norm_indicator_eq_indicator_norm], + refl }, + rw [this, ← integral_indicator], + swap, { exact hnm _ hs }, + refine (integral_abs_condexp_le _).trans (le_of_eq $ integral_congr_ae $ + eventually_of_forall $ λ x, _), + rw [← real.norm_eq_abs, norm_indicator_eq_indicator_norm], + refl, +end + +/-- If the real valued function `f` is bounded almost everywhere by `R`, then so is its conditional +expectation. -/ +lemma ae_bdd_condexp_of_ae_bdd {R : ℝ≥0} {f : α → ℝ} + (hbdd : ∀ᵐ x ∂μ, |f x| ≤ R) : + ∀ᵐ x ∂μ, |μ[f | m] x| ≤ R := +begin + by_cases hnm : m ≤ m0, + swap, + { simp_rw [condexp_of_not_le hnm, pi.zero_apply, abs_zero], + refine eventually_of_forall (λ x, R.coe_nonneg) }, + by_cases hfint : integrable f μ, + swap, + { simp_rw [condexp_undef hfint], + filter_upwards [hbdd] with x hx, + rw [pi.zero_apply, abs_zero], + exact (abs_nonneg _).trans hx }, + by_contra h, + change μ _ ≠ 0 at h, + simp only [← zero_lt_iff, set.compl_def, set.mem_set_of_eq, not_le] at h, + suffices : (μ {x | ↑R < |μ[f | m] x|}).to_real * ↑R < (μ {x | ↑R < |μ[f | m] x|}).to_real * ↑R, + { exact this.ne rfl }, + refine lt_of_lt_of_le (set_integral_gt_gt R.coe_nonneg _ _ h.ne.symm) _, + { simp_rw [← real.norm_eq_abs], + exact (strongly_measurable_condexp.mono hnm).measurable.norm }, + { exact integrable_condexp.abs.integrable_on }, + refine (set_integral_abs_condexp_le _ _).trans _, + { simp_rw [← real.norm_eq_abs], + exact @measurable_set_lt _ _ _ _ _ m _ _ _ _ _ measurable_const + strongly_measurable_condexp.norm.measurable }, + simp only [← smul_eq_mul, ← set_integral_const, nnreal.val_eq_coe, + is_R_or_C.coe_real_eq_id, id.def], + refine set_integral_mono_ae hfint.abs.integrable_on _ _, + { refine ⟨ae_strongly_measurable_const, lt_of_le_of_lt _ + (integrable_condexp.integrable_on : integrable_on (μ[f | m]) {x | ↑R < |μ[f | m] x|} μ).2⟩, + refine set_lintegral_mono (measurable.nnnorm _).coe_nnreal_ennreal + (strongly_measurable_condexp.mono hnm).measurable.nnnorm.coe_nnreal_ennreal (λ x hx, _), + { exact measurable_const }, + { rw [ennreal.coe_le_coe, real.nnnorm_of_nonneg R.coe_nonneg], + exact subtype.mk_le_mk.2 (le_of_lt hx) } }, + { exact hbdd }, +end + /-- Given a integrable function `g`, the conditional expectations of `g` with respect to a sequence of sub-σ-algebras is uniformly integrable. -/ lemma integrable.uniform_integrable_condexp {ι : Type*} [is_finite_measure μ] diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 3270c6a05ec81..3cc528844a562 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -1611,4 +1611,65 @@ lemma ae_le_trim_iff end integral_trim +section snorm_bound + +variables {m0 : measurable_space α} {μ : measure α} + +lemma snorm_one_le_of_le {r : ℝ≥0} {f : α → ℝ} + (hfint : integrable f μ) (hfint' : 0 ≤ ∫ x, f x ∂μ) (hf : ∀ᵐ ω ∂μ, f ω ≤ r) : + snorm f 1 μ ≤ 2 * μ set.univ * r := +begin + by_cases hr : r = 0, + { suffices : f =ᵐ[μ] 0, + { rw [snorm_congr_ae this, snorm_zero, hr, ennreal.coe_zero, mul_zero], + exact le_rfl }, + rw [hr, nonneg.coe_zero] at hf, + have hnegf : ∫ x, -f x ∂μ = 0, + { rw [integral_neg, neg_eq_zero], + exact le_antisymm (integral_nonpos_of_ae hf) hfint' }, + have := (integral_eq_zero_iff_of_nonneg_ae _ hfint.neg).1 hnegf, + { filter_upwards [this] with ω hω, + rwa [pi.neg_apply, pi.zero_apply, neg_eq_zero] at hω }, + { filter_upwards [hf] with ω hω, + rwa [pi.zero_apply, pi.neg_apply, right.nonneg_neg_iff] } }, + by_cases hμ : is_finite_measure μ, + swap, + { have : μ set.univ = ∞, + { by_contra hμ', + exact hμ (is_finite_measure.mk $ lt_top_iff_ne_top.2 hμ') }, + rw [this, ennreal.mul_top, if_neg, ennreal.top_mul, if_neg], + { exact le_top }, + { simp [hr] }, + { norm_num } }, + haveI := hμ, + rw [integral_eq_integral_pos_part_sub_integral_neg_part hfint, sub_nonneg] at hfint', + have hposbdd : ∫ ω, max (f ω) 0 ∂μ ≤ (μ set.univ).to_real • r, + { rw ← integral_const, + refine integral_mono_ae hfint.real_to_nnreal (integrable_const r) _, + filter_upwards [hf] with ω hω using real.to_nnreal_le_iff_le_coe.2 hω }, + rw [mem_ℒp.snorm_eq_integral_rpow_norm one_ne_zero ennreal.one_ne_top + (mem_ℒp_one_iff_integrable.2 hfint), + ennreal.of_real_le_iff_le_to_real (ennreal.mul_ne_top + (ennreal.mul_ne_top ennreal.two_ne_top $ @measure_ne_top _ _ _ hμ _) ennreal.coe_ne_top)], + simp_rw [ennreal.one_to_real, _root_.inv_one, real.rpow_one, real.norm_eq_abs, + ← max_zero_add_max_neg_zero_eq_abs_self, ← real.coe_to_nnreal'], + rw integral_add hfint.real_to_nnreal, + { simp only [real.coe_to_nnreal', ennreal.to_real_mul, ennreal.to_real_bit0, + ennreal.one_to_real, ennreal.coe_to_real] at hfint' ⊢, + refine (add_le_add_left hfint' _).trans _, + rwa [← two_mul, mul_assoc, mul_le_mul_left (two_pos : (0 : ℝ) < 2)] }, + { exact hfint.neg.sup (integrable_zero _ _ μ) } +end + +lemma snorm_one_le_of_le' {r : ℝ} {f : α → ℝ} + (hfint : integrable f μ) (hfint' : 0 ≤ ∫ x, f x ∂μ) (hf : ∀ᵐ ω ∂μ, f ω ≤ r) : + snorm f 1 μ ≤ 2 * μ set.univ * ennreal.of_real r := +begin + refine snorm_one_le_of_le hfint hfint' _, + simp only [real.coe_to_nnreal', le_max_iff], + filter_upwards [hf] with ω hω using or.inl hω, +end + +end snorm_bound + end measure_theory diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 0fa648f478874..a6ed1044d17f8 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -396,6 +396,30 @@ begin exact hfi.ae_strongly_measurable.ae_measurable.null_measurable (measurable_set_singleton 0).compl end +lemma set_integral_gt_gt {R : ℝ} {f : α → ℝ} (hR : 0 ≤ R) (hfm : measurable f) + (hfint : integrable_on f {x | ↑R < f x} μ) (hμ : μ {x | ↑R < f x} ≠ 0): + (μ {x | ↑R < f x}).to_real * R < ∫ x in {x | ↑R < f x}, f x ∂μ := +begin + have : integrable_on (λ x, R) {x | ↑R < f x} μ, + { refine ⟨ae_strongly_measurable_const, lt_of_le_of_lt _ hfint.2⟩, + refine set_lintegral_mono (measurable.nnnorm _).coe_nnreal_ennreal + hfm.nnnorm.coe_nnreal_ennreal (λ x hx, _), + { exact measurable_const }, + { simp only [ennreal.coe_le_coe, real.nnnorm_of_nonneg hR, + real.nnnorm_of_nonneg (hR.trans $ le_of_lt hx), subtype.mk_le_mk], + exact le_of_lt hx } }, + rw [← sub_pos, ← smul_eq_mul, ← set_integral_const, ← integral_sub hfint this, + set_integral_pos_iff_support_of_nonneg_ae], + { rw ← zero_lt_iff at hμ, + rwa set.inter_eq_self_of_subset_right, + exact λ x hx, ne.symm (ne_of_lt $ sub_pos.2 hx) }, + { change ∀ᵐ x ∂(μ.restrict _), _, + rw ae_restrict_iff, + { exact eventually_of_forall (λ x hx, sub_nonneg.2 $ le_of_lt hx) }, + { exact measurable_set_le measurable_zero (hfm.sub measurable_const) } }, + { exact integrable.sub hfint this }, +end + lemma set_integral_trim {α} {m m0 : measurable_space α} {μ : measure α} (hm : m ≤ m0) {f : α → E} (hf_meas : strongly_measurable[m] f) {s : set α} (hs : measurable_set[m] s) : ∫ x in s, f x ∂μ = ∫ x in s, f x ∂(μ.trim hm) := diff --git a/src/probability/martingale/basic.lean b/src/probability/martingale/basic.lean index af312b7e84da4..ba824d8012436 100644 --- a/src/probability/martingale/basic.lean +++ b/src/probability/martingale/basic.lean @@ -66,6 +66,16 @@ lemma martingale_const (ℱ : filtration ι m0) (μ : measure Ω) [is_finite_mea martingale (λ _ _, x) ℱ μ := ⟨adapted_const ℱ _, λ i j hij, by rw condexp_const (ℱ.le _)⟩ +lemma martingale_const_fun [order_bot ι] + (ℱ : filtration ι m0) (μ : measure Ω) [is_finite_measure μ] + {f : Ω → E} (hf : strongly_measurable[ℱ ⊥] f) (hfint : integrable f μ) : + martingale (λ _, f) ℱ μ := +begin + refine ⟨λ i, hf.mono $ ℱ.mono bot_le, λ i j hij, _⟩, + rw condexp_of_strongly_measurable (ℱ.le _) (hf.mono $ ℱ.mono bot_le) hfint, + apply_instance, +end + variables (E) lemma martingale_zero (ℱ : filtration ι m0) (μ : measure Ω) : martingale (0 : ι → Ω → E) ℱ μ := diff --git a/src/probability/martingale/borel_cantelli.lean b/src/probability/martingale/borel_cantelli.lean new file mode 100644 index 0000000000000..cd5664dcbcf35 --- /dev/null +++ b/src/probability/martingale/borel_cantelli.lean @@ -0,0 +1,397 @@ +/- +Copyright (c) 2022 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying +-/ +import probability.martingale.convergence +import probability.martingale.optional_stopping +import probability.martingale.centering +import probability.conditional_expectation + +/-! + +# Generalized Borel-Cantelli lemma + +This file proves Lévy's generalized Borel-Cantelli lemma which is a generalization of the +Borel-Cantelli lemmas. With this generalization, one can easily deduce the Borel-Cantelli lemmas +by choosing appropriate filtrations. This file also contains the one sided martingale bound which +is required to prove the generalized Borel-Cantelli. + +## Main results + +- `measure_theory.submartingale.bdd_above_iff_exists_tendsto`: the one sided martingale bound: given + a submartingale `f` with uniformly bounded differences, the set for which `f` converges is almost + everywhere equal to the set for which it is bounded. +- `measure_theory.ae_mem_limsup_at_top_iff`: Lévy's generalized Borel-Cantelli: + given a filtration `ℱ` and a sequence of sets `s` such that `s n ∈ ℱ n` for all `n`, + `limsup at_top s` is almost everywhere equal to the set for which `∑ ℙ[s (n + 1)∣ℱ n] = ∞`. + +## TODO + +Prove the missing second Borel-Cantelli lemma using this generalized version. + +-/ + +open filter +open_locale nnreal ennreal measure_theory probability_theory big_operators topological_space + +namespace measure_theory + +variables {Ω : Type*} {m0 : measurable_space Ω} {μ : measure Ω} + {ℱ : filtration ℕ m0} {f : ℕ → Ω → ℝ} {ω : Ω} + +/-! +### One sided martingale bound +-/ + +-- TODO: `least_ge` should be defined taking values in `with_top ℕ` once the `stopped_process` +-- refactor is complete +/-- `least_ge f r n` is the stopping time corresponding to the first time `f ≥ r`. -/ +noncomputable +def least_ge (f : ℕ → Ω → ℝ) (r : ℝ) (n : ℕ) := hitting f (set.Ici r) 0 n + +lemma adapted.is_stopping_time_least_ge (r : ℝ) (n : ℕ) (hf : adapted ℱ f) : + is_stopping_time ℱ (least_ge f r n) := +hitting_is_stopping_time hf measurable_set_Ici + +lemma least_ge_le {i : ℕ} {r : ℝ} (ω : Ω) : least_ge f r i ω ≤ i := +hitting_le ω + +-- The following four lemmas shows `least_ge` behaves like a stopped process. Ideally we should +-- define `least_ge` as a stopping time and take its stopped process. However, we can't do that +-- with our current definition since a stopping time takes only finite indicies. An upcomming +-- refactor should hopefully make it possible to have stopping times taking infinity as a value +lemma least_ge_mono {n m : ℕ} (hnm : n ≤ m) (r : ℝ) (ω : Ω) : + least_ge f r n ω ≤ least_ge f r m ω := +hitting_mono hnm + +lemma least_ge_eq_min (π : Ω → ℕ) (r : ℝ) (ω : Ω) + {n : ℕ} (hπn : ∀ ω, π ω ≤ n) : + least_ge f r (π ω) ω = min (π ω) (least_ge f r n ω) := +begin + classical, + refine le_antisymm (le_min (least_ge_le _) (least_ge_mono (hπn ω) r ω)) _, + by_cases hle : π ω ≤ least_ge f r n ω, + { rw [min_eq_left hle, least_ge], + by_cases h : ∃ j ∈ set.Icc 0 (π ω), f j ω ∈ set.Ici r, + { refine hle.trans (eq.le _), + rw [least_ge, ← hitting_eq_hitting_of_exists (hπn ω) h] }, + { simp only [hitting, if_neg h] } }, + { rw [min_eq_right (not_le.1 hle).le, least_ge, least_ge, + ← hitting_eq_hitting_of_exists (hπn ω) _], + rw [not_le, least_ge, hitting_lt_iff _ (hπn ω)] at hle, + exact let ⟨j, hj₁, hj₂⟩ := hle in ⟨j, ⟨hj₁.1, hj₁.2.le⟩, hj₂⟩ } +end + +lemma stopped_value_stopped_value_least_ge (f : ℕ → Ω → ℝ) (π : Ω → ℕ) (r : ℝ) + {n : ℕ} (hπn : ∀ ω, π ω ≤ n) : + stopped_value (λ i, stopped_value f (least_ge f r i)) π + = stopped_value (stopped_process f (least_ge f r n)) π := +by { ext1 ω, simp_rw [stopped_process, stopped_value], rw least_ge_eq_min _ _ _ hπn, } + +lemma submartingale.stopped_value_least_ge [is_finite_measure μ] + (hf : submartingale f ℱ μ) (r : ℝ) : + submartingale (λ i, stopped_value f (least_ge f r i)) ℱ μ := +begin + rw submartingale_iff_expected_stopped_value_mono, + { intros σ π hσ hπ hσ_le_π hπ_bdd, + obtain ⟨n, hπ_le_n⟩ := hπ_bdd, + simp_rw stopped_value_stopped_value_least_ge f σ r (λ i, (hσ_le_π i).trans (hπ_le_n i)), + simp_rw stopped_value_stopped_value_least_ge f π r hπ_le_n, + refine hf.expected_stopped_value_mono _ _ _ (λ ω, (min_le_left _ _).trans (hπ_le_n ω)), + { exact hσ.min (hf.adapted.is_stopping_time_least_ge _ _), }, + { exact hπ.min (hf.adapted.is_stopping_time_least_ge _ _), }, + { exact λ ω, min_le_min (hσ_le_π ω) le_rfl, }, }, + { exact λ i, strongly_measurable_stopped_value_of_le hf.adapted.prog_measurable_of_nat + (hf.adapted.is_stopping_time_least_ge _ _) least_ge_le, }, + { exact λ i, integrable_stopped_value _ ((hf.adapted.is_stopping_time_least_ge _ _)) + (hf.integrable) least_ge_le, }, +end + +variables {r : ℝ} {R : ℝ≥0} + +lemma norm_stopped_value_least_ge_le (hr : 0 ≤ r) (hf0 : f 0 = 0) + (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) (i : ℕ) : + ∀ᵐ ω ∂μ, stopped_value f (least_ge f r i) ω ≤ r + R := +begin + filter_upwards [hbdd] with ω hbddω, + change f (least_ge f r i ω) ω ≤ r + R, + by_cases heq : least_ge f r i ω = 0, + { rw [heq, hf0, pi.zero_apply], + exact add_nonneg hr R.coe_nonneg }, + { obtain ⟨k, hk⟩ := nat.exists_eq_succ_of_ne_zero heq, + rw [hk, add_comm, ← sub_le_iff_le_add], + have := not_mem_of_lt_hitting (hk.symm ▸ k.lt_succ_self : k < least_ge f r i ω) (zero_le _), + simp only [set.mem_union_eq, set.mem_Iic, set.mem_Ici, not_or_distrib, not_le] at this, + exact (sub_lt_sub_left this _).le.trans ((le_abs_self _).trans (hbddω _)) } +end + +lemma submartingale.stopped_value_least_ge_snorm_le [is_finite_measure μ] + (hf : submartingale f ℱ μ) (hr : 0 ≤ r) (hf0 : f 0 = 0) + (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) (i : ℕ) : + snorm (stopped_value f (least_ge f r i)) 1 μ ≤ 2 * μ set.univ * ennreal.of_real (r + R) := +begin + refine snorm_one_le_of_le' ((hf.stopped_value_least_ge r).integrable _) _ + (norm_stopped_value_least_ge_le hr hf0 hbdd i), + rw ← integral_univ, + refine le_trans _ ((hf.stopped_value_least_ge r).set_integral_le (zero_le _) + measurable_set.univ), + simp_rw [stopped_value, least_ge, hitting_of_le le_rfl, hf0, integral_zero'] +end + +lemma submartingale.stopped_value_least_ge_snorm_le' [is_finite_measure μ] + (hf : submartingale f ℱ μ) (hr : 0 ≤ r) (hf0 : f 0 = 0) + (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) (i : ℕ) : + snorm (stopped_value f (least_ge f r i)) 1 μ ≤ + ennreal.to_nnreal (2 * μ set.univ * ennreal.of_real (r + R)) := +begin + refine (hf.stopped_value_least_ge_snorm_le hr hf0 hbdd i).trans _, + simp [ennreal.coe_to_nnreal (measure_ne_top μ _), ennreal.coe_to_nnreal], +end + +/-- This lemma is superceded by `submartingale.bdd_above_iff_exists_tendsto`. -/ +lemma submartingale.exists_tendsto_of_abs_bdd_above_aux [is_finite_measure μ] + (hf : submartingale f ℱ μ) (hf0 : f 0 = 0) + (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : + ∀ᵐ ω ∂μ, bdd_above (set.range $ λ n, f n ω) → ∃ c, tendsto (λ n, f n ω) at_top (𝓝 c) := +begin + have ht : ∀ᵐ ω ∂μ, ∀ i : ℕ, ∃ c, tendsto (λ n, stopped_value f (least_ge f i n) ω) at_top (𝓝 c), + { rw ae_all_iff, + exact λ i, submartingale.exists_ae_tendsto_of_bdd (hf.stopped_value_least_ge i) + (hf.stopped_value_least_ge_snorm_le' i.cast_nonneg hf0 hbdd) }, + filter_upwards [ht] with ω hω hωb, + rw bdd_above at hωb, + obtain ⟨i, hi⟩ := exists_nat_gt hωb.some, + have hib : ∀ n, f n ω < i, + { intro n, + exact lt_of_le_of_lt ((mem_upper_bounds.1 hωb.some_mem) _ ⟨n, rfl⟩) hi }, + have heq : ∀ n, stopped_value f (least_ge f i n) ω = f n ω, + { intro n, + rw [least_ge, hitting, stopped_value], + simp only, + rw if_neg, + simp only [set.mem_Icc, set.mem_union, set.mem_Ici], + push_neg, + exact λ j _, hib j }, + simp only [← heq, hω i], +end + +lemma submartingale.bdd_above_iff_exists_tendsto_aux [is_finite_measure μ] + (hf : submartingale f ℱ μ) (hf0 : f 0 = 0) + (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : + ∀ᵐ ω ∂μ, bdd_above (set.range $ λ n, f n ω) ↔ ∃ c, tendsto (λ n, f n ω) at_top (𝓝 c) := +by filter_upwards [hf.exists_tendsto_of_abs_bdd_above_aux hf0 hbdd] with ω hω using + ⟨hω, λ ⟨c, hc⟩, hc.bdd_above_range⟩ + +/-- One sided martingale bound: If `f` is a submartingale which has uniformly bounded differences, +then for almost every `ω`, `f n ω` is bounded above (in `n`) if and only if it converges. -/ +lemma submartingale.bdd_above_iff_exists_tendsto [is_finite_measure μ] + (hf : submartingale f ℱ μ) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : + ∀ᵐ ω ∂μ, bdd_above (set.range $ λ n, f n ω) ↔ ∃ c, tendsto (λ n, f n ω) at_top (𝓝 c) := +begin + set g : ℕ → Ω → ℝ := λ n ω, f n ω - f 0 ω with hgdef, + have hg : submartingale g ℱ μ := + hf.sub_martingale (martingale_const_fun _ _ (hf.adapted 0) (hf.integrable 0)), + have hg0 : g 0 = 0, + { ext ω, + simp only [hgdef, sub_self, pi.zero_apply] }, + have hgbdd : ∀ᵐ ω ∂μ, ∀ (i : ℕ), |g (i + 1) ω - g i ω| ≤ ↑R, + { simpa only [sub_sub_sub_cancel_right] }, + filter_upwards [hg.bdd_above_iff_exists_tendsto_aux hg0 hgbdd] with ω hω, + convert hω using 1; rw eq_iff_iff, + { simp only [hgdef], + refine ⟨λ h, _, λ h, _⟩; + obtain ⟨b, hb⟩ := h; + refine ⟨b + |f 0 ω|, λ y hy, _⟩; + obtain ⟨n, rfl⟩ := hy, + { simp_rw [sub_eq_add_neg], + exact add_le_add (hb ⟨n, rfl⟩) (neg_le_abs_self _) }, + { exact sub_le_iff_le_add.1 (le_trans (sub_le_sub_left (le_abs_self _) _) (hb ⟨n, rfl⟩)) } }, + { simp only [hgdef], + refine ⟨λ h, _, λ h, _⟩; + obtain ⟨c, hc⟩ := h, + { exact ⟨c - f 0 ω, hc.sub_const _⟩ }, + { refine ⟨c + f 0 ω, _⟩, + have := hc.add_const (f 0 ω), + simpa only [sub_add_cancel] } } +end + +/-! +### Lévy's generalization of the Borel-Cantelli lemma + +Lévy's generalization of the Borel-Cantelli lemma states that: given a natural number indexed +filtration $(\mathcal{F}_n)$, and a sequence of sets $(s_n)$ such that for all +$n$, $s_n \in \mathcal{F}_n$, $limsup_n s_n$ is almost everywhere equal to the set for which +$\sum_n \mathbb{P}[s_n \mid \mathcal{F}_n] = \infty$. + +The proof strategy follows by constructing a martingale satisfying the one sided martingale bound. +In particular, we define +$$ + f_n := \sum_{k < n} \mathbf{1}_{s_{n + 1}} - \mathbb{P}[s_{n + 1} \mid \mathcal{F}_n]. +$$ +Then, as a martingale is both a sub and a super-martingale, the set for which it is unbounded from +above must agree with the set for which it is unbounded from below almost everywhere. Thus, it +can only converge to $\pm \infty$ with probability 0. Thus, by considering +$$ + \limsup_n s_n = \{\sum_n \mathbf{1}_{s_n} = \infty\} +$$ +almost everywhere, the result follows. +-/ + +lemma martingale.bdd_above_range_iff_bdd_below_range [is_finite_measure μ] + (hf : martingale f ℱ μ) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : + ∀ᵐ ω ∂μ, bdd_above (set.range (λ n, f n ω)) ↔ bdd_below (set.range (λ n, f n ω)) := +begin + have hbdd' : ∀ᵐ ω ∂μ, ∀ i, |(-f) (i + 1) ω - (-f) i ω| ≤ R, + { filter_upwards [hbdd] with ω hω i, + erw [← abs_neg, neg_sub, sub_neg_eq_add, neg_add_eq_sub], + exact hω i }, + have hup := hf.submartingale.bdd_above_iff_exists_tendsto hbdd, + have hdown := hf.neg.submartingale.bdd_above_iff_exists_tendsto hbdd', + filter_upwards [hup, hdown] with ω hω₁ hω₂, + have : (∃ c, tendsto (λ n, f n ω) at_top (𝓝 c)) ↔ ∃ c, tendsto (λ n, (-f) n ω) at_top (𝓝 c), + { split; rintro ⟨c, hc⟩, + { exact ⟨-c, hc.neg⟩ }, + { refine ⟨-c, _⟩, + convert hc.neg, + simp only [neg_neg, pi.neg_apply] } }, + rw [hω₁, this, ← hω₂], + split; rintro ⟨c, hc⟩; refine ⟨-c, λ ω hω, _⟩, + { rw mem_upper_bounds at hc, + refine neg_le.2 (hc _ _), + simpa only [pi.neg_apply, set.mem_range, neg_inj] }, + { rw mem_lower_bounds at hc, + simp_rw [set.mem_range, pi.neg_apply, neg_eq_iff_neg_eq, eq_comm] at hω, + refine le_neg.1 (hc _ _), + simpa only [set.mem_range] } +end + +lemma martingale.ae_not_tendsto_at_top_at_top [is_finite_measure μ] + (hf : martingale f ℱ μ) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : + ∀ᵐ ω ∂μ, ¬ tendsto (λ n, f n ω) at_top at_top := +by filter_upwards [hf.bdd_above_range_iff_bdd_below_range hbdd] with ω hω htop using + unbounded_of_tendsto_at_top htop (hω.2 $ bdd_below_range_of_tendsto_at_top_at_top htop) + +lemma martingale.ae_not_tendsto_at_top_at_bot [is_finite_measure μ] + (hf : martingale f ℱ μ) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : + ∀ᵐ ω ∂μ, ¬ tendsto (λ n, f n ω) at_top at_bot := +by filter_upwards [hf.bdd_above_range_iff_bdd_below_range hbdd] with ω hω htop using + unbounded_of_tendsto_at_bot htop (hω.1 $ bdd_above_range_of_tendsto_at_top_at_bot htop) + +namespace borel_cantelli + +/-- Auxiliary definition required to prove Lévy's generalization of the Borel-Cantelli lemmas for +which we will take the martingale part. -/ +noncomputable +def process (s : ℕ → set Ω) (n : ℕ) : Ω → ℝ := +∑ k in finset.range n, (s (k + 1)).indicator 1 + +variables {s : ℕ → set Ω} + +lemma process_zero : process s 0 = 0 := +by rw [process, finset.range_zero, finset.sum_empty] + +lemma adapted_process (hs : ∀ n, measurable_set[ℱ n] (s n)) : + adapted ℱ (process s) := +λ n, finset.strongly_measurable_sum' _ $ λ k hk, strongly_measurable_one.indicator $ + ℱ.mono (finset.mem_range.1 hk) _ $ hs _ + +lemma martingale_part_process_ae_eq (ℱ : filtration ℕ m0) (μ : measure Ω) (s : ℕ → set Ω) (n : ℕ) : + martingale_part ℱ μ (process s) n = + ∑ k in finset.range n, ((s (k + 1)).indicator 1 - μ[(s (k + 1)).indicator 1 | ℱ k]) := +begin + simp only [martingale_part_eq_sum, process_zero, zero_add], + refine finset.sum_congr rfl (λ k hk, _), + simp only [process, finset.sum_range_succ_sub_sum], +end + +lemma predictable_part_process_ae_eq (ℱ : filtration ℕ m0) (μ : measure Ω) (s : ℕ → set Ω) (n : ℕ) : + predictable_part ℱ μ (process s) n = + ∑ k in finset.range n, μ[(s (k + 1)).indicator (1 : Ω → ℝ) | ℱ k] := +begin + have := martingale_part_process_ae_eq ℱ μ s n, + simp_rw [martingale_part, process, finset.sum_sub_distrib] at this, + exact sub_right_injective this, +end + +lemma process_difference_le (s : ℕ → set Ω) (ω : Ω) (n : ℕ) : + |process s (n + 1) ω - process s n ω| ≤ (1 : ℝ≥0) := +begin + rw [nonneg.coe_one, process, process, finset.sum_apply, finset.sum_apply, + finset.sum_range_succ_sub_sum, ← real.norm_eq_abs, norm_indicator_eq_indicator_norm], + refine set.indicator_le' (λ _ _, _) (λ _ _, zero_le_one) _, + rw [pi.one_apply, norm_one] +end + +lemma integrable_process (μ : measure Ω) [is_finite_measure μ] + (hs : ∀ n, measurable_set[ℱ n] (s n)) (n : ℕ) : + integrable (process s n) μ := +integrable_finset_sum' _ $ λ k hk, integrable_on.indicator (integrable_const 1) $ ℱ.le _ _ $ hs _ + +end borel_cantelli + +open borel_cantelli + +/-- An a.e. monotone adapted process `f` with uniformly bounded differences converges to `+∞` if +and only if its predictable part also converges to `+∞`. -/ +lemma tendsto_sum_indicator_at_top_iff [is_finite_measure μ] + (hfmono : ∀ᵐ ω ∂μ, ∀ n, f n ω ≤ f (n + 1) ω) + (hf : adapted ℱ f) (hint : ∀ n, integrable (f n) μ) + (hbdd : ∀ᵐ ω ∂μ, ∀ n, |f (n + 1) ω - f n ω| ≤ R) : + ∀ᵐ ω ∂μ, tendsto (λ n, f n ω) at_top at_top ↔ + tendsto (λ n, predictable_part ℱ μ f n ω) at_top at_top := +begin + have h₁ := (martingale_martingale_part hf hint).ae_not_tendsto_at_top_at_top + (martingale_part_bdd_difference ℱ hbdd), + have h₂ := (martingale_martingale_part hf hint).ae_not_tendsto_at_top_at_bot + (martingale_part_bdd_difference ℱ hbdd), + have h₃ : ∀ᵐ ω ∂μ, ∀ n, 0 ≤ μ[f (n + 1) - f n | ℱ n] ω, + { refine ae_all_iff.2 (λ n, condexp_nonneg _), + filter_upwards [ae_all_iff.1 hfmono n] with ω hω using sub_nonneg.2 hω }, + filter_upwards [h₁, h₂, h₃, hfmono] with ω hω₁ hω₂ hω₃ hω₄, + split; intro ht, + { refine tendsto_at_top_at_top_of_monotone' _ _, + { intros n m hnm, + simp only [predictable_part, finset.sum_apply], + refine finset.sum_mono_set_of_nonneg hω₃ (finset.range_mono hnm) }, + rintro ⟨b, hbdd⟩, + rw ← tendsto_neg_at_bot_iff at ht, + simp only [martingale_part, sub_eq_add_neg] at hω₁, + exact hω₁ (tendsto_at_top_add_right_of_le _ (-b) + (tendsto_neg_at_bot_iff.1 ht) $ λ n, neg_le_neg (hbdd ⟨n, rfl⟩)) }, + { refine tendsto_at_top_at_top_of_monotone' (monotone_nat_of_le_succ hω₄) _, + rintro ⟨b, hbdd⟩, + exact hω₂ (tendsto_at_bot_add_left_of_ge _ b (λ n, hbdd ⟨n, rfl⟩) $ + tendsto_neg_at_bot_iff.2 ht) }, +end + +open borel_cantelli + +lemma tendsto_sum_indicator_at_top_iff' [is_finite_measure μ] + {s : ℕ → set Ω} (hs : ∀ n, measurable_set[ℱ n] (s n)) : + ∀ᵐ ω ∂μ, + tendsto (λ n, ∑ k in finset.range n, (s (k + 1)).indicator (1 : Ω → ℝ) ω) at_top at_top ↔ + tendsto (λ n, ∑ k in finset.range n, μ[(s (k + 1)).indicator (1 : Ω → ℝ) | ℱ k] ω) + at_top at_top := +begin + have := tendsto_sum_indicator_at_top_iff (eventually_of_forall $ λ ω n, _) (adapted_process hs) + (integrable_process μ hs) (eventually_of_forall $ process_difference_le s), + swap, + { rw [process, process, ← sub_nonneg, finset.sum_apply, finset.sum_apply, + finset.sum_range_succ_sub_sum], + exact set.indicator_nonneg (λ _ _, zero_le_one) _ }, + simp_rw [process, predictable_part_process_ae_eq] at this, + simpa using this, +end + +/-- **Lévy's generalization of the Borel-Cantelli lemma**: given a sequence of sets `s` and a +filtration `ℱ` such that for all `n`, `s n` is `ℱ n`-measurable, `at_top.limsup s` is almost +everywhere equal to the set for which `∑ k, ℙ(s (k + 1) | ℱ k) = ∞`. -/ +theorem ae_mem_limsup_at_top_iff [is_finite_measure μ] + {s : ℕ → set Ω} (hs : ∀ n, measurable_set[ℱ n] (s n)) : + ∀ᵐ ω ∂μ, ω ∈ limsup at_top s ↔ + tendsto (λ n, ∑ k in finset.range n, μ[(s (k + 1)).indicator (1 : Ω → ℝ) | ℱ k] ω) + at_top at_top := +(limsup_eq_tendsto_sum_indicator_at_top ℝ s).symm ▸ tendsto_sum_indicator_at_top_iff' hs + +end measure_theory diff --git a/src/probability/martingale/centering.lean b/src/probability/martingale/centering.lean index 9e244d8b878b1..67cbf109ecc9c 100644 --- a/src/probability/martingale/centering.lean +++ b/src/probability/martingale/centering.lean @@ -64,7 +64,7 @@ def martingale_part {m0 : measurable_space Ω} (ℱ : filtration ℕ m0) (μ : measure Ω) (f : ℕ → Ω → E) : ℕ → Ω → E := λ n, f n - predictable_part ℱ μ f n -lemma martingale_part_add_predictable_part (f : ℕ → Ω → E) : +lemma martingale_part_add_predictable_part (ℱ : filtration ℕ m0) (μ : measure Ω) (f : ℕ → Ω → E) : martingale_part ℱ μ f + predictable_part ℱ μ f = f := sub_add_cancel _ _ @@ -136,4 +136,29 @@ begin refl, end +section difference + +lemma predictable_part_bdd_difference {R : ℝ≥0} {f : ℕ → Ω → ℝ} + (ℱ : filtration ℕ m0) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : + ∀ᵐ ω ∂μ, ∀ i, |predictable_part ℱ μ f (i + 1) ω - predictable_part ℱ μ f i ω| ≤ R := +begin + simp_rw [predictable_part, finset.sum_apply, finset.sum_range_succ_sub_sum], + exact ae_all_iff.2 (λ i, ae_bdd_condexp_of_ae_bdd $ ae_all_iff.1 hbdd i), +end + +lemma martingale_part_bdd_difference {R : ℝ≥0} {f : ℕ → Ω → ℝ} + (ℱ : filtration ℕ m0) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : + ∀ᵐ ω ∂μ, ∀ i, |martingale_part ℱ μ f (i + 1) ω - martingale_part ℱ μ f i ω| ≤ ↑(2 * R) := +begin + filter_upwards [hbdd, predictable_part_bdd_difference ℱ hbdd] with ω hω₁ hω₂ i, + simp only [two_mul, martingale_part, pi.sub_apply], + have : |f (i + 1) ω - predictable_part ℱ μ f (i + 1) ω - (f i ω - predictable_part ℱ μ f i ω)| = + |(f (i + 1) ω - f i ω) - (predictable_part ℱ μ f (i + 1) ω - predictable_part ℱ μ f i ω)|, + { ring_nf }, -- `ring` suggests `ring_nf` despite proving the goal + rw this, + exact (abs_sub _ _).trans (add_le_add (hω₁ i) (hω₂ i)), +end + +end difference + end measure_theory diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index 75901a2fa3d0f..80a2df2c1a6b1 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov -/ import order.liminf_limsup import topology.algebra.order.basic +import order.filter.archimedean /-! # Lemmas about liminf and limsup in an order topology. @@ -43,6 +44,19 @@ lemma filter.tendsto.is_cobounded_under_ge {f : filter β} {u : β → α} {a : [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≥) u := h.is_bounded_under_le.is_cobounded_flip +lemma is_bounded_le_at_bot (α : Type*) [hα : nonempty α] [preorder α] : + (at_bot : filter α).is_bounded (≤) := +is_bounded_iff.2 ⟨set.Iic hα.some, mem_at_bot _, hα.some, λ x hx, hx⟩ + +lemma filter.tendsto.is_bounded_under_le_at_bot {α : Type*} [nonempty α] [preorder α] + {f : filter β} {u : β → α} (h : tendsto u f at_bot) : + f.is_bounded_under (≤) u := +(is_bounded_le_at_bot α).mono h + +lemma bdd_above_range_of_tendsto_at_top_at_bot {α : Type*} [nonempty α] [semilattice_sup α] + {u : ℕ → α} (hx : tendsto u at_top at_bot) : bdd_above (set.range u) := +(filter.tendsto.is_bounded_under_le_at_bot hx).bdd_above_range + end order_closed_topology section order_closed_topology @@ -69,6 +83,19 @@ lemma filter.tendsto.is_cobounded_under_le {f : filter β} {u : β → α} {a : [ne_bot f] (h : tendsto u f (𝓝 a)) : f.is_cobounded_under (≤) u := h.is_bounded_under_ge.is_cobounded_flip +lemma is_bounded_ge_at_top (α : Type*) [hα : nonempty α] [preorder α] : + (at_top : filter α).is_bounded (≥) := +is_bounded_le_at_bot αᵒᵈ + +lemma filter.tendsto.is_bounded_under_ge_at_top {α : Type*} [nonempty α] [preorder α] + {f : filter β} {u : β → α} (h : tendsto u f at_top) : + f.is_bounded_under (≥) u := +(is_bounded_ge_at_top α).mono h + +lemma bdd_below_range_of_tendsto_at_top_at_top {α : Type*} [nonempty α] [semilattice_inf α] + {u : ℕ → α} (hx : tendsto u at_top at_top) : bdd_below (set.range u) := +(filter.tendsto.is_bounded_under_ge_at_top hx).bdd_below_range + end order_closed_topology section conditionally_complete_linear_order @@ -291,3 +318,82 @@ lemma monotone.map_liminf_of_continuous_at f_incr.map_Liminf_of_continuous_at f_cont end monotone + +section indicator + +open_locale big_operators + +lemma limsup_eq_tendsto_sum_indicator_nat_at_top (s : ℕ → set α) : + limsup at_top s = + {ω | tendsto (λ n, ∑ k in finset.range n, (s (k + 1)).indicator (1 : α → ℕ) ω) at_top at_top} := +begin + ext ω, + simp only [limsup_eq_infi_supr_of_nat, ge_iff_le, set.supr_eq_Union, + set.infi_eq_Inter, set.mem_Inter, set.mem_Union, exists_prop], + split, + { intro hω, + refine tendsto_at_top_at_top_of_monotone' (λ n m hnm, finset.sum_mono_set_of_nonneg + (λ i, set.indicator_nonneg (λ _ _, zero_le_one) _) (finset.range_mono hnm)) _, + rintro ⟨i, h⟩, + simp only [mem_upper_bounds, set.mem_range, forall_exists_index, forall_apply_eq_imp_iff'] at h, + induction i with k hk, + { obtain ⟨j, hj₁, hj₂⟩ := hω 1, + refine not_lt.2 (h $ j + 1) (lt_of_le_of_lt + (finset.sum_const_zero.symm : 0 = ∑ k in finset.range (j + 1), 0).le _), + refine finset.sum_lt_sum (λ m _, set.indicator_nonneg (λ _ _, zero_le_one) _) + ⟨j - 1, finset.mem_range.2 (lt_of_le_of_lt (nat.sub_le _ _) j.lt_succ_self), _⟩, + rw [nat.sub_add_cancel hj₁, set.indicator_of_mem hj₂], + exact zero_lt_one }, + { rw imp_false at hk, + push_neg at hk, + obtain ⟨i, hi⟩ := hk, + obtain ⟨j, hj₁, hj₂⟩ := hω (i + 1), + replace hi : ∑ k in finset.range i, (s (k + 1)).indicator 1 ω = k + 1 := le_antisymm (h i) hi, + refine not_lt.2 (h $ j + 1) _, + rw [← finset.sum_range_add_sum_Ico _ (i.le_succ.trans (hj₁.trans j.le_succ)), hi], + refine lt_add_of_pos_right _ _, + rw (finset.sum_const_zero.symm : 0 = ∑ k in finset.Ico i (j + 1), 0), + refine finset.sum_lt_sum (λ m _, set.indicator_nonneg (λ _ _, zero_le_one) _) + ⟨j - 1, finset.mem_Ico.2 + ⟨(nat.le_sub_iff_right (le_trans ((le_add_iff_nonneg_left _).2 zero_le') hj₁)).2 hj₁, + lt_of_le_of_lt (nat.sub_le _ _) j.lt_succ_self⟩, _⟩, + rw [nat.sub_add_cancel (le_trans ((le_add_iff_nonneg_left _).2 zero_le') hj₁), + set.indicator_of_mem hj₂], + exact zero_lt_one } }, + { rintro hω i, + rw [set.mem_set_of_eq, tendsto_at_top_at_top] at hω, + by_contra hcon, + push_neg at hcon, + obtain ⟨j, h⟩ := hω (i + 1), + have : ∑ k in finset.range j, (s (k + 1)).indicator 1 ω ≤ i, + { have hle : ∀ j ≤ i, ∑ k in finset.range j, (s (k + 1)).indicator 1 ω ≤ i, + { refine λ j hij, (finset.sum_le_card_nsmul _ _ _ _ : _ ≤ (finset.range j).card • 1).trans _, + { exact λ m hm, set.indicator_apply_le' (λ _, le_rfl) (λ _, zero_le_one) }, + { simpa only [finset.card_range, smul_eq_mul, mul_one] } }, + by_cases hij : j < i, + { exact hle _ hij.le }, + { rw ← finset.sum_range_add_sum_Ico _ (not_lt.1 hij), + suffices : ∑ k in finset.Ico i j, (s (k + 1)).indicator 1 ω = 0, + { rw [this, add_zero], + exact hle _ le_rfl }, + rw finset.sum_eq_zero (λ m hm, _), + exact set.indicator_of_not_mem (hcon _ $ (finset.mem_Ico.1 hm).1.trans m.le_succ) _ } }, + exact not_le.2 (lt_of_lt_of_le i.lt_succ_self $ h _ le_rfl) this } +end + +lemma limsup_eq_tendsto_sum_indicator_at_top + (R : Type*) [ordered_semiring R] [nontrivial R] [archimedean R] (s : ℕ → set α) : + limsup at_top s = + {ω | tendsto (λ n, ∑ k in finset.range n, (s (k + 1)).indicator (1 : α → R) ω) at_top at_top} := +begin + rw limsup_eq_tendsto_sum_indicator_nat_at_top s, + ext ω, + simp only [set.mem_set_of_eq], + rw (_ : (λ n, ∑ k in finset.range n, (s (k + 1)).indicator (1 : α → R) ω) = + (λ n, ↑(∑ k in finset.range n, (s (k + 1)).indicator (1 : α → ℕ) ω))), + { exact tendsto_coe_nat_at_top_iff.symm }, + { ext n, + simp only [set.indicator, pi.one_apply, finset.sum_boole, nat.cast_id] } +end + +end indicator From b4502fd1865ffd0e4e5671503db0243d8f3c7883 Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Sun, 18 Sep 2022 22:14:10 +0000 Subject: [PATCH 300/828] =?UTF-8?q?feat(topology):=20is=5Fopen=20and=20is?= =?UTF-8?q?=5Fclosed=20in=20term=20of=20=E2=88=80=E1=B6=A0=20and=20?= =?UTF-8?q?=E2=88=83=E1=B6=A0=20(#16545)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/topology/basic.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 190b273fb2a6a..4967dbfe8b830 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -965,6 +965,10 @@ calc is_open s ↔ s ⊆ interior s : subset_interior_iff_open.symm lemma is_open_iff_mem_nhds {s : set α} : is_open s ↔ ∀a∈s, s ∈ 𝓝 a := is_open_iff_nhds.trans $ forall_congr $ λ _, imp_congr_right $ λ _, le_principal_iff +/-- A set `s` is open iff for every point `x` in `s` and every `y` close to `x`, `y` is in `s`. -/ +lemma is_open_iff_eventually {s : set α} : is_open s ↔ ∀ x, x ∈ s → ∀ᶠ y in 𝓝 x, y ∈ s := +is_open_iff_mem_nhds + theorem is_open_iff_ultrafilter {s : set α} : is_open s ↔ (∀ (x ∈ s) (l : ultrafilter α), ↑l ≤ 𝓝 x → s ∈ l) := by simp_rw [is_open_iff_mem_nhds, ← mem_iff_ultrafilter] @@ -987,6 +991,15 @@ by rw [filter.frequently, filter.eventually, ← mem_interior_iff_mem_nhds, alias mem_closure_iff_frequently ↔ _ filter.frequently.mem_closure +/-- A set `s` is closed iff for every point `x`, if there is a point `y` close to `x` that belongs +to `s` then `x` is in `s`. -/ +lemma is_closed_iff_frequently {s : set α} : is_closed s ↔ ∀ x, (∃ᶠ y in 𝓝 x, y ∈ s) → x ∈ s := +begin + rw ← closure_subset_iff_is_closed, + apply forall_congr (λ x, _), + rw mem_closure_iff_frequently +end + /-- The set of cluster points of a filter is closed. In particular, the set of limit points of a sequence is closed. -/ lemma is_closed_set_of_cluster_pt {f : filter α} : is_closed {x | cluster_pt x f} := From 489bc58fdd8a957eda4486aec7a354c5afce4270 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Mon, 19 Sep 2022 00:39:06 +0000 Subject: [PATCH 301/828] feat(topology/homeomorph): add `homeomorph.Pi_congr_right` (#16527) Co-authored-by: Yury G. Kudryashov --- src/topology/homeomorph.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 8e9afbe7afd0f..7e77c453e5ec2 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -387,6 +387,18 @@ def punit_prod : punit × α ≃ₜ α := end +/-- If each `β₁ i` is homeomorphic to `β₂ i`, then `Π i, β₁ i` is homeomorphic to `Π i, β₂ i`. -/ +@[simps apply to_equiv] def Pi_congr_right {ι : Type*} {β₁ β₂ : ι → Type*} + [Π i, topological_space (β₁ i)] [Π i, topological_space (β₂ i)] (F : Π i, β₁ i ≃ₜ β₂ i) : + (Π i, β₁ i) ≃ₜ (Π i, β₂ i) := +{ continuous_to_fun := continuous_pi (λ i, (F i).continuous.comp $ continuous_apply i), + continuous_inv_fun := continuous_pi (λ i, (F i).symm.continuous.comp $ continuous_apply i), + to_equiv := equiv.Pi_congr_right (λ i, (F i).to_equiv) } + +@[simp] lemma Pi_congr_right_symm {ι : Type*} {β₁ β₂ : ι → Type*} [Π i, topological_space (β₁ i)] + [Π i, topological_space (β₂ i)] (F : Π i, β₁ i ≃ₜ β₂ i) : + (Pi_congr_right F).symm = Pi_congr_right (λ i, (F i).symm) := rfl + /-- `ulift α` is homeomorphic to `α`. -/ def {u v} ulift {α : Type u} [topological_space α] : ulift.{v u} α ≃ₜ α := { continuous_to_fun := continuous_ulift_down, From 7ec8f8a279506860fd20a56ff5696180c04970bb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Matias=20Heikkil=C3=A4?= Date: Mon, 19 Sep 2022 02:22:47 +0000 Subject: [PATCH 302/828] feat(analysis/inner_product_space/basic): `ext_inner_map` (#16388) Equality of operators by equality of inner products. An easy corollary of `inner_map_self_eq_zero`. --- src/analysis/inner_product_space/basic.lean | 14 +++++++++++++- 1 file changed, 13 insertions(+), 1 deletion(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 3e0ed29ece738..7a89331e777c6 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -1157,7 +1157,7 @@ begin end /-- -If `⟪T x, x⟫_ℂ = 0` for all x, then T = 0. +A linear map `T` is zero, if and only if the identity `⟪T x, x⟫_ℂ = 0` holds for all `x`. -/ lemma inner_map_self_eq_zero (T : V →ₗ[ℂ] V) : (∀ (x : V), ⟪T x, x⟫_ℂ = 0) ↔ T = 0 := @@ -1171,6 +1171,18 @@ begin simp only [linear_map.zero_apply, inner_zero_left] } end +/-- +Two linear maps `S` and `T` are equal, if and only if the identity `⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ` holds +for all `x`. +-/ +lemma ext_inner_map (S T : V →ₗ[ℂ] V) : + (∀ (x : V), ⟪S x, x⟫_ℂ = ⟪T x, x⟫_ℂ) ↔ S = T := +begin + rw [←sub_eq_zero, ←inner_map_self_eq_zero], + refine forall_congr (λ x, _), + rw [linear_map.sub_apply, inner_sub_left, sub_eq_zero], +end + end complex section From 08899973335b295e6d648fc1ee44ff5469f89cef Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 19 Sep 2022 03:33:51 +0000 Subject: [PATCH 303/828] feat(group_theory/coset): Add `quotient_equiv_of_eq` (#16439) This PR adds `quotient_equiv_of_eq` and renames `equiv_quotient_of_eq` to `quotient_mul_equiv_of_eq`. --- src/data/zmod/quotient.lean | 4 ++-- src/group_theory/coset.lean | 12 ++++++++++++ src/group_theory/quotient_group.lean | 15 ++++++--------- src/ring_theory/valuation/valuation_subring.lean | 2 +- 4 files changed, 21 insertions(+), 12 deletions(-) diff --git a/src/data/zmod/quotient.lean b/src/data/zmod/quotient.lean index 2f757275d72cb..123bcc2d128bd 100644 --- a/src/data/zmod/quotient.lean +++ b/src/data/zmod/quotient.lean @@ -39,13 +39,13 @@ namespace int /-- `ℤ` modulo multiples of `n : ℕ` is `zmod n`. -/ def quotient_zmultiples_nat_equiv_zmod : ℤ ⧸ add_subgroup.zmultiples (n : ℤ) ≃+ zmod n := -(equiv_quotient_of_eq (zmod.ker_int_cast_add_hom _)).symm.trans $ +(quotient_add_equiv_of_eq (zmod.ker_int_cast_add_hom _)).symm.trans $ quotient_ker_equiv_of_right_inverse (int.cast_add_hom (zmod n)) coe int_cast_zmod_cast /-- `ℤ` modulo multiples of `a : ℤ` is `zmod a.nat_abs`. -/ def quotient_zmultiples_equiv_zmod (a : ℤ) : ℤ ⧸ add_subgroup.zmultiples a ≃+ zmod a.nat_abs := -(equiv_quotient_of_eq (zmultiples_nat_abs a)).symm.trans +(quotient_add_equiv_of_eq (zmultiples_nat_abs a)).symm.trans (quotient_zmultiples_nat_equiv_zmod a.nat_abs) /-- `ℤ` modulo the ideal generated by `n : ℕ` is `zmod n`. -/ diff --git a/src/group_theory/coset.lean b/src/group_theory/coset.lean index c0441368fb8ee..7abe2d26fd57f 100644 --- a/src/group_theory/coset.lean +++ b/src/group_theory/coset.lean @@ -441,6 +441,18 @@ calc α ≃ Σ L : α ⧸ s, {x : α // (x : α ⧸ s) = L} : variables {t : subgroup α} +/-- If two subgroups `M` and `N` of `G` are equal, their quotients are in bijection. -/ +@[to_additive "If two subgroups `M` and `N` of `G` are equal, their quotients are in bijection."] +def quotient_equiv_of_eq (h : s = t) : α ⧸ s ≃ α ⧸ t := +{ to_fun := quotient.map' id (λ a b h', h ▸ h'), + inv_fun := quotient.map' id (λ a b h', h.symm ▸ h'), + left_inv := λ q, induction_on' q (λ g, rfl), + right_inv := λ q, induction_on' q (λ g, rfl) } + +lemma quotient_equiv_of_eq_mk (h : s = t) (a : α) : + quotient_equiv_of_eq h (quotient_group.mk a) = (quotient_group.mk a) := +rfl + /-- If `H ≤ K`, then `G/H ≃ G/K × K/H` constructively, using the provided right inverse of the quotient map `G → G/K`. The classical version is `quotient_equiv_prod_of_le`. -/ @[to_additive "If `H ≤ K`, then `G/H ≃ G/K × K/H` constructively, using the provided right inverse diff --git a/src/group_theory/quotient_group.lean b/src/group_theory/quotient_group.lean index 9aeb88ef126d9..32ff90e709fb5 100644 --- a/src/group_theory/quotient_group.lean +++ b/src/group_theory/quotient_group.lean @@ -275,17 +275,14 @@ quotient_ker_equiv_of_right_inverse φ _ hφ.has_right_inverse.some_spec isomorphic. -/ @[to_additive "If two normal subgroups `M` and `N` of `G` are the same, their quotient groups are isomorphic."] -def equiv_quotient_of_eq {M N : subgroup G} [M.normal] [N.normal] (h : M = N) : +def quotient_mul_equiv_of_eq {M N : subgroup G} [M.normal] [N.normal] (h : M = N) : G ⧸ M ≃* G ⧸ N := -{ to_fun := (lift M (mk' N) (λ m hm, quotient_group.eq.mpr (by simpa [← h] using M.inv_mem hm))), - inv_fun := (lift N (mk' M) (λ n hn, quotient_group.eq.mpr (by simpa [← h] using N.inv_mem hn))), - left_inv := λ x, x.induction_on' $ by { intro, refl }, - right_inv := λ x, x.induction_on' $ by { intro, refl }, - map_mul' := λ x y, by rw monoid_hom.map_mul } +{ map_mul' := λ q r, quotient.induction_on₂' q r (λ g h, rfl), + .. subgroup.quotient_equiv_of_eq h } @[simp, to_additive] -lemma equiv_quotient_of_eq_mk {M N : subgroup G} [M.normal] [N.normal] (h : M = N) (x : G) : - quotient_group.equiv_quotient_of_eq h (quotient_group.mk x) = (quotient_group.mk x) := +lemma quotient_mul_equiv_of_eq_mk {M N : subgroup G} [M.normal] [N.normal] (h : M = N) (x : G) : + quotient_group.quotient_mul_equiv_of_eq h (quotient_group.mk x) = (quotient_group.mk x) := rfl /-- Let `A', A, B', B` be subgroups of `G`. If `A' ≤ B'` and `A ≤ B`, @@ -405,7 +402,7 @@ have φ_surjective : function.surjective φ := λ x, x.induction_on' $ change h⁻¹ * (h * n) ∈ N, rwa [←mul_assoc, inv_mul_self, one_mul], end, -(equiv_quotient_of_eq (by simp [comap_comap, ←comap_ker])).trans +(quotient_mul_equiv_of_eq (by simp [comap_comap, ←comap_ker])).trans (quotient_ker_equiv_of_surjective φ φ_surjective) end snd_isomorphism_thm diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index 3e05886ec1cea..94acc78a0b841 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -619,7 +619,7 @@ the units of the residue field of `A`. -/ def units_mod_principal_units_equiv_residue_field_units : (A.unit_group ⧸ (A.principal_unit_group.comap A.unit_group.subtype)) ≃* (local_ring.residue_field A)ˣ := -mul_equiv.trans (quotient_group.equiv_quotient_of_eq A.ker_unit_group_to_residue_field_units.symm) +(quotient_group.quotient_mul_equiv_of_eq A.ker_unit_group_to_residue_field_units.symm).trans (quotient_group.quotient_ker_equiv_of_surjective _ A.surjective_unit_group_to_residue_field_units) @[simp] From f59c26dab4bf4b97bb7404e4dd0f8aa52fee3fe9 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 19 Sep 2022 05:48:18 +0000 Subject: [PATCH 304/828] feat(analysis/complex/poincare_metric): new file (#14703) Introduce a `metric_space` structure on the upper half plane. --- .../complex/upper_half_plane/basic.lean | 32 ++ .../complex/upper_half_plane/metric.lean | 352 ++++++++++++++++++ src/topology/metric_space/basic.lean | 6 + 3 files changed, 390 insertions(+) create mode 100644 src/analysis/complex/upper_half_plane/metric.lean diff --git a/src/analysis/complex/upper_half_plane/basic.lean b/src/analysis/complex/upper_half_plane/basic.lean index 2832a2fbd3a0b..b9451fd70a07c 100644 --- a/src/analysis/complex/upper_half_plane/basic.lean +++ b/src/analysis/complex/upper_half_plane/basic.lean @@ -48,6 +48,8 @@ namespace upper_half_plane instance : inhabited ℍ := ⟨⟨complex.I, by simp⟩⟩ +instance : can_lift ℂ ℍ := subtype.can_lift (λ z, 0 < z.im) + /-- Imaginary part -/ def im (z : ℍ) := (z : ℂ).im @@ -267,4 +269,34 @@ lemma denom_apply (g : SL(2, ℤ)) (z : ℍ) : denom g z = (↑g : matrix (fin 2 end SL_modular_action +section pos_real_action + +instance pos_real_action : mul_action {x : ℝ // 0 < x} ℍ := +{ smul := λ x z, mk ((x : ℝ) • z) $ by simpa using mul_pos x.2 z.2, + one_smul := λ z, subtype.ext $ one_smul _ _, + mul_smul := λ x y z, subtype.ext $ mul_smul (x : ℝ) y (z : ℂ) } + +variables (x : {x : ℝ // 0 < x}) (z : ℍ) + +@[simp] lemma coe_pos_real_smul : ↑(x • z) = (x : ℝ) • (z : ℂ) := rfl +@[simp] lemma pos_real_im : (x • z).im = x * z.im := complex.smul_im _ _ +@[simp] lemma pos_real_re : (x • z).re = x * z.re := complex.smul_re _ _ + +end pos_real_action + +section real_add_action + +instance : add_action ℝ ℍ := +{ vadd := λ x z, mk (x + z) $ by simpa using z.im_pos, + zero_vadd := λ z, subtype.ext $ by simp, + add_vadd := λ x y z, subtype.ext $ by simp [add_assoc] } + +variables (x : ℝ) (z : ℍ) + +@[simp] lemma coe_vadd : ↑(x +ᵥ z) = (x + z : ℂ) := rfl +@[simp] lemma vadd_re : (x +ᵥ z).re = x + z.re := rfl +@[simp] lemma vadd_im : (x +ᵥ z).im = z.im := zero_add _ + +end real_add_action + end upper_half_plane diff --git a/src/analysis/complex/upper_half_plane/metric.lean b/src/analysis/complex/upper_half_plane/metric.lean new file mode 100644 index 0000000000000..f72a9e66dc13e --- /dev/null +++ b/src/analysis/complex/upper_half_plane/metric.lean @@ -0,0 +1,352 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import analysis.complex.upper_half_plane.topology +import analysis.special_functions.arsinh +import geometry.euclidean.inversion + +/-! +# Metric on the upper half-plane + +In this file we define a `metric_space` structure on the `upper_half_plane`. We use hyperbolic +(Poincaré) distance given by +`dist z w = 2 * arsinh (dist (z : ℂ) w / (2 * real.sqrt (z.im * w.im)))` instead of the induced +Euclidean distance because the hyperbolic distance is invariant under holomorphic automorphisms of +the upper half-plane. However, we ensure that the projection to `topological_space` is +definitionally equal to the induced topological space structure. + +We also prove that a metric ball/closed ball/sphere in Poincaré metric is a Euclidean ball/closed +ball/sphere with another center and radius. + +-/ + +noncomputable theory + +open_locale upper_half_plane complex_conjugate nnreal topological_space +open set metric filter real + +variables {z w : ℍ} {r R : ℝ} + +namespace upper_half_plane + +instance : has_dist ℍ := +⟨λ z w, 2 * arsinh (dist (z : ℂ) w / (2 * sqrt (z.im * w.im)))⟩ + +lemma dist_eq (z w : ℍ) : dist z w = 2 * arsinh (dist (z : ℂ) w / (2 * sqrt (z.im * w.im))) := +rfl + +lemma sinh_half_dist (z w : ℍ) : + sinh (dist z w / 2) = dist (z : ℂ) w / (2 * sqrt (z.im * w.im)) := +by rw [dist_eq, mul_div_cancel_left (arsinh _) two_ne_zero, sinh_arsinh] + +lemma cosh_half_dist (z w : ℍ) : + cosh (dist z w / 2) = dist (z : ℂ) (conj (w : ℂ)) / (2 * sqrt (z.im * w.im)) := +begin + have H₁ : (2 ^ 2 : ℝ) = 4, by norm_num1, + have H₂ : 0 < z.im * w.im, from mul_pos z.im_pos w.im_pos, + have H₃ : 0 < 2 * sqrt (z.im * w.im), from mul_pos two_pos (sqrt_pos.2 H₂), + rw [← sq_eq_sq (cosh_pos _).le (div_nonneg dist_nonneg H₃.le), cosh_sq', sinh_half_dist, div_pow, + div_pow, one_add_div (pow_ne_zero 2 H₃.ne'), mul_pow, sq_sqrt H₂.le, H₁], + congr' 1, + simp only [complex.dist_eq, complex.sq_abs, complex.norm_sq_sub, complex.norm_sq_conj, + complex.conj_conj, complex.mul_re, complex.conj_re, complex.conj_im, coe_im], + ring +end + +lemma tanh_half_dist (z w : ℍ) : + tanh (dist z w / 2) = dist (z : ℂ) w / dist (z : ℂ) (conj ↑w) := +begin + rw [tanh_eq_sinh_div_cosh, sinh_half_dist, cosh_half_dist, div_div_div_comm, div_self, div_one], + exact (mul_pos two_pos (sqrt_pos.2 $ mul_pos z.im_pos w.im_pos)).ne' +end + +lemma exp_half_dist (z w : ℍ) : + exp (dist z w / 2) = (dist (z : ℂ) w + dist (z : ℂ) (conj ↑w)) / (2 * sqrt (z.im * w.im)) := +by rw [← sinh_add_cosh, sinh_half_dist, cosh_half_dist, add_div] + +lemma cosh_dist (z w : ℍ) : cosh (dist z w) = 1 + dist (z : ℂ) w ^ 2 / (2 * z.im * w.im) := +by rw [dist_eq, cosh_two_mul, cosh_sq', add_assoc, ← two_mul, sinh_arsinh, div_pow, mul_pow, + sq_sqrt (mul_pos z.im_pos w.im_pos).le, sq (2 : ℝ), mul_assoc, ← mul_div_assoc, + mul_assoc, mul_div_mul_left _ _ (@two_ne_zero ℝ _ _)] + +lemma sinh_half_dist_add_dist (a b c : ℍ) : + sinh ((dist a b + dist b c) / 2) = + (dist (a : ℂ) b * dist (c : ℂ) (conj ↑b) + dist (b : ℂ) c * dist (a : ℂ) (conj ↑b)) / + (2 * sqrt (a.im * c.im) * dist (b : ℂ) (conj ↑b)) := +begin + simp only [add_div _ _ (2 : ℝ), sinh_add, sinh_half_dist, cosh_half_dist, div_mul_div_comm], + rw [← add_div, complex.dist_self_conj, coe_im, abs_of_pos b.im_pos, mul_comm (dist ↑b _), + dist_comm (b : ℂ), complex.dist_conj_comm, mul_mul_mul_comm, mul_mul_mul_comm _ _ _ b.im], + congr' 2, + rw [sqrt_mul, sqrt_mul, sqrt_mul, mul_comm (sqrt a.im), mul_mul_mul_comm, mul_self_sqrt, + mul_comm]; exact (im_pos _).le +end + +protected lemma dist_comm (z w : ℍ) : dist z w = dist w z := +by simp only [dist_eq, dist_comm (z : ℂ), mul_comm] + +lemma dist_le_iff_le_sinh : + dist z w ≤ r ↔ dist (z : ℂ) w / (2 * sqrt (z.im * w.im)) ≤ sinh (r / 2) := +by rw [← div_le_div_right (@two_pos ℝ _ _), ← sinh_le_sinh, sinh_half_dist] + +lemma dist_eq_iff_eq_sinh : + dist z w = r ↔ dist (z : ℂ) w / (2 * sqrt (z.im * w.im)) = sinh (r / 2) := +by rw [← div_left_inj' (@two_ne_zero ℝ _ _), ← sinh_inj, sinh_half_dist] + +lemma dist_eq_iff_eq_sq_sinh (hr : 0 ≤ r) : + dist z w = r ↔ dist (z : ℂ) w ^ 2 / (4 * z.im * w.im) = sinh (r / 2) ^ 2 := +begin + rw [dist_eq_iff_eq_sinh, ← sq_eq_sq, div_pow, mul_pow, sq_sqrt, mul_assoc], + { norm_num }, + { exact (mul_pos z.im_pos w.im_pos).le }, + { exact div_nonneg dist_nonneg (mul_nonneg zero_le_two $ sqrt_nonneg _) }, + { exact sinh_nonneg_iff.2 (div_nonneg hr zero_le_two) } +end + +protected lemma dist_triangle (a b c : ℍ) : dist a c ≤ dist a b + dist b c := +begin + rw [dist_le_iff_le_sinh, sinh_half_dist_add_dist, + div_mul_eq_div_div _ _ (dist _ _), le_div_iff, div_mul_eq_mul_div], + { exact div_le_div_of_le (mul_nonneg zero_le_two (sqrt_nonneg _)) + (euclidean_geometry.mul_dist_le_mul_dist_add_mul_dist (a : ℂ) b c (conj ↑b)) }, + { rw [dist_comm, dist_pos, ne.def, complex.eq_conj_iff_im], + exact b.im_ne_zero } +end + +lemma dist_le_dist_coe_div_sqrt (z w : ℍ) : + dist z w ≤ dist (z : ℂ) w / sqrt (z.im * w.im) := +begin + rw [dist_le_iff_le_sinh, ← div_mul_eq_div_div_swap, self_le_sinh_iff], + exact div_nonneg dist_nonneg (mul_nonneg zero_le_two (sqrt_nonneg _)) +end + +/-- An auxiliary `metric_space` instance on the upper half-plane. This instance has bad projection +to `topological_space`. We replace it later. -/ +def metric_space_aux : metric_space ℍ := +{ dist := dist, + dist_self := λ z, by rw [dist_eq, dist_self, zero_div, arsinh_zero, mul_zero], + dist_comm := upper_half_plane.dist_comm, + dist_triangle := upper_half_plane.dist_triangle, + eq_of_dist_eq_zero := λ z w h, + by simpa [dist_eq, real.sqrt_eq_zero', (mul_pos z.im_pos w.im_pos).not_le, subtype.coe_inj] + using h } + +open complex + +lemma cosh_dist' (z w : ℍ) : + real.cosh (dist z w) = ((z.re - w.re) ^ 2 + z.im ^ 2 + w.im ^ 2) / (2 * z.im * w.im) := +have H : 0 < 2 * z.im * w.im, from mul_pos (mul_pos two_pos z.im_pos) w.im_pos, +by { field_simp [cosh_dist, complex.dist_eq, complex.sq_abs, norm_sq_apply, H, H.ne'], ring } + +/-- Euclidean center of the circle with center `z` and radius `r` in the hyperbolic metric. -/ +def center (z : ℍ) (r : ℝ) : ℍ := ⟨⟨z.re, z.im * cosh r⟩, mul_pos z.im_pos (cosh_pos _)⟩ + +@[simp] lemma center_re (z r) : (center z r).re = z.re := rfl +@[simp] lemma center_im (z r) : (center z r).im = z.im * cosh r := rfl + +@[simp] lemma center_zero (z : ℍ) : center z 0 = z := +subtype.ext $ ext rfl $ by rw [coe_im, coe_im, center_im, real.cosh_zero, mul_one] + +lemma dist_coe_center_sq (z w : ℍ) (r : ℝ) : + dist (z : ℂ) (w.center r) ^ 2 = + 2 * z.im * w.im * (cosh (dist z w) - cosh r) + (w.im * sinh r) ^ 2 := +begin + have H : 2 * z.im * w.im ≠ 0, by apply_rules [mul_ne_zero, two_ne_zero, im_ne_zero], + simp only [complex.dist_eq, complex.sq_abs, norm_sq_apply, coe_re, coe_im, center_re, center_im, + cosh_dist', mul_div_cancel' _ H, sub_sq z.im, mul_pow, real.cosh_sq, sub_re, sub_im, mul_sub, + ← sq], + ring +end + +lemma dist_coe_center (z w : ℍ) (r : ℝ) : + dist (z : ℂ) (w.center r) = + sqrt (2 * z.im * w.im * (cosh (dist z w) - cosh r) + (w.im * sinh r) ^ 2) := +by rw [← sqrt_sq dist_nonneg, dist_coe_center_sq] + +lemma cmp_dist_eq_cmp_dist_coe_center (z w : ℍ) (r : ℝ) : + cmp (dist z w) r = cmp (dist (z : ℂ) (w.center r)) (w.im * sinh r) := +begin + letI := metric_space_aux, + cases lt_or_le r 0 with hr₀ hr₀, + { transitivity ordering.gt, + exacts [(hr₀.trans_le dist_nonneg).cmp_eq_gt, + ((mul_neg_of_pos_of_neg w.im_pos (sinh_neg_iff.2 hr₀)).trans_le + dist_nonneg).cmp_eq_gt.symm] }, + have hr₀' : 0 ≤ w.im * sinh r, from mul_nonneg w.im_pos.le (sinh_nonneg_iff.2 hr₀), + have hzw₀ : 0 < 2 * z.im * w.im, from mul_pos (mul_pos two_pos z.im_pos) w.im_pos, + simp only [← cosh_strict_mono_on.cmp_map_eq dist_nonneg hr₀, + ← (@strict_mono_on_pow ℝ _ _ two_pos).cmp_map_eq dist_nonneg hr₀', dist_coe_center_sq], + rw [← cmp_mul_pos_left hzw₀, ← cmp_sub_zero, ← mul_sub, ← cmp_add_right, zero_add], +end + +lemma dist_eq_iff_dist_coe_center_eq : dist z w = r ↔ dist (z : ℂ) (w.center r) = w.im * sinh r := +eq_iff_eq_of_cmp_eq_cmp (cmp_dist_eq_cmp_dist_coe_center z w r) + +@[simp] lemma dist_self_center (z : ℍ) (r : ℝ) : dist (z : ℂ) (z.center r) = z.im * (cosh r - 1) := +begin + rw [dist_of_re_eq (z.center_re r).symm, dist_comm, real.dist_eq, mul_sub, mul_one], + exact abs_of_nonneg (sub_nonneg.2 $ le_mul_of_one_le_right z.im_pos.le (one_le_cosh _)) +end + +@[simp] lemma dist_center_dist (z w : ℍ) : + dist (z : ℂ) (w.center (dist z w)) = w.im * sinh (dist z w) := +dist_eq_iff_dist_coe_center_eq.1 rfl + +lemma dist_lt_iff_dist_coe_center_lt : + dist z w < r ↔ dist (z : ℂ) (w.center r) < w.im * sinh r := +lt_iff_lt_of_cmp_eq_cmp (cmp_dist_eq_cmp_dist_coe_center z w r) + +lemma lt_dist_iff_lt_dist_coe_center : + r < dist z w ↔ w.im * sinh r < dist (z : ℂ) (w.center r) := +lt_iff_lt_of_cmp_eq_cmp (cmp_eq_cmp_symm.1 $ cmp_dist_eq_cmp_dist_coe_center z w r) + +lemma dist_le_iff_dist_coe_center_le : + dist z w ≤ r ↔ dist (z : ℂ) (w.center r) ≤ w.im * sinh r := +le_iff_le_of_cmp_eq_cmp (cmp_dist_eq_cmp_dist_coe_center z w r) + +lemma le_dist_iff_le_dist_coe_center : + r < dist z w ↔ w.im * sinh r < dist (z : ℂ) (w.center r) := +lt_iff_lt_of_cmp_eq_cmp (cmp_eq_cmp_symm.1 $ cmp_dist_eq_cmp_dist_coe_center z w r) + +/-- For two points on the same vertical line, the distance is equal to the distance between the +logarithms of their imaginary parts. -/ +lemma dist_of_re_eq (h : z.re = w.re) : dist z w = dist (log z.im) (log w.im) := +begin + have h₀ : 0 < z.im / w.im, from div_pos z.im_pos w.im_pos, + rw [dist_eq_iff_dist_coe_center_eq, real.dist_eq, ← abs_sinh, ← log_div z.im_ne_zero w.im_ne_zero, + sinh_log h₀, dist_of_re_eq, coe_im, coe_im, center_im, cosh_abs, cosh_log h₀, inv_div]; + [skip, exact h], + nth_rewrite 3 [← abs_of_pos w.im_pos], + simp only [← _root_.abs_mul, coe_im, real.dist_eq], + congr' 1, + field_simp [z.im_pos, w.im_pos, z.im_ne_zero, w.im_ne_zero], + ring +end + +/-- Hyperbolic distance between two points is greater than or equal to the distance between the +logarithms of their imaginary parts. -/ +lemma dist_log_im_le (z w : ℍ) : dist (log z.im) (log w.im) ≤ dist z w := +calc dist (log z.im) (log w.im) = @dist ℍ _ ⟨⟨0, z.im⟩, z.im_pos⟩ ⟨⟨0, w.im⟩, w.im_pos⟩ : + eq.symm $ @dist_of_re_eq ⟨⟨0, z.im⟩, z.im_pos⟩ ⟨⟨0, w.im⟩, w.im_pos⟩ rfl +... ≤ dist z w : + mul_le_mul_of_nonneg_left (arsinh_le_arsinh.2 $ div_le_div_of_le + (mul_nonneg zero_le_two (sqrt_nonneg _)) $ + by simpa [sqrt_sq_eq_abs] using complex.abs_im_le_abs (z - w)) zero_le_two + +lemma im_le_im_mul_exp_dist (z w : ℍ) : z.im ≤ w.im * exp (dist z w) := +begin + rw [← div_le_iff' w.im_pos, ← exp_log z.im_pos, ← exp_log w.im_pos, ← real.exp_sub, exp_le_exp], + exact (le_abs_self _).trans (dist_log_im_le z w) +end + +lemma im_div_exp_dist_le (z w : ℍ) : z.im / exp (dist z w) ≤ w.im := +(div_le_iff (exp_pos _)).2 (im_le_im_mul_exp_dist z w) + +/-- An upper estimate on the complex distance between two points in terms of the hyperbolic distance +and the imaginary part of one of the points. -/ +lemma dist_coe_le (z w : ℍ) : dist (z : ℂ) w ≤ w.im * (exp (dist z w) - 1) := +calc dist (z : ℂ) w ≤ dist (z : ℂ) (w.center (dist z w)) + dist (w : ℂ) (w.center (dist z w)) : + dist_triangle_right _ _ _ +... = w.im * (exp (dist z w) - 1) : + by rw [dist_center_dist, dist_self_center, ← mul_add, ← add_sub_assoc, real.sinh_add_cosh] + +/-- An upper estimate on the complex distance between two points in terms of the hyperbolic distance +and the imaginary part of one of the points. -/ +lemma le_dist_coe (z w : ℍ) : w.im * (1 - exp (-dist z w)) ≤ dist (z : ℂ) w := +calc w.im * (1 - exp (-dist z w)) + = dist (z : ℂ) (w.center (dist z w)) - dist (w : ℂ) (w.center (dist z w)) : + by { rw [dist_center_dist, dist_self_center, ← real.cosh_sub_sinh], ring } +... ≤ dist (z : ℂ) w : sub_le_iff_le_add.2 $ dist_triangle _ _ _ + +/-- The hyperbolic metric on the upper half plane. We ensure that the projection to +`topological_space` is definitionally equal to the subtype topology. -/ +instance : metric_space ℍ := metric_space_aux.replace_topology $ +begin + refine le_antisymm (continuous_id_iff_le.1 _) _, + { refine (@continuous_iff_continuous_dist _ _ metric_space_aux.to_pseudo_metric_space _ _).2 _, + have : ∀ (x : ℍ × ℍ), 2 * real.sqrt (x.1.im * x.2.im) ≠ 0, + from λ x, mul_ne_zero two_ne_zero (real.sqrt_pos.2 $ mul_pos x.1.im_pos x.2.im_pos).ne', + -- `continuity` fails to apply `continuous.div` + apply_rules [continuous.div, continuous.mul, continuous_const, continuous.arsinh, + continuous.dist, continuous_coe.comp, continuous_fst, continuous_snd, + real.continuous_sqrt.comp, continuous_im.comp] }, + { letI : metric_space ℍ := metric_space_aux, + refine le_of_nhds_le_nhds (λ z, _), + rw [nhds_induced], + refine (nhds_basis_ball.le_basis_iff (nhds_basis_ball.comap _)).2 (λ R hR, _), + have h₁ : 1 < R / im z + 1, from lt_add_of_pos_left _ (div_pos hR z.im_pos), + have h₀ : 0 < R / im z + 1, from one_pos.trans h₁, + refine ⟨log (R / im z + 1), real.log_pos h₁, _⟩, + refine λ w hw, (dist_coe_le w z).trans_lt _, + rwa [← lt_div_iff' z.im_pos, sub_lt_iff_lt_add, ← real.lt_log_iff_exp_lt h₀] } +end + +lemma im_pos_of_dist_center_le {z : ℍ} {r : ℝ} {w : ℂ} (h : dist w (center z r) ≤ z.im * sinh r) : + 0 < w.im := +calc 0 < z.im * (cosh r - sinh r) : mul_pos z.im_pos (sub_pos.2 $ sinh_lt_cosh _) +... = (z.center r).im - z.im * sinh r : mul_sub _ _ _ +... ≤ (z.center r).im - dist (z.center r : ℂ) w : sub_le_sub_left (by rwa [dist_comm]) _ +... ≤ w.im : sub_le.1 $ (le_abs_self _).trans (abs_im_le_abs $ z.center r - w) + +lemma image_coe_closed_ball (z : ℍ) (r : ℝ) : + (coe : ℍ → ℂ) '' closed_ball z r = closed_ball (z.center r) (z.im * sinh r) := +begin + ext w, split, + { rintro ⟨w, hw, rfl⟩, + exact dist_le_iff_dist_coe_center_le.1 hw }, + { intro hw, + lift w to ℍ using im_pos_of_dist_center_le hw, + exact mem_image_of_mem _ (dist_le_iff_dist_coe_center_le.2 hw) }, +end + +lemma image_coe_ball (z : ℍ) (r : ℝ) : + (coe : ℍ → ℂ) '' ball z r = ball (z.center r) (z.im * sinh r) := +begin + ext w, split, + { rintro ⟨w, hw, rfl⟩, + exact dist_lt_iff_dist_coe_center_lt.1 hw }, + { intro hw, + lift w to ℍ using im_pos_of_dist_center_le (ball_subset_closed_ball hw), + exact mem_image_of_mem _ (dist_lt_iff_dist_coe_center_lt.2 hw) }, +end + +lemma image_coe_sphere (z : ℍ) (r : ℝ) : + (coe : ℍ → ℂ) '' sphere z r = sphere (z.center r) (z.im * sinh r) := +begin + ext w, split, + { rintro ⟨w, hw, rfl⟩, + exact dist_eq_iff_dist_coe_center_eq.1 hw }, + { intro hw, + lift w to ℍ using im_pos_of_dist_center_le (sphere_subset_closed_ball hw), + exact mem_image_of_mem _ (dist_eq_iff_dist_coe_center_eq.2 hw) }, +end + +instance : proper_space ℍ := +begin + refine ⟨λ z r, _⟩, + rw [← inducing_coe.is_compact_iff, image_coe_closed_ball], + apply is_compact_closed_ball +end + +lemma isometry_vertical_line (a : ℝ) : isometry (λ y, mk ⟨a, exp y⟩ (exp_pos y)) := +begin + refine isometry.of_dist_eq (λ y₁ y₂, _), + rw [dist_of_re_eq], + exacts [congr_arg2 _ (log_exp _) (log_exp _), rfl] +end + +lemma isometry_real_vadd (a : ℝ) : isometry ((+ᵥ) a : ℍ → ℍ) := +isometry.of_dist_eq $ λ y₁ y₂, by simp only [dist_eq, coe_vadd, vadd_im, dist_add_left] + +lemma isometry_pos_mul (a : {x : ℝ // 0 < x}) : isometry ((•) a : ℍ → ℍ) := +begin + refine isometry.of_dist_eq (λ y₁ y₂, _), + simp only [dist_eq, coe_pos_real_smul, pos_real_im], congr' 2, + rw [dist_smul, mul_mul_mul_comm, real.sqrt_mul (mul_self_nonneg _), real.sqrt_mul_self_eq_abs, + real.norm_eq_abs, mul_left_comm], + exact mul_div_mul_left _ _ (mt _root_.abs_eq_zero.1 a.2.ne') +end + +end upper_half_plane diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 197a82f8ebdf3..c2bc8c35a9825 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1608,6 +1608,12 @@ lemma tendsto_iff_dist_tendsto_zero {f : β → α} {x : filter β} {a : α} : (tendsto f x (𝓝 a)) ↔ (tendsto (λb, dist (f b) a) x (𝓝 0)) := by rw [← nhds_comap_dist a, tendsto_comap_iff] +lemma continuous_iff_continuous_dist [topological_space β] {f : β → α} : + continuous f ↔ continuous (λ x : β × β, dist (f x.1) (f x.2)) := +⟨λ h, (h.comp continuous_fst).dist (h.comp continuous_snd), λ h, continuous_iff_continuous_at.2 $ + λ x, tendsto_iff_dist_tendsto_zero.2 $ + (h.comp (continuous_id.prod_mk continuous_const)).tendsto' _ _ $ dist_self _⟩ + lemma uniform_continuous_nndist : uniform_continuous (λp:α×α, nndist p.1 p.2) := uniform_continuous_dist.subtype_mk _ From d673bafe8294989cbc2fc1d31e1865ca10bfbe84 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 19 Sep 2022 07:18:09 +0000 Subject: [PATCH 305/828] feat(topology/algebra/ordered): add lemmas, golf (#16343) * add lemmas about `nhds_within_Ici`/`nhds_within_Iic`; * move `order_separated` and `order_topology.to_order_closed_topology` up; * use new lemmas to golf some proofs. --- src/topology/algebra/order/basic.lean | 236 +++++++++++++------------- 1 file changed, 119 insertions(+), 117 deletions(-) diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index 70e0be39a3c5e..442c1e88eba91 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -7,6 +7,7 @@ import algebra.group_with_zero.power import data.set.intervals.pi import order.filter.interval import topology.algebra.field +import topology.algebra.order.left_right import tactic.linarith import tactic.tfae import tactic.positivity @@ -76,7 +77,7 @@ definition `preorder.topology α` though, that can be registered as an instance for specific types. -/ -open classical set filter topological_space +open set filter topological_space open function open order_dual (to_dual of_dual) open_locale topological_space classical filter @@ -899,6 +900,47 @@ begin exact λ hx, ht.out a.2 y.2 ⟨le_of_lt h, le_of_not_gt hx⟩ } } end +lemma nhds_within_Ici_eq'' [topological_space α] [partial_order α] [order_topology α] (a : α) : + 𝓝[≥] a = (⨅ u (hu : a < u), 𝓟 (Iio u)) ⊓ 𝓟 (Ici a) := +begin + rw [nhds_within, nhds_eq_order], + refine le_antisymm (inf_le_inf_right _ inf_le_right) (le_inf (le_inf _ inf_le_left) inf_le_right), + exact inf_le_right.trans (le_infi₂ $ λ l hl, principal_mono.2 $ Ici_subset_Ioi.2 hl) +end + +lemma nhds_within_Iic_eq'' [topological_space α] [partial_order α] [order_topology α] (a : α) : + 𝓝[≤] a = (⨅ l < a, 𝓟 (Ioi l)) ⊓ 𝓟 (Iic a) := +nhds_within_Ici_eq'' (to_dual a) + +lemma nhds_within_Ici_eq' [topological_space α] [partial_order α] [order_topology α] {a : α} + (ha : ∃ u, a < u) : + 𝓝[≥] a = ⨅ u (hu : a < u), 𝓟 (Ico a u) := +by simp only [nhds_within_Ici_eq'', binfi_inf ha, inf_principal, Iio_inter_Ici] + +lemma nhds_within_Iic_eq' [topological_space α] [partial_order α] [order_topology α] {a : α} + (ha : ∃ l, l < a) : + 𝓝[≤] a = ⨅ l < a, 𝓟 (Ioc l a) := +by simp only [nhds_within_Iic_eq'', binfi_inf ha, inf_principal, Ioi_inter_Iic] + +lemma nhds_within_Ici_basis' [topological_space α] [linear_order α] [order_topology α] {a : α} + (ha : ∃ u, a < u) : (𝓝[≥] a).has_basis (λ u, a < u) (λ u, Ico a u) := +(nhds_within_Ici_eq' ha).symm ▸ has_basis_binfi_principal (λ b hb c hc, + ⟨min b c, lt_min hb hc, Ico_subset_Ico_right (min_le_left _ _), + Ico_subset_Ico_right (min_le_right _ _)⟩) ha + +lemma nhds_within_Iic_basis' [topological_space α] [linear_order α] [order_topology α] {a : α} + (ha : ∃ l, l < a) : (𝓝[≤] a).has_basis (λ l, l < a) (λ l, Ioc l a) := +by { convert @nhds_within_Ici_basis' αᵒᵈ _ _ _ (to_dual a) ha, + exact funext (λ x, (@dual_Ico _ _ _ _).symm) } + +lemma nhds_within_Ici_basis [topological_space α] [linear_order α] [order_topology α] + [no_max_order α] (a : α) : (𝓝[≥] a).has_basis (λ u, a < u) (λ u, Ico a u) := +nhds_within_Ici_basis' (exists_gt a) + +lemma nhds_within_Iic_basis [topological_space α] [linear_order α] [order_topology α] + [no_min_order α] (a : α) : (𝓝[≤] a).has_basis (λ l, l < a) (λ l, Ioc l a) := +nhds_within_Iic_basis' (exists_lt a) + lemma nhds_top_order [topological_space α] [partial_order α] [order_top α] [order_topology α] : 𝓝 (⊤:α) = (⨅l (h₂ : l < ⊤), 𝓟 (Ioi l)) := by simp [nhds_eq_order (⊤:α)] @@ -910,16 +952,8 @@ by simp [nhds_eq_order (⊥:α)] lemma nhds_top_basis [topological_space α] [linear_order α] [order_top α] [order_topology α] [nontrivial α] : (𝓝 ⊤).has_basis (λ a : α, a < ⊤) (λ a : α, Ioi a) := -⟨ begin - simp only [nhds_top_order], - refine @filter.mem_binfi_of_directed α α (λ a, 𝓟 (Ioi a)) (λ a, a < ⊤) _ _, - { rintros a (ha : a < ⊤) b (hb : b < ⊤), - use a ⊔ b, - simp only [filter.le_principal_iff, ge_iff_le, order.preimage], - exact ⟨sup_lt_iff.mpr ⟨ha, hb⟩, Ioi_subset_Ioi le_sup_left, Ioi_subset_Ioi le_sup_right⟩ }, - { obtain ⟨a, ha⟩ : ∃ a : α, a ≠ ⊤ := exists_ne ⊤, - exact ⟨a, lt_top_iff_ne_top.mpr ha⟩ } - end ⟩ +have ∃ x : α, x < ⊤, from (exists_ne ⊤).imp $ λ x hx, hx.lt_top, +by simpa only [Iic_top, nhds_within_univ, Ioc_top] using nhds_within_Iic_basis' this lemma nhds_bot_basis [topological_space α] [linear_order α] [order_bot α] [order_topology α] [nontrivial α] : @@ -985,42 +1019,75 @@ end order_closed_topology section order_topology variables [order_topology α] +lemma order_separated {a₁ a₂ : α} (h : a₁ < a₂) : + ∃u v : set α, is_open u ∧ is_open v ∧ a₁ ∈ u ∧ a₂ ∈ v ∧ (∀b₁∈u, ∀b₂∈v, b₁ < b₂) := +match dense_or_discrete a₁ a₂ with +| or.inl ⟨a, ha₁, ha₂⟩ := ⟨{a' | a' < a}, {a' | a < a'}, is_open_gt' a, is_open_lt' a, ha₁, ha₂, + assume b₁ h₁ b₂ h₂, lt_trans h₁ h₂⟩ +| or.inr ⟨h₁, h₂⟩ := ⟨{a | a < a₂}, {a | a₁ < a}, is_open_gt' a₂, is_open_lt' a₁, h, h, + assume b₁ hb₁ b₂ hb₂, + calc b₁ ≤ a₁ : h₂ _ hb₁ + ... < a₂ : h + ... ≤ b₂ : h₁ _ hb₂⟩ +end + +@[priority 100] -- see Note [lower instance priority] +instance order_topology.to_order_closed_topology : order_closed_topology α := +{ is_closed_le' := + is_open_compl_iff.1 $ is_open_prod_iff.mpr $ assume a₁ a₂ (h : ¬ a₁ ≤ a₂), + have h : a₂ < a₁, from lt_of_not_ge h, + let ⟨u, v, hu, hv, ha₁, ha₂, h⟩ := order_separated h in + ⟨v, u, hv, hu, ha₂, ha₁, assume ⟨b₁, b₂⟩ ⟨h₁, h₂⟩, not_le_of_gt $ h b₂ h₂ b₁ h₁⟩ } + +lemma exists_Ioc_subset_of_mem_nhds {a : α} {s : set α} (hs : s ∈ 𝓝 a) (h : ∃ l, l < a) : + ∃ l < a, Ioc l a ⊆ s := +(nhds_within_Iic_basis' h).mem_iff.mp (nhds_within_le_nhds hs) + lemma exists_Ioc_subset_of_mem_nhds' {a : α} {s : set α} (hs : s ∈ 𝓝 a) {l : α} (hl : l < a) : ∃ l' ∈ Ico l a, Ioc l' a ⊆ s := -begin - rw [nhds_eq_order a] at hs, - rcases hs with ⟨t₁, ht₁, t₂, ht₂, rfl⟩, - -- First we show that `t₂` includes `(-∞, a]`, so it suffices to show `(l', ∞) ⊆ t₁` - suffices : ∃ l' ∈ Ico l a, Ioi l' ⊆ t₁, - { have A : 𝓟 (Iic a) ≤ ⨅ b ∈ Ioi a, 𝓟 (Iio b), - from (le_infi $ λ b, le_infi $ λ hb, principal_mono.2 $ Iic_subset_Iio.2 hb), - have B : t₁ ∩ Iic a ⊆ t₁ ∩ t₂, - from inter_subset_inter_right _ (A ht₂), - from this.imp (λ l', Exists.imp $ λ hl' hl x hx, B ⟨hl hx.1, hx.2⟩) }, - clear ht₂ t₂, - -- Now we find `l` such that `(l', ∞) ⊆ t₁` - rw [mem_binfi_of_directed] at ht₁, - { rcases ht₁ with ⟨b, hb, hb'⟩, - exact ⟨max b l, ⟨le_max_right _ _, max_lt hb hl⟩, - λ x hx, hb' $ Ioi_subset_Ioi (le_max_left _ _) hx⟩ }, - { intros b hb b' hb', simp only [mem_Iio] at hb hb', - use [max b b', max_lt hb hb'], - simp [le_refl] }, - exact ⟨l, hl⟩ -end +let ⟨l', hl'a, hl's⟩ := exists_Ioc_subset_of_mem_nhds hs ⟨l, hl⟩ +in ⟨max l l', ⟨le_max_left _ _, max_lt hl hl'a⟩, + (Ioc_subset_Ioc_left $ le_max_right _ _).trans hl's⟩ lemma exists_Ico_subset_of_mem_nhds' {a : α} {s : set α} (hs : s ∈ 𝓝 a) {u : α} (hu : a < u) : ∃ u' ∈ Ioc a u, Ico a u' ⊆ s := by simpa only [order_dual.exists, exists_prop, dual_Ico, dual_Ioc] - using exists_Ioc_subset_of_mem_nhds' (show of_dual ⁻¹' s ∈ 𝓝 (to_dual a), from hs) hu.dual - -lemma exists_Ioc_subset_of_mem_nhds {a : α} {s : set α} (hs : s ∈ 𝓝 a) (h : ∃ l, l < a) : - ∃ l < a, Ioc l a ⊆ s := -let ⟨l', hl'⟩ := h in let ⟨l, hl⟩ := exists_Ioc_subset_of_mem_nhds' hs hl' in ⟨l, hl.fst.2, hl.snd⟩ + using exists_Ioc_subset_of_mem_nhds' (show of_dual ⁻¹' s ∈ 𝓝 (to_dual a), from hs) hu.dual lemma exists_Ico_subset_of_mem_nhds {a : α} {s : set α} (hs : s ∈ 𝓝 a) (h : ∃ u, a < u) : ∃ u (_ : a < u), Ico a u ⊆ s := -let ⟨l', hl'⟩ := h in let ⟨l, hl⟩ := exists_Ico_subset_of_mem_nhds' hs hl' in ⟨l, hl.fst.1, hl.snd⟩ +let ⟨l', hl'⟩ := h, ⟨l, hl⟩ := exists_Ico_subset_of_mem_nhds' hs hl' in ⟨l, hl.fst.1, hl.snd⟩ + +lemma exists_Icc_mem_subset_of_mem_nhds_within_Ici {a : α} {s : set α} (hs : s ∈ 𝓝[≥] a) : + ∃ b (_ : a ≤ b), Icc a b ∈ 𝓝[≥] a ∧ Icc a b ⊆ s := +begin + rcases (em (is_max a)).imp_right not_is_max_iff.mp with ha|ha, + { use a, simpa [ha.Ici_eq] using hs }, + { rcases (nhds_within_Ici_basis' ha).mem_iff.mp hs with ⟨b, hab, hbs⟩, + rcases eq_empty_or_nonempty (Ioo a b) with H|⟨c, hac, hcb⟩, + { have : Ico a b = Icc a a, + { rw [← Icc_union_Ioo_eq_Ico le_rfl hab, H, union_empty] }, + exact ⟨a, le_rfl, this ▸ ⟨Ico_mem_nhds_within_Ici $ left_mem_Ico.2 hab, hbs⟩⟩ }, + { refine ⟨c, hac.le, Icc_mem_nhds_within_Ici $ left_mem_Ico.mpr hac, _⟩, + exact (Icc_subset_Ico_right hcb).trans hbs } } +end + +lemma exists_Icc_mem_subset_of_mem_nhds_within_Iic {a : α} {s : set α} (hs : s ∈ 𝓝[≤] a) : + ∃ b ≤ a, Icc b a ∈ 𝓝[≤] a ∧ Icc b a ⊆ s := +by simpa only [dual_Icc, to_dual.surjective.exists] + using @exists_Icc_mem_subset_of_mem_nhds_within_Ici αᵒᵈ _ _ _ (to_dual a) _ hs + +lemma exists_Icc_mem_subset_of_mem_nhds {a : α} {s : set α} (hs : s ∈ 𝓝 a) : + ∃ b c, a ∈ Icc b c ∧ Icc b c ∈ 𝓝 a ∧ Icc b c ⊆ s := +begin + rcases exists_Icc_mem_subset_of_mem_nhds_within_Iic (nhds_within_le_nhds hs) + with ⟨b, hba, hb_nhds, hbs⟩, + rcases exists_Icc_mem_subset_of_mem_nhds_within_Ici (nhds_within_le_nhds hs) + with ⟨c, hac, hc_nhds, hcs⟩, + refine ⟨b, c, ⟨hba, hac⟩, _⟩, + rw [← Icc_union_Icc_eq_Icc hba hac, ← nhds_left_sup_nhds_right], + exact ⟨union_mem_sup hb_nhds hc_nhds, union_subset hbs hcs⟩ +end lemma is_open.exists_Ioo_subset [nontrivial α] {s : set α} (hs : is_open s) (h : s.nonempty) : ∃ a b, a < b ∧ Ioo a b ⊆ s := @@ -1037,26 +1104,6 @@ begin exact ⟨l, x, lx, Ioo_subset_Ioc_self.trans hl⟩ } end -lemma order_separated {a₁ a₂ : α} (h : a₁ < a₂) : - ∃u v : set α, is_open u ∧ is_open v ∧ a₁ ∈ u ∧ a₂ ∈ v ∧ (∀b₁∈u, ∀b₂∈v, b₁ < b₂) := -match dense_or_discrete a₁ a₂ with -| or.inl ⟨a, ha₁, ha₂⟩ := ⟨{a' | a' < a}, {a' | a < a'}, is_open_gt' a, is_open_lt' a, ha₁, ha₂, - assume b₁ h₁ b₂ h₂, lt_trans h₁ h₂⟩ -| or.inr ⟨h₁, h₂⟩ := ⟨{a | a < a₂}, {a | a₁ < a}, is_open_gt' a₂, is_open_lt' a₁, h, h, - assume b₁ hb₁ b₂ hb₂, - calc b₁ ≤ a₁ : h₂ _ hb₁ - ... < a₂ : h - ... ≤ b₂ : h₁ _ hb₂⟩ -end - -@[priority 100] -- see Note [lower instance priority] -instance order_topology.to_order_closed_topology : order_closed_topology α := -{ is_closed_le' := - is_open_compl_iff.1 $ is_open_prod_iff.mpr $ assume a₁ a₂ (h : ¬ a₁ ≤ a₂), - have h : a₂ < a₁, from lt_of_not_ge h, - let ⟨u, v, hu, hv, ha₁, ha₂, h⟩ := order_separated h in - ⟨v, u, hv, hu, ha₂, ha₁, assume ⟨b₁, b₂⟩ ⟨h₁, h₂⟩, not_le_of_gt $ h b₂ h₂ b₁ h₁⟩ } - lemma dense_of_exists_between [nontrivial α] {s : set α} (h : ∀ ⦃a b⦄, a < b → ∃ c ∈ s, a < c ∧ c < b) : dense s := begin @@ -1073,52 +1120,15 @@ lemma dense_iff_exists_between [densely_ordered α] [nontrivial α] {s : set α} dense s ↔ ∀ a b, a < b → ∃ c ∈ s, a < c ∧ c < b := ⟨λ h a b hab, h.exists_between hab, dense_of_exists_between⟩ -lemma order_topology.t2_space : t2_space α := by apply_instance - @[priority 100] -- see Note [lower instance priority] instance order_topology.t3_space : t3_space α := -{ regular := assume s a hs ha, - have hs' : sᶜ ∈ 𝓝 a, from is_open.mem_nhds hs.is_open_compl ha, - have ∃t:set α, is_open t ∧ (∀l∈ s, l < a → l ∈ t) ∧ 𝓝[t] a = ⊥, - from by_cases - (assume h : ∃l, l < a, - let ⟨l, hl, h⟩ := exists_Ioc_subset_of_mem_nhds hs' h in - match dense_or_discrete l a with - | or.inl ⟨b, hb₁, hb₂⟩ := ⟨{a | a < b}, is_open_gt' _, - assume c hcs hca, show c < b, - from lt_of_not_ge $ assume hbc, h ⟨lt_of_lt_of_le hb₁ hbc, le_of_lt hca⟩ hcs, - inf_principal_eq_bot.2 $ (𝓝 a).sets_of_superset ((is_open_lt' _).mem_nhds hb₂) $ - assume x (hx : b < x), show ¬ x < b, from not_lt.2 $ le_of_lt hx⟩ - | or.inr ⟨h₁, h₂⟩ := ⟨{a' | a' < a}, is_open_gt' _, assume b hbs hba, hba, - inf_principal_eq_bot.2 $ (𝓝 a).sets_of_superset ((is_open_lt' _).mem_nhds hl) $ - assume x (hx : l < x), show ¬ x < a, from not_lt.2 $ h₁ _ hx⟩ - end) - (assume : ¬ ∃l, l < a, ⟨∅, is_open_empty, assume l _ hl, (this ⟨l, hl⟩).elim, - nhds_within_empty _⟩), - let ⟨t₁, ht₁o, ht₁s, ht₁a⟩ := this in - have ∃t:set α, is_open t ∧ (∀u∈ s, u>a → u ∈ t) ∧ 𝓝[t] a = ⊥, - from by_cases - (assume h : ∃u, u > a, - let ⟨u, hu, h⟩ := exists_Ico_subset_of_mem_nhds hs' h in - match dense_or_discrete a u with - | or.inl ⟨b, hb₁, hb₂⟩ := ⟨{a | b < a}, is_open_lt' _, - assume c hcs hca, show c > b, - from lt_of_not_ge $ assume hbc, h ⟨le_of_lt hca, lt_of_le_of_lt hbc hb₂⟩ hcs, - inf_principal_eq_bot.2 $ (𝓝 a).sets_of_superset ((is_open_gt' _).mem_nhds hb₁) $ - assume x (hx : b > x), show ¬ x > b, from not_lt.2 $ le_of_lt hx⟩ - | or.inr ⟨h₁, h₂⟩ := ⟨{a' | a' > a}, is_open_lt' _, assume b hbs hba, hba, - inf_principal_eq_bot.2 $ (𝓝 a).sets_of_superset ((is_open_gt' _).mem_nhds hu) $ - assume x (hx : u > x), show ¬ x > a, from not_lt.2 $ h₂ _ hx⟩ - end) - (assume : ¬ ∃u, u > a, ⟨∅, is_open_empty, assume l _ hl, (this ⟨l, hl⟩).elim, - nhds_within_empty _⟩), - let ⟨t₂, ht₂o, ht₂s, ht₂a⟩ := this in - ⟨t₁ ∪ t₂, is_open.union ht₁o ht₂o, - assume x hx, - have x ≠ a, from assume eq, ha $ eq ▸ hx, - (ne_iff_lt_or_gt.mp this).imp (ht₁s _ hx) (ht₂s _ hx), - by rw [nhds_within_union, ht₁a, ht₂a, bot_sup_eq]⟩, - ..order_topology.t2_space } +begin + refine ⟨λ s a hs ha, _⟩, + have : sᶜ ∈ 𝓝 a, from hs.is_open_compl.mem_nhds ha, + rcases exists_Icc_mem_subset_of_mem_nhds this with ⟨b, c, -, hmem, hsub⟩, + refine ⟨(Icc b c)ᶜ, is_closed_Icc.is_open_compl, subset_compl_comm.2 hsub, _⟩, + rwa [nhds_within, inf_principal_eq_bot, compl_compl] +end /-- A set is a neighborhood of `a` if and only if it contains an interval `(l, u)` containing `a`, provided `a` is neither a bottom element nor a top element. -/ @@ -1129,12 +1139,9 @@ begin { assume h, rcases exists_Ico_subset_of_mem_nhds h hu with ⟨u, au, hu⟩, rcases exists_Ioc_subset_of_mem_nhds h hl with ⟨l, la, hl⟩, - refine ⟨l, u, ⟨la, au⟩, λx hx, _⟩, - cases le_total a x with hax hax, - { exact hu ⟨hax, hx.2⟩ }, - { exact hl ⟨hx.1, hax⟩ } }, + exact ⟨l, u, ⟨la, au⟩, Ioc_union_Ico_eq_Ioo la au ▸ union_subset hl hu⟩ }, { rintros ⟨l, u, ha, h⟩, - apply mem_of_superset (is_open.mem_nhds is_open_Ioo ha) h } + apply mem_of_superset (Ioo_mem_nhds ha.1 ha.2) h } end /-- A set is a neighborhood of `a` if and only if it contains an interval `(l, u)` containing `a`. @@ -1489,17 +1496,12 @@ lemma tfae_mem_nhds_within_Ici {a b : α} (hab : a < b) (s : set α) : begin tfae_have : 1 ↔ 2, by rw [nhds_within_Icc_eq_nhds_within_Ici hab], tfae_have : 1 ↔ 3, by rw [nhds_within_Ico_eq_nhds_within_Ici hab], + tfae_have : 1 ↔ 5, from (nhds_within_Ici_basis' ⟨b, hab⟩).mem_iff, tfae_have : 4 → 5, from λ ⟨u, umem, hu⟩, ⟨u, umem.1, hu⟩, - tfae_have : 5 → 1, - { rintros ⟨u, hau, hu⟩, - exact mem_of_superset (Ico_mem_nhds_within_Ici ⟨le_refl a, hau⟩) hu }, - tfae_have : 1 → 4, - { assume h, - rcases mem_nhds_within_iff_exists_mem_nhds_inter.1 h with ⟨v, va, hv⟩, - rcases exists_Ico_subset_of_mem_nhds' va hab with ⟨u, au, hu⟩, - refine ⟨u, au, λx hx, _⟩, - refine hv ⟨hu ⟨hx.1, hx.2⟩, _⟩, - exact hx.1 }, + tfae_have : 5 → 4, + { rintro ⟨u, hua, hus⟩, + exact ⟨min u b, ⟨lt_min hua hab, min_le_right _ _⟩, + (Ico_subset_Ico_right $ min_le_left _ _).trans hus⟩, }, tfae_finish end From bb05277f061e343f6adbb74d08d807907836197b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 19 Sep 2022 07:18:10 +0000 Subject: [PATCH 306/828] refactor(topology/algebra/with_zero_topology): review API, golf (#16359) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * use `γ ≠ 0` instead of `units`; * deduplicate, golf --- src/algebra/order/with_zero.lean | 1 + src/topology/algebra/valued_field.lean | 45 ++- src/topology/algebra/with_zero_topology.lean | 302 +++++++------------ src/topology/order.lean | 15 + src/topology/separation.lean | 13 + 5 files changed, 149 insertions(+), 227 deletions(-) diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index 888b1a41efa30..9d019991854ea 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -25,6 +25,7 @@ in another file. However, the lemmas about it are stated here. set_option old_structure_cmd true /-- A linearly ordered commutative group with a zero element. -/ +@[protect_proj] class linear_ordered_comm_group_with_zero (α : Type*) extends linear_ordered_comm_monoid_with_zero α, comm_group_with_zero α diff --git a/src/topology/algebra/valued_field.lean b/src/topology/algebra/valued_field.lean index a1477dd844ac0..469308e6f535e 100644 --- a/src/topology/algebra/valued_field.lean +++ b/src/topology/algebra/valued_field.lean @@ -122,18 +122,13 @@ lemma valued.continuous_valuation [valued K Γ₀] : continuous (v : K → Γ₀ begin rw continuous_iff_continuous_at, intro x, - classical, - by_cases h : x = 0, - { rw h, - change tendsto _ _ (𝓝 (v (0 : K))), - erw valuation.map_zero, - rw linear_ordered_comm_group_with_zero.tendsto_zero, - intro γ, - rw valued.mem_nhds_zero, - use [γ, set.subset.refl _] }, - { change tendsto _ _ _, - have v_ne : (v x : Γ₀) ≠ 0, from (valuation.ne_zero_iff _).mpr h, - rw linear_ordered_comm_group_with_zero.tendsto_of_ne_zero v_ne, + rcases eq_or_ne x 0 with rfl|h, + { rw [continuous_at, map_zero, linear_ordered_comm_group_with_zero.tendsto_zero], + intros γ hγ, + rw [filter.eventually, valued.mem_nhds_zero], + use [units.mk0 γ hγ, subset.rfl] }, + { have v_ne : (v x : Γ₀) ≠ 0, from (valuation.ne_zero_iff _).mpr h, + rw [continuous_at, linear_ordered_comm_group_with_zero.tendsto_of_ne_zero v_ne], apply valued.loc_const v_ne }, end end @@ -208,13 +203,10 @@ lemma continuous_extension : continuous (valued.extension : hat K → Γ₀) := begin refine completion.dense_inducing_coe.continuous_extend _, intro x₀, - by_cases h : x₀ = coe 0, + rcases eq_or_ne x₀ 0 with rfl|h, { refine ⟨0, _⟩, - erw [h, ← completion.dense_inducing_coe.to_inducing.nhds_eq_comap]; try { apply_instance }, - rw linear_ordered_comm_group_with_zero.tendsto_zero, - intro γ₀, - rw valued.mem_nhds, - exact ⟨γ₀, by simp⟩ }, + erw [← completion.dense_inducing_coe.to_inducing.nhds_eq_comap], + exact valued.continuous_valuation.tendsto' 0 0 (map_zero v) }, { have preimage_one : v ⁻¹' {(1 : Γ₀)} ∈ 𝓝 (1 : K), { have : (v (1 : K) : Γ₀) ≠ 0, { rw valuation.map_one, exact zero_ne_one.symm }, convert valued.loc_const this, @@ -263,27 +255,24 @@ lemma continuous_extension : continuous (valued.extension : hat K → Γ₀) := rcases this with ⟨z₀, y₀, y₀_in, hz₀, z₀_ne⟩, have vz₀_ne: (v z₀ : Γ₀) ≠ 0 := by rwa valuation.ne_zero_iff, refine ⟨v z₀, _⟩, - rw [linear_ordered_comm_group_with_zero.tendsto_of_ne_zero vz₀_ne, mem_comap], - use [(λ x, x*x₀) '' V', nhds_right], - intros x x_in, - rcases mem_preimage.1 x_in with ⟨y, y_in, hy⟩, clear x_in, - change y*x₀ = coe x at hy, - have : (v (x*z₀⁻¹) : Γ₀) = 1, + rw [linear_ordered_comm_group_with_zero.tendsto_of_ne_zero vz₀_ne, eventually_comap], + filter_upwards [nhds_right] with x x_in a ha, + rcases x_in with ⟨y, y_in, rfl⟩, + have : (v (a * z₀⁻¹) : Γ₀) = 1, { apply hV, have : ((z₀⁻¹ : K) : hat K) = z₀⁻¹, from map_inv₀ (completion.coe_ring_hom : K →+* hat K) z₀, - rw [completion.coe_mul, this, ← hy, hz₀, mul_inv, mul_comm y₀⁻¹, ← mul_assoc, mul_assoc y, + rw [completion.coe_mul, this, ha, hz₀, mul_inv, mul_comm y₀⁻¹, ← mul_assoc, mul_assoc y, mul_inv_cancel h, mul_one], solve_by_elim }, - calc v x = v (x*z₀⁻¹*z₀) : by rw [mul_assoc, inv_mul_cancel z₀_ne, mul_one] - ... = v (x*z₀⁻¹)*v z₀ : valuation.map_mul _ _ _ + calc v a = v (a * z₀⁻¹ * z₀) : by rw [mul_assoc, inv_mul_cancel z₀_ne, mul_one] + ... = v (a * z₀⁻¹) * v z₀ : valuation.map_mul _ _ _ ... = v z₀ : by rw [this, one_mul] }, end @[simp, norm_cast] lemma extension_extends (x : K) : extension (x : hat K) = v x := begin - haveI : t2_space Γ₀ := t3_space.t2_space _, refine completion.dense_inducing_coe.extend_eq_of_tendsto _, rw ← completion.dense_inducing_coe.nhds_eq_comap, exact valued.continuous_valuation.continuous_at, diff --git a/src/topology/algebra/with_zero_topology.lean b/src/topology/algebra/with_zero_topology.lean index f93a17272a78f..10c417963430d 100644 --- a/src/topology/algebra/with_zero_topology.lean +++ b/src/topology/algebra/with_zero_topology.lean @@ -15,8 +15,7 @@ Neighborhoods of zero are sets containing `{γ | γ < γ₀}` for some invertibl and every invertible element is open. In particular the topology is the following: "a subset `U ⊆ Γ₀` is open if `0 ∉ U` or if there is an invertible -`γ₀ ∈ Γ₀ such that {γ | γ < γ₀} ⊆ U`", but this fact is not proven here since the neighborhoods -description is what is actually useful. +`γ₀ ∈ Γ₀` such that `{γ | γ < γ₀} ⊆ U`", see `linear_ordered_comm_group_with_zero.is_open_iff`. We prove this topology is ordered and T₃ (in addition to be compatible with the monoid structure). @@ -33,253 +32,158 @@ All other instances will (`ordered_topology`, `t3_space`, `has_continuous_mul`) -/ -open_locale topological_space -open topological_space filter set +open_locale topological_space filter +open topological_space filter set function namespace linear_ordered_comm_group_with_zero -variables (Γ₀ : Type*) [linear_ordered_comm_group_with_zero Γ₀] - -/-- The neighbourhoods around γ ∈ Γ₀, used in the definition of the topology on Γ₀. -These neighbourhoods are defined as follows: -A set s is a neighbourhood of 0 if there is an invertible γ₀ ∈ Γ₀ such that {γ | γ < γ₀} ⊆ s. -If γ ≠ 0, then every set that contains γ is a neighbourhood of γ. -/ -def nhds_fun (x : Γ₀) : filter Γ₀ := -if x = 0 then ⨅ (γ₀ : Γ₀ˣ), principal {γ | γ < γ₀} else pure x +variables {α Γ₀ : Type*} [linear_ordered_comm_group_with_zero Γ₀] {γ γ₁ γ₂ : Γ₀} {l : filter α} + {f : α → Γ₀} /-- The topology on a linearly ordered commutative group with a zero element adjoined. A subset U is open if 0 ∉ U or if there is an invertible element γ₀ such that {γ | γ < γ₀} ⊆ U. -/ protected def topological_space : topological_space Γ₀ := -topological_space.mk_of_nhds (nhds_fun Γ₀) +topological_space.mk_of_nhds $ update pure 0 $ ⨅ γ ≠ 0, 𝓟 (Iio γ) local attribute [instance] linear_ordered_comm_group_with_zero.topological_space -/-- The neighbourhoods {γ | γ < γ₀} of 0 form a directed set indexed by the invertible -elements γ₀. -/ -lemma directed_lt : directed (≥) (λ γ₀ : Γ₀ˣ, principal {γ : Γ₀ | γ < γ₀}) := -begin - intros γ₁ γ₂, - use linear_order.min γ₁ γ₂ ; dsimp only, - split ; rw [ge_iff_le, principal_mono] ; intros x x_in, - { calc x < ↑(linear_order.min γ₁ γ₂) : x_in - ... ≤ γ₁ : min_le_left γ₁ γ₂ }, - { calc x < ↑(linear_order.min γ₁ γ₂) : x_in - ... ≤ γ₂ : min_le_right γ₁ γ₂ } -end +lemma nhds_eq_update : (𝓝 : Γ₀ → filter Γ₀) = update pure 0 (⨅ γ ≠ 0, 𝓟 (Iio γ)) := +funext $ nhds_mk_of_nhds_single $ le_infi₂ $ λ γ h₀, le_principal_iff.2 $ zero_lt_iff.2 h₀ --- We need two auxilliary lemmas to show that nhds_fun accurately describes the neighbourhoods --- coming from the topology (that is defined in terms of nhds_fun). +/-! +### Neighbourhoods of zero +-/ -/-- At all points of a linearly ordered commutative group with a zero element adjoined, -the pure filter is smaller than the filter given by nhds_fun. -/ -lemma pure_le_nhds_fun : pure ≤ nhds_fun Γ₀ := -λ x, by { by_cases hx : x = 0; simp [hx, nhds_fun] } +lemma nhds_zero : 𝓝 (0 : Γ₀) = ⨅ γ ≠ 0, 𝓟 (Iio γ) := by rw [nhds_eq_update, update_same] -/-- For every point Γ₀, and every “neighbourhood” s of it (described by nhds_fun), there is a -smaller “neighbourhood” t ⊆ s, such that s is a “neighbourhood“ of all the points in t. -/ -lemma nhds_fun_ok (x : Γ₀) {s} (s_in : s ∈ nhds_fun Γ₀ x) : - (∃ t ∈ nhds_fun Γ₀ x, t ⊆ s ∧ ∀ y ∈ t, s ∈ nhds_fun Γ₀ y) := +/-- 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`. -/ +lemma has_basis_nhds_zero : (𝓝 (0 : Γ₀)).has_basis (λ γ : Γ₀, γ ≠ 0) Iio := begin - by_cases hx : x = 0, - { simp only [hx, nhds_fun, exists_prop, if_true, eq_self_iff_true] at s_in ⊢, - cases (mem_infi_of_directed (directed_lt Γ₀) _).mp s_in with γ₀ h, - use {γ : Γ₀ | γ < γ₀}, - rw mem_principal at h, - split, - { apply mem_infi_of_mem γ₀, - rw mem_principal }, - { refine ⟨h, λ y y_in, _⟩, - by_cases hy : y = 0, - { simp only [hy, if_true, eq_self_iff_true], - apply mem_infi_of_mem γ₀, - rwa mem_principal }, - { simp [hy, h y_in] } } }, - { simp only [hx, nhds_fun, exists_prop, if_false, mem_pure] at s_in ⊢, - refine ⟨{x}, mem_singleton _, singleton_subset_iff.2 s_in, λ y y_in, _⟩, - simpa [mem_singleton_iff.mp y_in, hx] } + rw [nhds_zero], + refine has_basis_binfi_principal _ ⟨1, one_ne_zero⟩, + exact directed_on_iff_directed.2 (directed_of_inf $ λ a b hab, Iio_subset_Iio hab) end -variables {Γ₀} +lemma Iio_mem_nhds_zero (hγ : γ ≠ 0) : Iio γ ∈ 𝓝 (0 : Γ₀) := has_basis_nhds_zero.mem_of_mem hγ -/-- The neighbourhood filter of an invertible element consists of all sets containing that -element. -/ -lemma nhds_coe_units (γ : Γ₀ˣ) : 𝓝 (γ : Γ₀) = pure (γ : Γ₀) := -calc 𝓝 (γ : Γ₀) = nhds_fun Γ₀ γ : nhds_mk_of_nhds (nhds_fun Γ₀) γ (pure_le_nhds_fun Γ₀) - (nhds_fun_ok Γ₀) - ... = pure (γ : Γ₀) : if_neg γ.ne_zero +/-- If `γ` is an invertible element of a linearly ordered group with zero element adjoined, then +`Iio (γ : Γ₀)` is a neighbourhood of `0`. -/ +lemma nhds_zero_of_units (γ : Γ₀ˣ) : Iio ↑γ ∈ 𝓝 (0 : Γ₀) := Iio_mem_nhds_zero γ.ne_zero + +lemma tendsto_zero : tendsto f l (𝓝 (0 : Γ₀)) ↔ ∀ γ₀ ≠ 0, ∀ᶠ x in l, f x < γ₀ := by simp [nhds_zero] + +/-! +### Neighbourhoods of non-zero elements +-/ /-- The neighbourhood filter of a nonzero element consists of all sets containing that element. -/ -@[simp] lemma nhds_of_ne_zero (γ : Γ₀) (h : γ ≠ 0) : - 𝓝 γ = pure γ := -nhds_coe_units (units.mk0 _ h) - -/-- If γ is an invertible element of a linearly ordered group with zero element adjoined, -then {γ} is a neighbourhood of γ. -/ -lemma singleton_nhds_of_units (γ : Γ₀ˣ) : ({γ} : set Γ₀) ∈ 𝓝 (γ : Γ₀) := -by simp - -/-- If γ is a nonzero element of a linearly ordered group with zero element adjoined, -then {γ} is a neighbourhood of γ. -/ -lemma singleton_nhds_of_ne_zero (γ : Γ₀) (h : γ ≠ 0) : ({γ} : set Γ₀) ∈ 𝓝 (γ : Γ₀) := -by simp [h] - -/-- If U is a neighbourhood of 0 in a linearly ordered group with zero element adjoined, -then there exists an invertible element γ₀ such that {γ | γ < γ₀} ⊆ U. -/ -lemma has_basis_nhds_zero : - has_basis (𝓝 (0 : Γ₀)) (λ _, true) (λ γ₀ : Γ₀ˣ, {γ : Γ₀ | γ < γ₀}) := -⟨begin - intro U, - rw nhds_mk_of_nhds (nhds_fun Γ₀) 0 (pure_le_nhds_fun Γ₀) (nhds_fun_ok Γ₀), - simp only [nhds_fun, if_true, eq_self_iff_true, exists_true_left], - simp_rw [mem_infi_of_directed (directed_lt Γ₀), mem_principal] -end⟩ +@[simp] lemma nhds_of_ne_zero {γ : Γ₀} (h₀ : γ ≠ 0) : 𝓝 γ = pure γ := +by rw [nhds_eq_update, update_noteq h₀] -/-- If γ is an invertible element of a linearly ordered group with zero element adjoined, -then {x | x < γ} is a neighbourhood of 0. -/ -lemma nhds_zero_of_units (γ : Γ₀ˣ) : {x : Γ₀ | x < γ} ∈ 𝓝 (0 : Γ₀) := -by { rw has_basis_nhds_zero.mem_iff, use γ, simp } +/-- The neighbourhood filter of an invertible element consists of all sets containing that +element. -/ +lemma nhds_coe_units (γ : Γ₀ˣ) : 𝓝 (γ : Γ₀) = pure (γ : Γ₀) := nhds_of_ne_zero γ.ne_zero + +/-- If `γ` is an invertible element of a linearly ordered group with zero element adjoined, then +`{γ}` is a neighbourhood of `γ`. -/ +lemma singleton_mem_nhds_of_units (γ : Γ₀ˣ) : ({γ} : set Γ₀) ∈ 𝓝 (γ : Γ₀) := by simp -lemma tendsto_zero {α : Type*} {F : filter α} {f : α → Γ₀} : - tendsto f F (𝓝 (0 : Γ₀)) ↔ ∀ γ₀ : Γ₀ˣ, { x : α | f x < γ₀ } ∈ F := -by simpa using has_basis_nhds_zero.tendsto_right_iff +/-- If `γ` is a nonzero element of a linearly ordered group with zero element adjoined, then `{γ}` +is a neighbourhood of `γ`. -/ +lemma singleton_mem_nhds_of_ne_zero (h : γ ≠ 0) : ({γ} : set Γ₀) ∈ 𝓝 (γ : Γ₀) := by simp [h] -/-- If γ is a nonzero element of a linearly ordered group with zero element adjoined, -then {x | x < γ} is a neighbourhood of 0. -/ -lemma nhds_zero_of_ne_zero (γ : Γ₀) (h : γ ≠ 0) : {x : Γ₀ | x < γ} ∈ 𝓝 (0 : Γ₀) := -nhds_zero_of_units (units.mk0 _ h) +lemma has_basis_nhds_of_ne_zero {x : Γ₀} (h : x ≠ 0) : + has_basis (𝓝 x) (λ i : unit, true) (λ i, {x}) := +by { rw [nhds_of_ne_zero h], exact has_basis_pure _ } lemma has_basis_nhds_units (γ : Γ₀ˣ) : has_basis (𝓝 (γ : Γ₀)) (λ i : unit, true) (λ i, {γ}) := -begin - rw nhds_of_ne_zero _ γ.ne_zero, - exact has_basis_pure γ -end +has_basis_nhds_of_ne_zero γ.ne_zero -lemma has_basis_nhds_of_ne_zero {x : Γ₀} (h : x ≠ 0) : - has_basis (𝓝 x) (λ i : unit, true) (λ i, {x}) := -has_basis_nhds_units (units.mk0 x h) +lemma tendsto_of_ne_zero {γ : Γ₀} (h : γ ≠ 0) : tendsto f l (𝓝 γ) ↔ ∀ᶠ x in l, f x = γ := +by rw [nhds_of_ne_zero h, tendsto_pure] -lemma singleton_mem_nhds_of_ne_zero {x : Γ₀} (h : x ≠ 0) : {x} ∈ 𝓝 x := -begin - apply (has_basis_nhds_of_ne_zero h).mem_of_mem true.intro, - exact unit.star, -end +lemma tendsto_units {γ₀ : Γ₀ˣ} : tendsto f l (𝓝 (γ₀ : Γ₀)) ↔ ∀ᶠ x in l, f x = γ₀ := +tendsto_of_ne_zero γ₀.ne_zero + +lemma Iio_mem_nhds (h : γ₁ < γ₂) : Iio γ₂ ∈ 𝓝 γ₁ := +by rcases eq_or_ne γ₁ 0 with rfl|h₀; simp [*, h.ne', Iio_mem_nhds_zero] + +/-! +### Open/closed sets +-/ -lemma tendsto_units {α : Type*} {F : filter α} {f : α → Γ₀} {γ₀ : Γ₀ˣ} : - tendsto f F (𝓝 (γ₀ : Γ₀)) ↔ { x : α | f x = γ₀ } ∈ F := +lemma is_open_iff {s : set Γ₀} : is_open s ↔ (0 : Γ₀) ∉ s ∨ ∃ γ ≠ 0, Iio γ ⊆ s := begin - rw (has_basis_nhds_units γ₀).tendsto_right_iff, - simpa + rw [is_open_iff_mem_nhds, ← and_forall_ne (0 : Γ₀)], + simp [nhds_of_ne_zero, imp_iff_not_or, has_basis_nhds_zero.mem_iff] { contextual := tt } end -lemma tendsto_of_ne_zero {α : Type*} {F : filter α} {f : α → Γ₀} {γ : Γ₀} (h : γ ≠ 0) : - tendsto f F (𝓝 γ) ↔ { x : α | f x = γ } ∈ F := -@tendsto_units _ _ _ F f (units.mk0 γ h) +lemma is_closed_iff {s : set Γ₀} : is_closed s ↔ (0 : Γ₀) ∈ s ∨ ∃ γ ≠ 0, s ⊆ Ici γ := +by simp only [← is_open_compl_iff, is_open_iff, mem_compl_iff, not_not, ← compl_Ici, + compl_subset_compl] -variable (Γ₀) +lemma is_open_Iio {a : Γ₀} : is_open (Iio a) := +is_open_iff.mpr $ imp_iff_not_or.mp $ λ ha, ⟨a, ne_of_gt ha, subset.rfl⟩ -/-- The topology on a linearly ordered group with zero element adjoined -is compatible with the order structure. -/ +/-! +### Instances +-/ + +/-- The topology on a linearly ordered group with zero element adjoined is compatible with the order +structure: the set `{p : Γ₀ × Γ₀ | p.1 ≤ p.2}` is closed. -/ @[priority 100] -instance ordered_topology : order_closed_topology Γ₀ := +instance order_closed_topology : order_closed_topology Γ₀ := { is_closed_le' := begin - rw ← is_open_compl_iff, - show is_open {p : Γ₀ × Γ₀ | ¬p.fst ≤ p.snd}, - simp only [not_le], - rw is_open_iff_mem_nhds, - rintros ⟨a,b⟩ hab, - change b < a at hab, - have ha : a ≠ 0 := ne_zero_of_lt hab, - rw [nhds_prod_eq, mem_prod_iff], - by_cases hb : b = 0, - { subst b, - use [{a}, singleton_nhds_of_ne_zero _ ha, {x : Γ₀ | x < a}, nhds_zero_of_ne_zero _ ha], - intros p p_in, - cases mem_prod.1 p_in with h1 h2, - rw mem_singleton_iff at h1, - change p.2 < p.1, - rwa h1 }, - { use [{a}, singleton_nhds_of_ne_zero _ ha, {b}, singleton_nhds_of_ne_zero _ hb], - intros p p_in, - cases mem_prod.1 p_in with h1 h2, - rw mem_singleton_iff at h1 h2, - change p.2 < p.1, - rwa [h1, h2] } + simp only [← is_open_compl_iff, compl_set_of, not_le, is_open_iff_mem_nhds], + rintros ⟨a, b⟩ (hab : b < a), + rw [nhds_prod_eq, nhds_of_ne_zero (zero_le'.trans_lt hab).ne', pure_prod], + exact Iio_mem_nhds hab end } /-- The topology on a linearly ordered group with zero element adjoined is T₃. -/ @[priority 100] instance t3_space : t3_space Γ₀ := -begin - haveI : t1_space Γ₀ := t2_space.t1_space, - split, - intros s x s_closed x_not_in_s, - by_cases hx : x = 0, - { refine ⟨s, _, subset.rfl, _⟩, - { subst x, - rw is_open_iff_mem_nhds, - intros y hy, - by_cases hy' : y = 0, { subst y, contradiction }, - simpa [hy'] }, - { erw inf_eq_bot_iff, - use sᶜ, - simp only [exists_prop, mem_principal], - exact ⟨s_closed.compl_mem_nhds x_not_in_s, ⟨s, subset.refl s, by simp⟩⟩ } }, - { simp only [nhds_within, inf_eq_bot_iff, exists_prop, mem_principal], - exact ⟨{x}ᶜ, is_open_compl_iff.mpr is_closed_singleton, by rwa subset_compl_singleton_iff, - {x}, singleton_nhds_of_ne_zero x hx, {x}ᶜ, by simp [subset.refl]⟩ } -end +t3_space.of_lift'_closure $ λ γ, + begin + rcases ne_or_eq γ 0 with h₀|rfl, + { rw [nhds_of_ne_zero h₀, lift'_pure (monotone_closure Γ₀), closure_singleton, + principal_singleton] }, + { exact has_basis_nhds_zero.lift'_closure_eq_self + (λ x hx, is_closed_iff.2 $ or.inl $ zero_lt_iff.2 hx) }, + end /-- The topology on a linearly ordered group with zero element adjoined makes it a topological monoid. -/ @[priority 100] instance : has_continuous_mul Γ₀ := ⟨begin - have common : ∀ y ≠ (0 : Γ₀), continuous_at (λ (p : Γ₀ × Γ₀), p.fst * p.snd) (0, y), - { intros y hy, - set γ := units.mk0 y hy, - suffices : tendsto (λ (p : Γ₀ × Γ₀), p.fst * p.snd) ((𝓝 0).prod (𝓝 γ)) (𝓝 0), - by simpa [continuous_at, nhds_prod_eq], - suffices : ∀ (γ' : Γ₀ˣ), ∃ (γ'' : Γ₀ˣ), ∀ (a b : Γ₀), a < γ'' → b = y → a * b < γ', - { rw (has_basis_nhds_zero.prod $ has_basis_nhds_units γ).tendsto_iff has_basis_nhds_zero, - simpa }, - intros γ', - use γ⁻¹*γ', - rintros a b ha hb, - rw [hb, mul_comm], - rw [units.coe_mul] at ha, - simpa using inv_mul_lt_of_lt_mul₀ ha }, rw continuous_iff_continuous_at, rintros ⟨x, y⟩, - by_cases hx : x = 0; by_cases hy : y = 0, - { suffices : tendsto (λ (p : Γ₀ × Γ₀), p.fst * p.snd) (𝓝 (0, 0)) (𝓝 0), - by simpa [hx, hy, continuous_at], - suffices : ∀ (γ : Γ₀ˣ), ∃ (γ' : Γ₀ˣ), ∀ (a b : Γ₀), a < γ' → b < γ' → a * b < γ, - by simpa [nhds_prod_eq, has_basis_nhds_zero.prod_self.tendsto_iff has_basis_nhds_zero], - intros γ, - rcases exists_square_le γ with ⟨γ', h⟩, - use γ', - intros a b ha hb, - calc a*b < γ'*γ' : mul_lt_mul₀ ha hb - ... ≤ γ : by exact_mod_cast h }, - { rw hx, - exact common y hy }, - { rw hy, - have : (λ (p : Γ₀ × Γ₀), p.fst * p.snd) = - (λ (p : Γ₀ × Γ₀), p.fst * p.snd) ∘ (λ p : Γ₀ × Γ₀, (p.2, p.1)), - by { ext, rw [mul_comm] }, - rw this, - apply continuous_at.comp _ continuous_swap.continuous_at, - exact common x hx }, - { change tendsto _ _ _, - rw [nhds_prod_eq], - rw ((has_basis_nhds_of_ne_zero hx).prod (has_basis_nhds_of_ne_zero hy)).tendsto_iff - (has_basis_nhds_of_ne_zero $ mul_ne_zero hx hy), - suffices : ∀ (a b : Γ₀), a = x → b = y → a * b = x * y, by simpa, - rintros a b rfl rfl, - refl }, + wlog hle : x ≤ y := le_total x y using [x y, y x] tactic.skip, swap, + { simpa only [mul_comm, (∘), prod.swap] + using tendsto.comp this (continuous_swap.tendsto (x, y)) }, + rcases eq_or_ne x 0 with rfl|hx; [rcases eq_or_ne y 0 with rfl|hy, skip], + { rw [continuous_at, zero_mul], + refine ((has_basis_nhds_zero.prod_nhds has_basis_nhds_zero).tendsto_iff has_basis_nhds_zero).2 + (λ γ hγ, ⟨(γ, 1), ⟨hγ, one_ne_zero⟩, _⟩), + rintro ⟨x, y⟩ ⟨hx : x < γ, hy : y < 1⟩, + exact (mul_lt_mul₀ hx hy).trans_eq (mul_one γ) }, + { rw [continuous_at, zero_mul, nhds_prod_eq, nhds_of_ne_zero hy, prod_pure, tendsto_map'_iff], + refine (has_basis_nhds_zero.tendsto_iff has_basis_nhds_zero).2 (λ γ hγ, _), + refine ⟨γ / y, div_ne_zero hγ hy, λ x hx, _⟩, + calc x * y < γ / y * y : mul_lt_right₀ _ hx hy + ... = γ : div_mul_cancel _ hy }, + { have hy : y ≠ 0, from ((zero_lt_iff.mpr hx).trans_le hle).ne', + rw [continuous_at, nhds_prod_eq, nhds_of_ne_zero hx, nhds_of_ne_zero hy, prod_pure_pure], + exact pure_le_nhds (x * y) } end⟩ +@[priority 100] +instance : has_continuous_inv₀ Γ₀ := +⟨λ γ h, by { rw [continuous_at, nhds_of_ne_zero h], exact pure_le_nhds γ⁻¹ }⟩ + end linear_ordered_comm_group_with_zero diff --git a/src/topology/order.lean b/src/topology/order.lean index b9f26affdd12c..a8d5e6d96db3d 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -116,6 +116,21 @@ begin exact (n a).sets_of_superset (ht _ hat) hts }, end +lemma nhds_mk_of_nhds_single [decidable_eq α] {a₀ : α} {l : filter α} (h : pure a₀ ≤ l) (b : α) : + @nhds α (topological_space.mk_of_nhds $ update pure a₀ l) b = + (update pure a₀ l : α → filter α) b := +begin + refine nhds_mk_of_nhds _ _ (le_update_iff.mpr ⟨h, λ _ _, le_rfl⟩) (λ a s hs, _), + rcases eq_or_ne a a₀ with rfl|ha, + { refine ⟨s, hs, subset.rfl, λ b hb, _⟩, + rcases eq_or_ne b a with rfl|hb, + { exact hs }, + { rwa [update_noteq hb] } }, + { have hs' := hs, + rw [update_noteq ha] at hs ⊢, + exact ⟨{a}, rfl, singleton_subset_iff.mpr hs, forall_eq.2 hs'⟩ } +end + lemma nhds_mk_of_nhds_filter_basis (B : α → filter_basis α) (a : α) (h₀ : ∀ x (n ∈ B x), x ∈ n) (h₁ : ∀ x (n ∈ B x), ∃ n₁ ∈ B x, n₁ ⊆ n ∧ ∀ x' ∈ n₁, ∃ n₂ ∈ B x', n₂ ⊆ n) : @nhds α (topological_space.mk_of_nhds (λ x, (B x).filter)) a = (B a).filter := diff --git a/src/topology/separation.lean b/src/topology/separation.lean index e47e5bb58dddc..73ec638f59bb9 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -1282,6 +1282,19 @@ section t3 class t3_space (α : Type u) [topological_space α] extends t0_space α : Prop := (regular : ∀{s:set α} {a}, is_closed s → a ∉ s → ∃t, is_open t ∧ s ⊆ t ∧ 𝓝[t] a = ⊥) +lemma t3_space.of_lift'_closure [t0_space α] (h : ∀ x : α, (𝓝 x).lift' closure = 𝓝 x) : + t3_space α := +begin + refine ⟨λ s a hs ha, _⟩, + have : sᶜ ∈ (𝓝 a).lift' closure, + { rw [h], exact hs.is_open_compl.mem_nhds ha }, + rcases (𝓝 a).basis_sets.lift'_closure.mem_iff.mp this with ⟨U, haU, hU⟩, + refine ⟨(closure U)ᶜ, is_closed_closure.is_open_compl, subset_compl_comm.mp hU, not_not.mp _⟩, + rw [← ne, ← ne_bot_iff, ← mem_closure_iff_nhds_within_ne_bot, closure_compl, mem_compl_iff, + not_not, mem_interior_iff_mem_nhds], + exact mem_of_superset haU subset_closure +end + @[priority 100] -- see Note [lower instance priority] instance t3_space.t1_space [t3_space α] : t1_space α := begin From 85c08a609ab9abfc8a519786da796d896c306a06 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 19 Sep 2022 07:18:11 +0000 Subject: [PATCH 307/828] feat(algebra/order/ring): `linear_ordered_comm_semiring` (#16456) Add `linear_ordered_comm_semiring`, a typeclass that was missing from the algebraic order hierarchy for some reason and which acquired a few possible uses over time. --- src/algebra/algebra/subalgebra/basic.lean | 4 +- src/algebra/order/field_defs.lean | 2 +- src/algebra/order/positive/ring.lean | 10 ++--- src/algebra/order/ring.lean | 40 +++++++++++++++++-- src/combinatorics/pigeonhole.lean | 48 ++++++----------------- src/data/nat/basic.lean | 33 +++++++--------- src/ring_theory/subsemiring/basic.lean | 12 +++++- 7 files changed, 80 insertions(+), 69 deletions(-) diff --git a/src/algebra/algebra/subalgebra/basic.lean b/src/algebra/algebra/subalgebra/basic.lean index 1d0f93eba1835..620ae80cf9b58 100644 --- a/src/algebra/algebra/subalgebra/basic.lean +++ b/src/algebra/algebra/subalgebra/basic.lean @@ -201,7 +201,9 @@ instance to_ordered_comm_ring {R A} instance to_linear_ordered_semiring {R A} [comm_semiring R] [linear_ordered_semiring A] [algebra R A] (S : subalgebra R A) : linear_ordered_semiring S := S.to_subsemiring.to_linear_ordered_semiring -/-! There is no `linear_ordered_comm_semiring`. -/ +instance to_linear_ordered_comm_semiring {R A} + [comm_semiring R] [linear_ordered_comm_semiring A] [algebra R A] (S : subalgebra R A) : + linear_ordered_comm_semiring S := S.to_subsemiring.to_linear_ordered_comm_semiring instance to_linear_ordered_ring {R A} [comm_ring R] [linear_ordered_ring A] [algebra R A] (S : subalgebra R A) : linear_ordered_ring S := S.to_subring.to_linear_ordered_ring diff --git a/src/algebra/order/field_defs.lean b/src/algebra/order/field_defs.lean index 4b7b97c0ea97f..ee0bc5d2c7b0f 100644 --- a/src/algebra/order/field_defs.lean +++ b/src/algebra/order/field_defs.lean @@ -36,7 +36,7 @@ variables {α : Type*} /-- A linear ordered semifield is a field with a linear order respecting the operations. -/ @[protect_proj] class linear_ordered_semifield (α : Type*) - extends linear_ordered_semiring α, semifield α + extends linear_ordered_comm_semiring α, semifield α /-- A linear ordered field is a field with a linear order respecting the operations. -/ @[protect_proj] class linear_ordered_field (α : Type*) extends linear_ordered_comm_ring α, field α diff --git a/src/algebra/order/positive/ring.lean b/src/algebra/order/positive/ring.lean index 623a60ea8f757..3a8990a88684f 100644 --- a/src/algebra/order/positive/ring.lean +++ b/src/algebra/order/positive/ring.lean @@ -102,14 +102,10 @@ instance [ordered_comm_semiring R] [nontrivial R] : ordered_comm_monoid {x : R / .. subtype.coe_injective.comm_monoid (coe : {x : R // 0 < x} → R) coe_one coe_mul coe_pow } /-- If `R` is a nontrivial linear ordered commutative semiring, then `{x : R // 0 < x}` is a linear -ordered cancellative commutative monoid. We don't have a typeclass for linear ordered commutative -semirings, so we assume `[linear_ordered_semiring R] [is_commutative R (*)] instead. -/ -instance [linear_ordered_semiring R] [is_commutative R (*)] [nontrivial R] : - linear_ordered_cancel_comm_monoid {x : R // 0 < x} := +ordered cancellative commutative monoid. -/ +instance [linear_ordered_comm_semiring R] : linear_ordered_cancel_comm_monoid {x : R // 0 < x} := { le_of_mul_le_mul_left := λ a b c h, subtype.coe_le_coe.1 $ (mul_le_mul_left a.2).1 h, - .. subtype.linear_order _, - .. @positive.subtype.ordered_comm_monoid R - { mul_comm := is_commutative.comm, .. ‹linear_ordered_semiring R› } _ } + .. subtype.linear_order _, .. positive.subtype.ordered_comm_monoid } end mul_comm diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 21da4334cc8e9..9eb3aca53cfe9 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -34,6 +34,8 @@ For short, `*` respects `<`. * `linear_ordered_semiring`: Semiring with a linear order such that `+` respects `≤` and `*` respects `<`. +* `linear_ordered_comm_semiring`: Commutative semiring with a linear order such that `+` respects + `≤` and `*` respects `<`. * `linear_ordered_ring`: Ring with a linear order such that `+` respects `≤` and `*` respects `<`. * `linear_ordered_comm_ring`: Commutative ring with a linear order such that `+` respects `≤` and `*` respects `<`. @@ -66,6 +68,9 @@ immediate predecessors and what conditions are added to each of them. - `comm_ring` & partial order structure & `+` respects `≤` & `*` respects `<` * `linear_ordered_semiring` - `ordered_semiring` & totality of the order & nontriviality +* `linear_ordered_comm_semiring` + - `ordered_comm_semiring` & totality of the order & nontriviality + - `linear_ordered_semiring` & commutativity of multiplication - `linear_ordered_add_comm_monoid` & multiplication & nontriviality & `*` respects `<` * `linear_ordered_ring` - `ordered_ring` & totality of the order & nontriviality @@ -75,6 +80,7 @@ immediate predecessors and what conditions are added to each of them. * `linear_ordered_comm_ring` - `ordered_comm_ring` & totality of the order & nontriviality - `linear_ordered_ring` & commutativity of multiplication + - `linear_ordered_comm_semiring` & additive inverses - `is_domain` & linear order structure * `canonically_ordered_comm_semiring` - `canonically_ordered_add_monoid` & multiplication & `*` respects `<` & no zero divisors @@ -83,11 +89,12 @@ immediate predecessors and what conditions are added to each of them. ## TODO We're still missing some typeclasses, like -* `linear_ordered_comm_semiring` * `canonically_ordered_semiring` They have yet to come up in practice. -/ +open function + set_option old_structure_cmd true universe u @@ -582,7 +589,7 @@ class ordered_comm_semiring (α : Type u) extends ordered_semiring α, comm_semi See note [reducible non-instances]. -/ @[reducible] def function.injective.ordered_comm_semiring [ordered_comm_semiring α] {β : Type*} - [add_monoid_with_one β] [has_mul β] [has_pow β ℕ] + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) @@ -782,6 +789,31 @@ lemma min_mul_of_nonneg (a b : α) (hc : 0 ≤ c) : min a b * c = min (a * c) (b end linear_ordered_semiring +/-- A `linear_ordered_comm_semiring` is a nontrivial commutative semiring with a linear order such +that addition is monotone and multiplication by a positive number is strictly monotone. -/ +@[protect_proj, ancestor ordered_comm_semiring linear_ordered_semiring] +class linear_ordered_comm_semiring (α : Type*) + extends ordered_comm_semiring α, linear_ordered_semiring α + +@[priority 100] -- See note [lower instance priority] +instance linear_ordered_comm_semiring.to_linear_ordered_cancel_add_comm_monoid + [linear_ordered_comm_semiring α] : linear_ordered_cancel_add_comm_monoid α := +{ ..‹linear_ordered_comm_semiring α› } + +/-- Pullback a `linear_ordered_semiring` under an injective map. +See note [reducible non-instances]. -/ +@[reducible] +def function.injective.linear_ordered_comm_semiring [linear_ordered_comm_semiring α] {β : Type*} + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] + [has_sup β] [has_inf β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) + (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_comm_semiring β := +{ ..hf.linear_ordered_semiring f zero one add mul nsmul npow nat_cast hsup hinf, + ..hf.ordered_comm_semiring f zero one add mul nsmul npow nat_cast } + /-- An `ordered_ring α` is a ring `α` with a partial order such that addition is monotone and multiplication by a positive number is strictly monotone. -/ @[protect_proj] @@ -1237,8 +1269,8 @@ instance linear_ordered_comm_ring.to_ordered_comm_ring [d : linear_ordered_comm_ { ..d } @[priority 100] -- see Note [lower instance priority] -instance linear_ordered_comm_ring.to_linear_ordered_semiring [d : linear_ordered_comm_ring α] : - linear_ordered_semiring α := +instance linear_ordered_comm_ring.to_linear_ordered_comm_semiring [d : linear_ordered_comm_ring α] : + linear_ordered_comm_semiring α := { .. d, ..linear_ordered_ring.to_linear_ordered_semiring } section linear_ordered_comm_ring diff --git a/src/combinatorics/pigeonhole.lean b/src/combinatorics/pigeonhole.lean index e40201be6f968..08307cf5dddab 100644 --- a/src/combinatorics/pigeonhole.lean +++ b/src/combinatorics/pigeonhole.lean @@ -54,13 +54,6 @@ docstrings instead of the names. `measure_theory.exists_nonempty_inter_of_measure_univ_lt_sum_measure`: pigeonhole principle in a measure space. -## TODO - -The `_nsmul` lemmas could be generalized from `linear_ordered_comm_ring` to -`linear_ordered_comm_semiring` if the latter existed (or some combination of -`covariant`/`contravariant` classes once the refactor has gone deep enough). This would allow -deriving the `_mul` lemmas from the `_nsmul` ones. - ## Tags pigeonhole principle @@ -198,7 +191,7 @@ lemma exists_sum_fiber_le_of_sum_fiber_nonneg_of_sum_le_nsmul end -variables [linear_ordered_comm_ring M] +variables [linear_ordered_comm_semiring M] /-! ### The pigeonhole principles on `finset`s, pigeons counted by heads @@ -237,11 +230,7 @@ elements. -/ lemma exists_lt_card_fiber_of_mul_lt_card_of_maps_to (hf : ∀ a ∈ s, f a ∈ t) (hn : t.card * n < s.card) : ∃ y ∈ t, n < (s.filter (λ x, f x = y)).card := -begin - simp only [card_eq_sum_ones], - apply exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum hf, - simpa -end +exists_lt_card_fiber_of_nsmul_lt_card_of_maps_to hf hn /-- The pigeonhole principle for finitely many pigeons counted by heads: there is a pigeonhole with at most as many pigeons as the floor of the average number of pigeons across all pigeonholes. -/ @@ -261,12 +250,8 @@ More formally, given a function `f`, a finite sets `s` in its domain, a finite s codomain, and a natural number `n` such that `card s < card t * n`, there exists `y ∈ t` such that its preimage in `s` has less than `n` elements. -/ lemma exists_card_fiber_lt_of_card_lt_mul (hn : s.card < t.card * n) : - ∃ y ∈ t, (s.filter (λ x, f x = y)).card < n:= -begin - simp only [card_eq_sum_ones], - apply exists_sum_fiber_lt_of_sum_fiber_nonneg_of_sum_lt_nsmul (λ _ _, nat.zero_le _), - simpa -end + ∃ y ∈ t, (s.filter (λ x, f x = y)).card < n := +exists_card_fiber_lt_of_card_lt_nsmul hn /-- The pigeonhole principle for finitely many pigeons counted by heads: given a function between finite sets `s` and `t` and a number `b` such that `card t • b ≤ card s`, there exists `y ∈ t` such @@ -287,11 +272,7 @@ t` such that its preimage in `s` has at least `n` elements. See also lemma exists_le_card_fiber_of_mul_le_card_of_maps_to (hf : ∀ a ∈ s, f a ∈ t) (ht : t.nonempty) (hn : t.card * n ≤ s.card) : ∃ y ∈ t, n ≤ (s.filter (λ x, f x = y)).card := -begin - simp only [card_eq_sum_ones], - apply exists_le_sum_fiber_of_maps_to_of_nsmul_le_sum hf ht, - simpa -end +exists_le_card_fiber_of_nsmul_le_card_of_maps_to hf ht hn /-- The pigeonhole principle for finitely many pigeons counted by heads: given a function `f`, a finite sets `s` and `t`, and a number `b` such that `card s ≤ card t • b`, there exists `y ∈ t` such @@ -310,12 +291,8 @@ finite sets `s` in its domain, a finite set `t` in its codomain, and a natural n `card s ≤ card t * n`, there exists `y ∈ t` such that its preimage in `s` has no more than `n` elements. See also `finset.exists_card_fiber_lt_of_card_lt_mul` for a stronger statement. -/ lemma exists_card_fiber_le_of_card_le_mul (ht : t.nonempty) (hn : s.card ≤ t.card * n) : - ∃ y ∈ t, (s.filter (λ x, f x = y)).card ≤ n:= -begin - simp only [card_eq_sum_ones], - apply exists_sum_fiber_le_of_sum_fiber_nonneg_of_sum_le_nsmul (λ _ _, nat.zero_le _) ht, - simpa -end + ∃ y ∈ t, (s.filter (λ x, f x = y)).card ≤ n := +exists_card_fiber_le_of_card_le_nsmul ht hn end finset @@ -368,7 +345,7 @@ lemma exists_sum_fiber_le_of_sum_le_nsmul [nonempty β] (hb : (∑ x, w x) ≤ c end -variables [linear_ordered_comm_ring M] +variables [linear_ordered_comm_semiring M] /-- The strong pigeonhole principle for finitely many pigeons and pigeonholes. There is a pigeonhole @@ -389,7 +366,7 @@ More formally, given a function `f` between finite types `α` and `β` and a num elements. -/ lemma exists_lt_card_fiber_of_mul_lt_card (hn : card β * n < card α) : ∃ y : β, n < (univ.filter (λ x, f x = y)).card := -let ⟨y, _, h⟩ := exists_lt_card_fiber_of_mul_lt_card_of_maps_to (λ _ _, mem_univ _) hn in ⟨y, h⟩ +exists_lt_card_fiber_of_nsmul_lt_card _ hn /-- The strong pigeonhole principle for finitely many pigeons and pigeonholes. There is a pigeonhole with at most as many pigeons as the floor of the average number of pigeons across all pigeonholes. @@ -409,7 +386,7 @@ More formally, given a function `f` between finite types `α` and `β` and a num elements. -/ lemma exists_card_fiber_lt_of_card_lt_mul (hn : card α < card β * n) : ∃ y : β, (univ.filter (λ x, f x = y)).card < n := -let ⟨y, _, h⟩ := exists_card_fiber_lt_of_card_lt_mul hn in ⟨y, h⟩ +exists_card_fiber_lt_of_card_lt_nsmul _ hn /-- The strong pigeonhole principle for finitely many pigeons and pigeonholes. Given a function `f` between finite types `α` and `β` and a number `b` such that `card β • b ≤ card α`, there exists an @@ -426,8 +403,7 @@ element `y : β` such that its preimage has at least `n` elements. See also `fintype.exists_lt_card_fiber_of_mul_lt_card` for a stronger statement. -/ lemma exists_le_card_fiber_of_mul_le_card [nonempty β] (hn : card β * n ≤ card α) : ∃ y : β, n ≤ (univ.filter (λ x, f x = y)).card := -let ⟨y, _, h⟩ := exists_le_card_fiber_of_mul_le_card_of_maps_to (λ _ _, mem_univ _) univ_nonempty hn -in ⟨y, h⟩ +exists_le_card_fiber_of_nsmul_le_card _ hn /-- The strong pigeonhole principle for finitely many pigeons and pigeonholes. Given a function `f` between finite types `α` and `β` and a number `b` such that `card α ≤ card β • b`, there exists an @@ -443,7 +419,7 @@ element `y : β` such that its preimage has at most `n` elements. See also `fintype.exists_card_fiber_lt_of_card_lt_mul` for a stronger statement. -/ lemma exists_card_fiber_le_of_card_le_mul [nonempty β] (hn : card α ≤ card β * n) : ∃ y : β, (univ.filter (λ x, f x = y)).card ≤ n := -let ⟨y, _, h⟩ := exists_card_fiber_le_of_card_le_mul univ_nonempty hn in ⟨y, h⟩ +exists_card_fiber_le_of_card_le_nsmul _ hn end fintype diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index 69396d6710e40..e1eaeba817a19 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -53,7 +53,7 @@ instance : comm_semiring ℕ := nsmul_succ' := λ n x, by rw [nat.succ_eq_add_one, nat.add_comm, nat.right_distrib, nat.one_mul] } -instance : linear_ordered_semiring nat := +instance : linear_ordered_comm_semiring ℕ := { lt := nat.lt, add_le_add_left := @nat.add_le_add_left, le_of_add_le_add_left := @nat.le_of_add_le_add_left, @@ -64,28 +64,25 @@ instance : linear_ordered_semiring nat := exists_pair_ne := ⟨0, 1, ne_of_lt nat.zero_lt_one⟩, ..nat.comm_semiring, ..nat.linear_order } --- all the fields are already included in the linear_ordered_semiring instance -instance : linear_ordered_cancel_add_comm_monoid ℕ := { ..nat.linear_ordered_semiring } - instance : linear_ordered_comm_monoid_with_zero ℕ := { mul_le_mul_left := λ a b h c, nat.mul_le_mul_left c h, - ..nat.linear_ordered_semiring, + ..nat.linear_ordered_comm_semiring, ..(infer_instance : comm_monoid_with_zero ℕ)} -instance : ordered_comm_semiring ℕ := { .. nat.comm_semiring, .. nat.linear_ordered_semiring } - /-! Extra instances to short-circuit type class resolution -/ -instance : add_comm_monoid nat := by apply_instance -instance : add_monoid nat := by apply_instance -instance : monoid nat := by apply_instance -instance : comm_monoid nat := by apply_instance -instance : comm_semigroup nat := by apply_instance -instance : semigroup nat := by apply_instance -instance : add_comm_semigroup nat := by apply_instance -instance : add_semigroup nat := by apply_instance -instance : distrib nat := by apply_instance -instance : semiring nat := by apply_instance -instance : ordered_semiring nat := by apply_instance +instance : add_comm_monoid ℕ := infer_instance +instance : add_monoid ℕ := infer_instance +instance : monoid ℕ := infer_instance +instance : comm_monoid ℕ := infer_instance +instance : comm_semigroup ℕ := infer_instance +instance : semigroup ℕ := infer_instance +instance : add_comm_semigroup ℕ := infer_instance +instance : add_semigroup ℕ := infer_instance +instance : distrib ℕ := infer_instance +instance : semiring ℕ := infer_instance +instance : ordered_semiring ℕ := infer_instance +instance : ordered_comm_semiring ℕ := infer_instance +instance : linear_ordered_cancel_add_comm_monoid ℕ := infer_instance instance nat.order_bot : order_bot ℕ := { bot := 0, bot_le := nat.zero_le } diff --git a/src/ring_theory/subsemiring/basic.lean b/src/ring_theory/subsemiring/basic.lean index e693a75715df0..41640480f0cf6 100644 --- a/src/ring_theory/subsemiring/basic.lean +++ b/src/ring_theory/subsemiring/basic.lean @@ -133,7 +133,11 @@ instance to_linear_ordered_semiring {R} [linear_ordered_semiring R] [set_like S subtype.coe_injective.linear_ordered_semiring coe rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) -/-! Note: currently, there is no `linear_ordered_comm_semiring`. -/ +/-- A subsemiring of a `linear_ordered_comm_semiring` is a `linear_ordered_comm_semiring`. -/ +instance to_linear_ordered_comm_semiring {R} [linear_ordered_comm_semiring R] [set_like S R] + [subsemiring_class S R] : linear_ordered_comm_semiring s := +subtype.coe_injective.linear_ordered_comm_semiring coe + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) end subsemiring_class @@ -351,7 +355,11 @@ instance to_linear_ordered_semiring {R} [linear_ordered_semiring R] (s : subsemi subtype.coe_injective.linear_ordered_semiring coe rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) -/-! Note: currently, there is no `linear_ordered_comm_semiring`. -/ +/-- A subsemiring of a `linear_ordered_comm_semiring` is a `linear_ordered_comm_semiring`. -/ +instance to_linear_ordered_comm_semiring {R} [linear_ordered_comm_semiring R] (s : subsemiring R) : + linear_ordered_comm_semiring s := +subtype.coe_injective.linear_ordered_comm_semiring coe + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) protected lemma nsmul_mem {x : R} (hx : x ∈ s) (n : ℕ) : n • x ∈ s := nsmul_mem hx n From e3fac0b5a10a00a63f28505d9a5bea62d6869929 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 19 Sep 2022 07:18:13 +0000 Subject: [PATCH 308/828] refactor(topology/uniform_space/separated): drop `is_separated` (#16458) This predicate is no longer used outside of this file. If we'll need it in the future, then we can redefine it for any topological space in terms of `inseparable`. Also rename `topology.uniform_space.compact_separated` to `topology.uniform_space.compact`. --- src/analysis/box_integral/basic.lean | 2 +- src/topology/algebra/uniform_group.lean | 2 +- src/topology/continuous_function/compact.lean | 2 +- .../{compact_separated.lean => compact.lean} | 0 src/topology/uniform_space/separation.lean | 97 +------------------ 5 files changed, 8 insertions(+), 95 deletions(-) rename src/topology/uniform_space/{compact_separated.lean => compact.lean} (100%) diff --git a/src/analysis/box_integral/basic.lean b/src/analysis/box_integral/basic.lean index 2245ab9942ac9..2e0b1f8b82283 100644 --- a/src/analysis/box_integral/basic.lean +++ b/src/analysis/box_integral/basic.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov -/ import analysis.box_integral.partition.filter import analysis.box_integral.partition.measure -import topology.uniform_space.compact_separated +import topology.uniform_space.compact /-! # Integrals of Riemann, Henstock-Kurzweil, and McShane diff --git a/src/topology/algebra/uniform_group.lean b/src/topology/algebra/uniform_group.lean index df29665f1b193..eaba7aa396336 100644 --- a/src/topology/algebra/uniform_group.lean +++ b/src/topology/algebra/uniform_group.lean @@ -6,7 +6,7 @@ Authors: Patrick Massot, Johannes Hölzl import topology.uniform_space.uniform_convergence import topology.uniform_space.uniform_embedding import topology.uniform_space.complete_separated -import topology.uniform_space.compact_separated +import topology.uniform_space.compact import topology.algebra.group import tactic.abel diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index 0ec59004b2c0f..85f7602caab89 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import topology.continuous_function.bounded -import topology.uniform_space.compact_separated +import topology.uniform_space.compact import topology.compact_open import topology.sets.compacts diff --git a/src/topology/uniform_space/compact_separated.lean b/src/topology/uniform_space/compact.lean similarity index 100% rename from src/topology/uniform_space/compact_separated.lean rename to src/topology/uniform_space/compact.lean diff --git a/src/topology/uniform_space/separation.lean b/src/topology/uniform_space/separation.lean index 8ea80cca1bf87..627f8a10718c7 100644 --- a/src/topology/uniform_space/separation.lean +++ b/src/topology/uniform_space/separation.lean @@ -37,7 +37,6 @@ is equivalent to asking that the uniform structure induced on `s` is separated. * `separation_relation X : set (X × X)`: the separation relation * `separated_space X`: a predicate class asserting that `X` is separated -* `is_separated s`: a predicate asserting that `s : set X` is separated * `separation_quotient X`: the maximal separated quotient of `X`. * `separation_quotient.lift f`: factors a map `f : X → Y` through the separation quotient of `X`. * `separation_quotient.map f`: turns a map `f : X → Y` into a map between the separation quotients @@ -136,6 +135,11 @@ lemma eq_of_forall_symmetric {α : Type*} [uniform_space α] [separated_space α (h : ∀ {V}, V ∈ 𝓤 α → symmetric_rel V → (x, y) ∈ V) : x = y := eq_of_uniformity_basis has_basis_symmetric (by simpa [and_imp] using λ _, h) +lemma eq_of_cluster_pt_uniformity [separated_space α] {x y : α} (h : cluster_pt (x, y) (𝓤 α)) : + x = y := +eq_of_uniformity_basis uniformity_has_basis_closed $ λ V ⟨hV, hVc⟩, + is_closed_iff_cluster_pt.1 hVc _ $ h.mono $ le_principal_iff.2 hV + lemma id_rel_sub_separation_relation (α : Type*) [uniform_space α] : id_rel ⊆ 𝓢 α := begin unfold separation_rel, @@ -235,86 +239,6 @@ lemma is_closed_range_of_spaced_out {ι} [separated_space α] {V₀ : set (α × is_closed_of_spaced_out V₀_in $ by { rintro _ ⟨x, rfl⟩ _ ⟨y, rfl⟩ h, exact hf x y (ne_of_apply_ne f h) } -/-! -### Separated sets --/ - -/-- A set `s` in a uniform space `α` is separated if the separation relation `𝓢 α` -induces the trivial relation on `s`. -/ -def is_separated (s : set α) : Prop := ∀ x y ∈ s, (x, y) ∈ 𝓢 α → x = y - -lemma is_separated_def (s : set α) : is_separated s ↔ ∀ x y ∈ s, (x, y) ∈ 𝓢 α → x = y := -iff.rfl - -lemma is_separated_def' (s : set α) : is_separated s ↔ (s ×ˢ s) ∩ 𝓢 α ⊆ id_rel := -begin - rw is_separated_def, - split, - { rintros h ⟨x, y⟩ ⟨⟨x_in, y_in⟩, H⟩, - simp [h x x_in y y_in H] }, - { intros h x x_in y y_in xy_in, - rw ← mem_id_rel, - exact h ⟨mk_mem_prod x_in y_in, xy_in⟩ } -end - -lemma is_separated.mono {s t : set α} (hs : is_separated s) (hts : t ⊆ s) : is_separated t := -λ x hx y hy, hs x (hts hx) y (hts hy) - -lemma univ_separated_iff : is_separated (univ : set α) ↔ separated_space α := -begin - simp only [is_separated, mem_univ, true_implies_iff, separated_space_iff], - split, - { intro h, - exact subset.antisymm (λ ⟨x, y⟩ xy_in, h x y xy_in) (id_rel_sub_separation_relation α), }, - { intros h x y xy_in, - rwa h at xy_in }, -end - -lemma is_separated_of_separated_space [separated_space α] (s : set α) : is_separated s := -begin - rw [is_separated, separated_space.out], - tauto, -end - -lemma is_separated_iff_induced {s : set α} : is_separated s ↔ separated_space s := -begin - rw separated_space_iff, - change _ ↔ 𝓢 {x // x ∈ s} = _, - rw [separation_rel_comap rfl, is_separated_def'], - split; intro h, - { ext ⟨⟨x, x_in⟩, ⟨y, y_in⟩⟩, - suffices : (x, y) ∈ 𝓢 α ↔ x = y, by simpa only [mem_id_rel], - refine ⟨λ H, h ⟨mk_mem_prod x_in y_in, H⟩, _⟩, - rintro rfl, - exact id_rel_sub_separation_relation α rfl }, - { rintros ⟨x, y⟩ ⟨⟨x_in, y_in⟩, hS⟩, - have A : (⟨⟨x, x_in⟩, ⟨y, y_in⟩⟩ : ↥s × ↥s) ∈ prod.map (coe : s → α) (coe : s → α) ⁻¹' 𝓢 α, - from hS, - simpa using h.subset A } -end - -lemma eq_of_uniformity_inf_nhds_of_is_separated {s : set α} (hs : is_separated s) : - ∀ {x y : α}, x ∈ s → y ∈ s → cluster_pt (x, y) (𝓤 α) → x = y := -begin - intros x y x_in y_in H, - have : ∀ V ∈ 𝓤 α, (x, y) ∈ closure V, - { intros V V_in, - rw mem_closure_iff_cluster_pt, - have : 𝓤 α ≤ 𝓟 V, by rwa le_principal_iff, - exact H.mono this }, - apply hs x x_in y y_in, - simpa [separation_rel_eq_inter_closure], -end - -lemma eq_of_uniformity_inf_nhds [separated_space α] : - ∀ {x y : α}, cluster_pt (x, y) (𝓤 α) → x = y := -begin - have : is_separated (univ : set α), - { rw univ_separated_iff, - assumption }, - introv, - simpa using eq_of_uniformity_inf_nhds_of_is_separated this, -end /-! ### Separation quotient @@ -440,11 +364,6 @@ lemma eq_of_separated_of_uniform_continuous [separated_space β] {f : α → β} (H : uniform_continuous f) (h : x ≈ y) : f x = f y := separated_def.1 (by apply_instance) _ _ $ separated_of_uniform_continuous H h -lemma _root_.is_separated.eq_of_uniform_continuous {f : α → β} {x y : α} {s : set β} - (hs : is_separated s) (hxs : f x ∈ s) (hys : f y ∈ s) (H : uniform_continuous f) (h : x ≈ y) : - f x = f y := -(is_separated_def _).mp hs _ hxs _ hys $ λ _ h', h _ (H h') - /-- The maximal separated quotient of a uniform space `α`. -/ def separation_quotient (α : Type*) [uniform_space α] := quotient (separation_setoid α) @@ -519,10 +438,4 @@ separated_def.2 $ assume x y H, prod.ext (eq_of_separated_of_uniform_continuous uniform_continuous_fst H) (eq_of_separated_of_uniform_continuous uniform_continuous_snd H) -lemma _root_.is_separated.prod {s : set α} {t : set β} (hs : is_separated s) (ht : is_separated t) : - is_separated (s ×ˢ t) := -(is_separated_def _).mpr $ λ x hx y hy H, prod.ext - (hs.eq_of_uniform_continuous hx.1 hy.1 uniform_continuous_fst H) - (ht.eq_of_uniform_continuous hx.2 hy.2 uniform_continuous_snd H) - end uniform_space From 7cfeb2fa488fd58b8667ac5ff96e5f163e5e428d Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 19 Sep 2022 07:18:14 +0000 Subject: [PATCH 309/828] feat(data/real/cau_seq): define sup and inf (#16494) Showing that these coincide with the max and min (respectively) of `real` is left to a follow-up PR. Note that these do not form a `distrib_lattice` as `cau_seq` does not form a `partial_order`. --- src/data/real/cau_seq.lean | 153 +++++++++++++++++++++++++++++++++++++ 1 file changed, 153 insertions(+) diff --git a/src/data/real/cau_seq.lean b/src/data/real/cau_seq.lean index 6d5d9ec430435..82fee058380cf 100644 --- a/src/data/real/cau_seq.lean +++ b/src/data/real/cau_seq.lean @@ -603,6 +603,159 @@ theorem exists_lt (f : cau_seq α abs) : ∃ a : α, const a < f := let ⟨a, h⟩ := (-f).exists_gt in ⟨-a, show pos _, by rwa [const_neg, sub_neg_eq_add, add_comm, ← sub_neg_eq_add]⟩ +-- so named to match `rat_add_continuous_lemma` +theorem _root_.rat_sup_continuous_lemma {ε : α} {a₁ a₂ b₁ b₂ : α} : + abs (a₁ - b₁) < ε → abs (a₂ - b₂) < ε → abs (a₁ ⊔ a₂ - (b₁ ⊔ b₂)) < ε := +λ h₁ h₂, (abs_max_sub_max_le_max _ _ _ _).trans_lt (max_lt h₁ h₂) + +-- so named to match `rat_add_continuous_lemma` +theorem _root_.rat_inf_continuous_lemma {ε : α} {a₁ a₂ b₁ b₂ : α} : + abs (a₁ - b₁) < ε → abs (a₂ - b₂) < ε → abs (a₁ ⊓ a₂ - (b₁ ⊓ b₂)) < ε := +λ h₁ h₂, (abs_min_sub_min_le_max _ _ _ _).trans_lt (max_lt h₁ h₂) + +instance : has_sup (cau_seq α abs) := +⟨λ f g, ⟨f ⊔ g, λ ε ε0, + (exists_forall_ge_and (f.cauchy₃ ε0) (g.cauchy₃ ε0)).imp $ λ i H j ij, + let ⟨H₁, H₂⟩ := H _ le_rfl in rat_sup_continuous_lemma (H₁ _ ij) (H₂ _ ij)⟩⟩ + +instance : has_inf (cau_seq α abs) := +⟨λ f g, ⟨f ⊓ g, λ ε ε0, + (exists_forall_ge_and (f.cauchy₃ ε0) (g.cauchy₃ ε0)).imp $ λ i H j ij, + let ⟨H₁, H₂⟩ := H _ le_rfl in rat_inf_continuous_lemma (H₁ _ ij) (H₂ _ ij)⟩⟩ + +@[simp, norm_cast] lemma coe_sup (f g : cau_seq α abs) : ⇑(f ⊔ g) = f ⊔ g := rfl + +@[simp, norm_cast] lemma coe_inf (f g : cau_seq α abs) : ⇑(f ⊓ g) = f ⊓ g := rfl + +theorem sup_lim_zero {f g : cau_seq α abs} + (hf : lim_zero f) (hg : lim_zero g) : lim_zero (f ⊔ g) +| ε ε0 := (exists_forall_ge_and (hf _ ε0) (hg _ ε0)).imp $ + λ i H j ij, let ⟨H₁, H₂⟩ := H _ ij in begin + rw abs_lt at H₁ H₂ ⊢, + exact ⟨lt_sup_iff.mpr (or.inl H₁.1), sup_lt_iff.mpr ⟨H₁.2, H₂.2⟩⟩ + end + +theorem inf_lim_zero {f g : cau_seq α abs} + (hf : lim_zero f) (hg : lim_zero g) : lim_zero (f ⊓ g) +| ε ε0 := (exists_forall_ge_and (hf _ ε0) (hg _ ε0)).imp $ + λ i H j ij, let ⟨H₁, H₂⟩ := H _ ij in begin + rw abs_lt at H₁ H₂ ⊢, + exact ⟨lt_inf_iff.mpr ⟨H₁.1, H₂.1⟩, inf_lt_iff.mpr (or.inl H₁.2), ⟩ + end + +lemma sup_equiv_sup {a₁ b₁ a₂ b₂ : cau_seq α abs} (ha : a₁ ≈ a₂) (hb : b₁ ≈ b₂) : + a₁ ⊔ b₁ ≈ a₂ ⊔ b₂ := +begin + intros ε ε0, + obtain ⟨ai, hai⟩ := ha ε ε0, + obtain ⟨bi, hbi⟩ := hb ε ε0, + exact ⟨ai ⊔ bi, λ i hi, + (abs_max_sub_max_le_max (a₁ i) (b₁ i) (a₂ i) (b₂ i)).trans_lt + (max_lt (hai i (sup_le_iff.mp hi).1) (hbi i (sup_le_iff.mp hi).2))⟩, +end + +lemma inf_equiv_inf {a₁ b₁ a₂ b₂ : cau_seq α abs} (ha : a₁ ≈ a₂) (hb : b₁ ≈ b₂) : + a₁ ⊓ b₁ ≈ a₂ ⊓ b₂ := +begin + intros ε ε0, + obtain ⟨ai, hai⟩ := ha ε ε0, + obtain ⟨bi, hbi⟩ := hb ε ε0, + exact ⟨ai ⊔ bi, λ i hi, + (abs_min_sub_min_le_max (a₁ i) (b₁ i) (a₂ i) (b₂ i)).trans_lt + (max_lt (hai i (sup_le_iff.mp hi).1) (hbi i (sup_le_iff.mp hi).2))⟩, +end + +protected lemma sup_lt {a b c : cau_seq α abs} (ha : a < c) (hb : b < c) : a ⊔ b < c := +begin + obtain ⟨⟨εa, εa0, ia, ha⟩, ⟨εb, εb0, ib, hb⟩⟩ := ⟨ha, hb⟩, + refine ⟨εa ⊓ εb, lt_inf_iff.mpr ⟨εa0, εb0⟩, ia ⊔ ib, λ i hi, _⟩, + have := min_le_min (ha _ (sup_le_iff.mp hi).1) (hb _ (sup_le_iff.mp hi).2), + exact this.trans_eq (min_sub_sub_left _ _ _) +end + +protected lemma lt_inf {a b c : cau_seq α abs} (hb : a < b) (hc : a < c) : a < b ⊓ c := +begin + obtain ⟨⟨εb, εb0, ib, hb⟩, ⟨εc, εc0, ic, hc⟩⟩ := ⟨hb, hc⟩, + refine ⟨εb ⊓ εc, lt_inf_iff.mpr ⟨εb0, εc0⟩, ib ⊔ ic, λ i hi, _⟩, + have := min_le_min (hb _ (sup_le_iff.mp hi).1) (hc _ (sup_le_iff.mp hi).2), + exact this.trans_eq (min_sub_sub_right _ _ _), +end + +@[simp] protected lemma sup_idem (a : cau_seq α abs) : a ⊔ a = a := subtype.ext sup_idem + +@[simp] protected lemma inf_idem (a : cau_seq α abs) : a ⊓ a = a := subtype.ext inf_idem + +protected lemma sup_comm (a b : cau_seq α abs) : a ⊔ b = b ⊔ a := subtype.ext sup_comm + +protected lemma inf_comm (a b : cau_seq α abs) : a ⊓ b = b ⊓ a := subtype.ext inf_comm + +protected lemma sup_eq_right {a b : cau_seq α abs} (h : a ≤ b) : + a ⊔ b ≈ b := +begin + obtain ⟨ε, ε0 : _ < _, i, h⟩ | h := h, + { intros _ _, + refine ⟨i, λ j hj, _⟩, + dsimp, + erw ←max_sub_sub_right, + rwa [sub_self, max_eq_right, abs_zero], + rw [sub_nonpos, ←sub_nonneg], + exact ε0.le.trans (h _ hj) }, + { refine setoid.trans (sup_equiv_sup h (setoid.refl _)) _, + rw cau_seq.sup_idem, + exact setoid.refl _ }, +end + +protected lemma inf_eq_right {a b : cau_seq α abs} (h : b ≤ a) : + a ⊓ b ≈ b := +begin + obtain ⟨ε, ε0 : _ < _, i, h⟩ | h := h, + { intros _ _, + refine ⟨i, λ j hj, _⟩, + dsimp, + erw ←min_sub_sub_right, + rwa [sub_self, min_eq_right, abs_zero], + exact ε0.le.trans (h _ hj) }, + { refine setoid.trans (inf_equiv_inf (setoid.symm h) (setoid.refl _)) _, + rw cau_seq.inf_idem, + exact setoid.refl _ }, +end + +protected lemma sup_eq_left {a b : cau_seq α abs} (h : b ≤ a) : + a ⊔ b ≈ a := +by simpa only [cau_seq.sup_comm] using cau_seq.sup_eq_right h + +protected lemma inf_eq_left {a b : cau_seq α abs} (h : a ≤ b) : + a ⊓ b ≈ a := +by simpa only [cau_seq.inf_comm] using cau_seq.inf_eq_right h + +protected lemma sup_le {a b c : cau_seq α abs} (ha : a ≤ c) (hb : b ≤ c) : a ⊔ b ≤ c := +begin + cases ha with ha ha, + { cases hb with hb hb, + { exact or.inl (cau_seq.sup_lt ha hb) }, + { replace ha := le_of_le_of_eq ha.le (setoid.symm hb), + refine le_of_le_of_eq (or.inr _) hb, + exact cau_seq.sup_eq_right ha }, }, + { replace hb := le_of_le_of_eq hb (setoid.symm ha), + refine le_of_le_of_eq (or.inr _) ha, + exact cau_seq.sup_eq_left hb } +end + +protected lemma le_inf {a b c : cau_seq α abs} (hb : a ≤ b) (hc : a ≤ c) : a ≤ b ⊓ c := +begin + cases hb with hb hb, + { cases hc with hc hc, + { exact or.inl (cau_seq.lt_inf hb hc) }, + { replace hb := le_of_eq_of_le (setoid.symm hc) hb.le, + refine le_of_eq_of_le hc (or.inr _), + exact setoid.symm (cau_seq.inf_eq_right hb) }, }, + { replace hc := le_of_eq_of_le (setoid.symm hb) hc, + refine le_of_eq_of_le hb (or.inr _), + exact setoid.symm (cau_seq.inf_eq_left hc) } +end + +/-! Note that `distrib_lattice (cau_seq α abs)` is not true because there is no `partial_order`. -/ + end abs end cau_seq From 77fa4dcf7de23a501666b6053cd78f3ecc57f0bc Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 19 Sep 2022 07:18:17 +0000 Subject: [PATCH 310/828] feat(data/set/intervals/ord_connected): lemmas about `order_dual` etc (#16534) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * add `dual_interval`; * add `ord_connected_preimage` for an `order_hom_class`, use it to golf `ord_connected_image`; * add `dual_ord_connected_iff` and `dual_ord_connected`; * one implication from `ord_connected_iff_interval_subset_left` doesn't need `(hx : x ∈ s)`. --- src/data/set/intervals/ord_connected.lean | 38 ++++++++++++------- .../set/intervals/unordered_interval.lean | 3 ++ 2 files changed, 28 insertions(+), 13 deletions(-) diff --git a/src/data/set/intervals/ord_connected.lean b/src/data/set/intervals/ord_connected.lean index 358937c1d2b2c..72beaf11b0721 100644 --- a/src/data/set/intervals/ord_connected.lean +++ b/src/data/set/intervals/ord_connected.lean @@ -19,6 +19,7 @@ that all standard intervals are `ord_connected`. -/ open_locale interval +open order_dual (to_dual of_dual) namespace set section preorder @@ -144,19 +145,29 @@ instance [densely_ordered α] {s : set α} [hs : ord_connected s] : ⟨λ a b (h : (a : α) < b), let ⟨x, H⟩ := exists_between h in ⟨⟨x, (hs.out a.2 b.2) (Ioo_subset_Icc_self H)⟩, H⟩ ⟩ +@[instance] lemma ord_connected_preimage {F : Type*} [order_hom_class F α β] (f : F) {s : set β} + [hs : ord_connected s] : ord_connected (f ⁻¹' s) := +⟨λ x hx y hy z hz, hs.out hx hy ⟨order_hom_class.mono _ hz.1, order_hom_class.mono _ hz.2⟩⟩ + @[instance] lemma ord_connected_image {E : Type*} [order_iso_class E α β] (e : E) {s : set α} [hs : ord_connected s] : ord_connected (e '' s) := -begin - constructor, - rintro _ ⟨x, hx, rfl⟩ _ ⟨y, hy, rfl⟩ z ⟨hxz, hzy⟩, - exact ⟨equiv_like.inv e z, hs.out hx hy ⟨(le_map_inv_iff e).mpr hxz, (map_inv_le_iff e).mpr hzy⟩, - equiv_like.right_inv e z⟩ -end +by { erw [(e : α ≃o β).image_eq_preimage], apply ord_connected_preimage } @[instance] lemma ord_connected_range {E : Type*} [order_iso_class E α β] (e : E) : ord_connected (range e) := by simp_rw [← image_univ, ord_connected_image e] +@[simp] lemma dual_ord_connected_iff {s : set α} : + ord_connected (of_dual ⁻¹' s) ↔ ord_connected s := +begin + simp_rw [ord_connected_def, to_dual.surjective.forall, dual_Icc, subtype.forall'], + exact forall_swap +end + +@[instance] lemma dual_ord_connected {s : set α} [ord_connected s] : + ord_connected (of_dual ⁻¹' s) := +dual_ord_connected_iff.2 ‹_› + end preorder section linear_order @@ -177,15 +188,16 @@ lemma ord_connected_iff_interval_subset : ord_connected s ↔ ∀ ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s), [x, y] ⊆ s := ⟨λ h, h.interval_subset, λ H, ⟨λ x hx y hy, Icc_subset_interval.trans $ H hx hy⟩⟩ +lemma ord_connected_of_interval_subset_left (h : ∀ y ∈ s, [x, y] ⊆ s) : + ord_connected s := +ord_connected_iff_interval_subset.2 $ λ y hy z hz, +calc [y, z] ⊆ [y, x] ∪ [x, z] : interval_subset_interval_union_interval +... = [x, y] ∪ [x, z] : by rw [interval_swap] +... ⊆ s : union_subset (h y hy) (h z hz) + lemma ord_connected_iff_interval_subset_left (hx : x ∈ s) : ord_connected s ↔ ∀ ⦃y⦄, y ∈ s → [x, y] ⊆ s := -begin - refine ⟨λ hs, hs.interval_subset hx, λ hs, ord_connected_iff_interval_subset.2 $ λ y hy z hz, _⟩, - suffices h : [y, x] ∪ [x, z] ⊆ s, - { exact interval_subset_interval_union_interval.trans h }, - rw [interval_swap, union_subset_iff], - exact ⟨hs hy, hs hz⟩, -end +⟨λ hs, hs.interval_subset hx, ord_connected_of_interval_subset_left⟩ lemma ord_connected_iff_interval_subset_right (hx : x ∈ s) : ord_connected s ↔ ∀ ⦃y⦄, y ∈ s → [y, x] ⊆ s := diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index eac3814af5b4a..9461ef2701ef5 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -27,6 +27,7 @@ make the notation available. universe u open_locale pointwise +open order_dual (to_dual of_dual) namespace set @@ -39,6 +40,8 @@ def interval (a b : α) := Icc (min a b) (max a b) localized "notation (name := set.interval) `[`a `, ` b `]` := set.interval a b" in interval +@[simp] lemma dual_interval (a b : α) : [to_dual a, to_dual b] = of_dual ⁻¹' [a, b] := dual_Icc + @[simp] lemma interval_of_le (h : a ≤ b) : [a, b] = Icc a b := by rw [interval, min_eq_left h, max_eq_right h] From 499ab57072a09de5e4c11528957a3125007a47e2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 19 Sep 2022 07:18:19 +0000 Subject: [PATCH 311/828] refactor(topology/constructions): golf, add `@[simp]`/`iff` variants (#16539) --- src/topology/constructions.lean | 118 ++++++++++++++------------------ src/topology/homeomorph.lean | 3 +- 2 files changed, 53 insertions(+), 68 deletions(-) diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index 1a069cdb859cb..757cb51a8ff06 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -699,67 +699,29 @@ continuous_sup_rng_left continuous_coinduced_rng @[continuity] lemma continuous_inr : continuous (@inr α β) := continuous_sup_rng_right continuous_coinduced_rng -@[continuity] lemma continuous.sum_elim {f : α → γ} {g : β → γ} - (hf : continuous f) (hg : continuous g) : continuous (sum.elim f g) := -by simp only [continuous_sup_dom, continuous_coinduced_dom, sum.elim_comp_inl, sum.elim_comp_inr, - true_and, *] - -@[continuity] lemma continuous.sum_map {f : α → β} {g : γ → δ} - (hf : continuous f) (hg : continuous g) : continuous (sum.map f g) := -(continuous_inl.comp hf).sum_elim (continuous_inr.comp hg) - lemma is_open_sum_iff {s : set (α ⊕ β)} : is_open s ↔ is_open (inl ⁻¹' s) ∧ is_open (inr ⁻¹' s) := iff.rfl -lemma is_open_map_sum {f : α ⊕ β → γ} - (h₁ : is_open_map (λ a, f (inl a))) (h₂ : is_open_map (λ b, f (inr b))) : - is_open_map f := -begin - intros u hu, - rw is_open_sum_iff at hu, - cases hu with hu₁ hu₂, - have : u = inl '' (inl ⁻¹' u) ∪ inr '' (inr ⁻¹' u), - { ext (_|_); simp }, - rw [this, set.image_union, set.image_image, set.image_image], - exact is_open.union (h₁ _ hu₁) (h₂ _ hu₂) -end +lemma is_open_map_inl : is_open_map (@inl α β) := +λ u hu, by simpa [is_open_sum_iff, preimage_image_eq u sum.inl_injective] -lemma embedding_inl : embedding (@inl α β) := -{ induced := begin - unfold sum.topological_space, - apply le_antisymm, - { rw ← coinduced_le_iff_le_induced, exact le_sup_left }, - { intros u hu, existsi (inl '' u), - change - (is_open (inl ⁻¹' (@inl α β '' u)) ∧ - is_open (inr ⁻¹' (@inl α β '' u))) ∧ - inl ⁻¹' (inl '' u) = u, - rw [preimage_image_eq u sum.inl_injective, preimage_inr_image_inl], - exact ⟨⟨hu, is_open_empty⟩, rfl⟩ } - end, - inj := λ _ _, inl.inj_iff.mp } - -lemma embedding_inr : embedding (@inr α β) := -{ induced := begin - unfold sum.topological_space, - apply le_antisymm, - { rw ← coinduced_le_iff_le_induced, exact le_sup_right }, - { intros u hu, existsi (inr '' u), - change - (is_open (inl ⁻¹' (@inr α β '' u)) ∧ - is_open (inr ⁻¹' (@inr α β '' u))) ∧ - inr ⁻¹' (inr '' u) = u, - rw [preimage_inl_image_inr, preimage_image_eq u sum.inr_injective], - exact ⟨⟨is_open_empty, hu⟩, rfl⟩ } - end, - inj := λ _ _, inr.inj_iff.mp } +lemma is_open_map_inr : is_open_map (@inr α β) := +λ u hu, by simpa [is_open_sum_iff, preimage_image_eq u sum.inr_injective] + +lemma open_embedding_inl : open_embedding (@inl α β) := +open_embedding_of_continuous_injective_open continuous_inl inl_injective is_open_map_inl + +lemma open_embedding_inr : open_embedding (@inr α β) := +open_embedding_of_continuous_injective_open continuous_inr inr_injective is_open_map_inr -lemma is_open_range_inl : is_open (range (inl : α → α ⊕ β)) := -is_open_sum_iff.2 $ by simp +lemma embedding_inl : embedding (@inl α β) := open_embedding_inl.1 -lemma is_open_range_inr : is_open (range (inr : β → α ⊕ β)) := -is_open_sum_iff.2 $ by simp +lemma embedding_inr : embedding (@inr α β) := open_embedding_inr.1 + +lemma is_open_range_inl : is_open (range (inl : α → α ⊕ β)) := open_embedding_inl.2 + +lemma is_open_range_inr : is_open (range (inr : β → α ⊕ β)) := open_embedding_inr.2 lemma is_closed_range_inl : is_closed (range (inl : α → α ⊕ β)) := by { rw [← is_open_compl_iff, compl_range_inl], exact is_open_range_inr } @@ -767,21 +729,45 @@ by { rw [← is_open_compl_iff, compl_range_inl], exact is_open_range_inr } lemma is_closed_range_inr : is_closed (range (inr : β → α ⊕ β)) := by { rw [← is_open_compl_iff, compl_range_inr], exact is_open_range_inl } -lemma open_embedding_inl : open_embedding (inl : α → α ⊕ β) := -{ open_range := is_open_range_inl, - .. embedding_inl } - -lemma open_embedding_inr : open_embedding (inr : β → α ⊕ β) := -{ open_range := is_open_range_inr, - .. embedding_inr } - lemma closed_embedding_inl : closed_embedding (inl : α → α ⊕ β) := -{ closed_range := is_closed_range_inl, - .. embedding_inl } +⟨embedding_inl, is_closed_range_inl⟩ lemma closed_embedding_inr : closed_embedding (inr : β → α ⊕ β) := -{ closed_range := is_closed_range_inr, - .. embedding_inr } +⟨embedding_inr, is_closed_range_inr⟩ + +lemma nhds_inl (x : α) : 𝓝 (inl x : α ⊕ β) = map inl (𝓝 x) := +(open_embedding_inl.map_nhds_eq _).symm + +lemma nhds_inr (x : β) : 𝓝 (inr x : α ⊕ β) = map inr (𝓝 x) := +(open_embedding_inr.map_nhds_eq _).symm + +lemma continuous_sum_elim {f : α → γ} {g : β → γ} : + continuous (sum.elim f g) ↔ continuous f ∧ continuous g := +by simp only [continuous_sup_dom, continuous_coinduced_dom, sum.elim_comp_inl, sum.elim_comp_inr] + +@[continuity] lemma continuous.sum_elim {f : α → γ} {g : β → γ} + (hf : continuous f) (hg : continuous g) : continuous (sum.elim f g) := +continuous_sum_elim.2 ⟨hf, hg⟩ + +@[simp] lemma continuous_sum_map {f : α → β} {g : γ → δ} : + continuous (sum.map f g) ↔ continuous f ∧ continuous g := +continuous_sum_elim.trans $ embedding_inl.continuous_iff.symm.and embedding_inr.continuous_iff.symm + +@[continuity] lemma continuous.sum_map {f : α → β} {g : γ → δ} (hf : continuous f) + (hg : continuous g) : continuous (sum.map f g) := +continuous_sum_map.2 ⟨hf, hg⟩ + +lemma is_open_map_sum {f : α ⊕ β → γ} : + is_open_map f ↔ is_open_map (λ a, f (inl a)) ∧ is_open_map (λ b, f (inr b)) := +by simp only [is_open_map_iff_nhds_le, sum.forall, nhds_inl, nhds_inr, filter.map_map] + +@[simp] lemma is_open_map_sum_elim {f : α → γ} {g : β → γ} : + is_open_map (sum.elim f g) ↔ is_open_map f ∧ is_open_map g := +by simp only [is_open_map_sum, elim_inl, elim_inr] + +lemma is_open_map.sum_elim {f : α → γ} {g : β → γ} (hf : is_open_map f) (hg : is_open_map g) : + is_open_map (sum.elim f g) := +is_open_map_sum_elim.2 ⟨hf, hg⟩ end sum diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 7e77c453e5ec2..1bf50e868bb41 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -411,8 +411,7 @@ section distrib def sum_prod_distrib : (α ⊕ β) × γ ≃ₜ α × γ ⊕ β × γ := homeomorph.symm $ homeomorph_of_continuous_open (equiv.sum_prod_distrib α β γ).symm ((continuous_inl.prod_map continuous_id).sum_elim (continuous_inr.prod_map continuous_id)) $ - is_open_map_sum (open_embedding_inl.is_open_map.prod is_open_map.id) - (open_embedding_inr.is_open_map.prod is_open_map.id) + (is_open_map_inl.prod is_open_map.id).sum_elim (is_open_map_inr.prod is_open_map.id) /-- `α × (β ⊕ γ)` is homeomorphic to `α × β ⊕ α × γ`. -/ def prod_sum_distrib : α × (β ⊕ γ) ≃ₜ α × β ⊕ α × γ := From dcf3c6b632a857a0d8a24e8857b28fe0cc29f769 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 19 Sep 2022 09:57:53 +0000 Subject: [PATCH 312/828] feat(group_theory/commutator): Add `commutator_element_self` (#16437) This PR adds `commutator_element_self`, similar to `commutator_element_one_left` and `commutator_element_one_right`. --- src/group_theory/commutator.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/group_theory/commutator.lean b/src/group_theory/commutator.lean index c8f5200998a4a..8ad98a85d3bc5 100644 --- a/src/group_theory/commutator.lean +++ b/src/group_theory/commutator.lean @@ -40,6 +40,9 @@ variables (g₁ g₂ g₃ g) @[simp] lemma commutator_element_one_left : ⁅(1 : G), g⁆ = 1 := (commute.one_left g).commutator_eq +@[simp] lemma commutator_element_self : ⁅g, g⁆ = 1 := +(commute.refl g).commutator_eq + @[simp] lemma commutator_element_inv : ⁅g₁, g₂⁆⁻¹ = ⁅g₂, g₁⁆ := by simp_rw [commutator_element_def, mul_inv_rev, inv_inv, mul_assoc] From 3f498b03157744b41bf1d9526470b2a6267bd1cd Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 19 Sep 2022 11:49:37 +0000 Subject: [PATCH 313/828] chore(analysis,topology): replace `noncomputable theory` with `noncomputable` (#16552) The purpose of this change is to make it easier to parse the diff of a future PR that makes a significant fraction of these definitions computable. --- src/analysis/normed/field/basic.lean | 30 +++++++++++++++------------- src/topology/instances/rat.lean | 3 +-- 2 files changed, 17 insertions(+), 16 deletions(-) diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 68fea94d1ad80..56be2fce5eb61 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -17,7 +17,6 @@ definitions. variables {α : Type*} {β : Type*} {γ : Type*} {ι : Type*} -noncomputable theory open filter metric open_locale topological_space big_operators nnreal ennreal uniformity pointwise @@ -185,7 +184,7 @@ instance : non_unital_semi_normed_ring (ulift α) := /-- Non-unital seminormed ring structure on the product of two non-unital seminormed rings, using the sup norm. -/ -instance prod.non_unital_semi_normed_ring [non_unital_semi_normed_ring β] : +noncomputable instance prod.non_unital_semi_normed_ring [non_unital_semi_normed_ring β] : non_unital_semi_normed_ring (α × β) := { norm_mul := assume x y, calc @@ -202,7 +201,7 @@ instance prod.non_unital_semi_normed_ring [non_unital_semi_normed_ring β] : /-- Non-unital seminormed ring structure on the product of finitely many non-unital seminormed rings, using the sup norm. -/ -instance pi.non_unital_semi_normed_ring {π : ι → Type*} [fintype ι] +noncomputable instance pi.non_unital_semi_normed_ring {π : ι → Type*} [fintype ι] [Π i, non_unital_semi_normed_ring (π i)] : non_unital_semi_normed_ring (Π i, π i) := { norm_mul := λ x y, nnreal.coe_mono $ @@ -315,14 +314,15 @@ instance : semi_normed_ring (ulift α) := /-- Seminormed ring structure on the product of two seminormed rings, using the sup norm. -/ -instance prod.semi_normed_ring [semi_normed_ring β] : +noncomputable instance prod.semi_normed_ring [semi_normed_ring β] : semi_normed_ring (α × β) := { ..prod.non_unital_semi_normed_ring, ..prod.seminormed_add_comm_group, } /-- Seminormed ring structure on the product of finitely many seminormed rings, using the sup norm. -/ -instance pi.semi_normed_ring {π : ι → Type*} [fintype ι] [Π i, semi_normed_ring (π i)] : +noncomputable instance pi.semi_normed_ring {π : ι → Type*} [fintype ι] + [Π i, semi_normed_ring (π i)] : semi_normed_ring (Π i, π i) := { ..pi.non_unital_semi_normed_ring, ..pi.seminormed_add_comm_group, } @@ -338,13 +338,15 @@ instance : non_unital_normed_ring (ulift α) := /-- Non-unital normed ring structure on the product of two non-unital normed rings, using the sup norm. -/ -instance prod.non_unital_normed_ring [non_unital_normed_ring β] : non_unital_normed_ring (α × β) := +noncomputable instance prod.non_unital_normed_ring [non_unital_normed_ring β] : + non_unital_normed_ring (α × β) := { norm_mul := norm_mul_le, ..prod.seminormed_add_comm_group } /-- Normed ring structure on the product of finitely many non-unital normed rings, using the sup norm. -/ -instance pi.non_unital_normed_ring {π : ι → Type*} [fintype ι] [Π i, non_unital_normed_ring (π i)] : +noncomputable instance pi.non_unital_normed_ring {π : ι → Type*} [fintype ι] + [Π i, non_unital_normed_ring (π i)] : non_unital_normed_ring (Π i, π i) := { norm_mul := norm_mul_le, ..pi.normed_add_comm_group } @@ -366,12 +368,12 @@ instance : normed_ring (ulift α) := .. ulift.normed_add_comm_group } /-- Normed ring structure on the product of two normed rings, using the sup norm. -/ -instance prod.normed_ring [normed_ring β] : normed_ring (α × β) := +noncomputable instance prod.normed_ring [normed_ring β] : normed_ring (α × β) := { norm_mul := norm_mul_le, ..prod.normed_add_comm_group } /-- Normed ring structure on the product of finitely many normed rings, using the sup norm. -/ -instance pi.normed_ring {π : ι → Type*} [fintype ι] [Π i, normed_ring (π i)] : +noncomputable instance pi.normed_ring {π : ι → Type*} [fintype ι] [Π i, normed_ring (π i)] : normed_ring (Π i, π i) := { norm_mul := norm_mul_le, ..pi.normed_add_comm_group } @@ -616,11 +618,11 @@ end densely end normed_field -instance : normed_field ℝ := +noncomputable instance : normed_field ℝ := { norm_mul' := abs_mul, .. real.normed_add_comm_group } -instance : densely_normed_field ℝ := +noncomputable instance : densely_normed_field ℝ := { lt_norm_lt := λ _ _ h₀ hr, let ⟨x, h⟩ := exists_between hr in ⟨x, by rwa [real.norm_eq_abs, abs_of_nonneg (h₀.trans h.1.le)]⟩ } @@ -698,7 +700,7 @@ lemma normed_add_comm_group.tendsto_at_top' [nonempty α] [semilattice_sup α] [ tendsto f at_top (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N < n → ∥f n - b∥ < ε := (at_top_basis_Ioi.tendsto_iff metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) -instance : normed_comm_ring ℤ := +noncomputable instance : normed_comm_ring ℤ := { norm := λ n, ∥(n : ℝ)∥, norm_mul := λ m n, le_of_eq $ by simp only [norm, int.cast_mul, abs_mul], dist_eq := λ m n, by simp only [int.dist_eq, norm, int.cast_sub], @@ -724,12 +726,12 @@ end instance : norm_one_class ℤ := ⟨by simp [← int.norm_cast_real]⟩ -instance : normed_field ℚ := +noncomputable instance : normed_field ℚ := { norm := λ r, ∥(r : ℝ)∥, norm_mul' := λ r₁ r₂, by simp only [norm, rat.cast_mul, abs_mul], dist_eq := λ r₁ r₂, by simp only [rat.dist_eq, norm, rat.cast_sub] } -instance : densely_normed_field ℚ := +noncomputable instance : densely_normed_field ℚ := { lt_norm_lt := λ r₁ r₂ h₀ hr, let ⟨q, h⟩ := exists_rat_btwn hr in ⟨q, by { unfold norm, rwa abs_of_pos (h₀.trans_lt h.1) } ⟩ } diff --git a/src/topology/instances/rat.lean b/src/topology/instances/rat.lean index 21933125a4015..766fc107b595d 100644 --- a/src/topology/instances/rat.lean +++ b/src/topology/instances/rat.lean @@ -13,12 +13,11 @@ import topology.instances.real The structure of a metric space on `ℚ` is introduced in this file, induced from `ℝ`. -/ -noncomputable theory open metric set filter namespace rat -instance : metric_space ℚ := +noncomputable instance : metric_space ℚ := metric_space.induced coe rat.cast_injective real.metric_space theorem dist_eq (x y : ℚ) : dist x y = |x - y| := rfl From d742555c730e5ddf75b9edb4362769406835dc8a Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 19 Sep 2022 13:33:53 +0000 Subject: [PATCH 314/828] feat(topology/instances/ennreal): diameter of intervals (#16556) --- src/topology/instances/ennreal.lean | 14 +++++++++++++- 1 file changed, 13 insertions(+), 1 deletion(-) diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index f4ea915ed9763..89a7b6f46ff5f 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -1361,7 +1361,7 @@ is_closed_le (continuous_id.edist continuous_const) continuous_const begin refine le_antisymm (diam_le $ λ x hx y hy, _) (diam_mono subset_closure), have : edist x y ∈ closure (Iic (diam s)), - from map_mem_closure₂ continuous_edist hx hy (λ x hx y hy, edist_le_diam_of_mem hx hy), + from map_mem_closure₂ continuous_edist hx hy (λ x hx y hy, edist_le_diam_of_mem hx hy), rwa closure_Iic at this end @@ -1436,6 +1436,18 @@ le_antisymm (ediam_Icc a b ▸ diam_mono Ico_subset_Icc_self) le_antisymm (ediam_Icc a b ▸ diam_mono Ioc_subset_Icc_self) (ediam_Ioo a b ▸ diam_mono Ioo_subset_Ioc_self) +lemma diam_Icc {a b : ℝ} (h : a ≤ b) : metric.diam (Icc a b) = b - a := +by simp [metric.diam, ennreal.to_real_of_real, sub_nonneg.2 h] + +lemma diam_Ico {a b : ℝ} (h : a ≤ b) : metric.diam (Ico a b) = b - a := +by simp [metric.diam, ennreal.to_real_of_real, sub_nonneg.2 h] + +lemma diam_Ioc {a b : ℝ} (h : a ≤ b) : metric.diam (Ioc a b) = b - a := +by simp [metric.diam, ennreal.to_real_of_real, sub_nonneg.2 h] + +lemma diam_Ioo {a b : ℝ} (h : a ≤ b) : metric.diam (Ioo a b) = b - a := +by simp [metric.diam, ennreal.to_real_of_real, sub_nonneg.2 h] + end real /-- If `edist (f n) (f (n+1))` is bounded above by a function `d : ℕ → ℝ≥0∞`, From d7e97106dce717a99d6bd0a46f0a59ba50e80caf Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 19 Sep 2022 15:19:05 +0000 Subject: [PATCH 315/828] feat(ring_theory/polynomial/symmetric): restore `finset.powerset_len` version of `prod_X_sub_C_coeff` (#16424) In #15008, [`prod_X_add_C_coeff`](https://github.com/leanprover-community/mathlib/pull/15008/files#diff-08ead07a5e3b20f4db52e932b309a2ff767e486e4a0bf7d90b5520d25d95dc57L79-L81) was changed to use `multiset.esymm` in its RHS, which is defined in terms of `multiset.powerset_len` and not defeq to the original version which involves `finset.powerset_len` instead. This PR restores the `finset` version by introducing `finset.esymm_map_val`, a generalized version of `mv_polynomial.esymm_eq_multiset_esymm` (this lemma is renamed from `mv_polynomial.esymm_eq_multiset.esymm` to better conform with mathlib convention). Co-authored-by: negiizhao [Zulip](https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/Help.20with.20mv_polynomial/near/297702031) Co-authored-by: Junyan Xu --- src/ring_theory/polynomial/symmetric.lean | 18 ++++++++++-------- src/ring_theory/polynomial/vieta.lean | 16 ++++++++++++---- 2 files changed, 22 insertions(+), 12 deletions(-) diff --git a/src/ring_theory/polynomial/symmetric.lean b/src/ring_theory/polynomial/symmetric.lean index 73dfdb5f2d340..02860e4fe58d6 100644 --- a/src/ring_theory/polynomial/symmetric.lean +++ b/src/ring_theory/polynomial/symmetric.lean @@ -51,6 +51,10 @@ variables {R : Type*} [comm_semiring R] /-- The `n`th elementary symmetric function evaluated at the elements of `s` -/ def esymm (s : multiset R) (n : ℕ) : R := ((s.powerset_len n).map multiset.prod).sum +lemma _root_.finset.esymm_map_val {σ} (f : σ → R) (s : finset σ) (n : ℕ) : + (s.val.map f).esymm n = (s.powerset_len n).sum (λ t, t.prod f) := +by simpa only [esymm, powerset_len_map, ← finset.map_val_val_powerset_len, map_map] + end multiset namespace mv_polynomial @@ -132,14 +136,12 @@ def esymm (n : ℕ) : mv_polynomial σ R := /-- The `n`th elementary symmetric `mv_polynomial σ R` is obtained by evaluating the `n`th elementary symmetric at the `multiset` of the monomials -/ -lemma esymm_eq_multiset.esymm (n : ℕ) : - esymm σ R n = - multiset.esymm (multiset.map (λ i : σ, (X i : mv_polynomial σ R)) finset.univ.val) n := -begin - rw [esymm, multiset.esymm, finset.sum_eq_multiset_sum], - conv_lhs { congr, congr, funext, rw finset.prod_eq_multiset_prod }, - rw [multiset.powerset_len_map, ←map_val_val_powerset_len, multiset.map_map, multiset.map_map], -end +lemma esymm_eq_multiset_esymm : esymm σ R = (finset.univ.val.map X).esymm := +funext $ λ n, (finset.univ.esymm_map_val X n).symm + +lemma aeval_esymm_eq_multiset_esymm [algebra R S] (f : σ → S) (n : ℕ) : + aeval f (esymm σ R n) = (finset.univ.val.map f).esymm n := +by simp_rw [esymm, aeval_sum, aeval_prod, aeval_X, esymm_map_val] /-- We can define `esymm σ R n` by summing over a subtype instead of over `powerset_len`. -/ lemma esymm_eq_sum_subtype (n : ℕ) : esymm σ R n = diff --git a/src/ring_theory/polynomial/vieta.lean b/src/ring_theory/polynomial/vieta.lean index b35a1aa17d0de..9c4512d5551ed 100644 --- a/src/ring_theory/polynomial/vieta.lean +++ b/src/ring_theory/polynomial/vieta.lean @@ -68,6 +68,14 @@ begin exact nat.sub_lt_succ s.card k } end +lemma prod_X_add_C_coeff' {σ} (s : multiset σ) (r : σ → R) {k : ℕ} (h : k ≤ s.card) : + (s.map (λ i, X + C (r i))).prod.coeff k = (s.map r).esymm (s.card - k) := +by rw [← map_map (λ r, X + C r) r, prod_X_add_C_coeff]; rwa s.card_map r + +lemma _root_.finset.prod_X_add_C_coeff {σ} (s : finset σ) (r : σ → R) {k : ℕ} (h : k ≤ s.card) : + (∏ i in s, (X + C (r i))).coeff k = ∑ t in s.powerset_len (s.card - k), ∏ i in t, r i := +by { rw [finset.prod, prod_X_add_C_coeff' _ r h, finset.esymm_map_val], refl } + end semiring section ring @@ -142,10 +150,10 @@ lemma mv_polynomial.prod_C_add_X_eq_sum_esymm : begin let s := finset.univ.val.map (λ i : σ, mv_polynomial.X i), rw (_ : card σ = s.card), - { simp_rw [mv_polynomial.esymm_eq_multiset.esymm σ R _, finset.prod_eq_multiset_prod], + { simp_rw [mv_polynomial.esymm_eq_multiset_esymm σ R, finset.prod_eq_multiset_prod], convert multiset.prod_X_add_C_eq_sum_esymm s, rwa multiset.map_map, }, - { rw multiset.card_map, exact rfl, } + { rw multiset.card_map, refl, } end lemma mv_polynomial.prod_X_add_C_coeff (k : ℕ) (h : k ≤ card σ) : @@ -153,10 +161,10 @@ lemma mv_polynomial.prod_X_add_C_coeff (k : ℕ) (h : k ≤ card σ) : begin let s := finset.univ.val.map (λ i, (mv_polynomial.X i : mv_polynomial σ R)), rw (_ : card σ = s.card) at ⊢ h, - { rw [mv_polynomial.esymm_eq_multiset.esymm σ R (s.card - k), finset.prod_eq_multiset_prod], + { rw [mv_polynomial.esymm_eq_multiset_esymm σ R, finset.prod_eq_multiset_prod], convert multiset.prod_X_add_C_coeff s h, rwa multiset.map_map }, - repeat { rw multiset.card_map, exact rfl, }, + repeat { rw multiset.card_map, refl, }, end end mv_polynomial From 1ae5da9d90a97e93425baa5e18cec4128fd90797 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 19 Sep 2022 17:37:19 +0000 Subject: [PATCH 316/828] first attempt (#16557) --- .github/workflows/maintainer_merge.yml | 56 ++++++++++++++++++++++++++ 1 file changed, 56 insertions(+) create mode 100644 .github/workflows/maintainer_merge.yml diff --git a/.github/workflows/maintainer_merge.yml b/.github/workflows/maintainer_merge.yml new file mode 100644 index 0000000000000..6cf323687d1c1 --- /dev/null +++ b/.github/workflows/maintainer_merge.yml @@ -0,0 +1,56 @@ +name: Maintainer merge + +on: + issue_comment: + types: [created, edited] + pull_request_review: + types: [submitted] + +jobs: + ping_zulip: + name: Ping maintainers on Zulip + if: (github.event.issue.pull_request != 'null') && + (startsWith(github.event.comment.body, 'maintainer merge') || contains(github.event.comment.body, '\r\nmaintainer merge') || + (startsWith(github.event.review.body, 'maintainer merge') || contains(github.event.review.body, '\r\nmaintainer merge')) + runs-on: ubuntu-latest + steps: + - uses: octokit/request-action@v2.x + name: Get PR head + id: get_pr_head + with: + route: GET /repos/:repository/pulls/:pull_number + repository: ${{ github.repository }} + pull_number: ${{ github.event.issue.number }} + env: + GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} + + # Parse steps.get_pr_head.outputs.data, since it is a string + - id: parse_pr_head + name: Parse PR head + uses: gr2m/get-json-paths-action@v1.x + with: + json: ${{ steps.get_pr_head.outputs.data }} + head_user: 'head.user.login' + + - if: (steps.parse_pr_head.outputs.head_user == 'jcommelin') + name: Send message on Zulip + uses: zulip/github-actions-zulip/send-message@v1 + with: + api-key: ${{ secrets.ZULIP_API_KEY }} + email: 'github-bot@leanprover.zulipchat.com' + organization-url: 'https://leanprover.zulipchat.com' + to: 'mathlib maintainers' + type: 'stream' + topic: 'maintainer merge' + content: | + ${{ format('{0} requested a maintainer merge on PR #{1}:', github.event.comment.user.login, github.event.issue.number) }} + + > ${{ github.event.issue.title }} + + - if: (steps.parse_pr_head.outputs.head_user == 'jcommelin') + name: Add reaction + uses: GrantBirki/comment@v2.0.1 + with: + comment-id: ${{ github.event.comment.id }} + reactions: rocket + From 5ea93be09766d28ccd07a764e40fedf30052f5c6 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 19 Sep 2022 19:43:44 +0200 Subject: [PATCH 317/828] Revert "first attempt (#16557)" This reverts commit 1ae5da9d90a97e93425baa5e18cec4128fd90797. --- .github/workflows/maintainer_merge.yml | 56 -------------------------- 1 file changed, 56 deletions(-) delete mode 100644 .github/workflows/maintainer_merge.yml diff --git a/.github/workflows/maintainer_merge.yml b/.github/workflows/maintainer_merge.yml deleted file mode 100644 index 6cf323687d1c1..0000000000000 --- a/.github/workflows/maintainer_merge.yml +++ /dev/null @@ -1,56 +0,0 @@ -name: Maintainer merge - -on: - issue_comment: - types: [created, edited] - pull_request_review: - types: [submitted] - -jobs: - ping_zulip: - name: Ping maintainers on Zulip - if: (github.event.issue.pull_request != 'null') && - (startsWith(github.event.comment.body, 'maintainer merge') || contains(github.event.comment.body, '\r\nmaintainer merge') || - (startsWith(github.event.review.body, 'maintainer merge') || contains(github.event.review.body, '\r\nmaintainer merge')) - runs-on: ubuntu-latest - steps: - - uses: octokit/request-action@v2.x - name: Get PR head - id: get_pr_head - with: - route: GET /repos/:repository/pulls/:pull_number - repository: ${{ github.repository }} - pull_number: ${{ github.event.issue.number }} - env: - GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} - - # Parse steps.get_pr_head.outputs.data, since it is a string - - id: parse_pr_head - name: Parse PR head - uses: gr2m/get-json-paths-action@v1.x - with: - json: ${{ steps.get_pr_head.outputs.data }} - head_user: 'head.user.login' - - - if: (steps.parse_pr_head.outputs.head_user == 'jcommelin') - name: Send message on Zulip - uses: zulip/github-actions-zulip/send-message@v1 - with: - api-key: ${{ secrets.ZULIP_API_KEY }} - email: 'github-bot@leanprover.zulipchat.com' - organization-url: 'https://leanprover.zulipchat.com' - to: 'mathlib maintainers' - type: 'stream' - topic: 'maintainer merge' - content: | - ${{ format('{0} requested a maintainer merge on PR #{1}:', github.event.comment.user.login, github.event.issue.number) }} - - > ${{ github.event.issue.title }} - - - if: (steps.parse_pr_head.outputs.head_user == 'jcommelin') - name: Add reaction - uses: GrantBirki/comment@v2.0.1 - with: - comment-id: ${{ github.event.comment.id }} - reactions: rocket - From 043a74e25f6f75bd149dd9e3a3442bd63d17f857 Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Mon, 19 Sep 2022 18:02:44 +0000 Subject: [PATCH 318/828] feat(model_theory/satisfiability): Maximally consistent theories (#16547) Defines `first_order.language.Theory.is_maximal` to denote maximally consistent theories Shows that such theories are complete, and that the complete theory of a structure, as constructed, is maximal Co-authored-by: Aaron Anderson <65780815+awainverse@users.noreply.github.com> --- src/model_theory/bundled.lean | 10 ++++++ src/model_theory/satisfiability.lean | 46 ++++++++++++++++++++++++++-- 2 files changed, 54 insertions(+), 2 deletions(-) diff --git a/src/model_theory/bundled.lean b/src/model_theory/bundled.lean index 3d8b9f7fcacaf..c306a4dd25132 100644 --- a/src/model_theory/bundled.lean +++ b/src/model_theory/bundled.lean @@ -130,6 +130,16 @@ instance right_Structure {L' : language} {T : (L.sum L').Theory} (M : T.Model) : L'.Structure M := (Lhom.sum_inr : L' →ᴸ L.sum L').reduct M +/-- A model of a theory is also a model of any subtheory. -/ +@[simps] def subtheory_Model (M : T.Model) {T' : L.Theory} (h : T' ⊆ T) : + T'.Model := +{ carrier := M, + is_model := ⟨λ φ hφ, realize_sentence_of_mem T (h hφ)⟩ } + +instance subtheory_Model_models (M : T.Model) {T' : L.Theory} (h : T' ⊆ T) : + M.subtheory_Model h ⊨ T := +M.is_model + end Model variables {T} diff --git a/src/model_theory/satisfiability.lean b/src/model_theory/satisfiability.lean index 1930002c68f48..f6f69d596a505 100644 --- a/src/model_theory/satisfiability.lean +++ b/src/model_theory/satisfiability.lean @@ -277,10 +277,50 @@ lemma models_sentence_of_mem {φ : L.sentence} (h : φ ∈ T) : T ⊨ φ := models_sentence_iff.2 (λ _, realize_sentence_of_mem T h) +lemma models_iff_not_satisfiable (φ : L.sentence) : + T ⊨ φ ↔ ¬ is_satisfiable (T ∪ {φ.not}) := +begin + rw [models_sentence_iff, is_satisfiable], + refine ⟨λ h1 h2, (sentence.realize_not _).1 (realize_sentence_of_mem (T ∪ {formula.not φ}) + (set.subset_union_right _ _ (set.mem_singleton _))) + (h1 (h2.some.subtheory_Model (set.subset_union_left _ _))), λ h M, _⟩, + contrapose! h, + rw ← sentence.realize_not at h, + refine ⟨{ carrier := M, + is_model := ⟨λ ψ hψ, hψ.elim (realize_sentence_of_mem _) (λ h', _)⟩, }⟩, + rw set.mem_singleton_iff.1 h', + exact h, +end + /-- A theory is complete when it is satisfiable and models each sentence or its negation. -/ def is_complete (T : L.Theory) : Prop := T.is_satisfiable ∧ ∀ (φ : L.sentence), (T ⊨ φ) ∨ (T ⊨ φ.not) +/-- A theory is maximal when it is satisfiable and contains each sentence or its negation. + Maximal theories are complete. -/ +def is_maximal (T : L.Theory) : Prop := +T.is_satisfiable ∧ ∀ (φ : L.sentence), φ ∈ T ∨ φ.not ∈ T + +lemma is_maximal.is_complete (h : T.is_maximal) : T.is_complete := +h.imp_right (forall_imp (λ _, or.imp models_sentence_of_mem models_sentence_of_mem)) + +lemma is_maximal.mem_or_not_mem (h : T.is_maximal) (φ : L.sentence) : + φ ∈ T ∨ φ.not ∈ T := +h.2 φ + +lemma is_maximal.mem_of_models (h : T.is_maximal) {φ : L.sentence} + (hφ : T ⊨ φ) : + φ ∈ T := +begin + refine (h.mem_or_not_mem φ).resolve_right (λ con, _), + rw [models_iff_not_satisfiable, set.union_singleton, set.insert_eq_of_mem con] at hφ, + exact hφ h.1, +end + +lemma is_maximal.mem_iff_models (h : T.is_maximal) (φ : L.sentence) : + φ ∈ T ↔ T ⊨ φ := +⟨models_sentence_of_mem, h.mem_of_models⟩ + /-- Two (bounded) formulas are semantically equivalent over a theory `T` when they have the same interpretation in every model of `T`. (This is also known as logical equivalence, which also has a proof-theoretic definition.) -/ @@ -375,9 +415,11 @@ lemma mem_or_not_mem (φ : L.sentence) : φ ∈ L.complete_theory M ∨ φ.not ∈ L.complete_theory M := by simp_rw [complete_theory, set.mem_set_of_eq, sentence.realize, formula.realize_not, or_not] +lemma is_maximal [nonempty M] : (L.complete_theory M).is_maximal := +⟨is_satisfiable L M, mem_or_not_mem L M⟩ + lemma is_complete [nonempty M] : (L.complete_theory M).is_complete := -⟨is_satisfiable L M, - λ φ, ((mem_or_not_mem L M φ).imp Theory.models_sentence_of_mem Theory.models_sentence_of_mem)⟩ +(complete_theory.is_maximal L M).is_complete end complete_theory From eb7d26641c404b7275340d71541298eaf79b6c1c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 19 Sep 2022 18:53:17 +0000 Subject: [PATCH 319/828] feat(data/real/cau_seq): more lemmas about sup and inf (#16553) These will be needed to provide `distrib_lattice real` in terms of these operators later. --- src/data/real/cau_seq.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/src/data/real/cau_seq.lean b/src/data/real/cau_seq.lean index 82fee058380cf..7887be24e4149 100644 --- a/src/data/real/cau_seq.lean +++ b/src/data/real/cau_seq.lean @@ -728,6 +728,18 @@ protected lemma inf_eq_left {a b : cau_seq α abs} (h : a ≤ b) : a ⊓ b ≈ a := by simpa only [cau_seq.inf_comm] using cau_seq.inf_eq_right h +protected lemma le_sup_left {a b : cau_seq α abs} : a ≤ a ⊔ b := +le_of_exists ⟨0, λ j hj, le_sup_left⟩ + +protected lemma inf_le_left {a b : cau_seq α abs} : a ⊓ b ≤ a := +le_of_exists ⟨0, λ j hj, inf_le_left⟩ + +protected lemma le_sup_right {a b : cau_seq α abs} : b ≤ a ⊔ b := +le_of_exists ⟨0, λ j hj, le_sup_right⟩ + +protected lemma inf_le_right {a b : cau_seq α abs} : a ⊓ b ≤ b := +le_of_exists ⟨0, λ j hj, inf_le_right⟩ + protected lemma sup_le {a b c : cau_seq α abs} (ha : a ≤ c) (hb : b ≤ c) : a ⊔ b ≤ c := begin cases ha with ha ha, @@ -756,6 +768,12 @@ end /-! Note that `distrib_lattice (cau_seq α abs)` is not true because there is no `partial_order`. -/ +protected lemma sup_inf_distrib_left (a b c : cau_seq α abs) : a ⊔ (b ⊓ c) = (a ⊔ b) ⊓ (a ⊔ c) := +subtype.ext $ funext $ λ i, max_min_distrib_left + +protected lemma sup_inf_distrib_right (a b c : cau_seq α abs) : (a ⊓ b) ⊔ c = (a ⊔ c) ⊓ (b ⊔ c) := +subtype.ext $ funext $ λ i, max_min_distrib_right + end abs end cau_seq From c41e17127f8bdcf18a091900e0e2e5d2d9d4f650 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Mon, 19 Sep 2022 20:38:29 +0000 Subject: [PATCH 320/828] chore(probability/process): split `probability/stopping.lean` into three files, create `process` folder (#16521) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Split the content of stopping.lean into the following three files: * filtration.lean: definition and properties of filtrations * adapted.lean: adapted and progressively measurable processes (and soon predictable as well) * stopping.lean: stopping times, stopped values and stopped processes Co-authored-by: Rémy Degenne --- src/probability/martingale/basic.lean | 2 +- src/probability/martingale/upcrossing.lean | 1 - src/probability/process/adapted.lean | 219 +++++++++ src/probability/process/filtration.lean | 284 +++++++++++ .../{ => process}/hitting_time.lean | 2 +- src/probability/{ => process}/stopping.lean | 450 +----------------- 6 files changed, 511 insertions(+), 447 deletions(-) create mode 100644 src/probability/process/adapted.lean create mode 100644 src/probability/process/filtration.lean rename src/probability/{ => process}/hitting_time.lean (99%) rename src/probability/{ => process}/stopping.lean (72%) diff --git a/src/probability/martingale/basic.lean b/src/probability/martingale/basic.lean index ba824d8012436..8ec1911258fbb 100644 --- a/src/probability/martingale/basic.lean +++ b/src/probability/martingale/basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Kexing Ying -/ import probability.notation -import probability.hitting_time +import probability.process.hitting_time /-! # Martingales diff --git a/src/probability/martingale/upcrossing.lean b/src/probability/martingale/upcrossing.lean index 473fa98d7aa0c..0bd6b97f63d16 100644 --- a/src/probability/martingale/upcrossing.lean +++ b/src/probability/martingale/upcrossing.lean @@ -3,7 +3,6 @@ Copyright (c) 2022 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ -import probability.hitting_time import probability.martingale.basic /-! diff --git a/src/probability/process/adapted.lean b/src/probability/process/adapted.lean new file mode 100644 index 0000000000000..9e5095e49b3ef --- /dev/null +++ b/src/probability/process/adapted.lean @@ -0,0 +1,219 @@ +/- +Copyright (c) 2021 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying, Rémy Degenne +-/ +import probability.process.filtration + +/-! +# Adapted and progressively measurable processes + +This file defines some standard definition from the theory of stochastic processes including +filtrations and stopping times. These definitions are used to model the amount of information +at a specific time and are the first step in formalizing stochastic processes. + +## Main definitions + +* `measure_theory.adapted`: a sequence of functions `u` is said to be adapted to a + filtration `f` if at each point in time `i`, `u i` is `f i`-strongly measurable +* `measure_theory.prog_measurable`: a sequence of functions `u` is said to be progressively + measurable with respect to a filtration `f` if at each point in time `i`, `u` restricted to + `set.Iic i × Ω` is strongly measurable with respect to the product `measurable_space` structure + where the σ-algebra used for `Ω` is `f i`. + +## Main results + +* `adapted.prog_measurable_of_continuous`: a continuous adapted process is progressively measurable. + +## Tags + +adapted, progressively measurable + +-/ + +open filter order topological_space +open_locale classical measure_theory nnreal ennreal topological_space big_operators + +namespace measure_theory + +variables {Ω β ι : Type*} {m : measurable_space Ω} [topological_space β] [preorder ι] + {u v : ι → Ω → β} {f : filtration ι m} + +/-- A sequence of functions `u` is adapted to a filtration `f` if for all `i`, +`u i` is `f i`-measurable. -/ +def adapted (f : filtration ι m) (u : ι → Ω → β) : Prop := +∀ i : ι, strongly_measurable[f i] (u i) + +namespace adapted + +@[protected, to_additive] lemma mul [has_mul β] [has_continuous_mul β] + (hu : adapted f u) (hv : adapted f v) : + adapted f (u * v) := +λ i, (hu i).mul (hv i) + +@[protected, to_additive] lemma div [has_div β] [has_continuous_div β] + (hu : adapted f u) (hv : adapted f v) : + adapted f (u / v) := +λ i, (hu i).div (hv i) + +@[protected, to_additive] lemma inv [group β] [topological_group β] (hu : adapted f u) : + adapted f u⁻¹ := +λ i, (hu i).inv + +@[protected] lemma smul [has_smul ℝ β] [has_continuous_smul ℝ β] (c : ℝ) (hu : adapted f u) : + adapted f (c • u) := +λ i, (hu i).const_smul c + +@[protected] lemma strongly_measurable {i : ι} (hf : adapted f u) : + strongly_measurable[m] (u i) := +(hf i).mono (f.le i) + +lemma strongly_measurable_le {i j : ι} (hf : adapted f u) (hij : i ≤ j) : + strongly_measurable[f j] (u i) := +(hf i).mono (f.mono hij) + +end adapted + +lemma adapted_const (f : filtration ι m) (x : β) : adapted f (λ _ _, x) := +λ i, strongly_measurable_const + +variable (β) +lemma adapted_zero [has_zero β] (f : filtration ι m) : adapted f (0 : ι → Ω → β) := +λ i, @strongly_measurable_zero Ω β (f i) _ _ +variable {β} + +lemma filtration.adapted_natural [metrizable_space β] [mβ : measurable_space β] [borel_space β] + {u : ι → Ω → β} (hum : ∀ i, strongly_measurable[m] (u i)) : + adapted (filtration.natural u hum) u := +begin + assume i, + refine strongly_measurable.mono _ (le_supr₂_of_le i (le_refl i) le_rfl), + rw strongly_measurable_iff_measurable_separable, + exact ⟨measurable_iff_comap_le.2 le_rfl, (hum i).is_separable_range⟩ +end + +/-- Progressively measurable process. A sequence of functions `u` is said to be progressively +measurable with respect to a filtration `f` if at each point in time `i`, `u` restricted to +`set.Iic i × Ω` is measurable with respect to the product `measurable_space` structure where the +σ-algebra used for `Ω` is `f i`. +The usual definition uses the interval `[0,i]`, which we replace by `set.Iic i`. We recover the +usual definition for index types `ℝ≥0` or `ℕ`. -/ +def prog_measurable [measurable_space ι] (f : filtration ι m) (u : ι → Ω → β) : Prop := +∀ i, strongly_measurable[subtype.measurable_space.prod (f i)] (λ p : set.Iic i × Ω, u p.1 p.2) + +lemma prog_measurable_const [measurable_space ι] (f : filtration ι m) (b : β) : + prog_measurable f ((λ _ _, b) : ι → Ω → β) := +λ i, @strongly_measurable_const _ _ (subtype.measurable_space.prod (f i)) _ _ + +namespace prog_measurable + +variables [measurable_space ι] + +protected lemma adapted (h : prog_measurable f u) : adapted f u := +begin + intro i, + have : u i = (λ p : set.Iic i × Ω, u p.1 p.2) ∘ (λ x, (⟨i, set.mem_Iic.mpr le_rfl⟩, x)) := rfl, + rw this, + exact (h i).comp_measurable measurable_prod_mk_left, +end + +protected lemma comp {t : ι → Ω → ι} [topological_space ι] [borel_space ι] [metrizable_space ι] + (h : prog_measurable f u) (ht : prog_measurable f t) + (ht_le : ∀ i ω, t i ω ≤ i) : + prog_measurable f (λ i ω, u (t i ω) ω) := +begin + intro i, + have : (λ p : ↥(set.Iic i) × Ω, u (t (p.fst : ι) p.snd) p.snd) + = (λ p : ↥(set.Iic i) × Ω, u (p.fst : ι) p.snd) ∘ (λ p : ↥(set.Iic i) × Ω, + (⟨t (p.fst : ι) p.snd, set.mem_Iic.mpr ((ht_le _ _).trans p.fst.prop)⟩, p.snd)) := rfl, + rw this, + exact (h i).comp_measurable ((ht i).measurable.subtype_mk.prod_mk measurable_snd), +end + +section arithmetic + +@[to_additive] protected lemma mul [has_mul β] [has_continuous_mul β] + (hu : prog_measurable f u) (hv : prog_measurable f v) : + prog_measurable f (λ i ω, u i ω * v i ω) := +λ i, (hu i).mul (hv i) + +@[to_additive] protected lemma finset_prod' {γ} [comm_monoid β] [has_continuous_mul β] + {U : γ → ι → Ω → β} {s : finset γ} (h : ∀ c ∈ s, prog_measurable f (U c)) : + prog_measurable f (∏ c in s, U c) := +finset.prod_induction U (prog_measurable f) (λ _ _, prog_measurable.mul) + (prog_measurable_const _ 1) h + +@[to_additive] protected lemma finset_prod {γ} [comm_monoid β] [has_continuous_mul β] + {U : γ → ι → Ω → β} {s : finset γ} (h : ∀ c ∈ s, prog_measurable f (U c)) : + prog_measurable f (λ i a, ∏ c in s, U c i a) := +by { convert prog_measurable.finset_prod' h, ext i a, simp only [finset.prod_apply], } + +@[to_additive] protected lemma inv [group β] [topological_group β] (hu : prog_measurable f u) : + prog_measurable f (λ i ω, (u i ω)⁻¹) := +λ i, (hu i).inv + +@[to_additive] protected lemma div [group β] [topological_group β] + (hu : prog_measurable f u) (hv : prog_measurable f v) : + prog_measurable f (λ i ω, u i ω / v i ω) := +λ i, (hu i).div (hv i) + +end arithmetic + +end prog_measurable + +lemma prog_measurable_of_tendsto' {γ} [measurable_space ι] [metrizable_space β] + (fltr : filter γ) [fltr.ne_bot] [fltr.is_countably_generated] {U : γ → ι → Ω → β} + (h : ∀ l, prog_measurable f (U l)) (h_tendsto : tendsto U fltr (𝓝 u)) : + prog_measurable f u := +begin + assume i, + apply @strongly_measurable_of_tendsto (set.Iic i × Ω) β γ (measurable_space.prod _ (f i)) + _ _ fltr _ _ _ _ (λ l, h l i), + rw tendsto_pi_nhds at h_tendsto ⊢, + intro x, + specialize h_tendsto x.fst, + rw tendsto_nhds at h_tendsto ⊢, + exact λ s hs h_mem, h_tendsto {g | g x.snd ∈ s} (hs.preimage (continuous_apply x.snd)) h_mem, +end + +lemma prog_measurable_of_tendsto [measurable_space ι] [metrizable_space β] + {U : ℕ → ι → Ω → β} + (h : ∀ l, prog_measurable f (U l)) (h_tendsto : tendsto U at_top (𝓝 u)) : + prog_measurable f u := +prog_measurable_of_tendsto' at_top h h_tendsto + +/-- A continuous and adapted process is progressively measurable. -/ +theorem adapted.prog_measurable_of_continuous + [topological_space ι] [metrizable_space ι] [measurable_space ι] + [second_countable_topology ι] [opens_measurable_space ι] [metrizable_space β] + (h : adapted f u) (hu_cont : ∀ ω, continuous (λ i, u i ω)) : + prog_measurable f u := +λ i, @strongly_measurable_uncurry_of_continuous_of_strongly_measurable _ _ (set.Iic i) _ _ _ _ _ _ _ + (f i) _ (λ ω, (hu_cont ω).comp continuous_induced_dom) (λ j, (h j).mono (f.mono j.prop)) + +/-- For filtrations indexed by `ℕ`, `adapted` and `prog_measurable` are equivalent. This lemma +provides `adapted f u → prog_measurable f u`. See `prog_measurable.adapted` for the reverse +direction, which is true more generally. -/ +lemma adapted.prog_measurable_of_nat {f : filtration ℕ m} {u : ℕ → Ω → β} + [add_comm_monoid β] [has_continuous_add β] + (h : adapted f u) : prog_measurable f u := +begin + intro i, + have : (λ p : ↥(set.Iic i) × Ω, u ↑(p.fst) p.snd) + = λ p : ↥(set.Iic i) × Ω, ∑ j in finset.range (i + 1), if ↑p.fst = j then u j p.snd else 0, + { ext1 p, + rw finset.sum_ite_eq, + have hp_mem : (p.fst : ℕ) ∈ finset.range (i + 1) := finset.mem_range_succ_iff.mpr p.fst.prop, + simp only [hp_mem, if_true], }, + rw this, + refine finset.strongly_measurable_sum _ (λ j hj, strongly_measurable.ite _ _ _), + { suffices h_meas : measurable[measurable_space.prod _ (f i)] + (λ a : ↥(set.Iic i) × Ω, (a.fst : ℕ)), + from h_meas (measurable_set_singleton j), + exact measurable_fst.subtype_coe, }, + { have h_le : j ≤ i, from finset.mem_range_succ_iff.mp hj, + exact (strongly_measurable.mono (h j) (f.mono h_le)).comp_measurable measurable_snd, }, + { exact strongly_measurable_const, }, +end + +end measure_theory diff --git a/src/probability/process/filtration.lean b/src/probability/process/filtration.lean new file mode 100644 index 0000000000000..46252b0eaea80 --- /dev/null +++ b/src/probability/process/filtration.lean @@ -0,0 +1,284 @@ +/- +Copyright (c) 2021 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying, Rémy Degenne +-/ +import measure_theory.function.conditional_expectation.real + +/-! +# Filtrations + +This file defines filtrations of a measurable space and σ-finite filtrations. + +## Main definitions + +* `measure_theory.filtration`: a filtration on a measurable space. That is, a monotone sequence of + sub-σ-algebras. +* `measure_theory.sigma_finite_filtration`: a filtration `f` is σ-finite with respect to a measure + `μ` if for all `i`, `μ.trim (f.le i)` is σ-finite. +* `measure_theory.filtration.natular`: the smallest filtration that makes a process adapted. That + notion `adapted` is not defined yet in this file. See `measure_theory.adapted`. + +## Main results + +* `measure_theory.filtration.complete_lattice`: filtrations are a complete lattice. + +## Tags + +filtration, stochastic process + +-/ + +open filter order topological_space +open_locale classical measure_theory nnreal ennreal topological_space big_operators + +namespace measure_theory + + +/-- A `filtration` on a measurable space `Ω` with σ-algebra `m` is a monotone +sequence of sub-σ-algebras of `m`. -/ +structure filtration {Ω : Type*} (ι : Type*) [preorder ι] (m : measurable_space Ω) := +(seq : ι → measurable_space Ω) +(mono' : monotone seq) +(le' : ∀ i : ι, seq i ≤ m) + +variables {Ω β ι : Type*} {m : measurable_space Ω} + +instance [preorder ι] : has_coe_to_fun (filtration ι m) (λ _, ι → measurable_space Ω) := +⟨λ f, f.seq⟩ + +namespace filtration +variables [preorder ι] + +protected lemma mono {i j : ι} (f : filtration ι m) (hij : i ≤ j) : f i ≤ f j := f.mono' hij + +protected lemma le (f : filtration ι m) (i : ι) : f i ≤ m := f.le' i + +@[ext] protected lemma ext {f g : filtration ι m} (h : (f : ι → measurable_space Ω) = g) : f = g := +by { cases f, cases g, simp only, exact h, } + +variable (ι) +/-- The constant filtration which is equal to `m` for all `i : ι`. -/ +def const (m' : measurable_space Ω) (hm' : m' ≤ m) : filtration ι m := +⟨λ _, m', monotone_const, λ _, hm'⟩ +variable {ι} + +@[simp] +lemma const_apply {m' : measurable_space Ω} {hm' : m' ≤ m} (i : ι) : const ι m' hm' i = m' := rfl + +instance : inhabited (filtration ι m) := ⟨const ι m le_rfl⟩ + +instance : has_le (filtration ι m) := ⟨λ f g, ∀ i, f i ≤ g i⟩ + +instance : has_bot (filtration ι m) := ⟨const ι ⊥ bot_le⟩ + +instance : has_top (filtration ι m) := ⟨const ι m le_rfl⟩ + +instance : has_sup (filtration ι m) := ⟨λ f g, +{ seq := λ i, f i ⊔ g i, + mono' := λ i j hij, sup_le ((f.mono hij).trans le_sup_left) ((g.mono hij).trans le_sup_right), + le' := λ i, sup_le (f.le i) (g.le i) }⟩ + +@[norm_cast] lemma coe_fn_sup {f g : filtration ι m} : ⇑(f ⊔ g) = f ⊔ g := rfl + +instance : has_inf (filtration ι m) := ⟨λ f g, +{ seq := λ i, f i ⊓ g i, + mono' := λ i j hij, le_inf (inf_le_left.trans (f.mono hij)) (inf_le_right.trans (g.mono hij)), + le' := λ i, inf_le_left.trans (f.le i) }⟩ + +@[norm_cast] lemma coe_fn_inf {f g : filtration ι m} : ⇑(f ⊓ g) = f ⊓ g := rfl + +instance : has_Sup (filtration ι m) := ⟨λ s, +{ seq := λ i, Sup ((λ f : filtration ι m, f i) '' s), + mono' := λ i j hij, + begin + refine Sup_le (λ m' hm', _), + rw [set.mem_image] at hm', + obtain ⟨f, hf_mem, hfm'⟩ := hm', + rw ← hfm', + refine (f.mono hij).trans _, + have hfj_mem : f j ∈ ((λ g : filtration ι m, g j) '' s), from ⟨f, hf_mem, rfl⟩, + exact le_Sup hfj_mem, + end, + le' := λ i, + begin + refine Sup_le (λ m' hm', _), + rw [set.mem_image] at hm', + obtain ⟨f, hf_mem, hfm'⟩ := hm', + rw ← hfm', + exact f.le i, + end, }⟩ + +lemma Sup_def (s : set (filtration ι m)) (i : ι) : + Sup s i = Sup ((λ f : filtration ι m, f i) '' s) := +rfl + +noncomputable +instance : has_Inf (filtration ι m) := ⟨λ s, +{ seq := λ i, if set.nonempty s then Inf ((λ f : filtration ι m, f i) '' s) else m, + mono' := λ i j hij, + begin + by_cases h_nonempty : set.nonempty s, + swap, { simp only [h_nonempty, set.nonempty_image_iff, if_false, le_refl], }, + simp only [h_nonempty, if_true, le_Inf_iff, set.mem_image, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂], + refine λ f hf_mem, le_trans _ (f.mono hij), + have hfi_mem : f i ∈ ((λ g : filtration ι m, g i) '' s), from ⟨f, hf_mem, rfl⟩, + exact Inf_le hfi_mem, + end, + le' := λ i, + begin + by_cases h_nonempty : set.nonempty s, + swap, { simp only [h_nonempty, if_false, le_refl], }, + simp only [h_nonempty, if_true], + obtain ⟨f, hf_mem⟩ := h_nonempty, + exact le_trans (Inf_le ⟨f, hf_mem, rfl⟩) (f.le i), + end, }⟩ + +lemma Inf_def (s : set (filtration ι m)) (i : ι) : + Inf s i = if set.nonempty s then Inf ((λ f : filtration ι m, f i) '' s) else m := +rfl + +noncomputable +instance : complete_lattice (filtration ι m) := +{ le := (≤), + le_refl := λ f i, le_rfl, + le_trans := λ f g h h_fg h_gh i, (h_fg i).trans (h_gh i), + le_antisymm := λ f g h_fg h_gf, filtration.ext $ funext $ λ i, (h_fg i).antisymm (h_gf i), + sup := (⊔), + le_sup_left := λ f g i, le_sup_left, + le_sup_right := λ f g i, le_sup_right, + sup_le := λ f g h h_fh h_gh i, sup_le (h_fh i) (h_gh _), + inf := (⊓), + inf_le_left := λ f g i, inf_le_left, + inf_le_right := λ f g i, inf_le_right, + le_inf := λ f g h h_fg h_fh i, le_inf (h_fg i) (h_fh i), + Sup := Sup, + le_Sup := λ s f hf_mem i, le_Sup ⟨f, hf_mem, rfl⟩, + Sup_le := λ s f h_forall i, Sup_le $ λ m' hm', + begin + obtain ⟨g, hg_mem, hfm'⟩ := hm', + rw ← hfm', + exact h_forall g hg_mem i, + end, + Inf := Inf, + Inf_le := λ s f hf_mem i, + begin + have hs : s.nonempty := ⟨f, hf_mem⟩, + simp only [Inf_def, hs, if_true], + exact Inf_le ⟨f, hf_mem, rfl⟩, + end, + le_Inf := λ s f h_forall i, + begin + by_cases hs : s.nonempty, + swap, { simp only [Inf_def, hs, if_false], exact f.le i, }, + simp only [Inf_def, hs, if_true, le_Inf_iff, set.mem_image, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂], + exact λ g hg_mem, h_forall g hg_mem i, + end, + top := ⊤, + bot := ⊥, + le_top := λ f i, f.le' i, + bot_le := λ f i, bot_le, } + +end filtration + +lemma measurable_set_of_filtration [preorder ι] {f : filtration ι m} {s : set Ω} {i : ι} + (hs : measurable_set[f i] s) : measurable_set[m] s := +f.le i s hs + +/-- A measure is σ-finite with respect to filtration if it is σ-finite with respect +to all the sub-σ-algebra of the filtration. -/ +class sigma_finite_filtration [preorder ι] (μ : measure Ω) (f : filtration ι m) : Prop := +(sigma_finite : ∀ i : ι, sigma_finite (μ.trim (f.le i))) + +instance sigma_finite_of_sigma_finite_filtration [preorder ι] (μ : measure Ω) (f : filtration ι m) + [hf : sigma_finite_filtration μ f] (i : ι) : + sigma_finite (μ.trim (f.le i)) := +by apply hf.sigma_finite -- can't exact here + +@[priority 100] +instance is_finite_measure.sigma_finite_filtration [preorder ι] (μ : measure Ω) (f : filtration ι m) + [is_finite_measure μ] : + sigma_finite_filtration μ f := +⟨λ n, by apply_instance⟩ + +/-- Given a integrable function `g`, the conditional expectations of `g` with respect to a +filtration is uniformly integrable. -/ +lemma integrable.uniform_integrable_condexp_filtration + [preorder ι] {μ : measure Ω} [is_finite_measure μ] {f : filtration ι m} + {g : Ω → ℝ} (hg : integrable g μ) : + uniform_integrable (λ i, μ[g | f i]) 1 μ := +hg.uniform_integrable_condexp f.le + + +namespace filtration +variables [topological_space β] [metrizable_space β] [mβ : measurable_space β] [borel_space β] + [preorder ι] + +include mβ + +/-- Given a sequence of functions, the natural filtration is the smallest sequence +of σ-algebras such that that sequence of functions is measurable with respect to +the filtration. -/ +def natural (u : ι → Ω → β) (hum : ∀ i, strongly_measurable (u i)) : filtration ι m := +{ seq := λ i, ⨆ j ≤ i, measurable_space.comap (u j) mβ, + mono' := λ i j hij, bsupr_mono $ λ k, ge_trans hij, + le' := λ i, + begin + refine supr₂_le _, + rintros j hj s ⟨t, ht, rfl⟩, + exact (hum j).measurable ht, + end } + +section limit + +omit mβ + +variables {E : Type*} [has_zero E] [topological_space E] + {ℱ : filtration ι m} {f : ι → Ω → E} {μ : measure Ω} + +/-- Given a process `f` and a filtration `ℱ`, if `f` converges to some `g` almost everywhere and +`g` is `⨆ n, ℱ n`-measurable, then `limit_process f ℱ μ` chooses said `g`, else it returns 0. + +This definition is used to phrase the a.e. martingale convergence theorem +`submartingale.ae_tendsto_limit_process` where an L¹-bounded submartingale `f` adapted to `ℱ` +converges to `limit_process f ℱ μ` `μ`-almost everywhere. -/ +noncomputable +def limit_process (f : ι → Ω → E) (ℱ : filtration ι m) (μ : measure Ω . volume_tac) := +if h : ∃ g : Ω → E, strongly_measurable[⨆ n, ℱ n] g ∧ + ∀ᵐ ω ∂μ, tendsto (λ n, f n ω) at_top (𝓝 (g ω)) then classical.some h else 0 + +lemma strongly_measurable_limit_process : + strongly_measurable[⨆ n, ℱ n] (limit_process f ℱ μ) := +begin + rw limit_process, + split_ifs with h h, + exacts [(classical.some_spec h).1, strongly_measurable_zero] +end + +lemma strongly_measurable_limit_process' : + strongly_measurable[m] (limit_process f ℱ μ) := +strongly_measurable_limit_process.mono (Sup_le (λ m ⟨n, hn⟩, hn ▸ ℱ.le _)) + +lemma mem_ℒp_limit_process_of_snorm_bdd {R : ℝ≥0} {p : ℝ≥0∞} + {F : Type*} [normed_add_comm_group F] {ℱ : filtration ℕ m} {f : ℕ → Ω → F} + (hfm : ∀ n, ae_strongly_measurable (f n) μ) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : + mem_ℒp (limit_process f ℱ μ) p μ := +begin + rw limit_process, + split_ifs with h, + { refine ⟨strongly_measurable.ae_strongly_measurable + ((classical.some_spec h).1.mono (Sup_le (λ m ⟨n, hn⟩, hn ▸ ℱ.le _))), + lt_of_le_of_lt (Lp.snorm_lim_le_liminf_snorm hfm _ (classical.some_spec h).2) + (lt_of_le_of_lt _ (ennreal.coe_lt_top : ↑R < ∞))⟩, + simp_rw [liminf_eq, eventually_at_top], + exact Sup_le (λ b ⟨a, ha⟩, (ha a le_rfl).trans (hbdd _)) }, + { exact zero_mem_ℒp } +end + +end limit + +end filtration + +end measure_theory diff --git a/src/probability/hitting_time.lean b/src/probability/process/hitting_time.lean similarity index 99% rename from src/probability/hitting_time.lean rename to src/probability/process/hitting_time.lean index 62171ead4962e..e2e66876a9324 100644 --- a/src/probability/hitting_time.lean +++ b/src/probability/process/hitting_time.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying, Rémy Degenne -/ -import probability.stopping +import probability.process.stopping /-! # Hitting time diff --git a/src/probability/stopping.lean b/src/probability/process/stopping.lean similarity index 72% rename from src/probability/stopping.lean rename to src/probability/process/stopping.lean index 93263f9bc8810..1c49d729de55f 100644 --- a/src/probability/stopping.lean +++ b/src/probability/process/stopping.lean @@ -1,29 +1,19 @@ /- Copyright (c) 2021 Kexing Ying. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kexing Ying +Authors: Kexing Ying, Rémy Degenne -/ -import measure_theory.function.conditional_expectation.real + +import probability.process.adapted import topology.instances.discrete /-! -# Filtration and stopping time +# Stopping times, stopped processes and stopped values -This file defines some standard definition from the theory of stochastic processes including -filtrations and stopping times. These definitions are used to model the amount of information -at a specific time and is the first step in formalizing stochastic processes. +Definition and properties of stopping times. ## Main definitions -* `measure_theory.filtration`: a filtration on a measurable space -* `measure_theory.adapted`: a sequence of functions `u` is said to be adapted to a - filtration `f` if at each point in time `i`, `u i` is `f i`-strongly measurable -* `measure_theory.prog_measurable`: a sequence of functions `u` is said to be progressively - measurable with respect to a filtration `f` if at each point in time `i`, `u` restricted to - `set.Iic i × Ω` is strongly measurable with respect to the product `measurable_space` structure - where the σ-algebra used for `Ω` is `f i`. -* `measure_theory.filtration.natural`: the natural filtration with respect to a sequence of - measurable functions is the smallest filtration to which it is adapted to * `measure_theory.is_stopping_time`: a stopping time with respect to some filtration `f` is a function `τ` such that for all `i`, the preimage of `{j | j ≤ i}` along `τ` is `f i`-measurable @@ -31,7 +21,6 @@ at a specific time and is the first step in formalizing stochastic processes. ## Main results -* `adapted.prog_measurable_of_continuous`: a continuous adapted process is progressively measurable. * `prog_measurable.stopped_process`: the stopped process of a progressively measurable process is progressively measurable. * `mem_ℒp_stopped_process`: if a process belongs to `ℒp` at every time in `ℕ`, then its stopped @@ -39,7 +28,7 @@ at a specific time and is the first step in formalizing stochastic processes. ## Tags -filtration, stopping time, stochastic process +stopping time, stochastic process -/ @@ -48,411 +37,8 @@ open_locale classical measure_theory nnreal ennreal topological_space big_operat namespace measure_theory -/-! ### Filtrations -/ - -/-- A `filtration` on measurable space `Ω` with σ-algebra `m` is a monotone -sequence of sub-σ-algebras of `m`. -/ -structure filtration {Ω : Type*} (ι : Type*) [preorder ι] (m : measurable_space Ω) := -(seq : ι → measurable_space Ω) -(mono' : monotone seq) -(le' : ∀ i : ι, seq i ≤ m) - variables {Ω β ι : Type*} {m : measurable_space Ω} -instance [preorder ι] : has_coe_to_fun (filtration ι m) (λ _, ι → measurable_space Ω) := -⟨λ f, f.seq⟩ - -namespace filtration -variables [preorder ι] - -protected lemma mono {i j : ι} (f : filtration ι m) (hij : i ≤ j) : f i ≤ f j := f.mono' hij - -protected lemma le (f : filtration ι m) (i : ι) : f i ≤ m := f.le' i - -@[ext] protected lemma ext {f g : filtration ι m} (h : (f : ι → measurable_space Ω) = g) : f = g := -by { cases f, cases g, simp only, exact h, } - -variable (ι) -/-- The constant filtration which is equal to `m` for all `i : ι`. -/ -def const (m' : measurable_space Ω) (hm' : m' ≤ m) : filtration ι m := -⟨λ _, m', monotone_const, λ _, hm'⟩ -variable {ι} - -@[simp] -lemma const_apply {m' : measurable_space Ω} {hm' : m' ≤ m} (i : ι) : const ι m' hm' i = m' := rfl - -instance : inhabited (filtration ι m) := ⟨const ι m le_rfl⟩ - -instance : has_le (filtration ι m) := ⟨λ f g, ∀ i, f i ≤ g i⟩ - -instance : has_bot (filtration ι m) := ⟨const ι ⊥ bot_le⟩ - -instance : has_top (filtration ι m) := ⟨const ι m le_rfl⟩ - -instance : has_sup (filtration ι m) := ⟨λ f g, -{ seq := λ i, f i ⊔ g i, - mono' := λ i j hij, sup_le ((f.mono hij).trans le_sup_left) ((g.mono hij).trans le_sup_right), - le' := λ i, sup_le (f.le i) (g.le i) }⟩ - -@[norm_cast] lemma coe_fn_sup {f g : filtration ι m} : ⇑(f ⊔ g) = f ⊔ g := rfl - -instance : has_inf (filtration ι m) := ⟨λ f g, -{ seq := λ i, f i ⊓ g i, - mono' := λ i j hij, le_inf (inf_le_left.trans (f.mono hij)) (inf_le_right.trans (g.mono hij)), - le' := λ i, inf_le_left.trans (f.le i) }⟩ - -@[norm_cast] lemma coe_fn_inf {f g : filtration ι m} : ⇑(f ⊓ g) = f ⊓ g := rfl - -instance : has_Sup (filtration ι m) := ⟨λ s, -{ seq := λ i, Sup ((λ f : filtration ι m, f i) '' s), - mono' := λ i j hij, - begin - refine Sup_le (λ m' hm', _), - rw [set.mem_image] at hm', - obtain ⟨f, hf_mem, hfm'⟩ := hm', - rw ← hfm', - refine (f.mono hij).trans _, - have hfj_mem : f j ∈ ((λ g : filtration ι m, g j) '' s), from ⟨f, hf_mem, rfl⟩, - exact le_Sup hfj_mem, - end, - le' := λ i, - begin - refine Sup_le (λ m' hm', _), - rw [set.mem_image] at hm', - obtain ⟨f, hf_mem, hfm'⟩ := hm', - rw ← hfm', - exact f.le i, - end, }⟩ - -lemma Sup_def (s : set (filtration ι m)) (i : ι) : - Sup s i = Sup ((λ f : filtration ι m, f i) '' s) := -rfl - -noncomputable -instance : has_Inf (filtration ι m) := ⟨λ s, -{ seq := λ i, if set.nonempty s then Inf ((λ f : filtration ι m, f i) '' s) else m, - mono' := λ i j hij, - begin - by_cases h_nonempty : set.nonempty s, - swap, { simp only [h_nonempty, set.nonempty_image_iff, if_false, le_refl], }, - simp only [h_nonempty, if_true, le_Inf_iff, set.mem_image, forall_exists_index, and_imp, - forall_apply_eq_imp_iff₂], - refine λ f hf_mem, le_trans _ (f.mono hij), - have hfi_mem : f i ∈ ((λ g : filtration ι m, g i) '' s), from ⟨f, hf_mem, rfl⟩, - exact Inf_le hfi_mem, - end, - le' := λ i, - begin - by_cases h_nonempty : set.nonempty s, - swap, { simp only [h_nonempty, if_false, le_refl], }, - simp only [h_nonempty, if_true], - obtain ⟨f, hf_mem⟩ := h_nonempty, - exact le_trans (Inf_le ⟨f, hf_mem, rfl⟩) (f.le i), - end, }⟩ - -lemma Inf_def (s : set (filtration ι m)) (i : ι) : - Inf s i = if set.nonempty s then Inf ((λ f : filtration ι m, f i) '' s) else m := -rfl - -noncomputable -instance : complete_lattice (filtration ι m) := -{ le := (≤), - le_refl := λ f i, le_rfl, - le_trans := λ f g h h_fg h_gh i, (h_fg i).trans (h_gh i), - le_antisymm := λ f g h_fg h_gf, filtration.ext $ funext $ λ i, (h_fg i).antisymm (h_gf i), - sup := (⊔), - le_sup_left := λ f g i, le_sup_left, - le_sup_right := λ f g i, le_sup_right, - sup_le := λ f g h h_fh h_gh i, sup_le (h_fh i) (h_gh _), - inf := (⊓), - inf_le_left := λ f g i, inf_le_left, - inf_le_right := λ f g i, inf_le_right, - le_inf := λ f g h h_fg h_fh i, le_inf (h_fg i) (h_fh i), - Sup := Sup, - le_Sup := λ s f hf_mem i, le_Sup ⟨f, hf_mem, rfl⟩, - Sup_le := λ s f h_forall i, Sup_le $ λ m' hm', - begin - obtain ⟨g, hg_mem, hfm'⟩ := hm', - rw ← hfm', - exact h_forall g hg_mem i, - end, - Inf := Inf, - Inf_le := λ s f hf_mem i, - begin - have hs : s.nonempty := ⟨f, hf_mem⟩, - simp only [Inf_def, hs, if_true], - exact Inf_le ⟨f, hf_mem, rfl⟩, - end, - le_Inf := λ s f h_forall i, - begin - by_cases hs : s.nonempty, - swap, { simp only [Inf_def, hs, if_false], exact f.le i, }, - simp only [Inf_def, hs, if_true, le_Inf_iff, set.mem_image, forall_exists_index, and_imp, - forall_apply_eq_imp_iff₂], - exact λ g hg_mem, h_forall g hg_mem i, - end, - top := ⊤, - bot := ⊥, - le_top := λ f i, f.le' i, - bot_le := λ f i, bot_le, } - -end filtration - -lemma measurable_set_of_filtration [preorder ι] {f : filtration ι m} {s : set Ω} {i : ι} - (hs : measurable_set[f i] s) : measurable_set[m] s := -f.le i s hs - -/-- A measure is σ-finite with respect to filtration if it is σ-finite with respect -to all the sub-σ-algebra of the filtration. -/ -class sigma_finite_filtration [preorder ι] (μ : measure Ω) (f : filtration ι m) : Prop := -(sigma_finite : ∀ i : ι, sigma_finite (μ.trim (f.le i))) - -instance sigma_finite_of_sigma_finite_filtration [preorder ι] (μ : measure Ω) (f : filtration ι m) - [hf : sigma_finite_filtration μ f] (i : ι) : - sigma_finite (μ.trim (f.le i)) := -by apply hf.sigma_finite -- can't exact here - -@[priority 100] -instance is_finite_measure.sigma_finite_filtration [preorder ι] (μ : measure Ω) (f : filtration ι m) - [is_finite_measure μ] : - sigma_finite_filtration μ f := -⟨λ n, by apply_instance⟩ - -/-- Given a integrable function `g`, the conditional expectations of `g` with respect to a -filtration is uniformly integrable. -/ -lemma integrable.uniform_integrable_condexp_filtration - [preorder ι] {μ : measure Ω} [is_finite_measure μ] {f : filtration ι m} - {g : Ω → ℝ} (hg : integrable g μ) : - uniform_integrable (λ i, μ[g | f i]) 1 μ := -hg.uniform_integrable_condexp f.le - -section adapted_process - -variables [topological_space β] [preorder ι] - {u v : ι → Ω → β} {f : filtration ι m} - -/-- A sequence of functions `u` is adapted to a filtration `f` if for all `i`, -`u i` is `f i`-measurable. -/ -def adapted (f : filtration ι m) (u : ι → Ω → β) : Prop := -∀ i : ι, strongly_measurable[f i] (u i) - -namespace adapted - -@[protected, to_additive] lemma mul [has_mul β] [has_continuous_mul β] - (hu : adapted f u) (hv : adapted f v) : - adapted f (u * v) := -λ i, (hu i).mul (hv i) - -@[protected, to_additive] lemma div [has_div β] [has_continuous_div β] - (hu : adapted f u) (hv : adapted f v) : - adapted f (u / v) := -λ i, (hu i).div (hv i) - -@[protected, to_additive] lemma inv [group β] [topological_group β] (hu : adapted f u) : - adapted f u⁻¹ := -λ i, (hu i).inv - -@[protected] lemma smul [has_smul ℝ β] [has_continuous_smul ℝ β] (c : ℝ) (hu : adapted f u) : - adapted f (c • u) := -λ i, (hu i).const_smul c - -@[protected] lemma strongly_measurable {i : ι} (hf : adapted f u) : - strongly_measurable[m] (u i) := -(hf i).mono (f.le i) - -lemma strongly_measurable_le {i j : ι} (hf : adapted f u) (hij : i ≤ j) : - strongly_measurable[f j] (u i) := -(hf i).mono (f.mono hij) - -end adapted - -lemma adapted_const (f : filtration ι m) (x : β) : adapted f (λ _ _, x) := -λ i, strongly_measurable_const - -variable (β) -lemma adapted_zero [has_zero β] (f : filtration ι m) : adapted f (0 : ι → Ω → β) := -λ i, @strongly_measurable_zero Ω β (f i) _ _ -variable {β} - -/-- Progressively measurable process. A sequence of functions `u` is said to be progressively -measurable with respect to a filtration `f` if at each point in time `i`, `u` restricted to -`set.Iic i × Ω` is measurable with respect to the product `measurable_space` structure where the -σ-algebra used for `Ω` is `f i`. -The usual definition uses the interval `[0,i]`, which we replace by `set.Iic i`. We recover the -usual definition for index types `ℝ≥0` or `ℕ`. -/ -def prog_measurable [measurable_space ι] (f : filtration ι m) (u : ι → Ω → β) : Prop := -∀ i, strongly_measurable[subtype.measurable_space.prod (f i)] (λ p : set.Iic i × Ω, u p.1 p.2) - -lemma prog_measurable_const [measurable_space ι] (f : filtration ι m) (b : β) : - prog_measurable f ((λ _ _, b) : ι → Ω → β) := -λ i, @strongly_measurable_const _ _ (subtype.measurable_space.prod (f i)) _ _ - -namespace prog_measurable - -variables [measurable_space ι] - -protected lemma adapted (h : prog_measurable f u) : adapted f u := -begin - intro i, - have : u i = (λ p : set.Iic i × Ω, u p.1 p.2) ∘ (λ x, (⟨i, set.mem_Iic.mpr le_rfl⟩, x)) := rfl, - rw this, - exact (h i).comp_measurable measurable_prod_mk_left, -end - -protected lemma comp {t : ι → Ω → ι} [topological_space ι] [borel_space ι] [metrizable_space ι] - (h : prog_measurable f u) (ht : prog_measurable f t) - (ht_le : ∀ i ω, t i ω ≤ i) : - prog_measurable f (λ i ω, u (t i ω) ω) := -begin - intro i, - have : (λ p : ↥(set.Iic i) × Ω, u (t (p.fst : ι) p.snd) p.snd) - = (λ p : ↥(set.Iic i) × Ω, u (p.fst : ι) p.snd) ∘ (λ p : ↥(set.Iic i) × Ω, - (⟨t (p.fst : ι) p.snd, set.mem_Iic.mpr ((ht_le _ _).trans p.fst.prop)⟩, p.snd)) := rfl, - rw this, - exact (h i).comp_measurable ((ht i).measurable.subtype_mk.prod_mk measurable_snd), -end - -section arithmetic - -@[to_additive] protected lemma mul [has_mul β] [has_continuous_mul β] - (hu : prog_measurable f u) (hv : prog_measurable f v) : - prog_measurable f (λ i ω, u i ω * v i ω) := -λ i, (hu i).mul (hv i) - -@[to_additive] protected lemma finset_prod' {γ} [comm_monoid β] [has_continuous_mul β] - {U : γ → ι → Ω → β} {s : finset γ} (h : ∀ c ∈ s, prog_measurable f (U c)) : - prog_measurable f (∏ c in s, U c) := -finset.prod_induction U (prog_measurable f) (λ _ _, prog_measurable.mul) - (prog_measurable_const _ 1) h - -@[to_additive] protected lemma finset_prod {γ} [comm_monoid β] [has_continuous_mul β] - {U : γ → ι → Ω → β} {s : finset γ} (h : ∀ c ∈ s, prog_measurable f (U c)) : - prog_measurable f (λ i a, ∏ c in s, U c i a) := -by { convert prog_measurable.finset_prod' h, ext i a, simp only [finset.prod_apply], } - -@[to_additive] protected lemma inv [group β] [topological_group β] (hu : prog_measurable f u) : - prog_measurable f (λ i ω, (u i ω)⁻¹) := -λ i, (hu i).inv - -@[to_additive] protected lemma div [group β] [topological_group β] - (hu : prog_measurable f u) (hv : prog_measurable f v) : - prog_measurable f (λ i ω, u i ω / v i ω) := -λ i, (hu i).div (hv i) - -end arithmetic - -end prog_measurable - -lemma prog_measurable_of_tendsto' {γ} [measurable_space ι] [metrizable_space β] - (fltr : filter γ) [fltr.ne_bot] [fltr.is_countably_generated] {U : γ → ι → Ω → β} - (h : ∀ l, prog_measurable f (U l)) (h_tendsto : tendsto U fltr (𝓝 u)) : - prog_measurable f u := -begin - assume i, - apply @strongly_measurable_of_tendsto (set.Iic i × Ω) β γ (measurable_space.prod _ (f i)) - _ _ fltr _ _ _ _ (λ l, h l i), - rw tendsto_pi_nhds at h_tendsto ⊢, - intro x, - specialize h_tendsto x.fst, - rw tendsto_nhds at h_tendsto ⊢, - exact λ s hs h_mem, h_tendsto {g | g x.snd ∈ s} (hs.preimage (continuous_apply x.snd)) h_mem, -end - -lemma prog_measurable_of_tendsto [measurable_space ι] [metrizable_space β] - {U : ℕ → ι → Ω → β} - (h : ∀ l, prog_measurable f (U l)) (h_tendsto : tendsto U at_top (𝓝 u)) : - prog_measurable f u := -prog_measurable_of_tendsto' at_top h h_tendsto - - -/-- A continuous and adapted process is progressively measurable. -/ -theorem adapted.prog_measurable_of_continuous - [topological_space ι] [metrizable_space ι] [measurable_space ι] - [second_countable_topology ι] [opens_measurable_space ι] [metrizable_space β] - (h : adapted f u) (hu_cont : ∀ ω, continuous (λ i, u i ω)) : - prog_measurable f u := -λ i, @strongly_measurable_uncurry_of_continuous_of_strongly_measurable _ _ (set.Iic i) _ _ _ _ _ _ _ - (f i) _ (λ ω, (hu_cont ω).comp continuous_induced_dom) (λ j, (h j).mono (f.mono j.prop)) - -end adapted_process - -namespace filtration -variables [topological_space β] [metrizable_space β] [mβ : measurable_space β] [borel_space β] - [preorder ι] - -include mβ - -/-- Given a sequence of functions, the natural filtration is the smallest sequence -of σ-algebras such that that sequence of functions is measurable with respect to -the filtration. -/ -def natural (u : ι → Ω → β) (hum : ∀ i, strongly_measurable (u i)) : filtration ι m := -{ seq := λ i, ⨆ j ≤ i, measurable_space.comap (u j) mβ, - mono' := λ i j hij, bsupr_mono $ λ k, ge_trans hij, - le' := λ i, - begin - refine supr₂_le _, - rintros j hj s ⟨t, ht, rfl⟩, - exact (hum j).measurable ht, - end } - -lemma adapted_natural {u : ι → Ω → β} (hum : ∀ i, strongly_measurable[m] (u i)) : - adapted (natural u hum) u := -begin - assume i, - refine strongly_measurable.mono _ (le_supr₂_of_le i (le_refl i) le_rfl), - rw strongly_measurable_iff_measurable_separable, - exact ⟨measurable_iff_comap_le.2 le_rfl, (hum i).is_separable_range⟩ -end - -section limit - -omit mβ - -variables {E : Type*} [has_zero E] [topological_space E] - {ℱ : filtration ι m} {f : ι → Ω → E} {μ : measure Ω} - -/-- Given a process `f` and a filtration `ℱ`, if `f` converges to some `g` almost everywhere and -`g` is `⨆ n, ℱ n`-measurable, then `limit_process f ℱ μ` chooses said `g`, else it returns 0. - -This definition is used to phrase the a.e. martingale convergence theorem -`submartingale.ae_tendsto_limit_process` where an L¹-bounded submartingale `f` adapted to `ℱ` -converges to `limit_process f ℱ μ` `μ`-almost everywhere. -/ -noncomputable -def limit_process (f : ι → Ω → E) (ℱ : filtration ι m) (μ : measure Ω . volume_tac) := -if h : ∃ g : Ω → E, strongly_measurable[⨆ n, ℱ n] g ∧ - ∀ᵐ ω ∂μ, tendsto (λ n, f n ω) at_top (𝓝 (g ω)) then classical.some h else 0 - -lemma strongly_measurable_limit_process : - strongly_measurable[⨆ n, ℱ n] (limit_process f ℱ μ) := -begin - rw limit_process, - split_ifs with h h, - exacts [(classical.some_spec h).1, strongly_measurable_zero] -end - -lemma strongly_measurable_limit_process' : - strongly_measurable[m] (limit_process f ℱ μ) := -strongly_measurable_limit_process.mono (Sup_le (λ m ⟨n, hn⟩, hn ▸ ℱ.le _)) - -lemma mem_ℒp_limit_process_of_snorm_bdd {R : ℝ≥0} {p : ℝ≥0∞} - {F : Type*} [normed_add_comm_group F] {ℱ : filtration ℕ m} {f : ℕ → Ω → F} - (hfm : ∀ n, ae_strongly_measurable (f n) μ) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : - mem_ℒp (limit_process f ℱ μ) p μ := -begin - rw limit_process, - split_ifs with h, - { refine ⟨strongly_measurable.ae_strongly_measurable - ((classical.some_spec h).1.mono (Sup_le (λ m ⟨n, hn⟩, hn ▸ ℱ.le _))), - lt_of_le_of_lt (Lp.snorm_lim_le_liminf_snorm hfm _ (classical.some_spec h).2) - (lt_of_le_of_lt _ (ennreal.coe_lt_top : ↑R < ∞))⟩, - simp_rw [liminf_eq, eventually_at_top], - exact Sup_le (λ b ⟨a, ha⟩, (ha a le_rfl).trans (hbdd _)) }, - { exact zero_mem_ℒp } -end - -end limit - -end filtration /-! ### Stopping times -/ @@ -1444,30 +1030,6 @@ section add_comm_monoid variables [add_comm_monoid β] -/-- For filtrations indexed by `ℕ`, `adapted` and `prog_measurable` are equivalent. This lemma -provides `adapted f u → prog_measurable f u`. See `prog_measurable.adapted` for the reverse -direction, which is true more generally. -/ -lemma adapted.prog_measurable_of_nat [topological_space β] [has_continuous_add β] - (h : adapted f u) : prog_measurable f u := -begin - intro i, - have : (λ p : ↥(set.Iic i) × Ω, u ↑(p.fst) p.snd) - = λ p : ↥(set.Iic i) × Ω, ∑ j in finset.range (i + 1), if ↑p.fst = j then u j p.snd else 0, - { ext1 p, - rw finset.sum_ite_eq, - have hp_mem : (p.fst : ℕ) ∈ finset.range (i + 1) := finset.mem_range_succ_iff.mpr p.fst.prop, - simp only [hp_mem, if_true], }, - rw this, - refine finset.strongly_measurable_sum _ (λ j hj, strongly_measurable.ite _ _ _), - { suffices h_meas : measurable[measurable_space.prod _ (f i)] - (λ a : ↥(set.Iic i) × Ω, (a.fst : ℕ)), - from h_meas (measurable_set_singleton j), - exact measurable_fst.subtype_coe, }, - { have h_le : j ≤ i, from finset.mem_range_succ_iff.mp hj, - exact (strongly_measurable.mono (h j) (f.mono h_le)).comp_measurable measurable_snd, }, - { exact strongly_measurable_const, }, -end - /-- For filtrations indexed by `ℕ`, the stopped process obtained from an adapted process is adapted. -/ lemma adapted.stopped_process_of_nat [topological_space β] [has_continuous_add β] From 9824a558f8323af6f4388a266443cad4b4ab4e41 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Mon, 19 Sep 2022 20:38:31 +0000 Subject: [PATCH 321/828] feat(measure_theory/measurable_space): pi has measurable singletons (#16542) This generalises a result used for the birthday problem. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- .../100-theorems-list/93_birthday_problem.lean | 16 ++++------------ src/measure_theory/measurable_space.lean | 9 +++++++++ 2 files changed, 13 insertions(+), 12 deletions(-) diff --git a/archive/100-theorems-list/93_birthday_problem.lean b/archive/100-theorems-list/93_birthday_problem.lean index 921d534c83eab..6d9c2d8bf3b1d 100644 --- a/archive/100-theorems-list/93_birthday_problem.lean +++ b/archive/100-theorems-list/93_birthday_problem.lean @@ -38,25 +38,17 @@ variables {n m : ℕ} /- In order for Lean to understand that we can take probabilities in `fin 23 → fin 365`, we must tell Lean that there is a `measurable_space` structure on the space. Note that this instance is only for `fin m` - Lean automatically figures out that the function space `fin n → fin m` -is _also_ measurable, by using `measurable_space.pi` -/ - +is _also_ measurable, by using `measurable_space.pi`, and furthermore that all sets are measurable, +from `measurable_singleton_class.pi`. -/ instance : measurable_space (fin m) := ⊤ +instance : measurable_singleton_class (fin m) := ⟨λ _, trivial⟩ /- We then endow the space with a canonical measure, which is called ℙ. We define this to be the conditional counting measure. -/ noncomputable instance : measure_space (fin n → fin m) := ⟨cond_count set.univ⟩ --- Singletons are measurable; therefore, as `fin n → fin m` is finite, all sets are measurable. -instance : measurable_singleton_class (fin n → fin m) := -⟨λ f, begin - convert measurable_set.pi set.finite_univ.countable - (show ∀ i, i ∈ set.univ → measurable_set ({f i} : set (fin m)), from λ _ _, trivial), - ext g, - simp only [function.funext_iff, set.mem_singleton_iff, set.mem_univ_pi], -end⟩ - /- The canonical measure on `fin n → fin m` is a probability measure (except on an empty space). -/ -example : is_probability_measure (ℙ : measure (fin n → fin (m + 1))) := +instance : is_probability_measure (ℙ : measure (fin n → fin (m + 1))) := cond_count_is_probability_measure set.finite_univ set.univ_nonempty lemma fin_fin.measure_apply {s : set $ fin n → fin m} : diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index b7c4d0ab4c33a..bc6166e6fb754 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -697,6 +697,15 @@ begin { simp [measurable_set_pi_of_nonempty hs, h, ← not_nonempty_iff_eq_empty] } end +instance measurable_singleton_class.pi [countable δ] [Π a, measurable_singleton_class (π a)] : + measurable_singleton_class (Π a, π a) := +⟨λ f, begin + convert measurable_set.univ_pi + (show ∀ t, measurable_set {f t}, from λ t, measurable_set_singleton (f t)), + ext g, + simp only [function.funext_iff, mem_singleton_iff, mem_univ_pi], +end⟩ + variable (π) @[measurability] From b3951c65c6e797ff162ae8b69eab0063bcfb3d73 Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Mon, 19 Sep 2022 20:38:32 +0000 Subject: [PATCH 322/828] feat(model_theory/*): Lemmas about satisfiability (#16546) Proves that applying an injective language map to a theory does not change its satisfiability (`first_order.language.Theory.is_satisfiable_of_is_satisfiable_on_Theory`) Proves that a union of theories is satisfiable iff any finite union of the theories is (`first_order.language.Theory.is_satisfiable_Union_iff_is_satisfiable_Union_finset`) Proves a few other minor satisfiability and language map lemmas. Co-authored-by: Aaron Anderson <65780815+awainverse@users.noreply.github.com> --- src/model_theory/basic.lean | 2 +- src/model_theory/bundled.lean | 19 ++++++++++++--- src/model_theory/language_map.lean | 36 ++++++++++++++++++++++++++++ src/model_theory/satisfiability.lean | 34 ++++++++++++++++++++++++++ 4 files changed, 87 insertions(+), 4 deletions(-) diff --git a/src/model_theory/basic.lean b/src/model_theory/basic.lean index 047ac9598fa70..8cfc16ffd5ae2 100644 --- a/src/model_theory/basic.lean +++ b/src/model_theory/basic.lean @@ -238,7 +238,7 @@ variables (N : Type w') [L.Structure M] [L.Structure N] open Structure /-- Used for defining `first_order.language.Theory.Model.inhabited`. -/ -def trivial_unit_structure : L.Structure unit := ⟨default, default⟩ +def inhabited.trivial_structure {α : Type*} [inhabited α] : L.Structure α := ⟨default, default⟩ /-! ### Maps -/ diff --git a/src/model_theory/bundled.lean b/src/model_theory/bundled.lean index c306a4dd25132..bcbbb232f22c6 100644 --- a/src/model_theory/bundled.lean +++ b/src/model_theory/bundled.lean @@ -86,10 +86,10 @@ instance (M : T.Model) : nonempty M := infer_instance section inhabited -local attribute [instance] trivial_unit_structure +local attribute [instance] inhabited.trivial_structure -instance : inhabited (Model (∅ : L.Theory)) := -⟨Model.of _ unit⟩ +instance : inhabited (Model.{u v w} (∅ : L.Theory)) := +⟨Model.of _ punit⟩ end inhabited @@ -122,6 +122,19 @@ def ulift (M : Model.{u v w} T) : Model.{u v (max w w')} T := nonempty' := M.nonempty', is_model := (@Lhom.on_Theory_model L L' M (φ.reduct M) _ φ _ T).1 M.is_model, } +/-- When `φ` is injective, `default_expansion` expands a model of `T` to a model of `φ.on_Theory T` + arbitrarily. -/ +@[simps] noncomputable def default_expansion {L' : language} {φ : L →ᴸ L'} (h : φ.injective) + [∀ n (f : L'.functions n), decidable (f ∈ set.range (λ (f : L.functions n), φ.on_function f))] + [∀ n (r : L'.relations n), decidable (r ∈ set.range (λ (r : L.relations n), φ.on_relation r))] + (M : T.Model) [inhabited M] : + (φ.on_Theory T).Model := +{ carrier := M, + struc := φ.default_expansion M, + nonempty' := M.nonempty', + is_model := (@Lhom.on_Theory_model L L' M _ (φ.default_expansion M) φ + (h.is_expansion_on_default M) T).2 M.is_model, } + instance left_Structure {L' : language} {T : (L.sum L').Theory} (M : T.Model) : L.Structure M := (Lhom.sum_inl : L →ᴸ L.sum L').reduct M diff --git a/src/model_theory/language_map.lean b/src/model_theory/language_map.lean index e0eb3c7b02482..7407b211d7e5a 100644 --- a/src/model_theory/language_map.lean +++ b/src/model_theory/language_map.lean @@ -171,6 +171,17 @@ protected structure injective : Prop := (on_function {n} : function.injective (λ f : L.functions n, on_function ϕ f)) (on_relation {n} : function.injective (λ R : L.relations n, on_relation ϕ R)) +/-- Pulls a `L`-structure along a language map `ϕ : L →ᴸ L'`, and then expands it + to an `L'`-structure arbitrarily. -/ +noncomputable def default_expansion (ϕ : L →ᴸ L') + [∀ n (f : L'.functions n), decidable (f ∈ set.range (λ (f : L.functions n), on_function ϕ f))] + [∀ n (r : L'.relations n), decidable (r ∈ set.range (λ (r : L.relations n), on_relation ϕ r))] + (M : Type*) [inhabited M] [L.Structure M] : L'.Structure M := +{ fun_map := λ n f xs, if h' : f ∈ set.range (λ (f : L.functions n), on_function ϕ f) then + fun_map h'.some xs else default, + rel_map := λ n r xs, if h' : r ∈ set.range (λ (r : L.relations n), on_relation ϕ r) then + rel_map h'.some xs else default } + /-- A language homomorphism is an expansion on a structure if it commutes with the interpretation of all symbols on that structure. -/ class is_expansion_on (M : Type*) [L.Structure M] [L'.Structure M] : Prop := @@ -233,6 +244,12 @@ instance sum_inr_is_expansion_on (M : Type*) @fun_map (L'.sum L) M _ n (sum.inr f) x = fun_map f x := (Lhom.sum_inr : L →ᴸ L'.sum L).map_on_function f x +lemma sum_inl_injective : (Lhom.sum_inl : L →ᴸ L.sum L').injective := +⟨λ n, sum.inl_injective, λ n, sum.inl_injective⟩ + +lemma sum_inr_injective : (Lhom.sum_inr : L' →ᴸ L.sum L').injective := +⟨λ n, sum.inr_injective, λ n, sum.inr_injective⟩ + @[priority 100] instance is_expansion_on_reduct (ϕ : L →ᴸ L') (M : Type*) [L'.Structure M] : @is_expansion_on L L' ϕ M (ϕ.reduct M) _ := begin @@ -240,6 +257,22 @@ begin exact ⟨λ _ f _, rfl, λ _ R _, rfl⟩, end +lemma injective.is_expansion_on_default {ϕ : L →ᴸ L'} + [∀ n (f : L'.functions n), decidable (f ∈ set.range (λ (f : L.functions n), on_function ϕ f))] + [∀ n (r : L'.relations n), decidable (r ∈ set.range (λ (r : L.relations n), on_relation ϕ r))] + (h : ϕ.injective) (M : Type*) [inhabited M] [L.Structure M] : + @is_expansion_on L L' ϕ M _ (ϕ.default_expansion M) := +begin + letI := ϕ.default_expansion M, + refine ⟨λ n f xs, _, λ n r xs, _⟩, + { have hf : ϕ.on_function f ∈ set.range (λ (f : L.functions n), ϕ.on_function f) := ⟨f, rfl⟩, + refine (dif_pos hf).trans _, + rw h.on_function hf.some_spec }, + { have hr : ϕ.on_relation r ∈ set.range (λ (r : L.relations n), ϕ.on_relation r) := ⟨r, rfl⟩, + refine (dif_pos hr).trans _, + rw h.on_relation hr.some_spec }, +end + end Lhom /-- A language equivalence maps the symbols of one language to symbols of another bijectively. -/ @@ -344,6 +377,9 @@ by rw [with_constants, card_sum, card_constants_on] /-- The language map adding constants. -/ @[simps] def Lhom_with_constants : L →ᴸ L[[α]] := Lhom.sum_inl +lemma Lhom_with_constants_injective : (L.Lhom_with_constants α).injective := +Lhom.sum_inl_injective + variables {α} /-- The constant symbol indexed by a particular element. -/ diff --git a/src/model_theory/satisfiability.lean b/src/model_theory/satisfiability.lean index f6f69d596a505..76f92280902e8 100644 --- a/src/model_theory/satisfiability.lean +++ b/src/model_theory/satisfiability.lean @@ -70,6 +70,26 @@ lemma is_satisfiable.mono (h : T'.is_satisfiable) (hs : T ⊆ T') : T.is_satisfiable := ⟨(Theory.model.mono (Model.is_model h.some) hs).bundled⟩ +lemma is_satisfiable_empty (L : language.{u v}) : + is_satisfiable (∅ : L.Theory) := +⟨default⟩ + +lemma is_satisfiable_of_is_satisfiable_on_Theory {L' : language.{w w'}} (φ : L →ᴸ L') + (h : (φ.on_Theory T).is_satisfiable) : + T.is_satisfiable := +model.is_satisfiable (h.some.reduct φ) + +lemma is_satisfiable_on_Theory_iff {L' : language.{w w'}} {φ : L →ᴸ L'} + (h : φ.injective) : + (φ.on_Theory T).is_satisfiable ↔ T.is_satisfiable := +begin + classical, + refine ⟨is_satisfiable_of_is_satisfiable_on_Theory φ, + λ h', _⟩, + haveI : inhabited (h'.some) := classical.inhabited_of_nonempty', + exact model.is_satisfiable (h'.some.default_expansion h), +end + lemma is_satisfiable.is_finitely_satisfiable (h : T.is_satisfiable) : T.is_finitely_satisfiable := λ _, h.mono @@ -156,6 +176,20 @@ begin rw lift_lift, end +lemma is_satisfiable_Union_iff_is_satisfiable_Union_finset {ι : Type*} (T : ι → L.Theory) : + is_satisfiable (⋃ i, T i) ↔ ∀ (s : finset ι), is_satisfiable (⋃ (i ∈ s), T i) := +begin + classical, + refine ⟨λ h s, h.mono (set.Union_mono (λ _, set.Union_subset_iff.2 (λ _, refl _))), λ h, _⟩, + rw is_satisfiable_iff_is_finitely_satisfiable, + intros s hs, + rw set.Union_eq_Union_finset at hs, + obtain ⟨t, ht⟩ := directed.exists_mem_subset_of_finset_subset_bUnion _ hs, + { exact (h t).mono ht }, + { exact (monotone.directed_le (λ t1 t2 h, set.Union_mono (λ _, set.Union_mono' + (λ h1, ⟨h h1, refl _⟩)))) }, +end + end Theory variables (L) From ca827bae37707b5b3b3352760d78a349a69447db Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 19 Sep 2022 23:44:49 +0000 Subject: [PATCH 323/828] feat(analysis/special_functions/compare_exp): new file (#16543) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove `(λ z, z ^ a * exp (b * z)) =o[l] λ z, z ^ a' * exp (b' * z)` for an appropriate filter `l`, any complex `a`, `a'`, and real `b < b'`. --- .../asymptotics/asymptotic_equivalent.lean | 9 +- .../special_functions/compare_exp.lean | 182 ++++++++++++++++++ .../special_functions/complex/arg.lean | 2 + .../special_functions/polynomials.lean | 8 +- src/analysis/special_functions/pow.lean | 64 +++++- src/order/filter/at_top_bot.lean | 4 + src/order/liminf_limsup.lean | 3 + 7 files changed, 258 insertions(+), 14 deletions(-) create mode 100644 src/analysis/special_functions/compare_exp.lean diff --git a/src/analysis/asymptotics/asymptotic_equivalent.lean b/src/analysis/asymptotics/asymptotic_equivalent.lean index d3574393f76be..d881bf3941da1 100644 --- a/src/analysis/asymptotics/asymptotic_equivalent.lean +++ b/src/analysis/asymptotics/asymptotic_equivalent.lean @@ -147,11 +147,14 @@ end lemma is_equivalent.tendsto_nhds_iff {c : β} (huv : u ~[l] v) : tendsto u l (𝓝 c) ↔ tendsto v l (𝓝 c) := ⟨huv.tendsto_nhds, huv.symm.tendsto_nhds⟩ -lemma is_equivalent.add_is_o (huv : u ~[l] v) (hwv : w =o[l] v) : (w + u) ~[l] v := -by simpa only [is_equivalent, pi.sub_apply, add_sub] using hwv.add huv +lemma is_equivalent.add_is_o (huv : u ~[l] v) (hwv : w =o[l] v) : (u + w) ~[l] v := +by simpa only [is_equivalent, add_sub_right_comm] using huv.add hwv + +lemma is_equivalent.sub_is_o (huv : u ~[l] v) (hwv : w =o[l] v) : (u - w) ~[l] v := +by simpa only [sub_eq_add_neg] using huv.add_is_o hwv.neg_left lemma is_o.add_is_equivalent (hu : u =o[l] w) (hv : v ~[l] w) : (u + v) ~[l] w := -add_comm u v ▸ hv.add_is_o hu +add_comm v u ▸ hv.add_is_o hu lemma is_o.is_equivalent (huv : (u - v) =o[l] v) : u ~[l] v := huv diff --git a/src/analysis/special_functions/compare_exp.lean b/src/analysis/special_functions/compare_exp.lean new file mode 100644 index 0000000000000..9f7ec46b1b750 --- /dev/null +++ b/src/analysis/special_functions/compare_exp.lean @@ -0,0 +1,182 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import analysis.special_functions.pow +import analysis.asymptotics.specific_asymptotics + +/-! +# Growth estimates on `x ^ y` for complex `x`, `y` + +Let `l` be a filter on `ℂ` such that `complex.re` tends to infinity along `l` and `complex.im z` +grows at a subexponential rate compared to `complex.re z`. Then + +- `complex.is_o_log_abs_re_of_subexponential_im_re`: `real.log ∘ complex.abs` is `o`-small of + `complex.re` along `l`; + +- `complex.is_o_cpow_mul_exp`: $z^{a_1}e^{b_1 * z} = o\left(z^{a_1}e^{b_1 * z}\right)$ along `l` + for any complex `a₁`, `a₂` and real `b₁ < b₂`. + +We use these assumptions on `l` for two reasons. First, these are the assumptions that naturally +appear in the proof. Second, in some applications (e.g., in Ilyashenko's proof of the individual +finiteness theorem for limit cycles of polynomial ODEs with hyperbolic singularities only) natural +stronger assumptions (e.g., `im z` is bounded from below and from above) are not available. + +-/ + +open asymptotics filter function +open_locale topological_space + +namespace complex + +/-- We say that `l : filter ℂ` is an *exponential comparison filter* if the real part tends to +infinity along `l` and the imaginary part grows subexponentially compared to the real part. These +properties guarantee that `(λ z, z ^ a₁ * exp (b₁ * z)) =o[l] (λ z, z ^ a₂ * exp (b₂ * z))` for any +complex `a₁`, `a₂` and real `b₁ < b₂`. + +In particular, the second property is automatically satisfied if the imaginary part is bounded along +`l`. -/ +structure is_exp_cmp_filter (l : filter ℂ) : Prop := +(tendsto_re : tendsto re l at_top) +(is_O_im_pow_re : ∀ n : ℕ, (λ z : ℂ, z.im ^ n) =O[l] (λ z, real.exp z.re)) + +namespace is_exp_cmp_filter + +variables {l : filter ℂ} + +/-! +### Alternative constructors +-/ + +lemma of_is_O_im_re_rpow (hre : tendsto re l at_top) (r : ℝ) (hr : im =O[l] (λ z, z.re ^ r)) : + is_exp_cmp_filter l := +⟨hre, λ n, is_o.is_O $ + calc (λ z : ℂ, z.im ^ n) =O[l] (λ z, (z.re ^ r) ^ n) : hr.pow n + ... =ᶠ[l] (λ z, z.re ^ (r * n)) : (hre.eventually_ge_at_top 0).mono $ + λ z hz, by simp only [real.rpow_mul hz r n, real.rpow_nat_cast] + ... =o[l] (λ z, real.exp z.re) : (is_o_rpow_exp_at_top _).comp_tendsto hre⟩ + +lemma of_is_O_im_re_pow (hre : tendsto re l at_top) (n : ℕ) (hr : im =O[l] (λ z, z.re ^ n)) : + is_exp_cmp_filter l := +of_is_O_im_re_rpow hre n $ by simpa only [real.rpow_nat_cast] + +lemma of_bounded_under_abs_im (hre : tendsto re l at_top) + (him : is_bounded_under (≤) l (λ z, |z.im|)) : + is_exp_cmp_filter l := +of_is_O_im_re_pow hre 0 $ + by simpa only [pow_zero] using @is_bounded_under.is_O_const ℂ ℝ ℝ _ _ _ l him 1 one_ne_zero + +lemma of_bounded_under_im (hre : tendsto re l at_top) (him_le : is_bounded_under (≤) l im) + (him_ge : is_bounded_under (≥) l im) : + is_exp_cmp_filter l := +of_bounded_under_abs_im hre $ is_bounded_under_le_abs.2 ⟨him_le, him_ge⟩ + +/-! +### Preliminary lemmas +-/ + +lemma eventually_ne (hl : is_exp_cmp_filter l) : ∀ᶠ w : ℂ in l, w ≠ 0 := +hl.tendsto_re.eventually_ne_at_top' _ + +lemma tendsto_abs_re (hl : is_exp_cmp_filter l) : tendsto (λ z : ℂ, |z.re|) l at_top := +tendsto_abs_at_top_at_top.comp hl.tendsto_re + +lemma tendsto_abs (hl : is_exp_cmp_filter l) : tendsto abs l at_top := +tendsto_at_top_mono abs_re_le_abs hl.tendsto_abs_re + +lemma is_o_log_re_re (hl : is_exp_cmp_filter l) : (λ z, real.log z.re) =o[l] re := +real.is_o_log_id_at_top.comp_tendsto hl.tendsto_re + +lemma is_o_im_pow_exp_re (hl : is_exp_cmp_filter l) (n : ℕ) : + (λ z : ℂ, z.im ^ n) =o[l] (λ z, real.exp z.re) := +flip is_o.of_pow two_ne_zero $ + calc (λ z : ℂ, (z.im ^ n) ^ 2) = (λ z, z.im ^ (2 * n)) : by simp only [pow_mul'] + ... =O[l] (λ z, real.exp z.re) : hl.is_O_im_pow_re _ + ... = (λ z, (real.exp z.re) ^ 1) : by simp only [pow_one] + ... =o[l] (λ z, (real.exp z.re) ^ 2) : (is_o_pow_pow_at_top_of_lt one_lt_two).comp_tendsto $ + real.tendsto_exp_at_top.comp hl.tendsto_re + +lemma abs_im_pow_eventually_le_exp_re (hl : is_exp_cmp_filter l) (n : ℕ) : + (λ z : ℂ, |z.im| ^ n) ≤ᶠ[l] (λ z, real.exp z.re) := +by simpa using (hl.is_o_im_pow_exp_re n).bound zero_lt_one + +/-- If `l : filter ℂ` is an "exponential comparison filter", then $\log |z| =o(ℜ z)$ along `l`. +This is the main lemma in the proof of `complex.is_exp_cmp_filter.is_o_cpow_exp` below. +-/ +lemma is_o_log_abs_re (hl : is_exp_cmp_filter l) : (λ z, real.log (abs z)) =o[l] re := +calc (λ z, real.log (abs z)) =O[l] (λ z, real.log (real.sqrt 2) + real.log (max z.re (|z.im|))) : + is_O.of_bound 1 $ (hl.tendsto_re.eventually_ge_at_top 1).mono $ λ z hz, + begin + have h2 : 0 < real.sqrt 2, by simp, + have hz' : 1 ≤ abs z, from hz.trans (re_le_abs z), + have hz₀ : 0 < abs z, from one_pos.trans_le hz', + have hm₀ : 0 < max z.re (|z.im|), from lt_max_iff.2 (or.inl $ one_pos.trans_le hz), + rw [one_mul, real.norm_eq_abs, _root_.abs_of_nonneg (real.log_nonneg hz')], + refine le_trans _ (le_abs_self _), + rw [← real.log_mul, real.log_le_log, ← _root_.abs_of_nonneg (le_trans zero_le_one hz)], + exacts [abs_le_sqrt_two_mul_max z, one_pos.trans_le hz', (mul_pos h2 hm₀), h2.ne', hm₀.ne'] + end +... =o[l] re : is_o.add (is_o_const_left.2 $ or.inr $ hl.tendsto_abs_re) $ is_o_iff_nat_mul_le.2 $ + λ n, begin + filter_upwards [is_o_iff_nat_mul_le.1 hl.is_o_log_re_re n, hl.abs_im_pow_eventually_le_exp_re n, + hl.tendsto_re.eventually_gt_at_top 1] with z hre him h₁, + cases le_total (|z.im|) z.re with hle hle, + { rwa [max_eq_left hle] }, + { have H : 1 < |z.im|, from h₁.trans_le hle, + rwa [max_eq_right hle, real.norm_eq_abs, real.norm_eq_abs, abs_of_pos (real.log_pos H), + ← real.log_pow, real.log_le_iff_le_exp (pow_pos (one_pos.trans H) _), + abs_of_pos (one_pos.trans h₁)] } + end + +/-! +### Main results +-/ + +/-- If `l : filter ℂ` is an "exponential comparison filter", then for any complex `a` and any +positive real `b`, we have `(λ z, z ^ a) =o[l] (λ z, exp (b * z))`. -/ +lemma is_o_cpow_exp (hl : is_exp_cmp_filter l) (a : ℂ) {b : ℝ} (hb : 0 < b) : + (λ z, z ^ a) =o[l] (λ z, exp (b * z)) := +calc (λ z, z ^ a) =Θ[l] λ z, abs z ^ re a : is_Theta_cpow_const_rpow $ λ _ _, hl.eventually_ne +... =ᶠ[l] λ z, real.exp (re a * real.log (abs z)) : hl.eventually_ne.mono $ + λ z hz, by simp only [real.rpow_def_of_pos (abs_pos.2 hz), mul_comm] +... =o[l] λ z, exp (b * z) : is_o.of_norm_right $ + begin + simp only [norm_eq_abs, abs_exp, of_real_mul_re, real.is_o_exp_comp_exp_comp], + refine (is_equivalent.refl.sub_is_o _).symm.tendsto_at_top (hl.tendsto_re.const_mul_at_top hb), + exact (hl.is_o_log_abs_re.const_mul_left _).const_mul_right hb.ne' + end + +/-- If `l : filter ℂ` is an "exponential comparison filter", then for any complex `a₁`, `a₂` and any +real `b₁ < b₂`, we have `(λ z, z ^ a₁ * exp (b₁ * z)) =o[l] (λ z, z ^ a₂ * exp (b₂ * z))`. -/ +lemma is_o_cpow_mul_exp {b₁ b₂ : ℝ} (hl : is_exp_cmp_filter l) (hb : b₁ < b₂) (a₁ a₂ : ℂ) : + (λ z, z ^ a₁ * exp (b₁ * z)) =o[l] (λ z, z ^ a₂ * exp (b₂ * z)) := +calc (λ z, z ^ a₁ * exp (b₁ * z)) =ᶠ[l] (λ z, z ^ a₂ * exp (b₁ * z) * z ^ (a₁ - a₂)) : + hl.eventually_ne.mono $ λ z hz, + by { simp only, rw [mul_right_comm, ← cpow_add _ _ hz, add_sub_cancel'_right] } +... =o[l] λ z, z ^ a₂ * exp (b₁ * z) * exp (↑(b₂ - b₁) * z) : + (is_O_refl (λ z, z ^ a₂ * exp (b₁ * z)) l).mul_is_o $ hl.is_o_cpow_exp _ (sub_pos.2 hb) +... =ᶠ[l] λ z, z ^ a₂ * exp (b₂ * z) : + by simp only [of_real_sub, sub_mul, mul_assoc, ← exp_add, add_sub_cancel'_right] + +/-- If `l : filter ℂ` is an "exponential comparison filter", then for any complex `a` and any +negative real `b`, we have `(λ z, exp (b * z)) =o[l] (λ z, z ^ a)`. -/ +lemma is_o_exp_cpow (hl : is_exp_cmp_filter l) (a : ℂ) {b : ℝ} (hb : b < 0) : + (λ z, exp (b * z)) =o[l] (λ z, z ^ a) := +by simpa using hl.is_o_cpow_mul_exp hb 0 a + +/-- If `l : filter ℂ` is an "exponential comparison filter", then for any complex `a₁`, `a₂` and any +natural `b₁ < b₂`, we have `(λ z, z ^ a₁ * exp (b₁ * z)) =o[l] (λ z, z ^ a₂ * exp (b₂ * z))`. -/ +lemma is_o_pow_mul_exp {b₁ b₂ : ℝ} (hl : is_exp_cmp_filter l) (hb : b₁ < b₂) (m n : ℕ) : + (λ z, z ^ m * exp (b₁ * z)) =o[l] (λ z, z ^ n * exp (b₂ * z)) := +by simpa only [cpow_nat_cast] using hl.is_o_cpow_mul_exp hb m n + +/-- If `l : filter ℂ` is an "exponential comparison filter", then for any complex `a₁`, `a₂` and any +integer `b₁ < b₂`, we have `(λ z, z ^ a₁ * exp (b₁ * z)) =o[l] (λ z, z ^ a₂ * exp (b₂ * z))`. -/ +lemma is_o_zpow_mul_exp {b₁ b₂ : ℝ} (hl : is_exp_cmp_filter l) (hb : b₁ < b₂) (m n : ℤ) : + (λ z, z ^ m * exp (b₁ * z)) =o[l] (λ z, z ^ n * exp (b₂ * z)) := +by simpa only [cpow_int_cast] using hl.is_o_cpow_mul_exp hb m n + +end is_exp_cmp_filter + +end complex diff --git a/src/analysis/special_functions/complex/arg.lean b/src/analysis/special_functions/complex/arg.lean index c08f7fb3cc240..911713e79f06a 100644 --- a/src/analysis/special_functions/complex/arg.lean +++ b/src/analysis/special_functions/complex/arg.lean @@ -139,6 +139,8 @@ lemma arg_le_pi (x : ℂ) : arg x ≤ π := lemma neg_pi_lt_arg (x : ℂ) : -π < arg x := (arg_mem_Ioc x).1 +lemma abs_arg_le_pi (z : ℂ) : |arg z| ≤ π := abs_le.2 ⟨(neg_pi_lt_arg z).le, arg_le_pi z⟩ + @[simp] lemma arg_nonneg_iff {z : ℂ} : 0 ≤ arg z ↔ 0 ≤ z.im := begin rcases eq_or_ne z 0 with (rfl|h₀), { simp }, diff --git a/src/analysis/special_functions/polynomials.lean b/src/analysis/special_functions/polynomials.lean index 5255b6ed4b627..b9c190472cf71 100644 --- a/src/analysis/special_functions/polynomials.lean +++ b/src/analysis/special_functions/polynomials.lean @@ -39,12 +39,10 @@ lemma is_equivalent_at_top_lead : begin by_cases h : P = 0, { simp [h] }, - { conv_rhs - { funext, - rw [polynomial.eval_eq_sum_range, sum_range_succ] }, - exact is_equivalent.refl.add_is_o (is_o.sum $ λ i hi, is_o.const_mul_left + { simp only [polynomial.eval_eq_sum_range, sum_range_succ], + exact is_o.add_is_equivalent (is_o.sum $ λ i hi, is_o.const_mul_left (is_o.const_mul_right (λ hz, h $ leading_coeff_eq_zero.mp hz) $ - is_o_pow_pow_at_top_of_lt (mem_range.mp hi)) _) } + is_o_pow_pow_at_top_of_lt (mem_range.mp hi)) _) is_equivalent.refl } end lemma tendsto_at_top_of_leading_coeff_nonneg (hdeg : 0 < P.degree) (hnng : 0 ≤ P.leading_coeff) : diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 014d93414a158..4ee23bb0f116e 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -20,7 +20,7 @@ We also prove basic properties of these functions. noncomputable theory -open_locale classical real topological_space nnreal ennreal filter big_operators +open_locale classical real topological_space nnreal ennreal filter big_operators asymptotics open filter finset set namespace complex @@ -109,7 +109,7 @@ by simpa using cpow_neg x 1 by { rw ← cpow_nat_cast, simp only [nat.cast_bit0, nat.cast_one] } @[simp, norm_cast] lemma cpow_int_cast (x : ℂ) : ∀ (n : ℤ), x ^ (n : ℂ) = x ^ n -| (n : ℕ) := by simp; refl +| (n : ℕ) := by simp | -[1+ n] := by rw zpow_neg_succ_of_nat; simp only [int.neg_succ_of_nat_coe, int.cast_neg, complex.cpow_neg, inv_eq_one_div, int.cast_coe_nat, cpow_nat_cast] @@ -413,12 +413,64 @@ lemma abs_cpow_of_ne_zero {z : ℂ} (hz : z ≠ 0) (w : ℂ) : by rw [cpow_def_of_ne_zero hz, abs_exp, mul_re, log_re, log_im, real.exp_sub, real.rpow_def_of_pos (abs_pos.2 hz)] +lemma abs_cpow_of_imp {z w : ℂ} (h : z = 0 → w.re = 0 → w = 0) : + abs (z ^ w) = abs z ^ w.re / real.exp (arg z * im w) := +begin + rcases ne_or_eq z 0 with hz|rfl; [exact (abs_cpow_of_ne_zero hz w), rw abs_zero], + cases eq_or_ne w.re 0 with hw hw, + { simp [hw, h rfl hw] }, + { rw [real.zero_rpow hw, zero_div, zero_cpow, abs_zero], + exact ne_of_apply_ne re hw } +end + lemma abs_cpow_le (z w : ℂ) : abs (z ^ w) ≤ abs z ^ w.re / real.exp (arg z * im w) := begin - rcases ne_or_eq z 0 with hz|rfl; [exact (abs_cpow_of_ne_zero hz w).le, rw abs_zero], - rcases eq_or_ne w 0 with rfl|hw, { simp }, - rw [zero_cpow hw, abs_zero], - exact div_nonneg (real.rpow_nonneg_of_nonneg le_rfl _) (real.exp_pos _).le + by_cases h : z = 0 → w.re = 0 → w = 0, + { exact (abs_cpow_of_imp h).le }, + { push_neg at h, + rw [h.1, zero_cpow h.2.2, abs_zero], + exact div_nonneg (real.rpow_nonneg_of_nonneg le_rfl _) (real.exp_pos _).le }, +end + +section + +variables {α : Type*} {l : filter α} {f g : α → ℂ} + +open asymptotics + +lemma is_Theta_exp_arg_mul_im (hl : is_bounded_under (≤) l (λ x, |(g x).im|)) : + (λ x, real.exp (arg (f x) * im (g x))) =Θ[l] (λ x, (1 : ℝ)) := +begin + rcases hl with ⟨b, hb⟩, + refine real.is_Theta_exp_comp_one.2 ⟨π * b, _⟩, + rw eventually_map at hb ⊢, + refine hb.mono (λ x hx, _), + rw [_root_.abs_mul], + exact mul_le_mul (abs_arg_le_pi _) hx (_root_.abs_nonneg _) real.pi_pos.le +end + +lemma is_O_cpow_rpow (hl : is_bounded_under (≤) l (λ x, |(g x).im|)) : + (λ x, f x ^ g x) =O[l] (λ x, abs (f x) ^ (g x).re) := +calc (λ x, f x ^ g x) =O[l] (λ x, abs (f x) ^ (g x).re / real.exp (arg (f x) * im (g x))) : + is_O_of_le _ $ λ x, (abs_cpow_le _ _).trans (le_abs_self _) +... =Θ[l] (λ x, abs (f x) ^ (g x).re / (1 : ℝ)) : + (is_Theta_refl _ _).div (is_Theta_exp_arg_mul_im hl) +... =ᶠ[l] (λ x, abs (f x) ^ (g x).re) : by simp only [of_real_one, div_one] + +lemma is_Theta_cpow_rpow (hl_im : is_bounded_under (≤) l (λ x, |(g x).im|)) + (hl : ∀ᶠ x in l, f x = 0 → re (g x) = 0 → g x = 0): + (λ x, f x ^ g x) =Θ[l] (λ x, abs (f x) ^ (g x).re) := +calc (λ x, f x ^ g x) =Θ[l] (λ x, abs (f x) ^ (g x).re / real.exp (arg (f x) * im (g x))) : + is_Theta_of_norm_eventually_eq' $ hl.mono $ λ x, abs_cpow_of_imp +... =Θ[l] (λ x, abs (f x) ^ (g x).re / (1 : ℝ)) : + (is_Theta_refl _ _).div (is_Theta_exp_arg_mul_im hl_im) +... =ᶠ[l] (λ x, abs (f x) ^ (g x).re) : by simp only [of_real_one, div_one] + +lemma is_Theta_cpow_const_rpow {b : ℂ} (hl : b.re = 0 → b ≠ 0 → ∀ᶠ x in l, f x ≠ 0) : + (λ x, f x ^ b) =Θ[l] (λ x, abs (f x) ^ b.re) := +is_Theta_cpow_rpow is_bounded_under_const $ by simpa only [eventually_imp_distrib_right, ne.def, + ← not_frequently, not_imp_not, imp.swap] using hl + end @[simp] lemma abs_cpow_real (x : ℂ) (y : ℝ) : abs (x ^ (y : ℂ)) = x.abs ^ y := diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index ed7c0db7127a0..6932ee9c0856d 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -164,6 +164,10 @@ lemma tendsto.eventually_ne_at_top [preorder β] [no_max_order β] {f : α → (hf : tendsto f l at_top) (c : β) : ∀ᶠ x in l, f x ≠ c := hf.eventually (eventually_ne_at_top c) +lemma tendsto.eventually_ne_at_top' [preorder β] [no_max_order β] {f : α → β} {l : filter α} + (hf : tendsto f l at_top) (c : α) : ∀ᶠ x in l, x ≠ c := +(hf.eventually_ne_at_top (f c)).mono $ λ x, ne_of_apply_ne f + lemma eventually_lt_at_bot [preorder α] [no_min_order α] (a : α) : ∀ᶠ x in at_bot, x < a := Iio_mem_at_bot a diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index 73977d65726ed..e7b6f47ea369a 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -93,6 +93,9 @@ lemma is_bounded_under.mono_ge [preorder β] {l : filter α} {u v : α → β} (hu : is_bounded_under (≥) l u) (hv : u ≤ᶠ[l] v) : is_bounded_under (≥) l v := @is_bounded_under.mono_le α βᵒᵈ _ _ _ _ hu hv +lemma is_bounded_under_const [is_refl α r] {l : filter β} {a : α} : is_bounded_under r l (λ _, a) := +⟨a, eventually_map.2 $ eventually_of_forall $ λ _, refl _⟩ + lemma is_bounded.is_bounded_under {q : β → β → Prop} {u : α → β} (hf : ∀a₀ a₁, r a₀ a₁ → q (u a₀) (u a₁)) : f.is_bounded r → f.is_bounded_under q u | ⟨b, h⟩ := ⟨u b, show ∀ᶠ x in f, q (u x) (u b), from h.mono (λ x, hf x b)⟩ From 9019542b3af93db4076c0c9d7b1476c73b8efa3a Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 20 Sep 2022 01:29:13 +0000 Subject: [PATCH 324/828] =?UTF-8?q?feat(data/complex/module):=20get=20`smu?= =?UTF-8?q?l=5Fcomm=5Fclass=20=E2=84=9D=20M=20E`=20from=20the=20one=20for?= =?UTF-8?q?=20`=E2=84=82`=20(#16513)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This makes it so all the pieces are in places to realize a non-unital `ℂ`-algebra as a non-unital `ℝ`-algebra. Co-authored-by: Eric Wieser --- src/data/complex/module.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index b9246960022e5..fef643df9c16b 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -181,6 +181,13 @@ restrict_scalars.is_scalar_tower ℝ ℂ E (x : ℂ) • y = x • y := rfl +/-- The scalar action of `ℝ` on a `ℂ`-module `E` induced by `module.complex_to_real` commutes with +another scalar action of `M` on `E` whenever the action of `ℂ` commutes with the action of `M`. -/ +@[priority 900] +instance smul_comm_class.complex_to_real {M E : Type*} + [add_comm_group E] [module ℂ E] [has_smul M E] [smul_comm_class ℂ M E] : smul_comm_class ℝ M E := +{ smul_comm := λ r _ _, (smul_comm (r : ℂ) _ _ : _) } + @[priority 100] instance finite_dimensional.complex_to_real (E : Type*) [add_comm_group E] [module ℂ E] [finite_dimensional ℂ E] : finite_dimensional ℝ E := From 0a67ddd43ae4bd413e27535f34af1969ba633929 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 20 Sep 2022 05:26:47 +0000 Subject: [PATCH 325/828] feat(topology/separation): define `regular_space` (#16360) --- .../locally_convex/balanced_core_hull.lean | 2 +- src/order/filter/bases.lean | 9 + src/topology/algebra/group.lean | 35 ++- src/topology/algebra/infinite_sum.lean | 30 +- src/topology/algebra/order/basic.lean | 10 +- src/topology/algebra/order/extend_from.lean | 6 +- src/topology/algebra/with_zero_topology.lean | 16 +- src/topology/dense_embedding.lean | 4 +- src/topology/extend_from.lean | 4 +- src/topology/inseparable.lean | 5 + src/topology/order.lean | 7 + src/topology/separation.lean | 282 ++++++++++++------ src/topology/uniform_space/separation.lean | 32 +- 13 files changed, 267 insertions(+), 175 deletions(-) diff --git a/src/analysis/locally_convex/balanced_core_hull.lean b/src/analysis/locally_convex/balanced_core_hull.lean index 65e6cf2b31a42..bb3689336ffd1 100644 --- a/src/analysis/locally_convex/balanced_core_hull.lean +++ b/src/analysis/locally_convex/balanced_core_hull.lean @@ -245,7 +245,7 @@ filter.has_basis_self.mpr (λ s hs, ⟨balanced_core 𝕜 s, balanced_core_mem_nhds_zero hs, balanced_core_balanced s, balanced_core_subset s⟩) -lemma nhds_basis_closed_balanced [t3_space E] : (𝓝 (0 : E)).has_basis +lemma nhds_basis_closed_balanced [regular_space E] : (𝓝 (0 : E)).has_basis (λ (s : set E), s ∈ 𝓝 (0 : E) ∧ is_closed s ∧ balanced 𝕜 s) id := begin refine (closed_nhds_basis 0).to_has_basis (λ s hs, _) (λ s hs, ⟨s, ⟨hs.1, hs.2.1⟩, rfl.subset⟩), diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index 6f87ab1839f77..b8042969d46eb 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -606,6 +606,15 @@ by simp only [← principal_singleton, disjoint_principal_principal, disjoint_si (diagonal α)ᶜ ∈ l₁ ×ᶠ l₂ ↔ disjoint l₁ l₂ := by simp only [mem_prod_iff, filter.disjoint_iff, prod_subset_compl_diagonal_iff_disjoint] +lemma has_basis.disjoint_iff_left (h : l.has_basis p s) : + disjoint l l' ↔ ∃ i (hi : p i), (s i)ᶜ ∈ l' := +by simp only [h.disjoint_iff l'.basis_sets, exists_prop, id, ← disjoint_principal_left, + (has_basis_principal _).disjoint_iff l'.basis_sets, unique.exists_iff] + +lemma has_basis.disjoint_iff_right (h : l.has_basis p s) : + disjoint l' l ↔ ∃ i (hi : p i), (s i)ᶜ ∈ l' := +disjoint.comm.trans h.disjoint_iff_left + lemma le_iff_forall_inf_principal_compl {f g : filter α} : f ≤ g ↔ ∀ V ∈ g, f ⊓ 𝓟 Vᶜ = ⊥ := forall₂_congr $ λ _ _, mem_iff_inf_principal_compl diff --git a/src/topology/algebra/group.lean b/src/topology/algebra/group.lean index ef95bcf08edb2..42c9f6b5f9e42 100644 --- a/src/topology/algebra/group.lean +++ b/src/topology/algebra/group.lean @@ -996,28 +996,27 @@ variables (G) [topological_space G] [group G] [topological_group G] lemma topological_group.t1_space (h : @is_closed G _ {1}) : t1_space G := ⟨assume x, by { convert is_closed_map_mul_right x _ h, simp }⟩ +@[priority 100, to_additive] +instance topological_group.regular_space : regular_space G := +begin + refine regular_space.of_exists_mem_nhds_is_closed_subset (λ a s hs, _), + have : tendsto (λ p : G × G, p.1 * p.2) (𝓝 (a, 1)) (𝓝 a), + from continuous_mul.tendsto' _ _ (mul_one a), + rcases mem_nhds_prod_iff.mp (this hs) with ⟨U, hU, V, hV, hUV⟩, + rw [← image_subset_iff, image_prod] at hUV, + refine ⟨closure U, mem_of_superset hU subset_closure, is_closed_closure, _⟩, + calc closure U ⊆ closure U * interior V : subset_mul_left _ (mem_interior_iff_mem_nhds.2 hV) + ... = U * interior V : is_open_interior.closure_mul U + ... ⊆ U * V : mul_subset_mul_left interior_subset + ... ⊆ s : hUV +end + @[to_additive] -lemma topological_group.t3_space [t1_space G] : t3_space G := -⟨assume s a hs ha, - let f := λ p : G × G, p.1 * (p.2)⁻¹ in - have hf : continuous f := continuous_fst.mul continuous_snd.inv, - -- a ∈ -s implies f (a, 1) ∈ -s, and so (a, 1) ∈ f⁻¹' (-s); - -- and so can find t₁ t₂ open such that a ∈ t₁ × t₂ ⊆ f⁻¹' (-s) - let ⟨t₁, t₂, ht₁, ht₂, a_mem_t₁, one_mem_t₂, t_subset⟩ := - is_open_prod_iff.1 ((is_open_compl_iff.2 hs).preimage hf) a (1:G) (by simpa [f]) in - begin - use [s * t₂, ht₂.mul_left, λ x hx, ⟨x, 1, hx, one_mem_t₂, mul_one _⟩], - rw [nhds_within, inf_principal_eq_bot, mem_nhds_iff], - refine ⟨t₁, _, ht₁, a_mem_t₁⟩, - rintros x hx ⟨y, z, hy, hz, yz⟩, - have : x * z⁻¹ ∈ sᶜ := (prod_subset_iff.1 t_subset) x hx z hz, - have : x * z⁻¹ ∈ s, rw ← yz, simpa, - contradiction - end⟩ +lemma topological_group.t3_space [t1_space G] : t3_space G := ⟨⟩ @[to_additive] lemma topological_group.t2_space [t1_space G] : t2_space G := -@t3_space.t2_space G _ (topological_group.t3_space G) +by { haveI := topological_group.t3_space G, apply_instance } variables {G} (S : subgroup G) [subgroup.normal S] [is_closed (S : set G)] diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 635e1ba677e82..118bca4c834b9 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -369,7 +369,7 @@ lemma summable.even_add_odd {f : ℕ → α} (he : summable (λ k, f (2 * k))) summable f := (he.has_sum.even_add_odd ho.has_sum).summable -lemma has_sum.sigma [t3_space α] {γ : β → Type*} {f : (Σ b:β, γ b) → α} {g : β → α} {a : α} +lemma has_sum.sigma [regular_space α] {γ : β → Type*} {f : (Σ b:β, γ b) → α} {g : β → α} {a : α} (ha : has_sum f a) (hf : ∀b, has_sum (λc, f ⟨b, c⟩) (g b)) : has_sum g a := begin refine (at_top_basis.tendsto_iff (closed_nhds_basis a)).mpr _, @@ -390,12 +390,12 @@ end /-- If a series `f` on `β × γ` has sum `a` and for each `b` the restriction of `f` to `{b} × γ` has sum `g b`, then the series `g` has sum `a`. -/ -lemma has_sum.prod_fiberwise [t3_space α] {f : β × γ → α} {g : β → α} {a : α} +lemma has_sum.prod_fiberwise [regular_space α] {f : β × γ → α} {g : β → α} {a : α} (ha : has_sum f a) (hf : ∀b, has_sum (λc, f (b, c)) (g b)) : has_sum g a := has_sum.sigma ((equiv.sigma_equiv_prod β γ).has_sum_iff.2 ha) hf -lemma summable.sigma' [t3_space α] {γ : β → Type*} {f : (Σb:β, γ b) → α} +lemma summable.sigma' [regular_space α] {γ : β → Type*} {f : (Σb:β, γ b) → α} (ha : summable f) (hf : ∀b, summable (λc, f ⟨b, c⟩)) : summable (λb, ∑'c, f ⟨b, c⟩) := (ha.has_sum.sigma (assume b, (hf b).has_sum)).summable @@ -541,22 +541,22 @@ lemma tsum_sum {f : γ → β → α} {s : finset γ} (hf : ∀i∈s, summable ( ∑'b, ∑ i in s, f i b = ∑ i in s, ∑'b, f i b := (has_sum_sum $ assume i hi, (hf i hi).has_sum).tsum_eq -lemma tsum_sigma' [t3_space α] {γ : β → Type*} {f : (Σb:β, γ b) → α} - (h₁ : ∀b, summable (λc, f ⟨b, c⟩)) (h₂ : summable f) : ∑'p, f p = ∑'b c, f ⟨b, c⟩ := +variables [add_comm_monoid δ] [topological_space δ] [t3_space δ] [has_continuous_add δ] + +lemma tsum_sigma' {γ : β → Type*} {f : (Σb:β, γ b) → δ} (h₁ : ∀b, summable (λc, f ⟨b, c⟩)) + (h₂ : summable f) : ∑'p, f p = ∑'b c, f ⟨b, c⟩ := (h₂.has_sum.sigma (assume b, (h₁ b).has_sum)).tsum_eq.symm -lemma tsum_prod' [t3_space α] {f : β × γ → α} (h : summable f) - (h₁ : ∀b, summable (λc, f (b, c))) : +lemma tsum_prod' {f : β × γ → δ} (h : summable f) (h₁ : ∀b, summable (λc, f (b, c))) : ∑'p, f p = ∑'b c, f (b, c) := (h.has_sum.prod_fiberwise (assume b, (h₁ b).has_sum)).tsum_eq.symm -lemma tsum_comm' [t3_space α] {f : β → γ → α} (h : summable (function.uncurry f)) - (h₁ : ∀b, summable (f b)) (h₂ : ∀ c, summable (λ b, f b c)) : +lemma tsum_comm' {f : β → γ → δ} (h : summable (function.uncurry f)) (h₁ : ∀b, summable (f b)) + (h₂ : ∀ c, summable (λ b, f b c)) : ∑' c b, f b c = ∑' b c, f b c := begin - erw [← tsum_prod' h h₁, ← tsum_prod' h.prod_symm h₂, ← (equiv.prod_comm β γ).tsum_eq], - refl, - assumption + erw [← tsum_prod' h h₁, ← tsum_prod' h.prod_symm h₂, ← (equiv.prod_comm γ β).tsum_eq (uncurry f)], + refl end end has_continuous_add @@ -1228,12 +1228,12 @@ local attribute [instance] topological_add_group.t3_space /-- The sum over the complement of a finset tends to `0` when the finset grows to cover the whole space. This does not need a summability assumption, as otherwise all sums are zero. -/ -lemma tendsto_tsum_compl_at_top_zero [t1_space α] (f : β → α) : +lemma tendsto_tsum_compl_at_top_zero (f : β → α) : tendsto (λ (s : finset β), ∑' b : {x // x ∉ s}, f b) at_top (𝓝 0) := begin by_cases H : summable f, { assume e he, - rcases nhds_is_closed he with ⟨o, ho, oe, o_closed⟩, + rcases exists_mem_nhds_is_closed_subset he with ⟨o, ho, o_closed, oe⟩, simp only [le_eq_subset, set.mem_preimage, mem_at_top_sets, filter.mem_map, ge_iff_le], obtain ⟨s, hs⟩ : ∃ (s : finset β), ∀ (t : finset β), disjoint t s → ∑ (b : β) in t, f b ∈ o := cauchy_seq_finset_iff_vanishing.1 (tendsto.cauchy_seq H.has_sum) o ho, @@ -1302,7 +1302,7 @@ lemma summable.sigma_factor {γ : β → Type*} {f : (Σb:β, γ b) → α} (ha : summable f) (b : β) : summable (λc, f ⟨b, c⟩) := ha.comp_injective sigma_mk_injective -lemma summable.sigma [t1_space α] {γ : β → Type*} {f : (Σb:β, γ b) → α} +lemma summable.sigma {γ : β → Type*} {f : (Σb:β, γ b) → α} (ha : summable f) : summable (λb, ∑'c, f ⟨b, c⟩) := ha.sigma' (λ b, ha.sigma_factor b) diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index 442c1e88eba91..66bea3cc3f564 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -1122,13 +1122,9 @@ lemma dense_iff_exists_between [densely_ordered α] [nontrivial α] {s : set α} @[priority 100] -- see Note [lower instance priority] instance order_topology.t3_space : t3_space α := -begin - refine ⟨λ s a hs ha, _⟩, - have : sᶜ ∈ 𝓝 a, from hs.is_open_compl.mem_nhds ha, - rcases exists_Icc_mem_subset_of_mem_nhds this with ⟨b, c, -, hmem, hsub⟩, - refine ⟨(Icc b c)ᶜ, is_closed_Icc.is_open_compl, subset_compl_comm.2 hsub, _⟩, - rwa [nhds_within, inf_principal_eq_bot, compl_compl] -end +{ to_regular_space := regular_space.of_exists_mem_nhds_is_closed_subset $ + λ a s hs, let ⟨b, c, ha, hmem, hs⟩ := exists_Icc_mem_subset_of_mem_nhds hs + in ⟨Icc b c, hmem, is_closed_Icc, hs⟩ } /-- A set is a neighborhood of `a` if and only if it contains an interval `(l, u)` containing `a`, provided `a` is neither a bottom element nor a top element. -/ diff --git a/src/topology/algebra/order/extend_from.lean b/src/topology/algebra/order/extend_from.lean index 29a052da4e0eb..9578d72199b64 100644 --- a/src/topology/algebra/order/extend_from.lean +++ b/src/topology/algebra/order/extend_from.lean @@ -17,7 +17,7 @@ universes u v variables {α : Type u} {β : Type v} lemma continuous_on_Icc_extend_from_Ioo [topological_space α] [linear_order α] [densely_ordered α] - [order_topology α] [topological_space β] [t3_space β] {f : α → β} {a b : α} + [order_topology α] [topological_space β] [regular_space β] {f : α → β} {a b : α} {la lb : β} (hab : a ≠ b) (hf : continuous_on f (Ioo a b)) (ha : tendsto f (𝓝[>] a) (𝓝 la)) (hb : tendsto f (𝓝[<] b) (𝓝 lb)) : continuous_on (extend_from (Ioo a b) f) (Icc a b) := @@ -55,7 +55,7 @@ end lemma continuous_on_Ico_extend_from_Ioo [topological_space α] [linear_order α] [densely_ordered α] [order_topology α] [topological_space β] - [t3_space β] {f : α → β} {a b : α} {la : β} (hab : a < b) (hf : continuous_on f (Ioo a b)) + [regular_space β] {f : α → β} {a b : α} {la : β} (hab : a < b) (hf : continuous_on f (Ioo a b)) (ha : tendsto f (𝓝[>] a) (𝓝 la)) : continuous_on (extend_from (Ioo a b) f) (Ico a b) := begin @@ -70,7 +70,7 @@ end lemma continuous_on_Ioc_extend_from_Ioo [topological_space α] [linear_order α] [densely_ordered α] [order_topology α] [topological_space β] - [t3_space β] {f : α → β} {a b : α} {lb : β} (hab : a < b) (hf : continuous_on f (Ioo a b)) + [regular_space β] {f : α → β} {a b : α} {lb : β} (hab : a < b) (hf : continuous_on f (Ioo a b)) (hb : tendsto f (𝓝[<] b) (𝓝 lb)) : continuous_on (extend_from (Ioo a b) f) (Ioc a b) := begin diff --git a/src/topology/algebra/with_zero_topology.lean b/src/topology/algebra/with_zero_topology.lean index 10c417963430d..7a6bad1552bc5 100644 --- a/src/topology/algebra/with_zero_topology.lean +++ b/src/topology/algebra/with_zero_topology.lean @@ -147,14 +147,14 @@ instance order_closed_topology : order_closed_topology Γ₀ := /-- The topology on a linearly ordered group with zero element adjoined is T₃. -/ @[priority 100] instance t3_space : t3_space Γ₀ := -t3_space.of_lift'_closure $ λ γ, - begin - rcases ne_or_eq γ 0 with h₀|rfl, - { rw [nhds_of_ne_zero h₀, lift'_pure (monotone_closure Γ₀), closure_singleton, - principal_singleton] }, - { exact has_basis_nhds_zero.lift'_closure_eq_self - (λ x hx, is_closed_iff.2 $ or.inl $ zero_lt_iff.2 hx) }, - end +{ to_regular_space := regular_space.of_lift'_closure $ λ γ, + begin + rcases ne_or_eq γ 0 with h₀|rfl, + { rw [nhds_of_ne_zero h₀, lift'_pure (monotone_closure Γ₀), closure_singleton, + principal_singleton] }, + { exact has_basis_nhds_zero.lift'_closure_eq_self + (λ x hx, is_closed_iff.2 $ or.inl $ zero_lt_iff.2 hx) }, + end } /-- The topology on a linearly ordered group with zero element adjoined makes it a topological monoid. -/ diff --git a/src/topology/dense_embedding.lean b/src/topology/dense_embedding.lean index cb8bba18d0a98..1a4348cb76eea 100644 --- a/src/topology/dense_embedding.lean +++ b/src/topology/dense_embedding.lean @@ -348,13 +348,13 @@ begin rw filter.has_basis_iff at h ⊢, intros T, refine ⟨λ hT, _, λ hT, _⟩, - { obtain ⟨T', hT₁, hT₂, hT₃⟩ := nhds_is_closed hT, + { obtain ⟨T', hT₁, hT₂, hT₃⟩ := exists_mem_nhds_is_closed_subset hT, have hT₄ : f⁻¹' T' ∈ 𝓝 x, { rw hf.to_inducing.nhds_eq_comap x, exact ⟨T', hT₁, subset.rfl⟩, }, obtain ⟨i, hi, hi'⟩ := (h _).mp hT₄, exact ⟨i, hi, (closure_mono (image_subset f hi')).trans (subset.trans (closure_minimal - (image_subset_iff.mpr subset.rfl) hT₃) hT₂)⟩, }, + (image_subset_iff.mpr subset.rfl) hT₂) hT₃)⟩, }, { obtain ⟨i, hi, hi'⟩ := hT, suffices : closure (f '' s i) ∈ 𝓝 (f x), { filter_upwards [this] using hi', }, replace h := (h (s i)).mpr ⟨i, hi, subset.rfl⟩, diff --git a/src/topology/extend_from.lean b/src/topology/extend_from.lean index 10b4b2f69bb0a..474a27784ae76 100644 --- a/src/topology/extend_from.lean +++ b/src/topology/extend_from.lean @@ -54,7 +54,7 @@ lemma extend_from_extends [t2_space Y] {f : X → Y} {A : set X} (hf : continuou /-- If `f` is a function to a T₃ space `Y` which has a limit within `A` at any point of a set `B ⊆ closure A`, then `extend_from A f` is continuous on `B`. -/ -lemma continuous_on_extend_from [t3_space Y] {f : X → Y} {A B : set X} (hB : B ⊆ closure A) +lemma continuous_on_extend_from [regular_space Y] {f : X → Y} {A B : set X} (hB : B ⊆ closure A) (hf : ∀ x ∈ B, ∃ y, tendsto f (𝓝[A] x) (𝓝 y)) : continuous_on (extend_from A f) B := begin set φ := extend_from A f, @@ -79,7 +79,7 @@ end /-- If a function `f` to a T₃ space `Y` has a limit within a dense set `A` for any `x`, then `extend_from A f` is continuous. -/ -lemma continuous_extend_from [t3_space Y] {f : X → Y} {A : set X} (hA : dense A) +lemma continuous_extend_from [regular_space Y] {f : X → Y} {A : set X} (hA : dense A) (hf : ∀ x, ∃ y, tendsto f (𝓝[A] x) (𝓝 y)) : continuous (extend_from A f) := begin rw continuous_iff_continuous_on_univ, diff --git a/src/topology/inseparable.lean b/src/topology/inseparable.lean index 934f2346f436d..fb588e33b7e88 100644 --- a/src/topology/inseparable.lean +++ b/src/topology/inseparable.lean @@ -121,6 +121,11 @@ lemma specializes_iff_closure_subset : alias specializes_iff_closure_subset ↔ specializes.closure_subset _ +lemma filter.has_basis.specializes_iff {ι} {p : ι → Prop} {s : ι → set X} + (h : (𝓝 y).has_basis p s) : + x ⤳ y ↔ ∀ i, p i → x ∈ s i := +specializes_iff_pure.trans h.ge_iff + lemma specializes_rfl : x ⤳ x := le_rfl @[refl] lemma specializes_refl (x : X) : x ⤳ x := specializes_rfl diff --git a/src/topology/order.lean b/src/topology/order.lean index a8d5e6d96db3d..d6749a2c08290 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -226,6 +226,13 @@ instance : has_le (topological_space α) := protected lemma topological_space.le_def {α} {t s : topological_space α} : t ≤ s ↔ s.is_open ≤ t.is_open := iff.rfl +lemma is_open.mono {α} {t₁ t₂ : topological_space α} {s : set α} (hs : @is_open α t₂ s) + (h : t₁ ≤ t₂) : @is_open α t₁ s := h s hs + +lemma is_closed.mono {α} {t₁ t₂ : topological_space α} {s : set α} (hs : @is_closed α t₂ s) + (h : t₁ ≤ t₂) : @is_closed α t₁ s := +(@is_open_compl_iff α t₁ s).mp $ hs.is_open_compl.mono h + /-- The ordering on topologies on the type `α`. `t ≤ s` if every set open in `s` is also open in `t` (`t` is finer than `s`). -/ instance : partial_order (topological_space α) := diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 73ec638f59bb9..0efba82ace1d9 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -774,6 +774,16 @@ end @[simp] lemma disjoint_nhds_nhds [t2_space α] {x y : α} : disjoint (𝓝 x) (𝓝 y) ↔ x ≠ y := ⟨λ hd he, by simpa [he, nhds_ne_bot.ne] using hd, t2_space_iff_disjoint_nhds.mp ‹_› x y⟩ +lemma is_open_set_of_disjoint_nhds_nhds : + is_open {p : α × α | disjoint (𝓝 p.1) (𝓝 p.2)} := +begin + simp only [is_open_iff_mem_nhds, prod.forall, mem_set_of_eq], + intros x y h, + obtain ⟨U, hU, V, hV, hd⟩ := ((nhds_basis_opens x).disjoint_iff (nhds_basis_opens y)).mp h, + exact mem_nhds_prod_iff.mpr ⟨U, hU.2.mem_nhds hU.1, V, hV.2.mem_nhds hV.1, + λ ⟨x', y'⟩ ⟨hx', hy'⟩, disjoint_of_disjoint_of_mem hd (hU.2.mem_nhds hx') (hV.2.mem_nhds hy')⟩ +end + /-- A finite set can be separated by open sets. -/ lemma t2_separation_finset [t2_space α] (s : finset α) : ∃ f : α → set α, set.pairwise_disjoint ↑s f ∧ ∀ x ∈ s, x ∈ f x ∧ is_open (f x) := @@ -1275,109 +1285,197 @@ by rw [is_irreducible, is_preirreducible_iff_subsingleton, end separation -section t3 - -/-- A T₃ space is a T₀ space in which for every closed `C` and `x ∉ C`, there exist - disjoint open sets containing `x` and `C` respectively. -/ -class t3_space (α : Type u) [topological_space α] extends t0_space α : Prop := -(regular : ∀{s:set α} {a}, is_closed s → a ∉ s → ∃t, is_open t ∧ s ⊆ t ∧ 𝓝[t] a = ⊥) - -lemma t3_space.of_lift'_closure [t0_space α] (h : ∀ x : α, (𝓝 x).lift' closure = 𝓝 x) : - t3_space α := +section regular_space + +/-- A topological space is called a *regular space* if for any closed set `s` and `a ∉ s`, there +exist disjoint open sets `U ⊇ s` and `V ∋ a`. We formulate this condition in terms of `disjoint`ness +of filters `𝓝ˢ s` and `𝓝 a`. -/ +@[mk_iff] class regular_space (X : Type u) [topological_space X] : Prop := +(regular : ∀ {s : set X} {a}, is_closed s → a ∉ s → disjoint (𝓝ˢ s) (𝓝 a)) + +lemma regular_space_tfae (X : Type u) [topological_space X] : + tfae [regular_space X, + ∀ (s : set X) (a ∉ closure s), disjoint (𝓝ˢ s) (𝓝 a), + ∀ (a : X) (s : set X), disjoint (𝓝ˢ s) (𝓝 a) ↔ a ∉ closure s, + ∀ (a : X) (s ∈ 𝓝 a), ∃ t ∈ 𝓝 a, is_closed t ∧ t ⊆ s, + ∀ a : X, (𝓝 a).lift' closure ≤ 𝓝 a, + ∀ a : X, (𝓝 a).lift' closure = 𝓝 a] := begin - refine ⟨λ s a hs ha, _⟩, - have : sᶜ ∈ (𝓝 a).lift' closure, - { rw [h], exact hs.is_open_compl.mem_nhds ha }, - rcases (𝓝 a).basis_sets.lift'_closure.mem_iff.mp this with ⟨U, haU, hU⟩, - refine ⟨(closure U)ᶜ, is_closed_closure.is_open_compl, subset_compl_comm.mp hU, not_not.mp _⟩, - rw [← ne, ← ne_bot_iff, ← mem_closure_iff_nhds_within_ne_bot, closure_compl, mem_compl_iff, - not_not, mem_interior_iff_mem_nhds], - exact mem_of_superset haU subset_closure + tfae_have : 1 ↔ 5, + { rw [regular_space_iff, (@compl_surjective (set X) _).forall, forall_swap], + simp only [is_closed_compl_iff, mem_compl_iff, not_not, @and_comm (_ ∈ _), + (nhds_basis_opens _).lift'_closure.le_basis_iff (nhds_basis_opens _), and_imp, + (nhds_basis_opens _).disjoint_iff_right, exists_prop, ← subset_interior_iff_mem_nhds_set, + interior_compl, compl_subset_compl] }, + tfae_have : 5 → 6, from λ h a, (h a).antisymm (𝓝 _).le_lift'_closure, + tfae_have : 6 → 4, + { intros H a s hs, + rw [← H] at hs, + rcases (𝓝 a).basis_sets.lift'_closure.mem_iff.mp hs with ⟨U, hU, hUs⟩, + exact ⟨closure U, mem_of_superset hU subset_closure, is_closed_closure, hUs⟩ }, + tfae_have : 4 → 2, + { intros H s a ha, + have ha' : sᶜ ∈ 𝓝 a, by rwa [← mem_interior_iff_mem_nhds, interior_compl], + rcases H _ _ ha' with ⟨U, hU, hUc, hUs⟩, + refine disjoint_of_disjoint_of_mem disjoint_compl_left _ hU, + rwa [← subset_interior_iff_mem_nhds_set, hUc.is_open_compl.interior_eq, subset_compl_comm] }, + tfae_have : 2 → 3, + { refine λ H a s, ⟨λ hd has, mem_closure_iff_nhds_ne_bot.mp has _, H s a⟩, + exact (hd.symm.mono_right $ @principal_le_nhds_set _ _ s).eq_bot }, + tfae_have : 3 → 1, from λ H, ⟨λ s a hs ha, (H _ _).mpr $ hs.closure_eq.symm ▸ ha⟩, + tfae_finish end -@[priority 100] -- see Note [lower instance priority] -instance t3_space.t1_space [t3_space α] : t1_space α := -begin - rw t1_space_iff_exists_open, - intros x y hxy, - obtain ⟨U, hU, h⟩ := exists_is_open_xor_mem hxy, - cases h, - { exact ⟨U, hU, h⟩ }, - { obtain ⟨R, hR, hh⟩ := t3_space.regular (is_closed_compl_iff.mpr hU) (not_not.mpr h.1), - obtain ⟨V, hV, hhh⟩ := mem_nhds_iff.1 (filter.inf_principal_eq_bot.1 hh.2), - exact ⟨R, hR, hh.1 (mem_compl h.2), hV hhh.2⟩ } -end +lemma regular_space.of_lift'_closure (h : ∀ a : α, (𝓝 a).lift' closure = 𝓝 a) : regular_space α := +iff.mpr ((regular_space_tfae α).out 0 5) h + +lemma regular_space.of_basis {ι : α → Sort*} {p : Π a, ι a → Prop} {s : Π a, ι a → set α} + (h₁ : ∀ a, (𝓝 a).has_basis (p a) (s a)) (h₂ : ∀ a i, p a i → is_closed (s a i)) : + regular_space α := +regular_space.of_lift'_closure $ λ a, (h₁ a).lift'_closure_eq_self (h₂ a) + +lemma regular_space.of_exists_mem_nhds_is_closed_subset + (h : ∀ (a : α) (s ∈ 𝓝 a), ∃ t ∈ 𝓝 a, is_closed t ∧ t ⊆ s) : regular_space α := +iff.mpr ((regular_space_tfae α).out 0 3) h + +variables [regular_space α] {a : α} {s : set α} + +lemma disjoint_nhds_set_nhds : disjoint (𝓝ˢ s) (𝓝 a) ↔ a ∉ closure s := +iff.mp ((regular_space_tfae α).out 0 2) ‹_› _ _ -lemma nhds_is_closed [t3_space α] {a : α} {s : set α} (h : s ∈ 𝓝 a) : - ∃ t ∈ 𝓝 a, t ⊆ s ∧ is_closed t := -let ⟨s', h₁, h₂, h₃⟩ := mem_nhds_iff.mp h in -have ∃t, is_open t ∧ s'ᶜ ⊆ t ∧ 𝓝[t] a = ⊥, - from t3_space.regular h₂.is_closed_compl (not_not_intro h₃), -let ⟨t, ht₁, ht₂, ht₃⟩ := this in -⟨tᶜ, - mem_of_eq_bot $ by rwa [compl_compl], - subset.trans (compl_subset_comm.1 ht₂) h₁, - is_closed_compl_iff.mpr ht₁⟩ - -lemma closed_nhds_basis [t3_space α] (a : α) : - (𝓝 a).has_basis (λ s : set α, s ∈ 𝓝 a ∧ is_closed s) id := -⟨λ t, ⟨λ t_in, let ⟨s, s_in, h_st, h⟩ := nhds_is_closed t_in in ⟨s, ⟨s_in, h⟩, h_st⟩, - λ ⟨s, ⟨s_in, hs⟩, hst⟩, mem_of_superset s_in hst⟩⟩ - -lemma topological_space.is_topological_basis.exists_closure_subset [t3_space α] +lemma disjoint_nhds_nhds_set : disjoint (𝓝 a) (𝓝ˢ s) ↔ a ∉ closure s := +disjoint.comm.trans disjoint_nhds_set_nhds + +lemma exists_mem_nhds_is_closed_subset {a : α} {s : set α} (h : s ∈ 𝓝 a) : + ∃ t ∈ 𝓝 a, is_closed t ∧ t ⊆ s := +iff.mp ((regular_space_tfae α).out 0 3) ‹_› _ _ h + +lemma closed_nhds_basis (a : α) : (𝓝 a).has_basis (λ s : set α, s ∈ 𝓝 a ∧ is_closed s) id := +has_basis_self.2 (λ _, exists_mem_nhds_is_closed_subset) + +lemma lift'_nhds_closure (a : α) : (𝓝 a).lift' closure = 𝓝 a := +(closed_nhds_basis a).lift'_closure_eq_self (λ s hs, hs.2) + +lemma filter.has_basis.nhds_closure {ι : Sort*} {a : α} {p : ι → Prop} {s : ι → set α} + (h : (𝓝 a).has_basis p s) : (𝓝 a).has_basis p (λ i, closure (s i)) := +lift'_nhds_closure a ▸ h.lift'_closure + +lemma has_basis_nhds_closure (a : α) : (𝓝 a).has_basis (λ s, s ∈ 𝓝 a) closure := +(𝓝 a).basis_sets.nhds_closure + +lemma has_basis_opens_closure (a : α) : (𝓝 a).has_basis (λ s, a ∈ s ∧ is_open s) closure := +(nhds_basis_opens a).nhds_closure + +lemma topological_space.is_topological_basis.nhds_basis_closure + {B : set (set α)} (hB : topological_space.is_topological_basis B) (a : α) : + (𝓝 a).has_basis (λ s : set α, a ∈ s ∧ s ∈ B) closure := +by simpa only [and_comm] using hB.nhds_has_basis.nhds_closure + +lemma topological_space.is_topological_basis.exists_closure_subset {B : set (set α)} (hB : topological_space.is_topological_basis B) {a : α} {s : set α} (h : s ∈ 𝓝 a) : ∃ t ∈ B, a ∈ t ∧ closure t ⊆ s := +by simpa only [exists_prop, and.assoc] using hB.nhds_has_basis.nhds_closure.mem_iff.mp h + +lemma disjoint_nhds_nhds_iff_not_specializes {a b : α} : + disjoint (𝓝 a) (𝓝 b) ↔ ¬a ⤳ b := +by rw [← nhds_set_singleton, disjoint_nhds_set_nhds, specializes_iff_mem_closure] + +lemma specializes_comm {a b : α} : a ⤳ b ↔ b ⤳ a := +by simp only [← disjoint_nhds_nhds_iff_not_specializes.not_left, disjoint.comm] + +alias specializes_comm ↔ specializes.symm _ + +lemma specializes_iff_inseparable {a b : α} : a ⤳ b ↔ inseparable a b := +⟨λ h, h.antisymm h.symm, le_of_eq⟩ + +lemma is_closed_set_of_specializes : is_closed {p : α × α | p.1 ⤳ p.2} := +by simp only [← is_open_compl_iff, compl_set_of, ← disjoint_nhds_nhds_iff_not_specializes, + is_open_set_of_disjoint_nhds_nhds] + +lemma is_closed_set_of_inseparable : is_closed {p : α × α | inseparable p.1 p.2} := +by simp only [← specializes_iff_inseparable, is_closed_set_of_specializes] + +protected lemma inducing.regular_space [topological_space β] {f : β → α} (hf : inducing f) : + regular_space β := +regular_space.of_basis (λ b, by { rw [hf.nhds_eq_comap b], exact (closed_nhds_basis _).comap _ }) $ + λ b s hs, hs.2.preimage hf.continuous + +lemma regular_space_induced (f : β → α) : @regular_space β (induced f ‹_›) := +by { letI := induced f ‹_›, exact inducing.regular_space ⟨rfl⟩ } + +lemma regular_space_Inf {X} {T : set (topological_space X)} (h : ∀ t ∈ T, @regular_space X t) : + @regular_space X (Inf T) := begin - rcases nhds_is_closed h with ⟨t, hat, hts, htc⟩, - rcases hB.mem_nhds_iff.1 hat with ⟨u, huB, hau, hut⟩, - exact ⟨u, huB, hau, (closure_minimal hut htc).trans hts⟩ + letI := Inf T, + have : ∀ a, (𝓝 a).has_basis + (λ If : Σ I : set T, I → set X, + If.1.finite ∧ ∀ i : If.1, If.2 i ∈ @nhds X i a ∧ @is_closed X i (If.2 i)) + (λ If, ⋂ i : If.1, If.snd i), + { intro a, + rw [nhds_Inf, ← infi_subtype''], + exact has_basis_infi (λ t : T, @closed_nhds_basis X t (h t t.2) a) }, + refine regular_space.of_basis this (λ a If hIf, is_closed_Inter $ λ i, _), + exact (hIf.2 i).2.mono (Inf_le (i : T).2) end -lemma topological_space.is_topological_basis.nhds_basis_closure [t3_space α] - {B : set (set α)} (hB : topological_space.is_topological_basis B) (a : α) : - (𝓝 a).has_basis (λ s : set α, a ∈ s ∧ s ∈ B) closure := -⟨λ s, ⟨λ h, let ⟨t, htB, hat, hts⟩ := hB.exists_closure_subset h in ⟨t, ⟨hat, htB⟩, hts⟩, - λ ⟨t, ⟨hat, htB⟩, hts⟩, mem_of_superset (hB.mem_nhds htB hat) (subset_closure.trans hts)⟩⟩ +lemma regular_space_infi {ι X} {t : ι → topological_space X} (h : ∀ i, @regular_space X (t i)) : + @regular_space X (infi t) := +regular_space_Inf $ forall_range_iff.mpr h + +lemma regular_space.inf {X} {t₁ t₂ : topological_space X} (h₁ : @regular_space X t₁) + (h₂ : @regular_space X t₂) : @regular_space X (t₁ ⊓ t₂) := +by { rw [inf_eq_infi], exact regular_space_infi (bool.forall_bool.2 ⟨h₂, h₁⟩) } + +instance {p : α → Prop} : regular_space (subtype p) := +embedding_subtype_coe.to_inducing.regular_space + +instance [topological_space β] [regular_space β] : regular_space (α × β) := +(regular_space_induced prod.fst).inf (regular_space_induced prod.snd) + +instance {ι : Type*} {π : ι → Type*} [Π i, topological_space (π i)] [∀ i, regular_space (π i)] : + regular_space (Π i, π i) := +regular_space_infi $ λ i, regular_space_induced _ + +end regular_space + +section t3 + +/-- A T₃ space is a T₀ space which is a regular space. Any T₃ space is a T₁ space, a T₂ space, and +a T₂.₅ space. -/ +class t3_space (α : Type u) [topological_space α] extends t0_space α, regular_space α : Prop + +@[priority 100] -- see Note [lower instance priority] +instance t3_space.t2_5_space [t3_space α] : t2_5_space α := +begin + haveI : t2_space α, + { refine t2_space_iff_disjoint_nhds.mpr (λ x y hne, _), + have : x ∉ closure {y} ∨ y ∉ closure {x}, + from (t0_space_iff_or_not_mem_closure α).mp infer_instance x y hne, + wlog H : x ∉ closure {y} := this using [x y, y x] tactic.skip, + { rwa [← disjoint_nhds_nhds_set, nhds_set_singleton] at H }, + { exact λ h, (this h.symm).symm } }, + -- TODO: reformulate `t2_5_space` in terms of `(𝓝 x).lift' closure` + refine ⟨λ x y hne, _⟩, + rcases ((closed_nhds_basis x).disjoint_iff (closed_nhds_basis y)).1 + (disjoint_nhds_nhds.mpr hne) with ⟨U, ⟨hxU, hUc⟩, V, ⟨hyV, hVc⟩, hd⟩, + exact ⟨interior U, interior V, is_open_interior, is_open_interior, + hd.mono (closure_minimal interior_subset hUc) (closure_minimal interior_subset hVc), + mem_interior_iff_mem_nhds.2 hxU, mem_interior_iff_mem_nhds.2 hyV⟩ +end protected lemma embedding.t3_space [topological_space β] [t3_space β] {f : α → β} (hf : embedding f) : t3_space α := { to_t0_space := hf.t0_space, - regular := - begin - intros s a hs ha, - rcases hf.to_inducing.is_closed_iff.1 hs with ⟨s, hs', rfl⟩, - rcases t3_space.regular hs' ha with ⟨t, ht, hst, hat⟩, - refine ⟨f ⁻¹' t, ht.preimage hf.continuous, preimage_mono hst, _⟩, - rw [nhds_within, hf.to_inducing.nhds_eq_comap, ← comap_principal, ← comap_inf, - ← nhds_within, hat, comap_bot] - end } + to_regular_space := hf.to_inducing.regular_space } instance subtype.t3_space [t3_space α] {p : α → Prop} : t3_space (subtype p) := embedding_subtype_coe.t3_space -variable (α) -@[priority 100] -- see Note [lower instance priority] -instance t3_space.t2_space [t3_space α] : t2_space α := -⟨λ x y hxy, -let ⟨s, hs, hys, hxs⟩ := t3_space.regular is_closed_singleton - (mt mem_singleton_iff.1 hxy), - ⟨t, hxt, u, hsu, htu⟩ := empty_mem_iff_bot.2 hxs, - ⟨v, hvt, hv, hxv⟩ := mem_nhds_iff.1 hxt in -⟨v, s, hv, hs, hxv, singleton_subset_iff.1 hys, - (disjoint_iff_inter_eq_empty.2 htu.symm).mono hvt hsu⟩⟩ +instance [topological_space β] [t3_space α] [t3_space β] : t3_space (α × β) := ⟨⟩ -@[priority 100] -- see Note [lower instance priority] -instance t3_space.t2_5_space [t3_space α] : t2_5_space α := -⟨λ x y hxy, -let ⟨U, V, hU, hV, hh_1, hh_2, hUV⟩ := t2_separation hxy, - hxcV := not_not.mpr (interior_maximal hUV.subset_compl_right hU hh_1), - ⟨R, hR, hh⟩ := t3_space.regular is_closed_closure (by rwa closure_eq_compl_interior_compl), - ⟨A, hA, hhh⟩ := mem_nhds_iff.1 (filter.inf_principal_eq_bot.1 hh.2) in -⟨A, V, hhh.1, hV, disjoint_compl_left.mono_left ((closure_minimal hA hR.is_closed_compl).trans $ - compl_subset_compl.mpr hh.1), hhh.2, hh_2⟩⟩ - -variable {α} +instance {ι : Type*} {π : ι → Type*} [Π i, topological_space (π i)] [Π i, t3_space (π i)] : + t3_space (Π i, π i) := ⟨⟩ /-- Given two points `x ≠ y`, we can find neighbourhoods `x ∈ V₁ ⊆ U₁` and `y ∈ V₂ ⊆ U₂`, with the `Vₖ` closed and the `Uₖ` open, such that the `Uₖ` are disjoint. -/ @@ -1386,11 +1484,10 @@ lemma disjoint_nested_nhds [t3_space α] {x y : α} (h : x ≠ y) : V₁ ⊆ U₁ ∧ V₂ ⊆ U₂ ∧ disjoint U₁ U₂ := begin rcases t2_separation h with ⟨U₁, U₂, U₁_op, U₂_op, x_in, y_in, H⟩, - rcases nhds_is_closed (is_open.mem_nhds U₁_op x_in) with ⟨V₁, V₁_in, h₁, V₁_closed⟩, - rcases nhds_is_closed (is_open.mem_nhds U₂_op y_in) with ⟨V₂, V₂_in, h₂, V₂_closed⟩, - use [U₁, mem_of_superset V₁_in h₁, V₁, V₁_in, - U₂, mem_of_superset V₂_in h₂, V₂, V₂_in], - tauto + rcases exists_mem_nhds_is_closed_subset (U₁_op.mem_nhds x_in) with ⟨V₁, V₁_in, V₁_closed, h₁⟩, + rcases exists_mem_nhds_is_closed_subset (U₂_op.mem_nhds y_in) with ⟨V₂, V₂_in, V₂_closed, h₂⟩, + exact ⟨U₁, mem_of_superset V₁_in h₁, V₁, V₁_in, U₂, mem_of_superset V₂_in h₂, V₂, V₂_in, + V₁_closed, V₂_closed, U₁_op, U₂_op, h₁, h₂, H⟩ end end t3 @@ -1424,11 +1521,8 @@ end @[priority 100] -- see Note [lower instance priority] instance normal_space.t3_space [normal_space α] : t3_space α := { regular := λ s x hs hxs, let ⟨u, v, hu, hv, hsu, hxv, huv⟩ := - normal_separation hs is_closed_singleton - (λ _ ⟨hx, hy⟩, hxs $ mem_of_eq_of_mem (eq_of_mem_singleton hy).symm hx) in - ⟨u, hu, hsu, filter.empty_mem_iff_bot.1 $ filter.mem_inf_iff.2 - ⟨v, is_open.mem_nhds hv (singleton_subset_iff.1 hxv), u, filter.mem_principal_self u, - by rwa [eq_comm, inter_comm, ← disjoint_iff_inter_eq_empty]⟩⟩ } + normal_separation hs is_closed_singleton (disjoint_singleton_right.mpr hxs) in + disjoint_of_disjoint_of_mem huv (hu.mem_nhds_set.2 hsu) (hv.mem_nhds $ hxv rfl) } -- We can't make this an instance because it could cause an instance loop. lemma normal_of_compact_t2 [compact_space α] [t2_space α] : normal_space α := diff --git a/src/topology/uniform_space/separation.lean b/src/topology/uniform_space/separation.lean index 627f8a10718c7..a9e4be4296b81 100644 --- a/src/topology/uniform_space/separation.lean +++ b/src/topology/uniform_space/separation.lean @@ -80,6 +80,12 @@ variables [uniform_space α] [uniform_space β] [uniform_space γ] ### Separated uniform spaces -/ +@[priority 100] +instance uniform_space.to_regular_space : regular_space α := +regular_space.of_basis + (λ a, by { rw [nhds_eq_comap_uniformity], exact uniformity_has_basis_closed.comap _ }) + (λ a V hV, hV.2.preimage $ continuous_const.prod_mk continuous_id) + /-- The separation relation is the intersection of all entourages. Two points which are related by the separation relation are "indistinguishable" according to the uniform structure. -/ @@ -190,31 +196,7 @@ end @[priority 100] -- see Note [lower instance priority] instance separated_t3 [separated_space α] : t3_space α := -{ to_t0_space := by { haveI := separated_iff_t2.mp ‹_›, exact t1_space.t0_space }, - regular := λs a hs ha, - have sᶜ ∈ 𝓝 a, - from is_open.mem_nhds hs.is_open_compl ha, - have {p : α × α | p.1 = a → p.2 ∈ sᶜ} ∈ 𝓤 α, - from mem_nhds_uniformity_iff_right.mp this, - let ⟨d, hd, h⟩ := comp_mem_uniformity_sets this in - let e := {y:α| (a, y) ∈ d} in - have hae : a ∈ closure e, from subset_closure $ refl_mem_uniformity hd, - have closure e ×ˢ closure e ⊆ comp_rel d (comp_rel (e ×ˢ e) d), - begin - rw [←closure_prod_eq, closure_eq_inter_uniformity], - change (⨅d' ∈ 𝓤 α, _) ≤ comp_rel d (comp_rel _ d), - exact (infi_le_of_le d $ infi_le_of_le hd $ le_rfl) - end, - have e_subset : closure e ⊆ sᶜ, - from assume a' ha', - let ⟨x, (hx : (a, x) ∈ d), y, ⟨hx₁, hx₂⟩, (hy : (y, _) ∈ d)⟩ := @this ⟨a, a'⟩ ⟨hae, ha'⟩ in - have (a, a') ∈ comp_rel d d, from ⟨y, hx₂, hy⟩, - h this rfl, - have closure e ∈ 𝓝 a, from (𝓝 a).sets_of_superset (mem_nhds_left a hd) subset_closure, - have 𝓝 a ⊓ 𝓟 (closure e)ᶜ = ⊥, - from (is_compl_principal (closure e)).inf_right_eq_bot_iff.2 (le_principal_iff.2 this), - ⟨(closure e)ᶜ, is_closed_closure.is_open_compl, assume x h₁ h₂, @e_subset x h₂ h₁, this⟩, - ..@t2_space.t1_space _ _ (separated_iff_t2.mp ‹_›) } +by { haveI := separated_iff_t2.mp ‹_›, exact ⟨⟩ } lemma is_closed_of_spaced_out [separated_space α] {V₀ : set (α × α)} (V₀_in : V₀ ∈ 𝓤 α) {s : set α} (hs : s.pairwise (λ x y, (x, y) ∉ V₀)) : is_closed s := From 560891c425c743b1a25d4f8447cce6dd60947c1a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 20 Sep 2022 07:09:33 +0000 Subject: [PATCH 326/828] =?UTF-8?q?refactor(data/real/basic):=20make=20`?= =?UTF-8?q?=E2=8A=93`=20and=20`=E2=8A=94`=20computable=20on=20reals=20(#16?= =?UTF-8?q?463)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This also makes `norm` and `dist` computable on `real`, `nnreal`, `rat`, and `int`. From a mathematical point of view that doesn't care about computability, this still contains the (mildly) interesting result that the maximum of two reals is equal to the real produced by taking the elementwise maximum of the two cauchy sequences. This change causes some unpleasant elaboration pain in `modular.lean` for reasons presumably linked to unification of dependent typeclasses. It's not clear to me precisely what the cause if. Co-authored-by: Yaël Dillies --- src/algebraic_topology/simplicial_set.lean | 2 +- src/analysis/normed/field/basic.lean | 11 ++- src/analysis/normed/group/basic.lean | 7 +- src/analysis/normed/group/seminorm.lean | 8 +- .../normed_space/lattice_ordered_group.lean | 3 +- src/analysis/normed_space/ordered.lean | 1 - src/analysis/seminorm.lean | 4 +- src/data/real/basic.lean | 76 +++++++++++++++---- src/data/real/nnreal.lean | 13 ++-- src/number_theory/modular.lean | 13 +++- src/topology/instances/rat.lean | 5 +- src/topology/metric_space/basic.lean | 8 +- src/topology/uniform_space/compare_reals.lean | 1 - 13 files changed, 106 insertions(+), 46 deletions(-) diff --git a/src/algebraic_topology/simplicial_set.lean b/src/algebraic_topology/simplicial_set.lean index 36c6280682286..9dd4523e29d54 100644 --- a/src/algebraic_topology/simplicial_set.lean +++ b/src/algebraic_topology/simplicial_set.lean @@ -123,7 +123,7 @@ instance {n} : inhabited (sSet.truncated n) := ⟨(sk n).obj $ Δ[0]⟩ end sSet /-- The functor associating the singular simplicial set to a topological space. -/ -noncomputable def Top.to_sSet : Top ⥤ sSet := +def Top.to_sSet : Top ⥤ sSet := colimit_adj.restricted_yoneda simplex_category.to_Top /-- The geometric realization functor. -/ diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 56be2fce5eb61..f3a458a9f9261 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -618,6 +618,11 @@ end densely end normed_field +instance : normed_comm_ring ℝ := +{ norm_mul := λ x y, (abs_mul x y).le, + .. real.normed_add_comm_group, + .. real.comm_ring } + noncomputable instance : normed_field ℝ := { norm_mul' := abs_mul, .. real.normed_add_comm_group } @@ -700,7 +705,7 @@ lemma normed_add_comm_group.tendsto_at_top' [nonempty α] [semilattice_sup α] [ tendsto f at_top (𝓝 b) ↔ ∀ ε, 0 < ε → ∃ N, ∀ n, N < n → ∥f n - b∥ < ε := (at_top_basis_Ioi.tendsto_iff metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) -noncomputable instance : normed_comm_ring ℤ := +instance : normed_comm_ring ℤ := { norm := λ n, ∥(n : ℝ)∥, norm_mul := λ m n, le_of_eq $ by simp only [norm, int.cast_mul, abs_mul], dist_eq := λ m n, by simp only [int.dist_eq, norm, int.cast_sub], @@ -726,12 +731,12 @@ end instance : norm_one_class ℤ := ⟨by simp [← int.norm_cast_real]⟩ -noncomputable instance : normed_field ℚ := +instance : normed_field ℚ := { norm := λ r, ∥(r : ℝ)∥, norm_mul' := λ r₁ r₂, by simp only [norm, rat.cast_mul, abs_mul], dist_eq := λ r₁ r₂, by simp only [rat.dist_eq, norm, rat.cast_sub] } -noncomputable instance : densely_normed_field ℚ := +instance : densely_normed_field ℚ := { lt_norm_lt := λ r₁ r₂ h₀ hr, let ⟨q, h⟩ := exists_rat_btwn hr in ⟨q, by { unfold norm, rwa abs_of_pos (h₀.trans_lt h.1) } ⟩ } diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 67af706a829f8..ed247f9f9280a 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -110,12 +110,13 @@ instance : normed_add_comm_group punit := @[simp] lemma punit.norm_eq_zero (r : punit) : ∥r∥ = 0 := rfl -noncomputable instance : normed_add_comm_group ℝ := -{ norm := λ x, |x|, - dist_eq := assume x y, rfl } +instance : has_norm ℝ := { norm := λ x, |x| } @[simp] lemma real.norm_eq_abs (r : ℝ) : ∥r∥ = |r| := rfl +instance : normed_add_comm_group ℝ := +{ dist_eq := assume x y, rfl } + section seminormed_add_comm_group variables [seminormed_add_comm_group E] [seminormed_add_comm_group F] [seminormed_add_comm_group G] diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index abf47e5688a05..e30db0570bde8 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -208,7 +208,7 @@ variables (p q) (f : F →* E) -- TODO: define `has_Sup` too, from the skeleton at -- https://github.com/leanprover-community/mathlib/pull/11329#issuecomment-1008915345 -@[to_additive] noncomputable instance : has_sup (group_seminorm E) := +@[to_additive] instance : has_sup (group_seminorm E) := ⟨λ p q, { to_fun := p ⊔ q, map_one' := @@ -221,7 +221,7 @@ variables (p q) (f : F →* E) @[simp, to_additive] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl @[simp, to_additive] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl -@[to_additive] noncomputable instance : semilattice_sup (group_seminorm E) := +@[to_additive] instance : semilattice_sup (group_seminorm E) := fun_like.coe_injective.semilattice_sup _ coe_sup /-- Composition of a group seminorm with a monoid homomorphism as a group seminorm. -/ @@ -434,7 +434,7 @@ variables (p q) (f : F →* E) @[simp, to_additive] lemma add_apply (x : E) : (p + q) x = p x + q x := rfl -- TODO: define `has_Sup` -@[to_additive] noncomputable instance : has_sup (group_norm E) := +@[to_additive] instance : has_sup (group_norm E) := ⟨λ p q, { eq_one_of_map_eq_zero' := λ x hx, of_not_not $ λ h, hx.not_gt $ lt_sup_iff.2 $ or.inl $ map_pos_of_ne_one p h, @@ -443,7 +443,7 @@ variables (p q) (f : F →* E) @[simp, to_additive] lemma coe_sup : ⇑(p ⊔ q) = p ⊔ q := rfl @[simp, to_additive] lemma sup_apply (x : E) : (p ⊔ q) x = p x ⊔ q x := rfl -@[to_additive] noncomputable instance : semilattice_sup (group_norm E) := +@[to_additive] instance : semilattice_sup (group_norm E) := fun_like.coe_injective.semilattice_sup _ coe_sup end group diff --git a/src/analysis/normed_space/lattice_ordered_group.lean b/src/analysis/normed_space/lattice_ordered_group.lean index ecce0fa34c0c1..594b54a8bf7de 100644 --- a/src/analysis/normed_space/lattice_ordered_group.lean +++ b/src/analysis/normed_space/lattice_ordered_group.lean @@ -49,9 +49,10 @@ class normed_lattice_add_comm_group (α : Type*) lemma solid {α : Type*} [normed_lattice_add_comm_group α] {a b : α} (h : |a| ≤ |b|) : ∥a∥ ≤ ∥b∥ := normed_lattice_add_comm_group.solid a b h -noncomputable instance : normed_lattice_add_comm_group ℝ := +instance : normed_lattice_add_comm_group ℝ := { add_le_add_left := λ _ _ h _, add_le_add le_rfl h, solid := λ _ _, id, } + /-- A normed lattice ordered group is an ordered additive commutative group -/ diff --git a/src/analysis/normed_space/ordered.lean b/src/analysis/normed_space/ordered.lean index 218891f82576e..ca6af5e73b2d0 100644 --- a/src/analysis/normed_space/ordered.lean +++ b/src/analysis/normed_space/ordered.lean @@ -41,7 +41,6 @@ extends linear_ordered_field α, has_norm α, metric_space α := [normed_linear_ordered_field α] : normed_linear_ordered_group α := ⟨normed_linear_ordered_field.dist_eq⟩ -noncomputable instance : normed_linear_ordered_field ℚ := ⟨dist_eq_norm, norm_mul⟩ diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 651205d77c805..94acc9b2d1c93 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -194,7 +194,7 @@ instance [semiring R] [module R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ -- TODO: define `has_Sup` too, from the skeleton at -- https://github.com/leanprover-community/mathlib/pull/11329#issuecomment-1008915345 -noncomputable instance : has_sup (seminorm 𝕜 E) := +instance : has_sup (seminorm 𝕜 E) := { sup := λ p q, { to_fun := p ⊔ q, smul' := λ x v, (congr_arg2 max (map_smul_eq_mul p x v) (map_smul_eq_mul q x v)).trans $ @@ -218,7 +218,7 @@ instance : partial_order (seminorm 𝕜 E) := lemma le_def (p q : seminorm 𝕜 E) : p ≤ q ↔ (p : E → ℝ) ≤ q := iff.rfl lemma lt_def (p q : seminorm 𝕜 E) : p < q ↔ (p : E → ℝ) < q := iff.rfl -noncomputable instance : semilattice_sup (seminorm 𝕜 E) := +instance : semilattice_sup (seminorm 𝕜 E) := function.injective.semilattice_sup _ fun_like.coe_injective coe_sup end has_smul diff --git a/src/data/real/basic.lean b/src/data/real/basic.lean index 113e0afb158df..bca0c89a201bf 100644 --- a/src/data/real/basic.lean +++ b/src/data/real/basic.lean @@ -257,17 +257,70 @@ instance : ordered_cancel_add_comm_monoid ℝ := by apply_instance instance : ordered_add_comm_monoid ℝ := by apply_instance instance : nontrivial ℝ := ⟨⟨0, 1, ne_of_lt real.zero_lt_one⟩⟩ +@[irreducible] +private def sup : ℝ → ℝ → ℝ | ⟨x⟩ ⟨y⟩ := +⟨quotient.map₂ (⊔) (λ x₁ x₂ hx y₁ y₂ hy, sup_equiv_sup hx hy) x y⟩ + +instance : has_sup ℝ := ⟨sup⟩ + +lemma of_cauchy_sup (a b) : (⟨⟦a ⊔ b⟧⟩ : ℝ) = ⟨⟦a⟧⟩ ⊔ ⟨⟦b⟧⟩ := show _ = sup _ _, by { rw sup, refl } +@[simp] lemma mk_sup (a b) : (mk (a ⊔ b) : ℝ) = mk a ⊔ mk b := of_cauchy_sup _ _ + +@[irreducible] +private def inf : ℝ → ℝ → ℝ | ⟨x⟩ ⟨y⟩ := +⟨quotient.map₂ (⊓) (λ x₁ x₂ hx y₁ y₂ hy, inf_equiv_inf hx hy) x y⟩ + +instance : has_inf ℝ := ⟨inf⟩ + +lemma of_cauchy_inf (a b) : (⟨⟦a ⊓ b⟧⟩ : ℝ) = ⟨⟦a⟧⟩ ⊓ ⟨⟦b⟧⟩ := show _ = inf _ _, by { rw inf, refl } +@[simp] lemma mk_inf (a b) : (mk (a ⊓ b) : ℝ) = mk a ⊓ mk b := of_cauchy_inf _ _ + +instance : distrib_lattice ℝ := +{ sup := (⊔), + le := (≤), + le_sup_left := λ a, real.ind_mk a $ λ a b, real.ind_mk b $ λ b, begin + rw [←mk_sup, mk_le], + exact cau_seq.le_sup_left, + end, + le_sup_right := λ a, real.ind_mk a $ λ a b, real.ind_mk b $ λ b, begin + rw [←mk_sup, mk_le], + exact cau_seq.le_sup_right, + end, + sup_le := λ a, real.ind_mk a $ λ a b, real.ind_mk b $ λ b c, real.ind_mk c $ λ c, begin + simp_rw [←mk_sup, mk_le], + exact cau_seq.sup_le, + end, + inf := (⊓), + inf_le_left := λ a, real.ind_mk a $ λ a b, real.ind_mk b $ λ b, begin + rw [←mk_inf, mk_le], + exact cau_seq.inf_le_left, + end, + inf_le_right := λ a, real.ind_mk a $ λ a b, real.ind_mk b $ λ b, begin + rw [←mk_inf, mk_le], + exact cau_seq.inf_le_right, + end, + le_inf := λ a, real.ind_mk a $ λ a b, real.ind_mk b $ λ b c, real.ind_mk c $ λ c, begin + simp_rw [←mk_inf, mk_le], + exact cau_seq.le_inf, + end, + le_sup_inf := λ a, real.ind_mk a $ λ a b, real.ind_mk b $ λ b c, real.ind_mk c $ λ c, eq.le begin + simp only [←mk_sup, ←mk_inf], + exact congr_arg mk (cau_seq.sup_inf_distrib_left _ _ _).symm + end, + .. real.partial_order } + +/- Extra instances to short-circuit type class resolution -/ +instance : lattice ℝ := infer_instance +instance : semilattice_inf ℝ := infer_instance +instance : semilattice_sup ℝ := infer_instance + open_locale classical +instance : is_total ℝ (≤) := +⟨λ a, real.ind_mk a $ λ a b, real.ind_mk b $ λ b, by simpa using le_total a b⟩ + noncomputable instance : linear_order ℝ := -{ le_total := begin - intros a b, - induction a using real.ind_mk with a, - induction b using real.ind_mk with b, - simpa using le_total a b, - end, - decidable_le := by apply_instance, - .. real.partial_order } +lattice.to_linear_order _ noncomputable instance : linear_ordered_comm_ring ℝ := { .. real.nontrivial, .. real.ordered_ring, .. real.comm_ring, .. real.linear_order } @@ -278,7 +331,6 @@ noncomputable instance : linear_ordered_semiring ℝ := by apply_instance instance : is_domain ℝ := { .. real.nontrivial, .. real.comm_ring, .. linear_ordered_ring.is_domain } - noncomputable instance : linear_ordered_field ℝ := { inv := has_inv.inv, mul_inv_cancel := begin @@ -299,12 +351,6 @@ noncomputable instance : linear_ordered_field ℝ := noncomputable instance : linear_ordered_add_comm_group ℝ := by apply_instance noncomputable instance field : field ℝ := by apply_instance noncomputable instance : division_ring ℝ := by apply_instance -noncomputable instance : distrib_lattice ℝ := by apply_instance -noncomputable instance : lattice ℝ := by apply_instance -noncomputable instance : semilattice_inf ℝ := by apply_instance -noncomputable instance : semilattice_sup ℝ := by apply_instance -noncomputable instance : has_inf ℝ := by apply_instance -noncomputable instance : has_sup ℝ := by apply_instance noncomputable instance decidable_lt (a b : ℝ) : decidable (a < b) := by apply_instance noncomputable instance decidable_le (a b : ℝ) : decidable (a ≤ b) := by apply_instance noncomputable instance decidable_eq (a b : ℝ) : decidable (a = b) := by apply_instance diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 7af447cae0bee..5190b4779d803 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -54,7 +54,8 @@ open_locale classical big_operators @[derive [ ordered_semiring, comm_monoid_with_zero, -- to ensure these instance are computable floor_semiring, comm_semiring, semiring, - semilattice_inf, densely_ordered, order_bot, + semilattice_inf, semilattice_sup, + distrib_lattice, densely_ordered, order_bot, canonically_linear_ordered_semifield, linear_ordered_comm_group_with_zero, archimedean, linear_ordered_semiring, ordered_comm_semiring, canonically_ordered_comm_semiring, has_sub, has_ordered_sub, has_div, inhabited]] @@ -269,15 +270,15 @@ noncomputable def gi : galois_insertion real.to_nnreal coe := galois_insertion.monotone_intro nnreal.coe_mono real.to_nnreal_mono real.le_coe_to_nnreal (λ _, real.to_nnreal_coe) --- note that anything involving the (decidability of the) linear order, including `⊔`/`⊓` (min, max) +-- note that anything involving the (decidability of the) linear order, -- will be noncomputable, everything else should not be. example : order_bot ℝ≥0 := by apply_instance example : partial_order ℝ≥0 := by apply_instance noncomputable example : canonically_linear_ordered_add_monoid ℝ≥0 := by apply_instance noncomputable example : linear_ordered_add_comm_monoid ℝ≥0 := by apply_instance -noncomputable example : distrib_lattice ℝ≥0 := by apply_instance -noncomputable example : semilattice_inf ℝ≥0 := by apply_instance -noncomputable example : semilattice_sup ℝ≥0 := by apply_instance +example : distrib_lattice ℝ≥0 := by apply_instance +example : semilattice_inf ℝ≥0 := by apply_instance +example : semilattice_sup ℝ≥0 := by apply_instance noncomputable example : linear_ordered_semiring ℝ≥0 := by apply_instance example : ordered_comm_semiring ℝ≥0 := by apply_instance noncomputable example : linear_ordered_comm_monoid ℝ≥0 := by apply_instance @@ -819,7 +820,7 @@ end set namespace real /-- The absolute value on `ℝ` as a map to `ℝ≥0`. -/ -@[pp_nodot] noncomputable def nnabs : ℝ →*₀ ℝ≥0 := +@[pp_nodot] def nnabs : ℝ →*₀ ℝ≥0 := { to_fun := λ x, ⟨|x|, abs_nonneg x⟩, map_zero' := by { ext, simp }, map_one' := by { ext, simp }, diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index 79919a734a557..d8ecd9a6ea780 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -117,6 +117,9 @@ lemma tendsto_norm_sq_coprime_pair : filter.tendsto (λ p : fin 2 → ℤ, ((p 0 : ℂ) * z + p 1).norm_sq) cofinite at_top := begin + -- using this instance rather than the automatic `function.module` makes unification issues in + -- `linear_equiv.closed_embedding_of_injective` less bad later in the proof. + letI : module ℝ (fin 2 → ℝ) := normed_space.to_module, let π₀ : (fin 2 → ℝ) →ₗ[ℝ] ℝ := linear_map.proj 0, let π₁ : (fin 2 → ℝ) →ₗ[ℝ] ℝ := linear_map.proj 1, let f : (fin 2 → ℝ) →ₗ[ℝ] ℂ := π₀.smul_right (z:ℂ) + π₁.smul_right 1, @@ -152,15 +155,19 @@ begin conj_of_real, conj_of_real, ← of_real_mul, add_im, of_real_im, zero_add, inv_mul_eq_iff_eq_mul₀ hz], simp only [of_real_im, of_real_re, mul_im, zero_add, mul_zero] } }, - have h₁ := (linear_equiv.closed_embedding_of_injective hf).tendsto_cocompact, + have hf' : closed_embedding f, + { -- for some reason we get a timeout if we try and apply this lemma in a more sensible way + have := @linear_equiv.closed_embedding_of_injective ℝ _ (fin 2 → ℝ) _ (id _) ℂ _ _ _ _, + rotate 2, + exact f, + exact this hf }, have h₂ : tendsto (λ p : fin 2 → ℤ, (coe : ℤ → ℝ) ∘ p) cofinite (cocompact _), { convert tendsto.pi_map_Coprod (λ i, int.tendsto_coe_cofinite), { rw Coprod_cofinite }, { rw Coprod_cocompact } }, - exact tendsto_norm_sq_cocompact_at_top.comp (h₁.comp h₂) + exact tendsto_norm_sq_cocompact_at_top.comp (hf'.tendsto_cocompact.comp h₂), end - /-- Given `coprime_pair` `p=(c,d)`, the matrix `[[a,b],[*,*]]` is sent to `a*c+b*d`. This is the linear map version of this operation. -/ diff --git a/src/topology/instances/rat.lean b/src/topology/instances/rat.lean index 766fc107b595d..7cf21dd48f177 100644 --- a/src/topology/instances/rat.lean +++ b/src/topology/instances/rat.lean @@ -17,8 +17,9 @@ open metric set filter namespace rat -noncomputable instance : metric_space ℚ := -metric_space.induced coe rat.cast_injective real.metric_space +-- without the `by exact` this is noncomputable +instance : metric_space ℚ := +metric_space.induced coe (by exact rat.cast_injective) real.metric_space theorem dist_eq (x y : ℚ) : dist x y = |x - y| := rfl diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index c2bc8c35a9825..cd65e4f401b2e 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1191,7 +1191,7 @@ emetric.complete_of_cauchy_seq_tendsto section real /-- Instantiate the reals as a pseudometric space. -/ -noncomputable instance real.pseudo_metric_space : pseudo_metric_space ℝ := +instance real.pseudo_metric_space : pseudo_metric_space ℝ := { dist := λx y, |x - y|, dist_self := by simp [abs_zero], dist_comm := assume x y, abs_sub_comm _ _, @@ -1485,7 +1485,7 @@ end mul_opposite section nnreal -noncomputable instance : pseudo_metric_space ℝ≥0 := subtype.pseudo_metric_space +instance : pseudo_metric_space ℝ≥0 := subtype.pseudo_metric_space lemma nnreal.dist_eq (a b : ℝ≥0) : dist a b = |(a:ℝ) - b| := rfl @@ -2762,7 +2762,7 @@ instance : metric_space punit.{u + 1} := section real /-- Instantiate the reals as a metric space. -/ -noncomputable instance real.metric_space : metric_space ℝ := +instance real.metric_space : metric_space ℝ := { eq_of_dist_eq_zero := λ x y h, by simpa [dist, sub_eq_zero] using h, ..real.pseudo_metric_space } @@ -2770,7 +2770,7 @@ end real section nnreal -noncomputable instance : metric_space ℝ≥0 := subtype.metric_space +instance : metric_space ℝ≥0 := subtype.metric_space end nnreal diff --git a/src/topology/uniform_space/compare_reals.lean b/src/topology/uniform_space/compare_reals.lean index cab27db711aa9..c09c511925a79 100644 --- a/src/topology/uniform_space/compare_reals.lean +++ b/src/topology/uniform_space/compare_reals.lean @@ -75,7 +75,6 @@ begin end /-- Cauchy reals packaged as a completion of ℚ using the absolute value route. -/ -noncomputable def rational_cau_seq_pkg : @abstract_completion ℚ $ is_absolute_value.uniform_space (abs : ℚ → ℚ) := { space := ℝ, coe := (coe : ℚ → ℝ), From 95d67152446e6e4bfa32a824501083ec762551e9 Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Tue, 20 Sep 2022 07:09:34 +0000 Subject: [PATCH 327/828] feat(model_theory/types): Complete types over a theory (#16548) Defines the space of complete types over a particular theory in a particular type of variables --- src/model_theory/semantics.lean | 3 + src/model_theory/types.lean | 156 ++++++++++++++++++++++++++++++++ 2 files changed, 159 insertions(+) create mode 100644 src/model_theory/types.lean diff --git a/src/model_theory/semantics.lean b/src/model_theory/semantics.lean index 4a946f2219829..cb4f158b9eeb5 100644 --- a/src/model_theory/semantics.lean +++ b/src/model_theory/semantics.lean @@ -735,6 +735,9 @@ theorem model_iff_subset_complete_theory : M ⊨ T ↔ T ⊆ L.complete_theory M := T.model_iff +theorem complete_theory.subset [MT : M ⊨ T] : T ⊆ L.complete_theory M := +model_iff_subset_complete_theory.1 MT + end Theory instance model_complete_theory : M ⊨ L.complete_theory M := diff --git a/src/model_theory/types.lean b/src/model_theory/types.lean new file mode 100644 index 0000000000000..e16f938895b33 --- /dev/null +++ b/src/model_theory/types.lean @@ -0,0 +1,156 @@ +/- +Copyright (c) 2022 Aaron Anderson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Aaron Anderson +-/ +import model_theory.satisfiability + +/-! +# Type Spaces +This file defines the space of complete types over a first-order theory. +(Note that types in model theory are different from types in type theory.) + +## Main Definitions +* `first_order.language.Theory.complete_type`: + `T.complete_type α` consists of complete types over the theory `T` with variables `α`. + +## Main Results +* `first_order.language.Theory.complete_type.nonempty_iff`: + The space `T.complete_type α` is nonempty exactly when `T` is satisfiable. + +## Implementation Notes +* Complete types are implemented as maximal consistent theories in an expanded language. +More frequently they are described as maximal consistent sets of formulas, but this is equivalent. + +## TODO +* Connect `T.complete_type α` to sets of formulas `L.formula α`. + +-/ + +universes u v w w' + +open cardinal set +open_locale cardinal first_order classical + +namespace first_order +namespace language +namespace Theory + +variables {L : language.{u v}} (T : L.Theory) (α : Type w) + +/-- A complete type over a given theory in a certain type of variables is a maximally + consistent (with the theory) set of formulas in that type. -/ +structure complete_type := +(to_Theory : L[[α]].Theory) +(subset' : (L.Lhom_with_constants α).on_Theory T ⊆ to_Theory) +(is_maximal' : to_Theory.is_maximal) + +variables {T α} + +namespace complete_type + +instance : set_like (T.complete_type α) (L[[α]].sentence) := +⟨λ p, p.to_Theory, λ p q h, begin + cases p, + cases q, + congr', +end⟩ + +lemma is_maximal (p : T.complete_type α) : is_maximal (p : L[[α]].Theory) := +p.is_maximal' + +lemma subset (p : T.complete_type α) : + (L.Lhom_with_constants α).on_Theory T ⊆ (p : L[[α]].Theory) := +p.subset' + +lemma mem_or_not_mem (p : T.complete_type α) (φ : L[[α]].sentence) : φ ∈ p ∨ φ.not ∈ p := +p.is_maximal.mem_or_not_mem φ + +lemma mem_of_models (p : T.complete_type α) {φ : L[[α]].sentence} + (h : (L.Lhom_with_constants α).on_Theory T ⊨ φ) : + φ ∈ p := +(p.mem_or_not_mem φ).resolve_right (λ con, ((models_iff_not_satisfiable _).1 h) + (p.is_maximal.1.mono (union_subset p.subset (singleton_subset_iff.2 con)))) + +lemma not_mem_iff (p : T.complete_type α) (φ : L[[α]].sentence) : + φ.not ∈ p ↔ ¬ φ ∈ p := +⟨λ hf ht, begin + have h : ¬ is_satisfiable ({φ, φ.not} : L[[α]].Theory), + { rintros ⟨⟨_, _, h, _⟩⟩, + simp only [model_iff, mem_insert_iff, mem_singleton_iff, forall_eq_or_imp, + forall_eq] at h, + exact h.2 h.1 }, + refine h (p.is_maximal.1.mono _), + rw [insert_subset, singleton_subset_iff], + exact ⟨ht, hf⟩, +end, (p.mem_or_not_mem φ).resolve_left⟩ + +@[simp] lemma compl_set_of_mem {φ : L[[α]].sentence} : + {p : T.complete_type α | φ ∈ p}ᶜ = {p : T.complete_type α | φ.not ∈ p} := +ext (λ _, (not_mem_iff _ _).symm) + +lemma set_of_subset_eq_empty_iff (S : L[[α]].Theory) : + {p : T.complete_type α | S ⊆ ↑p} = ∅ ↔ + ¬ ((L.Lhom_with_constants α).on_Theory T ∪ S).is_satisfiable := +begin + rw [iff_not_comm, ← not_nonempty_iff_eq_empty, not_not, set.nonempty], + refine ⟨λ h, ⟨⟨L[[α]].complete_theory h.some, (subset_union_left _ S).trans + complete_theory.subset, complete_theory.is_maximal _ _⟩, (subset_union_right + ((L.Lhom_with_constants α).on_Theory T) _).trans complete_theory.subset⟩, _⟩, + rintro ⟨p, hp⟩, + exact p.is_maximal.1.mono (union_subset p.subset hp), +end + +lemma set_of_mem_eq_univ_iff (φ : L[[α]].sentence) : + {p : T.complete_type α | φ ∈ p} = univ ↔ (L.Lhom_with_constants α).on_Theory T ⊨ φ := +begin + rw [models_iff_not_satisfiable, ← compl_empty_iff, compl_set_of_mem, + ← set_of_subset_eq_empty_iff], + simp, +end + +lemma set_of_subset_eq_univ_iff (S : L[[α]].Theory) : + {p : T.complete_type α | S ⊆ ↑p} = univ ↔ + (∀ φ, φ ∈ S → (L.Lhom_with_constants α).on_Theory T ⊨ φ) := +begin + have h : {p : T.complete_type α | S ⊆ ↑p} = ⋂₀ ((λ φ, {p | φ ∈ p}) '' S), + { ext, + simp [subset_def] }, + simp_rw [h, sInter_eq_univ, ← set_of_mem_eq_univ_iff], + refine ⟨λ h φ φS, h _ ⟨_, φS, rfl⟩, _⟩, + rintro h _ ⟨φ, h1, rfl⟩, + exact h _ h1, +end + +lemma nonempty_iff : nonempty (T.complete_type α) ↔ + T.is_satisfiable := +begin + rw ← is_satisfiable_on_Theory_iff (Lhom_with_constants_injective L α), + rw [nonempty_iff_univ_nonempty, ← ne_empty_iff_nonempty, ne.def, not_iff_comm, + ← union_empty ((L.Lhom_with_constants α).on_Theory T), ← set_of_subset_eq_empty_iff], + simp, +end + +instance : nonempty (complete_type ∅ α) := +nonempty_iff.2 (is_satisfiable_empty L) + +lemma Inter_set_of_subset {ι : Type*} (S : ι → L[[α]].Theory) : + (⋂ (i : ι), {p : T.complete_type α | S i ⊆ p}) = {p | (⋃ (i : ι), S i) ⊆ p} := +begin + ext, + simp only [mem_Inter, mem_set_of_eq, Union_subset_iff], +end + +lemma to_list_foldr_inf_mem {p : T.complete_type α} {t : finset (L[[α]]).sentence} : + t.to_list.foldr (⊓) ⊤ ∈ p ↔ (t : L[[α]].Theory) ⊆ ↑p := +begin + simp_rw [subset_def, ← set_like.mem_coe, p.is_maximal.mem_iff_models, models_sentence_iff, + sentence.realize, formula.realize, bounded_formula.realize_foldr_inf, finset.mem_to_list], + exact ⟨λ h φ hφ M, h _ _ hφ, λ h M φ hφ, h _ hφ _⟩, +end + +end complete_type + +end Theory +end language +end first_order From b76d2b40a425b5f296bb8d3dfe915f6b05313b2d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 20 Sep 2022 07:09:35 +0000 Subject: [PATCH 328/828] feat(measure_theory/measurable_space): add `prod.measurable_singleton_class` (#16559) * golf `pi.measurable_singleton_class`, rename from `measurable_singleton_class.pi`; * add `prod.measurable_singleton_class`. --- src/measure_theory/measurable_space.lean | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index bc6166e6fb754..f3c4d2e7383da 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -573,6 +573,11 @@ lemma measurable_set_swap_iff {s : set (α × β)} : measurable_set (prod.swap ⁻¹' s) ↔ measurable_set s := ⟨λ hs, by { convert measurable_swap hs, ext ⟨x, y⟩, refl }, λ hs, measurable_swap hs⟩ +instance [measurable_singleton_class α] [measurable_singleton_class β] : + measurable_singleton_class (α × β) := +⟨λ ⟨a, b⟩, @singleton_prod_singleton _ _ a b ▸ + (measurable_set_singleton a).prod (measurable_set_singleton b)⟩ + lemma measurable_from_prod_countable [countable β] [measurable_singleton_class β] {mγ : measurable_space γ} {f : α × β → γ} (hf : ∀ y, measurable (λ x, f (x, y))) : measurable f := @@ -697,14 +702,9 @@ begin { simp [measurable_set_pi_of_nonempty hs, h, ← not_nonempty_iff_eq_empty] } end -instance measurable_singleton_class.pi [countable δ] [Π a, measurable_singleton_class (π a)] : +instance [countable δ] [Π a, measurable_singleton_class (π a)] : measurable_singleton_class (Π a, π a) := -⟨λ f, begin - convert measurable_set.univ_pi - (show ∀ t, measurable_set {f t}, from λ t, measurable_set_singleton (f t)), - ext g, - simp only [function.funext_iff, mem_singleton_iff, mem_univ_pi], -end⟩ +⟨λ f, univ_pi_singleton f ▸ measurable_set.univ_pi (λ t, measurable_set_singleton (f t))⟩ variable (π) From dc680a80ef7631557cacadb38aad6d863a59d411 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Tue, 20 Sep 2022 09:35:03 +0000 Subject: [PATCH 329/828] =?UTF-8?q?feat(topology/uniform=5Fspace):=20Every?= =?UTF-8?q?=20Cauchy=20sequence=20over=20=E2=84=95=20is=20totally=20bounde?= =?UTF-8?q?d=20(#16563)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR proves that the image of every Cauchy sequence `ℕ → α` is totally bounded. Co-authored-by: Yury G. Kudryashov --- src/topology/uniform_space/cauchy.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/src/topology/uniform_space/cauchy.lean b/src/topology/uniform_space/cauchy.lean index ac31a3cdfac01..86372d64aa777 100644 --- a/src/topology/uniform_space/cauchy.lean +++ b/src/topology/uniform_space/cauchy.lean @@ -560,6 +560,20 @@ lemma compact_of_totally_bounded_is_closed [complete_space α] {s : set α} (ht : totally_bounded s) (hc : is_closed s) : is_compact s := (@compact_iff_totally_bounded_complete α _ s).2 ⟨ht, hc.is_complete⟩ +/-- Every Cauchy sequence over `ℕ` is totally bounded. -/ +lemma cauchy_seq.totally_bounded_range {s : ℕ → α} (hs : cauchy_seq s) : + totally_bounded (range s) := +begin + refine totally_bounded_iff_subset.2 (λ a ha, _), + cases cauchy_seq_iff.1 hs a ha with n hn, + refine ⟨s '' {k | k ≤ n}, image_subset_range _ _, (finite_le_nat _).image _, _⟩, + rw [range_subset_iff, bUnion_image], + intro m, + rw [mem_Union₂], + cases le_total m n with hm hm, + exacts [⟨m, hm, refl_mem_uniformity ha⟩, ⟨n, le_refl n, hn m hm n le_rfl⟩] +end + /-! ### Sequentially complete space From 744170a92f4d81f3db943312f5c01af296266308 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 20 Sep 2022 11:17:56 +0000 Subject: [PATCH 330/828] feat(data/pfun): Product of partial functions (#15389) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `pfun.prod : (α →. γ) → (β →. δ) → α × β →. γ × δ`. --- src/data/pfun.lean | 56 +++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 55 insertions(+), 1 deletion(-) diff --git a/src/data/pfun.lean b/src/data/pfun.lean index ca938727bc5cf..7505bb55575f4 100644 --- a/src/data/pfun.lean +++ b/src/data/pfun.lean @@ -58,7 +58,7 @@ def pfun (α β : Type*) := α → part β infixr ` →. `:25 := pfun namespace pfun -variables {α β γ δ : Type*} +variables {α β γ δ ε ι : Type*} instance : inhabited (α →. β) := ⟨λ a, part.none⟩ @@ -454,4 +454,58 @@ ext $ λ _ _, by simp only [comp_apply, part.bind_comp] lemma coe_comp (g : β → γ) (f : α → β) : ((g ∘ f : α → γ) : α →. γ) = (g : β →. γ).comp f := ext $ λ _ _, by simp only [coe_val, comp_apply, part.bind_some] +/-- Product of partial functions. -/ +def prod_lift (f : α →. β) (g : α →. γ) : α →. β × γ := +λ x, ⟨(f x).dom ∧ (g x).dom, λ h, ((f x).get h.1, (g x).get h.2)⟩ + +@[simp] lemma dom_prod_lift (f : α →. β) (g : α →. γ) : + (f.prod_lift g).dom = {x | (f x).dom ∧ (g x).dom} := rfl + +lemma get_prod_lift (f : α →. β) (g : α →. γ) (x : α) (h) : + (f.prod_lift g x).get h = ((f x).get h.1, (g x).get h.2) := rfl + +@[simp] lemma prod_lift_apply (f : α →. β) (g : α →. γ) (x : α) : + f.prod_lift g x = ⟨(f x).dom ∧ (g x).dom, λ h, ((f x).get h.1, (g x).get h.2)⟩ := rfl + +lemma mem_prod_lift {f : α →. β} {g : α →. γ} {x : α} {y : β × γ} : + y ∈ f.prod_lift g x ↔ y.1 ∈ f x ∧ y.2 ∈ g x := +begin + transitivity ∃ hp hq, (f x).get hp = y.1 ∧ (g x).get hq = y.2, + { simp only [prod_lift, part.mem_mk_iff, and.exists, prod.ext_iff] }, + { simpa only [exists_and_distrib_left, exists_and_distrib_right] } +end + +/-- Product of partial functions. -/ +def prod_map (f : α →. γ) (g : β →. δ) : α × β →. γ × δ := +λ x, ⟨(f x.1).dom ∧ (g x.2).dom, λ h, ((f x.1).get h.1, (g x.2).get h.2)⟩ + +@[simp] lemma dom_prod_map (f : α →. γ) (g : β →. δ) : + (f.prod_map g).dom = {x | (f x.1).dom ∧ (g x.2).dom} := rfl + +lemma get_prod_map (f : α →. γ) (g : β →. δ) (x : α × β) (h) : + (f.prod_map g x).get h = ((f x.1).get h.1, (g x.2).get h.2) := rfl + +@[simp] lemma prod_map_apply (f : α →. γ) (g : β →. δ) (x : α × β) : + f.prod_map g x = ⟨(f x.1).dom ∧ (g x.2).dom, λ h, ((f x.1).get h.1, (g x.2).get h.2)⟩ := rfl + +lemma mem_prod_map {f : α →. γ} {g : β →. δ} {x : α × β} {y : γ × δ} : + y ∈ f.prod_map g x ↔ y.1 ∈ f x.1 ∧ y.2 ∈ g x.2 := +begin + transitivity ∃ hp hq, (f x.1).get hp = y.1 ∧ (g x.2).get hq = y.2, + { simp only [prod_map, part.mem_mk_iff, and.exists, prod.ext_iff] }, + { simpa only [exists_and_distrib_left, exists_and_distrib_right] } +end + +@[simp] lemma prod_lift_fst_comp_snd_comp (f : α →. γ) (g : β →. δ) : + prod_lift (f.comp ((prod.fst : α × β → α) : α × β →. α)) + (g.comp ((prod.snd : α × β → β) : α × β →. β)) = prod_map f g := +ext $ λ a, by simp + +@[simp] lemma prod_map_id_id : (pfun.id α).prod_map (pfun.id β) = pfun.id _ := +ext $ λ _ _, by simp [eq_comm] + +@[simp] lemma prod_map_comp_comp (f₁ : α →. β) (f₂ : β →. γ) (g₁ : δ →. ε) (g₂ : ε →. ι) : + (f₂.comp f₁).prod_map (g₂.comp g₁) = (f₂.prod_map g₂).comp (f₁.prod_map g₁) := +ext $ λ _ _, by tidy + end pfun From 4aab2abced69a9e579b1e6dc2856ed3db48e2cbd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 20 Sep 2022 11:17:57 +0000 Subject: [PATCH 331/828] =?UTF-8?q?feat(combinatorics/additive/pluennecke?= =?UTF-8?q?=5Fruzsa):=20The=20Pl=C3=BCnnecke-Ruzsa=20inequality=20(#15440)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove the Ruzsa triangle inequalities, the Plünnecke-Petridis lemma, and the Plünnecke-Ruzsa inequality. Co-authored-by: georgeshakan --- docs/references.bib | 30 +++ .../additive/pluennecke_ruzsa.lean | 241 ++++++++++++++++++ 2 files changed, 271 insertions(+) create mode 100644 src/combinatorics/additive/pluennecke_ruzsa.lean diff --git a/docs/references.bib b/docs/references.bib index 601d8fff59768..56e03ff295fe8 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1606,6 +1606,21 @@ @Article{ orosi2018faulhaber zbl = {1411.41023} } +@InCollection{ petridis2014, + author = {Petridis, G.}, + title = {The {Pl{\"u}nnecke}-{Ruzsa} inequality: an overview}, + booktitle = {Combinatorial and additive number theory. Selected papers based on the presentations at the conferences CANT 2011 and 2012, New York, NY, USA, May 2011 and May 2012}, + isbn = {978-1-4939-1600-9; 978-1-4939-1601-6}, + pages = {229--241}, + year = {2014}, + publisher = {New York, NY: Springer}, + language = {English}, + doi = {10.1007/978-1-4939-1601-6_16}, + keywords = {11B30}, + zblath = {6463830}, + zbl = {1371.11029} +} + @Article{ phillips1940, author = {Phillips, Ralph S.}, title = {Integration in a convex linear topological space}, @@ -1659,6 +1674,21 @@ @Book{ rudin2006real isbn = {0-07-100276-6} } +@Book{ tao-vu, + author = {Tao, Terence and Vu, Van H.}, + title = {Additive combinatorics}, + fseries = {Cambridge Studies in Advanced Mathematics}, + series = {Camb. Stud. Adv. Math.}, + volume = {105}, + isbn = {0-521-85386-9}, + year = {2006}, + publisher = {Cambridge: Cambridge University Press}, + language = {English}, + keywords = {11-02,05-02,05D10,05D40,11B75,11B13,11N13,11P70,11K31,11P82,28D05,37A45}, + zbmath = {5066399}, + zbl = {1127.11002} +} + @Book{ samuel1967, author = {Samuel, Pierre}, title = {Th\'{e}orie alg\'{e}brique des nombres}, diff --git a/src/combinatorics/additive/pluennecke_ruzsa.lean b/src/combinatorics/additive/pluennecke_ruzsa.lean new file mode 100644 index 0000000000000..16eb9cd950f8d --- /dev/null +++ b/src/combinatorics/additive/pluennecke_ruzsa.lean @@ -0,0 +1,241 @@ +/- +Copyright (c) 2022 Yaël Dillies, George Shakan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, George Shakan +-/ +import combinatorics.double_counting +import data.finset.pointwise +import data.rat.nnrat + +/-! +# The Plünnecke-Ruzsa inequality + +This file proves Ruzsa's triangle inequality, the Plünnecke-Petridis lemma, and the Plünnecke-Ruzsa +inequality. + +## Main declarations + +* `finset.card_sub_mul_le_card_sub_mul_card_sub`: Ruzsa's triangle inequality, difference version. +* `finset.card_add_mul_le_card_add_mul_card_add`: Ruzsa's triangle inequality, sum version. +* `finset.pluennecke_petridis`: The Plünnecke-Petridis lemma. +* `finset.card_smul_div_smul_le`: The Plünnecke-Ruzsa inequality. + +## References + +* [Giorgis Petridis, *The Plünnecke-Ruzsa inequality: an overview*][petridis2014] +* [Terrence Tao, Van Vu, *Additive Combinatorics][tao-vu] +-/ + +open nat +open_locale nnrat pointwise + +namespace finset +variables {α : Type*} [comm_group α] [decidable_eq α] {A B C : finset α} + +/-- **Ruzsa's triangle inequality**. Division version. -/ +@[to_additive card_sub_mul_le_card_sub_mul_card_sub +"**Ruzsa's triangle inequality**. Subtraction version."] +lemma card_div_mul_le_card_div_mul_card_div (A B C : finset α) : + (A / C).card * B.card ≤ (A / B).card * (B / C).card := +begin + rw [←card_product (A / B), ←mul_one ((finset.product _ _).card)], + refine card_mul_le_card_mul (λ b ac, ac.1 * ac.2 = b) (λ x hx, _) + (λ x hx, card_le_one_iff.2 $ λ u v hu hv, + ((mem_bipartite_below _).1 hu).2.symm.trans ((mem_bipartite_below _).1 hv).2), + obtain ⟨a, c, ha, hc, rfl⟩ := mem_div.1 hx, + refine card_le_card_of_inj_on (λ b, (a / b, b / c)) (λ b hb, _) (λ b₁ _ b₂ _ h, _), + { rw mem_bipartite_above, + exact ⟨mk_mem_product (div_mem_div ha hb) (div_mem_div hb hc), div_mul_div_cancel' _ _ _⟩ }, + { exact div_right_injective (prod.ext_iff.1 h).1 } +end + +/-- **Ruzsa's triangle inequality**. Div-mul-mul version. -/ +@[to_additive card_sub_mul_le_card_add_mul_card_add +"**Ruzsa's triangle inequality**. Sub-add-add version."] +lemma card_div_mul_le_card_mul_mul_card_mul (A B C : finset α) : + (A / C).card * B.card ≤ (A * B).card * (B * C).card := +begin + rw [←div_inv_eq_mul, ←card_inv B, ←card_inv (B * C), mul_inv, ←div_eq_mul_inv], + exact card_div_mul_le_card_div_mul_card_div _ _ _, +end + +/-- **Ruzsa's triangle inequality**. Mul-div-div version. -/ +@[to_additive card_add_mul_le_card_sub_mul_card_add +"**Ruzsa's triangle inequality**. Add-sub-sub version."] +lemma card_mul_mul_le_card_div_mul_card_mul (A B C : finset α) : + (A * C).card * B.card ≤ (A / B).card * (B * C).card := +by { rw [←div_inv_eq_mul, ←div_inv_eq_mul B], exact card_div_mul_le_card_div_mul_card_div _ _ _ } + +/-- **Ruzsa's triangle inequality**. Mul-mul-div version. -/ +@[to_additive card_add_mul_le_card_add_mul_card_sub +"**Ruzsa's triangle inequality**. Add-add-sub version."] +lemma card_mul_mul_le_card_mul_mul_card_div (A B C : finset α) : + (A * C).card * B.card ≤ (A * B).card * (B / C).card := +by { rw [←div_inv_eq_mul, div_eq_mul_inv B], exact card_div_mul_le_card_mul_mul_card_mul _ _ _ } + +@[to_additive] +lemma mul_pluennecke_petridis (C : finset α) + (hA : ∀ A' ⊆ A, (A * B).card * A'.card ≤ (A' * B).card * A.card) : + (A * B * C).card * A.card ≤ (A * B).card * (A * C).card := +begin + induction C using finset.induction_on with x C hc ih, + { simp }, + set A' := A ∩ (A * C / {x}) with hA', + set C' := insert x C with hC', + have h₀ : A' * {x} = (A * {x}) ∩ (A * C), + { rw [hA', inter_mul_singleton, (is_unit_singleton x).div_mul_cancel] }, + have h₁ : A * B * C' = (A * B * C) ∪ (A * B * {x}) \ (A' * B * {x}), + { rw [hC', insert_eq, union_comm, mul_union], + refine (sup_sdiff_eq_sup _).symm, + rw [mul_right_comm, mul_right_comm A, h₀], + exact mul_subset_mul_right (inter_subset_right _ _) }, + have h₂ : A' * B * {x} ⊆ A * B * {x} := + mul_subset_mul_right (mul_subset_mul_right $ inter_subset_left _ _), + have h₃ : (A * B * C').card ≤ (A * B * C).card + (A * B).card - (A' * B).card, + { rw h₁, + refine (card_union_le _ _).trans_eq _, + rw [card_sdiff h₂, ←add_tsub_assoc_of_le (card_le_of_subset h₂), card_mul_singleton, + card_mul_singleton] }, + refine (mul_le_mul_right' h₃ _).trans _, + rw [tsub_mul, add_mul], + refine (tsub_le_tsub (add_le_add_right ih _) $ hA _ $ inter_subset_left _ _).trans_eq _, + rw [←mul_add, ←mul_tsub, ←hA', insert_eq, mul_union, ←card_mul_singleton A x, + ←card_mul_singleton A' x, add_comm (card _), h₀, + eq_tsub_of_add_eq (card_union_add_card_inter _ _)], +end + +/-! ### Sum triangle inequality -/ + +-- Auxiliary lemma for Ruzsa's triangle sum inequality, and the Plünnecke-Ruzsa inequality. +@[to_additive] +private lemma mul_aux (hA : A.nonempty) (hAB : A ⊆ B) + (h : ∀ A' ∈ B.powerset.erase ∅, ((A * C).card : ℚ≥0) / ↑(A.card) ≤ ((A' * C).card) / ↑(A'.card)) : + ∀ A' ⊆ A, (A * C).card * A'.card ≤ (A' * C).card * A.card := +begin + rintro A' hAA', + obtain rfl | hA' := A'.eq_empty_or_nonempty, + { simp }, + have hA₀ : (0 : ℚ≥0) < A.card := cast_pos.2 hA.card_pos, + have hA₀' : (0 : ℚ≥0) < A'.card := cast_pos.2 hA'.card_pos, + exact_mod_cast (div_le_div_iff hA₀ hA₀').1 (h _ $ mem_erase_of_ne_of_mem hA'.ne_empty $ + mem_powerset.2 $ hAA'.trans hAB), +end + +/-- **Ruzsa's triangle inequality**. Multiplication version. -/ +@[to_additive card_add_mul_card_le_card_add_mul_card_add +"**Ruzsa's triangle inequality**. Addition version."] +lemma card_mul_mul_card_le_card_mul_mul_card_mul (A B C : finset α) : + (A * C).card * B.card ≤ (A * B).card * (B * C).card := +begin + obtain rfl | hB := B.eq_empty_or_nonempty, + { simp }, + have hB' : B ∈ B.powerset.erase ∅ := mem_erase_of_ne_of_mem hB.ne_empty (mem_powerset_self _), + obtain ⟨U, hU, hUA⟩ := exists_min_image (B.powerset.erase ∅) (λ U, (U * A).card/U.card : _ → ℚ≥0) + ⟨B, hB'⟩, + rw [mem_erase, mem_powerset, ←nonempty_iff_ne_empty] at hU, + refine cast_le.1 (_ : (_ : ℚ≥0) ≤ _), + push_cast, + refine (le_div_iff $ by exact cast_pos.2 hB.card_pos).1 _, + rw [mul_div_right_comm, mul_comm _ B], + refine (cast_le.2 $ card_le_card_mul_left _ hU.1).trans _, + refine le_trans _ (mul_le_mul (hUA _ hB') (cast_le.2 $ card_le_of_subset $ + mul_subset_mul_right hU.2) (zero_le _) $ zero_le _), + rw [←mul_div_right_comm, ←mul_assoc], + refine (le_div_iff $ by exact cast_pos.2 hU.1.card_pos).2 _, + exact_mod_cast mul_pluennecke_petridis C (mul_aux hU.1 hU.2 hUA), +end + +/-- **Ruzsa's triangle inequality**. Add-sub-sub version. -/ +lemma card_mul_mul_le_card_div_mul_card_div (A B C : finset α) : + (A * C).card * B.card ≤ (A / B).card * (B / C).card := +begin + rw [div_eq_mul_inv, ←card_inv B, ←card_inv (B / C), inv_div', div_inv_eq_mul], + exact card_mul_mul_card_le_card_mul_mul_card_mul _ _ _, +end + +/-- **Ruzsa's triangle inequality**. Sub-add-sub version. -/ +lemma card_div_mul_le_card_mul_mul_card_div (A B C : finset α) : + (A / C).card * B.card ≤ (A * B).card * (B / C).card := +by { rw [div_eq_mul_inv, div_eq_mul_inv], exact card_mul_mul_card_le_card_mul_mul_card_mul _ _ _ } + +/-- **Ruzsa's triangle inequality**. Sub-sub-add version. -/ +lemma card_div_mul_le_card_div_mul_card_mul (A B C : finset α) : + (A / C).card * B.card ≤ (A / B).card * (B * C).card := +by { rw [←div_inv_eq_mul, div_eq_mul_inv], exact card_mul_mul_le_card_div_mul_card_div _ _ _ } + + +lemma card_add_nsmul_le {α : Type*} [add_comm_group α] [decidable_eq α] {A B : finset α} + (hAB : ∀ A' ⊆ A, (A + B).card * A'.card ≤ (A' + B).card * A.card) (n : ℕ) : + ((A + n • B).card : ℚ≥0) ≤ ((A + B).card / A.card) ^ n * A.card := +begin + obtain rfl | hA := A.eq_empty_or_nonempty, + { simp }, + induction n with n ih, + { simp }, + rw [succ_nsmul, ←add_assoc, pow_succ, mul_assoc, ←mul_div_right_comm, le_div_iff, ←cast_mul], + swap, exact (cast_pos.2 hA.card_pos), + refine (cast_le.2 $ add_pluennecke_petridis _ hAB).trans _, + rw cast_mul, + exact mul_le_mul_of_nonneg_left ih (zero_le _), +end + +@[to_additive] +lemma card_mul_pow_le (hAB : ∀ A' ⊆ A, (A * B).card * A'.card ≤ (A' * B).card * A.card) (n : ℕ) : + ((A * B ^ n).card : ℚ≥0) ≤ ((A * B).card / A.card) ^ n * A.card := +begin + obtain rfl | hA := A.eq_empty_or_nonempty, + { simp }, + induction n with n ih, + { simp }, + rw [pow_succ, ←mul_assoc, pow_succ, @mul_assoc ℚ≥0, ←mul_div_right_comm, le_div_iff, ←cast_mul], + swap, exact (cast_pos.2 hA.card_pos), + refine (cast_le.2 $ mul_pluennecke_petridis _ hAB).trans _, + rw cast_mul, + exact mul_le_mul_of_nonneg_left ih (zero_le _), +end + +/-- The **Plünnecke-Ruzsa inequality**. Multiplication version. Note that this is genuinely harder +than the division version because we cannot use a double counting argument. -/ +@[to_additive "The **Plünnecke-Ruzsa inequality**. Addition version. Note that this is genuinely +harder than the subtraction version because we cannot use a double counting argument."] +lemma card_pow_div_pow_le (hA : A.nonempty) (B : finset α) (m n : ℕ) : + ((B ^ m / B ^ n).card : ℚ≥0) ≤ ((A * B).card / A.card) ^ (m + n) * A.card := +begin + have hA' : A ∈ A.powerset.erase ∅ := mem_erase_of_ne_of_mem hA.ne_empty (mem_powerset_self _), + obtain ⟨C, hC, hCA⟩ := exists_min_image (A.powerset.erase ∅) (λ C, (C * B).card/C.card : _ → ℚ≥0) + ⟨A, hA'⟩, + rw [mem_erase, mem_powerset, ←nonempty_iff_ne_empty] at hC, + refine (mul_le_mul_right $ cast_pos.2 hC.1.card_pos).1 _, + norm_cast, + refine (cast_le.2 $ card_div_mul_le_card_mul_mul_card_mul _ _ _).trans _, + push_cast, + rw mul_comm _ C, + refine (mul_le_mul (card_mul_pow_le (mul_aux hC.1 hC.2 hCA) _) + (card_mul_pow_le (mul_aux hC.1 hC.2 hCA) _) (zero_le _) $ zero_le _).trans _, + rw [mul_mul_mul_comm, ←pow_add, ←mul_assoc], + exact mul_le_mul_of_nonneg_right (mul_le_mul (pow_le_pow_of_le_left (zero_le _) (hCA _ hA') _) + (cast_le.2 $ card_le_of_subset hC.2) (zero_le _) $ zero_le _) (zero_le _), +end + +/-- The **Plünnecke-Ruzsa inequality**. Subtraction version. -/ +@[to_additive "The **Plünnecke-Ruzsa inequality**. Subtraction version."] +lemma card_pow_div_pow_le' (hA : A.nonempty) (B : finset α) (m n : ℕ) : + ((B ^ m / B ^ n).card : ℚ≥0) ≤ ((A / B).card / A.card) ^ (m + n) * A.card := +begin + rw [←card_inv, inv_div', ←inv_pow, ←inv_pow, div_eq_mul_inv A], + exact card_pow_div_pow_le hA _ _ _, +end + +/-- Special case of the **Plünnecke-Ruzsa inequality**. Multiplication version. -/ +@[to_additive "Special case of the **Plünnecke-Ruzsa inequality**. Addition version."] +lemma card_pow_le (hA : A.nonempty) (B : finset α) (n : ℕ) : + ((B ^ n).card : ℚ≥0) ≤ ((A * B).card / A.card) ^ n * A.card := +by simpa only [pow_zero, div_one] using card_pow_div_pow_le hA _ _ 0 + +/-- Special case of the **Plünnecke-Ruzsa inequality**. Division version. -/ +@[to_additive "Special case of the **Plünnecke-Ruzsa inequality**. Subtraction version."] +lemma card_pow_le' (hA : A.nonempty) (B : finset α) (n : ℕ) : + ((B ^ n).card : ℚ≥0) ≤ ((A / B).card / A.card) ^ n * A.card := +by simpa only [pow_zero, div_one] using card_pow_div_pow_le' hA _ _ 0 + +end finset From b7cc093e766ddc17be3fca21a60697cacf065bd6 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 20 Sep 2022 14:59:59 +0000 Subject: [PATCH 332/828] chore(topology/constructions): use `variables` (#16562) Also move `continuous_pi_iff` up and golf its proof. --- src/topology/constructions.lean | 100 ++++++++++++++------------------ 1 file changed, 44 insertions(+), 56 deletions(-) diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index 757cb51a8ff06..f0401e581ef2a 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -950,122 +950,109 @@ lemma continuous.quotient_map' {t : setoid β} {f : α → β} (hf : continuous end quotient section pi -variables {ι : Type*} {π : ι → Type*} +variables {ι : Type*} {π : ι → Type*} {κ : Type*} + [topological_space α] [∀i, topological_space (π i)] {f : α → Πi:ι, π i} -@[continuity] -lemma continuous_pi [topological_space α] [∀i, topological_space (π i)] {f : α → Πi:ι, π i} - (h : ∀i, continuous (λa, f a i)) : continuous f := -continuous_infi_rng.2 $ assume i, continuous_induced_rng.2 $ h i +lemma continuous_pi_iff : continuous f ↔ ∀ i, continuous (λ a, f a i) := +by simp only [continuous_infi_rng, continuous_induced_rng] -@[continuity] -lemma continuous_apply [∀i, topological_space (π i)] (i : ι) : - continuous (λp:Πi, π i, p i) := +@[continuity] lemma continuous_pi (h : ∀ i, continuous (λ a, f a i)) : continuous f := +continuous_pi_iff.2 h + +@[continuity] lemma continuous_apply (i : ι) : continuous (λp:Πi, π i, p i) := continuous_infi_dom continuous_induced_dom @[continuity] -lemma continuous_apply_apply {κ : Type*} {ρ : κ → ι → Type*} - [∀ j i, topological_space (ρ j i)] (j : κ) (i : ι) : - continuous (λ p : (Π j, Π i, ρ j i), p j i) := +lemma continuous_apply_apply {ρ : κ → ι → Type*} [∀ j i, topological_space (ρ j i)] + (j : κ) (i : ι) : continuous (λ p : (Π j, Π i, ρ j i), p j i) := (continuous_apply i).comp (continuous_apply j) -lemma continuous_at_apply [∀i, topological_space (π i)] (i : ι) (x : Π i, π i) : - continuous_at (λ p : Π i, π i, p i) x := +lemma continuous_at_apply (i : ι) (x : Π i, π i) : continuous_at (λ p : Π i, π i, p i) x := (continuous_apply i).continuous_at -lemma filter.tendsto.apply [∀i, topological_space (π i)] {l : filter α} {f : α → Π i, π i} +lemma filter.tendsto.apply {l : filter β} {f : β → Π i, π i} {x : Π i, π i} (h : tendsto f l (𝓝 x)) (i : ι) : tendsto (λ a, f a i) l (𝓝 $ x i) := (continuous_at_apply i _).tendsto.comp h -lemma continuous_pi_iff [topological_space α] [∀ i, topological_space (π i)] {f : α → Π i, π i} : - continuous f ↔ ∀ i, continuous (λ y, f y i) := -iff.intro (λ h i, (continuous_apply i).comp h) continuous_pi - -lemma nhds_pi [t : ∀i, topological_space (π i)] {a : Πi, π i} : - 𝓝 a = pi (λ i, 𝓝 (a i)) := -calc 𝓝 a = (⨅i, @nhds _ (@topological_space.induced _ _ (λx:Πi, π i, x i) (t i)) a) : nhds_infi - ... = (⨅i, comap (λx, x i) (𝓝 (a i))) : by simp [nhds_induced] +lemma nhds_pi {a : Πi, π i} : 𝓝 a = pi (λ i, 𝓝 (a i)) := +by simp only [nhds_infi, nhds_induced, filter.pi] -lemma tendsto_pi_nhds [t : ∀i, topological_space (π i)] {f : α → Πi, π i} {g : Πi, π i} - {u : filter α} : +lemma tendsto_pi_nhds {f : β → Πi, π i} {g : Πi, π i} {u : filter β} : tendsto f u (𝓝 g) ↔ ∀ x, tendsto (λ i, f i x) u (𝓝 (g x)) := by rw [nhds_pi, filter.tendsto_pi] -lemma continuous_at_pi [∀ i, topological_space (π i)] [topological_space α] {f : α → Π i, π i} - {x : α} : +lemma continuous_at_pi {f : α → Π i, π i} {x : α} : continuous_at f x ↔ ∀ i, continuous_at (λ y, f y i) x := tendsto_pi_nhds -lemma filter.tendsto.update [∀i, topological_space (π i)] [decidable_eq ι] - {l : filter α} {f : α → Π i, π i} {x : Π i, π i} (hf : tendsto f l (𝓝 x)) (i : ι) - {g : α → π i} {xi : π i} (hg : tendsto g l (𝓝 xi)) : +lemma filter.tendsto.update [decidable_eq ι] + {l : filter β} {f : β → Π i, π i} {x : Π i, π i} (hf : tendsto f l (𝓝 x)) (i : ι) + {g : β → π i} {xi : π i} (hg : tendsto g l (𝓝 xi)) : tendsto (λ a, update (f a) i (g a)) l (𝓝 $ update x i xi) := tendsto_pi_nhds.2 $ λ j, by { rcases em (j = i) with rfl|hj; simp [*, hf.apply] } -lemma continuous_at.update [∀i, topological_space (π i)] [topological_space α] [decidable_eq ι] - {f : α → Π i, π i} {a : α} (hf : continuous_at f a) (i : ι) {g : α → π i} +lemma continuous_at.update [decidable_eq ι] {a : α} (hf : continuous_at f a) (i : ι) {g : α → π i} (hg : continuous_at g a) : continuous_at (λ a, update (f a) i (g a)) a := hf.update i hg -lemma continuous.update [∀i, topological_space (π i)] [topological_space α] [decidable_eq ι] - {f : α → Π i, π i} (hf : continuous f) (i : ι) {g : α → π i} (hg : continuous g) : +lemma continuous.update [decidable_eq ι] (hf : continuous f) (i : ι) {g : α → π i} + (hg : continuous g) : continuous (λ a, update (f a) i (g a)) := continuous_iff_continuous_at.2 $ λ x, hf.continuous_at.update i hg.continuous_at -/-- `function.update f i x` is continuous in `(f, x)`. -/ -@[continuity] lemma continuous_update [∀i, topological_space (π i)] [decidable_eq ι] (i : ι) : +/-- `update f i x` is continuous in `(f, x)`. -/ +@[continuity] lemma continuous_update [decidable_eq ι] (i : ι) : continuous (λ f : (Π j, π j) × π i, update f.1 i f.2) := continuous_fst.update i continuous_snd lemma filter.tendsto.fin_insert_nth {n} {π : fin (n + 1) → Type*} [Π i, topological_space (π i)] - (i : fin (n + 1)) {f : α → π i} {l : filter α} {x : π i} (hf : tendsto f l (𝓝 x)) - {g : α → Π j : fin n, π (i.succ_above j)} {y : Π j, π (i.succ_above j)} (hg : tendsto g l (𝓝 y)) : + (i : fin (n + 1)) {f : β → π i} {l : filter β} {x : π i} (hf : tendsto f l (𝓝 x)) + {g : β → Π j : fin n, π (i.succ_above j)} {y : Π j, π (i.succ_above j)} (hg : tendsto g l (𝓝 y)) : tendsto (λ a, i.insert_nth (f a) (g a)) l (𝓝 $ i.insert_nth x y) := tendsto_pi_nhds.2 (λ j, fin.succ_above_cases i (by simpa) (by simpa using tendsto_pi_nhds.1 hg) j) lemma continuous_at.fin_insert_nth {n} {π : fin (n + 1) → Type*} [Π i, topological_space (π i)] - [topological_space α] (i : fin (n + 1)) {f : α → π i} {a : α} (hf : continuous_at f a) + (i : fin (n + 1)) {f : α → π i} {a : α} (hf : continuous_at f a) {g : α → Π j : fin n, π (i.succ_above j)} (hg : continuous_at g a) : continuous_at (λ a, i.insert_nth (f a) (g a)) a := hf.fin_insert_nth i hg lemma continuous.fin_insert_nth {n} {π : fin (n + 1) → Type*} [Π i, topological_space (π i)] - [topological_space α] (i : fin (n + 1)) {f : α → π i} (hf : continuous f) + (i : fin (n + 1)) {f : α → π i} (hf : continuous f) {g : α → Π j : fin n, π (i.succ_above j)} (hg : continuous g) : continuous (λ a, i.insert_nth (f a) (g a)) := continuous_iff_continuous_at.2 $ λ a, hf.continuous_at.fin_insert_nth i hg.continuous_at -lemma is_open_set_pi [∀a, topological_space (π a)] {i : set ι} {s : Πa, set (π a)} - (hi : i.finite) (hs : ∀a∈i, is_open (s a)) : is_open (pi i s) := +lemma is_open_set_pi {i : set ι} {s : Πa, set (π a)} (hi : i.finite) (hs : ∀a∈i, is_open (s a)) : + is_open (pi i s) := by rw [pi_def]; exact (is_open_bInter hi $ assume a ha, (hs _ ha).preimage (continuous_apply _)) -lemma is_closed_set_pi [∀a, topological_space (π a)] {i : set ι} {s : Πa, set (π a)} - (hs : ∀a∈i, is_closed (s a)) : is_closed (pi i s) := +lemma is_closed_set_pi {i : set ι} {s : Πa, set (π a)} (hs : ∀a∈i, is_closed (s a)) : + is_closed (pi i s) := by rw [pi_def]; exact (is_closed_Inter $ λ a, is_closed_Inter $ λ ha, (hs _ ha).preimage (continuous_apply _)) -lemma mem_nhds_of_pi_mem_nhds {ι : Type*} {α : ι → Type*} [Π (i : ι), topological_space (α i)] - {I : set ι} {s : Π i, set (α i)} (a : Π i, α i) (hs : I.pi s ∈ 𝓝 a) {i : ι} (hi : i ∈ I) : +lemma mem_nhds_of_pi_mem_nhds {I : set ι} {s : Π i, set (π i)} (a : Π i, π i) (hs : I.pi s ∈ 𝓝 a) + {i : ι} (hi : i ∈ I) : s i ∈ 𝓝 (a i) := by { rw nhds_pi at hs, exact mem_of_pi_mem_pi hs hi } -lemma set_pi_mem_nhds [Π a, topological_space (π a)] {i : set ι} {s : Π a, set (π a)} +lemma set_pi_mem_nhds {i : set ι} {s : Π a, set (π a)} {x : Π a, π a} (hi : i.finite) (hs : ∀ a ∈ i, s a ∈ 𝓝 (x a)) : pi i s ∈ 𝓝 x := by { rw [pi_def, bInter_mem hi], exact λ a ha, (continuous_apply a).continuous_at (hs a ha) } -lemma set_pi_mem_nhds_iff {α : ι → Type*} [Π (i : ι), topological_space (α i)] - {I : set ι} (hI : I.finite) {s : Π i, set (α i)} (a : Π i, α i) : +lemma set_pi_mem_nhds_iff {I : set ι} (hI : I.finite) {s : Π i, set (π i)} (a : Π i, π i) : I.pi s ∈ 𝓝 a ↔ ∀ (i : ι), i ∈ I → s i ∈ 𝓝 (a i) := by { rw [nhds_pi, pi_mem_pi_iff hI], apply_instance } -lemma interior_pi_set {α : ι → Type*} [Π i, topological_space (α i)] - {I : set ι} (hI : I.finite) {s : Π i, set (α i)} : +lemma interior_pi_set {I : set ι} (hI : I.finite) {s : Π i, set (π i)} : interior (pi I s) = I.pi (λ i, interior (s i)) := by { ext a, simp only [set.mem_pi, mem_interior_iff_mem_nhds, set_pi_mem_nhds_iff hI] } -lemma exists_finset_piecewise_mem_of_mem_nhds [decidable_eq ι] [Π i, topological_space (π i)] +lemma exists_finset_piecewise_mem_of_mem_nhds [decidable_eq ι] {s : set (Π a, π a)} {x : Π a, π a} (hs : s ∈ 𝓝 x) (y : Π a, π a) : ∃ I : finset ι, I.piecewise x y ∈ s := begin @@ -1075,7 +1062,7 @@ begin simpa [finset.mem_coe.1 hi] using mem_of_mem_nhds (htx i) end -lemma pi_eq_generate_from [∀a, topological_space (π a)] : +lemma pi_eq_generate_from : Pi.topological_space = generate_from {g | ∃(s:Πa, set (π a)) (i : finset ι), (∀a∈i, is_open (s a)) ∧ g = pi ↑i s} := le_antisymm @@ -1083,7 +1070,7 @@ le_antisymm (le_infi $ assume a s ⟨t, ht, s_eq⟩, generate_open.basic _ $ ⟨update (λa, univ) a t, {a}, by simpa using ht, s_eq ▸ by ext f; simp [set.pi]⟩) -lemma pi_generate_from_eq {g : Πa, set (set (π a))} : +lemma pi_generate_from_eq {π : ι → Type*} {g : Πa, set (set (π a))} : @Pi.topological_space ι π (λa, generate_from (g a)) = generate_from {t | ∃(s:Πa, set (π a)) (i : finset ι), (∀a∈i, s a ∈ g a) ∧ t = pi ↑i s} := let G := {t | ∃(s:Πa, set (π a)) (i : finset ι), (∀a∈i, s a ∈ g a) ∧ t = pi ↑i s} in @@ -1099,7 +1086,8 @@ begin exact assume s hs, generate_open.basic _ ⟨update (λa, univ) a s, {a}, by simp [hs]⟩ } end -lemma pi_generate_from_eq_finite {g : Πa, set (set (π a))} [finite ι] (hg : ∀a, ⋃₀ g a = univ) : +lemma pi_generate_from_eq_finite {π : ι → Type*} {g : Πa, set (set (π a))} [finite ι] + (hg : ∀a, ⋃₀ g a = univ) : @Pi.topological_space ι π (λa, generate_from (g a)) = generate_from {t | ∃(s:Πa, set (π a)), (∀a, s a ∈ g a) ∧ t = pi univ s} := begin @@ -1125,7 +1113,7 @@ endowed with a family of maps `f i : X → π i` for every `i : ι`, hence induc map `g : X → Π i, π i`. This lemma shows that infimum of the topologies on `X` induced by the `f i` as `i : ι` varies is simply the topology on `X` induced by `g : X → Π i, π i` where `Π i, π i` is endowed with the usual product topology. -/ -lemma inducing_infi_to_pi {X : Type*} [∀ i, topological_space (π i)] (f : Π i, X → π i) : +lemma inducing_infi_to_pi {X : Type*} (f : Π i, X → π i) : @inducing X (Π i, π i) (⨅ i, induced (f i) infer_instance) _ (λ x i, f i x) := begin constructor, @@ -1135,7 +1123,7 @@ begin erw induced_compose, end -variables [finite ι] [∀ i, topological_space (π i)] [∀ i, discrete_topology (π i)] +variables [finite ι] [∀ i, discrete_topology (π i)] /-- A finite product of discrete spaces is discrete. -/ instance Pi.discrete_topology : discrete_topology (Π i, π i) := From a5ed4088012701ca99edec5f9a042997c35c1197 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Tue, 20 Sep 2022 16:43:27 +0000 Subject: [PATCH 333/828] =?UTF-8?q?feat(group=5Ftheory/coset):=20Construct?= =?UTF-8?q?=20the=20embedding=20`H=20=E2=A7=B8=20=E2=A8=85=20i,=20f=20i=20?= =?UTF-8?q?=20=E2=86=92=20=CE=A0=20i,=20H=20=E2=A7=B8=20f=20i`=20(#16560)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR constructs the embedding `H ⧸ ⨅ i, f i → Π i, H ⧸ f i`, which is needed for #16393. Co-authored-by: Thomas Browning --- src/group_theory/coset.lean | 52 ++++++++++++++++++++++++++++--------- 1 file changed, 40 insertions(+), 12 deletions(-) diff --git a/src/group_theory/coset.lean b/src/group_theory/coset.lean index 7abe2d26fd57f..639d5237d5555 100644 --- a/src/group_theory/coset.lean +++ b/src/group_theory/coset.lean @@ -491,18 +491,46 @@ noncomputable def quotient_equiv_prod_of_le (h_le : s ≤ t) : α ⧸ s ≃ (α ⧸ t) × (t ⧸ s.subgroup_of t) := quotient_equiv_prod_of_le' h_le quotient.out' quotient.out_eq' -/-- If `K ≤ L`, then there is an embedding `K ⧸ (H.subgroup_of K) ↪ L ⧸ (H.subgroup_of L)`. -/ -@[to_additive "If `K ≤ L`, then there is an embedding - `K ⧸ (H.add_subgroup_of K) ↪ L ⧸ (H.add_subgroup_of L)`."] -def quotient_subgroup_of_embedding_of_le (H : subgroup α) {K L : subgroup α} (h : K ≤ L) : - K ⧸ (H.subgroup_of K) ↪ L ⧸ (H.subgroup_of L) := -{ to_fun := quotient.map' (set.inclusion h) (λ a b, by { simp [left_rel_apply], exact id }), - inj' := begin - refine quotient.ind₂' (λ a b, _), - refine λ h, (quotient.eq'.mpr ∘ left_rel_apply.mpr) _, - have := left_rel_apply.mp (quotient.eq'.mp h), - exact this, - end } +/-- If `s ≤ t`, then there is an embedding `s ⧸ H.subgroup_of s ↪ t ⧸ H.subgroup_of t`. -/ +@[to_additive "If `s ≤ t`, then there is an embedding + `s ⧸ H.add_subgroup_of s ↪ t ⧸ H.add_subgroup_of t`."] +def quotient_subgroup_of_embedding_of_le (H : subgroup α) (h : s ≤ t) : + s ⧸ H.subgroup_of s ↪ t ⧸ H.subgroup_of t := +{ to_fun := quotient.map' (inclusion h) (λ a b, by { simp_rw left_rel_eq, exact id }), + inj' := quotient.ind₂' $ by + { intros a b h, simpa only [quotient.map'_mk', quotient_group.eq'] using h } } + +@[simp, to_additive] +lemma quotient_subgroup_of_embedding_of_le_apply_mk (H : subgroup α) (h : s ≤ t) (g : s) : + quotient_subgroup_of_embedding_of_le H h (quotient_group.mk g) = + quotient_group.mk (inclusion h g) := +rfl + +/-- If `s ≤ t`, then there is a map `H ⧸ s.subgroup_of H → H ⧸ t.subgroup_of H`. -/ +@[to_additive "If `s ≤ t`, then there is an map + `H ⧸ s.add_subgroup_of H → H ⧸ t.add_subgroup_of H`."] +def quotient_subgroup_of_map_of_le (H : subgroup α) (h : s ≤ t) : + H ⧸ s.subgroup_of H → H ⧸ t.subgroup_of H := +quotient.map' id (λ a b, by { simp_rw [left_rel_eq], apply h }) + +@[simp, to_additive] +lemma quotient_subgroup_of_map_of_le_apply_mk (H : subgroup α) (h : s ≤ t) (g : H) : + quotient_subgroup_of_map_of_le H h (quotient_group.mk g) = quotient_group.mk g := +rfl + +/-- The natural embedding `H ⧸ (⨅ i, f i).subgroup_of H ↪ Π i, H ⧸ (f i).subgroup_of H`. -/ +@[to_additive "There is an embedding + `H ⧸ (⨅ i, f i).add_subgroup_of H) ↪ Π i, H ⧸ (f i).add_subgroup_of H`."] +def quotient_infi_embedding {ι : Type*} (f : ι → subgroup α) (H : subgroup α) : + H ⧸ (⨅ i, f i).subgroup_of H ↪ Π i, H ⧸ (f i).subgroup_of H := +{ to_fun := λ q i, quotient_subgroup_of_map_of_le H (infi_le f i) q, + inj' := quotient.ind₂' $ by simp_rw [funext_iff, quotient_subgroup_of_map_of_le_apply_mk, + quotient_group.eq', mem_subgroup_of, mem_infi, imp_self, forall_const] } + +@[simp, to_additive] lemma quotient_infi_embedding_apply_mk + {ι : Type*} (f : ι → subgroup α) (H : subgroup α) (g : H) (i : ι) : + quotient_infi_embedding f H (quotient_group.mk g) i = quotient_group.mk g := +rfl @[to_additive] lemma card_eq_card_quotient_mul_card_subgroup [fintype α] (s : subgroup α) [fintype s] [decidable_pred (λ a, a ∈ s)] : From a547b3b380ba689c0f2dc32036ab802813c62859 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 20 Sep 2022 19:59:00 +0000 Subject: [PATCH 334/828] feat(order/bounded_order): distrib_lattice instances for with_top and with_bot (#16569) --- src/order/bounded_order.lean | 25 +++++++++++++++++++++++++ 1 file changed, 25 insertions(+) diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index 629e1d1335d0d..ede566f7fef49 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -760,6 +760,20 @@ lemma coe_inf [semilattice_inf α] (a b : α) : ((a ⊓ b : α) : with_bot α) = instance [lattice α] : lattice (with_bot α) := { ..with_bot.semilattice_sup, ..with_bot.semilattice_inf } +instance [distrib_lattice α] : distrib_lattice (with_bot α) := +{ le_sup_inf := λ o₁ o₂ o₃, + match o₁, o₂, o₃ with + | ⊥, ⊥, ⊥ := le_rfl + | ⊥, ⊥, (a₁ : α) := le_rfl + | ⊥, (a₁ : α), ⊥ := le_rfl + | ⊥, (a₁ : α), (a₃ : α) := le_rfl + | (a₁ : α), ⊥, ⊥ := inf_le_left + | (a₁ : α), ⊥, (a₃ : α) := inf_le_left + | (a₁ : α), (a₂ : α), ⊥ := inf_le_right + | (a₁ : α), (a₂ : α), (a₃ : α) := coe_le_coe.mpr le_sup_inf + end, + ..with_bot.lattice } + instance decidable_le [has_le α] [@decidable_rel α (≤)] : @decidable_rel (with_bot α) (≤) | none x := is_true $ λ a h, option.no_confusion h | (some x) (some y) := @@ -1195,6 +1209,17 @@ lemma coe_sup [semilattice_sup α] (a b : α) : ((a ⊔ b : α) : with_top α) = instance [lattice α] : lattice (with_top α) := { ..with_top.semilattice_sup, ..with_top.semilattice_inf } +instance [distrib_lattice α] : distrib_lattice (with_top α) := +{ le_sup_inf := λ o₁ o₂ o₃, + match o₁, o₂, o₃ with + | ⊤, o₂, o₃ := le_rfl + | (a₁ : α), ⊤, ⊤ := le_rfl + | (a₁ : α), ⊤, (a₃ : α) := le_rfl + | (a₁ : α), (a₂ : α), ⊤ := le_rfl + | (a₁ : α), (a₂ : α), (a₃ : α) := coe_le_coe.mpr le_sup_inf + end, + ..with_top.lattice } + instance decidable_le [has_le α] [@decidable_rel α (≤)] : @decidable_rel (with_top α) (≤) := λ _ _, decidable_of_decidable_of_iff (with_bot.decidable_le _ _) (to_dual_le_to_dual_iff) From 236a0af465b94c00b97ab76d08c706d6f84e6b82 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 21 Sep 2022 06:16:26 +0000 Subject: [PATCH 335/828] =?UTF-8?q?feat(topology/algebra/order):=20a=20lin?= =?UTF-8?q?ear=20order=20is=20T=E2=82=85=20(#16540)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../intervals/ord_connected_component.lean | 189 ++++++++++++++++++ src/topology/algebra/order/basic.lean | 6 - src/topology/algebra/order/t5.lean | 92 +++++++++ 3 files changed, 281 insertions(+), 6 deletions(-) create mode 100644 src/data/set/intervals/ord_connected_component.lean create mode 100644 src/topology/algebra/order/t5.lean diff --git a/src/data/set/intervals/ord_connected_component.lean b/src/data/set/intervals/ord_connected_component.lean new file mode 100644 index 0000000000000..7d2fa5011140d --- /dev/null +++ b/src/data/set/intervals/ord_connected_component.lean @@ -0,0 +1,189 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import data.set.intervals.ord_connected + +/-! +# Order connected components of a set + +In this file we define `set.ord_connected_component s x` to be the set of `y` such that +`set.interval x y ⊆ s` and prove some basic facts about this definition. At the moment of writing, +this construction is used only to prove that any linear order with order topology is a T₅ space, +so we only add API needed for this lemma. +-/ + +open function order_dual +open_locale interval + +namespace set + +variables {α : Type*} [linear_order α] {s t : set α} {x y z : α} + +/-- Order-connected component of a point `x` in a set `s`. It is defined as the set of `y` such that +`set.interval x y ⊆ s`. Note that it is empty if and only if `x ∉ s`. -/ +def ord_connected_component (s : set α) (x : α) : set α := {y | [x, y] ⊆ s} + +lemma mem_ord_connected_component : y ∈ ord_connected_component s x ↔ [x, y] ⊆ s := iff.rfl + +lemma dual_ord_connected_component : + ord_connected_component (of_dual ⁻¹' s) (to_dual x) = of_dual ⁻¹' (ord_connected_component s x) := +ext $ to_dual.surjective.forall.2 $ λ x, + by { rw [mem_ord_connected_component, dual_interval], refl } + +lemma ord_connected_component_subset : ord_connected_component s x ⊆ s := +λ y hy, hy right_mem_interval + +lemma subset_ord_connected_component {t} [h : ord_connected s] (hs : x ∈ s) (ht : s ⊆ t) : + s ⊆ ord_connected_component t x := +λ y hy, (h.interval_subset hs hy).trans ht + +@[simp] lemma self_mem_ord_connected_component : x ∈ ord_connected_component s x ↔ x ∈ s := +by rw [mem_ord_connected_component, interval_self, singleton_subset_iff] + +@[simp] lemma nonempty_ord_connected_component : (ord_connected_component s x).nonempty ↔ x ∈ s := +⟨λ ⟨y, hy⟩, hy $ left_mem_interval, λ h, ⟨x, self_mem_ord_connected_component.2 h⟩⟩ + +@[simp] lemma ord_connected_component_eq_empty : ord_connected_component s x = ∅ ↔ x ∉ s := +by rw [← not_nonempty_iff_eq_empty, nonempty_ord_connected_component] + +@[simp] lemma ord_connected_component_empty : ord_connected_component ∅ x = ∅ := +ord_connected_component_eq_empty.2 (not_mem_empty x) + +@[simp] lemma ord_connected_component_univ : ord_connected_component univ x = univ := +by simp [ord_connected_component] + +lemma ord_connected_component_inter (s t : set α) (x : α) : + ord_connected_component (s ∩ t) x = ord_connected_component s x ∩ ord_connected_component t x := +by simp [ord_connected_component, set_of_and] + +lemma mem_ord_connected_component_comm : + y ∈ ord_connected_component s x ↔ x ∈ ord_connected_component s y := +by rw [mem_ord_connected_component, mem_ord_connected_component, interval_swap] + +lemma mem_ord_connected_component_trans (hxy : y ∈ ord_connected_component s x) + (hyz : z ∈ ord_connected_component s y) : z ∈ ord_connected_component s x := +calc [x, z] ⊆ [x, y] ∪ [y, z] : interval_subset_interval_union_interval +... ⊆ s : union_subset hxy hyz + +lemma ord_connected_component_eq (h : [x, y] ⊆ s) : + ord_connected_component s x = ord_connected_component s y := +ext $ λ z, ⟨mem_ord_connected_component_trans (mem_ord_connected_component_comm.2 h), + mem_ord_connected_component_trans h⟩ + +instance : ord_connected (ord_connected_component s x) := +ord_connected_of_interval_subset_left $ λ y hy z hz, (interval_subset_interval_left hz).trans hy + +/-- Projection from `s : set α` to `α` sending each order connected component of `s` to a single +point of this component. -/ +noncomputable def ord_connected_proj (s : set α) : s → α := +λ x : s, (nonempty_ord_connected_component.2 x.prop).some + +lemma ord_connected_proj_mem_ord_connected_component (s : set α) (x : s) : + ord_connected_proj s x ∈ ord_connected_component s x := +nonempty.some_mem _ + +lemma mem_ord_connected_component_ord_connected_proj (s : set α) (x : s) : + ↑x ∈ ord_connected_component s (ord_connected_proj s x) := +mem_ord_connected_component_comm.2 $ ord_connected_proj_mem_ord_connected_component s x + +@[simp] lemma ord_connected_component_ord_connected_proj (s : set α) (x : s) : + ord_connected_component s (ord_connected_proj s x) = ord_connected_component s x := +ord_connected_component_eq $ mem_ord_connected_component_ord_connected_proj _ _ + +@[simp] lemma ord_connected_proj_eq {x y : s} : + ord_connected_proj s x = ord_connected_proj s y ↔ [(x : α), y] ⊆ s := +begin + split; intro h, + { rw [← mem_ord_connected_component, ← ord_connected_component_ord_connected_proj, h, + ord_connected_component_ord_connected_proj, self_mem_ord_connected_component], + exact y.2 }, + { simp only [ord_connected_proj], + congr' 1, + exact ord_connected_component_eq h } +end + +/-- A set that intersects each order connected component of a set by a single point. Defined as the +range of `set.ord_connected_proj s`. -/ +def ord_connected_section (s : set α) : set α := range $ ord_connected_proj s + +lemma dual_ord_connected_section (s : set α) : + ord_connected_section (of_dual ⁻¹' s) = of_dual ⁻¹' (ord_connected_section s) := +begin + simp only [ord_connected_section, ord_connected_proj], + congr' 1 with x, simp only, congr' 1, + exact dual_ord_connected_component +end + +lemma ord_connected_section_subset : ord_connected_section s ⊆ s := +range_subset_iff.2 $ λ x, ord_connected_component_subset $ nonempty.some_mem _ + +lemma eq_of_mem_ord_connected_section_of_interval_subset (hx : x ∈ ord_connected_section s) + (hy : y ∈ ord_connected_section s) (h : [x, y] ⊆ s) : x = y := +begin + rcases hx with ⟨x, rfl⟩, rcases hy with ⟨y, rfl⟩, + exact ord_connected_proj_eq.2 (mem_ord_connected_component_trans + (mem_ord_connected_component_trans (ord_connected_proj_mem_ord_connected_component _ _) h) + (mem_ord_connected_component_ord_connected_proj _ _)) +end + +/-- Given two sets `s t : set α`, the set `set.order_separating_set s t` is the set of points that +belong both to some `set.ord_connected_component tᶜ x`, `x ∈ s`, and to some +`set.ord_connected_component sᶜ x`, `x ∈ t`. In the case of two disjoint closed sets, this is the +union of all open intervals $(a, b)$ such that their endpoints belong to different sets. -/ +def ord_separating_set (s t : set α) : set α := +(⋃ x ∈ s, ord_connected_component tᶜ x) ∩ (⋃ x ∈ t, ord_connected_component sᶜ x) + +lemma ord_separating_set_comm (s t : set α) : + ord_separating_set s t = ord_separating_set t s := +inter_comm _ _ + +lemma disjoint_left_ord_separating_set : disjoint s (ord_separating_set s t) := +disjoint.inter_right' _ $ disjoint_Union₂_right.2 $ λ x hx, disjoint_compl_right.mono_right $ + ord_connected_component_subset + +lemma disjoint_right_ord_separating_set : disjoint t (ord_separating_set s t) := +ord_separating_set_comm t s ▸ disjoint_left_ord_separating_set + +lemma dual_ord_separating_set : + ord_separating_set (of_dual ⁻¹' s) (of_dual ⁻¹' t) = of_dual ⁻¹' (ord_separating_set s t) := +by simp only [ord_separating_set, mem_preimage, ← to_dual.surjective.Union_comp, of_dual_to_dual, + dual_ord_connected_component, ← preimage_compl, preimage_inter, preimage_Union] + +/-- An auxiliary neighborhood that will be used in the proof of `order_topology.t5_space`. -/ +def ord_t5_nhd (s t : set α) : set α := +⋃ x ∈ s, ord_connected_component (tᶜ ∩ (ord_connected_section $ ord_separating_set s t)ᶜ) x + +lemma disjoint_ord_t5_nhd : disjoint (ord_t5_nhd s t) (ord_t5_nhd t s) := +begin + rintro x ⟨hx₁, hx₂⟩, + rcases mem_Union₂.1 hx₁ with ⟨a, has, ha⟩, clear hx₁, + rcases mem_Union₂.1 hx₂ with ⟨b, hbt, hb⟩, clear hx₂, + rw [mem_ord_connected_component, subset_inter_iff] at ha hb, + wlog hab : a ≤ b := le_total a b using [a b s t, b a t s] tactic.skip, + rotate, from λ h₁ h₂ h₃ h₄, this h₂ h₁ h₄ h₃, + cases ha with ha ha', cases hb with hb hb', + have hsub : [a, b] ⊆ (ord_separating_set s t).ord_connected_sectionᶜ, + { rw [ord_separating_set_comm, interval_swap] at hb', + calc [a, b] ⊆ [a, x] ∪ [x, b] : interval_subset_interval_union_interval + ... ⊆ (ord_separating_set s t).ord_connected_sectionᶜ : union_subset ha' hb' }, + clear ha' hb', + cases le_total x a with hxa hax, + { exact hb (Icc_subset_interval' ⟨hxa, hab⟩) has }, + cases le_total b x with hbx hxb, + { exact ha (Icc_subset_interval ⟨hab, hbx⟩) hbt }, + have : x ∈ ord_separating_set s t, + { exact ⟨mem_Union₂.2 ⟨a, has, ha⟩, mem_Union₂.2 ⟨b, hbt, hb⟩⟩ }, + lift x to ord_separating_set s t using this, + suffices : ord_connected_component (ord_separating_set s t) x ⊆ [a, b], + from hsub (this $ ord_connected_proj_mem_ord_connected_component _ _) (mem_range_self _), + rintros y (hy : [↑x, y] ⊆ ord_separating_set s t), + rw [interval_of_le hab, mem_Icc, ← not_lt, ← not_lt], + exact ⟨λ hya, disjoint_left.1 disjoint_left_ord_separating_set has + (hy $ Icc_subset_interval' ⟨hya.le, hax⟩), + λ hyb, disjoint_left.1 disjoint_right_ord_separating_set hbt + (hy $ Icc_subset_interval ⟨hxb, hyb.le⟩)⟩ +end + +end set diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index 66bea3cc3f564..480c85d277af1 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -1120,12 +1120,6 @@ lemma dense_iff_exists_between [densely_ordered α] [nontrivial α] {s : set α} dense s ↔ ∀ a b, a < b → ∃ c ∈ s, a < c ∧ c < b := ⟨λ h a b hab, h.exists_between hab, dense_of_exists_between⟩ -@[priority 100] -- see Note [lower instance priority] -instance order_topology.t3_space : t3_space α := -{ to_regular_space := regular_space.of_exists_mem_nhds_is_closed_subset $ - λ a s hs, let ⟨b, c, ha, hmem, hs⟩ := exists_Icc_mem_subset_of_mem_nhds hs - in ⟨Icc b c, hmem, is_closed_Icc, hs⟩ } - /-- A set is a neighborhood of `a` if and only if it contains an interval `(l, u)` containing `a`, provided `a` is neither a bottom element nor a top element. -/ lemma mem_nhds_iff_exists_Ioo_subset' {a : α} {s : set α} (hl : ∃ l, l < a) (hu : ∃ u, a < u) : diff --git a/src/topology/algebra/order/t5.lean b/src/topology/algebra/order/t5.lean new file mode 100644 index 0000000000000..1faf616de80a7 --- /dev/null +++ b/src/topology/algebra/order/t5.lean @@ -0,0 +1,92 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import topology.algebra.order.basic +import data.set.intervals.ord_connected_component + +/-! +# Linear order is a completely normal Hausdorff topological space + +In this file we prove that a linear order with order topology is a completely normal Hausdorff +topological space. +-/ + +open filter set function order_dual +open_locale topological_space filter interval + +variables {X : Type*} [linear_order X] [topological_space X] [order_topology X] + {a b c : X} {s t : set X} + +namespace set + +@[simp] lemma ord_connected_component_mem_nhds : + ord_connected_component s a ∈ 𝓝 a ↔ s ∈ 𝓝 a := +begin + refine ⟨λ h, mem_of_superset h ord_connected_component_subset, λ h, _⟩, + rcases exists_Icc_mem_subset_of_mem_nhds h with ⟨b, c, ha, ha', hs⟩, + exact mem_of_superset ha' (subset_ord_connected_component ha hs) +end + +lemma compl_section_ord_separating_set_mem_nhds_within_Ici (hd : disjoint s (closure t)) + (ha : a ∈ s) : + (ord_connected_section $ ord_separating_set s t)ᶜ ∈ 𝓝[≥] a := +begin + have hmem : tᶜ ∈ 𝓝[≥] a, + { refine mem_nhds_within_of_mem_nhds _, + rw [← mem_interior_iff_mem_nhds, interior_compl], + exact disjoint_left.1 hd ha }, + rcases exists_Icc_mem_subset_of_mem_nhds_within_Ici hmem with ⟨b, hab, hmem', hsub⟩, + by_cases H : disjoint (Icc a b) (ord_connected_section $ ord_separating_set s t), + { exact mem_of_superset hmem' (disjoint_left.1 H) }, + { simp only [set.disjoint_left, not_forall, not_not] at H, + rcases H with ⟨c, ⟨hac, hcb⟩, hc⟩, + have hsub' : Icc a b ⊆ ord_connected_component tᶜ a, + from subset_ord_connected_component (left_mem_Icc.2 hab) hsub, + replace hac : a < c := hac.lt_of_ne (ne.symm $ ne_of_mem_of_not_mem hc $ disjoint_left.1 + (disjoint_left_ord_separating_set.mono_right ord_connected_section_subset) ha), + refine mem_of_superset (Ico_mem_nhds_within_Ici (left_mem_Ico.2 hac)) (λ x hx hx', _), + refine hx.2.ne (eq_of_mem_ord_connected_section_of_interval_subset hx' hc _), + refine subset_inter (subset_Union₂_of_subset a ha _) _, + { exact ord_connected.interval_subset infer_instance (hsub' ⟨hx.1, hx.2.le.trans hcb⟩) + (hsub' ⟨hac.le, hcb⟩) }, + { rcases mem_Union₂.1 (ord_connected_section_subset hx').2 with ⟨y, hyt, hxy⟩, + refine subset_Union₂_of_subset y hyt (ord_connected.interval_subset infer_instance hxy _), + refine subset_ord_connected_component left_mem_interval hxy _, + suffices : c < y, + { rw [interval_of_ge (hx.2.trans this).le], + exact ⟨hx.2.le, this.le⟩ }, + refine lt_of_not_le (λ hyc, _), + have hya : y < a, from not_le.1 (λ hay, hsub ⟨hay, hyc.trans hcb⟩ hyt), + exact hxy (Icc_subset_interval ⟨hya.le, hx.1⟩) ha } } +end + +lemma compl_section_ord_separating_set_mem_nhds_within_Iic (hd : disjoint s (closure t)) + (ha : a ∈ s) : (ord_connected_section $ ord_separating_set s t)ᶜ ∈ 𝓝[≤] a := +have hd' : disjoint (of_dual ⁻¹' s) (closure $ of_dual ⁻¹' t) := hd, +have ha' : to_dual a ∈ of_dual ⁻¹' s := ha, +by simpa only [dual_ord_separating_set, dual_ord_connected_section] + using compl_section_ord_separating_set_mem_nhds_within_Ici hd' ha' + +lemma compl_section_ord_separating_set_mem_nhds (hd : disjoint s (closure t)) (ha : a ∈ s) : + (ord_connected_section $ ord_separating_set s t)ᶜ ∈ 𝓝 a := +begin + rw [← nhds_left_sup_nhds_right, mem_sup], + exact ⟨compl_section_ord_separating_set_mem_nhds_within_Iic hd ha, + compl_section_ord_separating_set_mem_nhds_within_Ici hd ha⟩ +end + +lemma ord_t5_nhd_mem_nhds_set (hd : disjoint s (closure t)) : ord_t5_nhd s t ∈ 𝓝ˢ s := +bUnion_mem_nhds_set $ λ x hx, ord_connected_component_mem_nhds.2 $ + inter_mem (by { rw [← mem_interior_iff_mem_nhds, interior_compl], exact disjoint_left.1 hd hx }) + (compl_section_ord_separating_set_mem_nhds hd hx) + +end set + +open set + +/-- A linear order with order topology is a completely normal Hausdorff topological space. -/ +@[priority 100] instance order_topology.t5_space : t5_space X := +⟨λ s t h₁ h₂, filter.disjoint_iff.2 ⟨ord_t5_nhd s t, ord_t5_nhd_mem_nhds_set h₂, ord_t5_nhd t s, + ord_t5_nhd_mem_nhds_set h₁.symm, disjoint_ord_t5_nhd⟩⟩ From 30553f13e1dca90d5018a45b273d27d295fd1262 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Wed, 21 Sep 2022 08:15:14 +0000 Subject: [PATCH 336/828] feat(algebra/hom/centroid): Introduce the centroid of a (non-unital, non-associative) semiring (#12746) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the centroid of a (non-unital, non-associative) semiring and shows that it forms a semiring. The centroid of a (non-unital, non-associative) ring is a ring. When the ring is prime, the centroid is commutative. Co-authored-by: Yaël Dillies Co-authored-by: Christopher Hoskin Co-authored-by: Christopher Hoskin --- docs/references.bib | 15 ++ src/algebra/hom/centroid.lean | 316 ++++++++++++++++++++++++++++++++++ 2 files changed, 331 insertions(+) create mode 100644 src/algebra/hom/centroid.lean diff --git a/docs/references.bib b/docs/references.bib index 56e03ff295fe8..96d52b84954e8 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -1128,6 +1128,21 @@ @Article{ izhakian2016 author = {Zur Izhakian and Manfred Knebusch and Louis Rowen} } +@Book{ Jacobson1956, + author = {Jacobson, Nathan}, + title = {Structure of rings}, + fseries = {Colloquium Publications. American Mathematical Society}, + series = {Colloq. Publ., Am. Math. Soc.}, + issn = {0065-9258}, + volume = {37}, + year = {1956}, + publisher = {American Mathematical Society (AMS), Providence, RI}, + language = {English}, + keywords = {16-02}, + zbmath = {3121681}, + zbl = {0073.02002} +} + @Book{ james1999, author = {James, Ioan}, title = {Topologies and uniformities}, diff --git a/src/algebra/hom/centroid.lean b/src/algebra/hom/centroid.lean new file mode 100644 index 0000000000000..aca9317f908ab --- /dev/null +++ b/src/algebra/hom/centroid.lean @@ -0,0 +1,316 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Christopher Hoskin +-/ +import algebra.group.pi +import algebra.group_power.lemmas +import algebra.hom.group_instances + +/-! +# Centroid homomorphisms + +Let `A` be a (non unital, non associative) algebra. The centroid of `A` is the set of linear maps +`T` on `A` such that `T` commutes with left and right multiplication, that is to say, for all `a` +and `b` in `A`, +$$ +T(ab) = (Ta)b, T(ab) = a(Tb). +$$ +In mathlib we call elements of the centroid "centroid homomorphisms" (`centroid_hom`) in keeping +with `add_monoid_hom` etc. + +We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to +be satisfied by itself and all stricter types. + +## Types of morphisms + +* `centroid_hom`: Maps which preserve left and right multiplication. + +## Typeclasses + +* `centroid_hom_class` + +## References + +* [Jacobson, Structure of Rings][Jacobson1956] +* [McCrimmon, A taste of Jordan algebras][mccrimmon2004] + +## Tags + +centroid +-/ + +open function + +variables {F α : Type*} + +-- Making `centroid_hom` an old structure will allow the lemma `to_add_monoid_hom_eq_coe` +-- to be true by `rfl`. After upgrading to Lean 4, this should no longer be needed +-- because eta for structures should provide the same result. +set_option old_structure_cmd true + +/-- The type of centroid homomorphisms from `α` to `α`. -/ +structure centroid_hom (α : Type*) [non_unital_non_assoc_semiring α] extends α →+ α := +(map_mul_left' (a b : α) : to_fun (a * b) = a * to_fun b) +(map_mul_right' (a b : α) : to_fun (a * b) = to_fun a * b) + +attribute [nolint doc_blame] centroid_hom.to_add_monoid_hom + +/-- `centroid_hom_class F α` states that `F` is a type of centroid homomorphisms. + +You should extend this class when you extend `centroid_hom`. -/ +class centroid_hom_class (F : Type*) (α : out_param $ Type*) [non_unital_non_assoc_semiring α] + extends add_monoid_hom_class F α α := +(map_mul_left (f : F) (a b : α) : f (a * b) = a * f b) +(map_mul_right (f : F) (a b : α) : f (a * b) = f a * b) + +export centroid_hom_class (map_mul_left map_mul_right) + +instance [non_unital_non_assoc_semiring α] [centroid_hom_class F α] : + has_coe_t F (centroid_hom α) := +⟨λ f, { to_fun := f, map_mul_left' := map_mul_left f, map_mul_right' := map_mul_right f, + ..(f : α →+ α) }⟩ + +/-! ### Centroid homomorphisms -/ + +namespace centroid_hom + +section non_unital_non_assoc_semiring + +variables [non_unital_non_assoc_semiring α] + +instance : centroid_hom_class (centroid_hom α) α := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by { cases f, cases g, congr' }, + map_zero := λ f, f.map_zero', + map_add := λ f, f.map_add', + map_mul_left := λ f, f.map_mul_left', + map_mul_right := λ f, f.map_mul_right' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun` +directly. -/ +instance : has_coe_to_fun (centroid_hom α) (λ _, α → α) := fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe {f : centroid_hom α} : f.to_fun = (f : α → α) := rfl + +@[ext] lemma ext {f g : centroid_hom α} (h : ∀ a, f a = g a) : f = g := fun_like.ext f g h + +@[simp, norm_cast] lemma coe_to_add_monoid_hom (f : centroid_hom α) : ⇑(f : α →+ α) = f := rfl +@[simp] lemma to_add_monoid_hom_eq_coe (f : centroid_hom α) : f.to_add_monoid_hom = f := rfl + +lemma coe_to_add_monoid_hom_injective : injective (coe : centroid_hom α → α →+ α) := +λ f g h, ext $ λ a, by { have := fun_like.congr_fun h a, exact this } + +/-- Turn a centroid homomorphism into an additive monoid endomorphism. -/ +def to_End (f : centroid_hom α) : add_monoid.End α := (f : α →+ α) + +lemma to_End_injective : injective (centroid_hom.to_End : centroid_hom α → add_monoid.End α) := +coe_to_add_monoid_hom_injective + +/-- Copy of a `centroid_hom` with a new `to_fun` equal to the old one. Useful to fix +definitional equalities. -/ +protected def copy (f : centroid_hom α) (f' : α → α) (h : f' = f) : + centroid_hom α := +{ to_fun := f', + map_mul_left' := λ a b, by simp_rw [h, map_mul_left], + map_mul_right' := λ a b, by simp_rw [h, map_mul_right], + ..f.to_add_monoid_hom.copy f' $ by exact h } + +variables (α) + +/-- `id` as a `centroid_hom`. -/ +protected def id : centroid_hom α := +{ map_mul_left' := λ _ _, rfl, + map_mul_right' := λ _ _, rfl, + .. add_monoid_hom.id α } + +instance : inhabited (centroid_hom α) := ⟨centroid_hom.id α⟩ + +@[simp, norm_cast] lemma coe_id : ⇑(centroid_hom.id α) = id := rfl + +@[simp, norm_cast] lemma coe_to_add_monoid_hom_id : + (centroid_hom.id α : α →+ α) = add_monoid_hom.id α := rfl + +variables {α} + +@[simp] lemma id_apply (a : α) : centroid_hom.id α a = a := rfl + +/-- Composition of `centroid_hom`s as a `centroid_hom`. -/ +def comp (g f : centroid_hom α) : centroid_hom α := +{ map_mul_left' := λ a b, (congr_arg g $ f.map_mul_left' _ _).trans $ g.map_mul_left' _ _, + map_mul_right' := λ a b, (congr_arg g $ f.map_mul_right' _ _).trans $ g.map_mul_right' _ _, + .. g.to_add_monoid_hom.comp f.to_add_monoid_hom } + +@[simp, norm_cast] lemma coe_comp (g f : centroid_hom α) : ⇑(g.comp f) = g ∘ f := rfl +@[simp] lemma comp_apply (g f : centroid_hom α) (a : α) : g.comp f a = g (f a) := rfl +@[simp, norm_cast] lemma coe_comp_add_monoid_hom (g f : centroid_hom α) : + (g.comp f : α →+ α) = (g : α →+ α).comp f := rfl +@[simp] lemma comp_assoc (h g f : centroid_hom α) : (h.comp g).comp f = h.comp (g.comp f) := rfl +@[simp] lemma comp_id (f : centroid_hom α) : f.comp (centroid_hom.id α) = f := ext $ λ a, rfl +@[simp] lemma id_comp (f : centroid_hom α) : (centroid_hom.id α).comp f = f := ext $ λ a, rfl + +lemma cancel_right {g₁ g₂ f : centroid_hom α} (hf : surjective f) : + g₁.comp f = g₂.comp f ↔ g₁ = g₂ := +⟨λ h, ext $ hf.forall.2 $ fun_like.ext_iff.1 h, congr_arg _⟩ + +lemma cancel_left {g f₁ f₂ : centroid_hom α} (hg : injective g) : g.comp f₁ = g.comp f₂ ↔ f₁ = f₂ := +⟨λ h, ext $ λ a, hg $ by rw [←comp_apply, h, comp_apply], congr_arg _⟩ + +instance : has_zero (centroid_hom α) := +⟨{ map_mul_left' := λ a b, (mul_zero _).symm, + map_mul_right' := λ a b, (zero_mul _).symm, + ..(0 : α →+ α) }⟩ + +instance : has_one (centroid_hom α) := ⟨centroid_hom.id α⟩ + +instance : has_add (centroid_hom α) := +⟨λ f g, + { map_mul_left' := λ a b, by simp [map_mul_left, mul_add], + map_mul_right' := λ a b, by simp [map_mul_right, add_mul], + ..(f + g : α →+ α) } ⟩ + +instance : has_mul (centroid_hom α) := ⟨comp⟩ + +instance has_nsmul : has_smul ℕ (centroid_hom α) := +⟨λ n f, + { map_mul_left' := λ a b, + by { change n • f (a * b) = a * n • f b, rw [map_mul_left f, ←mul_smul_comm] }, + map_mul_right' := λ a b, + by { change n • f (a * b) = n • f a * b, rw [map_mul_right f, ←smul_mul_assoc] }, + .. (n • f : α →+ α) }⟩ + +instance has_npow_nat : has_pow (centroid_hom α) ℕ := +⟨λ f n, +{ map_mul_left' := λ a b, begin + induction n with n ih, + { simp }, + { rw pow_succ, + exact (congr_arg f.to_End ih).trans (f.map_mul_left' _ _) } + end, + map_mul_right' := λ a b, begin + induction n with n ih, + { simp }, + { rw pow_succ, + exact (congr_arg f.to_End ih).trans (f.map_mul_right' _ _) } + end, + ..(f.to_End ^ n : add_monoid.End α) }⟩ + +@[simp, norm_cast] lemma coe_zero : ⇑(0 : centroid_hom α) = 0 := rfl +@[simp, norm_cast] lemma coe_one : ⇑(1 : centroid_hom α) = id := rfl +@[simp, norm_cast] lemma coe_add (f g : centroid_hom α) : ⇑(f + g) = f + g := rfl +@[simp, norm_cast] lemma coe_mul (f g : centroid_hom α) : ⇑(f * g) = f ∘ g := rfl +-- Eligible for `dsimp` +@[simp, norm_cast, nolint simp_nf] lemma coe_nsmul (f : centroid_hom α) (n : ℕ) : + ⇑(n • f) = n • f := rfl + +@[simp] lemma zero_apply (a : α) : (0 : centroid_hom α) a = 0 := rfl +@[simp] lemma one_apply (a : α) : (1 : centroid_hom α) a = a := rfl +@[simp] lemma add_apply (f g : centroid_hom α) (a : α) : (f + g) a = f a + g a := rfl +@[simp] lemma mul_apply (f g : centroid_hom α) (a : α) : (f * g) a = f (g a) := rfl +-- Eligible for `dsimp` +@[simp, nolint simp_nf] +lemma nsmul_apply (f : centroid_hom α) (n : ℕ) (a : α) : (n • f) a = n • f a := rfl + +@[simp] lemma to_End_zero : (0 : centroid_hom α).to_End = 0 := rfl +@[simp] lemma to_End_add (x y : centroid_hom α) : (x + y).to_End = x.to_End + y.to_End := rfl +lemma to_End_nsmul (x : centroid_hom α) (n : ℕ) : (n • x).to_End = n • x.to_End := rfl + +-- cf.`add_monoid_hom.add_comm_monoid` +instance : add_comm_monoid (centroid_hom α) := +coe_to_add_monoid_hom_injective.add_comm_monoid _ to_End_zero to_End_add to_End_nsmul + +instance : has_nat_cast (centroid_hom α) := +{ nat_cast := λ n, n • 1 } + +@[simp, norm_cast] lemma coe_nat_cast (n : ℕ) : ⇑(n : centroid_hom α) = n • id := rfl + +lemma nat_cast_apply (n : ℕ) (m : α): + (n : centroid_hom α) m = n • m := rfl + +@[simp] lemma to_End_one : (1 : centroid_hom α).to_End = 1 := rfl +@[simp] lemma to_End_mul (x y : centroid_hom α) : (x * y).to_End = x.to_End * y.to_End := rfl +@[simp] lemma to_End_pow (x : centroid_hom α) (n : ℕ) : (x ^ n).to_End = x.to_End ^ n := +by { ext, refl } +@[simp, norm_cast] lemma to_End_nat_cast (n : ℕ) : (n : centroid_hom α).to_End = ↑n := rfl + +-- cf `add_monoid.End.semiring` +instance : semiring (centroid_hom α) := +to_End_injective.semiring _ to_End_zero to_End_one to_End_add to_End_mul + to_End_nsmul to_End_pow to_End_nat_cast + +lemma comp_mul_comm (T S : centroid_hom α) (a b : α) : (T ∘ S) (a * b) = (S ∘ T) (a * b) := +by rw [comp_app, map_mul_right, map_mul_left, ←map_mul_right, ←map_mul_left] + +end non_unital_non_assoc_semiring + +section non_unital_non_assoc_ring +variables [non_unital_non_assoc_ring α] + +/-- Negation of `centroid_hom`s as a `centroid_hom`. -/ +instance : has_neg (centroid_hom α) := +⟨λ f, + { map_mul_left' := by simp [map_mul_left], + map_mul_right' := by simp [map_mul_right], + .. (-f : α →+ α) }⟩ +instance : has_sub (centroid_hom α) := +⟨λ f g, +{ map_mul_left' := λ a b, by simp [map_mul_left, mul_sub], + map_mul_right' := λ a b, by simp [map_mul_right, sub_mul], + .. (f - g : α →+ α) }⟩ + +instance has_zsmul : has_smul ℤ (centroid_hom α) := +⟨λ n f, + { map_mul_left' := λ a b, + by { change n • f (a * b) = a * n • f b, rw [map_mul_left f, ←mul_smul_comm] }, + map_mul_right' := λ a b, + by { change n • f (a * b) = n • f a * b, rw [map_mul_right f, ←smul_mul_assoc] }, + .. (n • f : α →+ α) }⟩ + +instance : has_int_cast (centroid_hom α) := +{ int_cast := λ z, z • 1 } + +@[simp, norm_cast] lemma coe_int_cast (z : ℤ) : ⇑(z : centroid_hom α) = z • id := rfl + +lemma int_cast_apply (z : ℤ) (m : α) : + (z : centroid_hom α) m = z • m := rfl + +@[simp] lemma to_End_neg (x : centroid_hom α) : (-x).to_End = -x.to_End := rfl +@[simp] lemma to_End_sub (x y : centroid_hom α) : (x - y).to_End = x.to_End - y.to_End := rfl +lemma to_End_zsmul (x : centroid_hom α) (n : ℤ) : (n • x).to_End = n • x.to_End := rfl + +instance : add_comm_group (centroid_hom α) := +to_End_injective.add_comm_group _ to_End_zero to_End_add to_End_neg to_End_sub + to_End_nsmul to_End_zsmul + + +@[simp, norm_cast] lemma coe_neg (f : centroid_hom α) : ⇑(-f) = -f := rfl +@[simp, norm_cast] lemma coe_sub (f g : centroid_hom α) : ⇑(f - g) = f - g := rfl + +@[simp] lemma neg_apply (f : centroid_hom α) (a : α) : (-f) a = - f a := rfl +@[simp] lemma sub_apply (f g : centroid_hom α) (a : α) : (f - g) a = f a - g a := rfl + +@[simp, norm_cast] lemma to_End_int_cast (z : ℤ) : (z : centroid_hom α).to_End = ↑z := rfl + +instance : ring (centroid_hom α) := to_End_injective.ring _ to_End_zero to_End_one + to_End_add to_End_mul to_End_neg to_End_sub to_End_nsmul to_End_zsmul + to_End_pow to_End_nat_cast to_End_int_cast + +end non_unital_non_assoc_ring + +section non_unital_ring +variables [non_unital_ring α] + +/-- A prime associative ring has commutative centroid. -/ +@[reducible] -- See note [reducible non instances] +def comm_ring (h : ∀ a b : α, (∀ r : α, a * r * b = 0) → a = 0 ∨ b = 0) : + comm_ring (centroid_hom α) := +{ mul_comm := λ f g, begin + ext, + refine sub_eq_zero.1 ((or_self _).1 $ h _ _ $ λ r, _), + rw [mul_assoc, sub_mul, sub_eq_zero, ← map_mul_right, ← map_mul_right, coe_mul, coe_mul, + comp_mul_comm], + end, + ..centroid_hom.ring } + +end non_unital_ring +end centroid_hom From 9936c3dfc04e5876f4368aeb2e60f8d8358d095a Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 21 Sep 2022 08:15:15 +0000 Subject: [PATCH 337/828] feat(ring_theory/graded_algebra/basic): lemmas about rings graded by canonically ordered monoids (#16083) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When the `canonically_ordered_add_monoid` `has_ordered_sub` and is cancellative (equivalent to the `[contravariant_class ι ι (+) (≤)]` assumption): * `(decompose 𝒜 (a * b) n : A) = if i ≤ n then a * decompose 𝒜 b (n - i) else 0` for homogeneous `a` of degree `i` * `(decompose 𝒜 (a * b) n : A) = if i ≤ n then (decompose 𝒜 a (n - i)) * b else 0` for homogeneous `b` of degree `i` Examples of such monoids include ℕ and finitely supported functions to ℕ. Co-authored-by: Junyan Xu --- src/algebra/direct_sum/internal.lean | 131 +++++++++++++++++++--- src/algebra/direct_sum/ring.lean | 18 +-- src/ring_theory/graded_algebra/basic.lean | 103 +++++++++-------- 3 files changed, 183 insertions(+), 69 deletions(-) diff --git a/src/algebra/direct_sum/internal.lean b/src/algebra/direct_sum/internal.lean index 87dbb2e9fb016..2ceb49d35e645 100644 --- a/src/algebra/direct_sum/internal.lean +++ b/src/algebra/direct_sum/internal.lean @@ -137,28 +137,133 @@ instance gcomm_ring [add_comm_monoid ι] [comm_ring R] [set_like σ R] end set_like +namespace direct_sum + +section coe + +variables [semiring R] [set_like σ R] [add_submonoid_class σ R] (A : ι → σ) + /-- The canonical ring isomorphism between `⨁ i, A i` and `R`-/ -def direct_sum.coe_ring_hom [add_monoid ι] [semiring R] [set_like σ R] - [add_submonoid_class σ R] (A : ι → σ) [set_like.graded_monoid A] : (⨁ i, A i) →+* R := +def coe_ring_hom [add_monoid ι] [set_like.graded_monoid A] : (⨁ i, A i) →+* R := direct_sum.to_semiring (λ i, add_submonoid_class.subtype (A i)) rfl (λ _ _ _ _, rfl) /-- The canonical ring isomorphism between `⨁ i, A i` and `R`-/ -@[simp] lemma direct_sum.coe_ring_hom_of [add_monoid ι] [semiring R] [set_like σ R] - [add_submonoid_class σ R] (A : ι → σ) [set_like.graded_monoid A] (i : ι) (x : A i) : - direct_sum.coe_ring_hom A (direct_sum.of (λ i, A i) i x) = x := +@[simp] lemma coe_ring_hom_of [add_monoid ι] [set_like.graded_monoid A] (i : ι) + (x : A i) : (coe_ring_hom A : _ →+* R) (of (λ i, A i) i x) = x := direct_sum.to_semiring_of _ _ _ _ _ -lemma direct_sum.coe_mul_apply [add_monoid ι] [semiring R] [set_like σ R] - [add_submonoid_class σ R] (A : ι → σ) [set_like.graded_monoid A] - [Π (i : ι) (x : A i), decidable (x ≠ 0)] (r r' : ⨁ i, A i) (i : ι) : - ((r * r') i : R) = - ∑ ij in (r.support ×ˢ r'.support).filter (λ ij : ι × ι, ij.1 + ij.2 = i), r ij.1 * r' ij.2 := +lemma coe_mul_apply [add_monoid ι] [set_like.graded_monoid A] + [Π (i : ι) (x : A i), decidable (x ≠ 0)] (r r' : ⨁ i, A i) (n : ι) : + ((r * r') n : R) = + ∑ ij in (r.support ×ˢ r'.support).filter (λ ij : ι × ι, ij.1 + ij.2 = n), r ij.1 * r' ij.2 := +begin + rw [mul_eq_sum_support_ghas_mul, dfinsupp.finset_sum_apply, add_submonoid_class.coe_finset_sum], + simp_rw [coe_of_apply, ←finset.sum_filter, set_like.coe_ghas_mul], +end + +lemma coe_mul_apply_eq_dfinsupp_sum [add_monoid ι] [set_like.graded_monoid A] + [Π (i : ι) (x : A i), decidable (x ≠ 0)] (r r' : ⨁ i, A i) (n : ι) : + ((r * r') n : R) = r.sum (λ i ri, r'.sum (λ j rj, if i + j = n then ri * rj else 0)) := +begin + simp only [mul_eq_dfinsupp_sum, dfinsupp.sum_apply], + iterate 2 { rw [dfinsupp.sum, add_submonoid_class.coe_finset_sum], congr, ext }, + dsimp only, split_ifs, + { subst h, rw of_eq_same, refl }, + { rw of_eq_of_ne _ _ _ _ h, refl }, +end + +lemma coe_of_mul_apply_aux [add_monoid ι] [set_like.graded_monoid A] {i : ι} + (r : A i) (r' : ⨁ i, A i) {j n : ι} (H : ∀ (x : ι), i + x = n ↔ x = j) : + ((of _ i r * r') n : R) = r * r' j := +begin + classical, + rw coe_mul_apply_eq_dfinsupp_sum, + apply (dfinsupp.sum_single_index _).trans, swap, + { simp_rw [zero_mem_class.coe_zero, zero_mul, if_t_t], exact dfinsupp.sum_zero }, + simp_rw [dfinsupp.sum, H, finset.sum_ite_eq'], + split_ifs, refl, + rw [dfinsupp.not_mem_support_iff.mp h, zero_mem_class.coe_zero, mul_zero], +end + +lemma coe_mul_of_apply_aux [add_monoid ι] [set_like.graded_monoid A] + (r : ⨁ i, A i) {i : ι} (r' : A i) {j n : ι} (H : ∀ (x : ι), x + i = n ↔ x = j) : + ((r * of _ i r') n : R) = r j * r' := +begin + classical, + rw [coe_mul_apply_eq_dfinsupp_sum, dfinsupp.sum_comm], + apply (dfinsupp.sum_single_index _).trans, swap, + { simp_rw [zero_mem_class.coe_zero, mul_zero, if_t_t], exact dfinsupp.sum_zero }, + simp_rw [dfinsupp.sum, H, finset.sum_ite_eq'], + split_ifs, refl, + rw [dfinsupp.not_mem_support_iff.mp h, zero_mem_class.coe_zero, zero_mul], +end + +lemma coe_of_mul_apply_add [add_left_cancel_monoid ι] [set_like.graded_monoid A] + {i : ι} (r : A i) (r' : ⨁ i, A i) (j : ι) : + ((of _ i r * r') (i + j) : R) = r * r' j := +coe_of_mul_apply_aux _ _ _ (λ x, ⟨λ h, add_left_cancel h, λ h, h ▸ rfl⟩) + +lemma coe_mul_of_apply_add [add_right_cancel_monoid ι] [set_like.graded_monoid A] + (r : ⨁ i, A i) {i : ι} (r' : A i) (j : ι) : + ((r * of _ i r') (j + i) : R) = r j * r' := +coe_mul_of_apply_aux _ _ _ (λ x, ⟨λ h, add_right_cancel h, λ h, h ▸ rfl⟩) + +end coe + +section canonically_ordered_add_monoid + +variables [semiring R] [set_like σ R] [add_submonoid_class σ R] (A : ι → σ) +variables [canonically_ordered_add_monoid ι] [set_like.graded_monoid A] + +lemma coe_of_mul_apply_of_not_le + {i : ι} (r : A i) (r' : ⨁ i, A i) (n : ι) + (h : ¬ i ≤ n) : ((of _ i r * r') n : R) = 0 := begin - rw [direct_sum.mul_eq_sum_support_ghas_mul, dfinsupp.finset_sum_apply, - add_submonoid_class.coe_finset_sum], - simp_rw [direct_sum.coe_of_apply, ←finset.sum_filter, set_like.coe_ghas_mul], + classical, + rw coe_mul_apply_eq_dfinsupp_sum, + apply (dfinsupp.sum_single_index _).trans, swap, + { simp_rw [zero_mem_class.coe_zero, zero_mul, if_t_t], exact dfinsupp.sum_zero }, + { rw [dfinsupp.sum, finset.sum_ite_of_false _ _ (λ x _ H, _), finset.sum_const_zero], + exact h ((self_le_add_right i x).trans_eq H) }, end +lemma coe_mul_of_apply_of_not_le + (r : ⨁ i, A i) {i : ι} (r' : A i) (n : ι) + (h : ¬ i ≤ n) : ((r * of _ i r') n : R) = 0 := +begin + classical, + rw [coe_mul_apply_eq_dfinsupp_sum, dfinsupp.sum_comm], + apply (dfinsupp.sum_single_index _).trans, swap, + { simp_rw [zero_mem_class.coe_zero, mul_zero, if_t_t], exact dfinsupp.sum_zero }, + { rw [dfinsupp.sum, finset.sum_ite_of_false _ _ (λ x _ H, _), finset.sum_const_zero], + exact h ((self_le_add_left i x).trans_eq H) }, +end + +variables [has_sub ι] [has_ordered_sub ι] [contravariant_class ι ι (+) (≤)] + +/- The following two lemmas only require the same hypotheses as `eq_tsub_iff_add_eq_of_le`, but we + state them for `canonically_ordered_add_monoid` + the above three typeclasses for convenience. -/ + +lemma coe_mul_of_apply_of_le (r : ⨁ i, A i) {i : ι} (r' : A i) (n : ι) + (h : i ≤ n) : ((r * of _ i r') n : R) = r (n - i) * r' := +coe_mul_of_apply_aux _ _ _ (λ x, (eq_tsub_iff_add_eq_of_le h).symm) + +lemma coe_of_mul_apply_of_le {i : ι} (r : A i) (r' : ⨁ i, A i) (n : ι) + (h : i ≤ n) : ((of _ i r * r') n : R) = r * r' (n - i) := +coe_of_mul_apply_aux _ _ _ (λ x, by rw [eq_tsub_iff_add_eq_of_le h, add_comm]) + +lemma coe_mul_of_apply (r : ⨁ i, A i) {i : ι} (r' : A i) (n : ι) [decidable (i ≤ n)] : + ((r * of _ i r') n : R) = if i ≤ n then r (n - i) * r' else 0 := +by { split_ifs, exacts [coe_mul_of_apply_of_le _ _ _ n h, coe_mul_of_apply_of_not_le _ _ _ n h] } + +lemma coe_of_mul_apply {i : ι} (r : A i) (r' : ⨁ i, A i) (n : ι) [decidable (i ≤ n)] : + ((of _ i r * r') n : R) = if i ≤ n then r * r' (n - i) else 0 := +by { split_ifs, exacts [coe_of_mul_apply_of_le _ _ _ n h, coe_of_mul_apply_of_not_le _ _ _ n h] } + +end canonically_ordered_add_monoid + +end direct_sum + /-! #### From `submodule`s -/ namespace submodule diff --git a/src/algebra/direct_sum/ring.lean b/src/algebra/direct_sum/ring.lean index b8c27d87566bc..5d47cb3600cae 100644 --- a/src/algebra/direct_sum/ring.lean +++ b/src/algebra/direct_sum/ring.lean @@ -265,21 +265,21 @@ by rw [list.of_fn_eq_map, of_list_dprod] open_locale big_operators +lemma mul_eq_dfinsupp_sum [Π (i : ι) (x : A i), decidable (x ≠ 0)] (a a' : ⨁ i, A i) : + a * a' = a.sum (λ i ai, a'.sum $ λ j aj, direct_sum.of _ _ $ graded_monoid.ghas_mul.mul ai aj) := +begin + change mul_hom _ a a' = _, + simpa only [mul_hom, to_add_monoid, dfinsupp.lift_add_hom_apply, dfinsupp.sum_add_hom_apply, + add_monoid_hom.dfinsupp_sum_apply, flip_apply, add_monoid_hom.dfinsupp_sum_add_hom_apply], +end + /-- A heavily unfolded version of the definition of multiplication -/ lemma mul_eq_sum_support_ghas_mul [Π (i : ι) (x : A i), decidable (x ≠ 0)] (a a' : ⨁ i, A i) : a * a' = ∑ ij in dfinsupp.support a ×ˢ dfinsupp.support a', direct_sum.of _ _ (graded_monoid.ghas_mul.mul (a ij.fst) (a' ij.snd)) := -begin - change direct_sum.mul_hom _ a a' = _, - dsimp [direct_sum.mul_hom, direct_sum.to_add_monoid, dfinsupp.lift_add_hom_apply], - simp only [dfinsupp.sum_add_hom_apply, dfinsupp.sum, dfinsupp.finset_sum_apply, - add_monoid_hom.coe_finset_sum, finset.sum_apply, add_monoid_hom.flip_apply, - add_monoid_hom.comp_hom_apply_apply, add_monoid_hom.comp_apply, - direct_sum.gmul_hom_apply_apply], - rw finset.sum_product, -end +by simp only [mul_eq_dfinsupp_sum, dfinsupp.sum, finset.sum_product] end semiring diff --git a/src/ring_theory/graded_algebra/basic.lean b/src/ring_theory/graded_algebra/basic.lean index 725c8a72f9271..da28a42dff2ea 100644 --- a/src/ring_theory/graded_algebra/basic.lean +++ b/src/ring_theory/graded_algebra/basic.lean @@ -105,59 +105,36 @@ end graded_ring section add_cancel_monoid -open direct_sum dfinsupp finset function +open direct_sum + +variables [decidable_eq ι] [semiring A] [set_like σ A] [add_submonoid_class σ A] (𝒜 : ι → σ) +variables {i j : ι} -lemma direct_sum.coe_decompose_mul_add_of_left_mem {ι σ A} - [decidable_eq ι] [add_left_cancel_monoid ι] [semiring A] - [set_like σ A] [add_submonoid_class σ A] (𝒜 : ι → σ) [graded_ring 𝒜] - {a b : A} {i j : ι} (a_mem : a ∈ 𝒜 i) : +namespace direct_sum + +lemma coe_decompose_mul_add_of_left_mem + [add_left_cancel_monoid ι] [graded_ring 𝒜] {a b : A} (a_mem : a ∈ 𝒜 i) : (decompose 𝒜 (a * b) (i + j) : A) = a * decompose 𝒜 b j := -begin - obtain rfl | ha := eq_or_ne a 0, - { simp }, - classical, - lift a to (𝒜 i) using a_mem, - erw [decompose_mul, coe_mul_apply, decompose_coe, support_of _ i a (λ r,by subst r; exact ha rfl), - singleton_product, map_filter, sum_map], - simp_rw [comp, embedding.coe_fn_mk, add_left_cancel_iff, filter_eq'], - refine dite (decompose 𝒜 b j = 0) (λ h, by simp [if_neg (not_mem_support_iff.mpr h), h]) (λ h, _), - erw [if_pos (mem_support_iff.mpr h), finset.sum_singleton, of_eq_same], - refl, -end - -lemma direct_sum.coe_decompose_mul_add_of_right_mem {ι σ A} - [decidable_eq ι] [add_right_cancel_monoid ι] [semiring A] - [set_like σ A] [add_submonoid_class σ A] (𝒜 : ι → σ) [graded_ring 𝒜] - {a b : A} {i j : ι} (b_mem : b ∈ 𝒜 j) : - (decompose 𝒜 (a * b) (i + j) : A) = (decompose 𝒜 a i) * b := -begin - obtain rfl | hb := eq_or_ne b 0, - { simp }, - classical, - lift b to (𝒜 j) using b_mem, - erw [decompose_mul, coe_mul_apply, decompose_coe, support_of _ j b (λ r,by subst r; exact hb rfl), - product_singleton, map_filter, sum_map], - simp_rw [comp, embedding.coe_fn_mk, add_right_cancel_iff, filter_eq'], - refine dite (decompose 𝒜 a i = 0) (λ h, by simp [if_neg (not_mem_support_iff.mpr h), h]) (λ h, _), - erw [if_pos (mem_support_iff.mpr h), finset.sum_singleton, of_eq_same], - refl, -end - -lemma direct_sum.decompose_mul_add_left {ι σ A} - [decidable_eq ι] [add_left_cancel_monoid ι] [semiring A] - [set_like σ A] [add_submonoid_class σ A] (𝒜 : ι → σ) [graded_ring 𝒜] - {i j : ι} (a : 𝒜 i) {b : A} : +by { lift a to 𝒜 i using a_mem, rw [decompose_mul, decompose_coe, coe_of_mul_apply_add] } + +lemma coe_decompose_mul_add_of_right_mem + [add_right_cancel_monoid ι] [graded_ring 𝒜] {a b : A} (b_mem : b ∈ 𝒜 j) : + (decompose 𝒜 (a * b) (i + j) : A) = decompose 𝒜 a i * b := +by { lift b to 𝒜 j using b_mem, rw [decompose_mul, decompose_coe, coe_mul_of_apply_add] } + +lemma decompose_mul_add_left + [add_left_cancel_monoid ι] [graded_ring 𝒜] (a : 𝒜 i) {b : A} : decompose 𝒜 (↑a * b) (i + j) = @graded_monoid.ghas_mul.mul ι (λ i, 𝒜 i) _ _ _ _ a (decompose 𝒜 b j) := -subtype.ext $ direct_sum.coe_decompose_mul_add_of_left_mem 𝒜 a.2 +subtype.ext $ coe_decompose_mul_add_of_left_mem 𝒜 a.2 -lemma direct_sum.decompose_mul_add_right {ι σ A} - [decidable_eq ι] [add_right_cancel_monoid ι] [semiring A] - [set_like σ A] [add_submonoid_class σ A] (𝒜 : ι → σ) [graded_ring 𝒜] - {i j : ι} {a : A} (b : 𝒜 j) : +lemma decompose_mul_add_right + [add_right_cancel_monoid ι] [graded_ring 𝒜] {a : A} (b : 𝒜 j) : decompose 𝒜 (a * ↑b) (i + j) = @graded_monoid.ghas_mul.mul ι (λ i, 𝒜 i) _ _ _ _ (decompose 𝒜 a i) b := -subtype.ext $ direct_sum.coe_decompose_mul_add_of_right_mem 𝒜 b.2 +subtype.ext $ coe_decompose_mul_add_of_right_mem 𝒜 b.2 + +end direct_sum end add_cancel_monoid @@ -228,7 +205,7 @@ end graded_algebra section canonical_order -open graded_ring set_like.graded_monoid direct_sum +open set_like.graded_monoid direct_sum variables [semiring A] [decidable_eq ι] variables [canonically_ordered_add_monoid ι] @@ -266,4 +243,36 @@ def graded_ring.proj_zero_ring_hom : A →+* A := simp only [add_mul, decompose_add, add_apply, add_mem_class.coe_add, ha, hb], }, end } +variables {a b : A} {n i : ι} + +namespace direct_sum + +lemma coe_decompose_mul_of_left_mem_of_not_le + (a_mem : a ∈ 𝒜 i) (h : ¬ i ≤ n) : (decompose 𝒜 (a * b) n : A) = 0 := +by { lift a to 𝒜 i using a_mem, rwa [decompose_mul, decompose_coe, coe_of_mul_apply_of_not_le] } + +lemma coe_decompose_mul_of_right_mem_of_not_le + (b_mem : b ∈ 𝒜 i) (h : ¬ i ≤ n) : (decompose 𝒜 (a * b) n : A) = 0 := +by { lift b to 𝒜 i using b_mem, rwa [decompose_mul, decompose_coe, coe_mul_of_apply_of_not_le] } + +variables [has_sub ι] [has_ordered_sub ι] [contravariant_class ι ι (+) (≤)] + +lemma coe_decompose_mul_of_left_mem_of_le + (a_mem : a ∈ 𝒜 i) (h : i ≤ n) : (decompose 𝒜 (a * b) n : A) = a * decompose 𝒜 b (n - i) := +by { lift a to 𝒜 i using a_mem, rwa [decompose_mul, decompose_coe, coe_of_mul_apply_of_le] } + +lemma coe_decompose_mul_of_right_mem_of_le + (b_mem : b ∈ 𝒜 i) (h : i ≤ n) : (decompose 𝒜 (a * b) n : A) = decompose 𝒜 a (n - i) * b := +by { lift b to 𝒜 i using b_mem, rwa [decompose_mul, decompose_coe, coe_mul_of_apply_of_le] } + +lemma coe_decompose_mul_of_left_mem (n) [decidable (i ≤ n)] (a_mem : a ∈ 𝒜 i) : + (decompose 𝒜 (a * b) n : A) = if i ≤ n then a * decompose 𝒜 b (n - i) else 0 := +by { lift a to 𝒜 i using a_mem, rwa [decompose_mul, decompose_coe, coe_of_mul_apply] } + +lemma coe_decompose_mul_of_right_mem (n) [decidable (i ≤ n)] (b_mem : b ∈ 𝒜 i) : + (decompose 𝒜 (a * b) n : A) = if i ≤ n then decompose 𝒜 a (n - i) * b else 0 := +by { lift b to 𝒜 i using b_mem, rwa [decompose_mul, decompose_coe, coe_mul_of_apply] } + +end direct_sum + end canonical_order From e2ce3fa08a9688e7af4d316d845119133e1394de Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 21 Sep 2022 08:15:16 +0000 Subject: [PATCH 338/828] feat(linear_algebra/affine_space/affine_subspace): `affine_subspace.subtype` (#16372) Define `affine_subspace.subtype`, the embedding map from an affine subspace to the ambient affine space as a bundled affine map, analogous to `submodule.subtype`. --- .../affine_space/affine_subspace.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index 19d7d019d5e01..648e9f5dca7be 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -407,6 +407,25 @@ rfl ↑(a +ᵥ b) = (a:V) +ᵥ (b:P) := rfl +/-- Embedding of an affine subspace to the ambient space, as an affine map. -/ +protected def subtype (s : affine_subspace k P) [nonempty s] : s →ᵃ[k] P := +{ to_fun := coe, + linear := s.direction.subtype, + map_vadd' := λ p v, rfl } + +@[simp] lemma subtype_linear (s : affine_subspace k P) [nonempty s] : + s.subtype.linear = s.direction.subtype := +rfl + +lemma subtype_apply (s : affine_subspace k P) [nonempty s] (p : s) : s.subtype p = p := +rfl + +@[simp] lemma coe_subtype (s : affine_subspace k P) [nonempty s] : (s.subtype : s → P) = coe := +rfl + +lemma injective_subtype (s : affine_subspace k P) [nonempty s] : function.injective s.subtype := +subtype.coe_injective + /-- Two affine subspaces with nonempty intersection are equal if and only if their directions are equal. -/ lemma eq_iff_direction_eq_of_mem {s₁ s₂ : affine_subspace k P} {p : P} (h₁ : p ∈ s₁) From ae9b867e7625ff064843ad64eb4ad4ad3e18ba0b Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Wed, 21 Sep 2022 08:15:18 +0000 Subject: [PATCH 339/828] feat(analysis/seminorm): continuity criterion for seminorms (#16402) --- src/analysis/normed/group/seminorm.lean | 6 ++ src/analysis/seminorm.lean | 74 +++++++++++++++++++++++++ 2 files changed, 80 insertions(+) diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index e30db0570bde8..2e9f7f120b00a 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -128,6 +128,12 @@ by { rw [div_eq_mul_inv, ←map_inv_eq_map f y], exact map_mul_le_add _ _ _ } @[to_additive] lemma le_map_add_map_div' : f x ≤ f y + f (y / x) := by simpa only [add_comm, map_div_rev, div_mul_cancel'] using map_mul_le_add f (x / y) y +@[to_additive] lemma abs_sub_map_le_div : |f x - f y| ≤ f (x / y) := +begin + rw [abs_sub_le_iff, sub_le_iff_le_add', sub_le_iff_le_add'], + exact ⟨le_map_add_map_div _ _ _, le_map_add_map_div' _ _ _⟩ +end + end group_seminorm_class @[to_additive, priority 100] -- See note [lower instance priority] diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 94acc9b2d1c93..a802fe7173c20 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -320,6 +320,9 @@ begin { exact nnreal.coe_pos.mpr ha }, end +lemma norm_sub_map_le_sub (p : seminorm 𝕜 E) (x y : E) : ∥p x - p y∥ ≤ p (x - y) := +abs_sub_map_le_sub p x y + end module end semi_normed_ring @@ -404,6 +407,8 @@ variables {x y : E} {r : ℝ} @[simp] lemma mem_ball : y ∈ ball p x r ↔ p (y - x) < r := iff.rfl +lemma mem_ball_self (hr : 0 < r) : x ∈ ball p x r := by simp [hr] + lemma mem_ball_zero : y ∈ ball p 0 r ↔ p y < r := by rw [mem_ball, sub_zero] lemma ball_zero_eq : ball p 0 r = { y : E | p y < r } := set.ext $ λ x, p.mem_ball_zero @@ -652,6 +657,75 @@ rfl end restrict_scalars +/-! ### Continuity criterions for seminorms -/ + +section continuity + +variables [semi_normed_ring 𝕜] [add_comm_group E] + [module 𝕜 E] + +lemma continuous_at_zero [norm_one_class 𝕜] [normed_algebra ℝ 𝕜] [module ℝ E] + [is_scalar_tower ℝ 𝕜 E] [topological_space E] [has_continuous_const_smul ℝ E] {p : seminorm 𝕜 E} + (hp : p.ball 0 1 ∈ (𝓝 0 : filter E)) : + continuous_at p 0 := +begin + change continuous_at (p.restrict_scalars ℝ) 0, + rw ← p.restrict_scalars_ball ℝ at hp, + refine metric.nhds_basis_ball.tendsto_right_iff.mpr _, + intros ε hε, + rw map_zero, + suffices : (p.restrict_scalars ℝ).ball 0 ε ∈ (𝓝 0 : filter E), + { rwa seminorm.ball_zero_eq_preimage_ball at this }, + have := (set_smul_mem_nhds_zero_iff hε.ne.symm).mpr hp, + rwa [seminorm.smul_ball_zero (norm_pos_iff.mpr hε.ne.symm), + real.norm_of_nonneg hε.le, mul_one] at this +end + +protected lemma uniform_continuous_of_continuous_at_zero [uniform_space E] [uniform_add_group E] + {p : seminorm 𝕜 E} (hp : continuous_at p 0) : + uniform_continuous p := +begin + have hp : filter.tendsto p (𝓝 0) (𝓝 0) := map_zero p ▸ hp, + rw [uniform_continuous, uniformity_eq_comap_nhds_zero_swapped, + metric.uniformity_eq_comap_nhds_zero, filter.tendsto_comap_iff], + exact tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds + (hp.comp filter.tendsto_comap) (λ xy, dist_nonneg) (λ xy, p.norm_sub_map_le_sub _ _) +end + +protected lemma continuous_of_continuous_at_zero [topological_space E] [topological_add_group E] + {p : seminorm 𝕜 E} (hp : continuous_at p 0) : + continuous p := +begin + letI := topological_add_group.to_uniform_space E, + haveI : uniform_add_group E := topological_add_comm_group_is_uniform, + exact (seminorm.uniform_continuous_of_continuous_at_zero hp).continuous +end + +protected lemma uniform_continuous [norm_one_class 𝕜] [normed_algebra ℝ 𝕜] [module ℝ E] + [is_scalar_tower ℝ 𝕜 E] [uniform_space E] [uniform_add_group E] [has_continuous_const_smul ℝ E] + {p : seminorm 𝕜 E} (hp : p.ball 0 1 ∈ (𝓝 0 : filter E)) : + uniform_continuous p := +seminorm.uniform_continuous_of_continuous_at_zero (continuous_at_zero hp) + +protected lemma continuous [norm_one_class 𝕜] [normed_algebra ℝ 𝕜] [module ℝ E] + [is_scalar_tower ℝ 𝕜 E] [topological_space E] [topological_add_group E] + [has_continuous_const_smul ℝ E] {p : seminorm 𝕜 E} (hp : p.ball 0 1 ∈ (𝓝 0 : filter E)) : + continuous p := +seminorm.continuous_of_continuous_at_zero (continuous_at_zero hp) + +lemma continuous_of_le [norm_one_class 𝕜] [normed_algebra ℝ 𝕜] [module ℝ E] + [is_scalar_tower ℝ 𝕜 E] [topological_space E] [topological_add_group E] + [has_continuous_const_smul ℝ E] {p q : seminorm 𝕜 E} (hq : continuous q) (hpq : p ≤ q) : + continuous p := +begin + refine seminorm.continuous (filter.mem_of_superset + (is_open.mem_nhds _ $ q.mem_ball_self zero_lt_one) (ball_antitone hpq)), + rw ball_zero_eq, + exact is_open_lt hq continuous_const +end + +end continuity + end seminorm /-! ### The norm as a seminorm -/ From a6d28ae9d5123fada05c7972875711a5e1451deb Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 21 Sep 2022 08:15:19 +0000 Subject: [PATCH 340/828] feat(topology/algebra/mul_action): `add_torsor.connected_space` (#16471) Add the lemma that an `add_torsor` for a connected space is itself a connected space, given `has_continuous_vadd`. --- src/topology/algebra/mul_action.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/src/topology/algebra/mul_action.lean b/src/topology/algebra/mul_action.lean index b5719e3c8d659..c38409f627ac9 100644 --- a/src/topology/algebra/mul_action.lean +++ b/src/topology/algebra/mul_action.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import algebra.add_torsor import topology.algebra.constructions import group_theory.group_action.prod import group_theory.group_action.basic @@ -169,3 +170,22 @@ has_continuous_smul_Inf $ set.forall_range_iff.mpr h by { rw inf_eq_infi, refine has_continuous_smul_infi (λ b, _), cases b; assumption } end lattice_ops + +section add_torsor + +variables (G : Type*) (P : Type*) [add_group G] [add_torsor G P] [topological_space G] +variables [preconnected_space G] [topological_space P] [has_continuous_vadd G P] +include G + +/-- An `add_torsor` for a connected space is a connected space. This is not an instance because +it loops for a group as a torsor over itself. -/ +protected lemma add_torsor.connected_space : connected_space P := +{ is_preconnected_univ := + begin + convert is_preconnected_univ.image ((equiv.vadd_const (classical.arbitrary P)) : G → P) + (continuous_id.vadd continuous_const).continuous_on, + rw [set.image_univ, equiv.range_eq_univ] + end, + to_nonempty := infer_instance } + +end add_torsor From d20572713694429af7d0774b580eeeae9d9bb3cc Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 21 Sep 2022 08:15:20 +0000 Subject: [PATCH 341/828] feat(topology/constructions): add/golf lemmas about `sigma` (#16571) * add some `iff` lemmas; * golf some proofs. --- src/topology/category/Top/limits.lean | 2 +- src/topology/constructions.lean | 128 ++++++++++++-------------- src/topology/homeomorph.lean | 9 +- 3 files changed, 61 insertions(+), 78 deletions(-) diff --git a/src/topology/category/Top/limits.lean b/src/topology/category/Top/limits.lean index e227bd9f53d48..ff8728bce2a7c 100644 --- a/src/topology/category/Top/limits.lean +++ b/src/topology/category/Top/limits.lean @@ -191,7 +191,7 @@ cofan.mk (Top.of (Σ i, α i)) (sigma_ι α) /-- The constructed cofan is indeed a colimit -/ def sigma_cofan_is_colimit {ι : Type v} (α : ι → Top.{max v u}) : is_colimit (sigma_cofan α) := { desc := λ S, { to_fun := λ s, S.ι.app ⟨s.1⟩ s.2, - continuous_to_fun := by { continuity, dsimp only, continuity } }, + continuous_to_fun := continuous_sigma $ λ i, map_continuous (S.ι.app ⟨i⟩) }, uniq' := by { intros S m h, ext ⟨i, x⟩, simp [← h ⟨i⟩] }, fac' := λ s j, by { cases j, tidy, }, } diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index f0401e581ef2a..392bb4919e27b 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -1137,7 +1137,8 @@ end) end pi section sigma -variables {ι : Type*} {σ : ι → Type*} [Π i, topological_space (σ i)] +variables {ι κ : Type*} {σ : ι → Type*} {τ : κ → Type*} + [Π i, topological_space (σ i)] [Π k, topological_space (τ k)] [topological_space α] @[continuity] lemma continuous_sigma_mk {i : ι} : continuous (@sigma.mk ι σ i) := @@ -1154,13 +1155,10 @@ begin intros s hs, rw is_open_sigma_iff, intro j, - rcases eq_or_ne i j with (rfl|hne), + rcases eq_or_ne j i with (rfl|hne), { rwa set.preimage_image_eq _ sigma_mk_injective }, - { convert is_open_empty, - apply set.eq_empty_of_subset_empty, - rintro x ⟨y, _, hy⟩, - have : i = j, by cc, - contradiction } + { rw [preimage_image_sigma_mk_of_ne hne], + exact is_open_empty } end lemma is_open_range_sigma_mk {i : ι} : is_open (set.range (@sigma.mk ι σ i)) := @@ -1171,17 +1169,14 @@ begin intros s hs, rw is_closed_sigma_iff, intro j, - rcases eq_or_ne i j with (rfl|hne), + rcases eq_or_ne j i with (rfl|hne), { rwa set.preimage_image_eq _ sigma_mk_injective }, - { convert is_closed_empty, - apply set.eq_empty_of_subset_empty, - rintro x ⟨y, _, hy⟩, - have : i = j, by cc, - contradiction } + { rw [preimage_image_sigma_mk_of_ne hne], + exact is_closed_empty } end -lemma is_closed_sigma_mk {i : ι} : is_closed (set.range (@sigma.mk ι σ i)) := -by { rw ←set.image_univ, exact is_closed_map_sigma_mk _ is_closed_univ } +lemma is_closed_range_sigma_mk {i : ι} : is_closed (set.range (@sigma.mk ι σ i)) := +is_closed_map_sigma_mk.closed_range lemma open_embedding_sigma_mk {i : ι} : open_embedding (@sigma.mk ι σ i) := open_embedding_of_continuous_injective_open @@ -1194,6 +1189,15 @@ closed_embedding_of_continuous_injective_closed lemma embedding_sigma_mk {i : ι} : embedding (@sigma.mk ι σ i) := closed_embedding_sigma_mk.1 +lemma sigma.nhds_mk (i : ι) (x : σ i) : 𝓝 (⟨i, x⟩ : sigma σ) = map (sigma.mk i) (𝓝 x) := +(open_embedding_sigma_mk.map_nhds_eq x).symm + +lemma sigma.nhds_eq (x : sigma σ) : 𝓝 x = map (sigma.mk x.1) (𝓝 x.2) := +by { cases x, apply sigma.nhds_mk } + +lemma comap_sigma_mk_nhds (i : ι) (x : σ i) : comap (sigma.mk i) (𝓝 ⟨i, x⟩) = 𝓝 x := +(embedding_sigma_mk.to_inducing.nhds_eq_comap _).symm + lemma is_open_sigma_fst_preimage (s : set ι) : is_open (sigma.fst ⁻¹' s : set (Σ a, σ a)) := begin rw [← bUnion_of_singleton s, preimage_Union₂], @@ -1201,65 +1205,47 @@ begin exact is_open_bUnion (λ _ _, is_open_range_sigma_mk) end +/-- A map out of a sum type is continuous iff its restriction to each summand is. -/ +@[simp] lemma continuous_sigma_iff {f : sigma σ → α} : + continuous f ↔ ∀ i, continuous (λ a, f ⟨i, a⟩) := +by simp only [continuous_supr_dom, continuous_coinduced_dom] + /-- A map out of a sum type is continuous if its restriction to each summand is. -/ -@[continuity] -lemma continuous_sigma [topological_space β] {f : sigma σ → β} - (h : ∀ i, continuous (λ a, f ⟨i, a⟩)) : continuous f := -continuous_supr_dom.2 (λ i, continuous_coinduced_dom.2 (h i)) +@[continuity] lemma continuous_sigma {f : sigma σ → α} (hf : ∀ i, continuous (λ a, f ⟨i, a⟩)) : + continuous f := +continuous_sigma_iff.2 hf -@[continuity] -lemma continuous_sigma_map {κ : Type*} {τ : κ → Type*} [Π k, topological_space (τ k)] - {f₁ : ι → κ} {f₂ : Π i, σ i → τ (f₁ i)} (hf : ∀ i, continuous (f₂ i)) : - continuous (sigma.map f₁ f₂) := -continuous_sigma $ λ i, - show continuous (λ a, sigma.mk (f₁ i) (f₂ i a)), - from continuous_sigma_mk.comp (hf i) +@[simp] lemma continuous_sigma_map {f₁ : ι → κ} {f₂ : Π i, σ i → τ (f₁ i)} : + continuous (sigma.map f₁ f₂) ↔ ∀ i, continuous (f₂ i) := +continuous_sigma_iff.trans $ by simp only [sigma.map, embedding_sigma_mk.continuous_iff] -lemma is_open_map_sigma [topological_space β] {f : sigma σ → β} - (h : ∀ i, is_open_map (λ a, f ⟨i, a⟩)) : is_open_map f := -begin - intros s hs, - rw is_open_sigma_iff at hs, - rw [← Union_image_preimage_sigma_mk_eq_self s, image_Union], - apply is_open_Union, - intro i, - rw [image_image], - exact h i _ (hs i) -end - -/-- The sum of embeddings is an embedding. -/ -lemma embedding_sigma_map {τ : ι → Type*} [Π i, topological_space (τ i)] - {f : Π i, σ i → τ i} (hf : ∀ i, embedding (f i)) : embedding (sigma.map id f) := -begin - refine ⟨⟨_⟩, function.injective_id.sigma_map (λ i, (hf i).inj)⟩, - refine le_antisymm - (continuous_iff_le_induced.mp (continuous_sigma_map (λ i, (hf i).continuous))) _, - intros s hs, - replace hs := is_open_sigma_iff.mp hs, - have : ∀ i, ∃ t, is_open t ∧ f i ⁻¹' t = sigma.mk i ⁻¹' s, - { intro i, - apply is_open_induced_iff.mp, - convert hs i, - exact (hf i).induced.symm }, - choose t ht using this, - apply is_open_induced_iff.mpr, - refine ⟨⋃ i, sigma.mk i '' t i, is_open_Union (λ i, is_open_map_sigma_mk _ (ht i).1), _⟩, - ext ⟨i, x⟩, - change (sigma.mk i (f i x) ∈ ⋃ (i : ι), sigma.mk i '' t i) ↔ x ∈ sigma.mk i ⁻¹' s, - rw [←(ht i).2, mem_Union], - split, - { rintro ⟨j, hj⟩, - rw mem_image at hj, - rcases hj with ⟨y, hy₁, hy₂⟩, - rcases sigma.mk.inj_iff.mp hy₂ with ⟨rfl, hy⟩, - replace hy := eq_of_heq hy, - subst y, - exact hy₁ }, - { intro hx, - use i, - rw mem_image, - exact ⟨f i x, hx, rfl⟩ } -end +@[continuity] lemma continuous.sigma_map {f₁ : ι → κ} {f₂ : Π i, σ i → τ (f₁ i)} + (hf : ∀ i, continuous (f₂ i)) : + continuous (sigma.map f₁ f₂) := +continuous_sigma_map.2 hf + +lemma is_open_map_sigma {f : sigma σ → α} : is_open_map f ↔ ∀ i, is_open_map (λ a, f ⟨i, a⟩) := +by simp only [is_open_map_iff_nhds_le, sigma.forall, sigma.nhds_eq, map_map] + +lemma is_open_map_sigma_map {f₁ : ι → κ} {f₂ : Π i, σ i → τ (f₁ i)} : + is_open_map (sigma.map f₁ f₂) ↔ ∀ i, is_open_map (f₂ i) := +is_open_map_sigma.trans $ forall_congr $ + λ i, (@open_embedding_sigma_mk _ _ _ (f₁ i)).is_open_map_iff.symm + +lemma inducing_sigma_map {f₁ : ι → κ} {f₂ : Π i, σ i → τ (f₁ i)} (h₁ : injective f₁) : + inducing (sigma.map f₁ f₂) ↔ ∀ i, inducing (f₂ i) := +by simp only [inducing_iff_nhds, sigma.forall, sigma.nhds_mk, sigma.map, ← map_sigma_mk_comap h₁, + map_inj sigma_mk_injective] + +lemma embedding_sigma_map {f₁ : ι → κ} {f₂ : Π i, σ i → τ (f₁ i)} (h : injective f₁) : + embedding (sigma.map f₁ f₂) ↔ ∀ i, embedding (f₂ i) := +by simp only [embedding_iff, injective.sigma_map, inducing_sigma_map h, forall_and_distrib, + h.sigma_map_iff] + +lemma open_embedding_sigma_map {f₁ : ι → κ} {f₂ : Π i, σ i → τ (f₁ i)} (h : injective f₁) : + open_embedding (sigma.map f₁ f₂) ↔ ∀ i, open_embedding (f₂ i) := +by simp only [open_embedding_iff_embedding_open, is_open_map_sigma_map, embedding_sigma_map h, + forall_and_distrib] end sigma diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 1bf50e868bb41..dbaf081171be3 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -423,12 +423,9 @@ variables {ι : Type*} {σ : ι → Type*} [Π i, topological_space (σ i)] /-- `(Σ i, σ i) × β` is homeomorphic to `Σ i, (σ i × β)`. -/ def sigma_prod_distrib : ((Σ i, σ i) × β) ≃ₜ (Σ i, (σ i × β)) := -homeomorph.symm $ -homeomorph_of_continuous_open (equiv.sigma_prod_distrib σ β).symm - (continuous_sigma $ λ i, - (continuous_sigma_mk.comp continuous_fst).prod_mk continuous_snd) - (is_open_map_sigma $ λ i, - (open_embedding_sigma_mk.prod open_embedding_id).is_open_map) +homeomorph.symm $ homeomorph_of_continuous_open (equiv.sigma_prod_distrib σ β).symm + (continuous_sigma $ λ i, continuous_sigma_mk.fst'.prod_mk continuous_snd) + (is_open_map_sigma.2 $ λ i, is_open_map_sigma_mk.prod is_open_map.id) end distrib From aae2a82f21bc70568d9eb1543d51531104fbab16 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 21 Sep 2022 10:15:01 +0000 Subject: [PATCH 342/828] feat(algebra/order/hom/basic): `positivity` extension for nonnegative homs (#16228) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add `positivity_map`, a `positivity` extension for expressions of the form `f x` where `f : F` where `nonneg_hom_class F α β`. Golf existing proofs using it. --- src/algebra/order/hom/basic.lean | 12 ++++++++++++ src/analysis/convex/gauge.lean | 9 ++++----- src/analysis/normed/group/seminorm.lean | 5 ++--- src/analysis/seminorm.lean | 7 +++---- test/positivity.lean | 3 +++ 5 files changed, 24 insertions(+), 12 deletions(-) diff --git a/src/algebra/order/hom/basic.lean b/src/algebra/order/hom/basic.lean index b36d1b7a08760..8e86895e0e375 100644 --- a/src/algebra/order/hom/basic.lean +++ b/src/algebra/order/hom/basic.lean @@ -6,6 +6,7 @@ Authors: Yaël Dillies import algebra.hom.group import algebra.order.with_zero import order.hom.basic +import tactic.positivity /-! # Algebraic order homomorphism classes @@ -58,3 +59,14 @@ attribute [simp] map_nonneg @[to_additive] lemma le_map_add_map_div [group α] [add_comm_semigroup β] [has_le β] [mul_le_add_hom_class F α β] (f : F) (a b : α) : f a ≤ f b + f (a / b) := by simpa only [add_comm, div_mul_cancel'] using map_mul_le_add f (a / b) b + +namespace tactic +open positivity + +/-- Extension for the `positivity` tactic: nonnegative maps take nonnegative values. -/ +@[positivity] +meta def positivity_map : expr → tactic strictness +| (expr.app `(⇑%%f) `(%%a)) := nonnegative <$> mk_app ``map_nonneg [f, a] +| _ := failed + +end tactic diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index 2b1855b092fd1..d2c2f50bb6135 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -136,8 +136,7 @@ begin suffices : (r⁻¹ * δ) • δ⁻¹ • x ∈ s, { rwa [smul_smul, mul_inv_cancel_right₀ δ_pos.ne'] at this }, rw mem_smul_set_iff_inv_smul_mem₀ δ_pos.ne' at hδ, - refine hs₁.smul_mem_of_zero_mem hs₀ hδ - ⟨mul_nonneg (inv_nonneg.2 hr'.le) δ_pos.le, _⟩, + refine hs₁.smul_mem_of_zero_mem hs₀ hδ ⟨by positivity, _⟩, rw [inv_mul_le_iff hr', mul_one], exact hδr.le }, { have hε' := (lt_add_iff_pos_right a).2 (half_pos hε), @@ -200,7 +199,7 @@ begin rintro b ⟨hb, x, hx', rfl⟩, refine not_lt.1 (λ hba, hx _), have ha := hb.trans hba, - refine ⟨(a⁻¹ * b) • x, hs₀ hx' (mul_nonneg (inv_nonneg.2 ha.le) hb.le) _, _⟩, + refine ⟨(a⁻¹ * b) • x, hs₀ hx' (by positivity) _, _⟩, { rw ←div_eq_inv_mul, exact div_le_one_of_le hba.le ha.le }, { rw [←mul_smul, mul_inv_cancel_left₀ ha.ne'] } @@ -416,7 +415,7 @@ begin rw p.mem_ball_zero at hy, rw [map_smul_eq_mul, real.norm_eq_abs, abs_of_pos hr], exact mul_le_of_le_one_right hr.le hy.le }, - { have hpε : 0 < p x + ε := add_pos_of_nonneg_of_pos (map_nonneg _ _) hε, + { have hpε : 0 < p x + ε := by positivity, refine hr ⟨hpε, (p x + ε)⁻¹ • x, _, smul_inv_smul₀ hpε.ne' _⟩, rw [p.mem_ball_zero, map_smul_eq_mul, real.norm_eq_abs, abs_of_pos (inv_pos.2 hpε), inv_mul_lt_iff hpε, mul_one], @@ -437,7 +436,7 @@ begin obtain rfl | hx := eq_or_ne x 0, { rw [norm_zero, gauge_zero] }, refine (le_of_forall_pos_le_add $ λ ε hε, _).antisymm _, - { have := add_pos_of_nonneg_of_pos (norm_nonneg x) hε, + { have : 0 < ∥x∥ + ε := by positivity, refine gauge_le_of_mem this.le _, rw [smul_ball this.ne', smul_zero, real.norm_of_nonneg this.le, mul_one, mem_ball_zero_iff], exact lt_add_of_pos_right _ hε }, diff --git a/src/analysis/normed/group/seminorm.lean b/src/analysis/normed/group/seminorm.lean index 2e9f7f120b00a..a182d15fe6b0b 100644 --- a/src/analysis/normed/group/seminorm.lean +++ b/src/analysis/normed/group/seminorm.lean @@ -264,13 +264,12 @@ variables [comm_group E] [comm_group F] (p q : group_seminorm E) (x y : E) @[to_additive] lemma mul_bdd_below_range_add {p q : group_seminorm E} {x : E} : bdd_below (range $ λ y, p y + q (x / y)) := -⟨0, by { rintro _ ⟨x, rfl⟩, exact add_nonneg (map_nonneg p _) (map_nonneg q _) }⟩ +⟨0, by { rintro _ ⟨x, rfl⟩, dsimp, positivity }⟩ @[to_additive] noncomputable instance : has_inf (group_seminorm E) := ⟨λ p q, { to_fun := λ x, ⨅ y, p y + q (x / y), - map_one' := cinfi_eq_of_forall_ge_of_forall_gt_exists_lt - (λ x, add_nonneg (map_nonneg p _) (map_nonneg q _)) + map_one' := cinfi_eq_of_forall_ge_of_forall_gt_exists_lt (λ x, by positivity) (λ r hr, ⟨1, by rwa [div_one, map_one_eq_zero p, map_one_eq_zero q, add_zero]⟩), mul_le' := λ x y, le_cinfi_add_cinfi $ λ u v, begin refine cinfi_le_of_le mul_bdd_below_range_add (u * v) _, diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index a802fe7173c20..984c5cc823e2d 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -344,7 +344,7 @@ variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] {p q : seminorm /-- Auxiliary lemma to show that the infimum of seminorms is well-defined. -/ lemma bdd_below_range_add : bdd_below (range $ λ u, p u + q (x - u)) := -⟨0, by { rintro _ ⟨x, rfl⟩, exact add_nonneg (map_nonneg p _) (map_nonneg q _) }⟩ +⟨0, by { rintro _ ⟨x, rfl⟩, dsimp, positivity }⟩ noncomputable instance : has_inf (seminorm 𝕜 E) := { inf := λ p q, @@ -354,8 +354,7 @@ noncomputable instance : has_inf (seminorm 𝕜 E) := intros a x, obtain rfl | ha := eq_or_ne a 0, { rw [norm_zero, zero_mul, zero_smul], - refine cinfi_eq_of_forall_ge_of_forall_gt_exists_lt - (λ i, add_nonneg (map_nonneg p _) (map_nonneg q _)) + refine cinfi_eq_of_forall_ge_of_forall_gt_exists_lt (λ i, by positivity) (λ x hx, ⟨0, by rwa [map_zero, sub_zero, map_zero, add_zero]⟩) }, simp_rw [real.mul_infi_of_nonneg (norm_nonneg a), mul_add, ←map_smul_eq_mul p, ←map_smul_eq_mul q, smul_sub], @@ -569,7 +568,7 @@ protected lemma absorbent_ball_zero (hr : 0 < r) : absorbent 𝕜 (ball p (0 : E begin rw absorbent_iff_nonneg_lt, rintro x, - have hxr : 0 ≤ p x/r := div_nonneg (map_nonneg p _) hr.le, + have hxr : 0 ≤ p x / r := by positivity, refine ⟨p x/r, hxr, λ a ha, _⟩, have ha₀ : 0 < ∥a∥ := hxr.trans_lt ha, refine ⟨a⁻¹ • x, _, smul_inv_smul₀ (norm_pos_iff.1 ha₀) x⟩, diff --git a/test/positivity.lean b/test/positivity.lean index 782184eeabf47..39428bf6ca2df 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -100,6 +100,9 @@ example {V : Type*} [normed_add_comm_group V] (x : V) : 0 ≤ ∥x∥ := by posi example {X : Type*} [metric_space X] (x y : X) : 0 ≤ dist x y := by positivity +example {E : Type*} [add_group E] {p : add_group_seminorm E} {x : E} : 0 ≤ p x := by positivity +example {E : Type*} [group E] {p : group_seminorm E} {x : E} : 0 ≤ p x := by positivity + /- ## Tests that the tactic is agnostic on reversed inequalities -/ example {a : ℤ} (ha : a > 0) : 0 ≤ a := by positivity From 84b27545049b007531353fbffdb449d94862b919 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 21 Sep 2022 12:36:51 +0000 Subject: [PATCH 343/828] feat(algebra/field/opposite): Missing instances (#16564) A few missing field-like instances for `mul_opposite` and `add_opposite`. --- src/algebra/field/opposite.lean | 20 +++++++++++++++++--- 1 file changed, 17 insertions(+), 3 deletions(-) diff --git a/src/algebra/field/opposite.lean b/src/algebra/field/opposite.lean index 2a6f100dcd73b..ea50691bd26db 100644 --- a/src/algebra/field/opposite.lean +++ b/src/algebra/field/opposite.lean @@ -7,17 +7,31 @@ import algebra.field.basic import algebra.ring.opposite /-! -# Field structure on the multiplicative opposite +# Field structure on the multiplicative/additive opposite -/ variables (α : Type*) -namespace mul_opposite +instance [division_semiring α] : division_semiring αᵐᵒᵖ := +{ .. mul_opposite.group_with_zero α, .. mul_opposite.semiring α } instance [division_ring α] : division_ring αᵐᵒᵖ := { .. mul_opposite.group_with_zero α, .. mul_opposite.ring α } +instance [semifield α] : semifield αᵐᵒᵖ := +{ .. mul_opposite.division_semiring α, .. mul_opposite.comm_semiring α } + instance [field α] : field αᵐᵒᵖ := { .. mul_opposite.division_ring α, .. mul_opposite.comm_ring α } -end mul_opposite +instance [division_semiring α] : division_semiring αᵃᵒᵖ := +{ ..add_opposite.group_with_zero α, ..add_opposite.semiring α } + +instance [division_ring α] : division_ring αᵃᵒᵖ := +{ ..add_opposite.group_with_zero α, ..add_opposite.ring α } + +instance [semifield α] : semifield αᵃᵒᵖ := +{ ..add_opposite.division_semiring α, ..add_opposite.comm_semiring α } + +instance [field α] : field αᵃᵒᵖ := +{ ..add_opposite.division_ring α, ..add_opposite.comm_ring α } From b1539919749b1cba60666bed196b359a5819a8c4 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Wed, 21 Sep 2022 15:26:28 +0000 Subject: [PATCH 344/828] doc(metric_space/basic): modify doc for metric spaces adding pseudo metric spaces (#16577) The doc preceeding the `pseudo_metric_space` instance spoke only about `metric spaces`. I have added the `pseudo_metric` version. --- src/topology/metric_space/basic.lean | 16 +++++++++------- 1 file changed, 9 insertions(+), 7 deletions(-) diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index cd65e4f401b2e..252166c1afc17 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -149,13 +149,15 @@ used. -/ protected meta def pseudo_metric_space.edist_dist_tac : tactic unit := tactic.intros >> `[exact (ennreal.of_real_eq_coe_nnreal _).symm <|> control_laws_tac] -/-- Metric space - -Each metric space induces a canonical `uniform_space` and hence a canonical `topological_space`. -This is enforced in the type class definition, by extending the `uniform_space` structure. When -instantiating a `metric_space` structure, the uniformity fields are not necessary, they will be -filled in by default. In the same way, each metric space induces an emetric space structure. -It is included in the structure, but filled in by default. +/-- Pseudo metric and Metric spaces + +A pseudo metric space is endowed with a distance for which the requirement `d(x,y)=0 → x = y` might +not hold. A metric space is a pseudo metric space such that `d(x,y)=0 → x = y`. +Each pseudo metric space induces a canonical `uniform_space` and hence a canonical +`topological_space` This is enforced in the type class definition, by extending the `uniform_space` +structure. When instantiating a `pseudo_metric_space` structure, the uniformity fields are not +necessary, they will be filled in by default. In the same way, each (pseudo) metric space induces a +(pseudo) emetric space structure. It is included in the structure, but filled in by default. -/ class pseudo_metric_space (α : Type u) extends has_dist α : Type u := (dist_self : ∀ x : α, dist x x = 0) From 88f41bb0bd5c895bca82a7900b0b6a205fd0e937 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 21 Sep 2022 20:56:31 +0000 Subject: [PATCH 345/828] feat(data/complex/exponential): `positivity` extension for `real.exp` (#16492) Add `positivity_exp`, a `positivity` extension to prove `0 < real.exp r` for all `r`. --- src/data/complex/exponential.lean | 11 +++++++++++ test/positivity.lean | 3 +++ 2 files changed, 14 insertions(+) diff --git a/src/data/complex/exponential.lean b/src/data/complex/exponential.lean index 45b60b8ae7c96..5d564c6b500b2 100644 --- a/src/data/complex/exponential.lean +++ b/src/data/complex/exponential.lean @@ -1630,6 +1630,17 @@ end end real +namespace tactic +open positivity real + +/-- Extension for the `positivity` tactic: `real.exp` is always positive. -/ +@[positivity] +meta def positivity_exp : expr → tactic strictness +| `(real.exp %%a) := positive <$> mk_app `real.exp_pos [a] +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `real.exp r`" + +end tactic + namespace complex @[simp] lemma abs_cos_add_sin_mul_I (x : ℝ) : abs (cos x + sin x * I) = 1 := diff --git a/test/positivity.lean b/test/positivity.lean index 39428bf6ca2df..507c74efdd0ae 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -1,5 +1,6 @@ import algebra.order.smul import analysis.normed.group.basic +import data.complex.exponential import data.real.sqrt import tactic.positivity @@ -96,6 +97,8 @@ example : 0 ≤ max (-3 : ℤ) 5 := by positivity example {α β : Type*} [ordered_semiring α] [ordered_add_comm_monoid β] [smul_with_zero α β] [ordered_smul α β] {a : α} (ha : 0 < a) {b : β} (hb : 0 < b) : 0 ≤ a • b := by positivity +example {r : ℝ} : 0 < real.exp r := by positivity + example {V : Type*} [normed_add_comm_group V] (x : V) : 0 ≤ ∥x∥ := by positivity example {X : Type*} [metric_space X] (x y : X) : 0 ≤ dist x y := by positivity From ec5f9adc44ba5559369652f24b5b1e8231dbc981 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Thu, 22 Sep 2022 05:58:38 +0000 Subject: [PATCH 346/828] feat(analysis/convolution): relax typeclasses and add lemma (#16587) * Relax type-class conditions on `measure_theory.integrable.integrable_convolution` and `convolution_flip` * Add `integral_convolution` --- src/analysis/convolution.lean | 47 ++++++++++++++++++++++++++--------- 1 file changed, 35 insertions(+), 12 deletions(-) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 626e60bb616c5..cbdf8be0e01fa 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -437,6 +437,15 @@ begin { exact (h $ sub_add_cancel x t).elim } end +section +variables [has_measurable_add₂ G] [has_measurable_neg G] [sigma_finite μ] [is_add_right_invariant μ] + +lemma measure_theory.integrable.integrable_convolution (hf : integrable f μ) (hg : integrable g μ) : + integrable (f ⋆[L, μ] g) μ := +(hf.convolution_integrand L hg).integral_prod_left + +end + variables [topological_space G] variables [topological_add_group G] @@ -500,12 +509,6 @@ lemma has_compact_support.continuous_convolution_right_of_integrable (hg.norm.bdd_above_range_of_has_compact_support hcg.norm).continuous_convolution_right_of_integrable L hf hg -variables [sigma_finite μ] [is_add_right_invariant μ] - -lemma measure_theory.integrable.integrable_convolution (hf : integrable f μ) (hg : integrable g μ) : - integrable (f ⋆[L, μ] g) μ := -(hf.convolution_integrand L hg).integral_prod_left - end group section comm_group @@ -515,10 +518,11 @@ variables [add_comm_group G] lemma support_convolution_subset : support (f ⋆[L, μ] g) ⊆ support f + support g := (support_convolution_subset_swap L).trans (add_comm _ _).subset -variables [topological_space G] -variables [topological_add_group G] -variables [borel_space G] -variables [is_add_left_invariant μ] [is_neg_invariant μ] +variables [is_add_left_invariant μ] [is_neg_invariant μ] + +section measurable +variables [has_measurable_neg G] +variables [has_measurable_add G] variable (L) /-- Commutativity of convolution -/ @@ -544,6 +548,11 @@ lemma convolution_lmul_swap [normed_space ℝ 𝕜] [complete_space 𝕜] {f : G (f ⋆[lmul 𝕜 𝕜, μ] g) x = ∫ t, f (x - t) * g t ∂μ := convolution_eq_swap _ +end measurable + +variables [topological_space G] +variables [topological_add_group G] +variables [borel_space G] variables [second_countable_topology G] lemma has_compact_support.continuous_convolution_left [locally_compact_space G] [t2_space G] @@ -772,9 +781,23 @@ variables {k : G → E''} variables (L₂ : F →L[𝕜] E'' →L[𝕜] F') variables (L₃ : E →L[𝕜] F'' →L[𝕜] F') variables (L₄ : E' →L[𝕜] E'' →L[𝕜] F'') -variables [add_group G] [has_measurable_add G] +variables [add_group G] variables [sigma_finite μ] -variables {ν : measure G} [sigma_finite ν] [is_add_right_invariant ν] + +lemma integral_convolution + [has_measurable_add₂ G] [has_measurable_neg G] [is_add_right_invariant μ] + [normed_space ℝ E] [normed_space ℝ E'] + [complete_space E] [complete_space E'] + (hf : integrable f μ) (hg : integrable g μ) : + ∫ x, (f ⋆[L, μ] g) x ∂μ = L (∫ x, f x ∂μ) (∫ x, g x ∂μ) := +begin + refine (integral_integral_swap (by apply hf.convolution_integrand L hg)).trans _, + simp_rw [integral_comp_comm _ (hg.comp_sub_right _), integral_sub_right_eq_self], + exact (L.flip (∫ x, g x ∂μ)).integral_comp_comm hf, +end + +variables [has_measurable_add G] {ν : measure G} [sigma_finite ν] [is_add_right_invariant ν] + /-- Convolution is associative. To do: prove that `hi` follows from simpler conditions. -/ From f4210e94a65a1b045fdfa2ab806e9bac56d1fa8a Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 22 Sep 2022 06:54:03 +0000 Subject: [PATCH 347/828] feat(data/polynomial/ring_division): move lemmas, add lemmas (#16432) Rename `count_map_roots` to `count_map_roots_of_injective`. Add 8 lemmas: ``` le_root_multiplicity_iff root_multiplicity_le_iff pow_root_multiplicity_not_dvd _root_.multiset.prod_X_sub_C_dvd_iff_le_roots count_map_roots map_roots_le map_roots_le_of_injective card_roots_le_map ``` Remove `root_multiplicity_of_dvd`, it's just the left direction of `le_root_multiplicity_iff`. Move and golf 3 lemmas: ``` le_root_multiplicity_map (8 -> 5) eq_root_multiplicity_map (10 -> 7) roots_map_of_injective_card_eq_total_degree (7 -> 4) ``` - [x] depends on: #16440 Co-authored-by: alreadydone Co-authored-by: Junyan Xu --- src/data/polynomial/ring_division.lean | 191 ++++++++++++++----------- src/field_theory/splitting_field.lean | 2 +- 2 files changed, 112 insertions(+), 81 deletions(-) diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index 2dfa0d0812559..062ebffad0d13 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -225,7 +225,37 @@ instance : is_domain R[X] := end ring section comm_ring -variables [comm_ring R] [is_domain R] {p q : R[X]} +variable [comm_ring R] + +/-- The multiplicity of `a` as root of a nonzero polynomial `p` is at least `n` iff + `(X - a) ^ n` divides `p`. -/ +lemma le_root_multiplicity_iff {p : R[X]} (p0 : p ≠ 0) {a : R} {n : ℕ} : + n ≤ root_multiplicity a p ↔ (X - C a) ^ n ∣ p := +begin + simp_rw [root_multiplicity, dif_neg p0, nat.le_find_iff, not_not], + refine ⟨λ h, _, λ h m hm, (pow_dvd_pow _ hm).trans h⟩, + cases n, { rw pow_zero, apply one_dvd }, { exact h n n.lt_succ_self }, +end + +lemma root_multiplicity_le_iff {p : R[X]} (p0 : p ≠ 0) (a : R) (n : ℕ) : + root_multiplicity a p ≤ n ↔ ¬ (X - C a) ^ (n + 1) ∣ p := +by rw [← (le_root_multiplicity_iff p0).not, not_le, nat.lt_add_one_iff] + +lemma pow_root_multiplicity_not_dvd {p : R[X]} (p0 : p ≠ 0) (a : R) : + ¬ (X - C a) ^ (root_multiplicity a p + 1) ∣ p := +by rw [← root_multiplicity_le_iff p0] + +/-- The multiplicity of `p + q` is at least the minimum of the multiplicities. -/ +lemma root_multiplicity_add {p q : R[X]} (a : R) (hzero : p + q ≠ 0) : + min (root_multiplicity a p) (root_multiplicity a q) ≤ root_multiplicity a (p + q) := +begin + rw le_root_multiplicity_iff hzero, + have hdivp : (X - C a) ^ root_multiplicity a p ∣ p := pow_root_multiplicity_dvd p a, + have hdivq : (X - C a) ^ root_multiplicity a q ∣ q := pow_root_multiplicity_dvd q a, + exact min_pow_dvd_add hdivp hdivq +end + +variables [is_domain R] {p q : R[X]} section roots @@ -315,27 +345,6 @@ begin simp only [root_multiplicity_mul hzero, root_multiplicity_X_sub_C_self, hn, nat.one_add] end -/-- If `(X - a) ^ n` divides a polynomial `p` then the multiplicity of `a` as root of `p` is at -least `n`. -/ -lemma root_multiplicity_of_dvd {p : R[X]} {a : R} {n : ℕ} - (hzero : p ≠ 0) (h : (X - C a) ^ n ∣ p) : n ≤ root_multiplicity a p := -begin - obtain ⟨q, hq⟩ := exists_eq_mul_right_of_dvd h, - rw hq at hzero, - simp only [hq, root_multiplicity_mul hzero, root_multiplicity_X_sub_C_pow, - ge_iff_le, _root_.zero_le, le_add_iff_nonneg_right], -end - -/-- The multiplicity of `p + q` is at least the minimum of the multiplicities. -/ -lemma root_multiplicity_add {p q : R[X]} (a : R) (hzero : p + q ≠ 0) : - min (root_multiplicity a p) (root_multiplicity a q) ≤ root_multiplicity a (p + q) := -begin - refine root_multiplicity_of_dvd hzero _, - have hdivp : (X - C a) ^ root_multiplicity a p ∣ p := pow_root_multiplicity_dvd p a, - have hdivq : (X - C a) ^ root_multiplicity a q ∣ q := pow_root_multiplicity_dvd q a, - exact min_pow_dvd_add hdivp hdivq -end - lemma exists_multiset_roots : ∀ {p : R[X]} (hp : p ≠ 0), ∃ s : multiset R, (s.card : with_bot ℕ) ≤ degree p ∧ ∀ a, s.count a = root_multiplicity a p | p := λ hp, by haveI := classical.prop_decidable (∃ x, is_root p x); exact @@ -541,61 +550,6 @@ calc ((roots ((X : R[X]) ^ n - C a)).card : with_bot ℕ) ≤ degree ((X : R[X]) ^ n - C a) : card_roots (X_pow_sub_C_ne_zero hn a) ... = n : degree_X_pow_sub_C hn a -section - -variables {A B : Type*} [comm_ring A] [comm_ring B] - -lemma le_root_multiplicity_map {p : A[X]} {f : A →+* B} (hmap : map f p ≠ 0) (a : A) : - root_multiplicity a p ≤ root_multiplicity (f a) (map f p) := -begin - have hp0 : p ≠ 0 := λ h, hmap (h.symm ▸ polynomial.map_zero f), - rw [root_multiplicity, root_multiplicity, dif_neg hp0, dif_neg hmap], - simp only [not_not, nat.lt_find_iff, nat.le_find_iff], - intros m hm, - have := (map_ring_hom f).map_dvd (hm m le_rfl), - simpa only [coe_map_ring_hom, map_pow, map_sub, map_X, map_C], -end - -lemma eq_root_multiplicity_map {p : A[X]} {f : A →+* B} (hf : function.injective f) - (a : A) : root_multiplicity a p = root_multiplicity (f a) (map f p) := -begin - by_cases hp0 : p = 0, { simp only [hp0, root_multiplicity_zero, polynomial.map_zero], }, - have hmap : map f p ≠ 0, { simpa only [polynomial.map_zero] using (map_injective f hf).ne hp0, }, - apply le_antisymm (le_root_multiplicity_map hmap a), - rw [root_multiplicity, root_multiplicity, dif_neg hp0, dif_neg hmap], - simp only [not_not, nat.lt_find_iff, nat.le_find_iff], - intros m hm, rw ← map_dvd_map f hf ((monic_X_sub_C a).pow _), - convert hm m le_rfl, - simp only [polynomial.map_pow, polynomial.map_sub, map_pow, map_sub, map_X, map_C], -end - -lemma count_map_roots [is_domain A] (p : A[X]) {f : A →+* B} (hf : function.injective f) - (a : B) : count a (p.roots.map f) ≤ root_multiplicity a (map f p) := -begin - by_cases h : ∃ t, f t = a, - { rcases h with ⟨h_w, rfl⟩, - rw [multiset.count_map_eq_count' f _ hf, count_roots], - exact (eq_root_multiplicity_map hf h_w).le }, - { suffices : (multiset.map f p.roots).count a = 0, - { rw this, exact zero_le _, }, - rw [multiset.count_map, multiset.card_eq_zero, multiset.filter_eq_nil], - rintro k hk rfl, - exact h ⟨k, rfl⟩, }, -end - -lemma roots_map_of_injective_card_eq_total_degree [is_domain A] [is_domain B] {p : A[X]} - {f : A →+* B} (hf : function.injective f) (hroots : p.roots.card = p.nat_degree) : - p.roots.map f = (map f p).roots := -begin - by_cases hp0 : p = 0, { simp only [hp0, roots_zero, multiset.map_zero, polynomial.map_zero], }, - have hmap : map f p ≠ 0, { simpa only [polynomial.map_zero] using (map_injective f hf).ne hp0, }, - apply multiset.eq_of_le_of_card_le, - { simpa only [multiset.le_iff_count, count_roots] using count_map_roots p hf }, - { simpa only [multiset.card_map, hroots] using (card_roots' _).trans (nat_degree_map_le f p) }, -end - -end - section nth_roots /-- `nth_roots n a` noncomputably returns the solutions to `x ^ n = a`-/ @@ -847,12 +801,21 @@ begin { exact polynomial.map_dvd _ (pow_root_multiplicity_dvd p a) }, end +/-- A Galois connection. -/ +lemma _root_.multiset.prod_X_sub_C_dvd_iff_le_roots {p : R[X]} (hp : p ≠ 0) (s : multiset R) : + (s.map (λ a, X - C a)).prod ∣ p ↔ s ≤ p.roots := +⟨λ h, multiset.le_iff_count.2 $ λ r, begin + rw [count_roots, le_root_multiplicity_iff hp, ← multiset.prod_repeat, + ← multiset.map_repeat (λ a, X - C a), ← multiset.filter_eq], + exact (multiset.prod_dvd_prod_of_le $ multiset.map_le_map $ s.filter_le _).trans h, +end, λ h, (multiset.prod_dvd_prod_of_le $ multiset.map_le_map h).trans p.prod_multiset_X_sub_C_dvd⟩ + lemma exists_prod_multiset_X_sub_C_mul (p : R[X]) : ∃ q, (p.roots.map (λ a, X - C a)).prod * q = p ∧ p.roots.card + q.nat_degree = p.nat_degree ∧ q.roots = 0 := begin - obtain ⟨q, he⟩ := prod_multiset_X_sub_C_dvd p, + obtain ⟨q, he⟩ := p.prod_multiset_X_sub_C_dvd, use [q, he.symm], obtain (rfl|hq) := eq_or_ne q 0, { rw mul_zero at he, subst he, simp }, @@ -869,8 +832,7 @@ can be written `p = p.leading_coeff * ∏(X - a)`, for `a` in `p.roots`. -/ lemma C_leading_coeff_mul_prod_multiset_X_sub_C (hroots : p.roots.card = p.nat_degree) : C p.leading_coeff * (p.roots.map (λ a, X - C a)).prod = p := (eq_leading_coeff_mul_of_monic_of_dvd_of_nat_degree_le monic_prod_multiset_X_sub_C - (prod_multiset_X_sub_C_dvd p) $ - ((nat_degree_multiset_prod_X_sub_C_eq_card _).trans hroots).ge).symm + p.prod_multiset_X_sub_C_dvd ((nat_degree_multiset_prod_X_sub_C_eq_card _).trans hroots).ge).symm /-- A monic polynomial `p` that has as many roots as its degree can be written `p = ∏(X - a)`, for `a` in `p.roots`. -/ @@ -881,6 +843,75 @@ by { convert C_leading_coeff_mul_prod_multiset_X_sub_C hroots, rw [hp.leading_co end comm_ring +section +variables {A B : Type*} [comm_ring A] [comm_ring B] + +lemma le_root_multiplicity_map {p : A[X]} {f : A →+* B} (hmap : map f p ≠ 0) (a : A) : + root_multiplicity a p ≤ root_multiplicity (f a) (p.map f) := +begin + rw [le_root_multiplicity_iff hmap], + refine trans _ ((map_ring_hom f).map_dvd (pow_root_multiplicity_dvd p a)), + rw [map_pow, map_sub, coe_map_ring_hom, map_X, map_C], +end + +lemma eq_root_multiplicity_map {p : A[X]} {f : A →+* B} (hf : function.injective f) + (a : A) : root_multiplicity a p = root_multiplicity (f a) (p.map f) := +begin + by_cases hp0 : p = 0, { simp only [hp0, root_multiplicity_zero, polynomial.map_zero], }, + apply le_antisymm (le_root_multiplicity_map ((polynomial.map_ne_zero_iff hf).mpr hp0) a), + rw [le_root_multiplicity_iff hp0, ← map_dvd_map f hf ((monic_X_sub_C a).pow _), + polynomial.map_pow, polynomial.map_sub, map_X, map_C], + apply pow_root_multiplicity_dvd, +end + +lemma count_map_roots [is_domain A] {p : A[X]} {f : A →+* B} (hmap : map f p ≠ 0) (b : B) : + (p.roots.map f).count b ≤ root_multiplicity b (p.map f) := +begin + rw [le_root_multiplicity_iff hmap, ← multiset.prod_repeat, ← multiset.map_repeat (λ a, X - C a)], + rw ← multiset.filter_eq, + refine (multiset.prod_dvd_prod_of_le $ multiset.map_le_map $ multiset.filter_le _ _).trans _, + convert polynomial.map_dvd _ p.prod_multiset_X_sub_C_dvd, + simp only [polynomial.map_multiset_prod, multiset.map_map], + congr, ext1, + simp only [function.comp_app, polynomial.map_sub, map_X, map_C], +end + +lemma count_map_roots_of_injective [is_domain A] (p : A[X]) {f : A →+* B} + (hf : function.injective f) (b : B) : + (p.roots.map f).count b ≤ root_multiplicity b (p.map f) := +begin + by_cases hp0 : p = 0, + { simp only [hp0, roots_zero, multiset.map_zero, + multiset.count_zero, polynomial.map_zero, root_multiplicity_zero] }, + { exact count_map_roots ((polynomial.map_ne_zero_iff hf).mpr hp0) b }, +end + +lemma map_roots_le [is_domain A] [is_domain B] {p : A[X]} {f : A →+* B} (h : p.map f ≠ 0) : + p.roots.map f ≤ (p.map f).roots := +multiset.le_iff_count.2 $ λ b, by { rw count_roots, apply count_map_roots h } + +lemma map_roots_le_of_injective [is_domain A] [is_domain B] (p : A[X]) + {f : A →+* B} (hf : function.injective f) : + p.roots.map f ≤ (p.map f).roots := +begin + by_cases hp0 : p = 0, { simp only [hp0, roots_zero, multiset.map_zero, polynomial.map_zero], }, + exact map_roots_le ((polynomial.map_ne_zero_iff hf).mpr hp0), +end + +lemma card_roots_le_map [is_domain A] [is_domain B] {p : A[X]} {f : A →+* B} (h : p.map f ≠ 0) : + p.roots.card ≤ (p.map f).roots.card := +by { rw ← p.roots.card_map f, exact multiset.card_le_of_le (map_roots_le h) } + +lemma roots_map_of_injective_of_card_eq_nat_degree [is_domain A] [is_domain B] {p : A[X]} + {f : A →+* B} (hf : function.injective f) (hroots : p.roots.card = p.nat_degree) : + p.roots.map f = (p.map f).roots := +begin + apply multiset.eq_of_le_of_card_le (map_roots_le_of_injective p hf), + simpa only [multiset.card_map, hroots] using (card_roots' _).trans (nat_degree_map_le f p), +end + +end + section variables [semiring R] [comm_ring S] [is_domain S] (φ : R →+* S) diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field.lean index f4522aa6a7dcc..15420e98482cf 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field.lean @@ -273,7 +273,7 @@ by rw [degree_eq_nat_degree p_ne_zero, nat_degree_eq_card_roots hsplit] theorem roots_map {f : K[X]} (hf : f.splits $ ring_hom.id K) : (f.map i).roots = f.roots.map i := -(roots_map_of_injective_card_eq_total_degree i.injective $ +(roots_map_of_injective_of_card_eq_nat_degree i.injective $ by { convert (nat_degree_eq_card_roots hf).symm, rw map_id }).symm lemma image_root_set [algebra F K] [algebra F L] {p : F[X]} (h : p.splits (algebra_map F K)) From e18fa7b91131401311474bde7c82aa48919f34de Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Thu, 22 Sep 2022 10:57:53 +0000 Subject: [PATCH 348/828] feat(probability/martingale/centering): uniqueness of Doob's decomposition (#16532) --- src/probability/martingale/basic.lean | 35 +++++++++++++++++++++++ src/probability/martingale/centering.lean | 35 +++++++++++++++++++++++ src/probability/process/adapted.lean | 9 ++++++ 3 files changed, 79 insertions(+) diff --git a/src/probability/martingale/basic.lean b/src/probability/martingale/basic.lean index 8ec1911258fbb..e9e8381497f03 100644 --- a/src/probability/martingale/basic.lean +++ b/src/probability/martingale/basic.lean @@ -490,6 +490,41 @@ begin simpa only [pi.zero_apply, pi.neg_apply, zero_eq_neg], end +-- Note that one cannot use `submartingale.zero_le_of_predictable` to prove the other two +-- corresponding lemmas without imposing more restrictions to the ordering of `E` +/-- A predictable submartingale is a.e. greater equal than its initial state. -/ +lemma submartingale.zero_le_of_predictable [preorder E] [sigma_finite_filtration μ 𝒢] + {f : ℕ → Ω → E} (hfmgle : submartingale f 𝒢 μ) (hfadp : adapted 𝒢 (λ n, f (n + 1))) (n : ℕ) : + f 0 ≤ᵐ[μ] f n := +begin + induction n with k ih, + { refl }, + { exact ih.trans ((hfmgle.2.1 k (k + 1) k.le_succ).trans_eq $ germ.coe_eq.mp $ congr_arg coe $ + condexp_of_strongly_measurable (𝒢.le _) (hfadp _) $ hfmgle.integrable _) } +end + +/-- A predictable supermartingale is a.e. less equal than its initial state. -/ +lemma supermartingale.le_zero_of_predictable [preorder E] [sigma_finite_filtration μ 𝒢] + {f : ℕ → Ω → E} (hfmgle : supermartingale f 𝒢 μ) (hfadp : adapted 𝒢 (λ n, f (n + 1))) (n : ℕ) : + f n ≤ᵐ[μ] f 0 := +begin + induction n with k ih, + { refl }, + { exact ((germ.coe_eq.mp $ congr_arg coe $ condexp_of_strongly_measurable (𝒢.le _) (hfadp _) $ + hfmgle.integrable _).symm.trans_le (hfmgle.2.1 k (k + 1) k.le_succ)).trans ih } +end + +/-- A predictable martingale is a.e. equal to its initial state. -/ +lemma martingale.eq_zero_of_predicatable [sigma_finite_filtration μ 𝒢] + {f : ℕ → Ω → E} (hfmgle : martingale f 𝒢 μ) (hfadp : adapted 𝒢 (λ n, f (n + 1))) (n : ℕ) : + f n =ᵐ[μ] f 0 := +begin + induction n with k ih, + { refl }, + { exact ((germ.coe_eq.mp (congr_arg coe $ condexp_of_strongly_measurable (𝒢.le _) (hfadp _) + (hfmgle.integrable _))).symm.trans (hfmgle.2 k (k + 1) k.le_succ)).trans ih } +end + namespace submartingale @[protected] diff --git a/src/probability/martingale/centering.lean b/src/probability/martingale/centering.lean index 67cbf109ecc9c..7bda4b917d673 100644 --- a/src/probability/martingale/centering.lean +++ b/src/probability/martingale/centering.lean @@ -136,6 +136,41 @@ begin refl, end +-- The following two lemmas demonstrate the essential uniqueness of the decomposition +lemma martingale_part_add_ae_eq [sigma_finite_filtration μ ℱ] {f g : ℕ → Ω → E} + (hf : martingale f ℱ μ) (hg : adapted ℱ (λ n, g (n + 1))) (hg0 : g 0 = 0) + (hgint : ∀ n, integrable (g n) μ) (n : ℕ) : + martingale_part ℱ μ (f + g) n =ᵐ[μ] f n := +begin + set h := f - martingale_part ℱ μ (f + g) with hhdef, + have hh : h = predictable_part ℱ μ (f + g) - g, + { rw [hhdef, sub_eq_sub_iff_add_eq_add, add_comm (predictable_part ℱ μ (f + g)), + martingale_part_add_predictable_part] }, + have hhpred : adapted ℱ (λ n, h (n + 1)), + { rw hh, + exact adapted_predictable_part.sub hg }, + have hhmgle : martingale h ℱ μ := hf.sub (martingale_martingale_part + (hf.adapted.add $ predictable.adapted hg $ hg0.symm ▸ strongly_measurable_zero) $ + λ n, (hf.integrable n).add $ hgint n), + refine (eventually_eq_iff_sub.2 _).symm, + filter_upwards [hhmgle.eq_zero_of_predicatable hhpred n] with ω hω, + rw [hhdef, pi.sub_apply] at hω, + rw [hω, pi.sub_apply, martingale_part], + simp [hg0], +end + +lemma predictable_part_add_ae_eq [sigma_finite_filtration μ ℱ] {f g : ℕ → Ω → E} + (hf : martingale f ℱ μ) (hg : adapted ℱ (λ n, g (n + 1))) (hg0 : g 0 = 0) + (hgint : ∀ n, integrable (g n) μ) (n : ℕ) : + predictable_part ℱ μ (f + g) n =ᵐ[μ] g n := +begin + filter_upwards [martingale_part_add_ae_eq hf hg hg0 hgint n] with ω hω, + rw ← add_right_inj (f n ω), + conv_rhs { rw [← pi.add_apply, ← pi.add_apply, + ← martingale_part_add_predictable_part ℱ μ (f + g)] }, + rw [pi.add_apply, pi.add_apply, hω], +end + section difference lemma predictable_part_bdd_difference {R : ℝ≥0} {f : ℕ → Ω → ℝ} diff --git a/src/probability/process/adapted.lean b/src/probability/process/adapted.lean index 9e5095e49b3ef..cce063dfd46fd 100644 --- a/src/probability/process/adapted.lean +++ b/src/probability/process/adapted.lean @@ -216,4 +216,13 @@ begin { exact strongly_measurable_const, }, end +-- this dot notation will make more sense once we have a more general definition for predictable +lemma predictable.adapted {f : filtration ℕ m} {u : ℕ → Ω → β} + (hu : adapted f (λ n, u (n + 1))) (hu0 : strongly_measurable[f 0] (u 0)) : + adapted f u := +λ n, match n with + | 0 := hu0 + | n + 1 := (hu n).mono (f.mono n.le_succ) +end + end measure_theory From 690714361e87bd440445d2a10f52bc7b9d87119d Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 22 Sep 2022 13:12:11 +0000 Subject: [PATCH 349/828] feat(geometry/manifold/instances/sphere): the differential of the embedding of the sphere into its ambient space (#16466) Formalized as part of the Sphere Eversion project. --- src/geometry/manifold/instances/sphere.lean | 105 +++++++++++++++++++- 1 file changed, 104 insertions(+), 1 deletion(-) diff --git a/src/geometry/manifold/instances/sphere.lean b/src/geometry/manifold/instances/sphere.lean index 7424b08e368f3..cc216dc0997ef 100644 --- a/src/geometry/manifold/instances/sphere.lean +++ b/src/geometry/manifold/instances/sphere.lean @@ -59,7 +59,7 @@ variables {E : Type*} [inner_product_space ℝ E] noncomputable theory -open metric finite_dimensional +open metric finite_dimensional function open_locale manifold local attribute [instance] fact_finite_dimensional_of_finrank_eq_succ @@ -129,6 +129,36 @@ begin ring end +lemma has_fderiv_at_stereo_inv_fun_aux (v : E) : + has_fderiv_at (stereo_inv_fun_aux v) (continuous_linear_map.id ℝ E) 0 := +begin + have h₀ : has_fderiv_at (λ w : E, ∥w∥ ^ 2) (0 : E →L[ℝ] ℝ) 0, + { convert (has_strict_fderiv_at_norm_sq _).has_fderiv_at, + simp }, + have h₁ : has_fderiv_at (λ w : E, (∥w∥ ^ 2 + 4)⁻¹) (0 : E →L[ℝ] ℝ) 0, + { convert (has_fderiv_at_inv _).comp _ (h₀.add (has_fderiv_at_const 4 0)); simp }, + have h₂ : has_fderiv_at (λ w, (4:ℝ) • w + (∥w∥ ^ 2 - 4) • v) + ((4:ℝ) • continuous_linear_map.id ℝ E) 0, + { convert ((has_fderiv_at_const (4:ℝ) 0).smul (has_fderiv_at_id 0)).add + ((h₀.sub (has_fderiv_at_const (4:ℝ) 0)).smul (has_fderiv_at_const v 0)), + ext w, + simp, }, + convert h₁.smul h₂, + ext w, + simp, +end + +lemma has_fderiv_at_stereo_inv_fun_aux_comp_coe (v : E) : + has_fderiv_at (stereo_inv_fun_aux v ∘ (coe : (ℝ ∙ v)ᗮ → E)) (ℝ ∙ v)ᗮ.subtypeL 0 := +begin + have : has_fderiv_at + (stereo_inv_fun_aux v) + (continuous_linear_map.id ℝ E) + ((ℝ ∙ v)ᗮ.subtypeL 0) := + has_fderiv_at_stereo_inv_fun_aux v, + convert this.comp (0 : (ℝ ∙ v)ᗮ) (by apply continuous_linear_map.has_fderiv_at), +end + lemma cont_diff_stereo_inv_fun_aux : cont_diff ℝ ⊤ (stereo_inv_fun_aux v) := begin have h₀ : cont_diff ℝ ⊤ (λ w : E, ∥w∥ ^ 2) := cont_diff_norm_sq, @@ -265,6 +295,18 @@ rfl @[simp] lemma stereographic_target (hv : ∥v∥ = 1) : (stereographic hv).target = set.univ := rfl +@[simp] lemma stereographic_apply_neg (v : sphere (0:E) 1) : + stereographic (norm_eq_of_mem_sphere v) (-v) = 0 := +by simp [stereographic_apply, orthogonal_projection_orthogonal_complement_singleton_eq_zero] + +@[simp] lemma stereographic_neg_apply (v : sphere (0:E) 1) : + stereographic (norm_eq_of_mem_sphere (-v)) v = 0 := +begin + convert stereographic_apply_neg (-v), + ext1, + simp, +end + end stereographic_projection section charted_space @@ -414,6 +456,67 @@ begin exact cont_mdiff_coe_sphere, end +/-- Consider the differential of the inclusion of the sphere in `E` at the point `v` as a continuous +linear map from `tangent_space (𝓡 n) v` to `E`. The range of this map is the orthogonal complement +of `v` in `E`. + +Note that there is an abuse here of the defeq between `E` and the tangent space to `E` at `(v:E`). +In general this defeq is not canonical, but in this case (the tangent space of a vector space) it is +canonical. -/ +lemma range_mfderiv_coe_sphere {n : ℕ} [fact (finrank ℝ E = n + 1)] (v : sphere (0:E) 1) : + (mfderiv (𝓡 n) 𝓘(ℝ, E) (coe : sphere (0:E) 1 → E) v : tangent_space (𝓡 n) v →L[ℝ] E).range + = (ℝ ∙ (v:E))ᗮ := +begin + rw ((cont_mdiff_coe_sphere v).mdifferentiable_at le_top).mfderiv, + simp only [chart_at, stereographic', stereographic_neg_apply, fderiv_within_univ, + linear_isometry_equiv.to_homeomorph_symm, linear_isometry_equiv.coe_to_homeomorph, + linear_isometry_equiv.map_zero] with mfld_simps, + let U := + (orthonormal_basis.from_orthogonal_span_singleton n (ne_zero_of_mem_unit_sphere (-v))).repr, + change (fderiv ℝ ((stereo_inv_fun_aux (-v : E) ∘ coe) ∘ U.symm) 0).range = (ℝ ∙ (v:E))ᗮ, + have : has_fderiv_at + (stereo_inv_fun_aux (-v : E) ∘ (coe : (ℝ ∙ (↑(-v):E))ᗮ → E)) + (ℝ ∙ (↑(-v):E))ᗮ.subtypeL + (U.symm 0), + { convert has_fderiv_at_stereo_inv_fun_aux_comp_coe (-v:E), + simp }, + rw (this.comp 0 U.symm.to_continuous_linear_equiv.has_fderiv_at).fderiv, + convert (U.symm : euclidean_space ℝ (fin n) ≃ₗᵢ[ℝ] (ℝ ∙ (↑(-v):E))ᗮ).range_comp + (ℝ ∙ (↑(-v):E))ᗮ.subtype using 1, + simp only [submodule.range_subtype, coe_neg_sphere], + congr' 1, + -- we must show `submodule.span ℝ {v} = submodule.span ℝ {-v}` + apply submodule.span_eq_span, + { simp only [set.singleton_subset_iff, set_like.mem_coe], + rw ← submodule.neg_mem_iff, + exact submodule.mem_span_singleton_self (-v) }, + { simp only [set.singleton_subset_iff, set_like.mem_coe], + rw submodule.neg_mem_iff, + exact submodule.mem_span_singleton_self v }, +end + +/-- Consider the differential of the inclusion of the sphere in `E` at the point `v` as a continuous +linear map from `tangent_space (𝓡 n) v` to `E`. This map is injective. -/ +lemma mfderiv_coe_sphere_injective {n : ℕ} [fact (finrank ℝ E = n + 1)] (v : sphere (0:E) 1) : + injective (mfderiv (𝓡 n) 𝓘(ℝ, E) (coe : sphere (0:E) 1 → E) v) := +begin + rw ((cont_mdiff_coe_sphere v).mdifferentiable_at le_top).mfderiv, + simp only [chart_at, stereographic', stereographic_neg_apply, fderiv_within_univ, + linear_isometry_equiv.to_homeomorph_symm, linear_isometry_equiv.coe_to_homeomorph, + linear_isometry_equiv.map_zero] with mfld_simps, + let U := + (orthonormal_basis.from_orthogonal_span_singleton n (ne_zero_of_mem_unit_sphere (-v))).repr, + change injective (fderiv ℝ ((stereo_inv_fun_aux (-v : E) ∘ coe) ∘ U.symm) 0), + have : has_fderiv_at + (stereo_inv_fun_aux (-v : E) ∘ (coe : (ℝ ∙ (↑(-v):E))ᗮ → E)) + (ℝ ∙ (↑(-v):E))ᗮ.subtypeL + (U.symm 0), + { convert has_fderiv_at_stereo_inv_fun_aux_comp_coe (-v:E), + simp }, + rw (this.comp 0 U.symm.to_continuous_linear_equiv.has_fderiv_at).fderiv, + simpa using subtype.coe_injective, +end + end smooth_manifold section circle From 3dc8592b78ea1725eeebd5412678e20aa859feb0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 22 Sep 2022 16:25:53 +0000 Subject: [PATCH 350/828] feat(tactic/positivity): Extension for `coe` (#16141) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add: * a case in the main `positivity` tactic to show that `0 ≤ a` in a `canonically_ordered_add_monoid` * `positivity_coe`, an extension to `positivity`, to turn `0 ≤ a`/`0 < a` into `0 ≤ ↑a`/`0 < ↑a` when `a` is of type `ℕ`, `ℤ`, `ℚ`. --- src/data/real/ennreal.lean | 19 ++++++- src/data/real/ereal.lean | 57 ++++++++++++++++++-- src/data/real/hyperreal.lean | 24 ++++++++- src/data/real/nnreal.lean | 26 +++++++-- src/measure_theory/function/lp_space.lean | 4 +- src/tactic/positivity.lean | 64 ++++++++++++++++++++--- test/positivity.lean | 36 +++++++++++++ 7 files changed, 209 insertions(+), 21 deletions(-) diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 99a303554d9a8..c976cbe79a34a 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -211,7 +211,7 @@ lemma coe_mono : monotone (coe : ℝ≥0 → ℝ≥0∞) := λ _ _, coe_le_coe.2 @[simp, norm_cast] lemma zero_eq_coe : 0 = (↑r : ℝ≥0∞) ↔ 0 = r := coe_eq_coe @[simp, norm_cast] lemma coe_eq_one : (↑r : ℝ≥0∞) = 1 ↔ r = 1 := coe_eq_coe @[simp, norm_cast] lemma one_eq_coe : 1 = (↑r : ℝ≥0∞) ↔ 1 = r := coe_eq_coe -@[simp, norm_cast] lemma coe_nonneg : 0 ≤ (↑r : ℝ≥0∞) ↔ 0 ≤ r := coe_le_coe +@[simp] lemma coe_nonneg : 0 ≤ (↑r : ℝ≥0∞) := coe_le_coe.2 $ zero_le _ @[simp, norm_cast] lemma coe_pos : 0 < (↑r : ℝ≥0∞) ↔ 0 < r := coe_lt_coe lemma coe_ne_zero : (r : ℝ≥0∞) ≠ 0 ↔ r ≠ 0 := not_congr coe_eq_coe @@ -1931,3 +1931,20 @@ by simpa only [image_image] using h.image_real_to_nnreal.image_coe_nnreal_ennrea end ord_connected end set + +namespace tactic +open positivity + +private lemma nnreal_coe_pos {r : ℝ≥0} : 0 < r → 0 < (r : ℝ≥0∞) := ennreal.coe_pos.2 + +/-- Extension for the `positivity` tactic: cast from `ℝ≥0` to `ℝ≥0∞`. -/ +@[positivity] +meta def positivity_coe_nnreal_ennreal : expr → tactic strictness +| `(@coe _ _ %%inst %%a) := do + unify inst `(@coe_to_lift _ _ $ @coe_base _ _ ennreal.has_coe), + positive p ← core a, -- We already know `0 ≤ r` for all `r : ℝ≥0∞` + positive <$> mk_app ``nnreal_coe_pos [p] +| e := pp e >>= fail ∘ format.bracket "The expression " + " is not of the form `(r : ℝ≥0∞)` for `r : ℝ≥0`" + +end tactic diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index 3d989b14c8591..9ccc89c3ae9dd 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -154,7 +154,19 @@ by { apply with_top.coe_lt_coe.2, exact with_bot.bot_lt_coe _ } @[simp, norm_cast] lemma coe_add (x y : ℝ) : ((x + y : ℝ) : ereal) = (x : ereal) + (y : ereal) := rfl -@[simp] lemma coe_zero : ((0 : ℝ) : ereal) = 0 := rfl +@[simp, norm_cast] lemma coe_zero : ((0 : ℝ) : ereal) = 0 := rfl + +@[simp, norm_cast] protected lemma coe_nonneg {x : ℝ} : (0 : ereal) ≤ x ↔ 0 ≤ x := +ereal.coe_le_coe_iff + +@[simp, norm_cast] protected lemma coe_nonpos {x : ℝ} : (x : ereal) ≤ 0 ↔ x ≤ 0 := +ereal.coe_le_coe_iff + +@[simp, norm_cast] protected lemma coe_pos {x : ℝ} : (0 : ereal) < x ↔ 0 < x := +ereal.coe_lt_coe_iff + +@[simp, norm_cast] protected lemma coe_neg' {x : ℝ} : (x : ereal) < 0 ↔ x < 0 := +ereal.coe_lt_coe_iff lemma to_real_le_to_real {x y : ereal} (h : x ≤ y) (hx : x ≠ ⊥) (hy : y ≠ ⊤) : x.to_real ≤ y.to_real := @@ -214,7 +226,8 @@ end lemma coe_nnreal_eq_coe_real (x : ℝ≥0) : ((x : ℝ≥0∞) : ereal) = (x : ℝ) := rfl -@[simp] lemma coe_ennreal_top : ((⊤ : ℝ≥0∞) : ereal) = ⊤ := rfl +@[simp, norm_cast] lemma coe_ennreal_zero : ((0 : ℝ≥0∞) : ereal) = 0 := rfl +@[simp, norm_cast] lemma coe_ennreal_top : ((⊤ : ℝ≥0∞) : ereal) = ⊤ := rfl @[simp] lemma coe_ennreal_eq_top_iff : ∀ {x : ℝ≥0∞}, (x : ereal) = ⊤ ↔ x = ⊤ | ⊤ := by simp @@ -247,6 +260,9 @@ lemma coe_nnreal_ne_top (x : ℝ≥0) : ((x : ℝ≥0∞) : ereal) ≠ ⊤ := de lemma coe_ennreal_nonneg (x : ℝ≥0∞) : (0 : ereal) ≤ x := coe_ennreal_le_coe_ennreal_iff.2 (zero_le x) +@[simp, norm_cast] lemma coe_ennreal_pos {x : ℝ≥0∞} : (0 : ereal) < x ↔ 0 < x := +by rw [←coe_ennreal_zero, coe_ennreal_lt_coe_ennreal_iff] + @[simp] lemma bot_lt_coe_ennreal (x : ℝ≥0∞) : (⊥ : ereal) < x := (bot_lt_coe 0).trans_le (coe_ennreal_nonneg _) @@ -257,8 +273,6 @@ coe_ennreal_le_coe_ennreal_iff.2 (zero_le x) | x ⊤ := by simp | (some x) (some y) := rfl -@[simp] lemma coe_ennreal_zero : ((0 : ℝ≥0∞) : ereal) = 0 := rfl - /-! ### Order -/ @@ -514,3 +528,38 @@ lemma to_real_mul : ∀ {x y : ereal}, to_real (x * y) = to_real x * to_real y | ⊥ ⊥ := by simp end ereal + +namespace tactic +open positivity + +private lemma ereal_coe_nonneg {r : ℝ} : 0 ≤ r → 0 ≤ (r : ereal) := ereal.coe_nonneg.2 +private lemma ereal_coe_pos {r : ℝ} : 0 < r → 0 < (r : ereal) := ereal.coe_pos.2 +private lemma ereal_coe_ennreal_pos {r : ℝ≥0∞} : 0 < r → 0 < (r : ereal) := ereal.coe_ennreal_pos.2 + +/-- Extension for the `positivity` tactic: cast from `ℝ` to `ereal`. -/ +@[positivity] +meta def positivity_coe_real_ereal : expr → tactic strictness +| `(@coe _ _ %%inst %%a) := do + unify inst `(@coe_to_lift _ _ $ @coe_base _ _ ereal.has_coe), + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``ereal_coe_pos [p] + | nonnegative p := nonnegative <$> mk_mapp ``ereal_coe_nonneg [a, p] + end +| e := pp e >>= fail ∘ format.bracket "The expression " + " is not of the form `(r : ereal)` for `r : ℝ`" + +/-- Extension for the `positivity` tactic: cast from `ℝ≥0∞` to `ereal`. -/ +@[positivity] +meta def positivity_coe_ennreal_ereal : expr → tactic strictness +| `(@coe _ _ %%inst %%a) := do + unify inst `(@coe_to_lift _ _ $ @coe_base _ _ ereal.has_coe_ennreal), + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``ereal_coe_ennreal_pos [p] + | nonnegative _ := nonnegative <$> mk_mapp `ereal.coe_ennreal_nonneg [a] + end +| e := pp e >>= fail ∘ format.bracket "The expression " + " is not of the form `(r : ereal)` for `r : ℝ≥0∞`" + +end tactic diff --git a/src/data/real/hyperreal.lean b/src/data/real/hyperreal.lean index de5b5d7c49da3..5a6f648ed4a7d 100644 --- a/src/data/real/hyperreal.lean +++ b/src/data/real/hyperreal.lean @@ -42,9 +42,9 @@ germ.const_inj @[simp, norm_cast] lemma coe_sub (x y : ℝ) : ↑(x - y) = (x - y : ℝ*) := rfl @[simp, norm_cast] lemma coe_lt_coe {x y : ℝ} : (x : ℝ*) < y ↔ x < y := germ.const_lt -@[simp, norm_cast] lemma coe_pos {x : ℝ} : 0 < (x : ℝ*) ↔ 0 < x := -coe_lt_coe +@[simp, norm_cast] lemma coe_pos {x : ℝ} : 0 < (x : ℝ*) ↔ 0 < x := coe_lt_coe @[simp, norm_cast] lemma coe_le_coe {x y : ℝ} : (x : ℝ*) ≤ y ↔ x ≤ y := germ.const_le_iff +@[simp, norm_cast] lemma coe_nonneg {x : ℝ} : 0 ≤ (x : ℝ*) ↔ 0 ≤ x := coe_le_coe @[simp, norm_cast] lemma coe_abs (x : ℝ) : ((|x| : ℝ) : ℝ*) = |x| := begin convert const_abs x, @@ -789,3 +789,23 @@ lemma infinite_mul_infinite {x y : ℝ*} : infinite x → infinite y → infinit λ hx hy, infinite_mul_of_infinite_not_infinitesimal hx (not_infinitesimal_of_infinite hy) end hyperreal + +namespace tactic +open positivity + +private lemma hyperreal_coe_nonneg {r : ℝ} : 0 ≤ r → 0 ≤ (r : ℝ*) := hyperreal.coe_nonneg.2 +private lemma hyperreal_coe_pos {r : ℝ} : 0 < r → 0 < (r : ℝ*) := hyperreal.coe_pos.2 + +/-- Extension for the `positivity` tactic: cast from `ℝ` to `ℝ*`. -/ +@[positivity] +meta def positivity_coe_real_hyperreal : expr → tactic strictness +| `(@coe _ _ %%inst %%a) := do + unify inst `(@coe_to_lift _ _ hyperreal.has_coe_t), + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``hyperreal_coe_pos [p] + | nonnegative p := nonnegative <$> mk_app ``hyperreal_coe_nonneg [p] + end +| e := pp e >>= fail ∘ format.bracket "The expression " " is not of the form `(r : ℝ*)` for `r : ℝ`" + +end tactic diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 5190b4779d803..354b5efca9e9c 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -3,12 +3,10 @@ Copyright (c) 2018 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ -import algebra.big_operators.ring -import data.real.pointwise -import algebra.indicator_function import algebra.algebra.basic -import algebra.order.module import algebra.order.nonneg +import data.real.pointwise +import tactic.positivity /-! # Nonnegative real numbers @@ -840,3 +838,23 @@ lemma cast_nat_abs_eq_nnabs_cast (n : ℤ) : by { ext, rw [nnreal.coe_nat_cast, int.cast_nat_abs, real.coe_nnabs] } end real + +namespace tactic +open positivity + +private lemma nnreal_coe_pos {r : ℝ≥0} : 0 < r → 0 < (r : ℝ) := nnreal.coe_pos.2 + +/-- Extension for the `positivity` tactic: cast from `ℝ≥0` to `ℝ`. -/ +@[positivity] +meta def positivity_coe_nnreal_real : expr → tactic strictness +| `(@coe _ _ %%inst %%a) := do + unify inst `(@coe_to_lift _ _ $ @coe_base _ _ nnreal.real.has_coe), + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``nnreal_coe_pos [p] + | nonnegative _ := nonnegative <$> mk_app ``nnreal.coe_nonneg [a] + end +| e := pp e >>= fail ∘ format.bracket "The expression " + " is not of the form `(r : ℝ)` for `r : ℝ≥0`" + +end tactic diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 48991e413c3b5..9279ec7da8b48 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1816,8 +1816,8 @@ lemma snorm_ess_sup_indicator_const_le (s : set α) (c : G) : snorm_ess_sup (s.indicator (λ x : α , c)) μ ≤ ∥c∥₊ := begin by_cases hμ0 : μ = 0, - { rw [hμ0, snorm_ess_sup_measure_zero, ennreal.coe_nonneg], - exact zero_le', }, + { rw [hμ0, snorm_ess_sup_measure_zero], + exact ennreal.coe_nonneg }, { exact (snorm_ess_sup_indicator_le s (λ x, c)).trans (snorm_ess_sup_const c hμ0).le, }, end diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index db20574053c15..7bc39befbdf9e 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -49,9 +49,7 @@ introduce these operations. ## TODO -Implement extensions for other operations (raising to non-numeral powers, `exp`, `log`, coercions -from `ℕ` and `ℝ≥0`). - +Implement extensions for other operations (raising to non-numeral powers, `log`). -/ namespace tactic @@ -84,6 +82,11 @@ meta def norm_num.positivity (e : expr) : tactic strictness := do pure (nonnegative p') else failed +/-- Second base case of the `positivity` tactic: Any element of a canonically ordered additive +monoid is nonnegative. -/ +meta def positivity_canon : expr → tactic strictness +| `(%%a) := nonnegative <$> mk_app ``zero_le [a] + namespace positivity /-- Given two tactics whose result is `strictness`, report a `strictness`: @@ -102,7 +105,7 @@ meta def orelse' (tac1 tac2 : tactic strictness) : tactic strictness := do /-! ### Core logic of the `positivity` tactic -/ -/-- Second base case of the `positivity` tactic. Prove an expression `e` is positive/nonnegative by +/-- Third base case of the `positivity` tactic. Prove an expression `e` is positive/nonnegative by finding a hypothesis of the form `a < e` or `a ≤ e` in which `a` can be proved positive/nonnegative by `norm_num`. -/ meta def compare_hyp (e p₂ : expr) : tactic strictness := do @@ -140,9 +143,10 @@ meta def attr : user_attribute (expr → tactic strictness) unit := (λ _, failed), pure $ λ e, orelse' (t e) $ orelse' -- run all the extensions on `e` - (norm_num.positivity e) $ -- directly try `norm_num` on `e` - -- loop over hypotheses and try to compare with `e` - local_context >>= list.foldl (λ tac h, orelse' tac (compare_hyp e h)) failed }, + (norm_num.positivity e) $ orelse' -- directly try `norm_num` on `e` + (positivity_canon e) $ -- try showing nonnegativity from canonicity of the order + -- loop over hypotheses and try to compare with `e` + local_context >>= list.foldl (λ tac h, orelse' tac (compare_hyp e h)) failed }, dependencies := [] } } /-- Look for a proof of positivity/nonnegativity of an expression `e`; if found, return the proof @@ -200,7 +204,7 @@ add_tactic_doc end interactive -variables {R : Type*} +variables {α R : Type*} /-! ### `positivity` extensions for particular arithmetic operations -/ @@ -390,4 +394,48 @@ meta def positivity_abs : expr → tactic strictness nonnegative <$> mk_app ``abs_nonneg [a] -- else report nonnegativity | _ := failed +private lemma nat_cast_pos [ordered_semiring α] [nontrivial α] {n : ℕ} : 0 < n → 0 < (n : α) := +nat.cast_pos.2 + +private lemma int_coe_nat_nonneg (n : ℕ) : 0 ≤ (n : ℤ) := n.cast_nonneg +private lemma int_coe_nat_pos {n : ℕ} : 0 < n → 0 < (n : ℤ) := nat.cast_pos.2 + +private lemma int_cast_nonneg [ordered_ring α] {n : ℤ} (hn : 0 ≤ n) : 0 ≤ (n : α) := +by { rw ←int.cast_zero, exact int.cast_mono hn } +private lemma int_cast_pos [ordered_ring α] [nontrivial α] {n : ℤ} : 0 < n → 0 < (n : α) := +int.cast_pos.2 + +private lemma rat_cast_nonneg [linear_ordered_field α] {q : ℚ} : 0 ≤ q → 0 ≤ (q : α) := +rat.cast_nonneg.2 +private lemma rat_cast_pos [linear_ordered_field α] {q : ℚ} : 0 < q → 0 < (q : α) := rat.cast_pos.2 + +/-- Extension for the `positivity` tactic: casts from `ℕ`, `ℤ`, `ℚ`. -/ +@[positivity] +meta def positivity_coe : expr → tactic strictness +| `(@coe _ %%typ %%inst %%a) := do + -- TODO: Using `match` here might turn out too strict since we really want the instance to *unify* + -- with one of the instances below rather than being equal on the nose. + -- If this turns out to indeed be a problem, we should figure out the right way to pattern match + -- up to defeq rather than equality of expressions. + -- See also "Reflexive tactics for algebra, revisited" by Kazuhiko Sakaguchi at ITP 2022. + match inst with + | `(@coe_to_lift _ _ %%inst) := do + strictness_a ← core a, + match inst, strictness_a with -- `mk_mapp` is necessary in some places. Why? + | `(nat.cast_coe), positive p := positive <$> mk_app ``nat_cast_pos [p] + | `(nat.cast_coe), _ := nonnegative <$> mk_mapp ``nat.cast_nonneg [typ, none, a] + | `(int.cast_coe), positive p := positive <$> mk_mapp ``int_cast_pos [typ, none, none, none, p] + | `(int.cast_coe), nonnegative p := nonnegative <$> + mk_mapp ``int_cast_nonneg [typ, none, none, p] + | `(rat.cast_coe), positive p := positive <$> mk_mapp ``rat_cast_pos [typ, none, none, p] + | `(rat.cast_coe), nonnegative p := nonnegative <$> + mk_mapp ``rat_cast_nonneg [typ, none, none, p] + | `(@coe_base _ _ int.has_coe), positive p := positive <$> mk_app ``int_coe_nat_pos [p] + | `(@coe_base _ _ int.has_coe), _ := nonnegative <$> mk_app ``int_coe_nat_nonneg [a] + | _, _ := failed + end + | _ := failed + end +| _ := failed + end tactic diff --git a/test/positivity.lean b/test/positivity.lean index 507c74efdd0ae..804b2de9adab0 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -1,6 +1,9 @@ import algebra.order.smul import analysis.normed.group.basic import data.complex.exponential +import data.rat.nnrat +import data.real.ereal +import data.real.hyperreal import data.real.sqrt import tactic.positivity @@ -9,6 +12,8 @@ import tactic.positivity This tactic proves goals of the form `0 ≤ a` and `0 < a`. -/ +open_locale ennreal nnrat nnreal + /- ## Numeric goals -/ example : 0 ≤ 0 := by positivity @@ -106,6 +111,37 @@ example {X : Type*} [metric_space X] (x y : X) : 0 ≤ dist x y := by positivity example {E : Type*} [add_group E] {p : add_group_seminorm E} {x : E} : 0 ≤ p x := by positivity example {E : Type*} [group E] {p : group_seminorm E} {x : E} : 0 ≤ p x := by positivity +/- ### Canonical orders -/ + +example {a : ℕ} : 0 ≤ a := by positivity +example {a : ℚ≥0} : 0 ≤ a := by positivity +example {a : ℝ≥0} : 0 ≤ a := by positivity +example {a : ℝ≥0∞} : 0 ≤ a := by positivity + +/- ### Coercions -/ + +example {a : ℕ} : (0 : ℤ) ≤ a := by positivity +example {a : ℕ} (ha : 0 < a) : (0 : ℤ) < a := by positivity +example {a : ℤ} (ha : 0 ≤ a) : (0 : ℚ) ≤ a := by positivity +example {a : ℤ} (ha : 0 < a) : (0 : ℚ) < a := by positivity +example {a : ℚ} (ha : 0 ≤ a) : (0 : ℝ) ≤ a := by positivity +example {a : ℚ} (ha : 0 < a) : (0 : ℝ) < a := by positivity +example {r : ℝ≥0} : (0 : ℝ) ≤ r := by positivity +example {r : ℝ≥0} (hr : 0 < r) : (0 : ℝ) < r := by positivity +example {r : ℝ≥0} (hr : 0 < r) : (0 : ℝ≥0∞) < r := by positivity +-- example {r : ℝ≥0} : (0 : ereal) ≤ r := by positivity -- TODO: Handle `coe_trans` +-- example {r : ℝ≥0} (hr : 0 < r) : (0 : ereal) < r := by positivity +example {r : ℝ} (hr : 0 ≤ r) : (0 : ereal) ≤ r := by positivity +example {r : ℝ} (hr : 0 < r) : (0 : ereal) < r := by positivity +example {r : ℝ} (hr : 0 ≤ r) : (0 : hyperreal) ≤ r := by positivity +example {r : ℝ} (hr : 0 < r) : (0 : hyperreal) < r := by positivity +example {r : ℝ≥0∞} : (0 : ereal) ≤ r := by positivity +example {r : ℝ≥0∞} (hr : 0 < r) : (0 : ereal) < r := by positivity + +example {α : Type*} [ordered_ring α] {n : ℤ} : 0 ≤ ((n ^ 2 : ℤ) : α) := by positivity +example {r : ℝ≥0} : 0 ≤ ((r : ℝ) : ereal) := by positivity +example {r : ℝ≥0} : 0 < ((r + 1 : ℝ) : ereal) := by positivity + /- ## Tests that the tactic is agnostic on reversed inequalities -/ example {a : ℤ} (ha : a > 0) : 0 ≤ a := by positivity From 951bf1d9e98a2042979ced62c0620bcfb3587cf8 Mon Sep 17 00:00:00 2001 From: Wrenna Robson Date: Thu, 22 Sep 2022 18:41:27 +0000 Subject: [PATCH 351/828] refactor(data/set/basic): Remove _eq lemmas (#16572) This PR removes `_eq` lemmas from `data.set.basic` which asset a `rfl` propositional equality, in favour of the `iff` versions, and makes the latter `simp` lemmas if they weren't already. Co-authored-by: Wrenna Robson <34025592+linesthatinterlace@users.noreply.github.com> --- .../100-theorems-list/30_ballot_problem.lean | 10 ++--- counterexamples/phillips.lean | 2 +- src/algebra/indicator_function.lean | 4 +- src/algebra/order/floor.lean | 4 +- src/algebra/support.lean | 4 +- .../locally_ringed_space.lean | 6 +-- .../prime_spectrum/basic.lean | 8 ++-- .../prime_spectrum/is_open_comap_C.lean | 2 +- .../projective_spectrum/topology.lean | 6 +-- src/analysis/box_integral/basic.lean | 6 +-- .../box_integral/partition/basic.lean | 2 +- .../box_integral/partition/split.lean | 4 +- src/analysis/calculus/fderiv_measurable.lean | 8 ++-- src/analysis/complex/abs_max.lean | 4 +- src/analysis/convex/basic.lean | 2 +- src/analysis/convex/gauge.lean | 6 +-- src/analysis/convex/join.lean | 4 +- src/analysis/normed_space/lp_space.lean | 2 +- src/analysis/seminorm.lean | 2 +- .../special_functions/polar_coord.lean | 2 +- src/combinatorics/simple_graph/basic.lean | 6 +-- .../simple_graph/connectivity.lean | 2 +- .../simple_graph/inc_matrix.lean | 2 +- src/computability/halting.lean | 2 +- src/data/fin/tuple/basic.lean | 2 +- src/data/fintype/basic.lean | 4 +- src/data/nat/lattice.lean | 4 +- src/data/set/basic.lean | 37 ++++++---------- src/data/set/function.lean | 2 +- src/data/set/intervals/basic.lean | 2 +- .../set/intervals/unordered_interval.lean | 2 +- src/data/set/lattice.lean | 2 +- src/data/set/pairwise.lean | 2 +- src/data/set/prod.lean | 4 +- src/geometry/manifold/charted_space.lean | 2 +- src/geometry/manifold/cont_mdiff.lean | 2 +- src/geometry/manifold/diffeomorph.lean | 2 +- src/geometry/manifold/partition_of_unity.lean | 2 +- src/group_theory/perm/support.lean | 2 +- .../affine_space/independent.lean | 2 +- src/logic/equiv/local_equiv.lean | 6 +-- .../constructions/borel_space.lean | 10 ++--- .../covering/differentiation.lean | 6 +-- src/measure_theory/covering/vitali.lean | 2 +- .../covering/vitali_family.lean | 4 +- .../decomposition/lebesgue.lean | 4 +- .../decomposition/signed_hahn.lean | 4 +- .../function/ae_measurable_order.lean | 2 +- .../function/convergence_in_measure.lean | 6 +-- src/measure_theory/function/egorov.lean | 4 +- src/measure_theory/function/jacobian.lean | 2 +- .../function/simple_func_dense_lp.lean | 6 +-- .../function/strongly_measurable.lean | 10 ++--- src/measure_theory/group/arithmetic.lean | 2 +- src/measure_theory/integral/lebesgue.lean | 12 +++--- src/measure_theory/integral/set_integral.lean | 4 +- src/measure_theory/measurable_space.lean | 2 +- src/measure_theory/measure/ae_disjoint.lean | 2 +- src/measure_theory/measure/ae_measurable.lean | 6 +-- src/measure_theory/measure/haar.lean | 4 +- src/measure_theory/measure/measure_space.lean | 2 +- src/measure_theory/pi_system.lean | 2 +- src/model_theory/definability.lean | 4 +- src/model_theory/semantics.lean | 4 +- src/order/compactly_generated.lean | 2 +- src/order/filter/at_top_bot.lean | 2 +- src/order/filter/basic.lean | 2 +- src/order/filter/curry.lean | 2 +- src/probability/independence.lean | 2 +- .../martingale/borel_cantelli.lean | 2 +- src/probability/process/hitting_time.lean | 6 +-- src/probability/process/stopping.lean | 42 +++++++++---------- src/ring_theory/hahn_series.lean | 4 +- src/topology/algebra/order/basic.lean | 4 +- src/topology/basic.lean | 4 +- src/topology/category/Profinite/default.lean | 2 +- src/topology/category/Top/limits.lean | 2 +- src/topology/fiber_bundle.lean | 18 ++++---- src/topology/locally_constant/basic.lean | 2 +- src/topology/locally_finite.lean | 4 +- src/topology/metric_space/baire.lean | 2 +- src/topology/metric_space/basic.lean | 2 +- .../gromov_hausdorff_realized.lean | 2 +- .../metric_space/partition_of_unity.lean | 2 +- src/topology/paracompact.lean | 4 +- src/topology/shrinking_lemma.lean | 2 +- src/topology/sober.lean | 2 +- src/topology/subset_properties.lean | 14 +++---- src/topology/uniform_space/basic.lean | 4 +- .../uniform_space/compact_convergence.lean | 2 +- src/topology/vector_bundle/basic.lean | 4 +- 91 files changed, 206 insertions(+), 217 deletions(-) diff --git a/archive/100-theorems-list/30_ballot_problem.lean b/archive/100-theorems-list/30_ballot_problem.lean index 9e95d36c42195..4891ab851160e 100644 --- a/archive/100-theorems-list/30_ballot_problem.lean +++ b/archive/100-theorems-list/30_ballot_problem.lean @@ -320,7 +320,7 @@ lemma first_vote_pos : (nonempty_image_iff.2 (counted_sequence_nonempty _ _))], { have : list.cons (-1) '' counted_sequence (p + 1) q ∩ {l : list ℤ | l.head = 1} = ∅, { ext, - simp only [mem_inter_eq, mem_image, mem_set_of_eq, mem_empty_eq, iff_false, not_and, + simp only [mem_inter_iff, mem_image, mem_set_of_eq, mem_empty_iff_false, iff_false, not_and, forall_exists_index, and_imp], rintro l _ rfl, norm_num }, @@ -392,7 +392,7 @@ begin rw [counted_succ_succ, union_inter_distrib_right, (_ : list.cons (-1) '' counted_sequence (p + 1) q ∩ {l | l.head = 1} = ∅), union_empty]; { ext, - simp only [mem_inter_eq, mem_image, mem_set_of_eq, and_iff_left_iff_imp, mem_empty_eq, + simp only [mem_inter_iff, mem_image, mem_set_of_eq, and_iff_left_iff_imp, mem_empty_iff_false, iff_false, not_and, forall_exists_index, and_imp], rintro y hy rfl, norm_num } @@ -410,7 +410,7 @@ begin have : (counted_sequence p (q + 1)).image (list.cons 1) ∩ stays_positive = (counted_sequence p (q + 1) ∩ stays_positive).image (list.cons 1), { ext t, - simp only [mem_inter_eq, mem_image], + simp only [mem_inter_iff, mem_image], split, { simp only [and_imp, exists_imp_distrib], rintro l hl rfl t, @@ -433,7 +433,7 @@ begin rw [counted_succ_succ, union_inter_distrib_right, (_ : list.cons 1 '' counted_sequence p (q + 1) ∩ {l : list ℤ | l.head = 1}ᶜ = ∅), empty_union]; { ext, - simp only [mem_inter_eq, mem_image, mem_set_of_eq, and_iff_left_iff_imp, mem_empty_eq, + simp only [mem_inter_iff, mem_image, mem_set_of_eq, and_iff_left_iff_imp, mem_empty_iff_false, iff_false, not_and, forall_exists_index, and_imp], rintro y hy rfl, norm_num } @@ -451,7 +451,7 @@ begin have : (counted_sequence (p + 1) q).image (list.cons (-1)) ∩ stays_positive = ((counted_sequence (p + 1) q) ∩ stays_positive).image (list.cons (-1)), { ext t, - simp only [mem_inter_eq, mem_image], + simp only [mem_inter_iff, mem_image], split, { simp only [and_imp, exists_imp_distrib], rintro l hl rfl t, diff --git a/counterexamples/phillips.lean b/counterexamples/phillips.lean index 548429ff4b7f5..1fc2dfafe2b21 100644 --- a/counterexamples/phillips.lean +++ b/counterexamples/phillips.lean @@ -546,7 +546,7 @@ begin apply ae_restrict_of_ae, refine measure_mono_null _ ((countable_ne Hcont φ).measure_zero _), assume x, - simp only [imp_self, mem_set_of_eq, mem_compl_eq], + simp only [imp_self, mem_set_of_eq, mem_compl_iff], end lemma integrable_comp (Hcont : #ℝ = aleph 1) (φ : (discrete_copy ℝ →ᵇ ℝ) →L[ℝ] ℝ) : diff --git a/src/algebra/indicator_function.lean b/src/algebra/indicator_function.lean index 08f80f48aabd8..1e783b1b83f72 100644 --- a/src/algebra/indicator_function.lean +++ b/src/algebra/indicator_function.lean @@ -96,7 +96,7 @@ mul_indicator_eq_one @[to_additive] lemma mul_indicator_apply_ne_one {a : α} : s.mul_indicator f a ≠ 1 ↔ a ∈ s ∩ mul_support f := -by simp only [ne.def, mul_indicator_apply_eq_one, not_imp, mem_inter_eq, mem_mul_support] +by simp only [ne.def, mul_indicator_apply_eq_one, not_imp, mem_inter_iff, mem_mul_support] @[simp, to_additive] lemma mul_support_mul_indicator : function.mul_support (s.mul_indicator f) = s ∩ function.mul_support f := @@ -413,7 +413,7 @@ begin rw [finset.prod_insert haI, finset.set_bUnion_insert, mul_indicator_union_of_not_mem_inter, ih _], { assume i hi j hj hij, exact hI i (finset.mem_insert_of_mem hi) j (finset.mem_insert_of_mem hj) hij }, - simp only [not_exists, exists_prop, mem_Union, mem_inter_eq, not_and], + simp only [not_exists, exists_prop, mem_Union, mem_inter_iff, not_and], assume hx a' ha', refine disjoint_left.1 (hI a (finset.mem_insert_self _ _) a' (finset.mem_insert_of_mem ha') _) hx, exact (ne_of_mem_of_not_mem ha' haI).symm diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index 4c38e55e3fb12..eca1977d4102c 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -596,7 +596,7 @@ end lemma preimage_fract (s : set α) : fract ⁻¹' s = ⋃ m : ℤ, (λ x, x - m) ⁻¹' (s ∩ Ico (0 : α) 1) := begin ext x, - simp only [mem_preimage, mem_Union, mem_inter_eq], + simp only [mem_preimage, mem_Union, mem_inter_iff], refine ⟨λ h, ⟨⌊x⌋, h, fract_nonneg x, fract_lt_one x⟩, _⟩, rintro ⟨m, hms, hm0, hm1⟩, obtain rfl : ⌊x⌋ = m, from floor_eq_iff.2 ⟨sub_nonneg.1 hm0, sub_lt_iff_lt_add'.1 hm1⟩, @@ -606,7 +606,7 @@ end lemma image_fract (s : set α) : fract '' s = ⋃ m : ℤ, (λ x, x - m) '' s ∩ Ico 0 1 := begin ext x, - simp only [mem_image, mem_inter_eq, mem_Union], split, + simp only [mem_image, mem_inter_iff, mem_Union], split, { rintro ⟨y, hy, rfl⟩, exact ⟨⌊y⌋, ⟨y, hy, rfl⟩, fract_nonneg y, fract_lt_one y⟩ }, { rintro ⟨m, ⟨y, hys, rfl⟩, h0, h1⟩, diff --git a/src/algebra/support.lean b/src/algebra/support.lean index 07661bf500dbc..3f139fa4dd09d 100644 --- a/src/algebra/support.lean +++ b/src/algebra/support.lean @@ -149,7 +149,7 @@ rfl @[to_additive support_prod_mk] lemma mul_support_prod_mk (f : α → M) (g : α → N) : mul_support (λ x, (f x, g x)) = mul_support f ∪ mul_support g := -set.ext $ λ x, by simp only [mul_support, not_and_distrib, mem_union_eq, mem_set_of_eq, +set.ext $ λ x, by simp only [mul_support, not_and_distrib, mem_union, mem_set_of_eq, prod.mk_eq_one, ne.def] @[to_additive support_prod_mk'] lemma mul_support_prod_mk' (f : α → M × N) : @@ -200,7 +200,7 @@ end division_monoid @[simp] lemma support_mul [mul_zero_class R] [no_zero_divisors R] (f g : α → R) : support (λ x, f x * g x) = support f ∩ support g := -set.ext $ λ x, by simp only [mem_support, mul_ne_zero_iff, mem_inter_eq, not_or_distrib] +set.ext $ λ x, by simp only [mem_support, mul_ne_zero_iff, mem_inter_iff, not_or_distrib] @[simp] lemma support_mul_subset_left [mul_zero_class R] (f g : α → R) : support (λ x, f x * g x) ⊆ support f := diff --git a/src/algebraic_geometry/locally_ringed_space.lean b/src/algebraic_geometry/locally_ringed_space.lean index 036fc1011fc45..cdb31489ff5b8 100644 --- a/src/algebraic_geometry/locally_ringed_space.lean +++ b/src/algebraic_geometry/locally_ringed_space.lean @@ -250,9 +250,9 @@ end X.to_RingedSpace.basic_open (0 : X.presheaf.obj $ op U) = ∅ := begin ext, - simp only [set.mem_empty_eq, topological_space.opens.empty_eq, topological_space.opens.mem_coe, - opens.coe_bot, iff_false, RingedSpace.basic_open, is_unit_zero_iff, set.mem_set_of_eq, - map_zero], + simp only [set.mem_empty_iff_false, topological_space.opens.empty_eq, + topological_space.opens.mem_coe, opens.coe_bot, iff_false, RingedSpace.basic_open, + is_unit_zero_iff, set.mem_set_of_eq, map_zero], rintro ⟨⟨y, _⟩, h, e⟩, exact @zero_ne_one (X.presheaf.stalk y) _ _ h, end diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index 5e0b3631d604e..54352ff14e92e 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -332,7 +332,7 @@ end lemma mem_compl_zero_locus_iff_not_mem {f : R} {I : prime_spectrum R} : I ∈ (zero_locus {f} : set (prime_spectrum R))ᶜ ↔ f ∉ I.as_ideal := -by rw [set.mem_compl_eq, mem_zero_locus, set.singleton_subset_iff]; refl +by rw [set.mem_compl_iff, mem_zero_locus, set.singleton_subset_iff]; refl /-- The Zariski topology on the prime spectrum of a commutative ring is defined via the closed sets of the topology: @@ -659,7 +659,7 @@ lemma is_open_basic_open {a : R} : is_open ((basic_open a) : set (prime_spectrum @[simp] lemma basic_open_eq_zero_locus_compl (r : R) : (basic_open r : set (prime_spectrum R)) = (zero_locus {r})ᶜ := -set.ext $ λ x, by simpa only [set.mem_compl_eq, mem_zero_locus, set.singleton_subset_iff] +set.ext $ λ x, by simpa only [set.mem_compl_iff, mem_zero_locus, set.singleton_subset_iff] @[simp] lemma basic_open_one : basic_open (1 : R) = ⊤ := topological_space.opens.ext $ by simp @@ -692,7 +692,7 @@ begin { rintros _ ⟨r, rfl⟩, exact is_open_basic_open }, { rintros p U hp ⟨s, hs⟩, - rw [← compl_compl U, set.mem_compl_eq, ← hs, mem_zero_locus, set.not_subset] at hp, + rw [← compl_compl U, set.mem_compl_iff, ← hs, mem_zero_locus, set.not_subset] at hp, obtain ⟨f, hfs, hfp⟩ := hp, refine ⟨basic_open f, ⟨f, rfl⟩, hfp, _⟩, rw [← set.compl_subset_compl, ← hs, basic_open_eq_zero_locus_compl, compl_compl], @@ -747,7 +747,7 @@ begin rw localization_comap_range S (submonoid.powers r), ext, simp only [mem_zero_locus, basic_open_eq_zero_locus_compl, set_like.mem_coe, set.mem_set_of_eq, - set.singleton_subset_iff, set.mem_compl_eq], + set.singleton_subset_iff, set.mem_compl_iff], split, { intros h₁ h₂, exact h₁ ⟨submonoid.mem_powers r, h₂⟩ }, diff --git a/src/algebraic_geometry/prime_spectrum/is_open_comap_C.lean b/src/algebraic_geometry/prime_spectrum/is_open_comap_C.lean index 436f841b76d8f..219830a0d1aa7 100644 --- a/src/algebraic_geometry/prime_spectrum/is_open_comap_C.lean +++ b/src/algebraic_geometry/prime_spectrum/is_open_comap_C.lean @@ -49,7 +49,7 @@ lemma image_of_Df_eq_comap_C_compl_zero_locus : image_of_Df f = prime_spectrum.comap (C : R →+* R[X]) '' (zero_locus {f})ᶜ := begin refine ext (λ x, ⟨λ hx, ⟨⟨map C x.val, (is_prime_map_C_of_is_prime x.property)⟩, ⟨_, _⟩⟩, _⟩), - { rw [mem_compl_eq, mem_zero_locus, singleton_subset_iff], + { rw [mem_compl_iff, mem_zero_locus, singleton_subset_iff], cases hx with i hi, exact λ a, hi (mem_map_C_iff.mp a i) }, { refine subtype.ext (ext (λ x, ⟨λ h, _, λ h, subset_span (mem_image_of_mem C.1 h)⟩)), diff --git a/src/algebraic_geometry/projective_spectrum/topology.lean b/src/algebraic_geometry/projective_spectrum/topology.lean index cc3cc424344d9..7614c39336f9f 100644 --- a/src/algebraic_geometry/projective_spectrum/topology.lean +++ b/src/algebraic_geometry/projective_spectrum/topology.lean @@ -287,7 +287,7 @@ end lemma mem_compl_zero_locus_iff_not_mem {f : A} {I : projective_spectrum 𝒜} : I ∈ (zero_locus 𝒜 {f} : set (projective_spectrum 𝒜))ᶜ ↔ f ∉ I.as_homogeneous_ideal := -by rw [set.mem_compl_eq, mem_zero_locus, set.singleton_subset_iff]; refl +by rw [set.mem_compl_iff, mem_zero_locus, set.singleton_subset_iff]; refl /-- The Zariski topology on the prime spectrum of a commutative ring is defined via the closed sets of the topology: @@ -364,7 +364,7 @@ lemma is_open_basic_open {a : A} : is_open ((basic_open 𝒜 a) : @[simp] lemma basic_open_eq_zero_locus_compl (r : A) : (basic_open 𝒜 r : set (projective_spectrum 𝒜)) = (zero_locus 𝒜 {r})ᶜ := -set.ext $ λ x, by simpa only [set.mem_compl_eq, mem_zero_locus, set.singleton_subset_iff] +set.ext $ λ x, by simpa only [set.mem_compl_iff, mem_zero_locus, set.singleton_subset_iff] @[simp] lemma basic_open_one : basic_open 𝒜 (1 : A) = ⊤ := topological_space.opens.ext $ by simp @@ -408,7 +408,7 @@ begin { rintros _ ⟨r, rfl⟩, exact is_open_basic_open 𝒜 }, { rintros p U hp ⟨s, hs⟩, - rw [← compl_compl U, set.mem_compl_eq, ← hs, mem_zero_locus, set.not_subset] at hp, + rw [← compl_compl U, set.mem_compl_iff, ← hs, mem_zero_locus, set.not_subset] at hp, obtain ⟨f, hfs, hfp⟩ := hp, refine ⟨basic_open 𝒜 f, ⟨f, rfl⟩, hfp, _⟩, rw [← set.compl_subset_compl, ← hs, basic_open_eq_zero_locus_compl, compl_compl], diff --git a/src/analysis/box_integral/basic.lean b/src/analysis/box_integral/basic.lean index 2e0b1f8b82283..93a271902d31c 100644 --- a/src/analysis/box_integral/basic.lean +++ b/src/analysis/box_integral/basic.lean @@ -207,7 +207,7 @@ begin (l.has_basis_to_filter_Union_top _).prod_self.tendsto_iff uniformity_basis_dist_le], refine forall₂_congr (λ ε ε0, exists_congr $ λ r, _), simp only [exists_prop, prod.forall, set.mem_Union, exists_imp_distrib, - prod_mk_mem_set_prod_eq, and_imp, mem_inter_eq, mem_set_of_eq], + prod_mk_mem_set_prod_eq, and_imp, mem_inter_iff, mem_set_of_eq], exact and_congr iff.rfl ⟨λ H c₁ c₂ π₁ π₂ h₁ hU₁ h₂ hU₂, H π₁ π₂ c₁ h₁ hU₁ c₂ h₂ hU₂, λ H π₁ π₂ c₁ h₁ hU₁ c₂ h₂ hU₂, H c₁ c₂ π₁ π₂ h₁ hU₁ h₂ hU₂⟩ end @@ -593,7 +593,7 @@ lemma tendsto_integral_sum_sum_integral (h : integrable I l f vol) (π₀ : prep begin refine ((l.has_basis_to_filter_Union I π₀).tendsto_iff nhds_basis_closed_ball).2 (λ ε ε0, _), refine ⟨h.convergence_r ε, h.convergence_r_cond ε, _⟩, - simp only [mem_inter_eq, set.mem_Union, mem_set_of_eq], + simp only [mem_inter_iff, set.mem_Union, mem_set_of_eq], rintro π ⟨c, hc, hU⟩, exact h.dist_integral_sum_sum_integral_le_of_mem_base_set_of_Union_eq ε0 hc hU end @@ -707,7 +707,7 @@ begin rcases exists_pos_mul_lt ε0' (B I) with ⟨ε', ε'0, hεI⟩, set δ : ℝ≥0 → ℝⁿ → Ioi (0 : ℝ) := λ c x, if x ∈ s then δ₁ c x (εs x) else (δ₂ c) x ε', refine ⟨δ, λ c, l.r_cond_of_bRiemann_eq_ff hl, _⟩, - simp only [set.mem_Union, mem_inter_eq, mem_set_of_eq], + simp only [set.mem_Union, mem_inter_iff, mem_set_of_eq], rintro π ⟨c, hπδ, hπp⟩, /- Now we split the sum into two parts based on whether `π.tag J` belongs to `s` or not. -/ rw [← g.sum_partition_boxes le_rfl hπp, mem_closed_ball, integral_sum, diff --git a/src/analysis/box_integral/partition/basic.lean b/src/analysis/box_integral/partition/basic.lean index 81ee6e1645dbc..98b348d3ca48e 100644 --- a/src/analysis/box_integral/partition/basic.lean +++ b/src/analysis/box_integral/partition/basic.lean @@ -434,7 +434,7 @@ begin refine (eq_of_boxes_subset_Union_superset (λ J₁ h₁, _) _).symm, { refine (mem_restrict _).2 ⟨J₁, π.mem_bUnion.2 ⟨J, hJ, h₁⟩, (inf_of_le_right _).symm⟩, exact with_bot.coe_le_coe.2 (le_of_mem _ h₁) }, - { simp only [Union_restrict, Union_bUnion, set.subset_def, set.mem_inter_eq, set.mem_Union], + { simp only [Union_restrict, Union_bUnion, set.subset_def, set.mem_inter_iff, set.mem_Union], rintro x ⟨hxJ, J₁, h₁, hx⟩, obtain rfl : J = J₁, from π.eq_of_mem_of_mem hJ h₁ hxJ (Union_subset _ hx), exact hx } diff --git a/src/analysis/box_integral/partition/split.lean b/src/analysis/box_integral/partition/split.lean index 7a3dd3c13e81e..5401c4d92424a 100644 --- a/src/analysis/box_integral/partition/split.lean +++ b/src/analysis/box_integral/partition/split.lean @@ -58,7 +58,7 @@ mk' I.lower (update I.upper i (min x (I.upper i))) begin rw [split_lower, coe_mk'], ext y, - simp only [mem_univ_pi, mem_Ioc, mem_inter_eq, mem_coe, mem_set_of_eq, forall_and_distrib, + simp only [mem_univ_pi, mem_Ioc, mem_inter_iff, mem_coe, mem_set_of_eq, forall_and_distrib, ← pi.le_def, le_update_iff, le_min_iff, and_assoc, and_forall_ne i, mem_def], rw [and_comm (y i ≤ x), pi.le_def] end @@ -91,7 +91,7 @@ mk' (update I.lower i (max x (I.lower i))) I.upper begin rw [split_upper, coe_mk'], ext y, - simp only [mem_univ_pi, mem_Ioc, mem_inter_eq, mem_coe, mem_set_of_eq, forall_and_distrib, + simp only [mem_univ_pi, mem_Ioc, mem_inter_iff, mem_coe, mem_set_of_eq, forall_and_distrib, forall_update_iff I.lower (λ j z, z < y j), max_lt_iff, and_assoc (x < y i), and_forall_ne i, mem_def], exact and_comm _ _ diff --git a/src/analysis/calculus/fderiv_measurable.lean b/src/analysis/calculus/fderiv_measurable.lean index 4b00960f73588..0423f9a2bc0a4 100644 --- a/src/analysis/calculus/fderiv_measurable.lean +++ b/src/analysis/calculus/fderiv_measurable.lean @@ -211,7 +211,7 @@ begin rcases mem_A_of_differentiable this hx.1 with ⟨R, R_pos, hR⟩, obtain ⟨n, hn⟩ : ∃ (n : ℕ), (1/2) ^ n < R := exists_pow_lt_of_lt_one R_pos (by norm_num : (1 : ℝ)/2 < 1), - simp only [mem_Union, mem_Inter, B, mem_inter_eq], + simp only [mem_Union, mem_Inter, B, mem_inter_iff], refine ⟨n, λ p hp q hq, ⟨fderiv 𝕜 f x, hx.2, ⟨_, _⟩⟩⟩; { refine hR _ ⟨pow_pos (by norm_num) _, lt_of_le_of_lt _ hn⟩, exact pow_le_pow_of_le_one (by norm_num) (by norm_num) (by assumption) } @@ -474,9 +474,9 @@ lemma B_mem_nhds_within_Ioi {K : set F} {r s ε x : ℝ} (hx : x ∈ B f K r s B f K r s ε ∈ 𝓝[>] x := begin obtain ⟨L, LK, hL₁, hL₂⟩ : ∃ (L : F), L ∈ K ∧ x ∈ A f L r ε ∧ x ∈ A f L s ε, - by simpa only [B, mem_Union, mem_inter_eq, exists_prop] using hx, + by simpa only [B, mem_Union, mem_inter_iff, exists_prop] using hx, filter_upwards [A_mem_nhds_within_Ioi hL₁, A_mem_nhds_within_Ioi hL₂] with y hy₁ hy₂, - simp only [B, mem_Union, mem_inter_eq, exists_prop], + simp only [B, mem_Union, mem_inter_iff, exists_prop], exact ⟨L, LK, hy₁, hy₂⟩ end @@ -567,7 +567,7 @@ begin rcases mem_A_of_differentiable this hx.1 with ⟨R, R_pos, hR⟩, obtain ⟨n, hn⟩ : ∃ (n : ℕ), (1/2) ^ n < R := exists_pow_lt_of_lt_one R_pos (by norm_num : (1 : ℝ)/2 < 1), - simp only [mem_Union, mem_Inter, B, mem_inter_eq], + simp only [mem_Union, mem_Inter, B, mem_inter_iff], refine ⟨n, λ p hp q hq, ⟨deriv_within f (Ici x) x, hx.2, ⟨_, _⟩⟩⟩; { refine hR _ ⟨pow_pos (by norm_num) _, lt_of_le_of_lt _ hn⟩, exact pow_le_pow_of_le_one (by norm_num) (by norm_num) (by assumption) } diff --git a/src/analysis/complex/abs_max.lean b/src/analysis/complex/abs_max.lean index 9c42071ab16b7..151176781c868 100644 --- a/src/analysis/complex/abs_max.lean +++ b/src/analysis/complex/abs_max.lean @@ -369,7 +369,7 @@ begin have hc : is_compact (closure U), from hb.is_compact_closure, obtain ⟨w, hwU, hle⟩ : ∃ w ∈ closure U, is_max_on (norm ∘ f) (closure U) w, from hc.exists_forall_ge hne.closure hd.continuous_on.norm, - rw [closure_eq_interior_union_frontier, mem_union_eq] at hwU, + rw [closure_eq_interior_union_frontier, mem_union] at hwU, cases hwU, rotate, { exact ⟨w, hwU, hle⟩ }, have : interior U ≠ univ, from ne_top_of_le_ne_top hc.ne_univ interior_subset_closure, rcases exists_mem_frontier_inf_dist_compl_eq_dist hwU this with ⟨z, hzU, hzw⟩, @@ -386,7 +386,7 @@ lemma norm_le_of_forall_mem_frontier_norm_le {f : E → F} {U : set E} (hU : bou {z : E} (hz : z ∈ closure U) : ∥f z∥ ≤ C := begin - rw [closure_eq_self_union_frontier, union_comm, mem_union_eq] at hz, + rw [closure_eq_self_union_frontier, union_comm, mem_union] at hz, cases hz, { exact hC z hz }, /- In case of a finite dimensional domain, one can just apply `complex.exists_mem_frontier_is_max_on_norm`. To make it work in any Banach space, we restrict diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index baea8f43976b9..0ebac40c09dc4 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -542,7 +542,7 @@ def std_simplex : set (ι → 𝕜) := lemma std_simplex_eq_inter : std_simplex 𝕜 ι = (⋂ x, {f | 0 ≤ f x}) ∩ {f | ∑ x, f x = 1} := -by { ext f, simp only [std_simplex, set.mem_inter_eq, set.mem_Inter, set.mem_set_of_eq] } +by { ext f, simp only [std_simplex, set.mem_inter_iff, set.mem_Inter, set.mem_set_of_eq] } lemma convex_std_simplex : convex 𝕜 (std_simplex 𝕜 ι) := begin diff --git a/src/analysis/convex/gauge.lean b/src/analysis/convex/gauge.lean index d2c2f50bb6135..bc3874fe379c9 100644 --- a/src/analysis/convex/gauge.lean +++ b/src/analysis/convex/gauge.lean @@ -102,7 +102,7 @@ begin end @[simp] lemma gauge_empty : gauge (∅ : set E) = 0 := -by { ext, simp only [gauge_def', real.Inf_empty, mem_empty_eq, pi.zero_apply, sep_false] } +by { ext, simp only [gauge_def', real.Inf_empty, mem_empty_iff_false, pi.zero_apply, sep_false] } lemma gauge_of_subset_zero (h : s ⊆ 0) : gauge s = 0 := by { obtain rfl | rfl := subset_singleton_iff_eq.1 h, exacts [gauge_empty, gauge_zero'] } @@ -221,7 +221,7 @@ begin rw [gauge_def', gauge_def', ←real.Inf_smul_of_nonneg ha], congr' 1, ext r, - simp_rw [set.mem_smul_set, set.mem_sep_eq], + simp_rw [set.mem_smul_set, set.mem_sep_iff], split, { rintro ⟨hr, hx⟩, simp_rw mem_Ioi at ⊢ hr, @@ -249,7 +249,7 @@ begin rw [gauge_def', pi.smul_apply, gauge_def', ←real.Inf_smul_of_nonneg (inv_nonneg.2 ha)], congr' 1, ext r, - simp_rw [set.mem_smul_set, set.mem_sep_eq], + simp_rw [set.mem_smul_set, set.mem_sep_iff], split, { rintro ⟨hr, y, hy, h⟩, simp_rw [mem_Ioi] at ⊢ hr, diff --git a/src/analysis/convex/join.lean b/src/analysis/convex/join.lean index 4461311991b0b..1080ba9973642 100644 --- a/src/analysis/convex/join.lean +++ b/src/analysis/convex/join.lean @@ -57,11 +57,11 @@ by simp [convex_join] @[simp] lemma convex_join_union_left (s₁ s₂ t : set E) : convex_join 𝕜 (s₁ ∪ s₂) t = convex_join 𝕜 s₁ t ∪ convex_join 𝕜 s₂ t := -by simp_rw [convex_join, mem_union_eq, Union_or, Union_union_distrib] +by simp_rw [convex_join, mem_union, Union_or, Union_union_distrib] @[simp] lemma convex_join_union_right (s t₁ t₂ : set E) : convex_join 𝕜 s (t₁ ∪ t₂) = convex_join 𝕜 s t₁ ∪ convex_join 𝕜 s t₂ := -by simp_rw [convex_join, mem_union_eq, Union_or, Union_union_distrib] +by simp_rw [convex_join, mem_union, Union_or, Union_union_distrib] @[simp] lemma convex_join_Union_left (s : ι → set E) (t : set E) : convex_join 𝕜 (⋃ i, s i) t = ⋃ i, convex_join 𝕜 (s i) t := diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index 63ebebb327537..aa19fce311d62 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -212,7 +212,7 @@ begin rcases p.trichotomy with rfl | rfl | hp, { apply mem_ℓp_zero, refine (hf.finite_dsupport.union hg.finite_dsupport).subset (λ i, _), - simp only [pi.add_apply, ne.def, set.mem_union_eq, set.mem_set_of_eq], + simp only [pi.add_apply, ne.def, set.mem_union, set.mem_set_of_eq], contrapose!, rintros ⟨hf', hg'⟩, simp [hf', hg'] }, diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 984c5cc823e2d..e686006c2c223 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -517,7 +517,7 @@ end @[simp] lemma ball_eq_emptyset (p : seminorm 𝕜 E) {x : E} {r : ℝ} (hr : r ≤ 0) : p.ball x r = ∅ := begin ext, - rw [seminorm.mem_ball, set.mem_empty_eq, iff_false, not_lt], + rw [seminorm.mem_ball, set.mem_empty_iff_false, iff_false, not_lt], exact hr.trans (map_nonneg p _), end diff --git a/src/analysis/special_functions/polar_coord.lean b/src/analysis/special_functions/polar_coord.lean index 7bf5c29748baa..f2679d0c748f5 100644 --- a/src/analysis/special_functions/polar_coord.lean +++ b/src/analysis/special_functions/polar_coord.lean @@ -111,7 +111,7 @@ lemma polar_coord_source_ae_eq_univ : begin have A : polar_coord.sourceᶜ ⊆ (linear_map.snd ℝ ℝ ℝ).ker, { assume x hx, - simp only [polar_coord_source, compl_union, mem_inter_eq, mem_compl_eq, mem_set_of_eq, not_lt, + simp only [polar_coord_source, compl_union, mem_inter_iff, mem_compl_iff, mem_set_of_eq, not_lt, not_not] at hx, exact hx.2 }, have B : volume ((linear_map.snd ℝ ℝ ℝ).ker : set (ℝ × ℝ)) = 0, diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index f3c745efcdc02..1676bc9ca8183 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -447,7 +447,7 @@ by rwa [←mk_mem_incidence_set_left_iff, lemma incidence_set_inter_incidence_set_of_not_adj (h : ¬ G.adj a b) (hn : a ≠ b) : G.incidence_set a ∩ G.incidence_set b = ∅ := begin - simp_rw [set.eq_empty_iff_forall_not_mem, set.mem_inter_eq, not_and], + simp_rw [set.eq_empty_iff_forall_not_mem, set.mem_inter_iff, not_and], intros u ha hb, exact h (G.adj_of_mem_incidence_set hn ha hb), end @@ -493,7 +493,7 @@ lemma adj_incidence_set_inter {v : V} {e : sym2 V} (he : e ∈ G.edge_set) (h : G.incidence_set v ∩ G.incidence_set h.other = {e} := begin ext e', - simp only [incidence_set, set.mem_sep_eq, set.mem_inter_eq, set.mem_singleton_iff], + simp only [incidence_set, set.mem_sep_iff, set.mem_inter_iff, set.mem_singleton_iff], refine ⟨λ h', _, _⟩, { rw ←sym2.other_spec h, exact (sym2.mem_and_mem_iff (edge_other_ne G he h).symm).mp ⟨h'.1.2, h'.2.2⟩ }, @@ -515,7 +515,7 @@ lemma neighbor_set_union_compl_neighbor_set_eq (G : simple_graph V) (v : V) : begin ext w, have h := @ne_of_adj _ G, - simp_rw [set.mem_union, mem_neighbor_set, compl_adj, set.mem_compl_eq, set.mem_singleton_iff], + simp_rw [set.mem_union, mem_neighbor_set, compl_adj, set.mem_compl_iff, set.mem_singleton_iff], tauto, end diff --git a/src/combinatorics/simple_graph/connectivity.lean b/src/combinatorics/simple_graph/connectivity.lean index 4b5e0aec79048..d8f1b83be76f3 100644 --- a/src/combinatorics/simple_graph/connectivity.lean +++ b/src/combinatorics/simple_graph/connectivity.lean @@ -1327,7 +1327,7 @@ lemma set_walk_length_zero_eq_of_ne {u v : V} (h : u ≠ v) : {p : G.walk u v | p.length = 0} = ∅ := begin ext p, - simp only [set.mem_set_of_eq, set.mem_empty_eq, iff_false], + simp only [set.mem_set_of_eq, set.mem_empty_iff_false, iff_false], exact λ h', absurd (walk.eq_of_length_eq_zero h') h, end diff --git a/src/combinatorics/simple_graph/inc_matrix.lean b/src/combinatorics/simple_graph/inc_matrix.lean index 359048d787636..f7cf9d646cda0 100644 --- a/src/combinatorics/simple_graph/inc_matrix.lean +++ b/src/combinatorics/simple_graph/inc_matrix.lean @@ -70,7 +70,7 @@ lemma inc_matrix_apply_mul_inc_matrix_apply : begin classical, simp only [inc_matrix, set.indicator_apply, ←ite_and_mul_zero, - pi.one_apply, mul_one, set.mem_inter_eq], + pi.one_apply, mul_one, set.mem_inter_iff], end lemma inc_matrix_apply_mul_inc_matrix_apply_of_not_adj (hab : a ≠ b) (h : ¬ G.adj a b) : diff --git a/src/computability/halting.lean b/src/computability/halting.lean index 0f0f670810b66..d53b1ea727d88 100644 --- a/src/computability/halting.lean +++ b/src/computability/halting.lean @@ -208,7 +208,7 @@ from λ f, ⟨set.mem_image_of_mem _, λ ⟨g, hg, e⟩, (H _ _ e).1 hg⟩, (partrec.nat_iff.1 $ eval_part.comp (const cf) computable.id) (partrec.nat_iff.1 $ eval_part.comp (const cg) computable.id) ((hC _).1 fC), -λ h, by obtain rfl | rfl := h; simp [computable_pred, set.mem_empty_eq]; +λ h, by obtain rfl | rfl := h; simp [computable_pred, set.mem_empty_iff_false]; exact ⟨by apply_instance, computable.const _⟩⟩ theorem halting_problem_re (n) : re_pred (λ c, (eval c n).dom) := diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index 58987a34a94ed..a1777cb57645e 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -569,7 +569,7 @@ set.ext $ λ p, by simp only [mem_preimage, insert_nth_mem_Icc, hx, true_and] lemma preimage_insert_nth_Icc_of_not_mem {i : fin (n + 1)} {x : α i} {q₁ q₂ : Π j, α j} (hx : x ∉ Icc (q₁ i) (q₂ i)) : i.insert_nth x ⁻¹' (Icc q₁ q₂) = ∅ := -set.ext $ λ p, by simp only [mem_preimage, insert_nth_mem_Icc, hx, false_and, mem_empty_eq] +set.ext $ λ p, by simp only [mem_preimage, insert_nth_mem_Icc, hx, false_and, mem_empty_iff_false] end insert_nth diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 4b43f316e564e..6a667aaa9b7dd 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -1293,7 +1293,7 @@ instance subtype.fintype (p : α → Prop) [decidable_pred p] [fintype α] : fin fintype.subtype (univ.filter p) (by simp) @[simp] lemma set.to_finset_eq_empty_iff {s : set α} [fintype s] : s.to_finset = ∅ ↔ s = ∅ := -by simp only [ext_iff, set.ext_iff, set.mem_to_finset, not_mem_empty, set.mem_empty_eq] +by simp only [ext_iff, set.ext_iff, set.mem_to_finset, not_mem_empty, set.mem_empty_iff_false] @[simp] lemma set.to_finset_empty : (∅ : set α).to_finset = ∅ := set.to_finset_eq_empty_iff.mpr rfl @@ -1564,7 +1564,7 @@ begin classical, rw [fintype.card_of_subtype (set.to_finset pᶜ), set.to_finset_compl p, finset.card_compl, fintype.card_of_subtype (set.to_finset p)]; - intro; simp only [set.mem_to_finset, set.mem_compl_eq]; refl, + intro; simp only [set.mem_to_finset, set.mem_compl_iff]; refl, end theorem fintype.card_subtype_mono (p q : α → Prop) (h : p ≤ q) diff --git a/src/data/nat/lattice.lean b/src/data/nat/lattice.lean index cdc59f7f06f1c..38a3d9f4d692a 100644 --- a/src/data/nat/lattice.lean +++ b/src/data/nat/lattice.lean @@ -40,7 +40,7 @@ dif_neg $ λ ⟨n, hn⟩, let ⟨k, hks, hk⟩ := h.exists_nat_lt n in (hn k hks begin cases eq_empty_or_nonempty s, { subst h, simp only [or_true, eq_self_iff_true, iff_true, Inf, has_Inf.Inf, - mem_empty_eq, exists_false, dif_neg, not_false_iff] }, + mem_empty_iff_false, exists_false, dif_neg, not_false_iff] }, { have := ne_empty_iff_nonempty.mpr h, simp only [this, or_false, nat.Inf_def, h, nat.find_eq_zero] } end @@ -104,7 +104,7 @@ noncomputable instance : conditionally_complete_linear_order_bot ℕ := cInf_le := assume s a hb ha, by rw [Inf_def ⟨a, ha⟩]; exact nat.find_min' _ ha, cSup_empty := begin - simp only [Sup_def, set.mem_empty_eq, forall_const, forall_prop_of_false, not_false_iff, + simp only [Sup_def, set.mem_empty_iff_false, forall_const, forall_prop_of_false, not_false_iff, exists_const], apply bot_unique (nat.find_min' _ _), trivial diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 4e3af611b712c..75d7a3d5e7bd6 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -195,7 +195,7 @@ nevertheless be useful for various reasons, e.g. to apply further projection not argument to `simp`. -/ lemma _root_.has_mem.mem.out {p : α → Prop} {a : α} (h : a ∈ {x | p x}) : p a := h -theorem nmem_set_of_eq {a : α} {p : α → Prop} : a ∉ {x | p x} = ¬ p a := rfl +theorem nmem_set_of_iff {a : α} {p : α → Prop} : a ∉ {x | p x} ↔ ¬ p a := iff.rfl @[simp] theorem set_of_mem_eq {s : set α} : {x | x ∈ s} = s := rfl @@ -358,7 +358,7 @@ nonempty_subtype.mp ‹_› theorem empty_def : (∅ : set α) = {x | false} := rfl -@[simp] theorem mem_empty_eq (x : α) : x ∈ (∅ : set α) = false := rfl +@[simp] theorem mem_empty_iff_false (x : α) : x ∈ (∅ : set α) ↔ false := iff.rfl @[simp] theorem set_of_false : {a : α | false} = ∅ := rfl @@ -468,9 +468,7 @@ theorem mem_union.elim {x : α} {a b : set α} {P : Prop} (H₁ : x ∈ a ∪ b) (H₂ : x ∈ a → P) (H₃ : x ∈ b → P) : P := or.elim H₁ H₂ H₃ -theorem mem_union (x : α) (a b : set α) : x ∈ a ∪ b ↔ x ∈ a ∨ x ∈ b := iff.rfl - -@[simp] theorem mem_union_eq (x : α) (a b : set α) : x ∈ a ∪ b = (x ∈ a ∨ x ∈ b) := rfl +@[simp] theorem mem_union (x : α) (a b : set α) : x ∈ a ∪ b ↔ (x ∈ a ∨ x ∈ b) := iff.rfl @[simp] theorem union_self (a : set α) : a ∪ a = a := ext $ λ x, or_self _ @@ -546,9 +544,7 @@ by simp only [← subset_empty_iff]; exact union_subset_iff theorem inter_def {s₁ s₂ : set α} : s₁ ∩ s₂ = {a | a ∈ s₁ ∧ a ∈ s₂} := rfl -theorem mem_inter_iff (x : α) (a b : set α) : x ∈ a ∩ b ↔ x ∈ a ∧ x ∈ b := iff.rfl - -@[simp] theorem mem_inter_eq (x : α) (a b : set α) : x ∈ a ∩ b = (x ∈ a ∧ x ∈ b) := rfl +@[simp] theorem mem_inter_iff (x : α) (a b : set α) : x ∈ a ∩ b ↔ (x ∈ a ∧ x ∈ b) := iff.rfl theorem mem_inter {x : α} {a b : set α} (ha : x ∈ a) (hb : x ∈ b) : x ∈ a ∩ b := ⟨ha, hb⟩ @@ -794,7 +790,7 @@ theorem set_compr_eq_eq_singleton {a : α} : {b | b = a} = {a} := rfl @[simp] theorem union_singleton : s ∪ {a} = insert a s := union_comm _ _ @[simp] theorem singleton_inter_nonempty : ({a} ∩ s).nonempty ↔ a ∈ s := -by simp only [set.nonempty, mem_inter_eq, mem_singleton_iff, exists_eq_left] +by simp only [set.nonempty, mem_inter_iff, mem_singleton_iff, exists_eq_left] @[simp] theorem inter_singleton_nonempty : (s ∩ {a}).nonempty ↔ a ∈ s := by rw [inter_comm, singleton_inter_nonempty] @@ -827,10 +823,7 @@ theorem mem_sep {s : set α} {p : α → Prop} {x : α} (xs : x ∈ s) (px : p x @[simp] theorem sep_mem_eq {s t : set α} : {x ∈ s | x ∈ t} = s ∩ t := rfl -@[simp] theorem mem_sep_eq {s : set α} {p : α → Prop} {x : α} : - x ∈ {x ∈ s | p x} = (x ∈ s ∧ p x) := rfl - -theorem mem_sep_iff {s : set α} {p : α → Prop} {x : α} : x ∈ {x ∈ s | p x} ↔ x ∈ s ∧ p x := +@[simp] theorem mem_sep_iff {s : set α} {p : α → Prop} {x : α} : x ∈ {x ∈ s | p x} ↔ x ∈ s ∧ p x := iff.rfl theorem eq_sep_of_subset {s t : set α} (h : s ⊆ t) : s = {x ∈ t | x ∈ s} := @@ -901,9 +894,7 @@ lemma compl_set_of {α} (p : α → Prop) : {a | p a}ᶜ = { a | ¬ p a } := rfl theorem not_mem_of_mem_compl {s : set α} {x : α} (h : x ∈ sᶜ) : x ∉ s := h -@[simp] theorem mem_compl_eq (s : set α) (x : α) : x ∈ sᶜ = (x ∉ s) := rfl - -theorem mem_compl_iff (s : set α) (x : α) : x ∈ sᶜ ↔ x ∉ s := iff.rfl +@[simp] theorem mem_compl_iff (s : set α) (x : α) : x ∈ sᶜ ↔ (x ∉ s) := iff.rfl lemma not_mem_compl_iff {x : α} : x ∉ sᶜ ↔ x ∈ s := not_not @@ -1275,7 +1266,7 @@ ite_same t (s ∩ s') ▸ ite_mono _ (inter_subset_left _ _) (inter_subset_right lemma ite_inter_inter (t s₁ s₂ s₁' s₂' : set α) : t.ite (s₁ ∩ s₂) (s₁' ∩ s₂') = t.ite s₁ s₁' ∩ t.ite s₂ s₂' := -by { ext x, simp only [set.ite, set.mem_inter_eq, set.mem_diff, set.mem_union_eq], itauto } +by { ext x, simp only [set.ite, set.mem_inter_iff, set.mem_diff, set.mem_union], itauto } lemma ite_inter (t s₁ s₂ s : set α) : t.ite (s₁ ∩ s) (s₂ ∩ s) = t.ite s₁ s₂ ∩ s := @@ -1379,8 +1370,6 @@ infix ` '' `:80 := image theorem mem_image_iff_bex {f : α → β} {s : set α} {y : β} : y ∈ f '' s ↔ ∃ x (_ : x ∈ s), f x = y := bex_def.symm -theorem mem_image_eq (f : α → β) (s : set α) (y: β) : y ∈ f '' s = ∃ x, x ∈ s ∧ f x = y := rfl - @[simp] theorem mem_image (f : α → β) (s : set α) (y : β) : y ∈ f '' s ↔ ∃ x, x ∈ s ∧ f x = y := iff.rfl @@ -1447,7 +1436,7 @@ h.set_image /-- Image is monotone with respect to `⊆`. See `set.monotone_image` for the statement in terms of `≤`. -/ theorem image_subset {a b : set α} (f : α → β) (h : a ⊆ b) : f '' a ⊆ f '' b := -by { simp only [subset_def, mem_image_eq], exact λ x, λ ⟨w, h1, h2⟩, ⟨w, h h1, h2⟩ } +by { simp only [subset_def, mem_image], exact λ x, λ ⟨w, h1, h2⟩, ⟨w, h h1, h2⟩ } theorem image_union (f : α → β) (s t : set α) : f '' (s ∪ t) = f '' s ∪ f '' t := @@ -2598,8 +2587,6 @@ variables {s s' : set α} {t t' : set β} {u u' : set γ} {a a' : α} {b b' : β def image2 (f : α → β → γ) (s : set α) (t : set β) : set γ := {c | ∃ a b, a ∈ s ∧ b ∈ t ∧ f a b = c } -lemma mem_image2_eq : c ∈ image2 f s t = ∃ a b, a ∈ s ∧ b ∈ t ∧ f a b = c := rfl - @[simp] lemma mem_image2 : c ∈ image2 f s t ↔ ∃ a b, a ∈ s ∧ b ∈ t ∧ f a b = c := iff.rfl lemma mem_image2_of_mem (h1 : a ∈ s) (h2 : b ∈ t) : f a b ∈ image2 f s t := @@ -2636,14 +2623,16 @@ lemma image2_union_left : image2 f (s ∪ s') t = image2 f s t ∪ image2 f s' t begin ext c, split, { rintros ⟨a, b, h1a|h2a, hb, rfl⟩;[left, right]; exact ⟨_, _, ‹_›, ‹_›, rfl⟩ }, - { rintro (⟨_, _, _, _, rfl⟩|⟨_, _, _, _, rfl⟩); refine ⟨_, _, _, ‹_›, rfl⟩; simp [mem_union, *] } + { rintro (⟨_, _, _, _, rfl⟩|⟨_, _, _, _, rfl⟩); refine ⟨_, _, _, ‹_›, rfl⟩; + simp [mem_union, *] } end lemma image2_union_right : image2 f s (t ∪ t') = image2 f s t ∪ image2 f s t' := begin ext c, split, { rintros ⟨a, b, ha, h1b|h2b, rfl⟩;[left, right]; exact ⟨_, _, ‹_›, ‹_›, rfl⟩ }, - { rintro (⟨_, _, _, _, rfl⟩|⟨_, _, _, _, rfl⟩); refine ⟨_, _, ‹_›, _, rfl⟩; simp [mem_union, *] } + { rintro (⟨_, _, _, _, rfl⟩|⟨_, _, _, _, rfl⟩); refine ⟨_, _, ‹_›, _, rfl⟩; + simp [mem_union, *] } end @[simp] lemma image2_empty_left : image2 f ∅ t = ∅ := ext $ by simp diff --git a/src/data/set/function.lean b/src/data/set/function.lean index b88e3aaada131..938aa1b4ce16b 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -1052,7 +1052,7 @@ by { intros i ht, by_cases hs : i ∈ s; simp [hf i ht, hg i ht, hs] } pi s (s'.piecewise t t') = pi (s ∩ s') t ∩ pi (s \ s') t' := begin ext x, - simp only [mem_pi, mem_inter_eq, ← forall_and_distrib], + simp only [mem_pi, mem_inter_iff, ← forall_and_distrib], refine forall_congr (λ i, _), by_cases hi : i ∈ s'; simp * end diff --git a/src/data/set/intervals/basic.lean b/src/data/set/intervals/basic.lean index af601cd479b82..7f8ead612f168 100644 --- a/src/data/set/intervals/basic.lean +++ b/src/data/set/intervals/basic.lean @@ -912,7 +912,7 @@ begin ext x, cases lt_or_le x b with hba hba, { simp [hba, h₁] }, - { simp only [mem_Iio, mem_union_eq, mem_Ioo, lt_max_iff], + { simp only [mem_Iio, mem_union, mem_Ioo, lt_max_iff], refine or_congr iff.rfl ⟨and.right, _⟩, exact λ h₂, ⟨h₁.trans_le hba, h₂⟩ }, end diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index 9461ef2701ef5..33d7df4ff9b6f 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -153,7 +153,7 @@ by cases le_total a b; simp [interval_oc, *] lemma forall_interval_oc_iff {P : α → Prop} : (∀ x ∈ Ι a b, P x) ↔ (∀ x ∈ Ioc a b, P x) ∧ (∀ x ∈ Ioc b a, P x) := -by simp only [interval_oc_eq_union, mem_union_eq, or_imp_distrib, forall_and_distrib] +by simp only [interval_oc_eq_union, mem_union, or_imp_distrib, forall_and_distrib] lemma interval_oc_subset_interval_oc_of_interval_subset_interval {a b c d : α} (h : [a, b] ⊆ [c, d]) : Ι a b ⊆ Ι c d := diff --git a/src/data/set/lattice.lean b/src/data/set/lattice.lean index 6c38cce8a7b6e..dd78286c6fb12 100644 --- a/src/data/set/lattice.lean +++ b/src/data/set/lattice.lean @@ -1696,7 +1696,7 @@ by simpa using h.preimage f lemma preimage_eq_empty_iff {s : set β} : f ⁻¹' s = ∅ ↔ disjoint s (range f) := ⟨λ h, begin simp only [eq_empty_iff_forall_not_mem, disjoint_iff_inter_eq_empty, not_exists, - mem_inter_eq, not_and, mem_range, mem_preimage] at h ⊢, + mem_inter_iff, not_and, mem_range, mem_preimage] at h ⊢, assume y hy x hx, rw ← hx at hy, exact h x hy, diff --git a/src/data/set/pairwise.lean b/src/data/set/pairwise.lean index e5fee97a3f25b..dbdeca2ae35d1 100644 --- a/src/data/set/pairwise.lean +++ b/src/data/set/pairwise.lean @@ -151,7 +151,7 @@ lemma pairwise_union : (s ∪ t).pairwise r ↔ s.pairwise r ∧ t.pairwise r ∧ ∀ (a ∈ s) (b ∈ t), a ≠ b → r a b ∧ r b a := begin - simp only [set.pairwise, mem_union_eq, or_imp_distrib, forall_and_distrib], + simp only [set.pairwise, mem_union, or_imp_distrib, forall_and_distrib], exact ⟨λ H, ⟨H.1.1, H.2.2, H.2.1, λ x hx y hy hne, H.1.2 y hy x hx hne.symm⟩, λ H, ⟨⟨H.1, λ x hx y hy hne, H.2.2.2 y hy x hx hne.symm⟩, H.2.2.1, H.2.1⟩⟩ end diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index 9bdeb92e401a6..222851ed0ee69 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -355,13 +355,13 @@ lemma pi_mono (h : ∀ i ∈ s, t₁ i ⊆ t₂ i) : pi s t₁ ⊆ pi s t₂ := λ x hx i hi, (h i hi $ hx i hi) lemma pi_inter_distrib : s.pi (λ i, t i ∩ t₁ i) = s.pi t ∩ s.pi t₁ := -ext $ λ x, by simp only [forall_and_distrib, mem_pi, mem_inter_eq] +ext $ λ x, by simp only [forall_and_distrib, mem_pi, mem_inter_iff] lemma pi_congr (h : s₁ = s₂) (h' : ∀ i ∈ s₁, t₁ i = t₂ i) : s₁.pi t₁ = s₂.pi t₂ := h ▸ (ext $ λ x, forall₂_congr $ λ i hi, h' i hi ▸ iff.rfl) lemma pi_eq_empty (hs : i ∈ s) (ht : t i = ∅) : s.pi t = ∅ := -by { ext f, simp only [mem_empty_eq, not_forall, iff_false, mem_pi, not_imp], +by { ext f, simp only [mem_empty_iff_false, not_forall, iff_false, mem_pi, not_imp], exact ⟨i, hs, by simp [ht]⟩ } lemma univ_pi_eq_empty (ht : t i = ∅) : pi univ t = ∅ := pi_eq_empty (mem_univ i) ht diff --git a/src/geometry/manifold/charted_space.lean b/src/geometry/manifold/charted_space.lean index 9e7705a5ed866..10d6c976af42a 100644 --- a/src/geometry/manifold/charted_space.lean +++ b/src/geometry/manifold/charted_space.lean @@ -262,7 +262,7 @@ instance : order_bot (structure_groupoid H) := apply u.id_mem }, { apply u.locality, assume x hx, - rw [hf, mem_empty_eq] at hx, + rw [hf, mem_empty_iff_false] at hx, exact hx.elim } end } diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index 4b7650eb57e3d..b9202c2bced10 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -948,7 +948,7 @@ begin (inter_mem _ self_mem_nhds_within)).congr_of_eventually_eq _ _, { filter_upwards [A], rintro x' ⟨hx', ht, hfx', hgfx'⟩, - simp only [*, mem_preimage, written_in_ext_chart_at, (∘), mem_inter_eq, e'.left_inv, true_and], + simp only [*, mem_preimage, written_in_ext_chart_at, (∘), mem_inter_iff, e'.left_inv, true_and], exact mem_range_self _ }, { filter_upwards [A], rintro x' ⟨hx', ht, hfx', hgfx'⟩, diff --git a/src/geometry/manifold/diffeomorph.lean b/src/geometry/manifold/diffeomorph.lean index cead8e7582a07..244e0eaa36750 100644 --- a/src/geometry/manifold/diffeomorph.lean +++ b/src/geometry/manifold/diffeomorph.lean @@ -411,7 +411,7 @@ def to_trans_diffeomorph (e : E ≃ₘ[𝕜] F) : M ≃ₘ⟮I, I.trans_diffeomo begin refine cont_mdiff_within_at_iff'.2 ⟨continuous_within_at_id, _⟩, refine e.symm.cont_diff.cont_diff_within_at.congr' (λ y hy, _) _, - { simp only [mem_inter_eq, I.ext_chart_at_trans_diffeomorph_target] at hy, + { simp only [mem_inter_iff, I.ext_chart_at_trans_diffeomorph_target] at hy, simp only [equiv.coe_refl, equiv.refl_symm, id, (∘), I.coe_ext_chart_at_trans_diffeomorph_symm, (ext_chart_at I x).right_inv hy.1] }, exact ⟨(ext_chart_at _ x).map_source (mem_ext_chart_source _ x), trivial, diff --git a/src/geometry/manifold/partition_of_unity.lean b/src/geometry/manifold/partition_of_unity.lean index 8d129b488da36..ba080f814e3e4 100644 --- a/src/geometry/manifold/partition_of_unity.lean +++ b/src/geometry/manifold/partition_of_unity.lean @@ -512,7 +512,7 @@ lemma emetric.exists_smooth_forall_closed_ball_subset {M} [emetric_space M] [cha (hfin : locally_finite K) : ∃ δ : C^∞⟮I, M; 𝓘(ℝ, ℝ), ℝ⟯, (∀ x, 0 < δ x) ∧ ∀ i (x ∈ K i), emetric.closed_ball x (ennreal.of_real (δ x)) ⊆ U i := -by simpa only [mem_inter_eq, forall_and_distrib, mem_preimage, mem_Inter, @forall_swap ι M] +by simpa only [mem_inter_iff, forall_and_distrib, mem_preimage, mem_Inter, @forall_swap ι M] using exists_smooth_forall_mem_convex_of_local_const I emetric.exists_forall_closed_ball_subset_aux₂ (emetric.exists_forall_closed_ball_subset_aux₁ hK hU hKU hfin) diff --git a/src/group_theory/perm/support.lean b/src/group_theory/perm/support.lean index 05851de4298be..602a4d84feb29 100644 --- a/src/group_theory/perm/support.lean +++ b/src/group_theory/perm/support.lean @@ -237,7 +237,7 @@ lemma set_support_mul_subset : {x | (p * q) x ≠ x} ⊆ {x | p x ≠ x} ∪ {x | q x ≠ x} := begin intro x, - simp only [perm.coe_mul, function.comp_app, ne.def, set.mem_union_eq, set.mem_set_of_eq], + simp only [perm.coe_mul, function.comp_app, ne.def, set.mem_union, set.mem_set_of_eq], by_cases hq : q x = x; simp [hq] end diff --git a/src/linear_algebra/affine_space/independent.lean b/src/linear_algebra/affine_space/independent.lean index 3128935830d75..542d5b51c80d2 100644 --- a/src/linear_algebra/affine_space/independent.lean +++ b/src/linear_algebra/affine_space/independent.lean @@ -534,7 +534,7 @@ begin rw [set.image_insert_eq, ← set.image_comp], simp, }, { use p, - simp only [equiv.coe_vadd_const, set.singleton_union, set.mem_inter_eq, coe_affine_span], + simp only [equiv.coe_vadd_const, set.singleton_union, set.mem_inter_iff, coe_affine_span], exact ⟨mem_span_points k _ _ (set.mem_insert p _), mem_span_points k _ _ hp⟩, }, }, end diff --git a/src/logic/equiv/local_equiv.lean b/src/logic/equiv/local_equiv.lean index ffdc7f1f2075f..558e1f788c99f 100644 --- a/src/logic/equiv/local_equiv.lean +++ b/src/logic/equiv/local_equiv.lean @@ -81,7 +81,7 @@ enough. -- with manifolds attribute [mfld_simps] id.def function.comp.left_id set.mem_set_of_eq set.image_eq_empty set.univ_inter set.preimage_univ set.prod_mk_mem_set_prod_eq and_true set.mem_univ -set.mem_image_of_mem true_and set.mem_inter_eq set.mem_preimage function.comp_app +set.mem_image_of_mem true_and set.mem_inter_iff set.mem_preimage function.comp_app set.inter_subset_left set.mem_prod set.range_id set.range_prod_map and_self set.mem_range_self eq_self_iff_true forall_const forall_true_iff set.inter_univ set.preimage_id function.comp.right_id not_false_iff and_imp set.prod_inter_prod set.univ_prod_univ true_or or_true prod.map_mk @@ -290,7 +290,7 @@ lemma symm_image_eq (h : e.is_image s t) : e.symm '' (e.target ∩ t) = e.source h.symm.image_eq lemma iff_preimage_eq : e.is_image s t ↔ e.source ∩ e ⁻¹' t = e.source ∩ s := -by simp only [is_image, set.ext_iff, mem_inter_eq, and.congr_right_iff, mem_preimage] +by simp only [is_image, set.ext_iff, mem_inter_iff, and.congr_right_iff, mem_preimage] alias iff_preimage_eq ↔ preimage_eq of_preimage_eq @@ -623,7 +623,7 @@ begin split, { simp [he.1] }, { assume x hx, - simp only [mem_inter_eq, restr_source] at hx, + simp only [mem_inter_iff, restr_source] at hx, exact he.2 hx.1 } end diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index aff33ca9e319e..10052329d7aa0 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -1035,7 +1035,7 @@ begin have h_empty : ∀ x, {a : α | ∃ (i : ι), f i x = a} = ∅, { intro x, ext1 y, - rw [set.mem_set_of_eq, set.mem_empty_eq, iff_false], + rw [set.mem_set_of_eq, set.mem_empty_iff_false, iff_false], exact λ hi, hι (nonempty_of_exists hi), }, simp_rw h_empty at hg, exact ⟨hg.exists.some, hg.mono (λ y hy, is_lub.unique hy hg.exists.some_spec)⟩, @@ -1092,7 +1092,7 @@ begin have h_empty : ∀ x, {a : α | ∃ (i : ι), f i x = a} = ∅, { intro x, ext1 y, - rw [set.mem_set_of_eq, set.mem_empty_eq, iff_false], + rw [set.mem_set_of_eq, set.mem_empty_iff_false, iff_false], exact λ hi, hι (nonempty_of_exists hi), }, simp_rw h_empty at hg, exact ⟨hg.exists.some, hg.mono (λ y hy, is_glb.unique hy hg.exists.some_spec)⟩, @@ -1356,8 +1356,8 @@ begin ext x, have : 0 = f x ∨ 0 < f x := eq_or_lt_of_le bot_le, rw eq_comm at this, - simp only [←and_or_distrib_left, this, mem_singleton_iff, mem_inter_eq, and_true, - mem_union_eq, mem_Ioi, mem_preimage], }, + simp only [←and_or_distrib_left, this, mem_singleton_iff, mem_inter_iff, and_true, + mem_union, mem_Ioi, mem_preimage], }, { apply disjoint_left.2 (λ x hx h'x, _), have : 0 < f x := h'x.2, exact lt_irrefl 0 (this.trans_le hx.2.le) }, @@ -1367,7 +1367,7 @@ begin { rw ← inter_union_distrib_left, congr, ext x, - simp only [mem_singleton_iff, mem_union_eq, mem_Ioo, mem_Ioi, mem_preimage], + simp only [mem_singleton_iff, mem_union, mem_Ioo, mem_Ioi, mem_preimage], have H : f x = ∞ ∨ f x < ∞ := eq_or_lt_of_le le_top, cases H, { simp only [H, eq_self_iff_true, or_false, with_top.zero_lt_top, not_top_lt, and_false] }, diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index db35ceea73f92..c5671d1179530 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -161,7 +161,7 @@ begin refine v.measure_le_of_frequently_le ρ ((measure.absolutely_continuous.refl μ).smul ε) _ _, assume x hx, rw hs at hx, - simp only [mem_inter_eq, not_lt, not_eventually, mem_set_of_eq] at hx, + simp only [mem_inter_iff, not_lt, not_eventually, mem_set_of_eq] at hx, exact hx.1 end ... ≤ ε⁻¹ * ρ o : ennreal.mul_le_mul le_rfl (measure_mono (inter_subset_right _ _)) @@ -228,13 +228,13 @@ begin lift d to ℝ≥0 using I d hd, apply v.null_of_frequently_le_of_frequently_ge hρ (ennreal.coe_lt_coe.1 hcd), { simp only [and_imp, exists_prop, not_frequently, not_and, not_lt, not_le, not_eventually, - mem_set_of_eq, mem_compl_eq, not_forall], + mem_set_of_eq, mem_compl_iff, not_forall], assume x h1x h2x, apply h1x.mono (λ a ha, _), refine (ennreal.div_le_iff_le_mul _ (or.inr (bot_le.trans_lt ha).ne')).1 ha.le, simp only [ennreal.coe_ne_top, ne.def, or_true, not_false_iff] }, { simp only [and_imp, exists_prop, not_frequently, not_and, not_lt, not_le, not_eventually, - mem_set_of_eq, mem_compl_eq, not_forall], + mem_set_of_eq, mem_compl_iff, not_forall], assume x h1x h2x, apply h2x.mono (λ a ha, _), exact ennreal.mul_le_of_le_div ha.le } }, diff --git a/src/measure_theory/covering/vitali.lean b/src/measure_theory/covering/vitali.lean index 7836842778ca8..6a692820e1aee 100644 --- a/src/measure_theory/covering/vitali.lean +++ b/src/measure_theory/covering/vitali.lean @@ -373,7 +373,7 @@ begin have k_closed : is_closed k := is_closed_bUnion w.finite_to_set (λ i hi, h't _ (ut (vu i.2))), have z_notmem_k : z ∉ k, - { simp only [not_exists, exists_prop, mem_Union, mem_sep_eq, forall_exists_index, + { simp only [not_exists, exists_prop, mem_Union, mem_sep_iff, forall_exists_index, set_coe.exists, not_and, exists_and_distrib_right, subtype.coe_mk], assume b hbv h'b h'z, have : z ∈ (s \ ⋃ (a : set α) (H : a ∈ u), a) ∩ (⋃ (a : set α) (H : a ∈ u), a) := diff --git a/src/measure_theory/covering/vitali_family.lean b/src/measure_theory/covering/vitali_family.lean index bd0d9afdbf04a..b8d1c2c9a38be 100644 --- a/src/measure_theory/covering/vitali_family.lean +++ b/src/measure_theory/covering/vitali_family.lean @@ -172,7 +172,7 @@ lemma mem_filter_at_iff {x : α} {s : set (set α)} : begin simp only [filter_at, exists_prop, gt_iff_lt], rw mem_binfi_of_directed, - { simp only [subset_def, and_imp, exists_prop, mem_sep_eq, mem_Ioi, mem_principal] }, + { simp only [subset_def, and_imp, exists_prop, mem_sep_iff, mem_Ioi, mem_principal] }, { simp only [directed_on, exists_prop, ge_iff_le, le_principal_iff, mem_Ioi, order.preimage, mem_principal], assume x hx y hy, @@ -185,7 +185,7 @@ end instance filter_at_ne_bot (x : α) : (v.filter_at x).ne_bot := begin simp only [ne_bot_iff, ←empty_mem_iff_bot, mem_filter_at_iff, not_exists, exists_prop, - mem_empty_eq, and_true, gt_iff_lt, not_and, ne.def, not_false_iff, not_forall], + mem_empty_iff_false, and_true, gt_iff_lt, not_and, ne.def, not_false_iff, not_forall], assume ε εpos, obtain ⟨w, w_sets, hw⟩ : ∃ (w ∈ v.sets_at x), w ⊆ closed_ball x ε := v.nontrivial x ε εpos, exact ⟨w, w_sets, hw⟩ diff --git a/src/measure_theory/decomposition/lebesgue.lean b/src/measure_theory/decomposition/lebesgue.lean index d42dfab5e70c2..a1eafd435fe4a 100644 --- a/src/measure_theory/decomposition/lebesgue.lean +++ b/src/measure_theory/decomposition/lebesgue.lean @@ -738,8 +738,8 @@ instance have_lebesgue_decomposition_of_sigma_finite { by_contra hij, exact h₂ i j hij ⟨hi₁, hj₁⟩ }, exact hj₂ (this ▸ hi₂) }, { intros x hx, - simp only [mem_Union, sup_eq_union, mem_inter_eq, - mem_union_eq, mem_compl_eq, or_iff_not_imp_left], + simp only [mem_Union, sup_eq_union, mem_inter_iff, + mem_union, mem_compl_iff, or_iff_not_imp_left], intro h, push_neg at h, rw [top_eq_univ, ← S.spanning, mem_Union] at hx, obtain ⟨i, hi⟩ := hx, diff --git a/src/measure_theory/decomposition/signed_hahn.lean b/src/measure_theory/decomposition/signed_hahn.lean index 5d8d72f2f2856..75d7867590196 100644 --- a/src/measure_theory/decomposition/signed_hahn.lean +++ b/src/measure_theory/decomposition/signed_hahn.lean @@ -262,11 +262,11 @@ begin rw [set.mem_Union, exists_prop, and_iff_right_iff_imp], exact λ _, h }, { convert le_of_eq s.empty.symm, - ext, simp only [exists_prop, set.mem_empty_eq, set.mem_Union, not_and, iff_false], + ext, simp only [exists_prop, set.mem_empty_iff_false, set.mem_Union, not_and, iff_false], exact λ h', false.elim (h h') } }, { intro, exact measurable_set.Union (λ _, restrict_nonpos_seq_measurable_set _) }, { intros a b hab x hx, - simp only [exists_prop, set.mem_Union, set.mem_inter_eq, set.inf_eq_inter] at hx, + simp only [exists_prop, set.mem_Union, set.mem_inter_iff, set.inf_eq_inter] at hx, exact let ⟨⟨_, hx₁⟩, _, hx₂⟩ := hx in restrict_nonpos_seq_disjoint a b hab ⟨hx₁, hx₂⟩ }, { apply set.Union_subset, intros a x, diff --git a/src/measure_theory/function/ae_measurable_order.lean b/src/measure_theory/function/ae_measurable_order.lean index db4ba98bf7821..2c5766576668d 100644 --- a/src/measure_theory/function/ae_measurable_order.lean +++ b/src/measure_theory/function/ae_measurable_order.lean @@ -77,7 +77,7 @@ begin change μ _ = 0, convert this, ext y, - simp only [not_exists, exists_prop, mem_set_of_eq, mem_compl_eq, not_not_mem] }, + simp only [not_exists, exists_prop, mem_set_of_eq, mem_compl_iff, not_not_mem] }, filter_upwards [this] with x hx, apply (infi_eq_of_forall_ge_of_forall_gt_exists_lt _ _).symm, { assume i, diff --git a/src/measure_theory/function/convergence_in_measure.lean b/src/measure_theory/function/convergence_in_measure.lean index 060a8a5bb4a54..49f9868e46f8f 100644 --- a/src/measure_theory/function/convergence_in_measure.lean +++ b/src/measure_theory/function/convergence_in_measure.lean @@ -115,7 +115,7 @@ begin suffices : {x : α | ε ≤ dist (f n x) (g x)} ⊆ t, from (measure_mono this).trans ht, rw ← set.compl_subset_compl, intros x hx, - rw [set.mem_compl_eq, set.nmem_set_of_eq, dist_comm, not_le], + rw [set.mem_compl_iff, set.nmem_set_of_iff, dist_comm, not_le], exact hN n hn x hx, end @@ -217,7 +217,7 @@ begin { refine λ x hx, metric.tendsto_at_top.mpr (λ ε hε, _), rw [hs, limsup_eq_infi_supr_of_nat] at hx, simp only [set.supr_eq_Union, set.infi_eq_Inter, set.compl_Inter, set.compl_Union, - set.mem_Union, set.mem_Inter, set.mem_compl_eq, set.mem_set_of_eq, not_le] at hx, + set.mem_Union, set.mem_Inter, set.mem_compl_iff, set.mem_set_of_eq, not_le] at hx, obtain ⟨N, hNx⟩ := hx, obtain ⟨k, hk_lt_ε⟩ := h_lt_ε_real ε hε, refine ⟨max N (k - 1), λ n hn_ge, lt_of_le_of_lt _ hk_lt_ε⟩, @@ -345,7 +345,7 @@ begin refine ((le_of_eq _).trans (ae_lt_of_ess_sup_lt this).le).trans hε.le, congr' with x, simp only [ennreal.of_real_le_iff_le_to_real ennreal.coe_lt_top.ne, ennreal.coe_to_real, - not_lt, coe_nnnorm, set.mem_set_of_eq, set.mem_compl_eq], + not_lt, coe_nnnorm, set.mem_set_of_eq, set.mem_compl_iff], rw ← dist_eq_norm (f n x) (g x), refl end diff --git a/src/measure_theory/function/egorov.lean b/src/measure_theory/function/egorov.lean index 9dbdad09ed008..3f26c31f95699 100644 --- a/src/measure_theory/function/egorov.lean +++ b/src/measure_theory/function/egorov.lean @@ -56,7 +56,7 @@ begin simp_rw [metric.tendsto_at_top, ae_iff] at hfg, rw [← nonpos_iff_eq_zero, ← hfg], refine measure_mono (λ x, _), - simp only [mem_inter_eq, mem_Inter, ge_iff_le, mem_not_convergent_seq_iff], + simp only [mem_inter_iff, mem_Inter, ge_iff_le, mem_not_convergent_seq_iff], push_neg, rintro ⟨hmem, hx⟩, refine ⟨hmem, 1 / (n + 1 : ℝ), nat.one_div_pos_of_nat, λ N, _⟩, @@ -179,7 +179,7 @@ begin obtain ⟨N, hN⟩ := exists_nat_one_div_lt hδ, rw eventually_at_top, refine ⟨egorov.not_convergent_seq_lt_index (half_pos hε) hf hg hsm hs hfg N, λ n hn x hx, _⟩, - simp only [mem_diff, egorov.Union_not_convergent_seq, not_exists, mem_Union, mem_inter_eq, + simp only [mem_diff, egorov.Union_not_convergent_seq, not_exists, mem_Union, mem_inter_iff, not_and, exists_and_distrib_left] at hx, obtain ⟨hxs, hx⟩ := hx, specialize hx hxs N, diff --git a/src/measure_theory/function/jacobian.lean b/src/measure_theory/function/jacobian.lean index c36c5707b59b6..c099536c0d4cc 100644 --- a/src/measure_theory/function/jacobian.lean +++ b/src/measure_theory/function/jacobian.lean @@ -239,7 +239,7 @@ begin obtain ⟨q, hq⟩ : ∃ q, F q = (n, z, p) := hF _, -- then `x` belongs to `t q`. apply mem_Union.2 ⟨q, _⟩, - simp only [hq, subset_closure hnz, hp, mem_inter_eq, and_self], + simp only [hq, subset_closure hnz, hp, mem_inter_iff, and_self], end variables [measurable_space E] [borel_space E] (μ : measure E) [is_add_haar_measure μ] diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index 93bf695847081..d37cbc611c38d 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -278,7 +278,7 @@ begin { suffices h_empty : f ⁻¹' {y} = ∅, by { rw [h_empty, measure_empty], exact ennreal.coe_lt_top, }, ext1 x, - rw [set.mem_preimage, set.mem_singleton_iff, mem_empty_eq, iff_false], + rw [set.mem_preimage, set.mem_singleton_iff, mem_empty_iff_false, iff_false], refine λ hxy, hyf _, rw [mem_range, set.mem_range], exact ⟨x, hxy⟩, }, @@ -745,7 +745,7 @@ begin rw mem_compl_iff at hxs, have hx' : x ∉ {a : α | ¬0 ≤ simple_func.to_simple_func f a}, from λ h, hxs (subset_to_measurable μ _ h), - rwa [set.nmem_set_of_eq, not_not] at hx', }, + rwa [set.nmem_set_of_iff, not_not] at hx', }, let f' := simple_func.piecewise s (measurable_set_to_measurable μ _).compl (simple_func.to_simple_func f) (simple_func.const α (0 : G)), refine ⟨f', λ x, _, _⟩, @@ -778,7 +778,7 @@ begin rw mem_closure_iff_seq_limit, have hg_mem_ℒp : mem_ℒp g p μ := Lp.mem_ℒp g, have zero_mem : (0 : G) ∈ (range g ∪ {0} : set G) ∩ {y | 0 ≤ y}, by simp only [union_singleton, - mem_inter_eq, mem_insert_iff, eq_self_iff_true, true_or, mem_set_of_eq, le_refl, and_self], + mem_inter_iff, mem_insert_iff, eq_self_iff_true, true_or, mem_set_of_eq, le_refl, and_self], haveI : separable_space (((range g ∪ {0}) ∩ {y | 0 ≤ y}) : set G), { apply is_separable.separable_space, apply is_separable.mono _ (set.inter_subset_left _ _), diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index e24a038c63a6a..6011f1f1e1461 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -730,7 +730,7 @@ begin ((ht.subtype_image ((hd.approx n).measurable_set_fiber x)).diff hs), ext1 y, - simp only [mem_union_eq, mem_preimage, mem_singleton_iff, mem_image, set_coe.exists, + simp only [mem_union, mem_preimage, mem_singleton_iff, mem_image, set_coe.exists, subtype.coe_mk, exists_and_distrib_right, exists_eq_right, mem_diff], by_cases hy : y ∈ s, { rw dif_pos hy, @@ -904,8 +904,8 @@ begin exact ⟨λ h, h.2, λ h, ⟨hg_seq_zero y h n, h⟩⟩, }, { suffices : (g_seq_s n) ⁻¹' {x} ∩ sᶜ = ∅, by { rw this, exact measurable_set.empty, }, ext1 y, - simp only [mem_inter_eq, mem_preimage, mem_singleton_iff, mem_compl_eq, mem_empty_eq, - iff_false, not_and, not_not_mem], + simp only [mem_inter_iff, mem_preimage, mem_singleton_iff, mem_compl_iff, + mem_empty_iff_false, iff_false, not_and, not_not_mem], refine imp_of_not_imp_not _ _ (λ hys, _), rw hg_seq_zero y hys n, exact ne.symm hx, }, @@ -1489,7 +1489,7 @@ begin refine ⟨λ H, ⟨H.ae_measurable, H.is_separable_ae_range⟩, _⟩, rintros ⟨H, ⟨t, t_sep, ht⟩⟩, rcases eq_empty_or_nonempty t with rfl|h₀, - { simp only [mem_empty_eq, eventually_false_iff_eq_bot, ae_eq_bot] at ht, + { simp only [mem_empty_iff_false, eventually_false_iff_eq_bot, ae_eq_bot] at ht, rw ht, exact ae_strongly_measurable_zero_measure f }, { obtain ⟨g, g_meas, gt, fg⟩ : ∃ (g : α → β), measurable g ∧ range g ⊆ t ∧ f =ᵐ[μ] g := @@ -1706,7 +1706,7 @@ begin have : (f a : ℝ≥0∞) ≠ 0, by simpa only [ne.def, ennreal.coe_eq_zero] using h'a, rw ha this }, { filter_upwards [ae_restrict_mem A.compl] with x hx, - simp only [not_not, mem_set_of_eq, mem_compl_eq] at hx, + simp only [not_not, mem_set_of_eq, mem_compl_iff] at hx, simp [hx] } }, { rintros ⟨g', g'meas, hg'⟩, refine ⟨λ x, (f x : ℝ)⁻¹ • g' x, hf.coe_nnreal_real.inv.strongly_measurable.smul g'meas, _⟩, diff --git a/src/measure_theory/group/arithmetic.lean b/src/measure_theory/group/arithmetic.lean index 4445b7fd061a3..a2ede97a8e05a 100644 --- a/src/measure_theory/group/arithmetic.lean +++ b/src/measure_theory/group/arithmetic.lean @@ -342,7 +342,7 @@ lemma measurable_set_eq_fun_of_countable {m : measurable_space α} {E} [measurab measurable_set {x | f x = g x} := begin have : {x | f x = g x} = ⋃ j, {x | f x = j} ∩ {x | g x = j}, - { ext1 x, simp only [set.mem_set_of_eq, set.mem_Union, set.mem_inter_eq, exists_eq_right'], }, + { ext1 x, simp only [set.mem_set_of_eq, set.mem_Union, set.mem_inter_iff, exists_eq_right'], }, rw this, refine measurable_set.Union (λ j, measurable_set.inter _ _), { exact hf (measurable_set_singleton j), }, diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 3348bf5a9124c..ff220f69465d5 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -2446,14 +2446,14 @@ begin simp only [pi.zero_apply, mem_set_of_eq, filter.mem_mk] at A, convert A, ext x, - simp only [and_comm, exists_prop, mem_inter_eq, iff_self, mem_set_of_eq, mem_compl_eq, + simp only [and_comm, exists_prop, mem_inter_iff, iff_self, mem_set_of_eq, mem_compl_iff, not_forall] }, { assume hs, let t := to_measurable μ ({x | f x ≠ 0} ∩ s), have A : s ⊆ t ∪ {x | f x = 0}, { assume x hx, rcases eq_or_ne (f x) 0 with fx|fx, - { simp only [fx, mem_union_eq, mem_set_of_eq, eq_self_iff_true, or_true] }, + { simp only [fx, mem_union, mem_set_of_eq, eq_self_iff_true, or_true] }, { left, apply subset_to_measurable _ _, exact ⟨fx, hx⟩ } }, @@ -2472,7 +2472,7 @@ begin rw [ae_iff, ae_iff, with_density_apply_eq_zero hf], congr', ext x, - simp only [exists_prop, mem_inter_eq, iff_self, mem_set_of_eq, not_forall], + simp only [exists_prop, mem_inter_iff, iff_self, mem_set_of_eq, not_forall], end lemma ae_with_density_iff_ae_restrict {p : α → Prop} {f : α → ℝ≥0∞} (hf : measurable f) : @@ -2501,7 +2501,7 @@ begin rw ha this }, { filter_upwards [ae_restrict_mem A.compl], assume x hx, - simp only [not_not, mem_set_of_eq, mem_compl_eq] at hx, + simp only [not_not, mem_set_of_eq, mem_compl_iff] at hx, simp [hx] } }, { rintros ⟨g', g'meas, hg'⟩, refine ⟨λ x, (f x : ℝ)⁻¹ • g' x, hf.coe_nnreal_real.inv.smul g'meas, _⟩, @@ -2530,7 +2530,7 @@ begin rw ha this }, { filter_upwards [ae_restrict_mem A.compl], assume x hx, - simp only [not_not, mem_set_of_eq, mem_compl_eq] at hx, + simp only [not_not, mem_set_of_eq, mem_compl_iff] at hx, simp [hx] } }, { rintros ⟨g', g'meas, hg'⟩, refine ⟨λ x, (f x)⁻¹ * g' x, hf.coe_nnreal_ennreal.inv.smul g'meas, _⟩, @@ -2631,7 +2631,7 @@ begin (hf.measurable_mk (measurable_set_singleton 0).compl).compl, filter_upwards [ae_restrict_mem M], assume x hx, - simp only [not_not, mem_set_of_eq, mem_compl_eq] at hx, + simp only [not_not, mem_set_of_eq, mem_compl_iff] at hx, simp only [hx, zero_mul, pi.mul_apply] } end ... = ∫⁻ (a : α), (f * g) a ∂μ : diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index a6ed1044d17f8..31ca3cc04c9a0 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -247,7 +247,7 @@ lemma set_integral_neg_eq_set_integral_nonpos [linear_order E] [order_closed_top ∫ x in {x | f x < 0}, f x ∂μ = ∫ x in {x | f x ≤ 0}, f x ∂μ := begin have h_union : {x | f x ≤ 0} = {x | f x < 0} ∪ {x | f x = 0}, - by { ext, simp_rw [set.mem_union_eq, set.mem_set_of_eq], exact le_iff_lt_or_eq, }, + by { ext, simp_rw [set.mem_union, set.mem_set_of_eq], exact le_iff_lt_or_eq, }, rw h_union, exact (set_integral_union_eq_left hf hfi (hf.measurable_set_lt strongly_measurable_const) (λ x hx, hx)).symm, @@ -274,7 +274,7 @@ begin refine set_integral_congr h_meas.compl (λ x hx, _), dsimp only, rw [real.norm_eq_abs, abs_eq_neg_self.mpr _], - rw [set.mem_compl_iff, set.nmem_set_of_eq] at hx, + rw [set.mem_compl_iff, set.nmem_set_of_iff] at hx, linarith, end ... = ∫ x in {x | 0 ≤ f x}, f x ∂μ - ∫ x in {x | f x ≤ 0}, f x ∂μ : diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index f3c4d2e7383da..0c4493e754935 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -625,7 +625,7 @@ begin { have B : x ∈ t (nat.find (P x)) ∪ (⋃ k, t k)ᶜ := nat.find_spec (P x), have B' : (∀ (i : ℕ), x ∉ t i) ↔ false, { simp only [iff_false, not_forall, not_not_mem], exact ⟨n, hx⟩ }, - simpa only [B', mem_union_eq, mem_Inter, or_false, compl_Union, mem_compl_eq] using B }, + simpa only [B', mem_union, mem_Inter, or_false, compl_Union, mem_compl_iff] using B }, congr, by_contra h, exact t_disj n (nat.find (P x)) (ne.symm h) ⟨hx, this⟩ diff --git a/src/measure_theory/measure/ae_disjoint.lean b/src/measure_theory/measure/ae_disjoint.lean index 992a83907378f..766c271f36b0a 100644 --- a/src/measure_theory/measure/ae_disjoint.lean +++ b/src/measure_theory/measure/ae_disjoint.lean @@ -38,7 +38,7 @@ begin { simp only [pairwise, disjoint_left, on_fun, mem_diff, not_and, and_imp, not_not], intros i j hne x hi hU hj, replace hU : x ∉ s i ∩ ⋃ j ≠ i, s j := λ h, hU (subset_to_measurable _ _ h), - simp only [mem_inter_eq, mem_Union, not_and, not_exists] at hU, + simp only [mem_inter_iff, mem_Union, not_and, not_exists] at hU, exact (hU hi j hne.symm hj).elim } end diff --git a/src/measure_theory/measure/ae_measurable.lean b/src/measure_theory/measure/ae_measurable.lean index 99b01758f417e..582015d62901a 100644 --- a/src/measure_theory/measure/ae_measurable.lean +++ b/src/measure_theory/measure/ae_measurable.lean @@ -77,7 +77,7 @@ begin refine ⟨⋃ i, ((h i).mk f ⁻¹' t) ∩ (s i)ᶜ, measurable_set.Union $ λ i, (measurable_mk _ ht).inter (measurable_set_to_measurable _ _).compl, _⟩, ext ⟨x, hx⟩, - simp only [mem_preimage, mem_Union, subtype.coe_mk, set.restrict, mem_inter_eq, + simp only [mem_preimage, mem_Union, subtype.coe_mk, set.restrict, mem_inter_iff, mem_compl_iff] at hx ⊢, split, { rintro ⟨i, hxt, hxs⟩, rwa hs _ _ hxs }, @@ -171,7 +171,7 @@ begin { simp only [g, hx, piecewise_eq_of_not_mem, not_false_iff], contrapose! hx, apply subset_to_measurable, - simp only [hx, mem_compl_eq, mem_set_of_eq, not_and, not_false_iff, implies_true_iff] + simp only [hx, mem_compl_iff, mem_set_of_eq, not_and, not_false_iff, implies_true_iff] {contextual := tt} } }, { have A : μ (to_measurable μ {x | f x = H.mk f x ∧ f x ∈ t}ᶜ) = 0, { rw [measure_to_measurable, ← compl_mem_ae_iff, compl_compl], @@ -181,7 +181,7 @@ begin simp only [g, hx, piecewise_eq_of_not_mem, not_false_iff], contrapose! hx, apply subset_to_measurable, - simp only [hx, mem_compl_eq, mem_set_of_eq, false_and, not_false_iff] } + simp only [hx, mem_compl_iff, mem_set_of_eq, false_and, not_false_iff] } end lemma exists_measurable_nonneg {β} [preorder β] [has_zero β] {mβ : measurable_space β} {f : α → β} diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar.lean index 0c07712eb7659..8a3501d6c7147 100644 --- a/src/measure_theory/measure/haar.lean +++ b/src/measure_theory/measure/haar.lean @@ -208,7 +208,7 @@ begin simp only [mem_preimage] at h2g₀, simp only [mem_Union], use g₀, split, { simp only [finset.mem_filter, h1g₀, true_and], use g, - simp only [hg, h2g₀, mem_inter_eq, mem_preimage, and_self] }, + simp only [hg, h2g₀, mem_inter_iff, mem_preimage, and_self] }, exact h2g₀ }, refine le_trans (add_le_add (this K₁.1 $ subset.trans (subset_union_left _ _) h1s) (this K₂.1 $ subset.trans (subset_union_right _ _) h1s)) _, @@ -432,7 +432,7 @@ begin let V := V₁ ∩ V₂, apply mem_of_subset_of_mem _ (chaar_mem_cl_prehaar K₀ ⟨V⁻¹, (is_open.inter h2V₁ h2V₂).preimage continuous_inv, - by simp only [mem_inv, inv_one, h3V₁, h3V₂, V, mem_inter_eq, true_and]⟩), + by simp only [mem_inv, inv_one, h3V₁, h3V₂, V, mem_inter_iff, true_and]⟩), unfold cl_prehaar, rw is_closed.closure_subset_iff, { rintro _ ⟨U, ⟨h1U, h2U, h3U⟩, rfl⟩, simp only [mem_preimage, eval, sub_eq_zero, mem_singleton_iff], rw [eq_comm], diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 74128d22e0b4f..167c142ab0a88 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -106,7 +106,7 @@ instance ae_is_measurably_generated : is_measurably_generated μ.ae := /-- See also `measure_theory.ae_restrict_interval_oc_iff`. -/ lemma ae_interval_oc_iff [linear_order α] {a b : α} {P : α → Prop} : (∀ᵐ x ∂μ, x ∈ Ι a b → P x) ↔ (∀ᵐ x ∂μ, x ∈ Ioc a b → P x) ∧ (∀ᵐ x ∂μ, x ∈ Ioc b a → P x) := -by simp only [interval_oc_eq_union, mem_union_eq, or_imp_distrib, eventually_and] +by simp only [interval_oc_eq_union, mem_union, or_imp_distrib, eventually_and] lemma measure_union (hd : disjoint s₁ s₂) (h : measurable_set s₂) : μ (s₁ ∪ s₂) = μ s₁ + μ s₂ := diff --git a/src/measure_theory/pi_system.lean b/src/measure_theory/pi_system.lean index 0af61def42504..78675f24a168c 100644 --- a/src/measure_theory/pi_system.lean +++ b/src/measure_theory/pi_system.lean @@ -335,7 +335,7 @@ begin { rw [ht1_eq, ht2_eq], simp_rw [← set.inf_eq_inter, g], ext1 x, - simp only [inf_eq_inter, mem_inter_eq, mem_Inter, finset.mem_union], + simp only [inf_eq_inter, mem_inter_iff, mem_Inter, finset.mem_union], refine ⟨λ h i hi_mem_union, _, λ h, ⟨λ i hi1, _, λ i hi2, _⟩⟩, { split_ifs, exacts [⟨h.1 i h_1, h.2 i h_2⟩, ⟨h.1 i h_1, set.mem_univ _⟩, diff --git a/src/model_theory/definability.lean b/src/model_theory/definability.lean index a4ddd753f5059..75839df113271 100644 --- a/src/model_theory/definability.lean +++ b/src/model_theory/definability.lean @@ -100,7 +100,7 @@ begin rcases hg with ⟨θ, hθ⟩, refine ⟨φ ⊔ θ, _⟩, ext, - rw [hφ, hθ, mem_set_of_eq, formula.realize_sup, mem_union_eq, mem_set_of_eq, + rw [hφ, hθ, mem_set_of_eq, formula.realize_sup, mem_union, mem_set_of_eq, mem_set_of_eq], end @@ -233,7 +233,7 @@ begin refine (congr rfl (ext _)).mp (definable_finset_bInter h' finset.univ), simp }, refine (congr rfl (ext (λ x, _))).mp (h.inter h'), - simp only [equiv.coe_trans, mem_inter_eq, mem_preimage, mem_image, + simp only [equiv.coe_trans, mem_inter_iff, mem_preimage, mem_image, exists_exists_and_eq_and, mem_set_of_eq], split, { rintro ⟨⟨y, ys, hy⟩, hx⟩, diff --git a/src/model_theory/semantics.lean b/src/model_theory/semantics.lean index cb4f158b9eeb5..7d6ecbc00524c 100644 --- a/src/model_theory/semantics.lean +++ b/src/model_theory/semantics.lean @@ -719,7 +719,7 @@ lemma model.mono {T' : L.Theory} (h : M ⊨ T') (hs : T ⊆ T') : lemma model.union {T' : L.Theory} (h : M ⊨ T) (h' : M ⊨ T') : M ⊨ T ∪ T' := begin - simp only [model_iff, set.mem_union_eq] at *, + simp only [model_iff, set.mem_union] at *, exact λ φ hφ, hφ.elim (h _) (h' _), end @@ -932,7 +932,7 @@ lemma model_distinct_constants_theory {M : Type w} [L[[α]].Structure M] (s : se M ⊨ L.distinct_constants_theory s ↔ set.inj_on (λ (i : α), (L.con i : M)) s := begin simp only [distinct_constants_theory, Theory.model_iff, set.mem_image, - set.mem_inter_eq, set.mem_prod, set.mem_compl_eq, prod.exists, forall_exists_index, and_imp], + set.mem_inter_iff, set.mem_prod, set.mem_compl_iff, prod.exists, forall_exists_index, and_imp], refine ⟨λ h a as b bs ab, _, _⟩, { contrapose! ab, have h' := h _ a b as bs ab rfl, diff --git a/src/order/compactly_generated.lean b/src/order/compactly_generated.lean index 10087eaa9b514..47a3120645fbd 100644 --- a/src/order/compactly_generated.lean +++ b/src/order/compactly_generated.lean @@ -425,7 +425,7 @@ theorem Iic_coatomic_of_compact_element {k : α} (h : is_compact_element k) : by_cases hS : S.nonempty, { exact ⟨Sup S, h.directed_Sup_lt_of_lt hS cC.directed_on SC, λ _, le_Sup⟩, }, exact ⟨b, lt_of_le_of_ne hbk htriv, by simp only [set.not_nonempty_iff_eq_empty.mp hS, - set.mem_empty_eq, forall_const, forall_prop_of_false, not_false_iff]⟩, }, + set.mem_empty_iff_false, forall_const, forall_prop_of_false, not_false_iff]⟩, }, end⟩ lemma coatomic_of_top_compact (h : is_compact_element (⊤ : α)) : is_coatomic α := diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index 6932ee9c0856d..e76e002fb8699 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -1654,7 +1654,7 @@ begin have hms_freq : ∀ (n : ℕ), x (y (ms n)) ∉ s, from λ n, hy_freq (ms n), have h_empty : (λ (n : ℕ), x (y (ms n))) ⁻¹' s = ∅, { ext1 n, - simp only [set.mem_preimage, set.mem_empty_eq, iff_false], + simp only [set.mem_preimage, set.mem_empty_iff_false, iff_false], exact hms_freq n, }, rw h_empty at hms_tendsto, exact empty_not_mem at_top hms_tendsto, diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 60e4c891ce7bd..f695218ebbd63 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1324,7 +1324,7 @@ eventually_eq_set.trans $ by simp lemma inter_eventually_eq_left {s t : set α} {l : filter α} : (s ∩ t : set α) =ᶠ[l] s ↔ ∀ᶠ x in l, x ∈ s → x ∈ t := -by simp only [eventually_eq_set, mem_inter_eq, and_iff_left_iff_imp] +by simp only [eventually_eq_set, mem_inter_iff, and_iff_left_iff_imp] lemma inter_eventually_eq_right {s t : set α} {l : filter α} : (s ∩ t : set α) =ᶠ[l] t ↔ ∀ᶠ x in l, x ∈ t → x ∈ s := diff --git a/src/order/filter/curry.lean b/src/order/filter/curry.lean index e5aeb997aca00..d0ad3eee57f57 100644 --- a/src/order/filter/curry.lean +++ b/src/order/filter/curry.lean @@ -58,7 +58,7 @@ def curry (f : filter α) (g : filter β) : filter (α × β) := end, inter_sets := begin intros x y hx hy, - simp only [set.mem_set_of_eq, set.mem_inter_eq] at hx hy ⊢, + simp only [set.mem_set_of_eq, set.mem_inter_iff] at hx hy ⊢, exact (hx.and hy).mono (λ a ha, (ha.1.and ha.2).mono (λ b hb, hb)), end, } diff --git a/src/probability/independence.lean b/src/probability/independence.lean index 889236043b328..4cc0e046c3894 100644 --- a/src/probability/independence.lean +++ b/src/probability/independence.lean @@ -631,7 +631,7 @@ begin have h_Inter_inter : (⋂ i ∈ S, (f i) ⁻¹' (sets_s' i)) ∩ (⋂ i ∈ T, (f i) ⁻¹' (sets_t' i)) = ⋂ i ∈ (S ∪ T), (f i) ⁻¹' (sets_s' i ∩ sets_t' i), { ext1 x, - simp only [set.mem_inter_eq, set.mem_Inter, set.mem_preimage, finset.mem_union], + simp only [set.mem_inter_iff, set.mem_Inter, set.mem_preimage, finset.mem_union], split; intro h, { intros i hi, cases hi, diff --git a/src/probability/martingale/borel_cantelli.lean b/src/probability/martingale/borel_cantelli.lean index cd5664dcbcf35..1af4452345f8f 100644 --- a/src/probability/martingale/borel_cantelli.lean +++ b/src/probability/martingale/borel_cantelli.lean @@ -122,7 +122,7 @@ begin { obtain ⟨k, hk⟩ := nat.exists_eq_succ_of_ne_zero heq, rw [hk, add_comm, ← sub_le_iff_le_add], have := not_mem_of_lt_hitting (hk.symm ▸ k.lt_succ_self : k < least_ge f r i ω) (zero_le _), - simp only [set.mem_union_eq, set.mem_Iic, set.mem_Ici, not_or_distrib, not_le] at this, + simp only [set.mem_union, set.mem_Iic, set.mem_Ici, not_or_distrib, not_le] at this, exact (sub_lt_sub_left this _).le.trans ((le_abs_self _).trans (hbddω _)) } end diff --git a/src/probability/process/hitting_time.lean b/src/probability/process/hitting_time.lean index e2e66876a9324..2e46740b5d738 100644 --- a/src/probability/process/hitting_time.lean +++ b/src/probability/process/hitting_time.lean @@ -59,7 +59,7 @@ begin { push_neg, intro j, rw set.Icc_eq_empty_of_lt h, - simp only [set.mem_empty_eq, is_empty.forall_iff], }, + simp only [set.mem_empty_iff_false, is_empty.forall_iff], }, simp only [h_not, if_false], end @@ -283,8 +283,8 @@ begin simp [← exists_or_distrib, ← or_and_distrib_right, le_or_lt] }, have h₂ : (⋃ i > n, {x | τ x = i} ∩ {x | hitting u s i N x ≤ n}) = ∅, { ext x, - simp only [gt_iff_lt, set.mem_Union, set.mem_inter_eq, set.mem_set_of_eq, - exists_prop, set.mem_empty_eq, iff_false, not_exists, not_and, not_le], + simp only [gt_iff_lt, set.mem_Union, set.mem_inter_iff, set.mem_set_of_eq, + exists_prop, set.mem_empty_iff_false, iff_false, not_exists, not_and, not_le], rintro m hm rfl, exact lt_of_lt_of_le hm (le_hitting (hτbdd _) _) }, rw [h₁, h₂, set.union_empty], diff --git a/src/probability/process/stopping.lean b/src/probability/process/stopping.lean index 1c49d729de55f..a0041b99663ce 100644 --- a/src/probability/process/stopping.lean +++ b/src/probability/process/stopping.lean @@ -71,7 +71,7 @@ begin by_cases hi_min : is_min i, { suffices : {ω : Ω | τ ω < i} = ∅, by { rw this, exact @measurable_set.empty _ (f i), }, ext1 ω, - simp only [set.mem_set_of_eq, set.mem_empty_eq, iff_false], + simp only [set.mem_set_of_eq, set.mem_empty_iff_false, iff_false], rw is_min_iff_forall_not_lt at hi_min, exact hi_min (τ ω), }, have : {ω : Ω | τ ω < i} = τ ⁻¹' (set.Iio i) := rfl, @@ -136,7 +136,7 @@ protected lemma measurable_set_ge_of_countable_range {ι} [linear_order ι] {τ measurable_set[f i] {ω | i ≤ τ ω} := begin have : {ω | i ≤ τ ω} = {ω | τ ω < i}ᶜ, - { ext1 ω, simp only [set.mem_set_of_eq, set.mem_compl_eq, not_lt], }, + { ext1 ω, simp only [set.mem_set_of_eq, set.mem_compl_iff, not_lt], }, rw this, exact (hτ.measurable_set_lt_of_countable_range h_countable i).compl, end @@ -157,7 +157,7 @@ lemma is_stopping_time.measurable_set_gt (hτ : is_stopping_time f τ) (i : ι) measurable_set[f i] {ω | i < τ ω} := begin have : {ω | i < τ ω} = {ω | τ ω ≤ i}ᶜ, - { ext1 ω, simp only [set.mem_set_of_eq, set.mem_compl_eq, not_le], }, + { ext1 ω, simp only [set.mem_set_of_eq, set.mem_compl_iff, not_le], }, rw this, exact (hτ.measurable_set_le i).compl, end @@ -174,7 +174,7 @@ begin by_cases hi_min : is_min i, { suffices : {ω | τ ω < i} = ∅, by { rw this, exact @measurable_set.empty _ (f i), }, ext1 ω, - simp only [set.mem_set_of_eq, set.mem_empty_eq, iff_false], + simp only [set.mem_set_of_eq, set.mem_empty_iff_false, iff_false], exact is_min_iff_forall_not_lt.mp hi_min (τ ω), }, obtain ⟨seq, -, -, h_tendsto, h_bound⟩ : ∃ seq : ℕ → ι, monotone seq ∧ (∀ j, seq j ≤ i) ∧ tendsto seq at_top (𝓝 i) ∧ (∀ j, seq j < i), @@ -214,7 +214,7 @@ lemma is_stopping_time.measurable_set_ge (hτ : is_stopping_time f τ) (i : ι) measurable_set[f i] {ω | i ≤ τ ω} := begin have : {ω | i ≤ τ ω} = {ω | τ ω < i}ᶜ, - { ext1 ω, simp only [set.mem_set_of_eq, set.mem_compl_eq, not_lt], }, + { ext1 ω, simp only [set.mem_set_of_eq, set.mem_compl_iff, not_lt], }, rw this, exact (hτ.measurable_set_lt i).compl, end @@ -223,7 +223,7 @@ lemma is_stopping_time.measurable_set_eq (hτ : is_stopping_time f τ) (i : ι) measurable_set[f i] {ω | τ ω = i} := begin have : {ω | τ ω = i} = {ω | τ ω ≤ i} ∩ {ω | τ ω ≥ i}, - { ext1 ω, simp only [set.mem_set_of_eq, ge_iff_le, set.mem_inter_eq, le_antisymm_iff], }, + { ext1 ω, simp only [set.mem_set_of_eq, ge_iff_le, set.mem_inter_iff, le_antisymm_iff], }, rw this, exact (hτ.measurable_set_le i).inter (hτ.measurable_set_ge i), end @@ -307,7 +307,7 @@ begin { rw not_le at hij, convert measurable_set.empty, ext ω, - simp only [set.mem_empty_eq, iff_false], + simp only [set.mem_empty_iff_false, iff_false], rintro (hx : τ ω + i = j), linarith }, end @@ -322,7 +322,7 @@ begin { exact measurable_set.Union (λ k, measurable_set.Union (λ hk, (hπ.measurable_set_eq_le hk).inter (hτ.add_const_nat i))) }, ext ω, - simp only [pi.add_apply, set.mem_set_of_eq, set.mem_Union, set.mem_inter_eq, exists_prop], + simp only [pi.add_apply, set.mem_set_of_eq, set.mem_Union, set.mem_inter_iff, exists_prop], refine ⟨λ h, ⟨π ω, by linarith, rfl, h⟩, _⟩, rintro ⟨j, hj, rfl, h⟩, assumption @@ -367,7 +367,7 @@ begin rw (_ : s ∩ {ω | π ω ≤ i} = s ∩ {ω | τ ω ≤ i} ∩ {ω | π ω ≤ i}), { exact (hs i).inter (hπ i) }, { ext, - simp only [set.mem_inter_eq, iff_self_and, and.congr_left_iff, set.mem_set_of_eq], + simp only [set.mem_inter_iff, iff_self_and, and.congr_left_iff, set.mem_set_of_eq], intros hle' _, exact le_trans (hle _) hle' }, end @@ -446,7 +446,7 @@ begin have : ∀ j, ({ω : Ω | τ ω = i} ∩ {ω : Ω | τ ω ≤ j}) = {ω : Ω | τ ω = i} ∩ {ω | i ≤ j}, { intro j, ext1 ω, - simp only [set.mem_inter_eq, set.mem_set_of_eq, and.congr_right_iff], + simp only [set.mem_inter_iff, set.mem_set_of_eq, and.congr_right_iff], intro hxi, rw hxi, }, split; intro h, @@ -506,7 +506,7 @@ protected lemma measurable_set_le' (hτ : is_stopping_time f τ) (i : ι) : begin intro j, have : {ω : Ω | τ ω ≤ i} ∩ {ω : Ω | τ ω ≤ j} = {ω : Ω | τ ω ≤ min i j}, - { ext1 ω, simp only [set.mem_inter_eq, set.mem_set_of_eq, le_min_iff], }, + { ext1 ω, simp only [set.mem_inter_iff, set.mem_set_of_eq, le_min_iff], }, rw this, exact f.mono (min_le_right i j) _ (hτ _), end @@ -535,7 +535,7 @@ protected lemma measurable_set_ge' [topological_space ι] [order_topology ι] begin have : {ω | i ≤ τ ω} = {ω | τ ω = i} ∪ {ω | i < τ ω}, { ext1 ω, - simp only [le_iff_lt_or_eq, set.mem_set_of_eq, set.mem_union_eq], + simp only [le_iff_lt_or_eq, set.mem_set_of_eq, set.mem_union], rw [@eq_comm _ i, or_comm], }, rw this, exact (hτ.measurable_set_eq' i).union (hτ.measurable_set_gt' i), @@ -573,7 +573,7 @@ protected lemma measurable_set_ge_of_countable_range' begin have : {ω | i ≤ τ ω} = {ω | τ ω = i} ∪ {ω | i < τ ω}, { ext1 ω, - simp only [le_iff_lt_or_eq, set.mem_set_of_eq, set.mem_union_eq], + simp only [le_iff_lt_or_eq, set.mem_set_of_eq, set.mem_union], rw [@eq_comm _ i, or_comm], }, rw this, exact (hτ.measurable_set_eq_of_countable_range' h_countable i).union (hτ.measurable_set_gt' i), @@ -610,7 +610,7 @@ begin split; rw set.mem_Union, { exact λ hx, ⟨τ ω, by simpa using hx⟩,}, { rintro ⟨i, hx⟩, - simp only [set.mem_range, set.Union_exists, set.mem_Union, set.mem_inter_eq, + simp only [set.mem_range, set.Union_exists, set.mem_Union, set.mem_inter_iff, set.mem_set_of_eq, exists_prop, exists_and_distrib_right] at hx, exact hx.1.2, } } end @@ -671,7 +671,7 @@ begin have : (s ∩ {ω | τ ω ≤ π ω} ∩ {ω | min (τ ω) (π ω) ≤ i}) = (s ∩ {ω | τ ω ≤ i}) ∩ {ω | min (τ ω) (π ω) ≤ i} ∩ {ω | min (τ ω) i ≤ min (min (τ ω) (π ω)) i}, { ext1 ω, - simp only [min_le_iff, set.mem_inter_eq, set.mem_set_of_eq, le_min_iff, le_refl, true_and, + simp only [min_le_iff, set.mem_inter_iff, set.mem_set_of_eq, le_min_iff, le_refl, true_and, and_true, true_or, or_true], by_cases hτi : τ ω ≤ i, { simp only [hτi, true_or, and_true, and.congr_right_iff], @@ -728,7 +728,7 @@ begin intro j, have : {ω | τ ω ≤ π ω} ∩ {ω | τ ω ≤ j} = {ω | min (τ ω) j ≤ min (π ω) j} ∩ {ω | τ ω ≤ j}, { ext1 ω, - simp only [set.mem_inter_eq, set.mem_set_of_eq, min_le_iff, le_min_iff, le_refl, and_true, + simp only [set.mem_inter_iff, set.mem_set_of_eq, min_le_iff, le_min_iff, le_refl, and_true, and.congr_left_iff], intro h, simp only [h, or_self, and_true], @@ -764,7 +764,7 @@ begin have : {ω | τ ω = π ω} ∩ {ω | τ ω ≤ j} = {ω | min (τ ω) j = min (π ω) j} ∩ {ω | τ ω ≤ j} ∩ {ω | π ω ≤ j}, { ext1 ω, - simp only [set.mem_inter_eq, set.mem_set_of_eq], + simp only [set.mem_inter_iff, set.mem_set_of_eq], refine ⟨λ h, ⟨⟨_, h.2⟩, _⟩, λ h, ⟨_, h.1.2⟩⟩, { rw h.1, }, { rw ← h.1, exact h.2, }, @@ -790,7 +790,7 @@ begin have : {ω | τ ω = π ω} ∩ {ω | τ ω ≤ j} = {ω | min (τ ω) j = min (π ω) j} ∩ {ω | τ ω ≤ j} ∩ {ω | π ω ≤ j}, { ext1 ω, - simp only [set.mem_inter_eq, set.mem_set_of_eq], + simp only [set.mem_inter_iff, set.mem_set_of_eq], refine ⟨λ h, ⟨⟨_, h.2⟩, _⟩, λ h, ⟨_, h.1.2⟩⟩, { rw h.1, }, { rw ← h.1, exact h.2, }, @@ -885,7 +885,7 @@ begin have hx_fst_le : ↑(ω : set.Iic i × Ω).fst ≤ i, from (ω : set.Iic i × Ω).fst.prop, refine hx_fst_le.trans (le_of_lt _), convert ω.prop, - simp only [not_le, set.mem_compl_eq, set.mem_set_of_eq], }, + simp only [not_le, set.mem_compl_iff, set.mem_set_of_eq], }, end lemma prog_measurable.stopped_process [metrizable_space ι] @@ -926,7 +926,7 @@ begin = stopped_value u (λ ω, min (τ ω) i) ⁻¹' t ∩ {ω : Ω | τ ω ≤ i}, by { rw this, exact ((h_str_meas i).measurable ht).inter (hτ.measurable_set_le i), }, ext1 ω, - simp only [stopped_value, set.mem_inter_eq, set.mem_preimage, set.mem_set_of_eq, + simp only [stopped_value, set.mem_inter_iff, set.mem_preimage, set.mem_set_of_eq, and.congr_left_iff], intro h, rw min_eq_left h, @@ -1129,7 +1129,7 @@ begin intro n, have : {ω | s.piecewise τ η ω ≤ n} = (s ∩ {ω | τ ω ≤ n}) ∪ (sᶜ ∩ {ω | η ω ≤ n}), { ext1 ω, - simp only [set.piecewise, set.mem_inter_eq, set.mem_set_of_eq, and.congr_right_iff], + simp only [set.piecewise, set.mem_inter_iff, set.mem_set_of_eq, and.congr_right_iff], by_cases hx : ω ∈ s; simp [hx], }, rw this, by_cases hin : i ≤ n, diff --git a/src/ring_theory/hahn_series.lean b/src/ring_theory/hahn_series.lean index a622d351842cf..b12508fac023a 100644 --- a/src/ring_theory/hahn_series.lean +++ b/src/ring_theory/hahn_series.lean @@ -548,7 +548,7 @@ instance [non_unital_non_assoc_semiring R] : distrib (hahn_series Γ R) := mul_coeff_right' hwf (set.subset_union_left _ _)], { simp only [add_coeff, mul_add, sum_add_distrib] }, { intro b, - simp only [add_coeff, ne.def, set.mem_union_eq, set.mem_set_of_eq, mem_support], + simp only [add_coeff, ne.def, set.mem_union, set.mem_set_of_eq, mem_support], contrapose!, intro h, rw [h.1, h.2, add_zero], } @@ -560,7 +560,7 @@ instance [non_unital_non_assoc_semiring R] : distrib (hahn_series Γ R) := mul_coeff_left' hwf (set.subset_union_left _ _)], { simp only [add_coeff, add_mul, sum_add_distrib] }, { intro b, - simp only [add_coeff, ne.def, set.mem_union_eq, set.mem_set_of_eq, mem_support], + simp only [add_coeff, ne.def, set.mem_union, set.mem_set_of_eq, mem_support], contrapose!, intro h, rw [h.1, h.2, add_zero], }, diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index 480c85d277af1..b845c9e7603c8 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -1182,8 +1182,8 @@ begin {x | x ∈ s ∧ x ∈ a ∧ y x ∉ a ∧ ¬(is_bot x)} ∪ {x | is_bot x}, { assume x hx, by_cases h'x : is_bot x, - { simp only [h'x, mem_set_of_eq, mem_union_eq, not_true, and_false, false_or] }, - { simpa only [h'x, hx.2.1, hx.2.2, mem_set_of_eq, mem_union_eq, + { simp only [h'x, mem_set_of_eq, mem_union, not_true, and_false, false_or] }, + { simpa only [h'x, hx.2.1, hx.2.2, mem_set_of_eq, mem_union, not_false_iff, and_true, or_false] using hx.left } }, exact countable.mono this (H.union (subsingleton_is_bot α).countable) }, let t := {x | x ∈ s ∧ x ∈ a ∧ y x ∉ a ∧ ¬(is_bot x)}, diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 4967dbfe8b830..8e23d770d3e54 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -1067,7 +1067,7 @@ theorem mem_closure_iff_nhds_basis {a : α} {p : ι → Prop} {s : ι → set α {t : set α} : a ∈ closure t ↔ ∀ i, p i → ∃ y ∈ t, y ∈ s i := (mem_closure_iff_nhds_basis' h).trans $ - by simp only [set.nonempty, mem_inter_eq, exists_prop, and_comm] + by simp only [set.nonempty, mem_inter_iff, exists_prop, and_comm] /-- `x` belongs to the closure of `s` if and only if some ultrafilter supported on `s` converges to `x`. -/ @@ -1262,7 +1262,7 @@ h ht lemma eventually_eq_zero_nhds {M₀} [has_zero M₀] {a : α} {f : α → M₀} : f =ᶠ[𝓝 a] 0 ↔ a ∉ closure (function.support f) := -by rw [← mem_compl_eq, ← interior_compl, mem_interior_iff_mem_nhds, function.compl_support]; refl +by rw [← mem_compl_iff, ← interior_compl, mem_interior_iff_mem_nhds, function.compl_support]; refl lemma cluster_pt.map {x : α} {la : filter α} {lb : filter β} (H : cluster_pt x la) {f : α → β} (hfc : continuous_at f x) (hf : tendsto f la lb) : diff --git a/src/topology/category/Profinite/default.lean b/src/topology/category/Profinite/default.lean index 8b99832f1c294..855d2ba1e657c 100644 --- a/src/topology/category/Profinite/default.lean +++ b/src/topology/category/Profinite/default.lean @@ -274,7 +274,7 @@ begin ext x, dsimp [locally_constant.of_clopen], rw if_neg, { refl }, refine mt (λ α, hVU α) _, - simp only [set.mem_range_self, not_true, not_false_iff, set.mem_compl_eq], }, + simp only [set.mem_range_self, not_true, not_false_iff, set.mem_compl_iff], }, apply_fun (λ e, (e y).down) at H, dsimp [locally_constant.of_clopen] at H, rw if_pos hyV at H, diff --git a/src/topology/category/Top/limits.lean b/src/topology/category/Top/limits.lean index ff8728bce2a7c..2cc567431158b 100644 --- a/src/topology/category/Top/limits.lean +++ b/src/topology/category/Top/limits.lean @@ -310,7 +310,7 @@ begin ext, split, { rintros ⟨y, rfl⟩, - simp only [set.mem_preimage, set.mem_range, set.mem_inter_eq, ←comp_apply], + simp only [set.mem_preimage, set.mem_range, set.mem_inter_iff, ←comp_apply], simp only [limits.prod.map_fst, limits.prod.map_snd, exists_apply_eq_apply, comp_apply, and_self] }, { rintros ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩, diff --git a/src/topology/fiber_bundle.lean b/src/topology/fiber_bundle.lean index b3350fdf80eaf..02e43717d75b6 100644 --- a/src/topology/fiber_bundle.lean +++ b/src/topology/fiber_bundle.lean @@ -237,7 +237,7 @@ end begin ext ⟨x, y⟩, suffices : x ∈ e.base_set → (proj (e.to_local_equiv.symm (x, y)) ∈ s ↔ x ∈ s), - by simpa only [prod_mk_mem_set_prod_eq, mem_inter_eq, and_true, mem_univ, and.congr_left_iff], + by simpa only [prod_mk_mem_set_prod_eq, mem_inter_iff, and_true, mem_univ, and.congr_left_iff], intro h, rw [e.proj_symm_apply' h] end @@ -935,14 +935,14 @@ def triv_change (i j : ι) : local_homeomorph (B × F) (B × F) := map_target' := λp hp, by simpa using hp, left_inv' := begin rintros ⟨x, v⟩ hx, - simp only [prod_mk_mem_set_prod_eq, mem_inter_eq, and_true, mem_univ] at hx, + simp only [prod_mk_mem_set_prod_eq, mem_inter_iff, and_true, mem_univ] at hx, rw [Z.coord_change_comp, Z.coord_change_self], { exact hx.1 }, { simp [hx] } end, right_inv' := begin rintros ⟨x, v⟩ hx, - simp only [prod_mk_mem_set_prod_eq, mem_inter_eq, and_true, mem_univ] at hx, + simp only [prod_mk_mem_set_prod_eq, mem_inter_iff, and_true, mem_univ] at hx, rw [Z.coord_change_comp, Z.coord_change_self], { exact hx.2 }, { simp [hx] }, @@ -983,14 +983,14 @@ def local_triv_as_local_equiv (i : ι) : local_equiv Z.total_space (B × F) := dsimp only, rw [Z.coord_change_comp, Z.coord_change_self], { exact Z.mem_base_set_at _ }, - { simp only [hx, mem_inter_eq, and_self, mem_base_set_at] } + { simp only [hx, mem_inter_iff, and_self, mem_base_set_at] } end, right_inv' := begin rintros ⟨x, v⟩ hx, simp only [prod_mk_mem_set_prod_eq, and_true, mem_univ] at hx, rw [Z.coord_change_comp, Z.coord_change_self], { exact hx }, - { simp only [hx, mem_inter_eq, and_self, mem_base_set_at] } + { simp only [hx, mem_inter_iff, and_self, mem_base_set_at] } end } variable (i : ι) @@ -1015,10 +1015,10 @@ begin { ext x, simp only [mem_local_triv_as_local_equiv_target] with mfld_simps, refl, }, { rintros ⟨x, v⟩ hx, simp only [triv_change, local_triv_as_local_equiv, local_equiv.symm, true_and, prod.mk.inj_iff, - prod_mk_mem_set_prod_eq, local_equiv.trans_source, mem_inter_eq, and_true, mem_preimage, proj, - mem_univ, local_equiv.coe_mk, eq_self_iff_true, local_equiv.coe_trans, + prod_mk_mem_set_prod_eq, local_equiv.trans_source, mem_inter_iff, and_true, mem_preimage, + proj, mem_univ, local_equiv.coe_mk, eq_self_iff_true, local_equiv.coe_trans, total_space.proj] at hx ⊢, - simp only [Z.coord_change_comp, hx, mem_inter_eq, and_self, mem_base_set_at], } + simp only [Z.coord_change_comp, hx, mem_inter_iff, and_self, mem_base_set_at], } end variable (ι) @@ -1037,7 +1037,7 @@ begin simp only [exists_prop, mem_Union, mem_singleton_iff], refine ⟨i, Z.base_set i ×ˢ univ, (Z.is_open_base_set i).prod is_open_univ, _⟩, ext p, - simp only [local_triv_as_local_equiv_apply, prod_mk_mem_set_prod_eq, mem_inter_eq, and_self, + simp only [local_triv_as_local_equiv_apply, prod_mk_mem_set_prod_eq, mem_inter_iff, and_self, mem_local_triv_as_local_equiv_source, and_true, mem_univ, mem_preimage], end diff --git a/src/topology/locally_constant/basic.lean b/src/topology/locally_constant/basic.lean index 78df3d95d6ba8..a5ef8903ab912 100644 --- a/src/topology/locally_constant/basic.lean +++ b/src/topology/locally_constant/basic.lean @@ -317,7 +317,7 @@ begin ext, simp only [of_clopen, nat.one_ne_zero, mem_singleton_iff, coe_mk, fin.zero_eq_one_iff, mem_preimage, ite_eq_right_iff, - mem_compl_eq], + mem_compl_iff], tauto, end diff --git a/src/topology/locally_finite.lean b/src/topology/locally_finite.lean index 02bbec2d1364b..58eaf95870161 100644 --- a/src/topology/locally_finite.lean +++ b/src/topology/locally_finite.lean @@ -92,7 +92,7 @@ begin have : t ∩ (⋂ i ∈ {i | (f i ∩ t).nonempty}, (f i)ᶜ) ∈ 𝓝 a, from inter_mem h_nhds ((bInter_mem h_fin).2 (λ i _, ha i)), filter_upwards [this], - simp only [mem_inter_eq, mem_Inter], + simp only [mem_inter_iff, mem_Inter], rintros b ⟨hbt, hn⟩ i hfb, exact hn i ⟨b, hfb, hbt⟩ hfb, end @@ -129,7 +129,7 @@ begin replace hN : ∀ x (n > N x) (y ∈ U x), f (n + 1) y = f n y, from λ x n hn y hy, by_contra (λ hne, hn.lt.not_le $ hN x ⟨y, hne, hy⟩), replace hN : ∀ x (n ≥ N x + 1) (y ∈ U x), f n y = f (N x + 1) y, - from λ x n hn y hy, nat.le_induction rfl (λ k hle, (hN x _ hle _ hy).trans) n hn, + from λ x n hn y hy, nat.le_induction rfl (λ k hle, (hN x _ hle _ hy).trans) n hn, refine ⟨λ x, f (N x + 1) x, λ x, _⟩, filter_upwards [filter.prod_mem_prod (eventually_gt_at_top (N x)) (hUx x)], rintro ⟨n, y⟩ ⟨hn : N x < n, hy : y ∈ U x⟩, diff --git a/src/topology/metric_space/baire.lean b/src/topology/metric_space/baire.lean index 56ce034d0a0ba..543c96184f490 100644 --- a/src/topology/metric_space/baire.lean +++ b/src/topology/metric_space/baire.lean @@ -164,7 +164,7 @@ begin /- Prove that ̀`⋂ n : ℕ, K n` is inside `U ∩ ⋂ n : ℕ, (f n)`. -/ have hK_subset : (⋂ n, K n : set α) ⊆ U ∩ (⋂ n, f n), { intros x hx, - simp only [mem_inter_eq, mem_Inter] at hx ⊢, + simp only [mem_inter_iff, mem_Inter] at hx ⊢, exact ⟨hK₀ $ hx 0, λ n, (hK_decreasing n (hx (n + 1))).1⟩ }, /- Prove that `⋂ n : ℕ, K n` is not empty, as an intersection of a decreasing sequence of nonempty compact subsets.-/ diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 252166c1afc17..3baf72b5accb9 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1686,7 +1686,7 @@ subset.antisymm lemma dense_iff {s : set α} : dense s ↔ ∀ x, ∀ r > 0, (ball x r ∩ s).nonempty := -forall_congr $ λ x, by simp only [mem_closure_iff, set.nonempty, exists_prop, mem_inter_eq, +forall_congr $ λ x, by simp only [mem_closure_iff, set.nonempty, exists_prop, mem_inter_iff, mem_ball', and_comm] lemma dense_range_iff {f : β → α} : diff --git a/src/topology/metric_space/gromov_hausdorff_realized.lean b/src/topology/metric_space/gromov_hausdorff_realized.lean index c5f09a4aaed7d..1d8bfd5f457e4 100644 --- a/src/topology/metric_space/gromov_hausdorff_realized.lean +++ b/src/topology/metric_space/gromov_hausdorff_realized.lean @@ -244,7 +244,7 @@ begin ∩ (⋂x y z, {f : Cb X Y | f (x, z) ≤ f (x, y) + f (y, z)}) ∩ (⋂x, {f : Cb X Y | f (x, x) = 0}) ∩ (⋂x y, {f : Cb X Y | f (x, y) ≤ max_var X Y}), - { ext, simp only [candidates_b, candidates, mem_inter_eq, mem_Inter, mem_set_of_eq] }, + { ext, simp only [candidates_b, candidates, mem_inter_iff, mem_Inter, mem_set_of_eq] }, rw this, repeat { apply is_closed.inter _ _ <|> apply is_closed_Inter _ diff --git a/src/topology/metric_space/partition_of_unity.lean b/src/topology/metric_space/partition_of_unity.lean index 5fa49bc4e0c8a..22f30453a6e2b 100644 --- a/src/topology/metric_space/partition_of_unity.lean +++ b/src/topology/metric_space/partition_of_unity.lean @@ -83,7 +83,7 @@ we have `emetric.closed_ball x (ennreal.of_real (δ x)) ⊆ U i`. -/ lemma exists_continuous_real_forall_closed_ball_subset (hK : ∀ i, is_closed (K i)) (hU : ∀ i, is_open (U i)) (hKU : ∀ i, K i ⊆ U i) (hfin : locally_finite K) : ∃ δ : C(X, ℝ), (∀ x, 0 < δ x) ∧ ∀ i (x ∈ K i), closed_ball x (ennreal.of_real $ δ x) ⊆ U i := -by simpa only [mem_inter_eq, forall_and_distrib, mem_preimage, mem_Inter, @forall_swap ι X] +by simpa only [mem_inter_iff, forall_and_distrib, mem_preimage, mem_Inter, @forall_swap ι X] using exists_continuous_forall_mem_convex_of_local_const exists_forall_closed_ball_subset_aux₂ (exists_forall_closed_ball_subset_aux₁ hK hU hKU hfin) diff --git a/src/topology/paracompact.lean b/src/topology/paracompact.lean index cc09cecc8b346..2dac44fc41ac6 100644 --- a/src/topology/paracompact.lean +++ b/src/topology/paracompact.lean @@ -81,7 +81,7 @@ begin { simp only [eq_univ_iff_forall, mem_Union], exact λ x, ⟨ind (t_inv x), _, rfl, ht_inv _⟩ }, { refine λ x, ⟨U x, hxU x, ((hU x).image ind).subset _⟩, - simp only [subset_def, mem_Union, mem_set_of_eq, set.nonempty, mem_inter_eq], + simp only [subset_def, mem_Union, mem_set_of_eq, set.nonempty, mem_inter_iff], rintro i ⟨y, ⟨a, rfl, hya⟩, hyU⟩, exact mem_image_of_mem _ ⟨y, hya, hyU⟩ }, { simp only [subset_def, mem_Union], @@ -178,7 +178,7 @@ begin have : (⋃ k ≤ K'.find x + 2, (range $ sigma.mk k) : set (Σ n, T' n)).finite, from (finite_le_nat _).bUnion (λ k hk, finite_range _), apply this.subset, rintro ⟨k, c, hc⟩, - simp only [mem_Union, mem_set_of_eq, mem_image_eq, subtype.coe_mk], + simp only [mem_Union, mem_set_of_eq, mem_image, subtype.coe_mk], rintro ⟨x, hxB : x ∈ B c (r k c), hxK⟩, refine ⟨k, _, ⟨c, hc⟩, rfl⟩, have := (mem_compl_iff _ _).1 (hr k c hxB), diff --git a/src/topology/shrinking_lemma.lean b/src/topology/shrinking_lemma.lean index cc359edf4b3ce..9224bb9a6526a 100644 --- a/src/topology/shrinking_lemma.lean +++ b/src/topology/shrinking_lemma.lean @@ -171,7 +171,7 @@ lemma exists_gt (v : partial_refinement u s) (hs : is_closed s) (i : ι) (hi : i ∃ v' : partial_refinement u s, v < v' := begin have I : s ∩ (⋂ j ≠ i, (v j)ᶜ) ⊆ v i, - { simp only [subset_def, mem_inter_eq, mem_Inter, and_imp], + { simp only [subset_def, mem_inter_iff, mem_Inter, and_imp], intros x hxs H, rcases mem_Union.1 (v.subset_Union hxs) with ⟨j, hj⟩, exact (em (j = i)).elim (λ h, h ▸ hj) (λ h, (H j h hj).elim) }, diff --git a/src/topology/sober.lean b/src/topology/sober.lean index a4ac489f6758e..aedd885c1f831 100644 --- a/src/topology/sober.lean +++ b/src/topology/sober.lean @@ -186,7 +186,7 @@ begin set.image_singleton, (show _ = _, from hx)], apply set.image_injective.mpr hf.inj, ext z, - simp only [set.image_preimage_eq_inter_range, set.mem_inter_eq, and.congr_left_iff], + simp only [set.image_preimage_eq_inter_range, set.mem_inter_iff, and.congr_left_iff], exact λ hy, ⟨λ h, hT.closure_eq ▸ closure_mono (set.inter_subset_left _ _) h, λ h, subset_closure ⟨h, hy⟩⟩ end diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index 627d55a1d3044..c2d43ef7b82bb 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -211,10 +211,10 @@ lemma is_compact.elim_finite_subfamily_closed {s : set α} {ι : Type v} (hs : i ∃ t : finset ι, s ∩ (⋂ i ∈ t, Z i) = ∅ := let ⟨t, ht⟩ := hs.elim_finite_subcover (λ i, (Z i)ᶜ) (λ i, (hZc i).is_open_compl) (by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_Union, - exists_prop, mem_inter_eq, not_and, iff_self, mem_Inter, mem_compl_eq] using hsZ) + exists_prop, mem_inter_iff, not_and, iff_self, mem_Inter, mem_compl_iff] using hsZ) in ⟨t, by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_Union, - exists_prop, mem_inter_eq, not_and, iff_self, mem_Inter, mem_compl_eq] using ht⟩ + exists_prop, mem_inter_iff, not_and, iff_self, mem_Inter, mem_compl_iff] using ht⟩ /-- If `s` is a compact set in a topological space `α` and `f : ι → set α` is a locally finite family of sets, then `f i ∩ s` is nonempty only for a finitely many `i`. -/ @@ -333,10 +333,10 @@ is_compact_of_finite_subfamily_closed $ assume ι Z hZc hsZ, let ⟨t, ht⟩ := h (λ i, (Z i)ᶜ) (assume i, is_open_compl_iff.mpr $ hZc i) (by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_Union, - exists_prop, mem_inter_eq, not_and, iff_self, mem_Inter, mem_compl_eq] using hsZ) + exists_prop, mem_inter_iff, not_and, iff_self, mem_Inter, mem_compl_iff] using hsZ) in ⟨t, by simpa only [subset_def, not_forall, eq_empty_iff_forall_not_mem, mem_Union, - exists_prop, mem_inter_eq, not_and, iff_self, mem_Inter, mem_compl_eq] using ht⟩ + exists_prop, mem_inter_iff, not_and, iff_self, mem_Inter, mem_compl_iff] using ht⟩ /-- A set `s` is compact if and only if for every open cover of `s`, there exists a finite subcover. -/ @@ -455,7 +455,7 @@ begin (λ i, (hV_cpct i).inter_right W_op.is_closed_compl) (λ i, (hV_closed i).inter W_op.is_closed_compl), rcases hV i j with ⟨k, hki, hkj⟩, - refine ⟨k, ⟨λ x, _, λ x, _⟩⟩ ; simp only [and_imp, mem_inter_eq, mem_compl_eq] ; tauto }, + refine ⟨k, ⟨λ x, _, λ x, _⟩⟩ ; simp only [and_imp, mem_inter_iff, mem_compl_iff] ; tauto }, have : ¬ (⋂ (i : ι), V i) ⊆ W, by simpa [← Inter_inter, inter_compl_nonempty_iff], contradiction end @@ -951,7 +951,7 @@ begin rw compl_subset_comm at ⊢ hAt hBt', refine subset.trans _ (set.prod_mono hAt hBt'), intros x, - simp only [compl_union, mem_inter_eq, mem_prod, mem_preimage, mem_compl_eq], + simp only [compl_union, mem_inter_iff, mem_prod, mem_preimage, mem_compl_iff], tauto }, { rintros ⟨t, ht, htS⟩, refine ⟨⟨(prod.fst '' t)ᶜ, _, _⟩, ⟨(prod.snd '' t)ᶜ, _, _⟩⟩, @@ -1176,7 +1176,7 @@ begin { convert_to (⋂(U : {U // U ∈ c}), U.1ᶜ).nonempty, { ext, simp only [not_exists, exists_prop, not_and, set.mem_Inter, subtype.forall, mem_set_of_eq, - mem_compl_eq, mem_sUnion] }, + mem_compl_iff, mem_sUnion] }, apply is_compact.nonempty_Inter_of_directed_nonempty_compact_closed, { rintros ⟨U, hU⟩ ⟨U', hU'⟩, obtain ⟨V, hVc, hVU, hVU'⟩ := hz.directed_on U hU U' hU', diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index 22e1275bd01b6..b95814c22b205 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -805,7 +805,7 @@ lemma closure_eq_uniformity (s : set $ α × α) : begin ext ⟨x, y⟩, simp only [mem_closure_iff_nhds_basis (uniform_space.has_basis_nhds_prod x y), mem_Inter, - mem_set_of_eq, and_imp, mem_comp_comp, exists_prop, ← mem_inter_eq, inter_comm, set.nonempty] + mem_set_of_eq, and_imp, mem_comp_comp, exists_prop, ← mem_inter_iff, inter_comm, set.nonempty] { contextual := tt } end @@ -991,7 +991,7 @@ lemma filter.has_basis.uniform_continuous_on_iff [uniform_space β] {p : γ → uniform_continuous_on f S ↔ ∀ i (hi : q i), ∃ j (hj : p j), ∀ x y ∈ S, (x, y) ∈ s j → (f x, f y) ∈ t i := ((ha.inf_principal (S ×ˢ S)).tendsto_iff hb).trans $ -by simp [prod.forall, set.inter_comm (s _), ball_mem_comm] +by simp_rw [prod.forall, set.inter_comm (s _), ball_mem_comm, mem_inter_iff, mem_prod, and_imp] end uniform_space diff --git a/src/topology/uniform_space/compact_convergence.lean b/src/topology/uniform_space/compact_convergence.lean index 357eee4ebcf10..c2e801812db46 100644 --- a/src/topology/uniform_space/compact_convergence.lean +++ b/src/topology/uniform_space/compact_convergence.lean @@ -285,7 +285,7 @@ begin rintros ⟨K₁, V₁⟩ ⟨hK₁, hV₁⟩ ⟨K₂, V₂⟩ ⟨hK₂, hV₂⟩, refine ⟨⟨K₁ ∪ K₂, V₁ ∩ V₂⟩, ⟨hK₁.union hK₂, filter.inter_mem hV₁ hV₂⟩, _⟩, simp only [le_eq_subset, prod.forall, set_of_subset_set_of, ge_iff_le, order.preimage, - ← forall_and_distrib, mem_inter_eq, mem_union_eq], + ← forall_and_distrib, mem_inter_iff, mem_union], exact λ f g, forall_imp (λ x, by tauto!), end diff --git a/src/topology/vector_bundle/basic.lean b/src/topology/vector_bundle/basic.lean index a5152e0f1972f..ce7ea0ab66b69 100644 --- a/src/topology/vector_bundle/basic.lean +++ b/src/topology/vector_bundle/basic.lean @@ -887,9 +887,9 @@ instance : topological_vector_bundle R F Z.fiber := rw [preimage_inter, ←preimage_comp, function.comp], simp only [total_space_mk], refine ext_iff.mpr (λ a, ⟨λ ha, _, λ ha, ⟨Z.mem_base_set_at b, _⟩⟩), - { simp only [mem_prod, mem_preimage, mem_inter_eq, local_triv_at_apply_mk] at ha, + { simp only [mem_prod, mem_preimage, mem_inter_iff, local_triv_at_apply_mk] at ha, exact ha.2.2, }, - { simp only [mem_prod, mem_preimage, mem_inter_eq, local_triv_at_apply_mk], + { simp only [mem_prod, mem_preimage, mem_inter_iff, local_triv_at_apply_mk], exact ⟨Z.mem_base_set_at b, ha⟩, } } end⟩, trivialization_atlas := set.range Z.local_triv, trivialization_at := Z.local_triv_at, From f18afbe9ac56dfc067aa397a41d232b4ceee776b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Thu, 22 Sep 2022 22:45:35 +0000 Subject: [PATCH 352/828] feat(set_theory/game/nim): recursors for left/right nim moves (#15408) --- src/set_theory/game/nim.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index ef5a43a5db66f..07e25164cf376 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -96,6 +96,16 @@ lemma move_right_nim {o : ordinal} (i) : (nim o).move_right (to_right_moves_nim i) = nim i := by simp +/-- A recursion principle for left moves of a nim game. -/ +@[elab_as_eliminator] def left_moves_nim_rec_on {o : ordinal} {P : (nim o).left_moves → Sort*} + (i : (nim o).left_moves) (H : ∀ a < o, P $ to_left_moves_nim ⟨a, H⟩) : P i := +by { rw ←to_left_moves_nim.apply_symm_apply i, apply H } + +/-- A recursion principle for right moves of a nim game. -/ +@[elab_as_eliminator] def right_moves_nim_rec_on {o : ordinal} {P : (nim o).right_moves → Sort*} + (i : (nim o).right_moves) (H : ∀ a < o, P $ to_right_moves_nim ⟨a, H⟩) : P i := +by { rw ←to_right_moves_nim.apply_symm_apply i, apply H } + instance is_empty_nim_zero_left_moves : is_empty (nim 0).left_moves := by { rw nim_def, exact ordinal.is_empty_out_zero } From 1732f9cfcdc724e48c4aa30e35c46ea52e64c0d8 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 22 Sep 2022 22:45:36 +0000 Subject: [PATCH 353/828] chore(topology): golf (#16585) 2 minor golfs from an old branch --- src/topology/algebra/group.lean | 5 ++--- src/topology/constructions.lean | 2 +- 2 files changed, 3 insertions(+), 4 deletions(-) diff --git a/src/topology/algebra/group.lean b/src/topology/algebra/group.lean index 42c9f6b5f9e42..946bc3416240a 100644 --- a/src/topology/algebra/group.lean +++ b/src/topology/algebra/group.lean @@ -1427,10 +1427,9 @@ Inf {b : group_topology β | (topological_space.coinduced f t) ≤ b.to_topologi lemma coinduced_continuous {α β : Type*} [t : topological_space α] [group β] (f : α → β) : cont t (coinduced f).to_topological_space f := begin - rw continuous_iff_coinduced_le, - refine le_Inf _, + rw [continuous_Inf_rng], rintros _ ⟨t', ht', rfl⟩, - exact ht', + exact continuous_iff_coinduced_le.2 ht' end end group_topology diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index 392bb4919e27b..cb915182bc25a 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -676,7 +676,7 @@ begin rw [is_open_map_iff_nhds_le], rintros ⟨a, b⟩, rw [nhds_prod_eq, nhds_prod_eq, ← filter.prod_map_map_eq], - exact filter.prod_mono (is_open_map_iff_nhds_le.1 hf a) (is_open_map_iff_nhds_le.1 hg b) + exact filter.prod_mono (hf.nhds_le a) (hg.nhds_le b) end protected lemma open_embedding.prod {f : α → β} {g : γ → δ} From 9806b5c30f479bb9e02e4ab75905e26a09206f20 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Thu, 22 Sep 2022 22:45:37 +0000 Subject: [PATCH 354/828] doc(topology/algebra/filter_basis): modify doc of `add_group_filter_basis_of_comm` (#16597) Added the `additive` adjective to the doc for the @additive variation `add_group_filter_basis_of_comm` of `group_filter_basis_of_comm` Co-authored-by: Kyle Miller --- src/topology/algebra/filter_basis.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/topology/algebra/filter_basis.lean b/src/topology/algebra/filter_basis.lean index 0f68d87bbddbe..5ef959246ca60 100644 --- a/src/topology/algebra/filter_basis.lean +++ b/src/topology/algebra/filter_basis.lean @@ -67,7 +67,7 @@ attribute [to_additive] group_filter_basis.conj' attribute [to_additive] group_filter_basis.to_filter_basis /-- `group_filter_basis` constructor in the commutative group case. -/ -@[to_additive "`add_group_filter_basis` constructor in the commutative group case."] +@[to_additive "`add_group_filter_basis` constructor in the additive commutative group case."] def group_filter_basis_of_comm {G : Type*} [comm_group G] (sets : set (set G)) (nonempty : sets.nonempty) From 33075a0271a28e8a2965311eedd29e9266e5fe5f Mon Sep 17 00:00:00 2001 From: mcdoll Date: Thu, 22 Sep 2022 22:45:38 +0000 Subject: [PATCH 355/828] chore(analysis): golf a few proofs (#16600) --- src/analysis/calculus/taylor.lean | 10 ++++------ src/analysis/special_functions/stirling.lean | 6 +++--- 2 files changed, 7 insertions(+), 9 deletions(-) diff --git a/src/analysis/calculus/taylor.lean b/src/analysis/calculus/taylor.lean index 6afa45503665c..9af200b2e4c71 100644 --- a/src/analysis/calculus/taylor.lean +++ b/src/analysis/calculus/taylor.lean @@ -272,7 +272,7 @@ begin simp only [taylor_within_eval_self] at h, rw [mul_comm, ←div_left_inj' (g'_ne y hy), mul_div_cancel _ (g'_ne y hy)] at h, rw ←h, - field_simp [g'_ne y hy, nat.factorial_ne_zero n], + field_simp [g'_ne y hy, n.factorial_ne_zero], ring, end @@ -307,7 +307,7 @@ begin use [y, hy], simp only [sub_self, zero_pow', ne.def, nat.succ_ne_zero, not_false_iff, zero_sub, mul_neg] at h, rw [h, neg_div, ←div_neg, neg_mul, neg_neg], - field_simp [nat.cast_add_one_ne_zero n, nat.factorial_ne_zero n, xy_ne y hy], + field_simp [n.cast_add_one_ne_zero, n.factorial_ne_zero, xy_ne y hy], ring, end @@ -331,7 +331,7 @@ begin rcases taylor_mean_remainder hx hf hf' gcont gdiff (λ _ _, by simp) with ⟨y, hy, h⟩, use [y, hy], rw h, - field_simp [nat.factorial_ne_zero n], + field_simp [n.factorial_ne_zero], ring, end @@ -352,8 +352,6 @@ begin have hf' : differentiable_on ℝ (iterated_deriv_within n f (Icc a b)) (Icc a b) := hf.differentiable_on_iterated_deriv_within (with_top.coe_lt_coe.mpr n.lt_succ_self) (unique_diff_on_Icc h), - -- natural numbers are non-negative - have fac_nonneg : 0 ≤ (n! : ℝ) := n!.cast_nonneg, -- We can uniformly bound the derivative of the Taylor polynomial have h' : ∀ (y : ℝ) (hy : y ∈ Ico a x), ∥((n! : ℝ)⁻¹ * (x - y) ^ n) • iterated_deriv_within (n + 1) f (Icc a b) y∥ @@ -376,7 +374,7 @@ begin exact (has_deriv_within_taylor_within_eval_at_Icc x h (I ht) hf.of_succ hf').mono I }, have := norm_image_sub_le_of_norm_deriv_le_segment' A h' x (right_mem_Icc.2 hx.1), simp only [taylor_within_eval_self] at this, - refine le_trans this (le_of_eq _), + refine this.trans_eq _, -- The rest is a trivial calculation rw [abs_of_nonneg (sub_nonneg.mpr hx.1)], ring_exp, diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index fa59c375ca816..d2f08f2a4e514 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -61,7 +61,7 @@ We have the expression lemma log_stirling_seq_formula (n : ℕ) : log (stirling_seq n.succ) = log n.succ!- 1 / 2 * log (2 * n.succ) - n.succ * log (n.succ / exp 1) := begin - have h1 : (0 : ℝ) < n.succ!:= cast_pos.mpr n.succ.factorial_pos, + have h1 : (0 : ℝ) < n.succ! := cast_pos.mpr n.succ.factorial_pos, have h2 : (0 : ℝ) < (2 * n.succ) := mul_pos two_pos (cast_pos.mpr (succ_pos n)), have h3 := real.sqrt_pos.mpr h2, have h4 := pow_pos (div_pos (cast_pos.mpr n.succ_pos ) (exp_pos 1)) n.succ, @@ -132,8 +132,8 @@ We have the bound `log (stirling_seq n) - log (stirling_seq (n+1))` ≤ 1/(4 n^ lemma log_stirling_seq_sub_log_stirling_seq_succ (n : ℕ) : log (stirling_seq n.succ) - log (stirling_seq n.succ.succ) ≤ 1 / (4 * n.succ ^ 2) := begin - have h₁ : 0 < 4 * ((n : ℝ) + 1) ^ 2 := by nlinarith [@cast_nonneg ℝ _ n], - have h₃ : 0 < (2 * ((n : ℝ) + 1) + 1) ^ 2 := by nlinarith [@cast_nonneg ℝ _ n], + have h₁ : 0 < 4 * ((n : ℝ) + 1) ^ 2 := by positivity, + have h₃ : 0 < (2 * ((n : ℝ) + 1) + 1) ^ 2 := by positivity, have h₂ : 0 < 1 - (1 / (2 * ((n : ℝ) + 1) + 1)) ^ 2, { rw ← mul_lt_mul_right h₃, have H : 0 < (2 * ((n : ℝ) + 1) + 1) ^ 2 - 1 := by nlinarith [@cast_nonneg ℝ _ n], From 98cbfb459a053c5ca44aec69f0a5a932b84c0d67 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 23 Sep 2022 04:37:03 +0000 Subject: [PATCH 356/828] chore(tactic/monotonicity/lemmas): Move lemmas to the correct places (#16509) A few lemmas were gathered in `tactic.monotonicity.lemmas` rather than the files they belong to. Generalize them, move them and fix their names. --- src/algebra/order/ring.lean | 6 ++ src/algebra/order/sub.lean | 21 ++++--- src/analysis/calculus/taylor.lean | 9 ++- .../inner_product_space/projection.lean | 3 +- src/data/bitvec/basic.lean | 4 +- .../liouville/liouville_constant.lean | 3 +- src/tactic/monotonicity/lemmas.lean | 62 +------------------ 7 files changed, 30 insertions(+), 78 deletions(-) diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 9eb3aca53cfe9..b8d50b646ad02 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -1164,6 +1164,12 @@ lemma neg_one_lt_zero : -1 < (0:α) := neg_lt_zero.2 zero_lt_one @[simp] lemma mul_lt_mul_right_of_neg {a b c : α} (h : c < 0) : a * c < b * c ↔ b < a := (strict_anti_mul_right h).lt_iff_lt +lemma lt_of_mul_lt_mul_of_nonpos_left (h : c * a < c * b) (hc : c ≤ 0) : b < a := +lt_of_mul_lt_mul_left (by rwa [neg_mul, neg_mul, neg_lt_neg_iff]) $ neg_nonneg.2 hc + +lemma lt_of_mul_lt_mul_of_nonpos_right (h : a * c < b * c) (hc : c ≤ 0) : b < a := +lt_of_mul_lt_mul_right (by rwa [mul_neg, mul_neg, neg_lt_neg_iff]) $ neg_nonneg.2 hc + lemma cmp_mul_neg_left {a : α} (ha : a < 0) (b c : α) : cmp (a * b) (a * c) = cmp c b := (strict_anti_mul_left ha).cmp_map_eq b c diff --git a/src/algebra/order/sub.lean b/src/algebra/order/sub.lean index bf808bdee5cea..e1b90310e8cd6 100644 --- a/src/algebra/order/sub.lean +++ b/src/algebra/order/sub.lean @@ -500,6 +500,10 @@ protected lemma lt_tsub_iff_left_of_le (hc : add_le_cancellable c) (h : c ≤ b) a < b - c ↔ c + a < b := by { rw [add_comm], exact hc.lt_tsub_iff_right_of_le h } +protected lemma tsub_inj_right (hab : add_le_cancellable (a - b)) (h₁ : b ≤ a) (h₂ : c ≤ a) + (h₃ : a - b = a - c) : b = c := +by { rw ← hab.inj, rw [tsub_add_cancel_of_le h₁, h₃, tsub_add_cancel_of_le h₂] } + protected lemma lt_of_tsub_lt_tsub_left_of_le [contravariant_class α α (+) (<)] (hb : add_le_cancellable b) (hca : c ≤ a) (h : a - b < a - c) : c < b := begin @@ -507,22 +511,18 @@ begin exact lt_of_add_lt_add_left (hb.lt_add_of_tsub_lt_right h), end +protected lemma tsub_lt_tsub_left_of_le (hab : add_le_cancellable (a - b)) (h₁ : b ≤ a) + (h : c < b) : a - b < a - c := +(tsub_le_tsub_left h.le _).lt_of_ne $ λ h', h.ne' $ hab.tsub_inj_right h₁ (h.le.trans h₁) h' + protected lemma tsub_lt_tsub_right_of_le (hc : add_le_cancellable c) (h : c ≤ a) (h2 : a < b) : a - c < b - c := by { apply hc.lt_tsub_of_add_lt_left, rwa [add_tsub_cancel_of_le h] } -protected lemma tsub_inj_right (hab : add_le_cancellable (a - b)) (h₁ : b ≤ a) (h₂ : c ≤ a) - (h₃ : a - b = a - c) : b = c := -by { rw ← hab.inj, rw [tsub_add_cancel_of_le h₁, h₃, tsub_add_cancel_of_le h₂] } - protected lemma tsub_lt_tsub_iff_left_of_le_of_le [contravariant_class α α (+) (<)] (hb : add_le_cancellable b) (hab : add_le_cancellable (a - b)) (h₁ : b ≤ a) (h₂ : c ≤ a) : a - b < a - c ↔ c < b := -begin - refine ⟨hb.lt_of_tsub_lt_tsub_left_of_le h₂, _⟩, - intro h, refine (tsub_le_tsub_left h.le _).lt_of_ne _, - rintro h2, exact h.ne' (hab.tsub_inj_right h₁ h₂ h2) -end +⟨hb.lt_of_tsub_lt_tsub_left_of_le h₂, hab.tsub_lt_tsub_left_of_le h₁⟩ @[simp] protected lemma add_tsub_tsub_cancel (hac : add_le_cancellable (a - c)) (h : c ≤ a) : (a + b) - (a - c) = b + c := @@ -588,6 +588,9 @@ lemma lt_of_tsub_lt_tsub_left_of_le [contravariant_class α α (+) (<)] (hca : c ≤ a) (h : a - b < a - c) : c < b := contravariant.add_le_cancellable.lt_of_tsub_lt_tsub_left_of_le hca h +lemma tsub_lt_tsub_left_of_le : b ≤ a → c < b → a - b < a - c := +contravariant.add_le_cancellable.tsub_lt_tsub_left_of_le + lemma tsub_lt_tsub_right_of_le (h : c ≤ a) (h2 : a < b) : a - c < b - c := contravariant.add_le_cancellable.tsub_lt_tsub_right_of_le h h2 diff --git a/src/analysis/calculus/taylor.lean b/src/analysis/calculus/taylor.lean index 9af200b2e4c71..312cd3414e699 100644 --- a/src/analysis/calculus/taylor.lean +++ b/src/analysis/calculus/taylor.lean @@ -356,16 +356,15 @@ begin have h' : ∀ (y : ℝ) (hy : y ∈ Ico a x), ∥((n! : ℝ)⁻¹ * (x - y) ^ n) • iterated_deriv_within (n + 1) f (Icc a b) y∥ ≤ (n! : ℝ)⁻¹ * |(x - a)|^n * C, - { intros y hy, + { rintro y ⟨hay, hyx⟩, rw [norm_smul, real.norm_eq_abs], -- Estimate the iterated derivative by `C` - refine mul_le_mul _ (hC y ⟨hy.1, hy.2.le.trans hx.2⟩) (by positivity) (by positivity), + refine mul_le_mul _ (hC y ⟨hay, hyx.le.trans hx.2⟩) (by positivity) (by positivity), -- The rest is a trivial calculation rw [abs_mul, abs_pow, abs_inv, nat.abs_cast], - mono*, + mono* with [0 ≤ (n! : ℝ)⁻¹], any_goals { positivity }, - { exact hy.1 }, - { linarith [hx.1, hy.2] } }, + linarith [hx.1, hyx] }, -- Apply the mean value theorem for vector valued functions: have A : ∀ t ∈ Icc a x, has_deriv_within_at (λ y, taylor_within_eval f n (Icc a b) y x) (((↑n!)⁻¹ * (x - t) ^ n) • iterated_deriv_within (n + 1) f (Icc a b) t) (Icc a x) t, diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index 65d4ad647837b..a061b650cbc0d 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -123,7 +123,8 @@ begin repeat {exact le_of_lt one_half_pos}, exact add_halves 1 }, have eq₁ : 4 * δ * δ ≤ 4 * ∥u - half • (wq + wp)∥ * ∥u - half • (wq + wp)∥, - { mono, mono, norm_num, apply mul_nonneg, norm_num, exact norm_nonneg _ }, + { simp_rw mul_assoc, + exact mul_le_mul_of_nonneg_left (mul_self_le_mul_self zero_le_δ eq) zero_le_four }, have eq₂ : ∥a∥ * ∥a∥ ≤ (δ + div) * (δ + div) := mul_self_le_mul_self (norm_nonneg _) (le_trans (le_of_lt $ hw q) (add_le_add_left (nat.one_div_le_one_div hq) _)), diff --git a/src/data/bitvec/basic.lean b/src/data/bitvec/basic.lean index 354ec842d1a93..11fd64c81bef0 100644 --- a/src/data/bitvec/basic.lean +++ b/src/data/bitvec/basic.lean @@ -46,8 +46,8 @@ begin rw [add_lsb_eq_twice_add_one], transitivity 2 * list.foldr (λ (x : bool) (y : ℕ), add_lsb y x) 0 ys_tl + 2 * 1, { ac_mono, rw two_mul, mono, cases ys_hd; simp }, - { rw ← left_distrib, ac_mono, norm_num, - apply ys_ih, refl } }, + { rw ← left_distrib, ac_mono, + exact ys_ih rfl, norm_num } } end lemma add_lsb_div_two {x b} : add_lsb x b / 2 = x := diff --git a/src/number_theory/liouville/liouville_constant.lean b/src/number_theory/liouville/liouville_constant.lean index 09f200e7d362d..6cebf03edd18f 100644 --- a/src/number_theory/liouville/liouville_constant.lean +++ b/src/number_theory/liouville/liouville_constant.lean @@ -123,8 +123,7 @@ lemma aux_calc (n : ℕ) {m : ℝ} (hm : 2 ≤ m) : calc (1 - 1 / m)⁻¹ * (1 / m ^ (n + 1)!) ≤ 2 * (1 / m ^ (n + 1)!) : -- the second factors coincide (and are non-negative), -- the first factors, satisfy the inequality `sub_one_div_inv_le_two` - mul_mono_nonneg (one_div_nonneg.mpr (pow_nonneg (zero_le_two.trans hm) _)) - (sub_one_div_inv_le_two hm) + mul_le_mul_of_nonneg_right (sub_one_div_inv_le_two hm) (by positivity) ... = 2 / m ^ (n + 1)! : mul_one_div 2 _ ... = 2 / m ^ (n! * (n + 1)) : congr_arg ((/) 2) (congr_arg (pow m) (mul_comm _ _)) ... ≤ 1 / m ^ (n! * n) : diff --git a/src/tactic/monotonicity/lemmas.lean b/src/tactic/monotonicity/lemmas.lean index 396611f573581..9f7be88c70997 100644 --- a/src/tactic/monotonicity/lemmas.lean +++ b/src/tactic/monotonicity/lemmas.lean @@ -4,70 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon -/ import algebra.order.ring -import data.nat.basic import data.set.lattice -import order.directed import tactic.monotonicity.basic variables {α : Type*} -@[mono] -lemma mul_mono_nonneg {x y z : α} [ordered_semiring α] - (h' : 0 ≤ z) - (h : x ≤ y) -: x * z ≤ y * z := -by apply mul_le_mul_of_nonneg_right; assumption - -lemma lt_of_mul_lt_mul_neg_right {a b c : α} [linear_ordered_ring α] - (h : a * c < b * c) (hc : c ≤ 0) : b < a := -have nhc : -c ≥ 0, from neg_nonneg_of_nonpos hc, -have h2 : -(b * c) < -(a * c), from neg_lt_neg h, -have h3 : b * (-c) < a * (-c), from calc - b * (-c) = - (b * c) : by rewrite neg_mul_eq_mul_neg - ... < - (a * c) : h2 - ... = a * (-c) : by rewrite neg_mul_eq_mul_neg, -lt_of_mul_lt_mul_right h3 nhc - -@[mono] -lemma mul_mono_nonpos {x y z : α} [linear_ordered_ring α] - (h' : z ≤ 0) (h : y ≤ x) : x * z ≤ y * z := -begin - classical, - by_contradiction h'', - revert h, - apply not_le_of_lt, - apply lt_of_mul_lt_mul_neg_right _ h', - apply lt_of_not_ge h'' -end - -@[mono] -lemma nat.sub_mono_left_strict {x y z : ℕ} - (h' : z ≤ x) - (h : x < y) -: x - z < y - z := -begin - have : z ≤ y, - { transitivity, assumption, apply le_of_lt h, }, - apply @nat.lt_of_add_lt_add_left z, - rw [add_tsub_cancel_of_le,add_tsub_cancel_of_le]; - solve_by_elim -end - -@[mono] -lemma nat.sub_mono_right_strict {x y z : ℕ} - (h' : x ≤ z) - (h : y < x) -: z - x < z - y := -begin - have h'' : y ≤ z, - { transitivity, apply le_of_lt h, assumption }, - apply @nat.lt_of_add_lt_add_right _ x, - rw [tsub_add_cancel_of_le h'], - apply @lt_of_le_of_lt _ _ _ (z - y + y), - rw [tsub_add_cancel_of_le h''], - apply nat.add_lt_add_left h -end - open set attribute [mono] inter_subset_inter union_subset_union @@ -81,7 +22,10 @@ attribute [mono] upper_bounds_mono_set lower_bounds_mono_set attribute [mono] add_le_add mul_le_mul neg_le_neg mul_lt_mul_of_pos_left mul_lt_mul_of_pos_right + mul_le_mul_of_nonneg_left mul_le_mul_of_nonneg_right + mul_le_mul_of_nonpos_left mul_le_mul_of_nonpos_right imp_imp_imp le_implies_le_of_le_of_le + tsub_lt_tsub_left_of_le tsub_lt_tsub_right_of_le tsub_le_tsub abs_le_abs sup_le_sup inf_le_inf attribute [mono left] add_lt_add_of_le_of_lt mul_lt_mul' From 79a2eb397f4c4d2fcd365d8346ee2c23706286f1 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 23 Sep 2022 04:37:04 +0000 Subject: [PATCH 357/828] chore(.github/workflows): `maintainer merge` workflow (#16555) This PR adds a workflow that will listen for PR comments containing the phrase `maintainer merge` at the beginning of a line. It is intended for use by mathlib reviewers to flag to mathlib maintainers that they have reviewed a PR and consider it ready for merging. - The workflow will send a message to the maintainer stream on zulip to ping maintainers. - The workflow will add a comment confirming that the ping has been sent. - The workflow will only work for messages from members of the "mathlib reviewers" team in the leanprover-community organization. Co-authored-by: Bryan Gin-ge Chen --- .../workflows/maintainer_merge_comment.yml | 42 +++++++++++++++++++ .github/workflows/maintainer_merge_review.yml | 40 ++++++++++++++++++ .../maintainer_merge_review_comment.yml | 40 ++++++++++++++++++ 3 files changed, 122 insertions(+) create mode 100644 .github/workflows/maintainer_merge_comment.yml create mode 100644 .github/workflows/maintainer_merge_review.yml create mode 100644 .github/workflows/maintainer_merge_review_comment.yml diff --git a/.github/workflows/maintainer_merge_comment.yml b/.github/workflows/maintainer_merge_comment.yml new file mode 100644 index 0000000000000..b49bb861f2556 --- /dev/null +++ b/.github/workflows/maintainer_merge_comment.yml @@ -0,0 +1,42 @@ +name: Maintainer merge (comment) + +on: + issue_comment: + types: [created, edited] + +jobs: + ping_zulip: + name: Ping maintainers on Zulip + if: (github.event.issue.pull_request != 'null') && (startsWith(github.event.comment.body, 'maintainer merge') || contains(toJSON(github.event.comment.body), '\r\nmaintainer merge')) + runs-on: ubuntu-latest + steps: + - name: Check whether user is part of mathlib-reviewers team + uses: TheModdingInquisition/actions-team-membership@v1.0 + with: + organization: 'leanprover-community' + team: 'mathlib-reviewers' # required. The team to check for + token: ${{ secrets.MATHLIB_REVIEWERS_TEAM_KEY }} # required. Personal Access Token with the `read:org` permission + comment: 'You seem to not be authorized' # optional. A comment to post if the user is not part of the team. + # This feature is only applicable in an issue (or PR) context + exit: true # optional. If the action should exit if the user is not part of the team. Defaults to true. + + - name: Send message on Zulip + uses: zulip/github-actions-zulip/send-message@v1 + with: + api-key: ${{ secrets.ZULIP_API_KEY }} + email: 'github-bot@leanprover.zulipchat.com' + organization-url: 'https://leanprover.zulipchat.com' + to: 'mathlib reviewers' + type: 'stream' + topic: 'maintainer merge' + content: | + ${{ format('{0} requested a maintainer merge on PR #{1}:', github.event.comment.user.login, github.event.issue.number) }} + + > ${{ github.event.issue.title }} + + - name: Add comment to PR + uses: GrantBirki/comment@v2.0.1 + with: + issue-number: ${{ github.event.issue.number }} + body: | + 🚀 Pull request has been placed on the maintainer queue by ${{ github.event.comment.user.login }}. diff --git a/.github/workflows/maintainer_merge_review.yml b/.github/workflows/maintainer_merge_review.yml new file mode 100644 index 0000000000000..60e48de15ff36 --- /dev/null +++ b/.github/workflows/maintainer_merge_review.yml @@ -0,0 +1,40 @@ +name: Maintainer merge (review) + +on: + pull_request_review: + types: [submitted, edited] + +jobs: + ping_zulip: + name: Ping maintainers on Zulip + if: (github.event.issue.pull_request != 'null') && (startsWith(github.event.review.body, 'maintainer merge') || contains(toJSON(github.event.review.body), '\r\nmaintainer merge')) + runs-on: ubuntu-latest + steps: + - name: Check whether user is part of mathlib-reviewers team + uses: TheModdingInquisition/actions-team-membership@v1.0 + with: + organization: 'leanprover-community' + team: 'mathlib-reviewers' # required. The team to check for + token: ${{ secrets.MATHLIB_REVIEWERS_TEAM_KEY }} # required. Personal Access Token with the `read:org` permission + exit: true # optional. If the action should exit if the user is not part of the team. Defaults to true. + + - name: Send message on Zulip + uses: zulip/github-actions-zulip/send-message@v1 + with: + api-key: ${{ secrets.ZULIP_API_KEY }} + email: 'github-bot@leanprover.zulipchat.com' + organization-url: 'https://leanprover.zulipchat.com' + to: 'mathlib reviewers' + type: 'stream' + topic: 'maintainer merge' + content: | + ${{ format('{0} requested a maintainer merge on PR #{1}:', github.event.review.user.login, github.event.pull_request.number) }} + + > ${{ github.event.pull_request.title }} + + - name: Add comment to PR + uses: GrantBirki/comment@v2.0.1 + with: + issue-number: ${{ github.event.pull_request.number }} + body: | + 🚀 Pull request has been placed on the maintainer queue by ${{ github.event.review.user.login }}. diff --git a/.github/workflows/maintainer_merge_review_comment.yml b/.github/workflows/maintainer_merge_review_comment.yml new file mode 100644 index 0000000000000..d3cb0c7a7d0df --- /dev/null +++ b/.github/workflows/maintainer_merge_review_comment.yml @@ -0,0 +1,40 @@ +name: Maintainer merge (review comment) + +on: + pull_request_review_comment: + types: [created, edited] + +jobs: + ping_zulip: + name: Ping maintainers on Zulip + if: (github.event.issue.pull_request != 'null') && (startsWith(github.event.comment.body, 'maintainer merge') || contains(toJSON(github.event.comment.body), '\r\nmaintainer merge')) + runs-on: ubuntu-latest + steps: + - name: Check whether user is part of mathlib-reviewers team + uses: TheModdingInquisition/actions-team-membership@v1.0 + with: + organization: 'leanprover-community' + team: 'mathlib-reviewers' # required. The team to check for + token: ${{ secrets.MATHLIB_REVIEWERS_TEAM_KEY }} # required. Personal Access Token with the `read:org` permission + exit: true # optional. If the action should exit if the user is not part of the team. Defaults to true. + + - name: Send message on Zulip + uses: zulip/github-actions-zulip/send-message@v1 + with: + api-key: ${{ secrets.ZULIP_API_KEY }} + email: 'github-bot@leanprover.zulipchat.com' + organization-url: 'https://leanprover.zulipchat.com' + to: 'mathlib reviewers' + type: 'stream' + topic: 'maintainer merge' + content: | + ${{ format('{0} requested a maintainer merge on PR #{1}:', github.event.comment.user.login, github.event.pull_request.number) }} + + > ${{ github.event.pull_request.title }} + + - name: Add comment to PR + uses: GrantBirki/comment@v2.0.1 + with: + issue-number: ${{ github.event.pull_request.number }} + body: | + 🚀 Pull request has been placed on the maintainer queue by ${{ github.event.comment.user.login }}. From 397230527de2ee2e7e0f3b7a9d45b0d2ad1eca01 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Fri, 23 Sep 2022 04:37:05 +0000 Subject: [PATCH 358/828] chore(topology/basic): rename lemmas (#16598) `interior_eq_iff_open` -> `interior_eq_iff_is_open`, it's dual lemma `closure_eq_iff_is_closed` renamed in #4 `subset_interior_iff_open` -> `subset_interior_iff_is_open` `subset_interior_iff_subset_of_open` -> `is_open.subset_interior_iff`, it's dual lemma `is_closed.closure_subset_iff` renamed in #3251 `closure_inter_open` -> `is_open.closure_inter` `closure_inter_open'` -> `is_open.closure_inter'` --- src/analysis/calculus/extend_deriv.lean | 2 +- src/analysis/convex/strict.lean | 2 +- src/topology/basic.lean | 34 +++++++++---------- src/topology/locally_finite.lean | 2 +- src/topology/maps.lean | 2 +- .../metric_space/hausdorff_distance.lean | 2 +- src/topology/nhds_set.lean | 2 +- src/topology/sets/opens.lean | 2 +- src/topology/subset_properties.lean | 2 +- src/topology/uniform_space/basic.lean | 2 +- 10 files changed, 26 insertions(+), 26 deletions(-) diff --git a/src/analysis/calculus/extend_deriv.lean b/src/analysis/calculus/extend_deriv.lean index 73f30665aba7b..469273820c714 100644 --- a/src/analysis/calculus/extend_deriv.lean +++ b/src/analysis/calculus/extend_deriv.lean @@ -58,7 +58,7 @@ begin { rw closure_prod_eq at this, intros y y_in, apply this ⟨x, y⟩, - have : B ∩ closure s ⊆ closure (B ∩ s), from closure_inter_open is_open_ball, + have : B ∩ closure s ⊆ closure (B ∩ s), from is_open_ball.closure_inter, exact ⟨this ⟨mem_ball_self δ_pos, hx⟩, this y_in⟩ }, have key : ∀ p : E × E, p ∈ (B ∩ s) ×ˢ (B ∩ s) → ∥f p.2 - f p.1 - (f' p.2 - f' p.1)∥ ≤ ε * ∥p.2 - p.1∥, diff --git a/src/analysis/convex/strict.lean b/src/analysis/convex/strict.lean index bd63c97e0952d..6d9659453d38e 100644 --- a/src/analysis/convex/strict.lean +++ b/src/analysis/convex/strict.lean @@ -144,7 +144,7 @@ variables [topological_space β] [linear_ordered_cancel_add_comm_monoid β] [ord lemma strict_convex_Iic (r : β) : strict_convex 𝕜 (Iic r) := begin rintro x (hx : x ≤ r) y (hy : y ≤ r) hxy a b ha hb hab, - refine (subset_interior_iff_subset_of_open is_open_Iio).2 Iio_subset_Iic_self _, + refine is_open_Iio.subset_interior_iff.2 Iio_subset_Iic_self _, rw ←convex.combo_self hab r, obtain rfl | hx := hx.eq_or_lt, { exact add_lt_add_left (smul_lt_smul_of_pos (hy.lt_of_ne hxy.symm) hb) _ }, diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 8e23d770d3e54..91a2f6f769f51 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -248,13 +248,13 @@ subset_sUnion_of_mem ⟨h₂, h₁⟩ lemma is_open.interior_eq {s : set α} (h : is_open s) : interior s = s := subset.antisymm interior_subset (interior_maximal (subset.refl s) h) -lemma interior_eq_iff_open {s : set α} : interior s = s ↔ is_open s := +lemma interior_eq_iff_is_open {s : set α} : interior s = s ↔ is_open s := ⟨assume h, h ▸ is_open_interior, is_open.interior_eq⟩ -lemma subset_interior_iff_open {s : set α} : s ⊆ interior s ↔ is_open s := -by simp only [interior_eq_iff_open.symm, subset.antisymm_iff, interior_subset, true_and] +lemma subset_interior_iff_is_open {s : set α} : s ⊆ interior s ↔ is_open s := +by simp only [interior_eq_iff_is_open.symm, subset.antisymm_iff, interior_subset, true_and] -lemma subset_interior_iff_subset_of_open {s t : set α} (h₁ : is_open s) : +lemma is_open.subset_interior_iff {s t : set α} (h₁ : is_open s) : s ⊆ interior t ↔ s ⊆ t := ⟨assume h, subset.trans h interior_subset, assume h₂, interior_maximal h₂ h₁⟩ @@ -305,7 +305,7 @@ have interior (s ∪ t) ⊆ s, from have u \ s ⊆ t, from assume x ⟨h₁, h₂⟩, or.resolve_left (hu₂ h₁) h₂, have u \ s ⊆ interior t, - by rwa subset_interior_iff_subset_of_open (is_open.sdiff hu₁ h₁), + by rwa (is_open.sdiff hu₁ h₁).subset_interior_iff, have u \ s ⊆ ∅, by rwa h₂ at this, this ⟨hx₁, hx₂⟩, @@ -314,7 +314,7 @@ subset.antisymm (interior_mono $ subset_union_left _ _) lemma is_open_iff_forall_mem_open : is_open s ↔ ∀ x ∈ s, ∃ t ⊆ s, is_open t ∧ x ∈ t := -by rw ← subset_interior_iff_open; simp only [subset_def, mem_interior] +by rw ← subset_interior_iff_is_open; simp only [subset_def, mem_interior] lemma interior_Inter_subset (s : ι → set α) : interior (⋂ i, s i) ⊆ ⋂ i, interior (s i) := subset_Inter $ λ i, interior_mono $ Inter_subset _ _ @@ -364,9 +364,9 @@ lemma is_closed.closure_subset_iff {s t : set α} (h₁ : is_closed t) : closure s ⊆ t ↔ s ⊆ t := ⟨subset.trans subset_closure, assume h, closure_minimal h h₁⟩ -lemma is_closed.mem_iff_closure_subset {α : Type*} [topological_space α] {U : set α} - (hU : is_closed U) {x : α} : x ∈ U ↔ closure ({x} : set α) ⊆ U := -(hU.closure_subset_iff.trans set.singleton_subset_iff).symm +lemma is_closed.mem_iff_closure_subset {s : set α} (hs : is_closed s) {x : α} : + x ∈ s ↔ closure ({x} : set α) ⊆ s := +(hs.closure_subset_iff.trans set.singleton_subset_iff).symm @[mono] lemma closure_mono {s t : set α} (h : s ⊆ t) : closure s ⊆ closure t := closure_minimal (subset.trans h subset_closure) is_closed_closure @@ -959,7 +959,7 @@ lemma subset_interior_iff_nhds {s V : set α} : s ⊆ interior V ↔ ∀ x ∈ s show (∀ x, x ∈ s → x ∈ _) ↔ _, by simp_rw mem_interior_iff_mem_nhds lemma is_open_iff_nhds {s : set α} : is_open s ↔ ∀a∈s, 𝓝 a ≤ 𝓟 s := -calc is_open s ↔ s ⊆ interior s : subset_interior_iff_open.symm +calc is_open s ↔ s ⊆ interior s : subset_interior_iff_is_open.symm ... ↔ (∀a∈s, 𝓝 a ≤ 𝓟 s) : by rw [interior_eq_nhds]; refl lemma is_open_iff_mem_nhds {s : set α} : is_open s ↔ ∀a∈s, s ∈ 𝓝 a := @@ -973,7 +973,7 @@ theorem is_open_iff_ultrafilter {s : set α} : is_open s ↔ (∀ (x ∈ s) (l : ultrafilter α), ↑l ≤ 𝓝 x → s ∈ l) := by simp_rw [is_open_iff_mem_nhds, ← mem_iff_ultrafilter] -lemma is_open_singleton_iff_nhds_eq_pure {α : Type*} [topological_space α] (a : α) : +lemma is_open_singleton_iff_nhds_eq_pure (a : α) : is_open ({a} : set α) ↔ 𝓝 a = pure a := begin split, @@ -1082,7 +1082,7 @@ calc is_closed s ↔ closure s ⊆ s : closure_subset_iff_is_closed.symm lemma is_closed_iff_nhds {s : set α} : is_closed s ↔ ∀ x, (∀ U ∈ 𝓝 x, (U ∩ s).nonempty) → x ∈ s := by simp_rw [is_closed_iff_cluster_pt, cluster_pt, inf_principal_ne_bot_iff] -lemma closure_inter_open {s t : set α} (h : is_open s) : s ∩ closure t ⊆ closure (s ∩ t) := +lemma is_open.closure_inter {s t : set α} (h : is_open s) : s ∩ closure t ⊆ closure (s ∩ t) := begin rintro a ⟨hs, ht⟩, have : s ∈ 𝓝 a := is_open.mem_nhds h hs, @@ -1090,13 +1090,13 @@ begin rwa [← inf_principal, ← inf_assoc, inf_eq_left.2 (le_principal_iff.2 this)], end -lemma closure_inter_open' {s t : set α} (h : is_open t) : closure s ∩ t ⊆ closure (s ∩ t) := -by simpa only [inter_comm] using closure_inter_open h +lemma is_open.closure_inter' {s t : set α} (h : is_open t) : closure s ∩ t ⊆ closure (s ∩ t) := +by simpa only [inter_comm] using h.closure_inter lemma dense.open_subset_closure_inter {s t : set α} (hs : dense s) (ht : is_open t) : t ⊆ closure (t ∩ s) := calc t = t ∩ closure s : by rw [hs.closure_eq, inter_univ] - ... ⊆ closure (t ∩ s) : closure_inter_open ht + ... ⊆ closure (t ∩ s) : ht.closure_inter lemma mem_closure_of_mem_closure_union {s₁ s₂ : set α} {x : α} (h : x ∈ closure (s₁ ∪ s₂)) (h₁ : s₁ᶜ ∈ 𝓝 x) : x ∈ closure s₂ := @@ -1112,7 +1112,7 @@ end /-- The intersection of an open dense set with a dense set is a dense set. -/ lemma dense.inter_of_open_left {s t : set α} (hs : dense s) (ht : dense t) (hso : is_open s) : dense (s ∩ t) := -λ x, (closure_minimal (closure_inter_open hso) is_closed_closure) $ +λ x, (closure_minimal hso.closure_inter is_closed_closure) $ by simp [hs.closure_eq, ht.closure_eq] /-- The intersection of a dense set with an open dense set is a dense set. -/ @@ -1127,7 +1127,7 @@ let ⟨U, hsub, ho, hx⟩ := mem_nhds_iff.1 ht in lemma closure_diff {s t : set α} : closure s \ closure t ⊆ closure (s \ t) := calc closure s \ closure t = (closure t)ᶜ ∩ closure s : by simp only [diff_eq, inter_comm] - ... ⊆ closure ((closure t)ᶜ ∩ s) : closure_inter_open $ is_open_compl_iff.mpr $ is_closed_closure + ... ⊆ closure ((closure t)ᶜ ∩ s) : (is_open_compl_iff.mpr $ is_closed_closure).closure_inter ... = closure (s \ closure t) : by simp only [diff_eq, inter_comm] ... ⊆ closure (s \ t) : closure_mono $ diff_subset_diff (subset.refl s) subset_closure diff --git a/src/topology/locally_finite.lean b/src/topology/locally_finite.lean index 58eaf95870161..ded3c28916569 100644 --- a/src/topology/locally_finite.lean +++ b/src/topology/locally_finite.lean @@ -78,7 +78,7 @@ begin intro x, rcases hf x with ⟨s, hsx, hsf⟩, refine ⟨interior s, interior_mem_nhds.2 hsx, hsf.subset $ λ i hi, _⟩, - exact (hi.mono (closure_inter_open' is_open_interior)).of_closure.mono + exact (hi.mono is_open_interior.closure_inter').of_closure.mono (inter_subset_inter_right _ interior_subset) end diff --git a/src/topology/maps.lean b/src/topology/maps.lean index 45c4b0b3318ef..a7252ab91a5ef 100644 --- a/src/topology/maps.lean +++ b/src/topology/maps.lean @@ -365,7 +365,7 @@ lemma is_open_map_iff_nhds_le [topological_space α] [topological_space β] {f : lemma is_open_map_iff_interior [topological_space α] [topological_space β] {f : α → β} : is_open_map f ↔ ∀ s, f '' (interior s) ⊆ interior (f '' s) := -⟨is_open_map.image_interior_subset, λ hs u hu, subset_interior_iff_open.mp $ +⟨is_open_map.image_interior_subset, λ hs u hu, subset_interior_iff_is_open.mp $ calc f '' u = f '' (interior u) : by rw hu.interior_eq ... ⊆ interior (f '' u) : hs u⟩ diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index 1e59651bbb2e9..a5d2c5fda17de 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -990,7 +990,7 @@ end lemma thickening_subset_interior_cthickening (δ : ℝ) (E : set α) : thickening δ E ⊆ interior (cthickening δ E) := -(subset_interior_iff_open.mpr (is_open_thickening)).trans +(subset_interior_iff_is_open.mpr (is_open_thickening)).trans (interior_mono (thickening_subset_cthickening δ E)) lemma closure_thickening_subset_cthickening (δ : ℝ) (E : set α) : diff --git a/src/topology/nhds_set.lean b/src/topology/nhds_set.lean index c50a4a94bb142..d6e72c0bad38d 100644 --- a/src/topology/nhds_set.lean +++ b/src/topology/nhds_set.lean @@ -54,7 +54,7 @@ lemma has_basis_nhds_set (s : set α) : (𝓝ˢ s).has_basis (λ U, is_open U ⟨λ t, by simp [mem_nhds_set_iff_exists, and_assoc]⟩ lemma is_open.mem_nhds_set (hU : is_open s) : s ∈ 𝓝ˢ t ↔ t ⊆ s := -by rw [← subset_interior_iff_mem_nhds_set, interior_eq_iff_open.mpr hU] +by rw [← subset_interior_iff_mem_nhds_set, interior_eq_iff_is_open.mpr hU] lemma principal_le_nhds_set : 𝓟 s ≤ 𝓝ˢ s := λ s hs, (subset_interior_iff_mem_nhds_set.mpr hs).trans interior_subset diff --git a/src/topology/sets/opens.lean b/src/topology/sets/opens.lean index 0a8fd73565891..63e011e115239 100644 --- a/src/topology/sets/opens.lean +++ b/src/topology/sets/opens.lean @@ -69,7 +69,7 @@ open order_dual (of_dual to_dual) /-- The galois coinsertion between sets and opens. -/ def gi : galois_coinsertion subtype.val (@interior α _) := -{ choice := λ s hs, ⟨s, interior_eq_iff_open.mp $ le_antisymm interior_subset hs⟩, +{ choice := λ s hs, ⟨s, interior_eq_iff_is_open.mp $ le_antisymm interior_subset hs⟩, gc := gc, u_l_le := λ _, interior_subset, choice_eq := λ s hs, le_antisymm hs interior_subset } diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index c2d43ef7b82bb..b8a1f62f23125 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -1390,7 +1390,7 @@ protected lemma is_clopen.is_closed (hs : is_clopen s) : is_closed s := hs.2 lemma is_clopen_iff_frontier_eq_empty {s : set α} : is_clopen s ↔ frontier s = ∅ := begin - rw [is_clopen, ← closure_eq_iff_is_closed, ← interior_eq_iff_open, frontier, diff_eq_empty], + rw [is_clopen, ← closure_eq_iff_is_closed, ← interior_eq_iff_is_open, frontier, diff_eq_empty], refine ⟨λ h, (h.2.trans h.1.symm).subset, λ h, _⟩, exact ⟨interior_subset.antisymm (subset_closure.trans h), (h.trans interior_subset).antisymm subset_closure⟩ diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index b95814c22b205..696fdecd164a2 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -848,7 +848,7 @@ le_antisymm let ⟨t, ht, hst, ht_comp⟩ := nhdset_of_mem_uniformity s hs in have s ⊆ interior d, from calc s ⊆ t : hst - ... ⊆ interior d : (subset_interior_iff_subset_of_open ht).mpr $ + ... ⊆ interior d : ht.subset_interior_iff.mpr $ λ x (hx : x ∈ t), let ⟨x, y, h₁, h₂, h₃⟩ := ht_comp hx in hs_comp ⟨x, h₁, y, h₂, h₃⟩, have interior d ∈ 𝓤 α, by filter_upwards [hs] using this, by simp [this]) From 028afe5d4f799a85883fe387105e4f2ba7a9fda3 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 23 Sep 2022 04:37:06 +0000 Subject: [PATCH 359/828] chore(category/endomorphism): remove unneeded ext attribute (#16603) Removes an unnecessary `@[ext]` attribute that used an extended syntax which we are not planning to port to mathlib4. Closes #16602. Co-authored-by: Scott Morrison --- src/category_theory/endomorphism.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/category_theory/endomorphism.lean b/src/category_theory/endomorphism.lean index 4bd5d4c486229..55057caac0b4c 100644 --- a/src/category_theory/endomorphism.lean +++ b/src/category_theory/endomorphism.lean @@ -99,8 +99,6 @@ The order of arguments in multiplication agrees with -/ def Aut (X : C) := X ≅ X -attribute [ext Aut] iso.ext - namespace Aut instance inhabited : inhabited (Aut X) := ⟨iso.refl X⟩ From 53322d0d876324ceabe21b1187a1de5799a6d7ac Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 23 Sep 2022 08:22:32 +0000 Subject: [PATCH 360/828] feat(analysis/special_functions/pow, set_theory/*): `positivity` extension for powers (#16462) Add the following `positivity` extensions to handle powers: * `positivity_rpow` for real powers * `positivity_opow` for ordinal powers * `positivity_cardinal_pow` for cardinal powers --- src/analysis/special_functions/pow.lean | 56 +++++++++++++++++++++++++ src/set_theory/cardinal/basic.lean | 19 +++++++++ src/set_theory/ordinal/arithmetic.lean | 17 ++++++++ src/tactic/positivity.lean | 51 +++++++++++++++------- test/positivity.lean | 27 +++++++++++- 5 files changed, 152 insertions(+), 18 deletions(-) diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 4ee23bb0f116e..2302dbe222696 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -2103,3 +2103,59 @@ prove_rpow' ``ennrpow_pos ``ennrpow_neg ``ennreal.rpow_zero `(ℝ≥0∞) `(ℝ) | _ := tactic.failed end norm_num + +namespace tactic +namespace positivity + +/-- Auxiliary definition for the `positivity` tactic to handle real powers of reals. -/ +meta def prove_rpow (a b : expr) : tactic strictness := +do + strictness_a ← core a, + match strictness_a with + | nonnegative p := nonnegative <$> mk_app ``real.rpow_nonneg_of_nonneg [p, b] + | positive p := positive <$> mk_app ``real.rpow_pos_of_pos [p, b] + end + +private lemma nnrpow_pos {a : ℝ≥0} (ha : 0 < a) (b : ℝ) : 0 < a ^ b := nnreal.rpow_pos ha + +/-- Auxiliary definition for the `positivity` tactic to handle real powers of nonnegative reals. -/ +meta def prove_nnrpow (a b : expr) : tactic strictness := +do + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``nnrpow_pos [p, b] + | _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0` + end + +private lemma ennrpow_pos {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 < b) : 0 < a ^ b := +ennreal.rpow_pos_of_nonneg ha hb.le + +/-- Auxiliary definition for the `positivity` tactic to handle real powers of extended nonnegative +reals. -/ +meta def prove_ennrpow (a b : expr) : tactic strictness := +do + strictness_a ← core a, + strictness_b ← core b, + match strictness_a, strictness_b with + | positive pa, positive pb := positive <$> mk_app ``ennrpow_pos [pa, pb] + | positive pa, nonnegative pb := positive <$> mk_app ``ennreal.rpow_pos_of_nonneg [pa, pb] + | _, _ := failed -- We already know `0 ≤ x` for all `x : ℝ≥0∞` + end + +end positivity + +open positivity + +/-- Extension for the `positivity` tactic: exponentiation by a real number is nonnegative when the +base is nonnegative and positive when the base is positive. -/ +@[positivity] +meta def positivity_rpow : expr → tactic strictness +| `(@has_pow.pow _ _ real.has_pow %%a %%b) := prove_rpow a b +| `(real.rpow %%a %%b) := prove_rpow a b +| `(@has_pow.pow _ _ nnreal.real.has_pow %%a %%b) := prove_nnrpow a b +| `(nnreal.rpow %%a %%b) := prove_nnrpow a b +| `(@has_pow.pow _ _ ennreal.real.has_pow %%a %%b) := prove_ennrpow a b +| `(ennreal.rpow %%a %%b) := prove_ennrpow a b +| _ := failed + +end tactic diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index e7b7eefa30b7c..6585a0ed2a0f7 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -11,6 +11,7 @@ import logic.small import order.conditionally_complete_lattice import order.succ_pred.basic import set_theory.cardinal.schroeder_bernstein +import tactic.positivity /-! # Cardinal Numbers @@ -518,6 +519,8 @@ end theorem power_le_power_right {a b c : cardinal} : a ≤ b → a ^ c ≤ b ^ c := induction_on₃ a b c $ λ α β γ ⟨e⟩, ⟨embedding.arrow_congr_right e⟩ +theorem power_pos {a : cardinal} (b) (ha : 0 < a) : 0 < a ^ b := (power_ne_zero _ ha.ne').bot_lt + end order_properties protected theorem lt_wf : @well_founded cardinal.{u} (<) := @@ -1580,3 +1583,19 @@ begin end end cardinal + +namespace tactic +open cardinal positivity + +/-- Extension for the `positivity` tactic: The cardinal power of a positive cardinal is positive. -/ +@[positivity] +meta def positivity_cardinal_pow : expr → tactic strictness +| `(@has_pow.pow _ _ %%inst %%a %%b) := do + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``power_pos [b, p] + | _ := failed -- We already know that `0 ≤ x` for all `x : cardinal` + end +| _ := failed + +end tactic diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 4a49974afe067..cde5a7d017013 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -2363,3 +2363,20 @@ begin end end ordinal + +namespace tactic +open ordinal positivity + +/-- Extension for the `positivity` tactic: `ordinal.opow` takes positive values on positive inputs. +-/ +@[positivity] +meta def positivity_opow : expr → tactic strictness +| `(@has_pow.pow _ _ %%inst %%a %%b) := do + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``opow_pos [b, p] + | _ := failed -- We already know that `0 ≤ x` for all `x : ordinal` + end +| _ := failed + +end tactic diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index 7bc39befbdf9e..1827ade0ca6de 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -358,29 +358,48 @@ meta def positivity_inv : expr → tactic strictness private lemma pow_zero_pos [ordered_semiring R] [nontrivial R] (a : R) : 0 < a ^ 0 := zero_lt_one.trans_le (pow_zero a).ge -/-- Extension for the `positivity` tactic: raising a number `a` to a natural number power `n` is -known to be positive if `n = 0` (since `a ^ 0 = 1`) or if `0 < a`, and is known to be nonnegative if -`n = 2` (squares are nonnegative) or if `0 ≤ a`. -/ +private lemma zpow_zero_pos [linear_ordered_semifield R] (a : R) : 0 < a ^ (0 : ℤ) := +zero_lt_one.trans_le (zpow_zero a).ge + +/-- Extension for the `positivity` tactic: raising a number `a` to a natural/integer power `n` is +positive if `n = 0` (since `a ^ 0 = 1`) or if `0 < a`, and is nonnegative if `n` is even (squares +are nonnegative) or if `0 ≤ a`. -/ @[positivity] meta def positivity_pow : expr → tactic strictness | `(%%a ^ %%n) := do - n_typ ← infer_type n, - match n_typ with - | `(ℕ) := do + typ ← infer_type n, + (do + unify typ `(ℕ), if n = `(0) then positive <$> mk_app ``pow_zero_pos [a] - else tactic.positivity.orelse' - -- squares are nonnegative (TODO: similar for any `bit0` exponent?) - (nonnegative <$> mk_app ``sq_nonneg [a]) - -- moreover `a ^ n` is positive if `a` is and nonnegative if `a` is - (do + else positivity.orelse' + (do -- even powers are nonnegative + match n with -- TODO: Decision procedure for parity + | `(bit0 %% n) := nonnegative <$> mk_app ``pow_bit0_nonneg [a, n] + | _ := failed + end) $ + do -- `a ^ n` is positive if `a` is, and nonnegative if `a` is strictness_a ← core a, match strictness_a with - | (positive pa) := positive <$> mk_app ``pow_pos [pa, n] - | (nonnegative pa) := nonnegative <$> mk_app ``pow_nonneg [pa, n] - end) - | _ := failed -- TODO handle integer powers, maybe even real powers - end + | positive p := positive <$> mk_app ``pow_pos [p, n] + | nonnegative p := nonnegative <$> mk_app `pow_nonneg [p, n] + end) <|> + (do + unify typ `(ℤ), + if n = `(0 : ℤ) then + positive <$> mk_app ``zpow_zero_pos [a] + else positivity.orelse' + (do -- even powers are nonnegative + match n with -- TODO: Decision procedure for parity + | `(bit0 %% n) := nonnegative <$> mk_app ``zpow_bit0_nonneg [a, n] + | _ := failed + end) $ + do -- `a ^ n` is positive if `a` is, and nonnegative if `a` is + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``zpow_pos_of_pos [p, n] + | nonnegative p := nonnegative <$> mk_app ``zpow_nonneg [p, n] + end) | _ := failed /-- Extension for the `positivity` tactic: an absolute value is nonnegative, and is strictly diff --git a/test/positivity.lean b/test/positivity.lean index 804b2de9adab0..2595bcb6f1f73 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -1,5 +1,6 @@ import algebra.order.smul import analysis.normed.group.basic +import analysis.special_functions.pow import data.complex.exponential import data.rat.nnrat import data.real.ereal @@ -14,6 +15,9 @@ This tactic proves goals of the form `0 ≤ a` and `0 < a`. open_locale ennreal nnrat nnreal +universe u +variables {α β : Type*} + /- ## Numeric goals -/ example : 0 ≤ 0 := by positivity @@ -56,7 +60,26 @@ example {a b : ℤ} (ha : 3 < a) (hb : 4 ≤ b) : 0 < 3 + a * b / 7 + b + 7 + 14 example {a : ℤ} (ha : 0 < a) : 0 < a / a := by positivity -example {a : ℕ} : 0 < a ^ 0 := by positivity +/-! ### Exponentiation -/ + +example [ordered_semiring α] [nontrivial α] (a : α) : 0 < a ^ 0 := by positivity +example [linear_ordered_ring α] (a : α) (n : ℕ) : 0 ≤ a ^ (bit0 n) := by positivity +example [ordered_semiring α] {a : α} {n : ℕ} (ha : 0 ≤ a) : 0 ≤ a ^ n := by positivity +example [ordered_semiring α] {a : α} {n : ℕ} (ha : 0 < a) : 0 < a ^ n := by positivity + +example [linear_ordered_semifield α] (a : α) : 0 < a ^ (0 : ℤ) := by positivity +example [linear_ordered_field α] (a : α) (n : ℤ) : 0 ≤ a ^ (bit0 n) := by positivity +example [linear_ordered_semifield α] {a : α} {n : ℤ} (ha : 0 ≤ a) : 0 ≤ a ^ n := by positivity +example [linear_ordered_semifield α] {a : α} {n : ℤ} (ha : 0 < a) : 0 < a ^ n := by positivity + +example {a b : cardinal.{u}} (ha : 0 < a) : 0 < a ^ b := by positivity +example {a b : ordinal.{u}} (ha : 0 < a) : 0 < a ^ b := by positivity + +example {a b : ℝ} (ha : 0 ≤ a) : 0 ≤ a ^ b := by positivity +example {a b : ℝ} (ha : 0 < a) : 0 < a ^ b := by positivity +example {a : ℝ≥0} {b : ℝ} (ha : 0 < a) : 0 < a ^ b := by positivity +example {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 ≤ b) : 0 < a ^ b := by positivity +example {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 < b) : 0 < a ^ b := by positivity example {a : ℤ} (ha : 3 < a) : 0 ≤ a ^ 2 + a := by positivity @@ -99,7 +122,7 @@ example : 0 ≤ max (0:ℤ) (-3) := by positivity example : 0 ≤ max (-3 : ℤ) 5 := by positivity -example {α β : Type*} [ordered_semiring α] [ordered_add_comm_monoid β] [smul_with_zero α β] +example [ordered_semiring α] [ordered_add_comm_monoid β] [smul_with_zero α β] [ordered_smul α β] {a : α} (ha : 0 < a) {b : β} (hb : 0 < b) : 0 ≤ a • b := by positivity example {r : ℝ} : 0 < real.exp r := by positivity From 02022cd63056d424b2859a2c185985d81dca6e33 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 23 Sep 2022 08:22:33 +0000 Subject: [PATCH 361/828] feat(topology/subset_properties): an infinite type with cofinite topology is irreducible (#16499) --- src/data/set/basic.lean | 3 +++ src/topology/separation.lean | 12 ++++++++++-- src/topology/sober.lean | 2 +- src/topology/subset_properties.lean | 18 +++++++++++++++++- 4 files changed, 31 insertions(+), 4 deletions(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 75d7a3d5e7bd6..93b1761e794dc 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1914,6 +1914,9 @@ by simp_rw [set.subsingleton, set.nontrivial, not_forall] @[simp] lemma not_nontrivial_iff : ¬ s.nontrivial ↔ s.subsingleton := iff.not_left not_subsingleton_iff.symm +alias not_nontrivial_iff ↔ _ subsingleton.not_nontrivial +alias not_subsingleton_iff ↔ _ nontrivial.not_subsingleton + theorem univ_eq_true_false : univ = ({true, false} : set Prop) := eq.symm $ eq_univ_of_forall $ classical.cases (by simp) (by simp) diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 0efba82ace1d9..573ae82e343be 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -1269,7 +1269,7 @@ begin compact_closure_of_subset_compact hV interior_subset⟩, end -lemma is_preirreducible_iff_subsingleton [t2_space α] (S : set α) : +lemma is_preirreducible_iff_subsingleton [t2_space α] {S : set α} : is_preirreducible S ↔ S.subsingleton := begin refine ⟨λ h x hx y hy, _, set.subsingleton.is_preirreducible⟩, @@ -1278,11 +1278,19 @@ begin exact ((h U V hU hV ⟨x, hx, hxU⟩ ⟨y, hy, hyV⟩).mono $ inter_subset_right _ _).not_disjoint h', end -lemma is_irreducible_iff_singleton [t2_space α] (S : set α) : +alias is_preirreducible_iff_subsingleton ↔ is_preirreducible.subsingleton _ +attribute [protected] is_preirreducible.subsingleton + +lemma is_irreducible_iff_singleton [t2_space α] {S : set α} : is_irreducible S ↔ ∃ x, S = {x} := by rw [is_irreducible, is_preirreducible_iff_subsingleton, exists_eq_singleton_iff_nonempty_subsingleton] +/-- There does not exist a nontrivial preirreducible T₂ space. -/ +lemma not_preirreducible_nontrivial_t2 (α) [topological_space α] [preirreducible_space α] + [nontrivial α] [t2_space α] : false := +(preirreducible_space.is_preirreducible_univ α).subsingleton.not_nontrivial nontrivial_univ + end separation section regular_space diff --git a/src/topology/sober.lean b/src/topology/sober.lean index aedd885c1f831..7f052ca9fca15 100644 --- a/src/topology/sober.lean +++ b/src/topology/sober.lean @@ -223,7 +223,7 @@ instance t2_space.quasi_sober [t2_space α] : quasi_sober α := begin constructor, rintro S h -, - obtain ⟨x, rfl⟩ := (is_irreducible_iff_singleton S).mp h, + obtain ⟨x, rfl⟩ := is_irreducible_iff_singleton.mp h, exact ⟨x, closure_singleton⟩ end diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index b8a1f62f23125..d13ee6737558d 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -1590,6 +1590,11 @@ theorem nonempty_preirreducible_inter [preirreducible_space α] {s t : set α} : by simpa only [univ_inter, univ_subset_iff] using @preirreducible_space.is_preirreducible_univ α _ _ s t +/-- In a (pre)irreducible space, a nonempty open set is dense. -/ +protected theorem is_open.dense [preirreducible_space α] {s : set α} (ho : is_open s) + (hne : s.nonempty) : dense s := +dense_iff_inter_open.2 $ λ t hto htne, nonempty_preirreducible_inter hto ho htne hne + theorem is_preirreducible.image {s : set α} (H : is_preirreducible s) (f : α → β) (hf : continuous_on f s) : is_preirreducible (f '' s) := begin @@ -1611,7 +1616,7 @@ end theorem is_irreducible.image {s : set α} (H : is_irreducible s) (f : α → β) (hf : continuous_on f s) : is_irreducible (f '' s) := -⟨nonempty_image_iff.mpr H.nonempty, H.is_preirreducible.image f hf⟩ +⟨H.nonempty.image _, H.is_preirreducible.image f hf⟩ lemma subtype.preirreducible_space {s : set α} (h : is_preirreducible s) : preirreducible_space s := @@ -1633,6 +1638,17 @@ lemma subtype.irreducible_space {s : set α} (h : is_irreducible s) : (subtype.preirreducible_space h.is_preirreducible).is_preirreducible_univ, to_nonempty := h.nonempty.to_subtype } +/-- An infinite type with cofinite topology is an irreducible topological space. -/ +@[priority 100] instance {α} [infinite α] : irreducible_space (cofinite_topology α) := +{ is_preirreducible_univ := λ u v, + begin + haveI : infinite (cofinite_topology α) := ‹_›, + simp only [cofinite_topology.is_open_iff, univ_inter], + intros hu hv hu' hv', + simpa only [compl_union, compl_compl] using ((hu hu').union (hv hv')).infinite_compl.nonempty + end, + to_nonempty := (infer_instance : nonempty α) } + /-- A set `s` is irreducible if and only if for every finite collection of open sets all of whose members intersect `s`, `s` also intersects the intersection of the entire collection From f89fa08bb6455c96c58e860902fadc8eb2c854ed Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 23 Sep 2022 08:22:35 +0000 Subject: [PATCH 362/828] chore(tactic): typos in docs (#16510) Co-authored-by: Scott Morrison --- src/tactic/core.lean | 4 ++-- src/tactic/interactive.lean | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/src/tactic/core.lean b/src/tactic/core.lean index 7524e2a649cc3..ddf187cdf0f64 100644 --- a/src/tactic/core.lean +++ b/src/tactic/core.lean @@ -996,7 +996,7 @@ end /-- `apply_list l`, for `l : list (tactic expr)`, -tries to apply the lemmas generated by the tactics in `l` on the first goal, and +tries to apply one of the lemmas generated by the tactics in `l` to the first goal, and fail if none succeeds. -/ meta def apply_list_expr (opt : apply_cfg) : list (tactic expr) → tactic unit @@ -1013,7 +1013,7 @@ meta def resolve_attribute_expr_list (attr_name : name) : tactic (list (tactic e /--`apply_rules args attrs n`: apply the lists of rules `args` (given as pexprs) and `attrs` (given -as names of attributes) and `the tactic assumption` on the first goal and the resulting subgoals, +as names of attributes) and the tactic `assumption` on the first goal and the resulting subgoals, iteratively, at most `n` times. Unlike `solve_by_elim`, `apply_rules` does not do any backtracking, and just greedily applies diff --git a/src/tactic/interactive.lean b/src/tactic/interactive.lean index ea73a5b041564..462543ebd069c 100644 --- a/src/tactic/interactive.lean +++ b/src/tactic/interactive.lean @@ -515,7 +515,7 @@ attribute from the list `attrs`, as well as the `assumption` tactic on the first goal and the resulting subgoals, iteratively, at most `n` times. `n` is optional, equal to 50 by default. You can pass an `apply_cfg` option argument as `apply_rules hs n opt`. -(A typical usage would be with `apply_rules hs n { md := reducible })`, +(A typical usage would be with `apply_rules hs n { md := reducible }`, which asks `apply_rules` to not unfold `semireducible` definitions (i.e. most) when checking if a lemma matches the goal.) From d0507e7be36876e9e2c84a710de50346ef187fd4 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Fri, 23 Sep 2022 08:22:36 +0000 Subject: [PATCH 363/828] =?UTF-8?q?feat(analysis/normed=5Fspace/affine=5Fi?= =?UTF-8?q?sometry):=20`subtype=E2=82=90=E1=B5=A2`=20(#16573)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Now that we have `affine_subspace.subtype`, add the corresponding definition as a bundled affine isometry (so resolving a comment about that being missing). The name uses subscript `ₐᵢ` rather than superscript `ᵃⁱ` as in the notation for `affine_isometry` because Lean does not accept the superscript in identifiers. --- .../normed_space/affine_isometry.lean | 28 +++++++++++++++++-- 1 file changed, 25 insertions(+), 3 deletions(-) diff --git a/src/analysis/normed_space/affine_isometry.lean b/src/analysis/normed_space/affine_isometry.lean index 269c09455175d..838114494fa23 100644 --- a/src/analysis/normed_space/affine_isometry.lean +++ b/src/analysis/normed_space/affine_isometry.lean @@ -197,9 +197,31 @@ instance : monoid (P →ᵃⁱ[𝕜] P) := end affine_isometry --- remark: by analogy with the `linear_isometry` file from which this is adapted, there should --- follow here a section defining an "inclusion" affine isometry from `p : affine_subspace 𝕜 P` --- into `P`; we omit this for now +namespace affine_subspace + +include V + +/-- `affine_subspace.subtype` as an `affine_isometry`. -/ +def subtypeₐᵢ (s : affine_subspace 𝕜 P) [nonempty s] : s →ᵃⁱ[𝕜] P := +{ norm_map := s.direction.subtypeₗᵢ.norm_map, + .. s.subtype } + +lemma subtypeₐᵢ_linear (s : affine_subspace 𝕜 P) [nonempty s] : + s.subtypeₐᵢ.linear = s.direction.subtype := +rfl + +@[simp] lemma subtypeₐᵢ_linear_isometry (s : affine_subspace 𝕜 P) [nonempty s] : + s.subtypeₐᵢ.linear_isometry = s.direction.subtypeₗᵢ := +rfl + +@[simp] lemma coe_subtypeₐᵢ (s : affine_subspace 𝕜 P) [nonempty s] : ⇑s.subtypeₐᵢ = s.subtype := +rfl + +@[simp] lemma subtypeₐᵢ_to_affine_map (s : affine_subspace 𝕜 P) [nonempty s] : + s.subtypeₐᵢ.to_affine_map = s.subtype := +rfl + +end affine_subspace variables (𝕜 P P₂) include V V₂ From 24b50bfcc4bffef8a9b5772df7646ea6048c97fd Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 23 Sep 2022 11:55:23 +0000 Subject: [PATCH 364/828] chore(data/finset/basic): move `disjoint` proofs earlier (#16436) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This avoids repeating work for `disj_union`, and takes the stance that `disjoint` is part of the lattice API. Co-authored-by: Kyle Miller Co-authored-by: Yaël Dillies --- src/data/finset/basic.lean | 221 ++++++++++++++++++------------------- 1 file changed, 108 insertions(+), 113 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index cc369ee8984ad..dd6ce0847cc36 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -764,6 +764,9 @@ instance : has_union (finset α) := ⟨λ s t, ⟨_, t.2.ndunion s.1⟩⟩ /-- `s ∩ t` is the set such that `a ∈ s ∩ t` iff `a ∈ s` and `a ∈ t`. -/ instance : has_inter (finset α) := ⟨λ s t, ⟨_, s.2.ndinter t.1⟩⟩ +instance {α : Type*} : order_bot (finset α) := +{ bot := ∅, bot_le := empty_subset } + instance : lattice (finset α) := { sup := (∪), sup_le := λ s t u hs ht a ha, (mem_ndunion.1 ha).elim (λ h, hs h) (λ h, ht h), @@ -775,10 +778,11 @@ instance : lattice (finset α) := inf_le_right := λ s t a h, (mem_ndinter.1 h).2, ..finset.partial_order } -/-! #### union -/ - @[simp] lemma sup_eq_union : ((⊔) : finset α → finset α → finset α) = (∪) := rfl @[simp] lemma inf_eq_inter : ((⊓) : finset α → finset α → finset α) = (∩) := rfl +@[simp] lemma bot_eq_empty {α : Type*} : (⊥ : finset α) = ∅ := rfl + +/-! #### union -/ lemma union_val_nd (s t : finset α) : (s ∪ t).1 = ndunion s.1 t.1 := rfl @@ -1009,11 +1013,6 @@ end lemma inter_subset_inter_left (h : t ⊆ u) : s ∩ t ⊆ s ∩ u := inter_subset_inter subset.rfl h lemma inter_subset_inter_right (h : s ⊆ t) : s ∩ u ⊆ t ∩ u := inter_subset_inter h subset.rfl -instance {α : Type u} : order_bot (finset α) := -{ bot := ∅, bot_le := empty_subset } - -@[simp] lemma bot_eq_empty {α : Type u} : (⊥ : finset α) = ∅ := rfl - instance : distrib_lattice (finset α) := { le_sup_inf := assume a b c, show (a ∪ b) ∩ (a ∪ c) ⊆ a ∪ b ∩ c, by simp only [subset_iff, mem_inter, mem_union, and_imp, or_imp_distrib] {contextual:=tt}; @@ -1073,6 +1072,67 @@ lemma ite_subset_union (s s' : finset α) (P : Prop) [decidable P] : lemma inter_subset_ite (s s' : finset α) (P : Prop) [decidable P] : s ∩ s' ⊆ ite P s s' := inf_le_ite s s' P +/-! ### disjoint -/ + +lemma disjoint_left : disjoint s t ↔ ∀ ⦃a⦄, a ∈ s → a ∉ t := +by simp only [_root_.disjoint, inf_eq_inter, le_iff_subset, subset_iff, mem_inter, not_and, + and_imp]; refl + +lemma disjoint_val : disjoint s t ↔ s.1.disjoint t.1 := disjoint_left +lemma disjoint_iff_inter_eq_empty : disjoint s t ↔ s ∩ t = ∅ := disjoint_iff + +instance decidable_disjoint (U V : finset α) : decidable (disjoint U V) := +decidable_of_decidable_of_iff (by apply_instance) eq_bot_iff + +lemma disjoint_right : disjoint s t ↔ ∀ ⦃a⦄, a ∈ t → a ∉ s := by rw [disjoint.comm, disjoint_left] +lemma disjoint_iff_ne : disjoint s t ↔ ∀ a ∈ s, ∀ b ∈ t, a ≠ b := +by simp only [disjoint_left, imp_not_comm, forall_eq'] + +lemma _root_.disjoint.forall_ne_finset (h : disjoint s t) (ha : a ∈ s) (hb : b ∈ t) : a ≠ b := +disjoint_iff_ne.1 h _ ha _ hb + +lemma not_disjoint_iff : ¬ disjoint s t ↔ ∃ a, a ∈ s ∧ a ∈ t := +not_forall.trans $ exists_congr $ λ a, not_not.trans mem_inter + +lemma disjoint_of_subset_left (h : s ⊆ u) (d : disjoint u t) : disjoint s t := +disjoint_left.2 (λ x m₁, (disjoint_left.1 d) (h m₁)) + +lemma disjoint_of_subset_right (h : t ⊆ u) (d : disjoint s u) : disjoint s t := +disjoint_right.2 (λ x m₁, (disjoint_right.1 d) (h m₁)) + +@[simp] theorem disjoint_empty_left (s : finset α) : disjoint ∅ s := disjoint_bot_left +@[simp] theorem disjoint_empty_right (s : finset α) : disjoint s ∅ := disjoint_bot_right + +@[simp] lemma disjoint_singleton_left : disjoint (singleton a) s ↔ a ∉ s := +by simp only [disjoint_left, mem_singleton, forall_eq] + +@[simp] lemma disjoint_singleton_right : disjoint s (singleton a) ↔ a ∉ s := +disjoint.comm.trans disjoint_singleton_left + +@[simp] lemma disjoint_singleton : disjoint ({a} : finset α) {b} ↔ a ≠ b := +by rw [disjoint_singleton_left, mem_singleton] + +@[simp] lemma disjoint_insert_left : disjoint (insert a s) t ↔ a ∉ t ∧ disjoint s t := +by simp only [disjoint_left, mem_insert, or_imp_distrib, forall_and_distrib, forall_eq] + +@[simp] lemma disjoint_insert_right : disjoint s (insert a t) ↔ a ∉ s ∧ disjoint s t := +disjoint.comm.trans $ by rw [disjoint_insert_left, disjoint.comm] + +@[simp] lemma disjoint_union_left : disjoint (s ∪ t) u ↔ disjoint s u ∧ disjoint t u := +by simp only [disjoint_left, mem_union, or_imp_distrib, forall_and_distrib] + +@[simp] lemma disjoint_union_right : disjoint s (t ∪ u) ↔ disjoint s t ∧ disjoint s u := +by simp only [disjoint_right, mem_union, or_imp_distrib, forall_and_distrib] + +lemma disjoint_self_iff_empty (s : finset α) : disjoint s s ↔ s = ∅ := disjoint_self + +@[simp, norm_cast] lemma disjoint_coe : disjoint (s : set α) t ↔ disjoint s t := +by { rw [finset.disjoint_left, set.disjoint_left], refl } + +@[simp, norm_cast] lemma pairwise_disjoint_coe {ι : Type*} {s : set ι} {f : ι → finset α} : + s.pairwise_disjoint (λ i, f i : ι → set α) ↔ s.pairwise_disjoint f := +forall₅_congr $ λ _ _ _ _ _, disjoint_coe + /-! ### erase -/ /-- `erase s a` is the set `s - {a}`, that is, the elements of `s` which are @@ -1339,6 +1399,16 @@ sup_eq_sdiff_sup_sdiff_sup_inf lemma erase_eq_empty_iff (s : finset α) (a : α) : s.erase a = ∅ ↔ s = ∅ ∨ s = {a} := by rw [←sdiff_singleton_eq_erase, sdiff_eq_empty_iff_subset, subset_singleton_iff] +--TODO@Yaël: Kill lemmas duplicate with `boolean_algebra` +lemma sdiff_disjoint : disjoint (t \ s) s := disjoint_left.2 $ assume a ha, (mem_sdiff.1 ha).2 +lemma disjoint_sdiff : disjoint s (t \ s) := sdiff_disjoint.symm + +lemma disjoint_sdiff_inter (s t : finset α) : disjoint (s \ t) (s ∩ t) := +disjoint_of_subset_right (inter_subset_right _ _) sdiff_disjoint + +lemma sdiff_eq_self_iff_disjoint : s \ t = s ↔ disjoint s t := sdiff_eq_self_iff_disjoint' +lemma sdiff_eq_self_of_disjoint (h : disjoint s t) : s \ t = s := sdiff_eq_self_iff_disjoint.2 h + /-! ### Symmetric difference -/ lemma mem_symm_diff : a ∈ s ∆ t ↔ a ∈ s ∧ a ∉ t ∨ a ∈ t ∧ a ∉ s := @@ -1761,6 +1831,19 @@ by { ext, simp only [mem_filter, mem_erase, ne.def], tauto } lemma filter_ne' [decidable_eq β] (s : finset β) (b : β) : s.filter (λ a, a ≠ b) = s.erase b := trans (filter_congr (λ _ _, ⟨ne.symm, ne.symm⟩)) (filter_ne s b) +lemma disjoint_filter {s : finset α} {p q : α → Prop} [decidable_pred p] [decidable_pred q] : + disjoint (s.filter p) (s.filter q) ↔ ∀ x ∈ s, p x → ¬ q x := +by split; simp [disjoint_left] {contextual := tt} + +lemma disjoint_filter_filter {s t : finset α} {p q : α → Prop} + [decidable_pred p] [decidable_pred q] : + disjoint s t → disjoint (s.filter p) (t.filter q) := +disjoint.mono (filter_subset _ _) (filter_subset _ _) + +lemma disjoint_filter_filter_neg (s : finset α) (p : α → Prop) [decidable_pred p] : + disjoint (s.filter p) (s.filter $ λ a, ¬ p a) := +(disjoint_filter.2 $ λ a _, id).symm + end filter /-! ### range -/ @@ -2072,6 +2155,10 @@ theorem map_disj_union_aux {f : α ↪ β} {s₁ s₂ : finset α} : (∀ a, a ∈ s₁ → a ∉ s₂) ↔ ∀ a, a ∈ map f s₁ → a ∉ map f s₂ := by simp_rw [forall_mem_map, mem_map'] +@[simp] lemma disjoint_map [decidable_eq α] [decidable_eq β] {s t : finset α} (f : α ↪ β) : + disjoint (s.map f) (t.map f) ↔ disjoint s t := +by rw [disjoint_left, disjoint_left, ←map_disj_union_aux] + theorem map_disj_union {f : α ↪ β} (s₁ s₂ : finset α) (h) (h' := map_disj_union_aux.1 h) : (s₁.disj_union s₂ h).map f = (s₁.map f).disj_union (s₂.map f) h' := eq_of_veq $ multiset.map_add _ _ _ @@ -2305,6 +2392,12 @@ end ⟨λ h, eq_empty_of_forall_not_mem $ λ a m, ne_empty_of_mem (mem_image_of_mem _ m) h, λ e, e.symm ▸ rfl⟩ +@[simp] lemma _root_.disjoint.of_image_finset [decidable_eq α] + {s t : finset α} {f : α → β} (h : disjoint (s.image f) (t.image f)) : + disjoint s t := +disjoint_iff_ne.2 $ λ a ha b hb, ne_of_apply_ne f $ h.forall_ne_finset + (mem_image_of_mem _ ha) (mem_image_of_mem _ hb) + lemma mem_range_iff_mem_finset_range_of_mod_eq' [decidable_eq α] {f : ℕ → α} {a : α} {n : ℕ} (hn : 0 < n) (h : ∀ i, f (i % n) = f i) : a ∈ set.range f ↔ a ∈ (finset.range n).image (λi, f i) := @@ -2352,6 +2445,11 @@ ext $ λ ⟨x, hx⟩, ⟨or.cases_on (mem_insert.1 hx) theorem map_eq_image (f : α ↪ β) (s : finset α) : s.map f = s.image f := eq_of_veq (s.map f).2.dedup.symm +@[simp] lemma disjoint_image [decidable_eq α] + {s t : finset α} {f : α → β} (hf : injective f) : + disjoint (s.image f) (t.image f) ↔ disjoint s t := +by convert disjoint_map ⟨_, hf⟩; simp [map_eq_image] + lemma image_const {s : finset α} (h : s.nonempty) (b : β) : s.image (λa, b) = singleton b := ext $ assume b', by simp only [mem_image, exists_prop, exists_and_distrib_right, h.bex, true_and, mem_singleton, eq_comm] @@ -2646,74 +2744,7 @@ by simp [finset.nonempty, ← exists_and_distrib_left, @exists_swap α] lemma nonempty.bUnion (hs : s.nonempty) (ht : ∀ x ∈ s, (t x).nonempty) : (s.bUnion t).nonempty := bUnion_nonempty.2 $ hs.imp $ λ x hx, ⟨hx, ht x hx⟩ -end bUnion - -/-! ### disjoint -/ ---TODO@Yaël: Kill lemmas duplicate with `boolean_algebra` -section disjoint -variables [decidable_eq α] [decidable_eq β] {f : α → β} {s t u : finset α} {a b : α} - -lemma disjoint_left : disjoint s t ↔ ∀ ⦃a⦄, a ∈ s → a ∉ t := -by simp only [_root_.disjoint, inf_eq_inter, le_iff_subset, subset_iff, mem_inter, not_and, - and_imp]; refl - -lemma disjoint_val : disjoint s t ↔ s.1.disjoint t.1 := disjoint_left -lemma disjoint_iff_inter_eq_empty : disjoint s t ↔ s ∩ t = ∅ := disjoint_iff - -instance decidable_disjoint (U V : finset α) : decidable (disjoint U V) := -decidable_of_decidable_of_iff (by apply_instance) eq_bot_iff - -lemma disjoint_right : disjoint s t ↔ ∀ ⦃a⦄, a ∈ t → a ∉ s := by rw [disjoint.comm, disjoint_left] -lemma disjoint_iff_ne : disjoint s t ↔ ∀ a ∈ s, ∀ b ∈ t, a ≠ b := -by simp only [disjoint_left, imp_not_comm, forall_eq'] - -lemma _root_.disjoint.forall_ne_finset (h : disjoint s t) (ha : a ∈ s) (hb : b ∈ t) : a ≠ b := -disjoint_iff_ne.1 h _ ha _ hb - -lemma not_disjoint_iff : ¬ disjoint s t ↔ ∃ a, a ∈ s ∧ a ∈ t := -not_forall.trans $ exists_congr $ λ a, not_not.trans mem_inter - -lemma disjoint_of_subset_left (h : s ⊆ u) (d : disjoint u t) : disjoint s t := -disjoint_left.2 (λ x m₁, (disjoint_left.1 d) (h m₁)) - -lemma disjoint_of_subset_right (h : t ⊆ u) (d : disjoint s u) : disjoint s t := -disjoint_right.2 (λ x m₁, (disjoint_right.1 d) (h m₁)) - -@[simp] theorem disjoint_empty_left (s : finset α) : disjoint ∅ s := disjoint_bot_left -@[simp] theorem disjoint_empty_right (s : finset α) : disjoint s ∅ := disjoint_bot_right - -@[simp] lemma disjoint_singleton_left : disjoint (singleton a) s ↔ a ∉ s := -by simp only [disjoint_left, mem_singleton, forall_eq] - -@[simp] lemma disjoint_singleton_right : disjoint s (singleton a) ↔ a ∉ s := -disjoint.comm.trans disjoint_singleton_left - -@[simp] lemma disjoint_singleton : disjoint ({a} : finset α) {b} ↔ a ≠ b := -by rw [disjoint_singleton_left, mem_singleton] - -@[simp] lemma disjoint_insert_left : disjoint (insert a s) t ↔ a ∉ t ∧ disjoint s t := -by simp only [disjoint_left, mem_insert, or_imp_distrib, forall_and_distrib, forall_eq] - -@[simp] lemma disjoint_insert_right : disjoint s (insert a t) ↔ a ∉ s ∧ disjoint s t := -disjoint.comm.trans $ by rw [disjoint_insert_left, disjoint.comm] - -@[simp] lemma disjoint_union_left : disjoint (s ∪ t) u ↔ disjoint s u ∧ disjoint t u := -by simp only [disjoint_left, mem_union, or_imp_distrib, forall_and_distrib] - -@[simp] lemma disjoint_union_right : disjoint s (t ∪ u) ↔ disjoint s t ∧ disjoint s u := -by simp only [disjoint_right, mem_union, or_imp_distrib, forall_and_distrib] - -lemma sdiff_disjoint : disjoint (t \ s) s := disjoint_left.2 $ assume a ha, (mem_sdiff.1 ha).2 -lemma disjoint_sdiff : disjoint s (t \ s) := sdiff_disjoint.symm - -lemma disjoint_sdiff_inter (s t : finset α) : disjoint (s \ t) (s ∩ t) := -disjoint_of_subset_right (inter_subset_right _ _) sdiff_disjoint - -lemma sdiff_eq_self_iff_disjoint : s \ t = s ↔ disjoint s t := sdiff_eq_self_iff_disjoint' -lemma sdiff_eq_self_of_disjoint (h : disjoint s t) : s \ t = s := sdiff_eq_self_iff_disjoint.2 h -lemma disjoint_self_iff_empty (s : finset α) : disjoint s s ↔ s = ∅ := disjoint_self - -lemma disjoint_bUnion_left {ι : Type*} (s : finset ι) (f : ι → finset α) (t : finset α) : +lemma disjoint_bUnion_left (s : finset α) (f : α → finset β) (t : finset β) : disjoint (s.bUnion f) t ↔ (∀ i ∈ s, disjoint (f i) t) := begin classical, @@ -2723,47 +2754,11 @@ begin simp only [disjoint_union_left, bUnion_insert, his, forall_mem_insert, ih] } end -lemma disjoint_bUnion_right {ι : Type*} (s : finset α) (t : finset ι) (f : ι → finset α) : +lemma disjoint_bUnion_right (s : finset β) (t : finset α) (f : α → finset β) : disjoint s (t.bUnion f) ↔ ∀ i ∈ t, disjoint s (f i) := by simpa only [disjoint.comm] using disjoint_bUnion_left t f s -lemma disjoint_filter {p q : α → Prop} [decidable_pred p] [decidable_pred q] : - disjoint (s.filter p) (s.filter q) ↔ ∀ x ∈ s, p x → ¬ q x := -by split; simp [disjoint_left] {contextual := tt} - -lemma disjoint_filter_filter {p q : α → Prop} [decidable_pred p] [decidable_pred q] : - (disjoint s t) → disjoint (s.filter p) (t.filter q) := -disjoint.mono (filter_subset _ _) (filter_subset _ _) - -lemma disjoint_filter_filter_neg (s : finset α) (p : α → Prop) [decidable_pred p] : - disjoint (s.filter p) (s.filter $ λ a, ¬ p a) := -(disjoint_filter.2 $ λ a _, id).symm - -@[simp, norm_cast] lemma disjoint_coe : disjoint (s : set α) t ↔ disjoint s t := -by { rw [finset.disjoint_left, set.disjoint_left], refl } - -@[simp, norm_cast] lemma pairwise_disjoint_coe {ι : Type*} {s : set ι} {f : ι → finset α} : - s.pairwise_disjoint (λ i, f i : ι → set α) ↔ s.pairwise_disjoint f := -forall₅_congr $ λ _ _ _ _ _, disjoint_coe - -@[simp] lemma _root_.disjoint.of_image_finset (h : disjoint (s.image f) (t.image f)) : - disjoint s t := -disjoint_iff_ne.2 $ λ a ha b hb, ne_of_apply_ne f $ h.forall_ne_finset - (mem_image_of_mem _ ha) (mem_image_of_mem _ hb) - -@[simp] lemma disjoint_image {f : α → β} (hf : injective f) : - disjoint (s.image f) (t.image f) ↔ disjoint s t := -begin - simp only [disjoint_iff_ne, mem_image, exists_prop, exists_imp_distrib, and_imp], - refine ⟨λ h a ha b hb hab, h _ _ ha rfl _ _ hb rfl $ congr_arg _ hab, _⟩, - rintro h _ a ha rfl _ b hb rfl, - exact hf.ne (h _ ha _ hb), -end - -@[simp] lemma disjoint_map {f : α ↪ β} : disjoint (s.map f) (t.map f) ↔ disjoint s t := -by { simp_rw map_eq_image, exact disjoint_image f.injective } - -end disjoint +end bUnion /-! ### choose -/ section choose From 313f36d4bf82a0ea51d9faa548c344abaaef1efd Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 23 Sep 2022 11:55:24 +0000 Subject: [PATCH 365/828] feat(order/filter/lift): replace some implications with iffs (#16452) --- src/order/filter/lift.lean | 56 +++++++++---------- src/topology/basic.lean | 2 +- src/topology/uniform_space/basic.lean | 2 +- src/topology/uniform_space/cauchy.lean | 6 +- src/topology/uniform_space/completion.lean | 4 +- .../uniform_space/uniform_embedding.lean | 13 +---- 6 files changed, 35 insertions(+), 48 deletions(-) diff --git a/src/order/filter/lift.lean b/src/order/filter/lift.lean index 11faccfa0d540..a635be5c441d0 100644 --- a/src/order/filter/lift.lean +++ b/src/order/filter/lift.lean @@ -72,6 +72,11 @@ lemma mem_lift_sets (hg : monotone g) {s : set β} : (f.basis_sets.mem_lift_iff (λ s, (g s).basis_sets) hg).trans $ by simp only [id, exists_mem_subset_iff] +lemma sInter_lift_sets (hg : monotone g) : + ⋂₀ {s | s ∈ f.lift g} = ⋂ s ∈ f, ⋂₀ {t | t ∈ g s} := +by simp only [sInter_eq_bInter, mem_set_of_eq, filter.mem_sets, mem_lift_sets hg, + Inter_exists, @Inter_comm _ (set β)] + lemma mem_lift {s : set β} {t : set α} (ht : t ∈ f) (hs : s ∈ g t) : s ∈ f.lift g := le_principal_iff.mp $ show f.lift g ≤ 𝓟 s, @@ -81,9 +86,9 @@ lemma lift_le {f : filter α} {g : set α → filter β} {h : filter β} {s : se (hs : s ∈ f) (hg : g s ≤ h) : f.lift g ≤ h := infi₂_le_of_le s hs hg -lemma le_lift {f : filter α} {g : set α → filter β} {h : filter β} - (hh : ∀s∈f, h ≤ g s) : h ≤ f.lift g := -le_infi₂ hh +lemma le_lift {f : filter α} {g : set α → filter β} {h : filter β} : + h ≤ f.lift g ↔ ∀ s ∈ f, h ≤ g s := +le_infi₂_iff lemma lift_mono (hf : f₁ ≤ f₂) (hg : g₁ ≤ g₂) : f₁.lift g₁ ≤ f₂.lift g₂ := infi_mono $ λ s, infi_mono' $ λ hs, ⟨hf hs, hg s⟩ @@ -141,43 +146,33 @@ le_antisymm lemma lift_lift_same_le_lift {g : set α → set α → filter β} : f.lift (λs, f.lift (g s)) ≤ f.lift (λs, g s s) := -le_infi $ assume s, le_infi $ assume hs, infi_le_of_le s $ infi_le_of_le hs $ infi_le_of_le s $ - infi_le _ hs +le_lift.2 $ λ s hs, lift_le hs $ lift_le hs le_rfl lemma lift_lift_same_eq_lift {g : set α → set α → filter β} (hg₁ : ∀s, monotone (λt, g s t)) (hg₂ : ∀t, monotone (λs, g s t)) : f.lift (λs, f.lift (g s)) = f.lift (λs, g s s) := -le_antisymm - lift_lift_same_le_lift - (le_infi $ assume s, le_infi $ assume hs, le_infi $ assume t, le_infi $ assume ht, - infi_le_of_le (s ∩ t) $ - infi_le_of_le (inter_mem hs ht) $ +lift_lift_same_le_lift.antisymm $ + le_lift.2 $ λ s hs, le_lift.2 $ λ t ht, lift_le (inter_mem hs ht) $ calc g (s ∩ t) (s ∩ t) ≤ g s (s ∩ t) : hg₂ (s ∩ t) (inter_subset_left _ _) - ... ≤ g s t : hg₁ s (inter_subset_right _ _)) + ... ≤ g s t : hg₁ s (inter_subset_right _ _) lemma lift_principal {s : set α} (hg : monotone g) : (𝓟 s).lift g = g s := -le_antisymm - (infi_le_of_le s $ infi_le _ $ subset.refl _) - (le_infi $ assume t, le_infi $ assume hi, hg hi) +(lift_le (mem_principal_self _) le_rfl).antisymm (le_lift.2 $ λ t ht, hg ht) theorem monotone_lift [preorder γ] {f : γ → filter α} {g : γ → set α → filter β} (hf : monotone f) (hg : monotone g) : monotone (λc, (f c).lift (g c)) := assume a b h, lift_mono (hf h) (hg h) lemma lift_ne_bot_iff (hm : monotone g) : (ne_bot $ f.lift g) ↔ (∀s∈f, ne_bot (g s)) := -begin - rw [filter.lift, infi_subtype', infi_ne_bot_iff_of_directed', subtype.forall'], - { rintros ⟨s, hs⟩ ⟨t, ht⟩, - exact ⟨⟨s ∩ t, inter_mem hs ht⟩, hm (inter_subset_left s t), hm (inter_subset_right s t)⟩ } -end +by simp only [ne_bot_iff, ne.def, ← empty_mem_iff_bot, mem_lift_sets hm, not_exists] @[simp] lemma lift_const {f : filter α} {g : filter β} : f.lift (λx, g) = g := -le_antisymm (lift_le univ_mem $ le_refl g) (le_lift $ assume s hs, le_refl g) +infi_subtype'.trans infi_const @[simp] lemma lift_inf {f : filter α} {g h : set α → filter β} : f.lift (λx, g x ⊓ h x) = f.lift g ⊓ f.lift h := -by simp only [filter.lift, infi_inf_eq, eq_self_iff_true] +by simp only [filter.lift, infi_inf_eq] @[simp] lemma lift_principal2 {f : filter α} : f.lift 𝓟 = f := le_antisymm @@ -250,13 +245,17 @@ begin simp only [exists_const] end -lemma mem_lift'_sets (hh : monotone h) {s : set β} : s ∈ (f.lift' h) ↔ (∃t∈f, h t ⊆ s) := +lemma mem_lift'_sets (hh : monotone h) {s : set β} : s ∈ f.lift' h ↔ ∃ t ∈ f, h t ⊆ s := mem_lift_sets $ monotone_principal.comp hh lemma eventually_lift'_iff (hh : monotone h) {p : β → Prop} : (∀ᶠ y in f.lift' h, p y) ↔ (∃ t ∈ f, ∀ y ∈ h t, p y) := mem_lift'_sets hh +lemma sInter_lift'_sets (hh : monotone h) : + ⋂₀ {s | s ∈ f.lift' h} = ⋂ s ∈ f, h s := +(sInter_lift_sets (monotone_principal.comp hh)).trans $ Inter₂_congr $ λ s hs, cInf_Ici + lemma lift'_le {f : filter α} {g : set α → set β} {h : filter β} {s : set α} (hs : s ∈ f) (hg : 𝓟 (g s) ≤ h) : f.lift' g ≤ h := lift_le hs hg @@ -298,9 +297,11 @@ by rw [← principal_singleton, lift'_principal hh] lemma lift'_bot (hh : monotone h) : (⊥ : filter α).lift' h = 𝓟 (h ∅) := by rw [← principal_empty, lift'_principal hh] -lemma principal_le_lift' {t : set β} (hh : ∀s∈f, t ⊆ h s) : - 𝓟 t ≤ f.lift' h := -le_infi $ assume s, le_infi $ assume hs, principal_mono.mpr (hh s hs) +lemma le_lift' {f : filter α} {h : set α → set β} {g : filter β} : + g ≤ f.lift' h ↔ ∀ s ∈ f, h s ∈ g := +le_lift.trans $ forall₂_congr $ λ s hs, le_principal_iff + +lemma principal_le_lift' {t : set β} : 𝓟 t ≤ f.lift' h ↔ ∀ s ∈ f, t ⊆ h s := le_lift' theorem monotone_lift' [preorder γ] {f : γ → filter α} {g : γ → set α → set β} (hf : monotone f) (hg : monotone g) : monotone (λc, (f c).lift' (g c)) := @@ -345,11 +346,6 @@ calc (ne_bot (f.lift' h)) ↔ (∀s∈f, ne_bot (𝓟 (h s))) : @[simp] lemma lift'_id {f : filter α} : f.lift' id = f := lift_principal2 -lemma le_lift' {f : filter α} {h : set α → set β} {g : filter β} - (h_le : ∀s∈f, h s ∈ g) : g ≤ f.lift' h := -le_infi $ assume s, le_infi $ assume hs, - by simpa only [h_le, le_principal_iff, function.comp_app] using h_le s hs - lemma lift'_infi [nonempty ι] {f : ι → filter α} {g : set α → set β} (hg : ∀ s t, g (s ∩ t) = g s ∩ g t) : (infi f).lift' g = (⨅ i, (f i).lift' g) := lift_infi $ λ s t, by rw [inf_principal, (∘), ← hg] diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 91a2f6f769f51..b485922e4669e 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -451,7 +451,7 @@ theorem mem_closure_iff {s : set α} {a : α} : let ⟨x, hc, hs⟩ := (H _ h₁.is_open_compl nc) in hc (h₂ hs)⟩ lemma filter.le_lift'_closure (l : filter α) : l ≤ l.lift' closure := -le_infi₂ $ λ s hs, le_principal_iff.2 $ mem_of_superset hs subset_closure +le_lift'.2 $ λ s hs, mem_of_superset hs subset_closure lemma filter.has_basis.lift'_closure {l : filter α} {p : ι → Prop} {s : ι → set α} (h : l.has_basis p s) : diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index 696fdecd164a2..0ec0af2ac8d0b 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -384,7 +384,7 @@ lemma filter.tendsto.uniformity_trans {l : filter β} {f₁ f₂ f₃ : β → (h₁₂ : tendsto (λ x, (f₁ x, f₂ x)) l (𝓤 α)) (h₂₃ : tendsto (λ x, (f₂ x, f₃ x)) l (𝓤 α)) : tendsto (λ x, (f₁ x, f₃ x)) l (𝓤 α) := begin - refine le_trans (le_lift' $ λ s hs, mem_map.2 _) comp_le_uniformity, + refine le_trans (le_lift'.2 $ λ s hs, mem_map.2 _) comp_le_uniformity, filter_upwards [h₁₂ hs, h₂₃ hs] with x hx₁₂ hx₂₃ using ⟨_, hx₁₂, hx₂₃⟩, end diff --git a/src/topology/uniform_space/cauchy.lean b/src/topology/uniform_space/cauchy.lean index 86372d64aa777..325d9efb2d448 100644 --- a/src/topology/uniform_space/cauchy.lean +++ b/src/topology/uniform_space/cauchy.lean @@ -365,10 +365,8 @@ instance complete_space.prod [uniform_space β] [complete_space α] [complete_sp let ⟨x1, hx1⟩ := complete_space.complete $ hf.map uniform_continuous_fst in let ⟨x2, hx2⟩ := complete_space.complete $ hf.map uniform_continuous_snd in ⟨(x1, x2), by rw [nhds_prod_eq, filter.prod_def]; - from filter.le_lift (λ s hs, filter.le_lift' $ λ t ht, - have H1 : prod.fst ⁻¹' s ∈ f.sets := hx1 hs, - have H2 : prod.snd ⁻¹' t ∈ f.sets := hx2 ht, - filter.inter_mem H1 H2)⟩ } + from filter.le_lift.2 (λ s hs, filter.le_lift'.2 $ λ t ht, + inter_mem (hx1 hs) (hx2 ht))⟩ } /--If `univ` is complete, the space is a complete space -/ lemma complete_space_of_is_complete_univ (h : is_complete (univ : set α)) : complete_space α := diff --git a/src/topology/uniform_space/completion.lean b/src/topology/uniform_space/completion.lean index d2a64d3752a0e..a566557cd2fb6 100644 --- a/src/topology/uniform_space/completion.lean +++ b/src/topology/uniform_space/completion.lean @@ -126,7 +126,7 @@ calc ((𝓤 α).lift' gen).lift' (λs, comp_rel s s) = instance : uniform_space (Cauchy α) := uniform_space.of_core { uniformity := (𝓤 α).lift' gen, - refl := principal_le_lift' $ assume s hs ⟨a, b⟩ (a_eq_b : a = b), + refl := principal_le_lift'.2 $ λ s hs ⟨a, b⟩ (a_eq_b : a = b), a_eq_b ▸ a.property.right hs, symm := symm_gen, comp := comp_gen } @@ -208,7 +208,7 @@ complete_space_extension assume f hf, let f' : Cauchy α := ⟨f, hf⟩ in have map pure_cauchy f ≤ (𝓤 $ Cauchy α).lift' (preimage (prod.mk f')), - from le_lift' $ assume s hs, + from le_lift'.2 $ assume s hs, let ⟨t, ht₁, (ht₂ : gen t ⊆ s)⟩ := (mem_lift'_sets monotone_gen).mp hs in let ⟨t', ht', (h : t' ×ˢ t' ⊆ t)⟩ := mem_prod_same_iff.mp (hf.right ht₁) in have t' ⊆ { y : α | (f', pure_cauchy y) ∈ gen t }, diff --git a/src/topology/uniform_space/uniform_embedding.lean b/src/topology/uniform_space/uniform_embedding.lean index 02fe5f8840471..76480ca943ace 100644 --- a/src/topology/uniform_space/uniform_embedding.lean +++ b/src/topology/uniform_space/uniform_embedding.lean @@ -413,16 +413,9 @@ end instance complete_space.sum [complete_space α] [complete_space β] : complete_space (α ⊕ β) := begin - rw complete_space_iff_is_complete_univ, - have A : is_complete (range (sum.inl : α → α ⊕ β)) := - uniform_embedding_inl.to_uniform_inducing.is_complete_range, - have B : is_complete (range (sum.inr : β → α ⊕ β)) := - uniform_embedding_inr.to_uniform_inducing.is_complete_range, - convert A.union B, - apply (eq_univ_of_forall (λ x, _)).symm, - cases x, - { left, exact mem_range_self _ }, - { right, exact mem_range_self _ } + rw [complete_space_iff_is_complete_univ, ← range_inl_union_range_inr], + exact uniform_embedding_inl.to_uniform_inducing.is_complete_range.union + uniform_embedding_inr.to_uniform_inducing.is_complete_range end end From 249a9c91f009e70faf0c8ce3a62fa224d3d59ba8 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 23 Sep 2022 14:45:51 +0000 Subject: [PATCH 366/828] feat(category_theory/limits/shapes/biproducts): more (co)kernel instances (#14188) Co-authored-by: Scott Morrison --- .../limits/shapes/biproducts.lean | 136 +++++++++++++++++- 1 file changed, 131 insertions(+), 5 deletions(-) diff --git a/src/category_theory/limits/shapes/biproducts.lean b/src/category_theory/limits/shapes/biproducts.lean index e103165ce5f4b..bc4b5f2cca98c 100644 --- a/src/category_theory/limits/shapes/biproducts.lean +++ b/src/category_theory/limits/shapes/biproducts.lean @@ -567,12 +567,13 @@ end end -variables (f : J → C) (i : J) [has_biproduct f] [has_biproduct (subtype.restrict (λ j, i ≠ j) f)] +section +variables (f : J → C) (i : J) [has_biproduct f] [has_biproduct (subtype.restrict (λ j, j ≠ i) f)] /-- 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.is_limit_from_subtype : is_limit - (kernel_fork.of_ι (biproduct.from_subtype f (λ j, i ≠ j)) + (kernel_fork.of_ι (biproduct.from_subtype f (λ j, j ≠ i)) (by simp) : kernel_fork (biproduct.π f i)) := fork.is_limit.mk' _ $ λ s, ⟨s.ι ≫ biproduct.to_subtype _ _, @@ -582,7 +583,7 @@ fork.is_limit.mk' _ $ λ s, biproduct.to_subtype_from_subtype_assoc, biproduct.map_π], rcases em (i = j) with (rfl|h), { rw [if_neg (not_not.2 rfl), comp_zero, comp_zero, kernel_fork.condition] }, - { rw [if_pos, category.comp_id], exact h, } + { rw [if_pos (ne.symm h), category.comp_id], } end, begin intros m hm, @@ -590,10 +591,18 @@ fork.is_limit.mk' _ $ λ s, exact (category.comp_id _).symm end⟩ +instance : has_kernel (biproduct.π f i) := +has_limit.mk ⟨_, biproduct.is_limit_from_subtype f i⟩ + +/-- The kernel of `biproduct.π f i` is `⨁ subtype.restrict {i}ᶜ f`. -/ +@[simps] +def kernel_biproduct_π_iso : kernel (biproduct.π f i) ≅ ⨁ subtype.restrict (λ j, j ≠ i) f := +limit.iso_limit_cone ⟨_, biproduct.is_limit_from_subtype f i⟩ + /-- The cokernel of `biproduct.ι f i` is the projection from the biproduct over the index set `J` onto the biproduct omitting `i`. -/ def biproduct.is_colimit_to_subtype : is_colimit - (cokernel_cofork.of_π (biproduct.to_subtype f (λ j, i ≠ j)) + (cokernel_cofork.of_π (biproduct.to_subtype f (λ j, j ≠ i)) (by simp) : cokernel_cofork (biproduct.ι f i)) := cofork.is_colimit.mk' _ $ λ s, ⟨biproduct.from_subtype _ _ ≫ s.π, @@ -603,7 +612,7 @@ cofork.is_colimit.mk' _ $ λ s, biproduct.ι_map_assoc], rcases em (i = j) with (rfl|h), { rw [if_neg (not_not.2 rfl), zero_comp, cokernel_cofork.condition] }, - { rw [if_pos, category.id_comp], exact h, } + { rw [if_pos (ne.symm h), category.id_comp], } end, begin intros m hm, @@ -611,6 +620,89 @@ cofork.is_colimit.mk' _ $ λ s, exact (category.id_comp _).symm end⟩ +instance : has_cokernel (biproduct.ι f i) := +has_colimit.mk ⟨_, biproduct.is_colimit_to_subtype f i⟩ + +/-- The cokernel of `biproduct.ι f i` is `⨁ subtype.restrict {i}ᶜ f`. -/ +@[simps] +def cokernel_biproduct_ι_iso : cokernel (biproduct.ι f i) ≅ ⨁ subtype.restrict (λ j, j ≠ i) f := +colimit.iso_colimit_cocone ⟨_, biproduct.is_colimit_to_subtype f i⟩ + +end + +section +open_locale classical + +-- Per #15067, we only allow indexing in `Type 0` here. +variables {K : Type} [fintype K] [has_finite_biproducts C] (f : K → C) + +/-- The limit cone exhibiting `⨁ subtype.restrict pᶜ f` as the kernel of +`biproduct.to_subtype f p` -/ +@[simps] +def kernel_fork_biproduct_to_subtype (p : set K) : + limit_cone (parallel_pair (biproduct.to_subtype f p) 0) := +{ cone := kernel_fork.of_ι (biproduct.from_subtype f pᶜ) begin + ext j k, + simp only [biproduct.ι_from_subtype_assoc, biproduct.ι_to_subtype, comp_zero, zero_comp], + erw [dif_neg j.2], + simp only [zero_comp], + end, + is_limit := kernel_fork.is_limit.of_ι _ _ (λ W g h, g ≫ biproduct.to_subtype f pᶜ) + begin + intros W' g' w, + ext j, + simp only [category.assoc, biproduct.to_subtype_from_subtype, pi.compl_apply, + biproduct.map_π], + split_ifs, + { simp, }, + { replace w := w =≫ biproduct.π _ ⟨j, not_not.mp h⟩, simpa using w.symm, }, + end + (by tidy), } + +instance (p : set K) : has_kernel (biproduct.to_subtype f p) := +has_limit.mk (kernel_fork_biproduct_to_subtype f p) + +/-- The kernel of `biproduct.to_subtype f p` is `⨁ subtype.restrict pᶜ f`. -/ +@[simps] +def kernel_biproduct_to_subtype_iso (p : set K) : + kernel (biproduct.to_subtype f p) ≅ ⨁ subtype.restrict pᶜ f := +limit.iso_limit_cone (kernel_fork_biproduct_to_subtype f p) + +/-- The colimit cocone exhibiting `⨁ subtype.restrict pᶜ f` as the cokernel of +`biproduct.from_subtype f p` -/ +@[simps] +def cokernel_cofork_biproduct_from_subtype (p : set K) : + colimit_cocone (parallel_pair (biproduct.from_subtype f p) 0) := +{ cocone := cokernel_cofork.of_π (biproduct.to_subtype f pᶜ) begin + ext j k, + simp only [pi.compl_apply, biproduct.ι_from_subtype_assoc, biproduct.ι_to_subtype, + comp_zero, zero_comp], + rw [dif_neg], + simp only [zero_comp], + exact not_not.mpr j.2, + end, + is_colimit := cokernel_cofork.is_colimit.of_π _ _ (λ W g h, biproduct.from_subtype f pᶜ ≫ g) + begin + intros W' g' w, + ext j, + simp only [biproduct.to_subtype_from_subtype_assoc, pi.compl_apply, biproduct.ι_map_assoc], + split_ifs, + { simp, }, + { replace w := biproduct.ι _ (⟨j, not_not.mp h⟩ : p) ≫= w, simpa using w.symm, }, + end + (by tidy), } + +instance (p : set K) : has_cokernel (biproduct.from_subtype f p) := +has_colimit.mk (cokernel_cofork_biproduct_from_subtype f p) + +/-- The cokernel of `biproduct.from_subtype f p` is `⨁ subtype.restrict pᶜ f`. -/ +@[simps] +def cokernel_biproduct_from_subtype_iso (p : set K) : + cokernel (biproduct.from_subtype f p) ≅ ⨁ subtype.restrict pᶜ f := +colimit.iso_colimit_cocone (cokernel_cofork_biproduct_from_subtype f p) + +end + end π_kernel end limits @@ -1380,6 +1472,40 @@ binary_bicone.is_colimit_inr_cokernel_cofork (binary_biproduct.is_colimit _ _) end has_binary_biproduct +variables {X Y : C} [has_binary_biproduct X Y] + +instance : has_kernel (biprod.fst : X ⊞ Y ⟶ X) := +has_limit.mk ⟨_, biprod.is_kernel_fst_kernel_fork X Y⟩ + +/-- The kernel of `biprod.fst : X ⊞ Y ⟶ X` is `Y`. -/ +@[simps] +def kernel_biprod_fst_iso : kernel (biprod.fst : X ⊞ Y ⟶ X) ≅ Y := +limit.iso_limit_cone ⟨_, biprod.is_kernel_fst_kernel_fork X Y⟩ + +instance : has_kernel (biprod.snd : X ⊞ Y ⟶ Y) := +has_limit.mk ⟨_, biprod.is_kernel_snd_kernel_fork X Y⟩ + +/-- The kernel of `biprod.snd : X ⊞ Y ⟶ Y` is `X`. -/ +@[simps] +def kernel_biprod_snd_iso : kernel (biprod.snd : X ⊞ Y ⟶ Y) ≅ X := +limit.iso_limit_cone ⟨_, biprod.is_kernel_snd_kernel_fork X Y⟩ + +instance : has_cokernel (biprod.inl : X ⟶ X ⊞ Y) := +has_colimit.mk ⟨_, biprod.is_cokernel_inl_cokernel_fork X Y⟩ + +/-- The cokernel of `biprod.inl : X ⟶ X ⊞ Y` is `Y`. -/ +@[simps] +def cokernel_biprod_inl_iso : cokernel (biprod.inl : X ⟶ X ⊞ Y) ≅ Y := +colimit.iso_colimit_cocone ⟨_, biprod.is_cokernel_inl_cokernel_fork X Y⟩ + +instance : has_cokernel (biprod.inr : Y ⟶ X ⊞ Y) := +has_colimit.mk ⟨_, biprod.is_cokernel_inr_cokernel_fork X Y⟩ + +/-- The cokernel of `biprod.inr : Y ⟶ X ⊞ Y` is `X`. -/ +@[simps] +def cokernel_biprod_inr_iso : cokernel (biprod.inr : Y ⟶ X ⊞ Y) ≅ X := +colimit.iso_colimit_cocone ⟨_, biprod.is_cokernel_inr_cokernel_fork X Y⟩ + end biprod_kernel section is_zero From 8f66240cab125b938b327d3850169d490cfbcdd8 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 23 Sep 2022 14:45:52 +0000 Subject: [PATCH 367/828] feat(group_theory/group_action): define `distrib_smul` and `smul_zero_class` (#16123) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit These are two new superclasses of `distrib_mul_action` that get rid of the `mul_action` part: * `smul_zero_class` is `has_smul` + `a • 0 = 0` * `distrib_smul` is `smul_zero_class` + `a • (x + y) = a • x + a • y`. The motivation for these classes is to instantiate `qsmul` on `splitting_field`: in general scalar multiplication with rational numbers is not a `distrib_mul_action` but it is a `distrib_smul`, and `distrib_smul` is sufficient to lift an action to the `splitting_field`. I set up both `distrib_mul_action` and `smul_with_zero` to be subclasses of the above classes, and unify `smul_zero` (depending on `distrib_mul_action`) and `smul_zero'` (depending on `smul_with_zero`) into one lemma. There are a few places where I need to help the elaborator because e.g. it's expecting `units.mk0 a ha • 0 = 0` (with `smul` coming from `distrib_mul_action`) and getting `a • 0 = 0` (with `smul` coming from `smul_zero_class`). Apparently having both the type and instance differ is too hard for the unifier. Because we don't have definitional eta for structures yet, setting up the inheritance is a bit more tricky than you might think; I added an `example` test case to ensure everything stays OK. Co-authored-by: Mauricio Collares Co-authored-by: Eric Rodriguez --- src/algebra/algebra/unitization.lean | 2 +- .../category/Module/change_of_rings.lean | 2 +- src/algebra/module/pi.lean | 4 +- src/algebra/module/prod.lean | 4 +- src/algebra/module/ulift.lean | 33 +++++++++++----- src/algebra/order/module.lean | 2 +- src/algebra/order/smul.lean | 8 ++-- src/algebra/smul_with_zero.lean | 13 ++----- src/algebra/triv_sq_zero_ext.lean | 2 +- src/analysis/convex/star.lean | 4 +- src/data/finset/pointwise.lean | 2 +- src/data/real/pointwise.lean | 12 +++--- src/data/set/pointwise.lean | 4 +- src/group_theory/group_action/defs.lean | 38 +++++++++++++++---- src/group_theory/group_action/group.lean | 6 +-- src/group_theory/group_action/pi.lean | 28 +++++++++++--- src/group_theory/group_action/prod.lean | 11 +++++- src/group_theory/group_action/units.lean | 10 ++++- src/information_theory/hamming.lean | 4 +- src/linear_algebra/matrix/diagonal.lean | 2 +- .../legendre_symbol/add_character.lean | 1 - src/order/filter/pointwise.lean | 2 +- src/ring_theory/localization/module.lean | 2 +- src/tactic/lint/type_classes.lean | 2 +- src/topology/support.lean | 2 +- 25 files changed, 132 insertions(+), 68 deletions(-) diff --git a/src/algebra/algebra/unitization.lean b/src/algebra/algebra/unitization.lean index 81dc0b2d05d02..ce86cba306cb8 100644 --- a/src/algebra/algebra/unitization.lean +++ b/src/algebra/algebra/unitization.lean @@ -213,7 +213,7 @@ ext neg_zero.symm rfl @[simp] lemma coe_smul [has_zero R] [has_zero S] [smul_with_zero S R] [has_smul S A] (r : S) (m : A) : (↑(r • m) : unitization R A) = r • m := -ext (smul_zero' _ _).symm rfl +ext (smul_zero _).symm rfl end diff --git a/src/algebra/category/Module/change_of_rings.lean b/src/algebra/category/Module/change_of_rings.lean index f09acdc2a3648..9d8c285ba5110 100644 --- a/src/algebra/category/Module/change_of_rings.lean +++ b/src/algebra/category/Module/change_of_rings.lean @@ -70,7 +70,7 @@ def restrict_scalars {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R @[simp] lemma restrict_scalars.smul_def {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) {M : Module.{v} S} (r : R) (m : (restrict_scalars f).obj M) : r • m = (f r • m : M) := rfl -@[simp] lemma restrict_scalars.smul_def' {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) +lemma restrict_scalars.smul_def' {R : Type u₁} {S : Type u₂} [ring R] [ring S] (f : R →+* S) {M : Module.{v} S} (r : R) (m : M) : (r • m : (restrict_scalars f).obj M) = (f r • m : M) := rfl @[priority 100] diff --git a/src/algebra/module/pi.lean b/src/algebra/module/pi.lean index 40d470756995d..727bacc31beb3 100644 --- a/src/algebra/module/pi.lean +++ b/src/algebra/module/pi.lean @@ -28,14 +28,14 @@ lemma _root_.is_smul_regular.pi {α : Type*} [Π i, has_smul α $ f i] {k : α} instance smul_with_zero (α) [has_zero α] [Π i, has_zero (f i)] [Π i, smul_with_zero α (f i)] : smul_with_zero α (Π i, f i) := -{ smul_zero := λ _, funext $ λ _, smul_zero' (f _) _, +{ smul_zero := λ _, funext $ λ _, smul_zero _, zero_smul := λ _, funext $ λ _, zero_smul _ _, ..pi.has_smul } instance smul_with_zero' {g : I → Type*} [Π i, has_zero (g i)] [Π i, has_zero (f i)] [Π i, smul_with_zero (g i) (f i)] : smul_with_zero (Π i, g i) (Π i, f i) := -{ smul_zero := λ _, funext $ λ _, smul_zero' (f _) _, +{ smul_zero := λ _, funext $ λ _, smul_zero _, zero_smul := λ _, funext $ λ _, zero_smul _ _, ..pi.has_smul' } diff --git a/src/algebra/module/prod.lean b/src/algebra/module/prod.lean index aa35b68ae03dc..020eb4accd320 100644 --- a/src/algebra/module/prod.lean +++ b/src/algebra/module/prod.lean @@ -18,13 +18,13 @@ namespace prod instance smul_with_zero [has_zero R] [has_zero M] [has_zero N] [smul_with_zero R M] [smul_with_zero R N] : smul_with_zero R (M × N) := -{ smul_zero := λ r, prod.ext (smul_zero' _ _) (smul_zero' _ _), +{ smul_zero := λ r, prod.ext (smul_zero _) (smul_zero _), zero_smul := λ ⟨m, n⟩, prod.ext (zero_smul _ _) (zero_smul _ _), ..prod.has_smul } instance mul_action_with_zero [monoid_with_zero R] [has_zero M] [has_zero N] [mul_action_with_zero R M] [mul_action_with_zero R N] : mul_action_with_zero R (M × N) := -{ smul_zero := λ r, prod.ext (smul_zero' _ _) (smul_zero' _ _), +{ smul_zero := λ r, prod.ext (smul_zero _) (smul_zero _), zero_smul := λ ⟨m, n⟩, prod.ext (zero_smul _ _) (zero_smul _ _), ..prod.mul_action } diff --git a/src/algebra/module/ulift.lean b/src/algebra/module/ulift.lean index 9d4bb1bfc1239..30f010e5c1d8d 100644 --- a/src/algebra/module/ulift.lean +++ b/src/algebra/module/ulift.lean @@ -58,19 +58,34 @@ instance mul_action' [monoid R] [mul_action R M] : mul_action R (ulift M) := { smul := (•), mul_smul := λ r s ⟨f⟩, ext _ _ $ mul_smul _ _ _, - one_smul := λ ⟨f⟩, ext _ _ $ one_smul _ _, - ..ulift.has_smul_left } + one_smul := λ ⟨f⟩, ext _ _ $ one_smul _ _ } + +instance smul_zero_class [has_zero M] [smul_zero_class R M] : + smul_zero_class (ulift R) M := +{ smul_zero := λ _, smul_zero _, + .. ulift.has_smul_left } + +instance smul_zero_class' [has_zero M] [smul_zero_class R M] : + smul_zero_class R (ulift M) := +{ smul_zero := λ c, by { ext, simp [smul_zero], } } + +instance distrib_smul [add_zero_class M] [distrib_smul R M] : + distrib_smul (ulift R) M := +{ smul_add := λ _, smul_add _ } + +instance distrib_smul' [add_zero_class M] [distrib_smul R M] : + distrib_smul R (ulift M) := +{ smul_add := λ c f g, by { ext, simp [smul_add], } } instance distrib_mul_action [monoid R] [add_monoid M] [distrib_mul_action R M] : distrib_mul_action (ulift R) M := -{ smul_zero := λ _, smul_zero _, - smul_add := λ _, smul_add _ } +{ ..ulift.mul_action, + ..ulift.distrib_smul } instance distrib_mul_action' [monoid R] [add_monoid M] [distrib_mul_action R M] : distrib_mul_action R (ulift M) := -{ smul_zero := λ c, by { ext, simp [smul_zero], }, - smul_add := λ c f g, by { ext, simp [smul_add], }, - ..ulift.mul_action' } +{ ..ulift.mul_action', + ..ulift.distrib_smul' } instance mul_distrib_mul_action [monoid R] [monoid M] [mul_distrib_mul_action R M] : mul_distrib_mul_action (ulift R) M := @@ -85,13 +100,13 @@ instance mul_distrib_mul_action' [monoid R] [monoid M] [mul_distrib_mul_action R instance smul_with_zero [has_zero R] [has_zero M] [smul_with_zero R M] : smul_with_zero (ulift R) M := -{ smul_zero := λ _, smul_zero' _ _, +{ smul_zero := λ _, smul_zero _, zero_smul := zero_smul _, ..ulift.has_smul_left } instance smul_with_zero' [has_zero R] [has_zero M] [smul_with_zero R M] : smul_with_zero R (ulift M) := -{ smul_zero := λ _, ulift.ext _ _ $ smul_zero' _ _, +{ smul_zero := λ _, ulift.ext _ _ $ smul_zero _, zero_smul := λ _, ulift.ext _ _ $ zero_smul _ _ } instance mul_action_with_zero [monoid_with_zero R] [has_zero M] [mul_action_with_zero R M] : diff --git a/src/algebra/order/module.lean b/src/algebra/order/module.lean index 191c1b490e5e4..60ab25692c66c 100644 --- a/src/algebra/order/module.lean +++ b/src/algebra/order/module.lean @@ -99,7 +99,7 @@ end lemma smul_nonpos_of_nonpos_of_nonneg (hc : c ≤ 0) (ha : 0 ≤ a) : c • a ≤ 0 := calc c • a ≤ c • 0 : smul_le_smul_of_nonpos ha hc - ... = 0 : smul_zero' M c + ... = 0 : smul_zero c lemma smul_nonneg_of_nonpos_of_nonpos (hc : c ≤ 0) (ha : a ≤ 0) : 0 ≤ c • a := @smul_nonpos_of_nonpos_of_nonneg k Mᵒᵈ _ _ _ _ _ _ hc ha diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index a0e0eb0d9b53f..7153c70c7d762 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -55,7 +55,7 @@ namespace order_dual instance [has_zero R] [add_zero_class M] [h : smul_with_zero R M] : smul_with_zero R Mᵒᵈ := { zero_smul := λ m, order_dual.rec (zero_smul _) m, - smul_zero := λ r, order_dual.rec (smul_zero' _) r, + smul_zero := λ r, order_dual.rec smul_zero r, ..order_dual.has_smul } instance [monoid R] [mul_action R M] : mul_action R Mᵒᵈ := @@ -70,7 +70,7 @@ instance [monoid_with_zero R] [add_monoid M] [mul_action_with_zero R M] : instance [monoid_with_zero R] [add_monoid M] [distrib_mul_action R M] : distrib_mul_action R Mᵒᵈ := { smul_add := λ k a, order_dual.rec (λ a' b, order_dual.rec (smul_add _ _) b) a, - smul_zero := λ r, order_dual.rec smul_zero r } + smul_zero := λ r, order_dual.rec (@smul_zero _ M _ _) r } instance [ordered_semiring R] [ordered_add_comm_monoid M] [smul_with_zero R M] [ordered_smul R M] : @@ -97,7 +97,7 @@ begin end lemma smul_nonneg (hc : 0 ≤ c) (ha : 0 ≤ a) : 0 ≤ c • a := -calc (0 : M) = c • (0 : M) : (smul_zero' M c).symm +calc (0 : M) = c • (0 : M) : (smul_zero c).symm ... ≤ c • a : smul_le_smul_of_nonneg ha hc lemma smul_nonpos_of_nonneg_of_nonpos (hc : 0 ≤ c) (ha : a ≤ 0) : c • a ≤ 0 := @@ -115,7 +115,7 @@ lemma smul_lt_smul_iff_of_pos (hc : 0 < c) : c • a < c • b ↔ a < b := ⟨λ h, lt_of_smul_lt_smul_of_nonneg h hc.le, λ h, smul_lt_smul_of_pos h hc⟩ lemma smul_pos_iff_of_pos (hc : 0 < c) : 0 < c • a ↔ 0 < a := -calc 0 < c • a ↔ c • 0 < c • a : by rw smul_zero' +calc 0 < c • a ↔ c • 0 < c • a : by rw smul_zero ... ↔ 0 < a : smul_lt_smul_iff_of_pos hc alias smul_pos_iff_of_pos ↔ _ smul_pos diff --git a/src/algebra/smul_with_zero.lean b/src/algebra/smul_with_zero.lean index 92cba278685ac..6a0c9db627f03 100644 --- a/src/algebra/smul_with_zero.lean +++ b/src/algebra/smul_with_zero.lean @@ -42,8 +42,7 @@ variables (R M) /-- `smul_with_zero` is a class consisting of a Type `R` with `0 ∈ R` and a scalar multiplication of `R` on a Type `M` with `0`, such that the equality `r • m = 0` holds if at least one among `r` or `m` equals `0`. -/ -class smul_with_zero [has_zero R] [has_zero M] extends has_smul R M := -(smul_zero : ∀ r : R, r • (0 : M) = 0) +class smul_with_zero [has_zero R] [has_zero M] extends smul_zero_class R M := (zero_smul : ∀ m : M, (0 : R) • m = 0) instance mul_zero_class.to_smul_with_zero [mul_zero_class R] : smul_with_zero R R := @@ -61,10 +60,6 @@ variables (R) {M} [has_zero R] [has_zero M] [smul_with_zero R M] @[simp] lemma zero_smul (m : M) : (0 : R) • m = 0 := smul_with_zero.zero_smul m -variables {R} (M) -/-- Note that this lemma has different typeclass assumptions to `smul_zero`. -/ -@[simp] lemma smul_zero' (r : R) : r • (0 : M) = 0 := smul_with_zero.smul_zero r - variables {R M} [has_zero R'] [has_zero M'] [has_smul R M'] /-- Pullback a `smul_with_zero` structure along an injective zero-preserving homomorphism. @@ -85,7 +80,7 @@ protected def function.surjective.smul_with_zero smul_with_zero R M' := { smul := (•), zero_smul := λ m, by { rcases hf m with ⟨x, rfl⟩, simp [←smul] }, - smul_zero := λ c, by simp only [← f.map_zero, ← smul, smul_zero'] } + smul_zero := λ c, by simp only [← f.map_zero, ← smul, smul_zero] } variables (M) @@ -172,7 +167,7 @@ begin obtain rfl | hc := eq_or_ne c 0, { simp only [inv_zero, zero_smul] }, obtain rfl | hx := eq_or_ne x 0, - { simp only [inv_zero, smul_zero'] }, + { simp only [inv_zero, smul_zero] }, { refine inv_eq_of_mul_eq_one_left _, rw [smul_mul_smul, inv_mul_cancel hc, inv_mul_cancel hx, one_smul] } end @@ -184,5 +179,5 @@ end group_with_zero def smul_monoid_with_zero_hom {α β : Type*} [monoid_with_zero α] [mul_zero_one_class β] [mul_action_with_zero α β] [is_scalar_tower α β β] [smul_comm_class α β β] : α × β →*₀ β := -{ map_zero' := smul_zero' _ _, +{ map_zero' := smul_zero _, .. smul_monoid_hom } diff --git a/src/algebra/triv_sq_zero_ext.lean b/src/algebra/triv_sq_zero_ext.lean index 5e8b40b4a58e4..d04a32e95c5aa 100644 --- a/src/algebra/triv_sq_zero_ext.lean +++ b/src/algebra/triv_sq_zero_ext.lean @@ -209,7 +209,7 @@ ext neg_zero.symm rfl @[simp] lemma inr_smul [has_zero R] [has_zero S] [smul_with_zero S R] [has_smul S M] (r : S) (m : M) : (inr (r • m) : tsze R M) = r • inr m := -ext (smul_zero' _ _).symm rfl +ext (smul_zero _).symm rfl end diff --git a/src/analysis/convex/star.lean b/src/analysis/convex/star.lean index 02e452ba5af2d..eab8812f833ba 100644 --- a/src/analysis/convex/star.lean +++ b/src/analysis/convex/star.lean @@ -301,9 +301,9 @@ lemma star_convex_zero_iff : star_convex 𝕜 0 s ↔ ∀ ⦃x : E⦄, x ∈ s → ∀ ⦃a : 𝕜⦄, 0 ≤ a → a ≤ 1 → a • x ∈ s := begin refine forall_congr (λ x, forall_congr $ λ hx, ⟨λ h a ha₀ ha₁, _, λ h a b ha hb hab, _⟩), - { simpa only [sub_add_cancel, eq_self_iff_true, forall_true_left, zero_add, smul_zero'] using + { simpa only [sub_add_cancel, eq_self_iff_true, forall_true_left, zero_add, smul_zero] using h (sub_nonneg_of_le ha₁) ha₀ }, - { rw [smul_zero', zero_add], + { rw [smul_zero, zero_add], exact h hb (by { rw ←hab, exact le_add_of_nonneg_left ha }) } end diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index 2a2078225347a..a33a2ee6f9404 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -1012,7 +1012,7 @@ lemma zero_smul_finset_subset (s : finset β) : (0 : α) • s ⊆ 0 := image_subset_iff.2 $ λ x _, mem_zero.2 $ zero_smul α x lemma zero_mem_smul_finset {t : finset β} {a : α} (h : (0 : β) ∈ t) : (0 : β) ∈ a • t := -mem_smul_finset.2 ⟨0, h, smul_zero' _ _⟩ +mem_smul_finset.2 ⟨0, h, smul_zero _⟩ variables [no_zero_smul_divisors α β] {a : α} diff --git a/src/data/real/pointwise.lean b/src/data/real/pointwise.lean index ba7b6512dacc0..f3b74ce52f747 100644 --- a/src/data/real/pointwise.lean +++ b/src/data/real/pointwise.lean @@ -31,14 +31,14 @@ variables [mul_action_with_zero α ℝ] [ordered_smul α ℝ] {a : α} lemma real.Inf_smul_of_nonneg (ha : 0 ≤ a) (s : set ℝ) : Inf (a • s) = a • Inf s := begin obtain rfl | hs := s.eq_empty_or_nonempty, - { rw [smul_set_empty, real.Inf_empty, smul_zero'] }, + { rw [smul_set_empty, real.Inf_empty, smul_zero] }, obtain rfl | ha' := ha.eq_or_lt, { rw [zero_smul_set hs, zero_smul], exact cInf_singleton 0 }, by_cases bdd_below s, { exact ((order_iso.smul_left ℝ ha').map_cInf' hs h).symm }, { rw [real.Inf_of_not_bdd_below (mt (bdd_below_smul_iff_of_pos ha').1 h), - real.Inf_of_not_bdd_below h, smul_zero'] } + real.Inf_of_not_bdd_below h, smul_zero] } end lemma real.smul_infi_of_nonneg (ha : 0 ≤ a) (f : ι → ℝ) : @@ -48,14 +48,14 @@ lemma real.smul_infi_of_nonneg (ha : 0 ≤ a) (f : ι → ℝ) : lemma real.Sup_smul_of_nonneg (ha : 0 ≤ a) (s : set ℝ) : Sup (a • s) = a • Sup s := begin obtain rfl | hs := s.eq_empty_or_nonempty, - { rw [smul_set_empty, real.Sup_empty, smul_zero'] }, + { rw [smul_set_empty, real.Sup_empty, smul_zero] }, obtain rfl | ha' := ha.eq_or_lt, { rw [zero_smul_set hs, zero_smul], exact cSup_singleton 0 }, by_cases bdd_above s, { exact ((order_iso.smul_left ℝ ha').map_cSup' hs h).symm }, { rw [real.Sup_of_not_bdd_above (mt (bdd_above_smul_iff_of_pos ha').1 h), - real.Sup_of_not_bdd_above h, smul_zero'] } + real.Sup_of_not_bdd_above h, smul_zero] } end lemma real.smul_supr_of_nonneg (ha : 0 ≤ a) (f : ι → ℝ) : @@ -70,14 +70,14 @@ variables [module α ℝ] [ordered_smul α ℝ] {a : α} lemma real.Inf_smul_of_nonpos (ha : a ≤ 0) (s : set ℝ) : Inf (a • s) = a • Sup s := begin obtain rfl | hs := s.eq_empty_or_nonempty, - { rw [smul_set_empty, real.Inf_empty, real.Sup_empty, smul_zero'] }, + { rw [smul_set_empty, real.Inf_empty, real.Sup_empty, smul_zero] }, obtain rfl | ha' := ha.eq_or_lt, { rw [zero_smul_set hs, zero_smul], exact cInf_singleton 0 }, by_cases bdd_above s, { exact ((order_iso.smul_left_dual ℝ ha').map_cSup' hs h).symm }, { rw [real.Inf_of_not_bdd_below (mt (bdd_below_smul_iff_of_neg ha').1 h), - real.Sup_of_not_bdd_above h, smul_zero'] } + real.Sup_of_not_bdd_above h, smul_zero] } end lemma real.smul_supr_of_nonpos (ha : a ≤ 0) (f : ι → ℝ) : diff --git a/src/data/set/pointwise.lean b/src/data/set/pointwise.lean index a10504bc3d005..a7a5e028cf786 100644 --- a/src/data/set/pointwise.lean +++ b/src/data/set/pointwise.lean @@ -1229,7 +1229,7 @@ lemma subsingleton_zero_smul_set (s : set β) : ((0 : α) • s).subsingleton := subsingleton_singleton.anti $ zero_smul_set_subset s lemma zero_mem_smul_set {t : set β} {a : α} (h : (0 : β) ∈ t) : (0 : β) ∈ a • t := -⟨0, h, smul_zero' _ _⟩ +⟨0, h, smul_zero _⟩ variables [no_zero_smul_divisors α β] {a : α} @@ -1242,7 +1242,7 @@ begin { exact or.inr ⟨hb, a, ha⟩ } }, { rintro (⟨hs, b, hb⟩ | ⟨ht, a, ha⟩), { exact ⟨0, b, hs, hb, zero_smul _ _⟩ }, - { exact ⟨a, 0, ha, ht, smul_zero' _ _⟩ } } + { exact ⟨a, 0, ha, ht, smul_zero _⟩ } } end lemma zero_mem_smul_set_iff (ha : a ≠ 0) : (0 : β) ∈ a • t ↔ (0 : β) ∈ t := diff --git a/src/group_theory/group_action/defs.lean b/src/group_theory/group_action/defs.lean index a11a76b7d6648..82865e043cd57 100644 --- a/src/group_theory/group_action/defs.lean +++ b/src/group_theory/group_action/defs.lean @@ -490,20 +490,44 @@ lemma is_scalar_tower.of_smul_one_mul {M N} [monoid N] [has_smul M N] end compatible_scalar +/-- Typeclass for scalar multiplication that preserves `0` on the right. -/ +class smul_zero_class (M A : Type*) [has_zero A] extends has_smul M A := +(smul_zero : ∀ (a : M), a • (0 : A) = 0) + +@[simp] theorem smul_zero [has_zero A] [smul_zero_class M A] (a : M) : a • (0 : A) = 0 := +smul_zero_class.smul_zero _ + +/-- Typeclass for scalar multiplication that preserves `0` and `+` on the right. + +This is exactly `distrib_mul_action` without the `mul_action` part. +-/ +@[ext] class distrib_smul (M A : Type*) [add_zero_class A] + extends smul_zero_class M A := +(smul_add : ∀ (a : M) (x y : A), a • (x + y) = a • x + a • y) + +theorem smul_add [add_zero_class A] [distrib_smul M A] (a : M) (b₁ b₂ : A) : + a • (b₁ + b₂) = a • b₁ + a • b₂ := +distrib_smul.smul_add _ _ _ + /-- Typeclass for multiplicative actions on additive structures. This generalizes group modules. -/ -@[ext] class distrib_mul_action (M : Type*) (A : Type*) [monoid M] [add_monoid A] +@[ext] class distrib_mul_action (M A : Type*) [monoid M] [add_monoid A] extends mul_action M A := -(smul_add : ∀(r : M) (x y : A), r • (x + y) = r • x + r • y) -(smul_zero : ∀(r : M), r • (0 : A) = 0) +(smul_zero : ∀ (a : M), a • (0 : A) = 0) +(smul_add : ∀ (a : M) (x y : A), a • (x + y) = a • x + a • y) section + variables [monoid M] [add_monoid A] [distrib_mul_action M A] -theorem smul_add (a : M) (b₁ b₂ : A) : a • (b₁ + b₂) = a • b₁ + a • b₂ := -distrib_mul_action.smul_add _ _ _ +@[priority 100] -- See note [lower instance priority] +instance distrib_mul_action.to_distrib_smul : distrib_smul M A := +{ ..‹distrib_mul_action M A› } -@[simp] theorem smul_zero (a : M) : a • (0 : A) = 0 := -distrib_mul_action.smul_zero _ +/-! Since Lean 3 does not have definitional eta for structures, we have to make sure +that the definition of `distrib_mul_action.to_distrib_smul` was done correctly, +and the two paths from `distrib_mul_action` to `has_smul` are indeed definitionally equal. -/ +example : (distrib_mul_action.to_mul_action.to_has_smul : has_smul M A) = + distrib_mul_action.to_distrib_smul.to_has_smul := rfl /-- Pullback a distributive multiplicative action along an injective additive monoid homomorphism. diff --git a/src/group_theory/group_action/group.lean b/src/group_theory/group_action/group.lean index e09d120823255..916cc6f5542f8 100644 --- a/src/group_theory/group_action/group.lean +++ b/src/group_theory/group_action/group.lean @@ -199,10 +199,10 @@ section gwz variables [group_with_zero α] [add_monoid β] [distrib_mul_action α β] theorem smul_eq_zero_iff_eq' {a : α} (ha : a ≠ 0) {x : β} : a • x = 0 ↔ x = 0 := -smul_eq_zero_iff_eq (units.mk0 a ha) +show units.mk0 a ha • x = 0 ↔ x = 0, from smul_eq_zero_iff_eq _ theorem smul_ne_zero_iff_ne' {a : α} (ha : a ≠ 0) {x : β} : a • x ≠ 0 ↔ x ≠ 0 := -smul_ne_zero_iff_ne (units.mk0 a ha) +show units.mk0 a ha • x ≠ 0 ↔ x ≠ 0, from smul_ne_zero_iff_ne _ end gwz @@ -279,7 +279,7 @@ variables [monoid α] [add_monoid β] [distrib_mul_action α β] @[simp] theorem smul_eq_zero {u : α} (hu : is_unit u) {x : β} : u • x = 0 ↔ x = 0 := -exists.elim hu $ λ u hu, hu ▸ smul_eq_zero_iff_eq u +exists.elim hu $ λ u hu, hu ▸ show u • x = 0 ↔ x = 0, from smul_eq_zero_iff_eq u end distrib_mul_action diff --git a/src/group_theory/group_action/pi.lean b/src/group_theory/group_action/pi.lean index 09a98569c6953..daf2aeb6f9ece 100644 --- a/src/group_theory/group_action/pi.lean +++ b/src/group_theory/group_action/pi.lean @@ -105,18 +105,36 @@ instance mul_action' {g : I → Type*} {m : Π i, monoid (f i)} [Π i, mul_actio mul_smul := λ r s f, funext $ λ i, mul_smul _ _ _, one_smul := λ f, funext $ λ i, one_smul _ _ } +instance smul_zero_class (α) {n : ∀ i, has_zero $ f i} + [∀ i, smul_zero_class α $ f i] : + @smul_zero_class α (Π i : I, f i) (@pi.has_zero I f n) := +{ smul_zero := λ c, funext $ λ i, smul_zero _ } + +instance smul_zero_class' {g : I → Type*} {n : Π i, has_zero $ g i} + [Π i, smul_zero_class (f i) (g i)] : + @smul_zero_class (Π i, f i) (Π i : I, g i) (@pi.has_zero I g n) := +{ smul_zero := by { intros, ext x, apply smul_zero } } + +instance distrib_smul (α) {n : ∀ i, add_zero_class $ f i} [∀ i, distrib_smul α $ f i] : + @distrib_smul α (Π i : I, f i) (@pi.add_zero_class I f n) := +{ smul_add := λ c f g, funext $ λ i, smul_add _ _ _ } + +instance distrib_smul' {g : I → Type*} {n : Π i, add_zero_class $ g i} + [Π i, distrib_smul (f i) (g i)] : + @distrib_smul (Π i, f i) (Π i : I, g i) (@pi.add_zero_class I g n) := +{ smul_add := by { intros, ext x, apply smul_add } } + instance distrib_mul_action (α) {m : monoid α} {n : ∀ i, add_monoid $ f i} [∀ i, distrib_mul_action α $ f i] : @distrib_mul_action α (Π i : I, f i) m (@pi.add_monoid I f n) := -{ smul_zero := λ c, funext $ λ i, smul_zero _, - smul_add := λ c f g, funext $ λ i, smul_add _ _ _, - ..pi.mul_action _ } +{ ..pi.mul_action _, + ..pi.distrib_smul _ } instance distrib_mul_action' {g : I → Type*} {m : Π i, monoid (f i)} {n : Π i, add_monoid $ g i} [Π i, distrib_mul_action (f i) (g i)] : @distrib_mul_action (Π i, f i) (Π i : I, g i) (@pi.monoid I f m) (@pi.add_monoid I g n) := -{ smul_add := by { intros, ext x, apply smul_add }, - smul_zero := by { intros, ext x, apply smul_zero } } +{ .. pi.mul_action', + .. pi.distrib_smul' } lemma single_smul {α} [monoid α] [Π i, add_monoid $ f i] [Π i, distrib_mul_action α $ f i] [decidable_eq I] (i : I) (r : α) (x : f i) : diff --git a/src/group_theory/group_action/prod.lean b/src/group_theory/group_action/prod.lean index 9b2a8d66b702a..42b31d7e8e08c 100644 --- a/src/group_theory/group_action/prod.lean +++ b/src/group_theory/group_action/prod.lean @@ -98,10 +98,17 @@ instance is_scalar_tower_both [has_mul N] [has_mul P] [has_smul M N] [has_smul M { mul_smul := λ a₁ a₂ p, mk.inj_iff.mpr ⟨mul_smul _ _ _, mul_smul _ _ _⟩, one_smul := λ ⟨b, c⟩, mk.inj_iff.mpr ⟨one_smul _ _, one_smul _ _⟩ } +instance {R M N : Type*} [has_zero M] [has_zero N] + [smul_zero_class R M] [smul_zero_class R N] : smul_zero_class R (M × N) := +{ smul_zero := λ a, mk.inj_iff.mpr ⟨smul_zero _, smul_zero _⟩ } + +instance {R M N : Type*} [add_zero_class M] [add_zero_class N] + [distrib_smul R M] [distrib_smul R N] : distrib_smul R (M × N) := +{ smul_add := λ a p₁ p₂, mk.inj_iff.mpr ⟨smul_add _ _ _, smul_add _ _ _⟩ } + instance {R M N : Type*} {r : monoid R} [add_monoid M] [add_monoid N] [distrib_mul_action R M] [distrib_mul_action R N] : distrib_mul_action R (M × N) := -{ smul_add := λ a p₁ p₂, mk.inj_iff.mpr ⟨smul_add _ _ _, smul_add _ _ _⟩, - smul_zero := λ a, mk.inj_iff.mpr ⟨smul_zero _, smul_zero _⟩ } +{ ..prod.distrib_smul } instance {R M N : Type*} {r : monoid R} [monoid M] [monoid N] [mul_distrib_mul_action R M] [mul_distrib_mul_action R N] : mul_distrib_mul_action R (M × N) := diff --git a/src/group_theory/group_action/units.lean b/src/group_theory/group_action/units.lean index e686cca619c39..108b93c18e669 100644 --- a/src/group_theory/group_action/units.lean +++ b/src/group_theory/group_action/units.lean @@ -49,9 +49,15 @@ instance [monoid M] [mul_action M α] : mul_action Mˣ α := { one_smul := (one_smul M : _), mul_smul := λ m n, mul_smul (m : M) n, } +instance [monoid M] [has_zero α] [smul_zero_class M α] : smul_zero_class Mˣ α := +{ smul := (•), + smul_zero := λ m, smul_zero m } + +instance [monoid M] [add_zero_class α] [distrib_smul M α] : distrib_smul Mˣ α := +{ smul_add := λ m, smul_add (m : M) } + instance [monoid M] [add_monoid α] [distrib_mul_action M α] : distrib_mul_action Mˣ α := -{ smul_add := λ m, smul_add (m : M), - smul_zero := λ m, smul_zero m, } +{ .. units.distrib_smul } instance [monoid M] [monoid α] [mul_distrib_mul_action M α] : mul_distrib_mul_action Mˣ α := { smul_mul := λ m, smul_mul' (m : M), diff --git a/src/information_theory/hamming.lean b/src/information_theory/hamming.lean index 541c080cd68dc..01ff0c312fa65 100644 --- a/src/information_theory/hamming.lean +++ b/src/information_theory/hamming.lean @@ -162,11 +162,11 @@ by {convert hamming_dist_comp f hf₁, simp_rw hf₂, refl} lemma hamming_norm_smul_le_hamming_norm [has_zero α] [Π i, smul_with_zero α (β i)] {k : α} {x : Π i, β i} : hamming_norm (k • x) ≤ hamming_norm x := -hamming_norm_comp_le_hamming_norm (λ i (c : β i), k • c) (λ i, by simp_rw smul_zero') +hamming_norm_comp_le_hamming_norm (λ i (c : β i), k • c) (λ i, by simp_rw smul_zero) lemma hamming_norm_smul [has_zero α] [Π i, smul_with_zero α (β i)] {k : α} (hk : ∀ i, is_smul_regular (β i) k) (x : Π i, β i) : hamming_norm (k • x) = hamming_norm x := -hamming_norm_comp (λ i (c : β i), k • c) hk (λ i, by simp_rw smul_zero') +hamming_norm_comp (λ i (c : β i), k • c) hk (λ i, by simp_rw smul_zero) end has_zero diff --git a/src/linear_algebra/matrix/diagonal.lean b/src/linear_algebra/matrix/diagonal.lean index 0f1d286f8c3ec..4adf3a8c42c18 100644 --- a/src/linear_algebra/matrix/diagonal.lean +++ b/src/linear_algebra/matrix/diagonal.lean @@ -37,7 +37,7 @@ linear_map.ext $ λ j, mul_vec_diagonal _ _ _ lemma diagonal_comp_std_basis (w : n → R) (i : n) : (diagonal w).to_lin'.comp (linear_map.std_basis R (λ_:n, R) i) = (w i) • linear_map.std_basis R (λ_:n, R) i := -linear_map.ext $ λ x, (diagonal_mul_vec_single w _ _).trans (pi.single_smul' i (w i) _) +linear_map.ext $ λ x, (diagonal_mul_vec_single w _ _).trans (pi.single_smul' i (w i) x) lemma diagonal_to_lin' (w : n → R) : (diagonal w).to_lin' = linear_map.pi (λi, w i • linear_map.proj i) := diff --git a/src/number_theory/legendre_symbol/add_character.lean b/src/number_theory/legendre_symbol/add_character.lean index 68737dcbd10bf..f7fdef8d303e0 100644 --- a/src/number_theory/legendre_symbol/add_character.lean +++ b/src/number_theory/legendre_symbol/add_character.lean @@ -313,7 +313,6 @@ def primitive_zmod_char (n : ℕ+) (F' : Type v) [field F'] (h : (n : F') ≠ 0) primitive_add_char (zmod n) F' := begin haveI : ne_zero ((n : ℕ) : F') := ⟨h⟩, - haveI : ne_zero ((n : ℕ) : cyclotomic_field n F') := ne_zero.of_no_zero_smul_divisors F' _ n, exact { n := n, char := zmod_char n (is_cyclotomic_extension.zeta_pow n F' _), diff --git a/src/order/filter/pointwise.lean b/src/order/filter/pointwise.lean index 0b3e90b529f1c..5ef1d336ca068 100644 --- a/src/order/filter/pointwise.lean +++ b/src/order/filter/pointwise.lean @@ -700,7 +700,7 @@ because `0 * ⊥ ≠ 0`. lemma ne_bot.smul_zero_nonneg (hf : f.ne_bot) : 0 ≤ f • (0 : filter β) := le_smul_iff.2 $ λ t₁ h₁ t₂ h₂, let ⟨a, ha⟩ := hf.nonempty_of_mem h₁ in - ⟨_, _, ha, h₂, smul_zero' _ _⟩ + ⟨_, _, ha, h₂, smul_zero _⟩ lemma ne_bot.zero_smul_nonneg (hg : g.ne_bot) : 0 ≤ (0 : filter α) • g := le_smul_iff.2 $ λ t₁ h₁ t₂ h₂, let ⟨b, hb⟩ := hg.nonempty_of_mem h₂ in ⟨_, _, h₁, hb, zero_smul _ _⟩ diff --git a/src/ring_theory/localization/module.lean b/src/ring_theory/localization/module.lean index bca0672c79990..f815e1b5facf3 100644 --- a/src/ring_theory/localization/module.lean +++ b/src/ring_theory/localization/module.lean @@ -43,7 +43,7 @@ begin choose a g' hg' using is_localization.exist_integer_multiples S s g, letI := λ i, classical.prop_decidable (i ∈ s), specialize hli s (λ i, if hi : i ∈ s then g' i hi else 0) _ i hi, - { rw [← @smul_zero _ M _ _ _ (a : R), ← hg, finset.smul_sum], + { rw [← @smul_zero _ M _ _ (a : R), ← hg, finset.smul_sum], refine finset.sum_congr rfl (λ i hi, _), dsimp only, rw [dif_pos hi, ← is_scalar_tower.algebra_map_smul Rₛ, hg' i hi, smul_assoc], diff --git a/src/tactic/lint/type_classes.lean b/src/tactic/lint/type_classes.lean index db71eccca7d85..0b0efafc654f0 100644 --- a/src/tactic/lint/type_classes.lean +++ b/src/tactic/lint/type_classes.lean @@ -262,7 +262,7 @@ Some instances take quite some time to fail, and we seem to run against the cach https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/odd.20repeated.20type.20class.20search -/ @[linter] meta def linter.fails_quickly : linter := -{ test := fails_quickly 20000, +{ test := fails_quickly 30000, auto_decls := tt, no_errors_found := "No type-class searches timed out.", errors_found := "TYPE CLASS SEARCHES TIMED OUT. diff --git a/src/topology/support.lean b/src/topology/support.lean index 191aae66dafee..bbf5d082b7f87 100644 --- a/src/topology/support.lean +++ b/src/topology/support.lean @@ -228,7 +228,7 @@ end lemma has_compact_support.smul_left' (hf : has_compact_support f') : has_compact_support (f • f') := begin rw [has_compact_support_iff_eventually_eq] at hf ⊢, - refine hf.mono (λ x hx, by simp_rw [pi.smul_apply', hx, pi.zero_apply, smul_zero']) + refine hf.mono (λ x hx, by simp_rw [pi.smul_apply', hx, pi.zero_apply, smul_zero]) end end smul_with_zero From e1783b0f3b1509a1db778551ba6c3ebbe439b02f Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Fri, 23 Sep 2022 17:33:58 +0000 Subject: [PATCH 368/828] doc(group_theory/sylow): update docstring to include card_eq_multiplicity as a main result (#14354) Sylow subgroups have been redefined to be maximal p-subgroups, instead of subgroups of cardinality `p^n` where `n` is the multiplicity of `p` in the cardinality of G. Sylow's first theorem basically proves that these two are equivalent. With the new definition the hard part of Sylow's first theorem is really proving the lemma `sylow.card_eq_multiplicity` and not proving the existence of a Sylow subgroup so I included it as a main result. --- src/group_theory/sylow.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index 279e1f5a59608..f88ccdd975903 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -32,6 +32,8 @@ The Sylow theorems are the following results for every finite group `G` and ever there exists a subgroup of `G` of order `pⁿ`. * `is_p_group.exists_le_sylow`: A generalization of Sylow's first theorem: Every `p`-subgroup is contained in a Sylow `p`-subgroup. +* `sylow.card_eq_multiplicity`: The cardinality of a Sylow group is `p ^ n` + where `n` is the multiplicity of `p` in the group order. * `sylow_conjugate`: A generalization of Sylow's second theorem: If the number of Sylow `p`-subgroups is finite, then all Sylow `p`-subgroups are conjugate. * `card_sylow_modeq_one`: A generalization of Sylow's third theorem: From 8deb2a6852b3d68fd590e0ea37329883317a9501 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Fri, 23 Sep 2022 17:33:59 +0000 Subject: [PATCH 369/828] feat(topology/separation): connected sets in a t1 space are infinite (#16584) Also rename `eq_univ_of_nonempty_clopen` to `is_clopen.eq_univ`. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/topology/connected.lean | 6 +++--- src/topology/path_connected.lean | 2 +- src/topology/separation.lean | 20 ++++++++++++++++++++ 3 files changed, 24 insertions(+), 4 deletions(-) diff --git a/src/topology/connected.lean b/src/topology/connected.lean index 02c407cd88282..509bd405dfcde 100644 --- a/src/topology/connected.lean +++ b/src/topology/connected.lean @@ -809,9 +809,9 @@ theorem is_clopen_iff [preconnected_space α] {s : set α} : is_clopen s ↔ s = h3 h2, by rintro (rfl | rfl); [exact is_clopen_empty, exact is_clopen_univ]⟩ -lemma eq_univ_of_nonempty_clopen [preconnected_space α] {s : set α} - (h : s.nonempty) (h' : is_clopen s) : s = univ := -by { rw is_clopen_iff at h', exact h'.resolve_left h.ne_empty } +lemma is_clopen.eq_univ [preconnected_space α] {s : set α} (h' : is_clopen s) (h : s.nonempty) : + s = univ := +(is_clopen_iff.mp h').resolve_left h.ne_empty lemma frontier_eq_empty_iff [preconnected_space α] {s : set α} : frontier s = ∅ ↔ s = ∅ ∨ s = univ := diff --git a/src/topology/path_connected.lean b/src/topology/path_connected.lean index b751dc4ea436d..e0178b81e2665 100644 --- a/src/topology/path_connected.lean +++ b/src/topology/path_connected.lean @@ -980,7 +980,7 @@ begin { introI hX, rw path_connected_space_iff_eq, use (classical.arbitrary X), - refine eq_univ_of_nonempty_clopen (by simp) ⟨_, _⟩, + refine is_clopen.eq_univ ⟨_, _⟩ (by simp), { rw is_open_iff_mem_nhds, intros y y_in, rcases (path_connected_basis y).ex_mem with ⟨U, ⟨U_in, hU⟩⟩, diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 573ae82e343be..8474385a6b42b 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -654,6 +654,26 @@ begin exact (set.to_finite _).is_closed end +lemma preconnected_space.trivial_of_discrete [preconnected_space α] [discrete_topology α] : + subsingleton α := +begin + rw ←not_nontrivial_iff_subsingleton, + rintro ⟨x, y, hxy⟩, + rw [ne.def, ←mem_singleton_iff, (is_clopen_discrete _).eq_univ $ singleton_nonempty y] at hxy, + exact hxy (mem_univ x) +end + +lemma is_preconnected.infinite_of_nontrivial [t1_space α] {s : set α} (h : is_preconnected s) + (hs : s.nontrivial) : s.infinite := +begin + refine mt (λ hf, (subsingleton_coe s).mp _) (not_subsingleton_iff.mpr hs), + haveI := @discrete_of_t1_of_finite s _ _ hf.to_subtype, + exact @preconnected_space.trivial_of_discrete _ _ (subtype.preconnected_space h) _ +end + +lemma connected_space.infinite [connected_space α] [nontrivial α] [t1_space α] : infinite α := +infinite_univ_iff.mp $ is_preconnected_univ.infinite_of_nontrivial nontrivial_univ + lemma singleton_mem_nhds_within_of_mem_discrete {s : set α} [discrete_topology s] {x : α} (hx : x ∈ s) : {x} ∈ 𝓝[s] x := From 01894cc4b367d646c82a658203d89491c93633a1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 23 Sep 2022 17:34:01 +0000 Subject: [PATCH 370/828] chore(docs): run `bibtool` on `references.bib` (#16606) It seems that this file is frequently extended in the wrong format. This PR just runs the suggested reformatter. At some point we should do this in CI. --- docs/references.bib | 89 +++++++++++++++++++++++---------------------- 1 file changed, 46 insertions(+), 43 deletions(-) diff --git a/docs/references.bib b/docs/references.bib index 96d52b84954e8..fffe6800f3a4e 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -241,7 +241,8 @@ @Book{ borceux-vol3 } @Book{ bosch-guntzer-remmert, - title = {Non-Archimedean Analysis : A Systematic Approach to Rigid Analytic Geometry}, + title = {Non-Archimedean Analysis : A Systematic Approach to Rigid + Analytic Geometry}, author = {S. Bosch and U. G{\"{u}}ntzer and R. Remmert}, series = {Grundlehren der mathematischen Wissenschaften}, volume = {261}, @@ -268,7 +269,7 @@ @Book{ bourbaki1966b year = {1966}, pages = {iv+363}, mrclass = {54-02 (00A05 54-01)}, - mrnumber = {979295}, + mrnumber = {979295} } @Book{ bourbaki1968, @@ -342,13 +343,13 @@ @Book{ bourbaki1987 url = {https://doi.org/10.1007/978-3-642-61715-7} } -@book{ boydVandenberghe2004, - author = {Stephen P. Boyd and Lieven Vandenberghe}, - title = {Convex Optimization}, - publisher = {Cambridge University Press}, - year = {2004}, - isbn = {978-0-521-83378-3}, - url = {https://web.stanford.edu/~boyd/cvxbook/bv_cvxbook.pdf} +@Book{ boydVandenberghe2004, + author = {Stephen P. Boyd and Lieven Vandenberghe}, + title = {Convex Optimization}, + publisher = {Cambridge University Press}, + year = {2004}, + isbn = {978-0-521-83378-3}, + url = {https://web.stanford.edu/~boyd/cvxbook/bv_cvxbook.pdf} } @Book{ cabreragarciarodriguezpalacios2014, @@ -1621,19 +1622,21 @@ @Article{ orosi2018faulhaber zbl = {1411.41023} } -@InCollection{ petridis2014, - author = {Petridis, G.}, - title = {The {Pl{\"u}nnecke}-{Ruzsa} inequality: an overview}, - booktitle = {Combinatorial and additive number theory. Selected papers based on the presentations at the conferences CANT 2011 and 2012, New York, NY, USA, May 2011 and May 2012}, - isbn = {978-1-4939-1600-9; 978-1-4939-1601-6}, - pages = {229--241}, - year = {2014}, - publisher = {New York, NY: Springer}, - language = {English}, - doi = {10.1007/978-1-4939-1601-6_16}, - keywords = {11B30}, - zblath = {6463830}, - zbl = {1371.11029} +@InCollection{ petridis2014, + author = {Petridis, G.}, + title = {The {Pl{\"u}nnecke}-{Ruzsa} inequality: an overview}, + booktitle = {Combinatorial and additive number theory. Selected papers + based on the presentations at the conferences CANT 2011 and + 2012, New York, NY, USA, May 2011 and May 2012}, + isbn = {978-1-4939-1600-9; 978-1-4939-1601-6}, + pages = {229--241}, + year = {2014}, + publisher = {New York, NY: Springer}, + language = {English}, + doi = {10.1007/978-1-4939-1601-6_16}, + keywords = {11B30}, + zblath = {6463830}, + zbl = {1371.11029} } @Article{ phillips1940, @@ -1689,21 +1692,6 @@ @Book{ rudin2006real isbn = {0-07-100276-6} } -@Book{ tao-vu, - author = {Tao, Terence and Vu, Van H.}, - title = {Additive combinatorics}, - fseries = {Cambridge Studies in Advanced Mathematics}, - series = {Camb. Stud. Adv. Math.}, - volume = {105}, - isbn = {0-521-85386-9}, - year = {2006}, - publisher = {Cambridge: Cambridge University Press}, - language = {English}, - keywords = {11-02,05-02,05D10,05D40,11B75,11B13,11N13,11P70,11K31,11P82,28D05,37A45}, - zbmath = {5066399}, - zbl = {1127.11002} -} - @Book{ samuel1967, author = {Samuel, Pierre}, title = {Th\'{e}orie alg\'{e}brique des nombres}, @@ -1851,6 +1839,21 @@ @Article{ Stone1979 mrreviewer = {J. Segal} } +@Book{ tao-vu, + author = {Tao, Terence and Vu, Van H.}, + title = {Additive combinatorics}, + fseries = {Cambridge Studies in Advanced Mathematics}, + series = {Camb. Stud. Adv. Math.}, + volume = {105}, + isbn = {0-521-85386-9}, + year = {2006}, + publisher = {Cambridge: Cambridge University Press}, + language = {English}, + keywords = {11-02,05-02,05D10,05D40,11B75,11B13,11N13,11P70,11K31,11P82,28D05,37A45}, + zbmath = {5066399}, + zbl = {1127.11002} +} + @Book{ tao2010, author = {Tao, Terence}, title = {An Epsilon of Room, I: Real Analysis: Pages from Year @@ -1947,7 +1950,7 @@ @Misc{ wedhorn_adic eprint = {arXiv:1910.05934} } -@book{ weidmann_linear, +@Book{ weidmann_linear, author = {Weidmann, Joachim}, title = {Linear operators in {H}ilbert spaces}, isbn = {0-387-90427-1}, @@ -1956,13 +1959,13 @@ @book{ weidmann_linear note = {Translated from the German by Joseph Sz\"{u}cs}, publisher = {Springer}, year = {1980}, - pages = {xiii+402}, + pages = {xiii+402} } -@misc{ welzl_garter, - author = {Emo Welzl and Bernd G\"{a}rtner}, - title = {Cone Programming}, - url = {https://ti.inf.ethz.ch/ew/lehre/ApproxSDP09/notes/conelp.pdf} +@Misc{ welzl_garter, + author = {Emo Welzl and Bernd G\"{a}rtner}, + title = {Cone Programming}, + url = {https://ti.inf.ethz.ch/ew/lehre/ApproxSDP09/notes/conelp.pdf} } @TechReport{ zaanen1966, From 27b903865678b50ed17a8c783487a0871c220875 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Fri, 23 Sep 2022 17:34:02 +0000 Subject: [PATCH 371/828] feat(algebra/punit_instances): linear_ordered_add_comm_monoid_with_top punit (#16609) --- src/algebra/punit_instances.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/algebra/punit_instances.lean b/src/algebra/punit_instances.lean index e0d96bc9658da..667a2a01d3f8d 100644 --- a/src/algebra/punit_instances.lean +++ b/src/algebra/punit_instances.lean @@ -80,6 +80,11 @@ instance : linear_ordered_cancel_add_comm_monoid punit := { le_of_add_le_add_left := λ _ _ _ _, trivial, .. punit.canonically_ordered_add_monoid, ..punit.linear_order } +instance : linear_ordered_add_comm_monoid_with_top punit := +{ top_add' := λ _, rfl, + ..punit.complete_boolean_algebra, + ..punit.linear_ordered_cancel_add_comm_monoid } + instance : has_smul R punit := { smul := λ _ _, star } From ac01c5383af7f6023df52aaf5a3413a2b081def8 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 23 Sep 2022 20:26:33 +0000 Subject: [PATCH 372/828] chore(data/finset/basic): Scope variables over smaller sections (#16613) Having everything in huge sections makes it harder to adjust typeclasses and variable names locally. --- src/data/finset/basic.lean | 23 +++++++++++++++++++++-- 1 file changed, 21 insertions(+), 2 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index dd6ce0847cc36..d38b2d4313cfd 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -585,7 +585,7 @@ by rw [disj_union_comm, singleton_disj_union] /-! ### insert -/ -section decidable_eq +section insert variables [decidable_eq α] {s t u v : finset α} {a b : α} /-- `insert a s` is the set `{a} ∪ s` containing `a` and the elements of `s`. -/ @@ -756,8 +756,13 @@ begin subtype.coe_mk] }, end +end insert + /-! ### Lattice structure -/ +section lattice +variables [decidable_eq α] {s t u v : finset α} {a b : α} + /-- `s ∪ t` is the set such that `a ∈ s ∪ t` iff `a ∈ s` or `a ∈ t`. -/ instance : has_union (finset α) := ⟨λ s t, ⟨_, t.2.ndunion s.1⟩⟩ @@ -1133,7 +1138,11 @@ by { rw [finset.disjoint_left, set.disjoint_left], refl } s.pairwise_disjoint (λ i, f i : ι → set α) ↔ s.pairwise_disjoint f := forall₅_congr $ λ _ _ _ _ _, disjoint_coe +end lattice + /-! ### erase -/ +section erase +variables [decidable_eq α] {s t u v : finset α} {a b : α} /-- `erase s a` is the set `s - {a}`, that is, the elements of `s` which are not equal to `a`. -/ @@ -1255,8 +1264,13 @@ lemma erase_inj_on (s : finset α) : set.inj_on s.erase s := λ _ _ _ _, (erase_ lemma erase_inj_on' (a : α) : {s : finset α | a ∈ s}.inj_on (λ s, erase s a) := λ s hs t ht (h : s.erase a = _), by rw [←insert_erase hs, ←insert_erase ht, h] +end erase + /-! ### sdiff -/ +section sdiff +variables [decidable_eq α] {s t u v : finset α} {a b : α} + /-- `s \ t` is the set consisting of the elements of `s` that are not in `t`. -/ instance : has_sdiff (finset α) := ⟨λs₁ s₂, ⟨s₁.1 - s₂.1, nodup_of_le tsub_le_self s₁.2⟩⟩ @@ -1409,14 +1423,19 @@ disjoint_of_subset_right (inter_subset_right _ _) sdiff_disjoint lemma sdiff_eq_self_iff_disjoint : s \ t = s ↔ disjoint s t := sdiff_eq_self_iff_disjoint' lemma sdiff_eq_self_of_disjoint (h : disjoint s t) : s \ t = s := sdiff_eq_self_iff_disjoint.2 h +end sdiff + /-! ### Symmetric difference -/ +section symm_diff +variables [decidable_eq α] {s t : finset α} {a b : α} + lemma mem_symm_diff : a ∈ s ∆ t ↔ a ∈ s ∧ a ∉ t ∨ a ∈ t ∧ a ∉ s := by simp_rw [symm_diff, sup_eq_union, mem_union, mem_sdiff] @[simp, norm_cast] lemma coe_symm_diff : (↑(s ∆ t) : set α) = s ∆ t := set.ext $ λ _, mem_symm_diff -end decidable_eq +end symm_diff /-! ### attach -/ From 44eb23a52cc38492aba78f6ad49c54f1e8407f17 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 23 Sep 2022 23:35:54 +0000 Subject: [PATCH 373/828] chore(linear_algebra/matrix/adjugate): remove unnecessary fin_cases with (#16605) This was the only use of `fin_cases ... with ...` in mathlib, and it wasn't even necessary. Removing it to avoid having to port unused tactic features. Co-authored-by: Scott Morrison --- src/linear_algebra/matrix/adjugate.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/linear_algebra/matrix/adjugate.lean b/src/linear_algebra/matrix/adjugate.lean index 9083adee09b76..1cad67fe7358d 100644 --- a/src/linear_algebra/matrix/adjugate.lean +++ b/src/linear_algebra/matrix/adjugate.lean @@ -346,7 +346,7 @@ lemma adjugate_fin_two (A : matrix (fin 2) (fin 2) α) : begin ext i j, rw [adjugate_apply, det_fin_two], - fin_cases i with [0, 1]; fin_cases j with [0, 1]; + fin_cases i; fin_cases j; simp only [nat.one_ne_zero, one_mul, fin.one_eq_zero_iff, pi.single_eq_same, zero_mul, fin.zero_eq_one_iff, sub_zero, pi.single_eq_of_ne, ne.def, not_false_iff, update_row_self, update_row_ne, cons_val_zero, mul_zero, mul_one, zero_sub, cons_val_one, head_cons, of_apply], From 5947fb69cc1fdfebaba1e1b1f0a04f26f0f612bf Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sat, 24 Sep 2022 05:06:37 +0000 Subject: [PATCH 374/828] chore(scripts): update nolints.txt (#16618) I am happy to remove some nolints for you! --- scripts/style-exceptions.txt | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index 0ee7f31fc65e1..cfeaa76b2fa8a 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -40,7 +40,7 @@ src/tactic/interactive.lean : line 10 : ERR_MOD : Module docstring missing, or t src/tactic/local_cache.lean : line 8 : ERR_MOD : Module docstring missing, or too late src/tactic/monotonicity/basic.lean : line 8 : ERR_MOD : Module docstring missing, or too late src/tactic/monotonicity/interactive.lean : line 11 : ERR_MOD : Module docstring missing, or too late -src/tactic/monotonicity/lemmas.lean : line 12 : ERR_MOD : Module docstring missing, or too late +src/tactic/monotonicity/lemmas.lean : line 10 : ERR_MOD : Module docstring missing, or too late src/tactic/noncomm_ring.lean : line 8 : ERR_MOD : Module docstring missing, or too late src/tactic/nth_rewrite/basic.lean : line 8 : ERR_MOD : Module docstring missing, or too late src/tactic/nth_rewrite/congr.lean : line 9 : ERR_MOD : Module docstring missing, or too late From ba80091d8268f1aa9c3e4d5c5b6d1df0dff68b19 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Sat, 24 Sep 2022 14:07:51 +0000 Subject: [PATCH 375/828] feat(data/complex/basic): bundle complex.abs (#16347) Doing this this way round makes the refactor in #16340 much easier. I will follow up this PR with similar PRs for `padic_norm` and `padic_norm_e`. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/algebra/order/absolute_value.lean | 4 + src/analysis/complex/basic.lean | 12 +- src/analysis/complex/circle.lean | 4 +- src/analysis/complex/isometry.lean | 4 +- src/analysis/complex/phragmen_lindelof.lean | 16 +- src/analysis/complex/polynomial.lean | 32 ++-- .../functions_bounded_at_infty.lean | 2 +- .../special_functions/compare_exp.lean | 2 +- .../special_functions/complex/arg.lean | 34 ++-- .../special_functions/complex/log.lean | 8 +- src/analysis/special_functions/gamma.lean | 10 +- .../special_functions/polar_coord.lean | 2 +- src/analysis/special_functions/pow.lean | 25 ++- src/data/complex/basic.lean | 148 +++++++++--------- src/data/complex/exponential.lean | 36 ++--- src/geometry/euclidean/oriented_angle.lean | 2 +- .../integral/circle_integral.lean | 15 +- .../integral/circle_transform.lean | 7 +- src/number_theory/l_series.lean | 17 +- src/number_theory/modular.lean | 2 +- .../polynomial/cyclotomic/eval.lean | 18 ++- 21 files changed, 206 insertions(+), 194 deletions(-) diff --git a/src/algebra/order/absolute_value.lean b/src/algebra/order/absolute_value.lean index e5036eb1c1d9c..3bfce7120fe8a 100644 --- a/src/algebra/order/absolute_value.lean +++ b/src/algebra/order/absolute_value.lean @@ -57,6 +57,8 @@ instance subadditive_hom_class : subadditive_hom_class (absolute_value R S) R S { map_add_le_add := λ f, f.add_le', ..absolute_value.zero_hom_class } +@[simp] lemma coe_mk (f : R →ₙ* S) {h₁ h₂ h₃} : ((absolute_value.mk f h₁ h₂ h₃) : R → S) = f := rfl + /-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun` directly. -/ instance : has_coe_to_fun (absolute_value R S) (λ f, R → S) := fun_like.has_coe_to_fun @@ -68,6 +70,8 @@ protected theorem nonneg (x : R) : 0 ≤ abv x := abv.nonneg' x protected theorem add_le (x y : R) : abv (x + y) ≤ abv x + abv y := abv.add_le' x y @[simp] protected theorem map_mul (x y : R) : abv (x * y) = abv x * abv y := abv.map_mul' x y +protected theorem ne_zero_iff {x : R} : abv x ≠ 0 ↔ x ≠ 0 := abv.eq_zero.not + protected theorem pos {x : R} (hx : x ≠ 0) : 0 < abv x := lt_of_le_of_ne (abv.nonneg x) (ne.symm $ mt abv.eq_zero.mp hx) diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index 13ccea2e8b15c..1e787adf61a56 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -40,14 +40,14 @@ instance : has_norm ℂ := ⟨abs⟩ instance : normed_add_comm_group ℂ := normed_add_comm_group.of_core ℂ -{ norm_eq_zero_iff := λ z, abs_eq_zero, - triangle := abs_add, - norm_neg := abs_neg } +{ norm_eq_zero_iff := λ x, abs.eq_zero, + triangle := abs.add_le, + norm_neg := abs.map_neg } instance : normed_field ℂ := { norm := abs, dist_eq := λ _ _, rfl, - norm_mul' := abs_mul, + norm_mul' := map_mul abs, .. complex.field, .. complex.normed_add_comm_group } instance : densely_normed_field ℂ := @@ -57,8 +57,8 @@ instance : densely_normed_field ℂ := instance {R : Type*} [normed_field R] [normed_algebra R ℝ] : normed_algebra R ℂ := { norm_smul_le := λ r x, begin - rw [norm_eq_abs, norm_eq_abs, ←algebra_map_smul ℝ r x, algebra.smul_def, abs_mul, - ←norm_algebra_map' ℝ r, coe_algebra_map, abs_of_real], + rw [norm_eq_abs, norm_eq_abs, ←algebra_map_smul ℝ r x, algebra.smul_def, map_mul, + ←norm_algebra_map' ℝ r, coe_algebra_map, abs_of_real], refl, end, to_algebra := complex.algebra } diff --git a/src/analysis/complex/circle.lean b/src/analysis/complex/circle.lean index c34547179b959..1b331d9c74203 100644 --- a/src/analysis/complex/circle.lean +++ b/src/analysis/complex/circle.lean @@ -47,7 +47,7 @@ lemma circle_def : ↑circle = {z : ℂ | abs z = 1} := set.ext $ λ z, mem_circ mem_circle_iff_abs.mp z.2 lemma mem_circle_iff_norm_sq {z : ℂ} : z ∈ circle ↔ norm_sq z = 1 := -by rw [mem_circle_iff_abs, complex.abs, real.sqrt_eq_one] +by simp [complex.abs] @[simp] lemma norm_sq_eq_of_mem_circle (z : circle) : norm_sq z = 1 := by simp [norm_sq_eq_abs] @@ -72,7 +72,7 @@ instance : topological_group circle := metric.sphere.topological_group /-- If `z` is a nonzero complex number, then `conj z / z` belongs to the unit circle. -/ @[simps] def circle.of_conj_div_self (z : ℂ) (hz : z ≠ 0) : circle := -⟨conj z / z, mem_circle_iff_abs.2 $ by rw [complex.abs_div, abs_conj, div_self (abs_ne_zero.2 hz)]⟩ +⟨conj z / z, mem_circle_iff_abs.2 $ by rw [map_div₀, abs_conj, div_self (complex.abs.ne_zero hz)]⟩ /-- The map `λ t, exp (t * I)` from `ℝ` to the unit circle in `ℂ`. -/ def exp_map_circle : C(ℝ, circle) := diff --git a/src/analysis/complex/isometry.lean b/src/analysis/complex/isometry.lean index 88de13fee8aa5..47ab817740365 100644 --- a/src/analysis/complex/isometry.lean +++ b/src/analysis/complex/isometry.lean @@ -35,7 +35,7 @@ local notation (name := complex.abs) `|` x `|` := complex.abs x rotation. -/ def rotation : circle →* (ℂ ≃ₗᵢ[ℝ] ℂ) := { to_fun := λ a, - { norm_map' := λ x, show |a * x| = |x|, by rw [complex.abs_mul, abs_coe_circle, one_mul], + { norm_map' := λ x, show |a * x| = |x|, by rw [map_mul, abs_coe_circle, one_mul], ..distrib_mul_action.to_linear_equiv ℝ ℂ a }, map_one' := linear_isometry_equiv.ext $ one_smul _, map_mul' := λ _ _, linear_isometry_equiv.ext $ mul_smul _ _ } @@ -82,7 +82,7 @@ lemma linear_isometry.im_apply_eq_im_or_neg_of_re_apply_eq_re {f : ℂ →ₗᵢ (f z).im = z.im ∨ (f z).im = -z.im := begin have h₁ := f.norm_map z, - simp only [complex.abs, norm_eq_abs] at h₁, + simp only [complex.abs_def, norm_eq_abs] at h₁, rwa [real.sqrt_inj (norm_sq_nonneg _) (norm_sq_nonneg _), norm_sq_apply (f z), norm_sq_apply z, h₂, add_left_cancel_iff, mul_self_eq_mul_self_iff] at h₁, end diff --git a/src/analysis/complex/phragmen_lindelof.lean b/src/analysis/complex/phragmen_lindelof.lean index afcff7594f579..d08247ab107d1 100644 --- a/src/analysis/complex/phragmen_lindelof.lean +++ b/src/analysis/complex/phragmen_lindelof.lean @@ -85,12 +85,12 @@ begin have : ∀ {c₁ c₂ B₁ B₂ : ℝ}, c₁ ≤ c₂ → 0 ≤ B₂ → B₁ ≤ B₂ → (λ z : ℂ, expR (B₁ * (abs z) ^ c₁)) =O[comap complex.abs at_top ⊓ l] (λ z, expR (B₂ * (abs z) ^ c₂)), - { have : ∀ᶠ z : ℂ in comap abs at_top ⊓ l, 1 ≤ abs z, + { have : ∀ᶠ z : ℂ in comap complex.abs at_top ⊓ l, 1 ≤ abs z, from ((eventually_ge_at_top 1).comap _).filter_mono inf_le_left, refine λ c₁ c₂ B₁ B₂ hc hB₀ hB, is_O.of_bound 1 (this.mono $ λ z hz, _), rw [one_mul, real.norm_eq_abs, real.norm_eq_abs, real.abs_exp, real.abs_exp, real.exp_le_exp], exact mul_le_mul hB (real.rpow_le_rpow_of_exponent_le hz hc) - (real.rpow_nonneg_of_nonneg (abs_nonneg _) _) hB₀ }, + (real.rpow_nonneg_of_nonneg (complex.abs.nonneg _) _) hB₀ }, rcases hBf with ⟨cf, hcf, Bf, hOf⟩, rcases hBg with ⟨cg, hcg, Bg, hOg⟩, refine ⟨max cf cg, max_lt hcf hcg, max 0 (max Bf Bg), _⟩, refine (hOf.trans $ this _ _ _).sub (hOg.trans $ this _ _ _), @@ -469,7 +469,7 @@ begin simpa only [mem_re_prod_im, mul_I_re, mul_I_im, neg_lt_zero, mem_Iio] using hw.symm }, refine quadrant_I (hd.comp (differentiable_id.mul_const _).diff_cont_on_cl H) (Exists₃.imp (λ c hc B hO, _) hB) him (λ x hx, _) hz_im hz_re, - { simpa only [(∘), complex.abs_mul, abs_I, mul_one] + { simpa only [(∘), map_mul, abs_I, mul_one] using hO.comp_tendsto ((tendsto_mul_right_cobounded I_ne_zero).inf H.tendsto) }, { rw [comp_app, mul_assoc, I_mul_I, mul_neg_one, ← of_real_neg], exact hre _ (neg_nonpos.2 hx) } @@ -534,7 +534,7 @@ begin refine quadrant_I (hd.comp differentiable_neg.diff_cont_on_cl H) _ (λ x hx, _) (λ x hx, _) hz_re hz_im, { refine Exists₃.imp (λ c hc B hO, _) hB, - simpa only [(∘), complex.abs_neg] + simpa only [(∘), complex.abs.map_neg] using hO.comp_tendsto (tendsto_neg_cobounded.inf H.tendsto) }, { rw [comp_app, ← of_real_neg], exact hre (-x) (neg_nonpos.2 hx) }, @@ -601,7 +601,7 @@ begin refine quadrant_II (hd.comp differentiable_neg.diff_cont_on_cl H) _ (λ x hx, _) (λ x hx, _) hz_re hz_im, { refine Exists₃.imp (λ c hc B hO, _) hB, - simpa only [(∘), complex.abs_neg] + simpa only [(∘), complex.abs.map_neg] using hO.comp_tendsto (tendsto_neg_cobounded.inf H.tendsto) }, { rw [comp_app, ← of_real_neg], exact hre (-x) (neg_nonneg.2 hx) }, @@ -814,7 +814,7 @@ begin ... = abs z ^ (1 : ℝ) : (real.rpow_one _).symm ... ≤ abs z ^ (max c 1) : real.rpow_le_rpow_of_exponent_le hr (le_max_right _ _) }, { exact mul_le_mul (le_max_left _ _) (real.rpow_le_rpow_of_exponent_le hr (le_max_left _ _)) - (real.rpow_nonneg_of_nonneg (abs_nonneg _) _) (le_max_right _ _) } }, + (real.rpow_nonneg_of_nonneg (complex.abs.nonneg _) _) (le_max_right _ _) } }, { rw tendsto_zero_iff_norm_tendsto_zero, simp only [hg], exact hre n }, { rw [hg, of_real_mul_re, I_re, mul_zero, real.exp_zero, one_pow, one_mul], @@ -845,7 +845,7 @@ begin suffices : eq_on (f - g) 0 {z : ℂ | 0 ≤ z.re}, by simpa only [eq_on, pi.sub_apply, pi.zero_apply, sub_eq_zero] using this, refine eq_zero_on_right_half_plane_of_superexponential_decay (hfd.sub hgd) _ hre _, - { set l : filter ℂ := comap abs at_top ⊓ 𝓟 {z : ℂ | 0 < z.re}, + { set l : filter ℂ := comap complex.abs at_top ⊓ 𝓟 {z : ℂ | 0 < z.re}, suffices : ∀ {c₁ c₂ B₁ B₂ : ℝ}, c₁ ≤ c₂ → B₁ ≤ B₂ → 0 ≤ B₂ → (λ z, expR (B₁ * abs z ^ c₁)) =O[l] (λ z, expR (B₂ * abs z ^ c₂)), { rcases hfexp with ⟨cf, hcf, Bf, hOf⟩, rcases hgexp with ⟨cg, hcg, Bg, hOg⟩, @@ -857,7 +857,7 @@ begin refine is_O.of_bound 1 (this.mono $ λ z hz, _), simp only [real.norm_of_nonneg (real.exp_pos _).le, real.exp_le_exp, one_mul], exact mul_le_mul hB (real.rpow_le_rpow_of_exponent_le hz hc) - (real.rpow_nonneg_of_nonneg (abs_nonneg _) _) hB₂ }, + (real.rpow_nonneg_of_nonneg (complex.abs.nonneg _) _) hB₂ }, { rcases hfim with ⟨Cf, hCf⟩, rcases hgim with ⟨Cg, hCg⟩, exact ⟨Cf + Cg, λ x, norm_sub_le_of_le (hCf x) (hCg x)⟩ } end diff --git a/src/analysis/complex/polynomial.lean b/src/analysis/complex/polynomial.lean index 5065f84e1f585..af6f3afcafdbd 100644 --- a/src/analysis/complex/polynomial.lean +++ b/src/analysis/complex/polynomial.lean @@ -15,7 +15,7 @@ This file proves that every nonconstant complex polynomial has a root. As a consequence, the complex numbers are algebraically closed. -/ -open complex polynomial metric filter is_absolute_value set +open complex polynomial metric filter set open_locale classical namespace complex @@ -37,10 +37,10 @@ have hg : g * (X - C z₀) ^ n = f - C (f.eval z₀), from div_by_monic_mul_pow_root_multiplicity_eq _ _, have hn0 : n ≠ 0, from λ hn0, by simpa [g, hn0] using hg0, let ⟨δ', hδ'₁, hδ'₂⟩ := continuous_iff.1 (polynomial.continuous g) z₀ - ((g.eval z₀).abs) (complex.abs_pos.2 hg0) in + ((g.eval z₀).abs) (abs.pos hg0) in let δ := min (min (δ' / 2) 1) (((f.eval z₀).abs / (g.eval z₀).abs) / 2) in -have hf0' : 0 < (f.eval z₀).abs, from complex.abs_pos.2 hf0, -have hg0' : 0 < abs (eval z₀ g), from complex.abs_pos.2 hg0, +have hf0' : 0 < (f.eval z₀).abs, from abs.pos hf0, +have hg0' : 0 < (eval z₀ g).abs, from abs.pos hg0, have hfg0 : 0 < (f.eval z₀).abs / abs (eval z₀ g), from div_pos hf0' hg0', have hδ0 : 0 < δ, from lt_min (lt_min (half_pos hδ'₁) (by norm_num)) (half_pos hfg0), have hδ : ∀ z : ℂ, abs (z - z₀) = δ → abs (g.eval z - g.eval z₀) < (g.eval z₀).abs, @@ -65,31 +65,31 @@ have hF₂ : (F.eval z').abs = (f.eval z₀).abs - (g.eval z₀).abs * δ ^ n, from calc (F.eval z').abs = (f.eval z₀ - f.eval z₀ * (g.eval z₀).abs * δ ^ n / (f.eval z₀).abs).abs : congr_arg abs hF₁ ... = abs (f.eval z₀) * complex.abs (1 - (g.eval z₀).abs * δ ^ n / - (f.eval z₀).abs : ℝ) : by rw [← complex.abs_mul]; + (f.eval z₀).abs : ℝ) : by rw [←map_mul]; exact congr_arg complex.abs - (by simp [mul_add, add_mul, mul_assoc, div_eq_mul_inv, sub_eq_add_neg]) + (by simp only [mul_add, mul_assoc, div_eq_mul_inv, sub_eq_add_neg, of_real_add, mul_one, + of_real_one, of_real_neg, of_real_mul, of_real_pow, of_real_inv, mul_neg]) ... = _ : by rw [complex.abs_of_nonneg (sub_nonneg.2 (le_of_lt hδs)), mul_sub, mul_div_cancel' _ (ne.symm (ne_of_lt hf0')), mul_one], have hef0 : abs (eval z₀ g) * (eval z₀ f).abs ≠ 0, - from mul_ne_zero (mt complex.abs_eq_zero.1 hg0) (mt complex.abs_eq_zero.1 hf0), + from mul_ne_zero (abs.ne_zero hg0) (abs.ne_zero hf0), have hz'z₀ : abs (z' - z₀) = δ, - by simp [z', mul_assoc, mul_left_comm _ (_ ^ n), mul_comm _ (_ ^ n), - mul_comm (eval z₀ f).abs, _root_.mul_div_cancel _ hef0, of_real_mul, - neg_mul, neg_div, is_absolute_value.abv_pow complex.abs, - complex.abs_of_nonneg hδ0.le, real.pow_nat_rpow_nat_inv hδ0.le hn0], + by simp only [z', mul_assoc, mul_left_comm _ (_ ^ n), mul_comm _ (_ ^ n), mul_comm (eval _ f).abs, + _root_.mul_div_cancel _ hef0, of_real_mul, neg_mul, neg_div, map_pow, abs_of_real, + add_sub_cancel, abs_cpow_inv_nat, absolute_value.map_neg, map_div₀, map_mul, + abs_abs, complex.abs_of_nonneg hδ0.le, real.pow_nat_rpow_nat_inv hδ0.le hn0], have hF₃ : (f.eval z' - F.eval z').abs < (g.eval z₀).abs * δ ^ n, from calc (f.eval z' - F.eval z').abs = (g.eval z' - g.eval z₀).abs * (z' - z₀).abs ^ n : - by rw [← eq_sub_iff_add_eq.1 hg, ← is_absolute_value.abv_pow complex.abs, - ← complex.abs_mul, sub_mul]; - simp [F, eval_pow, eval_add, eval_mul, eval_sub, eval_C, eval_X, eval_neg, add_sub_cancel, - sub_eq_add_neg, add_assoc] + by rw [← eq_sub_iff_add_eq.1 hg, ←map_pow abs, ←map_mul, sub_mul]; + simp only [eval_pow, eval_add, eval_mul, eval_C, eval_X, eval_neg, sub_eq_add_neg, + add_assoc, neg_add_rev, add_neg_cancel_comm_assoc] ... = (g.eval z' - g.eval z₀).abs * δ ^ n : by rw hz'z₀ ... < _ : (mul_lt_mul_right (pow_pos hδ0 _)).2 (hδ _ hz'z₀), lt_irrefl (f.eval z₀).abs $ calc (f.eval z₀).abs ≤ (f.eval z').abs : hz₀ _ ... = (F.eval z' + (f.eval z' - F.eval z')).abs : by simp - ... ≤ (F.eval z').abs + (f.eval z' - F.eval z').abs : complex.abs_add _ _ + ... ≤ (F.eval z').abs + (f.eval z' - F.eval z').abs : abs.add_le _ _ ... < (f.eval z₀).abs - (g.eval z₀).abs * δ ^ n + (g.eval z₀).abs * δ ^ n : add_lt_add_of_le_of_lt (by rw hF₂) hF₃ ... = (f.eval z₀).abs : sub_add_cancel _ _ diff --git a/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean index d8732563824f1..a56a5361b5f25 100644 --- a/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean +++ b/src/analysis/complex/upper_half_plane/functions_bounded_at_infty.lean @@ -63,7 +63,7 @@ bounded_filter_subalgebra at_im_infty lemma prod_of_bounded_is_bounded {f g : ℍ → ℂ} (hf : is_bounded_at_im_infty f) (hg : is_bounded_at_im_infty g) : is_bounded_at_im_infty (f * g) := -by simpa only [pi.one_apply, mul_one, norm_eq_abs, complex.abs_mul] using hf.mul hg +by simpa only [pi.one_apply, mul_one, norm_eq_abs] using hf.mul hg @[simp] lemma bounded_mem (f : ℍ → ℂ) : is_bounded_at_im_infty f ↔ ∃ (M A : ℝ), ∀ z : ℍ, A ≤ im z → abs (f z) ≤ M := diff --git a/src/analysis/special_functions/compare_exp.lean b/src/analysis/special_functions/compare_exp.lean index 9f7ec46b1b750..653f01558aa33 100644 --- a/src/analysis/special_functions/compare_exp.lean +++ b/src/analysis/special_functions/compare_exp.lean @@ -139,7 +139,7 @@ lemma is_o_cpow_exp (hl : is_exp_cmp_filter l) (a : ℂ) {b : ℝ} (hb : 0 < b) (λ z, z ^ a) =o[l] (λ z, exp (b * z)) := calc (λ z, z ^ a) =Θ[l] λ z, abs z ^ re a : is_Theta_cpow_const_rpow $ λ _ _, hl.eventually_ne ... =ᶠ[l] λ z, real.exp (re a * real.log (abs z)) : hl.eventually_ne.mono $ - λ z hz, by simp only [real.rpow_def_of_pos (abs_pos.2 hz), mul_comm] + λ z hz, by simp only [real.rpow_def_of_pos, abs.pos hz, mul_comm] ... =o[l] λ z, exp (b * z) : is_o.of_norm_right $ begin simp only [norm_eq_abs, abs_exp, of_real_mul_re, real.is_o_exp_comp_exp_comp], diff --git a/src/analysis/special_functions/complex/arg.lean b/src/analysis/special_functions/complex/arg.lean index 911713e79f06a..847a93fcd6814 100644 --- a/src/analysis/special_functions/complex/arg.lean +++ b/src/analysis/special_functions/complex/arg.lean @@ -40,9 +40,10 @@ by unfold arg; split_ifs; lemma cos_arg {x : ℂ} (hx : x ≠ 0) : real.cos (arg x) = x.re / x.abs := begin - have habs : 0 < abs x := abs_pos.2 hx, + have habs : 0 < abs x := abs.pos hx, have him : |im x / abs x| ≤ 1, - { rw [_root_.abs_div, abs_abs], exact div_le_one_of_le x.abs_im_le_abs x.abs_nonneg }, + { rw [_root_.abs_div, abs_abs], + exact div_le_one_of_le x.abs_im_le_abs (abs.nonneg x) }, rw abs_le at him, rw arg, split_ifs with h₁ h₂ h₂, { rw [real.cos_arcsin]; field_simp [real.sqrt_sq, habs.le, *] }, @@ -62,7 +63,7 @@ end begin rcases eq_or_ne x 0 with (rfl|hx), { simp }, - { have : abs x ≠ 0 := abs_ne_zero.2 hx, + { have : abs x ≠ 0 := abs.ne_zero hx, ext; field_simp [sin_arg, cos_arg hx, this, mul_comm (abs x)] } end @@ -85,8 +86,7 @@ by { ext x, simp only [mem_sphere_zero_iff_norm, norm_eq_abs, abs_eq_one_iff, me lemma arg_mul_cos_add_sin_mul_I {r : ℝ} (hr : 0 < r) {θ : ℝ} (hθ : θ ∈ Ioc (-π) π) : arg (r * (cos θ + sin θ * I)) = θ := begin - have hπ := real.pi_pos, - simp only [arg, abs_mul, abs_cos_add_sin_mul_I, abs_of_nonneg hr.le, mul_one], + simp only [arg, map_mul, abs_cos_add_sin_mul_I, abs_of_nonneg hr.le, mul_one], simp only [of_real_mul_re, of_real_mul_im, neg_im, ← of_real_cos, ← of_real_sin, ← mk_eq_add_mul_I, neg_div, mul_div_cancel_left _ hr.ne', mul_nonneg_iff_right_nonneg_of_pos hr], @@ -127,7 +127,7 @@ begin rw [← abs_mul_cos_add_sin_mul_I z, ← cos_add_int_mul_two_pi _ N, ← sin_add_int_mul_two_pi _ N], simp only [← of_real_one, ← of_real_bit0, ← of_real_mul, ← of_real_add, ← of_real_int_cast], - rwa [arg_mul_cos_add_sin_mul_I (abs_pos.2 hz) hN] + rwa [arg_mul_cos_add_sin_mul_I (abs.pos hz) hN] end @[simp] lemma range_arg : range arg = Ioc (-π) π := @@ -147,7 +147,7 @@ begin calc 0 ≤ arg z ↔ 0 ≤ real.sin (arg z) : ⟨λ h, real.sin_nonneg_of_mem_Icc ⟨h, arg_le_pi z⟩, by { contrapose!, intro h, exact real.sin_neg_of_neg_of_neg_pi_lt h (neg_pi_lt_arg _) }⟩ - ... ↔ _ : by rw [sin_arg, le_div_iff (abs_pos.2 h₀), zero_mul] + ... ↔ _ : by rw [sin_arg, le_div_iff (abs.pos h₀), zero_mul] end @[simp] lemma arg_neg_iff {z : ℂ} : arg z < 0 ↔ z.im < 0 := @@ -157,16 +157,16 @@ lemma arg_real_mul (x : ℂ) {r : ℝ} (hr : 0 < r) : arg (r * x) = arg x := begin rcases eq_or_ne x 0 with (rfl|hx), { rw mul_zero }, conv_lhs { rw [← abs_mul_cos_add_sin_mul_I x, ← mul_assoc, ← of_real_mul, - arg_mul_cos_add_sin_mul_I (mul_pos hr (abs_pos.2 hx)) x.arg_mem_Ioc] } + arg_mul_cos_add_sin_mul_I (mul_pos hr (abs.pos hx)) x.arg_mem_Ioc] } end lemma arg_eq_arg_iff {x y : ℂ} (hx : x ≠ 0) (hy : y ≠ 0) : arg x = arg y ↔ (abs y / abs x : ℂ) * x = y := begin - simp only [ext_abs_arg_iff, abs_mul, abs_div, abs_of_real, abs_abs, - div_mul_cancel _ (abs_ne_zero.2 hx), eq_self_iff_true, true_and], + simp only [ext_abs_arg_iff, map_mul, map_div₀, abs_of_real, abs_abs, + div_mul_cancel _ (abs.ne_zero hx), eq_self_iff_true, true_and], rw [← of_real_div, arg_real_mul], - exact div_pos (abs_pos.2 hy) (abs_pos.2 hx) + exact div_pos (abs.pos hy) (abs.pos hx) end @[simp] lemma arg_one : arg 1 = 0 := @@ -186,7 +186,7 @@ begin by_cases h : x = 0, { simp only [h, zero_div, complex.zero_im, complex.arg_zero, real.tan_zero, complex.zero_re] }, rw [real.tan_eq_sin_div_cos, sin_arg, cos_arg h, - div_div_div_cancel_right _ (abs_ne_zero.2 h)] + div_div_div_cancel_right _ (abs.ne_zero h)] end lemma arg_of_real_of_nonneg {x : ℝ} (hx : 0 ≤ x) : arg x = 0 := @@ -196,7 +196,7 @@ lemma arg_eq_zero_iff {z : ℂ} : arg z = 0 ↔ 0 ≤ z.re ∧ z.im = 0 := begin refine ⟨λ h, _, _⟩, { rw [←abs_mul_cos_add_sin_mul_I z, h], - simp [abs_nonneg] }, + simp [abs.nonneg] }, { cases z with x y, rintro ⟨h, rfl : y = 0⟩, exact arg_of_real_of_nonneg h } @@ -296,7 +296,7 @@ begin rw [iff_false, not_le, arg_of_re_neg_of_im_nonneg hre him, ← sub_lt_iff_lt_add, half_sub, real.neg_pi_div_two_lt_arcsin, neg_im, neg_div, neg_lt_neg_iff, div_lt_one, ← _root_.abs_of_nonneg him, abs_im_lt_abs], - exacts [hre.ne, abs_pos.2 $ ne_of_apply_ne re hre.ne] }, + exacts [hre.ne, abs.pos $ ne_of_apply_ne re hre.ne] }, { simp only [him], rw [iff_true, arg_of_re_neg_of_im_neg hre him], exact (sub_le_self _ real.pi_pos.le).trans (real.arcsin_le_pi_div_two _) } @@ -314,7 +314,7 @@ begin { simp only [him.not_le], rw [iff_false, not_le, arg_of_re_neg_of_im_neg hre him, sub_lt_iff_lt_add', ← sub_eq_add_neg, sub_half, real.arcsin_lt_pi_div_two, div_lt_one, neg_im, ← abs_of_neg him, abs_im_lt_abs], - exacts [hre.ne, abs_pos.2 $ ne_of_apply_ne re hre.ne] } + exacts [hre.ne, abs.pos $ ne_of_apply_ne re hre.ne] } end @[simp] lemma abs_arg_le_pi_div_two_iff {z : ℂ} : |arg z| ≤ π / 2 ↔ 0 ≤ re z := @@ -434,7 +434,7 @@ by rw [←one_mul (_ + _), ←of_real_one, arg_mul_cos_add_sin_mul_I_coe_angle z lemma arg_mul_coe_angle {x y : ℂ} (hx : x ≠ 0) (hy : y ≠ 0) : (arg (x * y) : real.angle) = arg x + arg y := begin - convert arg_mul_cos_add_sin_mul_I_coe_angle (mul_pos (abs_pos.2 hx) (abs_pos.2 hy)) + convert arg_mul_cos_add_sin_mul_I_coe_angle (mul_pos (abs.pos hx) (abs.pos hy)) (arg x + arg y : real.angle) using 3, simp_rw [←real.angle.coe_add, real.angle.sin_coe, real.angle.cos_coe, of_real_cos, of_real_sin, cos_add_sin_I, of_real_add, add_mul, exp_add, of_real_mul], @@ -496,7 +496,7 @@ lemma arg_eq_nhds_of_im_neg (hz : im z < 0) : lemma continuous_at_arg (h : 0 < x.re ∨ x.im ≠ 0) : continuous_at arg x := begin - have h₀ : abs x ≠ 0, { rw abs_ne_zero, rintro rfl, simpa using h }, + have h₀ : abs x ≠ 0, { rw abs.ne_zero_iff, rintro rfl, simpa using h }, rw [← lt_or_lt_iff_ne] at h, rcases h with (hx_re|hx_im|hx_im), exacts [(real.continuous_at_arcsin.comp (continuous_im.continuous_at.div diff --git a/src/analysis/special_functions/complex/log.lean b/src/analysis/special_functions/complex/log.lean index 8d583cf814b00..24ee1eb4a85b5 100644 --- a/src/analysis/special_functions/complex/log.lean +++ b/src/analysis/special_functions/complex/log.lean @@ -33,9 +33,9 @@ lemma log_im_le_pi (x : ℂ) : (log x).im ≤ π := by simp only [log_im, arg_le lemma exp_log {x : ℂ} (hx : x ≠ 0) : exp (log x) = x := by rw [log, exp_add_mul_I, ← of_real_sin, sin_arg, ← of_real_cos, cos_arg hx, - ← of_real_exp, real.exp_log (abs_pos.2 hx), mul_add, of_real_div, of_real_div, - mul_div_cancel' _ (of_real_ne_zero.2 (mt abs_eq_zero.1 hx)), ← mul_assoc, - mul_div_cancel' _ (of_real_ne_zero.2 (mt abs_eq_zero.1 hx)), re_add_im] + ← of_real_exp, real.exp_log (abs.pos hx), mul_add, of_real_div, of_real_div, + mul_div_cancel' _ (of_real_ne_zero.2 $ abs.ne_zero hx), ← mul_assoc, + mul_div_cancel' _ (of_real_ne_zero.2 $ abs.ne_zero hx), re_add_im] @[simp] lemma range_exp : range exp = {0}ᶜ := set.ext $ λ x, ⟨by { rintro ⟨x, rfl⟩, exact exp_ne_zero x }, λ hx, ⟨log x, exp_log hx⟩⟩ @@ -157,7 +157,7 @@ begin refine continuous_at.add _ _, { refine continuous_of_real.continuous_at.comp _, refine (real.continuous_at_log _).comp complex.continuous_abs.continuous_at, - rw abs_ne_zero, + rw complex.abs.ne_zero_iff, rintro rfl, simpa using h }, { have h_cont_mul : continuous (λ x : ℂ, x * I), from continuous_id'.mul continuous_const, diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index da22edf4b5bb9..a8e3acda7e0f8 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -102,7 +102,7 @@ begin refine has_finite_integral.congr (real.Gamma_integral_convergent hs).2 _, refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), dsimp only, - rw [norm_eq_abs, abs_mul, abs_of_nonneg $ le_of_lt $ exp_pos $ -x, + rw [norm_eq_abs, map_mul, abs_of_nonneg $ le_of_lt $ exp_pos $ -x, abs_cpow_eq_rpow_re_of_pos hx _], simp } end @@ -175,7 +175,7 @@ begin refine (_ : continuous_at (λ x:ℂ, x ^ (s - 1)) _).comp continuous_of_real.continuous_at, apply continuous_at_cpow_const, rw of_real_re, exact or.inl hx.1, }, rw ←has_finite_integral_norm_iff, - simp_rw [norm_eq_abs, complex.abs_mul], + simp_rw [norm_eq_abs, map_mul], refine (((real.Gamma_integral_convergent hs).mono_set Ioc_subset_Ioi_self).has_finite_integral.congr _).const_mul _, rw [eventually_eq, ae_restrict_iff'], @@ -244,7 +244,7 @@ begin have : (λ (e : ℝ), ∥-(e:ℂ) ^ s * (-e).exp∥ ) =ᶠ[at_top] (λ (e : ℝ), e ^ s.re * (-1 * e).exp ), { refine eventually_eq_of_mem (Ioi_mem_at_top 0) _, intros x hx, dsimp only, - rw [norm_eq_abs, abs_mul, abs_neg, abs_cpow_eq_rpow_re_of_pos hx, + rw [norm_eq_abs, map_mul, abs.map_neg, abs_cpow_eq_rpow_re_of_pos hx, abs_of_nonneg (exp_pos(-x)).le, neg_mul, one_mul],}, exact (tendsto_congr' this).mpr (tendsto_rpow_mul_exp_neg_mul_at_top_nhds_0 _ _ zero_lt_one), end @@ -411,7 +411,7 @@ begin rcases le_or_lt 1 x with h|h, { -- case 1 ≤ x refine le_add_of_nonneg_of_le (abs_nonneg _) _, - rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, complex.abs_mul, abs_mul, + rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, map_mul, abs_mul, ←complex.of_real_mul, complex.abs_of_real], refine mul_le_mul_of_nonneg_left _ (abs_nonneg _), rw complex.abs_cpow_eq_rpow_re_of_pos hx, @@ -419,7 +419,7 @@ begin apply rpow_le_rpow_of_exponent_le h, rw [complex.sub_re, complex.one_re], linarith, }, { refine le_add_of_le_of_nonneg _ (abs_nonneg _), - rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, complex.abs_mul, abs_mul, + rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, map_mul, abs_mul, ←complex.of_real_mul, complex.abs_of_real], refine mul_le_mul_of_nonneg_left _ (abs_nonneg _), rw complex.abs_cpow_eq_rpow_re_of_pos hx, diff --git a/src/analysis/special_functions/polar_coord.lean b/src/analysis/special_functions/polar_coord.lean index f2679d0c748f5..8b76bf9a13ab3 100644 --- a/src/analysis/special_functions/polar_coord.lean +++ b/src/analysis/special_functions/polar_coord.lean @@ -69,7 +69,7 @@ It is a homeomorphism between `ℝ^2 - (-∞, 0]` and `(0, +∞) × (-π, π)`. begin rintros ⟨x, y⟩ hxy, have A : sqrt (x ^ 2 + y ^ 2) = complex.abs (x + y * complex.I), - by simp only [complex.abs, complex.norm_sq, pow_two, monoid_with_zero_hom.coe_mk, + by simp only [complex.abs_def, complex.norm_sq, pow_two, monoid_with_zero_hom.coe_mk, complex.add_re, complex.of_real_re, complex.mul_re, complex.I_re, mul_zero, complex.of_real_im, complex.I_im, sub_self, add_zero, complex.add_im, complex.mul_im, mul_one, zero_add], diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 2302dbe222696..e9d02a03a042f 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -404,32 +404,31 @@ begin { rcases eq_or_ne y 0 with rfl|hy; simp * }, have hne : (x : ℂ) ≠ 0, from of_real_ne_zero.mpr hlt.ne, rw [cpow_def_of_ne_zero hne, cpow_def_of_ne_zero (neg_ne_zero.2 hne), ← exp_add, ← add_mul, - log, log, abs_neg, arg_of_real_of_neg hlt, ← of_real_neg, - arg_of_real_of_nonneg (neg_nonneg.2 hx), of_real_zero, zero_mul, add_zero] + log, log, abs.map_neg, arg_of_real_of_neg hlt, ← of_real_neg, + arg_of_real_of_nonneg (neg_nonneg.2 hx), of_real_zero, zero_mul, add_zero] end lemma abs_cpow_of_ne_zero {z : ℂ} (hz : z ≠ 0) (w : ℂ) : abs (z ^ w) = abs z ^ w.re / real.exp (arg z * im w) := by rw [cpow_def_of_ne_zero hz, abs_exp, mul_re, log_re, log_im, real.exp_sub, - real.rpow_def_of_pos (abs_pos.2 hz)] + real.rpow_def_of_pos (abs.pos hz)] lemma abs_cpow_of_imp {z w : ℂ} (h : z = 0 → w.re = 0 → w = 0) : abs (z ^ w) = abs z ^ w.re / real.exp (arg z * im w) := begin - rcases ne_or_eq z 0 with hz|rfl; [exact (abs_cpow_of_ne_zero hz w), rw abs_zero], + rcases ne_or_eq z 0 with hz|rfl; [exact (abs_cpow_of_ne_zero hz w), rw map_zero], cases eq_or_ne w.re 0 with hw hw, { simp [hw, h rfl hw] }, - { rw [real.zero_rpow hw, zero_div, zero_cpow, abs_zero], + { rw [real.zero_rpow hw, zero_div, zero_cpow, map_zero], exact ne_of_apply_ne re hw } end lemma abs_cpow_le (z w : ℂ) : abs (z ^ w) ≤ abs z ^ w.re / real.exp (arg z * im w) := begin - by_cases h : z = 0 → w.re = 0 → w = 0, - { exact (abs_cpow_of_imp h).le }, - { push_neg at h, - rw [h.1, zero_cpow h.2.2, abs_zero], - exact div_nonneg (real.rpow_nonneg_of_nonneg le_rfl _) (real.exp_pos _).le }, + rcases ne_or_eq z 0 with hz|rfl; [exact (abs_cpow_of_ne_zero hz w).le, rw map_zero], + rcases eq_or_ne w 0 with rfl|hw, { simp }, + rw [zero_cpow hw, map_zero], + exact div_nonneg (real.rpow_nonneg_of_nonneg le_rfl _) (real.exp_pos _).le end section @@ -445,8 +444,8 @@ begin refine real.is_Theta_exp_comp_one.2 ⟨π * b, _⟩, rw eventually_map at hb ⊢, refine hb.mono (λ x hx, _), - rw [_root_.abs_mul], - exact mul_le_mul (abs_arg_le_pi _) hx (_root_.abs_nonneg _) real.pi_pos.le + erw [abs_mul], + exact mul_le_mul (abs_arg_le_pi _) hx (abs_nonneg _) real.pi_pos.le end lemma is_O_cpow_rpow (hl : is_bounded_under (≤) l (λ x, |(g x).im|)) : @@ -488,7 +487,7 @@ lemma abs_cpow_eq_rpow_re_of_nonneg {x : ℝ} (hx : 0 ≤ x) {y : ℂ} (hy : re abs (x ^ y) = x ^ re y := begin rcases hx.eq_or_lt with rfl|hlt, - { rw [of_real_zero, zero_cpow, abs_zero, real.zero_rpow hy], + { rw [of_real_zero, zero_cpow, map_zero, real.zero_rpow hy], exact ne_of_apply_ne re hy }, { exact abs_cpow_eq_rpow_re_of_pos hlt y } end diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index b8c91c4bbf0d0..d648b26bb43f3 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -461,10 +461,56 @@ by simp only [sub_conj, of_real_mul, of_real_one, of_real_bit0, mul_right_comm, /-! ### Absolute value -/ +namespace abs_theory +-- We develop enough theory to bundle `abs` into an `absolute_value` before making things public; +-- this is so there's not two versions of it hanging around. + +local notation (name := abs) `abs` z := ((norm_sq z).sqrt) + +private lemma mul_self_abs (z : ℂ) : (abs z) * (abs z) = norm_sq z := +real.mul_self_sqrt (norm_sq_nonneg _) + +private lemma abs_nonneg' (z : ℂ) : 0 ≤ abs z := +real.sqrt_nonneg _ + +lemma abs_conj (z : ℂ) : (abs (conj z)) = abs z := +by simp + +private lemma abs_re_le_abs (z : ℂ) : |z.re| ≤ abs z := +begin + rw [mul_self_le_mul_self_iff (abs_nonneg z.re) (abs_nonneg' _), + abs_mul_abs_self, mul_self_abs], + apply re_sq_le_norm_sq +end + +private lemma re_le_abs (z : ℂ) : z.re ≤ abs z := +(abs_le.1 (abs_re_le_abs _)).2 + +private lemma abs_mul (z w : ℂ) : (abs (z * w)) = (abs z) * abs w := +by rw [norm_sq_mul, real.sqrt_mul (norm_sq_nonneg _)] + +private lemma abs_add (z w : ℂ) : (abs (z + w)) ≤ (abs z) + abs w := +(mul_self_le_mul_self_iff (abs_nonneg' (z + w)) + (add_nonneg (abs_nonneg' z) (abs_nonneg' w))).2 $ +begin + rw [mul_self_abs, add_mul_self_eq, mul_self_abs, mul_self_abs, add_right_comm, norm_sq_add, + add_le_add_iff_left, mul_assoc, mul_le_mul_left (@zero_lt_two ℝ _ _), + ←real.sqrt_mul $ norm_sq_nonneg z, ←norm_sq_conj w, ←map_mul], + exact re_le_abs (z * conj w) +end + /-- The complex absolute value function, defined as the square root of the norm squared. -/ -@[pp_nodot] noncomputable def abs (z : ℂ) : ℝ := (norm_sq z).sqrt +noncomputable def _root_.complex.abs : absolute_value ℂ ℝ := +{ to_fun := λ x, abs x, + map_mul' := abs_mul, + nonneg' := abs_nonneg', + eq_zero' := λ _, (real.sqrt_eq_zero $ norm_sq_nonneg _).trans norm_sq_eq_zero, + add_le' := abs_add } -local notation `abs'` := has_abs.abs +end abs_theory + +lemma abs_def : (abs : ℂ → ℝ) = λ z, (norm_sq z).sqrt := rfl +lemma abs_apply {z : ℂ} : abs z = (norm_sq z).sqrt := rfl @[simp, norm_cast] lemma abs_of_real (r : ℝ) : abs r = |r| := by simp [abs, norm_sq_of_real, real.sqrt_mul_self_eq_abs] @@ -488,56 +534,36 @@ by rw [sq_abs, norm_sq_apply, ← sq, ← sq, add_sub_cancel'] @[simp] lemma sq_abs_sub_sq_im (z : ℂ) : abs z ^ 2 - z.im ^ 2 = z.re ^ 2 := by rw [← sq_abs_sub_sq_re, sub_sub_cancel] -@[simp] lemma abs_zero : abs 0 = 0 := by simp [abs] -@[simp] lemma abs_one : abs 1 = 1 := by simp [abs] @[simp] lemma abs_I : abs I = 1 := by simp [abs] @[simp] lemma abs_two : abs 2 = 2 := calc abs 2 = abs (2 : ℝ) : by rw [of_real_bit0, of_real_one] ... = (2 : ℝ) : abs_of_nonneg (by norm_num) -lemma abs_nonneg (z : ℂ) : 0 ≤ abs z := -real.sqrt_nonneg _ - @[simp] lemma range_abs : range abs = Ici 0 := -subset.antisymm (range_subset_iff.2 abs_nonneg) $ λ x hx, ⟨x, abs_of_nonneg hx⟩ - -@[simp] lemma abs_eq_zero {z : ℂ} : abs z = 0 ↔ z = 0 := -(real.sqrt_eq_zero $ norm_sq_nonneg _).trans norm_sq_eq_zero - -lemma abs_ne_zero {z : ℂ} : abs z ≠ 0 ↔ z ≠ 0 := -not_congr abs_eq_zero - -@[simp] lemma abs_conj (z : ℂ) : abs (conj z) = abs z := -by simp [abs] - -@[simp] lemma abs_mul (z w : ℂ) : abs (z * w) = abs z * abs w := -by rw [abs, norm_sq_mul, real.sqrt_mul (norm_sq_nonneg _)]; refl +subset.antisymm (range_subset_iff.2 abs.nonneg) $ λ x hx, ⟨x, abs_of_nonneg hx⟩ -/-- `complex.abs` as a `monoid_with_zero_hom`. -/ -@[simps] noncomputable def abs_hom : ℂ →*₀ ℝ := -{ to_fun := abs, - map_zero' := abs_zero, - map_one' := abs_one, - map_mul' := abs_mul } +@[simp] lemma abs_conj (z : ℂ) : abs (conj z) = abs z := abs_theory.abs_conj z @[simp] lemma abs_prod {ι : Type*} (s : finset ι) (f : ι → ℂ) : abs (s.prod f) = s.prod (λ i, abs (f i)) := -map_prod abs_hom _ _ +map_prod abs _ _ @[simp] lemma abs_pow (z : ℂ) (n : ℕ) : abs (z ^ n) = abs z ^ n := -map_pow abs_hom z n +map_pow abs z n @[simp] lemma abs_zpow (z : ℂ) (n : ℤ) : abs (z ^ n) = abs z ^ n := -map_zpow₀ abs_hom z n +map_zpow₀ abs z n lemma abs_re_le_abs (z : ℂ) : |z.re| ≤ abs z := -by rw [mul_self_le_mul_self_iff (_root_.abs_nonneg z.re) (abs_nonneg _), - abs_mul_abs_self, mul_self_abs]; - apply re_sq_le_norm_sq +begin + rw [mul_self_le_mul_self_iff (abs_nonneg z.re) (abs.nonneg _), + abs_mul_abs_self, mul_self_abs], + apply re_sq_le_norm_sq +end lemma abs_im_le_abs (z : ℂ) : |z.im| ≤ abs z := -by rw [mul_self_le_mul_self_iff (_root_.abs_nonneg z.im) (abs_nonneg _), +by rw [mul_self_le_mul_self_iff (abs_nonneg z.im) (abs.nonneg _), abs_mul_abs_self, mul_self_abs]; apply im_sq_le_norm_sq @@ -548,67 +574,39 @@ lemma im_le_abs (z : ℂ) : z.im ≤ abs z := (abs_le.1 (abs_im_le_abs _)).2 @[simp] lemma abs_re_lt_abs {z : ℂ} : |z.re| < abs z ↔ z.im ≠ 0 := -by rw [abs, real.lt_sqrt (_root_.abs_nonneg _), norm_sq_apply, _root_.sq_abs, ← sq, - lt_add_iff_pos_right, mul_self_pos] +by rw [abs, absolute_value.coe_mk, mul_hom.coe_mk, real.lt_sqrt (abs_nonneg _), norm_sq_apply, + _root_.sq_abs, ← sq, lt_add_iff_pos_right, mul_self_pos] @[simp] lemma abs_im_lt_abs {z : ℂ} : |z.im| < abs z ↔ z.re ≠ 0 := by simpa using @abs_re_lt_abs (z * I) -/-- -The **triangle inequality** for complex numbers. --/ -lemma abs_add (z w : ℂ) : abs (z + w) ≤ abs z + abs w := -(mul_self_le_mul_self_iff (abs_nonneg _) - (add_nonneg (abs_nonneg _) (abs_nonneg _))).2 $ -begin - rw [mul_self_abs, add_mul_self_eq, mul_self_abs, mul_self_abs, - add_right_comm, norm_sq_add, add_le_add_iff_left, - mul_assoc, mul_le_mul_left (@zero_lt_two ℝ _ _)], - simpa [-mul_re] using re_le_abs (z * conj w) -end - -instance : is_absolute_value abs := -{ abv_nonneg := abs_nonneg, - abv_eq_zero := λ _, abs_eq_zero, - abv_add := abs_add, - abv_mul := abs_mul } -open is_absolute_value - @[simp] lemma abs_abs (z : ℂ) : |(abs z)| = abs z := -_root_.abs_of_nonneg (abs_nonneg _) - -@[simp] lemma abs_pos {z : ℂ} : 0 < abs z ↔ z ≠ 0 := abv_pos abs -@[simp] lemma abs_neg : ∀ z, abs (-z) = abs z := abv_neg abs -lemma abs_sub_comm : ∀ z w, abs (z - w) = abs (w - z) := abv_sub abs -lemma abs_sub_le : ∀ a b c, abs (a - c) ≤ abs (a - b) + abs (b - c) := abv_sub_le abs -@[simp] theorem abs_inv : ∀ z, abs z⁻¹ = (abs z)⁻¹ := abv_inv abs -@[simp] theorem abs_div : ∀ z w, abs (z / w) = abs z / abs w := abv_div abs - -lemma abs_abs_sub_le_abs_sub : ∀ z w, |abs z - abs w| ≤ abs (z - w) := -abs_abv_sub_le_abv_sub abs +_root_.abs_of_nonneg (abs.nonneg _) lemma abs_le_abs_re_add_abs_im (z : ℂ) : abs z ≤ |z.re| + |z.im| := -by simpa [re_add_im] using abs_add z.re (z.im * I) +by simpa [re_add_im] using abs.add_le z.re (z.im * I) lemma abs_le_sqrt_two_mul_max (z : ℂ) : abs z ≤ real.sqrt 2 * max (|z.re|) (|z.im|) := begin cases z with x y, simp only [abs, norm_sq_mk, ← sq], wlog hle : |x| ≤ |y| := le_total (|x|) (|y|) using [x y, y x] tactic.skip, - { calc real.sqrt (x ^ 2 + y ^ 2) ≤ real.sqrt (y ^ 2 + y ^ 2) : + { simp only [absolute_value.coe_mk, mul_hom.coe_mk, norm_sq_mk, ←sq], + calc real.sqrt (x ^ 2 + y ^ 2) ≤ real.sqrt (y ^ 2 + y ^ 2) : real.sqrt_le_sqrt (add_le_add_right (sq_le_sq.2 hle) _) ... = real.sqrt 2 * max (|x|) (|y|) : by rw [max_eq_right hle, ← two_mul, real.sqrt_mul two_pos.le, real.sqrt_sq_eq_abs] }, - { rwa [add_comm, max_comm] } + { dsimp, + rwa [add_comm, max_comm] } end lemma abs_re_div_abs_le_one (z : ℂ) : |z.re / z.abs| ≤ 1 := if hz : z = 0 then by simp [hz, zero_le_one] -else by { simp_rw [_root_.abs_div, abs_abs, div_le_iff (abs_pos.2 hz), one_mul, abs_re_le_abs] } +else by { simp_rw [_root_.abs_div, abs_abs, div_le_iff (abs.pos hz), one_mul, abs_re_le_abs] } lemma abs_im_div_abs_le_one (z : ℂ) : |z.im / z.abs| ≤ 1 := if hz : z = 0 then by simp [hz, zero_le_one] -else by { simp_rw [_root_.abs_div, abs_abs, div_le_iff (abs_pos.2 hz), one_mul, abs_im_le_abs] } +else by { simp_rw [_root_.abs_div, abs_abs, div_le_iff (abs.pos hz), one_mul, abs_im_le_abs] } @[simp, norm_cast] lemma abs_cast_nat (n : ℕ) : abs (n : ℂ) = n := by rw [← of_real_nat_cast, abs_of_nonneg (nat.cast_nonneg n)] @@ -617,7 +615,7 @@ by rw [← of_real_nat_cast, abs_of_nonneg (nat.cast_nonneg n)] by rw [← of_real_int_cast, abs_of_real, int.cast_abs] lemma norm_sq_eq_abs (x : ℂ) : norm_sq x = abs x ^ 2 := -by rw [abs, sq, real.mul_self_sqrt (norm_sq_nonneg _)] +by simp [abs, sq, real.mul_self_sqrt (norm_sq_nonneg _)] /-- We put a partial order on ℂ so that `z ≤ w` exactly if `w - z` is real and nonnegative. @@ -691,6 +689,8 @@ end complex_order /-! ### Cauchy sequences -/ +local notation `abs'` := has_abs.abs + theorem is_cau_seq_re (f : cau_seq ℂ abs) : is_cau_seq abs' (λ n, (f n).re) := λ ε ε0, (f.cauchy ε0).imp $ λ i H j ij, lt_of_le_of_lt (by simpa using abs_re_le_abs (f j - f i)) (H _ ij) @@ -710,7 +710,7 @@ noncomputable def cau_seq_im (f : cau_seq ℂ abs) : cau_seq ℝ abs' := lemma is_cau_seq_abs {f : ℕ → ℂ} (hf : is_cau_seq abs f) : is_cau_seq abs' (abs ∘ f) := λ ε ε0, let ⟨i, hi⟩ := hf ε ε0 in -⟨i, λ j hj, lt_of_le_of_lt (abs_abs_sub_le_abs_sub _ _) (hi j hj)⟩ +⟨i, λ j hj, lt_of_le_of_lt (abs.abs_abv_sub_le_abv_sub _ _) (hi j hj)⟩ /-- The limit of a Cauchy sequence of complex numbers. -/ noncomputable def lim_aux (f : cau_seq ℂ abs) : ℂ := @@ -765,7 +765,7 @@ noncomputable def cau_seq_abs (f : cau_seq ℂ abs) : cau_seq ℝ abs' := lemma lim_abs (f : cau_seq ℂ abs) : lim (cau_seq_abs f) = abs (lim f) := lim_eq_of_equiv_const (λ ε ε0, let ⟨i, hi⟩ := equiv_lim f ε ε0 in -⟨i, λ j hj, lt_of_le_of_lt (abs_abs_sub_le_abs_sub _ _) (hi j hj)⟩) +⟨i, λ j hj, lt_of_le_of_lt (abs.abs_abv_sub_le_abv_sub _ _) (hi j hj)⟩) variables {α : Type*} (s : finset α) diff --git a/src/data/complex/exponential.lean b/src/data/complex/exponential.lean index 5d564c6b500b2..f6103c1abf302 100644 --- a/src/data/complex/exponential.lean +++ b/src/data/complex/exponential.lean @@ -313,16 +313,16 @@ namespace complex lemma is_cau_abs_exp (z : ℂ) : is_cau_seq has_abs.abs (λ n, ∑ m in range n, abs (z ^ m / m!)) := let ⟨n, hn⟩ := exists_nat_gt (abs z) in -have hn0 : (0 : ℝ) < n, from lt_of_le_of_lt (abs_nonneg _) hn, -series_ratio_test n (complex.abs z / n) (div_nonneg (complex.abs_nonneg _) (le_of_lt hn0)) +have hn0 : (0 : ℝ) < n, from lt_of_le_of_lt (abs.nonneg _) hn, +series_ratio_test n (complex.abs z / n) (div_nonneg (abs.nonneg _) (le_of_lt hn0)) (by rwa [div_lt_iff hn0, one_mul]) (λ m hm, by rw [abs_abs, abs_abs, nat.factorial_succ, pow_succ, mul_comm m.succ, nat.cast_mul, ← div_div, mul_div_assoc, - mul_div_right_comm, abs_mul, abs_div, abs_cast_nat]; + mul_div_right_comm, abs.map_mul, map_div₀, abs_cast_nat]; exact mul_le_mul_of_nonneg_right - (div_le_div_of_le_left (abs_nonneg _) hn0 - (nat.cast_le.2 (le_trans hm (nat.le_succ _)))) (abs_nonneg _)) + (div_le_div_of_le_left (abs.nonneg _) hn0 + (nat.cast_le.2 (le_trans hm (nat.le_succ _)))) (abs.nonneg _)) noncomputable theory @@ -1296,17 +1296,17 @@ begin ... ≤ ∑ m in filter (λ k, n ≤ k) (range j), abs x ^ n * (1 / m!) : begin refine sum_le_sum (λ m hm, _), - rw [abs_mul, abv_pow abs, abs_div, abs_cast_nat], + rw [map_mul, map_pow, map_div₀, abs_cast_nat], refine mul_le_mul_of_nonneg_left ((div_le_div_right _).2 _) _, { exact nat.cast_pos.2 (nat.factorial_pos _), }, { rw abv_pow abs, - exact (pow_le_one _ (abs_nonneg _) hx), }, - { exact pow_nonneg (abs_nonneg _) _ }, + exact (pow_le_one _ (abs.nonneg _) hx), }, + { exact pow_nonneg (abs.nonneg _) _ }, end ... = abs x ^ n * (∑ m in (range j).filter (λ k, n ≤ k), (1 / m! : ℝ)) : by simp [abs_mul, abv_pow abs, abs_div, mul_sum.symm] ... ≤ abs x ^ n * (n.succ * (n! * n)⁻¹) : - mul_le_mul_of_nonneg_left (sum_div_factorial_le _ _ hn) (pow_nonneg (abs_nonneg _) _) + mul_le_mul_of_nonneg_left (sum_div_factorial_le _ _ hn) (pow_nonneg (abs.nonneg _) _) end lemma exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / (n.succ) ≤ 1 / 2) : @@ -1322,11 +1322,11 @@ begin calc abs (∑ (i : ℕ) in range k, x ^ (n + i) / ((n + i)! : ℂ)) ≤ ∑ (i : ℕ) in range k, abs (x ^ (n + i) / ((n + i)! : ℂ)) : abv_sum_le_sum_abv _ _ ... ≤ ∑ (i : ℕ) in range k, (abs x) ^ (n + i) / (n + i)! : - by simp only [complex.abs_cast_nat, complex.abs_div, abv_pow abs] + by simp only [complex.abs_cast_nat, map_div₀, abv_pow abs] ... ≤ ∑ (i : ℕ) in range k, (abs x) ^ (n + i) / (n! * n.succ ^ i) : _ ... = ∑ (i : ℕ) in range k, (abs x) ^ (n) / (n!) * ((abs x)^i / n.succ ^ i) : _ ... ≤ abs x ^ n / (↑n!) * 2 : _, - { refine sum_le_sum (λ m hm, div_le_div (pow_nonneg (abs_nonneg x) (n + m)) le_rfl _ _), + { refine sum_le_sum (λ m hm, div_le_div (pow_nonneg (abs.nonneg x) (n + m)) le_rfl _ _), { exact_mod_cast mul_pos n.factorial_pos (pow_pos n.succ_pos _), }, { exact_mod_cast (nat.factorial_mul_pow_le_factorial), }, }, { refine finset.sum_congr rfl (λ _ _, _), @@ -1338,11 +1338,11 @@ begin { transitivity (-1 : ℝ), { linarith }, { simp only [neg_le_sub_iff_le_add, div_pow, nat.cast_succ, le_add_iff_nonneg_left], - exact div_nonneg (pow_nonneg (abs_nonneg x) k) + exact div_nonneg (pow_nonneg (abs.nonneg x) k) (pow_nonneg (add_nonneg n.cast_nonneg zero_le_one) k) } }, { linarith }, { linarith }, }, - { exact div_nonneg (pow_nonneg (abs_nonneg x) n) (nat.cast_nonneg (n!)), }, }, + { exact div_nonneg (pow_nonneg (abs.nonneg x) n) (nat.cast_nonneg (n!)), }, }, end lemma abs_exp_sub_one_le {x : ℂ} (hx : abs x ≤ 1) : @@ -1473,10 +1473,10 @@ calc |cos x - (1 - x ^ 2 / 2)| = abs (complex.cos x - (1 - x ^ 2 / 2)) : end) ... ≤ abs ((complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m!) / 2) + abs ((complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m!) / 2) : - by rw add_div; exact abs_add _ _ + by rw add_div; exact complex.abs.add_le _ _ ... = (abs ((complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m!)) / 2 + abs ((complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m!)) / 2) : - by simp [complex.abs_div] + by simp [map_div₀] ... ≤ ((complex.abs (x * I) ^ 4 * (nat.succ 4 * (4! * (4 : ℕ))⁻¹)) / 2 + (complex.abs (-x * I) ^ 4 * (nat.succ 4 * (4! * (4 : ℕ))⁻¹)) / 2) : add_le_add ((div_le_div_right (by norm_num)).2 (complex.exp_bound (by simpa) dec_trivial)) @@ -1499,10 +1499,10 @@ calc |sin x - (x - x ^ 3 / 6)| = abs (complex.sin x - (x - x ^ 3 / 6)) : end) ... ≤ abs ((complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m!) * I / 2) + abs (-((complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m!) * I) / 2) : - by rw [sub_mul, sub_eq_add_neg, add_div]; exact abs_add _ _ + by rw [sub_mul, sub_eq_add_neg, add_div]; exact complex.abs.add_le _ _ ... = (abs ((complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m!)) / 2 + abs ((complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m!)) / 2) : - by simp [add_comm, complex.abs_div, complex.abs_mul] + by simp [add_comm, map_div₀] ... ≤ ((complex.abs (x * I) ^ 4 * (nat.succ 4 * (4! * (4 : ℕ))⁻¹)) / 2 + (complex.abs (-x * I) ^ 4 * (nat.succ 4 * (4! * (4 : ℕ))⁻¹)) / 2) : add_le_add ((div_le_div_right (by norm_num)).2 (complex.exp_bound (by simpa) dec_trivial)) @@ -1654,7 +1654,7 @@ by rw [← of_real_exp]; exact abs_of_nonneg (le_of_lt (real.exp_pos _)) by rw [exp_mul_I, abs_cos_add_sin_mul_I] lemma abs_exp (z : ℂ) : abs (exp z) = real.exp z.re := -by rw [exp_eq_exp_re_mul_sin_add_cos, abs_mul, abs_exp_of_real, abs_cos_add_sin_mul_I, mul_one] +by rw [exp_eq_exp_re_mul_sin_add_cos, map_mul, abs_exp_of_real, abs_cos_add_sin_mul_I, mul_one] lemma abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = y.re := by rw [abs_exp, abs_exp, real.exp_eq_exp] diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index c803acefa02d3..580adeea1ba38 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -869,7 +869,7 @@ begin by_cases hx : x = 0, { simp [hx] }, by_cases hy : y = 0, { simp [hy] }, rw [oangle, real.angle.cos_coe, complex.cos_arg], swap, { simp [hx, hy] }, - simp_rw [complex.abs_div, ←complex.norm_eq_abs, linear_isometry_equiv.norm_map, complex.div_re, + simp_rw [map_div₀, ←complex.norm_eq_abs, linear_isometry_equiv.norm_map, complex.div_re, ←complex.sq_abs, ←complex.norm_eq_abs, linear_isometry_equiv.norm_map, complex.isometry_of_orthonormal_symm_apply, complex.add_re, complex.add_im, is_R_or_C.I, complex.mul_I_re, complex.mul_I_im, complex.of_real_re, diff --git a/src/measure_theory/integral/circle_integral.lean b/src/measure_theory/integral/circle_integral.lean index 111cfaef670a9..be24aaf99983c 100644 --- a/src/measure_theory/integral/circle_integral.lean +++ b/src/measure_theory/integral/circle_integral.lean @@ -268,12 +268,17 @@ begin (deriv_circle_map_ne_zero hR)).eventually this, filter_upwards [self_mem_nhds_within, mem_nhds_within_of_mem_nhds (ball_mem_nhds _ zero_lt_one)], - simp [dist_eq, sub_eq_zero] { contextual := tt } }, + simp only [dist_eq, sub_eq_zero, mem_compl_iff, mem_singleton_iff, mem_ball, mem_diff, + mem_ball_zero_iff, norm_eq_abs, not_false_iff, and_self, implies_true_iff] + {contextual := tt} }, refine ((((has_deriv_at_circle_map c R θ).is_O_sub).mono inf_le_left).inv_rev (this.mono (λ θ' h₁ h₂, absurd h₂ h₁.2))).trans _, refine is_O.of_bound (|R|)⁻¹ (this.mono $ λ θ' hθ', _), set x := abs (f θ'), - suffices : x⁻¹ ≤ x ^ n, by simpa [inv_mul_cancel_left₀, mt _root_.abs_eq_zero.1 hR], + suffices : x⁻¹ ≤ x ^ n, + by simpa only [inv_mul_cancel_left₀, abs_eq_zero.not.2 hR, norm_eq_abs, map_inv₀, + algebra.id.smul_eq_mul, map_mul, abs_circle_map_zero, abs_I, mul_one, + abs_zpow, ne.def, not_false_iff] using this, have : x ∈ Ioo (0 : ℝ) 1, by simpa [and.comm, x] using hθ', rw ← zpow_neg_one, refine (zpow_strict_anti this.1 this.2).le_iff_le.2 (int.lt_add_one_iff.1 _), exact hn }, @@ -371,7 +376,7 @@ begin interval_integral.norm_integral_le_integral_norm real.two_pi_pos.le ... < ∫ θ in 0..2 * π, R * C : begin - simp only [norm_smul, deriv_circle_map, norm_eq_abs, complex.abs_mul, abs_I, mul_one, + simp only [norm_smul, deriv_circle_map, norm_eq_abs, map_mul, abs_I, mul_one, abs_circle_map_zero, abs_of_pos hR], refine interval_integral.integral_lt_integral_of_continuous_on_of_le_of_exists_lt real.two_pi_pos _ continuous_on_const (λ θ hθ, _) ⟨θ₀, Ioc_subset_Icc_self hmem, _⟩, @@ -513,9 +518,9 @@ lemma has_sum_two_pi_I_cauchy_power_series_integral {f : ℂ → E} {c : ℂ} {R has_sum (λ n : ℕ, ∮ z in C(c, R), (w / (z - c)) ^ n • (z - c)⁻¹ • f z) (∮ z in C(c, R), (z - (c + w))⁻¹ • f z) := begin - have hR : 0 < R := (abs_nonneg w).trans_lt hw, + have hR : 0 < R := (complex.abs.nonneg w).trans_lt hw, have hwR : abs w / R ∈ Ico (0 : ℝ) 1, - from ⟨div_nonneg (abs_nonneg w) hR.le, (div_lt_one hR).2 hw⟩, + from ⟨div_nonneg (complex.abs.nonneg w) hR.le, (div_lt_one hR).2 hw⟩, refine interval_integral.has_sum_integral_of_dominated_convergence (λ n θ, ∥f (circle_map c R θ)∥ * (abs w / R) ^ n) (λ n, _) (λ n, _) _ _ _, { simp only [deriv_circle_map], diff --git a/src/measure_theory/integral/circle_transform.lean b/src/measure_theory/integral/circle_transform.lean index 36b42bcf906b1..7086f0aef089c 100644 --- a/src/measure_theory/integral/circle_transform.lean +++ b/src/measure_theory/integral/circle_transform.lean @@ -160,10 +160,11 @@ begin have hy2: y1 ∈ [0, 2*π], by {convert (Ico_subset_Icc_self hy1), simp [interval_of_le real.two_pi_pos.le]}, have := mul_le_mul (hab ⟨⟨v, y1⟩, ⟨ball_subset_closed_ball (H hv), hy2⟩⟩) - (HX2 (circle_map z R y1) (circle_map_mem_sphere z hR.le y1)) (abs_nonneg _) (abs_nonneg _), - simp_rw hfun, + (HX2 (circle_map z R y1) (circle_map_mem_sphere z hR.le y1)) + (complex.abs.nonneg _) (complex.abs.nonneg _), + simp_rw hfun, simp only [circle_transform_bounding_function, circle_transform_deriv, V, norm_eq_abs, - algebra.id.smul_eq_mul, deriv_circle_map, abs_mul, abs_circle_map_zero, abs_I, mul_one, + algebra.id.smul_eq_mul, deriv_circle_map, map_mul, abs_circle_map_zero, abs_I, mul_one, ←mul_assoc, mul_inv_rev, inv_I, abs_neg, abs_inv, abs_of_real, one_mul, abs_two, abs_pow, mem_ball, gt_iff_lt, subtype.coe_mk, set_coe.forall, mem_prod, mem_closed_ball, and_imp, prod.forall, normed_space.sphere_nonempty, mem_sphere_iff_norm] at *, diff --git a/src/number_theory/l_series.lean b/src/number_theory/l_series.lean index b010d326c6b85..47fd7d2cd9d66 100644 --- a/src/number_theory/l_series.lean +++ b/src/number_theory/l_series.lean @@ -52,17 +52,17 @@ theorem l_series_summable_of_bounded_of_one_lt_real {f : arithmetic_function ℂ begin by_cases h0 : m = 0, { subst h0, - have hf : f = 0 := arithmetic_function.ext (λ n, complex.abs_eq_zero.1 - (le_antisymm (h n) (complex.abs_nonneg _))), + have hf : f = 0 := arithmetic_function.ext (λ n, complex.abs.eq_zero.1 + (le_antisymm (h n) (complex.abs.nonneg _))), simp [hf] }, refine summable_of_norm_bounded (λ (n : ℕ), m / (n ^ z)) _ _, { simp_rw [div_eq_mul_inv], exact (summable_mul_left_iff h0).1 (real.summable_nat_rpow_inv.2 hz) }, { intro n, - have hm : 0 ≤ m := le_trans (complex.abs_nonneg _) (h 0), + have hm : 0 ≤ m := le_trans (complex.abs.nonneg _) (h 0), cases n, { simp [hm, real.zero_rpow (ne_of_gt (lt_trans real.zero_lt_one hz))] }, - simp only [complex.abs_div, complex.norm_eq_abs], + simp only [map_div₀, complex.norm_eq_abs], apply div_le_div hm (h _) (real.rpow_pos_of_pos (nat.cast_pos.2 n.succ_pos) _) (le_of_eq _), rw [complex.abs_cpow_real, complex.abs_cast_nat] } end @@ -109,10 +109,11 @@ begin ext n, simp [n.succ_ne_zero] }, { apply congr rfl, - ext n, - cases n, { simp [h0] }, - simp only [n.succ_ne_zero, one_div, cast_one, nat_coe_apply, complex.abs_cpow_real, inv_inj, - complex.abs_inv, if_false, zeta_apply, complex.norm_eq_abs, complex.abs_of_nat] } + ext ⟨-|n⟩, + { simp [h0] }, + simp only [cast_zero, nat_coe_apply, zeta_apply, succ_ne_zero, if_false, cast_succ, one_div, + complex.norm_eq_abs, map_inv₀, complex.abs_cpow_real, inv_inj, zero_add], + rw [←cast_one, ←cast_add, complex.abs_of_nat, cast_add, cast_one] } end @[simp] theorem l_series_add {f g : arithmetic_function ℂ} {z : ℂ} diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index d8ecd9a6ea780..8a66693a5ff4c 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -62,7 +62,7 @@ we state lemmas in this file without spurious `coe_fn` terms. -/ local attribute [-instance] matrix.special_linear_group.has_coe_to_fun local attribute [-instance] matrix.general_linear_group.has_coe_to_fun -open complex (hiding abs_one abs_two abs_mul abs_add) +open complex (hiding abs_two) open matrix (hiding mul_smul) matrix.special_linear_group upper_half_plane noncomputable theory diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 79c5b3ed040c4..8cd844817de6a 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -199,7 +199,7 @@ begin { refine ⟨ζ, (mem_primitive_roots hn).mpr hζ, _⟩, suffices : ¬ same_ray ℝ (q : ℂ) ζ, { convert lt_norm_sub_of_not_same_ray this; - simp [real.norm_of_nonneg hq.le, hζ.norm'_eq_one hn.ne'] }, + simp only [hζ.norm'_eq_one hn.ne', real.norm_of_nonneg hq.le, complex.norm_real] }, rw complex.same_ray_iff, push_neg, refine ⟨by exact_mod_cast hq.ne', hζ.ne_zero hn.ne', _⟩, @@ -209,7 +209,8 @@ begin linarith [hζ.unique is_primitive_root.one] }, have : ¬eval ↑q (cyclotomic n ℂ) = 0, { erw cyclotomic.eval_apply q n (algebra_map ℝ ℂ), - simpa using (cyclotomic_pos' n hq').ne' }, + simpa only [complex.coe_algebra_map, complex.of_real_eq_zero] + using (cyclotomic_pos' n hq').ne' }, suffices : (units.mk0 (real.to_nnreal (q - 1)) (by simp [hq'])) ^ totient n < units.mk0 (∥(cyclotomic n ℂ).eval q∥₊) (by simp [this]), { simp only [←units.coe_lt_coe, units.coe_pow, units.coe_mk0, ← nnreal.coe_lt_coe, hq'.le, @@ -217,20 +218,21 @@ begin nnreal.coe_pow, real.coe_to_nnreal', max_eq_left, sub_nonneg] at this, convert this, erw [(cyclotomic.eval_apply q n (algebra_map ℝ ℂ)), eq_comm], - simp [cyclotomic_nonneg n hq'.le], }, + simp only [cyclotomic_nonneg n hq'.le, complex.coe_algebra_map, + complex.abs_of_real, abs_eq_self], }, simp only [cyclotomic_eq_prod_X_sub_primitive_roots hζ, eval_prod, eval_C, eval_X, eval_sub, nnnorm_prod, units.mk0_prod], convert finset.prod_lt_prod' _ _, swap, { exact λ _, units.mk0 (real.to_nnreal (q - 1)) (by simp [hq']) }, - { simp [complex.card_primitive_roots] }, + { simp only [complex.card_primitive_roots, prod_const, card_attach] }, { simp only [subtype.coe_mk, finset.mem_attach, forall_true_left, subtype.forall, - ←units.coe_le_coe, ← nnreal.coe_le_coe, complex.abs_nonneg, hq'.le, units.coe_mk0, + ←units.coe_le_coe, ← nnreal.coe_le_coe, complex.abs.nonneg, hq'.le, units.coe_mk0, real.coe_to_nnreal', coe_nnnorm, complex.norm_eq_abs, max_le_iff, tsub_le_iff_right], intros x hx, - simpa using hfor x hx, }, + simpa only [and_true, tsub_le_iff_right] using hfor x hx, }, { simp only [subtype.coe_mk, finset.mem_attach, exists_true_left, subtype.exists, ← nnreal.coe_lt_coe, ← units.coe_lt_coe, units.coe_mk0 _, coe_nnnorm], - simpa [hq'.le] using hex, }, + simpa only [hq'.le, real.coe_to_nnreal', max_eq_left, sub_nonneg] using hex }, end lemma cyclotomic_eval_lt_sub_one_pow_totient {n : ℕ} {q : ℝ} (hn' : 3 ≤ n) (hq' : 1 < q) : @@ -287,7 +289,7 @@ begin swap, { exact λ _, units.mk0 (real.to_nnreal (q + 1)) (by simp; linarith only [hq']) }, { simp [complex.card_primitive_roots], }, { simp only [subtype.coe_mk, finset.mem_attach, forall_true_left, subtype.forall, - ←units.coe_le_coe, ← nnreal.coe_le_coe, complex.abs_nonneg, hq'.le, units.coe_mk0, + ←units.coe_le_coe, ← nnreal.coe_le_coe, complex.abs.nonneg, hq'.le, units.coe_mk0, real.coe_to_nnreal, coe_nnnorm, complex.norm_eq_abs, max_le_iff], intros x hx, have : complex.abs _ ≤ _ := hfor x hx, From 2d915e4ef8f55de94a850f0e5363ba8b25dc4c29 Mon Sep 17 00:00:00 2001 From: damiano Date: Sat, 24 Sep 2022 20:23:47 +0000 Subject: [PATCH 376/828] chore(tactic/{core + compute_degree}): tightening up compute_degree_le (#15649) This PR instantiates meta-variables to fix [this bug](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/debugging.20.60compute_degree_le.60), using code of Floris. I also tightened up the main tactic: it uses focus1 and it has a version that does not throw errors, suitable for iterations. --- src/tactic/compute_degree.lean | 89 +++++++++++++++++----------------- src/tactic/core.lean | 11 +++++ test/compute_degree.lean | 7 +++ 3 files changed, 62 insertions(+), 45 deletions(-) diff --git a/src/tactic/compute_degree.lean b/src/tactic/compute_degree.lean index fcf643a7266cf..f6c9afef1a0d7 100644 --- a/src/tactic/compute_degree.lean +++ b/src/tactic/compute_degree.lean @@ -65,11 +65,8 @@ meta def guess_degree : expr → tactic expr pe ← to_expr ``(@nat_degree %%R %%inst) tt ff, pure $ expr.mk_app pe [e] -/-- `resolve_sum_step e` takes the type of the current goal `e` as input. -It tries to make progress on the goal `e` by reducing it to subgoals. -It assumes that `e` is of the form `f.nat_degree ≤ d`, failing otherwise. - -`resolve_sum_step` progresses into `f` if `f` is +/-- `resolve_sum_step` assumes that the current goal is of the form `f.nat_degree ≤ d`, failing +otherwise. It tries to make progress on the goal by progressing into `f` if `f` is * a sum, difference, opposite, product, or a power; * a monomial; * `C a`; @@ -80,33 +77,33 @@ or of the form `m ≤ n`, where `m n : ℕ`. If `d` is less than `guess_degree f`, this tactic will create unsolvable goals. -/ -meta def resolve_sum_step : expr → tactic unit -| `(nat_degree %%tl ≤ %%tr) := match tl with - | `(%%tl1 + %%tl2) := refine ``((nat_degree_add_le_iff_left _ _ _).mpr _) - | `(%%tl1 - %%tl2) := refine ``((nat_degree_sub_le_iff_left _).mpr _) - | `(%%tl1 * %%tl2) := do [d1, d2] ← [tl1, tl2].mmap guess_degree, - refine ``(nat_degree_mul_le.trans $ (add_le_add _ _).trans (_ : %%d1 + %%d2 ≤ %%tr)) - | `(- %%f) := refine ``((nat_degree_neg _).le.trans _) - | `(X ^ %%n) := refine ``((nat_degree_X_pow_le %%n).trans _) - | (app `(⇑(@monomial %%R %%inst %%n)) x) := refine ``((nat_degree_monomial_le %%x).trans _) - | (app `(⇑C) x) := refine ``((nat_degree_C %%x).le.trans (nat.zero_le %%tr)) - | `(X) := refine ``(nat_degree_X_le.trans _) - | `(has_zero.zero) := refine ``(nat_degree_zero.le.trans (nat.zero_le _)) - | `(has_one.one) := refine ``(nat_degree_one.le.trans (nat.zero_le _)) - | `(bit0 %%a) := refine ``((nat_degree_bit0 %%a).trans _) - | `(bit1 %%a) := refine ``((nat_degree_bit1 %%a).trans _) - | `(%%tl1 ^ %%n) := do - refine ``(nat_degree_pow_le.trans _), - refine ``(dite (%%n = 0) (λ (n0 : %%n = 0), (by simp only [n0, zero_mul, zero_le])) _), - n0 ← get_unused_name "n0" >>= intro, - refine ``((mul_comm _ _).le.trans ((nat.le_div_iff_mul_le' (nat.pos_of_ne_zero %%n0)).mp _)), - lem1 ← to_expr ``(nat.mul_div_cancel _ (nat.pos_of_ne_zero %%n0)) tt ff, - lem2 ← to_expr ``(nat.div_self (nat.pos_of_ne_zero %%n0)) tt ff, - focus1 (refine ``((%%n0 rfl).elim) <|> rewrite_target lem1 <|> rewrite_target lem2) <|> skip - | e := fail!"'{e}' is not supported" - end -| e := fail!("'resolve_sum_step' was called on\n" ++ - "{e}\nbut it expects `f.nat_degree ≤ d`") +meta def resolve_sum_step : tactic unit := do +t ← target >>= instantiate_mvars, +`(nat_degree %%tl ≤ %%tr) ← whnf t reducible | fail!("Goal is not of the form `f.nat_degree ≤ d`"), +match tl with +| `(%%tl1 + %%tl2) := refine ``((nat_degree_add_le_iff_left _ _ _).mpr _) +| `(%%tl1 - %%tl2) := refine ``((nat_degree_sub_le_iff_left _).mpr _) +| `(%%tl1 * %%tl2) := do [d1, d2] ← [tl1, tl2].mmap guess_degree, + refine ``(nat_degree_mul_le.trans $ (add_le_add _ _).trans (_ : %%d1 + %%d2 ≤ %%tr)) +| `(- %%f) := refine ``((nat_degree_neg _).le.trans _) +| `(X ^ %%n) := refine ``((nat_degree_X_pow_le %%n).trans _) +| (app `(⇑(@monomial %%R %%inst %%n)) x) := refine ``((nat_degree_monomial_le %%x).trans _) +| (app `(⇑C) x) := refine ``((nat_degree_C %%x).le.trans (nat.zero_le %%tr)) +| `(X) := refine ``(nat_degree_X_le.trans _) +| `(has_zero.zero) := refine ``(nat_degree_zero.le.trans (nat.zero_le _)) +| `(has_one.one) := refine ``(nat_degree_one.le.trans (nat.zero_le _)) +| `(bit0 %%a) := refine ``((nat_degree_bit0 %%a).trans _) +| `(bit1 %%a) := refine ``((nat_degree_bit1 %%a).trans _) +| `(%%tl1 ^ %%n) := do + refine ``(nat_degree_pow_le.trans _), + refine ``(dite (%%n = 0) (λ (n0 : %%n = 0), (by simp only [n0, zero_mul, zero_le])) _), + n0 ← get_unused_name "n0" >>= intro, + refine ``((mul_comm _ _).le.trans ((nat.le_div_iff_mul_le' (nat.pos_of_ne_zero %%n0)).mp _)), + lem1 ← to_expr ``(nat.mul_div_cancel _ (nat.pos_of_ne_zero %%n0)) tt ff, + lem2 ← to_expr ``(nat.div_self (nat.pos_of_ne_zero %%n0)) tt ff, + focus1 (refine ``((%%n0 rfl).elim) <|> rewrite_target lem1 <|> rewrite_target lem2) <|> skip +| e := fail!"'{e}' is not supported" +end /-- `norm_assum` simply tries `norm_num` and `assumption`. It is used to try to discharge as many as possible of the side-goals of `compute_degree_le`. @@ -130,6 +127,19 @@ meta def eval_guessing (n : ℕ) : expr → tactic ℕ | `(max %%a %%b) := max <$> eval_guessing a <*> eval_guessing b | e := eval_expr' ℕ e <|> pure n +/-- A general description of `compute_degree_le_aux` is in the doc-string of `compute_degree`. +The difference between the two is that `compute_degree_le_aux` makes no effort to close side-goals, +nor fails if the goal does not change. -/ +meta def compute_degree_le_aux : tactic unit := do +try $ refine ``(degree_le_nat_degree.trans (with_bot.coe_le_coe.mpr _)), +`(nat_degree %%tl ≤ %%tr) ← target | + fail "Goal is not of the form\n`f.nat_degree ≤ d` or `f.degree ≤ d`", +expected_deg ← guess_degree tl >>= eval_guessing 0, +deg_bound ← eval_expr' ℕ tr <|> pure expected_deg, +if deg_bound < expected_deg +then fail sformat!"the given polynomial has a term of expected degree\nat least '{expected_deg}'" +else repeat $ resolve_sum_step + end compute_degree namespace interactive @@ -162,19 +172,8 @@ by compute_degree_le ``` -/ meta def compute_degree_le : tactic unit := -do t ← target, - try $ refine ``(degree_le_nat_degree.trans (with_bot.coe_le_coe.mpr _)), - `(nat_degree %%tl ≤ %%tr) ← target | - fail "Goal is not of the form\n`f.nat_degree ≤ d` or `f.degree ≤ d`", - expected_deg ← guess_degree tl >>= eval_guessing 0, - deg_bound ← eval_expr' ℕ tr <|> pure expected_deg, - if deg_bound < expected_deg - then fail sformat!"the given polynomial has a term of expected degree\nat least '{expected_deg}'" - else - repeat $ target >>= resolve_sum_step, - (do gs ← get_goals >>= list.mmap infer_type, - success_if_fail $ gs.mfirst $ unify t) <|> fail "Goal did not change", - try $ any_goals' norm_assum +focus1 $ do check_target_changes compute_degree_le_aux, + try $ any_goals' norm_assum add_tactic_doc { name := "compute_degree_le", diff --git a/src/tactic/core.lean b/src/tactic/core.lean index ddf187cdf0f64..ee8b157133eb7 100644 --- a/src/tactic/core.lean +++ b/src/tactic/core.lean @@ -1219,6 +1219,17 @@ do r ← decorate_ex "iterate1 failed: tactic did not succeed" t, L ← iterate t, return (r :: L) +/-- A simple check: `check_target_changes tac` applies tactic `tac` and fails if the main target +before applying the tactic `tac` unifies with one of the goals produced by the tactic itself. +Useful to make sure that the tactic `tac` is actually making progress. -/ +meta def check_target_changes (tac : tactic α) : tactic α := +focus1 $ do + t ← target, + x ← tac, + gs ← get_goals >>= list.mmap infer_type, + (success_if_fail $ gs.mfirst $ unify t) <|> fail "Goal did not change", + pure x + /-- Introduces one or more variables and returns the new local constants. Fails if `intro` cannot be applied. -/ meta def intros1 : tactic (list expr) := diff --git a/test/compute_degree.lean b/test/compute_degree.lean index 495ed2537c49e..9157e368c480a 100644 --- a/test/compute_degree.lean +++ b/test/compute_degree.lean @@ -5,6 +5,13 @@ open_locale polynomial variables {R : Type*} [semiring R] {a b c d e : R} +example {R : Type*} [ring R] (h : ∀ {p q : R[X]}, p.nat_degree ≤ 0 → (p * q).nat_degree = 0) : + nat_degree (- 1 * 1 : R[X]) = 0 := +begin + apply h _, + compute_degree_le, +end + example {p : R[X]} {n : ℕ} {p0 : p.nat_degree = 0} : (p ^ n).nat_degree ≤ 0 := by compute_degree_le From 8cd94b84c249ef4ae2561a84936a081f00aeaa22 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 25 Sep 2022 08:17:38 +0000 Subject: [PATCH 377/828] feat(order/symm_diff): Heyting bi-implication (#16544) Define `bihimp`, the Heyting bi-implication operator. This is dual to `symm_diff` and generalizes `iff` on propositions. Delete `order.imp` as all the material there is now fully superseded by the Heyting algebra material (`himp`, defined in `order.heyting.basic`). --- src/order/boolean_algebra.lean | 9 +- src/order/bounded_order.lean | 3 + src/order/heyting/basic.lean | 2 + src/order/imp.lean | 94 ------------- src/order/symm_diff.lean | 248 ++++++++++++++++++++++++++++++--- 5 files changed, 237 insertions(+), 119 deletions(-) delete mode 100644 src/order/imp.lean diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index ba1b3f338cae5..468ed04798143 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -542,8 +542,13 @@ instance : boolean_algebra αᵒᵈ := @[simp] lemma sup_inf_inf_compl : (x ⊓ y) ⊔ (x ⊓ yᶜ) = x := by rw [← sdiff_eq, sup_inf_sdiff _ _] -@[simp] lemma compl_sdiff : (x \ y)ᶜ = xᶜ ⊔ y := -by rw [sdiff_eq, compl_inf, compl_compl] +@[simp] lemma compl_sdiff : (x \ y)ᶜ = x ⇨ y := +by rw [sdiff_eq, himp_eq, compl_inf, compl_compl, sup_comm] + +@[simp] lemma compl_himp : (x ⇨ y)ᶜ = x \ y := @compl_sdiff αᵒᵈ _ _ _ + +@[simp] lemma compl_sdiff_compl : xᶜ \ yᶜ = y \ x := by rw [sdiff_compl, sdiff_eq, inf_comm] +@[simp] lemma compl_himp_compl : xᶜ ⇨ yᶜ = y ⇨ x := @compl_sdiff_compl αᵒᵈ _ _ _ lemma disjoint_compl_left_iff : disjoint xᶜ y ↔ y ≤ x := by rw [←le_compl_iff_disjoint_left, compl_compl] diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index ede566f7fef49..0ceff65d56554 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -1660,6 +1660,9 @@ section is_compl (disjoint : disjoint x y) (codisjoint : codisjoint x y) +lemma is_compl_iff [lattice α] [bounded_order α] {a b : α} : + is_compl a b ↔ disjoint a b ∧ codisjoint a b := ⟨λ h, ⟨h.1, h.2⟩, λ h, ⟨h.1, h.2⟩⟩ + namespace is_compl section bounded_order diff --git a/src/order/heyting/basic.lean b/src/order/heyting/basic.lean index f295ebab1a3f3..2ff8c76ec2a4f 100644 --- a/src/order/heyting/basic.lean +++ b/src/order/heyting/basic.lean @@ -259,6 +259,8 @@ end -- `p → q → r ↔ q → p → r` lemma himp_left_comm (a b c : α) : a ⇨ b ⇨ c = b ⇨ a ⇨ c := by simp_rw [himp_himp, inf_comm] +@[simp] lemma himp_idem : b ⇨ b ⇨ a = b ⇨ a := by rw [himp_himp, inf_idem] + lemma himp_inf_distrib (a b c : α) : a ⇨ b ⊓ c = (a ⇨ b) ⊓ (a ⇨ c) := eq_of_forall_le_iff $ λ d, by simp_rw [le_himp_iff, le_inf_iff, le_himp_iff] diff --git a/src/order/imp.lean b/src/order/imp.lean deleted file mode 100644 index 30e3b3cc2fc3a..0000000000000 --- a/src/order/imp.lean +++ /dev/null @@ -1,94 +0,0 @@ -/- -Copyright (c) 2021 Floris van Doorn. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Floris van Doorn, Yury Kudryashov --/ -import order.symm_diff -import tactic.monotonicity.basic - -/-! -# Implication and equivalence as operations on a boolean algebra - -In this file we define `lattice.imp` (notation: `a ⇒ₒ b`) and `lattice.biimp` (notation: `a ⇔ₒ b`) -to be the implication and equivalence as operations on a boolean algebra. More precisely, we put -`a ⇒ₒ b = aᶜ ⊔ b` and `a ⇔ₒ b = (a ⇒ₒ b) ⊓ (b ⇒ₒ a)`. Equivalently, `a ⇒ₒ b = (a \ b)ᶜ` and -`a ⇔ₒ b = (a ∆ b)ᶜ`. For propositions these operations are equal to the usual implication and `iff`. --/ - -variables {α β : Type*} - -namespace lattice - -/-- Implication as a binary operation on a boolean algebra. -/ -def imp [has_compl α] [has_sup α] (a b : α) : α := aᶜ ⊔ b - -infix ` ⇒ₒ `:65 := lattice.imp - -/-- Equivalence as a binary operation on a boolean algebra. -/ -def biimp [has_compl α] [has_sup α] [has_inf α] (a b : α) : α := (a ⇒ₒ b) ⊓ (b ⇒ₒ a) - -infix ` ⇔ₒ `:60 := lattice.biimp - -@[simp] lemma imp_eq_arrow (p q : Prop) : p ⇒ₒ q = (p → q) := propext imp_iff_not_or.symm - -@[simp] lemma biimp_eq_iff (p q : Prop) : p ⇔ₒ q = (p ↔ q) := by simp [biimp, ← iff_def] - -variables [boolean_algebra α] {a b c d : α} - -@[simp] lemma compl_imp (a b : α) : (a ⇒ₒ b)ᶜ = a \ b := by simp [imp, sdiff_eq] - -lemma compl_sdiff (a b : α) : (a \ b)ᶜ = a ⇒ₒ b := by rw [← compl_imp, compl_compl] - -@[mono] lemma imp_mono (h₁ : a ≤ b) (h₂ : c ≤ d) : b ⇒ₒ c ≤ a ⇒ₒ d := -sup_le_sup (compl_le_compl h₁) h₂ - -lemma inf_imp_eq (a b c : α) : a ⊓ (b ⇒ₒ c) = (a ⇒ₒ b) ⇒ₒ (a ⊓ c) := -by unfold imp; simp [inf_sup_left] - -@[simp] lemma imp_eq_top_iff : (a ⇒ₒ b = ⊤) ↔ a ≤ b := -by rw [← compl_sdiff, compl_eq_top, sdiff_eq_bot_iff] - -@[simp] lemma imp_eq_bot_iff : (a ⇒ₒ b = ⊥) ↔ (a = ⊤ ∧ b = ⊥) := by simp [imp] - -@[simp] lemma imp_bot (a : α) : a ⇒ₒ ⊥ = aᶜ := sup_bot_eq - -@[simp] lemma top_imp (a : α) : ⊤ ⇒ₒ a = a := by simp [imp] - -@[simp] lemma bot_imp (a : α) : ⊥ ⇒ₒ a = ⊤ := imp_eq_top_iff.2 bot_le - -@[simp] lemma imp_top (a : α) : a ⇒ₒ ⊤ = ⊤ := imp_eq_top_iff.2 le_top - -@[simp] lemma imp_self (a : α) : a ⇒ₒ a = ⊤ := compl_sup_eq_top - -@[simp] lemma compl_imp_compl (a b : α) : aᶜ ⇒ₒ bᶜ = b ⇒ₒ a := by simp [imp, sup_comm] - -lemma imp_inf_le {α : Type*} [boolean_algebra α] (a b : α) : (a ⇒ₒ b) ⊓ a ≤ b := -by { unfold imp, rw [inf_sup_right], simp } - -lemma inf_imp_eq_imp_imp (a b c : α) : ((a ⊓ b) ⇒ₒ c) = (a ⇒ₒ (b ⇒ₒ c)) := by simp [imp, sup_assoc] - -lemma le_imp_iff : a ≤ (b ⇒ₒ c) ↔ a ⊓ b ≤ c := -by rw [imp, sup_comm, is_compl_compl.le_sup_right_iff_inf_left_le] - -lemma biimp_mp (a b : α) : (a ⇔ₒ b) ≤ (a ⇒ₒ b) := inf_le_left - -lemma biimp_mpr (a b : α) : (a ⇔ₒ b) ≤ (b ⇒ₒ a) := inf_le_right - -lemma biimp_comm (a b : α) : (a ⇔ₒ b) = (b ⇔ₒ a) := -by {unfold lattice.biimp, rw inf_comm} - -@[simp] lemma biimp_eq_top_iff : a ⇔ₒ b = ⊤ ↔ a = b := -by simp [biimp, ← le_antisymm_iff] - -@[simp] lemma biimp_self (a : α) : a ⇔ₒ a = ⊤ := biimp_eq_top_iff.2 rfl - -lemma biimp_symm : a ≤ (b ⇔ₒ c) ↔ a ≤ (c ⇔ₒ b) := by rw biimp_comm - -lemma compl_symm_diff (a b : α) : (a ∆ b)ᶜ = a ⇔ₒ b := -by simp only [biimp, imp, symm_diff, sdiff_eq, compl_sup, compl_inf, compl_compl] - -lemma compl_biimp (a b : α) : (a ⇔ₒ b)ᶜ = a ∆ b := by rw [← compl_symm_diff, compl_compl] - -@[simp] lemma compl_biimp_compl : aᶜ ⇔ₒ bᶜ = a ⇔ₒ b := by simp [biimp, inf_comm] - -end lattice diff --git a/src/order/symm_diff.lean b/src/order/symm_diff.lean index dd9529d82c243..a5036c065ea86 100644 --- a/src/order/symm_diff.lean +++ b/src/order/symm_diff.lean @@ -1,26 +1,30 @@ /- Copyright (c) 2021 Bryan Gin-ge Chen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Adam Topaz, Bryan Gin-ge Chen +Authors: Adam Topaz, Bryan Gin-ge Chen, Yaël Dillies -/ import order.boolean_algebra /-! -# Symmetric difference +# Symmetric difference and bi-implication -The symmetric difference or disjunctive union of sets `A` and `B` is the set of elements that are -in either `A` or `B` but not both. Translated into propositions, the symmetric difference is `xor`. +This file defines the symmetric difference and bi-implication operators in (co-)Heyting algebras. -The symmetric difference operator (`symm_diff`) is defined in this file for any type with `⊔` and -`\` via the formula `(A \ B) ⊔ (B \ A)`, however the theorems proved about it only hold for -`generalized_boolean_algebra`s and `boolean_algebra`s. +## Examples -The symmetric difference is the addition operator in the Boolean ring structure on Boolean algebras. +Some examples are +* The symmetric difference of two sets is the set of elements that are in either but not both. +* The symmetric difference on propositions is `xor`. +* The symmetric difference on `bool` is `bxor`. +* The equivalence of propositions. Two propositions are equivalent if they imply each other. +* The symmetric difference translates to addition when considering a Boolean algebra as a Boolean + ring. ## Main declarations -* `symm_diff`: the symmetric difference operator, defined as `(A \ B) ⊔ (B \ A)` +* `symm_diff`: The symmetric difference operator, defined as `(a \ b) ⊔ (b \ a)` +* `bihimp`: The bi-implication operator, defined as `(b ⇨ a) ⊓ (a ⇨ b)` In generalized Boolean algebras, the symmetric difference operator is: @@ -30,6 +34,7 @@ In generalized Boolean algebras, the symmetric difference operator is: ## Notations * `a ∆ b`: `symm_diff a b` +* `a ⇔ b`: `bihimp a b` ## References @@ -39,29 +44,42 @@ Proof from the Book" by John McCuan: * ## Tags -boolean ring, generalized boolean algebra, boolean algebra, symmetric differences + +boolean ring, generalized boolean algebra, boolean algebra, symmetric difference, bi-implication, +Heyting -/ -open function +open function order_dual -variables {α : Type*} +variables {ι α β : Type*} {π : ι → Type*} /-- The symmetric difference operator on a type with `⊔` and `\` is `(A \ B) ⊔ (B \ A)`. -/ def symm_diff [has_sup α] [has_sdiff α] (a b : α) : α := a \ b ⊔ b \ a +/-- The Heyting bi-implication is `(b ⇨ a) ⊓ (a ⇨ b)`. This generalizes equivalence of +propositions. -/ +def bihimp [has_inf α] [has_himp α] (a b : α) : α := (b ⇨ a) ⊓ (a ⇨ b) + /- This notation might conflict with the Laplacian once we have it. Feel free to put it in locale `order` or `symm_diff` if that happens. -/ infix ` ∆ `:100 := symm_diff +infix ` ⇔ `:100 := bihimp lemma symm_diff_def [has_sup α] [has_sdiff α] (a b : α) : a ∆ b = (a \ b) ⊔ (b \ a) := rfl +lemma bihimp_def [has_inf α] [has_himp α] (a b : α) : a ⇔ b = (b ⇨ a) ⊓ (a ⇨ b):= rfl lemma symm_diff_eq_xor (p q : Prop) : p ∆ q = xor p q := rfl +@[simp] lemma bihimp_iff_iff {p q : Prop} : p ⇔ q ↔ (p ↔ q) := +(iff_iff_implies_and_implies _ _).symm.trans iff.comm @[simp] lemma bool.symm_diff_eq_bxor : ∀ p q : bool, p ∆ q = bxor p q := dec_trivial section generalized_coheyting_algebra variables [generalized_coheyting_algebra α] (a b c d : α) +@[simp] lemma to_dual_symm_diff : to_dual (a ∆ b) = to_dual a ⇔ to_dual b := rfl +@[simp] lemma of_dual_bihimp (a b : αᵒᵈ) : of_dual (a ⇔ b) = of_dual a ∆ of_dual b := rfl + lemma symm_diff_comm : a ∆ b = b ∆ a := by simp only [(∆), sup_comm] instance symm_diff_is_comm : is_commutative α (∆) := ⟨symm_diff_comm⟩ @@ -70,6 +88,9 @@ instance symm_diff_is_comm : is_commutative α (∆) := ⟨symm_diff_comm⟩ @[simp] lemma symm_diff_bot : a ∆ ⊥ = a := by rw [(∆), sdiff_bot, bot_sdiff, sup_bot_eq] @[simp] lemma bot_symm_diff : ⊥ ∆ a = a := by rw [symm_diff_comm, symm_diff_bot] +@[simp] lemma symm_diff_eq_bot {a b : α} : a ∆ b = ⊥ ↔ a = b := +by simp_rw [symm_diff, sup_eq_bot_iff, sdiff_eq_bot_iff, le_antisymm_iff] + lemma symm_diff_of_le {a b : α} (h : a ≤ b) : a ∆ b = b \ a := by rw [symm_diff, sdiff_eq_bot_iff.2 h, bot_sup_eq] @@ -132,6 +153,60 @@ end end generalized_coheyting_algebra +section generalized_heyting_algebra +variables [generalized_heyting_algebra α] (a b c d : α) + +@[simp] lemma to_dual_bihimp : to_dual (a ⇔ b) = to_dual a ∆ to_dual b := rfl +@[simp] lemma of_dual_symm_diff (a b : αᵒᵈ) : of_dual (a ∆ b) = of_dual a ⇔ of_dual b := rfl + +lemma bihimp_comm : a ⇔ b = b ⇔ a := by simp only [(⇔), inf_comm] + +instance bihimp_is_comm : is_commutative α (⇔) := ⟨bihimp_comm⟩ + +@[simp] lemma bihimp_self : a ⇔ a = ⊤ := by rw [(⇔), inf_idem, himp_self] +@[simp] lemma bihimp_top : a ⇔ ⊤ = a := by rw [(⇔), himp_top, top_himp, inf_top_eq] +@[simp] lemma top_bihimp : ⊤ ⇔ a = a := by rw [bihimp_comm, bihimp_top] + +@[simp] lemma bihimp_eq_top {a b : α} : a ⇔ b = ⊤ ↔ a = b := @symm_diff_eq_bot αᵒᵈ _ _ _ + +lemma bihimp_of_le {a b : α} (h : a ≤ b) : a ⇔ b = b ⇨ a := +by rw [bihimp, himp_eq_top_iff.2 h, inf_top_eq] + +lemma bihimp_of_ge {a b : α} (h : b ≤ a) : a ⇔ b = a ⇨ b := +by rw [bihimp, himp_eq_top_iff.2 h, top_inf_eq] + +lemma le_bihimp {a b c : α} (hb : a ⊓ b ≤ c) (hc : a ⊓ c ≤ b) : a ≤ b ⇔ c := +le_inf (le_himp_iff.2 hc) $ le_himp_iff.2 hb + +lemma le_bihimp_iff {a b c : α} : a ≤ b ⇔ c ↔ a ⊓ b ≤ c ∧ a ⊓ c ≤ b := +by simp_rw [bihimp, le_inf_iff, le_himp_iff, and.comm] + +@[simp] lemma inf_le_bihimp {a b : α} : a ⊓ b ≤ a ⇔ b := inf_le_inf le_himp le_himp + +lemma bihimp_eq_inf_himp_inf : a ⇔ b = (a ⊔ b) ⇨ (a ⊓ b) := by simp [himp_inf_distrib, bihimp] + +lemma codisjoint.bihimp_eq_inf {a b : α} (h : codisjoint a b) : a ⇔ b = a ⊓ b := +by rw [(⇔), h.himp_eq_left, h.himp_eq_right] + +lemma himp_bihimp : a ⇨ (b ⇔ c) = ((a ⊓ c) ⇨ b) ⊓ ((a ⊓ b) ⇨ c) := +by rw [bihimp, himp_inf_distrib, himp_himp, himp_himp] + +@[simp] lemma sup_himp_bihimp : (a ⊔ b) ⇨ (a ⇔ b) = a ⇔ b := +by { rw himp_bihimp, simp [bihimp] } + +@[simp] lemma bihimp_himp_eq_inf : a ⇔ (a ⇨ b) = a ⊓ b := @symm_diff_sdiff_eq_sup αᵒᵈ _ _ _ +@[simp] lemma himp_bihimp_eq_inf : (b ⇨ a) ⇔ b = a ⊓ b := @sdiff_symm_diff_eq_sup αᵒᵈ _ _ _ + +@[simp] lemma bihimp_inf_sup : a ⇔ b ⊓ (a ⊔ b) = a ⊓ b := @symm_diff_sup_inf αᵒᵈ _ _ _ +@[simp] lemma sup_inf_bihimp : (a ⊔ b) ⊓ a ⇔ b = a ⊓ b := @inf_sup_symm_diff αᵒᵈ _ _ _ + +@[simp] lemma bihimp_bihimp_sup : a ⇔ b ⇔ (a ⊔ b) = a ⊓ b := @symm_diff_symm_diff_inf αᵒᵈ _ _ _ +@[simp] lemma sup_bihimp_bihimp : (a ⊔ b) ⇔ (a ⇔ b) = a ⊓ b := @inf_symm_diff_symm_diff αᵒᵈ _ _ _ + +lemma bihimp_triangle : a ⇔ b ⊓ b ⇔ c ≤ a ⇔ c := @symm_diff_triangle αᵒᵈ _ _ _ _ + +end generalized_heyting_algebra + section coheyting_algebra variables [coheyting_algebra α] (a : α) @@ -148,6 +223,20 @@ by rw [h.eq_hnot, hnot_symm_diff_self] end coheyting_algebra +section heyting_algebra +variables [heyting_algebra α] (a : α) + +@[simp] lemma bihimp_bot : a ⇔ ⊥ = aᶜ := by simp [bihimp] +@[simp] lemma bot_bihimp : ⊥ ⇔ a = aᶜ := by simp [bihimp] + +@[simp] lemma compl_bihimp_self : aᶜ ⇔ a = ⊥ := @hnot_symm_diff_self αᵒᵈ _ _ +@[simp] lemma bihimp_hnot_self : a ⇔ aᶜ = ⊥ := @symm_diff_hnot_self αᵒᵈ _ _ + +lemma is_compl.bihimp_eq_bot {a b : α} (h : is_compl a b) : a ⇔ b = ⊥ := +by rw [h.eq_compl, compl_bihimp_self] + +end heyting_algebra + section generalized_boolean_algebra variables [generalized_boolean_algebra α] (a b c d : α) @@ -179,7 +268,9 @@ by rw [symm_diff_def, sup_sdiff, sdiff_idem, sdiff_sdiff_self, bot_sup_eq] @[simp] lemma symm_diff_sdiff_right : (a ∆ b) \ b = a \ b := by rw [symm_diff_comm, symm_diff_sdiff_left] -@[simp] lemma sdiff_symm_diff_self : a \ (a ∆ b) = a ⊓ b := by simp [sdiff_symm_diff] +@[simp] lemma sdiff_symm_diff_left : a \ (a ∆ b) = a ⊓ b := by simp [sdiff_symm_diff] +@[simp] lemma sdiff_symm_diff_right : b \ (a ∆ b) = a ⊓ b := +by rw [symm_diff_comm, inf_comm, sdiff_symm_diff_left] lemma symm_diff_eq_sup : a ∆ b = a ⊔ b ↔ disjoint a b := begin @@ -256,10 +347,6 @@ calc a ∆ b = a ↔ a ∆ b = a ∆ ⊥ : by rw symm_diff_bot @[simp] lemma symm_diff_eq_right : a ∆ b = b ↔ a = ⊥ := by rw [symm_diff_comm, symm_diff_eq_left] -@[simp] lemma symm_diff_eq_bot : a ∆ b = ⊥ ↔ a = b := -calc a ∆ b = ⊥ ↔ a ∆ b = a ∆ a : by rw symm_diff_self - ... ↔ a = b : by rw [symm_diff_right_inj, eq_comm] - protected lemma disjoint.symm_diff_left (ha : disjoint a c) (hb : disjoint b c) : disjoint (a ∆ b) c := by { rw symm_diff_eq_sup_sdiff_inf, exact (ha.sup_left hb).disjoint_sdiff_left } @@ -277,18 +364,95 @@ end end generalized_boolean_algebra section boolean_algebra -variables [boolean_algebra α] (a b c : α) +variables [boolean_algebra α] (a b c d : α) + +/- `cogeneralized_boolean_algebra` isn't actually a typeclass, but the lemmas in here are dual to +the `generalized_boolean_algebra` ones -/ +section cogeneralized_boolean_algebra + +@[simp] lemma inf_himp_bihimp : (a ⇔ b) ⇨ (a ⊓ b) = a ⊔ b := @sup_sdiff_symm_diff αᵒᵈ _ _ _ + +lemma codisjoint_bihimp_sup : codisjoint (a ⇔ b) (a ⊔ b) := @disjoint_symm_diff_inf αᵒᵈ _ _ _ + +@[simp] lemma himp_bihimp_left : a ⇨ (a ⇔ b) = a ⇨ b := @symm_diff_sdiff_left αᵒᵈ _ _ _ +@[simp] lemma himp_bihimp_right : b ⇨ (a ⇔ b) = b ⇨ a := @symm_diff_sdiff_right αᵒᵈ _ _ _ + +@[simp] lemma bihimp_himp_left : (a ⇔ b) ⇨ a = a ⊔ b := @sdiff_symm_diff_left αᵒᵈ _ _ _ +@[simp] lemma bihimp_himp_right : (a ⇔ b) ⇨ b = a ⊔ b := @sdiff_symm_diff_right αᵒᵈ _ _ _ + +@[simp] lemma bihimp_eq_inf : a ⇔ b = a ⊓ b ↔ codisjoint a b := @symm_diff_eq_sup αᵒᵈ _ _ _ + +@[simp] lemma bihimp_le_iff_left : a ⇔ b ≤ a ↔ codisjoint a b := @le_symm_diff_iff_left αᵒᵈ _ _ _ +@[simp] lemma bihimp_le_iff_right : a ⇔ b ≤ b ↔ codisjoint a b := @le_symm_diff_iff_right αᵒᵈ _ _ _ + +lemma bihimp_assoc : a ⇔ b ⇔ c = a ⇔ (b ⇔ c) := @symm_diff_assoc αᵒᵈ _ _ _ _ + +instance bihimp_is_assoc : is_associative α (⇔) := ⟨bihimp_assoc⟩ + +lemma bihimp_left_comm : a ⇔ (b ⇔ c) = b ⇔ (a ⇔ c) := by simp_rw [←bihimp_assoc, bihimp_comm] +lemma bihimp_right_comm : a ⇔ b ⇔ c = a ⇔ c ⇔ b := by simp_rw [bihimp_assoc, bihimp_comm] + +lemma bihimp_bihimp_bihimp_comm : (a ⇔ b) ⇔ (c ⇔ d) = (a ⇔ c) ⇔ (b ⇔ d) := +by simp_rw [bihimp_assoc, bihimp_left_comm] + +@[simp] lemma bihimp_bihimp_cancel_left : a ⇔ (a ⇔ b) = b := by simp [←bihimp_assoc] +@[simp] lemma bihimp_bihimp_cancel_right : b ⇔ a ⇔ a = b := by simp [bihimp_assoc] + +@[simp] lemma bihimp_bihimp_self : a ⇔ b ⇔ a = b := by rw [bihimp_comm, bihimp_bihimp_cancel_left] + +lemma bihimp_left_involutive (a : α) : involutive (⇔ a) := bihimp_bihimp_cancel_right _ +lemma bihimp_right_involutive (a : α) : involutive ((⇔) a) := bihimp_bihimp_cancel_left _ +lemma bihimp_left_injective (a : α) : injective (⇔ a) := @symm_diff_left_injective αᵒᵈ _ _ +lemma bihimp_right_injective (a : α) : injective ((⇔) a) := @symm_diff_right_injective αᵒᵈ _ _ +lemma bihimp_left_surjective (a : α) : surjective (⇔ a) := @symm_diff_left_surjective αᵒᵈ _ _ +lemma bihimp_right_surjective (a : α) : surjective ((⇔) a) := @symm_diff_right_surjective αᵒᵈ _ _ + +variables {a b c} + +@[simp] lemma bihimp_left_inj : a ⇔ b = c ⇔ b ↔ a = c := (bihimp_left_injective _).eq_iff +@[simp] lemma bihimp_right_inj : a ⇔ b = a ⇔ c ↔ b = c := (bihimp_right_injective _).eq_iff + +@[simp] lemma bihimp_eq_left : a ⇔ b = a ↔ b = ⊤ := @symm_diff_eq_left αᵒᵈ _ _ _ +@[simp] lemma bihimp_eq_right : a ⇔ b = b ↔ a = ⊤ := @symm_diff_eq_right αᵒᵈ _ _ _ + +protected lemma codisjoint.bihimp_left (ha : codisjoint a c) (hb : codisjoint b c) : + codisjoint (a ⇔ b) c := +(ha.inf_left hb).mono_left inf_le_bihimp + +protected lemma codisjoint.bihimp_right (ha : codisjoint a b) (hb : codisjoint a c) : + codisjoint a (b ⇔ c) := +(ha.inf_right hb).mono_right inf_le_bihimp + +end cogeneralized_boolean_algebra lemma symm_diff_eq : a ∆ b = (a ⊓ bᶜ) ⊔ (b ⊓ aᶜ) := by simp only [(∆), sdiff_eq] +lemma bihimp_eq : a ⇔ b = (a ⊔ bᶜ) ⊓ (b ⊔ aᶜ) := by simp only [(⇔), himp_eq] + +lemma symm_diff_eq' : a ∆ b = (a ⊔ b) ⊓ (aᶜ ⊔ bᶜ) := +by rw [symm_diff_eq_sup_sdiff_inf, sdiff_eq, compl_inf] + +lemma bihimp_eq' : a ⇔ b = (a ⊓ b) ⊔ (aᶜ ⊓ bᶜ) := @symm_diff_eq' αᵒᵈ _ _ _ lemma symm_diff_top : a ∆ ⊤ = aᶜ := symm_diff_top' _ lemma top_symm_diff : ⊤ ∆ a = aᶜ := top_symm_diff' _ -lemma compl_symm_diff : (a ∆ b)ᶜ = (a ⊓ b) ⊔ (aᶜ ⊓ bᶜ) := -by simp only [←top_sdiff, sdiff_symm_diff, top_inf_eq] +@[simp] lemma compl_symm_diff : (a ∆ b)ᶜ = a ⇔ b := +by simp_rw [symm_diff, compl_sup_distrib, compl_sdiff, bihimp, inf_comm] + +@[simp] lemma compl_bihimp : (a ⇔ b)ᶜ = a ∆ b := @compl_symm_diff αᵒᵈ _ _ _ + +@[simp] lemma compl_symm_diff_compl : aᶜ ∆ bᶜ = a ∆ b := +sup_comm.trans $ by simp_rw [compl_sdiff_compl, sdiff_eq, symm_diff_eq] + +@[simp] lemma compl_bihimp_compl : aᶜ ⇔ bᶜ = a ⇔ b := @compl_symm_diff_compl αᵒᵈ _ _ _ + +@[simp] lemma symm_diff_eq_top : a ∆ b = ⊤ ↔ is_compl a b := +by rw [symm_diff_eq', ←compl_inf, inf_eq_top_iff, compl_eq_top, is_compl_iff, disjoint_iff, + codisjoint_iff, and.comm] -lemma symm_diff_eq_top_iff : a ∆ b = ⊤ ↔ is_compl a b := -by rw [symm_diff_eq_iff_sdiff_eq le_top, top_sdiff, compl_eq_iff_is_compl] +@[simp] lemma bihimp_eq_bot : a ⇔ b = ⊥ ↔ is_compl a b := +by rw [bihimp_eq', ←compl_sup, sup_eq_bot_iff, compl_eq_bot, is_compl_iff, disjoint_iff, + codisjoint_iff] @[simp] lemma compl_symm_diff_self : aᶜ ∆ a = ⊤ := hnot_symm_diff_self _ @[simp] lemma symm_diff_compl_self : a ∆ aᶜ = ⊤ := symm_diff_hnot_self _ @@ -297,7 +461,7 @@ lemma symm_diff_symm_diff_right' : a ∆ (b ∆ c) = (a ⊓ b ⊓ c) ⊔ (a ⊓ bᶜ ⊓ cᶜ) ⊔ (aᶜ ⊓ b ⊓ cᶜ) ⊔ (aᶜ ⊓ bᶜ ⊓ c) := calc a ∆ (b ∆ c) = (a ⊓ ((b ⊓ c) ⊔ (bᶜ ⊓ cᶜ))) ⊔ (((b ⊓ cᶜ) ⊔ (c ⊓ bᶜ)) ⊓ aᶜ) : by rw [symm_diff_eq, compl_symm_diff, - symm_diff_eq] + bihimp_eq', symm_diff_eq] ... = (a ⊓ b ⊓ c) ⊔ (a ⊓ bᶜ ⊓ cᶜ) ⊔ (b ⊓ cᶜ ⊓ aᶜ) ⊔ (c ⊓ bᶜ ⊓ aᶜ) : by rw [inf_sup_left, inf_sup_right, ←sup_assoc, ←inf_assoc, ←inf_assoc] @@ -310,3 +474,41 @@ calc a ∆ (b ∆ c) = (a ⊓ ((b ⊓ c) ⊔ (bᶜ ⊓ cᶜ))) ⊔ end end boolean_algebra + +/-! ### Prod -/ + +section prod + +@[simp] lemma symm_diff_fst [generalized_coheyting_algebra α] [generalized_coheyting_algebra β] + (a b : α × β) : + (a ∆ b).1 = a.1 ∆ b.1 := rfl +@[simp] lemma symm_diff_snd [generalized_coheyting_algebra α] [generalized_coheyting_algebra β] + (a b : α × β) : + (a ∆ b).2 = a.2 ∆ b.2 := rfl + +@[simp] lemma bihimp_fst [generalized_heyting_algebra α] [generalized_heyting_algebra β] + (a b : α × β) : + (a ⇔ b).1 = a.1 ⇔ b.1 := rfl +@[simp] lemma bihimp_snd [generalized_heyting_algebra α] [generalized_heyting_algebra β] + (a b : α × β) : + (a ⇔ b).2 = a.2 ⇔ b.2 := rfl + +end prod + +/-! ### Pi -/ + +namespace pi + +lemma symm_diff_def [Π i, generalized_coheyting_algebra (π i)] (a b : Π i, π i) : + a ∆ b = λ i, a i ∆ b i := rfl + +lemma bihimp_def [Π i, generalized_heyting_algebra (π i)] (a b : Π i, π i) : + a ⇔ b = λ i, a i ⇔ b i := rfl + +@[simp] lemma symm_diff_apply [Π i, generalized_coheyting_algebra (π i)] (a b : Π i, π i) (i : ι) : + (a ∆ b) i = a i ∆ b i := rfl + +@[simp] lemma bihimp_apply [Π i, generalized_heyting_algebra (π i)] (a b : Π i, π i) (i : ι) : + (a ⇔ b) i = a i ⇔ b i := rfl + +end pi From 8f09a40271f7da8f27e4939f516996a9d653ca8b Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sun, 25 Sep 2022 11:24:21 +0000 Subject: [PATCH 378/828] chore(category_theory/preadditive/biproducts): speed up biprod.of_components_eq (#16516) --- src/category_theory/preadditive/biproducts.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/src/category_theory/preadditive/biproducts.lean b/src/category_theory/preadditive/biproducts.lean index 8f8c976b4e853..eb707be966f12 100644 --- a/src/category_theory/preadditive/biproducts.lean +++ b/src/category_theory/preadditive/biproducts.lean @@ -134,7 +134,10 @@ lemma biprod.of_components_eq (f : X₁ ⊞ X₂ ⟶ Y₁ ⊞ Y₂) : biprod.of_components (biprod.inl ≫ f ≫ biprod.fst) (biprod.inl ≫ f ≫ biprod.snd) (biprod.inr ≫ f ≫ biprod.fst) (biprod.inr ≫ f ≫ biprod.snd) = f := begin - ext; simp, + ext; + simp only [category.comp_id, biprod.inr_fst, biprod.inr_snd, biprod.inl_snd, add_zero, zero_add, + biprod.inl_of_components, biprod.inr_of_components, eq_self_iff_true, category.assoc, comp_zero, + biprod.inl_fst, preadditive.add_comp], end @[simp] From 3b727dc2623011155a0677c21dcf3deafe8a0ce5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 25 Sep 2022 11:24:22 +0000 Subject: [PATCH 379/828] feat(algebra/module/basic): More general `smul_ne_zero` (#16608) Prove the one way implication of `smul_ne_zero` that holds without `smul_eq_zero` separately. Call it `smul_ne_zero` and rename `smul_ne_zero` to `smul_ne_zero_iff`. This matches `mul_ne_zero` and `mul_ne_zero_iff`. --- src/algebra/module/basic.lean | 22 +++++++++++++--------- src/algebra/support.lean | 12 ++++++------ src/geometry/euclidean/oriented_angle.lean | 2 +- 3 files changed, 20 insertions(+), 16 deletions(-) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index abe2872e1df65..e4f7ba125817f 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -512,19 +512,23 @@ instance no_zero_divisors.to_no_zero_smul_divisors [has_zero R] [has_mul R] [no_ no_zero_smul_divisors R R := ⟨λ c x, eq_zero_or_eq_zero_of_mul_eq_zero⟩ -section module +lemma smul_ne_zero [has_zero R] [has_zero M] [has_smul R M] [no_zero_smul_divisors R M] {c : R} + {x : M} (hc : c ≠ 0) (hx : x ≠ 0) : c • x ≠ 0 := +λ h, (eq_zero_or_eq_zero_of_smul_eq_zero h).elim hc hx -variables [semiring R] [add_comm_monoid M] [module R M] +section smul_with_zero +variables [has_zero R] [has_zero M] [smul_with_zero R M] [no_zero_smul_divisors R M] {c : R} {x : M} -@[simp] -theorem smul_eq_zero [no_zero_smul_divisors R M] {c : R} {x : M} : - c • x = 0 ↔ c = 0 ∨ x = 0 := +@[simp] lemma smul_eq_zero : c • x = 0 ↔ c = 0 ∨ x = 0 := ⟨eq_zero_or_eq_zero_of_smul_eq_zero, - λ h, h.elim (λ h, h.symm ▸ zero_smul R x) (λ h, h.symm ▸ smul_zero c)⟩ + λ h, h.elim (λ h, h.symm ▸ zero_smul R x) (λ h, h.symm ▸ smul_zero c)⟩ + +lemma smul_ne_zero_iff : c • x ≠ 0 ↔ c ≠ 0 ∧ x ≠ 0 := by rw [ne.def, smul_eq_zero, not_or_distrib] -theorem smul_ne_zero [no_zero_smul_divisors R M] {c : R} {x : M} : - c • x ≠ 0 ↔ c ≠ 0 ∧ x ≠ 0 := -by simp only [ne.def, smul_eq_zero, not_or_distrib] +end smul_with_zero + +section module +variables [semiring R] [add_comm_monoid M] [module R M] section nat diff --git a/src/algebra/support.lean b/src/algebra/support.lean index 3f139fa4dd09d..ed487de8161ed 100644 --- a/src/algebra/support.lean +++ b/src/algebra/support.lean @@ -198,9 +198,14 @@ mul_support_binop_subset (/) one_div_one f g end division_monoid +lemma support_smul [has_zero R] [has_zero M] [smul_with_zero R M] [no_zero_smul_divisors R M] + (f : α → R) (g : α → M) : + support (f • g) = support f ∩ support g := +ext $ λ x, smul_ne_zero_iff + @[simp] lemma support_mul [mul_zero_class R] [no_zero_divisors R] (f g : α → R) : support (λ x, f x * g x) = support f ∩ support g := -set.ext $ λ x, by simp only [mem_support, mul_ne_zero_iff, mem_inter_iff, not_or_distrib] +support_smul f g @[simp] lemma support_mul_subset_left [mul_zero_class R] (f g : α → R) : support (λ x, f x * g x) ⊆ support f := @@ -220,11 +225,6 @@ lemma support_smul_subset_left [has_zero M] [has_zero β] [smul_with_zero M β] support (f • g) ⊆ support f := λ x hfg hf, hfg $ by rw [pi.smul_apply', hf, zero_smul] -lemma support_smul [semiring R] [add_comm_monoid M] [module R M] - [no_zero_smul_divisors R M] (f : α → R) (g : α → M) : - support (f • g) = support f ∩ support g := -ext $ λ x, smul_ne_zero - lemma support_const_smul_of_ne_zero [semiring R] [add_comm_monoid M] [module R M] [no_zero_smul_divisors R M] (c : R) (g : α → M) (hc : c ≠ 0) : support (c • g) = support g := diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 580adeea1ba38..89a8bc684257d 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -333,7 +333,7 @@ begin by_cases hx : x = 0, { simp [hx] }, rcases lt_trichotomy r 0 with hr|hr|hr, { rw ←neg_smul, - exact or.inr ⟨hx, smul_ne_zero.2 ⟨hr.ne, hx⟩, + exact or.inr ⟨hx, smul_ne_zero hr.ne hx, same_ray_pos_smul_right x (left.neg_pos_iff.2 hr)⟩ }, { simp [hr] }, { exact or.inl (same_ray_pos_smul_right x hr) } } From 1313934bbc1f1b470c3e49bf2350e5e17364c97d Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Sun, 25 Sep 2022 15:14:30 +0000 Subject: [PATCH 380/828] feat(analysis/inner_product_space/l2_space): compute orthogonal projection on U given a hilbert basis of U (#15541) --- src/analysis/inner_product_space/l2_space.lean | 8 +++++++- src/analysis/inner_product_space/pi_L2.lean | 6 ++++++ 2 files changed, 13 insertions(+), 1 deletion(-) diff --git a/src/analysis/inner_product_space/l2_space.lean b/src/analysis/inner_product_space/l2_space.lean index 518ff11a4b9a9..ddb60eb9ced67 100644 --- a/src/analysis/inner_product_space/l2_space.lean +++ b/src/analysis/inner_product_space/l2_space.lean @@ -80,7 +80,7 @@ Hilbert space, Hilbert sum, l2, Hilbert basis, unitary equivalence, isometric is -/ open is_R_or_C submodule filter -open_locale big_operators nnreal ennreal classical complex_conjugate +open_locale big_operators nnreal ennreal classical complex_conjugate topological_space noncomputable theory @@ -485,6 +485,12 @@ end (b.to_orthonormal_basis : ι → E) = b := orthonormal_basis.coe_mk _ _ +protected lemma has_sum_orthogonal_projection {U : submodule 𝕜 E} + [complete_space U] (b : hilbert_basis ι 𝕜 U) (x : E) : + has_sum (λ i, ⟪(b i : E), x⟫ • b i) (orthogonal_projection U x) := +by simpa only [b.repr_apply_apply, inner_orthogonal_projection_eq_of_mem_left] + using b.has_sum_repr (orthogonal_projection U x) + lemma finite_spans_dense (b : hilbert_basis ι 𝕜 E) : (⨆ J : finset ι, span 𝕜 (J.image b : set E)).topological_closure = ⊤ := eq_top_iff.mpr $ b.dense_span.ge.trans diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 1a460c5986670..38dcf50b06a30 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -332,6 +332,12 @@ begin refl, end +protected lemma orthogonal_projection_eq_sum {U : submodule 𝕜 E} [complete_space U] + (b : orthonormal_basis ι 𝕜 U) (x : E) : + orthogonal_projection U x = ∑ i, ⟪(b i : E), x⟫ • b i := +by simpa only [b.repr_apply_apply, inner_orthogonal_projection_eq_of_mem_left] + using (b.sum_repr (orthogonal_projection U x)).symm + /-- Mapping an orthonormal basis along a `linear_isometry_equiv`. -/ protected def map {G : Type*} [inner_product_space 𝕜 G] (b : orthonormal_basis ι 𝕜 E) (L : E ≃ₗᵢ[𝕜] G) : From 4a927b25c3307fd39f82bca9feb05830c47f4e90 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sun, 25 Sep 2022 17:18:09 +0000 Subject: [PATCH 381/828] feat(algebra/order/hom/monoid): add `monotone_iff_map_nonneg/pos` + golf (#16601) If the source is an ordered add_group, a monoid hom is monotone if and only if it sends nonnegative elements to a nonnegative elements. This connects the notion of [positive linear functional](https://en.wikipedia.org/wiki/Positive_linear_functional) to monotonicity. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/Amenable.20Groups/near/300076886) --- src/algebra/order/complete_field.lean | 11 +--------- src/algebra/order/hom/monoid.lean | 31 +++++++++++++++++++++++++++ 2 files changed, 32 insertions(+), 10 deletions(-) diff --git a/src/algebra/order/complete_field.lean b/src/algebra/order/complete_field.lean index 3837dee54f365..9ab005145b761 100644 --- a/src/algebra/order/complete_field.lean +++ b/src/algebra/order/complete_field.lean @@ -316,16 +316,7 @@ section real variables {R S : Type*} [ordered_ring R] [linear_ordered_ring S] lemma ring_hom_monotone (hR : ∀ r : R, 0 ≤ r → ∃ s : R, s^2 = r) (f : R →+* S) : monotone f := -begin - intros a b hab, - apply le_of_sub_nonneg, - obtain ⟨s, hs⟩ := hR (b - a) (sub_nonneg.mpr hab), - calc - f b - f a = f (b - a) : (ring_hom.map_sub _ _ _).symm - ... = f(s^2) : by rw ←hs - ... = f(s)^2 : ring_hom.map_pow _ _ _ - ... ≥ 0 : sq_nonneg _ -end +(monotone_iff_map_nonneg f).2 $ λ r h, by { obtain ⟨s, rfl⟩ := hR r h, rw map_pow, apply sq_nonneg } /-- There exists no nontrivial ring homomorphism `ℝ →+* ℝ`. -/ instance real.ring_hom.unique : unique (ℝ →+* ℝ) := diff --git a/src/algebra/order/hom/monoid.lean b/src/algebra/order/hom/monoid.lean index 5d9ab16c0c6de..ad3d404337a62 100644 --- a/src/algebra/order/hom/monoid.lean +++ b/src/algebra/order/hom/monoid.lean @@ -171,6 +171,37 @@ lemma map_nonpos (ha : a ≤ 0) : f a ≤ 0 := by { rw ←map_zero f, exact orde end ordered_add_comm_monoid +section ordered_add_comm_group + +variables [ordered_add_comm_group α] + [ordered_add_comm_monoid β] [add_monoid_hom_class F α β] (f : F) + +lemma monotone_iff_map_nonneg : monotone (f : α → β) ↔ ∀ a, 0 ≤ a → 0 ≤ f a := +⟨λ h a, by { rw ←map_zero f, apply h }, λ h a b hl, + by { rw [←sub_add_cancel b a, map_add f], exact le_add_of_nonneg_left (h _ $ sub_nonneg.2 hl) }⟩ + +lemma antitone_iff_map_nonpos : antitone (f : α → β) ↔ ∀ a, 0 ≤ a → f a ≤ 0 := +monotone_to_dual_comp_iff.symm.trans $ monotone_iff_map_nonneg _ +lemma monotone_iff_map_nonpos : monotone (f : α → β) ↔ ∀ a ≤ 0, f a ≤ 0 := +antitone_comp_of_dual_iff.symm.trans $ antitone_iff_map_nonpos _ +lemma antitone_iff_map_nonneg : antitone (f : α → β) ↔ ∀ a ≤ 0, 0 ≤ f a := +monotone_comp_of_dual_iff.symm.trans $ monotone_iff_map_nonneg _ + +variable [covariant_class β β (+) (<)] + +lemma strict_mono_iff_map_pos : strict_mono (f : α → β) ↔ ∀ a, 0 < a → 0 < f a := +⟨λ h a, by { rw ←map_zero f, apply h }, λ h a b hl, + by { rw [←sub_add_cancel b a, map_add f], exact lt_add_of_pos_left _ (h _ $ sub_pos.2 hl) }⟩ + +lemma strict_anti_iff_map_neg : strict_anti (f : α → β) ↔ ∀ a, 0 < a → f a < 0 := +strict_mono_to_dual_comp_iff.symm.trans $ strict_mono_iff_map_pos _ +lemma strict_mono_iff_map_neg : strict_mono (f : α → β) ↔ ∀ a < 0, f a < 0 := +strict_anti_comp_of_dual_iff.symm.trans $ strict_anti_iff_map_neg _ +lemma strict_anti_iff_map_pos : strict_anti (f : α → β) ↔ ∀ a < 0, 0 < f a := +strict_mono_comp_of_dual_iff.symm.trans $ strict_mono_iff_map_pos _ + +end ordered_add_comm_group + namespace order_monoid_hom section preorder variables [preorder α] [preorder β] [preorder γ] [preorder δ] [mul_one_class α] From 5a8eded100433306ec0f59c3b6ae21a161b46c68 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Sun, 25 Sep 2022 18:26:49 +0000 Subject: [PATCH 382/828] refactor(linear_algebra/affine_space): remove `open_locale classical` (#16628) Use of `open_locale classical` in mathlib is liable to cause problems if it results in the classical decidability instances forming part of the type of a lemma, making that lemma harder to use in any context where typeclass inference finds a different decidability instance. Remove it from affine space files, adding explicit decidability instance parameters where needed for the type of a lemma and uses of the `classical` tactic when only needed for a proof. --- .../affine_space/affine_subspace.lean | 5 +++-- .../affine_space/combination.lean | 20 ++++++++++++------- .../affine_space/finite_dimensional.lean | 4 ++-- .../affine_space/independent.lean | 5 ++++- 4 files changed, 22 insertions(+), 12 deletions(-) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index 648e9f5dca7be..e5f72dfcde09e 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -52,7 +52,7 @@ Those depending on analysis or topology are defined elsewhere; see -/ noncomputable theory -open_locale big_operators classical affine +open_locale big_operators affine open set @@ -1032,7 +1032,8 @@ end /-- The `vector_span` is the span of the pairwise subtractions with a given point on the right, excluding the subtraction of that point from itself. -/ -lemma vector_span_eq_span_vsub_finset_right_ne {s : finset P} {p : P} (hp : p ∈ s) : +lemma vector_span_eq_span_vsub_finset_right_ne [decidable_eq P] [decidable_eq V] {s : finset P} + {p : P} (hp : p ∈ s) : vector_span k (s : set P) = submodule.span k ((s.erase p).image (-ᵥ p)) := by simp [vector_span_eq_span_vsub_set_right_ne _ (finset.mem_coe.mpr hp)] diff --git a/src/linear_algebra/affine_space/combination.lean b/src/linear_algebra/affine_space/combination.lean index e58d5cd4e5ae4..cbfaee749a4c1 100644 --- a/src/linear_algebra/affine_space/combination.lean +++ b/src/linear_algebra/affine_space/combination.lean @@ -40,7 +40,7 @@ These definitions are for sums over a `finset`; versions for a -/ noncomputable theory -open_locale big_operators classical affine +open_locale big_operators affine namespace finset @@ -122,7 +122,7 @@ end /-- The weighted sum is unaffected by removing the base point, if present, from the set of points. -/ -@[simp] lemma weighted_vsub_of_point_erase (w : ι → k) (p : ι → P) (i : ι) : +@[simp] lemma weighted_vsub_of_point_erase [decidable_eq ι] (w : ι → k) (p : ι → P) (i : ι) : (s.erase i).weighted_vsub_of_point p (p i) w = s.weighted_vsub_of_point p (p i) w := begin rw [weighted_vsub_of_point_apply, weighted_vsub_of_point_apply], @@ -298,7 +298,7 @@ lemma affine_combination_vsub (w₁ w₂ : ι → k) (p : ι → P) : s.affine_combination p w₁ -ᵥ s.affine_combination p w₂ = s.weighted_vsub p (w₁ - w₂) := by rw [←affine_map.linear_map_vsub, affine_combination_linear, vsub_eq_sub] -lemma attach_affine_combination_of_injective +lemma attach_affine_combination_of_injective [decidable_eq P] (s : finset P) (w : P → k) (f : s → P) (hf : function.injective f) : s.attach.affine_combination f (w ∘ f) = (image f univ).affine_combination id w := begin @@ -314,7 +314,7 @@ end lemma attach_affine_combination_coe (s : finset P) (w : P → k) : s.attach.affine_combination (coe : s → P) (w ∘ coe) = s.affine_combination id w := -by rw [attach_affine_combination_of_injective s w (coe : s → P) subtype.coe_injective, +by classical; rw [attach_affine_combination_of_injective s w (coe : s → P) subtype.coe_injective, univ_eq_attach, attach_image_coe] omit S @@ -406,6 +406,7 @@ lemma eq_weighted_vsub_of_point_subset_iff_eq_weighted_vsub_of_point_subtype {v ∃ (fs : finset s) (w : s → k) (hw : ∑ i in fs, w i = x), v = fs.weighted_vsub_of_point (λ (i : s), p i) b w := begin + classical, simp_rw weighted_vsub_of_point_apply, split, { rintros ⟨fs, hfs, w, rfl, rfl⟩, @@ -538,7 +539,7 @@ by simp [centroid_def, affine_combination_apply] /-- The centroid of two points, expressed directly as adding a vector to a point. -/ -lemma centroid_pair [invertible (2 : k)] (p : ι → P) (i₁ i₂ : ι) : +lemma centroid_pair [decidable_eq ι] [invertible (2 : k)] (p : ι → P) (i₁ i₂ : ι) : ({i₁, i₂} : finset ι).centroid k p = (2 ⁻¹ : k) • (p i₂ -ᵥ p i₁) +ᵥ p i₁ := begin by_cases h : i₁ = i₂, @@ -662,8 +663,8 @@ have the same centroid. -/ lemma centroid_eq_of_inj_on_of_image_eq {p : ι → P} (hi : ∀ i j ∈ s, p i = p j → i = j) {p₂ : ι₂ → P} (hi₂ : ∀ i j ∈ s₂, p₂ i = p₂ j → i = j) (he : p '' ↑s = p₂ '' ↑s₂) : s.centroid k p = s₂.centroid k p₂ := -by rw [s.centroid_eq_centroid_image_of_inj_on k hi rfl, - s₂.centroid_eq_centroid_image_of_inj_on k hi₂ he] +by classical; rw [s.centroid_eq_centroid_image_of_inj_on k hi rfl, + s₂.centroid_eq_centroid_image_of_inj_on k hi₂ he] end finset @@ -680,6 +681,7 @@ lemma weighted_vsub_mem_vector_span {s : finset ι} {w : ι → k} (h : ∑ i in s, w i = 0) (p : ι → P) : s.weighted_vsub p w ∈ vector_span k (set.range p) := begin + classical, rcases is_empty_or_nonempty ι with hι|⟨⟨i0⟩⟩, { resetI, simp [finset.eq_empty_of_is_empty s] }, { rw [vector_span_range_eq_span_range_vsub_right k p i0, ←set.image_univ, @@ -703,6 +705,7 @@ lemma affine_combination_mem_affine_span [nontrivial k] {s : finset ι} {w : ι (h : ∑ i in s, w i = 1) (p : ι → P) : s.affine_combination p w ∈ affine_span k (set.range p) := begin + classical, have hnz : ∑ i in s, w i ≠ 0 := h.symm ▸ one_ne_zero, have hn : s.nonempty := finset.nonempty_of_sum_ne_zero hnz, cases hn with i1 hi1, @@ -728,6 +731,7 @@ lemma mem_vector_span_iff_eq_weighted_vsub {v : V} {p : ι → P} : v ∈ vector_span k (set.range p) ↔ ∃ (s : finset ι) (w : ι → k) (h : ∑ i in s, w i = 0), v = s.weighted_vsub p w := begin + classical, split, { rcases is_empty_or_nonempty ι with hι|⟨⟨i0⟩⟩, swap, { rw [vector_span_range_eq_span_range_vsub_right k p i0, ←set.image_univ, @@ -772,6 +776,7 @@ lemma eq_affine_combination_of_mem_affine_span {p1 : P} {p : ι → P} (h : p1 ∈ affine_span k (set.range p)) : ∃ (s : finset ι) (w : ι → k) (hw : ∑ i in s, w i = 1), p1 = s.affine_combination p w := begin + classical, have hn : ((affine_span k (set.range p)) : set P).nonempty := ⟨p1, h⟩, rw [affine_span_nonempty, set.range_nonempty_iff_nonempty] at hn, cases hn with i0, @@ -805,6 +810,7 @@ lemma eq_affine_combination_of_mem_affine_span_of_fintype [fintype ι] {p1 : P} (h : p1 ∈ affine_span k (set.range p)) : ∃ (w : ι → k) (hw : ∑ i, w i = 1), p1 = finset.univ.affine_combination p w := begin + classical, obtain ⟨s, w, hw, rfl⟩ := eq_affine_combination_of_mem_affine_span h, refine ⟨(s : set ι).indicator w, _, finset.affine_combination_indicator_subset w p s.subset_univ⟩, simp only [finset.mem_coe, set.indicator_apply, ← hw], diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index 7c181394b7b5e..ce6230126093a 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -20,7 +20,7 @@ subspaces of affine spaces. -/ noncomputable theory -open_locale big_operators classical affine +open_locale big_operators affine section affine_space' @@ -70,7 +70,7 @@ finite_dimensional_direction_affine_span_of_finite k (set.to_finite _) /-- An affine-independent family of points in a finite-dimensional affine space is finite. -/ noncomputable def fintype_of_fin_dim_affine_independent [finite_dimensional k V] {p : ι → P} (hi : affine_independent k p) : fintype ι := -if hι : is_empty ι then (@fintype.of_is_empty _ hι) else +by classical; exact if hι : is_empty ι then (@fintype.of_is_empty _ hι) else begin let q := (not_is_empty_iff.mp hι).some, rw affine_independent_iff_linear_independent_vsub k p q at hi, diff --git a/src/linear_algebra/affine_space/independent.lean b/src/linear_algebra/affine_space/independent.lean index 542d5b51c80d2..94e3251552e58 100644 --- a/src/linear_algebra/affine_space/independent.lean +++ b/src/linear_algebra/affine_space/independent.lean @@ -32,7 +32,7 @@ This file defines affinely independent families of points. -/ noncomputable theory -open_locale big_operators classical affine +open_locale big_operators affine open function section affine_independent @@ -81,6 +81,7 @@ from a base point in that family are linearly independent. -/ lemma affine_independent_iff_linear_independent_vsub (p : ι → P) (i1 : ι) : affine_independent k p ↔ linear_independent k (λ i : {x // x ≠ i1}, (p i -ᵥ p i1 : V)) := begin + classical, split, { intro h, rw linear_independent_iff', @@ -177,6 +178,7 @@ lemma affine_independent_iff_indicator_eq_of_affine_combination_eq (p : ι → P ∑ i in s2, w2 i = 1 → s1.affine_combination p w1 = s2.affine_combination p w2 → set.indicator ↑s1 w1 = set.indicator ↑s2 w2 := begin + classical, split, { intros ha s1 s2 w1 w2 hw1 hw2 heq, ext i, @@ -271,6 +273,7 @@ family. -/ lemma affine_independent.comp_embedding {ι2 : Type*} (f : ι2 ↪ ι) {p : ι → P} (ha : affine_independent k p) : affine_independent k (p ∘ f) := begin + classical, intros fs w hw hs i0 hi0, let fs' := fs.map f, let w' := λ i, if h : ∃ i2, f i2 = i then w h.some else 0, From 868ee2c662c19d3795324558cf38650b5180be0d Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 25 Sep 2022 20:15:26 +0000 Subject: [PATCH 383/828] chore(.docker): fix build failure in docker images (#16590) Co-authored-by: Scott Morrison Co-authored-by: Eric Wieser --- .docker/debian/lean/Dockerfile | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/.docker/debian/lean/Dockerfile b/.docker/debian/lean/Dockerfile index e43de629063d2..c04ea2ba15661 100644 --- a/.docker/debian/lean/Dockerfile +++ b/.docker/debian/lean/Dockerfile @@ -26,7 +26,10 @@ RUN python3 -m pip install --user mathlibtools # now install `pyinstaller`, and run it on the installed copy of `leanproject` RUN python3 -m pip install --user pyinstaller WORKDIR /home/lean/.local/bin -RUN pyinstaller --onefile --noconfirm leanproject +# We need the `--hidden-import` flag here due to PyInstaller not knowing the dependencies +# of PyNaCl (https://github.com/pyinstaller/pyinstaller/issues/3525), which is itself a transitive +# dependency of mathlibtools via PyGithub. +RUN pyinstaller --onefile --noconfirm --hidden-import _cffi_backend leanproject # this has created `/home/lean/.local/bin/dist/leanproject`, # which we'll now copy to a fresh container From 74f6e95cf899e68bade839e0afa1de70dff95ff2 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Sun, 25 Sep 2022 21:25:26 +0000 Subject: [PATCH 384/828] feat(data/real/ennreal): make of_real_sub easier to rewrite with (#16621) A tiny edit to make this lemma more general for the purpose of rewriting - previously `q` was only for `nnreal` (even though it had the assumption of nonnegativity). Co-authored-by: Bhavik Mehta --- src/data/real/ennreal.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index c976cbe79a34a..994e7e8ce337c 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -1603,7 +1603,7 @@ eq_comm.trans of_real_eq_zero alias of_real_eq_zero ↔ _ of_real_of_nonpos -lemma of_real_sub (p : ℝ) (hq : 0 ≤ q) : +lemma of_real_sub (p : ℝ) {q : ℝ} (hq : 0 ≤ q) : ennreal.of_real (p - q) = ennreal.of_real p - ennreal.of_real q := begin obtain h | h := le_total p q, From a337782b19f1f0d2ddaacef8fcd19803b864b2f4 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Sun, 25 Sep 2022 21:25:27 +0000 Subject: [PATCH 385/828] chore(data/list/range): fix incorrect docstring (#16622) The description of `iota` was incorrect (this can be easily checked by viewing its definition, using `#eval` or looking at any of the lemmas about it) Co-authored-by: Bhavik Mehta --- src/data/list/range.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/list/range.lean b/src/data/list/range.lean index fe806451fca54..aa0b3fc912bdb 100644 --- a/src/data/list/range.lean +++ b/src/data/list/range.lean @@ -14,7 +14,7 @@ import data.list.zip This file shows basic results about `list.iota`, `list.range`, `list.range'` (all defined in `data.list.defs`) and defines `list.fin_range`. `fin_range n` is the list of elements of `fin n`. -`iota n = [1, ..., n]` and `range n = [0, ..., n - 1]` are basic list constructions used for +`iota n = [n, n - 1, ..., 1]` and `range n = [0, ..., n - 1]` are basic list constructions used for tactics. `range' a b = [a, ..., a + b - 1]` is there to help prove properties about them. Actual maths should use `list.Ico` instead. -/ From 765955fbb26d61ae8cf455d34c9760337fed5540 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 26 Sep 2022 00:54:26 +0000 Subject: [PATCH 386/828] feat(linear_algebra/ray): `iff` versions of some lemmas (#16642) Add `iff` versions of some `same_ray` lemmas. These are lemmas where we already have some form of both directions, but not the `iff` version (and the separate directions are still useful on their own, since one direction is true for weaker typeclass assumptions than the `iff` version and the other is available with dot notation). --- src/linear_algebra/ray.lean | 45 +++++++++++++++++++++++++++++++++++++ 1 file changed, 45 insertions(+) diff --git a/src/linear_algebra/ray.lean b/src/linear_algebra/ray.lean index d05573e8e4a44..ea44a62866284 100644 --- a/src/linear_algebra/ray.lean +++ b/src/linear_algebra/ray.lean @@ -577,3 +577,48 @@ lemma exists_eq_smul (h : same_ray R v₁ v₂) : ⟨v₁ + v₂, h.exists_eq_smul_add⟩ end same_ray + +section linear_ordered_field + +variables {R : Type*} [linear_ordered_field R] +variables {M : Type*} [add_comm_group M] [module R M] {x y : M} + +lemma exists_pos_left_iff_same_ray (hx : x ≠ 0) (hy : y ≠ 0) : + (∃ r : R, 0 < r ∧ r • x = y) ↔ same_ray R x y := +begin + refine ⟨λ h, _, λ h, h.exists_pos_left hx hy⟩, + rcases h with ⟨r, hr, rfl⟩, + exact same_ray_pos_smul_right x hr +end + +lemma exists_pos_left_iff_same_ray_and_ne_zero (hx : x ≠ 0) : + (∃ r : R, 0 < r ∧ r • x = y) ↔ (same_ray R x y ∧ y ≠ 0) := +begin + split, + { rintro ⟨r, hr, rfl⟩, + simp [hx, hr.le, hr.ne'] }, + { rintro ⟨hxy, hy⟩, + exact (exists_pos_left_iff_same_ray hx hy).2 hxy } +end + +lemma exists_nonneg_left_iff_same_ray (hx : x ≠ 0) : + (∃ r : R, 0 ≤ r ∧ r • x = y) ↔ same_ray R x y := +begin + refine ⟨λ h, _, λ h, h.exists_nonneg_left hx⟩, + rcases h with ⟨r, hr, rfl⟩, + exact same_ray_nonneg_smul_right x hr +end + +lemma exists_pos_right_iff_same_ray (hx : x ≠ 0) (hy : y ≠ 0) : + (∃ r : R, 0 < r ∧ x = r • y) ↔ same_ray R x y := +by simpa only [same_ray_comm, eq_comm] using exists_pos_left_iff_same_ray hy hx + +lemma exists_pos_right_iff_same_ray_and_ne_zero (hy : y ≠ 0) : + (∃ r : R, 0 < r ∧ x = r • y) ↔ (same_ray R x y ∧ x ≠ 0) := +by simpa only [same_ray_comm, eq_comm] using exists_pos_left_iff_same_ray_and_ne_zero hy + +lemma exists_nonneg_right_iff_same_ray (hy : y ≠ 0) : + (∃ r : R, 0 ≤ r ∧ x = r • y) ↔ same_ray R x y := +by simpa only [same_ray_comm, eq_comm] using exists_nonneg_left_iff_same_ray hy + +end linear_ordered_field From 653307acae41aeb0ff8e45a03979ba90c32815fa Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Mon, 26 Sep 2022 02:19:54 +0000 Subject: [PATCH 387/828] chore(number_theory/cyclotomic): fix typo in lemma name (#16643) --- src/number_theory/cyclotomic/rat.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/src/number_theory/cyclotomic/rat.lean b/src/number_theory/cyclotomic/rat.lean index 0b065fea48e53..e7dfbe00438df 100644 --- a/src/number_theory/cyclotomic/rat.lean +++ b/src/number_theory/cyclotomic/rat.lean @@ -13,7 +13,7 @@ We gather results about cyclotomic extensions of `ℚ`. In particular, we comput integers of a `p ^ n`-th cyclotomic extension of `ℚ`. ## Main results -* `is_cyclotomic_extension.rat.is_integral_closure_adjoing_singleton_of_prime_pow`: if `K` is a +* `is_cyclotomic_extension.rat.is_integral_closure_adjoin_singleton_of_prime_pow`: if `K` is a `p ^ k`-th cyclotomic extension of `ℚ`, then `(adjoin ℤ {ζ})` is the integral closure of `ℤ` in `K`. * `is_cyclotomic_extension.rat.cyclotomic_ring_is_integral_closure_of_prime_pow`: the integral @@ -75,7 +75,7 @@ end /-- If `K` is a `p ^ k`-th cyclotomic extension of `ℚ`, then `(adjoin ℤ {ζ})` is the integral closure of `ℤ` in `K`. -/ -lemma is_integral_closure_adjoing_singleton_of_prime_pow +lemma is_integral_closure_adjoin_singleton_of_prime_pow [hcycl : is_cyclotomic_extension {p ^ k} ℚ K] (hζ : is_primitive_root ζ ↑(p ^ k)) : is_integral_closure (adjoin ℤ ({ζ} : set K)) ℤ K := begin @@ -121,12 +121,12 @@ begin exact subalgebra.sub_mem _ (self_mem_adjoin_singleton ℤ _) (subalgebra.one_mem _) } end -lemma is_integral_closure_adjoing_singleton_of_prime [hcycl : is_cyclotomic_extension {p} ℚ K] +lemma is_integral_closure_adjoin_singleton_of_prime [hcycl : is_cyclotomic_extension {p} ℚ K] (hζ : is_primitive_root ζ ↑p) : is_integral_closure (adjoin ℤ ({ζ} : set K)) ℤ K := begin rw [← pow_one p] at hζ hcycl, - exactI is_integral_closure_adjoing_singleton_of_prime_pow hζ, + exactI is_integral_closure_adjoin_singleton_of_prime_pow hζ, end local attribute [-instance] cyclotomic_field.algebra @@ -143,7 +143,7 @@ begin { exact ne_zero.char_zero } }, have hζ := zeta_spec (p ^ k) ℚ (cyclotomic_field (p ^ k) ℚ), refine ⟨is_fraction_ring.injective _ _, λ x, ⟨λ h, ⟨⟨x, _⟩, rfl⟩, _⟩⟩, - { have := (is_integral_closure_adjoing_singleton_of_prime_pow hζ).is_integral_iff, + { have := (is_integral_closure_adjoin_singleton_of_prime_pow hζ).is_integral_iff, obtain ⟨y, rfl⟩ := this.1 h, convert adjoin_mono _ y.2, { simp only [eq_iff_true_of_subsingleton] }, @@ -178,7 +178,7 @@ unity and `K` is a `p ^ k`-th cyclotomic extension of `ℚ`. -/ @[simps] noncomputable def _root_.is_primitive_root.adjoin_equiv_ring_of_integers [hcycl : is_cyclotomic_extension {p ^ k} ℚ K] (hζ : is_primitive_root ζ ↑(p ^ k)) : adjoin ℤ ({ζ} : set K) ≃ₐ[ℤ] (𝓞 K) := -let _ := is_integral_closure_adjoing_singleton_of_prime_pow hζ in +let _ := is_integral_closure_adjoin_singleton_of_prime_pow hζ in by exactI (is_integral_closure.equiv ℤ (adjoin ℤ ({ζ} : set K)) K (𝓞 K)) /-- The integral `power_basis` of `𝓞 K` given by a primitive root of unity, where `K` is a `p ^ k` From 19deedcc6c6217ea91c0bd0f12d3ab0884a268d7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 26 Sep 2022 03:20:38 +0000 Subject: [PATCH 388/828] feat(combinatorics/simple_graph/density): `positivity` extension for `edge_density` (#16640) Add a `positivity` extension for `rel.edge_density` and `simple_graph.edge_density`. Also golf the file a little using `positivity`. --- src/combinatorics/simple_graph/density.lean | 21 +++++++++++++++++++-- test/positivity.lean | 6 ++++++ 2 files changed, 25 insertions(+), 2 deletions(-) diff --git a/src/combinatorics/simple_graph/density.lean b/src/combinatorics/simple_graph/density.lean index e46e30d36ad9e..1adf323b46a8a 100644 --- a/src/combinatorics/simple_graph/density.lean +++ b/src/combinatorics/simple_graph/density.lean @@ -5,6 +5,7 @@ Authors: Yaël Dillies, Bhavik Mehta -/ import combinatorics.simple_graph.basic import order.partition.finpartition +import tactic.positivity /-! # Edge density @@ -163,7 +164,7 @@ begin refine (sub_le_sub_left (mul_edge_density_le_edge_density r hs ht hs₂ ht₂) _).trans _, refine le_trans _ (mul_le_of_le_one_right _ (edge_density_le_one r s₂ t₂)), { rw [sub_mul, one_mul] }, - refine sub_nonneg_of_le (mul_le_one _ (div_nonneg (nat.cast_nonneg _) (nat.cast_nonneg _)) _); + refine sub_nonneg_of_le (mul_le_one _ (by positivity) _); exact div_le_one_of_le (nat.cast_le.2 (card_le_of_subset ‹_›)) (nat.cast_nonneg _), end @@ -203,7 +204,7 @@ begin apply sub_le_sub_left (mul_le_mul ((le_div_iff _).2 hs₂) ((le_div_iff _).2 ht₂) hδ₁.le _), { exact_mod_cast (hs₂'.mono hs).card_pos }, { exact_mod_cast (ht₂'.mono ht).card_pos }, - { exact div_nonneg (nat.cast_nonneg _) (nat.cast_nonneg _) } + { positivity } end /-- If `s₂ ⊆ s₁`, `t₂ ⊆ t₁` and they take up all but a `δ`-proportion, then the difference in edge @@ -349,3 +350,19 @@ lemma edge_density_comm (s t : finset α) : G.edge_density s t = G.edge_density edge_density_comm G.symm s t end simple_graph + +namespace tactic +open positivity + +/-- Extension for the `positivity` tactic: `rel.edge_density` and `simple_graph.edge_density` are +always nonnegative. -/ +@[positivity] +meta def positivity_edge_density : expr → tactic strictness +| `(rel.edge_density %%r %%s %%t) := nonnegative <$> + mk_mapp ``rel.edge_density_nonneg [none, none, r, none, s, t] +| `(simple_graph.edge_density %%G %%s %%t) := nonnegative <$> + mk_mapp ``simple_graph.edge_density_nonneg [none, G, none, s, t] +| e := pp e >>= fail ∘ format.bracket "The expression `" + "` isn't of the form `rel.edge_density r s t` nor `simple_graph.edge_density G s t`" + +end tactic diff --git a/test/positivity.lean b/test/positivity.lean index 2595bcb6f1f73..2fb84f91783f4 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -1,6 +1,7 @@ import algebra.order.smul import analysis.normed.group.basic import analysis.special_functions.pow +import combinatorics.simple_graph.density import data.complex.exponential import data.rat.nnrat import data.real.ereal @@ -134,6 +135,11 @@ example {X : Type*} [metric_space X] (x y : X) : 0 ≤ dist x y := by positivity example {E : Type*} [add_group E] {p : add_group_seminorm E} {x : E} : 0 ≤ p x := by positivity example {E : Type*} [group E] {p : group_seminorm E} {x : E} : 0 ≤ p x := by positivity +example {r : α → β → Prop} [Π a, decidable_pred (r a)] {s : finset α} {t : finset β} : + 0 ≤ rel.edge_density r s t := by positivity +example {G : simple_graph α} [decidable_rel G.adj] {s t : finset α} : + 0 ≤ G.edge_density s t := by positivity + /- ### Canonical orders -/ example {a : ℕ} : 0 ≤ a := by positivity From a48b8eff2178578ba7f0d9fe8b6fd382a6566a2b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 26 Sep 2022 05:32:49 +0000 Subject: [PATCH 389/828] feat(data/list/basic): `list.head`, `list.head'`, and `list.tail` are surjective (#16625) Also add `list.eq_cons_of_mem_head'`, a more specific version of `list.mem_of_mem_head'` --- src/data/list/basic.lean | 15 +++++++++++++-- 1 file changed, 13 insertions(+), 2 deletions(-) diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 1dca88a4290fc..bc1987346bfe1 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -828,9 +828,20 @@ by { cases l₂, { contradiction, }, { rw list.last'_append_cons, exact h } } theorem head_eq_head' [inhabited α] (l : list α) : head l = (head' l).iget := by cases l; refl -theorem mem_of_mem_head' {x : α} : ∀ {l : list α}, x ∈ l.head' → x ∈ l +theorem surjective_head [inhabited α] : surjective (@head α _) := λ x, ⟨[x], rfl⟩ + +theorem surjective_head' : surjective (@head' α) := option.forall.2 ⟨⟨[], rfl⟩, λ x, ⟨[x], rfl⟩⟩ + +theorem surjective_tail : surjective (@tail α) +| [] := ⟨[], rfl⟩ +| (a :: l) := ⟨a :: a :: l, rfl⟩ + +lemma eq_cons_of_mem_head' {x : α} : ∀ {l : list α}, x ∈ l.head' → l = x::tail l | [] h := (option.not_mem_none _ h).elim -| (a::l) h := by { simp only [head', option.mem_def] at h, exact h ▸ or.inl rfl } +| (a::l) h := by { simp only [head', option.mem_def] at h, exact h ▸ rfl } + +theorem mem_of_mem_head' {x : α} {l : list α} (h : x ∈ l.head') : x ∈ l := +(eq_cons_of_mem_head' h).symm ▸ mem_cons_self _ _ @[simp] theorem head_cons [inhabited α] (a : α) (l : list α) : head (a::l) = a := rfl From 3b4e9d58eae1978db9bd73d4c5f9c90fd4bb65ee Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 26 Sep 2022 05:32:50 +0000 Subject: [PATCH 390/828] feat(data/option/basic): `option.map` is injective (#16626) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * prove that `option.map : (α → β) → (option α → option β)` is injective; * add `iff` version of this lemma; * add `option.map_comp_some` and `option.map_eq_id`; * drop `option.map_id'`: it was the same as `option.map_id`. --- src/data/option/basic.lean | 13 ++++++++++++- 1 file changed, 12 insertions(+), 1 deletion(-) diff --git a/src/data/option/basic.lean b/src/data/option/basic.lean index f8587890de76a..16ff17eba180f 100644 --- a/src/data/option/basic.lean +++ b/src/data/option/basic.lean @@ -88,6 +88,8 @@ theorem map_injective {f : α → β} (Hf : function.injective f) : function.inj | none none H := rfl | (some a₁) (some a₂) H := by rw Hf (option.some.inj H) +@[simp] theorem map_comp_some (f : α → β) : option.map f ∘ some = some ∘ f := rfl + @[ext] theorem ext : ∀ {o₁ o₂ : option α}, (∀ a, a ∈ o₁ ↔ a ∈ o₂) → o₁ = o₂ | none none H := rfl | (some a) o H := ((H _).1 rfl).symm @@ -179,11 +181,20 @@ by { cases x; simp only [map_none, map_some, eq_self_iff_true] } x.map f = none ↔ x = none := by { cases x; simp only [map_none', map_some', eq_self_iff_true] } +/-- `option.map` as a function between functions is injective. -/ +theorem map_injective' : function.injective (@option.map α β) := +λ f g h, funext $ λ x, some_injective _ $ by simp only [← map_some', h] + +@[simp] theorem map_inj {f g : α → β} : option.map f = option.map g ↔ f = g := +map_injective'.eq_iff + lemma map_congr {f g : α → β} {x : option α} (h : ∀ a ∈ x, f a = g a) : option.map f x = option.map g x := by { cases x; simp only [map_none', map_some', h, mem_def] } -@[simp] theorem map_id' : option.map (@id α) = id := map_id +attribute [simp] map_id + +@[simp] theorem map_eq_id {f : α → α} : option.map f = id ↔ f = id := map_injective'.eq_iff' map_id @[simp] lemma map_map (h : β → γ) (g : α → β) (x : option α) : option.map h (option.map g x) = option.map (h ∘ g) x := From ebf8ff42db1e71dd0880e3f274dd5ac9303b8a2a Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 26 Sep 2022 09:37:05 +0000 Subject: [PATCH 391/828] feat(algebraic_geometry/projective_spectrum/Scheme): the function from Spec to Proj restricted to basic open set. (#15259) We want to have a homeomorphism between $\mathrm{Proj}|_{D(f)}$ to $\mathrm{Spec}{A^0_f}$, we have the forward direction already. This PR is the underlying function of backward direction. Continuity will be proved separately. --- .../projective_spectrum/scheme.lean | 176 +++++++++++++++++- 1 file changed, 170 insertions(+), 6 deletions(-) diff --git a/src/algebraic_geometry/projective_spectrum/scheme.lean b/src/algebraic_geometry/projective_spectrum/scheme.lean index 06965cb6ab393..19c57bcbd7859 100644 --- a/src/algebraic_geometry/projective_spectrum/scheme.lean +++ b/src/algebraic_geometry/projective_spectrum/scheme.lean @@ -5,6 +5,7 @@ Authors: Jujian Zhang -/ import algebraic_geometry.projective_spectrum.structure_sheaf import algebraic_geometry.Spec +import ring_theory.graded_algebra.radical /-! # Proj as a scheme @@ -21,7 +22,7 @@ This file is to prove that `Proj` is a scheme. * `Spec` : `Spec` as a locally ringed space * `Spec.T` : the underlying topological space of `Spec` * `sbo g` : basic open set at `g` in `Spec` -* `A⁰_x` : the degree zero part of localized ring `Aₓ` +* `A⁰_x` : the degree zero part of localized ring `Aₓ` ## Implementation @@ -31,13 +32,27 @@ equipped with this structure sheaf is a scheme. We achieve this by using an affi open sets in `Proj`, more specifically: 1. We prove that `Proj` can be covered by basic open sets at homogeneous element of positive degree. -2. We prove that for any `f : A`, `Proj.T | (pbo f)` is homeomorphic to `Spec.T A⁰_f`: +2. We prove that for any homogeneous element `f : A` of positive degree `m`, `Proj.T | (pbo f)` is + homeomorphic to `Spec.T A⁰_f`: - forward direction `to_Spec`: for any `x : pbo f`, i.e. a relevant homogeneous prime ideal `x`, send it to - `x ∩ span {g / 1 | g ∈ A}` (see `Proj_iso_Spec_Top_component.to_Spec.carrier`). This ideal is + `A⁰_f ∩ span {g / 1 | g ∈ x}` (see `Proj_iso_Spec_Top_component.to_Spec.carrier`). This ideal is prime, the proof is in `Proj_iso_Spec_Top_component.to_Spec.to_fun`. The fact that this function is continuous is found in `Proj_iso_Spec_Top_component.to_Spec` - - backward direction `from_Spec`: TBC + - backward direction `from_Spec`: + for any `q : Spec A⁰_f`, we send it to `{a | ∀ i, aᵢᵐ/fⁱ ∈ q}`; we need this to be a + homogeneous prime ideal that is relevant. + * This is in fact an ideal, the proof can be found in + `Proj_iso_Spec_Top_component.from_Spec.carrier.as_ideal`; + * This ideal is also homogeneous, the proof can be found in + `Proj_iso_Spec_Top_component.from_Spec.carrier.as_ideal.homogeneous`; + * This ideal is relevant, the proof can be found in + `Proj_iso_Spec_Top_component.from_Spec.carrier.relevant`; + * This ideal is prime, the proof can be found in + `Proj_iso_Spec_Top_component.from_Spec.carrier.prime`. + Hence we have a well defined function `Spec.T A⁰_f → Proj.T | (pbo f)`, this function is called + `Proj_iso_Spec_Top_component.from_Spec.to_fun`. But to prove the continuity of this function, + we need to prove `from_Spec ∘ to_Spec` and `to_Spec ∘ from_Spec` are both identities (TBC). ## Main Definitions and Statements @@ -155,6 +170,14 @@ x.2.some_spec.some_spec lemma degree_zero_part.coe_mul {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) (x y : A⁰_ f_deg) : (↑(x * y) : away f) = x * y := rfl +lemma degree_zero_part.coe_one {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) : + (↑(1 : A⁰_ f_deg) : away f) = 1 := rfl + +lemma degree_zero_part.coe_sum {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) {ι : Type*} + (s : finset ι) (g : ι → A⁰_ f_deg) : + (↑(∑ i in s, g i) : away f) = ∑ i in s, (g i : away f) := +by { classical, induction s using finset.induction_on with i s hi ih; simp } + end namespace Proj_iso_Spec_Top_component @@ -201,8 +224,6 @@ begin obtain ⟨⟨_, N, rfl⟩, hN⟩ := is_localization.exist_integer_multiples_of_finset (submonoid.powers f) (c.support.image c), choose acd hacd using hN, - have prop1 : ∀ i, i ∈ c.support → c i ∈ finset.image c c.support, - { intros i hi, rw finset.mem_image, refine ⟨_, hi, rfl⟩, }, refine ⟨c, N, acd, _⟩, rw [← eq1, smul_sum, map_sum, ← sum_attach], @@ -359,6 +380,149 @@ def to_Spec {f : A} (m : ℕ) (f_deg : f ∈ 𝒜 m) : end +namespace from_Spec + +open graded_algebra set_like finset (hiding mk_zero) + +variables {𝒜} {f : A} {m : ℕ} {f_deg : f ∈ 𝒜 m} + +private meta def mem_tac : tactic unit := +let b : tactic unit := + `[exact pow_mem_graded _ (submodule.coe_mem _) <|> exact nat_cast_mem_graded _ _] in +b <|> `[by repeat { all_goals { apply graded_monoid.mul_mem } }; b] + +/--The function from `Spec A⁰_f` to `Proj|D(f)` is defined by `q ↦ {a | aᵢᵐ/fⁱ ∈ q}`, i.e. sending +`q` a prime ideal in `A⁰_f` to the homogeneous prime relevant ideal containing only and all the +elements `a : A` such that for every `i`, the degree 0 element formed by dividing the `m`-th power +of the `i`-th projection of `a` by the `i`-th power of the degree-`m` homogeneous element `f`, +lies in `q`. + +The set `{a | aᵢᵐ/fⁱ ∈ q}` +* is an ideal, as proved in `carrier.as_ideal`; +* is homogeneous, as proved in `carrier.as_homogeneous_ideal`; +* is prime, as proved in `carrier.as_ideal.prime`; +* is relevant, as proved in `carrier.relevant`. +-/ +def carrier (q : Spec.T (A⁰_ f_deg)) : set A := +{a | ∀ i, (⟨mk (proj 𝒜 i a ^ m) ⟨_, _, rfl⟩, i, ⟨_, by mem_tac⟩, rfl⟩ : A⁰_ f_deg) ∈ q.1} + +lemma mem_carrier_iff (q : Spec.T (A⁰_ f_deg)) (a : A) : + a ∈ carrier q ↔ + ∀ i, (⟨mk (proj 𝒜 i a ^ m) ⟨_, _, rfl⟩, i, ⟨_, by mem_tac⟩, rfl⟩ : A⁰_ f_deg) ∈ q.1 := +iff.rfl + +lemma carrier.add_mem (q : Spec.T (A⁰_ f_deg)) {a b : A} (ha : a ∈ carrier q) (hb : b ∈ carrier q) : + a + b ∈ carrier q := +begin + refine λ i, (q.2.mem_or_mem _).elim id id, + change subtype.mk (localization.mk _ _ * mk _ _) _ ∈ q.1, + simp_rw [mk_mul, ← pow_add, map_add, add_pow, mk_sum, mul_comm, ← nsmul_eq_mul, ← smul_mk], + let g : ℕ → A⁰_ f_deg := λ j, (m + m).choose j • if h2 : m + m < j then 0 else if h1 : j ≤ m + then ⟨mk (proj 𝒜 i a ^ j * proj 𝒜 i b ^ (m - j)) ⟨_, i, rfl⟩, i, ⟨_, _⟩, rfl⟩ * + ⟨mk (proj 𝒜 i b ^ m) ⟨_, i, rfl⟩, i, ⟨_, by mem_tac⟩, rfl⟩ + else ⟨mk (proj 𝒜 i a ^ m) ⟨_, i, rfl⟩, i, ⟨_, by mem_tac⟩, rfl⟩ * + ⟨mk (proj 𝒜 i a ^ (j - m) * proj 𝒜 i b ^ (m + m - j)) ⟨_, i, rfl⟩, i, ⟨_, _⟩, rfl⟩, + rotate, + { rw (_ : m * i = _), mem_tac, rw [← add_smul, nat.add_sub_of_le h1], refl }, + { rw (_ : m * i = _), mem_tac, rw ← add_smul, congr, + zify [le_of_not_lt h2, le_of_not_le h1], abel }, + convert_to ∑ i in range (m + m + 1), g i ∈ q.1, swap, + { refine q.1.sum_mem (λ j hj, nsmul_mem _ _), split_ifs, + exacts [q.1.zero_mem, q.1.mul_mem_left _ (hb i), q.1.mul_mem_right _ (ha i)] }, + apply subtype.ext, + rw [degree_zero_part.coe_sum, subtype.coe_mk], + apply finset.sum_congr rfl (λ j hj, _), + congr' 1, split_ifs with h2 h1, + { exact ((finset.mem_range.1 hj).not_le h2).elim }, + all_goals { simp only [subtype.val_eq_coe, degree_zero_part.coe_mul, subtype.coe_mk, mk_mul] }, + { rw [mul_assoc, ← pow_add, add_comm (m - j), nat.add_sub_assoc h1] }, + { rw [← mul_assoc, ← pow_add, nat.add_sub_of_le (le_of_not_le h1)] }, +end + +variables (hm : 0 < m) (q : Spec.T (A⁰_ f_deg)) +include hm + +lemma carrier.zero_mem : (0 : A) ∈ carrier q := +λ i, by simpa only [linear_map.map_zero, zero_pow hm, mk_zero] using submodule.zero_mem _ + +lemma carrier.smul_mem (c x : A) (hx : x ∈ carrier q) : c • x ∈ carrier q := +begin + revert c, + refine direct_sum.decomposition.induction_on 𝒜 _ _ _, + { rw zero_smul, exact carrier.zero_mem hm _ }, + { rintros n ⟨a, ha⟩ i, + simp_rw [subtype.coe_mk, proj_apply, smul_eq_mul, coe_decompose_mul_of_left_mem 𝒜 i ha], + split_ifs, + { convert_to (⟨mk _ ⟨_, n, rfl⟩, n, ⟨_, pow_mem_graded m ha⟩, rfl⟩ : A⁰_ f_deg) * + ⟨mk _ ⟨_, i - n, rfl⟩, _, ⟨proj 𝒜 (i - n) x ^ m, by mem_tac⟩, rfl⟩ ∈ q.1, + { erw [subtype.ext_iff, subring.coe_mul, mk_mul, subtype.coe_mk, mul_pow], + congr, erw [← pow_add, nat.add_sub_of_le h] }, + { exact ideal.mul_mem_left _ _ (hx _) } }, + { simp_rw [zero_pow hm, mk_zero], exact q.1.zero_mem } }, + { simp_rw add_smul, exact λ _ _, carrier.add_mem q }, +end + +/-- +For a prime ideal `q` in `A⁰_f`, the set `{a | aᵢᵐ/fⁱ ∈ q}` as an ideal. +-/ +def carrier.as_ideal : ideal A := +{ carrier := carrier q, + zero_mem' := carrier.zero_mem hm q, + add_mem' := λ a b, carrier.add_mem q, + smul_mem' := carrier.smul_mem hm q } + +lemma carrier.as_ideal.homogeneous : (carrier.as_ideal hm q).is_homogeneous 𝒜 := +λ i a ha j, (em (i = j)).elim + (λ h, h ▸ by simpa only [proj_apply, decompose_coe, of_eq_same] using ha _) + (λ h, by simpa only [proj_apply, decompose_of_mem_ne 𝒜 (submodule.coe_mem (decompose 𝒜 a i)) h, + zero_pow hm, mk_zero] using submodule.zero_mem _) + +/-- +For a prime ideal `q` in `A⁰_f`, the set `{a | aᵢᵐ/fⁱ ∈ q}` as a homogeneous ideal. +-/ +def carrier.as_homogeneous_ideal : homogeneous_ideal 𝒜 := +⟨carrier.as_ideal hm q, carrier.as_ideal.homogeneous hm q⟩ + +lemma carrier.denom_not_mem : f ∉ carrier.as_ideal hm q := +λ rid, q.is_prime.ne_top $ (ideal.eq_top_iff_one _).mpr +begin + convert rid m, + simpa only [subtype.ext_iff, degree_zero_part.coe_one, subtype.coe_mk, proj_apply, + decompose_of_mem_same _ f_deg] using (mk_self (⟨_, m, rfl⟩ : submonoid.powers f)).symm, +end + +lemma carrier.relevant : ¬ homogeneous_ideal.irrelevant 𝒜 ≤ carrier.as_homogeneous_ideal hm q := +λ rid, carrier.denom_not_mem hm q $ rid $ direct_sum.decompose_of_mem_ne 𝒜 f_deg hm.ne' + +lemma carrier.as_ideal.ne_top : (carrier.as_ideal hm q) ≠ ⊤ := +λ rid, carrier.denom_not_mem hm q (rid.symm ▸ submodule.mem_top) + +lemma carrier.as_ideal.prime : (carrier.as_ideal hm q).is_prime := +(carrier.as_ideal.homogeneous hm q).is_prime_of_homogeneous_mem_or_mem + (carrier.as_ideal.ne_top hm q) $ λ x y ⟨nx, hnx⟩ ⟨ny, hny⟩ hxy, show (∀ i, _ ∈ _) ∨ ∀ i, _ ∈ _, +begin + rw [← and_forall_ne nx, and_iff_left, ← and_forall_ne ny, and_iff_left], + { apply q.2.mem_or_mem, convert hxy (nx + ny) using 1, + simp_rw [proj_apply, decompose_of_mem_same 𝒜 hnx, decompose_of_mem_same 𝒜 hny, + decompose_of_mem_same 𝒜 (mul_mem hnx hny), mul_pow, pow_add], + exact subtype.ext (mk_mul _ _ _ _) }, + all_goals { intros n hn, + convert q.1.zero_mem using 2, + rw [proj_apply, decompose_of_mem_ne 𝒜 _ hn.symm], + { rw [zero_pow hm, mk_zero] }, + { exact hnx <|> exact hny } }, +end + +variable (f_deg) +/-- +The function `Spec A⁰_f → Proj|D(f)` by sending `q` to `{a | aᵢᵐ/fⁱ ∈ q}`. +-/ +def to_fun : (Spec.T (A⁰_ f_deg)) → (Proj.T| (pbo f)) := +λ q, ⟨⟨carrier.as_homogeneous_ideal hm q, carrier.as_ideal.prime hm q, carrier.relevant hm q⟩, + (projective_spectrum.mem_basic_open _ f _).mp $ carrier.denom_not_mem hm q⟩ + +end from_Spec + end Proj_iso_Spec_Top_component end algebraic_geometry From 525ffd73d8a8c1a8b899b5969edd0c2c286bd050 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Mon, 26 Sep 2022 12:01:06 +0000 Subject: [PATCH 392/828] feat(group_theory/subgroup/basic): golf proof of instance normal_inf_normal (#16620) Removed some `rw` from the proof of the instance `normal_inf_normal`. --- src/group_theory/subgroup/basic.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index e48c5119b27b0..30b944cd279d4 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2989,8 +2989,7 @@ instance sup_normal (H K : subgroup G) [hH : H.normal] [hK : K.normal] : (H ⊔ @[to_additive] instance normal_inf_normal (H K : subgroup G) [hH : H.normal] [hK : K.normal] : (H ⊓ K).normal := -{ conj_mem := λ n hmem g, - by { rw mem_inf at *, exact ⟨hH.conj_mem n hmem.1 g, hK.conj_mem n hmem.2 g⟩ } } +⟨λ n hmem g, ⟨hH.conj_mem n hmem.1 g, hK.conj_mem n hmem.2 g⟩⟩ @[to_additive] lemma subgroup_of_sup (A A' B : subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : (A ⊔ A').subgroup_of B = A.subgroup_of B ⊔ A'.subgroup_of B := From d5e1961fa552d1d1e6f74989af3dc3e467ba476c Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 26 Sep 2022 14:37:53 +0000 Subject: [PATCH 393/828] feat(topology/sheaves/abelian): category of sheaves is abelian (#16403) and that sheafification functor is additive --- src/topology/sheaves/abelian.lean | 57 +++++++++++++++++++++++++++++++ 1 file changed, 57 insertions(+) create mode 100644 src/topology/sheaves/abelian.lean diff --git a/src/topology/sheaves/abelian.lean b/src/topology/sheaves/abelian.lean new file mode 100644 index 0000000000000..e8a2a18b3c144 --- /dev/null +++ b/src/topology/sheaves/abelian.lean @@ -0,0 +1,57 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Adam Topaz, Jujian Zhang +-/ +import category_theory.abelian.functor_category +import category_theory.preadditive.functor_category +import category_theory.abelian.transfer +import category_theory.sites.left_exact + +/-! +# Category of sheaves is abelian +Let `C, D` be categories and `J` be a grothendieck topology on `C`, when `D` is abelian and +sheafification is possible in `C`, `Sheaf J D` is abelian as well (`Sheaf_is_abelian`). + +Hence, `presheaf_to_Sheaf` is an additive functor (`presheaf_to_Sheaf_additive`). + +-/ + +noncomputable theory + +namespace category_theory + +open category_theory.limits + +section abelian + +universes w v u +variables {C : Type (max v u)} [category.{v} C] +variables {D : Type w} [category.{max v u} D] [abelian D] +variables {J : grothendieck_topology C} + +-- This needs to be specified manually because of universe level. +instance : abelian (Cᵒᵖ ⥤ D) := @abelian.functor_category_abelian.{v} Cᵒᵖ _ D _ _ + +-- This also needs to be specified manually, but I don't know why. +instance : has_finite_products (Sheaf J D) := +{ out := λ j ij, { has_limit := by { introI, apply_instance } } } + +-- sheafification assumptions +variables [∀ (P : Cᵒᵖ ⥤ D) (X : C) (S : J.cover X), has_multiequalizer (S.index P)] +variables [∀ (X : C), has_colimits_of_shape (J.cover X)ᵒᵖ D] +variables [concrete_category.{max v u} D] [preserves_limits (forget D)] +variables [∀ (X : C), preserves_colimits_of_shape (J.cover X)ᵒᵖ (forget D)] +variables [reflects_isomorphisms (forget D)] + +instance Sheaf_is_abelian [has_finite_limits D] : abelian (Sheaf J D) := +let adj := sheafification_adjunction J D in abelian_of_adjunction _ _ (as_iso adj.counit) adj + +local attribute [instance] preserves_binary_biproducts_of_preserves_binary_products + +instance presheaf_to_Sheaf_additive : (presheaf_to_Sheaf J D).additive := +(presheaf_to_Sheaf J D).additive_of_preserves_binary_biproducts + +end abelian + +end category_theory From a44ce4daea1daa8183717758d5f61b8158225665 Mon Sep 17 00:00:00 2001 From: Devon Tuma Date: Mon, 26 Sep 2022 17:06:58 +0000 Subject: [PATCH 394/828] feat(real/ennreal): basic lemmas about `ennreal.to_nnreal` and `ennreal.to_real` (#16318) Provides lemmas for when `to_nnreal` or `to_real` equal `1`, and when two instances of `to_nnreal` and `to_real` equal each other. --- src/data/real/ennreal.lean | 34 ++++++++++++++++++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 994e7e8ce337c..6f0d22f562e8e 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -190,6 +190,17 @@ by rintro (h | h); simp [h]⟩ lemma to_real_eq_zero_iff (x : ℝ≥0∞) : x.to_real = 0 ↔ x = 0 ∨ x = ∞ := by simp [ennreal.to_real, to_nnreal_eq_zero_iff] +lemma to_nnreal_eq_one_iff (x : ℝ≥0∞) : x.to_nnreal = 1 ↔ x = 1 := +begin + refine ⟨λ h, _, congr_arg _⟩, + cases x, + { exact false.elim (zero_ne_one $ ennreal.top_to_nnreal.symm.trans h) }, + { exact congr_arg _ h } +end + +lemma to_real_eq_one_iff (x : ℝ≥0∞) : x.to_real = 1 ↔ x = 1 := +by rw [ennreal.to_real, nnreal.coe_eq_one, ennreal.to_nnreal_eq_one_iff] + @[simp] lemma coe_ne_top : (r : ℝ≥0∞) ≠ ∞ := with_top.coe_ne_top @[simp] lemma top_ne_coe : ∞ ≠ (r : ℝ≥0∞) := with_top.top_ne_coe @[simp] lemma of_real_ne_top {r : ℝ} : ennreal.of_real r ≠ ∞ := by simp [ennreal.of_real] @@ -222,6 +233,29 @@ lemma coe_ne_zero : (r : ℝ≥0∞) ≠ 0 ↔ r ≠ 0 := not_congr coe_eq_coe @[simp, norm_cast] lemma coe_bit1 : (↑(bit1 r) : ℝ≥0∞) = bit1 r := by simp [bit1] lemma coe_two : ((2:ℝ≥0) : ℝ≥0∞) = 2 := by norm_cast +lemma to_nnreal_eq_to_nnreal_iff (x y : ℝ≥0∞) : + x.to_nnreal = y.to_nnreal ↔ x = y ∨ (x = 0 ∧ y = ⊤) ∨ (x = ⊤ ∧ y = 0) := +begin + cases x, + { simp only [@eq_comm ℝ≥0 _ y.to_nnreal, @eq_comm ℝ≥0∞ _ y, to_nnreal_eq_zero_iff, + none_eq_top, top_to_nnreal, top_ne_zero, false_and, eq_self_iff_true, + true_and, false_or, or_comm (y = ⊤)] }, + { cases y; simp } +end + +lemma to_real_eq_to_real_iff (x y : ℝ≥0∞) : + x.to_real = y.to_real ↔ x = y ∨ (x = 0 ∧ y = ⊤) ∨ (x = ⊤ ∧ y = 0) := +by simp only [ennreal.to_real, nnreal.coe_eq, to_nnreal_eq_to_nnreal_iff] + +lemma to_nnreal_eq_to_nnreal_iff' {x y : ℝ≥0∞} (hx : x ≠ ⊤) (hy : y ≠ ⊤) : + x.to_nnreal = y.to_nnreal ↔ x = y := +by simp only [ennreal.to_nnreal_eq_to_nnreal_iff x y, hx, hy, and_false, false_and, or_false] + +lemma to_real_eq_to_real_iff' {x y : ℝ≥0∞} (hx : x ≠ ⊤) (hy : y ≠ ⊤) : + x.to_real = y.to_real ↔ x = y := +by simp only [ennreal.to_real, nnreal.coe_eq, to_nnreal_eq_to_nnreal_iff' hx hy] + + protected lemma zero_lt_one : 0 < (1 : ℝ≥0∞) := canonically_ordered_comm_semiring.zero_lt_one From dfd17f2fb8ecb5a602eaa05c61c2e474b6f7e79e Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Mon, 26 Sep 2022 17:06:59 +0000 Subject: [PATCH 395/828] feat(data/fintype/basic): `set.to_finset_of_finset` (#16460) --- src/data/fintype/basic.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 6a667aaa9b7dd..ff867a1a59911 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -624,6 +624,12 @@ by simp [to_finset] @[simp] theorem mem_to_finset_val {s : set α} [fintype s] {a : α} : a ∈ s.to_finset.1 ↔ a ∈ s := mem_to_finset +/-- Many `fintype` instances for sets are defined using an extensionally equal `finset`. +Rewriting `s.to_finset` with `set.to_finset_of_finset` replaces the term with such a `finset`. -/ +theorem to_finset_of_finset {p : set α} (s : finset α) (H : ∀ x, x ∈ s ↔ x ∈ p) : + @set.to_finset _ p (fintype.of_finset s H) = s := +finset.ext (λ x, by rw [mem_to_finset, H]) + /-- Membership of a set with a `fintype` instance is decidable. Using this as an instance leads to potential loops with `subtype.fintype` under certain decidability From b41babaaebec680008cf246a0b7a2ea23273d580 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Mon, 26 Sep 2022 17:07:00 +0000 Subject: [PATCH 396/828] doc(analysis/special_functions/stirling): fix equation (#16651) --- src/analysis/special_functions/stirling.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index d2f08f2a4e514..272bd7a5d729c 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -42,7 +42,7 @@ namespace stirling -/ /-- -Define `stirling_seq n` as $\frac{n!}{\sqrt{2n}/(\frac{n}{e})^n$. +Define `stirling_seq n` as $\frac{n!}{\sqrt{2n}(\frac{n}{e})^n}$. Stirling's formula states that this sequence has limit $\sqrt(π)$. -/ noncomputable def stirling_seq (n : ℕ) : ℝ := From a12363e291959854e6f187567e8ab1eae3ca6d77 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 26 Sep 2022 17:07:01 +0000 Subject: [PATCH 397/828] perf(analysis/normed_space/continuous_affine_map): Speedup (#16652) This changes the compilation time of `continuous_affine_map.to_const_prod_continuous_linear_map` from 19s to 3s for me locally. --- src/analysis/normed_space/continuous_affine_map.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/analysis/normed_space/continuous_affine_map.lean b/src/analysis/normed_space/continuous_affine_map.lean index 330977e666af5..3d990b8987bfa 100644 --- a/src/analysis/normed_space/continuous_affine_map.lean +++ b/src/analysis/normed_space/continuous_affine_map.lean @@ -220,9 +220,9 @@ def to_const_prod_continuous_linear_map : (V →A[𝕜] W) ≃ₗᵢ[𝕜] W × inv_fun := λ p, p.2.to_continuous_affine_map + const 𝕜 V p.1, left_inv := λ f, by { ext, rw f.decomp, simp, }, right_inv := by { rintros ⟨v, f⟩, ext; simp, }, - map_add' := by simp, - map_smul' := by simp, - norm_map' := λ f, by simp [prod.norm_def, norm_def], } + map_add' := λ _ _, rfl, + map_smul' := λ _ _, rfl, + norm_map' := λ f, rfl } @[simp] lemma to_const_prod_continuous_linear_map_fst (f : V →A[𝕜] W) : (to_const_prod_continuous_linear_map 𝕜 V W f).fst = f 0 := From db24ba7cd284f91316bc556dbf5ebff9068a706c Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 26 Sep 2022 19:49:20 +0000 Subject: [PATCH 398/828] chore(number_theory/padics): stylistic changes (#15755) Extracted stylistic changes (backticks, brackets, linebreaks, naming, spacing) from #15221 as suggested in the comments. --- src/number_theory/padics/padic_integers.lean | 153 ++++++------- src/number_theory/padics/padic_norm.lean | 112 ++++------ src/number_theory/padics/padic_numbers.lean | 204 ++++++++--------- src/number_theory/padics/padic_val.lean | 218 ++++++++----------- src/number_theory/padics/ring_homs.lean | 2 +- 5 files changed, 287 insertions(+), 402 deletions(-) diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index 3a151d38e601a..599aeb437e32d 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -11,30 +11,30 @@ import topology.metric_space.cau_seq_filter /-! # p-adic integers -This file defines the p-adic integers `ℤ_p` as the subtype of `ℚ_p` with norm `≤ 1`. -We show that `ℤ_p` -* is complete -* is nonarchimedean -* is a normed ring -* is a local ring -* is a discrete valuation ring +This file defines the `p`-adic integers `ℤ_[p]` as the subtype of `ℚ_[p]` with norm `≤ 1`. +We show that `ℤ_[p]` +* is complete, +* is nonarchimedean, +* is a normed ring, +* is a local ring, and +* is a discrete valuation ring. The relation between `ℤ_[p]` and `zmod p` is established in another file. ## Important definitions -* `padic_int` : the type of p-adic numbers +* `padic_int` : the type of `p`-adic integers ## Notation -We introduce the notation `ℤ_[p]` for the p-adic integers. +We introduce the notation `ℤ_[p]` for the `p`-adic integers. ## Implementation notes Much, but not all, of this file assumes that `p` is prime. This assumption is inferred automatically -by taking `[fact (nat.prime p)] as a type class argument. +by taking `[fact p.prime]` as a type class argument. -Coercions into `ℤ_p` are set up to work with the `norm_cast` tactic. +Coercions into `ℤ_[p]` are set up to work with the `norm_cast` tactic. ## References @@ -51,8 +51,9 @@ open padic metric local_ring noncomputable theory open_locale classical -/-- The p-adic integers ℤ_p are the p-adic numbers with norm ≤ 1. -/ +/-- The `p`-adic integers `ℤ_[p]` are the `p`-adic numbers with norm `≤ 1`. -/ def padic_int (p : ℕ) [fact p.prime] := {x : ℚ_[p] // ∥x∥ ≤ 1} + notation `ℤ_[`p`]` := padic_int p namespace padic_int @@ -67,7 +68,7 @@ lemma ext {x y : ℤ_[p]} : (x : ℚ_[p]) = y → x = y := subtype.ext variables (p) -/-- The padic integers as a subring of the padics -/ +/-- The `p`-adic integers as a subring of `ℚ_[p]`. -/ def subring : subring (ℚ_[p]) := { carrier := {x : ℚ_[p] | ∥x∥ ≤ 1}, zero_mem' := by norm_num, @@ -80,26 +81,25 @@ def subring : subring (ℚ_[p]) := variables {p} -/-- Addition on ℤ_p is inherited from ℚ_p. -/ +/-- Addition on `ℤ_[p]` is inherited from `ℚ_[p]`. -/ instance : has_add ℤ_[p] := (by apply_instance : has_add (subring p)) -/-- Multiplication on ℤ_p is inherited from ℚ_p. -/ +/-- Multiplication on `ℤ_[p]` is inherited from `ℚ_[p]`. -/ instance : has_mul ℤ_[p] := (by apply_instance : has_mul (subring p)) -/-- Negation on ℤ_p is inherited from ℚ_p. -/ +/-- Negation on `ℤ_[p]` is inherited from `ℚ_[p]`. -/ instance : has_neg ℤ_[p] := (by apply_instance : has_neg (subring p)) -/-- Subtraction on ℤ_p is inherited from ℚ_p. -/ +/-- Subtraction on `ℤ_[p]` is inherited from `ℚ_[p]`. -/ instance : has_sub ℤ_[p] := (by apply_instance : has_sub (subring p)) -/-- Zero on ℤ_p is inherited from ℚ_p. -/ +/-- Zero on `ℤ_[p]` is inherited from `ℚ_[p]`. -/ instance : has_zero ℤ_[p] := (by apply_instance : has_zero (subring p)) instance : inhabited ℤ_[p] := ⟨0⟩ -/-- One on ℤ_p is inherited from ℚ_p. -/ -instance : has_one ℤ_[p] := -⟨⟨1, by norm_num⟩⟩ +/-- One on `ℤ_[p]` is inherited from `ℚ_[p]`. -/ +instance : has_one ℤ_[p] := ⟨⟨1, by norm_num⟩⟩ @[simp] lemma mk_zero {h} : (⟨0, h⟩ : ℤ_[p]) = (0 : ℤ_[p]) := rfl @@ -121,17 +121,17 @@ instance : comm_ring ℤ_[p] := @[simp, norm_cast] lemma coe_nat_cast (n : ℕ) : ((n : ℤ_[p]) : ℚ_[p]) = n := rfl @[simp, norm_cast] lemma coe_int_cast (z : ℤ) : ((z : ℤ_[p]) : ℚ_[p]) = z := rfl -/-- The coercion from ℤ[p] to ℚ[p] as a ring homomorphism. -/ +/-- The coercion from `ℤ_[p]` to `ℚ_[p]` as a ring homomorphism. -/ def coe.ring_hom : ℤ_[p] →+* ℚ_[p] := (subring p).subtype @[simp, norm_cast] lemma coe_pow (x : ℤ_[p]) (n : ℕ) : (↑(x^n) : ℚ_[p]) = (↑x : ℚ_[p])^n := rfl @[simp] lemma mk_coe (k : ℤ_[p]) : (⟨k, k.2⟩ : ℤ_[p]) = k := subtype.coe_eta _ _ -/-- The inverse of a p-adic integer with norm equal to 1 is also a p-adic integer. Otherwise, the -inverse is defined to be 0. -/ +/-- The inverse of a `p`-adic integer with norm equal to `1` is also a `p`-adic integer. +Otherwise, the inverse is defined to be `0`. -/ def inv : ℤ_[p] → ℤ_[p] -| ⟨k, _⟩ := if h : ∥k∥ = 1 then ⟨1/k, by simp [h]⟩ else 0 +| ⟨k, _⟩ := if h : ∥k∥ = 1 then ⟨k⁻¹, by simp [h]⟩ else 0 instance : char_zero ℤ_[p] := { cast_injective := @@ -142,10 +142,8 @@ instance : char_zero ℤ_[p] := suffices (z1 : ℚ_[p]) = z2 ↔ z1 = z2, from iff.trans (by norm_cast) this, by norm_cast -/-- -A sequence of integers that is Cauchy with respect to the `p`-adic norm -converges to a `p`-adic integer. --/ +/-- A sequence of integers that is Cauchy with respect to the `p`-adic norm converges to a `p`-adic +integer. -/ def of_int_seq (seq : ℕ → ℤ) (h : is_cau_seq (padic_norm p) (λ n, seq n)) : ℤ_[p] := ⟨⟦⟨_, h⟩⟧, show ↑(padic_seq.norm _) ≤ (1 : ℝ), begin @@ -194,12 +192,11 @@ instance is_absolute_value : is_absolute_value (λ z : ℤ_[p], ∥z∥) := { abv_nonneg := norm_nonneg, abv_eq_zero := λ ⟨_, _⟩, by simp [norm_eq_zero], abv_add := λ ⟨_,_⟩ ⟨_, _⟩, norm_add_le _ _, - abv_mul := λ _ _, by simp only [norm_def, padic_norm_e.mul, padic_int.coe_mul]} + abv_mul := λ _ _, by simp only [norm_def, padic_norm_e.mul, padic_int.coe_mul] } variables {p} -instance : is_domain ℤ_[p] := -function.injective.is_domain (subring p).subtype subtype.coe_injective +instance : is_domain ℤ_[p] := function.injective.is_domain (subring p).subtype subtype.coe_injective end padic_int @@ -211,45 +208,35 @@ variables {p : ℕ} [fact p.prime] lemma norm_le_one (z : ℤ_[p]) : ∥z∥ ≤ 1 := z.2 -@[simp] lemma norm_mul (z1 z2 : ℤ_[p]) : ∥z1 * z2∥ = ∥z1∥ * ∥z2∥ := -by simp [norm_def] +@[simp] lemma norm_mul (z1 z2 : ℤ_[p]) : ∥z1 * z2∥ = ∥z1∥ * ∥z2∥ := by simp [norm_def] -@[simp] lemma norm_pow (z : ℤ_[p]) : ∀ n : ℕ, ∥z^n∥ = ∥z∥^n +@[simp] lemma norm_pow (z : ℤ_[p]) : ∀ n : ℕ, ∥z ^ n∥ = ∥z∥ ^ n | 0 := by simp -| (k+1) := by { rw [pow_succ, pow_succ, norm_mul], congr, apply norm_pow } +| (k + 1) := by { rw [pow_succ, pow_succ, norm_mul], congr, apply norm_pow } theorem nonarchimedean (q r : ℤ_[p]) : ∥q + r∥ ≤ max (∥q∥) (∥r∥) := padic_norm_e.nonarchimedean _ _ theorem norm_add_eq_max_of_ne {q r : ℤ_[p]} : ∥q∥ ≠ ∥r∥ → ∥q+r∥ = max (∥q∥) (∥r∥) := padic_norm_e.add_eq_max_of_ne -lemma norm_eq_of_norm_add_lt_right {z1 z2 : ℤ_[p]} - (h : ∥z1 + z2∥ < ∥z2∥) : ∥z1∥ = ∥z2∥ := -by_contradiction $ λ hne, - not_lt_of_ge (by rw norm_add_eq_max_of_ne hne; apply le_max_right) h +lemma norm_eq_of_norm_add_lt_right {z1 z2 : ℤ_[p]} (h : ∥z1 + z2∥ < ∥z2∥) : ∥z1∥ = ∥z2∥ := +by_contradiction $ λ hne, not_lt_of_ge (by rw norm_add_eq_max_of_ne hne; apply le_max_right) h -lemma norm_eq_of_norm_add_lt_left {z1 z2 : ℤ_[p]} - (h : ∥z1 + z2∥ < ∥z1∥) : ∥z1∥ = ∥z2∥ := -by_contradiction $ λ hne, - not_lt_of_ge (by rw norm_add_eq_max_of_ne hne; apply le_max_left) h +lemma norm_eq_of_norm_add_lt_left {z1 z2 : ℤ_[p]} (h : ∥z1 + z2∥ < ∥z1∥) : ∥z1∥ = ∥z2∥ := +by_contradiction $ λ hne, not_lt_of_ge (by rw norm_add_eq_max_of_ne hne; apply le_max_left) h -@[simp] lemma padic_norm_e_of_padic_int (z : ℤ_[p]) : ∥(↑z : ℚ_[p])∥ = ∥z∥ := -by simp [norm_def] +@[simp] lemma padic_norm_e_of_padic_int (z : ℤ_[p]) : ∥(z : ℚ_[p])∥ = ∥z∥ := by simp [norm_def] -lemma norm_int_cast_eq_padic_norm (z : ℤ) : ∥(z : ℤ_[p])∥ = ∥(z : ℚ_[p])∥ := -by simp [norm_def] +lemma norm_int_cast_eq_padic_norm (z : ℤ) : ∥(z : ℤ_[p])∥ = ∥(z : ℚ_[p])∥ := by simp [norm_def] -@[simp] lemma norm_eq_padic_norm {q : ℚ_[p]} (hq : ∥q∥ ≤ 1) : - @norm ℤ_[p] _ ⟨q, hq⟩ = ∥q∥ := rfl +@[simp] lemma norm_eq_padic_norm {q : ℚ_[p]} (hq : ∥q∥ ≤ 1) : @norm ℤ_[p] _ ⟨q, hq⟩ = ∥q∥ := rfl @[simp] lemma norm_p : ∥(p : ℤ_[p])∥ = p⁻¹ := padic_norm_e.norm_p @[simp] lemma norm_p_pow (n : ℕ) : ∥(p : ℤ_[p])^n∥ = p^(-n:ℤ) := padic_norm_e.norm_p_pow n -private def cau_seq_to_rat_cau_seq (f : cau_seq ℤ_[p] norm) : - cau_seq ℚ_[p] (λ a, ∥a∥) := -⟨ λ n, f n, - λ _ hε, by simpa [norm, norm_def] using f.cauchy hε ⟩ +private def cau_seq_to_rat_cau_seq (f : cau_seq ℤ_[p] norm) : cau_seq ℚ_[p] (λ a, ∥a∥) := +⟨ λ n, f n, λ _ hε, by simpa [norm, norm_def] using f.cauchy hε ⟩ variables (p) @@ -257,18 +244,17 @@ instance complete : cau_seq.is_complete ℤ_[p] norm := ⟨ λ f, have hqn : ∥cau_seq.lim (cau_seq_to_rat_cau_seq f)∥ ≤ 1, from padic_norm_e_lim_le zero_lt_one (λ _, norm_le_one _), - ⟨ ⟨_, hqn⟩, - λ ε, by simpa [norm, norm_def] using cau_seq.equiv_lim (cau_seq_to_rat_cau_seq f) ε⟩⟩ + ⟨⟨_, hqn⟩, λ ε, by simpa [norm, norm_def] using cau_seq.equiv_lim (cau_seq_to_rat_cau_seq f) ε⟩⟩ end padic_int namespace padic_int variables (p : ℕ) [hp : fact p.prime] + include hp -lemma exists_pow_neg_lt {ε : ℝ} (hε : 0 < ε) : - ∃ (k : ℕ), ↑p ^ -((k : ℕ) : ℤ) < ε := +lemma exists_pow_neg_lt {ε : ℝ} (hε : 0 < ε) : ∃ k : ℕ, ↑p ^ -(k : ℤ) < ε := begin obtain ⟨k, hk⟩ := exists_nat_gt ε⁻¹, use k, @@ -282,8 +268,7 @@ begin { exact_mod_cast hp.1.pos } end -lemma exists_pow_neg_lt_rat {ε : ℚ} (hε : 0 < ε) : - ∃ (k : ℕ), ↑p ^ -((k : ℕ) : ℤ) < ε := +lemma exists_pow_neg_lt_rat {ε : ℚ} (hε : 0 < ε) : ∃ k : ℕ, ↑p ^ -(k : ℤ) < ε := begin obtain ⟨k, hk⟩ := @exists_pow_neg_lt p _ ε (by exact_mod_cast hε), use k, @@ -297,31 +282,27 @@ lemma norm_int_lt_one_iff_dvd (k : ℤ) : ∥(k : ℤ_[p])∥ < 1 ↔ (p : ℤ) suffices ∥(k : ℚ_[p])∥ < 1 ↔ ↑p ∣ k, by rwa norm_int_cast_eq_padic_norm, padic_norm_e.norm_int_lt_one_iff_dvd k -lemma norm_int_le_pow_iff_dvd {k : ℤ} {n : ℕ} : ∥(k : ℤ_[p])∥ ≤ p ^ (-n : ℤ) ↔ (p^n : ℤ) ∣ k := -suffices ∥(k : ℚ_[p])∥ ≤ ((↑p)^(-n : ℤ)) ↔ ↑(p^n) ∣ k, +lemma norm_int_le_pow_iff_dvd {k : ℤ} {n : ℕ} : ∥(k : ℤ_[p])∥ ≤ p ^ (-n : ℤ) ↔ (p ^ n : ℤ) ∣ k := +suffices ∥(k : ℚ_[p])∥ ≤ p ^ (-n : ℤ) ↔ ↑(p ^ n) ∣ k, by simpa [norm_int_cast_eq_padic_norm], padic_norm_e.norm_int_le_pow_iff_dvd _ _ /-! ### Valuation on `ℤ_[p]` -/ -/-- `padic_int.valuation` lifts the p-adic valuation on `ℚ` to `ℤ_[p]`. -/ +/-- `padic_int.valuation` lifts the `p`-adic valuation on `ℚ` to `ℤ_[p]`. -/ def valuation (x : ℤ_[p]) := padic.valuation (x : ℚ_[p]) -lemma norm_eq_pow_val {x : ℤ_[p]} (hx : x ≠ 0) : - ∥x∥ = p^(-x.valuation) := +lemma norm_eq_pow_val {x : ℤ_[p]} (hx : x ≠ 0) : ∥x∥ = (p : ℝ) ^ -x.valuation := begin convert padic.norm_eq_pow_val _, contrapose! hx, exact subtype.val_injective hx end -@[simp] lemma valuation_zero : valuation (0 : ℤ_[p]) = 0 := -padic.valuation_zero +@[simp] lemma valuation_zero : valuation (0 : ℤ_[p]) = 0 := padic.valuation_zero -@[simp] lemma valuation_one : valuation (1 : ℤ_[p]) = 0 := -padic.valuation_one +@[simp] lemma valuation_one : valuation (1 : ℤ_[p]) = 0 := padic.valuation_one -@[simp] lemma valuation_p : valuation (p : ℤ_[p]) = 1 := -by simp [valuation] +@[simp] lemma valuation_p : valuation (p : ℤ_[p]) = 1 := by simp [valuation] lemma valuation_nonneg (x : ℤ_[p]) : 0 ≤ x.valuation := begin @@ -368,8 +349,7 @@ lemma mul_inv : ∀ {z : ℤ_[p]}, ∥z∥ = 1 → z * z.inv = 1 simp [mul_inv_cancel hk] end -lemma inv_mul {z : ℤ_[p]} (hz : ∥z∥ = 1) : z.inv * z = 1 := -by rw [mul_comm, mul_inv hz] +lemma inv_mul {z : ℤ_[p]} (hz : ∥z∥ = 1) : z.inv * z = 1 := by rw [mul_comm, mul_inv hz] lemma is_unit_iff {z : ℤ_[p]} : is_unit z ↔ ∥z∥ = 1 := ⟨λ h, begin @@ -383,8 +363,8 @@ lemma norm_lt_one_add {z1 z2 : ℤ_[p]} (hz1 : ∥z1∥ < 1) (hz2 : ∥z2∥ < 1 lt_of_le_of_lt (nonarchimedean _ _) (max_lt hz1 hz2) lemma norm_lt_one_mul {z1 z2 : ℤ_[p]} (hz2 : ∥z2∥ < 1) : ∥z1 * z2∥ < 1 := -calc ∥z1 * z2∥ = ∥z1∥ * ∥z2∥ : by simp - ... < 1 : mul_lt_one_of_nonneg_of_lt_one_right (norm_le_one _) (norm_nonneg _) hz2 +calc ∥z1 * z2∥ = ∥z1∥ * ∥z2∥ : by simp + ... < 1 : mul_lt_one_of_nonneg_of_lt_one_right (norm_le_one _) (norm_nonneg _) hz2 @[simp] lemma mem_nonunits {z : ℤ_[p]} : z ∈ nonunits ℤ_[p] ↔ ∥z∥ < 1 := by rw lt_iff_le_and_ne; simp [norm_le_one z, nonunits, is_unit_iff] @@ -393,24 +373,21 @@ by rw lt_iff_le_and_ne; simp [norm_le_one z, nonunits, is_unit_iff] def mk_units {u : ℚ_[p]} (h : ∥u∥ = 1) : ℤ_[p]ˣ := let z : ℤ_[p] := ⟨u, le_of_eq h⟩ in ⟨z, z.inv, mul_inv h, inv_mul h⟩ -@[simp] -lemma mk_units_eq {u : ℚ_[p]} (h : ∥u∥ = 1) : ((mk_units h : ℤ_[p]) : ℚ_[p]) = u := -rfl +@[simp] lemma mk_units_eq {u : ℚ_[p]} (h : ∥u∥ = 1) : ((mk_units h : ℤ_[p]) : ℚ_[p]) = u := rfl -@[simp] lemma norm_units (u : ℤ_[p]ˣ) : ∥(u : ℤ_[p])∥ = 1 := -is_unit_iff.mp $ by simp +@[simp] lemma norm_units (u : ℤ_[p]ˣ) : ∥(u : ℤ_[p])∥ = 1 := is_unit_iff.mp $ by simp /-- `unit_coeff hx` is the unit `u` in the unique representation `x = u * p ^ n`. See `unit_coeff_spec`. -/ def unit_coeff {x : ℤ_[p]} (hx : x ≠ 0) : ℤ_[p]ˣ := -let u : ℚ_[p] := x*p^(-x.valuation) in +let u : ℚ_[p] := x * p ^ -x.valuation in have hu : ∥u∥ = 1, by simp [hx, nat.zpow_ne_zero_of_pos (by exact_mod_cast hp.1.pos) x.valuation, norm_eq_pow_val, zpow_neg, inv_mul_cancel], mk_units hu @[simp] lemma unit_coeff_coe {x : ℤ_[p]} (hx : x ≠ 0) : - (unit_coeff hx : ℚ_[p]) = x * p ^ (-x.valuation) := rfl + (unit_coeff hx : ℚ_[p]) = x * p ^ -x.valuation := rfl lemma unit_coeff_spec {x : ℤ_[p]} (hx : x ≠ 0) : x = (unit_coeff hx : ℤ_[p]) * p ^ int.nat_abs (valuation x) := @@ -472,14 +449,12 @@ begin rw [norm_le_pow_iff_le_valuation x hx, mem_span_pow_iff_le_valuation x hx] end -lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℤ_[p]) (n : ℤ) : - ∥x∥ ≤ p ^ n ↔ ∥x∥ < p ^ (n + 1) := +lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℤ_[p]) (n : ℤ) : ∥x∥ ≤ p ^ n ↔ ∥x∥ < p ^ (n + 1) := begin rw norm_def, exact padic.norm_le_pow_iff_norm_lt_pow_add_one _ _, end -lemma norm_lt_pow_iff_norm_le_pow_sub_one (x : ℤ_[p]) (n : ℤ) : - ∥x∥ < p ^ n ↔ ∥x∥ ≤ p ^ (n - 1) := +lemma norm_lt_pow_iff_norm_le_pow_sub_one (x : ℤ_[p]) (n : ℤ) : ∥x∥ < p ^ n ↔ ∥x∥ ≤ p ^ (n - 1) := by rw [norm_le_pow_iff_norm_lt_pow_add_one, sub_add_cancel] lemma norm_lt_one_iff_dvd (x : ℤ_[p]) : ∥x∥ < 1 ↔ ↑p ∣ x := @@ -523,16 +498,14 @@ begin { exact_mod_cast hp.1.ne_zero } end -lemma irreducible_p : irreducible (p : ℤ_[p]) := -prime.irreducible prime_p +lemma irreducible_p : irreducible (p : ℤ_[p]) := prime.irreducible prime_p instance : discrete_valuation_ring ℤ_[p] := discrete_valuation_ring.of_has_unit_mul_pow_irreducible_factorization ⟨p, irreducible_p, λ x hx, ⟨x.valuation.nat_abs, unit_coeff hx, by rw [mul_comm, ← unit_coeff_spec hx]⟩⟩ -lemma ideal_eq_span_pow_p {s : ideal ℤ_[p]} (hs : s ≠ ⊥) : - ∃ n : ℕ, s = ideal.span {p ^ n} := +lemma ideal_eq_span_pow_p {s : ideal ℤ_[p]} (hs : s ≠ ⊥) : ∃ n : ℕ, s = ideal.span {p ^ n} := discrete_valuation_ring.ideal_eq_span_pow_irreducible hs irreducible_p open cau_seq diff --git a/src/number_theory/padics/padic_norm.lean b/src/number_theory/padics/padic_norm.lean index e44df9e92c04e..0380f329ff94a 100644 --- a/src/number_theory/padics/padic_norm.lean +++ b/src/number_theory/padics/padic_norm.lean @@ -8,13 +8,13 @@ import number_theory.padics.padic_val /-! # p-adic norm -This file defines the p-adic norm on ℚ. +This file defines the `p`-adic norm on `ℚ`. -The p-adic valuation on ℚ is the difference of the multiplicities of `p` in the numerator and +The `p`-adic valuation on `ℚ` is the difference of the multiplicities of `p` in the numerator and denominator of `q`. This function obeys the standard properties of a valuation, with the appropriate -assumptions on p. +assumptions on `p`. -The valuation induces a norm on ℚ. This norm is a nonarchimedean absolute value. +The valuation induces a norm on `ℚ`. This norm is a nonarchimedean absolute value. It takes values in {0} ∪ {1/p^k | k ∈ ℤ}. ## Notations @@ -24,7 +24,7 @@ This file uses the local notation `/.` for `rat.mk`. ## Implementation notes Much, but not all, of this file assumes that `p` is prime. This assumption is inferred automatically -by taking `[fact (prime p)]` as a type class argument. +by taking `[fact p.prime]` as a type class argument. ## References @@ -37,24 +37,21 @@ by taking `[fact (prime p)]` as a type class argument. p-adic, p adic, padic, norm, valuation -/ -/-- -If `q ≠ 0`, the p-adic norm of a rational `q` is `p ^ (-(padic_val_rat p q))`. -If `q = 0`, the p-adic norm of `q` is 0. --/ -def padic_norm (p : ℕ) (q : ℚ) : ℚ := -if q = 0 then 0 else (↑p : ℚ) ^ (-(padic_val_rat p q)) +/-- If `q ≠ 0`, the `p`-adic norm of a rational `q` is `p ^ -padic_val_rat p q`. +If `q = 0`, the `p`-adic norm of `q` is `0`. -/ +def padic_norm (p : ℕ) (q : ℚ) : ℚ := if q = 0 then 0 else (p : ℚ) ^ -padic_val_rat p q namespace padic_norm open padic_val_rat variables {p : ℕ} -/-- Unfolds the definition of the p-adic norm of `q` when `q ≠ 0`. -/ +/-- Unfolds the definition of the `p`-adic norm of `q` when `q ≠ 0`. -/ @[simp] protected lemma eq_zpow_of_nonzero {q : ℚ} (hq : q ≠ 0) : - padic_norm p q = p ^ (-(padic_val_rat p q)) := + padic_norm p q = p ^ -padic_val_rat p q := by simp [hq, padic_norm] -/-- The p-adic norm is nonnegative. -/ +/-- The `p`-adic norm is nonnegative. -/ protected lemma nonneg (q : ℚ) : 0 ≤ padic_norm p q := if hq : q = 0 then by simp [hq, padic_norm] else @@ -64,29 +61,25 @@ else exact_mod_cast nat.zero_le _ end -/-- The p-adic norm of 0 is 0. -/ +/-- The `p`-adic norm of `0` is `0`. -/ @[simp] protected lemma zero : padic_norm p 0 = 0 := by simp [padic_norm] -/-- The p-adic norm of 1 is 1. -/ +/-- The `p`-adic norm of `1` is `1`. -/ @[simp] protected lemma one : padic_norm p 1 = 1 := by simp [padic_norm] -/-- -The p-adic norm of `p` is `1/p` if `p > 1`. +/-- The `p`-adic norm of `p` is `p⁻¹` if `p > 1`. -See also `padic_norm.padic_norm_p_of_prime` for a version that assumes `p` is prime. --/ -lemma padic_norm_p (hp : 1 < p) : padic_norm p p = 1 / p := +See also `padic_norm.padic_norm_p_of_prime` for a version assuming `p` is prime. -/ +lemma padic_norm_p (hp : 1 < p) : padic_norm p p = p⁻¹ := by simp [padic_norm, (pos_of_gt hp).ne', padic_val_nat.self hp] -/-- -The p-adic norm of `p` is `1/p` if `p` is prime. +/-- The `p`-adic norm of `p` is `p⁻¹` if `p` is prime. -See also `padic_norm.padic_norm_p` for a version that assumes `1 < p`. --/ -@[simp] lemma padic_norm_p_of_prime [fact p.prime] : padic_norm p p = 1 / p := +See also `padic_norm.padic_norm_p` for a version assuming `1 < p`. -/ +@[simp] lemma padic_norm_p_of_prime [fact p.prime] : padic_norm p p = p⁻¹ := padic_norm_p $ nat.prime.one_lt (fact.out _) -/-- The p-adic norm of `q` is `1` if `q` is prime and not equal to `p`. -/ +/-- The `p`-adic norm of `q` is `1` if `q` is prime and not equal to `p`. -/ lemma padic_norm_of_prime_of_ne {q : ℕ} [p_prime : fact p.prime] [q_prime : fact q.prime] (neq : p ≠ q) : padic_norm p q = 1 := begin @@ -95,28 +88,23 @@ begin simp [padic_norm, p, q_prime.1.1, q_prime.1.ne_zero], end -/-- -The p-adic norm of `p` is less than 1 if `1 < p`. +/-- The `p`-adic norm of `p` is less than `1` if `1 < p`. -See also `padic_norm.padic_norm_p_lt_one_of_prime` for a version assuming `prime p`. --/ +See also `padic_norm.padic_norm_p_lt_one_of_prime` for a version assuming `p` is prime. -/ lemma padic_norm_p_lt_one (hp : 1 < p) : padic_norm p p < 1 := begin - rw [padic_norm_p hp, div_lt_iff, one_mul], - { exact_mod_cast hp }, - { exact_mod_cast zero_lt_one.trans hp }, + rw [padic_norm_p hp, inv_lt_one_iff], + exact_mod_cast (or.inr hp) end -/-- -The p-adic norm of `p` is less than 1 if `p` is prime. +/-- The `p`-adic norm of `p` is less than `1` if `p` is prime. -See also `padic_norm.padic_norm_p_lt_one` for a version assuming `1 < p`. --/ +See also `padic_norm.padic_norm_p_lt_one` for a version assuming `1 < p`. -/ lemma padic_norm_p_lt_one_of_prime [fact p.prime] : padic_norm p p < 1 := padic_norm_p_lt_one $ nat.prime.one_lt (fact.out _) /-- `padic_norm p q` takes discrete values `p ^ -z` for `z : ℤ`. -/ -protected theorem values_discrete {q : ℚ} (hq : q ≠ 0) : ∃ z : ℤ, padic_norm p q = p ^ (-z) := +protected theorem values_discrete {q : ℚ} (hq : q ≠ 0) : ∃ z : ℤ, padic_norm p q = p ^ -z := ⟨ (padic_val_rat p q), by simp [padic_norm, hq] ⟩ /-- `padic_norm p` is symmetric. -/ @@ -135,7 +123,7 @@ begin exact_mod_cast ne_of_gt hp.1.pos end -/-- If the p-adic norm of `q` is 0, then `q` is 0. -/ +/-- If the `p`-adic norm of `q` is 0, then `q` is `0`. -/ lemma zero_of_padic_norm_eq_zero {q : ℚ} (h : padic_norm p q = 0) : q = 0 := begin apply by_contradiction, intro hq, @@ -145,24 +133,24 @@ begin exact_mod_cast hp.1.ne_zero end -/-- The p-adic norm is multiplicative. -/ -@[simp] protected theorem mul (q r : ℚ) : padic_norm p (q*r) = padic_norm p q * padic_norm p r := +/-- The `p`-adic norm is multiplicative. -/ +@[simp] protected theorem mul (q r : ℚ) : padic_norm p (q * r) = padic_norm p q * padic_norm p r := if hq : q = 0 then by simp [hq] else if hr : r = 0 then by simp [hr] else - have q*r ≠ 0, from mul_ne_zero hq hr, + have q * r ≠ 0, from mul_ne_zero hq hr, have (p : ℚ) ≠ 0, by simp [hp.1.ne_zero], by simp [padic_norm, *, padic_val_rat.mul, zpow_add₀ this, mul_comm] -/-- The p-adic norm respects division. -/ +/-- The `p`-adic norm respects division. -/ @[simp] protected theorem div (q r : ℚ) : padic_norm p (q / r) = padic_norm p q / padic_norm p r := if hr : r = 0 then by simp [hr] else eq_div_of_mul_eq (padic_norm.nonzero hr) (by rw [←padic_norm.mul, div_mul_cancel _ hr]) -/-- The p-adic norm of an integer is at most 1. -/ -protected theorem of_int (z : ℤ) : padic_norm p ↑z ≤ 1 := +/-- The `p`-adic norm of an integer is at most `1`. -/ +protected theorem of_int (z : ℤ) : padic_norm p z ≤ 1 := if hz : z = 0 then by simp [hz, zero_le_one] else begin unfold padic_norm, @@ -199,10 +187,8 @@ else apply min_le_padic_val_rat_add; assumption } end -/-- -The p-adic norm is nonarchimedean: the norm of `p + q` is at most the max of the norm of `p` and -the norm of `q`. --/ +/-- The `p`-adic norm is nonarchimedean: the norm of `p + q` is at most the max of the norm of `p` +and the norm of `q`. -/ protected theorem nonarchimedean {q r : ℚ} : padic_norm p (q + r) ≤ max (padic_norm p q) (padic_norm p r) := begin @@ -210,26 +196,20 @@ begin exact nonarchimedean_aux hle end -/-- -The p-adic norm respects the triangle inequality: the norm of `p + q` is at most the norm of `p` -plus the norm of `q`. --/ +/-- The `p`-adic norm respects the triangle inequality: the norm of `p + q` is at most the norm of +`p` plus the norm of `q`. -/ theorem triangle_ineq (q r : ℚ) : padic_norm p (q + r) ≤ padic_norm p q + padic_norm p r := calc padic_norm p (q + r) ≤ max (padic_norm p q) (padic_norm p r) : padic_norm.nonarchimedean ... ≤ padic_norm p q + padic_norm p r : max_le_add_of_nonneg (padic_norm.nonneg _) (padic_norm.nonneg _) -/-- -The p-adic norm of a difference is at most the max of each component. Restates the archimedean -property of the p-adic norm. --/ +/-- The `p`-adic norm of a difference is at most the max of each component. Restates the archimedean +property of the `p`-adic norm. -/ protected theorem sub {q r : ℚ} : padic_norm p (q - r) ≤ max (padic_norm p q) (padic_norm p r) := by rw [sub_eq_add_neg, ←padic_norm.neg r]; apply padic_norm.nonarchimedean -/-- -If the p-adic norms of `q` and `r` are different, then the norm of `q + r` is equal to the max of -the norms of `q` and `r`. --/ +/-- If the `p`-adic norms of `q` and `r` are different, then the norm of `q + r` is equal to the max +of the norms of `q` and `r`. -/ lemma add_eq_max_of_ne {q r : ℚ} (hne : padic_norm p q ≠ padic_norm p r) : padic_norm p (q + r) = max (padic_norm p q) (padic_norm p r) := begin @@ -251,10 +231,8 @@ begin { rwa max_eq_left_of_lt hlt } end -/-- -The p-adic norm is an absolute value: positive-definite and multiplicative, satisfying the triangle -inequality. --/ +/-- The `p`-adic norm is an absolute value: positive-definite and multiplicative, satisfying the +triangle inequality. -/ instance : is_absolute_value (padic_norm p) := { abv_nonneg := padic_norm.nonneg, abv_eq_zero := λ _, ⟨zero_of_padic_norm_eq_zero, λ hx, by simpa only [hx]⟩, @@ -265,7 +243,7 @@ lemma dvd_iff_norm_le {n : ℕ} {z : ℤ} : ↑(p ^ n) ∣ z ↔ padic_norm p z begin unfold padic_norm, split_ifs with hz, { norm_cast at hz, - have : 0 ≤ (p^n : ℚ), {apply pow_nonneg, exact_mod_cast le_of_lt hp.1.pos }, + have : 0 ≤ (p ^ n : ℚ), {apply pow_nonneg, exact_mod_cast le_of_lt hp.1.pos }, simp [hz, this] }, { rw [zpow_le_iff_le, neg_le_neg_iff, padic_val_rat.of_int, padic_val_int.of_ne_one_ne_zero hp.1.ne_one _], diff --git a/src/number_theory/padics/padic_numbers.lean b/src/number_theory/padics/padic_numbers.lean index c11594bed6308..139d6412590f8 100644 --- a/src/number_theory/padics/padic_numbers.lean +++ b/src/number_theory/padics/padic_numbers.lean @@ -9,43 +9,43 @@ import number_theory.padics.padic_norm /-! # p-adic numbers -This file defines the p-adic numbers (rationals) `ℚ_p` as -the completion of `ℚ` with respect to the p-adic norm. -We show that the p-adic norm on ℚ extends to `ℚ_p`, that `ℚ` is embedded in `ℚ_p`, -and that `ℚ_p` is Cauchy complete. +This file defines the `p`-adic numbers (rationals) `ℚ_[p]` as +the completion of `ℚ` with respect to the `p`-adic norm. +We show that the `p`-adic norm on `ℚ` extends to `ℚ_[p]`, that `ℚ` is embedded in `ℚ_[p]`, +and that `ℚ_[p]` is Cauchy complete. ## Important definitions -* `padic` : the type of p-adic numbers -* `padic_norm_e` : the rational valued p-adic norm on `ℚ_p` -* `padic.add_valuation` : the additive `p`-adic valuation on `ℚ_p`, with values in `with_top ℤ`. +* `padic` : the type of `p`-adic numbers +* `padic_norm_e` : the rational valued `p`-adic norm on `ℚ_[p]` +* `padic.add_valuation` : the additive `p`-adic valuation on `ℚ_[p]`, with values in `with_top ℤ` ## Notation -We introduce the notation `ℚ_[p]` for the p-adic numbers. +We introduce the notation `ℚ_[p]` for the `p`-adic numbers. ## Implementation notes Much, but not all, of this file assumes that `p` is prime. This assumption is inferred automatically -by taking `[fact (prime p)]` as a type class argument. +by taking `[fact p.prime]` as a type class argument. -We use the same concrete Cauchy sequence construction that is used to construct ℝ. -`ℚ_p` inherits a field structure from this construction. -The extension of the norm on ℚ to `ℚ_p` is *not* analogous to extending the absolute value to ℝ, -and hence the proof that `ℚ_p` is complete is different from the proof that ℝ is complete. +We use the same concrete Cauchy sequence construction that is used to construct `ℝ`. +`ℚ_[p]` inherits a field structure from this construction. +The extension of the norm on `ℚ` to `ℚ_[p]` is *not* analogous to extending the absolute value to +`ℝ` and hence the proof that `ℚ_[p]` is complete is different from the proof that ℝ is complete. A small special-purpose simplification tactic, `padic_index_simp`, is used to manipulate sequence indices in the proof that the norm extends. -`padic_norm_e` is the rational-valued p-adic norm on `ℚ_p`. -To instantiate `ℚ_p` as a normed field, we must cast this into a ℝ-valued norm. +`padic_norm_e` is the rational-valued `p`-adic norm on `ℚ_[p]`. +To instantiate `ℚ_[p]` as a normed field, we must cast this into a `ℝ`-valued norm. The `ℝ`-valued norm, using notation `∥ ∥` from normed spaces, is the canonical representation of this norm. `simp` prefers `padic_norm` to `padic_norm_e` when possible. Since `padic_norm_e` and `∥ ∥` have different types, `simp` does not rewrite one to the other. -Coercions from `ℚ` to `ℚ_p` are set up to work with the `norm_cast` tactic. +Coercions from `ℚ` to `ℚ_[p]` are set up to work with the `norm_cast` tactic. ## References @@ -63,7 +63,7 @@ open_locale classical open nat multiplicity padic_norm cau_seq cau_seq.completion metric -/-- The type of Cauchy sequences of rationals with respect to the p-adic norm. -/ +/-- The type of Cauchy sequences of rationals with respect to the `p`-adic norm. -/ @[reducible] def padic_seq (p : ℕ) := cau_seq _ (padic_norm p) namespace padic_seq @@ -71,7 +71,7 @@ namespace padic_seq section variables {p : ℕ} [fact p.prime] -/-- The p-adic norm of the entries of a nonzero Cauchy sequence of rationals is eventually +/-- The `p`-adic norm of the entries of a nonzero Cauchy sequence of rationals is eventually constant. -/ lemma stationary {f : cau_seq ℚ (padic_norm p)} (hf : ¬ f ≈ 0) : ∃ N, ∀ m n, N ≤ m → N ≤ n → padic_norm p (f n) = padic_norm p (f m) := @@ -95,9 +95,8 @@ let ⟨ε, hε, N1, hN1⟩ := this, apply _root_.lt_irrefl _ this end ⟩ -/-- For all n ≥ stationary_point f hf, the p-adic norm of f n is the same. -/ -def stationary_point {f : padic_seq p} (hf : ¬ f ≈ 0) : ℕ := -classical.some $ stationary hf +/-- For all `n ≥ stationary_point f hf`, the `p`-adic norm of `f n` is the same. -/ +def stationary_point {f : padic_seq p} (hf : ¬ f ≈ 0) : ℕ := classical.some $ stationary hf lemma stationary_point_spec {f : padic_seq p} (hf : ¬ f ≈ 0) : ∀ {m n}, stationary_point hf ≤ m → stationary_point hf ≤ n → @@ -106,8 +105,7 @@ classical.some_spec $ stationary hf /-- Since the norm of the entries of a Cauchy sequence is eventually stationary, we can lift the norm to sequences. -/ -def norm (f : padic_seq p) : ℚ := -if hf : f ≈ 0 then 0 else padic_norm p (f (stationary_point hf)) +def norm (f : padic_seq p) : ℚ := if hf : f ≈ 0 then 0 else padic_norm p (f (stationary_point hf)) lemma norm_zero_iff (f : padic_seq p) : f.norm = 0 ↔ f ≈ 0 := begin @@ -133,28 +131,24 @@ variables {p : ℕ} [fact p.prime] lemma equiv_zero_of_val_eq_of_equiv_zero {f g : padic_seq p} (h : ∀ k, padic_norm p (f k) = padic_norm p (g k)) (hf : f ≈ 0) : g ≈ 0 := -λ ε hε, let ⟨i, hi⟩ := hf _ hε in -⟨i, λ j hj, by simpa [h] using hi _ hj⟩ +λ ε hε, let ⟨i, hi⟩ := hf _ hε in ⟨i, λ j hj, by simpa [h] using hi _ hj⟩ -lemma norm_nonzero_of_not_equiv_zero {f : padic_seq p} (hf : ¬ f ≈ 0) : - f.norm ≠ 0 := +lemma norm_nonzero_of_not_equiv_zero {f : padic_seq p} (hf : ¬ f ≈ 0) : f.norm ≠ 0 := hf ∘ f.norm_zero_iff.1 lemma norm_eq_norm_app_of_nonzero {f : padic_seq p} (hf : ¬ f ≈ 0) : ∃ k, f.norm = padic_norm p k ∧ k ≠ 0 := have heq : f.norm = padic_norm p (f $ stationary_point hf), by simp [norm, hf], -⟨f $ stationary_point hf, heq, - λ h, norm_nonzero_of_not_equiv_zero hf (by simpa [h] using heq)⟩ +⟨f $ stationary_point hf, heq, λ h, norm_nonzero_of_not_equiv_zero hf (by simpa [h] using heq)⟩ lemma not_lim_zero_const_of_nonzero {q : ℚ} (hq : q ≠ 0) : ¬ lim_zero (const (padic_norm p) q) := λ h', hq $ const_lim_zero.1 h' -lemma not_equiv_zero_const_of_nonzero {q : ℚ} (hq : q ≠ 0) : ¬ (const (padic_norm p) q) ≈ 0 := +lemma not_equiv_zero_const_of_nonzero {q : ℚ} (hq : q ≠ 0) : ¬ const (padic_norm p) q ≈ 0 := λ h : lim_zero (const (padic_norm p) q - 0), not_lim_zero_const_of_nonzero hq $ by simpa using h lemma norm_nonneg (f : padic_seq p) : 0 ≤ f.norm := -if hf : f ≈ 0 then by simp [hf, norm] -else by simp [norm, hf, padic_norm.nonneg] +if hf : f ≈ 0 then by simp [hf, norm] else by simp [norm, hf, padic_norm.nonneg] /-- An auxiliary lemma for manipulating sequence indices. -/ lemma lift_index_left_left {f : padic_seq p} (hf : ¬ f ≈ 0) (v2 v3 : ℕ) : @@ -198,15 +192,12 @@ variables {p : ℕ} [fact p.prime] /-! ### Valuation on `padic_seq` -/ -/-- -The `p`-adic valuation on `ℚ` lifts to `padic_seq p`. -`valuation f` is defined to be the valuation of the (`ℚ`-valued) stationary point of `f`. --/ +/-- The `p`-adic valuation on `ℚ` lifts to `padic_seq p`. +`valuation f` is defined to be the valuation of the (`ℚ`-valued) stationary point of `f`. -/ def valuation (f : padic_seq p) : ℤ := if hf : f ≈ 0 then 0 else padic_val_rat p (f (stationary_point hf)) -lemma norm_eq_pow_val {f : padic_seq p} (hf : ¬ f ≈ 0) : - f.norm = p^(-f.valuation : ℤ) := +lemma norm_eq_pow_val {f : padic_seq p} (hf : ¬ f ≈ 0) : f.norm = p ^ (-f.valuation : ℤ) := begin rw [norm, valuation, dif_neg hf, dif_neg hf, padic_norm, if_neg], intro H, @@ -244,10 +235,8 @@ do [v1, v2, v3] ← [hh, hf, hg].mmap when at_.include_goal (tactic.simp_target sl >> tactic.skip), hs ← at_.get_locals, hs.mmap' (tactic.simp_hyp sl []) -/-- - This is a special-purpose tactic that lifts padic_norm (f (stationary_point f)) to - padic_norm (f (max _ _ _)). --/ +/-- This is a special-purpose tactic that lifts `padic_norm (f (stationary_point f))` to +`padic_norm (f (max _ _ _))`. -/ meta def tactic.interactive.padic_index_simp (l : interactive.parse interactive.types.pexpr_list) (at_ : interactive.parse interactive.types.location) : tactic unit := do [h, f, g] ← l.mmap tactic.i_to_expr, @@ -277,8 +266,7 @@ else apply padic_norm.mul end -lemma eq_zero_iff_equiv_zero (f : padic_seq p) : mk f = 0 ↔ f ≈ 0 := -mk_eq +lemma eq_zero_iff_equiv_zero (f : padic_seq p) : mk f = 0 ↔ f ≈ 0 := mk_eq lemma ne_zero_iff_nequiv_zero (f : padic_seq p) : mk f ≠ 0 ↔ ¬ f ≈ 0 := not_iff_not.2 (eq_zero_iff_equiv_zero _) @@ -292,8 +280,7 @@ else have ¬ (const (padic_norm p) q) ≈ 0, from not_equiv_zero_const_of_nonzero hq, by simp [norm, this] -lemma norm_values_discrete (a : padic_seq p) (ha : ¬ a ≈ 0) : - (∃ (z : ℤ), a.norm = ↑p ^ (-z)) := +lemma norm_values_discrete (a : padic_seq p) (ha : ¬ a ≈ 0) : ∃ z : ℤ, a.norm = p ^ -z := let ⟨k, hk, hk'⟩ := norm_eq_norm_app_of_nonzero ha in by simpa [hk] using padic_norm.values_discrete hk' @@ -303,8 +290,7 @@ by simp [h1, norm, hp.1.one_lt] private lemma norm_eq_of_equiv_aux {f g : padic_seq p} (hf : ¬ f ≈ 0) (hg : ¬ g ≈ 0) (hfg : f ≈ g) (h : padic_norm p (f (stationary_point hf)) ≠ padic_norm p (g (stationary_point hg))) - (hlt : padic_norm p (g (stationary_point hg)) < padic_norm p (f (stationary_point hf))) : - false := + (hlt : padic_norm p (g (stationary_point hg)) < padic_norm p (f (stationary_point hf))) : false := begin have hpn : 0 < padic_norm p (f (stationary_point hf)) - padic_norm p (g (stationary_point hg)), from sub_pos_of_lt hlt, @@ -346,16 +332,16 @@ else have hg : ¬ g ≈ 0, from hf ∘ setoid.trans hfg, by unfold norm; split_ifs; exact norm_eq_of_equiv hf hg hfg private lemma norm_nonarchimedean_aux {f g : padic_seq p} - (hfg : ¬ f + g ≈ 0) (hf : ¬ f ≈ 0) (hg : ¬ g ≈ 0) : (f + g).norm ≤ max (f.norm) (g.norm) := + (hfg : ¬ f + g ≈ 0) (hf : ¬ f ≈ 0) (hg : ¬ g ≈ 0) : (f + g).norm ≤ max f.norm g.norm := begin unfold norm, split_ifs, padic_index_simp [hfg, hf, hg], apply padic_norm.nonarchimedean end -theorem norm_nonarchimedean (f g : padic_seq p) : (f + g).norm ≤ max (f.norm) (g.norm) := +theorem norm_nonarchimedean (f g : padic_seq p) : (f + g).norm ≤ max f.norm g.norm := if hfg : f + g ≈ 0 then - have 0 ≤ max (f.norm) (g.norm), from le_max_of_le_left (norm_nonneg _), + have 0 ≤ max f.norm g.norm, from le_max_of_le_left (norm_nonneg _), by simpa only [hfg, norm, ne.def, le_max_iff, cau_seq.add_apply, not_true, dif_pos] else if hf : f ≈ 0 then have hfg' : f + g ≈ g, @@ -363,7 +349,7 @@ else if hf : f ≈ 0 then show lim_zero (f + g - g), by simpa only [sub_zero, add_sub_cancel] using hf }, have hcfg : (f + g).norm = g.norm, from norm_equiv hfg', have hcl : f.norm = 0, from (norm_zero_iff f).2 hf, - have max (f.norm) (g.norm) = g.norm, + have max f.norm g.norm = g.norm, by rw hcl; exact max_eq_right (norm_nonneg _), by rw [this, hcfg] else if hg : g ≈ 0 then @@ -372,7 +358,7 @@ else if hg : g ≈ 0 then show lim_zero (f + g - f), by simpa only [add_sub_cancel', sub_zero] using hg }, have hcfg : (f + g).norm = f.norm, from norm_equiv hfg', have hcl : g.norm = 0, from (norm_zero_iff g).2 hg, - have max (f.norm) (g.norm) = f.norm, + have max f.norm g.norm = f.norm, by rw hcl; exact max_eq_left (norm_nonneg _), by rw [this, hcfg] else norm_nonarchimedean_aux hfg hf hg @@ -395,8 +381,7 @@ else rw [hpf, hpg, h] end -lemma norm_neg (a : padic_seq p) : (-a).norm = a.norm := -norm_eq $ by simp +lemma norm_neg (a : padic_seq p) : (-a).norm = a.norm := norm_eq $ by simp lemma norm_eq_of_add_equiv_zero {f g : padic_seq p} (h : f + g ≈ 0) : f.norm = g.norm := have lim_zero (f + g - 0), from h, @@ -410,13 +395,13 @@ have hfg : ¬f + g ≈ 0, from mt norm_eq_of_add_equiv_zero hfgne, if hf : f ≈ 0 then have lim_zero (f - 0), from hf, have f + g ≈ g, from show lim_zero ((f + g) - g), by simpa only [sub_zero, add_sub_cancel], - have h1 : (f+g).norm = g.norm, from norm_equiv this, + have h1 : (f + g).norm = g.norm, from norm_equiv this, have h2 : f.norm = 0, from (norm_zero_iff _).2 hf, by rw [h1, h2]; rw max_eq_right (norm_nonneg _) else if hg : g ≈ 0 then have lim_zero (g - 0), from hg, have f + g ≈ f, from show lim_zero ((f + g) - f), by rw [add_sub_cancel']; simpa only [sub_zero], - have h1 : (f+g).norm = f.norm, from norm_equiv this, + have h1 : (f + g).norm = f.norm, from norm_equiv this, have h2 : g.norm = 0, from (norm_zero_iff _).2 hg, by rw [h1, h2]; rw max_eq_left (norm_nonneg _) else @@ -429,7 +414,8 @@ end end embedding end padic_seq -/-- The p-adic numbers `Q_[p]` are the Cauchy completion of `ℚ` with respect to the p-adic norm. -/ +/-- The `p`-adic numbers `ℚ_[p]` are the Cauchy completion of `ℚ` with respect to the `p`-adic norm. +-/ def padic (p : ℕ) [fact p.prime] := @cau_seq.completion.Cauchy _ _ _ _ (padic_norm p) _ notation `ℚ_[` p `]` := padic p @@ -468,8 +454,7 @@ lemma const_equiv {q r : ℚ} : const (padic_norm p) q ≈ const (padic_norm p) @[norm_cast] lemma coe_inj {q r : ℚ} : (↑q : ℚ_[p]) = ↑r ↔ q = r := ⟨(const_equiv p).1 ∘ quotient.eq.1, λ h, by rw h⟩ -instance : char_zero ℚ_[p] := -⟨λ m n, by { rw ← rat.cast_coe_nat, norm_cast, exact id }⟩ +instance : char_zero ℚ_[p] := ⟨λ m n, by { rw ← rat.cast_coe_nat, norm_cast, exact id }⟩ @[norm_cast] lemma coe_add : ∀ {x y : ℚ}, (↑(x + y) : ℚ_[p]) = ↑x + ↑y := rat.cast_add @[norm_cast] lemma coe_neg : ∀ {x : ℚ}, (↑(-x) : ℚ_[p]) = -↑x := rat.cast_neg @@ -482,7 +467,7 @@ instance : char_zero ℚ_[p] := end completion end padic -/-- The rational-valued p-adic norm on `ℚ_p` is lifted from the norm on Cauchy sequences. The +/-- The rational-valued `p`-adic norm on `ℚ_[p]` is lifted from the norm on Cauchy sequences. The canonical form of this function is the normed space instance, with notation `∥ ∥`. -/ def padic_norm_e {p : ℕ} [hp : fact p.prime] : ℚ_[p] → ℚ := quotient.lift padic_seq.norm $ @padic_seq.norm_equiv _ _ @@ -509,22 +494,18 @@ begin exact hN _ le_rfl _ hi } end -protected lemma nonneg (q : ℚ_[p]) : 0 ≤ padic_norm_e q := -quotient.induction_on q $ norm_nonneg +protected lemma nonneg (q : ℚ_[p]) : 0 ≤ padic_norm_e q := quotient.induction_on q norm_nonneg lemma zero_def : (0 : ℚ_[p]) = ⟦0⟧ := rfl lemma zero_iff (q : ℚ_[p]) : padic_norm_e q = 0 ↔ q = 0 := -quotient.induction_on q $ - by simpa only [zero_def, quotient.eq] using norm_zero_iff +quotient.induction_on q $ by simpa only [zero_def, quotient.eq] using norm_zero_iff -@[simp] protected lemma zero : padic_norm_e (0 : ℚ_[p]) = 0 := -(zero_iff _).2 rfl +@[simp] protected lemma zero : padic_norm_e (0 : ℚ_[p]) = 0 := (zero_iff _).2 rfl /-- Theorems about `padic_norm_e` are named with a `'` so the names do not conflict with the equivalent theorems about `norm` (`∥ ∥`). -/ -@[simp] protected lemma one' : padic_norm_e (1 : ℚ_[p]) = 1 := -norm_one +@[simp] protected lemma one' : padic_norm_e (1 : ℚ_[p]) = 1 := norm_one @[simp] protected lemma neg (q : ℚ_[p]) : padic_norm_e (-q) = padic_norm_e q := quotient.induction_on q $ norm_neg @@ -566,13 +547,13 @@ instance : is_absolute_value (@padic_norm_e p _) := @[simp] lemma eq_padic_norm' (q : ℚ) : padic_norm_e (q : ℚ_[p]) = padic_norm p q := norm_const _ -protected theorem image' {q : ℚ_[p]} : q ≠ 0 → ∃ n : ℤ, padic_norm_e q = p ^ (-n) := +protected theorem image' {q : ℚ_[p]} : q ≠ 0 → ∃ n : ℤ, padic_norm_e q = p ^ -n := quotient.induction_on q $ λ f hf, have ¬ f ≈ 0, from (ne_zero_iff_nequiv_zero f).1 hf, norm_values_discrete f this lemma sub_rev (q r : ℚ_[p]) : padic_norm_e (q - r) = padic_norm_e (r - q) := -by rw ←(padic_norm_e.neg); simp +by rw ← padic_norm_e.neg; simp end embedding end padic_norm_e @@ -582,7 +563,9 @@ namespace padic section complete open padic_seq padic -theorem rat_dense' {p : ℕ} [fact p.prime] (q : ℚ_[p]) {ε : ℚ} (hε : 0 < ε) : +variables {p : ℕ} [fact p.prime] (f : cau_seq _ (@padic_norm_e p _)) + +theorem rat_dense' (q : ℚ_[p]) {ε : ℚ} (hε : 0 < ε) : ∃ r : ℚ, padic_norm_e (q - r) < ε := quotient.induction_on q $ λ q', have ∃ N, ∀ m n ≥ N, padic_norm p (q' m - q' n) < ε, from cauchy₂ _ hε, @@ -602,20 +585,19 @@ quotient.induction_on q $ λ q', { exact hN _ (lt_of_not_ge hle).le _ le_rfl } } end⟩ -variables {p : ℕ} [fact p.prime] (f : cau_seq _ (@padic_norm_e p _)) open classical -private lemma div_nat_pos (n : ℕ) : 0 < (1 / ((n + 1): ℚ)) := +private lemma div_nat_pos (n : ℕ) : 0 < 1 / (n + 1 : ℚ) := div_pos zero_lt_one (by exact_mod_cast succ_pos _) -/-- `lim_seq f`, for `f` a Cauchy sequence of `p`-adic numbers, -is a sequence of rationals with the same limit point as `f`. -/ +/-- `lim_seq f`, for `f` a Cauchy sequence of `p`-adic numbers, is a sequence of rationals with the +same limit point as `f`. -/ def lim_seq : ℕ → ℚ := λ n, classical.some (rat_dense' (f n) (div_nat_pos n)) lemma exi_rat_seq_conv {ε : ℚ} (hε : 0 < ε) : ∃ N, ∀ i ≥ N, padic_norm_e (f i - (lim_seq f i : ℚ_[p])) < ε := begin - refine (exists_nat_gt (1/ε)).imp (λ N hN i hi, _), + refine (exists_nat_gt (1 / ε)).imp (λ N hN i hi, _), have h := classical.some_spec (rat_dense' (f i) (div_nat_pos i)), refine lt_of_lt_of_le h ((div_le_iff' $ by exact_mod_cast succ_pos _).mpr _), rw right_distrib, @@ -644,7 +626,7 @@ begin { field_simp [this], simp only [bit0, bit1, mul_add, mul_one] }, rw this, apply add_lt_add, - { suffices : padic_norm_e ((↑(lim_seq f j) - f j) + (f j - f (max N N2))) < ε / 3 + ε / 3, + { suffices : padic_norm_e ((lim_seq f j - f j) + (f j - f (max N N2))) < ε / 3 + ε / 3, by simpa only [sub_add_sub_cancel], apply lt_of_le_of_lt, { apply padic_norm_e.add }, @@ -673,7 +655,7 @@ theorem complete' : ∃ q : ℚ_[p], ∀ ε > 0, ∃ N, ∀ i ≥ N, padic_norm_ { ring_nf at this; exact this }, { apply lt_of_le_of_lt, { apply padic_norm_e.add }, - { have : ε = ε / 2 + ε / 2, by rw ←(add_self_div_two ε); simp, + { have : ε = ε / 2 + ε / 2, by rw ← add_self_div_two ε; simp, rw this, apply add_lt_add, { apply hN2, exact le_of_max_le_right hi }, @@ -719,8 +701,7 @@ instance is_absolute_value : is_absolute_value (λ a : ℚ_[p], ∥a∥) := abv_add := norm_add_le, abv_mul := by simp [has_norm.norm, padic_norm_e.mul'] } -theorem rat_dense {p : ℕ} {hp : fact p.prime} (q : ℚ_[p]) {ε : ℝ} (hε : 0 < ε) : - ∃ r : ℚ, ∥q - r∥ < ε := +theorem rat_dense (q : ℚ_[p]) {ε : ℝ} (hε : 0 < ε) : ∃ r : ℚ, ∥q - r∥ < ε := let ⟨ε', hε'l, hε'r⟩ := exists_rat_btwn hε, ⟨r, hr⟩ := rat_dense' q (by simpa using hε'l) in ⟨r, lt_trans (by simpa [has_norm.norm] using hr) hε'r⟩ @@ -744,7 +725,7 @@ begin exact_mod_cast nonarchimedean' _ _ end -theorem add_eq_max_of_ne {q r : ℚ_[p]} (h : ∥q∥ ≠ ∥r∥) : ∥q+r∥ = max (∥q∥) (∥r∥) := +theorem add_eq_max_of_ne {q r : ℚ_[p]} (h : ∥q∥ ≠ ∥r∥) : ∥q + r∥ = max (∥q∥) (∥r∥) := begin unfold has_norm.norm, apply_mod_cast add_eq_max_of_ne', @@ -754,7 +735,7 @@ begin exact_mod_cast h' end -@[simp] lemma eq_padic_norm (q : ℚ) : ∥(↑q : ℚ_[p])∥ = padic_norm p q := +@[simp] lemma eq_padic_norm (q : ℚ) : ∥(q : ℚ_[p])∥ = padic_norm p q := begin unfold has_norm.norm, rw [← padic_norm_e.eq_padic_norm'] @@ -776,7 +757,7 @@ begin exact_mod_cast hp.1.one_lt end -@[simp] lemma norm_p_pow (n : ℤ) : ∥(p^n : ℚ_[p])∥ = p^-n := +@[simp] lemma norm_p_pow (n : ℤ) : ∥(p ^ n : ℚ_[p])∥ = p ^ -n := by rw [norm_zpow, norm_p]; field_simp instance : nontrivially_normed_field ℚ_[p] := @@ -786,15 +767,14 @@ instance : nontrivially_normed_field ℚ_[p] := end⟩, .. padic.normed_field p } -protected theorem image {q : ℚ_[p]} : q ≠ 0 → ∃ n : ℤ, ∥q∥ = ↑((↑p : ℚ) ^ (-n)) := +protected theorem image {q : ℚ_[p]} : q ≠ 0 → ∃ n : ℤ, ∥q∥ = ↑((p : ℚ) ^ -n) := quotient.induction_on q $ λ f hf, have ¬ f ≈ 0, from (padic_seq.ne_zero_iff_nequiv_zero f).1 hf, let ⟨n, hn⟩ := padic_seq.norm_values_discrete f this in ⟨n, congr_arg coe hn⟩ -protected lemma is_rat (q : ℚ_[p]) : ∃ q' : ℚ, ∥q∥ = ↑q' := -if h : q = 0 then ⟨0, by simp [h]⟩ -else let ⟨n, hn⟩ := padic_norm_e.image h in ⟨_, hn⟩ +protected lemma is_rat (q : ℚ_[p]) : ∃ q' : ℚ, ∥q∥ = q' := +if h : q = 0 then ⟨0, by simp [h]⟩ else let ⟨n, hn⟩ := padic_norm_e.image h in ⟨_, hn⟩ /--`rat_norm q`, for a `p`-adic number `q` is the `p`-adic norm of `q`, as rational number. @@ -826,8 +806,7 @@ theorem norm_rat_le_one : ∀ {q : ℚ} (hq : ¬ p ∣ q.denom), ∥(q : ℚ_[p] end theorem norm_int_le_one (z : ℤ) : ∥(z : ℚ_[p])∥ ≤ 1 := -suffices ∥((z : ℚ) : ℚ_[p])∥ ≤ 1, by simpa, -norm_rat_le_one $ by simp [hp.1.ne_one] +suffices ∥((z : ℚ) : ℚ_[p])∥ ≤ 1, by simpa, norm_rat_le_one $ by simp [hp.1.ne_one] lemma norm_int_lt_one_iff_dvd (k : ℤ) : ∥(k : ℚ_[p])∥ < 1 ↔ ↑p ∣ k := begin @@ -862,21 +841,19 @@ begin exact_mod_cast hp.1.one_lt } } end -lemma norm_int_le_pow_iff_dvd (k : ℤ) (n : ℕ) : ∥(k : ℚ_[p])∥ ≤ (↑p)^(-n : ℤ) ↔ ↑(p^n) ∣ k := +lemma norm_int_le_pow_iff_dvd (k : ℤ) (n : ℕ) : ∥(k : ℚ_[p])∥ ≤ ↑p ^ (-n : ℤ) ↔ ↑(p ^ n) ∣ k := begin - have : (p : ℝ) ^ (-n : ℤ) = ↑((p ^ (-n : ℤ) : ℚ)), {simp}, + have : (p : ℝ) ^ (-n : ℤ) = ↑(p ^ (-n : ℤ) : ℚ), {simp}, rw [show (k : ℚ_[p]) = ((k : ℚ) : ℚ_[p]), by norm_cast, eq_padic_norm, this], norm_cast, rw ← padic_norm.dvd_iff_norm_le end -lemma eq_of_norm_add_lt_right {p : ℕ} {hp : fact p.prime} {z1 z2 : ℚ_[p]} - (h : ∥z1 + z2∥ < ∥z2∥) : ∥z1∥ = ∥z2∥ := +lemma eq_of_norm_add_lt_right {z1 z2 : ℚ_[p]} (h : ∥z1 + z2∥ < ∥z2∥) : ∥z1∥ = ∥z2∥ := by_contradiction $ λ hne, not_lt_of_ge (by rw padic_norm_e.add_eq_max_of_ne hne; apply le_max_right) h -lemma eq_of_norm_add_lt_left {p : ℕ} {hp : fact p.prime} {z1 z2 : ℚ_[p]} - (h : ∥z1 + z2∥ < ∥z1∥) : ∥z1∥ = ∥z2∥ := +lemma eq_of_norm_add_lt_left {z1 z2 : ℚ_[p]} (h : ∥z1 + z2∥ < ∥z1∥) : ∥z1∥ = ∥z2∥ := by_contradiction $ λ hne, not_lt_of_ge (by rw padic_norm_e.add_eq_max_of_ne hne; apply le_max_left) h @@ -909,8 +886,8 @@ begin exact_mod_cast hN i hi end -lemma padic_norm_e_lim_le {f : cau_seq ℚ_[p] norm} {a : ℝ} (ha : 0 < a) - (hf : ∀ i, ∥f i∥ ≤ a) : ∥f.lim∥ ≤ a := +lemma padic_norm_e_lim_le {f : cau_seq ℚ_[p] norm} {a : ℝ} (ha : 0 < a) (hf : ∀ i, ∥f i∥ ≤ a) : + ∥f.lim∥ ≤ a := let ⟨N, hN⟩ := setoid.symm (cau_seq.equiv_lim f) _ ha in calc ∥f.lim∥ = ∥f.lim - f N + f N∥ : by simp ... ≤ max (∥f.lim - f N∥) (∥f N∥) : padic_norm_e.nonarchimedean _ _ @@ -932,11 +909,8 @@ end /-! ### Valuation on `ℚ_[p]` -/ -/-- -`padic.valuation` lifts the p-adic valuation on rationals to `ℚ_[p]`. --/ -def valuation : ℚ_[p] → ℤ := -quotient.lift (@padic_seq.valuation p _) (λ f g h, +/-- `padic.valuation` lifts the `p`-adic valuation on rationals to `ℚ_[p]`. -/ +def valuation : ℚ_[p] → ℤ := quotient.lift (@padic_seq.valuation p _) (λ f g h, begin by_cases hf : f ≈ 0, { have hg : g ≈ 0, from setoid.trans (setoid.symm h) hf, @@ -946,8 +920,7 @@ begin exact padic_seq.norm_equiv h } end) -@[simp] lemma valuation_zero : valuation (0 : ℚ_[p]) = 0 := -dif_pos ((const_equiv p).2 rfl) +@[simp] lemma valuation_zero : valuation (0 : ℚ_[p]) = 0 := dif_pos ((const_equiv p).2 rfl) @[simp] lemma valuation_one : valuation (1 : ℚ_[p]) = 0 := begin @@ -958,7 +931,7 @@ begin simp end -lemma norm_eq_pow_val {x : ℚ_[p]} : x ≠ 0 → ∥x∥ = p^(-x.valuation) := +lemma norm_eq_pow_val {x : ℚ_[p]} : x ≠ 0 → ∥x∥ = p ^ -x.valuation := begin apply quotient.induction_on' x, clear x, intros f hf, @@ -1012,16 +985,14 @@ begin exact h_norm end -/-- The additive p-adic valuation on `ℚ_p`, with values in `with_top ℤ`. -/ -def add_valuation_def : ℚ_[p] → (with_top ℤ) := -λ x, if x = 0 then ⊤ else x.valuation +/-- The additive `p`-adic valuation on `ℚ_[p]`, with values in `with_top ℤ`. -/ +def add_valuation_def : ℚ_[p] → with_top ℤ := λ x, if x = 0 then ⊤ else x.valuation @[simp] lemma add_valuation.map_zero : add_valuation_def (0 : ℚ_[p]) = ⊤ := by simp only [add_valuation_def, if_pos (eq.refl _)] @[simp] lemma add_valuation.map_one : add_valuation_def (1 : ℚ_[p]) = 0 := -by simp only [add_valuation_def, if_neg one_ne_zero, valuation_one, - with_top.coe_zero] +by simp only [add_valuation_def, if_neg one_ne_zero, valuation_one, with_top.coe_zero] lemma add_valuation.map_mul (x y : ℚ_[p]) : add_valuation_def (x * y) = add_valuation_def x + add_valuation_def y := @@ -1050,21 +1021,19 @@ begin exact valuation_map_add hxy }}} end -/-- The additive `p`-adic valuation on `ℚ_p`, as an `add_valuation`. -/ +/-- The additive `p`-adic valuation on `ℚ_[p]`, as an `add_valuation`. -/ def add_valuation : add_valuation ℚ_[p] (with_top ℤ) := add_valuation.of add_valuation_def add_valuation.map_zero add_valuation.map_one add_valuation.map_add add_valuation.map_mul -@[simp] lemma add_valuation.apply {x : ℚ_[p]} (hx : x ≠ 0) : - x.add_valuation = x.valuation := +@[simp] lemma add_valuation.apply {x : ℚ_[p]} (hx : x ≠ 0) : x.add_valuation = x.valuation := by simp only [add_valuation, add_valuation.of_apply, add_valuation_def, if_neg hx] section norm_le_iff /-! ### Various characterizations of open unit balls -/ -lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℚ_[p]) (n : ℤ) : - ∥x∥ ≤ p ^ n ↔ ∥x∥ < p ^ (n + 1) := +lemma norm_le_pow_iff_norm_lt_pow_add_one (x : ℚ_[p]) (n : ℤ) : ∥x∥ ≤ p ^ n ↔ ∥x∥ < p ^ (n + 1) := begin have aux : ∀ n : ℤ, 0 < (p ^ n : ℝ), { apply nat.zpow_pos_of_pos, exact hp.1.pos }, @@ -1075,8 +1044,7 @@ begin rw [H.le_iff_le, H.lt_iff_lt, int.lt_add_one_iff] end -lemma norm_lt_pow_iff_norm_le_pow_sub_one (x : ℚ_[p]) (n : ℤ) : - ∥x∥ < p ^ n ↔ ∥x∥ ≤ p ^ (n - 1) := +lemma norm_lt_pow_iff_norm_le_pow_sub_one (x : ℚ_[p]) (n : ℤ) : ∥x∥ < p ^ n ↔ ∥x∥ ≤ p ^ (n - 1) := by rw [norm_le_pow_iff_norm_lt_pow_add_one, sub_add_cancel] end norm_le_iff diff --git a/src/number_theory/padics/padic_val.lean b/src/number_theory/padics/padic_val.lean index 55bd9cb0ea5ff..6da916c78d05e 100644 --- a/src/number_theory/padics/padic_val.lean +++ b/src/number_theory/padics/padic_val.lean @@ -11,13 +11,13 @@ import tactic.ring_exp /-! # p-adic Valuation -This file defines the p-adic valuation on ℕ, ℤ, and ℚ. +This file defines the `p`-adic valuation on `ℕ`, `ℤ`, and `ℚ`. -The p-adic valuation on ℚ is the difference of the multiplicities of `p` in the numerator and +The `p`-adic valuation on `ℚ` is the difference of the multiplicities of `p` in the numerator and denominator of `q`. This function obeys the standard properties of a valuation, with the appropriate -assumptions on p. The p-adic valuations on ℕ and ℤ agree with that on ℚ. +assumptions on `p`. The `p`-adic valuations on `ℕ` and `ℤ` agree with that on `ℚ`. -The valuation induces a norm on ℚ. This norm is defined in padic_norm.lean. +The valuation induces a norm on `ℚ`. This norm is defined in padic_norm.lean. ## Notations @@ -26,7 +26,7 @@ This file uses the local notation `/.` for `rat.mk`. ## Implementation notes Much, but not all, of this file assumes that `p` is prime. This assumption is inferred automatically -by taking `[fact (prime p)]` as a type class argument. +by taking `[fact p.prime]` as a type class argument. ## References @@ -47,34 +47,29 @@ open_locale rat open multiplicity -/-- -For `p ≠ 1`, the p-adic valuation of a natural `n ≠ 0` is the largest natural number `k` such that -p^k divides z. -If `n = 0` or `p = 1`, then `padic_val_nat p q` defaults to 0. --/ +/-- For `p ≠ 1`, the `p`-adic valuation of a natural `n ≠ 0` is the largest natural number `k` such +that `p^k` divides `z`. If `n = 0` or `p = 1`, then `padic_val_nat p q` defaults to `0`. -/ def padic_val_nat (p : ℕ) (n : ℕ) : ℕ := -if h : p ≠ 1 ∧ 0 < n -then (multiplicity p n).get (multiplicity.finite_nat_iff.2 h) -else 0 +if h : p ≠ 1 ∧ 0 < n then (multiplicity p n).get (multiplicity.finite_nat_iff.2 h) else 0 namespace padic_val_nat + open multiplicity + variables {p : ℕ} -/-- `padic_val_nat p 0` is 0 for any `p`. -/ -@[simp] protected lemma zero : padic_val_nat p 0 = 0 := -by simp [padic_val_nat] +/-- `padic_val_nat p 0` is `0` for any `p`. -/ +@[simp] protected lemma zero : padic_val_nat p 0 = 0 := by simp [padic_val_nat] -/-- `padic_val_nat p 1` is 0 for any `p`. -/ -@[simp] protected lemma one : padic_val_nat p 1 = 0 := -by { unfold padic_val_nat, split_ifs, simp } +/-- `padic_val_nat p 1` is `0` for any `p`. -/ +@[simp] protected lemma one : padic_val_nat p 1 = 0 := by { unfold padic_val_nat, split_ifs, simp } -/-- For `p ≠ 0, p ≠ 1, `padic_val_rat p p` is 1. -/ +/-- If `p ≠ 0` and `p ≠ 1`, then `padic_val_rat p p` is `1`. -/ @[simp] lemma self (hp : 1 < p) : padic_val_nat p p = 1 := begin have neq_one : (¬ p = 1) ↔ true, { exact iff_of_true ((ne_of_lt hp).symm) trivial }, - have eq_zero_false : (p = 0) ↔ false, + have eq_zero_false : p = 0 ↔ false, { exact iff_false_intro ((ne_of_lt (trans zero_lt_one hp)).symm) }, simp [padic_val_nat, neq_one, eq_zero_false] end @@ -84,16 +79,14 @@ by { rw padic_val_nat, split_ifs; simp [multiplicity_eq_zero_of_not_dvd h] } end padic_val_nat -/-- -For `p ≠ 1`, the p-adic valuation of an integer `z ≠ 0` is the largest natural number `k` such that -p^k divides z. -If `x = 0` or `p = 1`, then `padic_val_int p q` defaults to 0. --/ -def padic_val_int (p : ℕ) (z : ℤ) : ℕ := -padic_val_nat p (z.nat_abs) +/-- For `p ≠ 1`, the `p`-adic valuation of an integer `z ≠ 0` is the largest natural number `k` such +that `p^k` divides `z`. If `x = 0` or `p = 1`, then `padic_val_int p q` defaults to `0`. -/ +def padic_val_int (p : ℕ) (z : ℤ) : ℕ := padic_val_nat p z.nat_abs namespace padic_val_int + open multiplicity + variables {p : ℕ} lemma of_ne_one_ne_zero {z : ℤ} (hp : p ≠ 1) (hz : z ≠ 0) : padic_val_int p z = @@ -104,21 +97,17 @@ begin refl end -/-- `padic_val_int p 0` is 0 for any `p`. -/ -@[simp] protected lemma zero : padic_val_int p 0 = 0 := -by simp [padic_val_int] +/-- `padic_val_int p 0` is `0` for any `p`. -/ +@[simp] protected lemma zero : padic_val_int p 0 = 0 := by simp [padic_val_int] -/-- `padic_val_int p 1` is 0 for any `p`. -/ -@[simp] protected lemma one : padic_val_int p 1 = 0 := -by simp [padic_val_int] +/-- `padic_val_int p 1` is `0` for any `p`. -/ +@[simp] protected lemma one : padic_val_int p 1 = 0 := by simp [padic_val_int] -/-- The p-adic value of an natural is its p-adic_value as an integer -/ -@[simp] lemma of_nat {n : ℕ} : padic_val_int p (n : ℤ) = padic_val_nat p n := -by simp [padic_val_int] +/-- The `p`-adic value of a natural is its `p`-adic value as an integer. -/ +@[simp] lemma of_nat {n : ℕ} : padic_val_int p n = padic_val_nat p n := by simp [padic_val_int] -/-- For `p ≠ 0, p ≠ 1, `padic_val_int p p` is 1. -/ -lemma self (hp : 1 < p) : padic_val_int p p = 1 := -by simp [padic_val_nat.self hp] +/-- If `p ≠ 0` and `p ≠ 1`, then `padic_val_int p p` is `1`. -/ +lemma self (hp : 1 < p) : padic_val_int p p = 1 := by simp [padic_val_nat.self hp] lemma eq_zero_of_not_dvd {z : ℤ} (h : ¬ (p : ℤ) ∣ z) : padic_val_int p z = 0 := begin @@ -128,13 +117,9 @@ end end padic_val_int -/-- -`padic_val_rat` defines the valuation of a rational `q` to be the valuation of `q.num` minus the -valuation of `q.denom`. -If `q = 0` or `p = 1`, then `padic_val_rat p q` defaults to 0. --/ -def padic_val_rat (p : ℕ) (q : ℚ) : ℤ := -padic_val_int p q.num - padic_val_nat p q.denom +/-- `padic_val_rat` defines the valuation of a rational `q` to be the valuation of `q.num` minus the +valuation of `q.denom`. If `q = 0` or `p = 1`, then `padic_val_rat p q` defaults to `0`. -/ +def padic_val_rat (p : ℕ) (q : ℚ) : ℤ := padic_val_int p q.num - padic_val_nat p q.denom namespace padic_val_rat @@ -146,25 +131,23 @@ variables {p : ℕ} @[simp] protected lemma neg (q : ℚ) : padic_val_rat p (-q) = padic_val_rat p q := by simp [padic_val_rat, padic_val_int] -/-- `padic_val_rat p 0` is 0 for any `p`. -/ +/-- `padic_val_rat p 0` is `0` for any `p`. -/ @[simp] protected lemma zero : padic_val_rat p 0 = 0 := by simp [padic_val_rat] -/-- `padic_val_rat p 1` is 0 for any `p`. -/ +/-- `padic_val_rat p 1` is `0` for any `p`. -/ @[simp] protected lemma one : padic_val_rat p 1 = 0 := by simp [padic_val_rat] -/-- The p-adic value of an integer `z ≠ 0` is its p-adic_value as a rational -/ +/-- The `p`-adic value of an integer `z ≠ 0` is its `p`-adic_value as a rational. -/ @[simp] lemma of_int {z : ℤ} : padic_val_rat p z = padic_val_int p z := by simp [padic_val_rat] -/-- The p-adic value of an integer `z ≠ 0` is the multiplicity of `p` in `z`. -/ +/-- The `p`-adic value of an integer `z ≠ 0` is the multiplicity of `p` in `z`. -/ lemma of_int_multiplicity {z : ℤ} (hp : p ≠ 1) (hz : z ≠ 0) : - padic_val_rat p (z : ℚ) = (multiplicity (p : ℤ) z).get - (finite_int_iff.2 ⟨hp, hz⟩) := + padic_val_rat p (z : ℚ) = (multiplicity (p : ℤ) z).get (finite_int_iff.2 ⟨hp, hz⟩) := by rw [of_int, padic_val_int.of_ne_one_ne_zero hp hz] lemma multiplicity_sub_multiplicity {q : ℚ} (hp : p ≠ 1) (hq : q ≠ 0) : padic_val_rat p q = (multiplicity (p : ℤ) q.num).get (finite_int_iff.2 ⟨hp, rat.num_ne_zero_of_ne_zero hq⟩) - - (multiplicity p q.denom).get - (by { rw [← finite_iff_dom, finite_nat_iff], exact ⟨hp, q.pos⟩ }) := + (multiplicity p q.denom).get (by { rw [← finite_iff_dom, finite_nat_iff], exact ⟨hp, q.pos⟩ }) := begin rw [padic_val_rat, padic_val_int.of_ne_one_ne_zero hp, padic_val_nat, dif_pos], { refl }, @@ -172,11 +155,11 @@ begin { exact rat.num_ne_zero_of_ne_zero hq } end -/-- The p-adic value of an integer `z ≠ 0` is its p-adic_value as a rational -/ +/-- The `p`-adic value of an integer `z ≠ 0` is its `p`-adic value as a rational. -/ @[simp] lemma of_nat {n : ℕ} : padic_val_rat p n = padic_val_nat p n := by simp [padic_val_rat] -/-- For `p ≠ 0, p ≠ 1, `padic_val_rat p p` is 1. -/ -lemma self (hp : 1 < p) : padic_val_rat p p = 1 := by simp [of_nat, hp] +/-- If `p ≠ 0` and `p ≠ 1`, then `padic_val_rat p p` is `1`. -/ +lemma self (hp : 1 < p) : padic_val_rat p p = 1 := by simp [hp] end padic_val_rat @@ -187,17 +170,14 @@ variables {p : ℕ} lemma zero_le_padic_val_rat_of_nat (n : ℕ) : 0 ≤ padic_val_rat p n := by simp /-- `padic_val_rat` coincides with `padic_val_nat`. -/ -@[norm_cast] lemma padic_val_rat_of_nat (n : ℕ) : - ↑(padic_val_nat p n) = padic_val_rat p n := +@[norm_cast] lemma padic_val_rat_of_nat (n : ℕ) : ↑(padic_val_nat p n) = padic_val_rat p n := by simp -/-- -A simplification of `padic_val_nat` when one input is prime, by analogy with `padic_val_rat_def`. --/ +/-- A simplification of `padic_val_nat` when one input is prime, by analogy with +`padic_val_rat_def`. -/ lemma padic_val_nat_def [hp : fact p.prime] {n : ℕ} (hn : 0 < n) : - padic_val_nat p n = - (multiplicity p n).get - (multiplicity.finite_nat_iff.2 ⟨nat.prime.ne_one hp.1, hn⟩) := + padic_val_nat p n + = (multiplicity p n).get (multiplicity.finite_nat_iff.2 ⟨nat.prime.ne_one hp.1, hn⟩) := begin simp [padic_val_nat], split_ifs, @@ -263,32 +243,29 @@ end /-- A rewrite lemma for `padic_val_rat p (q * r)` with conditions `q ≠ 0`, `r ≠ 0`. -/ protected lemma mul {q r : ℚ} (hq : q ≠ 0) (hr : r ≠ 0) : padic_val_rat p (q * r) = padic_val_rat p q + padic_val_rat p r := -have q*r = (q.num * r.num) /. (q.denom * r.denom), by rw_mod_cast rat.mul_num_denom, +have q * r = (q.num * r.num) /. (q.denom * r.denom), by rw_mod_cast rat.mul_num_denom, have hq' : q.num /. q.denom ≠ 0, by rw rat.num_denom; exact hq, have hr' : r.num /. r.denom ≠ 0, by rw rat.num_denom; exact hr, have hp' : _root_.prime (p : ℤ), from nat.prime_iff_prime_int.1 hp.1, begin rw [padic_val_rat.defn p (mul_ne_zero hq hr) this], - conv_rhs { rw [←(@rat.num_denom q), padic_val_rat.defn p hq', - ←(@rat.num_denom r), padic_val_rat.defn p hr'] }, + conv_rhs { rw [← @rat.num_denom q, padic_val_rat.defn p hq', ← @rat.num_denom r, + padic_val_rat.defn p hr'] }, rw [multiplicity.mul' hp', multiplicity.mul' hp']; simp [add_comm, add_left_comm, sub_eq_add_neg] end /-- A rewrite lemma for `padic_val_rat p (q^k)` with condition `q ≠ 0`. -/ protected lemma pow {q : ℚ} (hq : q ≠ 0) {k : ℕ} : - padic_val_rat p (q ^ k) = k * padic_val_rat p q := + padic_val_rat p (q ^ k) = k * padic_val_rat p q := by induction k; simp [*, padic_val_rat.mul hq (pow_ne_zero _ hq), pow_succ, add_mul, add_comm] -/-- -A rewrite lemma for `padic_val_rat p (q⁻¹)` with condition `q ≠ 0`. --/ -protected lemma inv (q : ℚ) : - padic_val_rat p (q⁻¹) = -padic_val_rat p q := +/-- A rewrite lemma for `padic_val_rat p (q⁻¹)` with condition `q ≠ 0`. -/ +protected lemma inv (q : ℚ) : padic_val_rat p q⁻¹ = -padic_val_rat p q := begin by_cases hq : q = 0, { simp [hq] }, - { rw [eq_neg_iff_add_eq_zero, ← padic_val_rat.mul (inv_ne_zero hq) hq, - inv_mul_cancel hq, padic_val_rat.one], + { rw [eq_neg_iff_add_eq_zero, ← padic_val_rat.mul (inv_ne_zero hq) hq, inv_mul_cancel hq, + padic_val_rat.one], exact hp }, end @@ -300,14 +277,12 @@ begin all_goals { exact hp } end -/-- -A condition for `padic_val_rat p (n₁ / d₁) ≤ padic_val_rat p (n₂ / d₂), -in terms of divisibility by `p^n`. --/ +/-- A condition for `padic_val_rat p (n₁ / d₁) ≤ padic_val_rat p (n₂ / d₂)`, in terms of +divisibility by `p^n`. -/ lemma padic_val_rat_le_padic_val_rat_iff {n₁ n₂ d₁ d₂ : ℤ} (hn₁ : n₁ ≠ 0) (hn₂ : n₂ ≠ 0) (hd₁ : d₁ ≠ 0) (hd₂ : d₂ ≠ 0) : padic_val_rat p (n₁ /. d₁) ≤ padic_val_rat p (n₂ /. d₂) ↔ - ∀ (n : ℕ), ↑p ^ n ∣ n₁ * d₂ → ↑p ^ n ∣ n₂ * d₁ := + ∀ n : ℕ, ↑p ^ n ∣ n₁ * d₂ → ↑p ^ n ∣ n₂ * d₁ := have hf1 : finite (p : ℤ) (n₁ * d₂), from finite_int_prime_iff.2 (mul_ne_zero hn₁ hd₂), have hf2 : finite (p : ℤ) (n₂ * d₁), @@ -316,41 +291,35 @@ have hf2 : finite (p : ℤ) (n₂ * d₁), { to_lhs, rw [padic_val_rat.defn p (rat.mk_ne_zero_of_ne_zero hn₁ hd₁) rfl, padic_val_rat.defn p (rat.mk_ne_zero_of_ne_zero hn₂ hd₂) rfl, - sub_le_iff_le_add', - ← add_sub_assoc, - le_sub_iff_add_le], + sub_le_iff_le_add', ← add_sub_assoc, le_sub_iff_add_le], norm_cast, rw [← multiplicity.mul' (nat.prime_iff_prime_int.1 hp.1) hf1, add_comm, ← multiplicity.mul' (nat.prime_iff_prime_int.1 hp.1) hf2, part_enat.get_le_get, multiplicity_le_multiplicity_iff] } -/-- -Sufficient conditions to show that the p-adic valuation of `q` is less than or equal to the -p-adic vlauation of `q + r`. --/ -theorem le_padic_val_rat_add_of_le {q r : ℚ} - (hqr : q + r ≠ 0) - (h : padic_val_rat p q ≤ padic_val_rat p r) : - padic_val_rat p q ≤ padic_val_rat p (q + r) := +/-- Sufficient conditions to show that the `p`-adic valuation of `q` is less than or equal to the +`p`-adic valuation of `q + r`. -/ +theorem le_padic_val_rat_add_of_le {q r : ℚ} (hqr : q + r ≠ 0) + (h : padic_val_rat p q ≤ padic_val_rat p r) : padic_val_rat p q ≤ padic_val_rat p (q + r) := if hq : q = 0 then by simpa [hq] using h else if hr : r = 0 then by simp [hr] else have hqn : q.num ≠ 0, from rat.num_ne_zero_of_ne_zero hq, have hqd : (q.denom : ℤ) ≠ 0, by exact_mod_cast rat.denom_ne_zero _, have hrn : r.num ≠ 0, from rat.num_ne_zero_of_ne_zero hr, have hrd : (r.denom : ℤ) ≠ 0, by exact_mod_cast rat.denom_ne_zero _, -have hqreq : q + r = (((q.num * r.denom + q.denom * r.num : ℤ)) /. (↑q.denom * ↑r.denom : ℤ)), +have hqreq : q + r = (q.num * r.denom + q.denom * r.num) /. (q.denom * r.denom), from rat.add_num_denom _ _, -have hqrd : q.num * ↑(r.denom) + ↑(q.denom) * r.num ≠ 0, +have hqrd : q.num * r.denom + q.denom * r.num ≠ 0, from rat.mk_num_ne_zero_of_ne_zero hqr hqreq, begin - conv_lhs { rw ←(@rat.num_denom q) }, + conv_lhs { rw ← @rat.num_denom q }, rw [hqreq, padic_val_rat_le_padic_val_rat_iff hqn hqrd hqd (mul_ne_zero hqd hrd), ← multiplicity_le_multiplicity_iff, mul_left_comm, multiplicity.mul (nat.prime_iff_prime_int.1 hp.1), add_mul], - rw [←(@rat.num_denom q), ←(@rat.num_denom r), - padic_val_rat_le_padic_val_rat_iff hqn hrn hqd hrd, ← multiplicity_le_multiplicity_iff] at h, - calc _ ≤ min (multiplicity ↑p (q.num * ↑(r.denom) * ↑(q.denom))) - (multiplicity ↑p (↑(q.denom) * r.num * ↑(q.denom))) : (le_min + rw [← @rat.num_denom q, ← @rat.num_denom r, padic_val_rat_le_padic_val_rat_iff hqn hrn hqd hrd, + ← multiplicity_le_multiplicity_iff] at h, + calc _ ≤ min (multiplicity ↑p (q.num * ↑r.denom * ↑q.denom)) + (multiplicity ↑p (↑q.denom * r.num * ↑q.denom)) : (le_min (by rw [@multiplicity.mul _ _ _ _ (_ * _) _ (nat.prime_iff_prime_int.1 hp.1), add_comm]) (by rw [mul_assoc, @multiplicity.mul _ _ _ _ (q.denom : ℤ) (_ * _) (nat.prime_iff_prime_int.1 hp.1)]; exact add_le_add_left h _)) @@ -358,23 +327,19 @@ begin all_goals { exact hp } end -/-- -The minimum of the valuations of `q` and `r` is less than or equal to the valuation of `q + r`. --/ +/-- The minimum of the valuations of `q` and `r` is at most the valuation of `q + r`. -/ theorem min_le_padic_val_rat_add {q r : ℚ} (hqr : q + r ≠ 0) : min (padic_val_rat p q) (padic_val_rat p r) ≤ padic_val_rat p (q + r) := (le_total (padic_val_rat p q) (padic_val_rat p r)).elim (λ h, by rw [min_eq_left h]; exact le_padic_val_rat_add_of_le hqr h) - (λ h, by rw [min_eq_right h, add_comm]; exact le_padic_val_rat_add_of_le - (by rwa add_comm) h) + (λ h, by rw [min_eq_right h, add_comm]; exact le_padic_val_rat_add_of_le (by rwa add_comm) h) open_locale big_operators -/-- A finite sum of rationals with positive p-adic valuation has positive p-adic valuation - (if the sum is non-zero). -/ -theorem sum_pos_of_pos {n : ℕ} {F : ℕ → ℚ} - (hF : ∀ i, i < n → 0 < padic_val_rat p (F i)) (hn0 : ∑ i in finset.range n, F i ≠ 0) : - 0 < padic_val_rat p (∑ i in finset.range n, F i) := +/-- A finite sum of rationals with positive `p`-adic valuation has positive `p`-adic valuation +(if the sum is non-zero). -/ +theorem sum_pos_of_pos {n : ℕ} {F : ℕ → ℚ} (hF : ∀ i, i < n → 0 < padic_val_rat p (F i)) + (hn0 : ∑ i in finset.range n, F i ≠ 0) : 0 < padic_val_rat p (∑ i in finset.range n, F i) := begin induction n with d hd, { exact false.elim (hn0 rfl) }, @@ -418,7 +383,7 @@ begin exact hp end -/-- Dividing out by a prime factor reduces the padic_val_nat by 1. -/ +/-- Dividing out by a prime factor reduces the `padic_val_nat` by `1`. -/ protected lemma div (dvd : p ∣ b) : padic_val_nat p (b / p) = (padic_val_nat p b) - 1 := begin convert padic_val_nat.div_of_dvd dvd, @@ -426,7 +391,7 @@ begin exact hp end -/-- A version of `padic_val_rat.pow` for `padic_val_nat` -/ +/-- A version of `padic_val_rat.pow` for `padic_val_nat`. -/ protected lemma pow (n : ℕ) (ha : a ≠ 0) : padic_val_nat p (a ^ n) = n * padic_val_nat p a := by simpa only [← @nat.cast_inj ℤ] with push_cast using padic_val_rat.pow (cast_ne_zero.mpr ha) @@ -445,22 +410,23 @@ end padic_val_nat section padic_val_nat -lemma dvd_of_one_le_padic_val_nat {p n : ℕ} (hp : 1 ≤ padic_val_nat p n) : - p ∣ n := +variables {p : ℕ} + +lemma dvd_of_one_le_padic_val_nat {n : ℕ} (hp : 1 ≤ padic_val_nat p n) : p ∣ n := begin by_contra h, rw padic_val_nat.eq_zero_of_not_dvd h at hp, exact lt_irrefl 0 (lt_of_lt_of_le zero_lt_one hp) end -lemma pow_padic_val_nat_dvd {p n : ℕ} : p ^ (padic_val_nat p n) ∣ n := +lemma pow_padic_val_nat_dvd {n : ℕ} : p ^ padic_val_nat p n ∣ n := begin rcases n.eq_zero_or_pos with rfl | hn, { simp }, rcases eq_or_ne p 1 with rfl | hp, { simp }, rw [multiplicity.pow_dvd_iff_le_multiplicity, padic_val_nat_def']; assumption end -lemma pow_succ_padic_val_nat_not_dvd {p n : ℕ} [hp : fact (nat.prime p)] (hn : 0 < n) : +lemma pow_succ_padic_val_nat_not_dvd {n : ℕ} [hp : fact p.prime] (hn : 0 < n) : ¬ p ^ (padic_val_nat p n + 1) ∣ n := begin rw multiplicity.pow_dvd_iff_le_multiplicity, @@ -472,8 +438,8 @@ begin { apply_instance } end -lemma padic_val_nat_dvd_iff {p : ℕ} (n : ℕ) [hp : fact p.prime] (a : ℕ) : - p^n ∣ a ↔ a = 0 ∨ n ≤ padic_val_nat p a := +lemma padic_val_nat_dvd_iff (n : ℕ) [hp : fact p.prime] (a : ℕ) : + p ^ n ∣ a ↔ a = 0 ∨ n ≤ padic_val_nat p a := begin split, { rw [pow_dvd_iff_le_multiplicity, padic_val_nat], @@ -487,12 +453,12 @@ begin { exact dvd_trans (pow_dvd_pow p h) pow_padic_val_nat_dvd } } end -lemma padic_val_nat_primes {p q : ℕ} [hp : fact p.prime] [hq : fact q.prime] - (neq : p ≠ q) : padic_val_nat p q = 0 := +lemma padic_val_nat_primes {q : ℕ} [hp : fact p.prime] [hq : fact q.prime] (neq : p ≠ q) : + padic_val_nat p q = 0 := @padic_val_nat.eq_zero_of_not_dvd p q $ (not_congr (iff.symm (prime_dvd_prime_iff_eq hp.1 hq.1))).mp neq -protected lemma padic_val_nat.div' {p : ℕ} [hp : fact p.prime] : +protected lemma padic_val_nat.div' [hp : fact p.prime] : ∀ {m : ℕ} (cpm : coprime p m) {b : ℕ} (dvd : m ∣ b), padic_val_nat p (b / m) = padic_val_nat p b | 0 := λ cpm b dvd, by { rw zero_dvd_iff at dvd, rw [dvd, nat.zero_div] } | (n + 1) := @@ -502,7 +468,7 @@ protected lemma padic_val_nat.div' {p : ℕ} [hp : fact p.prime] : rw [mul_div_right c (nat.succ_pos _)],by_cases hc : c = 0, { rw [hc, mul_zero] }, { rw padic_val_nat.mul, - { suffices : ¬ p ∣ (n+1), + { suffices : ¬ p ∣ (n + 1), { rw [padic_val_nat.eq_zero_of_not_dvd this, zero_add] }, contrapose! cpm, exact hp.1.dvd_iff_not_coprime.mp cpm }, @@ -512,7 +478,7 @@ protected lemma padic_val_nat.div' {p : ℕ} [hp : fact p.prime] : open_locale big_operators -lemma range_pow_padic_val_nat_subset_divisors {n : ℕ} (p : ℕ) (hn : n ≠ 0) : +lemma range_pow_padic_val_nat_subset_divisors {n : ℕ} (hn : n ≠ 0) : (finset.range (padic_val_nat p n + 1)).image (pow p) ⊆ n.divisors := begin intros t ht, @@ -522,8 +488,8 @@ begin exact ⟨(pow_dvd_pow p $ by linarith).trans pow_padic_val_nat_dvd, hn⟩ end -lemma range_pow_padic_val_nat_subset_divisors' {p n : ℕ} [hp : fact p.prime] : - (finset.range (padic_val_nat p n)).image (λ t, p ^ (t + 1)) ⊆ (n.divisors \ {1}) := +lemma range_pow_padic_val_nat_subset_divisors' {n : ℕ} [hp : fact p.prime] : + (finset.range (padic_val_nat p n)).image (λ t, p ^ (t + 1)) ⊆ n.divisors \ {1} := begin rcases eq_or_ne n 0 with rfl | hn, { simp }, @@ -558,7 +524,7 @@ end lemma padic_val_int_self : padic_val_int p p = 1 := padic_val_int.self hp.out.one_lt lemma padic_val_int.mul {a b : ℤ} (ha : a ≠ 0) (hb : b ≠ 0) : - padic_val_int p (a*b) = padic_val_int p a + padic_val_int p b := + padic_val_int p (a * b) = padic_val_int p a + padic_val_int p b := begin simp_rw padic_val_int, rw [int.nat_abs_mul, padic_val_nat.mul]; diff --git a/src/number_theory/padics/ring_homs.lean b/src/number_theory/padics/ring_homs.lean index 3172279bbab27..d699e84a98519 100644 --- a/src/number_theory/padics/ring_homs.lean +++ b/src/number_theory/padics/ring_homs.lean @@ -172,7 +172,7 @@ variable (x : ℤ_[p]) lemma exists_mem_range : ∃ n : ℕ, n < p ∧ (x - n ∈ maximal_ideal ℤ_[p]) := begin simp only [maximal_ideal_eq_span_p, ideal.mem_span_singleton, ← norm_lt_one_iff_dvd], - obtain ⟨r, hr⟩ := rat_dense (x : ℚ_[p]) zero_lt_one, + obtain ⟨r, hr⟩ := rat_dense p (x : ℚ_[p]) zero_lt_one, have H : ∥(r : ℚ_[p])∥ ≤ 1, { rw norm_sub_rev at hr, calc _ = ∥(r : ℚ_[p]) - x + x∥ : by ring_nf From f38b4dfebca13bf736338a5b4cd44e7bd6c21929 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 26 Sep 2022 19:49:21 +0000 Subject: [PATCH 399/828] feat(topology/sheaves/skyscraper): define skyscraper presheaves and calculate stalks (#15934) --- src/category_theory/eq_to_hom.lean | 10 ++ src/topology/inseparable.lean | 8 ++ src/topology/sheaves/skyscraper.lean | 161 +++++++++++++++++++++++++++ 3 files changed, 179 insertions(+) create mode 100644 src/topology/sheaves/skyscraper.lean diff --git a/src/category_theory/eq_to_hom.lean b/src/category_theory/eq_to_hom.lean index c3c30dea4260a..c256bf5de9c76 100644 --- a/src/category_theory/eq_to_hom.lean +++ b/src/category_theory/eq_to_hom.lean @@ -45,6 +45,16 @@ def eq_to_hom {X Y : C} (p : X = Y) : X ⟶ Y := by rw p; exact 𝟙 _ eq_to_hom p ≫ eq_to_hom q = eq_to_hom (p.trans q) := by { cases p, cases q, simp, } +lemma comp_eq_to_hom_iff {X Y Y' : C} (p : Y = Y') (f : X ⟶ Y) (g : X ⟶ Y') : + f ≫ eq_to_hom p = g ↔ f = g ≫ eq_to_hom p.symm := +{ mp := λ h, h ▸ by simp, + mpr := λ h, by simp [eq_whisker h (eq_to_hom p)] } + +lemma eq_to_hom_comp_iff {X X' Y : C} (p : X = X') (f : X ⟶ Y) (g : X' ⟶ Y) : + eq_to_hom p ≫ g = f ↔ g = eq_to_hom p.symm ≫ f := +{ mp := λ h, h ▸ by simp, + mpr := λ h, h ▸ by simp [whisker_eq _ h] } + /-- If we (perhaps unintentionally) perform equational rewriting on the source object of a morphism, diff --git a/src/topology/inseparable.lean b/src/topology/inseparable.lean index fb588e33b7e88..8a1b60fcfac36 100644 --- a/src/topology/inseparable.lean +++ b/src/topology/inseparable.lean @@ -151,6 +151,14 @@ by simp only [specializes_iff_mem_closure, hf.closure_eq_preimage_closure_image, lemma subtype_specializes_iff {p : X → Prop} (x y : subtype p) : x ⤳ y ↔ (x : X) ⤳ y := inducing_coe.specializes_iff.symm +lemma not_specializes_iff_exists_open : ¬ x ⤳ y ↔ ∃ (S : set X), is_open S ∧ y ∈ S ∧ x ∉ S := +⟨λ h, by { contrapose! h, rwa specializes_iff_forall_open, }, + λ h, by { contrapose! h, rwa ←specializes_iff_forall_open, }⟩ + +lemma not_specializes_iff_exists_closed : ¬ x ⤳ y ↔ ∃ (S : set X), is_closed S ∧ x ∈ S ∧ y ∉ S := +⟨λ h, by { contrapose! h, rwa specializes_iff_forall_closed, }, + λ h, by { contrapose! h, rwa ←specializes_iff_forall_closed, }⟩ + variable (X) /-- Specialization forms a preorder on the topological space. -/ diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean new file mode 100644 index 0000000000000..6752445b8d277 --- /dev/null +++ b/src/topology/sheaves/skyscraper.lean @@ -0,0 +1,161 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import algebraic_geometry.sheafed_space +import topology.sheaves.sheaf_condition.opens_le_cover +import topology.sheaves.stalks +import category_theory.preadditive.injective + +/-! +# Skyscraper (pre)sheaves + +A skyscraper (pre)sheaf `𝓕 : (pre)sheaf C X` is the (pre)sheaf with value `A` at point `p₀` that is +supported only at open sets contain `p₀`, i.e. `𝓕(U) = A` if `p₀ ∈ U` and `𝓕(U) = *` if `p₀ ∉ U` +where `*` is a terminal object of `C`. In terms of stalks, `𝓕` is supported at all specializations +of `p₀`, i.e. if `p₀ ⤳ x` then `𝓕ₓ ≅ A` and if `¬ p₀ ⤳ x` then `𝓕ₓ ≅ *`. + +## Main definitions + +* `skyscraper_presheaf`: `skyscraper_presheaf p₀ A` is the skyscraper presheaf at point `p₀` with + value `A`. +* `skyscraper_sheaf`: the skyscraper presheaf satisfies the sheaf condition. + +## Main statements + +* `skyscraper_presheaf_stalk_of_specializes`: if `y ∈ closure {p₀}` then the stalk of + `skyscraper_presheaf p₀ A` at `y` is `A`. +* `skyscraper_presheaf_stalk_of_not_specializes`: if `y ∉ closure {p₀}` then the stalk of + `skyscraper_presheaf p₀ A` at `y` is `*` the terminal object. + +TODO: generalize universe level when calculating stalks, after generalizing universe level of stalk. +-/ + +noncomputable theory + +open topological_space Top category_theory category_theory.limits opposite + +universes u v w + +variables {X : Top.{u}} (p₀ : X) [Π (U : opens X), decidable (p₀ ∈ U)] + +section + +variables {C : Type v} [category.{w} C] (A : C) [has_terminal C] + +/-- +A skyscraper presheaf is a presheaf supported at a single point: if `p₀ ∈ X` is a specified +point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A` if `p₀ ∈ U` and +`U ↦ *` if `p₀ ∉ A` where `*` is some terminal object. +-/ +@[simps] def skyscraper_presheaf : presheaf C X := +{ obj := λ U, if p₀ ∈ unop U then A else terminal C, + map := λ U V i, if h : p₀ ∈ unop V + then eq_to_hom $ by erw [if_pos h, if_pos (le_of_hom i.unop h)] + else ((if_neg h).symm.rec terminal_is_terminal).from _, + map_id' := λ U, (em (p₀ ∈ U.unop)).elim (λ h, dif_pos h) + (λ h, ((if_neg h).symm.rec terminal_is_terminal).hom_ext _ _), + map_comp' := λ U V W iVU iWV, + begin + by_cases hW : p₀ ∈ unop W, + { have hV : p₀ ∈ unop V := le_of_hom iWV.unop hW, + simp only [dif_pos hW, dif_pos hV, eq_to_hom_trans] }, + { rw [dif_neg hW], apply ((if_neg hW).symm.rec terminal_is_terminal).hom_ext } + end } + +end + +section + +-- In this section, we calculate the stalks for skyscraper presheaves. +-- We need to restrict universe level. + +variables {C : Type v} [category.{u} C] (A : C) [has_terminal C] + +/-- +The cocone at `A` for the stalk functor of `skyscraper_presheaf p₀ A` when `y ∈ closure {p₀}` +-/ +@[simps] def skyscraper_presheaf_cocone_of_specializes {y : X} (h : p₀ ⤳ y) : + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ A) := +{ X := A, + ι := { app := λ U, eq_to_hom $ if_pos $ h.mem_open U.unop.1.2 U.unop.2, + naturality' := λ U V inc, begin + change dite _ _ _ ≫ _ = _, rw dif_pos, + { erw [category.comp_id, eq_to_hom_trans], refl }, + { exact h.mem_open V.unop.1.2 V.unop.2 }, + end } } + +/-- +The cocone at `A` for the stalk functor of `skyscraper_presheaf p₀ A` when `y ∈ closure {p₀}` is a +colimit +-/ +noncomputable def skyscraper_presheaf_cocone_is_colimit_of_specializes + {y : X} (h : p₀ ⤳ y) : is_colimit (skyscraper_presheaf_cocone_of_specializes p₀ A h) := +{ desc := λ c, eq_to_hom (if_pos trivial).symm ≫ c.ι.app (op ⊤), + fac' := λ c U, begin + rw ← c.w (hom_of_le $ (le_top : unop U ≤ _)).op, + change _ ≫ _ ≫ dite _ _ _ ≫ _ = _, + rw dif_pos, + { simpa only [skyscraper_presheaf_cocone_of_specializes_ι_app, + eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp] }, + { exact h.mem_open U.unop.1.2 U.unop.2 }, + end, + uniq' := λ c f h, by rw [← h, skyscraper_presheaf_cocone_of_specializes_ι_app, + eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp] } + +/-- +If `y ∈ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ A` at `y` is `A`. +-/ +noncomputable def skyscraper_presheaf_stalk_of_specializes [has_colimits C] + {y : X} (h : p₀ ⤳ y) : (skyscraper_presheaf p₀ A).stalk y ≅ A := +colimit.iso_colimit_cocone ⟨_, skyscraper_presheaf_cocone_is_colimit_of_specializes p₀ A h⟩ + +/-- +The cocone at `*` for the stalk functor of `skyscraper_presheaf p₀ A` when `y ∉ closure {p₀}` +-/ +@[simps] def skyscraper_presheaf_cocone (y : X) : + cocone ((open_nhds.inclusion y).op ⋙ skyscraper_presheaf p₀ A) := +{ X := terminal C, + ι := + { app := λ U, terminal.from _, + naturality' := λ U V inc, terminal_is_terminal.hom_ext _ _ } } + +/-- +The cocone at `*` for the stalk functor of `skyscraper_presheaf p₀ A` when `y ∉ closure {p₀}` is a +colimit +-/ +noncomputable def skyscraper_presheaf_cocone_is_colimit_of_not_specializes + {y : X} (h : ¬p₀ ⤳ y) : is_colimit (skyscraper_presheaf_cocone p₀ A y) := +let h1 : ∃ (U : open_nhds y), p₀ ∉ U.1 := + let ⟨U, ho, h₀, hy⟩ := not_specializes_iff_exists_open.mp h in ⟨⟨⟨U, ho⟩, h₀⟩, hy⟩ in +{ desc := λ c, eq_to_hom (if_neg h1.some_spec).symm ≫ c.ι.app (op h1.some), + fac' := λ c U, begin + change _ = c.ι.app (op U.unop), + simp only [← c.w (hom_of_le $ @inf_le_left _ _ h1.some U.unop).op, + ← c.w (hom_of_le $ @inf_le_right _ _ h1.some U.unop).op, ← category.assoc], + congr' 1, + refine ((if_neg _).symm.rec terminal_is_terminal).hom_ext _ _, + exact λ h, h1.some_spec h.1, + end, + uniq' := λ c f H, begin + rw [← category.id_comp f, ← H, ← category.assoc], + congr' 1, apply terminal_is_terminal.hom_ext, + end } + +/-- +If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ A` at `y` is isomorphic to a +terminal object. +-/ +noncomputable def skyscraper_presheaf_stalk_of_not_specializes [has_colimits C] + {y : X} (h : ¬p₀ ⤳ y) : (skyscraper_presheaf p₀ A).stalk y ≅ terminal C := +colimit.iso_colimit_cocone ⟨_, skyscraper_presheaf_cocone_is_colimit_of_not_specializes _ A h⟩ + +/-- +If `y ∉ closure {p₀}`, then the stalk of `skyscraper_presheaf p₀ A` at `y` is a terminal object +-/ +def skyscraper_presheaf_stalk_of_not_specializes_is_terminal + [has_colimits C] {y : X} (h : ¬p₀ ⤳ y) : is_terminal ((skyscraper_presheaf p₀ A).stalk y) := +is_terminal.of_iso terminal_is_terminal $ (skyscraper_presheaf_stalk_of_not_specializes _ _ h).symm + +end From 9f972c7eb9598ebf4159c65e7a7a776961d2d6f0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Mon, 26 Sep 2022 19:49:22 +0000 Subject: [PATCH 400/828] feat(topology/uniform_space/completion): add uniform_completion.complete_equiv_self (#16612) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - Change ```abstract_completion.compare_equiv``` to uniform bijection. - Define the ```abstract_completion α``` given by ```α``` when it is complete. - Use it to prove that there is a uniform bijection between a complete space and its ```uniform_completion```. - Upgrade the bijection between ```Bourbaki reals``` and ```Cauchy reals``` to a uniform bijection. - Add a new function ```function.dense_range_id``` (needed in one of the proofs) --- src/topology/basic.lean | 3 +++ .../uniform_space/abstract_completion.lean | 15 ++++++++++----- src/topology/uniform_space/compare_reals.lean | 4 ++-- src/topology/uniform_space/completion.lean | 5 +++++ 4 files changed, 20 insertions(+), 7 deletions(-) diff --git a/src/topology/basic.lean b/src/topology/basic.lean index b485922e4669e..cd10a32dfe068 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -1446,6 +1446,9 @@ variables {f} lemma function.surjective.dense_range (hf : function.surjective f) : dense_range f := λ x, by simp [hf.range_eq] +lemma dense_range_id : dense_range (id : α → α) := +function.surjective.dense_range function.surjective_id + lemma dense_range_iff_closure_range : dense_range f ↔ closure (range f) = univ := dense_iff_closure_eq diff --git a/src/topology/uniform_space/abstract_completion.lean b/src/topology/uniform_space/abstract_completion.lean index 75c54a385929b..06110a8d31994 100644 --- a/src/topology/uniform_space/abstract_completion.lean +++ b/src/topology/uniform_space/abstract_completion.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot -/ import topology.uniform_space.uniform_embedding +import topology.uniform_space.equiv /-! # Abstract theory of Hausdorff completions of uniform spaces @@ -67,6 +68,10 @@ variables {α : Type*} [uniform_space α] (pkg : abstract_completion α) local notation `hatα` := pkg.space local notation `ι` := pkg.coe +/-- If `α` is complete, then it is an abstract completion of itself. -/ +def of_complete [separated_space α] [complete_space α] : abstract_completion α := +mk α id infer_instance infer_instance infer_instance uniform_inducing_id dense_range_id + lemma closure_range : closure (range ι) = univ := pkg.dense.closure_range @@ -222,12 +227,14 @@ begin refl end -/-- The bijection between two completions of the same uniform space. -/ -def compare_equiv : pkg.space ≃ pkg'.space := +/-- The uniform bijection between two completions of the same uniform space. -/ +def compare_equiv : pkg.space ≃ᵤ pkg'.space := { to_fun := pkg.compare pkg', inv_fun := pkg'.compare pkg, left_inv := congr_fun (pkg'.inverse_compare pkg), - right_inv := congr_fun (pkg.inverse_compare pkg') } + right_inv := congr_fun (pkg.inverse_compare pkg'), + uniform_continuous_to_fun := uniform_continuous_compare _ _, + uniform_continuous_inv_fun := uniform_continuous_compare _ _, } lemma uniform_continuous_compare_equiv : uniform_continuous (pkg.compare_equiv pkg') := pkg.uniform_continuous_compare pkg' @@ -253,8 +260,6 @@ protected def prod : abstract_completion (α × β) := dense := pkg.dense.prod_map pkg'.dense } end prod - - section extension₂ variables (pkg' : abstract_completion β) local notation `hatβ` := pkg'.space diff --git a/src/topology/uniform_space/compare_reals.lean b/src/topology/uniform_space/compare_reals.lean index c09c511925a79..4da0b8014a4e1 100644 --- a/src/topology/uniform_space/compare_reals.lean +++ b/src/topology/uniform_space/compare_reals.lean @@ -103,8 +103,8 @@ instance bourbaki.uniform_space: uniform_space Bourbakiℝ := completion.uniform /-- Bourbaki reals packaged as a completion of Q using the general theory. -/ def Bourbaki_pkg : abstract_completion Q := completion.cpkg -/-- The equivalence between Bourbaki and Cauchy reals-/ -noncomputable def compare_equiv : Bourbakiℝ ≃ ℝ := +/-- The uniform bijection between Bourbaki and Cauchy reals. -/ +noncomputable def compare_equiv : Bourbakiℝ ≃ᵤ ℝ := Bourbaki_pkg.compare_equiv rational_cau_seq_pkg lemma compare_uc : uniform_continuous (compare_equiv) := diff --git a/src/topology/uniform_space/completion.lean b/src/topology/uniform_space/completion.lean index a566557cd2fb6..fdc70b7d2714b 100644 --- a/src/topology/uniform_space/completion.lean +++ b/src/topology/uniform_space/completion.lean @@ -407,6 +407,11 @@ lemma dense_inducing_coe : dense_inducing (coe : α → completion α) := { dense := dense_range_coe, ..(uniform_inducing_coe α).inducing } +/-- The uniform bijection between a complete space and its uniform completion. -/ +def uniform_completion.complete_equiv_self [complete_space α] [separated_space α]: + completion α ≃ᵤ α := +abstract_completion.compare_equiv completion.cpkg abstract_completion.of_complete + open topological_space instance separable_space_completion [separable_space α] : separable_space (completion α) := From b3d23d741a3f7e4e5d3b1538e460ba9deccd63f9 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Mon, 26 Sep 2022 19:49:24 +0000 Subject: [PATCH 401/828] chore(analysis/special_functions/pow): squeeze a nonterminal simp (#16657) --- src/analysis/special_functions/pow.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index e9d02a03a042f..fb147b521f216 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -60,7 +60,7 @@ lemma zero_cpow_eq_iff {x : ℂ} {a : ℂ} : 0 ^ x = a ↔ (x ≠ 0 ∧ a = 0) begin split, { intros hyp, - simp [cpow_def] at hyp, + simp only [cpow_def, eq_self_iff_true, if_true] at hyp, by_cases x = 0, { subst h, simp only [if_true, eq_self_iff_true] at hyp, right, exact ⟨rfl, hyp.symm⟩}, { rw if_neg h at hyp, left, exact ⟨h, hyp.symm⟩, }, }, From ab2e9fb6a0bf887fd1364b3ba9769e4d4bf591a4 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 26 Sep 2022 21:59:53 +0000 Subject: [PATCH 402/828] feat(algebra/category/Group/injective): divisible groups are injective in category of `AddCommGroup` (#16110) --- src/algebra/category/Group/injective.lean | 109 ++++++++++++++++++++++ 1 file changed, 109 insertions(+) create mode 100644 src/algebra/category/Group/injective.lean diff --git a/src/algebra/category/Group/injective.lean b/src/algebra/category/Group/injective.lean new file mode 100644 index 0000000000000..aa98c46120ec5 --- /dev/null +++ b/src/algebra/category/Group/injective.lean @@ -0,0 +1,109 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import algebra.category.Group.epi_mono +import algebra.category.Group.Z_Module_equivalence +import algebra.category.Module.epi_mono +import algebra.module.injective +import category_theory.preadditive.injective +import group_theory.divisible +import ring_theory.principal_ideal_domain + +/-! +# Injective objects in the category of abelian groups + +In this file we prove that divisible groups are injective object in category of (additive) abelian +groups. + +-/ + + +open category_theory +open_locale pointwise + +universe u + +variables (A : Type u) [add_comm_group A] + +namespace AddCommGroup + +lemma injective_of_injective_as_module [injective (⟨A⟩ : Module ℤ)] : + category_theory.injective (⟨A⟩ : AddCommGroup) := +{ factors := λ X Y g f m, + begin + resetI, + let G : (⟨X⟩ : Module ℤ) ⟶ ⟨A⟩ := + { map_smul' := by { intros, rw [ring_hom.id_apply, g.to_fun_eq_coe, map_zsmul], }, ..g }, + let F : (⟨X⟩ : Module ℤ) ⟶ ⟨Y⟩ := + { map_smul' := by { intros, rw [ring_hom.id_apply, f.to_fun_eq_coe, map_zsmul], }, ..f }, + haveI : mono F, + { refine ⟨λ Z α β eq1, _⟩, + let α' : AddCommGroup.of Z ⟶ X := α.to_add_monoid_hom, + let β' : AddCommGroup.of Z ⟶ X := β.to_add_monoid_hom, + have eq2 : α' ≫ f = β' ≫ f, + { ext, + simp only [category_theory.comp_apply, linear_map.to_add_monoid_hom_coe], + simpa only [Module.coe_comp, linear_map.coe_mk, + function.comp_app] using fun_like.congr_fun eq1 x }, + rw cancel_mono at eq2, + ext, simpa only using fun_like.congr_fun eq2 x, }, + refine ⟨(injective.factor_thru G F).to_add_monoid_hom, _⟩, + ext, convert fun_like.congr_fun (injective.comp_factor_thru G F) x, + end } + +lemma injective_as_module_of_injective_as_Ab [injective (⟨A⟩ : AddCommGroup)] : + injective (⟨A⟩ : Module ℤ) := +{ factors := λ X Y g f m, + begin + resetI, + let G : (⟨X⟩ : AddCommGroup) ⟶ ⟨A⟩ := g.to_add_monoid_hom, + let F : (⟨X⟩ : AddCommGroup) ⟶ ⟨Y⟩ := f.to_add_monoid_hom, + haveI : mono F, + { rw mono_iff_injective, intros _ _ h, exact ((Module.mono_iff_injective f).mp m) h, }, + refine ⟨{map_smul' := _, ..injective.factor_thru G F}, _⟩, + { intros m x, rw [add_monoid_hom.to_fun_eq_coe, ring_hom.id_apply], + induction m using int.induction_on with n hn n hn, + { rw [zero_smul], + convert map_zero _, + convert zero_smul _ x, }, + { simp only [add_smul, map_add, hn, one_smul], }, + { simp only [sub_smul, map_sub, hn, one_smul] }, }, + ext, convert fun_like.congr_fun (injective.comp_factor_thru G F) x, + end } + +instance injective_of_divisible [divisible_by A ℤ] : + category_theory.injective (⟨A⟩ : AddCommGroup) := +@@injective_of_injective_as_module A _ $ +@@module.injective_object_of_injective_module ℤ _ A _ _ $ +module.Baer.injective $ +λ I g, begin + rcases is_principal_ideal_ring.principal I with ⟨m, rfl⟩, + by_cases m_eq_zero : m = 0, + { subst m_eq_zero, + refine ⟨{ to_fun := _, map_add' := _, map_smul' := _ }, λ n hn, _⟩, + { intros n, exact g 0, }, + { intros n1 n2, + simp only [map_zero, add_zero] }, + { intros n1 n2, + simp only [map_zero, smul_zero], }, + { rw [submodule.span_singleton_eq_bot.mpr rfl, submodule.mem_bot] at hn, + simp only [hn, map_zero], + symmetry, + convert map_zero _, }, }, + { set gₘ := g ⟨m, submodule.subset_span (set.mem_singleton _)⟩ with gm_eq, + refine ⟨{ to_fun := _, map_add' := _, map_smul' := _ }, λ n hn, _⟩, + { intros n, + exact n • divisible_by.div gₘ m, }, + { intros n1 n2, simp only [add_smul], }, + { intros n1 n2, + rw [ring_hom.id_apply, smul_eq_mul, mul_smul], }, + { rw submodule.mem_span_singleton at hn, + rcases hn with ⟨n, rfl⟩, + simp only [gm_eq, algebra.id.smul_eq_mul, linear_map.coe_mk], + rw [mul_smul, divisible_by.div_cancel (g ⟨m, _⟩) m_eq_zero, ←linear_map.map_smul], + congr, }, }, +end + +end AddCommGroup From 1a4f92795e3b047b4de60ac2aa1f4547e913567c Mon Sep 17 00:00:00 2001 From: mcdoll Date: Mon, 26 Sep 2022 21:59:54 +0000 Subject: [PATCH 403/828] chore(analysis/special_functions/pow): squeezing some simps (#16660) --- src/analysis/special_functions/pow.lean | 21 +++++++++++---------- 1 file changed, 11 insertions(+), 10 deletions(-) diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index fb147b521f216..f0f6486f4bd76 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -80,7 +80,7 @@ else by rw [cpow_def, if_neg (one_ne_zero : (1 : ℂ) ≠ 0), if_neg hx, mul_one by rw cpow_def; split_ifs; simp [one_ne_zero, *] at * lemma cpow_add {x : ℂ} (y z : ℂ) (hx : x ≠ 0) : x ^ (y + z) = x ^ y * x ^ z := -by simp [cpow_def]; simp [*, exp_add, mul_add] at * +by simp only [cpow_def, ite_mul, boole_mul, mul_ite, mul_boole]; simp [*, exp_add, mul_add] at * lemma cpow_mul {x y : ℂ} (z : ℂ) (h₁ : -π < (log x * y).im) (h₂ : (log x * y).im ≤ π) : x ^ (y * z) = (x ^ y) ^ z := @@ -91,7 +91,7 @@ begin end lemma cpow_neg (x y : ℂ) : x ^ -y = (x ^ y)⁻¹ := -by simp [cpow_def]; split_ifs; simp [exp_neg] +by simp only [cpow_def, neg_eq_zero, mul_neg]; split_ifs; simp [exp_neg] lemma cpow_sub {x : ℂ} (y z : ℂ) (hx : x ≠ 0) : x ^ (y - z) = x ^ y / x ^ z := by rw [sub_eq_add_neg, cpow_add _ _ hx, cpow_neg, div_eq_mul_inv] @@ -335,7 +335,7 @@ lemma zero_rpow_eq_iff {x : ℝ} {a : ℝ} : 0 ^ x = a ↔ (x ≠ 0 ∧ a = 0) begin split, { intros hyp, - simp [rpow_def] at hyp, + simp only [rpow_def, complex.of_real_zero] at hyp, by_cases x = 0, { subst h, simp only [complex.one_re, complex.of_real_zero, complex.cpow_zero] at hyp, @@ -395,7 +395,8 @@ end real namespace complex lemma of_real_cpow {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : ((x ^ y : ℝ) : ℂ) = (x : ℂ) ^ (y : ℂ) := -by simp [real.rpow_def_of_nonneg hx, complex.cpow_def]; split_ifs; simp [complex.of_real_log hx] +by simp only [real.rpow_def_of_nonneg hx, complex.cpow_def, of_real_eq_zero]; split_ifs; + simp [complex.of_real_log hx] lemma of_real_cpow_of_nonpos {x : ℝ} (hx : x ≤ 0) (y : ℂ) : (x : ℂ) ^ y = ((-x) : ℂ) ^ y * exp (π * I * y) := @@ -1383,7 +1384,7 @@ begin rw this, refine continuous_real_to_nnreal.continuous_at.comp (continuous_at.comp _ _), { apply real.continuous_at_rpow, - simp at h, + simp only [ne.def] at h, rw ← (nnreal.coe_eq_zero x) at h, exact h }, { exact ((continuous_subtype_val.comp continuous_fst).prod_mk continuous_snd).continuous_at } @@ -1545,7 +1546,7 @@ begin cases x, { exact dif_pos zero_lt_one }, { change ite _ _ _ = _, - simp, + simp only [nnreal.rpow_one, some_eq_coe, ite_eq_right_iff, top_ne_coe, and_imp], exact λ _, zero_le_one.not_lt } end @@ -1775,7 +1776,7 @@ lemma rpow_lt_rpow_of_exponent_gt {x : ℝ≥0∞} {y z : ℝ} (hx0 : 0 < x) (hx x^y < x^z := begin lift x to ℝ≥0 using ne_of_lt (lt_of_lt_of_le hx1 le_top), - simp at hx0 hx1, + simp only [coe_lt_one_iff, coe_pos] at hx0 hx1, simp [coe_rpow_of_ne_zero (ne_of_gt hx0), nnreal.rpow_lt_rpow_of_exponent_gt hx0 hx1 hyz] end @@ -1788,7 +1789,7 @@ begin rcases lt_trichotomy z 0 with Hz|Hz|Hz; simp [Hy, Hz, h, zero_rpow_of_neg, zero_rpow_of_pos, le_refl]; linarith }, - { simp at hx1, + { rw [coe_le_one_iff] at hx1, simp [coe_rpow_of_ne_zero h, nnreal.rpow_le_rpow_of_exponent_ge (bot_lt_iff_ne_bot.mpr h) hx1 hyz] } end @@ -1819,7 +1820,7 @@ begin cases lt_or_le 0 p with hp_pos hp_nonpos, { exact rpow_pos_of_nonneg hx_pos (le_of_lt hp_pos), }, { rw [←neg_neg p, rpow_neg, inv_pos], - exact rpow_ne_top_of_nonneg (by simp [hp_nonpos]) hx_ne_top, }, + exact rpow_ne_top_of_nonneg (right.nonneg_neg_iff.mpr hp_nonpos) hx_ne_top, }, end lemma rpow_lt_one {x : ℝ≥0∞} {z : ℝ} (hx : x < 1) (hz : 0 < z) : x^z < 1 := @@ -1978,7 +1979,7 @@ begin apply continuous_iff_continuous_at.2 (λ x, _), rcases lt_trichotomy 0 y with hy|rfl|hy, { exact continuous_at_rpow_const_of_pos hy }, - { simp, exact continuous_at_const }, + { simp only [rpow_zero], exact continuous_at_const }, { obtain ⟨z, hz⟩ : ∃ z, y = -z := ⟨-y, (neg_neg _).symm⟩, have z_pos : 0 < z, by simpa [hz] using hy, simp_rw [hz, rpow_neg], From e46daa0abb2a4cb80c796e9104afa778976beb1e Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 27 Sep 2022 00:15:04 +0000 Subject: [PATCH 404/828] feat(analysis/normed_space/operator_norm): remove to_span_singleton_norm (#16654) Remove `continuous_linear_map.to_span_singleton_norm`. It is a duplicate of `continuous_linear_map.norm_to_span_singleton` (which has more explicit arguments and weaker type-class assumptions). --- src/analysis/normed_space/operator_norm.lean | 3 --- 1 file changed, 3 deletions(-) diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index 016744c507242..7a7cbf0936dab 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -1258,9 +1258,6 @@ begin simpa only [hf, hx, mul_le_mul_right] using f.le_op_norm x, end -lemma to_span_singleton_norm (x : E) : ∥to_span_singleton 𝕜 x∥ = ∥x∥ := -homothety_norm _ (to_span_singleton_homothety 𝕜 x) - variable (f) theorem uniform_embedding_of_bound {K : ℝ≥0} (hf : ∀ x, ∥x∥ ≤ K * ∥f x∥) : From e2621d935895abe70071ab828a4ee6e26a52afe4 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Tue, 27 Sep 2022 07:32:45 +0000 Subject: [PATCH 405/828] chore(number_theory/legendre_symbol/norm_num): remove now unnecessary instance (#16659) This just removes a shortcut instance in `number_theory/legendre_symnol/norm_num` that was necessary to avoid computability trouble with the meta code, but is no longer so, since the root cause was fixed by #16463. --- src/number_theory/legendre_symbol/norm_num.lean | 3 --- 1 file changed, 3 deletions(-) diff --git a/src/number_theory/legendre_symbol/norm_num.lean b/src/number_theory/legendre_symbol/norm_num.lean index 4d77458655448..00cfe5eb881ff 100644 --- a/src/number_theory/legendre_symbol/norm_num.lean +++ b/src/number_theory/legendre_symbol/norm_num.lean @@ -231,9 +231,6 @@ end lemmas section evaluation --- The following is to prevent strange error messages from occurring. -instance : ring ℚ := division_ring.to_ring ℚ - /-! ### Certified evaluation of the Jacobi symbol From 3397560e65278e5f31acefcdea63138bd53d1cd4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Dupuis?= Date: Tue, 27 Sep 2022 09:35:29 +0000 Subject: [PATCH 406/828] chore(topology/algebra/module/basic): remove `continuous_linear_map.ker` and `continuous_linear_map.range` (#16208) This PR removes `continuous_linear_map.ker` and `continuous_linear_map.range`, which are now obsolete since `linear_map.ker` and `linear_map.range` are defined for any `linear_map_class` morphism. --- src/analysis/calculus/implicit.lean | 98 ++++++++++--------- src/analysis/calculus/inverse.lean | 2 +- .../calculus/lagrange_multipliers.lean | 13 +-- src/analysis/inner_product_space/basic.lean | 7 +- src/analysis/inner_product_space/dual.lean | 6 +- .../inner_product_space/lax_milgram.lean | 12 +-- .../inner_product_space/projection.lean | 15 +-- src/analysis/normed_space/banach.lean | 48 ++++----- src/analysis/normed_space/complemented.lean | 18 ++-- .../normed_space/finite_dimension.lean | 2 +- src/geometry/manifold/mfderiv.lean | 4 +- src/geometry/manifold/whitney_embedding.lean | 2 +- src/linear_algebra/basic.lean | 21 ++-- src/linear_algebra/prod.lean | 2 +- src/linear_algebra/projection.lean | 2 +- .../function/continuous_map_dense.lean | 6 +- src/measure_theory/function/lp_space.lean | 6 +- src/topology/algebra/module/basic.lean | 94 ++++++++---------- .../algebra/module/finite_dimension.lean | 4 +- 19 files changed, 188 insertions(+), 174 deletions(-) diff --git a/src/analysis/calculus/implicit.lean b/src/analysis/calculus/implicit.lean index 80704de48ffb0..2a50efb0c3174 100644 --- a/src/analysis/calculus/implicit.lean +++ b/src/analysis/calculus/implicit.lean @@ -49,6 +49,7 @@ open_locale topological_space open filter open continuous_linear_map (fst snd smul_right ker_prod) open continuous_linear_equiv (of_bijective) +open linear_map (ker range) /-! ### General version @@ -99,9 +100,9 @@ structure implicit_function_data (𝕜 : Type*) [nontrivially_normed_field 𝕜] (pt : E) (left_has_deriv : has_strict_fderiv_at left_fun left_deriv pt) (right_has_deriv : has_strict_fderiv_at right_fun right_deriv pt) -(left_range : left_deriv.range = ⊤) -(right_range : right_deriv.range = ⊤) -(is_compl_ker : is_compl left_deriv.ker right_deriv.ker) +(left_range : range left_deriv = ⊤) +(right_range : range right_deriv = ⊤) +(is_compl_ker : is_compl (ker left_deriv) (ker right_deriv)) namespace implicit_function_data @@ -215,8 +216,8 @@ variables (f f') /-- Data used to apply the generic implicit function theorem to the case of a strictly differentiable map such that its derivative is surjective and has a complemented kernel. -/ @[simp] def implicit_function_data_of_complemented (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : - implicit_function_data 𝕜 E F f'.ker := + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : + implicit_function_data 𝕜 E F (ker f') := { left_fun := f, left_deriv := f', right_fun := λ x, classical.some hker (x - a), @@ -232,58 +233,58 @@ differentiable map such that its derivative is surjective and has a complemented /-- A local homeomorphism between `E` and `F × f'.ker` sending level surfaces of `f` to vertical subspaces. -/ def implicit_to_local_homeomorph_of_complemented (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : - local_homeomorph E (F × f'.ker) := + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : + local_homeomorph E (F × (ker f')) := (implicit_function_data_of_complemented f f' hf hf' hker).to_local_homeomorph /-- Implicit function `g` defined by `f (g z y) = z`. -/ def implicit_function_of_complemented (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : - F → f'.ker → E := + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : + F → (ker f') → E := (implicit_function_data_of_complemented f f' hf hf' hker).implicit_function end defs @[simp] lemma implicit_to_local_homeomorph_of_complemented_fst (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) (x : E) : + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) (x : E) : (hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker x).fst = f x := rfl lemma implicit_to_local_homeomorph_of_complemented_apply - (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) - (hker : f'.ker.closed_complemented) (y : E) : + (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) + (hker : (ker f').closed_complemented) (y : E) : hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker y = (f y, classical.some hker (y - a)) := rfl @[simp] lemma implicit_to_local_homeomorph_of_complemented_apply_ker - (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) - (hker : f'.ker.closed_complemented) (y : f'.ker) : + (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) + (hker : (ker f').closed_complemented) (y : ker f') : hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker (y + a) = (f (y + a), y) := by simp only [implicit_to_local_homeomorph_of_complemented_apply, add_sub_cancel, classical.some_spec hker] @[simp] lemma implicit_to_local_homeomorph_of_complemented_self - (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : + (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker a = (f a, 0) := by simp [hf.implicit_to_local_homeomorph_of_complemented_apply] lemma mem_implicit_to_local_homeomorph_of_complemented_source (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : a ∈ (hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker).source := mem_to_local_homeomorph_source _ lemma mem_implicit_to_local_homeomorph_of_complemented_target (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : - (f a, (0 : f'.ker)) ∈ (hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker).target := + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : + (f a, (0 : ker f')) ∈ (hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker).target := by simpa only [implicit_to_local_homeomorph_of_complemented_self] using ((hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker).map_source $ (hf.mem_implicit_to_local_homeomorph_of_complemented_source hf' hker)) /-- `implicit_function_of_complemented` sends `(z, y)` to a point in `f ⁻¹' z`. -/ lemma map_implicit_function_of_complemented_eq (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : - ∀ᶠ (p : F × f'.ker) in 𝓝 (f a, 0), + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : + ∀ᶠ (p : F × (ker f')) in 𝓝 (f a, 0), f (hf.implicit_function_of_complemented f f' hf' hker p.1 p.2) = p.1 := ((hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker).eventually_right_inverse $ hf.mem_implicit_to_local_homeomorph_of_complemented_target hf' hker).mono $ λ ⟨z, y⟩ h, @@ -292,13 +293,13 @@ lemma map_implicit_function_of_complemented_eq (hf : has_strict_fderiv_at f f' a /-- Any point in some neighborhood of `a` can be represented as `implicit_function` of some point. -/ lemma eq_implicit_function_of_complemented (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : ∀ᶠ x in 𝓝 a, hf.implicit_function_of_complemented f f' hf' hker (f x) (hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker x).snd = x := (implicit_function_data_of_complemented f f' hf hf' hker).implicit_function_apply_image @[simp] lemma implicit_function_of_complemented_apply_image (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : hf.implicit_function_of_complemented f f' hf' hker (f a) 0 = a := begin convert (hf.implicit_to_local_homeomorph_of_complemented f f' hf' hker).left_inv @@ -307,12 +308,21 @@ begin end lemma to_implicit_function_of_complemented (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (hker : f'.ker.closed_complemented) : + (hf' : range f' = ⊤) (hker : (ker f').closed_complemented) : has_strict_fderiv_at (hf.implicit_function_of_complemented f f' hf' hker (f a)) - f'.ker.subtypeL 0 := -by convert (implicit_function_data_of_complemented f f' hf hf' - hker).implicit_function_has_strict_fderiv_at f'.ker.subtypeL _ _; - simp only [implicit_function_data_of_complemented]; ext; simp [classical.some_spec hker] + (ker f').subtypeL 0 := +begin + convert (implicit_function_data_of_complemented f f' hf hf' + hker).implicit_function_has_strict_fderiv_at (ker f').subtypeL _ _, + swap, + { ext, simp only [classical.some_spec hker, implicit_function_data_of_complemented, + continuous_linear_map.coe_comp', submodule.coe_subtypeL', submodule.coe_subtype, + function.comp_app, continuous_linear_map.coe_id', id.def] }, + swap, + { ext, simp only [continuous_linear_map.coe_comp', submodule.coe_subtypeL', submodule.coe_subtype, + function.comp_app, linear_map.map_coe_ker, continuous_linear_map.zero_apply] }, + simp only [implicit_function_data_of_complemented, map_sub, sub_self], +end end complemented @@ -342,46 +352,46 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜] /-- Given a map `f : E → F` to a finite dimensional space with a surjective derivative `f'`, returns a local homeomorphism between `E` and `F × ker f'`. -/ -def implicit_to_local_homeomorph (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) : - local_homeomorph E (F × f'.ker) := +def implicit_to_local_homeomorph (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) : + local_homeomorph E (F × (ker f')) := by haveI := finite_dimensional.complete 𝕜 F; exact hf.implicit_to_local_homeomorph_of_complemented f f' hf' f'.ker_closed_complemented_of_finite_dimensional_range /-- Implicit function `g` defined by `f (g z y) = z`. -/ -def implicit_function (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) : - F → f'.ker → E := +def implicit_function (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) : + F → (ker f') → E := function.curry $ (hf.implicit_to_local_homeomorph f f' hf').symm variables {f f'} @[simp] lemma implicit_to_local_homeomorph_fst (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) (x : E) : + (hf' : range f' = ⊤) (x : E) : (hf.implicit_to_local_homeomorph f f' hf' x).fst = f x := rfl @[simp] lemma implicit_to_local_homeomorph_apply_ker - (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) (y : f'.ker) : + (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) (y : ker f') : hf.implicit_to_local_homeomorph f f' hf' (y + a) = (f (y + a), y) := by apply implicit_to_local_homeomorph_of_complemented_apply_ker @[simp] lemma implicit_to_local_homeomorph_self - (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) : + (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) : hf.implicit_to_local_homeomorph f f' hf' a = (f a, 0) := by apply implicit_to_local_homeomorph_of_complemented_self lemma mem_implicit_to_local_homeomorph_source (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) : + (hf' : range f' = ⊤) : a ∈ (hf.implicit_to_local_homeomorph f f' hf').source := mem_to_local_homeomorph_source _ lemma mem_implicit_to_local_homeomorph_target (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) : - (f a, (0 : f'.ker)) ∈ (hf.implicit_to_local_homeomorph f f' hf').target := + (hf' : range f' = ⊤) : + (f a, (0 : ker f')) ∈ (hf.implicit_to_local_homeomorph f f' hf').target := by apply mem_implicit_to_local_homeomorph_of_complemented_target lemma tendsto_implicit_function (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) {α : Type*} {l : filter α} {g₁ : α → F} {g₂ : α → f'.ker} + (hf' : range f' = ⊤) {α : Type*} {l : filter α} {g₁ : α → F} {g₂ : α → ker f'} (h₁ : tendsto g₁ l (𝓝 $ f a)) (h₂ : tendsto g₂ l (𝓝 0)) : tendsto (λ t, hf.implicit_function f f' hf' (g₁ t) (g₂ t)) l (𝓝 a) := begin @@ -394,25 +404,25 @@ end alias tendsto_implicit_function ← _root_.filter.tendsto.implicit_function /-- `implicit_function` sends `(z, y)` to a point in `f ⁻¹' z`. -/ -lemma map_implicit_function_eq (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) : - ∀ᶠ (p : F × f'.ker) in 𝓝 (f a, 0), f (hf.implicit_function f f' hf' p.1 p.2) = p.1 := +lemma map_implicit_function_eq (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) : + ∀ᶠ (p : F × (ker f')) in 𝓝 (f a, 0), f (hf.implicit_function f f' hf' p.1 p.2) = p.1 := by apply map_implicit_function_of_complemented_eq @[simp] lemma implicit_function_apply_image (hf : has_strict_fderiv_at f f' a) - (hf' : f'.range = ⊤) : + (hf' : range f' = ⊤) : hf.implicit_function f f' hf' (f a) 0 = a := by apply implicit_function_of_complemented_apply_image /-- Any point in some neighborhood of `a` can be represented as `implicit_function` of some point. -/ -lemma eq_implicit_function (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) : +lemma eq_implicit_function (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) : ∀ᶠ x in 𝓝 a, hf.implicit_function f f' hf' (f x) (hf.implicit_to_local_homeomorph f f' hf' x).snd = x := by apply eq_implicit_function_of_complemented -lemma to_implicit_function (hf : has_strict_fderiv_at f f' a) (hf' : f'.range = ⊤) : +lemma to_implicit_function (hf : has_strict_fderiv_at f f' a) (hf' : range f' = ⊤) : has_strict_fderiv_at (hf.implicit_function f f' hf' (f a)) - f'.ker.subtypeL 0 := + (ker f').subtypeL 0 := by apply to_implicit_function_of_complemented end finite_dimensional diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index 0d145396471ed..b9ea5c303d7f6 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -555,7 +555,7 @@ end lemma map_nhds_eq_of_surj [complete_space E] [complete_space F] {f : E → F} {f' : E →L[𝕜] F} {a : E} - (hf : has_strict_fderiv_at f (f' : E →L[𝕜] F) a) (h : f'.range = ⊤) : + (hf : has_strict_fderiv_at f (f' : E →L[𝕜] F) a) (h : linear_map.range f' = ⊤) : map f (𝓝 a) = 𝓝 (f a) := begin let f'symm := f'.nonlinear_right_inverse_of_surjective h, diff --git a/src/analysis/calculus/lagrange_multipliers.lean b/src/analysis/calculus/lagrange_multipliers.lean index 0e001fa1211cb..6d3f314434da8 100644 --- a/src/analysis/calculus/lagrange_multipliers.lean +++ b/src/analysis/calculus/lagrange_multipliers.lean @@ -38,7 +38,7 @@ a complete space, then the linear map `x ↦ (f' x, φ' x)` is not surjective. - lemma is_local_extr_on.range_ne_top_of_has_strict_fderiv_at (hextr : is_local_extr_on φ {x | f x = f x₀} x₀) (hf' : has_strict_fderiv_at f f' x₀) (hφ' : has_strict_fderiv_at φ φ' x₀) : - (f'.prod φ').range ≠ ⊤ := + linear_map.range (f'.prod φ') ≠ ⊤ := begin intro htop, set fφ := λ x, (f x, φ x), @@ -68,11 +68,12 @@ begin refine ⟨Λ, Λ₀, e.map_ne_zero_iff.1 h0, λ x, _⟩, convert linear_map.congr_fun (linear_map.range_le_ker_iff.1 hΛ') x using 1, -- squeezed `simp [mul_comm]` to speed up elaboration - simp only [linear_map.coprod_equiv_apply, linear_equiv.refl_apply, - linear_map.ring_lmap_equiv_self_symm_apply, linear_map.comp_apply, - continuous_linear_map.coe_coe, continuous_linear_map.prod_apply, - linear_equiv.trans_apply, linear_equiv.prod_apply, linear_map.coprod_apply, - linear_map.smul_right_apply, linear_map.one_apply, smul_eq_mul, mul_comm] + simp only [mul_comm, algebra.id.smul_eq_mul, linear_equiv.trans_apply, linear_equiv.prod_apply, + linear_equiv.refl_apply, linear_map.ring_lmap_equiv_self_symm_apply, + linear_map.coprod_equiv_apply, continuous_linear_map.to_linear_map_eq_coe, + continuous_linear_map.coe_prod, linear_map.coprod_comp_prod, linear_map.add_apply, + linear_map.coe_comp, continuous_linear_map.coe_coe, linear_map.coe_smul_right, + linear_map.one_apply] end /-- Lagrange multipliers theorem: if `φ : E → ℝ` has a local extremum on the set `{x | f x = f x₀}` diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 7a89331e777c6..b0375bc8dbd12 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -2245,7 +2245,7 @@ by simp [disjoint_iff, K.inf_orthogonal_eq_bot] /-- `Kᗮ` can be characterized as the intersection of the kernels of the operations of inner product with each of the elements of `K`. -/ -lemma orthogonal_eq_inter : Kᗮ = ⨅ v : K, (innerSL (v:E)).ker := +lemma orthogonal_eq_inter : Kᗮ = ⨅ v : K, linear_map.ker (innerSL (v:E) : E →L[𝕜] 𝕜) := begin apply le_antisymm, { rw le_infi_iff, @@ -2260,8 +2260,9 @@ end lemma submodule.is_closed_orthogonal : is_closed (Kᗮ : set E) := begin rw orthogonal_eq_inter K, - convert is_closed_Inter (λ v : K, (innerSL (v:E)).is_closed_ker), - simp + have := λ v : K, continuous_linear_map.is_closed_ker (innerSL (v:E) : E →L[𝕜] 𝕜), + convert is_closed_Inter this, + simp only [submodule.infi_coe], end /-- In a complete space, the orthogonal complement of any submodule `K` is complete. -/ diff --git a/src/analysis/inner_product_space/dual.lean b/src/analysis/inner_product_space/dual.lean index fbb0a733d6134..0771398a6366d 100644 --- a/src/analysis/inner_product_space/dual.lean +++ b/src/analysis/inner_product_space/dual.lean @@ -99,7 +99,7 @@ def to_dual : E ≃ₗᵢ⋆[𝕜] normed_space.dual 𝕜 E := linear_isometry_equiv.of_surjective (to_dual_map 𝕜 E) begin intros ℓ, - set Y := ker ℓ with hY, + set Y := linear_map.ker ℓ with hY, by_cases htriv : Y = ⊤, { have hℓ : ℓ = 0, { have h' := linear_map.ker_eq_top.mp htriv, @@ -114,8 +114,8 @@ begin refine ⟨((ℓ z)† / ⟪z, z⟫) • z, _⟩, ext x, have h₁ : (ℓ z) • x - (ℓ x) • z ∈ Y, - { rw [mem_ker, map_sub, continuous_linear_map.map_smul, continuous_linear_map.map_smul, - algebra.id.smul_eq_mul, algebra.id.smul_eq_mul, mul_comm], + { rw [linear_map.mem_ker, map_sub, continuous_linear_map.map_smul, + continuous_linear_map.map_smul, algebra.id.smul_eq_mul, algebra.id.smul_eq_mul, mul_comm], exact sub_self (ℓ x * ℓ z) }, have h₂ : (ℓ z) * ⟪z, x⟫ = (ℓ x) * ⟪z, z⟫, { have h₃ := calc diff --git a/src/analysis/inner_product_space/lax_milgram.lean b/src/analysis/inner_product_space/lax_milgram.lean index 9dff892cda211..ef9a303e8bb9b 100644 --- a/src/analysis/inner_product_space/lax_milgram.lean +++ b/src/analysis/inner_product_space/lax_milgram.lean @@ -34,7 +34,7 @@ dual, Lax-Milgram -/ noncomputable theory -open is_R_or_C linear_map continuous_linear_map inner_product_space +open is_R_or_C linear_map continuous_linear_map inner_product_space linear_map (ker range) open_locale real_inner_product_space nnreal universe u @@ -73,23 +73,23 @@ begin simpa using below_bound, end -lemma ker_eq_bot (coercive : is_coercive B) : B♯.ker = ⊥ := +lemma ker_eq_bot (coercive : is_coercive B) : ker B♯ = ⊥ := begin - rw [←ker_coe, linear_map.ker_eq_bot], + rw [linear_map_class.ker_eq_bot], rcases coercive.antilipschitz with ⟨_, _, antilipschitz⟩, exact antilipschitz.injective, end -lemma closed_range (coercive : is_coercive B) : is_closed (B♯.range : set V) := +lemma closed_range (coercive : is_coercive B) : is_closed (range B♯ : set V) := begin rcases coercive.antilipschitz with ⟨_, _, antilipschitz⟩, exact antilipschitz.is_closed_range B♯.uniform_continuous, end -lemma range_eq_top (coercive : is_coercive B) : B♯.range = ⊤ := +lemma range_eq_top (coercive : is_coercive B) : range B♯ = ⊤ := begin haveI := coercive.closed_range.complete_space_coe, - rw ← B♯.range.orthogonal_orthogonal, + rw ← (range B♯).orthogonal_orthogonal, rw submodule.eq_top_iff', intros v w mem_w_orthogonal, rcases coercive with ⟨C, C_pos, coercivity⟩, diff --git a/src/analysis/inner_product_space/projection.lean b/src/analysis/inner_product_space/projection.lean index a061b650cbc0d..4b1a9b17adf17 100644 --- a/src/analysis/inner_product_space/projection.lean +++ b/src/analysis/inner_product_space/projection.lean @@ -38,7 +38,7 @@ The Coq code is available at the following address: Date: Tue, 27 Sep 2022 11:52:59 +0000 Subject: [PATCH 407/828] feat(combinatorics/quiver/path): Turn a quiver path into a list (#15240) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `quiver.path.to_list : path a b → list V`. --- src/combinatorics/quiver/path.lean | 40 +++++++++++++++++++++++++++++- 1 file changed, 39 insertions(+), 1 deletion(-) diff --git a/src/combinatorics/quiver/path.lean b/src/combinatorics/quiver/path.lean index 430bbf2e3a708..ed87f1cafe5b7 100644 --- a/src/combinatorics/quiver/path.lean +++ b/src/combinatorics/quiver/path.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Wärn, Scott Morrison -/ import combinatorics.quiver.basic +import data.list.basic /-! # Paths in quivers @@ -12,6 +13,8 @@ Given a quiver `V`, we define the type of paths from `a : V` to `b : V` as an in family. We define composition of paths and the action of prefunctors on paths. -/ +open function + universes v v₁ v₂ u u₁ u₂ namespace quiver @@ -27,7 +30,7 @@ path.nil.cons e namespace path -variables {V : Type u} [quiver V] +variables {V : Type u} [quiver V] {a b : V} /-- The length of a path is the number of arrows it uses. -/ def length {a : V} : Π {b : V}, path a b → ℕ @@ -42,6 +45,9 @@ instance {a : V} : inhabited (path a a) := ⟨path.nil⟩ @[simp] lemma length_cons (a b c : V) (p : path a b) (e : b ⟶ c) : (p.cons e).length = p.length + 1 := rfl +lemma eq_of_length_zero (p : path a b) (hzero : p.length = 0) : a = b := +by { cases p, { refl }, { cases nat.succ_ne_zero _ hzero } } + /-- Composition of paths. -/ def comp {a b : V} : Π {c}, path a b → path b c → path a c | _ p (path.nil) := p @@ -59,6 +65,38 @@ def comp {a b : V} : Π {c}, path a b → path b c → path a c | c p q path.nil := rfl | d p q (path.cons r e) := by rw [comp_cons, comp_cons, comp_cons, comp_assoc] +/-- Turn a path into a list. The list contains `a` at its head, but not `b` a priori. -/ +@[simp] def to_list : Π {b : V}, path a b → list V +| b nil := [] +| b (@cons _ _ _ c _ p f) := c :: p.to_list + +/-- `quiver.path.to_list` is a contravariant functor. The inversion comes from `quiver.path` and +`list` having different preferred directions for adding elements. -/ +@[simp] lemma to_list_comp (p : path a b) : + ∀ {c} (q : path b c), (p.comp q).to_list = q.to_list ++ p.to_list +| c nil := by simp +| c (@cons _ _ _ d _ q f) := by simp [to_list_comp] + +lemma to_list_chain_nonempty : + ∀ {b} (p : path a b), p.to_list.chain (λ x y, nonempty (y ⟶ x)) b +| b nil := list.chain.nil +| b (cons p f) := p.to_list_chain_nonempty.cons ⟨f⟩ + +variables [∀ a b : V, subsingleton (a ⟶ b)] + +lemma to_list_injective (a : V) : ∀ b, injective (to_list : path a b → list V) +| b nil nil h := rfl +| b nil (@cons _ _ _ c _ p f) h := (list.cons_ne_nil _ _ h.symm).elim +| b (@cons _ _ _ c _ p f) nil h := (list.cons_ne_nil _ _ h).elim +| b (@cons _ _ _ c _ p f) (@cons _ _ s t u C D) h := begin + simp only [to_list] at h, + obtain ⟨rfl, hAC⟩ := h, + simp [to_list_injective _ hAC], +end + +@[simp] lemma to_list_inj {p q : path a b} : p.to_list = q.to_list ↔ p = q := +(to_list_injective _ _).eq_iff + end path end quiver From ac49df8144e602e3eb5c0e5eaa042ddb78f6a263 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Tue, 27 Sep 2022 11:53:00 +0000 Subject: [PATCH 408/828] =?UTF-8?q?feat(group=5Ftheory/exponent):=20`card?= =?UTF-8?q?=20G=20=E2=88=A3=20exponent=20G=20^=20rank=20G`=20(#16354)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds a lemma stating that `nat.card G ∣ monoid.exponent G ^ group.rank G`. --- src/group_theory/exponent.lean | 31 +++++++++++++++++++++++++++++++ 1 file changed, 31 insertions(+) diff --git a/src/group_theory/exponent.lean b/src/group_theory/exponent.lean index b7a8183ec2526..02da4ee73ae32 100644 --- a/src/group_theory/exponent.lean +++ b/src/group_theory/exponent.lean @@ -3,6 +3,8 @@ Copyright (c) 2021 Julian Kuelshammer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Julian Kuelshammer -/ +import data.zmod.quotient +import group_theory.noncomm_pi_coprod import group_theory.order_of_element import algebra.gcd_monoid.finset import algebra.punit_instances @@ -313,3 +315,32 @@ end end cancel_comm_monoid end monoid + +section comm_group + +open subgroup +open_locale big_operators + +variables (G) [comm_group G] [group.fg G] + +@[to_additive] lemma card_dvd_exponent_pow_rank : nat.card G ∣ monoid.exponent G ^ group.rank G := +begin + obtain ⟨S, hS1, hS2⟩ := group.rank_spec G, + rw [←hS1, ←fintype.card_coe, ←finset.card_univ, ←finset.prod_const], + let f : (Π g : S, zpowers (g : G)) →* G := noncomm_pi_coprod (λ s t h x y hx hy, mul_comm x y), + have hf : function.surjective f, + { rw [←monoid_hom.range_top_iff_surjective, eq_top_iff, ←hS2, closure_le], + exact λ g hg, ⟨pi.mul_single ⟨g, hg⟩ ⟨g, mem_zpowers g⟩, noncomm_pi_coprod_mul_single _ _⟩ }, + replace hf := nat_card_dvd_of_surjective f hf, + rw nat.card_pi at hf, + refine hf.trans (finset.prod_dvd_prod_of_dvd _ _ (λ g hg, _)), + rw ← order_eq_card_zpowers', + exact monoid.order_dvd_exponent (g : G), +end + +@[to_additive] lemma card_dvd_exponent_pow_rank' {n : ℕ} (hG : ∀ g : G, g ^ n = 1) : + nat.card G ∣ n ^ group.rank G := +(card_dvd_exponent_pow_rank G).trans + (pow_dvd_pow_of_dvd (monoid.exponent_dvd_of_forall_pow_eq_one G n hG) (group.rank G)) + +end comm_group From 72944168309a1ba69a94c18b9f09ec9b578bb3ec Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 27 Sep 2022 11:53:01 +0000 Subject: [PATCH 409/828] feat(topology/instances/real): classify discrete subgroups (#16592) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The subgroups aℤ (i.e. `zmultiples a`) of ℝ are discrete, in the sense of having finite intersection with any compact subset. Co-authored-by: Alex Kontorovich <58564076+AlexKontorovich@users.noreply.github.com> --- src/group_theory/subgroup/basic.lean | 9 ++++++ src/topology/algebra/field.lean | 16 +++++++++- src/topology/algebra/monoid.lean | 22 +++++++++++++ src/topology/instances/real.lean | 46 ++++++++++++++++++++++++++++ src/topology/metric_space/basic.lean | 14 --------- 5 files changed, 92 insertions(+), 15 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 30b944cd279d4..8ae72589440c1 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2035,6 +2035,15 @@ cod_restrict f _ $ λ x, ⟨x, rfl⟩ @[simp, to_additive] lemma coe_range_restrict (f : G →* N) (g : G) : (f.range_restrict g : N) = f g := rfl +@[to_additive] +lemma coe_comp_range_restrict (f : G →* N) : + (coe : f.range → N) ∘ (⇑(f.range_restrict) : G → f.range) = f := +rfl + +@[to_additive] +lemma subtype_comp_range_restrict (f : G →* N) : f.range.subtype.comp (f.range_restrict) = f := +ext $ f.coe_range_restrict + @[to_additive] lemma range_restrict_surjective (f : G →* N) : function.surjective f.range_restrict := λ ⟨_, g, rfl⟩, ⟨g, rfl⟩ diff --git a/src/topology/algebra/field.lean b/src/topology/algebra/field.lean index f394d0115f4f0..98745ad7625cc 100644 --- a/src/topology/algebra/field.lean +++ b/src/topology/algebra/field.lean @@ -63,7 +63,21 @@ instance top_monoid_units [topological_semiring R] [induced_units R] : end⟩ end topological_ring -variables (K : Type*) [division_ring K] [topological_space K] +variables {K : Type*} [division_ring K] [topological_space K] + +/-- Left-multiplication by a nonzero element of a topological division ring is proper, i.e., +inverse images of compact sets are compact. -/ +lemma filter.tendsto_cocompact_mul_left₀ [has_continuous_mul K] {a : K} (ha : a ≠ 0) : + filter.tendsto (λ x : K, a * x) (filter.cocompact K) (filter.cocompact K) := +filter.tendsto_cocompact_mul_left (inv_mul_cancel ha) + +/-- Right-multiplication by a nonzero element of a topological division ring is proper, i.e., +inverse images of compact sets are compact. -/ +lemma filter.tendsto_cocompact_mul_right₀ [has_continuous_mul K] {a : K} (ha : a ≠ 0) : + filter.tendsto (λ x : K, x * a) (filter.cocompact K) (filter.cocompact K) := +filter.tendsto_cocompact_mul_right (mul_inv_cancel ha) + +variables (K) /-- A topological division ring is a division ring with a topology where all operations are continuous, including inversion. -/ diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 5886701a9a683..55b2fc8b7bcba 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -417,6 +417,28 @@ lemma continuous_on.pow {f : X → M} {s : set X} (hf : continuous_on f s) (n : continuous_on (λ x, f x ^ n) s := λ x hx, (hf x hx).pow n +/-- Left-multiplication by a left-invertible element of a topological monoid is proper, i.e., +inverse images of compact sets are compact. -/ +lemma filter.tendsto_cocompact_mul_left {a b : M} (ha : b * a = 1) : + filter.tendsto (λ x : M, a * x) (filter.cocompact M) (filter.cocompact M) := +begin + refine filter.tendsto.of_tendsto_comp _ (filter.comap_cocompact_le (continuous_mul_left b)), + convert filter.tendsto_id, + ext x, + simp [ha], +end + +/-- Right-multiplication by a right-invertible element of a topological monoid is proper, i.e., +inverse images of compact sets are compact. -/ +lemma filter.tendsto_cocompact_mul_right {a b : M} (ha : a * b = 1) : + filter.tendsto (λ x : M, x * a) (filter.cocompact M) (filter.cocompact M) := +begin + refine filter.tendsto.of_tendsto_comp _ (filter.comap_cocompact_le (continuous_mul_right b)), + convert filter.tendsto_id, + ext x, + simp [ha], +end + /-- If `R` acts on `A` via `A`, then continuous multiplication implies continuous scalar multiplication by constants. diff --git a/src/topology/instances/real.lean b/src/topology/instances/real.lean index 226dbffb45a0f..74e29016255ff 100644 --- a/src/topology/instances/real.lean +++ b/src/topology/instances/real.lean @@ -216,6 +216,52 @@ end periodic section subgroups +namespace int +open metric + +/-- Under the coercion from `ℤ` to `ℝ`, inverse images of compact sets are finite. -/ +lemma tendsto_coe_cofinite : tendsto (coe : ℤ → ℝ) cofinite (cocompact ℝ) := +begin + refine tendsto_cocompact_of_tendsto_dist_comp_at_top (0 : ℝ) _, + simp only [filter.tendsto_at_top, eventually_cofinite, not_le, ← mem_ball], + change ∀ r : ℝ, (coe ⁻¹' (ball (0 : ℝ) r)).finite, + simp [real.ball_eq_Ioo, set.finite_Ioo], +end + +/-- For nonzero `a`, the "multiples of `a`" map `zmultiples_hom` from `ℤ` to `ℝ` is discrete, i.e. +inverse images of compact sets are finite. -/ +lemma tendsto_zmultiples_hom_cofinite {a : ℝ} (ha : a ≠ 0) : + tendsto (zmultiples_hom ℝ a) cofinite (cocompact ℝ) := +begin + convert (tendsto_cocompact_mul_right₀ ha).comp int.tendsto_coe_cofinite, + ext n, + simp, +end + +end int + +namespace add_subgroup + +/-- The subgroup "multiples of `a`" (`zmultiples a`) is a discrete subgroup of `ℝ`, i.e. its +intersection with compact sets is finite. -/ +lemma tendsto_zmultiples_subtype_cofinite (a : ℝ) : + tendsto (zmultiples a).subtype cofinite (cocompact ℝ) := +begin + rcases eq_or_ne a 0 with rfl | ha, + { rw add_subgroup.zmultiples_zero_eq_bot, + intros K hK, + rw [filter.mem_map, mem_cofinite], + apply set.to_finite }, + intros K hK, + have H := int.tendsto_zmultiples_hom_cofinite ha hK, + simp only [filter.mem_map, mem_cofinite, ← preimage_compl] at ⊢ H, + rw [← (zmultiples_hom ℝ a).range_restrict_surjective.image_preimage + ((zmultiples a).subtype ⁻¹' Kᶜ), ← preimage_comp, ← add_monoid_hom.coe_comp_range_restrict], + exact finite.image _ H, +end + +end add_subgroup + /-- Given a nontrivial subgroup `G ⊆ ℝ`, if `G ∩ ℝ_{>0}` has no minimum then `G` is dense. -/ lemma real.subgroup_dense_of_no_min {G : add_subgroup ℝ} {g₀ : ℝ} (g₀_in : g₀ ∈ G) (g₀_ne : g₀ ≠ 0) (H' : ¬ ∃ a : ℝ, is_least {g : ℝ | g ∈ G ∧ 0 < g} a) : diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 3baf72b5accb9..d480956f88b2a 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -2489,20 +2489,6 @@ lemma tendsto_cocompact_of_tendsto_dist_comp_at_top {f : β → α} {l : filter (h : tendsto (λ y, dist (f y) x) l at_top) : tendsto f l (cocompact α) := by { refine tendsto.mono_right _ (comap_dist_right_at_top_le_cocompact x), rwa tendsto_comap_iff } -namespace int -open metric - -/-- Under the coercion from `ℤ` to `ℝ`, inverse images of compact sets are finite. -/ -lemma tendsto_coe_cofinite : tendsto (coe : ℤ → ℝ) cofinite (cocompact ℝ) := -begin - refine tendsto_cocompact_of_tendsto_dist_comp_at_top (0 : ℝ) _, - simp only [filter.tendsto_at_top, eventually_cofinite, not_le, ← mem_ball], - change ∀ r : ℝ, (coe ⁻¹' (ball (0 : ℝ) r)).finite, - simp [real.ball_eq_Ioo, set.finite_Ioo], -end - -end int - /-- We now define `metric_space`, extending `pseudo_metric_space`. -/ class metric_space (α : Type u) extends pseudo_metric_space α : Type u := (eq_of_dist_eq_zero : ∀ {x y : α}, dist x y = 0 → x = y) From f48c65aa43ca315b6aad18ec2915896064df6d23 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 27 Sep 2022 11:53:03 +0000 Subject: [PATCH 410/828] feat(analysis/convex/topology): connectedness and `same_ray` (#16661) Add lemmas that the set of vectors in the same ray as a given vector, and the set of nonzero vectors in the same ray as a given nonzero vector, are connected (in a real normed space). --- src/analysis/convex/topology.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/analysis/convex/topology.lean b/src/analysis/convex/topology.lean index 1efbd7ec8b029..c1be6f7735b23 100644 --- a/src/analysis/convex/topology.lean +++ b/src/analysis/convex/topology.lean @@ -404,4 +404,20 @@ begin simpa only [sub_add_sub_cancel', norm_sub_rev] using h.norm_add.symm end +/-- The set of vectors in the same ray as `x` is connected. -/ +lemma is_connected_set_of_same_ray (x : E) : is_connected {y | same_ray ℝ x y} := +begin + by_cases hx : x = 0, { simpa [hx] using is_connected_univ }, + simp_rw ←exists_nonneg_left_iff_same_ray hx, + exact is_connected_Ici.image _ ((continuous_id.smul continuous_const).continuous_on) +end + +/-- The set of nonzero vectors in the same ray as the nonzero vector `x` is connected. -/ +lemma is_connected_set_of_same_ray_and_ne_zero {x : E} (hx : x ≠ 0) : + is_connected {y | same_ray ℝ x y ∧ y ≠ 0} := +begin + simp_rw ←exists_pos_left_iff_same_ray_and_ne_zero hx, + exact is_connected_Ioi.image _ ((continuous_id.smul continuous_const).continuous_on) +end + end normed_space From 801c013fed96f0d3d9c15e23fe3ddad4481960a8 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 27 Sep 2022 14:28:32 +0000 Subject: [PATCH 411/828] refactor(tactic/lift): remove attribute/simp-set and swap side goal (#16565) This PR accomplishes two things: * It simplifies the design of the `lift` tactic (while keeping roughly the same user interface, see below). Specifically, it does not need an attribute set to simplify expressions coming from `can_lift` instances. - This will make the tactic easier to port to Lean 4. - It accomplishes this by moving two fields of `can_lift` into `out_param`s. So writing the instances is slightly different from before. * If the `using h` clause is left out, then `lift` produces a subgoal. This subgoal used to come after the main goal. But I think it is natural to make it the first goal. So that you can write ``` lift n to nat with k, { linarith }, ``` instead of ``` lift n to nat using by { linarith } with k, ``` --- src/algebra/add_torsor.lean | 2 +- src/algebra/group/units.lean | 6 +- src/algebra/group/with_one.lean | 6 +- .../complex/upper_half_plane/basic.lean | 2 +- src/data/complex/basic.lean | 6 +- src/data/finset/basic.lean | 22 ++-- src/data/finsupp/defs.lean | 6 +- src/data/list/basic.lean | 7 +- src/data/multiset/basic.lean | 7 +- src/data/nat/enat.lean | 2 +- src/data/pnat/basic.lean | 8 +- src/data/rat/defs.lean | 4 +- src/data/rat/nnrat.lean | 6 +- src/data/real/ennreal.lean | 6 +- src/data/real/ereal.lean | 12 +- src/data/real/nnreal.lean | 2 +- src/data/set/basic.lean | 14 +-- src/data/set/finite.lean | 6 +- src/linear_algebra/quadratic_form/basic.lean | 7 +- src/logic/embedding.lean | 6 +- src/logic/equiv/basic.lean | 6 +- src/order/bounded_order.lean | 12 +- src/order/filter/basic.lean | 7 +- src/order/hom/basic.lean | 6 +- src/order/interval.lean | 3 +- src/set_theory/cardinal/basic.lean | 8 +- src/tactic/lift.lean | 103 ++++++++---------- src/topology/alexandroff.lean | 6 +- src/topology/continuous_function/units.lean | 7 +- src/topology/instances/nnreal.lean | 7 +- src/topology/sets/compacts.lean | 6 +- test/lift.lean | 37 +++---- 32 files changed, 138 insertions(+), 207 deletions(-) diff --git a/src/algebra/add_torsor.lean b/src/algebra/add_torsor.lean index 8e7e9a2fe9982..36ec9cde7bece 100644 --- a/src/algebra/add_torsor.lean +++ b/src/algebra/add_torsor.lean @@ -335,7 +335,7 @@ def const_vadd_hom : multiplicative G →* equiv.perm P := variable {P} -open function +open _root_.function /-- Point reflection in `x` as a permutation. -/ def point_reflection (x : P) : perm P := (const_vsub x).trans (vadd_const x) diff --git a/src/algebra/group/units.lean b/src/algebra/group/units.lean index 23f4851d08a2b..bf35b52e17926 100644 --- a/src/algebra/group/units.lean +++ b/src/algebra/group/units.lean @@ -353,10 +353,8 @@ lemma is_unit_of_subsingleton [monoid M] [subsingleton M] (a : M) : is_unit a := attribute [nontriviality] is_add_unit_of_subsingleton -@[to_additive] instance [monoid M] : can_lift M Mˣ := -{ coe := coe, - cond := is_unit, - prf := λ _, id } +@[to_additive] instance [monoid M] : can_lift M Mˣ coe is_unit := +{ prf := λ _, id } @[to_additive] instance [monoid M] [subsingleton M] : unique Mˣ := { default := 1, diff --git a/src/algebra/group/with_one.lean b/src/algebra/group/with_one.lean index 7a9feae33d333..c568d8b9acd00 100644 --- a/src/algebra/group/with_one.lean +++ b/src/algebra/group/with_one.lean @@ -94,10 +94,8 @@ lemma ne_one_iff_exists {x : with_one α} : x ≠ 1 ↔ ∃ (a : α), ↑a = x : option.ne_none_iff_exists @[to_additive] -instance : can_lift (with_one α) α := -{ coe := coe, - cond := λ a, a ≠ 1, - prf := λ a, ne_one_iff_exists.1 } +instance can_lift : can_lift (with_one α) α coe (λ a, a ≠ 1) := +{ prf := λ a, ne_one_iff_exists.1 } @[simp, norm_cast, to_additive] lemma coe_inj {a b : α} : (a : with_one α) = b ↔ a = b := diff --git a/src/analysis/complex/upper_half_plane/basic.lean b/src/analysis/complex/upper_half_plane/basic.lean index b9451fd70a07c..d21fd1cad8e23 100644 --- a/src/analysis/complex/upper_half_plane/basic.lean +++ b/src/analysis/complex/upper_half_plane/basic.lean @@ -48,7 +48,7 @@ namespace upper_half_plane instance : inhabited ℍ := ⟨⟨complex.I, by simp⟩⟩ -instance : can_lift ℂ ℍ := subtype.can_lift (λ z, 0 < z.im) +instance can_lift : can_lift ℂ ℍ coe (λ z, 0 < z.im) := subtype.can_lift (λ z, 0 < z.im) /-- Imaginary part -/ def im (z : ℍ) := (z : ℂ).im diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index d648b26bb43f3..a6ee5d7cbc964 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -66,10 +66,8 @@ lemma of_real_def (r : ℝ) : (r : ℂ) = ⟨r, 0⟩ := rfl theorem of_real_injective : function.injective (coe : ℝ → ℂ) := λ z w, congr_arg re -instance : can_lift ℂ ℝ := -{ cond := λ z, z.im = 0, - coe := coe, - prf := λ z hz, ⟨z.re, ext rfl hz.symm⟩ } +instance can_lift : can_lift ℂ ℝ coe (λ z, z.im = 0) := +{ prf := λ z hz, ⟨z.re, ext rfl hz.symm⟩ } /-- The product of a set on the real axis and a set on the imaginary axis of the complex plane, denoted by `s ×ℂ t`. -/ diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index d38b2d4313cfd..3f1252d41f0ba 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -212,18 +212,15 @@ instance {α : Type u} : has_coe_to_sort (finset α) (Type u) := ⟨λ s, {x // instance pi_finset_coe.can_lift (ι : Type*) (α : Π i : ι, Type*) [ne : Π i, nonempty (α i)] (s : finset ι) : -can_lift (Π i : s, α i) (Π i, α i) := -{ coe := λ f i, f i, - .. pi_subtype.can_lift ι α (∈ s) } + can_lift (Π i : s, α i) (Π i, α i) (λ f i, f i) (λ _, true) := +pi_subtype.can_lift ι α (∈ s) instance pi_finset_coe.can_lift' (ι α : Type*) [ne : nonempty α] (s : finset ι) : - can_lift (s → α) (ι → α) := + can_lift (s → α) (ι → α) (λ f i, f i) (λ _, true) := pi_finset_coe.can_lift ι (λ _, α) s -instance finset_coe.can_lift (s : finset α) : can_lift α s := -{ coe := coe, - cond := λ a, a ∈ s, - prf := λ a ha, ⟨⟨a, ha⟩, rfl⟩ } +instance finset_coe.can_lift (s : finset α) : can_lift α s coe (λ a, a ∈ s) := +{ prf := λ a ha, ⟨⟨a, ha⟩, rfl⟩ } @[simp, norm_cast] lemma coe_sort_coe (s : finset α) : ((s : set α) : Sort*) = s := rfl @@ -2269,10 +2266,9 @@ by { rw mem_image, simp only [exists_prop, const_apply, exists_and_distrib_right lemma mem_image_const_self : b ∈ s.image (const α b) ↔ s.nonempty := mem_image_const.trans $ and_iff_left rfl -instance [can_lift β α] : can_lift (finset β) (finset α) := -{ cond := λ s, ∀ x ∈ s, can_lift.cond α x, - coe := image can_lift.coe, - prf := +instance can_lift (c) (p) [can_lift β α c p] : + can_lift (finset β) (finset α) (image c) (λ s, ∀ x ∈ s, p x) := +{ prf := begin rintro ⟨⟨l⟩, hd : l.nodup⟩ hl, lift l to list α using hl, @@ -2562,7 +2558,7 @@ begin split, swap, { rintro ⟨t, ht, rfl⟩, rw [coe_image], exact set.image_subset f ht }, intro h, - letI : can_lift β s := ⟨f ∘ coe, λ y, y ∈ f '' s, λ y ⟨x, hxt, hy⟩, ⟨⟨x, hxt⟩, hy⟩⟩, + letI : can_lift β s (f ∘ coe) (λ y, y ∈ f '' s) := ⟨λ y ⟨x, hxt, hy⟩, ⟨⟨x, hxt⟩, hy⟩⟩, lift t to finset s using h, refine ⟨t.map (embedding.subtype _), map_subtype_subset _, _⟩, ext y, simp diff --git a/src/data/finsupp/defs.lean b/src/data/finsupp/defs.lean index 2513d43d6cb1f..f430aa3c51be9 100644 --- a/src/data/finsupp/defs.lean +++ b/src/data/finsupp/defs.lean @@ -524,10 +524,8 @@ noncomputable def of_support_finite lemma of_support_finite_coe {f : α → M} {hf : (function.support f).finite} : (of_support_finite f hf : α → M) = f := rfl -instance : can_lift (α → M) (α →₀ M) := -{ coe := coe_fn, - cond := λ f, (function.support f).finite, - prf := λ f hf, ⟨of_support_finite f hf, rfl⟩ } +instance can_lift : can_lift (α → M) (α →₀ M) coe_fn (λ f, (function.support f).finite) := +{ prf := λ f hf, ⟨of_support_finite f hf, rfl⟩ } end of_support_finite diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index bc1987346bfe1..427146bdeb6fa 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -208,10 +208,9 @@ by rw [range_map, subtype.range_coe] /-- If each element of a list can be lifted to some type, then the whole list can be lifted to this type. -/ -instance [h : can_lift α β] : can_lift (list α) (list β) := -{ coe := list.map h.coe, - cond := λ l, ∀ x ∈ l, can_lift.cond β x, - prf := λ l H, +instance can_lift (c) (p) [can_lift α β c p] : + can_lift (list α) (list β) (list.map c) (λ l, ∀ x ∈ l, p x) := +{ prf := λ l H, begin rw [← set.mem_range, range_map], exact λ a ha, can_lift.prf a (H a ha), diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 515aed2fffe64..b304b1e89c5d0 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -822,10 +822,9 @@ quotient.induction_on₂ s t $ λ l₁ l₂, congr_arg coe $ map_append _ _ _ /-- If each element of `s : multiset α` can be lifted to `β`, then `s` can be lifted to `multiset β`. -/ -instance [can_lift α β] : can_lift (multiset α) (multiset β) := -{ cond := λ s, ∀ x ∈ s, can_lift.cond β x, - coe := map can_lift.coe, - prf := by { rintro ⟨l⟩ hl, lift l to list β using hl, exact ⟨l, coe_map _ _⟩ } } +instance can_lift (c) (p) [can_lift α β c p] : + can_lift (multiset α) (multiset β) (map c) (λ s, ∀ x ∈ s, p x) := +{ prf := by { rintro ⟨l⟩ hl, lift l to list β using hl, exact ⟨l, coe_map _ _⟩ } } /-- `multiset.map` as an `add_monoid_hom`. -/ def map_add_monoid_hom (f : α → β) : multiset α →+ multiset β := diff --git a/src/data/nat/enat.lean b/src/data/nat/enat.lean index 7a2e3ea159366..b46f12583089a 100644 --- a/src/data/nat/enat.lean +++ b/src/data/nat/enat.lean @@ -35,7 +35,7 @@ variables {m n : ℕ∞} @[simp, norm_cast] lemma coe_sub (m n : ℕ) : ↑(m - n) = (m - n : ℕ∞) := rfl @[simp, norm_cast] lemma coe_mul (m n : ℕ) : ↑(m * n) = (m * n : ℕ∞) := with_top.coe_mul -instance : can_lift ℕ∞ ℕ := with_top.can_lift +instance can_lift : can_lift ℕ∞ ℕ coe (λ n, n ≠ ⊤) := with_top.can_lift /-- Conversion of `ℕ∞` to `ℕ` sending `∞` to `0`. -/ def to_nat : monoid_with_zero_hom ℕ∞ ℕ := diff --git a/src/data/pnat/basic.lean b/src/data/pnat/basic.lean index 231e05a3e7eab..707a7b280aeb6 100644 --- a/src/data/pnat/basic.lean +++ b/src/data/pnat/basic.lean @@ -422,11 +422,11 @@ end pnat section can_lift -instance nat.can_lift_pnat : can_lift ℕ ℕ+ := -⟨coe, λ n, 0 < n, λ n hn, ⟨nat.to_pnat' n, pnat.to_pnat'_coe hn⟩⟩ +instance nat.can_lift_pnat : can_lift ℕ ℕ+ coe ((<) 0) := +⟨λ n hn, ⟨nat.to_pnat' n, pnat.to_pnat'_coe hn⟩⟩ -instance int.can_lift_pnat : can_lift ℤ ℕ+ := -⟨coe, λ n, 0 < n, λ n hn, ⟨nat.to_pnat' (int.nat_abs n), +instance int.can_lift_pnat : can_lift ℤ ℕ+ coe ((<) 0) := +⟨λ n hn, ⟨nat.to_pnat' (int.nat_abs n), by rw [coe_coe, nat.to_pnat'_coe, if_pos (int.nat_abs_pos_of_ne_zero hn.ne'), int.nat_abs_of_nonneg hn.le]⟩⟩ diff --git a/src/data/rat/defs.lean b/src/data/rat/defs.lean index 4d46b52bcd7b9..67ffc9802c7d9 100644 --- a/src/data/rat/defs.lean +++ b/src/data/rat/defs.lean @@ -764,8 +764,8 @@ by { conv_rhs { rw [←(@num_denom q), hq] }, rw [coe_int_eq_mk], refl } lemma denom_eq_one_iff (r : ℚ) : r.denom = 1 ↔ ↑r.num = r := ⟨rat.coe_int_num_of_denom_eq_one, λ h, h ▸ rat.coe_int_denom r.num⟩ -instance : can_lift ℚ ℤ := -⟨coe, λ q, q.denom = 1, λ q hq, ⟨q.num, coe_int_num_of_denom_eq_one hq⟩⟩ +instance can_lift : can_lift ℚ ℤ coe (λ q, q.denom = 1) := +⟨λ q hq, ⟨q.num, coe_int_num_of_denom_eq_one hq⟩⟩ theorem coe_nat_eq_mk (n : ℕ) : ↑n = n /. 1 := by rw [← int.cast_coe_nat, coe_int_eq_mk] diff --git a/src/data/rat/nnrat.lean b/src/data/rat/nnrat.lean index 29e05d2d872e2..5e13d9d27ee6d 100644 --- a/src/data/rat/nnrat.lean +++ b/src/data/rat/nnrat.lean @@ -41,10 +41,8 @@ instance : has_coe ℚ≥0 ℚ := ⟨subtype.val⟩ /- Simp lemma to put back `n.val` into the normal form given by the coercion. -/ @[simp] lemma val_eq_coe (q : ℚ≥0) : q.val = q := rfl -instance : can_lift ℚ ℚ≥0 := -{ coe := coe, - cond := λ q, 0 ≤ q, - prf := λ q hq, ⟨⟨q, hq⟩, rfl⟩ } +instance can_lift : can_lift ℚ ℚ≥0 coe (λ q, 0 ≤ q) := +{ prf := λ q hq, ⟨⟨q, hq⟩, rfl⟩ } @[ext] lemma ext : (p : ℚ) = (q : ℚ) → p = q := subtype.ext diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 6f0d22f562e8e..98aa329bbfd1f 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -104,10 +104,8 @@ instance : inhabited ℝ≥0∞ := ⟨0⟩ instance : has_coe ℝ≥0 ℝ≥0∞ := ⟨ option.some ⟩ -instance : can_lift ℝ≥0∞ ℝ≥0 := -{ coe := coe, - cond := λ r, r ≠ ∞, - prf := λ x hx, ⟨option.get $ option.ne_none_iff_is_some.1 hx, option.some_get _⟩ } +instance can_lift : can_lift ℝ≥0∞ ℝ≥0 coe (λ r, r ≠ ∞) := +{ prf := λ x hx, ⟨option.get $ option.ne_none_iff_is_some.1 hx, option.some_get _⟩ } @[simp] lemma none_eq_top : (none : ℝ≥0∞) = ∞ := rfl @[simp] lemma some_eq_coe (a : ℝ≥0) : (some a : ℝ≥0∞) = (↑a : ℝ≥0∞) := rfl diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index 9ccc89c3ae9dd..f9371a704457d 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -101,10 +101,8 @@ protected def rec {C : ereal → Sort*} (h_bot : C ⊥) (h_real : Π a : ℝ, C /-! ### Real coercion -/ -instance : can_lift ereal ℝ := -{ coe := coe, - cond := λ r, r ≠ ⊤ ∧ r ≠ ⊥, - prf := λ x hx, +instance can_lift : can_lift ereal ℝ coe (λ r, r ≠ ⊤ ∧ r ≠ ⊥) := +{ prf := λ x hx, begin induction x using ereal.rec, { simpa using hx }, @@ -172,10 +170,10 @@ lemma to_real_le_to_real {x y : ereal} (h : x ≤ y) (hx : x ≠ ⊥) (hy : y x.to_real ≤ y.to_real := begin lift x to ℝ, + { simp [hx, (h.trans_lt (lt_top_iff_ne_top.2 hy)).ne], }, lift y to ℝ, - { simpa using h }, { simp [hy, ((bot_lt_iff_ne_bot.2 hx).trans_le h).ne'] }, - { simp [hx, (h.trans_lt (lt_top_iff_ne_top.2 hy)).ne], }, + simpa using h end lemma coe_to_real {x : ereal} (hx : x ≠ ⊤) (h'x : x ≠ ⊥) : (x.to_real : ereal) = x := @@ -301,8 +299,8 @@ def ne_top_bot_equiv_real : ({⊥, ⊤}ᶜ : set ereal) ≃ ℝ := inv_fun := λ x, ⟨x, by simp⟩, left_inv := λ ⟨x, hx⟩, subtype.eq $ begin lift x to ℝ, + { simpa [not_or_distrib, and_comm] using hx }, { simp }, - { simpa [not_or_distrib, and_comm] using hx } end, right_inv := λ x, by simp } diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 354b5efca9e9c..75ef1fe775d56 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -67,7 +67,7 @@ instance : has_coe ℝ≥0 ℝ := ⟨subtype.val⟩ /- Simp lemma to put back `n.val` into the normal form given by the coercion. -/ @[simp] lemma val_eq_coe (n : ℝ≥0) : n.val = n := rfl -instance : can_lift ℝ ℝ≥0 := subtype.can_lift _ +instance can_lift : can_lift ℝ ℝ≥0 coe (λ r, 0 ≤ r) := subtype.can_lift _ protected lemma eq {n m : ℝ≥0} : (n : ℝ) = (m : ℝ) → n = m := subtype.eq diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 93b1761e794dc..22fa723578ec0 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -115,12 +115,11 @@ instance {α : Type u} : has_coe_to_sort (set α) (Type u) := ⟨λ s, {x // x instance pi_set_coe.can_lift (ι : Type u) (α : Π i : ι, Type v) [ne : Π i, nonempty (α i)] (s : set ι) : - can_lift (Π i : s, α i) (Π i, α i) := -{ coe := λ f i, f i, - .. pi_subtype.can_lift ι α s } + can_lift (Π i : s, α i) (Π i, α i) (λ f i, f i) (λ _, true) := +pi_subtype.can_lift ι α s instance pi_set_coe.can_lift' (ι : Type u) (α : Type v) [ne : nonempty α] (s : set ι) : - can_lift (s → α) (ι → α) := + can_lift (s → α) (ι → α) (λ f i, f i) (λ _, true) := pi_set_coe.can_lift ι (λ _, α) s end set @@ -2151,10 +2150,9 @@ is_compl.compl_eq is_compl_range_inl_range_inr.symm @[simp] theorem range_quot_mk (r : α → α → Prop) : range (quot.mk r) = univ := (surjective_quot_mk r).range_eq -instance can_lift [can_lift α β] : can_lift (set α) (set β) := -{ coe := λ s, can_lift.coe '' s, - cond := λ s, ∀ x ∈ s, can_lift.cond β x, - prf := λ s hs, subset_range_iff_exists_image_eq.mp (λ x hx, can_lift.prf _ (hs x hx)) } +instance can_lift (c) (p) [can_lift α β c p] : + can_lift (set α) (set β) (('') c) (λ s, ∀ x ∈ s, p x) := +{ prf := λ s hs, subset_range_iff_exists_image_eq.mp (λ x hx, can_lift.prf _ (hs x hx)) } @[simp] theorem range_quotient_mk [setoid α] : range (λx : α, ⟦x⟧) = univ := range_quot_mk _ diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index ddc4def3d29cd..b444ccac7fe5e 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -104,10 +104,8 @@ theorem finite.exists_finset_coe {s : set α} (h : s.finite) : by { casesI h, exact ⟨s.to_finset, s.coe_to_finset⟩ } /-- Finite sets can be lifted to finsets. -/ -instance : can_lift (set α) (finset α) := -{ coe := coe, - cond := set.finite, - prf := λ s hs, hs.exists_finset_coe } +instance : can_lift (set α) (finset α) coe set.finite := +{ prf := λ s hs, hs.exists_finset_coe } /-- A set is infinite if it is not finite. diff --git a/src/linear_algebra/quadratic_form/basic.lean b/src/linear_algebra/quadratic_form/basic.lean index a934839a9de37..17aa943d65917 100644 --- a/src/linear_algebra/quadratic_form/basic.lean +++ b/src/linear_algebra/quadratic_form/basic.lean @@ -629,10 +629,9 @@ abbreviation associated' : quadratic_form R M →ₗ[ℤ] bilin_form R M := associated_hom ℤ /-- Symmetric bilinear forms can be lifted to quadratic forms -/ -instance : can_lift (bilin_form R M) (quadratic_form R M) := -{ coe := associated_hom ℕ, - cond := bilin_form.is_symm, - prf := λ B hB, ⟨B.to_quadratic_form, associated_left_inverse _ hB⟩ } +instance can_lift : + can_lift (bilin_form R M) (quadratic_form R M) (associated_hom ℕ) bilin_form.is_symm := +{ prf := λ B hB, ⟨B.to_quadratic_form, associated_left_inverse _ hB⟩ } /-- There exists a non-null vector with respect to any quadratic form `Q` whose associated bilinear form is non-zero, i.e. there exists `x` such that `Q x ≠ 0`. -/ diff --git a/src/logic/embedding.lean b/src/logic/embedding.lean index c7ee1e4a6625c..8e7063d287391 100644 --- a/src/logic/embedding.lean +++ b/src/logic/embedding.lean @@ -34,10 +34,8 @@ instance {α : Sort u} {β : Sort v} : embedding_like (α ↪ β) α β := injective' := embedding.inj', coe_injective' := λ f g h, by { cases f, cases g, congr' } } -instance {α β : Sort*} : can_lift (α → β) (α ↪ β) := -{ coe := coe_fn, - cond := injective, - prf := λ f hf, ⟨⟨f, hf⟩, rfl⟩ } +instance {α β : Sort*} : can_lift (α → β) (α ↪ β) coe_fn injective := +{ prf := λ f hf, ⟨⟨f, hf⟩, rfl⟩ } end function diff --git a/src/logic/equiv/basic.lean b/src/logic/equiv/basic.lean index af1fdde196ec5..7cc44e6d615ff 100644 --- a/src/logic/equiv/basic.lean +++ b/src/logic/equiv/basic.lean @@ -1657,10 +1657,8 @@ lemma of_bijective_apply_symm_apply (f : α → β) (hf : bijective f) (x : β) (of_bijective f hf).symm (f x) = x := (of_bijective f hf).symm_apply_apply x -instance : can_lift (α → β) (α ≃ β) := -{ coe := coe_fn, - cond := bijective, - prf := λ f hf, ⟨of_bijective f hf, rfl⟩ } +instance : can_lift (α → β) (α ≃ β) coe_fn bijective := +{ prf := λ f hf, ⟨of_bijective f hf, rfl⟩ } section diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index 0ceff65d56554..cf60b2f747c44 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -599,10 +599,8 @@ by { cases x, simpa using h, refl, } @[simp] lemma unbot_coe (x : α) (h : (x : with_bot α) ≠ ⊥ := coe_ne_bot) : (x : with_bot α).unbot h = x := rfl -instance : can_lift (with_bot α) α := -{ coe := coe, - cond := λ r, r ≠ ⊥, - prf := λ x h, ⟨x.unbot h, coe_unbot _ _⟩ } +instance can_lift : can_lift (with_bot α) α coe (λ r, r ≠ ⊥) := +{ prf := λ x h, ⟨x.unbot h, coe_unbot _ _⟩ } section has_le variables [has_le α] @@ -963,10 +961,8 @@ with_bot.coe_unbot x h @[simp] lemma untop_coe (x : α) (h : (x : with_top α) ≠ ⊤ := coe_ne_top) : (x : with_top α).untop h = x := rfl -instance : can_lift (with_top α) α := -{ coe := coe, - cond := λ r, r ≠ ⊤, - prf := λ x h, ⟨x.untop h, coe_untop _ _⟩ } +instance can_lift : can_lift (with_top α) α coe (λ r, r ≠ ⊤) := +{ prf := λ x h, ⟨x.untop h, coe_untop _ _⟩ } section has_le variables [has_le α] diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index f695218ebbd63..f475bd4fd9cb8 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1799,10 +1799,9 @@ end lemma map_comap_of_mem {f : filter β} {m : α → β} (hf : range m ∈ f) : (f.comap m).map m = f := by rw [map_comap, inf_eq_left.2 (le_principal_iff.2 hf)] -instance [can_lift α β] : can_lift (filter α) (filter β) := -{ coe := map can_lift.coe, - cond := λ f, ∀ᶠ x : α in f, can_lift.cond β x, - prf := λ f hf, ⟨comap can_lift.coe f, map_comap_of_mem $ hf.mono can_lift.prf⟩ } +instance can_lift (c) (p) [can_lift α β c p] : + can_lift (filter α) (filter β) (map c) (λ f, ∀ᶠ x : α in f, p x) := +{ prf := λ f hf, ⟨comap c f, map_comap_of_mem $ hf.mono can_lift.prf⟩ } lemma comap_le_comap_iff {f g : filter β} {m : α → β} (hf : range m ∈ f) : comap m f ≤ comap m g ↔ f ≤ g := diff --git a/src/order/hom/basic.lean b/src/order/hom/basic.lean index deab8d48d7eb2..e66c9933a9205 100644 --- a/src/order/hom/basic.lean +++ b/src/order/hom/basic.lean @@ -175,10 +175,8 @@ lemma ext (f g : α →o β) (h : (f : α → β) = g) : f = g := fun_like.coe_i lemma coe_eq (f : α →o β) : coe f = f := by ext ; refl /-- One can lift an unbundled monotone function to a bundled one. -/ -instance : can_lift (α → β) (α →o β) := -{ coe := coe_fn, - cond := monotone, - prf := λ f h, ⟨⟨f, h⟩, rfl⟩ } +instance : can_lift (α → β) (α →o β) coe_fn monotone := +{ prf := λ f h, ⟨⟨f, h⟩, rfl⟩ } /-- Copy of an `order_hom` with a new `to_fun` equal to the old one. Useful to fix definitional equalities. -/ diff --git a/src/order/interval.lean b/src/order/interval.lean index 76defaf048fa9..d26f7dabe297d 100644 --- a/src/order/interval.lean +++ b/src/order/interval.lean @@ -169,7 +169,8 @@ section has_le variables [has_le α] {s t : interval α} instance : has_coe_t (nonempty_interval α) (interval α) := with_bot.has_coe_t -instance : can_lift (interval α) (nonempty_interval α) := with_bot.can_lift +instance can_lift : can_lift (interval α) (nonempty_interval α) coe (λ r, r ≠ ⊥) := +with_bot.can_lift lemma coe_injective : injective (coe : nonempty_interval α → interval α) := with_bot.coe_injective @[simp, norm_cast] lemma coe_inj {s t : nonempty_interval α} : (s : interval α) = t ↔ s = t := diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index 6585a0ed2a0f7..58e2f04368d08 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -106,8 +106,8 @@ def mk : Type u → cardinal := quotient.mk localized "prefix (name := cardinal.mk) `#` := cardinal.mk" in cardinal -instance can_lift_cardinal_Type : can_lift cardinal.{u} (Type u) := -⟨mk, λ c, true, λ c _, quot.induction_on c $ λ α, ⟨α, rfl⟩⟩ +instance can_lift_cardinal_Type : can_lift cardinal.{u} (Type u) mk (λ _, true) := +⟨λ c _, quot.induction_on c $ λ α, ⟨α, rfl⟩⟩ @[elab_as_eliminator] lemma induction_on {p : cardinal → Prop} (c : cardinal) (h : ∀ α, p (#α)) : p c := @@ -1005,8 +1005,8 @@ alias le_aleph_0_iff_set_countable ↔ _ _root_.set.countable.le_aleph_0 #{x // p x} ≤ ℵ₀ ↔ {x | p x}.countable := le_aleph_0_iff_set_countable -instance can_lift_cardinal_nat : can_lift cardinal ℕ := -⟨ coe, λ x, x < ℵ₀, λ x hx, let ⟨n, hn⟩ := lt_aleph_0.mp hx in ⟨n, hn.symm⟩⟩ +instance can_lift_cardinal_nat : can_lift cardinal ℕ coe (λ x, x < ℵ₀) := +⟨λ x hx, let ⟨n, hn⟩ := lt_aleph_0.mp hx in ⟨n, hn.symm⟩⟩ theorem add_lt_aleph_0 {a b : cardinal} (ha : a < ℵ₀) (hb : b < ℵ₀) : a + b < ℵ₀ := match a, b, lt_aleph_0.1 ha, lt_aleph_0.1 hb with diff --git a/src/tactic/lift.lean b/src/tactic/lift.lean index 457223a8c8bef..6fa9bdf5ad5e2 100644 --- a/src/tactic/lift.lean +++ b/src/tactic/lift.lean @@ -3,7 +3,9 @@ Copyright (c) 2019 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ + import tactic.rcases + /-! # lift tactic @@ -17,20 +19,18 @@ lift, tactic /-- A class specifying that you can lift elements from `α` to `β` assuming `cond` is true. Used by the tactic `lift`. -/ -class can_lift (α β : Sort*) := -(coe : β → α) -(cond : α → Prop) +class can_lift (α β : Sort*) (coe : out_param $ β → α) (cond : out_param $ α → Prop) := (prf : ∀(x : α), cond x → ∃(y : β), coe y = x) -instance : can_lift ℤ ℕ := -⟨coe, λ n, 0 ≤ n, λ n hn, ⟨n.nat_abs, int.nat_abs_of_nonneg hn⟩⟩ +instance : can_lift ℤ ℕ coe ((≤) 0) := +⟨λ n hn, ⟨n.nat_abs, int.nat_abs_of_nonneg hn⟩⟩ /-- Enable automatic handling of pi types in `can_lift`. -/ -instance pi.can_lift (ι : Sort*) (α β : ι → Sort*) [Π i : ι, can_lift (α i) (β i)] : - can_lift (Π i : ι, α i) (Π i : ι, β i) := -{ coe := λ f i, can_lift.coe (f i), - cond := λ f, ∀ i, can_lift.cond (β i) (f i), - prf := λ f hf, ⟨λ i, classical.some (can_lift.prf (f i) (hf i)), funext $ λ i, +instance pi.can_lift (ι : Sort*) (α β : ι → Sort*) + (coe : Π i, β i → α i) (P : Π i, α i → Prop) + [Π i : ι, can_lift (α i) (β i) (coe i) (P i)] : + can_lift (Π i : ι, α i) (Π i : ι, β i) (λ f i, coe i (f i)) (λ f, ∀ i, P i (f i)) := +{ prf := λ f hf, ⟨λ i, classical.some (can_lift.prf (f i) (hf i)), funext $ λ i, classical.some_spec (can_lift.prf (f i) (hf i))⟩ } lemma subtype.exists_pi_extension {ι : Sort*} {α : ι → Sort*} [ne : Π i, nonempty (α i)] @@ -45,39 +45,18 @@ end instance pi_subtype.can_lift (ι : Sort*) (α : ι → Sort*) [ne : Π i, nonempty (α i)] (p : ι → Prop) : - can_lift (Π i : subtype p, α i) (Π i, α i) := -{ coe := λ f i, f i, - cond := λ _, true, - prf := λ f _, subtype.exists_pi_extension f } + can_lift (Π i : subtype p, α i) (Π i, α i) (λ f i, f i) (λ _, true) := +{ prf := λ f _, subtype.exists_pi_extension f } instance pi_subtype.can_lift' (ι : Sort*) (α : Sort*) [ne : nonempty α] (p : ι → Prop) : - can_lift (subtype p → α) (ι → α) := + can_lift (subtype p → α) (ι → α) (λ f i, f i) (λ _, true) := pi_subtype.can_lift ι (λ _, α) p -instance subtype.can_lift {α : Sort*} (p : α → Prop) : can_lift α {x // p x} := -{ coe := coe, - cond := p, - prf := λ a ha, ⟨⟨a, ha⟩, rfl⟩ } +instance subtype.can_lift {α : Sort*} (p : α → Prop) : can_lift α {x // p x} coe p := +{ prf := λ a ha, ⟨⟨a, ha⟩, rfl⟩ } open tactic -/-- -A user attribute used internally by the `lift` tactic. -This should not be applied by hand. --/ -@[user_attribute] -meta def can_lift_attr : user_attribute (list name) := -{ name := "_can_lift", - descr := "internal attribute used by the lift tactic", - parser := failed, - cache_cfg := - { mk_cache := λ _, - do { ls ← attribute.get_instances `instance, - ls.mfilter $ λ l, - do { (_,t) ← mk_const l >>= infer_type >>= open_pis, - return $ t.is_app_of `can_lift } }, - dependencies := [`instance] } } - namespace tactic /-- @@ -87,23 +66,24 @@ Construct the proof of `cond x` in the lift tactic. * `s` and `to_unfold` contain the information of the simp set used to simplify. If the proof was specified, we check whether it has the correct type. -If it doesn't have the correct type, we display an error message -(but first call dsimp on the expression in the message). +If it doesn't have the correct type, we display an error message. If the proof was not specified, we create assert it as a local constant. (The name of this local constant doesn't matter, since `lift` will remove it from the context.) -/ -meta def get_lift_prf (h : option pexpr) (old_tp new_tp inst e : expr) - (s : simp_lemmas) (to_unfold : list name) : tactic expr := do - expected_prf_ty ← mk_app `can_lift.cond [old_tp, new_tp, inst, e], - expected_prf_ty ← s.dsimplify to_unfold expected_prf_ty, - if h_some : h.is_some then - decorate_error "lift tactic failed." $ i_to_expr ``((%%(option.get h_some) : %%expected_prf_ty)) - else do - prf_nm ← get_unused_name, - prf ← assert prf_nm expected_prf_ty, - swap, - return prf +meta def get_lift_prf (h : option pexpr) (e P : expr) : tactic (expr × bool) := do + let expected_prf_ty := P.app e, + expected_prf_ty ← simp_lemmas.mk.dsimplify [] expected_prf_ty {fail_if_unchanged := ff}, + match h with + | some h := do + e ← decorate_error "lift tactic failed." (i_to_expr ``((%%h : %%expected_prf_ty))), + return (e, tt) + | none := do + prf_nm ← get_unused_name, + prf ← assert prf_nm expected_prf_ty, + swap, + return (prf, ff) + end /-- Lift the expression `p` to the type `t`, with proof obligation given by `h`. The list `n` is used for the two newly generated names, and to specify whether `h` should @@ -116,19 +96,21 @@ do e ← i_to_expr p, old_tp ← infer_type e, new_tp ← i_to_expr ``(%%t : Sort*), - inst_type ← mk_app ``can_lift [old_tp, new_tp], + coe ← i_to_expr (``(%%new_tp → %%old_tp)) >>= mk_meta_var, + P ← i_to_expr (``(%%old_tp → Prop)) >>= mk_meta_var, + inst_type ← mk_app ``can_lift [old_tp, new_tp, coe, P], inst ← mk_instance inst_type <|> pformat!"Failed to find a lift from {old_tp} to {new_tp}. Provide an instance of\n {inst_type}" >>= fail, - /- make the simp set to get rid of `can_lift` projections -/ - can_lift_instances ← can_lift_attr.get_cache >>= λ l, l.mmap resolve_name, - (s, to_unfold) ← mk_simp_set tt [] $ can_lift_instances.map simp_arg_type.expr, - prf_cond ← get_lift_prf h old_tp new_tp inst e s to_unfold, + inst ← instantiate_mvars inst, + coe ← instantiate_mvars coe, + P ← instantiate_mvars P, + (prf_cond, b) ← get_lift_prf h e P, let prf_nm := if prf_cond.is_local_constant then some prf_cond.local_pp_name else none, /- We use mk_mapp to apply `can_lift.prf` to all but one argument, and then just use expr.app for the last argument. For some reason we get an error when applying mk_mapp it to all arguments. -/ - prf_ex0 ← mk_mapp `can_lift.prf [old_tp, new_tp, inst, e], + prf_ex0 ← mk_mapp `can_lift.prf [old_tp, new_tp, coe, P, inst, e], let prf_ex := prf_ex0 prf_cond, /- Find the name of the new variable -/ new_nm ← if n ≠ [] then return n.head @@ -138,11 +120,10 @@ do eq_nm ← if hn : 1 < n.length then return (n.nth_le 1 hn) else if e.is_local_constant then return `rfl else get_unused_name `h, - /- We add the proof of the existential statement to the context and then apply - `dsimp` to it, unfolding all `can_lift` instances. -/ + /- We add the proof of the existential statement to the context -/ temp_nm ← get_unused_name, temp_e ← note temp_nm none prf_ex, - dsimp_hyp temp_e s to_unfold {}, + dsimp_hyp temp_e none [] { fail_if_unchanged := ff }, /- We case on the existential. We use `rcases` because `eq_nm` could be `rfl`. -/ rcases none (pexpr.of_expr temp_e) $ rcases_patt.tuple ([new_nm, eq_nm].map rcases_patt.one), /- If the lifted variable is not a local constant, @@ -152,7 +133,8 @@ do /- If the proof `prf_cond` is a local constant, remove it from the context, unless `n` specifies to keep it. -/ if h_prf_nm : prf_nm.is_some ∧ n.nth 2 ≠ prf_nm then - get_local (option.get h_prf_nm.1) >>= clear else skip + get_local (option.get h_prf_nm.1) >>= clear else skip, + if b then skip else swap setup_tactic_parser @@ -175,9 +157,10 @@ Lift an expression to another type. (here `P` is some term of type `ℤ → Prop`). * The argument `using hn` is optional, the tactic `lift n to ℕ` does the same, but also creates a new subgoal that `n ≥ 0` (where `n` is the old variable). + This subgoal will be placed at the top of the goal list. + So for example the tactic `lift n to ℕ` transforms the goal `n : ℤ, h : P n ⊢ n = 3` to two goals - `n : ℕ, h : P ↑n ⊢ ↑n = 3` and `n : ℤ, h : P n ⊢ n ≥ 0`. + `n : ℤ, h : P n ⊢ n ≥ 0` and `n : ℕ, h : P ↑n ⊢ ↑n = 3`. * You can also use `lift n to ℕ using e` where `e` is any expression of type `n ≥ 0`. * Use `lift n to ℕ with k` to specify the name of the new variable. * Use `lift n to ℕ with k hk` to also specify the name of the equality `↑k = n`. In this case, `n` diff --git a/src/topology/alexandroff.lean b/src/topology/alexandroff.lean index dac07a9dacc39..581c82ff0c0a1 100644 --- a/src/topology/alexandroff.lean +++ b/src/topology/alexandroff.lean @@ -102,10 +102,8 @@ lemma ne_infty_iff_exists {x : alexandroff X} : x ≠ ∞ ↔ ∃ (y : X), (y : alexandroff X) = x := by induction x using alexandroff.rec; simp -instance : can_lift (alexandroff X) X := -{ coe := coe, - cond := λ x, x ≠ ∞, - prf := λ x, ne_infty_iff_exists.1 } +instance can_lift : can_lift (alexandroff X) X coe (λ x, x ≠ ∞) := +with_top.can_lift lemma not_mem_range_coe_iff {x : alexandroff X} : x ∉ range (coe : X → alexandroff X) ↔ x = ∞ := diff --git a/src/topology/continuous_function/units.lean b/src/topology/continuous_function/units.lean index 6552ba5f8a04d..ead31e4b73c7b 100644 --- a/src/topology/continuous_function/units.lean +++ b/src/topology/continuous_function/units.lean @@ -63,10 +63,9 @@ noncomputable def units_of_forall_is_unit {f : C(X, R)} (h : ∀ x, is_unit (f x { to_fun := λ x, (h x).unit, continuous_to_fun := normed_ring.is_unit_unit_continuous h } -instance : can_lift C(X, R) C(X, Rˣ) := -{ coe := λ f, ⟨λ x, f x, units.continuous_coe.comp f.continuous⟩, - cond := λ f, ∀ x, is_unit (f x), - prf := λ f h, ⟨units_of_forall_is_unit h, by { ext, refl }⟩ } +instance can_lift : can_lift C(X, R) C(X, Rˣ) + (λ f, ⟨λ x, f x, units.continuous_coe.comp f.continuous⟩) (λ f, ∀ x, is_unit (f x)) := +{ prf := λ f h, ⟨units_of_forall_is_unit h, by { ext, refl }⟩ } lemma is_unit_iff_forall_is_unit (f : C(X, R)) : is_unit f ↔ ∀ x, is_unit (f x) := diff --git a/src/topology/instances/nnreal.lean b/src/topology/instances/nnreal.lean index 0f8c94e0ba68a..c28df7fb6cc83 100644 --- a/src/topology/instances/nnreal.lean +++ b/src/topology/instances/nnreal.lean @@ -77,10 +77,9 @@ continuous_subtype_val @[simps { fully_applied := ff }] def _root_.continuous_map.coe_nnreal_real : C(ℝ≥0, ℝ) := ⟨coe, continuous_coe⟩ -instance {X : Type*} [topological_space X] : can_lift C(X, ℝ) C(X, ℝ≥0) := -{ coe := continuous_map.coe_nnreal_real.comp, - cond := λ f, ∀ x, 0 ≤ f x, - prf := λ f hf, ⟨⟨λ x, ⟨f x, hf x⟩, f.2.subtype_mk _⟩, fun_like.ext' rfl⟩ } +instance continuous_map.can_lift {X : Type*} [topological_space X] : + can_lift C(X, ℝ) C(X, ℝ≥0) continuous_map.coe_nnreal_real.comp (λ f, ∀ x, 0 ≤ f x) := +{ prf := λ f hf, ⟨⟨λ x, ⟨f x, hf x⟩, f.2.subtype_mk _⟩, fun_like.ext' rfl⟩ } @[simp, norm_cast] lemma tendsto_coe {f : filter α} {m : α → ℝ≥0} {x : ℝ≥0} : tendsto (λa, (m a : ℝ)) f (𝓝 (x : ℝ)) ↔ tendsto m f (𝓝 x) := diff --git a/src/topology/sets/compacts.lean b/src/topology/sets/compacts.lean index b1b27a1341619..f4f37e3489b03 100644 --- a/src/topology/sets/compacts.lean +++ b/src/topology/sets/compacts.lean @@ -45,10 +45,8 @@ lemma compact (s : compacts α) : is_compact (s : set α) := s.compact' instance (K : compacts α) : compact_space K := is_compact_iff_compact_space.1 K.compact -instance : can_lift (set α) (compacts α) := -{ coe := coe, - cond := is_compact, - prf := λ K hK, ⟨⟨K, hK⟩, rfl⟩ } +instance : can_lift (set α) (compacts α) coe is_compact := +{ prf := λ K hK, ⟨⟨K, hK⟩, rfl⟩ } @[ext] protected lemma ext {s t : compacts α} (h : (s : set α) = t) : s = t := set_like.ext' h diff --git a/test/lift.lean b/test/lift.lean index 2ff6db4df82ea..b48d2b0a69e93 100644 --- a/test/lift.lean +++ b/test/lift.lean @@ -1,8 +1,7 @@ -import data.int.basic +import data.set.basic import tactic.lift /-! Some tests of the `lift` tactic. -/ - example (n m k x z u : ℤ) (hn : 0 < n) (hk : 0 ≤ k + n) (hu : 0 ≤ u) (h : k + n = 2 + x) (f : false) : k + n = m + x := @@ -10,15 +9,15 @@ begin lift n to ℕ using le_of_lt hn, guard_target (k + ↑n = m + x), guard_hyp hn : (0 : ℤ) < ↑n, lift m to ℕ, - guard_target (k + ↑n = ↑m + x), tactic.swap, guard_target (0 ≤ m), tactic.swap, + guard_target (0 ≤ m), tactic.swap, guard_target (k + ↑n = ↑m + x), tactic.num_goals >>= λ n, guard (n = 2), lift (k + n) to ℕ using hk with l hl, guard_hyp l : ℕ, guard_hyp hl : ↑l = k + ↑n, guard_target (↑l = ↑m + x), tactic.success_if_fail (tactic.get_local `hk), lift x to ℕ with y hy, - guard_hyp y : ℕ, guard_hyp hy : ↑y = x, guard_target (↑l = ↑m + x), + tactic.swap, guard_hyp y : ℕ, guard_hyp hy : ↑y = x, guard_target (↑l = ↑m + x), tactic.swap, lift z to ℕ with w, - guard_hyp w : ℕ, tactic.success_if_fail (tactic.get_local `z), + tactic.swap, guard_hyp w : ℕ, tactic.success_if_fail (tactic.get_local `z), tactic.swap, lift u to ℕ using hu with u rfl hu, guard_hyp hu : (0 : ℤ) ≤ ↑u, @@ -31,7 +30,7 @@ begin lift f to α → ℕ using hf, guard_target ((0:ℤ) ≤ 2 * (λ i : α, (f i : ℤ)) a), guard_hyp hf' : ∀ a, ((λ i : α, (f i:ℤ)) a) < 1, - exact int.coe_nat_nonneg _ + constructor, end -- fail gracefully when the lifted variable is a local definition @@ -44,19 +43,16 @@ begin refl end -instance can_lift_unit : can_lift unit unit := -⟨id, λ x, true, λ x _, ⟨x, rfl⟩⟩ - -/- test whether new instances of `can_lift` are added as simp lemmas -/ -run_cmd do l ← can_lift_attr.get_cache, guard (`can_lift_unit ∈ l) +instance can_lift_unit : can_lift unit unit id (λ _, true) := ⟨λ x _, ⟨x, rfl⟩⟩ /- test error messages -/ example (n : ℤ) (hn : 0 < n) : true := begin - success_if_fail_with_msg {lift n to ℕ using hn} "lift tactic failed. -invalid type ascription, term has type\n 0 < n\nbut is expected to have type\n 0 ≤ n", + success_if_fail_with_msg {lift n to ℕ using hn} ("lift tactic failed.\n" ++ + "invalid type ascription, term has type\n 0 < n\nbut is expected to have type\n 0 ≤ n"), success_if_fail_with_msg {lift (n : option ℤ) to ℕ} - "Failed to find a lift from option ℤ to ℕ. Provide an instance of\n can_lift (option ℤ) ℕ", + ("Failed to find a lift from option ℤ to ℕ. " ++ + "Provide an instance of\n can_lift (option ℤ) ℕ ?m_1 ?m_2"), trivial end @@ -67,15 +63,8 @@ begin exact 0 end -instance can_lift_subtype (R : Type*) (P : R → Prop) : can_lift R {x // P x} := -{ coe := coe, - cond := λ x, P x, - prf := λ x hx, ⟨⟨x, hx⟩, rfl⟩ } - -instance can_lift_set (R : Type*) (s : set R) : can_lift R s := -{ coe := coe, - cond := λ x, x ∈ s, - prf := λ x hx, ⟨⟨x, hx⟩, rfl⟩ } +instance can_lift_set (R : Type*) (s : set R) : can_lift R s coe (λ x, x ∈ s) := +{ prf := λ x hx, ⟨⟨x, hx⟩, rfl⟩ } example {R : Type*} {P : R → Prop} (x : R) (hx : P x) : true := by { lift x to {x // P x} using hx with y, trivial } @@ -85,7 +74,7 @@ example {R : Type*} {s : set R} (x : R) (hx : x ∈ s) : true := by { lift x to s using hx with y, trivial } example (n : ℤ) (hn : 0 ≤ n) : true := -by { lift n to ℕ, trivial, exact hn } +by { lift n to ℕ, exact hn, trivial } example (n : ℤ) (hn : 0 ≤ n) : true := by { lift n to ℕ using hn, trivial } From dc1ac244487b298c6d71bad4bf112623311d21e9 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 27 Sep 2022 17:24:09 +0000 Subject: [PATCH 412/828] feat(topology/subset_properties): lemmas about `disjoint`, `nhds_set`, and `is_compact` (#16591) The set neighborhoods filter of a compact set is disjoint with a filter `l` if and only if the neighborhoods filter of each point of this set is disjoint with `l`. --- src/topology/nhds_set.lean | 7 +++++-- src/topology/separation.lean | 2 +- src/topology/subset_properties.lean | 21 +++++++++++++++++++++ 3 files changed, 27 insertions(+), 3 deletions(-) diff --git a/src/topology/nhds_set.lean b/src/topology/nhds_set.lean index d6e72c0bad38d..a4a7a1916354e 100644 --- a/src/topology/nhds_set.lean +++ b/src/topology/nhds_set.lean @@ -83,8 +83,11 @@ lemma mem_nhds_set_empty : s ∈ 𝓝ˢ (∅ : set α) := by simp @[simp] lemma nhds_set_univ : 𝓝ˢ (univ : set α) = ⊤ := by rw [is_open_univ.nhds_set_eq, principal_univ] -lemma monotone_nhds_set : monotone (𝓝ˢ : set α → filter α) := -λ s t hst, Sup_le_Sup $ image_subset _ hst +@[mono] lemma nhds_set_mono (h : s ⊆ t) : 𝓝ˢ s ≤ 𝓝ˢ t := Sup_le_Sup $ image_subset _ h + +lemma monotone_nhds_set : monotone (𝓝ˢ : set α → filter α) := λ s t, nhds_set_mono + +lemma nhds_le_nhds_set (h : x ∈ s) : 𝓝 x ≤ 𝓝ˢ s := le_Sup $ mem_image_of_mem _ h @[simp] lemma nhds_set_union (s t : set α) : 𝓝ˢ (s ∪ t) = 𝓝ˢ s ⊔ 𝓝ˢ t := by simp only [nhds_set, image_union, Sup_union] diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 8474385a6b42b..8e76f4de10663 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -573,7 +573,7 @@ lemma injective_nhds_set [t1_space α] : function.injective (𝓝ˢ : set α → lemma strict_mono_nhds_set [t1_space α] : strict_mono (𝓝ˢ : set α → filter α) := monotone_nhds_set.strict_mono_of_injective injective_nhds_set -@[simp] lemma nhds_le_nhds_set [t1_space α] {s : set α} {x : α} : 𝓝 x ≤ 𝓝ˢ s ↔ x ∈ s := +@[simp] lemma nhds_le_nhds_set_iff [t1_space α] {s : set α} {x : α} : 𝓝 x ≤ 𝓝ˢ s ↔ x ∈ s := by rw [← nhds_set_singleton, nhds_set_le_iff, singleton_subset_iff] /-- Removing a non-isolated point from a dense set, one still obtains a dense set. -/ diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index d13ee6737558d..b824ae53ca921 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -204,6 +204,27 @@ let ⟨t, ht⟩ := hs.elim_nhds_subcover' (λ x _, U x) hU in ⟨t.image coe, λ x hx, let ⟨y, hyt, hyx⟩ := finset.mem_image.1 hx in hyx ▸ y.2, by rwa finset.set_bUnion_finset_image⟩ +/-- The neighborhood filter of a compact set is disjoint with a filter `l` if and only if the +neighborhood filter of each point of this set is disjoint with `l`. -/ +lemma is_compact.disjoint_nhds_set_left {l : filter α} (hs : is_compact s) : + disjoint (𝓝ˢ s) l ↔ ∀ x ∈ s, disjoint (𝓝 x) l := +begin + refine ⟨λ h x hx, h.mono_left $ nhds_le_nhds_set hx, λ H, _⟩, + choose! U hxU hUl using λ x hx, (nhds_basis_opens x).disjoint_iff_left.1 (H x hx), + choose hxU hUo using hxU, + rcases hs.elim_nhds_subcover U (λ x hx, (hUo x hx).mem_nhds (hxU x hx)) with ⟨t, hts, hst⟩, + refine (has_basis_nhds_set _).disjoint_iff_left.2 + ⟨⋃ x ∈ t, U x, ⟨is_open_bUnion $ λ x hx, hUo x (hts x hx), hst⟩, _⟩, + rw [compl_Union₂, bInter_finset_mem], + exact λ x hx, hUl x (hts x hx) +end + +/-- A filter `l` is disjoint with the neighborhood filter of a compact set if and only if it is +disjoint with the neighborhood filter of each point of this set. -/ +lemma is_compact.disjoint_nhds_set_right {l : filter α} (hs : is_compact s) : + disjoint l (𝓝ˢ s) ↔ ∀ x ∈ s, disjoint l (𝓝 x) := +by simpa only [disjoint.comm] using hs.disjoint_nhds_set_left + /-- For every family of closed sets whose intersection avoids a compact set, there exists a finite subfamily whose intersection avoids this compact set. -/ lemma is_compact.elim_finite_subfamily_closed {s : set α} {ι : Type v} (hs : is_compact s) From 33c6eeae5bda762b558e6094798fe63f7333b3f2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 27 Sep 2022 17:24:10 +0000 Subject: [PATCH 413/828] chore(topology/separation): rename `separated` to `separated_nhds` (#16604) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit E.g., Wikipedia uses "separated" for `disjoint (closure s) t ∧ disjoint s (closure t)`. --- src/topology/separation.lean | 83 +++++++++++++++++------------------- 1 file changed, 40 insertions(+), 43 deletions(-) diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 8e76f4de10663..ef53e9bc19d4a 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -11,12 +11,13 @@ import topology.inseparable /-! # Separation properties of topological spaces. -This file defines the predicate `separated`, and common separation axioms +This file defines the predicate `separated_nhds`, and common separation axioms (under the Kolmogorov classification). ## Main definitions -* `separated`: Two `set`s are separated if they are contained in disjoint open sets. +* `separated_nhds`: Two `set`s are separated by neighbourhoods if they are contained in disjoint + open sets. * `t0_space`: A T₀/Kolmogorov space is a space where, for every two points `x ≠ y`, there is an open set that contains one, but not the other. * `t1_space`: A T₁/Fréchet space is a space where every singleton set is closed. @@ -52,7 +53,7 @@ This file defines the predicate `separated`, and common separation axioms * `t2_iff_nhds`: A space is T₂ iff the neighbourhoods of distinct points generate the bottom filter. * `t2_iff_is_closed_diagonal`: A space is T₂ iff the `diagonal` of `α` (that is, the set of all points of the form `(a, a) : α × α`) is closed under the product topology. -* `finset_disjoint_finset_opens_of_t2`: Any two disjoint finsets are `separated`. +* `finset_disjoint_finset_opens_of_t2`: Any two disjoint finsets are `separated_nhds`. * Most topological constructions preserve Hausdorffness; these results are part of the typeclass inference system (e.g. `embedding.t2_space`) * `set.eq_on.closure`: If two functions are equal on some set `s`, they are equal on its closure. @@ -99,62 +100,60 @@ variables {α : Type u} {β : Type v} [topological_space α] section separation /-- -`separated` is a predicate on pairs of sub`set`s of a topological space. It holds if the two +`separated_nhds` is a predicate on pairs of sub`set`s of a topological space. It holds if the two sub`set`s are contained in disjoint open sets. -/ -def separated : set α → set α → Prop := +def separated_nhds : set α → set α → Prop := λ (s t : set α), ∃ U V : (set α), (is_open U) ∧ is_open V ∧ (s ⊆ U) ∧ (t ⊆ V) ∧ disjoint U V -lemma separated_iff_disjoint {s t : set α} : - separated s t ↔ disjoint (𝓝ˢ s) (𝓝ˢ t) := -by simp only [(has_basis_nhds_set s).disjoint_iff (has_basis_nhds_set t), separated, exists_prop, - ← exists_and_distrib_left, and.assoc, and.comm, and.left_comm] +lemma separated_nhds_iff_disjoint {s t : set α} : + separated_nhds s t ↔ disjoint (𝓝ˢ s) (𝓝ˢ t) := +by simp only [(has_basis_nhds_set s).disjoint_iff (has_basis_nhds_set t), separated_nhds, + exists_prop, ← exists_and_distrib_left, and.assoc, and.comm, and.left_comm] -namespace separated +namespace separated_nhds -open separated +variables {s s₁ s₂ t t₁ t₂ u : set α} -@[symm] lemma symm {s t : set α} : separated s t → separated t s := +@[symm] lemma symm : separated_nhds s t → separated_nhds t s := λ ⟨U, V, oU, oV, aU, bV, UV⟩, ⟨V, U, oV, oU, bV, aU, disjoint.symm UV⟩ -lemma comm (s t : set α) : separated s t ↔ separated t s := -⟨symm, symm⟩ +lemma comm (s t : set α) : separated_nhds s t ↔ separated_nhds t s := ⟨symm, symm⟩ -lemma preimage [topological_space β] {f : α → β} {s t : set β} (h : separated s t) - (hf : continuous f) : separated (f ⁻¹' s) (f ⁻¹' t) := +lemma preimage [topological_space β] {f : α → β} {s t : set β} (h : separated_nhds s t) + (hf : continuous f) : separated_nhds (f ⁻¹' s) (f ⁻¹' t) := let ⟨U, V, oU, oV, sU, tV, UV⟩ := h in ⟨f ⁻¹' U, f ⁻¹' V, oU.preimage hf, oV.preimage hf, preimage_mono sU, preimage_mono tV, UV.preimage f⟩ -protected lemma disjoint {s t : set α} (h : separated s t) : disjoint s t := +protected lemma disjoint (h : separated_nhds s t) : disjoint s t := let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h in hd.mono hsU htV -lemma disjoint_closure_left {s t : set α} (h : separated s t) : disjoint (closure s) t := +lemma disjoint_closure_left (h : separated_nhds s t) : disjoint (closure s) t := let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h in (hd.closure_left hV).mono (closure_mono hsU) htV -lemma disjoint_closure_right {s t : set α} (h : separated s t) : disjoint s (closure t) := +lemma disjoint_closure_right (h : separated_nhds s t) : disjoint s (closure t) := h.symm.disjoint_closure_left.symm -lemma empty_right (a : set α) : separated a ∅ := +lemma empty_right (s : set α) : separated_nhds s ∅ := ⟨_, _, is_open_univ, is_open_empty, λ a h, mem_univ a, λ a h, by cases h, disjoint_empty _⟩ -lemma empty_left (a : set α) : separated ∅ a := +lemma empty_left (s : set α) : separated_nhds ∅ s := (empty_right _).symm -lemma mono {s₁ s₂ t₁ t₂ : set α} (h : separated s₂ t₂) (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : - separated s₁ t₁ := +lemma mono (h : separated_nhds s₂ t₂) (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : separated_nhds s₁ t₁ := let ⟨U, V, hU, hV, hsU, htV, hd⟩ := h in ⟨U, V, hU, hV, hs.trans hsU, ht.trans htV, hd⟩ -lemma union_left {a b c : set α} : separated a c → separated b c → separated (a ∪ b) c := -by simpa only [separated_iff_disjoint, nhds_set_union, disjoint_sup_left] using and.intro +lemma union_left : separated_nhds s u → separated_nhds t u → separated_nhds (s ∪ t) u := +by simpa only [separated_nhds_iff_disjoint, nhds_set_union, disjoint_sup_left] using and.intro -lemma union_right {a b c : set α} (ab : separated a b) (ac : separated a c) : - separated a (b ∪ c) := -(ab.symm.union_left ac.symm).symm +lemma union_right (ht : separated_nhds s t) (hu : separated_nhds s u) : + separated_nhds s (t ∪ u) := +(ht.symm.union_left hu.symm).symm -end separated +end separated_nhds /-- A T₀ space, also known as a Kolmogorov space, is a topological space such that for every pair `x ≠ y`, there is an open set containing one but not the other. We formulate the definition in terms @@ -873,10 +872,10 @@ t2_iff_is_closed_diagonal.mp ‹_› section separated -open separated finset +open separated_nhds finset lemma finset_disjoint_finset_opens_of_t2 [t2_space α] : - ∀ (s t : finset α), disjoint s t → separated (s : set α) t := + ∀ (s t : finset α), disjoint s t → separated_nhds (s : set α) t := begin refine induction_on_union _ (λ a b hi d, (hi d.symm).symm) (λ a d, empty_right a) (λ a b ab, _) _, { obtain ⟨U, V, oU, oV, aU, bV, UV⟩ := t2_separation (finset.disjoint_singleton.1 ab), @@ -888,7 +887,7 @@ begin end lemma point_disjoint_finset_opens_of_t2 [t2_space α] {x : α} {s : finset α} (h : x ∉ s) : - separated ({x} : set α) s := + separated_nhds ({x} : set α) s := by exact_mod_cast finset_disjoint_finset_opens_of_t2 {x} s (finset.disjoint_singleton_left.mpr h) end separated @@ -1135,8 +1134,8 @@ lemma function.left_inverse.closed_embedding [t2_space α] {f : α → β} {g : lemma compact_compact_separated [t2_space α] {s t : set α} (hs : is_compact s) (ht : is_compact t) (hst : disjoint s t) : - ∃ u v, is_open u ∧ is_open v ∧ s ⊆ u ∧ t ⊆ v ∧ disjoint u v := -by simp only [prod_subset_compl_diagonal_iff_disjoint.symm] at ⊢ hst; + separated_nhds s t := +by simp only [separated_nhds, prod_subset_compl_diagonal_iff_disjoint.symm] at ⊢ hst; exact generalized_tube_lemma hs ht is_closed_diagonal.is_open_compl hst /-- In a `t2_space`, every compact set is closed. -/ @@ -1526,12 +1525,11 @@ section normality omits T₂), is one in which for every pair of disjoint closed sets `C` and `D`, there exist disjoint open sets containing `C` and `D` respectively. -/ class normal_space (α : Type u) [topological_space α] extends t1_space α : Prop := -(normal : ∀ s t : set α, is_closed s → is_closed t → disjoint s t → - ∃ u v, is_open u ∧ is_open v ∧ s ⊆ u ∧ t ⊆ v ∧ disjoint u v) +(normal : ∀ s t : set α, is_closed s → is_closed t → disjoint s t → separated_nhds s t) theorem normal_separation [normal_space α] {s t : set α} (H1 : is_closed s) (H2 : is_closed t) (H3 : disjoint s t) : - ∃ u v, is_open u ∧ is_open v ∧ s ⊆ u ∧ t ⊆ v ∧ disjoint u v := + separated_nhds s t := normal_space.normal s t H1 H2 H3 theorem normal_exists_closure_subset [normal_space α] {s t : set α} (hs : is_closed s) @@ -1562,11 +1560,10 @@ protected lemma closed_embedding.normal_space [topological_space β] [normal_spa normal := begin intros s t hs ht hst, - rcases normal_space.normal (f '' s) (f '' t) (hf.is_closed_map s hs) (hf.is_closed_map t ht) - (disjoint_image_of_injective hf.inj hst) with ⟨u, v, hu, hv, hsu, htv, huv⟩, - rw image_subset_iff at hsu htv, - exact ⟨f ⁻¹' u, f ⁻¹' v, hu.preimage hf.continuous, hv.preimage hf.continuous, - hsu, htv, huv.preimage f⟩ + have H : separated_nhds (f '' s) (f '' t), + from normal_space.normal (f '' s) (f '' t) (hf.is_closed_map s hs) (hf.is_closed_map t ht) + (disjoint_image_of_injective hf.inj hst), + exact (H.preimage hf.continuous).mono (subset_preimage_image _ _) (subset_preimage_image _ _) end } variable (α) @@ -1652,7 +1649,7 @@ instance [t5_space α] {p : α → Prop} : t5_space {x // p x} := embedding_subt /-- A `T₅` space is a `T₄` space. -/ @[priority 100] -- see Note [lower instance priority] instance t5_space.to_normal_space [t5_space α] : normal_space α := -⟨λ s t hs ht hd, separated_iff_disjoint.2 $ +⟨λ s t hs ht hd, separated_nhds_iff_disjoint.2 $ completely_normal (by rwa [hs.closure_eq]) (by rwa [ht.closure_eq])⟩ end completely_normal From 89417b902085be3b15c1ee4f262fad5e566733a1 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 27 Sep 2022 17:24:11 +0000 Subject: [PATCH 414/828] feat(geometry/euclidean/basic): angles in subspaces (#16610) Add lemmas that the angle between two vectors or three points in a subspace equals the angle between those vectors or points in the whole space. --- src/geometry/euclidean/basic.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index 2fd08526903dd..fcfe5f33c0f0c 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -91,6 +91,10 @@ by rw [angle, angle, real_inner_smul_left, inner_smul_right, norm_smul, norm_smu angle (f u) (f v) = angle u v := by rw [angle, angle, f.inner_map_map, f.norm_map, f.norm_map] +@[simp, norm_cast] lemma _root_.submodule.angle_coe {s : submodule ℝ V} (x y : s) : + angle (x : V) (y : V) = angle x y := +s.subtypeₗᵢ.angle_map x y + lemma is_conformal_map.preserves_angle {E F : Type*} [inner_product_space ℝ E] [inner_product_space ℝ F] {f' : E →L[ℝ] F} (h : is_conformal_map f') (u v : E) : @@ -412,6 +416,11 @@ end ∠ (f p₁) (f p₂) (f p₃) = ∠ p₁ p₂ p₃ := by simp_rw [angle, ←affine_isometry.map_vsub, linear_isometry.angle_map] +@[simp, norm_cast] lemma _root_.affine_subspace.angle_coe {s : affine_subspace ℝ P} + (p₁ p₂ p₃ : s) : + by haveI : nonempty s := ⟨p₁⟩; exact ∠ (p₁ : P) (p₂ : P) (p₃ : P) = ∠ p₁ p₂ p₃ := +by haveI : nonempty s := ⟨p₁⟩; exact s.subtypeₐᵢ.angle_map p₁ p₂ p₃ + /-- The angle at a point does not depend on the order of the other two points. -/ lemma angle_comm (p1 p2 p3 : P) : ∠ p1 p2 p3 = ∠ p3 p2 p1 := From 7a4b066c3509d73d7d955a69d844ff6e2effcc5b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Tue, 27 Sep 2022 17:24:12 +0000 Subject: [PATCH 415/828] chore(probability/martingale): use volume_tac in definitions (#16619) Use the tactic `volume_tac` in the definitions `martingale`, `submartingale`, `supermartingale`, `martingale_part` and `predictable_part`. In order to do that, change the order of the arguments of `martingale_part` and `predictable_part`. --- src/probability/martingale/basic.lean | 8 ++- .../martingale/borel_cantelli.lean | 6 +- src/probability/martingale/centering.lean | 58 +++++++++---------- 3 files changed, 37 insertions(+), 35 deletions(-) diff --git a/src/probability/martingale/basic.lean b/src/probability/martingale/basic.lean index e9e8381497f03..241afacb1cdc2 100644 --- a/src/probability/martingale/basic.lean +++ b/src/probability/martingale/basic.lean @@ -47,19 +47,21 @@ variables {Ω E ι : Type*} [preorder ι] /-- A family of functions `f : ι → Ω → E` is a martingale with respect to a filtration `ℱ` if `f` is adapted with respect to `ℱ` and for all `i ≤ j`, `μ[f j | ℱ i] =ᵐ[μ] f i`. -/ -def martingale (f : ι → Ω → E) (ℱ : filtration ι m0) (μ : measure Ω) : Prop := +def martingale (f : ι → Ω → E) (ℱ : filtration ι m0) (μ : measure Ω . volume_tac) : Prop := adapted ℱ f ∧ ∀ i j, i ≤ j → μ[f j | ℱ i] =ᵐ[μ] f i /-- A family of integrable functions `f : ι → Ω → E` is a supermartingale with respect to a filtration `ℱ` if `f` is adapted with respect to `ℱ` and for all `i ≤ j`, `μ[f j | ℱ.le i] ≤ᵐ[μ] f i`. -/ -def supermartingale [has_le E] (f : ι → Ω → E) (ℱ : filtration ι m0) (μ : measure Ω) : Prop := +def supermartingale [has_le E] (f : ι → Ω → E) (ℱ : filtration ι m0) (μ : measure Ω . volume_tac) : + Prop := adapted ℱ f ∧ (∀ i j, i ≤ j → μ[f j | ℱ i] ≤ᵐ[μ] f i) ∧ ∀ i, integrable (f i) μ /-- A family of integrable functions `f : ι → Ω → E` is a submartingale with respect to a filtration `ℱ` if `f` is adapted with respect to `ℱ` and for all `i ≤ j`, `f i ≤ᵐ[μ] μ[f j | ℱ.le i]`. -/ -def submartingale [has_le E] (f : ι → Ω → E) (ℱ : filtration ι m0) (μ : measure Ω) : Prop := +def submartingale [has_le E] (f : ι → Ω → E) (ℱ : filtration ι m0) (μ : measure Ω . volume_tac) : + Prop := adapted ℱ f ∧ (∀ i j, i ≤ j → f i ≤ᵐ[μ] μ[f j | ℱ i]) ∧ ∀ i, integrable (f i) μ lemma martingale_const (ℱ : filtration ι m0) (μ : measure Ω) [is_finite_measure μ] (x : E) : diff --git a/src/probability/martingale/borel_cantelli.lean b/src/probability/martingale/borel_cantelli.lean index 1af4452345f8f..ba26caf71b6ad 100644 --- a/src/probability/martingale/borel_cantelli.lean +++ b/src/probability/martingale/borel_cantelli.lean @@ -297,7 +297,7 @@ lemma adapted_process (hs : ∀ n, measurable_set[ℱ n] (s n)) : ℱ.mono (finset.mem_range.1 hk) _ $ hs _ lemma martingale_part_process_ae_eq (ℱ : filtration ℕ m0) (μ : measure Ω) (s : ℕ → set Ω) (n : ℕ) : - martingale_part ℱ μ (process s) n = + martingale_part (process s) ℱ μ n = ∑ k in finset.range n, ((s (k + 1)).indicator 1 - μ[(s (k + 1)).indicator 1 | ℱ k]) := begin simp only [martingale_part_eq_sum, process_zero, zero_add], @@ -306,7 +306,7 @@ begin end lemma predictable_part_process_ae_eq (ℱ : filtration ℕ m0) (μ : measure Ω) (s : ℕ → set Ω) (n : ℕ) : - predictable_part ℱ μ (process s) n = + predictable_part (process s) ℱ μ n = ∑ k in finset.range n, μ[(s (k + 1)).indicator (1 : Ω → ℝ) | ℱ k] := begin have := martingale_part_process_ae_eq ℱ μ s n, @@ -339,7 +339,7 @@ lemma tendsto_sum_indicator_at_top_iff [is_finite_measure μ] (hf : adapted ℱ f) (hint : ∀ n, integrable (f n) μ) (hbdd : ∀ᵐ ω ∂μ, ∀ n, |f (n + 1) ω - f n ω| ≤ R) : ∀ᵐ ω ∂μ, tendsto (λ n, f n ω) at_top at_top ↔ - tendsto (λ n, predictable_part ℱ μ f n ω) at_top at_top := + tendsto (λ n, predictable_part f ℱ μ n ω) at_top at_top := begin have h₁ := (martingale_martingale_part hf hint).ae_not_tendsto_at_top_at_top (martingale_part_bdd_difference ℱ hbdd), diff --git a/src/probability/martingale/centering.lean b/src/probability/martingale/centering.lean index 7bda4b917d673..dd3190e9ebaec 100644 --- a/src/probability/martingale/centering.lean +++ b/src/probability/martingale/centering.lean @@ -12,20 +12,20 @@ import probability.martingale.basic Any `ℕ`-indexed stochastic process which is adapted and integrable can be written as the sum of a martingale and a predictable process. This result is also known as **Doob's decomposition theorem**. From a process `f`, a filtration `ℱ` and a measure `μ`, we define two processes -`martingale_part ℱ μ f` and `predictable_part ℱ μ f`. +`martingale_part f ℱ μ` and `predictable_part f ℱ μ`. ## Main definitions -* `measure_theory.predictable_part ℱ μ f`: a predictable process such that - `f = predictable_part ℱ μ f + martingale_part ℱ μ f` -* `measure_theory.martingale_part ℱ μ f`: a martingale such that - `f = predictable_part ℱ μ f + martingale_part ℱ μ f` +* `measure_theory.predictable_part f ℱ μ`: a predictable process such that + `f = predictable_part f ℱ μ + martingale_part f ℱ μ` +* `measure_theory.martingale_part f ℱ μ`: a martingale such that + `f = predictable_part f ℱ μ + martingale_part f ℱ μ` ## Main statements -* `measure_theory.adapted_predictable_part`: `(λ n, predictable_part ℱ μ f (n+1))` is adapted. That +* `measure_theory.adapted_predictable_part`: `(λ n, predictable_part f ℱ μ (n+1))` is adapted. That is, `predictable_part` is predictable. -* `measure_theory.martingale_martingale_part`: `martingale_part ℱ μ f` is a martingale. +* `measure_theory.martingale_martingale_part`: `martingale_part f ℱ μ` is a martingale. -/ @@ -43,17 +43,17 @@ variables {Ω E : Type*} {m0 : measurable_space Ω} {μ : measure Ω} process. This is the predictable process. See `martingale_part` for the martingale. -/ noncomputable def predictable_part {m0 : measurable_space Ω} - (ℱ : filtration ℕ m0) (μ : measure Ω) (f : ℕ → Ω → E) : ℕ → Ω → E := + (f : ℕ → Ω → E) (ℱ : filtration ℕ m0) (μ : measure Ω . volume_tac) : ℕ → Ω → E := λ n, ∑ i in finset.range n, μ[f (i+1) - f i | ℱ i] -@[simp] lemma predictable_part_zero : predictable_part ℱ μ f 0 = 0 := +@[simp] lemma predictable_part_zero : predictable_part f ℱ μ 0 = 0 := by simp_rw [predictable_part, finset.range_zero, finset.sum_empty] -lemma adapted_predictable_part : adapted ℱ (λ n, predictable_part ℱ μ f (n+1)) := +lemma adapted_predictable_part : adapted ℱ (λ n, predictable_part f ℱ μ (n+1)) := λ n, finset.strongly_measurable_sum' _ (λ i hin, strongly_measurable_condexp.mono (ℱ.mono (finset.mem_range_succ_iff.mp hin))) -lemma adapted_predictable_part' : adapted ℱ (λ n, predictable_part ℱ μ f n) := +lemma adapted_predictable_part' : adapted ℱ (λ n, predictable_part f ℱ μ n) := λ n, finset.strongly_measurable_sum' _ (λ i hin, strongly_measurable_condexp.mono (ℱ.mono (finset.mem_range_le hin))) @@ -61,15 +61,15 @@ lemma adapted_predictable_part' : adapted ℱ (λ n, predictable_part ℱ μ f n process. This is the martingale. See `predictable_part` for the predictable process. -/ noncomputable def martingale_part {m0 : measurable_space Ω} - (ℱ : filtration ℕ m0) (μ : measure Ω) (f : ℕ → Ω → E) : ℕ → Ω → E := -λ n, f n - predictable_part ℱ μ f n + (f : ℕ → Ω → E) (ℱ : filtration ℕ m0) (μ : measure Ω . volume_tac) : ℕ → Ω → E := +λ n, f n - predictable_part f ℱ μ n lemma martingale_part_add_predictable_part (ℱ : filtration ℕ m0) (μ : measure Ω) (f : ℕ → Ω → E) : - martingale_part ℱ μ f + predictable_part ℱ μ f = f := + martingale_part f ℱ μ + predictable_part f ℱ μ = f := sub_add_cancel _ _ lemma martingale_part_eq_sum : - martingale_part ℱ μ f = + martingale_part f ℱ μ = λ n, f 0 + ∑ i in finset.range n, (f (i+1) - f i - μ[f (i+1) - f i | ℱ i]) := begin rw [martingale_part, predictable_part], @@ -78,11 +78,11 @@ begin end lemma adapted_martingale_part (hf : adapted ℱ f) : - adapted ℱ (martingale_part ℱ μ f) := + adapted ℱ (martingale_part f ℱ μ) := adapted.sub hf adapted_predictable_part' lemma integrable_martingale_part (hf_int : ∀ n, integrable (f n) μ) (n : ℕ) : - integrable (martingale_part ℱ μ f n) μ := + integrable (martingale_part f ℱ μ n) μ := begin rw martingale_part_eq_sum, exact (hf_int 0).add @@ -91,11 +91,11 @@ end lemma martingale_martingale_part (hf : adapted ℱ f) (hf_int : ∀ n, integrable (f n) μ) [sigma_finite_filtration μ ℱ] : - martingale (martingale_part ℱ μ f) ℱ μ := + martingale (martingale_part f ℱ μ) ℱ μ := begin refine ⟨adapted_martingale_part hf, λ i j hij, _⟩, - -- ⊢ μ[martingale_part ℱ μ f j | ℱ i] =ᵐ[μ] martingale_part ℱ μ f i - have h_eq_sum : μ[martingale_part ℱ μ f j | ℱ i] + -- ⊢ μ[martingale_part f ℱ μ j | ℱ i] =ᵐ[μ] martingale_part f ℱ μ i + have h_eq_sum : μ[martingale_part f ℱ μ j | ℱ i] =ᵐ[μ] f 0 + ∑ k in finset.range j, (μ[f (k+1) - f k | ℱ i] - μ[μ[f (k+1) - f k | ℱ k] | ℱ i]), { rw martingale_part_eq_sum, refine (condexp_add (hf_int 0) _).trans _, @@ -140,11 +140,11 @@ end lemma martingale_part_add_ae_eq [sigma_finite_filtration μ ℱ] {f g : ℕ → Ω → E} (hf : martingale f ℱ μ) (hg : adapted ℱ (λ n, g (n + 1))) (hg0 : g 0 = 0) (hgint : ∀ n, integrable (g n) μ) (n : ℕ) : - martingale_part ℱ μ (f + g) n =ᵐ[μ] f n := + martingale_part (f + g) ℱ μ n =ᵐ[μ] f n := begin - set h := f - martingale_part ℱ μ (f + g) with hhdef, - have hh : h = predictable_part ℱ μ (f + g) - g, - { rw [hhdef, sub_eq_sub_iff_add_eq_add, add_comm (predictable_part ℱ μ (f + g)), + set h := f - martingale_part (f + g) ℱ μ with hhdef, + have hh : h = predictable_part (f + g) ℱ μ - g, + { rw [hhdef, sub_eq_sub_iff_add_eq_add, add_comm (predictable_part (f + g) ℱ μ), martingale_part_add_predictable_part] }, have hhpred : adapted ℱ (λ n, h (n + 1)), { rw hh, @@ -162,7 +162,7 @@ end lemma predictable_part_add_ae_eq [sigma_finite_filtration μ ℱ] {f g : ℕ → Ω → E} (hf : martingale f ℱ μ) (hg : adapted ℱ (λ n, g (n + 1))) (hg0 : g 0 = 0) (hgint : ∀ n, integrable (g n) μ) (n : ℕ) : - predictable_part ℱ μ (f + g) n =ᵐ[μ] g n := + predictable_part (f + g) ℱ μ n =ᵐ[μ] g n := begin filter_upwards [martingale_part_add_ae_eq hf hg hg0 hgint n] with ω hω, rw ← add_right_inj (f n ω), @@ -175,7 +175,7 @@ section difference lemma predictable_part_bdd_difference {R : ℝ≥0} {f : ℕ → Ω → ℝ} (ℱ : filtration ℕ m0) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : - ∀ᵐ ω ∂μ, ∀ i, |predictable_part ℱ μ f (i + 1) ω - predictable_part ℱ μ f i ω| ≤ R := + ∀ᵐ ω ∂μ, ∀ i, |predictable_part f ℱ μ (i + 1) ω - predictable_part f ℱ μ i ω| ≤ R := begin simp_rw [predictable_part, finset.sum_apply, finset.sum_range_succ_sub_sum], exact ae_all_iff.2 (λ i, ae_bdd_condexp_of_ae_bdd $ ae_all_iff.1 hbdd i), @@ -183,12 +183,12 @@ end lemma martingale_part_bdd_difference {R : ℝ≥0} {f : ℕ → Ω → ℝ} (ℱ : filtration ℕ m0) (hbdd : ∀ᵐ ω ∂μ, ∀ i, |f (i + 1) ω - f i ω| ≤ R) : - ∀ᵐ ω ∂μ, ∀ i, |martingale_part ℱ μ f (i + 1) ω - martingale_part ℱ μ f i ω| ≤ ↑(2 * R) := + ∀ᵐ ω ∂μ, ∀ i, |martingale_part f ℱ μ (i + 1) ω - martingale_part f ℱ μ i ω| ≤ ↑(2 * R) := begin filter_upwards [hbdd, predictable_part_bdd_difference ℱ hbdd] with ω hω₁ hω₂ i, simp only [two_mul, martingale_part, pi.sub_apply], - have : |f (i + 1) ω - predictable_part ℱ μ f (i + 1) ω - (f i ω - predictable_part ℱ μ f i ω)| = - |(f (i + 1) ω - f i ω) - (predictable_part ℱ μ f (i + 1) ω - predictable_part ℱ μ f i ω)|, + have : |f (i + 1) ω - predictable_part f ℱ μ (i + 1) ω - (f i ω - predictable_part f ℱ μ i ω)| = + |(f (i + 1) ω - f i ω) - (predictable_part f ℱ μ (i + 1) ω - predictable_part f ℱ μ i ω)|, { ring_nf }, -- `ring` suggests `ring_nf` despite proving the goal rw this, exact (abs_sub _ _).trans (add_le_add (hω₁ i) (hω₂ i)), From d0f92493cf6417acd8c853d0d7e22a2452b0b3a1 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 27 Sep 2022 17:24:13 +0000 Subject: [PATCH 416/828] feat(data/list/chain): add `list.chain'_is_infix` (#16627) --- src/data/list/chain.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/data/list/chain.lean b/src/data/list/chain.lean index 05e77e86c9f13..24b7a0bcc359e 100644 --- a/src/data/list/chain.lean +++ b/src/data/list/chain.lean @@ -187,6 +187,12 @@ theorem chain'.iff_mem : ∀ {l : list α}, chain' R l ↔ chain' (λ x y, x ∈ @[simp] theorem chain'_cons {x y l} : chain' R (x :: y :: l) ↔ R x y ∧ chain' R (y :: l) := chain_cons +theorem chain'_is_infix : ∀ l : list α, chain' (λ x y, [x, y] <:+: l) l +| [] := chain'_nil +| [a] := chain'_singleton _ +| (a::b::l) := chain'_cons.2 ⟨⟨[], l, by simp⟩, + (chain'_is_infix (b::l)).imp $ λ x y h, h.trans ⟨[a], [], by simp⟩⟩ + theorem chain'_split {a : α} : ∀ {l₁ l₂ : list α}, chain' R (l₁ ++ a :: l₂) ↔ chain' R (l₁ ++ [a]) ∧ chain' R (a :: l₂) | [] l₂ := (and_iff_right (chain'_singleton a)).symm From c433ac053fbbc11977462b21aa9cfd1b3594da40 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Tue, 27 Sep 2022 17:24:14 +0000 Subject: [PATCH 417/828] feat(analysis/calculus): add lemma for norm of zeroth iterated derivative (#16631) --- src/analysis/calculus/cont_diff.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 98f8bed33ad57..04f2c3399ab00 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -1437,6 +1437,13 @@ variable {𝕜} lemma iterated_fderiv_zero_eq_comp : iterated_fderiv 𝕜 0 f = (continuous_multilinear_curry_fin0 𝕜 E F).symm ∘ f := rfl +lemma norm_iterated_fderiv_zero : + ∥iterated_fderiv 𝕜 0 f x∥ = ∥f x∥ := +begin + rw [←continuous_multilinear_map.fin0_apply_norm, iterated_fderiv_zero_apply], + exact fin.elim0', +end + lemma iterated_fderiv_succ_apply_left {n : ℕ} (m : fin (n + 1) → E): (iterated_fderiv 𝕜 (n + 1) f x : (fin (n + 1) → E) → F) m = (fderiv 𝕜 (iterated_fderiv 𝕜 n f) x : E → (E [×n]→L[𝕜] F)) (m 0) (tail m) := rfl From 3bfcb4ce940082cecb21adde863b377393b55ccf Mon Sep 17 00:00:00 2001 From: mcdoll Date: Tue, 27 Sep 2022 17:24:15 +0000 Subject: [PATCH 418/828] feat(analysis/locally_convex): boundedness in normed spaces (#16636) Co-authored-by: Moritz Doll --- src/analysis/locally_convex/bounded.lean | 10 +++++++++- 1 file changed, 9 insertions(+), 1 deletion(-) diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index dfdb88768c29a..d14d5a29d2f6b 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -31,7 +31,7 @@ von Neumann-bounded sets. -/ -variables {𝕜 E F ι : Type*} +variables {𝕜 E E' F ι : Type*} open filter open_locale topological_space pointwise @@ -271,6 +271,14 @@ begin { exact λ ⟨C, hC⟩, (is_vonN_bounded_closed_ball 𝕜 E C).subset hC } end +lemma is_vonN_bounded_iff' (s : set E) : + bornology.is_vonN_bounded 𝕜 s ↔ ∃ r : ℝ, ∀ (x : E) (hx : x ∈ s), ∥x∥ ≤ r := +by rw [normed_space.is_vonN_bounded_iff, ←metric.bounded_iff_is_bounded, bounded_iff_forall_norm_le] + +lemma image_is_vonN_bounded_iff (f : E' → E) (s : set E') : + bornology.is_vonN_bounded 𝕜 (f '' s) ↔ ∃ r : ℝ, ∀ (x : E') (hx : x ∈ s), ∥f x∥ ≤ r := +by simp_rw [is_vonN_bounded_iff', set.ball_image_iff] + /-- In a normed space, the von Neumann bornology (`bornology.vonN_bornology`) is equal to the metric bornology. -/ lemma vonN_bornology_eq : bornology.vonN_bornology 𝕜 E = pseudo_metric_space.to_bornology := From 6f9edf5bc334c0d3b279ecbd1fd464a0fbc45188 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 27 Sep 2022 17:24:16 +0000 Subject: [PATCH 419/828] fix(tactic/positivity + test): instantiate meta-variables and add a test (#16647) Fix an issue with `positivity` reported by @YaelDillies [here](https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/New.20tactic.3A.20.60positivity.60/near/300639970). For the fix, it seems that instantiating meta-variables is enough. --- src/tactic/positivity.lean | 2 +- test/positivity.lean | 9 +++++++++ 2 files changed, 10 insertions(+), 1 deletion(-) diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index 1827ade0ca6de..dd495add2ede2 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -179,7 +179,7 @@ example {b : ℤ} : 0 ≤ max (-3) (b ^ 2) := by positivity ``` -/ meta def positivity : tactic unit := focus1 $ do - t ← target, + t ← target >>= instantiate_mvars, (rel_desired, a) ← match t with | `(0 ≤ %%e₂) := pure (ff, e₂) | `(%%e₂ ≥ 0) := pure (ff, e₂) diff --git a/test/positivity.lean b/test/positivity.lean index 2fb84f91783f4..7741b9493e608 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -14,6 +14,15 @@ import tactic.positivity This tactic proves goals of the form `0 ≤ a` and `0 < a`. -/ +/- Test for instantiating meta-variables. Reported on +https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/New.20tactic.3A.20.60positivity.60/near/300639970 +-/ +example : 0 ≤ 0 := +begin + apply le_trans _ le_rfl, + positivity, +end + open_locale ennreal nnrat nnreal universe u From a652850a283c4eeaefee07f39c728d80e427544e Mon Sep 17 00:00:00 2001 From: Mark Lavrentyev Date: Tue, 27 Sep 2022 17:24:17 +0000 Subject: [PATCH 420/828] feat(algebra/group/type_tags): finite & infinite instances for additive/multiplicative group type tags (#16662) - adds finite & infinite instances for additive/multiplicative group type tags --- src/algebra/group/type_tags.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/algebra/group/type_tags.lean b/src/algebra/group/type_tags.lean index 5ee1bd499e447..68fcbe0feda94 100644 --- a/src/algebra/group/type_tags.lean +++ b/src/algebra/group/type_tags.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import algebra.hom.group import logic.equiv.basic +import data.finite.defs /-! # Type tags that turn additive structures into multiplicative, and vice versa @@ -69,6 +70,12 @@ end multiplicative instance [inhabited α] : inhabited (additive α) := ⟨additive.of_mul default⟩ instance [inhabited α] : inhabited (multiplicative α) := ⟨multiplicative.of_add default⟩ +instance [finite α] : finite (additive α) := finite.of_equiv α (by refl) +instance [finite α] : finite (multiplicative α) := finite.of_equiv α (by refl) + +instance [infinite α] : infinite (additive α) := by tauto +instance [infinite α] : infinite (multiplicative α) := by tauto + instance [nontrivial α] : nontrivial (additive α) := additive.of_mul.injective.nontrivial From ec74a7cc52199e66a564a3c0de2f04a74700daa8 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 27 Sep 2022 17:24:18 +0000 Subject: [PATCH 421/828] feat(topology/algebra/[uniform_]group): more `ext` lemmas (#16667) --- src/topology/algebra/group.lean | 6 ++++++ src/topology/algebra/uniform_group.lean | 15 +++++++++++++++ 2 files changed, 21 insertions(+) diff --git a/src/topology/algebra/group.lean b/src/topology/algebra/group.lean index 946bc3416240a..96f6b41799c14 100644 --- a/src/topology/algebra/group.lean +++ b/src/topology/algebra/group.lean @@ -637,6 +637,12 @@ lemma topological_group.ext {G : Type*} [group G] {t t' : topological_space G} eq_of_nhds_eq_nhds $ λ x, by rw [← @nhds_translation_mul_inv G t _ _ x , ← @nhds_translation_mul_inv G t' _ _ x , ← h] +@[to_additive] +lemma topological_group.ext_iff {G : Type*} [group G] {t t' : topological_space G} + (tg : @topological_group G t _) (tg' : @topological_group G t' _) : + t = t' ↔ @nhds G t 1 = @nhds G t' 1 := +⟨λ h, h ▸ rfl, tg.ext tg'⟩ + @[to_additive] lemma topological_group.of_nhds_aux {G : Type*} [group G] [topological_space G] (hinv : tendsto (λ (x : G), x⁻¹) (𝓝 1) (𝓝 1)) diff --git a/src/topology/algebra/uniform_group.lean b/src/topology/algebra/uniform_group.lean index eaba7aa396336..22a5d140f3d09 100644 --- a/src/topology/algebra/uniform_group.lean +++ b/src/topology/algebra/uniform_group.lean @@ -209,6 +209,21 @@ end 𝓤 α = comap (λx:α×α, x.1 / x.2) (𝓝 (1:α)) := by { rw [← comap_swap_uniformity, uniformity_eq_comap_nhds_one, comap_comap, (∘)], refl } +@[to_additive] lemma uniform_group.ext {G : Type*} [group G] {u v : uniform_space G} + (hu : @uniform_group G u _) (hv : @uniform_group G v _) + (h : @nhds _ u.to_topological_space 1 = @nhds _ v.to_topological_space 1) : + u = v := +begin + refine uniform_space_eq _, + change @uniformity _ u = @uniformity _ v, + rw [@uniformity_eq_comap_nhds_one _ u _ hu, @uniformity_eq_comap_nhds_one _ v _ hv, h] +end + +@[to_additive] lemma uniform_group.ext_iff {G : Type*} [group G] {u v : uniform_space G} + (hu : @uniform_group G u _) (hv : @uniform_group G v _) : + u = v ↔ @nhds _ u.to_topological_space 1 = @nhds _ v.to_topological_space 1 := +⟨λ h, h ▸ rfl, hu.ext hv⟩ + variables {α} @[to_additive] theorem uniform_group.uniformity_countably_generated From 99e885b90b6593d27d62f61e44fac7a6cb50b113 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Tue, 27 Sep 2022 20:24:07 +0000 Subject: [PATCH 422/828] chore(ring_theory/(mv_)polynomial/{symmetric, homogeneous}): move files (#16414) --- src/ring_theory/{polynomial => mv_polynomial}/homogeneous.lean | 0 src/ring_theory/{polynomial => mv_polynomial}/symmetric.lean | 0 src/ring_theory/polynomial/vieta.lean | 2 +- 3 files changed, 1 insertion(+), 1 deletion(-) rename src/ring_theory/{polynomial => mv_polynomial}/homogeneous.lean (100%) rename src/ring_theory/{polynomial => mv_polynomial}/symmetric.lean (100%) diff --git a/src/ring_theory/polynomial/homogeneous.lean b/src/ring_theory/mv_polynomial/homogeneous.lean similarity index 100% rename from src/ring_theory/polynomial/homogeneous.lean rename to src/ring_theory/mv_polynomial/homogeneous.lean diff --git a/src/ring_theory/polynomial/symmetric.lean b/src/ring_theory/mv_polynomial/symmetric.lean similarity index 100% rename from src/ring_theory/polynomial/symmetric.lean rename to src/ring_theory/mv_polynomial/symmetric.lean diff --git a/src/ring_theory/polynomial/vieta.lean b/src/ring_theory/polynomial/vieta.lean index 9c4512d5551ed..2c7006fdcaf3c 100644 --- a/src/ring_theory/polynomial/vieta.lean +++ b/src/ring_theory/polynomial/vieta.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Hanting Zhang -/ import field_theory.splitting_field -import ring_theory.polynomial.symmetric +import ring_theory.mv_polynomial.symmetric /-! # Vieta's Formula From 8e8c7fd56d025de86d953dc008c8d2eea35a6312 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 27 Sep 2022 20:24:08 +0000 Subject: [PATCH 423/828] =?UTF-8?q?feat(analysis/normed=5Fspace/star/basic?= =?UTF-8?q?):=20in=20a=20C=E2=8B=86-ring=20`star=20x=20*=20x=20=3D=200=20?= =?UTF-8?q?=E2=86=94=20x=20=3D=200`=20(#16672)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds a few convenience lemmas for C⋆-rings regarding criteria for being equal, or not equal, to zero. --- src/analysis/normed_space/star/basic.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index 5051945464265..7675e92c88d35 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -110,6 +110,20 @@ by rw [norm_star_mul_self, norm_star] lemma nnnorm_star_mul_self {x : E} : ∥x⋆ * x∥₊ = ∥x∥₊ * ∥x∥₊ := subtype.ext norm_star_mul_self +@[simp] +lemma star_mul_self_eq_zero_iff (x : E) : star x * x = 0 ↔ x = 0 := +by { rw [←norm_eq_zero, norm_star_mul_self], exact mul_self_eq_zero.trans norm_eq_zero } + +lemma star_mul_self_ne_zero_iff (x : E) : star x * x ≠ 0 ↔ x ≠ 0 := +by simp only [ne.def, star_mul_self_eq_zero_iff] + +@[simp] +lemma mul_star_self_eq_zero_iff (x : E) : x * star x = 0 ↔ x = 0 := +by simpa only [star_eq_zero, star_star] using @star_mul_self_eq_zero_iff _ _ _ _ (star x) + +lemma mul_star_self_ne_zero_iff (x : E) : x * star x ≠ 0 ↔ x ≠ 0 := +by simp only [ne.def, mul_star_self_eq_zero_iff] + end non_unital section prod_pi From 999e092579418cc13547d204d57c8cec22adb27e Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Tue, 27 Sep 2022 22:00:34 +0000 Subject: [PATCH 424/828] feat(group_theory/finiteness): Add variants of `rank_closure_le_card` (#16364) This PR adds some more API lemmas for `group.rank`. `rank_congr` is useful since `rw h` and `simp only [h]` have run into trouble. --- src/group_theory/finiteness.lean | 59 ++++++++++++++++++++++++++++ src/group_theory/subgroup/basic.lean | 10 +++++ 2 files changed, 69 insertions(+) diff --git a/src/group_theory/finiteness.lean b/src/group_theory/finiteness.lean index 0e829a45dadf2..c0fd85f680ddf 100644 --- a/src/group_theory/finiteness.lean +++ b/src/group_theory/finiteness.lean @@ -9,6 +9,7 @@ import data.finset import group_theory.quotient_group import group_theory.submonoid.operations import group_theory.subgroup.basic +import set_theory.cardinal.finite /-! # Finitely generated monoids and groups @@ -161,6 +162,20 @@ lemma submonoid.powers_fg (r : M) : (submonoid.powers r).fg := instance monoid.powers_fg (r : M) : monoid.fg (submonoid.powers r) := (monoid.fg_iff_submonoid_fg _).mpr (submonoid.powers_fg r) +@[to_additive] instance monoid.closure_finset_fg (s : finset M) : + monoid.fg (submonoid.closure (s : set M)) := +begin + refine ⟨⟨s.preimage coe (subtype.coe_injective.inj_on _), _⟩⟩, + rw [finset.coe_preimage, submonoid.closure_closure_coe_preimage], +end + +@[to_additive] instance monoid.closure_finite_fg (s : set M) [finite s] : + monoid.fg (submonoid.closure s) := +begin + haveI := fintype.of_finite s, + exact s.coe_to_finset ▸ monoid.closure_finset_fg s.to_finset, +end + /-! ### Groups and subgroups -/ variables {G H : Type*} [group G] [add_group H] @@ -278,6 +293,20 @@ group.fg_iff_monoid.fg.mpr $ @monoid.fg_of_surjective G _ G' _ (group.fg_iff_mon instance group.fg_range {G' : Type*} [group G'] [group.fg G] (f : G →* G') : group.fg f.range := group.fg_of_surjective f.range_restrict_surjective +@[to_additive] instance group.closure_finset_fg (s : finset G) : + group.fg (subgroup.closure (s : set G)) := +begin + refine ⟨⟨s.preimage coe (subtype.coe_injective.inj_on _), _⟩⟩, + rw [finset.coe_preimage, ←subgroup.coe_subtype, subgroup.closure_preimage_eq_top], +end + +@[to_additive] instance group.closure_finite_fg (s : set G) [finite s] : + group.fg (subgroup.closure s) := +begin + haveI := fintype.of_finite s, + exact s.coe_to_finset ▸ group.closure_finset_fg s.to_finset, +end + variables (G) /-- The minimum number of generators of a group. -/ @@ -317,6 +346,36 @@ le_antisymm (group.rank_le_of_surjective f.symm f.symm.surjective) end group +namespace subgroup + +@[to_additive] lemma rank_congr {H K : subgroup G} [group.fg H] [group.fg K] (h : H = K) : + group.rank H = group.rank K := +by unfreezingI { subst h } + +@[to_additive] lemma rank_closure_finset_le_card (s : finset G) : + group.rank (closure (s : set G)) ≤ s.card := +begin + classical, + let t : finset (closure (s : set G)) := s.preimage coe (subtype.coe_injective.inj_on _), + have ht : closure (t : set (closure (s : set G))) = ⊤, + { rw finset.coe_preimage, + exact closure_preimage_eq_top s }, + apply (group.rank_le (closure (s : set G)) ht).trans, + rw [←finset.card_image_of_inj_on, finset.image_preimage], + { apply finset.card_filter_le }, + { apply subtype.coe_injective.inj_on }, +end + +@[to_additive] lemma rank_closure_finite_le_nat_card (s : set G) [finite s] : + group.rank (closure s) ≤ nat.card s := +begin + haveI := fintype.of_finite s, + rw [nat.card_eq_fintype_card, ←s.to_finset_card, ←rank_congr (congr_arg _ s.coe_to_finset)], + exact rank_closure_finset_le_card s.to_finset, +end + +end subgroup + section quotient_group @[to_additive] diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 8ae72589440c1..9272caefdd98c 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2387,6 +2387,16 @@ begin rwa [comap_map_eq, comap_map_eq, sup_of_le_left hH, sup_of_le_left hK] at hf, end +@[to_additive] lemma closure_preimage_eq_top (s : set G) : + closure ((closure s).subtype ⁻¹' s) = ⊤ := +begin + apply map_injective (show function.injective (closure s).subtype, from subtype.coe_injective), + rwa [monoid_hom.map_closure, ←monoid_hom.range_eq_map, subtype_range, + set.image_preimage_eq_of_subset], + rw [coe_subtype, subtype.range_coe_subtype], + exact subset_closure, +end + @[to_additive] lemma comap_sup_eq_of_le_range {H K : subgroup N} (hH : H ≤ f.range) (hK : K ≤ f.range) : comap f H ⊔ comap f K = comap f (H ⊔ K) := From c251522d50206cd5c5d8468839683a02e8b4b12a Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Tue, 27 Sep 2022 22:00:36 +0000 Subject: [PATCH 425/828] feat(order/atoms): link `set` to `is_atom` API (#16665) --- src/order/atoms.lean | 34 ++++++++++++++++++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/src/order/atoms.lean b/src/order/atoms.lean index 7435fec154fd2..5acd4f34dacb0 100644 --- a/src/order/atoms.lean +++ b/src/order/atoms.lean @@ -683,3 +683,37 @@ instance finite.to_is_atomic [partial_order α] [order_bot α] [finite α] : is_ is_coatomic_dual_iff_is_atomic.mp finite.to_is_coatomic end fintype + +namespace set + +lemma is_atom_singleton (x : α) : is_atom ({x} : set α) := +⟨(singleton_nonempty x).ne_empty, λ s hs, ssubset_singleton_iff.mp hs⟩ + +lemma is_atom_iff (s : set α) : is_atom s ↔ ∃ x, s = {x} := +begin + refine ⟨_, by { rintro ⟨x, rfl⟩, exact is_atom_singleton x }⟩, + rintro ⟨hs₁, hs₂⟩, + obtain ⟨x, hx⟩ := ne_empty_iff_nonempty.mp hs₁, + have := singleton_subset_iff.mpr hx, + refine ⟨x, subset.antisymm _ this⟩, + by_contra h, + exact (singleton_nonempty x).ne_empty (hs₂ {x} (ssubset_of_subset_not_subset this h)), +end + +lemma is_coatom_iff (s : set α) : is_coatom s ↔ ∃ x, s = {x}ᶜ := +by simp_rw [is_compl_compl.is_coatom_iff_is_atom, is_atom_iff, @eq_comm _ s, compl_eq_comm] + +lemma is_coatom_singleton_compl (x : α) : is_coatom ({x}ᶜ : set α) := +(is_coatom_iff {x}ᶜ).mpr ⟨x, rfl⟩ + +instance : is_atomistic (set α) := +{ eq_Sup_atoms := λ s, ⟨(λ x, {x}) '' s, + by rw [Sup_eq_sUnion, sUnion_image, bUnion_of_singleton], + by { rintro - ⟨x, hx, rfl⟩, exact is_atom_singleton x }⟩ } + +instance : is_coatomistic (set α) := +{ eq_Inf_coatoms := λ s, ⟨(λ x, {x}ᶜ) '' sᶜ, + by rw [Inf_eq_sInter, sInter_image, ←compl_Union₂, bUnion_of_singleton, compl_compl], + by { rintro - ⟨x, hx, rfl⟩, exact is_coatom_singleton_compl x }⟩ } + +end set From 6fe90f28236524f78871e2708ee33850581508fe Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Wed, 28 Sep 2022 05:29:41 +0000 Subject: [PATCH 426/828] feat(algebra/jordan/special): The symmetrization of an associative ring is a commutative Jordan multiplication (#11401) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A commutative multiplication on a real or complex space can be constructed from any multiplication by "symmetrisation" i.e ``` a∘b = 1/2(ab+ba). ``` When the original multiplication is associative, the symmetrised algebra is a commutative Jordan algebra. A commutative Jordan algebra which can be constructed in this way from an associative multiplication is said to be a special Jordan algebra. This PR shows more generally that for a ring where the scalar `2` is invertible, the symmetrised multiplication is a commutative Jordan multiplication. Co-authored-by: Christopher Hoskin --- src/algebra/jordan/basic.lean | 8 ++++---- src/algebra/symmetrized.lean | 31 ++++++++++++++++++++++++++++++- 2 files changed, 34 insertions(+), 5 deletions(-) diff --git a/src/algebra/jordan/basic.lean b/src/algebra/jordan/basic.lean index 4dd058d744d60..8e2ae3b1360c9 100644 --- a/src/algebra/jordan/basic.lean +++ b/src/algebra/jordan/basic.lean @@ -21,6 +21,10 @@ A more general concept of a (non-commutative) Jordan ring can also be defined, a (non-commutative, non-associative) ring `A` where, for each `a` in `A`, the operators of left and right multiplication by `a` and `a^2` commute. +Every associative algebra can be equipped with a symmetrized multiplication (characterized by +`sym_alg.sym_mul_sym`) making it into a commutative Jordan algebra (`sym_alg.is_comm_jordan`). +Jordan algebras arising this way are said to be special. + A real Jordan algebra `A` can be introduced by ```lean variables {A : Type*} [non_unital_non_assoc_ring A] [module ℝ A] [smul_comm_class ℝ A A] @@ -53,10 +57,6 @@ deep connections to mathematical physics, functional analysis and differential g information about these connections the interested reader is referred to [alfsenshultz2003], [chu2012], [friedmanscarr2005], [iordanescu2003] and [upmeier1987]. -Every associative algebra can be equipped with a symmetrized multiplication (characterized by -`sym_alg.sym_mul_sym`) making it into a commutative Jordan algebra. Jordan algebras arising this way -are said to be special. - There are also exceptional Jordan algebras which can be shown not to be the symmetrization of any associative algebra. The 3x3 matrices of octonions is the canonical example. diff --git a/src/algebra/symmetrized.lean b/src/algebra/symmetrized.lean index 3eccd1459a78b..f1016159ca8b7 100644 --- a/src/algebra/symmetrized.lean +++ b/src/algebra/symmetrized.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Christopher Hoskin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Christopher Hoskin -/ +import algebra.jordan.basic import algebra.module.basic -import tactic.abel /-! # Symmetrized algebra @@ -210,4 +210,33 @@ lemma mul_comm [has_mul α] [add_comm_semigroup α] [has_one α] [invertible (2 a * b = b * a := by rw [mul_def, mul_def, add_comm] + +instance [ring α] [invertible (2 : α)] : is_comm_jordan αˢʸᵐ := +{ mul_comm := sym_alg.mul_comm, + lmul_comm_rmul_rmul := λ a b, begin + -- Rearrange LHS + have commute_half_left := λ a : α, (commute.one_left a).bit0_left.inv_of_left.eq, + rw [mul_def, mul_def a b, unsym_sym, ← mul_assoc, ← commute_half_left (unsym (a*a)), mul_assoc, + mul_assoc, ← mul_add, ← mul_assoc, add_mul, mul_add (unsym (a * a)), ← add_assoc, ← mul_assoc, + ← mul_assoc], + + -- Rearrange RHS + nth_rewrite_rhs 0 [mul_def], + nth_rewrite_rhs 0 [mul_def], + nth_rewrite_rhs 2 [mul_def], + + rw [unsym_sym, sym_inj, ← mul_assoc, ← commute_half_left (unsym a), mul_assoc (⅟2) (unsym a), + mul_assoc (⅟2) _ (unsym a), ← mul_add, ← mul_assoc], + + nth_rewrite_rhs 0 mul_add (unsym a), + rw [add_mul, ← add_assoc, ← mul_assoc, ← mul_assoc], + + rw unsym_mul_self, + rw [← mul_assoc, ← mul_assoc, ← mul_assoc, ← mul_assoc, ← sub_eq_zero, ← mul_sub], + + convert mul_zero (⅟(2:α) * ⅟(2:α)), + rw [add_sub_add_right_eq_sub, add_assoc, add_assoc, add_sub_add_left_eq_sub, add_comm, + add_sub_add_right_eq_sub, sub_eq_zero], + end } + end sym_alg From 494b1f8e7c3f93318054b4778e42aa14c74b189d Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 28 Sep 2022 06:21:16 +0000 Subject: [PATCH 427/828] feat(linear_algebra/affine_space/affine_subspace): `parallel` (#16298) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the notion of two affine subspaces being parallel and set up associated API. This notion is very similar to the subspaces having equal `direction`, and in many cases I expect equality of direction will remain more convenient to use. However, the notions aren't exactly the same (the empty affine subspace and a single-point subspace have the same `direction` but aren't considered parallel), and having a definition of `parallel`, even if not used much, allows geometrical statements involving things being parallel to be made in a way corresponding more closely to the informal formulation. Note that the notation defined for `parallel` is based on what characters are available rather than on their proper Unicode semantics. There are at least four characters in Unicode with somewhat similar appearance: * U+01C1 LATIN LETTER LATERAL CLICK `ǁ` * U+2016 DOUBLE VERTICAL LINE `‖` (for which the Unicode Character Database says "used in pairs to indicate norm of a matrix") * U+2225 PARALLEL TO `∥` * U+23F8 DOUBLE VERTICAL BAR `⏸` Based on the Unicode descriptions, U+2225 would be natural to use for `parallel`, and U+2016 for norms. However, mathlib makes extensive use of U+2225 for norms; for geometry, both norms and `parallel` are of use and it doesn't work well to try to use the notation for both (unless there's some way to set the precedences of the different notations that will make Lean parse both uses correctly). (There's a local notation using U+2225 for `fuzzy` in the context of games, which seems a more appropriate use of U+2225 and probably doesn't cause problems because of games and norms not being used together.) So this PR uses U+2016 for `parallel` instead of the logical U+2225 (even if in principle the uses for `parallel` and norms should be swapped to correspond better to the Unicode semantics). There are other local uses of U+2016 for `fintype.card` in a few places, but I don't expect those to cause a problem. --- .../affine_space/affine_subspace.lean | 84 +++++++++++++++++++ 1 file changed, 84 insertions(+) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index e5f72dfcde09e..062a94428f653 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -23,6 +23,8 @@ This file defines affine subspaces (over modules) and the affine span of a set o various lemmas relating to the set of vectors in the `direction`, and relating the lattice structure on affine subspaces to that on their directions. +* `affine_subspace.parallel`, notation `‖`, gives the property of two affine subspaces being + parallel (one being a translate of the other). * `affine_span` gives the affine subspace spanned by a set of points, with `vector_span` giving its direction. `affine_span` is defined in terms of `span_points`, which gives an explicit description of @@ -1418,3 +1420,85 @@ lemma comap_supr {ι : Sort*} (f : P₁ →ᵃ[k] P₂) (s : ι → affine_subsp end affine_subspace end map_comap + +namespace affine_subspace + +open affine_equiv + +variables {k : Type*} {V : Type*} {P : Type*} [ring k] [add_comm_group V] [module k V] +variables [affine_space V P] +include V + +/-- Two affine subspaces are parallel if one is related to the other by adding the same vector +to all points. -/ +def parallel (s₁ s₂ : affine_subspace k P) : Prop := +∃ v : V, s₂ = s₁.map (const_vadd k P v) + +/- The notation should logically be U+2225 PARALLEL TO, but that is used globally for norms at +present, and norms and parallelism are both widely used in geometry, so use U+2016 DOUBLE +VERTICAL LINE (which is logically more appropriate for norms) instead here to avoid conflict. -/ +localized "infix (name := affine_subspace.parallel) ` ‖ `:50 := affine_subspace.parallel" in affine + +@[symm] lemma parallel.symm {s₁ s₂ : affine_subspace k P} (h : s₁ ‖ s₂) : s₂ ‖ s₁ := +begin + rcases h with ⟨v, rfl⟩, + refine ⟨-v, _⟩, + rw [map_map, ←coe_trans_to_affine_map, ←const_vadd_add, neg_add_self, const_vadd_zero, + coe_refl_to_affine_map, map_id] +end + +lemma parallel_comm {s₁ s₂ : affine_subspace k P} : s₁ ‖ s₂ ↔ s₂ ‖ s₁ := +⟨parallel.symm, parallel.symm⟩ + +@[refl] lemma parallel.refl (s : affine_subspace k P) : s ‖ s := +⟨0, by simp⟩ + +@[trans] lemma parallel.trans {s₁ s₂ s₃ : affine_subspace k P} (h₁₂ : s₁ ‖ s₂) (h₂₃ : s₂ ‖ s₃) : + s₁ ‖ s₃ := +begin + rcases h₁₂ with ⟨v₁₂, rfl⟩, + rcases h₂₃ with ⟨v₂₃, rfl⟩, + refine ⟨v₂₃ + v₁₂, _⟩, + rw [map_map, ←coe_trans_to_affine_map, ←const_vadd_add] +end + +lemma parallel.direction_eq {s₁ s₂ : affine_subspace k P} (h : s₁ ‖ s₂) : + s₁.direction = s₂.direction := +begin + rcases h with ⟨v, rfl⟩, + simp +end + +@[simp] lemma parallel_bot_iff_eq_bot {s : affine_subspace k P} : + s ‖ ⊥ ↔ s = ⊥ := +begin + refine ⟨λ h, _, λ h, h ▸ parallel.refl _⟩, + rcases h with ⟨v, h⟩, + rwa [eq_comm, map_eq_bot_iff] at h +end + +@[simp] lemma bot_parallel_iff_eq_bot {s : affine_subspace k P} : + ⊥ ‖ s ↔ s = ⊥ := +by rw [parallel_comm, parallel_bot_iff_eq_bot] + +lemma parallel_iff_direction_eq_and_eq_bot_iff_eq_bot {s₁ s₂ : affine_subspace k P} : + s₁ ‖ s₂ ↔ s₁.direction = s₂.direction ∧ (s₁ = ⊥ ↔ s₂ = ⊥) := +begin + refine ⟨λ h, ⟨h.direction_eq, _, _⟩, λ h, _⟩, + { rintro rfl, exact bot_parallel_iff_eq_bot.1 h }, + { rintro rfl, exact parallel_bot_iff_eq_bot.1 h }, + { rcases h with ⟨hd, hb⟩, + by_cases hs₁ : s₁ = ⊥, + { rw [hs₁, bot_parallel_iff_eq_bot], + exact hb.1 hs₁ }, + { have hs₂ : s₂ ≠ ⊥ := hb.not.1 hs₁, + rcases (nonempty_iff_ne_bot s₁).2 hs₁ with ⟨p₁, hp₁⟩, + rcases (nonempty_iff_ne_bot s₂).2 hs₂ with ⟨p₂, hp₂⟩, + refine ⟨p₂ -ᵥ p₁, (eq_iff_direction_eq_of_mem hp₂ _).2 _⟩, + { rw mem_map, + refine ⟨p₁, hp₁, _⟩, + simp }, + { simpa using hd.symm } } } +end + +end affine_subspace From 8ef6d8f247cef05e19016bebcb7340fd8c962d4c Mon Sep 17 00:00:00 2001 From: mcdoll Date: Wed, 28 Sep 2022 07:45:13 +0000 Subject: [PATCH 428/828] feat(analysis/schwartz_space): add lemmas for seminorms (#16634) --- src/analysis/schwartz_space.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index 8e573be7fa024..853527ca6dd2f 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -333,6 +333,20 @@ lemma le_seminorm (k n : ℕ) (f : 𝓢(E, F)) (x : E) : ∥x∥ ^ k * ∥iterated_fderiv ℝ n f x∥ ≤ seminorm 𝕜 k n f := f.le_seminorm_aux k n x +lemma norm_iterated_fderiv_le_seminorm (f : 𝓢(E, F)) (n : ℕ) (x₀ : E): + ∥iterated_fderiv ℝ n f x₀∥ ≤ (schwartz_map.seminorm 𝕜 0 n) f := +begin + have := schwartz_map.le_seminorm 𝕜 0 n f x₀, + rwa [pow_zero, one_mul] at this, +end + +lemma norm_pow_mul_le_seminorm (f : 𝓢(E, F)) (k : ℕ) (x₀ : E): + ∥x₀∥^k * ∥f x₀∥ ≤ (schwartz_map.seminorm 𝕜 k 0) f := +begin + have := schwartz_map.le_seminorm 𝕜 k 0 f x₀, + rwa norm_iterated_fderiv_zero at this, +end + end seminorms section topology From 60851c931abde886a2731523cfc6ddb29b06d83f Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 28 Sep 2022 14:15:46 +0000 Subject: [PATCH 429/828] feat(topology/order/lattice): add `topological_lattice` instance for linear orders with an order closed topology (#16664) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds a `topological_lattice` instance for and linear order with an order closed topology. In particular, this gives us previously nonexistent `topological_lattice` instances for `ℝ≥0`, `ℝ≥0∞` and things like `set.Icc (a : ℝ) b` as well as many others. In addition, it makes the `topological_lattice` instance for `ℝ` available earlier in the import hierarchy, because the previously existing instance was derived from `real.normed_lattice_add_comm_group` via `normed_lattice_add_comm_group_topological_lattice`. --- src/topology/order/lattice.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/topology/order/lattice.lean b/src/topology/order/lattice.lean index e42a6d380704e..a8880feb4642c 100644 --- a/src/topology/order/lattice.lean +++ b/src/topology/order/lattice.lean @@ -63,6 +63,11 @@ instance order_dual.topological_lattice (L : Type*) [topological_space L] [lattice L] [topological_lattice L] : topological_lattice Lᵒᵈ := {} +@[priority 100] -- see Note [lower instance priority] +instance linear_order.topological_lattice {L : Type*} [topological_space L] [linear_order L] + [order_closed_topology L] : topological_lattice L := +{ continuous_inf := continuous_min, continuous_sup := continuous_max } + variables {L : Type*} [topological_space L] variables {X : Type*} [topological_space X] From 7247a3cce1d85855956c54f5fed27a10ae1737ac Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 28 Sep 2022 15:55:10 +0000 Subject: [PATCH 430/828] refactor(algebra/group/defs): `inv_one_class`, `neg_zero_class` (#16187) Define typeclasses `inv_one_class` and `neg_zero_class`, to allow results depending on those properties to be proved more generally than for `division_monoid` and `subtraction_monoid` without requiring duplication. Also define `div_inv_one_monoid` and `sub_neg_zero_monoid`. The only instances added here are those deduced from `division_monoid` and `subtraction_monoid`, and the only lemmas generalized to these classes were previously proved for those classes. Additional instances intended for followups include: * `neg_zero_class` for the combination of `mul_zero_class` with `has_distrib_neg`, so eliminating the separate `neg_zero'` lemma. * `sub_neg_zero_monoid` for `ereal`. * `div_inv_one_monoid` for `ennreal`. * The usual `pointwise`, `pi` and `prod` instances. Additional lemmas intended to be generalized to use these typeclasses in followups include: * `antiperiodic.nat_mul_eq_of_eq_zero` and `antiperiodic.int_mul_eq_of_eq_zero`, which currently require the codomain of the antiperiodic function to be a `subtraction_monoid`. (The latter will also require involutive `neg`, as will some lemmas about inequalities.) * Given appropriate typeclasses for the interaction of inequalities with `inv` and `neg` (which will also enabling combining `left` and `right` variants, which is the main motivation of these changes), lemmas such as `left.inv_le_one_iff`, `left.one_le_inv_iff`, `left.one_lt_inv_iff`, `left.inv_lt_one_iff` and their `right` and additive variants. Some of these currently have duplicates for `ennreal`, for example. Zulip thread raising question of such typeclasses: https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/typeclasses.20for.20orders.20with.20neg.20and.20inv --- src/algebra/group/basic.lean | 20 ++++++++++++++++---- src/algebra/group/defs.lean | 32 ++++++++++++++++++++++++++++++++ src/tactic/abel.lean | 2 +- 3 files changed, 49 insertions(+), 5 deletions(-) diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index 3ac8a14a628c5..097f8c2fd4ba0 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -244,6 +244,17 @@ by rw [div_eq_mul_inv, one_div] end div_inv_monoid +section div_inv_one_monoid +variables [div_inv_one_monoid G] + +@[simp, to_additive] lemma div_one (a : G) : a / 1 = a := +by simp [div_eq_mul_inv] + +@[to_additive] lemma one_div_one : (1 : G) / 1 = 1 := +div_one _ + +end div_inv_one_monoid + section division_monoid variables [division_monoid α] {a b c : α} @@ -275,12 +286,13 @@ variables (a b c) @[to_additive] lemma inv_div_left : a⁻¹ / b = (b * a)⁻¹ := by simp @[simp, to_additive] lemma inv_div : (a / b)⁻¹ = b / a := by simp @[simp, to_additive] lemma one_div_div : 1 / (a / b) = b / a := by simp -@[simp, to_additive] lemma inv_one : (1 : α)⁻¹ = 1 := -by simpa only [one_div, inv_inv] using (inv_div (1 : α) 1).symm -@[simp, to_additive] lemma div_one : a / 1 = a := by simp -@[to_additive] lemma one_div_one : (1 : α) / 1 = 1 := div_one _ @[to_additive] lemma one_div_one_div : 1 / (1 / a) = a := by simp +@[priority 100, to_additive] instance division_monoid.to_div_inv_one_monoid : + div_inv_one_monoid α := +{ inv_one := by simpa only [one_div, inv_inv] using (inv_div (1 : α) 1).symm, + ..division_monoid.to_div_inv_monoid α } + variables {a b c} @[simp, to_additive] lemma inv_eq_one : a⁻¹ = 1 ↔ a = 1 := inv_injective.eq_iff' inv_one diff --git a/src/algebra/group/defs.lean b/src/algebra/group/defs.lean index 91b2e8dc4a856..edcd8148d4736 100644 --- a/src/algebra/group/defs.lean +++ b/src/algebra/group/defs.lean @@ -659,6 +659,38 @@ alias div_eq_mul_inv ← division_def end div_inv_monoid +section inv_one_class + +set_option extends_priority 50 + +/-- Typeclass for expressing that `-0 = 0`. -/ +class neg_zero_class (G : Type*) extends has_zero G, has_neg G := +(neg_zero : -(0 : G) = 0) + +/-- A `sub_neg_monoid` where `-0 = 0`. -/ +class sub_neg_zero_monoid (G : Type*) extends sub_neg_monoid G, neg_zero_class G + +/-- Typeclass for expressing that `1⁻¹ = 1`. -/ +@[to_additive] +class inv_one_class (G : Type*) extends has_one G, has_inv G := +(inv_one : (1 : G)⁻¹ = 1) + +attribute [to_additive neg_zero_class.to_has_neg] inv_one_class.to_has_inv +attribute [to_additive neg_zero_class.to_has_zero] inv_one_class.to_has_one + +/-- A `div_inv_monoid` where `1⁻¹ = 1`. -/ +@[to_additive sub_neg_zero_monoid] +class div_inv_one_monoid (G : Type*) extends div_inv_monoid G, inv_one_class G + +attribute [to_additive sub_neg_zero_monoid.to_sub_neg_monoid] div_inv_one_monoid.to_div_inv_monoid +attribute [to_additive sub_neg_zero_monoid.to_neg_zero_class] div_inv_one_monoid.to_inv_one_class + +variables [inv_one_class G] + +@[simp, to_additive] lemma inv_one : (1 : G)⁻¹ = 1 := inv_one_class.inv_one + +end inv_one_class + /-- A `subtraction_monoid` is a `sub_neg_monoid` with involutive negation and such that `-(a + b) = -b + -a` and `a + b = 0 → -a = b`. -/ @[protect_proj, ancestor sub_neg_monoid has_involutive_neg] diff --git a/src/tactic/abel.lean b/src/tactic/abel.lean index 8f7581d983617..8e97c0ab87822 100644 --- a/src/tactic/abel.lean +++ b/src/tactic/abel.lean @@ -181,7 +181,7 @@ by simp [h₂.symm, h₁.symm, termg]; ac_refl meta def eval_neg (c : context) : normal_expr → tactic (normal_expr × expr) | (zero e) := do - p ← c.mk_app ``neg_zero ``subtraction_monoid [], + p ← c.mk_app ``neg_zero ``neg_zero_class [], return (zero' c, p) | (nterm e n x a) := do (n', h₁) ← mk_app ``has_neg.neg [n.1] >>= norm_num.eval_field, From 4dc217e4191837ed19eb17ac8dc1124d67f1bd22 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 28 Sep 2022 15:55:11 +0000 Subject: [PATCH 431/828] feat(algebra/order/floor): Positivity extensions for `floor` and `ceil` (#16635) Add two `positivity` extensions: * `positivity_floor` for `int.floor` * `positivity_ceil` for `nat.ceil`, `int.ceil` --- src/algebra/order/floor.lean | 45 ++++++++++++++++++++++++++++++++++++ test/positivity.lean | 7 ++++++ 2 files changed, 52 insertions(+) diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index eca1977d4102c..9b56e2ef2613c 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Kevin Kappelmann -/ import tactic.abel import tactic.linarith +import tactic.positivity /-! # Floor and ceil @@ -885,6 +886,7 @@ by simp_rw [round_eq, ←map_floor _ hf, map_add, one_div, map_inv₀, map_bit0, end int +section floor_ring_to_semiring variables {α} [linear_ordered_ring α] [floor_ring α] /-! #### A floor ring as a floor semiring -/ @@ -919,6 +921,8 @@ by { rw [←int.ceil_to_nat, int.to_nat_of_nonneg (int.ceil_nonneg ha)] } lemma nat.cast_ceil_eq_cast_int_ceil (ha : 0 ≤ a) : (⌈a⌉₊ : α) = ⌈a⌉ := by rw [←nat.cast_ceil_eq_int_ceil ha, int.cast_coe_nat] +end floor_ring_to_semiring + /-- There exists at most one `floor_ring` structure on a given linear ordered ring. -/ lemma subsingleton_floor_ring {α} [linear_ordered_ring α] : subsingleton (floor_ring α) := @@ -928,3 +932,44 @@ begin have : H₁.ceil = H₂.ceil := funext (λ a, H₁.gc_ceil_coe.l_unique H₂.gc_ceil_coe $ λ _, rfl), cases H₁, cases H₂, congr; assumption end + +namespace tactic +open positivity + +private lemma int_floor_nonneg [linear_ordered_ring α] [floor_ring α] {a : α} (ha : 0 ≤ a) : + 0 ≤ ⌊a⌋ := int.floor_nonneg.2 ha +private lemma int_floor_nonneg_of_pos [linear_ordered_ring α] [floor_ring α] {a : α} (ha : 0 < a) : + 0 ≤ ⌊a⌋ := int_floor_nonneg ha.le + +/-- Extension for the `positivity` tactic: `int.floor` is nonnegative if its input is. -/ +@[positivity] +meta def positivity_floor : expr → tactic strictness +| `(⌊%%a⌋) := do + strictness_a ← core a, + match strictness_a with + | positive p := nonnegative <$> mk_app ``int_floor_nonneg_of_pos [p] + | nonnegative p := nonnegative <$> mk_app ``int_floor_nonneg [p] + end +| e := pp e >>= fail ∘ format.bracket "The expression `" "` is not of the form `⌊a⌋`" + +private lemma nat_ceil_pos [linear_ordered_semiring α] [floor_semiring α] {a : α} : + 0 < a → 0 < ⌈a⌉₊ := nat.ceil_pos.2 +private lemma int_ceil_pos [linear_ordered_ring α] [floor_ring α] {a : α} : 0 < a → 0 < ⌈a⌉ := +int.ceil_pos.2 + +/-- Extension for the `positivity` tactic: `ceil` and `int.ceil` are positive/nonnegative if +their input is. -/ +@[positivity] +meta def positivity_ceil : expr → tactic strictness +| `(⌈%%a⌉₊) := do + positive p ← core a, -- We already know `0 ≤ n` for all `n : ℕ` + positive <$> mk_app ``nat_ceil_pos [p] +| `(⌈%%a⌉) := do + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``int_ceil_pos [p] + | nonnegative p := nonnegative <$> mk_app ``int.ceil_nonneg [p] + end +| e := pp e >>= fail ∘ format.bracket "The expression `" "` is not of the form `⌈a⌉₊` nor `⌈a⌉`" + +end tactic diff --git a/test/positivity.lean b/test/positivity.lean index 7741b9493e608..c2c7d29fa0cb0 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -91,6 +91,13 @@ example {a : ℝ≥0} {b : ℝ} (ha : 0 < a) : 0 < a ^ b := by positivity example {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 ≤ b) : 0 < a ^ b := by positivity example {a : ℝ≥0∞} {b : ℝ} (ha : 0 < a) (hb : 0 < b) : 0 < a ^ b := by positivity +example {a : ℝ} (ha : 0 < a) : 0 ≤ ⌊a⌋ := by positivity +example {a : ℝ} (ha : 0 ≤ a) : 0 ≤ ⌊a⌋ := by positivity + +example {a : ℝ} (ha : 0 < a) : 0 < ⌈a⌉₊ := by positivity +example {a : ℝ} (ha : 0 < a) : 0 < ⌈a⌉ := by positivity +example {a : ℝ} (ha : 0 ≤ a) : 0 ≤ ⌈a⌉ := by positivity + example {a : ℤ} (ha : 3 < a) : 0 ≤ a ^ 2 + a := by positivity example {a : ℤ} (ha : 3 < a) : 0 ≤ a ^ 3 + a := by positivity From 35f0f1ba7176254c4474e4358f4d0044832b9698 Mon Sep 17 00:00:00 2001 From: Wrenna Robson Date: Wed, 28 Sep 2022 18:56:28 +0000 Subject: [PATCH 432/828] feat(data/set/basic): Refactor and additions of sep lemmas (#16566) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add and refactor sep lemmas. Co-authored-by: Yaël Dillies <[yael.dillies@gmail.com](mailto:yael.dillies@gmail.com)> Co-authored-by: Wrenna Robson <34025592+linesthatinterlace@users.noreply.github.com> --- src/analysis/convex/quasiconvex.lean | 2 +- src/data/set/basic.lean | 60 ++++++++++++++++------------ src/data/set/finite.lean | 2 +- 3 files changed, 37 insertions(+), 27 deletions(-) diff --git a/src/analysis/convex/quasiconvex.lean b/src/analysis/convex/quasiconvex.lean index d36aa2622e957..a38845b99c915 100644 --- a/src/analysis/convex/quasiconvex.lean +++ b/src/analysis/convex/quasiconvex.lean @@ -93,7 +93,7 @@ lemma quasiconvex_on.sup (hf : quasiconvex_on 𝕜 s f) (hg : quasiconvex_on quasiconvex_on 𝕜 s (f ⊔ g) := begin intro r, - simp_rw [pi.sup_def, sup_le_iff, ←set.sep_inter_sep], + simp_rw [pi.sup_def, sup_le_iff, set.sep_and], exact (hf r).inter (hg r), end diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 22fa723578ec0..3a6274a4176a9 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -209,8 +209,6 @@ lemma set_of_bijective : bijective (set_of : (α → Prop) → set α) := biject @[simp] theorem set_of_subset_set_of {p q : α → Prop} : {a | p a} ⊆ {a | q a} ↔ (∀a, p a → q a) := iff.rfl -@[simp] lemma sep_set_of {p q : α → Prop} : {a ∈ {a | p a } | q a} = {a | p a ∧ q a} := rfl - lemma set_of_and {p q : α → Prop} : {a | p a ∧ q a} = {a | p a} ∩ {a | q a} := rfl lemma set_of_or {p q : α → Prop} : {a | p a ∨ q a} = {a | p a} ∪ {a | q a} := rfl @@ -817,40 +815,52 @@ eq_singleton_iff_unique_mem.trans $ and_congr_left $ λ H, ⟨λ h', ⟨_, h'⟩ /-! ### Lemmas about sets defined as `{x ∈ s | p x}`. -/ -theorem mem_sep {s : set α} {p : α → Prop} {x : α} (xs : x ∈ s) (px : p x) : x ∈ {x ∈ s | p x} := -⟨xs, px⟩ +section sep +variables {p q : α → Prop} {x : α} -@[simp] theorem sep_mem_eq {s t : set α} : {x ∈ s | x ∈ t} = s ∩ t := rfl +theorem mem_sep (xs : x ∈ s) (px : p x) : x ∈ {x ∈ s | p x} := ⟨xs, px⟩ -@[simp] theorem mem_sep_iff {s : set α} {p : α → Prop} {x : α} : x ∈ {x ∈ s | p x} ↔ x ∈ s ∧ p x := -iff.rfl +@[simp] theorem sep_mem_eq : {x ∈ s | x ∈ t} = s ∩ t := rfl + +@[simp] theorem mem_sep_iff : x ∈ {x ∈ s | p x} ↔ x ∈ s ∧ p x := iff.rfl + +theorem sep_ext_iff : {x ∈ s | p x} = {x ∈ s | q x} ↔ ∀ x ∈ s, (p x ↔ q x) := +by simp_rw [ext_iff, mem_sep_iff, and.congr_right_iff] -theorem eq_sep_of_subset {s t : set α} (h : s ⊆ t) : s = {x ∈ t | x ∈ s} := -(inter_eq_self_of_subset_right h).symm +theorem sep_eq_of_subset (h : s ⊆ t) : {x ∈ t | x ∈ s} = s := +inter_eq_self_of_subset_right h @[simp] theorem sep_subset (s : set α) (p : α → Prop) : {x ∈ s | p x} ⊆ s := λ x, and.left -@[simp] lemma sep_empty (p : α → Prop) : {x ∈ (∅ : set α) | p x} = ∅ := -by { ext, exact false_and _ } +@[simp] lemma sep_eq_self_iff_mem_true : {x ∈ s | p x} = s ↔ ∀ x ∈ s, p x := +by simp_rw [ext_iff, mem_sep_iff, and_iff_left_iff_imp] -theorem forall_not_of_sep_empty {s : set α} {p : α → Prop} (H : {x ∈ s | p x} = ∅) - (x) : x ∈ s → ¬ p x := not_and.1 (eq_empty_iff_forall_not_mem.1 H x : _) +@[simp] lemma sep_eq_empty_iff_mem_false : {x ∈ s | p x} = ∅ ↔ ∀ x ∈ s, ¬ p x := +by simp_rw [ext_iff, mem_sep_iff, mem_empty_iff_false, iff_false, not_and] -@[simp] lemma sep_univ {α} {p : α → Prop} : {a ∈ (univ : set α) | p a} = {a | p a} := univ_inter _ +@[simp] lemma sep_true : {x ∈ s | true} = s := inter_univ s -@[simp] lemma sep_true : {a ∈ s | true} = s := -by { ext, simp } +@[simp] lemma sep_false : {x ∈ s | false} = ∅ := inter_empty s -@[simp] lemma sep_false : {a ∈ s | false} = ∅ := -by { ext, simp } +@[simp] lemma sep_empty (p : α → Prop) : {x ∈ (∅ : set α) | p x} = ∅ := empty_inter p -lemma sep_inter_sep {p q : α → Prop} : - {x ∈ s | p x} ∩ {x ∈ s | q x} = {x ∈ s | p x ∧ q x} := -begin - ext, - simp_rw [mem_inter_iff, mem_sep_iff], - rw [and_and_and_comm, and_self], -end +@[simp] lemma sep_univ : {x ∈ (univ : set α) | p x} = {x | p x} := univ_inter p + +@[simp] lemma sep_union : {x ∈ s ∪ t | p x} = {x ∈ s | p x} ∪ {x ∈ t | p x} := +union_inter_distrib_right + +@[simp] lemma sep_inter : {x ∈ s ∩ t | p x} = {x ∈ s | p x} ∩ {x ∈ t | p x} := +inter_inter_distrib_right s t p + +@[simp] lemma sep_and : {x ∈ s | p x ∧ q x} = {x ∈ s | p x} ∩ {x ∈ s | q x} := +inter_inter_distrib_left s p q + +@[simp] lemma sep_or : {x ∈ s | p x ∨ q x} = {x ∈ s | p x} ∪ {x ∈ s | q x} := +inter_union_distrib_left + +@[simp] lemma sep_set_of : {x ∈ {y | p y} | q x} = {x | p x ∧ q x} := rfl + +end sep @[simp] lemma subset_singleton_iff {α : Type*} {s : set α} {x : α} : s ⊆ {x} ↔ ∀ y ∈ s, y = x := iff.rfl diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index b444ccac7fe5e..91ab26af7a1e3 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -387,7 +387,7 @@ instance finite_sep (s : set α) (p : α → Prop) [finite s] : by { casesI nonempty_fintype s, apply_instance } protected lemma subset (s : set α) {t : set α} [finite s] (h : t ⊆ s) : finite t := -by { rw eq_sep_of_subset h, apply_instance } +by { rw ←sep_eq_of_subset h, apply_instance } instance finite_inter_of_right (s t : set α) [finite t] : finite (s ∩ t : set α) := finite.set.subset t (inter_subset_right s t) From fe840598819df283a6066b12a65d1a480c25bf2f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 28 Sep 2022 21:58:08 +0000 Subject: [PATCH 433/828] feat(set_theory/game/pgame): ditch `restricted` (#15037) We ditch `pgame.restricted`. The docstring erroneously claimed it to be something other than what it was, and in its current form, it really only served as a worse form of `le_def`. It was barely used anyways. See also my comments on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Well-founded.20recursion.20for.20pgames/near/287531824). --- src/set_theory/game/pgame.lean | 38 +++++----------------------------- 1 file changed, 5 insertions(+), 33 deletions(-) diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index e9bf0c246ede7..1cd894b03d91b 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -674,29 +674,6 @@ end /-! ### Relabellings -/ -/-- `restricted x y` says that Left always has no more moves in `x` than in `y`, - and Right always has no more moves in `y` than in `x` -/ -inductive restricted : pgame.{u} → pgame.{u} → Type (u+1) -| mk : Π {x y : pgame} (L : x.left_moves → y.left_moves) (R : y.right_moves → x.right_moves), - (∀ i, restricted (x.move_left i) (y.move_left (L i))) → - (∀ j, restricted (x.move_right (R j)) (y.move_right j)) → restricted x y - -/-- The identity restriction. -/ -@[refl] def restricted.refl : Π (x : pgame), restricted x x -| x := ⟨_, _, λ i, restricted.refl _, λ j, restricted.refl _⟩ -using_well_founded { dec_tac := pgame_wf_tac } - -instance (x : pgame) : inhabited (restricted x x) := ⟨restricted.refl _⟩ - -/-- Transitivity of restriction. -/ -def restricted.trans : Π {x y z : pgame} (r : restricted x y) (s : restricted y z), restricted x z -| x y z ⟨L₁, R₁, hL₁, hR₁⟩ ⟨L₂, R₂, hL₂, hR₂⟩ := -⟨_, _, λ i, (hL₁ i).trans (hL₂ _), λ j, (hR₁ _).trans (hR₂ j)⟩ - -theorem restricted.le : Π {x y : pgame} (r : restricted x y), x ≤ y -| x y ⟨L, R, hL, hR⟩ := -le_def.2 ⟨λ i, or.inl ⟨L i, (hL i).le⟩, λ i, or.inr ⟨R i, (hR i).le⟩⟩ - /-- `relabelling x y` says that `x` and `y` are really the same game, just dressed up differently. Specifically, there is a bijection between the moves for Left in `x` and in `y`, and similarly @@ -757,14 +734,6 @@ def move_right_symm : ∀ (r : x ≡r y) (i : y.right_moves), x.move_right (r.right_moves_equiv.symm i) ≡r y.move_right i | ⟨L, R, hL, hR⟩ i := by simpa using hR (R.symm i) -/-- If `x` is a relabelling of `y`, then `x` is a restriction of `y`. -/ -def restricted : Π {x y : pgame} (r : x ≡r y), restricted x y -| x y r := ⟨_, _, λ i, (r.move_left i).restricted, λ j, (r.move_right_symm j).restricted⟩ -using_well_founded { dec_tac := pgame_wf_tac } - -/-! It's not the case that `restricted x y → restricted y x → x ≡r y`, but if we insisted that the -maps in a restriction were injective, then one could use Schröder-Bernstein for do this. -/ - /-- The identity relabelling. -/ @[refl] def refl : Π (x : pgame), x ≡r x | x := ⟨equiv.refl _, equiv.refl _, λ i, refl _, λ j, refl _⟩ @@ -776,8 +745,11 @@ instance (x : pgame) : inhabited (x ≡r x) := ⟨refl _⟩ @[symm] def symm : Π {x y : pgame}, x ≡r y → y ≡r x | x y ⟨L, R, hL, hR⟩ := mk' L R (λ i, (hL i).symm) (λ j, (hR j).symm) -theorem le (r : x ≡r y) : x ≤ y := r.restricted.le -theorem ge (r : x ≡r y) : y ≤ x := r.symm.restricted.le +theorem le : ∀ {x y : pgame} (r : x ≡r y), x ≤ y +| x y r := le_def.2 ⟨λ i, or.inl ⟨_, (r.move_left i).le⟩, λ j, or.inr ⟨_, (r.move_right_symm j).le⟩⟩ +using_well_founded { dec_tac := pgame_wf_tac } + +theorem ge {x y : pgame} (r : x ≡r y) : y ≤ x := r.symm.le /-- A relabelling lets us prove equivalence of games. -/ theorem equiv (r : x ≡r y) : x ≈ y := ⟨r.le, r.ge⟩ From 326722b3209cf23f28167b6a5b4355fc67881611 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 28 Sep 2022 21:58:09 +0000 Subject: [PATCH 434/828] feat(set_theory/zfc/basic): nonempty predicate (#15546) We define `Set.nonempty` matching `set.nonempty` and prove the basic results. --- src/set_theory/zfc/basic.lean | 36 +++++++++++++++++++++++++++++++++++ 1 file changed, 36 insertions(+) diff --git a/src/set_theory/zfc/basic.lean b/src/set_theory/zfc/basic.lean index df5bbe6c3c911..00a3d2406577e 100644 --- a/src/set_theory/zfc/basic.lean +++ b/src/set_theory/zfc/basic.lean @@ -222,6 +222,21 @@ def to_set (u : pSet.{u}) : set pSet.{u} := {x | x ∈ u} @[simp] theorem mem_to_set (a u : pSet.{u}) : a ∈ u.to_set ↔ a ∈ u := iff.rfl +/-- A nonempty set is one that contains some element. -/ +protected def nonempty (u : pSet) : Prop := u.to_set.nonempty + +theorem nonempty_def (u : pSet) : u.nonempty ↔ ∃ x, x ∈ u := iff.rfl + +theorem nonempty_of_mem {x u : pSet} (h : x ∈ u) : u.nonempty := ⟨x, h⟩ + +@[simp] theorem nonempty_to_set_iff {u : pSet} : u.to_set.nonempty ↔ u.nonempty := iff.rfl + +theorem nonempty_type_iff_nonempty {x : pSet} : nonempty x.type ↔ pSet.nonempty x := +⟨λ ⟨i⟩, ⟨_, func_mem _ i⟩, λ ⟨i, j, h⟩, ⟨j⟩⟩ + +theorem nonempty_of_nonempty_type (x : pSet) [h : nonempty x.type] : pSet.nonempty x := +nonempty_type_iff_nonempty.1 h + /-- Two pre-sets are equivalent iff they have the same members. -/ theorem equiv.eq {x y : pSet} : equiv x y ↔ to_set x = to_set y := equiv_iff_mem.trans set.ext_iff.symm @@ -243,6 +258,8 @@ instance : is_empty (type (∅)) := pempty.is_empty @[simp] theorem empty_subset (x : pSet.{u}) : (∅ : pSet) ⊆ x := λ x, x.elim +@[simp] theorem not_nonempty_empty : ¬ pSet.nonempty ∅ := by simp [pSet.nonempty] + protected theorem equiv_empty (x : pSet) [is_empty x.type] : equiv x ∅ := pSet.equiv_of_is_empty x _ @@ -465,6 +482,15 @@ quotient.induction_on x $ λ a, begin exact ⟨i, subtype.coe_injective (quotient.sound h.symm)⟩ end +/-- A nonempty set is one that contains some element. -/ +protected def nonempty (u : Set) : Prop := u.to_set.nonempty + +theorem nonempty_def (u : Set) : u.nonempty ↔ ∃ x, x ∈ u := iff.rfl + +theorem nonempty_of_mem {x u : Set} (h : x ∈ u) : u.nonempty := ⟨x, h⟩ + +@[simp] theorem nonempty_to_set_iff {u : Set} : u.to_set.nonempty ↔ u.nonempty := iff.rfl + /-- `x ⊆ y` as ZFC sets means that all members of `x` are members of `y`. -/ protected def subset (x y : Set.{u}) := ∀ ⦃z⦄, z ∈ x → z ∈ y @@ -510,6 +536,16 @@ quotient.induction_on x pSet.mem_empty @[simp] theorem empty_subset (x : Set.{u}) : (∅ : Set) ⊆ x := quotient.induction_on x $ λ y, subset_iff.2 $ pSet.empty_subset y +@[simp] theorem not_nonempty_empty : ¬ Set.nonempty ∅ := by simp [Set.nonempty] + +@[simp] theorem nonempty_mk_iff {x : pSet} : (mk x).nonempty ↔ x.nonempty := +begin + refine ⟨_, λ ⟨a, h⟩, ⟨mk a, h⟩⟩, + rintro ⟨a, h⟩, + induction a using quotient.induction_on, + exact ⟨a, h⟩ +end + theorem eq_empty (x : Set.{u}) : x = ∅ ↔ ∀ y : Set.{u}, y ∉ x := ⟨λ h y, (h.symm ▸ mem_empty y), λ h, ext (λ y, ⟨λ yx, absurd yx (h y), λ y0, absurd y0 (mem_empty _)⟩)⟩ From 76747cab11a423d089dcadce1d1ae158bde35a16 Mon Sep 17 00:00:00 2001 From: damiano Date: Wed, 28 Sep 2022 21:58:10 +0000 Subject: [PATCH 435/828] feat(set_theory/cardinal/ordinal): some lemmas about adding finite cardinals and (in)equalities (#16262) These lemmas arose from [this Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/cardinal.2Ele_of_add_le_add). Comments are welcome! --- src/set_theory/cardinal/ordinal.lean | 32 ++++++++++++++++++++++++++++ 1 file changed, 32 insertions(+) diff --git a/src/set_theory/cardinal/ordinal.lean b/src/set_theory/cardinal/ordinal.lean index 7503ee6a185ad..792112ce3328a 100644 --- a/src/set_theory/cardinal/ordinal.lean +++ b/src/set_theory/cardinal/ordinal.lean @@ -644,6 +644,9 @@ end lemma add_eq_right_iff {a b : cardinal} : a + b = b ↔ (max ℵ₀ a ≤ b ∨ a = 0) := by { rw [add_comm, add_eq_left_iff] } +lemma add_nat_eq {a : cardinal} (n : ℕ) (ha : ℵ₀ ≤ a) : a + n = a := +add_eq_left ha ((nat_lt_aleph_0 _).le.trans ha) + lemma add_one_eq {a : cardinal} (ha : ℵ₀ ≤ a) : a + 1 = a := add_eq_left ha (one_le_aleph_0.trans ha) @@ -679,6 +682,35 @@ theorem principal_add_ord {c : cardinal} (hc : ℵ₀ ≤ c) : ordinal.principal theorem principal_add_aleph (o : ordinal) : ordinal.principal (+) (aleph o).ord := principal_add_ord $ aleph_0_le_aleph o +lemma add_right_inj_of_lt_aleph_0 {α β γ : cardinal} (γ₀ : γ < aleph_0) : + α + γ = β + γ ↔ α = β := +⟨λ h, cardinal.eq_of_add_eq_add_right h γ₀, λ h, congr_fun (congr_arg (+) h) γ⟩ + +@[simp] lemma add_nat_inj {α β : cardinal} (n : ℕ) : + α + n = β + n ↔ α = β := +add_right_inj_of_lt_aleph_0 (nat_lt_aleph_0 _) + +@[simp] lemma add_one_inj {α β : cardinal} : + α + 1 = β + 1 ↔ α = β := +add_right_inj_of_lt_aleph_0 one_lt_aleph_0 + +lemma add_le_add_iff_of_lt_aleph_0 {α β γ : cardinal} (γ₀ : γ < cardinal.aleph_0) : + α + γ ≤ β + γ ↔ α ≤ β := +begin + refine ⟨λ h, _, λ h, add_le_add_right h γ⟩, + contrapose h, + rw [not_le, lt_iff_le_and_ne, ne] at h ⊢, + exact ⟨add_le_add_right h.1 γ, mt (add_right_inj_of_lt_aleph_0 γ₀).1 h.2⟩, +end + +@[simp] lemma add_nat_le_add_nat_iff_of_lt_aleph_0 {α β : cardinal} (n : ℕ) : + α + n ≤ β + n ↔ α ≤ β := +add_le_add_iff_of_lt_aleph_0 (nat_lt_aleph_0 n) + +@[simp] lemma add_one_le_add_one_iff_of_lt_aleph_0 {α β : cardinal} : + α + 1 ≤ β + 1 ↔ α ≤ β := +add_le_add_iff_of_lt_aleph_0 one_lt_aleph_0 + /-! ### Properties about power -/ theorem pow_le {κ μ : cardinal.{u}} (H1 : ℵ₀ ≤ κ) (H2 : μ < ℵ₀) : κ ^ μ ≤ κ := From 25e7fe6b8ecdb5b1a155ec46bf8d2a33cddf27c0 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 28 Sep 2022 21:58:11 +0000 Subject: [PATCH 436/828] chore(algebra/order/monoid_lemmas_zero_lt): reorder, create aliases (#16522) The first part of #16449 --- src/algebra/order/monoid_lemmas_zero_lt.lean | 64 ++++++++++---------- 1 file changed, 31 insertions(+), 33 deletions(-) diff --git a/src/algebra/order/monoid_lemmas_zero_lt.lean b/src/algebra/order/monoid_lemmas_zero_lt.lean index 78c077cb399e3..0df10df40ffd6 100644 --- a/src/algebra/order/monoid_lemmas_zero_lt.lean +++ b/src/algebra/order/monoid_lemmas_zero_lt.lean @@ -116,56 +116,54 @@ instance mul_pos_reflect_lt.to_contravariant_class_pos_mul_lt [mul_pos_reflect_l contravariant_class α>0 α (λ x y, y * x) (<) := ⟨λ a b c bc, @contravariant_class.elim α≥0 α (λ x y, y * x) (<) _ ⟨_, a.2.le⟩ _ _ bc⟩ -lemma mul_le_mul_of_nonneg_left [pos_mul_mono α] (h : b ≤ c) (ha : 0 ≤ a) : a * b ≤ a * c := -@covariant_class.elim α≥0 α (λ x y, x * y) (≤) _ ⟨a, ha⟩ _ _ h +lemma mul_le_mul_of_nonneg_left [pos_mul_mono α] (h : b ≤ c) (a0 : 0 ≤ a) : a * b ≤ a * c := +@covariant_class.elim α≥0 α (λ x y, x * y) (≤) _ ⟨a, a0⟩ _ _ h -lemma mul_le_mul_of_nonneg_right [mul_pos_mono α] (h : b ≤ c) (ha : 0 ≤ a) : b * a ≤ c * a := -@covariant_class.elim α≥0 α (λ x y, y * x) (≤) _ ⟨a, ha⟩ _ _ h +lemma mul_le_mul_of_nonneg_right [mul_pos_mono α] (h : b ≤ c) (a0 : 0 ≤ a) : b * a ≤ c * a := +@covariant_class.elim α≥0 α (λ x y, y * x) (≤) _ ⟨a, a0⟩ _ _ h -lemma mul_lt_mul_of_pos_left [pos_mul_strict_mono α] (bc : b < c) (ha : 0 < a) : a * b < a * c := -@covariant_class.elim α>0 α (λ x y, x * y) (<) _ ⟨a, ha⟩ _ _ bc +lemma mul_lt_mul_of_pos_left [pos_mul_strict_mono α] (bc : b < c) (a0 : 0 < a) : a * b < a * c := +@covariant_class.elim α>0 α (λ x y, x * y) (<) _ ⟨a, a0⟩ _ _ bc -lemma mul_lt_mul_of_pos_right [mul_pos_strict_mono α] (bc : b < c) (ha : 0 < a) : b * a < c * a := -@covariant_class.elim α>0 α (λ x y, y * x) (<) _ ⟨a, ha⟩ _ _ bc +lemma mul_lt_mul_of_pos_right [mul_pos_strict_mono α] (bc : b < c) (a0 : 0 < a) : b * a < c * a := +@covariant_class.elim α>0 α (λ x y, y * x) (<) _ ⟨a, a0⟩ _ _ bc -lemma lt_of_mul_lt_mul_left [pos_mul_reflect_lt α] (h : a * b < a * c) (ha : 0 ≤ a) : b < c := -@contravariant_class.elim α≥0 α (λ x y, x * y) (<) _ ⟨a, ha⟩ _ _ h +lemma lt_of_mul_lt_mul_left [pos_mul_reflect_lt α] (h : a * b < a * c) (a0 : 0 ≤ a) : b < c := +@contravariant_class.elim α≥0 α (λ x y, x * y) (<) _ ⟨a, a0⟩ _ _ h -lemma lt_of_mul_lt_mul_right [mul_pos_reflect_lt α] (h : b * a < c * a) (ha : 0 ≤ a) : b < c := -@contravariant_class.elim α≥0 α (λ x y, y * x) (<) _ ⟨a, ha⟩ _ _ h +lemma lt_of_mul_lt_mul_right [mul_pos_reflect_lt α] (h : b * a < c * a) (a0 : 0 ≤ a) : b < c := +@contravariant_class.elim α≥0 α (λ x y, y * x) (<) _ ⟨a, a0⟩ _ _ h -@[simp] -lemma mul_lt_mul_left [pos_mul_strict_mono α] [pos_mul_reflect_lt α] +lemma le_of_mul_le_mul_left [pos_mul_mono_rev α] (bc : a * b ≤ a * c) (a0 : 0 < a) : b ≤ c := +@contravariant_class.elim α>0 α (λ x y, x * y) (≤) _ ⟨a, a0⟩ _ _ bc + +lemma le_of_mul_le_mul_right [mul_pos_mono_rev α] (bc : b * a ≤ c * a) (a0 : 0 < a) : b ≤ c := +@contravariant_class.elim α>0 α (λ x y, y * x) (≤) _ ⟨a, a0⟩ _ _ bc + +alias lt_of_mul_lt_mul_left ← lt_of_mul_lt_mul_of_nonneg_left +alias lt_of_mul_lt_mul_right ← lt_of_mul_lt_mul_of_nonneg_right +alias le_of_mul_le_mul_left ← le_of_mul_le_mul_of_pos_left +alias le_of_mul_le_mul_right ← le_of_mul_le_mul_of_pos_right + +@[simp] lemma mul_lt_mul_left [pos_mul_strict_mono α] [pos_mul_reflect_lt α] (a0 : 0 < a) : a * b < a * c ↔ b < c := @rel_iff_cov α>0 α (λ x y, x * y) (<) _ _ ⟨a, a0⟩ _ _ -@[simp] -lemma mul_lt_mul_right [mul_pos_strict_mono α] [mul_pos_reflect_lt α] +@[simp] lemma mul_lt_mul_right [mul_pos_strict_mono α] [mul_pos_reflect_lt α] (a0 : 0 < a) : b * a < c * a ↔ b < c := @rel_iff_cov α>0 α (λ x y, y * x) (<) _ _ ⟨a, a0⟩ _ _ -lemma le_of_mul_le_mul_of_pos_left [pos_mul_mono_rev α] - (bc : a * b ≤ a * c) (a0 : 0 < a) : - b ≤ c := -@contravariant_class.elim α>0 α (λ x y, x * y) (≤) _ ⟨a, a0⟩ _ _ bc - -lemma le_of_mul_le_mul_of_pos_right [mul_pos_mono_rev α] - (bc : b * a ≤ c * a) (a0 : 0 < a) : - b ≤ c := -@contravariant_class.elim α>0 α (λ x y, y * x) (≤) _ ⟨a, a0⟩ _ _ bc - -alias le_of_mul_le_mul_of_pos_left ← le_of_mul_le_mul_left -alias le_of_mul_le_mul_of_pos_right ← le_of_mul_le_mul_right - -@[simp] lemma mul_le_mul_left [pos_mul_mono α] [pos_mul_mono_rev α] (ha : 0 < a) : +@[simp] lemma mul_le_mul_left [pos_mul_mono α] [pos_mul_mono_rev α] + (a0 : 0 < a) : a * b ≤ a * c ↔ b ≤ c := -@rel_iff_cov α>0 α (λ x y, x * y) (≤) _ _ ⟨a, ha⟩ _ _ +@rel_iff_cov α>0 α (λ x y, x * y) (≤) _ _ ⟨a, a0⟩ _ _ -@[simp] lemma mul_le_mul_right [mul_pos_mono α] [mul_pos_mono_rev α] (ha : 0 < a) : +@[simp] lemma mul_le_mul_right [mul_pos_mono α] [mul_pos_mono_rev α] + (a0 : 0 < a) : b * a ≤ c * a ↔ b ≤ c := -@rel_iff_cov α>0 α (λ x y, y * x) (≤) _ _ ⟨a, ha⟩ _ _ +@rel_iff_cov α>0 α (λ x y, y * x) (≤) _ _ ⟨a, a0⟩ _ _ lemma mul_lt_mul_of_pos_of_nonneg [pos_mul_strict_mono α] [mul_pos_mono α] (h₁ : a ≤ b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 ≤ d) : a * c < b * d := From f769858878281250254534ba937f774ffed37fbb Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 28 Sep 2022 21:58:12 +0000 Subject: [PATCH 437/828] =?UTF-8?q?feat(tactic/push=5Fneg):=20option=20for?= =?UTF-8?q?=20an=20alternate=20normal=20form=20of=20`=C2=AC=20(P=20?= =?UTF-8?q?=E2=88=A7=20Q)`=20(#16586)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Backport a feature of the [mathlib4 version](https://github.com/leanprover-community/mathlib4/pull/344) of `push_neg`: an option to make `¬ (P ∧ Q)` be normalized to `¬ P ∨ ¬ Q`, rather than `P → ¬ Q`. That was actually the original behaviour, but it was changed in #3362. I have implemented this as a global option `trace.push_neg.use_distrib` (using the [tracing option hack](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/custom.20options)) rather than as a piece of configuration information which is passed when using the tactic, because I imagine this feature will mostly be used in teaching (that was both my motivation and also @PatrickMassot's when he wrote the original version), where it is convenient to "set it and forget it". This is also how it is implemented in the mathlib4 version (cc @dupuisf @j-loreaux). Zulip: https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/alternative.20normal.20form.20for.20push_neg --- src/tactic/push_neg.lean | 14 ++++++++++++-- test/push_neg.lean | 20 ++++++++++++++++++++ 2 files changed, 32 insertions(+), 2 deletions(-) diff --git a/src/tactic/push_neg.lean b/src/tactic/push_neg.lean index 0feea63b832a0..e97097ab8a34e 100644 --- a/src/tactic/push_neg.lean +++ b/src/tactic/push_neg.lean @@ -10,6 +10,10 @@ import logic.basic open tactic expr +/- Enable the option `trace.push_neg.use_distrib` in order to have `¬ (p ∧ q)` normalized to +`¬ p ∨ ¬ q`, rather than the default `p → ¬ q`. -/ +declare_trace push_neg.use_distrib + namespace push_neg section @@ -21,6 +25,7 @@ variable (s : α → Prop) local attribute [instance, priority 10] classical.prop_decidable theorem not_not_eq : (¬ ¬ p) = p := propext not_not theorem not_and_eq : (¬ (p ∧ q)) = (p → ¬ q) := propext not_and +theorem not_and_distrib_eq : (¬ (p ∧ q)) = (¬ p ∨ ¬ q) := propext not_and_distrib theorem not_or_eq : (¬ (p ∨ q)) = (¬ p ∧ ¬ q) := propext not_or_distrib theorem not_forall_eq : (¬ ∀ x, s x) = (∃ x, ¬ s x) := propext not_forall theorem not_exists_eq : (¬ ∃ x, s x) = (∀ x, ¬ s x) := propext not_exists @@ -48,8 +53,13 @@ do e ← whnf_reducible e, match ne with | `(¬ %%a) := do pr ← mk_app ``not_not_eq [a], return (some (a, pr)) - | `(%%a ∧ %%b) := do pr ← mk_app ``not_and_eq [a, b], - return (some (`((%%a : Prop) → ¬ %%b), pr)) + | `(%%a ∧ %%b) := do distrib ← get_bool_option `trace.push_neg.use_distrib ff, + if distrib then do + pr ← mk_app ``not_and_distrib_eq [a, b], + return (some (`(¬ (%%a : Prop) ∨ ¬ %%b), pr)) + else do + pr ← mk_app ``not_and_eq [a, b], + return (some (`((%%a : Prop) → ¬ %%b), pr)) | `(%%a ∨ %%b) := do pr ← mk_app ``not_or_eq [a, b], return (some (`(¬ %%a ∧ ¬ %%b), pr)) | `(%%a ≤ %%b) := do e ← to_expr ``(%%b < %%a), diff --git a/test/push_neg.lean b/test/push_neg.lean index 887ab382dd276..d12558b63a46e 100644 --- a/test/push_neg.lean +++ b/test/push_neg.lean @@ -89,3 +89,23 @@ begin guard (ht = h1t) ), exact hf end + +/-! Test the option `trace.push_neg.use_distrib` for changing the normal form of `¬(P ∧ Q)`. -/ + +section + +example (a b : ℤ) (h : ¬ (∃ x, (a < x ∧ x < b))) : ∀ x, a < x → b ≤ x := +begin + push_neg at h, + exact h, +end + +set_option trace.push_neg.use_distrib true + +example (a b : ℤ) (h : ¬ (∃ x, (a < x ∧ x < b))) : ∀ x, x ≤ a ∨ b ≤ x := +begin + push_neg at h, + exact h, +end + +end From 0abedaf891cba6d85adca21085388b0f4c4e36c3 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 28 Sep 2022 21:58:14 +0000 Subject: [PATCH 438/828] feat(data/nat/parity): iterations of involutive functions (#16630) --- src/data/bool/basic.lean | 2 ++ src/data/nat/parity.lean | 26 ++++++++++++++++++++++++++ src/logic/function/basic.lean | 2 ++ 3 files changed, 30 insertions(+) diff --git a/src/data/bool/basic.lean b/src/data/bool/basic.lean index 3dbbacec5a89d..1997c7369338d 100644 --- a/src/data/bool/basic.lean +++ b/src/data/bool/basic.lean @@ -98,6 +98,8 @@ by by_cases p; simp * @[simp] theorem cond_bnot {α} (b : bool) (t e : α) : cond (!b) t e = cond b e t := by cases b; refl +theorem bnot_ne_id : bnot ≠ id := λ h, ff_ne_tt $ congr_fun h tt + theorem coe_bool_iff : ∀ {a b : bool}, (a ↔ b) ↔ a = b := dec_trivial theorem eq_tt_of_ne_ff : ∀ {a : bool}, a ≠ ff → a = tt := dec_trivial diff --git a/src/data/nat/parity.lean b/src/data/nat/parity.lean index 9851b36c3f407..ee1e8e91211e2 100644 --- a/src/data/nat/parity.lean +++ b/src/data/nat/parity.lean @@ -235,6 +235,32 @@ end nat open nat +namespace function +namespace involutive + +variables {α : Type*} {f : α → α} {n : ℕ} + +theorem iterate_bit0 (hf : involutive f) (n : ℕ) : f^[bit0 n] = id := +by rw [bit0, ← two_mul, iterate_mul, involutive_iff_iter_2_eq_id.1 hf, iterate_id] + +theorem iterate_bit1 (hf : involutive f) (n : ℕ) : f^[bit1 n] = f := +by rw [bit1, iterate_succ, hf.iterate_bit0, comp.left_id] + +theorem iterate_even (hf : involutive f) (hn : even n) : f^[n] = id := +let ⟨m, hm⟩ := hn in hm.symm ▸ hf.iterate_bit0 m + +theorem iterate_odd (hf : involutive f) (hn : odd n) : f^[n] = f := +let ⟨m, hm⟩ := odd_iff_exists_bit1.mp hn in hm.symm ▸ hf.iterate_bit1 m + +theorem iterate_eq_self (hf : involutive f) (hne : f ≠ id) : f^[n] = f ↔ odd n := +⟨λ H, odd_iff_not_even.2 $ λ hn, hne $ by rwa [hf.iterate_even hn, eq_comm] at H, hf.iterate_odd⟩ + +theorem iterate_eq_id (hf : involutive f) (hne : f ≠ id) : f^[n] = id ↔ even n := +⟨λ H, even_iff_not_odd.2 $ λ hn, hne $ by rwa [hf.iterate_odd hn] at H, hf.iterate_even⟩ + +end involutive +end function + variables {R : Type*} [monoid R] [has_distrib_neg R] {n : ℕ} lemma neg_one_pow_eq_one_iff_even (h : (-1 : R) ≠ 1) : (-1 : R) ^ n = 1 ↔ even n := diff --git a/src/logic/function/basic.lean b/src/logic/function/basic.lean index 427d6f658d8db..754f71f3975ca 100644 --- a/src/logic/function/basic.lean +++ b/src/logic/function/basic.lean @@ -662,6 +662,8 @@ def involutive {α} (f : α → α) : Prop := ∀ x, f (f x) = x lemma involutive_iff_iter_2_eq_id {α} {f : α → α} : involutive f ↔ (f^[2] = id) := funext_iff.symm +lemma _root_.bool.involutive_bnot : involutive bnot := bnot_bnot + namespace involutive variables {α : Sort u} {f : α → α} (h : involutive f) include h From acc2a10bcdf762de4d6369ab9190f8a403eb01ba Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 28 Sep 2022 21:58:15 +0000 Subject: [PATCH 439/828] feat(algebra/module/basic): weaken assumption (#16673) --- src/algebra/module/basic.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index e4f7ba125817f..2ea00c13212df 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -615,15 +615,16 @@ end smul_injective end module -section division_ring +section group_with_zero -variables [division_ring R] [add_comm_group M] [module R M] +variables [group_with_zero R] [add_monoid M] [distrib_mul_action R M] +/-- This instance applies to `division_semiring`s, in particular `nnreal` and `nnrat`. -/ @[priority 100] -- see note [lower instance priority] -instance division_ring.to_no_zero_smul_divisors : no_zero_smul_divisors R M := +instance group_with_zero.to_no_zero_smul_divisors : no_zero_smul_divisors R M := ⟨λ c x h, or_iff_not_imp_left.2 $ λ hc, (smul_eq_zero_iff_eq' hc).1 h⟩ -end division_ring +end group_with_zero @[priority 100] -- see note [lower instance priority] instance rat_module.no_zero_smul_divisors [add_comm_group M] [module ℚ M] : From 345b38d652d628167edcdf8f043024e9f7a98587 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 28 Sep 2022 21:58:16 +0000 Subject: [PATCH 440/828] feat(measure_theory/measure/lebesgue): deduce that a property is almost sure from a localized version in intervals (#16684) --- src/measure_theory/measure/lebesgue.lean | 84 ++++++++++++++++++- src/measure_theory/measure/measure_space.lean | 17 ++++ 2 files changed, 100 insertions(+), 1 deletion(-) diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index d853b83c1f69b..9c643bbbf012a 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -25,7 +25,7 @@ are proved more generally for any additive Haar measure on a finite-dimensional -/ noncomputable theory -open classical set filter measure_theory measure_theory.measure +open classical set filter measure_theory measure_theory.measure topological_space open ennreal (of_real) open_locale big_operators ennreal nnreal topological_space @@ -527,3 +527,85 @@ volume_region_between_eq_integral' f_int g_int hs ((ae_restrict_iff' hs).mpr (eventually_of_forall hfg)) end region_between + +/-- Consider a real set `s`. If a property is true almost everywhere in `s ∩ (a, b)` for +all `a, b ∈ s`, then it is true almost everywhere in `s`. Formulated with `μ.restrict`. +See also `ae_of_mem_of_ae_of_mem_inter_Ioo`. -/ +lemma ae_restrict_of_ae_restrict_inter_Ioo + {μ : measure ℝ} [has_no_atoms μ] {s : set ℝ} {p : ℝ → Prop} + (h : ∀ a b, a ∈ s → b ∈ s → a < b → ∀ᵐ x ∂(μ.restrict (s ∩ Ioo a b)), p x) : + ∀ᵐ x ∂(μ.restrict s), p x := +begin + /- By second-countability, we cover `s` by countably many intervals `(a, b)` (except maybe for + two endpoints, which don't matter since `μ` does not have any atom). -/ + let T : s × s → set ℝ := λ p, Ioo p.1 p.2, + let u := ⋃ (i : ↥s × ↥s), T i, + have hfinite : (s \ u).finite, + { refine set.finite_of_forall_between_eq_endpoints (s \ u) (λ x hx y hy z hz hxy hyz, _), + by_contra' h, + apply hy.2, + exact mem_Union_of_mem (⟨x, hx.1⟩, ⟨z, hz.1⟩) + ⟨lt_of_le_of_ne hxy h.1, lt_of_le_of_ne hyz h.2⟩ }, + obtain ⟨A, A_count, hA⟩ : + ∃ (A : set (↥s × ↥s)), A.countable ∧ (⋃ (i ∈ A), T i) = ⋃ (i : ↥s × ↥s), T i := + is_open_Union_countable _ (λ p, is_open_Ioo), + have : s ⊆ (s \ u) ∪ (⋃ p ∈ A, s ∩ T p), + { assume x hx, + by_cases h'x : x ∈ ⋃ (i : ↥s × ↥s), T i, + { rw ← hA at h'x, + obtain ⟨p, pA, xp⟩ : ∃ (p : ↥s × ↥s), p ∈ A ∧ x ∈ T p, + by simpa only [mem_Union, exists_prop, set_coe.exists, exists_and_distrib_right] using h'x, + right, + exact mem_bUnion pA ⟨hx, xp⟩ }, + { exact or.inl ⟨hx, h'x⟩ } }, + apply ae_restrict_of_ae_restrict_of_subset this, + rw [ae_restrict_union_iff, ae_restrict_bUnion_iff _ A_count], + split, + { have : μ.restrict (s \ u) = 0, by simp only [restrict_eq_zero, hfinite.measure_zero], + simp only [this, ae_zero] }, + { rintros ⟨⟨a, as⟩, ⟨b, bs⟩⟩ -, + dsimp [T], + rcases le_or_lt b a with hba|hab, + { simp only [Ioo_eq_empty_of_le hba, inter_empty, restrict_empty, ae_zero] }, + { exact h a b as bs hab } } +end + +/-- Consider a real set `s`. If a property is true almost everywhere in `s ∩ (a, b)` for +all `a, b ∈ s`, then it is true almost everywhere in `s`. Formulated with bare membership. +See also `ae_restrict_of_ae_restrict_inter_Ioo`. -/ +lemma ae_of_mem_of_ae_of_mem_inter_Ioo + {μ : measure ℝ} [has_no_atoms μ] {s : set ℝ} {p : ℝ → Prop} + (h : ∀ a b, a ∈ s → b ∈ s → a < b → ∀ᵐ x ∂μ, x ∈ s ∩ Ioo a b → p x) : + ∀ᵐ x ∂μ, x ∈ s → p x := +begin + /- By second-countability, we cover `s` by countably many intervals `(a, b)` (except maybe for + two endpoints, which don't matter since `μ` does not have any atom). -/ + let T : s × s → set ℝ := λ p, Ioo p.1 p.2, + let u := ⋃ (i : ↥s × ↥s), T i, + have hfinite : (s \ u).finite, + { refine set.finite_of_forall_between_eq_endpoints (s \ u) (λ x hx y hy z hz hxy hyz, _), + by_contra' h, + apply hy.2, + exact mem_Union_of_mem (⟨x, hx.1⟩, ⟨z, hz.1⟩) + ⟨lt_of_le_of_ne hxy h.1, lt_of_le_of_ne hyz h.2⟩ }, + obtain ⟨A, A_count, hA⟩ : + ∃ (A : set (↥s × ↥s)), A.countable ∧ (⋃ (i ∈ A), T i) = ⋃ (i : ↥s × ↥s), T i := + is_open_Union_countable _ (λ p, is_open_Ioo), + have M : ∀ᵐ x ∂μ, x ∉ s \ u, from hfinite.countable.ae_not_mem _, + have M' : ∀ᵐ x ∂μ, ∀ (i : ↥s × ↥s) (H : i ∈ A), x ∈ s ∩ T i → p x, + { rw ae_ball_iff A_count, + rintros ⟨⟨a, as⟩, ⟨b, bs⟩⟩ -, + change ∀ᵐ (x : ℝ) ∂μ, x ∈ s ∩ Ioo a b → p x, + rcases le_or_lt b a with hba|hab, + { simp only [Ioo_eq_empty_of_le hba, inter_empty, is_empty.forall_iff, eventually_true, + mem_empty_iff_false], }, + { exact h a b as bs hab } }, + filter_upwards [M, M'] with x hx h'x, + assume xs, + by_cases Hx : x ∈ ⋃ (i : ↥s × ↥s), T i, + { rw ← hA at Hx, + obtain ⟨p, pA, xp⟩ : ∃ (p : ↥s × ↥s), p ∈ A ∧ x ∈ T p, + by simpa only [mem_Union, exists_prop, set_coe.exists, exists_and_distrib_right] using Hx, + apply h'x p pA ⟨xs, xp⟩ }, + { exact false.elim (hx ⟨xs, Hx⟩) } +end diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 167c142ab0a88..ec1182d915da4 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -1994,6 +1994,23 @@ lemma ae_restrict_bUnion_finset_eq (s : ι → set α) (t : finset ι) : (μ.restrict (⋃ i ∈ t, s i)).ae = ⨆ i ∈ t, (μ.restrict (s i)).ae := ae_restrict_bUnion_eq s t.countable_to_set +lemma ae_restrict_Union_iff [countable ι] (s : ι → set α) (p : α → Prop) : + (∀ᵐ x ∂ (μ.restrict (⋃ i, s i)), p x) ↔ (∀ i, (∀ᵐ x ∂ (μ.restrict (s i)), p x)) := +by simp + +lemma ae_restrict_union_iff (s t : set α) (p : α → Prop) : + (∀ᵐ x ∂ (μ.restrict (s ∪ t)), p x) ↔ + ((∀ᵐ x ∂ (μ.restrict s), p x) ∧ (∀ᵐ x ∂ (μ.restrict t), p x)) := +by simp + +lemma ae_restrict_bUnion_iff (s : ι → set α) {t : set ι} (ht : t.countable) (p : α → Prop) : + (∀ᵐ x ∂(μ.restrict (⋃ i ∈ t, s i)), p x) ↔ ∀ i ∈ t, ∀ᵐ x ∂(μ.restrict (s i)), p x := +by simp_rw [filter.eventually, ae_restrict_bUnion_eq s ht, mem_supr] + +@[simp] lemma ae_restrict_bUnion_finset_iff (s : ι → set α) (t : finset ι) (p : α → Prop) : + (∀ᵐ x ∂(μ.restrict (⋃ i ∈ t, s i)), p x) ↔ ∀ i ∈ t, ∀ᵐ x ∂(μ.restrict (s i)), p x := +by simp_rw [filter.eventually, ae_restrict_bUnion_finset_eq s, mem_supr] + lemma ae_eq_restrict_Union_iff [countable ι] (s : ι → set α) (f g : α → δ) : f =ᵐ[μ.restrict (⋃ i, s i)] g ↔ ∀ i, f =ᵐ[μ.restrict (s i)] g := by simp_rw [eventually_eq, ae_restrict_Union_eq, eventually_supr] From e63a4d5e5e7348c58e4257e3354d9c07da9da2ad Mon Sep 17 00:00:00 2001 From: mcdoll Date: Thu, 29 Sep 2022 01:02:45 +0000 Subject: [PATCH 441/828] feat(analysis/special_functions/pow): sqrt and inequalities (#16515) This PR proves a few lemmas about `real.sqrt` and `real.rpow` as well as inequality lemmas for negative powers. --- src/analysis/special_functions/pow.lean | 45 +++++++++++++++++++++++++ 1 file changed, 45 insertions(+) diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index f0f6486f4bd76..f357c8a642a4b 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -648,6 +648,44 @@ lemma rpow_lt_rpow_iff (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z < y ^ lemma rpow_le_rpow_iff (hx : 0 ≤ x) (hy : 0 ≤ y) (hz : 0 < z) : x ^ z ≤ y ^ z ↔ x ≤ y := le_iff_le_iff_lt_iff_lt.2 $ rpow_lt_rpow_iff hy hx hz +lemma le_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : + x ≤ y ^ z⁻¹ ↔ y ≤ x ^ z := +begin + have hz' : 0 < -z := by rwa [lt_neg, neg_zero'], + have hxz : 0 < x ^ (-z) := real.rpow_pos_of_pos hx _, + have hyz : 0 < y ^ z⁻¹ := real.rpow_pos_of_pos hy _, + rw [←real.rpow_le_rpow_iff hx.le hyz.le hz', ←real.rpow_mul hy.le], + simp only [ne_of_lt hz, real.rpow_neg_one, mul_neg, inv_mul_cancel, ne.def, not_false_iff], + rw [le_inv hxz hy, ←real.rpow_neg_one, ←real.rpow_mul hx.le], + simp, +end + +lemma lt_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : + x < y ^ z⁻¹ ↔ y < x ^ z := +begin + have hz' : 0 < -z := by rwa [lt_neg, neg_zero'], + have hxz : 0 < x ^ (-z) := real.rpow_pos_of_pos hx _, + have hyz : 0 < y ^ z⁻¹ := real.rpow_pos_of_pos hy _, + rw [←real.rpow_lt_rpow_iff hx.le hyz.le hz', ←real.rpow_mul hy.le], + simp only [ne_of_lt hz, real.rpow_neg_one, mul_neg, inv_mul_cancel, ne.def, not_false_iff], + rw [lt_inv hxz hy, ←real.rpow_neg_one, ←real.rpow_mul hx.le], + simp, +end + +lemma rpow_inv_lt_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : + x ^ z⁻¹ < y ↔ y ^ z < x := +begin + convert lt_rpow_inv_iff_of_neg (real.rpow_pos_of_pos hx _) (real.rpow_pos_of_pos hy _) hz; + simp [←real.rpow_mul hx.le, ←real.rpow_mul hy.le, ne_of_lt hz], +end + +lemma rpow_inv_le_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : + x ^ z⁻¹ ≤ y ↔ y ^ z ≤ x := +begin + convert le_rpow_inv_iff_of_neg (real.rpow_pos_of_pos hx _) (real.rpow_pos_of_pos hy _) hz; + simp [←real.rpow_mul hx.le, ←real.rpow_mul hy.le, ne_of_lt hz], +end + lemma rpow_lt_rpow_of_exponent_lt (hx : 1 < x) (hyz : y < z) : x^y < x^z := begin repeat {rw [rpow_def_of_pos (lt_trans zero_lt_one hx)]}, @@ -959,6 +997,13 @@ begin rw [sqrt_eq_zero_of_nonpos h.le, rpow_def_of_neg h, this, cos_pi_div_two, mul_zero] } end +lemma rpow_div_two_eq_sqrt {x : ℝ} (r : ℝ) (hx : 0 ≤ x) : x ^ (r/2) = (sqrt x) ^ r := +begin + rw [sqrt_eq_rpow, ← rpow_mul hx], + congr, + ring, +end + end sqrt end real From 8f89631efda2364c5e4fe16f1307ab216aff047e Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 29 Sep 2022 01:02:46 +0000 Subject: [PATCH 442/828] refactor(algebra/periodic): use `neg_zero_class` for `antiperiodic.nat_mul_eq_of_eq_zero` (#16687) Weaken the typeclass assumption on the codomain of the antiperiodic function in `antiperiodic.nat_mul_eq_of_eq_zero` from `subtraction_monoid` to `neg_zero_class`. The corresponding `int` lemma also requires involutive `neg` so will be dealt with separately once we have a further typeclass for the combination of `neg_zero_class` with `has_involutive_neg`. --- src/algebra/periodic.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index 1ea42572a7c38..442c4cafb585b 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -408,14 +408,13 @@ lemma antiperiodic.int_odd_mul_antiperiodic [ring α] [has_involutive_neg β] antiperiodic f (n * (2 * c) + c) := λ x, by rw [← add_assoc, h, h.periodic.int_mul] -lemma antiperiodic.nat_mul_eq_of_eq_zero [comm_semiring α] [subtraction_monoid β] +lemma antiperiodic.nat_mul_eq_of_eq_zero [comm_semiring α] [neg_zero_class β] (h : antiperiodic f c) (hi : f 0 = 0) (n : ℕ) : f (n * c) = 0 := begin - rcases nat.even_or_odd n with ⟨k, rfl⟩ | ⟨k, rfl⟩; - have hk : (k : α) * (2 * c) = 2 * k * c := by rw [mul_left_comm, ← mul_assoc], - { simpa [← two_mul, hk, hi] using (h.nat_even_mul_periodic k).eq }, - { simpa [add_mul, hk, hi] using (h.nat_odd_mul_antiperiodic k).eq }, + induction n with k hk, + { simp [hi] }, + { simp [hk, add_mul, h (k * c)] } end lemma antiperiodic.int_mul_eq_of_eq_zero [comm_ring α] [subtraction_monoid β] From 0ec98ed0d3bd20fcb6096e46c995534e33246ec3 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 29 Sep 2022 01:02:47 +0000 Subject: [PATCH 443/828] refactor(data/real/ereal): `sub_neg_zero_monoid` instance (#16688) Add a `sub_neg_zero_monoid` instance for `ereal`, so eliminating its `neg_zero`, `sub_zero` and `sub_eq_add_neg` lemmas and allowing results for `neg_zero_class` to be applied to `ereal`. --- src/data/real/ereal.lean | 15 ++++++++------- .../integral/vitali_caratheodory.lean | 2 +- 2 files changed, 9 insertions(+), 8 deletions(-) diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index f9371a704457d..dffc98ea3e875 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -380,7 +380,6 @@ instance : has_neg ereal := ⟨ereal.neg⟩ @[simp] lemma neg_top : - (⊤ : ereal) = ⊥ := rfl @[simp] lemma neg_bot : - (⊥ : ereal) = ⊤ := rfl -@[simp] lemma neg_zero : - (0 : ereal) = 0 := by { change ((-0 : ℝ) : ereal) = 0, simp } instance : has_involutive_neg ereal := { neg := has_neg.neg, @@ -402,7 +401,7 @@ by { rw neg_eq_iff_neg_eq, simp [eq_comm] } by { rw neg_eq_iff_neg_eq, simp [eq_comm] } @[simp] lemma neg_eg_zero_iff {x : ereal} : - x = 0 ↔ x = 0 := -by { rw neg_eq_iff_neg_eq, simp [eq_comm] } +by { rw neg_eq_iff_neg_eq, change ((-0 : ℝ) : ereal) = _ ↔ _, simp [eq_comm] } /-- if `-a ≤ b` then `-b ≤ a` on `ereal`. -/ protected theorem neg_le_of_neg_le : ∀ {a b : ereal} (h : -a ≤ b), -b ≤ a @@ -448,11 +447,16 @@ lemma neg_lt_iff_neg_lt {a b : ereal} : -a < b ↔ -b < a := /-- Subtraction on `ereal`, defined by `x - y = x + (-y)`. Since addition is badly behaved at some points, so is subtraction. There is no standard algebraic typeclass involving subtraction that is -registered on `ereal` because of this bad behavior. -/ +registered on `ereal`, beyond `sub_neg_zero_monoid`, because of this bad behavior. -/ protected noncomputable def sub (x y : ereal) : ereal := x + (-y) noncomputable instance : has_sub ereal := ⟨ereal.sub⟩ +noncomputable instance : sub_neg_zero_monoid ereal := +{ neg_zero := by { change ((-0 : ℝ) : ereal) = 0, simp }, + ..(infer_instance : add_monoid ereal), + ..ereal.has_neg } + @[simp] lemma top_sub (x : ereal) : ⊤ - x = ⊤ := top_add x @[simp] lemma sub_bot (x : ereal) : x - ⊥ = ⊤ := add_top x @@ -460,11 +464,8 @@ noncomputable instance : has_sub ereal := ⟨ereal.sub⟩ @[simp] lemma bot_sub_coe (x : ℝ) : (⊥ : ereal) - x = ⊥ := rfl @[simp] lemma coe_sub_bot (x : ℝ) : (x : ereal) - ⊤ = ⊥ := rfl -@[simp] lemma sub_zero (x : ereal) : x - 0 = x := by { change x + (-0) = x, simp } @[simp] lemma zero_sub (x : ereal) : 0 - x = - x := by { change 0 + (-x) = - x, simp } -lemma sub_eq_add_neg (x y : ereal) : x - y = x + -y := rfl - lemma sub_le_sub {x y z t : ereal} (h : x ≤ y) (h' : t ≤ z) : x - z ≤ y - t := add_le_add h (neg_le_neg_iff.2 h') @@ -491,7 +492,7 @@ end lemma to_real_sub {x y : ereal} (hx : x ≠ ⊤) (h'x : x ≠ ⊥) (hy : y ≠ ⊤) (h'y : y ≠ ⊥) : to_real (x - y) = to_real x - to_real y := begin - rw [ereal.sub_eq_add_neg, to_real_add hx h'x, to_real_neg], + rw [sub_eq_add_neg, to_real_add hx h'x, to_real_neg], { refl }, { simpa using hy }, { simpa using h'y } diff --git a/src/measure_theory/integral/vitali_caratheodory.lean b/src/measure_theory/integral/vitali_caratheodory.lean index 6b3e57b920466..a5c8e8ae91696 100644 --- a/src/measure_theory/integral/vitali_caratheodory.lean +++ b/src/measure_theory/integral/vitali_caratheodory.lean @@ -480,7 +480,7 @@ begin by { congr' 1, field_simp [δ, mul_comm] }, show ∀ᵐ (x : α) ∂μ, g x < ⊤, { filter_upwards [gp_lt_top] with _ hx, - simp [g, ereal.sub_eq_add_neg, lt_top_iff_ne_top, lt_top_iff_ne_top.1 hx], }, + simp [g, sub_eq_add_neg, lt_top_iff_ne_top, lt_top_iff_ne_top.1 hx], }, show ∀ x, (f x : ereal) < g x, { assume x, rw ereal.coe_real_ereal_eq_coe_to_nnreal_sub_coe_to_nnreal (f x), From 787643752bfb164d029e26aa0e944c69078899bc Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 29 Sep 2022 01:02:48 +0000 Subject: [PATCH 444/828] refactor(data/real/ennreal): `div_inv_one_monoid` instance (#16689) Add a `div_inv_one_monoid` instance for `ennreal`, so eliminating its `inv_one` and `div_one` lemmas. (Once we have typeclasses for antitone `inv`, more separate `ennreal` lemmas can be eliminated.) --- src/analysis/specific_limits/basic.lean | 2 +- src/data/real/ennreal.lean | 8 +++----- src/measure_theory/integral/average.lean | 2 +- src/measure_theory/measure/haar.lean | 2 +- src/probability/cond_count.lean | 2 +- src/topology/metric_space/hausdorff_dimension.lean | 2 +- 6 files changed, 8 insertions(+), 10 deletions(-) diff --git a/src/analysis/specific_limits/basic.lean b/src/analysis/specific_limits/basic.lean index 0a964e07c6ab9..6b908a43b83c8 100644 --- a/src/analysis/specific_limits/basic.lean +++ b/src/analysis/specific_limits/basic.lean @@ -309,7 +309,7 @@ end /-- If `edist (f n) (f (n+1))` is bounded by `C * 2^-n`, then the distance from `f 0` to the limit of `f` is bounded above by `2 * C`. -/ lemma edist_le_of_edist_le_geometric_two_of_tendsto₀: edist (f 0) a ≤ 2 * C := -by simpa only [pow_zero, div_eq_mul_inv, ennreal.inv_one, mul_one] +by simpa only [pow_zero, div_eq_mul_inv, inv_one, mul_one] using edist_le_of_edist_le_geometric_two_of_tendsto C hu ha 0 end edist_le_geometric_two diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 98aa329bbfd1f..e0de967a50f89 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -1018,11 +1018,9 @@ by rw [div_eq_mul_inv, div_eq_mul_inv, coe_mul, coe_inv hr] lemma div_zero (h : a ≠ 0) : a / 0 = ∞ := by simp [div_eq_mul_inv, h] -@[simp] lemma inv_one : (1 : ℝ≥0∞)⁻¹ = 1 := -by simpa only [coe_inv one_ne_zero, coe_one] using coe_eq_coe.2 inv_one - -@[simp] lemma div_one {a : ℝ≥0∞} : a / 1 = a := -by rw [div_eq_mul_inv, inv_one, mul_one] +instance : div_inv_one_monoid ℝ≥0∞ := +{ inv_one := by simpa only [coe_inv one_ne_zero, coe_one] using coe_eq_coe.2 inv_one, + ..ennreal.div_inv_monoid } protected lemma inv_pow {n : ℕ} : (a^n)⁻¹ = (a⁻¹)^n := begin diff --git a/src/measure_theory/integral/average.lean b/src/measure_theory/integral/average.lean index 1e61e65b5d9cd..191635b0bd6cd 100644 --- a/src/measure_theory/integral/average.lean +++ b/src/measure_theory/integral/average.lean @@ -78,7 +78,7 @@ by rw [average_def, integral_smul_measure, ennreal.to_real_inv] lemma average_eq_integral [is_probability_measure μ] (f : α → E) : ⨍ x, f x ∂μ = ∫ x, f x ∂μ := -by rw [average, measure_univ, ennreal.inv_one, one_smul] +by rw [average, measure_univ, inv_one, one_smul] @[simp] lemma measure_smul_average [is_finite_measure μ] (f : α → E) : (μ univ).to_real • ⨍ x, f x ∂μ = ∫ x, f x ∂μ := diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar.lean index 8a3501d6c7147..f2d5b992b74f4 100644 --- a/src/measure_theory/measure/haar.lean +++ b/src/measure_theory/measure/haar.lean @@ -595,7 +595,7 @@ theorem haar_measure_unique (μ : measure G) [sigma_finite μ] [is_mul_left_inva (K₀ : positive_compacts G) : μ = μ K₀ • haar_measure K₀ := (measure_eq_div_smul μ (haar_measure K₀) K₀.compact.measurable_set (measure_pos_of_nonempty_interior _ K₀.interior_nonempty).ne' - K₀.compact.measure_lt_top.ne).trans (by rw [haar_measure_self, ennreal.div_one]) + K₀.compact.measure_lt_top.ne).trans (by rw [haar_measure_self, div_one]) example [locally_compact_space G] (μ : measure G) [is_haar_measure μ] (K₀ : positive_compacts G) : μ = μ K₀.1 • haar_measure K₀ := diff --git a/src/probability/cond_count.lean b/src/probability/cond_count.lean index 86f1aef08120e..9787b2d47b859 100644 --- a/src/probability/cond_count.lean +++ b/src/probability/cond_count.lean @@ -87,7 +87,7 @@ lemma cond_count_singleton (ω : Ω) (t : set Ω) [decidable (ω ∈ t)] : cond_count {ω} t = if ω ∈ t then 1 else 0 := begin rw [cond_count, cond_apply _ (measurable_set_singleton ω), measure.count_singleton, - ennreal.inv_one, one_mul], + inv_one, one_mul], split_ifs, { rw [(by simpa : ({ω} : set Ω) ∩ t = {ω}), measure.count_singleton] }, { rw [(by simpa : ({ω} : set Ω) ∩ t = ∅), measure.count_empty] }, diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index 6f326924b6214..9c5c3117657dd 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -345,7 +345,7 @@ lemma dimH_image_le_of_locally_lipschitz_on [second_countable_topology X] {f : X begin have : ∀ x ∈ s, ∃ (C : ℝ≥0) (t ∈ 𝓝[s] x), holder_on_with C 1 f t, by simpa only [holder_on_with_one] using hf, - simpa only [ennreal.coe_one, ennreal.div_one] + simpa only [ennreal.coe_one, div_one] using dimH_image_le_of_locally_holder_on zero_lt_one this end From d95851bcbf42b87d47645a10564c41116b444e7b Mon Sep 17 00:00:00 2001 From: negiizhao Date: Thu, 29 Sep 2022 06:46:39 +0000 Subject: [PATCH 445/828] chore(data/finset/lattice): use more common name, fix spaces (#16336) `coe_le_max_of_mem` -> `le_max` `le_max_of_mem` -> `le_max_of_eq` `min_le_coe_of_mem` -> `min_le` `min_le_of_mem` -> `min_le_of_eq` `coe_le_max_of_mem` is an analogue of `le_sup` and `min_le_coe_of_mem` is an analogue of `inf_le`, new names are more consistent with them. Co-authored-by: Floris van Doorn --- .../73_ascending_descending_sequences.lean | 6 +- src/combinatorics/simple_graph/basic.lean | 4 +- src/data/finset/lattice.lean | 68 +++++++++---------- 3 files changed, 39 insertions(+), 39 deletions(-) diff --git a/archive/100-theorems-list/73_ascending_descending_sequences.lean b/archive/100-theorems-list/73_ascending_descending_sequences.lean index 532713d49e552..7bcdec52f837d 100644 --- a/archive/100-theorems-list/73_ascending_descending_sequences.lean +++ b/archive/100-theorems-list/73_ascending_descending_sequences.lean @@ -115,16 +115,16 @@ begin rintros x ⟨rfl | _⟩ y ⟨rfl | _⟩ _, { apply (irrefl _ ‹j < j›).elim }, { exfalso, - apply not_le_of_lt (trans ‹i < j› ‹j < y›) (le_max_of_mem ‹y ∈ t› ‹t.max = i›) }, + apply not_le_of_lt (trans ‹i < j› ‹j < y›) (le_max_of_eq ‹y ∈ t› ‹t.max = i›) }, { apply lt_of_le_of_lt _ ‹f i < f j› <|> apply lt_of_lt_of_le ‹f j < f i› _, - rcases lt_or_eq_of_le (le_max_of_mem ‹x ∈ t› ‹t.max = i›) with _ | rfl, + rcases lt_or_eq_of_le (le_max_of_eq ‹x ∈ t› ‹t.max = i›) with _ | rfl, { apply le_of_lt (ht₁.2.2 ‹x ∈ t› (mem_of_max ‹t.max = i›) ‹x < i›) }, { refl } }, { apply ht₁.2.2 ‹x ∈ t› ‹y ∈ t› ‹x < y› } }, -- Finally show that this new subsequence is one longer than the old one. { rw [card_insert_of_not_mem, ht₂], intro _, - apply not_le_of_lt ‹i < j› (le_max_of_mem ‹j ∈ t› ‹t.max = i›) } } }, + apply not_le_of_lt ‹i < j› (le_max_of_eq ‹j ∈ t› ‹t.max = i›) } } }, -- Finished both goals! -- Now that we have uniqueness of each label, it remains to do some counting to finish off. -- Suppose all the labels are small. diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index 1676bc9ca8183..75fad3a99c2a3 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -924,7 +924,7 @@ end lemma min_degree_le_degree [decidable_rel G.adj] (v : V) : G.min_degree ≤ G.degree v := begin obtain ⟨t, ht⟩ := finset.min_of_mem (mem_image_of_mem (λ v, G.degree v) (mem_univ v)), - have := finset.min_le_of_mem (mem_image_of_mem _ (mem_univ v)) ht, + have := finset.min_le_of_eq (mem_image_of_mem _ (mem_univ v)) ht, rwa [min_degree, ht] end @@ -966,7 +966,7 @@ end lemma degree_le_max_degree [decidable_rel G.adj] (v : V) : G.degree v ≤ G.max_degree := begin obtain ⟨t, ht : _ = _⟩ := finset.max_of_mem (mem_image_of_mem (λ v, G.degree v) (mem_univ v)), - have := finset.le_max_of_mem (mem_image_of_mem _ (mem_univ v)) ht, + have := finset.le_max_of_eq (mem_image_of_mem _ (mem_univ v)) ht, rwa [max_degree, ht], end diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index dec7cb992bdec..85e35a175df44 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -62,10 +62,10 @@ begin exact sup_sup_sup_comm _ _ _ _ } end -theorem sup_congr {f g : β → α} (hs : s₁ = s₂) (hfg : ∀a∈s₂, f a = g a) : s₁.sup f = s₂.sup g := +theorem sup_congr {f g : β → α} (hs : s₁ = s₂) (hfg : ∀ a ∈ s₂, f a = g a) : s₁.sup f = s₂.sup g := by subst hs; exact finset.fold_congr hfg -@[simp] protected lemma sup_le_iff {a : α} : s.sup f ≤ a ↔ (∀b ∈ s, f b ≤ a) := +@[simp] protected lemma sup_le_iff {a : α} : s.sup f ≤ a ↔ (∀ b ∈ s, f b ≤ a) := begin apply iff.trans multiset.sup_le, simp only [multiset.mem_map, and_imp, exists_imp_distrib], @@ -91,13 +91,13 @@ lemma sup_ite (p : β → Prop) [decidable_pred p] : (s.filter p).sup f ⊔ (s.filter (λ i, ¬ p i)).sup g := fold_ite _ -lemma sup_le {a : α} : (∀b ∈ s, f b ≤ a) → s.sup f ≤ a := +lemma sup_le {a : α} : (∀ b ∈ s, f b ≤ a) → s.sup f ≤ a := finset.sup_le_iff.2 lemma le_sup {b : β} (hb : b ∈ s) : f b ≤ s.sup f := finset.sup_le_iff.1 le_rfl _ hb -lemma sup_mono_fun {g : β → α} (h : ∀b∈s, f b ≤ g b) : s.sup f ≤ s.sup g := +lemma sup_mono_fun {g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : s.sup f ≤ s.sup g := sup_le (λ b hb, le_trans (h b hb) (le_sup hb)) lemma sup_mono (h : s₁ ⊆ s₂) : s₁.sup f ≤ s₂.sup f := @@ -154,7 +154,7 @@ finset.cons_induction_on s bot (λ c t hc ih, by rw [sup_cons, sup_cons, g_sup, /-- Computing `sup` in a subtype (closed under `sup`) is the same as computing it in `α`. -/ lemma sup_coe {P : α → Prop} - {Pbot : P ⊥} {Psup : ∀{{x y}}, P x → P y → P (x ⊔ y)} + {Pbot : P ⊥} {Psup : ∀ {{x y}}, P x → P y → P (x ⊔ y)} (t : finset β) (f : β → {x : α // P x}) : (@sup _ _ (subtype.semilattice_sup Psup) (subtype.order_bot Pbot) t f : α) = t.sup (λ x, f x) := by { rw [comp_sup_eq_sup_comp coe]; intros; refl } @@ -229,7 +229,7 @@ end end sup -lemma sup_eq_supr [complete_lattice β] (s : finset α) (f : α → β) : s.sup f = (⨆a∈s, f a) := +lemma sup_eq_supr [complete_lattice β] (s : finset α) (f : α → β) : s.sup f = (⨆ a ∈ s, f a) := le_antisymm (finset.sup_le $ assume a ha, le_supr_of_le a $ le_supr _ ha) (supr_le $ assume a, supr_le $ assume ha, le_sup ha) @@ -287,7 +287,7 @@ lemma inf_union [decidable_eq β] : (s₁ ∪ s₂).inf f = s₁.inf f ⊓ s₂. lemma inf_inf : s.inf (f ⊓ g) = s.inf f ⊓ s.inf g := @sup_sup αᵒᵈ _ _ _ _ _ _ -theorem inf_congr {f g : β → α} (hs : s₁ = s₂) (hfg : ∀a∈s₂, f a = g a) : s₁.inf f = s₂.inf g := +theorem inf_congr {f g : β → α} (hs : s₁ = s₂) (hfg : ∀ a ∈ s₂, f a = g a) : s₁.inf f = s₂.inf g := by subst hs; exact finset.fold_congr hfg @[simp] lemma inf_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : @@ -305,10 +305,10 @@ lemma le_inf_iff {a : α} : a ≤ s.inf f ↔ ∀ b ∈ s, a ≤ f b := lemma inf_le {b : β} (hb : b ∈ s) : s.inf f ≤ f b := le_inf_iff.1 le_rfl _ hb -lemma le_inf {a : α} : (∀b ∈ s, a ≤ f b) → a ≤ s.inf f := +lemma le_inf {a : α} : (∀ b ∈ s, a ≤ f b) → a ≤ s.inf f := le_inf_iff.2 -lemma inf_mono_fun {g : β → α} (h : ∀b∈s, f b ≤ g b) : s.inf f ≤ s.inf g := +lemma inf_mono_fun {g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : s.inf f ≤ s.inf g := le_inf (λ b hb, le_trans (inf_le hb) (h b hb)) lemma inf_mono (h : s₁ ⊆ s₂) : s₂.inf f ≤ s₁.inf f := @@ -365,7 +365,7 @@ lemma comp_inf_eq_inf_comp [semilattice_inf γ] [order_top γ] {s : finset β} /-- Computing `inf` in a subtype (closed under `inf`) is the same as computing it in `α`. -/ lemma inf_coe {P : α → Prop} - {Ptop : P ⊤} {Pinf : ∀{{x y}}, P x → P y → P (x ⊓ y)} + {Ptop : P ⊤} {Pinf : ∀ {{x y}}, P x → P y → P (x ⊓ y)} (t : finset β) (f : β → {x : α // P x}) : (@inf _ _ (subtype.semilattice_inf Pinf) (subtype.order_top Ptop) t f : α) = t.inf (λ x, f x) := @sup_coe αᵒᵈ _ _ _ _ Ptop Pinf t f @@ -496,7 +496,7 @@ lemma inf_eq_infi [complete_lattice β] (s : finset α) (f : α → β) : s.inf lemma inf_id_eq_Inf [complete_lattice α] (s : finset α) : s.inf id = Inf s := @sup_id_eq_Sup αᵒᵈ _ _ -lemma inf_id_set_eq_sInter (s : finset (set α)) : s.inf id = ⋂₀(↑s) := +lemma inf_id_set_eq_sInter (s : finset (set α)) : s.inf id = ⋂₀ ↑s := inf_id_eq_Inf _ @[simp] lemma inf_set_eq_bInter (s : finset α) (f : α → set β) : s.inf f = ⋂ x ∈ s, f x := @@ -845,14 +845,14 @@ finset.induction_on s (λ _ H, by cases H) { exact mem_insert_of_mem (ih h) } } end) -lemma coe_le_max_of_mem {a : α} {s : finset α} (as : a ∈ s) : ↑a ≤ s.max := +lemma le_max {a : α} {s : finset α} (as : a ∈ s) : ↑a ≤ s.max := le_sup as lemma not_mem_of_max_lt_coe {a : α} {s : finset α} (h : s.max < a) : a ∉ s := -mt coe_le_max_of_mem h.not_le +mt le_max h.not_le -theorem le_max_of_mem {s : finset α} {a b : α} (h₁ : a ∈ s) (h₂ : s.max = b) : a ≤ b := -with_bot.coe_le_coe.mp $ (coe_le_max_of_mem h₁).trans h₂.le +theorem le_max_of_eq {s : finset α} {a b : α} (h₁ : a ∈ s) (h₂ : s.max = b) : a ≤ b := +with_bot.coe_le_coe.mp $ (le_max h₁).trans h₂.le theorem not_mem_of_max_lt {s : finset α} {a b : α} (h₁ : b < a) (h₂ : s.max = ↑b) : a ∉ s := finset.not_mem_of_max_lt_coe $ h₂.trans_lt $ with_bot.coe_lt_coe.mpr h₁ @@ -895,17 +895,17 @@ theorem min_eq_top {s : finset α} : s.min = ⊤ ↔ s = ∅ := theorem mem_of_min {s : finset α} : ∀ {a : α}, s.min = a → a ∈ s := @mem_of_max αᵒᵈ _ s -lemma min_le_coe_of_mem {a : α} {s : finset α} (as : a ∈ s) : s.min ≤ a := +lemma min_le {a : α} {s : finset α} (as : a ∈ s) : s.min ≤ a := inf_le as lemma not_mem_of_coe_lt_min {a : α} {s : finset α} (h : ↑a < s.min) : a ∉ s := -mt min_le_coe_of_mem h.not_le +mt min_le h.not_le -theorem min_le_of_mem {s : finset α} {a b : α} (h₁ : b ∈ s) (h₂ : s.min = a) : a ≤ b := -with_top.coe_le_coe.mp $ h₂.ge.trans (min_le_coe_of_mem h₁) +theorem min_le_of_eq {s : finset α} {a b : α} (h₁ : b ∈ s) (h₂ : s.min = a) : a ≤ b := +with_top.coe_le_coe.mp $ h₂.ge.trans (min_le h₁) theorem not_mem_of_lt_min {s : finset α} {a b : α} (h₁ : a < b) (h₂ : s.min = ↑b) : a ∉ s := -finset.not_mem_of_coe_lt_min $ ( with_top.coe_lt_coe.mpr h₁).trans_eq h₂.symm +finset.not_mem_of_coe_lt_min $ (with_top.coe_lt_coe.mpr h₁).trans_eq h₂.symm lemma min_mono {s t : finset α} (st : s ⊆ t) : t.min ≤ s.min := inf_mono st @@ -914,13 +914,13 @@ lemma le_min {m : with_top α} {s : finset α} (st : ∀ a : α, a ∈ s → m m ≤ s.min := le_inf st -/-- Given a nonempty finset `s` in a linear order `α `, then `s.min' h` is its minimum, as an +/-- Given a nonempty finset `s` in a linear order `α`, then `s.min' h` is its minimum, as an element of `α`, where `h` is a proof of nonemptiness. Without this assumption, use instead `s.min`, taking values in `with_top α`. -/ def min' (s : finset α) (H : s.nonempty) : α := inf' s H id -/-- Given a nonempty finset `s` in a linear order `α `, then `s.max' h` is its maximum, as an +/-- Given a nonempty finset `s` in a linear order `α`, then `s.max' h` is its maximum, as an element of `α`, where `h` is a proof of nonemptiness. Without this assumption, use instead `s.max`, taking values in `with_bot α`. -/ def max' (s : finset α) (H : s.nonempty) : α := @@ -931,7 +931,7 @@ variables (s : finset α) (H : s.nonempty) {x : α} theorem min'_mem : s.min' H ∈ s := mem_of_min $ by simp [min', finset.min] theorem min'_le (x) (H2 : x ∈ s) : s.min' ⟨x, H2⟩ ≤ x := -min_le_of_mem H2 (with_top.coe_untop _ _).symm +min_le_of_eq H2 (with_top.coe_untop _ _).symm theorem le_min' (x) (H2 : ∀ y ∈ s, x ≤ y) : x ≤ s.min' H := H2 _ $ min'_mem _ _ @@ -948,7 +948,7 @@ by simp [min'] theorem max'_mem : s.max' H ∈ s := mem_of_max $ by simp [max', finset.max] theorem le_max' (x) (H2 : x ∈ s) : x ≤ s.max' ⟨x, H2⟩ := -le_max_of_mem H2 (with_bot.coe_unbot _ _).symm +le_max_of_eq H2 (with_bot.coe_unbot _ _).symm theorem max'_le (x) (H2 : ∀ y ∈ s, y ≤ x) : s.max' H ≤ x := H2 _ $ max'_mem _ _ @@ -1165,7 +1165,7 @@ lemma exists_max_image (s : finset β) (f : β → α) (h : s.nonempty) : begin cases max_of_nonempty (h.image f) with y hy, rcases mem_image.mp (mem_of_max hy) with ⟨x, hx, rfl⟩, - exact ⟨x, hx, λ x' hx', le_max_of_mem (mem_image_of_mem f hx') hy⟩, + exact ⟨x, hx, λ x' hx', le_max_of_eq (mem_image_of_mem f hx') hy⟩, end lemma exists_min_image (s : finset β) (f : β → α) (h : s.nonempty) : @@ -1248,7 +1248,7 @@ variables {ι' : Sort*} [complete_lattice α] `⨆ i ∈ t, s i`. This version assumes `ι` is a `Type*`. See `supr_eq_supr_finset'` for a version that works for `ι : Sort*`. -/ lemma supr_eq_supr_finset (s : ι → α) : - (⨆i, s i) = (⨆t:finset ι, ⨆i∈t, s i) := + (⨆ i, s i) = (⨆ t : finset ι, ⨆ i ∈ t, s i) := begin classical, exact le_antisymm @@ -1261,7 +1261,7 @@ end `⨆ i ∈ t, s i`. This version works for `ι : Sort*`. See `supr_eq_supr_finset` for a version that assumes `ι : Type*` but has no `plift`s. -/ lemma supr_eq_supr_finset' (s : ι' → α) : - (⨆i, s i) = (⨆t:finset (plift ι'), ⨆i∈t, s (plift.down i)) := + (⨆ i, s i) = (⨆ t : finset (plift ι'), ⨆ i ∈ t, s (plift.down i)) := by rw [← supr_eq_supr_finset, ← equiv.plift.surjective.supr_comp]; refl /-- Infimum of `s i`, `i : ι`, is equal to the infimum over `t : finset ι` of infima @@ -1274,7 +1274,7 @@ lemma infi_eq_infi_finset (s : ι → α) : (⨅ i, s i) = ⨅ (t : finset ι) ( `⨅ i ∈ t, s i`. This version works for `ι : Sort*`. See `infi_eq_infi_finset` for a version that assumes `ι : Type*` but has no `plift`s. -/ lemma infi_eq_infi_finset' (s : ι' → α) : - (⨅i, s i) = (⨅t:finset (plift ι'), ⨅i∈t, s (plift.down i)) := + (⨅ i, s i) = (⨅ t : finset (plift ι'), ⨅ i ∈ t, s (plift.down i)) := @supr_eq_supr_finset' αᵒᵈ _ _ _ end lattice @@ -1286,21 +1286,21 @@ variables {ι' : Sort*} of finite subfamilies. This version assumes `ι : Type*`. See also `Union_eq_Union_finset'` for a version that works for `ι : Sort*`. -/ lemma Union_eq_Union_finset (s : ι → set α) : - (⋃i, s i) = (⋃t:finset ι, ⋃i∈t, s i) := + (⋃ i, s i) = (⋃ t : finset ι, ⋃ i ∈ t, s i) := supr_eq_supr_finset s /-- Union of an indexed family of sets `s : ι → set α` is equal to the union of the unions of finite subfamilies. This version works for `ι : Sort*`. See also `Union_eq_Union_finset` for a version that assumes `ι : Type*` but avoids `plift`s in the right hand side. -/ lemma Union_eq_Union_finset' (s : ι' → set α) : - (⋃i, s i) = (⋃t:finset (plift ι'), ⋃i∈t, s (plift.down i)) := + (⋃ i, s i) = (⋃ t : finset (plift ι'), ⋃ i ∈ t, s (plift.down i)) := supr_eq_supr_finset' s /-- Intersection of an indexed family of sets `s : ι → set α` is equal to the intersection of the intersections of finite subfamilies. This version assumes `ι : Type*`. See also `Inter_eq_Inter_finset'` for a version that works for `ι : Sort*`. -/ lemma Inter_eq_Inter_finset (s : ι → set α) : - (⋂i, s i) = (⋂t:finset ι, ⋂i∈t, s i) := + (⋂ i, s i) = (⋂ t : finset ι, ⋂ i ∈ t, s i) := infi_eq_infi_finset s /-- Intersection of an indexed family of sets `s : ι → set α` is equal to the intersection of the @@ -1308,7 +1308,7 @@ intersections of finite subfamilies. This version works for `ι : Sort*`. See al `Inter_eq_Inter_finset` for a version that assumes `ι : Type*` but avoids `plift`s in the right hand side. -/ lemma Inter_eq_Inter_finset' (s : ι' → set α) : - (⋂i, s i) = (⋂t:finset (plift ι'), ⋂i∈t, s (plift.down i)) := + (⋂ i, s i) = (⋂ t : finset (plift ι'), ⋂ i ∈ t, s (plift.down i)) := infi_eq_infi_finset' s end set @@ -1370,7 +1370,7 @@ lemma infi_option_to_finset (o : option α) (f : α → β) : (⨅ x ∈ o.to_fi variables [decidable_eq α] theorem supr_union {f : α → β} {s t : finset α} : - (⨆ x ∈ s ∪ t, f x) = (⨆x∈s, f x) ⊔ (⨆x∈t, f x) := + (⨆ x ∈ s ∪ t, f x) = (⨆ x ∈ s, f x) ⊔ (⨆ x ∈ t, f x) := by simp [supr_or, supr_sup_eq] theorem infi_union {f : α → β} {s t : finset α} : @@ -1473,7 +1473,7 @@ lemma set_bInter_insert (a : α) (s : finset α) (t : α → set β) : infi_insert a s t lemma set_bUnion_finset_image {f : γ → α} {g : α → set β} {s : finset γ} : - (⋃x ∈ s.image f, g x) = (⋃y ∈ s, g (f y)) := + (⋃ x ∈ s.image f, g x) = (⋃ y ∈ s, g (f y)) := supr_finset_image lemma set_bInter_finset_image {f : γ → α} {g : α → set β} {s : finset γ} : From 7944d185129638f4c2fa8b8a915537eb88c5ff1f Mon Sep 17 00:00:00 2001 From: Tian Chen <27919714+peakpoint@users.noreply.github.com> Date: Thu, 29 Sep 2022 10:48:58 +0000 Subject: [PATCH 446/828] feat(number_theory/multiplicity): lifting the exponent lemma (#8915) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `multiplicity.int.pow_sub_pow` is the [lifting the exponent lemma](https://en.wikipedia.org/wiki/Lifting-the-exponent_lemma) for odd primes. Some variations are also proved. Co-authored-by: Mantas Baksys Co-authored-by: YaelDillies Co-authored-by: MantasBaksys Co-authored-by: Mantas Bakšys <39908973+MantasBaksys@users.noreply.github.com> --- src/algebra/geom_sum.lean | 15 + src/number_theory/multiplicity.lean | 408 ++++++++++++++++++++++++++++ 2 files changed, 423 insertions(+) create mode 100644 src/number_theory/multiplicity.lean diff --git a/src/algebra/geom_sum.lean b/src/algebra/geom_sum.lean index d4e7a76f46407..666185670230d 100644 --- a/src/algebra/geom_sum.lean +++ b/src/algebra/geom_sum.lean @@ -197,6 +197,21 @@ lemma mul_neg_geom_sum [ring α] (x : α) (n : ℕ) : (1 - x) * (∑ i in range n, x ^ i) = 1 - x ^ n := op_injective $ by simpa using geom_sum_mul_neg (op x) n +protected lemma commute.geom_sum₂_comm {α : Type u} [semiring α] {x y : α} (n : ℕ) + (h : commute x y) : + ∑ i in range n, x ^ i * y ^ (n - 1 - i) = ∑ i in range n, y ^ i * x ^ (n - 1 - i) := +begin + cases n, { simp }, + simp only [nat.succ_eq_add_one, nat.add_sub_cancel], + rw ← finset.sum_flip, + refine finset.sum_congr rfl (λ i hi, _), + simpa [nat.sub_sub_self (nat.succ_le_succ_iff.mp (finset.mem_range.mp hi))] using h.pow_pow _ _ +end + +lemma geom_sum₂_comm {α : Type u} [comm_semiring α] (x y : α) (n : ℕ) : + ∑ i in range n, x ^ i * y ^ (n - 1 - i) = ∑ i in range n, y ^ i * x ^ (n - 1 - i) := +(commute.all x y).geom_sum₂_comm n + protected theorem commute.geom_sum₂ [division_ring α] {x y : α} (h' : commute x y) (h : x ≠ y) (n : ℕ) : (∑ i in range n, x ^ i * (y ^ (n - 1 - i))) = (x ^ n - y ^ n) / (x - y) := have x - y ≠ 0, by simp [*, -sub_eq_add_neg, sub_eq_iff_eq_add] at *, diff --git a/src/number_theory/multiplicity.lean b/src/number_theory/multiplicity.lean new file mode 100644 index 0000000000000..f8aa3ff4c324a --- /dev/null +++ b/src/number_theory/multiplicity.lean @@ -0,0 +1,408 @@ +/- +Copyright (c) 2022 Tian Chen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Tian Chen, Mantas Bakšys +-/ +import algebra.geom_sum +import data.int.parity +import number_theory.padics.padic_norm +import data.zmod.basic + +/-! +# Multiplicity in Number Theory + +This file contains results in number theory relating to multiplicity. + +## Main statements + +* `multiplicity.int.pow_sub_pow` is the lifting the exponent lemma for odd primes. + We also prove several variations of the lemma. + +## References + +* [Wikipedia, *Lifting-the-exponent lemma*] + (https://en.wikipedia.org/wiki/Lifting-the-exponent_lemma) +-/ + +open ideal ideal.quotient finset +open_locale big_operators + +variables {R : Type*} {n : ℕ} + + +section comm_ring +variables [comm_ring R] {a b x y : R} + + +lemma dvd_geom_sum₂_iff_of_dvd_sub {x y p : R} (h : p ∣ x - y) : + p ∣ ∑ i in range n, x ^ i * y ^ (n - 1 - i) ↔ p ∣ n * y ^ (n - 1) := +begin + rw [← mem_span_singleton, ← ideal.quotient.eq] at h, + simp only [← mem_span_singleton, ← eq_zero_iff_mem, ring_hom.map_geom_sum₂, h, geom_sum₂_self, + _root_.map_mul, map_pow, map_nat_cast] +end + +lemma dvd_geom_sum₂_iff_of_dvd_sub' {x y p : R} (h : p ∣ x - y) : + p ∣ ∑ i in range n, x ^ i * y ^ (n - 1 - i) ↔ p ∣ n * x ^ (n - 1) := +by rw [geom_sum₂_comm, dvd_geom_sum₂_iff_of_dvd_sub]; simpa using (dvd_neg _ _).mpr h + +lemma dvd_geom_sum₂_self {x y : R} (h : ↑n ∣ x - y) : ↑n ∣ ∑ i in range n, x ^ i * y ^ (n - 1 - i):= +(dvd_geom_sum₂_iff_of_dvd_sub h).mpr (dvd_mul_right _ _) + +lemma sq_dvd_add_pow_sub_sub (p x : R) (n : ℕ) : + p ^ 2 ∣ (x + p) ^ n - x ^ (n - 1) * p * n - x ^ n := +begin + cases n, + { simp only [pow_zero, nat.cast_zero, mul_zero, sub_zero, sub_self, dvd_zero]}, + { simp only [nat.succ_sub_succ_eq_sub, tsub_zero, nat.cast_succ, add_pow, + finset.sum_range_succ, nat.choose_self, nat.succ_sub _, tsub_self, pow_one, + nat.choose_succ_self_right, pow_zero, mul_one, nat.cast_zero, zero_add, nat.succ_eq_add_one], + suffices : p ^ 2 ∣ ∑ (i : ℕ) in range n, x ^ i * p ^ (n + 1 - i) * ↑((n + 1).choose i), + { convert this; abel }, + { apply finset.dvd_sum, + intros y hy, + calc p ^ 2 ∣ p ^ (n + 1 - y) : pow_dvd_pow p (le_tsub_of_add_le_left + (by linarith [finset.mem_range.mp hy])) + ... ∣ x ^ y * p ^ (n + 1 - y) * ↑((n + 1).choose y) : dvd_mul_of_dvd_left (dvd_mul_left _ _) + ((n + 1).choose y) }} +end + +lemma not_dvd_geom_sum₂ {p : R} (hp : prime p) + (hxy : p ∣ x - y) (hx : ¬p ∣ x) (hn : ¬p ∣ n) : + ¬p ∣ ∑ i in range n, x ^ i * y ^ (n - 1 - i) := +λ h, hx $ hp.dvd_of_dvd_pow $ +(hp.dvd_or_dvd $ (dvd_geom_sum₂_iff_of_dvd_sub' hxy).mp h).resolve_left hn + +variables {p : ℕ} (a b) + +lemma odd_sq_dvd_geom_sum₂_sub (hp : odd p) : + ↑p ^ 2 ∣ ∑ i in range p, (a + p * b) ^ i * a ^ (p - 1 - i) - p * a ^ (p - 1) := +begin + have h1 : ∀ i, ↑p ^ 2 ∣ (a + ↑p * b) ^ i - (a ^ (i - 1) * (↑p * b) * ↑i + a ^ i), + { intro i, + calc ↑p ^ 2 ∣ (↑p * b) ^ 2 : by simp only [mul_pow, dvd_mul_right] + ... ∣ (a + ↑p * b) ^ i - (a ^ (i - 1) * (↑p * b) * ↑i + a ^ i) : + by simp only [sq_dvd_add_pow_sub_sub (↑p * b) a i, ← sub_sub] }, + simp_rw [← mem_span_singleton, ← ideal.quotient.eq] at *, + calc ideal.quotient.mk (span {↑p ^ 2}) (∑ i in range p, (a + ↑p * b) ^ i * a ^ (p - 1 - i)) + = ∑ (i : ℕ) in finset.range p, mk (span {↑p ^ 2}) + ((a ^ (i - 1) * (↑p * b) * ↑i + a ^ i) * a ^ (p - 1 - i)) : + by simp_rw [ring_hom.map_geom_sum₂, ← map_pow, h1, ← _root_.map_mul] + ... = mk (span {↑p ^ 2}) (∑ (x : ℕ) in finset.range p, + a ^ (x - 1) * (a ^ (p - 1 - x) * (↑p * (b * ↑x)))) + + mk (span {↑p ^ 2}) (∑ (x : ℕ) in finset.range p, a ^ (x + (p - 1 - x))) : + by { ring_exp, + simp only [← pow_add, map_add, finset.sum_add_distrib, ← map_sum] } + ... = mk (span {↑p ^ 2}) (∑ (x : ℕ) in finset.range p, + a ^ (x - 1) * (a ^ (p - 1 - x) * (↑p * (b * ↑x)))) + + mk (span {↑ p ^ 2}) ∑ (x : ℕ) in finset.range p, a ^ (p - 1) : + by { rw [add_right_inj, finset.sum_congr rfl], + intros x hx, + rw [← nat.add_sub_assoc _ x, nat.add_sub_cancel_left], + exact nat.le_pred_of_lt (finset.mem_range.mp hx) } + ... = mk (span {↑p ^ 2}) (∑ (x : ℕ) in finset.range p, + a ^ (x - 1) * (a ^ (p - 1 - x) * (↑p * (b * ↑x)))) + + mk (span {↑ p ^ 2}) (↑p * a ^ (p - 1)) : + by simp only [add_right_inj, finset.sum_const, finset.card_range, nsmul_eq_mul] + ... = mk (span {↑p ^ 2}) (↑p * b * ∑ (x : ℕ) in finset.range p, a ^ (p - 2) * x) + + mk (span {↑p ^ 2}) (↑p * a ^ (p - 1)) : + by { simp only [finset.mul_sum, ← mul_assoc, ← pow_add], + rw finset.sum_congr rfl, + rintros (⟨⟩|⟨x⟩) hx, + { rw [nat.cast_zero, mul_zero, mul_zero] }, + { have : x.succ - 1 + (p - 1 - x.succ) = p - 2, + { rw ← nat.add_sub_assoc (nat.le_pred_of_lt (finset.mem_range.mp hx)), + exact congr_arg nat.pred (nat.add_sub_cancel_left _ _)}, + rw this, + ring_exp_eq }} + ... = mk (span {↑p ^ 2}) (↑p * a ^ (p - 1)) : + by { simp only [add_left_eq_self, ← finset.mul_sum], + norm_cast, + simp only [finset.sum_range_id, nat.cast_mul, _root_.map_mul, + nat.mul_div_assoc _ (even_iff_two_dvd.mp (nat.odd.sub_odd hp odd_one))], + ring_exp, + simp only [← map_pow, mul_eq_zero_of_left, ideal.quotient.eq_zero_iff_mem, + mem_span_singleton] } +end + + +namespace multiplicity + +section integral_domain +variables [is_domain R] [@decidable_rel R (∣)] + +lemma pow_sub_pow_of_prime {p : R} (hp : prime p) {x y : R} (hxy : p ∣ x - y) (hx : ¬p ∣ x) + {n : ℕ} (hn : ¬p ∣ n) : + multiplicity p (x ^ n - y ^ n) = multiplicity p (x - y) := +by rw [←geom_sum₂_mul, multiplicity.mul hp, + multiplicity_eq_zero_of_not_dvd (not_dvd_geom_sum₂ hp hxy hx hn), zero_add] + +variables (hp : prime (p : R)) (hp1 : odd p) (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) +include hp hp1 hxy hx + +lemma geom_sum₂_eq_one : multiplicity ↑p (∑ i in range p, x ^ i * y ^ (p - 1 - i)) = 1 := +begin + rw ← nat.cast_one, + refine multiplicity.eq_coe_iff.2 ⟨_, _⟩, + { rw pow_one, + exact dvd_geom_sum₂_self hxy }, + rw dvd_iff_dvd_of_dvd_sub hxy at hx, + cases hxy with k hk, + rw [one_add_one_eq_two, eq_add_of_sub_eq' hk], + refine mt (dvd_iff_dvd_of_dvd_sub (@odd_sq_dvd_geom_sum₂_sub _ _ y k _ hp1)).mp _, + rw [pow_two, mul_dvd_mul_iff_left hp.ne_zero], + exact mt hp.dvd_of_dvd_pow hx +end + +lemma pow_prime_sub_pow_prime : multiplicity ↑p (x ^ p - y ^ p) = multiplicity ↑p (x - y) + 1 := +by rw [←geom_sum₂_mul, multiplicity.mul hp, geom_sum₂_eq_one hp hp1 hxy hx, add_comm] + +lemma pow_prime_pow_sub_pow_prime_pow (a : ℕ) : + multiplicity ↑p (x ^ p ^ a - y ^ p ^ a) = multiplicity ↑p (x - y) + a := +begin + induction a with a h_ind, + { rw [nat.cast_zero, add_zero, pow_zero, pow_one, pow_one] }, + rw [←nat.add_one, nat.cast_add, nat.cast_one, ←add_assoc, ←h_ind, pow_succ', pow_mul, pow_mul], + apply pow_prime_sub_pow_prime hp hp1, + { rw ←geom_sum₂_mul, + exact dvd_mul_of_dvd_right hxy _ }, + { exact λ h, hx (hp.dvd_of_dvd_pow h) } +end + +end integral_domain + +section lifting_the_exponent +variables (hp : nat.prime p) (hp1 : odd p) +include hp hp1 + +/-- **Lifting the exponent lemma** for odd primes. -/ +lemma int.pow_sub_pow {x y : ℤ} (hxy : ↑p ∣ x - y) (hx : ¬↑p ∣ x) (n : ℕ) : + multiplicity ↑p (x ^ n - y ^ n) = multiplicity ↑p (x - y) + multiplicity p n := +begin + cases n, + { simp only [multiplicity.zero, add_top, pow_zero, sub_self] }, + have h : (multiplicity _ _).dom := finite_nat_iff.mpr ⟨hp.ne_one, n.succ_pos⟩, + rcases eq_coe_iff.mp (part_enat.coe_get h).symm with ⟨⟨k, hk⟩, hpn⟩, + conv_lhs { rw [hk, pow_mul, pow_mul] }, + rw nat.prime_iff_prime_int at hp, + rw [pow_sub_pow_of_prime hp, pow_prime_pow_sub_pow_prime_pow hp hp1 hxy hx, part_enat.coe_get], + { rw ←geom_sum₂_mul, + exact dvd_mul_of_dvd_right hxy _ }, + { exact λ h, hx (hp.dvd_of_dvd_pow h) }, + { rw int.coe_nat_dvd, + rintro ⟨c, rfl⟩, + refine hpn ⟨c, _⟩, + rwa [pow_succ', mul_assoc] } +end +lemma int.pow_add_pow {x y : ℤ} (hxy : ↑p ∣ x + y) (hx : ¬↑p ∣ x) {n : ℕ} (hn : odd n) : + multiplicity ↑p (x ^ n + y ^ n) = multiplicity ↑p (x + y) + multiplicity p n := +begin + rw ←sub_neg_eq_add at hxy, + rw [←sub_neg_eq_add, ←sub_neg_eq_add, ←odd.neg_pow hn], + exact int.pow_sub_pow hp hp1 hxy hx n +end + +lemma nat.pow_sub_pow {x y : ℕ} (hxy : p ∣ x - y) (hx : ¬p ∣ x) (n : ℕ) : + multiplicity p (x ^ n - y ^ n) = multiplicity p (x - y) + multiplicity p n := +begin + obtain hyx | hyx := le_total y x, + { iterate 2 { rw ← int.coe_nat_multiplicity }, + rw int.coe_nat_sub (nat.pow_le_pow_of_le_left hyx n), + rw ← int.coe_nat_dvd at hxy hx, + push_cast at *, + exact int.pow_sub_pow hp hp1 hxy hx n }, + { simp only [nat.sub_eq_zero_iff_le.mpr hyx, + nat.sub_eq_zero_iff_le.mpr (nat.pow_le_pow_of_le_left hyx n), multiplicity.zero, + part_enat.top_add] } +end + +lemma nat.pow_add_pow {x y : ℕ} (hxy : p ∣ x + y) (hx : ¬p ∣ x) {n : ℕ} (hn : odd n) : + multiplicity p (x ^ n + y ^ n) = multiplicity p (x + y) + multiplicity p n := +begin + iterate 2 { rw [←int.coe_nat_multiplicity] }, + rw ←int.coe_nat_dvd at hxy hx, + push_cast at *, + exact int.pow_add_pow hp hp1 hxy hx hn +end + +end lifting_the_exponent +end multiplicity +end comm_ring + +lemma pow_two_pow_sub_pow_two_pow [comm_ring R] {x y : R} (n : ℕ) : + x ^ (2 ^ n) - y ^ (2 ^ n) = (∏ i in finset.range n, (x ^ (2 ^ i) + y ^ (2 ^ i))) * (x - y) := +begin + induction n with d hd, + { simp only [pow_zero, pow_one, finset.range_zero, finset.prod_empty, one_mul] }, + { suffices : x ^ 2 ^ d.succ - y ^ 2 ^ d.succ = (x ^ 2 ^ d + y ^ 2 ^ d) * (x ^ 2 ^ d - y ^ 2 ^ d), + { rw [this, hd, finset.prod_range_succ, ← mul_assoc, mul_comm (x ^ 2 ^ d + y ^ 2 ^ d)] }, + { ring_exp_eq } } +end + +lemma _root_.int.sq_mod_four_eq_one_of_odd {x : ℤ} : odd x → x ^ 2 % 4 = 1 := +begin + intro hx, + -- Replace `x : ℤ` with `y : zmod 4` + replace hx : x % (2 : ℕ) = 1 % (2 : ℕ), { rw int.odd_iff at hx, norm_num [hx] }, + calc x^2 % (4 : ℕ) + = 1 % (4 : ℕ) : _ + ... = 1 : by norm_num, + rw ← zmod.int_coe_eq_int_coe_iff' at hx ⊢, + push_cast, + rw [← map_int_cast (zmod.cast_hom (show 2 ∣ 4, by norm_num) (zmod 2)) x] at hx, + set y : zmod 4 := x, + -- Now we can just consider each of the 4 possible values for y + fin_cases y using hy; + rw hy at ⊢ hx; revert hx; dec_trivial +end + +lemma int.two_pow_two_pow_add_two_pow_two_pow {x y : ℤ} + (hx : ¬ 2 ∣ x) (hxy : 4 ∣ (x - y)) + (i : ℕ) : multiplicity 2 (x ^ 2 ^ i + y ^ 2 ^ i) = ↑(1 : ℕ) := +begin + have hx_odd : odd x, { rwa [int.odd_iff_not_even, even_iff_two_dvd] }, + have hxy_even : even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by norm_num) hxy), + have hy_odd : odd y := by simpa using hx_odd.sub_even hxy_even, + refine multiplicity.eq_coe_iff.mpr ⟨_, _⟩, + { rw [pow_one, ← even_iff_two_dvd], + exact (hx_odd.pow).add_odd hy_odd.pow }, + cases i with i, + { intro hxy', + have : 2 * 2 ∣ 2 * x, { convert dvd_add hxy hxy', ring_exp }, + have : 2 ∣ x := (mul_dvd_mul_iff_left (by norm_num)).mp this, + contradiction }, + suffices : ∀ (x : ℤ), odd x → x ^ (2 ^ (i + 1)) % 4 = 1, + { rw [show (2 ^ (1 + 1) : ℤ) = 4, by norm_num, int.dvd_iff_mod_eq_zero, int.add_mod, + this _ hx_odd, this _ hy_odd], + norm_num }, + intros x hx, + rw [pow_succ, mul_comm, pow_mul, int.sq_mod_four_eq_one_of_odd hx.pow] +end + +lemma int.two_pow_two_pow_sub_pow_two_pow {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x - y) (hx : ¬ 2 ∣ x) : + multiplicity 2 (x ^ (2 ^ n) - y ^ (2 ^ n)) = multiplicity 2 (x - y) + n := +by simp only [pow_two_pow_sub_pow_two_pow n, multiplicity.mul int.prime_two, + multiplicity.finset.prod (int.prime_two), add_comm, nat.cast_one, finset.sum_const, + finset.card_range, nsmul_one, int.two_pow_two_pow_add_two_pow_two_pow hx hxy] + +lemma int.two_pow_sub_pow' {x y : ℤ} (n : ℕ) (hxy : 4 ∣ x - y) (hx : ¬ 2 ∣ x) : + multiplicity 2 (x ^ n - y ^ n) = multiplicity 2 (x - y) + multiplicity (2 : ℤ) n := +begin + have hx_odd : odd x, { rwa [int.odd_iff_not_even, even_iff_two_dvd] }, + have hxy_even : even (x - y) := even_iff_two_dvd.mpr (dvd_trans (by norm_num) hxy), + have hy_odd : odd y := by simpa using hx_odd.sub_even hxy_even, + cases n, + { simp only [pow_zero, sub_self, multiplicity.zero, int.coe_nat_zero, part_enat.add_top] }, + have h : (multiplicity 2 n.succ).dom := multiplicity.finite_nat_iff.mpr ⟨by norm_num, n.succ_pos⟩, + rcases multiplicity.eq_coe_iff.mp (part_enat.coe_get h).symm with ⟨⟨k, hk⟩, hpn⟩, + rw [hk, pow_mul, pow_mul, multiplicity.pow_sub_pow_of_prime, + int.two_pow_two_pow_sub_pow_two_pow _ hxy hx, + ← hk, part_enat.coe_get], + { norm_cast }, + { exact int.prime_two }, + { simpa only [even_iff_two_dvd] using hx_odd.pow.sub_odd hy_odd.pow }, + { simpa only [even_iff_two_dvd, int.odd_iff_not_even] using hx_odd.pow }, + erw [int.coe_nat_dvd], -- `erw` to deal with `2 : ℤ` vs `(2 : ℕ) : ℤ` + contrapose! hpn, + rw pow_succ', + conv_rhs { rw hk }, + exact mul_dvd_mul_left _ hpn +end + +/-- **Lifting the exponent lemma** for `p = 2` -/ +lemma int.two_pow_sub_pow {x y : ℤ} {n : ℕ} (hxy : 2 ∣ x - y) (hx : ¬ 2 ∣ x) (hn : even n) : + multiplicity 2 (x ^ n - y ^ n) + 1 = multiplicity 2 (x + y) + multiplicity 2 (x - y) + + multiplicity (2 : ℤ) n := +begin + have hy : odd y, + { rw [← even_iff_two_dvd, ← int.odd_iff_not_even] at hx, + replace hxy := (@even_neg _ _ (x - y)).mpr (even_iff_two_dvd.mpr hxy), + convert even.add_odd hxy hx, + abel }, + cases hn with d hd, + subst hd, + simp only [← two_mul, pow_mul], + have hxy4 : 4 ∣ x ^ 2 - y ^ 2, + { rw [int.dvd_iff_mod_eq_zero, int.sub_mod, int.sq_mod_four_eq_one_of_odd _, + int.sq_mod_four_eq_one_of_odd hy], + { norm_num }, + { simp only [int.odd_iff_not_even, even_iff_two_dvd, hx, not_false_iff] }}, + rw [int.two_pow_sub_pow' d hxy4 _, sq_sub_sq, ← int.coe_nat_mul_out, + multiplicity.mul (int.prime_two), multiplicity.mul (int.prime_two)], + suffices : multiplicity (2 : ℤ) ↑(2 : ℕ) = 1, + { rw [this, add_comm (1 : part_enat), ← add_assoc] }, + { norm_cast, + rw multiplicity.multiplicity_self _ _, + { apply prime.not_unit, + simp only [← nat.prime_iff, nat.prime_two] }, + { exact two_ne_zero }}, + { rw [← even_iff_two_dvd, ← int.odd_iff_not_even], + apply odd.pow, + simp only [int.odd_iff_not_even, even_iff_two_dvd, hx, not_false_iff] } +end + +lemma nat.two_pow_sub_pow {x y : ℕ} (hxy : 2 ∣ x - y) (hx : ¬2 ∣ x) {n : ℕ} (hn : even n) : + multiplicity 2 (x ^ n - y ^ n) + 1 = multiplicity 2 (x + y) + multiplicity 2 (x - y) + + multiplicity 2 n := +begin + obtain hyx | hyx := le_total y x, + { iterate 3 { rw ←multiplicity.int.coe_nat_multiplicity }, + have hxyn : y ^ n ≤ x ^ n := pow_le_pow_of_le_left' hyx _, + simp only [int.coe_nat_sub hyx, int.coe_nat_sub (pow_le_pow_of_le_left' hyx _), int.coe_nat_add, + int.coe_nat_pow], + rw ←int.coe_nat_dvd at hx, + rw [←int.coe_nat_dvd, int.coe_nat_sub hyx] at hxy, + convert int.two_pow_sub_pow hxy hx hn using 2, + rw ← multiplicity.int.coe_nat_multiplicity, + refl }, + { simp only [nat.sub_eq_zero_iff_le.mpr hyx, + nat.sub_eq_zero_iff_le.mpr (pow_le_pow_of_le_left' hyx n), multiplicity.zero, + part_enat.top_add, part_enat.add_top] } +end + +namespace padic_val_nat + +variables {x y : ℕ} + +lemma pow_two_sub_pow (hyx : y < x) (hxy : 2 ∣ x - y) (hx : ¬ 2 ∣ x) {n : ℕ} (hn : 0 < n) + (hneven : even n) : + padic_val_nat 2 (x ^ n - y ^ n) + 1 = + padic_val_nat 2 (x + y) + padic_val_nat 2 (x - y) + padic_val_nat 2 n := +begin + simp only [←part_enat.coe_inj, nat.cast_add], + iterate 4 { rw [padic_val_nat_def, part_enat.coe_get] }, + { convert nat.two_pow_sub_pow hxy hx hneven using 2 }, + { exact hn }, + { exact (nat.sub_pos_of_lt hyx) }, + { linarith }, + { simp only [tsub_pos_iff_lt, pow_lt_pow_of_lt_left hyx (@zero_le' _ y _) hn] } +end + +variables {p : ℕ} [hp : fact p.prime] (hp1 : odd p) +include hp hp1 + +lemma pow_sub_pow (hyx : y < x) (hxy : p ∣ x - y) (hx : ¬p ∣ x) {n : ℕ} (hn : 0 < n) : + padic_val_nat p (x ^ n - y ^ n) = padic_val_nat p (x - y) + padic_val_nat p n := +begin + rw [←part_enat.coe_inj, nat.cast_add], + iterate 3 { rw [padic_val_nat_def, part_enat.coe_get] }, + { exact multiplicity.nat.pow_sub_pow hp.out hp1 hxy hx n }, + { exact hn }, + { exact nat.sub_pos_of_lt hyx }, + { exact nat.sub_pos_of_lt (nat.pow_lt_pow_of_lt_left hyx hn) } +end + +lemma pow_add_pow (hxy : p ∣ x + y) (hx : ¬p ∣ x) {n : ℕ} (hn : odd n) : + padic_val_nat p (x ^ n + y ^ n) = padic_val_nat p (x + y) + padic_val_nat p n := +begin + cases y, + { have := dvd_zero p, contradiction }, + rw [←part_enat.coe_inj, nat.cast_add], + iterate 3 { rw [padic_val_nat_def, part_enat.coe_get] }, + { exact multiplicity.nat.pow_add_pow hp.out hp1 hxy hx hn }, + { exact (odd.pos hn) }, + { simp only [add_pos_iff, nat.succ_pos', or_true] }, + { exact (nat.lt_add_left _ _ _ (pow_pos y.succ_pos _)) } +end + +end padic_val_nat From a630444a2870b61fc4c98ef6c720ca91cf69e0ff Mon Sep 17 00:00:00 2001 From: Vincent Beffara Date: Thu, 29 Sep 2022 10:48:59 +0000 Subject: [PATCH 447/828] feat(analysis/analytic/isolated_zeros): the uniqueness theorem for analytic fns (#16489) --- src/analysis/analytic/basic.lean | 30 ++++++++++ src/analysis/analytic/isolated_zeros.lean | 60 ++++++++++++++++++- .../calculus/formal_multilinear_series.lean | 20 +++++++ src/topology/continuous_on.lean | 8 +++ src/topology/separation.lean | 11 ++++ 5 files changed, 126 insertions(+), 3 deletions(-) diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index d141ef6a92920..ac10a5351b276 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -151,6 +151,11 @@ lemma radius_eq_top_of_forall_image_add_eq_zero (n : ℕ) (hn : ∀ m, p (m + n) p.radius_eq_top_of_eventually_eq_zero $ mem_at_top_sets.2 ⟨n, λ k hk, tsub_add_cancel_of_le hk ▸ hn _⟩ +@[simp] lemma const_formal_multilinear_series_radius {v : F} : + (const_formal_multilinear_series 𝕜 E v).radius = ⊤ := +(const_formal_multilinear_series 𝕜 E v).radius_eq_top_of_forall_image_add_eq_zero 1 + (by simp [const_formal_multilinear_series]) + /-- For `r` strictly smaller than the radius of `p`, then `∥pₙ∥ rⁿ` tends to zero exponentially: for some `0 < a < 1`, `∥p n∥ rⁿ = o(aⁿ)`. -/ lemma is_o_of_lt_radius (h : ↑r < p.radius) : @@ -448,6 +453,23 @@ lemma has_fpower_series_at.eventually_eq_zero ∀ᶠ z in 𝓝 x, f z = 0 := let ⟨r, hr⟩ := hf in hr.eventually_eq_zero +lemma has_fpower_series_on_ball_const {c : F} {e : E} : + has_fpower_series_on_ball (λ _, c) (const_formal_multilinear_series 𝕜 E c) e ⊤ := +begin + refine ⟨by simp, with_top.zero_lt_top, λ y hy, has_sum_single 0 (λ n hn, _)⟩, + simp [const_formal_multilinear_series_apply hn] +end + +lemma has_fpower_series_at_const {c : F} {e : E} : + has_fpower_series_at (λ _, c) (const_formal_multilinear_series 𝕜 E c) e := +⟨⊤, has_fpower_series_on_ball_const⟩ + +lemma analytic_at_const {v : F} : analytic_at 𝕜 (λ _, v) x := +⟨const_formal_multilinear_series 𝕜 E v, has_fpower_series_at_const⟩ + +lemma analytic_on_const {v : F} {s : set E} : analytic_on 𝕜 (λ _, v) s := +λ z _, analytic_at_const + lemma has_fpower_series_on_ball.add (hf : has_fpower_series_on_ball f pf x r) (hg : has_fpower_series_on_ball g pg x r) : has_fpower_series_on_ball (f + g) (pf + pg) x r := @@ -494,6 +516,14 @@ lemma analytic_at.sub (hf : analytic_at 𝕜 f x) (hg : analytic_at 𝕜 g x) : analytic_at 𝕜 (f - g) x := by simpa only [sub_eq_add_neg] using hf.add hg.neg +lemma analytic_on.add {s : set E} (hf : analytic_on 𝕜 f s) (hg : analytic_on 𝕜 g s) : + analytic_on 𝕜 (f + g) s := +λ z hz, (hf z hz).add (hg z hz) + +lemma analytic_on.sub {s : set E} (hf : analytic_on 𝕜 f s) (hg : analytic_on 𝕜 g s) : + analytic_on 𝕜 (f - g) s := +λ z hz, (hf z hz).sub (hg z hz) + lemma has_fpower_series_on_ball.coeff_zero (hf : has_fpower_series_on_ball f pf x r) (v : fin 0 → E) : pf 0 v = f x := begin diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean index 92ff0d95fb527..d60bf934f46f4 100644 --- a/src/analysis/analytic/isolated_zeros.lean +++ b/src/analysis/analytic/isolated_zeros.lean @@ -14,19 +14,22 @@ import topology.algebra.infinite_sum # Principle of isolated zeros This file proves the fact that the zeros of a non-constant analytic function of one variable are -isolated. It also introduces a little bit of API in the `has_fpower_series_at` namespace that -is useful in this setup. +isolated. It also introduces a little bit of API in the `has_fpower_series_at` namespace that is +useful in this setup. ## Main results * `analytic_at.eventually_eq_zero_or_eventually_ne_zero` is the main statement that if a function is analytic at `z₀`, then either it is identically zero in a neighborhood of `z₀`, or it does not vanish in a punctured neighborhood of `z₀`. +* `analytic_on.eq_on_of_preconnected_of_frequently_eq` is the identity theorem for analytic + functions: if a function `f` is analytic on a connected set `U` and is zero on a set with an + accumulation point in `U` then `f` is identically `0` on `U`. -/ open_locale classical -open filter function nat formal_multilinear_series emetric +open filter function nat formal_multilinear_series emetric set open_locale topological_space big_operators variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] @@ -137,4 +140,55 @@ begin { exact or.inr (hp.locally_ne_zero h) } end +lemma frequently_zero_iff_eventually_zero {f : 𝕜 → E} {w : 𝕜} (hf : analytic_at 𝕜 f w) : + (∃ᶠ z in 𝓝[≠] w, f z = 0) ↔ (∀ᶠ z in 𝓝 w, f z = 0) := +⟨hf.eventually_eq_zero_or_eventually_ne_zero.resolve_right, + λ h, (h.filter_mono nhds_within_le_nhds).frequently⟩ + end analytic_at + +namespace analytic_on + +variables {U : set 𝕜} {w : 𝕜} + +theorem eq_on_of_preconnected_of_frequently_eq (hf : analytic_on 𝕜 f U) (hU : is_preconnected U) + (hw : w ∈ U) (hfw : ∃ᶠ z in 𝓝[≠] w, f z = 0) : + eq_on f 0 U := +begin + by_contra, + simp only [eq_on, not_forall] at h, + obtain ⟨x, hx1, hx2⟩ := h, + let u := {z | f =ᶠ[𝓝 z] 0}, + have hu : is_open u := is_open_set_of_eventually_nhds, + have hu' : (U ∩ u).nonempty := ⟨w, hw, (hf w hw).frequently_zero_iff_eventually_zero.mp hfw⟩, + let v := {z | ∀ᶠ w in 𝓝[≠] z, f w ≠ 0}, + have hv : is_open v := by apply is_open_set_of_eventually_nhds_within, + have hv' : (U ∩ v).nonempty, + from ⟨x, hx1, ((hf x hx1).continuous_at.eventually_ne hx2).filter_mono nhds_within_le_nhds⟩, + have huv : U ⊆ u ∪ v := λ z hz, (hf z hz).eventually_eq_zero_or_eventually_ne_zero, + have huv' : u ∩ v = ∅, + by { ext z, + simp only [mem_inter_iff, mem_empty_iff_false, iff_false, not_and], + exact λ h, (h.filter_mono nhds_within_le_nhds).frequently }, + simpa [huv'] using hU u v hu hv huv hu' hv' +end + +theorem eq_on_of_preconnected_of_mem_closure (hf : analytic_on 𝕜 f U) (hU : is_preconnected U) + (hw : w ∈ U) (hfw : w ∈ closure ({z | f z = 0} \ {w})) : + eq_on f 0 U := +hf.eq_on_of_preconnected_of_frequently_eq hU hw (mem_closure_ne_iff_frequently_within.mp hfw) + +theorem eq_on_of_preconnected_of_frequently_eq' (hf : analytic_on 𝕜 f U) (hg : analytic_on 𝕜 g U) + (hU : is_preconnected U) (hw : w ∈ U) (hfg : ∃ᶠ z in 𝓝[≠] w, f z = g z) : + eq_on f g U := +begin + have hfg' : ∃ᶠ z in 𝓝[≠] w, (f - g) z = 0 := hfg.mono (λ z h, by rw [pi.sub_apply, h, sub_self]), + simpa [sub_eq_zero] using λ z hz, (hf.sub hg).eq_on_of_preconnected_of_frequently_eq hU hw hfg' hz +end + +theorem eq_on_of_preconnected_of_mem_closure' (hf : analytic_on 𝕜 f U) (hg : analytic_on 𝕜 g U) + (hU : is_preconnected U) (hw : w ∈ U) (hfw : w ∈ closure ({z | f z = g z} \ {w})) : + eq_on f g U := +hf.eq_on_of_preconnected_of_frequently_eq' hg hU hw (mem_closure_ne_iff_frequently_within.mp hfw) + +end analytic_on diff --git a/src/analysis/calculus/formal_multilinear_series.lean b/src/analysis/calculus/formal_multilinear_series.lean index bcbd3eb378590..4cf449fdf27b6 100644 --- a/src/analysis/calculus/formal_multilinear_series.lean +++ b/src/analysis/calculus/formal_multilinear_series.lean @@ -289,3 +289,23 @@ by induction k with k ih generalizing p; refl <|> simpa [ih] end fslope end formal_multilinear_series + +section const + +/-- The formal multilinear series where all terms of positive degree are equal to zero, and the term +of degree zero is `c`. It is the power series expansion of the constant function equal to `c` +everywhere. -/ +def const_formal_multilinear_series (𝕜 : Type*) [nontrivially_normed_field 𝕜] + (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] [has_continuous_const_smul 𝕜 E] + [topological_add_group E] {F : Type*} [normed_add_comm_group F] [topological_add_group F] + [normed_space 𝕜 F] [has_continuous_const_smul 𝕜 F] (c : F) : formal_multilinear_series 𝕜 E F +| 0 := continuous_multilinear_map.curry0 _ _ c +| _ := 0 + +@[simp] lemma const_formal_multilinear_series_apply [nontrivially_normed_field 𝕜] + [normed_add_comm_group E] [normed_add_comm_group F] [normed_space 𝕜 E] [normed_space 𝕜 F] + {c : F} {n : ℕ} (hn : n ≠ 0) : + const_formal_multilinear_series 𝕜 E c n = 0 := +nat.cases_on n (λ hn, (hn rfl).elim) (λ _ _, rfl) hn + +end const diff --git a/src/topology/continuous_on.lean b/src/topology/continuous_on.lean index 492859e552e22..e4c16d4191367 100644 --- a/src/topology/continuous_on.lean +++ b/src/topology/continuous_on.lean @@ -44,6 +44,14 @@ lemma eventually_nhds_within_iff {a : α} {s : set α} {p : α → Prop} : (∀ᶠ x in 𝓝[s] a, p x) ↔ ∀ᶠ x in 𝓝 a, x ∈ s → p x := eventually_inf_principal +lemma frequently_nhds_within_iff {z : α} {p : α → Prop} : + (∃ᶠ x in 𝓝[≠] z, p x) ↔ (∃ᶠ x in 𝓝 z, p x ∧ x ≠ z) := +iff.not (by simp [eventually_nhds_within_iff, not_imp_not]) + +lemma mem_closure_ne_iff_frequently_within {z : α} {s : set α} : + z ∈ closure (s \ {z}) ↔ ∃ᶠ x in 𝓝[≠] z, x ∈ s := +by simp [mem_closure_iff_frequently, frequently_nhds_within_iff] + @[simp] lemma eventually_nhds_within_nhds_within {a : α} {s : set α} {p : α → Prop} : (∀ᶠ y in 𝓝[s] a, ∀ᶠ x in 𝓝[s] y, p x) ↔ ∀ᶠ x in 𝓝[s] a, p x := begin diff --git a/src/topology/separation.lean b/src/topology/separation.lean index ef53e9bc19d4a..9bb375982375d 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -333,6 +333,17 @@ begin exact mem_nhds_within_of_mem_nhds (is_open_ne.mem_nhds h) end +lemma is_open_set_of_eventually_nhds_within [t1_space α] {p : α → Prop} : + is_open {x | ∀ᶠ y in 𝓝[≠] x, p y} := +begin + refine is_open_iff_mem_nhds.mpr (λ a ha, _), + filter_upwards [eventually_nhds_nhds_within.mpr ha] with b hb, + by_cases a = b, + { subst h, exact hb }, + { rw (ne.symm h).nhds_within_compl_singleton at hb, + exact hb.filter_mono nhds_within_le_nhds } +end + protected lemma set.finite.is_closed [t1_space α] {s : set α} (hs : set.finite s) : is_closed s := begin From 37cfab4b55a0680c97eb6d9ef2b5b3e9d2745bce Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Thu, 29 Sep 2022 12:54:40 +0000 Subject: [PATCH 448/828] feat(topology/path_connected): add five lemmas (#16501) add five lemmas about `symm` and `trans` operations on paths --- src/topology/path_connected.lean | 46 +++++++++++++++++++++++++++++++- 1 file changed, 45 insertions(+), 1 deletion(-) diff --git a/src/topology/path_connected.lean b/src/topology/path_connected.lean index e0178b81e2665..7cad71b7f8c62 100644 --- a/src/topology/path_connected.lean +++ b/src/topology/path_connected.lean @@ -138,6 +138,32 @@ begin simp end +/-! #### Space of paths -/ + +open continuous_map + +instance : has_coe (path x y) C(I, X) := ⟨λ γ, γ.1⟩ + +/-- +The following instance defines the topology on the path space to be induced from the +compact-open topology on the space `C(I,X)` of continuous maps from `I` to `X`. +-/ +instance : topological_space (path x y) := +topological_space.induced (coe : _ → C(I, X)) continuous_map.compact_open + +lemma continuous_eval : continuous (λ p : path x y × I, p.1 p.2) := +continuous_eval'.comp $ continuous_induced_dom.prod_map continuous_id + +@[continuity] lemma _root_.continuous.path_eval {Y} [topological_space Y] + {f : Y → path x y} {g : Y → I} (hf : continuous f) (hg : continuous g) : + continuous (λ y, f y (g y)) := continuous.comp continuous_eval (hf.prod_mk hg) + +lemma continuous_uncurry_iff {Y} [topological_space Y] {g : Y → path x y} : + continuous ↿g ↔ continuous g := +iff.symm $ continuous_induced_rng.trans + ⟨λ h, continuous_uncurry_of_continuous ⟨_, h⟩, continuous_of_continuous_uncurry ↑g⟩ + + /-- A continuous map extending a path to `ℝ`, constant before `0` and after `1`. -/ def extend : ℝ → X := Icc_extend zero_le_one γ @@ -346,9 +372,13 @@ lemma symm_continuous_family {X ι : Type*} [topological_space X] [topological_s continuous ↿(λ t, (γ t).symm) := h.comp (continuous_id.prod_map continuous_symm) +@[continuity] +lemma continuous_symm : continuous (symm : path x y → path y x) := +continuous_uncurry_iff.mp $ symm_continuous_family _ (continuous_fst.path_eval continuous_snd) + @[continuity] lemma continuous_uncurry_extend_of_continuous_family {X ι : Type*} [topological_space X] - [topological_space ι] {a b : ι → X} (γ : Π (t : ι), path (a t) (b t)) (h : continuous ↿γ) : + [topological_space ι] {a b : ι → X} (γ : Π (t : ι), path (a t) (b t)) (h : continuous ↿γ) : continuous ↿(λ t, (γ t).extend) := h.comp (continuous_id.prod_map continuous_proj_Icc) @@ -372,6 +402,20 @@ begin simp [hst, mul_inv_cancel (@two_ne_zero ℝ _ _)] } end +@[continuity] +lemma _root_.continuous.path_trans {f : Y → path x y} {g : Y → path y z} : continuous f → + continuous g → continuous (λ t, (f t).trans (g t)) := +begin + intros hf hg, + apply continuous_uncurry_iff.mp, + exact trans_continuous_family _ (continuous_uncurry_iff.mpr hf) + _ (continuous_uncurry_iff.mpr hg), +end + +@[continuity] +lemma continuous_trans {x y z : X} : continuous (λ ρ : path x y × path y z, ρ.1.trans ρ.2) := + continuous_fst.path_trans continuous_snd + /-! #### Product of paths -/ section prod variables {a₁ a₂ a₃ : X} {b₁ b₂ b₃ : Y} From 2e09d79080f1d1538a2199ac77a6c3696878ea79 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 29 Sep 2022 12:54:41 +0000 Subject: [PATCH 449/828] feat(topology/algebra/algebra): `algebra_clm` does not need a normed space or field (#16690) This also clean up the variables in the `topology/algebra/algebra.lean` file, to ensure that the type variables come first (as this is a useful convention for when using `@` notation). --- src/analysis/normed_space/basic.lean | 19 --------------- src/analysis/normed_space/exponential.lean | 3 ++- src/analysis/normed_space/spectrum.lean | 2 +- src/topology/algebra/algebra.lean | 27 ++++++++++++++++------ 4 files changed, 23 insertions(+), 28 deletions(-) diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 462b1fb9034ba..3e5f384e385ef 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -455,25 +455,6 @@ begin rw [dist_eq_norm, dist_eq_norm, ← ring_hom.map_sub, norm_algebra_map'], end -/-- The inclusion of the base field in a normed algebra as a continuous linear map. -/ -@[simps] -def algebra_map_clm : 𝕜 →L[𝕜] 𝕜' := -{ to_fun := algebra_map 𝕜 𝕜', - map_add' := (algebra_map 𝕜 𝕜').map_add, - map_smul' := λ r x, by rw [algebra.id.smul_eq_mul, map_mul, ring_hom.id_apply, algebra.smul_def], - cont := - have lipschitz_with ∥(1 : 𝕜')∥₊ (algebra_map 𝕜 𝕜') := λ x y, begin - rw [edist_eq_coe_nnnorm_sub, edist_eq_coe_nnnorm_sub, ←map_sub, ←ennreal.coe_mul, - ennreal.coe_le_coe, mul_comm], - exact (nnnorm_algebra_map _ _).le, - end, this.continuous } - -lemma algebra_map_clm_coe : - (algebra_map_clm 𝕜 𝕜' : 𝕜 → 𝕜') = (algebra_map 𝕜 𝕜' : 𝕜 → 𝕜') := rfl - -lemma algebra_map_clm_to_linear_map : - (algebra_map_clm 𝕜 𝕜').to_linear_map = algebra.linear_map 𝕜 𝕜' := rfl - instance normed_algebra.id : normed_algebra 𝕜 𝕜 := { .. normed_field.to_normed_space, .. algebra.id 𝕜} diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 3135feba55d1b..81581cc4ea78e 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -8,6 +8,7 @@ import analysis.analytic.basic import analysis.complex.basic import data.nat.choose.cast import data.finset.noncomm_prod +import topology.algebra.algebra /-! # Exponential in a Banach algebra @@ -289,7 +290,7 @@ end complete_algebra lemma algebra_map_exp_comm_of_mem_ball [complete_space 𝕂] (x : 𝕂) (hx : x ∈ emetric.ball (0 : 𝕂) (exp_series 𝕂 𝕂).radius) : algebra_map 𝕂 𝔸 (exp 𝕂 x) = exp 𝕂 (algebra_map 𝕂 𝔸 x) := -map_exp_of_mem_ball _ (algebra_map_clm _ _).continuous _ hx +map_exp_of_mem_ball _ (continuous_algebra_map 𝕂 𝔸) _ hx end any_field_any_algebra diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 1529e7b246779..995ef43a4dc3b 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -71,7 +71,7 @@ not_not.mp $ λ hn, h.not_le $ le_supr₂ k hn variable [complete_space A] lemma is_open_resolvent_set (a : A) : is_open (ρ a) := -units.is_open.preimage ((algebra_map_clm 𝕜 A).continuous.sub continuous_const) +units.is_open.preimage ((continuous_algebra_map 𝕜 A).sub continuous_const) lemma is_closed (a : A) : is_closed (σ a) := (is_open_resolvent_set a).is_closed_compl diff --git a/src/topology/algebra/algebra.lean b/src/topology/algebra/algebra.lean index a339814d41144..f79c7f66f4cc1 100644 --- a/src/topology/algebra/algebra.lean +++ b/src/topology/algebra/algebra.lean @@ -28,11 +28,11 @@ open_locale classical universes u v w section topological_algebra -variables (R : Type*) [topological_space R] [comm_semiring R] -variables (A : Type u) [topological_space A] -variables [semiring A] +variables (R : Type*) (A : Type u) +variables [comm_semiring R] [semiring A] [algebra R A] +variables [topological_space R] [topological_space A] [topological_semiring A] -lemma continuous_algebra_map_iff_smul [algebra R A] [topological_semiring A] : +lemma continuous_algebra_map_iff_smul : continuous (algebra_map R A) ↔ continuous (λ p : R × A, p.1 • p.2) := begin refine ⟨λ h, _, λ h, _⟩, @@ -41,15 +41,28 @@ begin end @[continuity] -lemma continuous_algebra_map [algebra R A] [topological_semiring A] [has_continuous_smul R A] : +lemma continuous_algebra_map [has_continuous_smul R A] : continuous (algebra_map R A) := (continuous_algebra_map_iff_smul R A).2 continuous_smul -lemma has_continuous_smul_of_algebra_map [algebra R A] [topological_semiring A] - (h : continuous (algebra_map R A)) : +lemma has_continuous_smul_of_algebra_map (h : continuous (algebra_map R A)) : has_continuous_smul R A := ⟨(continuous_algebra_map_iff_smul R A).1 h⟩ +variables [has_continuous_smul R A] + +/-- The inclusion of the base ring in a topological algebra as a continuous linear map. -/ +@[simps] +def algebra_map_clm : R →L[R] A := +{ to_fun := algebra_map R A, + cont := continuous_algebra_map R A, + .. algebra.linear_map R A } + +lemma algebra_map_clm_coe : ⇑(algebra_map_clm R A) = algebra_map R A := rfl + +lemma algebra_map_clm_to_linear_map : + (algebra_map_clm R A).to_linear_map = algebra.linear_map R A := rfl + end topological_algebra section topological_algebra From 426042843d5e24816efacd77b899fc7355383b71 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 29 Sep 2022 14:41:49 +0000 Subject: [PATCH 450/828] refactor(*): `inv_one_class`, `neg_zero_class` instances replacing lemmas (#16699) Replace `inv_one` lemmas for `matrix`, `fractional_ideal`, `mv_power_series` and `power_series` with `inv_one_class` instances, and a `neg_zero` lemma for `pgame` with a `neg_zero_class` instance. --- .../matrix/nonsingular_inverse.lean | 6 ++- src/ring_theory/dedekind_domain/ideal.lean | 8 ++-- src/ring_theory/power_series/basic.lean | 9 +++-- src/set_theory/game/impartial.lean | 4 +- src/set_theory/game/pgame.lean | 37 ++++++++++--------- 5 files changed, 36 insertions(+), 28 deletions(-) diff --git a/src/linear_algebra/matrix/nonsingular_inverse.lean b/src/linear_algebra/matrix/nonsingular_inverse.lean index 48634f7d65856..5375985da2046 100644 --- a/src/linear_algebra/matrix/nonsingular_inverse.lean +++ b/src/linear_algebra/matrix/nonsingular_inverse.lean @@ -429,8 +429,10 @@ begin simp [hn] }, end -@[simp] lemma inv_one : (1 : matrix n n α)⁻¹ = 1 := -inv_eq_left_inv (by simp) +noncomputable instance : inv_one_class (matrix n n α) := +{ inv_one := inv_eq_left_inv (by simp), + ..matrix.has_one, + ..matrix.has_inv } lemma inv_smul (k : α) [invertible k] (h : is_unit A.det) : (k • A)⁻¹ = ⅟k • A⁻¹ := inv_eq_left_inv (by simp [h, smul_smul]) diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 78f6e62740ad2..2ebf8d471d1f6 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -187,8 +187,10 @@ begin ((generator (I : submodule R₁ K))⁻¹)) hI).symm end -@[simp] lemma inv_one : (1⁻¹ : fractional_ideal R₁⁰ K) = 1 := -fractional_ideal.div_one +noncomputable instance : inv_one_class (fractional_ideal R₁⁰ K) := +{ inv_one := fractional_ideal.div_one, + ..fractional_ideal.has_one, + ..fractional_ideal.has_inv K } end fractional_ideal @@ -431,7 +433,7 @@ lemma mul_inv_cancel_of_le_one [h : is_dedekind_domain A] begin -- Handle a few trivial cases. by_cases hI1 : I = ⊤, - { rw [hI1, coe_ideal_top, one_mul, fractional_ideal.inv_one] }, + { rw [hI1, coe_ideal_top, one_mul, inv_one] }, by_cases hNF : is_field A, { letI := hNF.to_field, rcases hI1 (I.eq_bot_or_top.resolve_left hI0) }, -- We'll show a contradiction with `exists_not_mem_one_of_ne_bot`: diff --git a/src/ring_theory/power_series/basic.lean b/src/ring_theory/power_series/basic.lean index 683aaf683c405..c0c865154df01 100644 --- a/src/ring_theory/power_series/basic.lean +++ b/src/ring_theory/power_series/basic.lean @@ -764,8 +764,10 @@ begin mv_power_series.inv_mul_cancel _ h.right] } end -@[simp] lemma inv_one : (1 : mv_power_series σ k)⁻¹ = 1 := -by { rw [mv_power_series.inv_eq_iff_mul_eq_one, mul_one], simp } +instance : inv_one_class (mv_power_series σ k) := +{ inv_one := by { rw [mv_power_series.inv_eq_iff_mul_eq_one, mul_one], simp }, + ..mv_power_series.has_one, + ..mv_power_series.has_inv } @[simp] lemma C_inv (r : k) : (C σ k r)⁻¹ = C σ k r⁻¹ := begin @@ -1603,8 +1605,7 @@ mv_power_series.inv_eq_iff_mul_eq_one h (φ * ψ)⁻¹ = ψ⁻¹ * φ⁻¹ := mv_power_series.mul_inv_rev _ _ -@[simp] lemma inv_one : (1 : power_series k)⁻¹ = 1 := -mv_power_series.inv_one +instance : inv_one_class (power_series k) := mv_power_series.inv_one_class @[simp] lemma C_inv (r : k) : (C k r)⁻¹ = C k r⁻¹ := mv_power_series.C_inv _ diff --git a/src/set_theory/game/impartial.lean b/src/set_theory/game/impartial.lean index 0b63c30ebc117..d650ca0fb92df 100644 --- a/src/set_theory/game/impartial.lean +++ b/src/set_theory/game/impartial.lean @@ -109,14 +109,14 @@ variables (G : pgame) [impartial G] lemma nonpos : ¬ 0 < G := λ h, begin have h' := neg_lt_neg_iff.2 h, - rw [pgame.neg_zero, lt_congr_left (neg_equiv_self G).symm] at h', + rw [neg_zero, lt_congr_left (neg_equiv_self G).symm] at h', exact (h.trans h').false end lemma nonneg : ¬ G < 0 := λ h, begin have h' := neg_lt_neg_iff.2 h, - rw [pgame.neg_zero, lt_congr_right (neg_equiv_self G).symm] at h', + rw [neg_zero, lt_congr_right (neg_equiv_self G).symm] at h', exact (h.trans h').false end diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 1cd894b03d91b..bd4e75ace9ff1 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -816,11 +816,14 @@ instance : has_involutive_neg pgame := end, ..pgame.has_neg } -@[simp] protected lemma neg_zero : -(0 : pgame) = 0 := -begin - dsimp [has_zero.zero, has_neg.neg, neg], - congr; funext i; cases i -end +instance : neg_zero_class pgame := +{ neg_zero := + begin + dsimp [has_zero.zero, has_neg.neg, neg], + congr; funext i; cases i + end, + ..pgame.has_zero, + ..pgame.has_neg } @[simp] lemma neg_of_lists (L R : list pgame) : -of_lists L R = of_lists (R.map (λ x, -x)) (L.map (λ x, -x)) := @@ -953,34 +956,34 @@ theorem lt_neg_iff {x y : pgame} : y < -x ↔ x < -y := by rw [←neg_neg x, neg_lt_neg_iff, neg_neg] @[simp] theorem neg_le_zero_iff {x : pgame} : -x ≤ 0 ↔ 0 ≤ x := -by rw [neg_le_iff, pgame.neg_zero] +by rw [neg_le_iff, neg_zero] @[simp] theorem zero_le_neg_iff {x : pgame} : 0 ≤ -x ↔ x ≤ 0 := -by rw [le_neg_iff, pgame.neg_zero] +by rw [le_neg_iff, neg_zero] @[simp] theorem neg_lf_zero_iff {x : pgame} : -x ⧏ 0 ↔ 0 ⧏ x := -by rw [neg_lf_iff, pgame.neg_zero] +by rw [neg_lf_iff, neg_zero] @[simp] theorem zero_lf_neg_iff {x : pgame} : 0 ⧏ -x ↔ x ⧏ 0 := -by rw [lf_neg_iff, pgame.neg_zero] +by rw [lf_neg_iff, neg_zero] @[simp] theorem neg_lt_zero_iff {x : pgame} : -x < 0 ↔ 0 < x := -by rw [neg_lt_iff, pgame.neg_zero] +by rw [neg_lt_iff, neg_zero] @[simp] theorem zero_lt_neg_iff {x : pgame} : 0 < -x ↔ x < 0 := -by rw [lt_neg_iff, pgame.neg_zero] +by rw [lt_neg_iff, neg_zero] @[simp] theorem neg_equiv_zero_iff {x : pgame} : -x ≈ 0 ↔ x ≈ 0 := -by rw [neg_equiv_iff, pgame.neg_zero] +by rw [neg_equiv_iff, neg_zero] @[simp] theorem neg_fuzzy_zero_iff {x : pgame} : -x ∥ 0 ↔ x ∥ 0 := -by rw [neg_fuzzy_iff, pgame.neg_zero] +by rw [neg_fuzzy_iff, neg_zero] @[simp] theorem zero_equiv_neg_iff {x : pgame} : 0 ≈ -x ↔ 0 ≈ x := -by rw [←neg_equiv_iff, pgame.neg_zero] +by rw [←neg_equiv_iff, neg_zero] @[simp] theorem zero_fuzzy_neg_iff {x : pgame} : 0 ∥ -x ↔ 0 ∥ x := -by rw [←neg_fuzzy_iff, pgame.neg_zero] +by rw [←neg_fuzzy_iff, neg_zero] /-! ### Addition and subtraction -/ @@ -1147,7 +1150,7 @@ using_well_founded { dec_tac := pgame_wf_tac } instance : has_sub pgame := ⟨λ x y, x + -y⟩ @[simp] theorem sub_zero (x : pgame) : x - 0 = x + 0 := -show x + -0 = x + 0, by rw pgame.neg_zero +show x + -0 = x + 0, by rw neg_zero /-- If `w` has the same moves as `x` and `y` has the same moves as `z`, then `w - y` has the same moves as `x - z`. -/ @@ -1218,7 +1221,7 @@ end theorem zero_le_add_left_neg (x : pgame) : 0 ≤ -x + x := begin - rw [←neg_le_neg_iff, pgame.neg_zero], + rw [←neg_le_neg_iff, neg_zero], exact neg_add_le.trans (add_left_neg_le_zero _) end From d755c086c32fa54ff08a1f5e2fe767a43ed829fb Mon Sep 17 00:00:00 2001 From: sgouezel Date: Thu, 29 Sep 2022 16:43:02 +0000 Subject: [PATCH 451/828] feat(data/set/intervals/monotone): extend a monotone function on a set to a globally monotone function (#16682) --- src/data/set/intervals/monotone.lean | 155 +++++++++++++++++++++++++-- 1 file changed, 145 insertions(+), 10 deletions(-) diff --git a/src/data/set/intervals/monotone.lean b/src/data/set/intervals/monotone.lean index bc11fba75cc52..c5ab722f8da6e 100644 --- a/src/data/set/intervals/monotone.lean +++ b/src/data/set/intervals/monotone.lean @@ -11,9 +11,15 @@ import tactic.field_simp # Monotonicity on intervals In this file we prove that a function is (strictly) monotone (or antitone) on a linear order `α` -provided that it is (strictly) monotone on `(-∞, a]` and on `[a, +∞)`. We also provide an order -isomorphism `order_iso_Ioo_neg_one_one` between the open interval `(-1, 1)` in a linear ordered -field and the whole field. +provided that it is (strictly) monotone on `(-∞, a]` and on `[a, +∞)`. This is a special case +of a more general statement where one deduces monotonicity on a union from monotonicity on each +set. + +We deduce in `monotone_on.exists_monotone_extension` that a function which is monotone on a set +with a smallest and a largest element admits a monotone extension to the whole space. + +We also provide an order isomorphism `order_iso_Ioo_neg_one_one` between the open +interval `(-1, 1)` in a linear ordered field and the whole field. -/ open set @@ -22,34 +28,163 @@ section variables {α β : Type*} [linear_order α] [preorder β] {a : α} {f : α → β} +/-- If `f` is strictly monotone both on `s` and `t`, with `s` to the left of `t` and the center +point belonging to both `s` and `t`, then `f` is strictly monotone on `s ∪ t` -/ +protected lemma strict_mono_on.union {s t : set α} {c : α} (h₁ : strict_mono_on f s) + (h₂ : strict_mono_on f t) (hs : is_greatest s c) (ht : is_least t c) : + strict_mono_on f (s ∪ t) := +begin + have A : ∀ x, x ∈ s ∪ t → x ≤ c → x ∈ s, + { assume x hx hxc, + cases hx, { exact hx }, + rcases eq_or_lt_of_le hxc with rfl|h'x, { exact hs.1 }, + exact (lt_irrefl _ (h'x.trans_le (ht.2 hx))).elim }, + have B : ∀ x, x ∈ s ∪ t → c ≤ x → x ∈ t, + { assume x hx hxc, + cases hx, swap, { exact hx }, + rcases eq_or_lt_of_le hxc with rfl|h'x, { exact ht.1 }, + exact (lt_irrefl _ (h'x.trans_le (hs.2 hx))).elim }, + assume x hx y hy hxy, + rcases lt_or_le x c with hxc|hcx, + { have xs : x ∈ s, from A _ hx hxc.le, + rcases lt_or_le y c with hyc|hcy, + { exact h₁ xs (A _ hy hyc.le) hxy }, + { exact (h₁ xs hs.1 hxc).trans_le (h₂.monotone_on ht.1 (B _ hy hcy) hcy) } }, + { have xt : x ∈ t, from B _ hx hcx, + have yt : y ∈ t, from B _ hy (hcx.trans hxy.le), + exact h₂ xt yt hxy } +end + /-- If `f` is strictly monotone both on `(-∞, a]` and `[a, ∞)`, then it is strictly monotone on the whole line. -/ protected lemma strict_mono_on.Iic_union_Ici (h₁ : strict_mono_on f (Iic a)) (h₂ : strict_mono_on f (Ici a)) : strict_mono f := begin - intros x y hxy, - cases lt_or_le a x with hax hxa; [skip, cases le_or_lt y a with hya hay], - exacts [h₂ hax.le (hax.trans hxy).le hxy, h₁ hxa hya hxy, - (h₁.monotone_on hxa le_rfl hxa).trans_lt (h₂ le_rfl hay.le hay)] + rw [← strict_mono_on_univ, ← @Iic_union_Ici _ _ a], + exact strict_mono_on.union h₁ h₂ is_greatest_Iic is_least_Ici, end +/-- If `f` is strictly antitone both on `s` and `t`, with `s` to the left of `t` and the center +point belonging to both `s` and `t`, then `f` is strictly antitone on `s ∪ t` -/ +protected lemma strict_anti_on.union {s t : set α} {c : α} (h₁ : strict_anti_on f s) + (h₂ : strict_anti_on f t) (hs : is_greatest s c) (ht : is_least t c) : + strict_anti_on f (s ∪ t) := +(h₁.dual_right.union h₂.dual_right hs ht).dual_right + /-- If `f` is strictly antitone both on `(-∞, a]` and `[a, ∞)`, then it is strictly antitone on the whole line. -/ protected lemma strict_anti_on.Iic_union_Ici (h₁ : strict_anti_on f (Iic a)) (h₂ : strict_anti_on f (Ici a)) : strict_anti f := (h₁.dual_right.Iic_union_Ici h₂.dual_right).dual_right + +/-- If `f` is monotone both on `s` and `t`, with `s` to the left of `t` and the center +point belonging to both `s` and `t`, then `f` is monotone on `s ∪ t` -/ +protected lemma monotone_on.union_right {s t : set α} {c : α} (h₁ : monotone_on f s) + (h₂ : monotone_on f t) (hs : is_greatest s c) (ht : is_least t c) : + monotone_on f (s ∪ t) := +begin + have A : ∀ x, x ∈ s ∪ t → x ≤ c → x ∈ s, + { assume x hx hxc, + cases hx, { exact hx }, + rcases eq_or_lt_of_le hxc with rfl|h'x, { exact hs.1 }, + exact (lt_irrefl _ (h'x.trans_le (ht.2 hx))).elim }, + have B : ∀ x, x ∈ s ∪ t → c ≤ x → x ∈ t, + { assume x hx hxc, + cases hx, swap, { exact hx }, + rcases eq_or_lt_of_le hxc with rfl|h'x, { exact ht.1 }, + exact (lt_irrefl _ (h'x.trans_le (hs.2 hx))).elim }, + assume x hx y hy hxy, + rcases lt_or_le x c with hxc|hcx, + { have xs : x ∈ s, from A _ hx hxc.le, + rcases lt_or_le y c with hyc|hcy, + { exact h₁ xs (A _ hy hyc.le) hxy }, + { exact (h₁ xs hs.1 hxc.le).trans (h₂ ht.1 (B _ hy hcy) hcy) } }, + { have xt : x ∈ t, from B _ hx hcx, + have yt : y ∈ t, from B _ hy (hcx.trans hxy), + exact h₂ xt yt hxy } +end + +/-- If `f` is monotone both on `(-∞, a]` and `[a, ∞)`, then it is monotone on the whole line. -/ protected lemma monotone_on.Iic_union_Ici (h₁ : monotone_on f (Iic a)) (h₂ : monotone_on f (Ici a)) : monotone f := begin - intros x y hxy, - cases le_total x a with hxa hax; [cases le_total y a with hya hay, skip], - exacts [h₁ hxa hya hxy, (h₁ hxa le_rfl hxa).trans (h₂ le_rfl hay hay), h₂ hax (hax.trans hxy) hxy] + rw [← monotone_on_univ, ← @Iic_union_Ici _ _ a], + exact monotone_on.union_right h₁ h₂ is_greatest_Iic is_least_Ici end +/-- If `f` is antitone both on `s` and `t`, with `s` to the left of `t` and the center +point belonging to both `s` and `t`, then `f` is antitone on `s ∪ t` -/ +protected lemma antitone_on.union_right {s t : set α} {c : α} (h₁ : antitone_on f s) + (h₂ : antitone_on f t) (hs : is_greatest s c) (ht : is_least t c) : + antitone_on f (s ∪ t) := +(h₁.dual_right.union_right h₂.dual_right hs ht).dual_right + +/-- If `f` is antitone both on `(-∞, a]` and `[a, ∞)`, then it is antitone on the whole line. -/ protected lemma antitone_on.Iic_union_Ici (h₁ : antitone_on f (Iic a)) (h₂ : antitone_on f (Ici a)) : antitone f := (h₁.dual_right.Iic_union_Ici h₂.dual_right).dual_right +/-- If a function is monotone on a set `s`, then it admits a monotone extension to the whole space +provided `s` has a least element `a` and a greatest element `b`. -/ +lemma monotone_on.exists_monotone_extension {β : Type*} [conditionally_complete_linear_order β] + {f : α → β} {s : set α} (h : monotone_on f s) {a b : α} + (ha : is_least s a) (hb : is_greatest s b) : + ∃ g : α → β, monotone g ∧ eq_on f g s := +begin + /- The extension is defined by `f x = f a` for `x ≤ a`, and `f x` is the supremum of the values + of `f` to the left of `x` for `x ≥ a`. -/ + have aleb : a ≤ b := hb.2 ha.1, + have H : ∀ x ∈ s, f x = Sup (f '' (Icc a x ∩ s)), + { assume x xs, + have xmem : x ∈ Icc a x ∩ s := ⟨⟨ha.2 xs, le_rfl⟩, xs⟩, + have H : ∀ z, z ∈ f '' (Icc a x ∩ s) → z ≤ f x, + { rintros _ ⟨z, ⟨⟨az, zx⟩, zs⟩, rfl⟩, + exact h zs xs zx }, + apply le_antisymm, + { exact le_cSup ⟨f x, H⟩ (mem_image_of_mem _ xmem) }, + { exact cSup_le (nonempty_image_iff.2 ⟨x, xmem⟩) H } }, + let g := λ x, if x ≤ a then f a else Sup (f '' (Icc a x ∩ s)), + have hfg : eq_on f g s, + { assume x xs, + dsimp only [g], + by_cases hxa : x ≤ a, + { have : x = a, from le_antisymm hxa (ha.2 xs), + simp only [if_true, this, le_refl] }, + rw [if_neg hxa], + exact H x xs }, + have M1 : monotone_on g (Iic a), + { rintros x (hx : x ≤ a) y (hy : y ≤ a) hxy, + dsimp only [g], + simp only [hx, hy, if_true] }, + have g_eq : ∀ x ∈ Ici a, g x = Sup (f '' (Icc a x ∩ s)), + { rintros x ax, + dsimp only [g], + by_cases hxa : x ≤ a, + { have : x = a := le_antisymm hxa ax, + simp_rw [hxa, if_true, H a ha.1, this] }, + simp only [hxa, if_false], }, + have M2 : monotone_on g (Ici a), + { rintros x ax y ay hxy, + rw [g_eq x ax, g_eq y ay], + apply cSup_le_cSup, + { refine ⟨f b, _⟩, + rintros _ ⟨z, ⟨⟨az, zy⟩, zs⟩, rfl⟩, + exact h zs hb.1 (hb.2 zs) }, + { exact ⟨f a, mem_image_of_mem _ ⟨⟨le_rfl, ax⟩, ha.1⟩⟩ }, + { apply image_subset, + apply inter_subset_inter_left, + exact Icc_subset_Icc le_rfl hxy } }, + exact ⟨g, M1.Iic_union_Ici M2, hfg⟩, +end + +/-- If a function is antitone on a set `s`, then it admits an antitone extension to the whole space +provided `s` has a least element `a` and a greatest element `b`. -/ +lemma antitone_on.exists_antitone_extension {β : Type*} [conditionally_complete_linear_order β] + {f : α → β} {s : set α} (h : antitone_on f s) {a b : α} + (ha : is_least s a) (hb : is_greatest s b) : + ∃ g : α → β, antitone g ∧ eq_on f g s := +h.dual_right.exists_monotone_extension ha hb + end section ordered_group From c8760bde60264742b5c4e51e2d721f2290d735f4 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 29 Sep 2022 16:43:03 +0000 Subject: [PATCH 452/828] refactor(algebra/ring/basic): replace `neg_zero'` by `neg_zero_class` instance (#16686) Eliminate the separate `neg_zero'` lemma for the combination of `mul_zero_class` with `has_distrib_neg` by adding a `neg_zero_class` instance for that case. --- src/algebra/ring/basic.lean | 8 +++++--- src/analysis/special_functions/pow.lean | 4 ++-- src/data/polynomial/laurent.lean | 2 +- src/data/sign.lean | 2 +- src/linear_algebra/symplectic_group.lean | 2 +- src/probability/moments.lean | 4 ++-- 6 files changed, 12 insertions(+), 10 deletions(-) diff --git a/src/algebra/ring/basic.lean b/src/algebra/ring/basic.lean index cc544997fe7e1..956571775f59d 100644 --- a/src/algebra/ring/basic.lean +++ b/src/algebra/ring/basic.lean @@ -542,9 +542,11 @@ end mul_one_class section mul_zero_class variables [mul_zero_class α] [has_distrib_neg α] -/-- Prefer `neg_zero` if `subtraction_monoid` is available. -/ -@[simp] lemma neg_zero' : (-0 : α) = 0 := -by rw [←zero_mul (0 : α), ←neg_mul, mul_zero, mul_zero] +@[priority 100] +instance mul_zero_class.neg_zero_class : neg_zero_class α := +{ neg_zero := by rw [←zero_mul (0 : α), ←neg_mul, mul_zero, mul_zero], + ..mul_zero_class.to_has_zero α, + ..has_distrib_neg.to_has_involutive_neg α } end mul_zero_class diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index f357c8a642a4b..c87976de0f018 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -651,7 +651,7 @@ le_iff_le_iff_lt_iff_lt.2 $ rpow_lt_rpow_iff hy hx hz lemma le_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : x ≤ y ^ z⁻¹ ↔ y ≤ x ^ z := begin - have hz' : 0 < -z := by rwa [lt_neg, neg_zero'], + have hz' : 0 < -z := by rwa [lt_neg, neg_zero], have hxz : 0 < x ^ (-z) := real.rpow_pos_of_pos hx _, have hyz : 0 < y ^ z⁻¹ := real.rpow_pos_of_pos hy _, rw [←real.rpow_le_rpow_iff hx.le hyz.le hz', ←real.rpow_mul hy.le], @@ -663,7 +663,7 @@ end lemma lt_rpow_inv_iff_of_neg (hx : 0 < x) (hy : 0 < y) (hz : z < 0) : x < y ^ z⁻¹ ↔ y < x ^ z := begin - have hz' : 0 < -z := by rwa [lt_neg, neg_zero'], + have hz' : 0 < -z := by rwa [lt_neg, neg_zero], have hxz : 0 < x ^ (-z) := real.rpow_pos_of_pos hx _, have hyz : 0 < y ^ z⁻¹ := real.rpow_pos_of_pos hy _, rw [←real.rpow_lt_rpow_iff hx.le hyz.le hz', ←real.rpow_mul hy.le], diff --git a/src/data/polynomial/laurent.lean b/src/data/polynomial/laurent.lean index e72cf68ecf818..2d3617e6f9d0f 100644 --- a/src/data/polynomial/laurent.lean +++ b/src/data/polynomial/laurent.lean @@ -355,7 +355,7 @@ lemma reduce_to_polynomial_of_mul_T (f : R[T;T⁻¹]) {Q : R[T;T⁻¹] → Prop} begin induction f using laurent_polynomial.induction_on_mul_T with f n, induction n with n hn, - { simpa only [int.coe_nat_zero, neg_zero', T_zero, mul_one] using Qf _ }, + { simpa only [int.coe_nat_zero, neg_zero, T_zero, mul_one] using Qf _ }, { convert QT _ _, simpa using hn } end diff --git a/src/data/sign.lean b/src/data/sign.lean index 092565b849f62..236cfff561330 100644 --- a/src/data/sign.lean +++ b/src/data/sign.lean @@ -283,7 +283,7 @@ lemma sign_mul (x y : α) : sign (x * y) = sign x * sign y := begin rcases lt_trichotomy x 0 with hx | hx | hx; rcases lt_trichotomy y 0 with hy | hy | hy; simp only [sign_zero, mul_zero, zero_mul, sign_pos, sign_neg, hx, hy, mul_one, neg_one_mul, - neg_neg, one_mul, mul_pos_of_neg_of_neg, mul_neg_of_neg_of_pos, neg_zero', + neg_neg, one_mul, mul_pos_of_neg_of_neg, mul_neg_of_neg_of_pos, neg_zero, mul_neg_of_pos_of_neg, mul_pos] end diff --git a/src/linear_algebra/symplectic_group.lean b/src/linear_algebra/symplectic_group.lean index a9fab15586eb1..f223aa107a12a 100644 --- a/src/linear_algebra/symplectic_group.lean +++ b/src/linear_algebra/symplectic_group.lean @@ -47,7 +47,7 @@ variables [fintype l] lemma J_squared : (J l R) ⬝ (J l R) = -1 := begin rw [J, from_blocks_multiply], - simp only [matrix.zero_mul, matrix.neg_mul, zero_add, neg_zero', matrix.one_mul, add_zero], + simp only [matrix.zero_mul, matrix.neg_mul, zero_add, neg_zero, matrix.one_mul, add_zero], rw [← neg_zero, ← matrix.from_blocks_neg, ← from_blocks_one], end diff --git a/src/probability/moments.lean b/src/probability/moments.lean index bbbdc72197fdf..c1ef19a82846c 100644 --- a/src/probability/moments.lean +++ b/src/probability/moments.lean @@ -59,7 +59,7 @@ by simp only [moment, hp, zero_pow', ne.def, not_false_iff, pi.zero_apply, integ @[simp] lemma central_moment_zero (hp : p ≠ 0) : central_moment 0 p μ = 0 := by simp only [central_moment, hp, pi.zero_apply, integral_const, algebra.id.smul_eq_mul, - mul_zero, zero_sub, pi.pow_apply, pi.neg_apply, neg_zero', zero_pow', ne.def, not_false_iff] + mul_zero, zero_sub, pi.pow_apply, pi.neg_apply, neg_zero, zero_pow', ne.def, not_false_iff] lemma central_moment_one' [is_finite_measure μ] (h_int : integrable X μ) : central_moment X 1 μ = (1 - (μ set.univ).to_real) * μ[X] := @@ -307,7 +307,7 @@ lemma measure_ge_le_exp_mul_mgf [is_finite_measure μ] (ε : ℝ) (ht : 0 ≤ t) begin cases ht.eq_or_lt with ht_zero_eq ht_pos, { rw ht_zero_eq.symm, - simp only [neg_zero', zero_mul, exp_zero, mgf_zero', one_mul], + simp only [neg_zero, zero_mul, exp_zero, mgf_zero', one_mul], rw ennreal.to_real_le_to_real (measure_ne_top μ _) (measure_ne_top μ _), exact measure_mono (set.subset_univ _), }, calc (μ {ω | ε ≤ X ω}).to_real From 71e28e0cedb597249044f0a1491bc4af4e23304c Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 29 Sep 2022 19:48:11 +0000 Subject: [PATCH 453/828] =?UTF-8?q?feat(topology/uniform=5Fspace/uniform?= =?UTF-8?q?=5Fconvergence=5Ftopology):=20bases=20for=20uniform=20structure?= =?UTF-8?q?s=20of=20=F0=9D=94=96-convergence=20(#14778)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit By definition, the sets `S(V) := {(f, g) | ∀ x, (f x, g x) ∈ V}` for `V∈𝓤 β` form a basis for the uniformity of uniform convergence on `α → β`. We extend this result in the two following ways : - we show that it suffices to consider only the sets `V` in a basis of `𝓤 β` instead of all the entourages - we deduce a similar result for the uniformity of 𝔖-convergence for a directed 𝔖 : in that case, a basis is given by the sets `S'(A,V) := {(f, g) | ∀ x ∈ A, (f x, g x) ∈ V}` for `A ∈𝔖` and `V` in a basis of `𝓤 β` --- .../uniform_convergence_topology.lean | 130 +++++++++++++++++- 1 file changed, 124 insertions(+), 6 deletions(-) diff --git a/src/topology/uniform_space/uniform_convergence_topology.lean b/src/topology/uniform_space/uniform_convergence_topology.lean index 92fdd99b70693..9e887ae54d0cf 100644 --- a/src/topology/uniform_space/uniform_convergence_topology.lean +++ b/src/topology/uniform_space/uniform_convergence_topology.lean @@ -231,16 +231,32 @@ protected lemma has_basis_uniformity : (uniform_convergence.gen α β) := (uniform_convergence.is_basis_gen α β (𝓤 β)).has_basis +/-- The uniformity of `α → β` endowed with the uniform structure of uniform convergence on admits +the family `{(f, g) | ∀ x, (f x, g x) ∈ V}` for `V ∈ 𝓑` as a filter basis, for any basis +`𝓑` of `𝓤 β` (in the case `𝓑 = (𝓤 β).as_basis` this is true by definition). -/ +protected lemma has_basis_uniformity_of_basis {ι : Sort*} {p : ι → Prop} {s : ι → set (β × β)} + (h : (𝓤 β).has_basis p s) : + (𝓤 (α → β)).has_basis p (uniform_convergence.gen α β ∘ s) := +(uniform_convergence.has_basis_uniformity α β).to_has_basis + (λ U hU, let ⟨i, hi, hiU⟩ := h.mem_iff.mp hU in ⟨i, hi, λ uv huv x, hiU (huv x)⟩) + (λ i hi, ⟨s i, h.mem_of_mem hi, subset_refl _⟩) + /-- Topology of uniform convergence. -/ protected def topological_space : topological_space (α → β) := (𝒰(α, β, infer_instance)).to_topological_space +/-- If `α → β` is endowed with the topology of uniform convergence, `𝓝 f` admits the family +`{g | ∀ x, (f x, g x) ∈ V}` for `V ∈ 𝓑` as a filter basis, for any basis `𝓑` of `𝓤 β`. -/ +protected lemma has_basis_nhds_of_basis (f) {p : ι → Prop} {s : ι → set (β × β)} + (h : has_basis (𝓤 β) p s) : + (𝓝 f).has_basis p (λ i, {g | (f, g) ∈ uniform_convergence.gen α β (s i)}) := +nhds_basis_uniformity' (uniform_convergence.has_basis_uniformity_of_basis α β h) + /-- If `α → β` is endowed with the topology of uniform convergence, `𝓝 f` admits the family `{g | ∀ x, (f x, g x) ∈ V}` for `V ∈ 𝓤 β` as a filter basis. -/ -protected lemma has_basis_nhds : - (𝓝 f).has_basis (λ V, V ∈ 𝓤 β) - (λ V, {g | (f, g) ∈ uniform_convergence.gen α β V}) := -nhds_basis_uniformity' (uniform_convergence.has_basis_uniformity α β) +protected lemma has_basis_nhds (f) : + (𝓝 f).has_basis (λ V, V ∈ 𝓤 β) (λ V, {g | (f, g) ∈ uniform_convergence.gen α β V}) := +uniform_convergence.has_basis_nhds_of_basis α β f (filter.basis_sets _) variables {α} @@ -407,7 +423,7 @@ end protected lemma tendsto_iff_tendsto_uniformly : tendsto F p (𝓝 f) ↔ tendsto_uniformly F f p := begin letI : uniform_space (α → β) := 𝒰(α, β, _), - rw [(uniform_convergence.has_basis_nhds α β).tendsto_right_iff, tendsto_uniformly], + rw [(uniform_convergence.has_basis_nhds α β f).tendsto_right_iff, tendsto_uniformly], exact iff.rfl, end @@ -469,11 +485,49 @@ end uniform_convergence namespace uniform_convergence_on -variables (α β : Type*) {γ ι : Type*} [uniform_space β] (𝔖 : set (set α)) +variables {α β : Type*} {γ ι : Type*} variables {F : ι → α → β} {f : α → β} {s s' : set α} {x : α} {p : filter ι} {g : ι → α} local notation `𝒰(`α`, `β`, `u`)` := @uniform_convergence.uniform_space α β u +/-- Basis sets for the uniformity of `𝔖`-convergence: for `S : set α` and `V : set (β × β)`, +`gen S V` is the set of pairs `(f, g)` of functions `α → β` such that `∀ x ∈ S, (f x, g x) ∈ V`. -/ +protected def gen (S : set α) (V : set (β × β)) : set ((α → β) × (α → β)) := + {uv : (α → β) × (α → β) | ∀ x ∈ S, (uv.1 x, uv.2 x) ∈ V} + +/-- For `S : set α` and `V : set (β × β)`, we have +`uniform_convergence_on.gen S V = (S.restrict × S.restrict) ⁻¹' (uniform_convergence.gen S β V)`. +This is the crucial fact for proving that the family `uniform_convergence_on.gen S V` for +`S ∈ 𝔖` and `V ∈ 𝓤 β` is indeed a basis for the uniformity `α → β` endowed with `𝒱(α, β, 𝔖, uβ)` +the uniform structure of `𝔖`-convergence, as defined in `uniform_convergence_on.uniform_space`. -/ +protected lemma gen_eq_preimage_restrict (S : set α) (V : set (β × β)) : + uniform_convergence_on.gen S V = + (prod.map S.restrict S.restrict) ⁻¹' (uniform_convergence.gen S β V) := +begin + ext uv, + exact ⟨λ h ⟨x, hx⟩, h x hx, λ h x hx, h ⟨x, hx⟩⟩ +end + +/-- `uniform_convergence_on.gen` is antitone in the first argument and monotone in the second. -/ +protected lemma gen_mono {S S' : set α} {V V' : set (β × β)} (hS : S' ⊆ S) (hV : V ⊆ V') : + uniform_convergence_on.gen S V ⊆ uniform_convergence_on.gen S' V' := +λ uv h x hx, hV (h x $ hS hx) + +/-- If `𝔖 : set (set α)` is nonempty and directed and `𝓑` is a filter basis on `β × β`, then the +family `uniform_convergence_on.gen S V` for `S ∈ 𝔖` and `V ∈ 𝓑` is a filter basis. +We will show in `has_basis_uniformity_of_basis` that, if `𝓑` is a basis for `𝓤 β`, then the +corresponding filter is the uniformity of `(α → β, 𝒱(α, β, 𝔖, uβ))`. -/ +protected lemma is_basis_gen (𝔖 : set (set α)) (h : 𝔖.nonempty) (h' : directed_on (⊆) 𝔖) + (𝓑 : filter_basis $ β × β) : + is_basis (λ SV : set α × set (β × β), SV.1 ∈ 𝔖 ∧ SV.2 ∈ 𝓑) + (λ SV, uniform_convergence_on.gen SV.1 SV.2) := +⟨h.prod 𝓑.nonempty, λ U₁V₁ U₂V₂ h₁ h₂, + let ⟨U₃, hU₃, hU₁₃, hU₂₃⟩ := h' U₁V₁.1 h₁.1 U₂V₂.1 h₂.1 in + let ⟨V₃, hV₃, hV₁₂₃⟩ := 𝓑.inter_sets h₁.2 h₂.2 in ⟨⟨U₃, V₃⟩, ⟨⟨hU₃, hV₃⟩, λ uv huv, + ⟨(λ x hx, (hV₁₂₃ $ huv x $ hU₁₃ hx).1), (λ x hx, (hV₁₂₃ $ huv x $ hU₂₃ hx).2)⟩⟩⟩⟩ + +variables (α β) [uniform_space β] (𝔖 : set (set α)) + /-- Uniform structure of `𝔖`-convergence, i.e uniform convergence on the elements of `𝔖`. It is defined as the infimum, for `S ∈ 𝔖`, of the pullback of `𝒰 S β` by `S.restrict`, the map of restriction to `S`. We will denote it `𝒱(α, β, 𝔖, uβ)`, where `uβ` is the uniform structure @@ -499,6 +553,70 @@ begin refl end +protected lemma has_basis_uniformity_of_basis_aux₁ {p : ι → Prop} {s : ι → set (β × β)} + (hb : has_basis (𝓤 β) p s) (S : set α) : + (@uniformity (α → β) ((uniform_convergence.uniform_space S β).comap S.restrict)).has_basis + p (λ i, uniform_convergence_on.gen S (s i)) := +begin + simp_rw [uniform_convergence_on.gen_eq_preimage_restrict, uniformity_comap rfl], + exact (uniform_convergence.has_basis_uniformity_of_basis S β hb).comap _ +end + +protected lemma has_basis_uniformity_of_basis_aux₂ (h : directed_on (⊆) 𝔖) {p : ι → Prop} + {s : ι → set (β × β)} (hb : has_basis (𝓤 β) p s) : + directed_on ((λ s : set α, (uniform_convergence.uniform_space s β).comap + (s.restrict : (α → β) → s → β)) ⁻¹'o ge) 𝔖 := +h.mono $ λ s t hst, + ((uniform_convergence_on.has_basis_uniformity_of_basis_aux₁ α β hb _).le_basis_iff + (uniform_convergence_on.has_basis_uniformity_of_basis_aux₁ α β hb _)).mpr + (λ V hV, ⟨V, hV, uniform_convergence_on.gen_mono hst subset_rfl⟩) + +/-- If `𝔖 : set (set α)` is nonempty and directed and `𝓑` is a filter basis of `𝓤 β`, then the +uniformity of `(α → β, 𝒱(α, β, 𝔖, uβ))` admits the family `{(f, g) | ∀ x ∈ S, (f x, g x) ∈ V}` for +`S ∈ 𝔖` and `V ∈ 𝓑` as a filter basis. -/ +protected lemma has_basis_uniformity_of_basis (h : 𝔖.nonempty) (h' : directed_on (⊆) 𝔖) + {p : ι → Prop} {s : ι → set (β × β)} (hb : has_basis (𝓤 β) p s) : + (@uniformity (α → β) (uniform_convergence_on.uniform_space α β 𝔖)).has_basis + (λ Si : set α × ι, Si.1 ∈ 𝔖 ∧ p Si.2) + (λ Si, uniform_convergence_on.gen Si.1 (s Si.2)) := +begin + simp only [infi_uniformity'], + exact has_basis_binfi_of_directed h (λ S, (@uniform_convergence_on.gen α β S) ∘ s) _ + (λ S hS, uniform_convergence_on.has_basis_uniformity_of_basis_aux₁ α β hb S) + (uniform_convergence_on.has_basis_uniformity_of_basis_aux₂ α β 𝔖 h' hb) +end + +/-- If `𝔖 : set (set α)` is nonempty and directed, then the uniformity of +`(α → β, 𝒱(α, β, 𝔖, uβ))` admits the family `{(f, g) | ∀ x ∈ S, (f x, g x) ∈ V}` for `S ∈ 𝔖` and +`V ∈ 𝓤 β` as a filter basis. -/ +protected lemma has_basis_uniformity (h : 𝔖.nonempty) (h' : directed_on (⊆) 𝔖) : + (@uniformity (α → β) (uniform_convergence_on.uniform_space α β 𝔖)).has_basis + (λ SV : set α × set (β × β), SV.1 ∈ 𝔖 ∧ SV.2 ∈ 𝓤 β) + (λ SV, uniform_convergence_on.gen SV.1 SV.2) := +uniform_convergence_on.has_basis_uniformity_of_basis α β 𝔖 h h' (𝓤 β).basis_sets + +/-- If `α → β` is endowed with the topology of `𝔖`-convergence, where `𝔖 : set (set α)` is +nonempty and directed, then `𝓝 f` admits the family `{g | ∀ x ∈ S, (f x, g x) ∈ V}` for `S ∈ 𝔖` +and `V ∈ 𝓑` as a filter basis, for any basis `𝓑` of `𝓤 β`. -/ +protected lemma has_basis_nhds_of_basis (f) (h : 𝔖.nonempty) (h' : directed_on (⊆) 𝔖) + {p : ι → Prop} {s : ι → set (β × β)} (hb : has_basis (𝓤 β) p s) : + (@nhds (α → β) (uniform_convergence_on.topological_space α β 𝔖) f).has_basis + (λ Si : set α × ι, Si.1 ∈ 𝔖 ∧ p Si.2) + (λ Si, {g | (g, f) ∈ uniform_convergence_on.gen Si.1 (s Si.2)}) := +begin + letI : uniform_space (α → β) := uniform_convergence_on.uniform_space α β 𝔖, + exact nhds_basis_uniformity (uniform_convergence_on.has_basis_uniformity_of_basis α β 𝔖 h h' hb) +end + +/-- If `α → β` is endowed with the topology of `𝔖`-convergence, where `𝔖 : set (set α)` is +nonempty and directed, then `𝓝 f` admits the family `{g | ∀ x ∈ S, (f x, g x) ∈ V}` for `S ∈ 𝔖` +and `V ∈ 𝓤 β` as a filter basis. -/ +protected lemma has_basis_nhds (f) (h : 𝔖.nonempty) (h' : directed_on (⊆) 𝔖) : + (@nhds (α → β) (uniform_convergence_on.topological_space α β 𝔖) f).has_basis + (λ SV : set α × set (β × β), SV.1 ∈ 𝔖 ∧ SV.2 ∈ 𝓤 β) + (λ SV, {g | (g, f) ∈ uniform_convergence_on.gen SV.1 SV.2}) := +uniform_convergence_on.has_basis_nhds_of_basis α β 𝔖 f h h' (filter.basis_sets _) + /-- If `S ∈ 𝔖`, then the restriction to `S` is a uniformly continuous map from `𝒱(α, β, 𝔖, uβ)` to `𝒰(↥S, β, uβ)`. -/ protected lemma uniform_continuous_restrict (h : s ∈ 𝔖) : From ed33fcf8c852a6b8b18c20e81894f9c20c2ff0bb Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 29 Sep 2022 21:54:17 +0000 Subject: [PATCH 454/828] feat(order/filter/*): a family of pairwise disjoint filters has a family of pairwise disjoint members (#16504) --- src/order/filter/bases.lean | 25 +++++++++++++++++++++ src/order/filter/basic.lean | 25 +++++++++++++++++++++ src/topology/separation.lean | 42 ++++++++++-------------------------- 3 files changed, 61 insertions(+), 31 deletions(-) diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index b8042969d46eb..2599b0fa333f7 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -564,6 +564,31 @@ lemma has_basis.disjoint_iff (hl : l.has_basis p s) (hl' : l'.has_basis p' s') : not_iff_not.mp $ by simp only [disjoint_iff, ← ne.def, ← ne_bot_iff, hl.inf_basis_ne_bot_iff hl', not_exists, bot_eq_empty, ne_empty_iff_nonempty, inf_eq_inter] +lemma _root_.disjoint.exists_mem_filter_basis (h : disjoint l l') (hl : l.has_basis p s) + (hl' : l'.has_basis p' s') : + ∃ i (hi : p i) i' (hi' : p' i'), disjoint (s i) (s' i') := +(hl.disjoint_iff hl').1 h + +lemma _root_.pairwise.exists_mem_filter_basis_of_disjoint {I : Type*} [finite I] + {l : I → filter α} {ι : I → Sort*} {p : Π i, ι i → Prop} {s : Π i, ι i → set α} + (hd : pairwise (disjoint on l)) (h : ∀ i, (l i).has_basis (p i) (s i)) : + ∃ ind : Π i, ι i, (∀ i, p i (ind i)) ∧ pairwise (disjoint on λ i, s i (ind i)) := +begin + rcases hd.exists_mem_filter_of_disjoint with ⟨t, htl, hd⟩, + choose ind hp ht using λ i, (h i).mem_iff.1 (htl i), + exact ⟨ind, hp, hd.mono $ λ i j hij, hij.mono (ht _) (ht _)⟩ +end + +lemma _root_.set.pairwise_disjoint.exists_mem_filter_basis {I : Type*} {l : I → filter α} + {ι : I → Sort*} {p : Π i, ι i → Prop} {s : Π i, ι i → set α} {S : set I} + (hd : S.pairwise_disjoint l) (hS : S.finite) (h : ∀ i, (l i).has_basis (p i) (s i)) : + ∃ ind : Π i, ι i, (∀ i, p i (ind i)) ∧ S.pairwise_disjoint (λ i, s i (ind i)) := +begin + rcases hd.exists_mem_filter hS with ⟨t, htl, hd⟩, + choose ind hp ht using λ i, (h i).mem_iff.1 (htl i), + exact ⟨ind, hp, hd.mono ht⟩ +end + lemma inf_ne_bot_iff : ne_bot (l ⊓ l') ↔ ∀ ⦃s : set α⦄ (hs : s ∈ l) ⦃s'⦄ (hs' : s' ∈ l'), (s ∩ s').nonempty := l.basis_sets.inf_ne_bot_iff diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index f475bd4fd9cb8..70763cb535b72 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -630,6 +630,31 @@ lemma inf_eq_bot_iff {f g : filter α} : f ⊓ g = ⊥ ↔ ∃ (U ∈ f) (V ∈ g), U ∩ V = ∅ := by simpa only [disjoint_iff] using filter.disjoint_iff +lemma _root_.pairwise.exists_mem_filter_of_disjoint {ι : Type*} [finite ι] + {l : ι → filter α} (hd : pairwise (disjoint on l)) : + ∃ s : ι → set α, (∀ i, s i ∈ l i) ∧ pairwise (disjoint on s) := +begin + simp only [pairwise, function.on_fun, filter.disjoint_iff, subtype.exists'] at hd, + choose! s t hst using hd, + refine ⟨λ i, ⋂ j, s i j ∩ t j i, λ i, _, λ i j hij, _⟩, + exacts [Inter_mem.2 (λ j, inter_mem (s i j).2 (t j i).2), + (hst i j hij).mono ((Inter_subset _ j).trans (inter_subset_left _ _)) + ((Inter_subset _ i).trans (inter_subset_right _ _))] +end + +lemma _root_.set.pairwise_disjoint.exists_mem_filter {ι : Type*} {l : ι → filter α} {t : set ι} + (hd : t.pairwise_disjoint l) (ht : t.finite) : + ∃ s : ι → set α, (∀ i, s i ∈ l i) ∧ t.pairwise_disjoint s := +begin + casesI ht, + obtain ⟨s, hd⟩ : ∃ s : Π i : t, {s : set α // s ∈ l i}, pairwise (disjoint on λ i, (s i : set α)), + { rcases (hd.subtype _ _).exists_mem_filter_of_disjoint with ⟨s, hsl, hsd⟩, + exact ⟨λ i, ⟨s i, hsl i⟩, hsd⟩ }, + -- TODO: Lean fails to find `can_lift` instance and fails to use an instance supplied by `letI` + rcases @subtype.exists_pi_extension ι (λ i, {s // s ∈ l i}) _ _ s with ⟨s, rfl⟩, + exact ⟨λ i, s i, λ i, (s i).2, pairwise.set_of_subtype _ _ hd⟩ +end + /-- There is exactly one filter on an empty type. --/ instance unique [is_empty α] : unique (filter α) := { default := ⊥, uniq := filter_eq_bot_of_is_empty } diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 9bb375982375d..7588d5d98802d 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -804,6 +804,17 @@ end @[simp] lemma disjoint_nhds_nhds [t2_space α] {x y : α} : disjoint (𝓝 x) (𝓝 y) ↔ x ≠ y := ⟨λ hd he, by simpa [he, nhds_ne_bot.ne] using hd, t2_space_iff_disjoint_nhds.mp ‹_› x y⟩ +lemma pairwise_disjoint_nhds [t2_space α] : pairwise (disjoint on (𝓝 : α → filter α)) := +λ x y, disjoint_nhds_nhds.2 + +protected lemma set.pairwise_disjoint_nhds [t2_space α] (s : set α) : s.pairwise_disjoint 𝓝 := +pairwise_disjoint_nhds.set_pairwise s + +/-- Points of a finite set can be separated by open sets from each other. -/ +lemma set.finite.t2_separation [t2_space α] {s : set α} (hs : s.finite) : + ∃ U : α → set α, (∀ x, x ∈ U x ∧ is_open (U x)) ∧ s.pairwise_disjoint U := +s.pairwise_disjoint_nhds.exists_mem_filter_basis hs nhds_basis_opens + lemma is_open_set_of_disjoint_nhds_nhds : is_open {p : α × α | disjoint (𝓝 p.1) (𝓝 p.2)} := begin @@ -814,37 +825,6 @@ begin λ ⟨x', y'⟩ ⟨hx', hy'⟩, disjoint_of_disjoint_of_mem hd (hU.2.mem_nhds hx') (hV.2.mem_nhds hy')⟩ end -/-- A finite set can be separated by open sets. -/ -lemma t2_separation_finset [t2_space α] (s : finset α) : - ∃ f : α → set α, set.pairwise_disjoint ↑s f ∧ ∀ x ∈ s, x ∈ f x ∧ is_open (f x) := -finset.induction_on s (by simp) begin - rintros t s ht ⟨f, hf, hf'⟩, - have hty : ∀ y : s, t ≠ y := by { rintros y rfl, exact ht y.2 }, - choose u v hu hv htu hxv huv using λ {x} (h : t ≠ x), t2_separation h, - refine ⟨λ x, if ht : t = x then ⋂ y : s, u (hty y) else f x ∩ v ht, _, _⟩, - { rintros x hx₁ y hy₁ hxy a ⟨hx, hy⟩, - rw [finset.mem_coe, finset.mem_insert, eq_comm] at hx₁ hy₁, - rcases eq_or_ne t x with rfl | hx₂; - rcases eq_or_ne t y with rfl | hy₂, - { exact hxy rfl }, - { simp_rw [dif_pos rfl, mem_Inter] at hx, - simp_rw [dif_neg hy₂] at hy, - exact huv hy₂ ⟨hx ⟨y, hy₁.resolve_left hy₂⟩, hy.2⟩ }, - { simp_rw [dif_neg hx₂] at hx, - simp_rw [dif_pos rfl, mem_Inter] at hy, - exact huv hx₂ ⟨hy ⟨x, hx₁.resolve_left hx₂⟩, hx.2⟩ }, - { simp_rw [dif_neg hx₂] at hx, - simp_rw [dif_neg hy₂] at hy, - exact hf (hx₁.resolve_left hx₂) (hy₁.resolve_left hy₂) hxy ⟨hx.1, hy.1⟩ } }, - { intros x hx, - split_ifs with ht, - { refine ⟨mem_Inter.2 (λ y, _), is_open_Inter (λ y, hu (hty y))⟩, - rw ←ht, - exact htu (hty y) }, - { have hx := hf' x ((finset.mem_insert.1 hx).resolve_left (ne.symm ht)), - exact ⟨⟨hx.1, hxv ht⟩, is_open.inter hx.2 (hv ht)⟩ } } -end - @[priority 100] -- see Note [lower instance priority] instance t2_space.t1_space [t2_space α] : t1_space α := t1_space_iff_disjoint_pure_nhds.mpr $ λ x y hne, (disjoint_nhds_nhds.2 hne).mono_left $ From 6fc2df65f8d551779ba6ec3a0630cecd44d69ba3 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 30 Sep 2022 00:01:06 +0000 Subject: [PATCH 455/828] feat(order/filter/{prod,pi}): add `filter.prod_le_prod`, `filter.pi_le_pi` etc (#16468) --- src/data/set/prod.lean | 6 ++++-- src/order/filter/pi.lean | 21 +++++++++++++++++++++ src/order/filter/prod.lean | 26 ++++++++++++++++++++++++++ src/topology/inseparable.lean | 34 +++++++++++++++++++++++++++------- 4 files changed, 78 insertions(+), 9 deletions(-) diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index 222851ed0ee69..06acc5f6fbb53 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -456,15 +456,17 @@ image_subset_iff.2 $ λ f hf, hf i hs lemma eval_image_univ_pi_subset : eval i '' pi univ t ⊆ t i := eval_image_pi_subset (mem_univ i) -lemma eval_image_pi (hs : i ∈ s) (ht : (s.pi t).nonempty) : eval i '' s.pi t = t i := +lemma subset_eval_image_pi (ht : (s.pi t).nonempty) (i : ι) : t i ⊆ eval i '' s.pi t := begin - refine (eval_image_pi_subset hs).antisymm _, classical, obtain ⟨f, hf⟩ := ht, refine λ y hy, ⟨update f i y, λ j hj, _, update_same _ _ _⟩, obtain rfl | hji := eq_or_ne j i; simp [*, hf _ hj] end +lemma eval_image_pi (hs : i ∈ s) (ht : (s.pi t).nonempty) : eval i '' s.pi t = t i := +(eval_image_pi_subset hs).antisymm (subset_eval_image_pi ht i) + @[simp] lemma eval_image_univ_pi (ht : (pi univ t).nonempty) : (λ f : Π i, α i, f i) '' pi univ t = t i := eval_image_pi (mem_univ i) ht diff --git a/src/order/filter/pi.lean b/src/order/filter/pi.lean index f1c6e3a7ee21f..ee86f301ba4ef 100644 --- a/src/order/filter/pi.lean +++ b/src/order/filter/pi.lean @@ -143,6 +143,27 @@ by simpa using @pi_inf_principal_univ_pi_eq_bot ι α f (λ _, univ) instance [∀ i, ne_bot (f i)] : ne_bot (pi f) := pi_ne_bot.2 ‹_› +@[simp] lemma map_eval_pi (f : Π i, filter (α i)) [∀ i, ne_bot (f i)] (i : ι) : + map (eval i) (pi f) = f i := +begin + refine le_antisymm (tendsto_eval_pi f i) (λ s hs, _), + rcases mem_pi.1 (mem_map.1 hs) with ⟨I, hIf, t, htf, hI⟩, + rw [← image_subset_iff] at hI, + refine mem_of_superset (htf i) ((subset_eval_image_pi _ _).trans hI), + exact nonempty_of_mem (pi_mem_pi hIf (λ i hi, htf i)) +end + +@[simp] lemma pi_le_pi [∀ i, ne_bot (f₁ i)] : pi f₁ ≤ pi f₂ ↔ ∀ i, f₁ i ≤ f₂ i := +⟨λ h i, map_eval_pi f₁ i ▸ (tendsto_eval_pi _ _).mono_left h, pi_mono⟩ + +@[simp] lemma pi_inj [∀ i, ne_bot (f₁ i)] : pi f₁ = pi f₂ ↔ f₁ = f₂ := +begin + refine ⟨λ h, _, congr_arg pi⟩, + have hle : f₁ ≤ f₂ := pi_le_pi.1 h.le, + haveI : ∀ i, ne_bot (f₂ i) := λ i, ne_bot_of_le (hle i), + exact hle.antisymm (pi_le_pi.1 h.ge) +end + end pi /-! ### `n`-ary coproducts of filters -/ diff --git a/src/order/filter/prod.lean b/src/order/filter/prod.lean index c674f22486f75..2d8134c16d6b6 100644 --- a/src/order/filter/prod.lean +++ b/src/order/filter/prod.lean @@ -225,6 +225,32 @@ by simp only [filter.prod, comap_comap, (∘), inf_comm, prod.fst_swap, lemma prod_comm : f ×ᶠ g = map (λ p : β×α, (p.2, p.1)) (g ×ᶠ f) := by { rw [prod_comm', ← map_swap_eq_comap_swap], refl } +@[simp] lemma map_fst_prod (f : filter α) (g : filter β) [ne_bot g] : map prod.fst (f ×ᶠ g) = f := +begin + refine le_antisymm tendsto_fst (λ s hs, _), + rw [mem_map, mem_prod_iff] at hs, + rcases hs with ⟨t₁, h₁, t₂, h₂, hs⟩, + rw [← image_subset_iff, fst_image_prod] at hs, + exacts [mem_of_superset h₁ hs, nonempty_of_mem h₂] +end + +@[simp] lemma map_snd_prod (f : filter α) (g : filter β) [ne_bot f] : map prod.snd (f ×ᶠ g) = g := +by rw [prod_comm, map_map, (∘), map_fst_prod] + +@[simp] lemma prod_le_prod {f₁ f₂ : filter α} {g₁ g₂ : filter β} [ne_bot f₁] [ne_bot g₁] : + f₁ ×ᶠ g₁ ≤ f₂ ×ᶠ g₂ ↔ f₁ ≤ f₂ ∧ g₁ ≤ g₂ := +⟨λ h, ⟨map_fst_prod f₁ g₁ ▸ tendsto_fst.mono_left h, map_snd_prod f₁ g₁ ▸ tendsto_snd.mono_left h⟩, + λ h, prod_mono h.1 h.2⟩ + +@[simp] lemma prod_inj {f₁ f₂ : filter α} {g₁ g₂ : filter β} [ne_bot f₁] [ne_bot g₁] : + f₁ ×ᶠ g₁ = f₂ ×ᶠ g₂ ↔ f₁ = f₂ ∧ g₁ = g₂ := +begin + refine ⟨λ h, _, λ h, h.1 ▸ h.2 ▸ rfl⟩, + have hle : f₁ ≤ f₂ ∧ g₁ ≤ g₂ := prod_le_prod.1 h.le, + haveI := ne_bot_of_le hle.1, haveI := ne_bot_of_le hle.2, + exact ⟨hle.1.antisymm $ (prod_le_prod.1 h.ge).1, hle.2.antisymm $ (prod_le_prod.1 h.ge).2⟩ +end + lemma eventually_swap_iff {p : (α × β) → Prop} : (∀ᶠ (x : α × β) in (f ×ᶠ g), p x) ↔ ∀ᶠ (y : β × α) in (g ×ᶠ f), p y.swap := by { rw [prod_comm, eventually_map], simpa, } diff --git a/src/topology/inseparable.lean b/src/topology/inseparable.lean index 8a1b60fcfac36..77338de425f79 100644 --- a/src/topology/inseparable.lean +++ b/src/topology/inseparable.lean @@ -35,10 +35,10 @@ topological space, separation setoid -/ open set filter function -open_locale topological_space +open_locale topological_space filter -variables {X Y : Type*} [topological_space X] [topological_space Y] {x y z : X} - {s : set X} {f : X → Y} +variables {X Y Z α ι : Type*} {π : ι → Type*} [topological_space X] [topological_space Y] + [topological_space Z] [∀ i, topological_space (π i)] {x y z : X} {s : set X} {f : X → Y} /-! ### `specializes` relation @@ -151,13 +151,22 @@ by simp only [specializes_iff_mem_closure, hf.closure_eq_preimage_closure_image, lemma subtype_specializes_iff {p : X → Prop} (x y : subtype p) : x ⤳ y ↔ (x : X) ⤳ y := inducing_coe.specializes_iff.symm +@[simp] lemma specializes_prod {x₁ x₂ : X} {y₁ y₂ : Y} : + (x₁, y₁) ⤳ (x₂, y₂) ↔ x₁ ⤳ x₂ ∧ y₁ ⤳ y₂ := +by simp only [specializes, nhds_prod_eq, prod_le_prod] + +lemma specializes.prod {x₁ x₂ : X} {y₁ y₂ : Y} (hx : x₁ ⤳ x₂) (hy : y₁ ⤳ y₂) : + (x₁, y₁) ⤳ (x₂, y₂) := +specializes_prod.2 ⟨hx, hy⟩ + +@[simp] lemma specializes_pi {f g : Π i, π i} : f ⤳ g ↔ ∀ i, f i ⤳ g i := +by simp only [specializes, nhds_pi, pi_le_pi] + lemma not_specializes_iff_exists_open : ¬ x ⤳ y ↔ ∃ (S : set X), is_open S ∧ y ∈ S ∧ x ∉ S := -⟨λ h, by { contrapose! h, rwa specializes_iff_forall_open, }, - λ h, by { contrapose! h, rwa ←specializes_iff_forall_open, }⟩ +by { rw [specializes_iff_forall_open], push_neg, refl } lemma not_specializes_iff_exists_closed : ¬ x ⤳ y ↔ ∃ (S : set X), is_closed S ∧ x ∈ S ∧ y ∉ S := -⟨λ h, by { contrapose! h, rwa specializes_iff_forall_closed, }, - λ h, by { contrapose! h, rwa ←specializes_iff_forall_closed, }⟩ +by { rw [specializes_iff_forall_closed], push_neg, refl } variable (X) @@ -230,6 +239,17 @@ by simp only [inseparable_iff_specializes_and, hf.specializes_iff] lemma subtype_inseparable_iff {p : X → Prop} (x y : subtype p) : x ~ y ↔ (x : X) ~ y := inducing_coe.inseparable_iff.symm +@[simp] lemma inseparable_prod {x₁ x₂ : X} {y₁ y₂ : Y} : + (x₁, y₁) ~ (x₂, y₂) ↔ x₁ ~ x₂ ∧ y₁ ~ y₂ := +by simp only [inseparable, nhds_prod_eq, prod_inj] + +lemma inseparable.prod {x₁ x₂ : X} {y₁ y₂ : Y} (hx : x₁ ~ x₂) (hy : y₁ ~ y₂) : + (x₁, y₁) ~ (x₂, y₂) := +inseparable_prod.2 ⟨hx, hy⟩ + +@[simp] lemma inseparable_pi {f g : Π i, π i} : f ~ g ↔ ∀ i, f i ~ g i := +by simp only [inseparable, nhds_pi, funext_iff, pi_inj] + namespace inseparable @[refl] lemma refl (x : X) : x ~ x := eq.refl (𝓝 x) From 4ed50443f7db3a946135461873349a7869eeda6d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 30 Sep 2022 00:01:07 +0000 Subject: [PATCH 456/828] feat(tactic/positivity): Extension for `finset.card` (#16637) A best effort `positivity` extension for `finset.card`. This looks for an assumption of the form `s.nonempty` in context to prove `0 < s.card`. --- src/tactic/positivity.lean | 16 +++++++++++++++- test/positivity.lean | 3 +++ 2 files changed, 18 insertions(+), 1 deletion(-) diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index dd495add2ede2..2039db395f4f4 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -441,7 +441,7 @@ meta def positivity_coe : expr → tactic strictness | `(@coe_to_lift _ _ %%inst) := do strictness_a ← core a, match inst, strictness_a with -- `mk_mapp` is necessary in some places. Why? - | `(nat.cast_coe), positive p := positive <$> mk_app ``nat_cast_pos [p] + | `(nat.cast_coe), positive p := positive <$> mk_mapp ``nat_cast_pos [typ, none, none, none, p] | `(nat.cast_coe), _ := nonnegative <$> mk_mapp ``nat.cast_nonneg [typ, none, a] | `(int.cast_coe), positive p := positive <$> mk_mapp ``int_cast_pos [typ, none, none, none, p] | `(int.cast_coe), nonnegative p := nonnegative <$> @@ -457,4 +457,18 @@ meta def positivity_coe : expr → tactic strictness end | _ := failed +private lemma card_univ_pos (α : Type*) [fintype α] [nonempty α] : + 0 < (finset.univ : finset α).card := +finset.univ_nonempty.card_pos + +/-- Extension for the `positivity` tactic: `finset.card s` is positive if `s` is nonempty. -/ +@[positivity] +meta def positivity_finset_card : expr → tactic strictness +| `(finset.card %%s) := do -- TODO: Partial decision procedure for `finset.nonempty` + p ← to_expr ``(finset.nonempty %%s) >>= find_assumption, + positive <$> mk_app ``finset.nonempty.card_pos [p] +| `(@fintype.card %%α %%i) := positive <$> mk_mapp ``fintype.card_pos [α, i, none] +| e := pp e >>= fail ∘ format.bracket "The expression `" + "` isn't of the form `finset.card s` or `fintype.card α`" + end tactic diff --git a/test/positivity.lean b/test/positivity.lean index c2c7d29fa0cb0..1b41fa5b0343d 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -142,6 +142,9 @@ example : 0 ≤ max (-3 : ℤ) 5 := by positivity example [ordered_semiring α] [ordered_add_comm_monoid β] [smul_with_zero α β] [ordered_smul α β] {a : α} (ha : 0 < a) {b : β} (hb : 0 < b) : 0 ≤ a • b := by positivity +example {α : Type*} (s : finset α) (hs : s.nonempty) : 0 < s.card := by positivity +example {α : Type*} [fintype α] [nonempty α] : 0 < fintype.card α := by positivity + example {r : ℝ} : 0 < real.exp r := by positivity example {V : Type*} [normed_add_comm_group V] (x : V) : 0 ≤ ∥x∥ := by positivity From 07e46bcf438abab18a1146f896eff02064a4951e Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Fri, 30 Sep 2022 03:07:25 +0000 Subject: [PATCH 457/828] feat(data/fin): iff on add or sub across last-0 break (#15916) Co-authored-by: Yakov Pechersky --- src/data/fin/basic.lean | 65 +++++++++++++++++++++++++++++++++++++++++ 1 file changed, 65 insertions(+) diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index bd02a922c3d66..6b1d5f3728209 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -583,6 +583,43 @@ begin { rw [←succ_zero_eq_one, succ_lt_succ_iff], exact succ_pos a } end +@[simp] lemma add_one_lt_iff {n : ℕ} {k : fin (n + 2)} : + k + 1 < k ↔ k = last _ := +begin + simp only [lt_iff_coe_lt_coe, coe_add, coe_last, ext_iff], + cases k with k hk, + rcases (le_of_lt_succ hk).eq_or_lt with rfl|hk', + { simp }, + { simp [hk'.ne, mod_eq_of_lt (succ_lt_succ hk'), le_succ _] } +end + +@[simp] lemma add_one_le_iff {n : ℕ} {k : fin (n + 1)} : + k + 1 ≤ k ↔ k = last _ := +begin + cases n, + { simp [subsingleton.elim (k + 1) k, subsingleton.elim (fin.last _) k] }, + rw [←not_iff_not, ←add_one_lt_iff, lt_iff_le_and_ne, not_and'], + refine ⟨λ h _, h, λ h, h _⟩, + rw [ne.def, ext_iff, coe_add_one], + split_ifs with hk hk; + simp [hk, eq_comm], +end + +@[simp] lemma last_le_iff {n : ℕ} {k : fin (n + 1)} : + last n ≤ k ↔ k = last n := +top_le_iff + +@[simp] lemma lt_add_one_iff {n : ℕ} {k : fin (n + 1)} : + k < k + 1 ↔ k < last n := +begin + rw ←not_iff_not, + simp +end + +@[simp] lemma le_zero_iff {n : ℕ} {k : fin (n + 1)} : + k ≤ 0 ↔ k = 0 := +le_bot_iff + lemma succ_succ_ne_one (a : fin n) : fin.succ (fin.succ a) ≠ 1 := ne_of_gt (one_lt_succ_succ a) /-- `cast_lt i h` embeds `i` into a `fin` where `h` proves it belongs into. -/ @@ -1294,6 +1331,34 @@ begin nat.mod_eq_of_lt (tsub_lt_self (nat.succ_pos _) (tsub_pos_of_lt h)), h] } end +@[simp] lemma lt_sub_one_iff {n : ℕ} {k : fin (n + 2)} : + k < k - 1 ↔ k = 0 := +begin + rcases k with ⟨(_|k), hk⟩, + simp [lt_iff_coe_lt_coe], + have : (k + 1 + (n + 1)) % (n + 2) = k % (n + 2), + { rw [add_right_comm, add_assoc, add_mod_right] }, + simp [lt_iff_coe_lt_coe, ext_iff, fin.coe_sub, succ_eq_add_one, this, + mod_eq_of_lt ((lt_succ_self _).trans hk)] +end + +@[simp] lemma le_sub_one_iff {n : ℕ} {k : fin (n + 1)} : + k ≤ k - 1 ↔ k = 0 := +begin + cases n, + { simp [subsingleton.elim (k - 1) k, subsingleton.elim 0 k] }, + rw [←lt_sub_one_iff, le_iff_lt_or_eq, lt_sub_one_iff, or_iff_left_iff_imp, eq_comm, + sub_eq_iff_eq_add], + simp +end + +lemma sub_one_lt_iff {n : ℕ} {k : fin (n + 1)} : + k - 1 < k ↔ 0 < k := +begin + rw ←not_iff_not, + simp +end + /-- By sending `x` to `last n - x`, `fin n` is order-equivalent to its `order_dual`. -/ def _root_.order_iso.fin_equiv : ∀ {n}, (fin n)ᵒᵈ ≃o fin n | 0 := ⟨⟨elim0, elim0, elim0, elim0⟩, elim0⟩ From 667f2a62bef17369b83860bce0f71eac52dca21f Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 30 Sep 2022 06:01:26 +0000 Subject: [PATCH 458/828] feat(analysis/normed_space/basic): add `norm_algebra_map_nnreal` (#16709) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds `simp` lemmas saying that `∥algebra_map ℝ≥0 𝕜 x∥ = x` and similarly for `∥⬝∥₊` whenever `𝕜` is a normed `ℝ`-algebra and satisfies `norm_one_class`. These are needed separately from `norm_algebra_map'` and `nnnorm_algebra_map'` because `𝕜` cannot be a normed `ℝ≥0`-algebra for the simple reason that `ℝ≥0` is not a normed field. --- src/analysis/normed_space/basic.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 3e5f384e385ef..d9a8445abf35f 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -446,6 +446,18 @@ by rw [norm_algebra_map, norm_one, mul_one] @[simp] lemma nnnorm_algebra_map' [norm_one_class 𝕜'] (x : 𝕜) : ∥algebra_map 𝕜 𝕜' x∥₊ = ∥x∥₊ := subtype.ext $ norm_algebra_map' _ _ +section nnreal + +variables [norm_one_class 𝕜'] [normed_algebra ℝ 𝕜'] + +@[simp] lemma norm_algebra_map_nnreal (x : ℝ≥0) : ∥algebra_map ℝ≥0 𝕜' x∥ = x := +(norm_algebra_map' 𝕜' (x : ℝ)).symm ▸ real.norm_of_nonneg x.prop + +@[simp] lemma nnnorm_algebra_map_nnreal (x : ℝ≥0) : ∥algebra_map ℝ≥0 𝕜' x∥₊ = x := +subtype.ext $ norm_algebra_map_nnreal 𝕜' x + +end nnreal + variables (𝕜 𝕜') /-- In a normed algebra, the inclusion of the base field in the extended field is an isometry. -/ From 9b1e92043c32b8aafa5495d6a68abb7e04f46041 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 30 Sep 2022 06:01:27 +0000 Subject: [PATCH 459/828] feat(topology/continuous_function/{basic, compact}): add a few missing lemmas (#16714) --- src/topology/continuous_function/basic.lean | 3 +++ src/topology/continuous_function/compact.lean | 7 +++++++ 2 files changed, 10 insertions(+) diff --git a/src/topology/continuous_function/basic.lean b/src/topology/continuous_function/basic.lean index 09d7365e936a1..f411c809ca579 100644 --- a/src/topology/continuous_function/basic.lean +++ b/src/topology/continuous_function/basic.lean @@ -80,6 +80,9 @@ instance : has_coe_to_fun (C(α, β)) (λ _, α → β) := fun_like.has_coe_to_f -- this must come after the coe_to_fun definition initialize_simps_projections continuous_map (to_fun → apply) +@[protected, simp, norm_cast] +lemma coe_coe {F : Type*} [continuous_map_class F α β] (f : F) : ⇑(f : C(α, β)) = f := rfl + @[ext] lemma ext {f g : C(α, β)} (h : ∀ a, f a = g a) : f = g := fun_like.ext _ _ h /-- Copy of a `continuous_map` with a new `to_fun` equal to the old one. Useful to fix definitional diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index 85f7602caab89..6d94f0acc0e9b 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -185,10 +185,17 @@ lemma norm_le_of_nonempty [nonempty α] {M : ℝ} : ∥f∥ ≤ M ↔ ∀ x, ∥ lemma norm_lt_iff {M : ℝ} (M0 : 0 < M) : ∥f∥ < M ↔ ∀ x, ∥f x∥ < M := @bounded_continuous_function.norm_lt_iff_of_compact _ _ _ _ _ (mk_of_compact f) _ M0 +theorem nnnorm_lt_iff {M : ℝ≥0} (M0 : 0 < M) : ∥f∥₊ < M ↔ ∀ (x : α), ∥f x∥₊ < M := +f.norm_lt_iff M0 + lemma norm_lt_iff_of_nonempty [nonempty α] {M : ℝ} : ∥f∥ < M ↔ ∀ x, ∥f x∥ < M := @bounded_continuous_function.norm_lt_iff_of_nonempty_compact _ _ _ _ _ _ (mk_of_compact f) _ +lemma nnnorm_lt_iff_of_nonempty [nonempty α] {M : ℝ≥0} : + ∥f∥₊ < M ↔ ∀ x, ∥f x∥₊ < M := +f.norm_lt_iff_of_nonempty + lemma apply_le_norm (f : C(α, ℝ)) (x : α) : f x ≤ ∥f∥ := le_trans (le_abs.mpr (or.inl (le_refl (f x)))) (f.norm_coe_le_norm x) From 3a0be826ae674464b20e7437016b967e19de04bb Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 30 Sep 2022 08:02:38 +0000 Subject: [PATCH 460/828] feat(topology/sheaves): presheaf on indiscrete space is sheaf iff value at empty is terminal (#16694) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit + Show that the indiscrete topology (⊤ : topological_space α), defined to be the topology generated by nothing, consists of exactly the empty set and the whole space. + Show that a presheaf on an indiscrete space (in particular the one point space) is a sheaf if its value at the empty set is a terminal object. + Generalize universe level in the converse `is_terminal_of_empty` (which holds for any space). Co-authored-by: Jujian Zhang --- src/topology/order.lean | 13 +++++ src/topology/sets/opens.lean | 4 ++ src/topology/sheaves/punit.lean | 50 +++++++++++++++++++ .../sheaves/sheaf_condition/sites.lean | 2 +- 4 files changed, 68 insertions(+), 1 deletion(-) create mode 100644 src/topology/sheaves/punit.lean diff --git a/src/topology/order.lean b/src/topology/order.lean index d6749a2c08290..5ab285f5c4d75 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -256,6 +256,19 @@ lemma is_open_implies_is_open_iff {a b : topological_space α} : (∀ s, a.is_open s → b.is_open s) ↔ b ≤ a := iff.rfl +/-- The only open sets in the indiscrete topology are the empty set and the whole space. -/ +lemma topological_space.is_open_top_iff {α} (U : set α) : + (⊤ : topological_space α).is_open U ↔ U = ∅ ∨ U = univ := +⟨λ h, begin + induction h with V h _ _ _ _ ih₁ ih₂ _ _ ih, + { cases h }, { exact or.inr rfl }, + { obtain ⟨rfl|rfl, rfl|rfl⟩ := ⟨ih₁, ih₂⟩; simp }, + { rw [sUnion_eq_empty, or_iff_not_imp_left], + intro h, push_neg at h, obtain ⟨U, hU, hne⟩ := h, + have := (ih U hU).resolve_left hne, subst this, + refine sUnion_eq_univ_iff.2 (λ a, ⟨_, hU, trivial⟩) }, +end, by { rintro (rfl|rfl), exacts [@is_open_empty _ ⊤, @is_open_univ _ ⊤] }⟩ + /-- A topological space is discrete if every set is open, that is, its topology equals the discrete topology `⊥`. -/ class discrete_topology (α : Type*) [t : topological_space α] : Prop := diff --git a/src/topology/sets/opens.lean b/src/topology/sets/opens.lean index 63e011e115239..527ff91f7fcf2 100644 --- a/src/topology/sets/opens.lean +++ b/src/topology/sets/opens.lean @@ -151,6 +151,10 @@ by rw [← subtype.coe_injective.eq_iff, opens.coe_bot, ← set.not_nonempty_iff lemma ne_bot_iff_nonempty (U : opens α) : U ≠ ⊥ ↔ set.nonempty (U : set α) := by rw [ne.def, ← opens.not_nonempty_iff_eq_bot, not_not] +/-- An open set in the indiscrete topology is either empty or the whole space. -/ +lemma eq_bot_or_top {α} [t : topological_space α] (h : t = ⊤) (U : opens α) : U = ⊥ ∨ U = ⊤ := +by { simp_rw ← ext_iff, unfreezingI { subst h }, exact (is_open_top_iff U.1).1 U.2 } + /-- A set of `opens α` is a basis if the set of corresponding sets is a topological basis. -/ def is_basis (B : set (opens α)) : Prop := is_topological_basis ((coe : _ → set α) '' B) diff --git a/src/topology/sheaves/punit.lean b/src/topology/sheaves/punit.lean new file mode 100644 index 0000000000000..c6962e31579cc --- /dev/null +++ b/src/topology/sheaves/punit.lean @@ -0,0 +1,50 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import topology.sheaves.sheaf_condition.sites + +/-! +# Presheaves on punit + +Presheaves on punit satisfy sheaf condition iff its value at empty set is a terminal object. +-/ + +namespace Top.presheaf + +universes u v w + +open category_theory category_theory.limits Top opposite + +variables {C : Type u} [category.{v} C] + +lemma is_sheaf_of_is_terminal_of_indiscrete {X : Top.{w}} (hind : X.str = ⊤) (F : presheaf C X) + (it : is_terminal $ F.obj $ op ⊥) : F.is_sheaf := +λ c U s hs, begin + obtain rfl | hne := eq_or_ne U ⊥, + { intros _ _, rw @exists_unique_iff_exists _ ⟨λ _ _, _⟩, + { refine ⟨it.from _, λ U hU hs, is_terminal.hom_ext _ _ _⟩, rwa le_bot_iff.1 hU.le }, + { apply it.hom_ext } }, + { convert presieve.is_sheaf_for_top_sieve _, rw ←sieve.id_mem_iff_eq_top, + have := (U.eq_bot_or_top hind).resolve_left hne, subst this, + obtain he | ⟨⟨x⟩⟩ := is_empty_or_nonempty X, + { exact (hne $ topological_space.opens.ext_iff.1 $ set.univ_eq_empty_iff.2 he).elim }, + obtain ⟨U, f, hf, hm⟩ := hs x trivial, + obtain rfl | rfl := U.eq_bot_or_top hind, + { cases hm }, { convert hf } }, +end + +lemma is_sheaf_iff_is_terminal_of_indiscrete {X : Top.{w}} (hind : X.str = ⊤) + (F : presheaf C X) : F.is_sheaf ↔ nonempty (is_terminal $ F.obj $ op ⊥) := +⟨λ h, ⟨sheaf.is_terminal_of_empty ⟨F, h⟩⟩, λ ⟨it⟩, is_sheaf_of_is_terminal_of_indiscrete hind F it⟩ + +lemma is_sheaf_on_punit_of_is_terminal (F : presheaf C (Top.of punit)) + (it : is_terminal $ F.obj $ op ⊥) : F.is_sheaf := +is_sheaf_of_is_terminal_of_indiscrete (@subsingleton.elim (topological_space punit) _ _ _) F it + +lemma is_sheaf_on_punit_iff_is_terminal (F : presheaf C (Top.of punit)) : + F.is_sheaf ↔ nonempty (is_terminal $ F.obj $ op ⊥) := +⟨λ h, ⟨sheaf.is_terminal_of_empty ⟨F, h⟩⟩, λ ⟨it⟩, is_sheaf_on_punit_of_is_terminal F it⟩ + +end Top.presheaf diff --git a/src/topology/sheaves/sheaf_condition/sites.lean b/src/topology/sheaves/sheaf_condition/sites.lean index 7cb8f2bfb00d0..2214a7e4172e6 100644 --- a/src/topology/sheaves/sheaf_condition/sites.lean +++ b/src/topology/sheaves/sheaf_condition/sites.lean @@ -487,7 +487,7 @@ namespace Top.sheaf open category_theory topological_space Top opposite variables {C : Type u} [category.{v} C] -variables {X : Top.{v}} {ι : Type*} {B : ι → opens X} +variables {X : Top.{w}} {ι : Type*} {B : ι → opens X} variables (F : X.presheaf C) (F' : sheaf C X) (h : opens.is_basis (set.range B)) /-- The empty component of a sheaf is terminal -/ From acd02bbdc99d4acbbe9e2dbd9129e89e38ddeafc Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 30 Sep 2022 08:02:39 +0000 Subject: [PATCH 461/828] feat(algebra/hom/equiv): two little `simp` lemmas for `units.map_equiv` (#16701) Two lemmas I needed for working with the class group of rings. Co-authored-by: Anne Baanen --- src/algebra/hom/equiv.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/algebra/hom/equiv.lean b/src/algebra/hom/equiv.lean index 6c97f007cdda4..4664e0d76730d 100644 --- a/src/algebra/hom/equiv.lean +++ b/src/algebra/hom/equiv.lean @@ -515,6 +515,14 @@ def map_equiv (h : M ≃* N) : Mˣ ≃* Nˣ := right_inv := λ u, ext $ h.right_inv u, .. map h.to_monoid_hom } +@[simp] +lemma map_equiv_symm (h : M ≃* N) : (map_equiv h).symm = map_equiv h.symm := +rfl + +@[simp] +lemma coe_map_equiv (h : M ≃* N) (x : Mˣ) : (map_equiv h x : N) = h x := +rfl + /-- Left multiplication by a unit of a monoid is a permutation of the underlying type. -/ @[to_additive "Left addition of an additive unit is a permutation of the underlying type.", simps apply {fully_applied := ff}] From 001ba509ee4eb3d38b4c34500010c8761e9d5754 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 30 Sep 2022 08:02:40 +0000 Subject: [PATCH 462/828] feat(ring_theory/fractional_ideal): `simp` lemmas for `fractional_ideal.canonical_equiv` (#16702) Some lemmas I needed for working with the class group of a ring of integers. --- src/ring_theory/fractional_ideal.lean | 33 ++++++++++++++++++++++++++- 1 file changed, 32 insertions(+), 1 deletion(-) diff --git a/src/ring_theory/fractional_ideal.lean b/src/ring_theory/fractional_ideal.lean index a74b7ba832514..f112c4ea6179e 100644 --- a/src/ring_theory/fractional_ideal.lean +++ b/src/ring_theory/fractional_ideal.lean @@ -734,10 +734,41 @@ by { rw [mem_canonical_equiv_apply, canonical_equiv, map_equiv_symm, map_equiv, ring_equiv.coe_mk, mem_map], exact ⟨λ ⟨y, mem, eq⟩, ⟨y, mem, eq⟩, λ ⟨y, mem, eq⟩, ⟨y, mem, eq⟩⟩ } -@[simp] lemma canonical_equiv_flip (I) : +lemma canonical_equiv_flip (I) : canonical_equiv S P P' (canonical_equiv S P' P I) = I := by rw [←canonical_equiv_symm, ring_equiv.symm_apply_apply] +@[simp] +lemma canonical_equiv_canonical_equiv (P'' : Type*) [comm_ring P''] [algebra R P''] + [is_localization S P''] (I : fractional_ideal S P) : + canonical_equiv S P' P'' (canonical_equiv S P P' I) = canonical_equiv S P P'' I := +begin + ext, + simp only [is_localization.map_map, ring_hom_inv_pair.comp_eq₂, mem_canonical_equiv_apply, + exists_prop, exists_exists_and_eq_and], + refl +end + +lemma canonical_equiv_trans_canonical_equiv (P'' : Type*) [comm_ring P''] + [algebra R P''] [is_localization S P''] : + (canonical_equiv S P P').trans (canonical_equiv S P' P'') = canonical_equiv S P P'' := +ring_equiv.ext (canonical_equiv_canonical_equiv S P P' P'') + +@[simp] +lemma canonical_equiv_coe_ideal (I : ideal R) : + canonical_equiv S P P' I = I := +by { ext, simp [is_localization.map_eq] } + +omit loc' + +@[simp] +lemma canonical_equiv_self : canonical_equiv S P P = ring_equiv.refl _ := +begin + rw ← canonical_equiv_trans_canonical_equiv S P P, + convert (canonical_equiv S P P).symm_trans_self, + exact (canonical_equiv_symm S P P).symm +end + end semiring section is_fraction_ring From d610cec7c4fd04973116a99e5906d2b7f5fd33b6 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 30 Sep 2022 08:02:41 +0000 Subject: [PATCH 463/828] feat(group_theory/quotient_group): `simp` lemmas for `quotient_group.map` (#16703) Little lemmas that I needed to work with the class group of a ring of integers. Co-authored-by: Anne Baanen --- src/group_theory/quotient_group.lean | 34 ++++++++++++++++++++++++++++ src/group_theory/subgroup/basic.lean | 4 ++++ 2 files changed, 38 insertions(+) diff --git a/src/group_theory/quotient_group.lean b/src/group_theory/quotient_group.lean index 32ff90e709fb5..b45d3a0feec7e 100644 --- a/src/group_theory/quotient_group.lean +++ b/src/group_theory/quotient_group.lean @@ -185,6 +185,40 @@ lift_mk' _ _ x map N M f h (mk' _ x) = ↑(f x) := quotient_group.lift_mk' _ _ x +@[to_additive] +lemma map_id_apply (h : N ≤ subgroup.comap (monoid_hom.id _) N := (subgroup.comap_id N).le) (x) : + map N N (monoid_hom.id _) h x = x := +begin + refine induction_on' x (λ x, _), + simp only [map_coe, monoid_hom.id_apply] +end + +@[simp, to_additive] +lemma map_id (h : N ≤ subgroup.comap (monoid_hom.id _) N := (subgroup.comap_id N).le) : + map N N (monoid_hom.id _) h = monoid_hom.id _ := +monoid_hom.ext (map_id_apply N h) + +@[simp, to_additive] +lemma map_map {I : Type*} [group I] (M : subgroup H) (O : subgroup I) + [M.normal] [O.normal] + (f : G →* H) (g : H →* I) (hf : N ≤ subgroup.comap f M) (hg : M ≤ subgroup.comap g O) + (hgf : N ≤ subgroup.comap (g.comp f) O := + hf.trans ((subgroup.comap_mono hg).trans_eq (subgroup.comap_comap _ _ _))) (x : G ⧸ N) : + map M O g hg (map N M f hf x) = map N O (g.comp f) hgf x := +begin + refine induction_on' x (λ x, _), + simp only [map_coe, monoid_hom.comp_apply] +end + +@[simp, to_additive] +lemma map_comp_map {I : Type*} [group I] (M : subgroup H) (O : subgroup I) + [M.normal] [O.normal] + (f : G →* H) (g : H →* I) (hf : N ≤ subgroup.comap f M) (hg : M ≤ subgroup.comap g O) + (hgf : N ≤ subgroup.comap (g.comp f) O := + hf.trans ((subgroup.comap_mono hg).trans_eq (subgroup.comap_comap _ _ _))) : + (map M O g hg).comp (map N M f hf) = map N O (g.comp f) hgf := +monoid_hom.ext (map_map N M O f g hf hg hgf) + omit nN variables (φ : G →* H) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 9272caefdd98c..721e922aa759a 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1134,6 +1134,10 @@ lemma comap_comap (K : subgroup P) (g : N →* P) (f : G →* N) : (K.comap g).comap f = K.comap (g.comp f) := rfl +@[simp, to_additive] lemma comap_id (K : subgroup N) : + K.comap (monoid_hom.id _) = K := +by { ext, refl } + /-- The image of a subgroup along a monoid homomorphism is a subgroup. -/ @[to_additive "The image of an `add_subgroup` along an `add_monoid` homomorphism is an `add_subgroup`."] From af90fef9f821b2701b1fb5788d8c9c712a6634e6 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 30 Sep 2022 08:02:42 +0000 Subject: [PATCH 464/828] feat(data/fin): golf, add lemmas (#16711) * add `fin.range_succ`, `fin.exists_succ_eq_iff`, and `fin.range_fin_succ`; * golf `fin.eq_succ_of_ne_zero` and `fin.range_cons`. --- src/data/fin/basic.lean | 14 ++++++++------ src/data/fin/tuple/basic.lean | 20 ++++++-------------- 2 files changed, 14 insertions(+), 20 deletions(-) diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index 6b1d5f3728209..b17c1d80e005a 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -271,11 +271,7 @@ begin end lemma eq_succ_of_ne_zero {n : ℕ} {i : fin (n + 1)} (hi : i ≠ 0) : ∃ j : fin n, i = j.succ := -begin - cases eq_zero_or_eq_succ i, - { exfalso, exact hi h, }, - { exact h, }, -end +(eq_zero_or_eq_succ i).resolve_left hi /-- The greatest value of `fin (n+1)` -/ def last (n : ℕ) : fin (n+1) := ⟨_, n.lt_succ_self⟩ @@ -1120,7 +1116,7 @@ lemma forall_fin_succ {P : fin (n+1) → Prop} : lemma exists_fin_succ {P : fin (n+1) → Prop} : (∃ i, P i) ↔ P 0 ∨ (∃i:fin n, P i.succ) := ⟨λ ⟨i, h⟩, fin.cases or.inl (λ i hi, or.inr ⟨i, hi⟩) i h, - λ h, or.elim h (λ h, ⟨0, h⟩) $ λ⟨i, hi⟩, ⟨i.succ, hi⟩⟩ + λ h, h.elim (λ h, ⟨0, h⟩) $ λ⟨i, hi⟩, ⟨i.succ, hi⟩⟩ lemma forall_fin_one {p : fin 1 → Prop} : (∀ i, p i) ↔ p 0 := @unique.forall_iff (fin 1) _ p lemma exists_fin_one {p : fin 1 → Prop} : (∃ i, p i) ↔ p 0 := @unique.exists_iff (fin 1) _ p @@ -1525,6 +1521,12 @@ end @[simp] lemma range_succ_above (p : fin (n + 1)) : set.range (p.succ_above) = {p}ᶜ := set.ext $ λ _, exists_succ_above_eq_iff +@[simp] lemma range_succ (n : ℕ) : set.range (fin.succ : fin n → fin (n + 1)) = {0}ᶜ := +range_succ_above 0 + +@[simp] lemma exists_succ_eq_iff {x : fin (n + 1)} : (∃ y, fin.succ y = x) ↔ x ≠ 0 := +@exists_succ_above_eq_iff n 0 x + /-- Given a fixed pivot `x : fin (n + 1)`, `x.succ_above` is injective -/ lemma succ_above_right_injective {x : fin (n + 1)} : injective (succ_above x) := (succ_above x).injective diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index a1777cb57645e..4384a1b1bb651 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -198,21 +198,13 @@ begin simp [and_assoc, exists_and_distrib_left], end -@[simp] -lemma range_cons {α : Type*} {n : ℕ} (x : α) (b : fin n → α) : +lemma range_fin_succ {α} (f : fin (n + 1) → α) : + set.range f = insert (f 0) (set.range (fin.tail f)) := +set.ext $ λ y, exists_fin_succ.trans $ eq_comm.or iff.rfl + +@[simp] lemma range_cons {α : Type*} {n : ℕ} (x : α) (b : fin n → α) : set.range (fin.cons x b : fin n.succ → α) = insert x (set.range b) := -begin - ext y, - simp only [set.mem_range, set.mem_insert_iff], - split, - { rintros ⟨i, rfl⟩, - refine cases (or.inl (cons_zero _ _)) (λ i, or.inr ⟨i, _⟩) i, - rw cons_succ }, - { rintros (rfl | ⟨i, hi⟩), - { exact ⟨0, fin.cons_zero _ _⟩ }, - { refine ⟨i.succ, _⟩, - rw [cons_succ, hi] } } -end +by rw [range_fin_succ, cons_zero, tail_cons] /-- `fin.append ho u v` appends two vectors of lengths `m` and `n` to produce one of length `o = m + n`. `ho` provides control of definitional equality From b2a657206ddbdcba8b6356eac3a032c22b19ebeb Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 30 Sep 2022 08:02:43 +0000 Subject: [PATCH 465/828] feat(topology/instances/nnreal): generalize `has_continuous_smul` instance (#16713) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This generalizes the `has_continuous_smul ℝ≥0 ℝ` instance to `has_continuous_smul ℝ≥0 α` whenever there is a `mul_action ℝ α` with `has_continuous_smul R α` (the `mul_action` is the minimal assumption to get an induced action of `ℝ≥0` on `α` in mathlib). --- src/topology/instances/nnreal.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/topology/instances/nnreal.lean b/src/topology/instances/nnreal.lean index c28df7fb6cc83..758a7ded46ff4 100644 --- a/src/topology/instances/nnreal.lean +++ b/src/topology/instances/nnreal.lean @@ -21,7 +21,7 @@ Instances for the following typeclasses are defined: * `order_topology ℝ≥0` * `has_continuous_sub ℝ≥0` * `has_continuous_inv₀ ℝ≥0` (continuity of `x⁻¹` away from `0`) -* `has_continuous_smul ℝ≥0 ℝ` +* `has_continuous_smul ℝ≥0 α` (whenever `α` has a continuous `mul_action ℝ α`) Everything is inherited from the corresponding structures on the reals. @@ -116,9 +116,9 @@ instance : has_continuous_inv₀ ℝ≥0 := ⟨λ x hx, tendsto_coe.1 $ (real.tendsto_inv $ nnreal.coe_ne_zero.2 hx).comp continuous_coe.continuous_at⟩ -instance : has_continuous_smul ℝ≥0 ℝ := -{ continuous_smul := real.continuous_mul.comp $ - (continuous_subtype_val.comp continuous_fst).prod_mk continuous_snd } +instance [topological_space α] [mul_action ℝ α] [has_continuous_smul ℝ α] : + has_continuous_smul ℝ≥0 α := +{ continuous_smul := (continuous_induced_dom.comp continuous_fst).smul continuous_snd } @[norm_cast] lemma has_sum_coe {f : α → ℝ≥0} {r : ℝ≥0} : has_sum (λa, (f a : ℝ)) (r : ℝ) ↔ has_sum f r := From 699c2cabc9130c0eee01a57d55ba8449b695f1c7 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Fri, 30 Sep 2022 13:10:15 +0000 Subject: [PATCH 466/828] feat(analysis/convex/between): betweenness in affine spaces (#16191) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the notions of (weak and strict) betweenness for points in affine spaces over an ordered ring, for use in describing geometrical configurations. Until convexity is refactored to support abstract affine combination spaces, this means having a definition `affine_segment` that duplicates `segment` in the affine space case (and is proved to equal `segment` when the affine space is a module considered as an affine space over itself). However, the bulk of results concern betweenness, not `affine_segment`, and so would be just as relevant after such a refactor, even if some of the proofs would change, and indeed most of the things stated about `affine_segment` involve `+ᵥ` and `-ᵥ` and so would still be meaningful results, distinct from those already present for `segment`, after such a refactor (at which point they might apply for whatever typeclass describes an `add_torsor` for a module that is also an abstract affine combination space where the two affine combination structures agree). So I think the actual duplication here is minimal and defining `affine_segment` is a reasonable approach to allow betweenness to be handled in affine spaces now rather than making it depend on a possible future refactor. There are certainly more things that could sensibly be stated about betweenness (e.g. various forms of Pasch's axiom), but I think this is a reasonable starting point. One thing I definitely intend to add in a followup is notions of two points being (weakly or strictly) on the same side or opposite sides of an affine subspace (e.g. a line); I think it will probably be most convenient to define those notions in terms of `same_ray` and then prove appropriate results about how they relate to betweenness for points. --- src/analysis/convex/between.lean | 500 +++++++++++++++++++++++++++++++ 1 file changed, 500 insertions(+) create mode 100644 src/analysis/convex/between.lean diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean new file mode 100644 index 0000000000000..4b03b63fea532 --- /dev/null +++ b/src/analysis/convex/between.lean @@ -0,0 +1,500 @@ +/- +Copyright (c) 2022 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers +-/ +import analysis.convex.segment +import linear_algebra.affine_space.finite_dimensional + +/-! +# Betweenness in affine spaces + +This file defines notions of a point in an affine space being between two given points. + +## Main definitions + +* `affine_segment R x y`: The segment of points weakly between `x` and `y`. +* `wbtw R x y z`: The point `y` is weakly between `x` and `z`. +* `sbtw R x y z`: The point `y` is strictly between `x` and `z`. + +-/ + +variables (R : Type*) {V V' P P' : Type*} + +open affine_map + +section ordered_ring + +variables [ordered_ring R] [add_comm_group V] [module R V] [add_torsor V P] +variables [add_comm_group V'] [module R V'] [add_torsor V' P'] + +include V + +/-- The segment of points weakly between `x` and `y`. When convexity is refactored to support +abstract affine combination spaces, this will no longer need to be a separate definition from +`segment`. However, lemmas involving `+ᵥ` or `-ᵥ` will still be relevant after such a +refactoring, as distinct from versions involving `+` or `-` in a module. -/ +def affine_segment (x y : P) := line_map x y '' (set.Icc (0 : R) 1) + +lemma affine_segment_eq_segment (x y : V) : affine_segment R x y = segment R x y := +by rw [segment_eq_image_line_map, affine_segment] + +lemma affine_segment_comm (x y : P) : affine_segment R x y = affine_segment R y x := +begin + refine set.ext (λ z, _), + split; + { rintro ⟨t, ht, hxy⟩, + refine ⟨1 - t, _, _⟩, + { rwa [set.sub_mem_Icc_iff_right, sub_self, sub_zero] }, + { rwa [line_map_apply_one_sub] } }, +end + +lemma left_mem_affine_segment (x y : P) : x ∈ affine_segment R x y := +⟨0, set.left_mem_Icc.2 zero_le_one, line_map_apply_zero _ _⟩ + +lemma right_mem_affine_segment (x y : P) : y ∈ affine_segment R x y := +⟨1, set.right_mem_Icc.2 zero_le_one, line_map_apply_one _ _⟩ + +include V' + +variables {R} + +@[simp] lemma affine_segment_image (f : P →ᵃ[R] P') (x y : P) : + f '' affine_segment R x y = affine_segment R (f x) (f y) := +begin + rw [affine_segment, affine_segment, set.image_image, ←comp_line_map], + refl +end + +omit V' + +variables (R) + +@[simp] lemma affine_segment_const_vadd_image (x y : P) (v : V) : + ((+ᵥ) v) '' affine_segment R x y = affine_segment R (v +ᵥ x) (v +ᵥ y) := +affine_segment_image (affine_equiv.const_vadd R P v : P →ᵃ[R] P) x y + +@[simp] lemma affine_segment_vadd_const_image (x y : V) (p : P) : + (+ᵥ p) '' affine_segment R x y = affine_segment R (x +ᵥ p) (y +ᵥ p) := +affine_segment_image (affine_equiv.vadd_const R p : V →ᵃ[R] P) x y + +@[simp] lemma affine_segment_const_vsub_image (x y p : P) : + ((-ᵥ) p) '' affine_segment R x y = affine_segment R (p -ᵥ x) (p -ᵥ y) := +affine_segment_image (affine_equiv.const_vsub R p : P →ᵃ[R] V) x y + +@[simp] lemma affine_segment_vsub_const_image (x y p : P) : + (-ᵥ p) '' affine_segment R x y = affine_segment R (x -ᵥ p) (y -ᵥ p) := +affine_segment_image ((affine_equiv.vadd_const R p).symm : P →ᵃ[R] V) x y + +variables {R} + +@[simp] lemma mem_const_vadd_affine_segment {x y z : P} (v : V) : + v +ᵥ z ∈ affine_segment R (v +ᵥ x) (v +ᵥ y) ↔ z ∈ affine_segment R x y := +by rw [←affine_segment_const_vadd_image, (add_action.injective v).mem_set_image] + +@[simp] lemma mem_vadd_const_affine_segment {x y z : V} (p : P) : + z +ᵥ p ∈ affine_segment R (x +ᵥ p) (y +ᵥ p) ↔ z ∈ affine_segment R x y := +by rw [←affine_segment_vadd_const_image, (vadd_right_injective p).mem_set_image] +variables {R} + +@[simp] lemma mem_const_vsub_affine_segment {x y z : P} (p : P) : + p -ᵥ z ∈ affine_segment R (p -ᵥ x) (p -ᵥ y) ↔ z ∈ affine_segment R x y := +by rw [←affine_segment_const_vsub_image, (vsub_right_injective p).mem_set_image] + +@[simp] lemma mem_vsub_const_affine_segment {x y z : P} (p : P) : + z -ᵥ p ∈ affine_segment R (x -ᵥ p) (y -ᵥ p) ↔ z ∈ affine_segment R x y := +by rw [←affine_segment_vsub_const_image, (vsub_left_injective p).mem_set_image] + +variables (R) + +/-- The point `y` is weakly between `x` and `z`. -/ +def wbtw (x y z : P) : Prop := y ∈ affine_segment R x z + +/-- The point `y` is strictly between `x` and `z`. -/ +def sbtw (x y z : P) : Prop := wbtw R x y z ∧ y ≠ x ∧ y ≠ z + +variables {R} + +include V' + +lemma wbtw.map {x y z : P} (h : wbtw R x y z) (f : P →ᵃ[R] P') : wbtw R (f x) (f y) (f z) := +begin + rw [wbtw, ←affine_segment_image], + exact set.mem_image_of_mem _ h +end + +lemma function.injective.wbtw_map_iff {x y z : P} {f : P →ᵃ[R] P'} (hf : function.injective f) : + wbtw R (f x) (f y) (f z) ↔ wbtw R x y z := +begin + refine ⟨λ h, _, λ h, h.map _⟩, + rwa [wbtw, ←affine_segment_image, hf.mem_set_image] at h +end + +lemma function.injective.sbtw_map_iff {x y z : P} {f : P →ᵃ[R] P'} (hf : function.injective f) : + sbtw R (f x) (f y) (f z) ↔ sbtw R x y z := +by simp_rw [sbtw, hf.wbtw_map_iff, hf.ne_iff] + +@[simp] lemma affine_equiv.wbtw_map_iff {x y z : P} (f : P ≃ᵃ[R] P') : + wbtw R (f x) (f y) (f z) ↔ wbtw R x y z := +begin + refine function.injective.wbtw_map_iff (_ : function.injective f.to_affine_map), + exact f.injective +end + +@[simp] lemma affine_equiv.sbtw_map_iff {x y z : P} (f : P ≃ᵃ[R] P') : + sbtw R (f x) (f y) (f z) ↔ sbtw R x y z := +begin + refine function.injective.sbtw_map_iff (_ : function.injective f.to_affine_map), + exact f.injective +end + +omit V' + +@[simp] lemma wbtw_const_vadd_iff {x y z : P} (v : V) : + wbtw R (v +ᵥ x) (v +ᵥ y) (v +ᵥ z) ↔ wbtw R x y z := +mem_const_vadd_affine_segment _ + +@[simp] lemma wbtw_vadd_const_iff {x y z : V} (p : P) : + wbtw R (x +ᵥ p) (y +ᵥ p) (z +ᵥ p) ↔ wbtw R x y z := +mem_vadd_const_affine_segment _ + +@[simp] lemma wbtw_const_vsub_iff {x y z : P} (p : P) : + wbtw R (p -ᵥ x) (p -ᵥ y) (p -ᵥ z) ↔ wbtw R x y z := +mem_const_vsub_affine_segment _ + +@[simp] lemma wbtw_vsub_const_iff {x y z : P} (p : P) : + wbtw R (x -ᵥ p) (y -ᵥ p) (z -ᵥ p) ↔ wbtw R x y z := +mem_vsub_const_affine_segment _ + +@[simp] lemma sbtw_const_vadd_iff {x y z : P} (v : V) : + sbtw R (v +ᵥ x) (v +ᵥ y) (v +ᵥ z) ↔ sbtw R x y z := +by simp_rw [sbtw, wbtw_const_vadd_iff, (add_action.injective v).ne_iff] + +@[simp] lemma sbtw_vadd_const_iff {x y z : V} (p : P) : + sbtw R (x +ᵥ p) (y +ᵥ p) (z +ᵥ p) ↔ sbtw R x y z := +by simp_rw [sbtw, wbtw_vadd_const_iff, (vadd_right_injective p).ne_iff] + +@[simp] lemma sbtw_const_vsub_iff {x y z : P} (p : P) : + sbtw R (p -ᵥ x) (p -ᵥ y) (p -ᵥ z) ↔ sbtw R x y z := +by simp_rw [sbtw, wbtw_const_vsub_iff, (vsub_right_injective p).ne_iff] + +@[simp] lemma sbtw_vsub_const_iff {x y z : P} (p : P) : + sbtw R (x -ᵥ p) (y -ᵥ p) (z -ᵥ p) ↔ sbtw R x y z := +by simp_rw [sbtw, wbtw_vsub_const_iff, (vsub_left_injective p).ne_iff] + +lemma sbtw.wbtw {x y z : P} (h : sbtw R x y z) : wbtw R x y z := +h.1 + +lemma sbtw.ne_left {x y z : P} (h : sbtw R x y z) : y ≠ x := +h.2.1 + +lemma sbtw.left_ne {x y z : P} (h : sbtw R x y z) : x ≠ y := +h.2.1.symm + +lemma sbtw.ne_right {x y z : P} (h : sbtw R x y z) : y ≠ z := +h.2.2 + +lemma sbtw.right_ne {x y z : P} (h : sbtw R x y z) : z ≠ y := +h.2.2.symm + +lemma sbtw.mem_image_Ioo {x y z : P} (h : sbtw R x y z) : + y ∈ line_map x z '' (set.Ioo (0 : R) 1) := +begin + rcases h with ⟨⟨t, ht, rfl⟩, hyx, hyz⟩, + rcases set.eq_endpoints_or_mem_Ioo_of_mem_Icc ht with rfl|rfl|ho, + { exfalso, simpa using hyx }, + { exfalso, simpa using hyz }, + { exact ⟨t, ho, rfl⟩ } +end + +lemma wbtw_comm {x y z : P} : wbtw R x y z ↔ wbtw R z y x := +by rw [wbtw, wbtw, affine_segment_comm] + +alias wbtw_comm ↔ wbtw.symm _ + +lemma sbtw_comm {x y z : P} : sbtw R x y z ↔ sbtw R z y x := +by rw [sbtw, sbtw, wbtw_comm, ←and_assoc, ←and_assoc, and.right_comm] + +alias sbtw_comm ↔ sbtw.symm _ + +variables (R) + +@[simp] lemma wbtw_self_left (x y : P) : wbtw R x x y := +left_mem_affine_segment _ _ _ + +@[simp] lemma wbtw_self_right (x y : P) : wbtw R x y y := +right_mem_affine_segment _ _ _ + +@[simp] lemma wbtw_self_iff {x y : P} : wbtw R x y x ↔ y = x := +begin + refine ⟨λ h, _, λ h, _⟩, + { simpa [wbtw, affine_segment] using h }, + { rw h, + exact wbtw_self_left R x x } +end + +@[simp] lemma not_sbtw_self_left (x y : P) : ¬ sbtw R x x y := +λ h, h.ne_left rfl + +@[simp] lemma not_sbtw_self_right (x y : P) : ¬ sbtw R x y y := +λ h, h.ne_right rfl + +variables {R} + +lemma sbtw.left_ne_right {x y z : P} (h : sbtw R x y z) : x ≠ z := +begin + rintro rfl, + rw [sbtw, wbtw_self_iff] at h, + exact h.2.1 h.1 +end + +lemma sbtw_iff_mem_image_Ioo_and_ne [no_zero_smul_divisors R V] {x y z : P} : + sbtw R x y z ↔ y ∈ line_map x z '' (set.Ioo (0 : R) 1) ∧ x ≠ z := +begin + refine ⟨λ h, ⟨h.mem_image_Ioo, h.left_ne_right⟩, λ h, _⟩, + rcases h with ⟨⟨t, ht, rfl⟩, hxz⟩, + refine ⟨⟨t, set.mem_Icc_of_Ioo ht, rfl⟩, _⟩, + rw [line_map_apply, ←@vsub_ne_zero V, ←@vsub_ne_zero V _ _ _ _ z, vadd_vsub_assoc, + vadd_vsub_assoc, ←neg_vsub_eq_vsub_rev z x, ←@neg_one_smul R, ←add_smul, + ←sub_eq_add_neg], + simp [smul_ne_zero, hxz.symm, sub_eq_zero, ht.1.ne.symm, ht.2.ne] +end + +variables (R) + +@[simp] lemma not_sbtw_self (x y : P) : ¬ sbtw R x y x := +λ h, h.left_ne_right rfl + +lemma wbtw_swap_left_iff [no_zero_smul_divisors R V] {x y : P} (z : P) : + (wbtw R x y z ∧ wbtw R y x z) ↔ x = y := +begin + split, + { rintro ⟨hxyz, hyxz⟩, + rcases hxyz with ⟨ty, hty, rfl⟩, + rcases hyxz with ⟨tx, htx, hx⟩, + simp_rw [line_map_apply, ←add_vadd] at hx, + rw [←@vsub_eq_zero_iff_eq V, vadd_vsub, vsub_vadd_eq_vsub_sub, smul_sub, smul_smul, + ←sub_smul, ←add_smul, smul_eq_zero] at hx, + rcases hx with h|h, + { nth_rewrite 0 ←mul_one tx at h, + rw [←mul_sub, add_eq_zero_iff_neg_eq] at h, + have h' : ty = 0, + { refine le_antisymm _ hty.1, + rw [←h, left.neg_nonpos_iff], + exact mul_nonneg htx.1 (sub_nonneg.2 hty.2) }, + simp [h'] }, + { rw vsub_eq_zero_iff_eq at h, + simp [h] } }, + { rintro rfl, + exact ⟨wbtw_self_left _ _ _, wbtw_self_left _ _ _⟩ } +end + +lemma wbtw_swap_right_iff [no_zero_smul_divisors R V] (x : P) {y z : P} : + (wbtw R x y z ∧ wbtw R x z y) ↔ y = z := +begin + nth_rewrite 0 wbtw_comm, + nth_rewrite 1 wbtw_comm, + rw eq_comm, + exact wbtw_swap_left_iff R x +end + +lemma wbtw_rotate_iff [no_zero_smul_divisors R V] (x : P) {y z : P} : + (wbtw R x y z ∧ wbtw R z x y) ↔ x = y := +by rw [wbtw_comm, wbtw_swap_right_iff, eq_comm] + +variables {R} + +lemma wbtw.swap_left_iff [no_zero_smul_divisors R V] {x y z : P} (h : wbtw R x y z) : + wbtw R y x z ↔ x = y := +by rw [←wbtw_swap_left_iff R z, and_iff_right h] + +lemma wbtw.swap_right_iff [no_zero_smul_divisors R V] {x y z : P} (h : wbtw R x y z) : + wbtw R x z y ↔ y = z := +by rw [←wbtw_swap_right_iff R x, and_iff_right h] + +lemma wbtw.rotate_iff [no_zero_smul_divisors R V] {x y z : P} (h : wbtw R x y z) : + wbtw R z x y ↔ x = y := +by rw [←wbtw_rotate_iff R x, and_iff_right h] + +lemma sbtw.not_swap_left [no_zero_smul_divisors R V] {x y z : P} (h : sbtw R x y z) : + ¬ wbtw R y x z := +λ hs, h.left_ne (h.wbtw.swap_left_iff.1 hs) + +lemma sbtw.not_swap_right [no_zero_smul_divisors R V] {x y z : P} (h : sbtw R x y z) : + ¬ wbtw R x z y := +λ hs, h.ne_right (h.wbtw.swap_right_iff.1 hs) + +lemma sbtw.not_rotate [no_zero_smul_divisors R V] {x y z : P} (h : sbtw R x y z) : + ¬ wbtw R z x y := +λ hs, h.left_ne (h.wbtw.rotate_iff.1 hs) + +lemma wbtw.trans_left {w x y z : P} (h₁ : wbtw R w y z) (h₂ : wbtw R w x y) : wbtw R w x z := +begin + rcases h₁ with ⟨t₁, ht₁, rfl⟩, + rcases h₂ with ⟨t₂, ht₂, rfl⟩, + refine ⟨t₂ * t₁, ⟨mul_nonneg ht₂.1 ht₁.1, mul_le_one ht₂.2 ht₁.1 ht₁.2⟩, _⟩, + simp [line_map_apply, smul_smul] +end + +lemma wbtw.trans_right {w x y z : P} (h₁ : wbtw R w x z) (h₂ : wbtw R x y z) : wbtw R w y z := +begin + rw wbtw_comm at *, + exact h₁.trans_left h₂ +end + +lemma wbtw.trans_sbtw_left [no_zero_smul_divisors R V] {w x y z : P} (h₁ : wbtw R w y z) + (h₂ : sbtw R w x y) : sbtw R w x z := +begin + refine ⟨h₁.trans_left h₂.wbtw, h₂.ne_left, _⟩, + rintro rfl, + exact h₂.right_ne ((wbtw_swap_right_iff R w).1 ⟨h₁, h₂.wbtw⟩) +end + +lemma wbtw.trans_sbtw_right [no_zero_smul_divisors R V] {w x y z : P} (h₁ : wbtw R w x z) + (h₂ : sbtw R x y z) : sbtw R w y z := +begin + rw wbtw_comm at *, + rw sbtw_comm at *, + exact h₁.trans_sbtw_left h₂ +end + +lemma sbtw.trans_left [no_zero_smul_divisors R V] {w x y z : P} (h₁ : sbtw R w y z) + (h₂ : sbtw R w x y) : sbtw R w x z := +h₁.wbtw.trans_sbtw_left h₂ + +lemma sbtw.trans_right [no_zero_smul_divisors R V] {w x y z : P} (h₁ : sbtw R w x z) + (h₂ : sbtw R x y z) : sbtw R w y z := +h₁.wbtw.trans_sbtw_right h₂ + +end ordered_ring + +section linear_ordered_field + +variables [linear_ordered_field R] [add_comm_group V] [module R V] [add_torsor V P] + +include V + +variables {R} + +lemma wbtw_smul_vadd_smul_vadd_of_nonneg_of_le (x : P) (v : V) {r₁ r₂ : R} (hr₁ : 0 ≤ r₁) + (hr₂ : r₁ ≤ r₂) : wbtw R x (r₁ • v +ᵥ x) (r₂ • v +ᵥ x) := +begin + refine ⟨r₁ / r₂, ⟨div_nonneg hr₁ (hr₁.trans hr₂), div_le_one_of_le hr₂ (hr₁.trans hr₂)⟩, _⟩, + by_cases h : r₁ = 0, { simp [h] }, + simp [line_map_apply, smul_smul, ((hr₁.lt_of_ne' h).trans_le hr₂).ne.symm] +end + +lemma wbtw_or_wbtw_smul_vadd_of_nonneg (x : P) (v : V) {r₁ r₂ : R} (hr₁ : 0 ≤ r₁) (hr₂ : 0 ≤ r₂) : + wbtw R x (r₁ • v +ᵥ x) (r₂ • v +ᵥ x) ∨ wbtw R x (r₂ • v +ᵥ x) (r₁ • v +ᵥ x) := +begin + rcases le_total r₁ r₂ with h|h, + { exact or.inl (wbtw_smul_vadd_smul_vadd_of_nonneg_of_le x v hr₁ h) }, + { exact or.inr (wbtw_smul_vadd_smul_vadd_of_nonneg_of_le x v hr₂ h) } +end + +lemma wbtw_smul_vadd_smul_vadd_of_nonpos_of_le (x : P) (v : V) {r₁ r₂ : R} (hr₁ : r₁ ≤ 0) + (hr₂ : r₂ ≤ r₁) : wbtw R x (r₁ • v +ᵥ x) (r₂ • v +ᵥ x) := +begin + convert wbtw_smul_vadd_smul_vadd_of_nonneg_of_le x (-v) (left.nonneg_neg_iff.2 hr₁) + (neg_le_neg_iff.2 hr₂) using 1; + rw neg_smul_neg +end + +lemma wbtw_or_wbtw_smul_vadd_of_nonpos (x : P) (v : V) {r₁ r₂ : R} (hr₁ : r₁ ≤ 0) (hr₂ : r₂ ≤ 0) : + wbtw R x (r₁ • v +ᵥ x) (r₂ • v +ᵥ x) ∨ wbtw R x (r₂ • v +ᵥ x) (r₁ • v +ᵥ x) := +begin + rcases le_total r₁ r₂ with h|h, + { exact or.inr (wbtw_smul_vadd_smul_vadd_of_nonpos_of_le x v hr₂ h) }, + { exact or.inl (wbtw_smul_vadd_smul_vadd_of_nonpos_of_le x v hr₁ h) } +end + +lemma wbtw_smul_vadd_smul_vadd_of_nonpos_of_nonneg (x : P) (v : V) {r₁ r₂ : R} (hr₁ : r₁ ≤ 0) + (hr₂ : 0 ≤ r₂) : wbtw R (r₁ • v +ᵥ x) x (r₂ • v +ᵥ x) := +begin + convert wbtw_smul_vadd_smul_vadd_of_nonneg_of_le (r₁ • v +ᵥ x) v (left.nonneg_neg_iff.2 hr₁) + (neg_le_sub_iff_le_add.2 ((le_add_iff_nonneg_left r₁).2 hr₂)) using 1; + simp [sub_smul, ←add_vadd] +end + +lemma wbtw_smul_vadd_smul_vadd_of_nonneg_of_nonpos (x : P) (v : V) {r₁ r₂ : R} (hr₁ : 0 ≤ r₁) + (hr₂ : r₂ ≤ 0) : wbtw R (r₁ • v +ᵥ x) x (r₂ • v +ᵥ x) := +begin + rw wbtw_comm, + exact wbtw_smul_vadd_smul_vadd_of_nonpos_of_nonneg x v hr₂ hr₁ +end + +lemma wbtw.trans_left_right {w x y z : P} (h₁ : wbtw R w y z) (h₂ : wbtw R w x y) : wbtw R x y z := +begin + rcases h₁ with ⟨t₁, ht₁, rfl⟩, + rcases h₂ with ⟨t₂, ht₂, rfl⟩, + refine ⟨(t₁ - t₂ * t₁) / (1 - t₂ * t₁), + ⟨div_nonneg (sub_nonneg.2 (mul_le_of_le_one_left ht₁.1 ht₂.2)) + (sub_nonneg.2 (mul_le_one ht₂.2 ht₁.1 ht₁.2)), + div_le_one_of_le (sub_le_sub_right ht₁.2 _) + (sub_nonneg.2 (mul_le_one ht₂.2 ht₁.1 ht₁.2))⟩, _⟩, + simp only [line_map_apply, smul_smul, ←add_vadd, vsub_vadd_eq_vsub_sub, smul_sub, ←sub_smul, + ←add_smul, vadd_vsub, vadd_right_cancel_iff, div_mul_eq_mul_div, div_sub_div_same], + nth_rewrite 0 [←mul_one (t₁ - t₂ * t₁)], + rw [←mul_sub, mul_div_assoc], + by_cases h : 1 - t₂ * t₁ = 0, + { rw [sub_eq_zero, eq_comm] at h, + rw h, + suffices : t₁ = 1, by simp [this], + exact eq_of_le_of_not_lt ht₁.2 + (λ ht₁lt, (mul_lt_one_of_nonneg_of_lt_one_right ht₂.2 ht₁.1 ht₁lt).ne h) }, + { rw div_self h, + ring_nf } +end + +lemma wbtw.trans_right_left {w x y z : P} (h₁ : wbtw R w x z) (h₂ : wbtw R x y z) : wbtw R w x y := +begin + rw wbtw_comm at *, + exact h₁.trans_left_right h₂ +end + +lemma sbtw.trans_left_right {w x y z : P} (h₁ : sbtw R w y z) (h₂ : sbtw R w x y) : sbtw R x y z := +⟨h₁.wbtw.trans_left_right h₂.wbtw, h₂.right_ne, h₁.ne_right⟩ + +lemma sbtw.trans_right_left {w x y z : P} (h₁ : sbtw R w x z) (h₂ : sbtw R x y z) : sbtw R w x y := +⟨h₁.wbtw.trans_right_left h₂.wbtw, h₁.ne_left, h₂.left_ne⟩ + +lemma wbtw.collinear {x y z : P} (h : wbtw R x y z) : collinear R ({x, y, z} : set P) := +begin + rw collinear_iff_exists_forall_eq_smul_vadd, + refine ⟨x, z -ᵥ x, _⟩, + intros p hp, + simp_rw [set.mem_insert_iff, set.mem_singleton_iff] at hp, + rcases hp with rfl|rfl|rfl, + { refine ⟨0, _⟩, simp }, + { rcases h with ⟨t, -, rfl⟩, + exact ⟨t, rfl⟩ }, + { refine ⟨1, _⟩, simp } +end + +lemma collinear.wbtw_or_wbtw_or_wbtw {x y z : P} (h : collinear R ({x, y, z} : set P)) : + wbtw R x y z ∨ wbtw R y z x ∨ wbtw R z x y := +begin + rw collinear_iff_of_mem R (set.mem_insert _ _) at h, + rcases h with ⟨v, h⟩, + simp_rw [set.mem_insert_iff, set.mem_singleton_iff] at h, + have hy := h y (or.inr (or.inl rfl)), + have hz := h z (or.inr (or.inr rfl)), + rcases hy with ⟨ty, rfl⟩, + rcases hz with ⟨tz, rfl⟩, + rcases lt_trichotomy ty 0 with hy0|rfl|hy0, + { rcases lt_trichotomy tz 0 with hz0|rfl|hz0, + { nth_rewrite 1 wbtw_comm, + rw ←or_assoc, + exact or.inl (wbtw_or_wbtw_smul_vadd_of_nonpos _ _ hy0.le hz0.le) }, + { simp }, + { exact or.inr (or.inr (wbtw_smul_vadd_smul_vadd_of_nonneg_of_nonpos _ _ hz0.le hy0.le)) } }, + { simp }, + { rcases lt_trichotomy tz 0 with hz0|rfl|hz0, + { refine or.inr (or.inr (wbtw_smul_vadd_smul_vadd_of_nonpos_of_nonneg _ _ hz0.le hy0.le)) }, + { simp }, + { nth_rewrite 1 wbtw_comm, + rw ←or_assoc, + exact or.inl (wbtw_or_wbtw_smul_vadd_of_nonneg _ _ hy0.le hz0.le) } } +end + +end linear_ordered_field From 98ed2a204009eaa0ef3e634669574091dc619199 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 30 Sep 2022 13:10:16 +0000 Subject: [PATCH 467/828] move(algebra/order/ring_lemmas): Rename file (#16520) Rename `algebra.order.monoid_lemmas_zero_lt` to `algebra.order.ring_lemmas` because `algebra.order.monoid_lemmas_zero_lt` is to `algebra.order.ring` what `algebra.order.monoid_lemmas` is to `algebra.order.monoid`. --- src/algebra/order/ring.lean | 5 ++--- .../order/{monoid_lemmas_zero_lt.lean => ring_lemmas.lean} | 0 2 files changed, 2 insertions(+), 3 deletions(-) rename src/algebra/order/{monoid_lemmas_zero_lt.lean => ring_lemmas.lean} (100%) diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index b8d50b646ad02..207719e1f75cb 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -4,10 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro -/ import algebra.char_zero.defs -import algebra.order.group -import algebra.order.monoid_lemmas_zero_lt -import algebra.order.sub import algebra.hom.ring +import algebra.order.group +import algebra.order.ring_lemmas import data.set.intervals.basic /-! diff --git a/src/algebra/order/monoid_lemmas_zero_lt.lean b/src/algebra/order/ring_lemmas.lean similarity index 100% rename from src/algebra/order/monoid_lemmas_zero_lt.lean rename to src/algebra/order/ring_lemmas.lean From 3df3c6179d5c3a89507411ceea18ff523dfea679 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 30 Sep 2022 13:10:18 +0000 Subject: [PATCH 468/828] feat(linear_algebra/quadratic_form/basic): algebraic lemmas about `bilin_form.to_quadratic_form` (#16616) Following the usual pattern, we defined the bundle additive morphism so that we can copy across the salient lemmas about sums. The `polar_to_quadratic_form` lemma in the diff was an existing lemma that has just been moved below the new `semiring` section. --- src/linear_algebra/quadratic_form/basic.lean | 62 ++++++++++++++++---- 1 file changed, 49 insertions(+), 13 deletions(-) diff --git a/src/linear_algebra/quadratic_form/basic.lean b/src/linear_algebra/quadratic_form/basic.lean index 17aa943d65917..f6d660d7d6d1c 100644 --- a/src/linear_algebra/quadratic_form/basic.lean +++ b/src/linear_algebra/quadratic_form/basic.lean @@ -68,6 +68,8 @@ universes u v w variables {S : Type*} variables {R R₁: Type*} {M : Type*} +open_locale big_operators + section polar variables [ring R] [comm_ring R₁] [add_comm_group M] @@ -162,7 +164,6 @@ protected def copy (Q : quadratic_form R M) (Q' : M → R) (h : Q' = ⇑Q) : qua end fun_like - section semiring variables [semiring R] [add_comm_monoid M] [module R M] variables (Q : quadratic_form R M) @@ -374,7 +375,6 @@ def eval_add_monoid_hom (m : M) : quadratic_form R M →+ R := (pi.eval_add_monoid_hom _ m).comp coe_fn_add_monoid_hom section sum -open_locale big_operators @[simp] lemma coe_fn_sum {ι : Type*} (Q : ι → quadratic_form R M) (s : finset ι) : ⇑(∑ i in s, Q i) = ∑ i in s, Q i := @@ -521,15 +521,7 @@ quadratic form. namespace bilin_form open quadratic_form -section ring -variables [ring R] [add_comm_group M] [module R M] -variables {B : bilin_form R M} - -lemma polar_to_quadratic_form (x y : M) : polar (λ x, B x x) x y = B x y + B y x := -by { simp only [add_assoc, add_sub_cancel', add_right, polar, add_left_inj, add_neg_cancel_left, - add_left, sub_eq_add_neg _ (B y y), add_comm (B y x) _] } - -end ring +section semiring variables [semiring R] [add_comm_monoid M] [module R M] variables {B : bilin_form R M} @@ -550,6 +542,52 @@ variables (R M) @[simp] lemma to_quadratic_form_zero : (0 : bilin_form R M).to_quadratic_form = 0 := rfl end +@[simp] lemma to_quadratic_form_add (B₁ B₂ : bilin_form R M) : + (B₁ + B₂).to_quadratic_form = B₁.to_quadratic_form + B₂.to_quadratic_form := rfl + +@[simp] lemma to_quadratic_form_smul [monoid S] [distrib_mul_action S R] [smul_comm_class S R R] + (a : S) (B : bilin_form R M) : + (a • B).to_quadratic_form = a • B.to_quadratic_form := rfl + +section +variables (R M) +/-- `bilin_form.to_quadratic_form` as an additive homomorphism -/ +@[simps] def to_quadratic_form_add_monoid_hom : bilin_form R M →+ quadratic_form R M := +{ to_fun := to_quadratic_form, + map_zero' := to_quadratic_form_zero _ _, + map_add' := to_quadratic_form_add } +end + +@[simp] lemma to_quadratic_form_list_sum (B : list (bilin_form R M)) : + B.sum.to_quadratic_form = (B.map to_quadratic_form).sum := +map_list_sum (to_quadratic_form_add_monoid_hom R M) B + +@[simp] lemma to_quadratic_form_multiset_sum (B : multiset (bilin_form R M)) : + B.sum.to_quadratic_form = (B.map to_quadratic_form).sum := +map_multiset_sum (to_quadratic_form_add_monoid_hom R M) B + +@[simp] lemma to_quadratic_form_sum {ι : Type*} (s : finset ι) (B : ι → bilin_form R M) : + (∑ i in s, B i).to_quadratic_form = ∑ i in s, (B i).to_quadratic_form := +map_sum (to_quadratic_form_add_monoid_hom R M) B s + +end semiring + +section ring +variables [ring R] [add_comm_group M] [module R M] +variables {B : bilin_form R M} + +lemma polar_to_quadratic_form (x y : M) : polar (λ x, B x x) x y = B x y + B y x := +by { simp only [add_assoc, add_sub_cancel', add_right, polar, add_left_inj, add_neg_cancel_left, + add_left, sub_eq_add_neg _ (B y y), add_comm (B y x) _] } + +@[simp] lemma to_quadratic_form_neg (B : bilin_form R M) : + (-B).to_quadratic_form = -B.to_quadratic_form := rfl + +@[simp] lemma to_quadratic_form_sub (B₁ B₂ : bilin_form R M) : + (B₁ - B₂).to_quadratic_form = B₁.to_quadratic_form - B₂.to_quadratic_form := rfl + +end ring + end bilin_form namespace quadratic_form @@ -898,8 +936,6 @@ end bilin_form namespace quadratic_form -open_locale big_operators - open finset bilin_form variables {M₁ : Type*} [semiring R] [comm_semiring R₁] [add_comm_monoid M] [add_comm_monoid M₁] From c804ede606e2652bc3193d0e9468d0a44325cb25 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 30 Sep 2022 13:10:19 +0000 Subject: [PATCH 469/828] feat(category_theory/groupoid): simplify groupoid.inv to category_theory.inv (#16624) Add simp lemma `groupoid.inv_eq_inv` to simplify `groupoid.inv` to `category_theory.inv` (which uses the `is_iso` instance) to gain access to the developed API around `category_theory.inv`. This isn't a defeq though so I can imagine sometimes we may want to `simp [-groupoid.inv_eq_inv]`, but most of the times this simp lemma makes things smoother. --- src/category_theory/groupoid.lean | 11 ++++++++--- 1 file changed, 8 insertions(+), 3 deletions(-) diff --git a/src/category_theory/groupoid.lean b/src/category_theory/groupoid.lean index 73b3692f5381b..0140965cd0bdd 100644 --- a/src/category_theory/groupoid.lean +++ b/src/category_theory/groupoid.lean @@ -40,8 +40,6 @@ class groupoid (obj : Type u) extends category.{v} obj : Type (max u (v+1)) := restate_axiom groupoid.inv_comp' restate_axiom groupoid.comp_inv' -attribute [simp] groupoid.inv_comp groupoid.comp_inv - /-- A `large_groupoid` is a groupoid where the objects live in `Type (u+1)` while the morphisms live in `Type u`. @@ -59,7 +57,14 @@ variables {C : Type u} [groupoid.{v} C] {X Y : C} @[priority 100] -- see Note [lower instance priority] instance is_iso.of_groupoid (f : X ⟶ Y) : is_iso f := -⟨⟨groupoid.inv f, by simp⟩⟩ +⟨⟨groupoid.inv f, groupoid.comp_inv f, groupoid.inv_comp f⟩⟩ + +@[simp] lemma groupoid.inv_eq_inv (f : X ⟶ Y) : groupoid.inv f = inv f := +is_iso.eq_inv_of_hom_inv_id $ groupoid.comp_inv f + +/-- `groupoid.inv` is involutive. -/ +@[simps] def groupoid.inv_equiv : (X ⟶ Y) ≃ (Y ⟶ X) := +⟨groupoid.inv, groupoid.inv, λ f, by simp, λ f, by simp⟩ variables (X Y) From a52be2a5d88ffc603e7f442cfe7ce48b9ca51e98 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Fri, 30 Sep 2022 13:10:20 +0000 Subject: [PATCH 470/828] feat(category_theory/localization): whiskering_left_equivalence and definition of the predicate (#16646) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this PR, an equivalence `localization.construction.whiskering_left_equivalence W D : (W.localization ⥤ D) ≌ W.functors_inverting D` is obtained and the predicate `L.is_localization W` is defined for a functor `L : C ⥤ D`. --- .../localization/construction.lean | 112 ++++++++++++++---- .../localization/predicate.lean | 44 +++++++ src/category_theory/morphism_property.lean | 14 +++ 3 files changed, 148 insertions(+), 22 deletions(-) create mode 100644 src/category_theory/localization/predicate.lean diff --git a/src/category_theory/localization/construction.lean b/src/category_theory/localization/construction.lean index 7e86810b73ce3..ba6baf39f8ae9 100644 --- a/src/category_theory/localization/construction.lean +++ b/src/category_theory/localization/construction.lean @@ -30,27 +30,15 @@ The expected property of `lift G hG` if expressed by the lemma `fac` and the uniqueness is expressed by `uniq`. TODO: -1) show that for any category `E`, the composition of functors gives -an equivalence of categories between `W.localization ⥤ E` and the full -subcategory of `C ⥤ E` consisting of functors inverting `W`. (This only -requires an extension property for natural transformations of functors.) - -2) define a predicate `is_localization L W` for a functor `L : C ⥤ D` and -a class of morphisms `W` in `C` expressing that it is a localization with respect -to `W`, i.e. that it inverts `W` and that the obvious functor `W.localization ⥤ D` -induced by `L` is an equivalence of categories. (It is more straightforward -to define this predicate this way rather than by using a universal property which -may imply attempting at quantifying on all universes.) - -3) implement a constructor for `is_localization L W` which would take +1) implement a constructor for `is_localization L W` which would take as an input a *strict* universal property (`lift`/`fac`/`uniq`) similar to what is obtained here for `W.localization`. (Practically speaking, this is the easiest way to show that a functor is a localization.) -4) when we have `is_localization L W`, then show that `D ⥤ E` identifies +2) when we have `is_localization L W`, then show that `D ⥤ E` identifies to the full subcategory of `C ⥤ E` consisting of `W`-inverting functors. -5) provide an API for the lifting of functors `C ⥤ E`, for which +3) provide an API for the lifting of functors `C ⥤ E`, for which `fac`/`uniq` assertions would be expressed as isomorphisms rather than by equalities of functors. @@ -66,7 +54,7 @@ open category_theory.category namespace category_theory -variables {C D : Type*} [category C] [category D] (W : morphism_property C) +variables {C : Type*} [category C] (W : morphism_property C) {D : Type*} [category D] namespace localization @@ -143,13 +131,10 @@ def Wiso {X Y : C} (w : X ⟶ Y) (hw : W w) : iso (W.Q.obj X) (W.Q.obj Y) := /-- The formal inverse in `W.localization` of a morphism `w` in `W`. -/ abbreviation Winv {X Y : C} (w : X ⟶ Y) (hw : W w) := (Wiso w hw).inv -end construction - -end localization - -namespace localization +variable (W) -namespace construction +lemma _root_.category_theory.morphism_property.Q_inverts : W.is_inverted_by W.Q := +λ X Y w hw, is_iso.of_iso (localization.construction.Wiso w hw) variables {W} (G : C ⥤ D) (hG : W.is_inverted_by G) @@ -315,6 +300,89 @@ begin nat_trans_extension_app, nat_trans_extension.app_eq], end +lemma nat_trans_hcomp_injective {F G : W.localization ⥤ D} {τ₁ τ₂ : F ⟶ G} + (h : 𝟙 W.Q ◫ τ₁ = 𝟙 W.Q ◫ τ₂) : τ₁ = τ₂ := +begin + ext X, + have eq := (obj_equiv W).right_inv X, + simp only [obj_equiv] at eq, + rw [← eq, ← nat_trans.id_hcomp_app, ← nat_trans.id_hcomp_app, h], +end + +variables (W D) + +namespace whiskering_left_equivalence + +/-- The functor `(W.localization ⥤ D) ⥤ (W.functors_inverting D)` induced by the +composition with `W.Q : C ⥤ W.localization`. -/ +@[simps] +def functor : (W.localization ⥤ D) ⥤ (W.functors_inverting D) := +full_subcategory.lift _ ((whiskering_left _ _ D).obj W.Q) + (λ F, morphism_property.is_inverted_by.of_comp W W.Q W.Q_inverts _) + +/-- The function `(W.functors_inverting D) ⥤ (W.localization ⥤ D)` induced by +`construction.lift`. -/ +@[simps] +def inverse : (W.functors_inverting D) ⥤ (W.localization ⥤ D) := +{ obj := λ G, lift G.obj G.property, + map := λ G₁ G₂ τ, nat_trans_extension (eq_to_hom (by rw fac) ≫ τ ≫ eq_to_hom (by rw fac)), + map_id' := λ G, nat_trans_hcomp_injective begin + rw nat_trans_extension_hcomp, + ext X, + simpa only [nat_trans.comp_app, eq_to_hom_app, eq_to_hom_refl, comp_id, id_comp, + nat_trans.hcomp_id_app, nat_trans.id_app, functor.map_id], + end, + map_comp' := λ G₁ G₂ G₃ τ₁ τ₂, nat_trans_hcomp_injective begin + ext X, + simpa only [nat_trans_extension_hcomp, nat_trans.comp_app, eq_to_hom_app, eq_to_hom_refl, + id_comp, comp_id, nat_trans.hcomp_app, nat_trans.id_app, functor.map_id, + nat_trans_extension_app, nat_trans_extension.app_eq], + end, } + +/-- The unit isomorphism of the equivalence of categories `whiskering_left_equivalence W D`. -/ +@[simps] +def unit_iso : 𝟭 (W.localization ⥤ D) ≅ functor W D ⋙ inverse W D := eq_to_iso +begin + refine functor.ext (λ G, _) (λ G₁ G₂ τ, _), + { apply uniq, + dsimp [functor], + rw fac, }, + { apply nat_trans_hcomp_injective, + ext X, + simp only [functor.id_map, nat_trans.hcomp_app, comp_id, functor.comp_map, + inverse_map, nat_trans.comp_app, eq_to_hom_app, eq_to_hom_refl, nat_trans_extension_app, + nat_trans_extension.app_eq, functor_map_app, id_comp], }, +end + +/-- The counit isomorphism of the equivalence of categories `whiskering_left_equivalence W D`. -/ +@[simps] +def counit_iso : inverse W D ⋙ functor W D ≅ 𝟭 (W.functors_inverting D) := eq_to_iso +begin + refine functor.ext _ _, + { rintro ⟨G, hG⟩, + ext1, + apply fac, }, + { rintros ⟨G₁, hG₁⟩ ⟨G₂, hG₂⟩ f, + ext X, + apply nat_trans_extension.app_eq, }, +end + +end whiskering_left_equivalence + +/-- The equivalence of categories `(W.localization ⥤ D) ≌ (W.functors_inverting D)` +induced by the composition with `W.Q : C ⥤ W.localization`. -/ +def whiskering_left_equivalence : (W.localization ⥤ D) ≌ W.functors_inverting D := +{ functor := whiskering_left_equivalence.functor W D, + inverse := whiskering_left_equivalence.inverse W D, + unit_iso := whiskering_left_equivalence.unit_iso W D, + counit_iso := whiskering_left_equivalence.counit_iso W D, + functor_unit_iso_comp' := λ F, begin + ext X, + simpa only [eq_to_hom_app, whiskering_left_equivalence.unit_iso_hom, + whiskering_left_equivalence.counit_iso_hom, eq_to_hom_map, eq_to_hom_trans, + eq_to_hom_refl], + end, } + end construction end localization diff --git a/src/category_theory/localization/predicate.lean b/src/category_theory/localization/predicate.lean new file mode 100644 index 0000000000000..dec36bd151ce2 --- /dev/null +++ b/src/category_theory/localization/predicate.lean @@ -0,0 +1,44 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import category_theory.localization.construction + +/-! + +# Predicate for localized categories + +In this file, a predicate `L.is_localization W` is introduced for a functor `L : C ⥤ D` +and `W : morphism_property C`. It expresses that `L` identifies `D` with the localized +category of `C` with respect to `W` (up to equivalence). + +-/ + +namespace category_theory + +variables {C D : Type*} [category C] [category D] +variables (L : C ⥤ D) (W : morphism_property C) + +namespace functor + +/-- The predicate expressing that, up to equivalence, a functor `L : C ⥤ D` +identifies the category `D` with the localized category of `C` with respect +to `W : morphism_property C`. -/ +class is_localization : Prop := +(inverts : W.is_inverted_by L) +(nonempty_is_equivalence : nonempty (is_equivalence (localization.construction.lift L inverts))) + +instance Q_is_localization : W.Q.is_localization W := +{ inverts := W.Q_inverts, + nonempty_is_equivalence := begin + suffices : localization.construction.lift W.Q W.Q_inverts = 𝟭 _, + { apply nonempty.intro, rw this, apply_instance, }, + apply localization.construction.uniq, + simpa only [localization.construction.fac], + end, } + +end functor + +end category_theory diff --git a/src/category_theory/morphism_property.lean b/src/category_theory/morphism_property.lean index 31947a267b336..c5da9502f0532 100644 --- a/src/category_theory/morphism_property.lean +++ b/src/category_theory/morphism_property.lean @@ -170,6 +170,11 @@ to isomorphisms in `D`. -/ def is_inverted_by (P : morphism_property C) (F : C ⥤ D) : Prop := ∀ ⦃X Y : C⦄ (f : X ⟶ Y) (hf : P f), is_iso (F.map f) +lemma is_inverted_by.of_comp {C₁ C₂ C₃ : Type*} [category C₁] [category C₂] [category C₃] + (W : morphism_property C₁) (F : C₁ ⥤ C₂) (hF : W.is_inverted_by F) (G : C₂ ⥤ C₃) : + W.is_inverted_by (F ⋙ G) := +λ X Y f hf, by { haveI := hF f hf, dsimp, apply_instance, } + /-- Given `app : Π X, F₁.obj X ⟶ F₂.obj X` where `F₁` and `F₂` are two functors, this is the `morphism_property C` satisfied by the morphisms in `C` with respect to whom `app` is natural. -/ @@ -200,6 +205,15 @@ end end naturality_property +/-- The full subcategory of `C ⥤ D` consisting of functors inverting morphisms in `W` -/ +@[derive category, nolint has_nonempty_instance] +def functors_inverting (W : morphism_property C) (D : Type*) [category D] := +full_subcategory (λ (F : C ⥤ D), W.is_inverted_by F) + +/-- A constructor for `W.functors_inverting D` -/ +def functors_inverting.mk {W : morphism_property C} {D : Type*} [category D] +(F : C ⥤ D) (hF : W.is_inverted_by F) : W.functors_inverting D := ⟨F, hF⟩ + end morphism_property end category_theory From 64ad902d75689c4263f0ffb229691c44ae8e0b61 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 30 Sep 2022 13:10:21 +0000 Subject: [PATCH 471/828] refactor(algebra/order/{smul,module}): Turn lemmas around (#16696) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Match the `mul` lemmas by having the `⁻¹` on the LHS in `inv_smul_le_iff`, `inv_smul_lt_iff`, etc... Also generalize for free to `ordered_add_comm_monoid`. --- src/algebra/order/module.lean | 20 ++++++------ src/algebra/order/smul.lean | 34 +++++++------------- src/linear_algebra/affine_space/ordered.lean | 9 +++--- 3 files changed, 25 insertions(+), 38 deletions(-) diff --git a/src/algebra/order/module.lean b/src/algebra/order/module.lean index 60ab25692c66c..8c3ef26904dd8 100644 --- a/src/algebra/order/module.lean +++ b/src/algebra/order/module.lean @@ -160,17 +160,17 @@ begin exact smul_le_smul_iff_of_pos (neg_pos_of_neg hc), end -lemma smul_lt_iff_of_neg (hc : c < 0) : c • a < b ↔ c⁻¹ • b < a := -begin - rw [←neg_neg c, ←neg_neg a, neg_smul_neg, inv_neg, neg_smul _ b, neg_lt_neg_iff], - exact smul_lt_iff_of_pos (neg_pos_of_neg hc), -end +lemma inv_smul_le_iff_of_neg (h : c < 0) : c⁻¹ • a ≤ b ↔ c • b ≤ a := +by { rw [←smul_le_smul_iff_of_neg h, smul_inv_smul₀ h.ne], apply_instance } -lemma lt_smul_iff_of_neg (hc : c < 0) : a < c • b ↔ b < c⁻¹ • a := -begin - rw [←neg_neg c, ←neg_neg b, neg_smul_neg, inv_neg, neg_smul _ a, neg_lt_neg_iff], - exact lt_smul_iff_of_pos (neg_pos_of_neg hc), -end +lemma inv_smul_lt_iff_of_neg (h : c < 0) : c⁻¹ • a < b ↔ c • b < a := +by { rw [←smul_lt_smul_iff_of_neg h, smul_inv_smul₀ h.ne], apply_instance } + +lemma smul_inv_le_iff_of_neg (h : c < 0) : a ≤ c⁻¹ • b ↔ b ≤ c • a := +by { rw [←smul_le_smul_iff_of_neg h, smul_inv_smul₀ h.ne], apply_instance } + +lemma smul_inv_lt_iff_of_neg (h : c < 0) : a < c⁻¹ • b ↔ b < c • a := +by { rw [←smul_lt_smul_iff_of_neg h, smul_inv_smul₀ h.ne], apply_instance } variables (M) diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index 7153c70c7d762..b47859245c111 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -174,11 +174,8 @@ instance linear_ordered_semiring.to_ordered_smul {R : Type*} [linear_ordered_sem ordered_smul.mk'' $ λ c, strict_mono_mul_left_of_pos section linear_ordered_semifield -variables [linear_ordered_semifield 𝕜] - -section ordered_add_comm_monoid -variables [ordered_add_comm_monoid M] [ordered_add_comm_monoid N] [mul_action_with_zero 𝕜 M] - [mul_action_with_zero 𝕜 N] +variables [linear_ordered_semifield 𝕜] [ordered_add_comm_monoid M] [ordered_add_comm_monoid N] + [mul_action_with_zero 𝕜 M] [mul_action_with_zero 𝕜 N] /-- To prove that a vector space over a linear ordered field is ordered, it suffices to verify only the first axiom of `ordered_smul`. -/ @@ -213,32 +210,24 @@ instance pi.ordered_smul' [ordered_smul 𝕜 M] : ordered_smul 𝕜 (ι → M) : /- Sometimes Lean fails to unify the module with the scalars, so we define another instance. -/ instance pi.ordered_smul'' : ordered_smul 𝕜 (ι → 𝕜) := @pi.ordered_smul' ι 𝕜 𝕜 _ _ _ _ -end ordered_add_comm_monoid - -section ordered_add_comm_group -variables [ordered_add_comm_group M] [mul_action_with_zero 𝕜 M] [ordered_smul 𝕜 M] {s : set M} - {a b : M} {c : 𝕜} +variables [ordered_smul 𝕜 M] {s : set M} {a b : M} {c : 𝕜} lemma smul_le_smul_iff_of_pos (hc : 0 < c) : c • a ≤ c • b ↔ a ≤ b := ⟨λ h, inv_smul_smul₀ hc.ne' a ▸ inv_smul_smul₀ hc.ne' b ▸ smul_le_smul_of_nonneg h (inv_nonneg.2 hc.le), λ h, smul_le_smul_of_nonneg h hc.le⟩ -lemma smul_lt_iff_of_pos (hc : 0 < c) : c • a < b ↔ a < c⁻¹ • b := -calc c • a < b ↔ c • a < c • c⁻¹ • b : by rw [smul_inv_smul₀ hc.ne'] -... ↔ a < c⁻¹ • b : smul_lt_smul_iff_of_pos hc +lemma inv_smul_le_iff (h : 0 < c) : c⁻¹ • a ≤ b ↔ a ≤ c • b := +by { rw [←smul_le_smul_iff_of_pos h, smul_inv_smul₀ h.ne'], apply_instance } -lemma lt_smul_iff_of_pos (hc : 0 < c) : a < c • b ↔ c⁻¹ • a < b := -calc a < c • b ↔ c • c⁻¹ • a < c • b : by rw [smul_inv_smul₀ hc.ne'] -... ↔ c⁻¹ • a < b : smul_lt_smul_iff_of_pos hc +lemma inv_smul_lt_iff (h : 0 < c) : c⁻¹ • a < b ↔ a < c • b := +by { rw [←smul_lt_smul_iff_of_pos h, smul_inv_smul₀ h.ne'], apply_instance } -lemma smul_le_iff_of_pos (hc : 0 < c) : c • a ≤ b ↔ a ≤ c⁻¹ • b := -calc c • a ≤ b ↔ c • a ≤ c • c⁻¹ • b : by rw [smul_inv_smul₀ hc.ne'] -... ↔ a ≤ c⁻¹ • b : smul_le_smul_iff_of_pos hc +lemma le_inv_smul_iff (h : 0 < c) : a ≤ c⁻¹ • b ↔ c • a ≤ b := +by { rw [←smul_le_smul_iff_of_pos h, smul_inv_smul₀ h.ne'], apply_instance } -lemma le_smul_iff_of_pos (hc : 0 < c) : a ≤ c • b ↔ c⁻¹ • a ≤ b := -calc a ≤ c • b ↔ c • c⁻¹ • a ≤ c • b : by rw [smul_inv_smul₀ hc.ne'] -... ↔ c⁻¹ • a ≤ b : smul_le_smul_iff_of_pos hc +lemma lt_inv_smul_iff (h : 0 < c) : a < c⁻¹ • b ↔ c • a < b := +by { rw [←smul_lt_smul_iff_of_pos h, smul_inv_smul₀ h.ne'], apply_instance } variables (M) @@ -264,7 +253,6 @@ variables {M} @[simp] lemma bdd_above_smul_iff_of_pos (hc : 0 < c) : bdd_above (c • s) ↔ bdd_above s := (order_iso.smul_left _ hc).bdd_above_image -end ordered_add_comm_group end linear_ordered_semifield namespace tactic diff --git a/src/linear_algebra/affine_space/ordered.lean b/src/linear_algebra/affine_space/ordered.lean index 3c3759bc516bb..048cefc28727e 100644 --- a/src/linear_algebra/affine_space/ordered.lean +++ b/src/linear_algebra/affine_space/ordered.lean @@ -203,7 +203,7 @@ begin vsub_eq_sub, vsub_eq_sub, vsub_eq_sub, vadd_eq_add, vadd_eq_add, smul_eq_mul, add_sub_cancel, smul_sub, smul_sub, smul_sub, sub_le_iff_le_add, mul_inv_rev, mul_smul, mul_smul, ←smul_sub, ←smul_sub, ←smul_add, smul_smul, - ← mul_inv_rev, smul_le_iff_of_pos (inv_pos.2 h), inv_inv, smul_smul, + ← mul_inv_rev, inv_smul_le_iff h, smul_smul, mul_inv_cancel_right₀ (right_ne_zero_of_mul h.ne'), smul_add, smul_inv_smul₀ (left_ne_zero_of_mul h.ne')], apply_instance @@ -236,11 +236,10 @@ begin rw [← line_map_apply_one_sub, ← line_map_apply_one_sub _ _ r], revert h, generalize : 1 - r = r', clear r, intro h, simp_rw [line_map_apply, slope, vsub_eq_sub, vadd_eq_add, smul_eq_mul], - rw [sub_add_eq_sub_sub_swap, sub_self, zero_sub, le_smul_iff_of_pos, inv_inv, smul_smul, - neg_mul_eq_mul_neg, neg_sub, mul_inv_cancel_right₀, le_sub, ← neg_sub (f b), smul_neg, - neg_add_eq_sub], + rw [sub_add_eq_sub_sub_swap, sub_self, zero_sub, neg_mul_eq_mul_neg, neg_sub, le_inv_smul_iff h, + smul_smul, mul_inv_cancel_right₀, le_sub, ← neg_sub (f b), smul_neg, neg_add_eq_sub], { exact right_ne_zero_of_mul h.ne' }, - { simpa [mul_sub] using h } + { apply_instance } end /-- Given `c = line_map a b r`, `c < b`, the point `(c, f c)` is non-strictly above the From cc967daad4a50323c2d9527b122ffe763777ad1d Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Fri, 30 Sep 2022 13:10:24 +0000 Subject: [PATCH 472/828] feat(algebra/hom/group): `simp` lemmas for applying generic morphism coercions (#16700) There are a bunch of random specific versions of these lemmas floating around, which can be made generic to apply to all `one_hom_class`/`mul_hom_class`/`monoid_hom_class` instances. Compare existing `ring_hom.coe_coe`. --- src/algebra/hom/group.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/algebra/hom/group.lean b/src/algebra/hom/group.lean index c67dbb5c7e05c..65e48f8fd5fa2 100644 --- a/src/algebra/hom/group.lean +++ b/src/algebra/hom/group.lean @@ -202,6 +202,9 @@ ne_of_apply_ne f $ ne_of_ne_of_eq hx (map_one f).symm instance [one_hom_class F M N] : has_coe_t F (one_hom M N) := ⟨λ f, { to_fun := f, map_one' := map_one f }⟩ +@[simp, to_additive] lemma one_hom.coe_coe [one_hom_class F M N] (f : F) : + ((f : one_hom M N) : M → N) = f := rfl + end one section mul @@ -246,6 +249,9 @@ mul_hom_class.map_mul f x y instance [mul_hom_class F M N] : has_coe_t F (M →ₙ* N) := ⟨λ f, { to_fun := f, map_mul' := map_mul f }⟩ +@[simp, to_additive] lemma mul_hom.coe_coe [mul_hom_class F M N] (f : F) : + ((f : mul_hom M N) : M → N) = f := rfl + end mul section mul_one @@ -289,6 +295,9 @@ instance monoid_hom.monoid_hom_class : monoid_hom_class (M →* N) M N := instance [monoid_hom_class F M N] : has_coe_t F (M →* N) := ⟨λ f, { to_fun := f, map_one' := map_one f, map_mul' := map_mul f }⟩ +@[simp, to_additive] lemma monoid_hom.coe_coe [monoid_hom_class F M N] (f : F) : + ((f : M →* N) : M → N) = f := rfl + @[to_additive] lemma map_mul_eq_one [monoid_hom_class F M N] (f : F) {a b : M} (h : a * b = 1) : f a * f b = 1 := @@ -402,6 +411,9 @@ instance monoid_with_zero_hom.monoid_with_zero_hom_class : instance [monoid_with_zero_hom_class F M N] : has_coe_t F (M →*₀ N) := ⟨λ f, { to_fun := f, map_one' := map_one f, map_zero' := map_zero f, map_mul' := map_mul f }⟩ +@[simp] lemma monoid_with_zero_hom.coe_coe [monoid_with_zero_hom_class F M N] (f : F) : + ((f : M →*₀ N) : M → N) = f := rfl + end mul_zero_one -- completely uninteresting lemmas about coercion to function, that all homs need From be05d53c17248523e2a52eac6d689a848898d6a7 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 30 Sep 2022 13:10:26 +0000 Subject: [PATCH 473/828] feat(algebra/free_monoid): add 2 lemmas (#16712) --- src/algebra/free_monoid.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/algebra/free_monoid.lean b/src/algebra/free_monoid.lean index 9141a95efce5c..cbe4d5d79e439 100644 --- a/src/algebra/free_monoid.lean +++ b/src/algebra/free_monoid.lean @@ -43,6 +43,10 @@ lemma one_def : (1 : free_monoid α) = [] := rfl lemma mul_def (xs ys : list α) : (xs * ys : free_monoid α) = (xs ++ ys : list α) := rfl +@[to_additive] +lemma prod_eq_join (xs : list (free_monoid α)) : xs.prod = xs.join := +by induction xs; simp [*, mul_def, list.join, one_def] + /-- Embeds an element of `α` into `free_monoid α` as a singleton list. -/ @[to_additive "Embeds an element of `α` into `free_add_monoid α` as a singleton list." ] def of (x : α) : free_monoid α := [x] @@ -54,6 +58,8 @@ lemma of_def (x : α) : of x = [x] := rfl lemma of_injective : function.injective (@of α) := λ a b, list.head_eq_of_cons_eq +@[to_additive] lemma of_mul_eq_cons (x : α) (l : free_monoid α) : of x * l = x :: l := rfl + /-- Recursor for `free_monoid` using `1` and `of x * xs` instead of `[]` and `x :: xs`. -/ @[elab_as_eliminator, to_additive "Recursor for `free_add_monoid` using `0` and `of x + xs` instead of `[]` and `x :: xs`."] From 6e5e3f6820d2340d144ca71ba741f7b599d03713 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Fri, 30 Sep 2022 18:24:32 +0000 Subject: [PATCH 474/828] feat(analysis/locally_convex): the topology of a locally convex space is generated by seminorms (#15035) This PR provides the proof that every locally convex space has a family of seminorms that induces the topology. This PR also adds a new simp-lemma `is_R_or_C.real_norm`, which calculates the norm of a real number `r` coerced into a `is_R_or_C` type as the norm of `r`. This made it necessary to change some proofs in a few places. Co-authored-by: Moritz Doll --- src/analysis/complex/basic.lean | 2 +- src/analysis/locally_convex/abs_convex.lean | 172 ++++++++++++++++++ src/analysis/locally_convex/basic.lean | 14 ++ .../locally_convex/with_seminorms.lean | 4 - src/analysis/normed_space/is_R_or_C.lean | 8 +- src/data/complex/is_R_or_C.lean | 6 +- .../algebra/module/locally_convex.lean | 1 - 7 files changed, 194 insertions(+), 13 deletions(-) create mode 100644 src/analysis/locally_convex/abs_convex.lean diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index 1e787adf61a56..c02a93d9ccbc3 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -112,7 +112,7 @@ by rw [nndist_comm, nndist_conj_self] @[simp] lemma comap_abs_nhds_zero : filter.comap abs (𝓝 0) = 𝓝 0 := comap_norm_nhds_zero -@[simp] lemma norm_real (r : ℝ) : ∥(r : ℂ)∥ = ∥r∥ := abs_of_real _ +lemma norm_real (r : ℝ) : ∥(r : ℂ)∥ = ∥r∥ := abs_of_real _ @[simp] lemma norm_rat (r : ℚ) : ∥(r : ℂ)∥ = |(r : ℝ)| := by { rw ← of_real_rat_cast, exact norm_real _ } diff --git a/src/analysis/locally_convex/abs_convex.lean b/src/analysis/locally_convex/abs_convex.lean new file mode 100644 index 0000000000000..d50eabc26d284 --- /dev/null +++ b/src/analysis/locally_convex/abs_convex.lean @@ -0,0 +1,172 @@ +/- +Copyright (c) 2022 Moritz Doll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Doll +-/ +import analysis.locally_convex.balanced_core_hull +import analysis.locally_convex.with_seminorms +import analysis.convex.gauge + +/-! +# Absolutely convex sets + +A set is called absolutely convex or disked if it is convex and balanced. +The importance of absolutely convex sets comes from the fact that every locally convex +topological vector space has a basis consisting of absolutely convex sets. + +## Main definitions + +* `gauge_seminorm_family`: the seminorm family induced by all open absolutely convex neighborhoods +of zero. + +## Main statements + +* `with_gauge_seminorm_family`: the topology of a locally convex space is induced by the family +`gauge_seminorm_family`. + +## Todo + +* Define the disked hull + +## Tags + +disks, convex, balanced +-/ + + +open normed_field set +open_locale big_operators nnreal pointwise topological_space + +variables {𝕜 E F G ι : Type*} + +section nontrivially_normed_field + +variables (𝕜 E) {s : set E} + +variables [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] +variables [module ℝ E] [smul_comm_class ℝ 𝕜 E] +variables [topological_space E] [locally_convex_space ℝ E] [has_continuous_smul 𝕜 E] + +lemma nhds_basis_abs_convex : (𝓝 (0 : E)).has_basis + (λ (s : set E), s ∈ 𝓝 (0 : E) ∧ balanced 𝕜 s ∧ convex ℝ s) id := +begin + refine (locally_convex_space.convex_basis_zero ℝ E).to_has_basis (λ s hs, _) + (λ s hs, ⟨s, ⟨hs.1, hs.2.2⟩, rfl.subset⟩), + refine ⟨convex_hull ℝ (balanced_core 𝕜 s), _, convex_hull_min (balanced_core_subset s) hs.2⟩, + refine ⟨filter.mem_of_superset (balanced_core_mem_nhds_zero hs.1) (subset_convex_hull ℝ _), _⟩, + refine ⟨balanced_convex_hull_of_balanced (balanced_core_balanced s), _⟩, + exact convex_convex_hull ℝ (balanced_core 𝕜 s), +end + +variables [has_continuous_smul ℝ E] [topological_add_group E] + +lemma nhds_basis_abs_convex_open : (𝓝 (0 : E)).has_basis + (λ (s : set E), (0 : E) ∈ s ∧ is_open s ∧ balanced 𝕜 s ∧ convex ℝ s) id := +begin + refine (nhds_basis_abs_convex 𝕜 E).to_has_basis _ _, + { rintros s ⟨hs_nhds, hs_balanced, hs_convex⟩, + refine ⟨interior s, _, interior_subset⟩, + exact ⟨mem_interior_iff_mem_nhds.mpr hs_nhds, is_open_interior, + hs_balanced.interior (mem_interior_iff_mem_nhds.mpr hs_nhds), hs_convex.interior⟩ }, + rintros s ⟨hs_zero, hs_open, hs_balanced, hs_convex⟩, + exact ⟨s, ⟨hs_open.mem_nhds hs_zero, hs_balanced, hs_convex⟩, rfl.subset⟩, +end + +end nontrivially_normed_field + +section absolutely_convex_sets + +variables [topological_space E] [add_comm_monoid E] [has_zero E] [semi_normed_ring 𝕜] +variables [has_smul 𝕜 E] [has_smul ℝ E] + +variables (𝕜 E) + +/-- The type of absolutely convex open sets. -/ +def abs_convex_open_sets := +{ s : set E // (0 : E) ∈ s ∧ is_open s ∧ balanced 𝕜 s ∧ convex ℝ s } + +instance abs_convex_open_sets.has_coe : has_coe (abs_convex_open_sets 𝕜 E) (set E) := ⟨subtype.val⟩ + +namespace abs_convex_open_sets + +variables {𝕜 E} + +lemma coe_zero_mem (s : abs_convex_open_sets 𝕜 E) : (0 : E) ∈ (s : set E) := s.2.1 + +lemma coe_is_open (s : abs_convex_open_sets 𝕜 E) : is_open (s : set E) := s.2.2.1 + +lemma coe_nhds (s : abs_convex_open_sets 𝕜 E) : (s : set E) ∈ 𝓝 (0 : E) := +s.coe_is_open.mem_nhds s.coe_zero_mem + +lemma coe_balanced (s : abs_convex_open_sets 𝕜 E) : balanced 𝕜 (s : set E) := s.2.2.2.1 + +lemma coe_convex (s : abs_convex_open_sets 𝕜 E) : convex ℝ (s : set E) := s.2.2.2.2 + +end abs_convex_open_sets + +instance : nonempty (abs_convex_open_sets 𝕜 E) := +begin + rw ←exists_true_iff_nonempty, + dunfold abs_convex_open_sets, + rw subtype.exists, + exact ⟨set.univ, ⟨mem_univ 0, is_open_univ, balanced_univ, convex_univ⟩, trivial⟩, +end + +end absolutely_convex_sets + +variables [is_R_or_C 𝕜] +variables [add_comm_group E] [topological_space E] +variables [module 𝕜 E] [module ℝ E] [is_scalar_tower ℝ 𝕜 E] +variables [has_continuous_smul ℝ E] + +variables (𝕜 E) + +/-- The family of seminorms defined by the gauges of absolute convex open sets. -/ +noncomputable +def gauge_seminorm_family : seminorm_family 𝕜 E (abs_convex_open_sets 𝕜 E) := +λ s, gauge_seminorm s.coe_balanced s.coe_convex (absorbent_nhds_zero s.coe_nhds) + +variables {𝕜 E} + +lemma gauge_seminorm_family_ball (s : abs_convex_open_sets 𝕜 E) : + (gauge_seminorm_family 𝕜 E s).ball 0 1 = (s : set E) := +begin + dunfold gauge_seminorm_family, + rw seminorm.ball_zero_eq, + simp_rw gauge_seminorm_to_fun, + exact gauge_lt_one_eq_self_of_open s.coe_convex s.coe_zero_mem s.coe_is_open, +end + +variables [topological_add_group E] [has_continuous_smul 𝕜 E] +variables [smul_comm_class ℝ 𝕜 E] [locally_convex_space ℝ E] + +/-- The topology of a locally convex space is induced by the gauge seminorm family. -/ +lemma with_gauge_seminorm_family : with_seminorms (gauge_seminorm_family 𝕜 E) := +begin + refine seminorm_family.with_seminorms_of_has_basis _ _, + refine filter.has_basis.to_has_basis (nhds_basis_abs_convex_open 𝕜 E) (λ s hs, _) (λ s hs, _), + { refine ⟨s, ⟨_, rfl.subset⟩⟩, + rw seminorm_family.basis_sets_iff, + refine ⟨{⟨s, hs⟩}, 1, one_pos, _⟩, + simp only [finset.sup_singleton], + rw gauge_seminorm_family_ball, + simp only [subtype.coe_mk] }, + refine ⟨s, ⟨_, rfl.subset⟩⟩, + rw seminorm_family.basis_sets_iff at hs, + rcases hs with ⟨t, r, hr, hs⟩, + rw seminorm.ball_finset_sup_eq_Inter _ _ _ hr at hs, + rw hs, + -- We have to show that the intersection contains zero, is open, balanced, and convex + refine ⟨mem_Inter₂.mpr (λ _ _, by simp [seminorm.mem_ball_zero, hr]), + is_open_bInter (to_finite _) (λ _ _, _), + balanced_Inter₂ (λ _ _, seminorm.balanced_ball_zero _ _), + convex_Inter₂ (λ _ _, seminorm.convex_ball _ _ _)⟩, + -- The only nontrivial part is to show that the ball is open + have hr' : r = ∥(r : 𝕜)∥ * 1 := by simp [abs_of_pos hr], + have hr'' : (r : 𝕜) ≠ 0 := by simp [ne_of_gt hr], + rw hr', + rw ←seminorm.smul_ball_zero (norm_pos_iff.mpr hr''), + refine is_open.smul₀ _ hr'', + rw gauge_seminorm_family_ball, + exact abs_convex_open_sets.coe_is_open _, +end diff --git a/src/analysis/locally_convex/basic.lean b/src/analysis/locally_convex/basic.lean index 098c89bb39681..0ef0716b1032a 100644 --- a/src/analysis/locally_convex/basic.lean +++ b/src/analysis/locally_convex/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jean Lo, Bhavik Mehta, Yaël Dillies -/ import analysis.convex.basic +import analysis.convex.hull import analysis.normed_space.lattice_ordered_group import analysis.normed_space.ordered @@ -342,6 +343,19 @@ end lemma absorbent.zero_mem (hs : absorbent 𝕜 s) : (0 : E) ∈ s := absorbs_zero_iff.1 $ absorbent_iff_forall_absorbs_singleton.1 hs _ +variables [module ℝ E] [smul_comm_class ℝ 𝕜 E] + +lemma balanced_convex_hull_of_balanced (hs : balanced 𝕜 s) : balanced 𝕜 (convex_hull ℝ s) := +begin + suffices : convex ℝ {x | ∀ a : 𝕜, ∥a∥ ≤ 1 → a • x ∈ convex_hull ℝ s}, + { rw balanced_iff_smul_mem at hs ⊢, + refine λ a ha x hx, convex_hull_min _ this hx a ha, + exact λ y hy a ha, subset_convex_hull ℝ s (hs ha hy) }, + intros x hx y hy u v hu hv huv a ha, + simp only [smul_add, ← smul_comm], + exact convex_convex_hull ℝ s (hx a ha) (hy a ha) hu hv huv +end + end nontrivially_normed_field section real diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 6afbaf2ddc9d2..709452696abbe 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -23,10 +23,6 @@ bounded by a finite number of seminorms in `E`. * `seminorm_family.to_locally_convex_space`: A space equipped with a family of seminorms is locally convex. -## TODO - -Show that for any locally convex space there exist seminorms that induce the topology. - ## Tags seminorm, locally convex diff --git a/src/analysis/normed_space/is_R_or_C.lean b/src/analysis/normed_space/is_R_or_C.lean index 5579d501b96b7..f8b89986bbe2c 100644 --- a/src/analysis/normed_space/is_R_or_C.lean +++ b/src/analysis/normed_space/is_R_or_C.lean @@ -28,11 +28,11 @@ This file exists mainly to avoid importing `is_R_or_C` in the main normed space open metric -@[simp, is_R_or_C_simps] lemma is_R_or_C.norm_coe_norm {𝕜 : Type*} [is_R_or_C 𝕜] - {E : Type*} [normed_add_comm_group E] {z : E} : ∥(∥z∥ : 𝕜)∥ = ∥z∥ := -by { unfold_coes, simp only [norm_algebra_map', ring_hom.to_fun_eq_coe, norm_norm], } +variables {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [normed_add_comm_group E] -variables {𝕜 : Type*} [is_R_or_C 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +lemma is_R_or_C.norm_coe_norm {z : E} : ∥(∥z∥ : 𝕜)∥ = ∥z∥ := by simp + +variables [normed_space 𝕜 E] /-- Lemma to normalize a vector in a normed space `E` over either `ℂ` or `ℝ` to unit length. -/ @[simp] lemma norm_smul_inv_norm {x : E} (hx : x ≠ 0) : ∥(∥x∥⁻¹ : 𝕜) • x∥ = 1 := diff --git a/src/data/complex/is_R_or_C.lean b/src/data/complex/is_R_or_C.lean index 3c3a368b4f371..d98c8f09f741b 100644 --- a/src/data/complex/is_R_or_C.lean +++ b/src/data/complex/is_R_or_C.lean @@ -185,6 +185,9 @@ by simp only [add_zero, of_real_im, zero_mul, of_real_re, mul_im] @[is_R_or_C_simps] lemma smul_im : ∀ (r : ℝ) (z : K), im (r • z) = r * (im z) := λ r z, by { rw algebra.smul_def, apply of_real_mul_im } +@[simp, is_R_or_C_simps] lemma norm_real (r : ℝ) : ∥(r : K)∥ = ∥r∥ := +by rw [is_R_or_C.of_real_alg, norm_smul, norm_one, mul_one] + /-! ### The imaginary unit, `I` -/ /-- The imaginary unit. -/ @@ -266,9 +269,6 @@ def norm_sq : K →*₀ ℝ := lemma norm_sq_eq_def {z : K} : ∥z∥^2 = (re z) * (re z) + (im z) * (im z) := norm_sq_eq_def_ax z lemma norm_sq_eq_def' (z : K) : norm_sq z = ∥z∥^2 := by { rw norm_sq_eq_def, refl } -@[simp, is_R_or_C_simps] lemma norm_sq_of_real (r : ℝ) : ∥(r : K)∥^2 = r * r := -by simp only [norm_sq_eq_def, add_zero, mul_zero] with is_R_or_C_simps - @[is_R_or_C_simps] lemma norm_sq_zero : norm_sq (0 : K) = 0 := norm_sq.map_zero @[is_R_or_C_simps] lemma norm_sq_one : norm_sq (1 : K) = 1 := norm_sq.map_one diff --git a/src/topology/algebra/module/locally_convex.lean b/src/topology/algebra/module/locally_convex.lean index 3f2133b60152f..f7f14efe4ed83 100644 --- a/src/topology/algebra/module/locally_convex.lean +++ b/src/topology/algebra/module/locally_convex.lean @@ -24,7 +24,6 @@ In a module, this is equivalent to `0` satisfying such properties. - define a structure `locally_convex_filter_basis`, extending `module_filter_basis`, for filter bases generating a locally convex topology -- show that any locally convex topology is generated by a family of seminorms -/ From 98c61b1ed9609277c5e57362379e32b75d9c15f0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 30 Sep 2022 18:24:34 +0000 Subject: [PATCH 475/828] feat(order/hom/heyting): Heyting homomorphisms (#15308) Define the type of Heyting homomorphisms, maps between Heyting algebras that preserve Heyting implication. --- src/algebra/ring/boolean_ring.lean | 4 +- src/data/fun_like/equiv.lean | 16 ++ src/order/heyting/basic.lean | 2 +- src/order/heyting/hom.lean | 408 ++++++++++++++++++++++++++++ src/order/hom/complete_lattice.lean | 4 + src/order/hom/lattice.lean | 12 +- 6 files changed, 439 insertions(+), 7 deletions(-) create mode 100644 src/order/heyting/hom.lean diff --git a/src/algebra/ring/boolean_ring.lean b/src/algebra/ring/boolean_ring.lean index 48516c021d212..eba5bcbbd83e2 100644 --- a/src/algebra/ring/boolean_ring.lean +++ b/src/algebra/ring/boolean_ring.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Bryan Gin-ge Chen, Yaël Dillies -/ import algebra.punit_instances -import order.hom.lattice +import order.heyting.hom import tactic.abel import tactic.ring @@ -359,7 +359,7 @@ from `α` to `β` considered as Boolean rings. -/ { to_fun := to_boolring ∘ f ∘ of_boolring, map_zero' := f.map_bot', map_one' := f.map_top', - map_add' := map_symm_diff f, + map_add' := map_symm_diff' f, map_mul' := f.map_inf' } @[simp] lemma bounded_lattice_hom.as_boolring_id : diff --git a/src/data/fun_like/equiv.lean b/src/data/fun_like/equiv.lean index 041968a82801b..7327b35a5416d 100644 --- a/src/data/fun_like/equiv.lean +++ b/src/data/fun_like/equiv.lean @@ -175,6 +175,22 @@ function.injective.of_comp_iff' f (equiv_like.bijective e) function.bijective (f ∘ e) ↔ function.bijective f := (equiv_like.bijective e).of_comp_iff f +/-- This lemma is only supposed to be used in the generic context, when working with instances +of classes extending `equiv_like`. +For concrete isomorphism types such as `equiv`, you should use `equiv.symm_apply_apply` +or its equivalent. + +TODO: define a generic form of `equiv.symm`. -/ +@[simp] lemma inv_apply_apply (e : E) (a : α) : equiv_like.inv e (e a) = a := left_inv _ _ + +/-- This lemma is only supposed to be used in the generic context, when working with instances +of classes extending `equiv_like`. +For concrete isomorphism types such as `equiv`, you should use `equiv.apply_symm_apply` +or its equivalent. + +TODO: define a generic form of `equiv.symm`. -/ +@[simp] lemma apply_inv_apply (e : E) (b : β) : e (equiv_like.inv e b) = b := right_inv _ _ + omit iE include iF diff --git a/src/order/heyting/basic.lean b/src/order/heyting/basic.lean index 2ff8c76ec2a4f..9bb5c83bb0e39 100644 --- a/src/order/heyting/basic.lean +++ b/src/order/heyting/basic.lean @@ -61,7 +61,7 @@ variables {ι α β : Type*} /-- Syntax typeclass for Heyting negation `¬`. -The difference between `has_hnot` and `has_compl` is that the former belongs to Heyting algebras, +The difference between `has_compl` and `has_hnot` is that the former belongs to Heyting algebras, while the latter belongs to co-Heyting algebras. They are both pseudo-complements, but `compl` underestimates while `hnot` overestimates. In boolean algebras, they are equal. See `hnot_eq_compl`. -/ diff --git a/src/order/heyting/hom.lean b/src/order/heyting/hom.lean new file mode 100644 index 0000000000000..e42f4892cd38c --- /dev/null +++ b/src/order/heyting/hom.lean @@ -0,0 +1,408 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.hom.lattice + +/-! +# Heyting algebra morphisms + +A Heyting homomorphism between two Heyting algebras is a bounded lattice homomorphism that preserves +Heyting implication. + +We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to +be satisfied by itself and all stricter types. + +## Types of morphisms + +* `heyting_hom`: Heyting homomorphisms. +* `coheyting_hom`: Co-Heyting homomorphisms. +* `biheyting_hom`: Bi-Heyting homomorphisms. + +## Typeclasses + +* `heyting_hom_class` +* `coheyting_hom_class` +* `biheyting_hom_class` +-/ + +open function + +variables {F α β γ δ : Type*} + +/-- The type of Heyting homomorphisms from `α` to `β`. Bounded lattice homomorphisms that preserve +Heyting implication. -/ +@[protect_proj] +structure heyting_hom (α β : Type*) [heyting_algebra α] [heyting_algebra β] + extends lattice_hom α β := +(map_bot' : to_fun ⊥ = ⊥) +(map_himp' : ∀ a b, to_fun (a ⇨ b) = to_fun a ⇨ to_fun b) + +/-- The type of co-Heyting homomorphisms from `α` to `β`. Bounded lattice homomorphisms that +preserve difference. -/ +@[protect_proj] +structure coheyting_hom (α β : Type*) [coheyting_algebra α] [coheyting_algebra β] + extends lattice_hom α β := +(map_top' : to_fun ⊤ = ⊤) +(map_sdiff' : ∀ a b, to_fun (a \ b) = to_fun a \ to_fun b) + +/-- The type of bi-Heyting homomorphisms from `α` to `β`. Bounded lattice homomorphisms that +preserve Heyting implication and difference. -/ +@[protect_proj] +structure biheyting_hom (α β : Type*) [biheyting_algebra α] [biheyting_algebra β] + extends lattice_hom α β := +(map_himp' : ∀ a b, to_fun (a ⇨ b) = to_fun a ⇨ to_fun b) +(map_sdiff' : ∀ a b, to_fun (a \ b) = to_fun a \ to_fun b) + +/-- `heyting_hom_class F α β` states that `F` is a type of Heyting homomorphisms. + +You should extend this class when you extend `heyting_hom`. -/ +class heyting_hom_class (F : Type*) (α β : out_param $ Type*) [heyting_algebra α] + [heyting_algebra β] extends lattice_hom_class F α β := +(map_bot (f : F) : f ⊥ = ⊥) +(map_himp (f : F) : ∀ a b, f (a ⇨ b) = f a ⇨ f b) + +/-- `coheyting_hom_class F α β` states that `F` is a type of co-Heyting homomorphisms. + +You should extend this class when you extend `coheyting_hom`. -/ +class coheyting_hom_class (F : Type*) (α β : out_param $ Type*) [coheyting_algebra α] + [coheyting_algebra β] extends lattice_hom_class F α β := +(map_top (f : F) : f ⊤ = ⊤) +(map_sdiff (f : F) : ∀ a b, f (a \ b) = f a \ f b) + +/-- `biheyting_hom_class F α β` states that `F` is a type of bi-Heyting homomorphisms. + +You should extend this class when you extend `biheyting_hom`. -/ +class biheyting_hom_class (F : Type*) (α β : out_param $ Type*) [biheyting_algebra α] + [biheyting_algebra β] extends lattice_hom_class F α β := +(map_himp (f : F) : ∀ a b, f (a ⇨ b) = f a ⇨ f b) +(map_sdiff (f : F) : ∀ a b, f (a \ b) = f a \ f b) + +export heyting_hom_class (map_himp) +export coheyting_hom_class (map_sdiff) + +attribute [simp] map_himp map_sdiff + +@[priority 100] -- See note [lower instance priority] +instance heyting_hom_class.to_bounded_lattice_hom_class [heyting_algebra α] [heyting_algebra β] + [heyting_hom_class F α β] : bounded_lattice_hom_class F α β := +{ map_top := λ f, by rw [←@himp_self α _ ⊥, ←himp_self, map_himp], + ..‹heyting_hom_class F α β› } + +@[priority 100] -- See note [lower instance priority] +instance coheyting_hom_class.to_bounded_lattice_hom_class [coheyting_algebra α] + [coheyting_algebra β] [coheyting_hom_class F α β] : bounded_lattice_hom_class F α β := +{ map_bot := λ f, by rw [←@sdiff_self α _ ⊤, ←sdiff_self, map_sdiff], + ..‹coheyting_hom_class F α β› } + +@[priority 100] -- See note [lower instance priority] +instance biheyting_hom_class.to_heyting_hom_class [biheyting_algebra α] [biheyting_algebra β] + [biheyting_hom_class F α β] : + heyting_hom_class F α β := +{ map_bot := λ f, by rw [←@sdiff_self α _ ⊤, ←sdiff_self, biheyting_hom_class.map_sdiff], + ..‹biheyting_hom_class F α β› } + +@[priority 100] -- See note [lower instance priority] +instance biheyting_hom_class.to_coheyting_hom_class [biheyting_algebra α] [biheyting_algebra β] + [biheyting_hom_class F α β] : + coheyting_hom_class F α β := +{ map_top := λ f, by rw [←@himp_self α _ ⊥, ←himp_self, map_himp], + ..‹biheyting_hom_class F α β› } + +@[priority 100] -- See note [lower instance priority] +instance order_iso_class.to_heyting_hom_class [heyting_algebra α] [heyting_algebra β] + [order_iso_class F α β] : + heyting_hom_class F α β := +{ map_himp := λ f a b, eq_of_forall_le_iff $ λ c, + by { simp only [←map_inv_le_iff, le_himp_iff], rw ←order_iso_class.map_le_map_iff f, simp }, + ..order_iso_class.to_bounded_lattice_hom_class } + +@[priority 100] -- See note [lower instance priority] +instance order_iso_class.to_coheyting_hom_class [coheyting_algebra α] [coheyting_algebra β] + [order_iso_class F α β] : + coheyting_hom_class F α β := +{ map_sdiff := λ f a b, eq_of_forall_ge_iff $ λ c, + by { simp only [←le_map_inv_iff, sdiff_le_iff], rw ←order_iso_class.map_le_map_iff f, simp }, + ..order_iso_class.to_bounded_lattice_hom_class } + +@[priority 100] -- See note [lower instance priority] +instance order_iso_class.to_biheyting_hom_class [biheyting_algebra α] [biheyting_algebra β] + [order_iso_class F α β] : + biheyting_hom_class F α β := +{ map_himp := λ f a b, eq_of_forall_le_iff $ λ c, + by { simp only [←map_inv_le_iff, le_himp_iff], rw ←order_iso_class.map_le_map_iff f, simp }, + map_sdiff := λ f a b, eq_of_forall_ge_iff $ λ c, + by { simp only [←le_map_inv_iff, sdiff_le_iff], rw ←order_iso_class.map_le_map_iff f, simp }, + ..order_iso_class.to_lattice_hom_class } + +/-- This can't be an instance because of typeclass loops. -/ +@[reducible] -- See note [reducible non instances] +def bounded_lattice_hom_class.to_biheyting_hom_class [boolean_algebra α] [boolean_algebra β] + [bounded_lattice_hom_class F α β] : + biheyting_hom_class F α β := +{ map_himp := λ f a b, by rw [himp_eq, himp_eq, map_sup, (is_compl_compl.map _).compl_eq], + map_sdiff := λ f a b, by rw [sdiff_eq, sdiff_eq, map_inf, (is_compl_compl.map _).compl_eq], + ..‹bounded_lattice_hom_class F α β› } + +section heyting_algebra +variables [heyting_algebra α] [heyting_algebra β] [heyting_hom_class F α β] (f : F) +include β + +@[simp] lemma map_compl (a : α) : f aᶜ = (f a)ᶜ := by rw [←himp_bot, ←himp_bot, map_himp, map_bot] + +@[simp] lemma map_bihimp (a b : α) : f (a ⇔ b) = f a ⇔ f b := +by simp_rw [bihimp, map_inf, map_himp] + +-- TODO: `map_bihimp` + +end heyting_algebra + +section coheyting_algebra +variables [coheyting_algebra α] [coheyting_algebra β] [coheyting_hom_class F α β] (f : F) +include β + +@[simp] lemma map_hnot (a : α) : f ¬a = ¬f a := +by rw [←top_sdiff', ←top_sdiff', map_sdiff, map_top] + +@[simp] lemma map_symm_diff (a b : α) : f (a ∆ b) = f a ∆ f b := +by simp_rw [symm_diff, map_sup, map_sdiff] + +end coheyting_algebra + +instance [heyting_algebra α] [heyting_algebra β] [heyting_hom_class F α β] : + has_coe_t F (heyting_hom α β) := +⟨λ f, { to_fun := f, + map_sup' := map_sup f, + map_inf' := map_inf f, + map_bot' := map_bot f, + map_himp' := map_himp f }⟩ + +instance [coheyting_algebra α] [coheyting_algebra β] [coheyting_hom_class F α β] : + has_coe_t F (coheyting_hom α β) := +⟨λ f, { to_fun := f, + map_sup' := map_sup f, + map_inf' := map_inf f, + map_top' := map_top f, + map_sdiff' := map_sdiff f }⟩ + +instance [biheyting_algebra α] [biheyting_algebra β] [biheyting_hom_class F α β] : + has_coe_t F (biheyting_hom α β) := +⟨λ f, { to_fun := f, + map_sup' := map_sup f, + map_inf' := map_inf f, + map_himp' := map_himp f, + map_sdiff' := map_sdiff f }⟩ + +namespace heyting_hom +variables [heyting_algebra α] [heyting_algebra β] [heyting_algebra γ] [heyting_algebra δ] + +instance : heyting_hom_class (heyting_hom α β) α β := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by obtain ⟨⟨⟨_, _⟩, _⟩, _⟩ := f; obtain ⟨⟨⟨_, _⟩, _⟩, _⟩ := g; congr', + map_sup := λ f, f.map_sup', + map_inf := λ f, f.map_inf', + map_bot := λ f, f.map_bot', + map_himp := heyting_hom.map_himp' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun` +directly. -/ +instance : has_coe_to_fun (heyting_hom α β) (λ _, α → β) := fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe {f : heyting_hom α β} : f.to_fun = (f : α → β) := rfl + +@[ext] lemma ext {f g : heyting_hom α β} (h : ∀ a, f a = g a) : f = g := fun_like.ext f g h + +/-- Copy of a `heyting_hom` with a new `to_fun` equal to the old one. Useful to fix definitional +equalities. -/ +protected def copy (f : heyting_hom α β) (f' : α → β) (h : f' = f) : heyting_hom α β := +{ to_fun := f', + map_sup' := by simpa only [h] using map_sup f, + map_inf' := by simpa only [h] using map_inf f, + map_bot' := by simpa only [h] using map_bot f, + map_himp' := by simpa only [h] using map_himp f } + +variables (α) + +/-- `id` as a `heyting_hom`. -/ +protected def id : heyting_hom α α := +{ to_lattice_hom := lattice_hom.id _, + map_himp' := λ a b, rfl, + ..bot_hom.id _ } + +@[simp] lemma coe_id : ⇑(heyting_hom.id α) = id := rfl + +variables {α} + +@[simp] lemma id_apply (a : α) : heyting_hom.id α a = a := rfl + +instance : inhabited (heyting_hom α α) := ⟨heyting_hom.id _⟩ + +instance : partial_order (heyting_hom α β) := partial_order.lift _ fun_like.coe_injective + +/-- Composition of `heyting_hom`s as a `heyting_hom`. -/ +def comp (f : heyting_hom β γ) (g : heyting_hom α β) : heyting_hom α γ := +{ to_fun := f ∘ g, + map_bot' := by simp, + map_himp' := λ a b, by simp, + ..f.to_lattice_hom.comp g.to_lattice_hom } + +variables {f f₁ f₂ : heyting_hom α β} {g g₁ g₂ : heyting_hom β γ} + +@[simp] lemma coe_comp (f : heyting_hom β γ) (g : heyting_hom α β) : ⇑(f.comp g) = f ∘ g := rfl +@[simp] lemma comp_apply (f : heyting_hom β γ) (g : heyting_hom α β) (a : α) : + f.comp g a = f (g a) := rfl +@[simp] lemma comp_assoc (f : heyting_hom γ δ) (g : heyting_hom β γ) (h : heyting_hom α β) : + (f.comp g).comp h = f.comp (g.comp h) := rfl +@[simp] lemma comp_id (f : heyting_hom α β) : f.comp (heyting_hom.id α) = f := ext $ λ a, rfl +@[simp] lemma id_comp (f : heyting_hom α β) : (heyting_hom.id β).comp f = f := ext $ λ a, rfl + +lemma cancel_right (hf : surjective f) : g₁.comp f = g₂.comp f ↔ g₁ = g₂ := +⟨λ h, ext $ hf.forall.2 $ fun_like.ext_iff.1 h, congr_arg _⟩ + +lemma cancel_left (hg : injective g) : g.comp f₁ = g.comp f₂ ↔ f₁ = f₂ := +⟨λ h, heyting_hom.ext $ λ a, hg $ by rw [←comp_apply, h, comp_apply], congr_arg _⟩ + +end heyting_hom + +namespace coheyting_hom +variables [coheyting_algebra α] [coheyting_algebra β] [coheyting_algebra γ] [coheyting_algebra δ] + +instance : coheyting_hom_class (coheyting_hom α β) α β := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by obtain ⟨⟨⟨_, _⟩, _⟩, _⟩ := f; obtain ⟨⟨⟨_, _⟩, _⟩, _⟩ := g; congr', + map_sup := λ f, f.map_sup', + map_inf := λ f, f.map_inf', + map_top := λ f, f.map_top', + map_sdiff := coheyting_hom.map_sdiff' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun` +directly. -/ +instance : has_coe_to_fun (coheyting_hom α β) (λ _, α → β) := fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe {f : coheyting_hom α β} : f.to_fun = (f : α → β) := rfl + +@[ext] lemma ext {f g : coheyting_hom α β} (h : ∀ a, f a = g a) : f = g := fun_like.ext f g h + +/-- Copy of a `coheyting_hom` with a new `to_fun` equal to the old one. Useful to fix definitional +equalities. -/ +protected def copy (f : coheyting_hom α β) (f' : α → β) (h : f' = f) : coheyting_hom α β := +{ to_fun := f', + map_sup' := by simpa only [h] using map_sup f, + map_inf' := by simpa only [h] using map_inf f, + map_top' := by simpa only [h] using map_top f, + map_sdiff' := by simpa only [h] using map_sdiff f } + +variables (α) + +/-- `id` as a `coheyting_hom`. -/ +protected def id : coheyting_hom α α := +{ to_lattice_hom := lattice_hom.id _, + map_sdiff' := λ a b, rfl, + ..top_hom.id _ } + +@[simp] lemma coe_id : ⇑(coheyting_hom.id α) = id := rfl + +variables {α} + +@[simp] lemma id_apply (a : α) : coheyting_hom.id α a = a := rfl + +instance : inhabited (coheyting_hom α α) := ⟨coheyting_hom.id _⟩ + +instance : partial_order (coheyting_hom α β) := partial_order.lift _ fun_like.coe_injective + +/-- Composition of `coheyting_hom`s as a `coheyting_hom`. -/ +def comp (f : coheyting_hom β γ) (g : coheyting_hom α β) : coheyting_hom α γ := +{ to_fun := f ∘ g, + map_top' := by simp, + map_sdiff' := λ a b, by simp, + ..f.to_lattice_hom.comp g.to_lattice_hom } + +variables {f f₁ f₂ : coheyting_hom α β} {g g₁ g₂ : coheyting_hom β γ} + +@[simp] lemma coe_comp (f : coheyting_hom β γ) (g : coheyting_hom α β) : ⇑(f.comp g) = f ∘ g := rfl +@[simp] lemma comp_apply (f : coheyting_hom β γ) (g : coheyting_hom α β) (a : α) : + f.comp g a = f (g a) := rfl +@[simp] lemma comp_assoc (f : coheyting_hom γ δ) (g : coheyting_hom β γ) (h : coheyting_hom α β) : + (f.comp g).comp h = f.comp (g.comp h) := rfl +@[simp] lemma comp_id (f : coheyting_hom α β) : f.comp (coheyting_hom.id α) = f := ext $ λ a, rfl +@[simp] lemma id_comp (f : coheyting_hom α β) : (coheyting_hom.id β).comp f = f := ext $ λ a, rfl + +lemma cancel_right (hf : surjective f) : g₁.comp f = g₂.comp f ↔ g₁ = g₂ := +⟨λ h, ext $ hf.forall.2 $ fun_like.ext_iff.1 h, congr_arg _⟩ + +lemma cancel_left (hg : injective g) : g.comp f₁ = g.comp f₂ ↔ f₁ = f₂ := +⟨λ h, coheyting_hom.ext $ λ a, hg $ by rw [←comp_apply, h, comp_apply], congr_arg _⟩ + +end coheyting_hom + + +namespace biheyting_hom +variables [biheyting_algebra α] [biheyting_algebra β] [biheyting_algebra γ] [biheyting_algebra δ] + +instance : biheyting_hom_class (biheyting_hom α β) α β := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by obtain ⟨⟨⟨_, _⟩, _⟩, _⟩ := f; obtain ⟨⟨⟨_, _⟩, _⟩, _⟩ := g; congr', + map_sup := λ f, f.map_sup', + map_inf := λ f, f.map_inf', + map_himp := λ f, f.map_himp', + map_sdiff := λ f, f.map_sdiff' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun` +directly. -/ +instance : has_coe_to_fun (biheyting_hom α β) (λ _, α → β) := fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe {f : biheyting_hom α β} : f.to_fun = (f : α → β) := rfl + +@[ext] lemma ext {f g : biheyting_hom α β} (h : ∀ a, f a = g a) : f = g := fun_like.ext f g h + +/-- Copy of a `biheyting_hom` with a new `to_fun` equal to the old one. Useful to fix definitional +equalities. -/ +protected def copy (f : biheyting_hom α β) (f' : α → β) (h : f' = f) : biheyting_hom α β := +{ to_fun := f', + map_sup' := by simpa only [h] using map_sup f, + map_inf' := by simpa only [h] using map_inf f, + map_himp' := by simpa only [h] using map_himp f, + map_sdiff' := by simpa only [h] using map_sdiff f } + +variables (α) + +/-- `id` as a `biheyting_hom`. -/ +protected def id : biheyting_hom α α := +{ to_lattice_hom := lattice_hom.id _, + ..heyting_hom.id _, ..coheyting_hom.id _ } + +@[simp] lemma coe_id : ⇑(biheyting_hom.id α) = id := rfl + +variables {α} + +@[simp] lemma id_apply (a : α) : biheyting_hom.id α a = a := rfl + +instance : inhabited (biheyting_hom α α) := ⟨biheyting_hom.id _⟩ + +instance : partial_order (biheyting_hom α β) := partial_order.lift _ fun_like.coe_injective + +/-- Composition of `biheyting_hom`s as a `biheyting_hom`. -/ +def comp (f : biheyting_hom β γ) (g : biheyting_hom α β) : biheyting_hom α γ := +{ to_fun := f ∘ g, + map_himp' := λ a b, by simp, + map_sdiff' := λ a b, by simp, + ..f.to_lattice_hom.comp g.to_lattice_hom } + +variables {f f₁ f₂ : biheyting_hom α β} {g g₁ g₂ : biheyting_hom β γ} + +@[simp] lemma coe_comp (f : biheyting_hom β γ) (g : biheyting_hom α β) : ⇑(f.comp g) = f ∘ g := rfl +@[simp] lemma comp_apply (f : biheyting_hom β γ) (g : biheyting_hom α β) (a : α) : + f.comp g a = f (g a) := rfl +@[simp] lemma comp_assoc (f : biheyting_hom γ δ) (g : biheyting_hom β γ) (h : biheyting_hom α β) : + (f.comp g).comp h = f.comp (g.comp h) := rfl +@[simp] lemma comp_id (f : biheyting_hom α β) : f.comp (biheyting_hom.id α) = f := ext $ λ a, rfl +@[simp] lemma id_comp (f : biheyting_hom α β) : (biheyting_hom.id β).comp f = f := ext $ λ a, rfl + +lemma cancel_right (hf : surjective f) : g₁.comp f = g₂.comp f ↔ g₁ = g₂ := +⟨λ h, ext $ hf.forall.2 $ fun_like.ext_iff.1 h, congr_arg _⟩ + +lemma cancel_left (hg : injective g) : g.comp f₁ = g.comp f₂ ↔ f₁ = f₂ := +⟨λ h, biheyting_hom.ext $ λ a, hg $ by rw [←comp_apply, h, comp_apply], congr_arg _⟩ + +end biheyting_hom diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index 596d043ba064f..dc5e26e305dc7 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -31,6 +31,10 @@ be satisfied by itself and all stricter types. ## Concrete homs * `complete_lattice.set_preimage`: `set.preimage` as a complete lattice homomorphism. + +## TODO + +Frame homs are Heyting homs. -/ open function order_dual set diff --git a/src/order/hom/lattice.lean b/src/order/hom/lattice.lean index 2843cd4bbb8d8..42fce255b6019 100644 --- a/src/order/hom/lattice.lean +++ b/src/order/hom/lattice.lean @@ -234,12 +234,16 @@ section boolean_algebra variables [boolean_algebra α] [boolean_algebra β] [bounded_lattice_hom_class F α β] (f : F) include β -lemma map_compl (a : α) : f aᶜ = (f a)ᶜ := (is_compl_compl.map _).compl_eq.symm +/-- Special case of `map_compl` for boolean algebras. -/ +lemma map_compl' (a : α) : f aᶜ = (f a)ᶜ := (is_compl_compl.map _).compl_eq.symm -lemma map_sdiff (a b : α) : f (a \ b) = f a \ f b := by rw [sdiff_eq, sdiff_eq, map_inf, map_compl] +/-- Special case of `map_sdiff` for boolean algebras. -/ +lemma map_sdiff' (a b : α) : f (a \ b) = f a \ f b := +by rw [sdiff_eq, sdiff_eq, map_inf, map_compl'] -lemma map_symm_diff (a b : α) : f (a ∆ b) = f a ∆ f b := -by rw [symm_diff, symm_diff, map_sup, map_sdiff, map_sdiff] +/-- Special case of `map_symm_diff` for boolean algebras. -/ +lemma map_symm_diff' (a b : α) : f (a ∆ b) = f a ∆ f b := +by rw [symm_diff, symm_diff, map_sup, map_sdiff', map_sdiff'] end boolean_algebra From 73afba48667a1b774b52688604f86ac25c7042be Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 30 Sep 2022 18:24:35 +0000 Subject: [PATCH 476/828] refactor(analysis/normed/group/basic): Replace `normed_add_comm_group.core` with group norms (#16238) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Delete `seminormed_add_comm_group.core`/`normed_add_comm_group.core` in favor of `add_group_norm`/`group_norm`. The former are unbundled prop versions of the latter, specialized to `∥ ∥`. --- src/algebra/order/hom/basic.lean | 9 +++ src/analysis/complex/basic.lean | 9 +-- src/analysis/inner_product_space/basic.lean | 29 +++----- src/analysis/normed/group/basic.lean | 68 ++++++------------- src/analysis/normed/group/hom.lean | 13 +++- .../normed_space/continuous_affine_map.lean | 24 +++---- src/analysis/normed_space/lp_space.lean | 12 ++-- src/analysis/normed_space/multilinear.lean | 20 +++--- src/analysis/normed_space/operator_norm.lean | 17 +++-- src/measure_theory/function/lp_space.lean | 13 ++-- 10 files changed, 103 insertions(+), 111 deletions(-) diff --git a/src/algebra/order/hom/basic.lean b/src/algebra/order/hom/basic.lean index 8e86895e0e375..c21ad5adfbbb5 100644 --- a/src/algebra/order/hom/basic.lean +++ b/src/algebra/order/hom/basic.lean @@ -56,6 +56,15 @@ export mul_le_add_hom_class (map_mul_le_add) attribute [simp] map_nonneg +@[to_additive] lemma le_map_mul_map_div [group α] [comm_semigroup β] [has_le β] + [submultiplicative_hom_class F α β] (f : F) (a b : α) : f a ≤ f b * f (a / b) := +by simpa only [mul_comm, div_mul_cancel'] using map_mul_le_mul f (a / b) b + +@[to_additive] +lemma le_map_div_mul_map_div [group α] [comm_semigroup β] [has_le β] + [submultiplicative_hom_class F α β] (f : F) (a b c: α) : f (a / c) ≤ f (a / b) * f (b / c) := +by simpa only [div_mul_div_cancel'] using map_mul_le_mul f (a / b) (b / c) + @[to_additive] lemma le_map_add_map_div [group α] [add_comm_semigroup β] [has_le β] [mul_le_add_hom_class F α β] (f : F) (a b : α) : f a ≤ f b + f (a / b) := by simpa only [add_comm, div_mul_cancel'] using map_mul_le_add f (a / b) b diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index c02a93d9ccbc3..91a1ad9d28ac1 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -39,10 +39,11 @@ instance : has_norm ℂ := ⟨abs⟩ @[simp] lemma norm_eq_abs (z : ℂ) : ∥z∥ = abs z := rfl instance : normed_add_comm_group ℂ := -normed_add_comm_group.of_core ℂ -{ norm_eq_zero_iff := λ x, abs.eq_zero, - triangle := abs.add_le, - norm_neg := abs.map_neg } +add_group_norm.to_normed_add_comm_group +{ map_zero' := map_zero abs, + neg' := abs.map_neg, + eq_zero_of_map_eq_zero' := λ _, abs.eq_zero.1, + ..abs } instance : normed_field ℂ := { norm := abs, diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index b0375bc8dbd12..6311244e8f4a2 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -325,22 +325,11 @@ end /-- Normed group structure constructed from an `inner_product_space.core` structure -/ def to_normed_add_comm_group : normed_add_comm_group F := -normed_add_comm_group.of_core F -{ norm_eq_zero_iff := assume x, - begin - split, - { intro H, - change sqrt (re ⟪x, x⟫) = 0 at H, - rw [sqrt_eq_zero inner_self_nonneg] at H, - apply (inner_self_eq_zero : ⟪x, x⟫ = 0 ↔ x = 0).mp, - rw ext_iff, - exact ⟨by simp [H], by simp [inner_self_im_zero]⟩ }, - { rintro rfl, - change sqrt (re ⟪0, 0⟫) = 0, - simp only [sqrt_zero, inner_zero_right, add_monoid_hom.map_zero] } - end, - triangle := assume x y, - begin +add_group_norm.to_normed_add_comm_group +{ to_fun := λ x, sqrt (re ⟪x, x⟫), + map_zero' := by simp only [sqrt_zero, inner_zero_right, map_zero], + neg' := λ x, by simp only [inner_neg_left, neg_neg, inner_neg_right], + add_le' := λ x y, begin have h₁ : abs ⟪x, y⟫ ≤ ∥x∥ * ∥y∥ := abs_inner_le_norm _ _, have h₂ : re ⟪x, y⟫ ≤ abs ⟪x, y⟫ := re_le_abs _, have h₃ : re ⟪x, y⟫ ≤ ∥x∥ * ∥y∥ := by linarith, @@ -348,9 +337,13 @@ normed_add_comm_group.of_core F have : ∥x + y∥ * ∥x + y∥ ≤ (∥x∥ + ∥y∥) * (∥x∥ + ∥y∥), { simp [←inner_self_eq_norm_mul_norm, inner_add_add_self, add_mul, mul_add, mul_comm], linarith }, - exact nonneg_le_nonneg_of_sq_le_sq (add_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) this + exact nonneg_le_nonneg_of_sq_le_sq (add_nonneg (sqrt_nonneg _) (sqrt_nonneg _)) this, end, - norm_neg := λ x, by simp only [norm, inner_neg_left, neg_neg, inner_neg_right] } + eq_zero_of_map_eq_zero' := λ x hx, (inner_self_eq_zero : ⟪x, x⟫ = 0 ↔ x = 0).1 $ begin + change sqrt (re ⟪x, x⟫) = 0 at hx, + rw [sqrt_eq_zero inner_self_nonneg] at hx, + exact ext (by simp [hx]) (by simp [inner_self_im_zero]), + end } local attribute [instance] to_normed_add_comm_group diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index ed247f9f9280a..93adf28b3f9c6 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -81,28 +81,18 @@ def seminormed_add_comm_group.of_add_dist' [has_norm E] [add_comm_group E] [pseu { rw [sub_eq_add_neg, ← add_right_neg y], apply H2 } end } -/-- A seminormed group can be built from a seminorm that satisfies algebraic properties. This is -formalised in this structure. -/ -structure seminormed_add_comm_group.core (E : Type*) [add_comm_group E] [has_norm E] : Prop := -(norm_zero : ∥(0 : E)∥ = 0) -(triangle : ∀ x y : E, ∥x + y∥ ≤ ∥x∥ + ∥y∥) -(norm_neg : ∀ x : E, ∥-x∥ = ∥x∥) - -/-- Constructing a seminormed group from core properties of a seminorm, i.e., registering the -pseudodistance and the pseudometric space structure from the seminorm properties. Note that in most -cases this instance creates bad definitional equalities (e.g., it does not take into account -a possibly existing `uniform_space` instance on `E`). -/ -def seminormed_add_comm_group.of_core (E : Type*) [add_comm_group E] [has_norm E] - (C : seminormed_add_comm_group.core E) : seminormed_add_comm_group E := -{ dist := λ x y, ∥x - y∥, - dist_eq := assume x y, by refl, - dist_self := assume x, by simp [C.norm_zero], - dist_triangle := assume x y z, - calc ∥x - z∥ = ∥x - y + (y - z)∥ : by rw sub_add_sub_cancel - ... ≤ ∥x - y∥ + ∥y - z∥ : C.triangle _ _, - dist_comm := assume x y, - calc ∥x - y∥ = ∥ -(y - x)∥ : by simp - ... = ∥y - x∥ : by { rw [C.norm_neg] } } +/-- Construct a seminormed group from a seminorm, i.e., registering the pseudodistance and the +pseudometric space structure from the seminorm properties. Note that in most cases this instance +creates bad definitional equalities (e.g., it does not take into account a possibly existing +`uniform_space` instance on `E`). -/ +def add_group_seminorm.to_seminormed_add_comm_group [add_comm_group E] (f : add_group_seminorm E) : + seminormed_add_comm_group E := +{ dist := λ x y, f (x - y), + norm := f, + dist_eq := λ x y, rfl, + dist_self := λ x, by simp only [sub_self, map_zero], + dist_triangle := le_map_sub_add_map_sub f, + dist_comm := map_sub_rev f } instance : normed_add_comm_group punit := { norm := function.const _ 0, @@ -1147,32 +1137,14 @@ def normed_add_comm_group.of_add_dist [has_norm E] [add_comm_group E] [metric_sp { have := H2 (x-y) 0 y, rwa [sub_add_cancel, zero_add] at this } end } -/-- A normed group can be built from a norm that satisfies algebraic properties. This is -formalised in this structure. -/ -structure normed_add_comm_group.core (E : Type*) [add_comm_group E] [has_norm E] : Prop := -(norm_eq_zero_iff : ∀ x : E, ∥x∥ = 0 ↔ x = 0) -(triangle : ∀ x y : E, ∥x + y∥ ≤ ∥x∥ + ∥y∥) -(norm_neg : ∀ x : E, ∥-x∥ = ∥x∥) - -/-- The `seminormed_add_comm_group.core` induced by a `normed_add_comm_group.core`. -/ -lemma normed_add_comm_group.core.to_seminormed_add_comm_group.core {E : Type*} [add_comm_group E] - [has_norm E] - (C : normed_add_comm_group.core E) : seminormed_add_comm_group.core E := -{ norm_zero := (C.norm_eq_zero_iff 0).2 rfl, - triangle := C.triangle, - norm_neg := C.norm_neg } - -/-- Constructing a normed group from core properties of a norm, i.e., registering the distance and -the metric space structure from the norm properties. -/ -def normed_add_comm_group.of_core (E : Type*) [add_comm_group E] [has_norm E] - (C : normed_add_comm_group.core E) : normed_add_comm_group E := -{ eq_of_dist_eq_zero := λ x y h, - begin - rw [dist_eq_norm] at h, - exact sub_eq_zero.mp ((C.norm_eq_zero_iff _).1 h) - end - ..seminormed_add_comm_group.of_core E - (normed_add_comm_group.core.to_seminormed_add_comm_group.core C) } +/-- Construct a normed group from a norm, i.e., registering the distance and the metric space +structure from the norm properties. Note that in most cases this instance creates bad definitional +equalities (e.g., it does not take into account a possibly existing `uniform_space` instance on +`E`). -/ +def add_group_norm.to_normed_add_comm_group [add_comm_group E] (f : add_group_norm E) : + normed_add_comm_group E := +{ eq_of_dist_eq_zero := λ x y h, sub_eq_zero.1 $ eq_zero_of_map_eq_zero f h, + ..f.to_add_group_seminorm.to_seminormed_add_comm_group } variables [normed_add_comm_group E] [normed_add_comm_group F] {x y : E} diff --git a/src/analysis/normed/group/hom.lean b/src/analysis/normed/group/hom.lean index 89bb35bccb42a..21172be652425 100644 --- a/src/analysis/normed/group/hom.lean +++ b/src/analysis/normed/group/hom.lean @@ -438,14 +438,23 @@ coe_injective.add_comm_group _ rfl (λ _ _, rfl) (λ _, rfl) (λ _ _, rfl) (λ _ /-- Normed group homomorphisms themselves form a seminormed group with respect to the operator norm. -/ instance to_seminormed_add_comm_group : seminormed_add_comm_group (normed_add_group_hom V₁ V₂) := -seminormed_add_comm_group.of_core _ ⟨op_norm_zero, op_norm_add_le, op_norm_neg⟩ +add_group_seminorm.to_seminormed_add_comm_group +{ to_fun := op_norm, + map_zero' := op_norm_zero, + neg' := op_norm_neg, + add_le' := op_norm_add_le } /-- Normed group homomorphisms themselves form a normed group with respect to the operator norm. -/ instance to_normed_add_comm_group {V₁ V₂ : Type*} [normed_add_comm_group V₁] [normed_add_comm_group V₂] : normed_add_comm_group (normed_add_group_hom V₁ V₂) := -normed_add_comm_group.of_core _ ⟨λ f, op_norm_zero_iff, op_norm_add_le, op_norm_neg⟩ +add_group_norm.to_normed_add_comm_group +{ to_fun := op_norm, + map_zero' := op_norm_zero, + neg' := op_norm_neg, + add_le' := op_norm_add_le, + eq_zero_of_map_eq_zero' := λ f, op_norm_zero_iff.1 } /-- Coercion of a `normed_add_group_hom` is an `add_monoid_hom`. Similar to `add_monoid_hom.coe_fn`. -/ diff --git a/src/analysis/normed_space/continuous_affine_map.lean b/src/analysis/normed_space/continuous_affine_map.lean index 3d990b8987bfa..9ab473114ca9d 100644 --- a/src/analysis/normed_space/continuous_affine_map.lean +++ b/src/analysis/normed_space/continuous_affine_map.lean @@ -159,11 +159,16 @@ calc ∥f∥ = (max ∥f 0∥ ∥f.cont_linear∥) : by rw norm_def ... = ∥f.cont_linear∥ : max_eq_right (norm_nonneg _) noncomputable instance : normed_add_comm_group (V →A[𝕜] W) := -normed_add_comm_group.of_core _ -{ norm_eq_zero_iff := λ f, - begin - rw norm_def, - refine ⟨λ h₀, _, by { rintros rfl, simp, }⟩, +add_group_norm.to_normed_add_comm_group +{ to_fun := λ f, max ∥f 0∥ ∥f.cont_linear∥, + map_zero' := by simp, + neg' := λ f, by simp, + add_le' := λ f g, begin + simp only [pi.add_apply, add_cont_linear, coe_add, max_le_iff], + exact ⟨(norm_add_le _ _).trans (add_le_add (le_max_left _ _) (le_max_left _ _)), + (norm_add_le _ _).trans (add_le_add (le_max_right _ _) (le_max_right _ _))⟩, + end, + eq_zero_of_map_eq_zero' := λ f h₀, begin rcases max_eq_iff.mp h₀ with ⟨h₁, h₂⟩ | ⟨h₁, h₂⟩; rw h₁ at h₂, { rw [norm_le_zero_iff, cont_linear_eq_zero_iff_exists_const] at h₂, @@ -176,14 +181,7 @@ normed_add_comm_group.of_core _ simp only [function.const_apply, coe_const, norm_le_zero_iff] at h₂, rw h₂, refl, }, - end, - triangle := λ f g, - begin - simp only [norm_def, pi.add_apply, add_cont_linear, coe_add, max_le_iff], - exact ⟨(norm_add_le _ _).trans (add_le_add (le_max_left _ _) (le_max_left _ _)), - (norm_add_le _ _).trans (add_le_add (le_max_right _ _) (le_max_right _ _))⟩, - end, - norm_neg := λ f, by simp [norm_def], } + end } instance : normed_space 𝕜 (V →A[𝕜] W) := { norm_smul_le := λ t f, by simp only [norm_def, smul_cont_linear, coe_smul, pi.smul_apply, diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index aa19fce311d62..6fe3fb7086574 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -413,7 +413,7 @@ begin simpa [real.zero_rpow hp.ne'] using real.zero_rpow hp' } end -lemma norm_eq_zero_iff ⦃f : lp E p⦄ : ∥f∥ = 0 ↔ f = 0 := +lemma norm_eq_zero_iff {f : lp E p} : ∥f∥ = 0 ↔ f = 0 := begin classical, refine ⟨λ h, _, by { rintros rfl, exact norm_zero }⟩, @@ -458,9 +458,11 @@ begin end instance [hp : fact (1 ≤ p)] : normed_add_comm_group (lp E p) := -normed_add_comm_group.of_core _ -{ norm_eq_zero_iff := norm_eq_zero_iff, - triangle := λ f g, begin +add_group_norm.to_normed_add_comm_group +{ to_fun := norm, + map_zero' := norm_zero, + neg' := norm_neg, + add_le' := λ f g, begin unfreezingI { rcases p.dichotomy with rfl | hp' }, { casesI is_empty_or_nonempty α, { simp [lp.eq_zero' f] }, @@ -483,7 +485,7 @@ normed_add_comm_group.of_core _ intros i, exact real.rpow_le_rpow (norm_nonneg _) (norm_add_le _ _) hp''.le }, end, - norm_neg := norm_neg } + eq_zero_of_map_eq_zero' := λ f, norm_eq_zero_iff.1 } -- TODO: define an `ennreal` version of `is_conjugate_exponent`, and then express this inequality -- in a better version which also covers the case `p = 1, q = ∞`. diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index 4415d8327d832..75da3346e2dd0 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -344,17 +344,12 @@ cInf_le bounds_bdd_below ⟨add_nonneg (op_norm_nonneg _) (op_norm_nonneg _), λ x, by { rw add_mul, exact norm_add_le_of_le (le_op_norm _ _) (le_op_norm _ _) }⟩ +lemma op_norm_zero : ∥(0 : continuous_multilinear_map 𝕜 E G)∥ = 0 := +(op_norm_nonneg _).antisymm' $ op_norm_le_bound 0 le_rfl $ λ m, by simp + /-- A continuous linear map is zero iff its norm vanishes. -/ theorem op_norm_zero_iff : ∥f∥ = 0 ↔ f = 0 := -begin - split, - { assume h, - ext m, - simpa [h] using f.le_op_norm m }, - { rintro rfl, - apply le_antisymm (op_norm_le_bound 0 le_rfl (λm, _)) (op_norm_nonneg _), - simp } -end +⟨λ h, by { ext m, simpa [h] using f.le_op_norm m }, by { rintro rfl, exact op_norm_zero }⟩ section variables {𝕜' : Type*} [normed_field 𝕜'] [normed_space 𝕜' G] [smul_comm_class 𝕜 𝕜' G] @@ -373,7 +368,12 @@ lemma op_norm_neg : ∥-f∥ = ∥f∥ := by { rw norm_def, apply congr_arg, ext /-- Continuous multilinear maps themselves form a normed space with respect to the operator norm. -/ instance normed_add_comm_group : normed_add_comm_group (continuous_multilinear_map 𝕜 E G) := -normed_add_comm_group.of_core _ ⟨op_norm_zero_iff, op_norm_add_le, op_norm_neg⟩ +add_group_norm.to_normed_add_comm_group +{ to_fun := norm, + map_zero' := op_norm_zero, + neg' := op_norm_neg, + add_le' := op_norm_add_le, + eq_zero_of_map_eq_zero' := λ f, f.op_norm_zero_iff.1 } /-- An alias of `continuous_multilinear_map.normed_add_comm_group` with non-dependent types to help typeclass search. -/ diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index 7a7cbf0936dab..f15433ba581cd 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -383,7 +383,11 @@ lemma op_norm_smul_le {𝕜' : Type*} [normed_field 𝕜'] [normed_space 𝕜' F /-- Continuous linear maps themselves form a seminormed space with respect to the operator norm. -/ instance to_seminormed_add_comm_group : seminormed_add_comm_group (E →SL[σ₁₂] F) := -seminormed_add_comm_group.of_core _ ⟨op_norm_zero, λ x y, op_norm_add_le x y, op_norm_neg⟩ +add_group_seminorm.to_seminormed_add_comm_group +{ to_fun := norm, + map_zero' := op_norm_zero, + add_le' := op_norm_add_le, + neg' := op_norm_neg } lemma nnnorm_def (f : E →SL[σ₁₂] F) : ∥f∥₊ = Inf {c | ∀ x, ∥f x∥₊ ≤ c * ∥x∥₊} := begin @@ -1222,9 +1226,7 @@ iff.intro (λ hn, continuous_linear_map.ext (λ x, norm_le_zero_iff.1 (calc _ ≤ ∥f∥ * ∥x∥ : le_op_norm _ _ ... = _ : by rw [hn, zero_mul]))) - (λ hf, le_antisymm (cInf_le bounds_bdd_below - ⟨le_rfl, λ _, le_of_eq (by { rw [zero_mul, hf], exact norm_zero })⟩) - (op_norm_nonneg _)) + (by { rintro rfl, exact op_norm_zero }) /-- If a normed space is non-trivial, then the norm of the identity equals `1`. -/ @[simp] lemma norm_id [nontrivial E] : ∥id 𝕜 E∥ = 1 := @@ -1239,7 +1241,12 @@ instance norm_one_class [nontrivial E] : norm_one_class (E →L[𝕜] E) := ⟨n /-- Continuous linear maps themselves form a normed space with respect to the operator norm. -/ instance to_normed_add_comm_group [ring_hom_isometric σ₁₂] : normed_add_comm_group (E →SL[σ₁₂] F) := -normed_add_comm_group.of_core _ ⟨λ f, op_norm_zero_iff f, op_norm_add_le, op_norm_neg⟩ +add_group_norm.to_normed_add_comm_group +{ to_fun := norm, + map_zero' := op_norm_zero, + neg' := op_norm_neg, + add_le' := op_norm_add_le, + eq_zero_of_map_eq_zero' := λ f, (op_norm_zero_iff f).1 } /-- Continuous linear maps form a normed ring with respect to the operator norm. -/ instance to_normed_ring : normed_ring (E →L[𝕜] E) := diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index c4a6480f96c74..5d38b024c5ccf 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -14,7 +14,7 @@ import topology.continuous_function.compact # ℒp space and Lp space This file describes properties of almost everywhere strongly measurable functions with finite -seminorm, denoted by `snorm f p μ` and defined for `p:ℝ≥0∞` as `0` if `p=0`, +seminorm, denoted by `snorm f p μ` and defined for `p:ℝ≥0∞` asmm_group (Lp E p μ) := `0` if `p=0`, `(∫ ∥f a∥^p ∂μ) ^ (1/p)` for `0 < p < ∞` and `ess_sup ∥f∥ μ` for `p=∞`. The Prop-valued `mem_ℒp f p μ` states that a function `f : α → E` has finite seminorm. @@ -1727,10 +1727,11 @@ instance [hp : fact (1 ≤ p)] : normed_add_comm_group (Lp E p μ) := edist_dist := λ f g, by rw [edist_def, dist_def, ←snorm_congr_ae (coe_fn_sub _ _), ennreal.of_real_to_real (snorm_ne_top (f - g))], - .. normed_add_comm_group.of_core (Lp E p μ) - { norm_eq_zero_iff := λ f, norm_eq_zero_iff (ennreal.zero_lt_one.trans_le hp.1), - triangle := begin - assume f g, + ..add_group_norm.to_normed_add_comm_group + { to_fun := (norm : Lp E p μ → ℝ), + map_zero' := norm_zero, + neg' := by simp, + add_le' := λ f g, begin simp only [norm_def], rw ← ennreal.to_real_add (snorm_ne_top f) (snorm_ne_top g), suffices h_snorm : snorm ⇑(f + g) p μ ≤ snorm ⇑f p μ + snorm ⇑g p μ, @@ -1739,7 +1740,7 @@ instance [hp : fact (1 ≤ p)] : normed_add_comm_group (Lp E p μ) := rw [snorm_congr_ae (coe_fn_add _ _)], exact snorm_add_le (Lp.ae_strongly_measurable f) (Lp.ae_strongly_measurable g) hp.1, end, - norm_neg := by simp } } + eq_zero_of_map_eq_zero' := λ f, (norm_eq_zero_iff $ ennreal.zero_lt_one.trans_le hp.1).1 } } -- check no diamond is created example [fact (1 ≤ p)] : From 2fb109f0938f26a6499b7793c5f1902ebfce6b8c Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 30 Sep 2022 21:28:58 +0000 Subject: [PATCH 477/828] feat(analysis/inner_product/orientation, geometry/euclidean/oriented_angle): use bundled orthonormal bases (#16475) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Bundled orthonormal bases (as opposed to bases with a mixin predicate `orthonormal`) were defined in #12060, and some of the use cases were switched over in #12253. This PR completes the job, switching to bundled orthonormal bases in `inner_product/orientation` and `euclidean/oriented_angle` as well as in one remaining construction (the standard ` ℝ`-orthonormal basis of `ℂ`) in `inner_product/pi_L2`. Formalized as part of the Sphere Eversion project. The part that I will be using in future PRs is the bundled version of the construction `orthonormal_basis.adjust_to_orientation` in `inner_product/orientation`. --- .../inner_product_space/orientation.lean | 64 ++- src/analysis/inner_product_space/pi_L2.lean | 73 ++- src/geometry/euclidean/oriented_angle.lean | 517 +++++++++--------- 3 files changed, 340 insertions(+), 314 deletions(-) diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index 9195a0b66ca8b..cdc5569c7e729 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -21,42 +21,64 @@ orientation. noncomputable theory variables {E : Type*} [inner_product_space ℝ E] -variables {ι : Type*} [fintype ι] [decidable_eq ι] open finite_dimensional -/-- `basis.adjust_to_orientation`, applied to an orthonormal basis, produces an orthonormal -basis. -/ -lemma orthonormal.orthonormal_adjust_to_orientation [nonempty ι] {e : basis ι ℝ E} - (h : orthonormal ℝ e) (x : orientation ℝ E ι) : orthonormal ℝ (e.adjust_to_orientation x) := -h.orthonormal_of_forall_eq_or_eq_neg (e.adjust_to_orientation_apply_eq_or_eq_neg x) +section adjust_to_orientation +variables {ι : Type*} [fintype ι] [decidable_eq ι] [nonempty ι] (e : orthonormal_basis ι ℝ E) + (x : orientation ℝ E ι) -/-- An orthonormal basis, indexed by `fin n`, with the given orientation. -/ -protected def orientation.fin_orthonormal_basis {n : ℕ} (hn : 0 < n) (h : finrank ℝ E = n) - (x : orientation ℝ E (fin n)) : basis (fin n) ℝ E := +/-- `orthonormal_basis.adjust_to_orientation`, applied to an orthonormal basis, produces an +orthonormal basis. -/ +lemma orthonormal_basis.orthonormal_adjust_to_orientation : + orthonormal ℝ (e.to_basis.adjust_to_orientation x) := begin - haveI := fin.pos_iff_nonempty.1 hn, - haveI := finite_dimensional_of_finrank (h.symm ▸ hn : 0 < finrank ℝ E), - exact (fin_std_orthonormal_basis h).to_basis.adjust_to_orientation x + apply e.orthonormal.orthonormal_of_forall_eq_or_eq_neg, + simpa using e.to_basis.adjust_to_orientation_apply_eq_or_eq_neg x end -/-- `orientation.fin_orthonormal_basis` is orthonormal. -/ -protected lemma orientation.fin_orthonormal_basis_orthonormal {n : ℕ} (hn : 0 < n) - (h : finrank ℝ E = n) (x : orientation ℝ E (fin n)) : - orthonormal ℝ (x.fin_orthonormal_basis hn h) := +/-- Given an orthonormal basis and an orientation, return an orthonormal basis giving that +orientation: either the original basis, or one constructed by negating a single (arbitrary) basis +vector. -/ +def orthonormal_basis.adjust_to_orientation : orthonormal_basis ι ℝ E := +(e.to_basis.adjust_to_orientation x).to_orthonormal_basis (e.orthonormal_adjust_to_orientation x) + +lemma orthonormal_basis.to_basis_adjust_to_orientation : + (e.adjust_to_orientation x).to_basis = e.to_basis.adjust_to_orientation x := +(e.to_basis.adjust_to_orientation x).to_basis_to_orthonormal_basis _ + +/-- `adjust_to_orientation` gives an orthonormal basis with the required orientation. -/ +@[simp] lemma orthonormal_basis.orientation_adjust_to_orientation : + (e.adjust_to_orientation x).to_basis.orientation = x := +begin + rw e.to_basis_adjust_to_orientation, + exact e.to_basis.orientation_adjust_to_orientation x, +end + +/-- Every basis vector from `adjust_to_orientation` is either that from the original basis or its +negation. -/ +lemma orthonormal_basis.adjust_to_orientation_apply_eq_or_eq_neg (i : ι) : + e.adjust_to_orientation x i = e i ∨ e.adjust_to_orientation x i = -(e i) := +by simpa [← e.to_basis_adjust_to_orientation] + using e.to_basis.adjust_to_orientation_apply_eq_or_eq_neg x i + +end adjust_to_orientation + +/-- An orthonormal basis, indexed by `fin n`, with the given orientation. -/ +protected def orientation.fin_orthonormal_basis {n : ℕ} (hn : 0 < n) (h : finrank ℝ E = n) + (x : orientation ℝ E (fin n)) : orthonormal_basis (fin n) ℝ E := begin haveI := fin.pos_iff_nonempty.1 hn, haveI := finite_dimensional_of_finrank (h.symm ▸ hn : 0 < finrank ℝ E), - exact (show orthonormal ℝ (fin_std_orthonormal_basis h).to_basis, -- Note sure how to format this - by simp only [orthonormal_basis.coe_to_basis, orthonormal_basis.orthonormal] - ).orthonormal_adjust_to_orientation _ + exact (fin_std_orthonormal_basis h).adjust_to_orientation x end /-- `orientation.fin_orthonormal_basis` gives a basis with the required orientation. -/ @[simp] lemma orientation.fin_orthonormal_basis_orientation {n : ℕ} (hn : 0 < n) (h : finrank ℝ E = n) (x : orientation ℝ E (fin n)) : - (x.fin_orthonormal_basis hn h).orientation = x := + (x.fin_orthonormal_basis hn h).to_basis.orientation = x := begin haveI := fin.pos_iff_nonempty.1 hn, - exact basis.orientation_adjust_to_orientation _ _ + haveI := finite_dimensional_of_finrank (h.symm ▸ hn : 0 < finrank ℝ E), + exact (fin_std_orthonormal_basis h).orientation_adjust_to_orientation x end diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 38dcf50b06a30..5cd0a253847e5 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -348,6 +348,11 @@ protected def map {G : Type*} [inner_product_space 𝕜 G] (b : orthonormal_basi (b : orthonormal_basis ι 𝕜 E) (L : E ≃ₗᵢ[𝕜] G) (i : ι) : b.map L i = L (b i) := rfl +@[simp] protected lemma to_basis_map {G : Type*} [inner_product_space 𝕜 G] + (b : orthonormal_basis ι 𝕜 E) (L : E ≃ₗᵢ[𝕜] G) : + (b.map L).to_basis = b.to_basis.map L.to_linear_equiv := +rfl + /-- A basis that is orthonormal is an orthonormal basis. -/ def _root_.basis.to_orthonormal_basis (v : basis ι 𝕜 E) (hv : orthonormal 𝕜 v) : orthonormal_basis ι 𝕜 E := @@ -470,66 +475,52 @@ by { classical, end orthonormal_basis -/-- If `f : E ≃ₗᵢ[𝕜] E'` is a linear isometry of inner product spaces then an orthonormal basis `v` -of `E` determines a linear isometry `e : E' ≃ₗᵢ[𝕜] euclidean_space 𝕜 ι`. This result states that -`e` may be obtained either by transporting `v` to `E'` or by composing with the linear isometry -`E ≃ₗᵢ[𝕜] euclidean_space 𝕜 ι` provided by `v`. -/ -@[simp] lemma basis.map_isometry_euclidean_of_orthonormal (v : basis ι 𝕜 E) (hv : orthonormal 𝕜 v) - (f : E ≃ₗᵢ[𝕜] E') : - ((v.map f.to_linear_equiv).to_orthonormal_basis (hv.map_linear_isometry_equiv f)).repr = - f.symm.trans (v.to_orthonormal_basis hv).repr := -linear_isometry_equiv.to_linear_equiv_injective $ v.map_equiv_fun _ - -/-- `ℂ` is isometric to `ℝ²` with the Euclidean inner product. -/ -def complex.isometry_euclidean : ℂ ≃ₗᵢ[ℝ] (euclidean_space ℝ (fin 2)) := +/-- `![1, I]` is an orthonormal basis for `ℂ` considered as a real inner product space. -/ +def complex.orthonormal_basis_one_I : orthonormal_basis (fin 2) ℝ ℂ := (complex.basis_one_I.to_orthonormal_basis begin rw orthonormal_iff_ite, intros i, fin_cases i; intros j; fin_cases j; simp [real_inner_eq_re_inner] -end).repr +end) -@[simp] lemma complex.isometry_euclidean_symm_apply (x : euclidean_space ℝ (fin 2)) : - complex.isometry_euclidean.symm x = (x 0) + (x 1) * I := -begin - convert complex.basis_one_I.equiv_fun_symm_apply x, - { simpa }, - { simp }, -end +@[simp] lemma complex.orthonormal_basis_one_I_repr_apply (z : ℂ) : + complex.orthonormal_basis_one_I.repr z = ![z.re, z.im] := +rfl -lemma complex.isometry_euclidean_proj_eq_self (z : ℂ) : - ↑(complex.isometry_euclidean z 0) + ↑(complex.isometry_euclidean z 1) * (I : ℂ) = z := -by rw [← complex.isometry_euclidean_symm_apply (complex.isometry_euclidean z), - complex.isometry_euclidean.symm_apply_apply z] +@[simp] lemma complex.orthonormal_basis_one_I_repr_symm_apply (x : euclidean_space ℝ (fin 2)) : + complex.orthonormal_basis_one_I.repr.symm x = (x 0) + (x 1) * I := +rfl -@[simp] lemma complex.isometry_euclidean_apply_zero (z : ℂ) : - complex.isometry_euclidean z 0 = z.re := -by { conv_rhs { rw ← complex.isometry_euclidean_proj_eq_self z }, simp } +@[simp] lemma complex.to_basis_orthonormal_basis_one_I : + complex.orthonormal_basis_one_I.to_basis = complex.basis_one_I := +basis.to_basis_to_orthonormal_basis _ _ -@[simp] lemma complex.isometry_euclidean_apply_one (z : ℂ) : - complex.isometry_euclidean z 1 = z.im := -by { conv_rhs { rw ← complex.isometry_euclidean_proj_eq_self z }, simp } +@[simp] lemma complex.coe_orthonormal_basis_one_I : + (complex.orthonormal_basis_one_I : (fin 2) → ℂ) = ![1, I] := +by simp [complex.orthonormal_basis_one_I] /-- The isometry between `ℂ` and a two-dimensional real inner product space given by a basis. -/ -def complex.isometry_of_orthonormal {v : basis (fin 2) ℝ F} (hv : orthonormal ℝ v) : ℂ ≃ₗᵢ[ℝ] F := -complex.isometry_euclidean.trans (v.to_orthonormal_basis hv).repr.symm +def complex.isometry_of_orthonormal (v : orthonormal_basis (fin 2) ℝ F) : ℂ ≃ₗᵢ[ℝ] F := +complex.orthonormal_basis_one_I.repr.trans v.repr.symm -@[simp] lemma complex.map_isometry_of_orthonormal {v : basis (fin 2) ℝ F} (hv : orthonormal ℝ v) +@[simp] lemma complex.map_isometry_of_orthonormal (v : orthonormal_basis (fin 2) ℝ F) (f : F ≃ₗᵢ[ℝ] F') : - complex.isometry_of_orthonormal (hv.map_linear_isometry_equiv f) = - (complex.isometry_of_orthonormal hv).trans f := -by simp [complex.isometry_of_orthonormal, linear_isometry_equiv.trans_assoc] + complex.isometry_of_orthonormal (v.map f) = + (complex.isometry_of_orthonormal v).trans f := +by simp [complex.isometry_of_orthonormal, linear_isometry_equiv.trans_assoc, orthonormal_basis.map] lemma complex.isometry_of_orthonormal_symm_apply - {v : basis (fin 2) ℝ F} (hv : orthonormal ℝ v) (f : F) : - (complex.isometry_of_orthonormal hv).symm f = (v.coord 0 f : ℂ) + (v.coord 1 f : ℂ) * I := + (v : orthonormal_basis (fin 2) ℝ F) (f : F) : + (complex.isometry_of_orthonormal v).symm f + = (v.to_basis.coord 0 f : ℂ) + (v.to_basis.coord 1 f : ℂ) * I := by simp [complex.isometry_of_orthonormal] lemma complex.isometry_of_orthonormal_apply - {v : basis (fin 2) ℝ F} (hv : orthonormal ℝ v) (z : ℂ) : - complex.isometry_of_orthonormal hv z = z.re • v 0 + z.im • v 1 := -by simp [complex.isometry_of_orthonormal, (dec_trivial : (finset.univ : finset (fin 2)) = {0, 1})] + (v : orthonormal_basis (fin 2) ℝ F) (z : ℂ) : + complex.isometry_of_orthonormal v z = z.re • v 0 + z.im • v 1 := +by simp [complex.isometry_of_orthonormal, ← v.sum_repr_symm] open finite_dimensional diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 89a8bc684257d..9d4937d23c5a4 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -43,43 +43,42 @@ noncomputable theory open_locale real open_locale real_inner_product_space -namespace orthonormal +namespace orthonormal_basis variables {V : Type*} [inner_product_space ℝ V] -variables {b : basis (fin 2) ℝ V} (hb : orthonormal ℝ b) -include hb +variables (b : orthonormal_basis (fin 2) ℝ V) /-- The oriented angle from `x` to `y`, modulo `2 * π`. If either vector is 0, this is 0. -/ def oangle (x y : V) : real.angle := -complex.arg ((complex.isometry_of_orthonormal hb).symm y / - (complex.isometry_of_orthonormal hb).symm x) +complex.arg ((complex.isometry_of_orthonormal b).symm y / + (complex.isometry_of_orthonormal b).symm x) /-- Oriented angles are continuous when the vectors involved are nonzero. -/ lemma continuous_at_oangle {x : V × V} (hx1 : x.1 ≠ 0) (hx2 : x.2 ≠ 0) : - continuous_at (λ y : V × V, hb.oangle y.1 y.2) x := + continuous_at (λ y : V × V, b.oangle y.1 y.2) x := (complex.continuous_at_arg_coe_angle (by simp [hx1, hx2])).comp $ continuous_at.div - ((complex.isometry_of_orthonormal hb).symm.continuous.comp continuous_snd).continuous_at - ((complex.isometry_of_orthonormal hb).symm.continuous.comp continuous_fst).continuous_at + ((complex.isometry_of_orthonormal b).symm.continuous.comp continuous_snd).continuous_at + ((complex.isometry_of_orthonormal b).symm.continuous.comp continuous_fst).continuous_at (by simp [hx1]) /-- If the first vector passed to `oangle` is 0, the result is 0. -/ -@[simp] lemma oangle_zero_left (x : V) : hb.oangle 0 x = 0 := +@[simp] lemma oangle_zero_left (x : V) : b.oangle 0 x = 0 := by simp [oangle] /-- If the second vector passed to `oangle` is 0, the result is 0. -/ -@[simp] lemma oangle_zero_right (x : V) : hb.oangle x 0 = 0 := +@[simp] lemma oangle_zero_right (x : V) : b.oangle x 0 = 0 := by simp [oangle] /-- If the two vectors passed to `oangle` are the same, the result is 0. -/ -@[simp] lemma oangle_self (x : V) : hb.oangle x x = 0 := +@[simp] lemma oangle_self (x : V) : b.oangle x x = 0 := begin by_cases h : x = 0; simp [oangle, h] end /-- Swapping the two vectors passed to `oangle` negates the angle. -/ -lemma oangle_rev (x y : V) : hb.oangle y x = -hb.oangle x y := +lemma oangle_rev (x y : V) : b.oangle y x = -b.oangle x y := begin simp only [oangle], convert complex.arg_inv_coe_angle _, @@ -87,12 +86,12 @@ begin end /-- Adding the angles between two vectors in each order results in 0. -/ -@[simp] lemma oangle_add_oangle_rev (x y : V) : hb.oangle x y + hb.oangle y x = 0 := -by simp [hb.oangle_rev y x] +@[simp] lemma oangle_add_oangle_rev (x y : V) : b.oangle x y + b.oangle y x = 0 := +by simp [b.oangle_rev y x] /-- Negating the first vector passed to `oangle` adds `π` to the angle. -/ lemma oangle_neg_left {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - hb.oangle (-x) y = hb.oangle x y + π := + b.oangle (-x) y = b.oangle x y + π := begin simp only [oangle, div_neg_eq_neg_div, map_neg], refine complex.arg_neg_coe_angle _, @@ -101,7 +100,7 @@ end /-- Negating the second vector passed to `oangle` adds `π` to the angle. -/ lemma oangle_neg_right {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - hb.oangle x (-y) = hb.oangle x y + π := + b.oangle x (-y) = b.oangle x y + π := begin simp only [oangle, neg_div, map_neg], refine complex.arg_neg_coe_angle _, @@ -110,51 +109,51 @@ end /-- Negating the first vector passed to `oangle` does not change twice the angle. -/ @[simp] lemma two_zsmul_oangle_neg_left (x y : V) : - (2 : ℤ) • hb.oangle (-x) y = (2 : ℤ) • hb.oangle x y := + (2 : ℤ) • b.oangle (-x) y = (2 : ℤ) • b.oangle x y := begin by_cases hx : x = 0, { simp [hx] }, { by_cases hy : y = 0, { simp [hy] }, - { simp [hb.oangle_neg_left hx hy] } } + { simp [b.oangle_neg_left hx hy] } } end /-- Negating the second vector passed to `oangle` does not change twice the angle. -/ @[simp] lemma two_zsmul_oangle_neg_right (x y : V) : - (2 : ℤ) • hb.oangle x (-y) = (2 : ℤ) • hb.oangle x y := + (2 : ℤ) • b.oangle x (-y) = (2 : ℤ) • b.oangle x y := begin by_cases hx : x = 0, { simp [hx] }, { by_cases hy : y = 0, { simp [hy] }, - { simp [hb.oangle_neg_right hx hy] } } + { simp [b.oangle_neg_right hx hy] } } end /-- Negating both vectors passed to `oangle` does not change the angle. -/ -@[simp] lemma oangle_neg_neg (x y : V) : hb.oangle (-x) (-y) = hb.oangle x y := +@[simp] lemma oangle_neg_neg (x y : V) : b.oangle (-x) (-y) = b.oangle x y := by simp [oangle, neg_div_neg_eq] /-- Negating the first vector produces the same angle as negating the second vector. -/ -lemma oangle_neg_left_eq_neg_right (x y : V) : hb.oangle (-x) y = hb.oangle x (-y) := +lemma oangle_neg_left_eq_neg_right (x y : V) : b.oangle (-x) y = b.oangle x (-y) := by rw [←neg_neg y, oangle_neg_neg, neg_neg] /-- The angle between the negation of a nonzero vector and that vector is `π`. -/ -@[simp] lemma oangle_neg_self_left {x : V} (hx : x ≠ 0) : hb.oangle (-x) x = π := +@[simp] lemma oangle_neg_self_left {x : V} (hx : x ≠ 0) : b.oangle (-x) x = π := by simp [oangle_neg_left, hx] /-- The angle between a nonzero vector and its negation is `π`. -/ -@[simp] lemma oangle_neg_self_right {x : V} (hx : x ≠ 0) : hb.oangle x (-x) = π := +@[simp] lemma oangle_neg_self_right {x : V} (hx : x ≠ 0) : b.oangle x (-x) = π := by simp [oangle_neg_right, hx] /-- Twice the angle between the negation of a vector and that vector is 0. -/ -@[simp] lemma two_zsmul_oangle_neg_self_left (x : V) : (2 : ℤ) • hb.oangle (-x) x = 0 := +@[simp] lemma two_zsmul_oangle_neg_self_left (x : V) : (2 : ℤ) • b.oangle (-x) x = 0 := begin by_cases hx : x = 0; simp [hx] end /-- Twice the angle between a vector and its negation is 0. -/ -@[simp] lemma two_zsmul_oangle_neg_self_right (x : V) : (2 : ℤ) • hb.oangle x (-x) = 0 := +@[simp] lemma two_zsmul_oangle_neg_self_right (x : V) : (2 : ℤ) • b.oangle x (-x) = 0 := begin by_cases hx : x = 0; simp [hx] @@ -163,19 +162,19 @@ end /-- Adding the angles between two vectors in each order, with the first vector in each angle negated, results in 0. -/ @[simp] lemma oangle_add_oangle_rev_neg_left (x y : V) : - hb.oangle (-x) y + hb.oangle (-y) x = 0 := + b.oangle (-x) y + b.oangle (-y) x = 0 := by rw [oangle_neg_left_eq_neg_right, oangle_rev, add_left_neg] /-- Adding the angles between two vectors in each order, with the second vector in each angle negated, results in 0. -/ @[simp] lemma oangle_add_oangle_rev_neg_right (x y : V) : - hb.oangle x (-y) + hb.oangle y (-x) = 0 := -by rw [hb.oangle_rev (-x), oangle_neg_left_eq_neg_right, add_neg_self] + b.oangle x (-y) + b.oangle y (-x) = 0 := +by rw [b.oangle_rev (-x), oangle_neg_left_eq_neg_right, add_neg_self] /-- Multiplying the first vector passed to `oangle` by a positive real does not change the angle. -/ @[simp] lemma oangle_smul_left_of_pos (x y : V) {r : ℝ} (hr : 0 < r) : - hb.oangle (r • x) y = hb.oangle x y := + b.oangle (r • x) y = b.oangle x y := begin simp only [oangle, linear_isometry_equiv.map_smul, complex.real_smul], rw [mul_comm, div_mul_eq_div_mul_one_div, one_div, mul_comm, ←complex.of_real_inv], @@ -186,7 +185,7 @@ end /-- Multiplying the second vector passed to `oangle` by a positive real does not change the angle. -/ @[simp] lemma oangle_smul_right_of_pos (x y : V) {r : ℝ} (hr : 0 < r) : - hb.oangle x (r • y) = hb.oangle x y := + b.oangle x (r • y) = b.oangle x y := begin simp only [oangle, linear_isometry_equiv.map_smul, complex.real_smul], congr' 1, @@ -197,18 +196,18 @@ end /-- Multiplying the first vector passed to `oangle` by a negative real produces the same angle as negating that vector. -/ @[simp] lemma oangle_smul_left_of_neg (x y : V) {r : ℝ} (hr : r < 0) : - hb.oangle (r • x) y = hb.oangle (-x) y := -by rw [←neg_neg r, neg_smul, ←smul_neg, hb.oangle_smul_left_of_pos _ _ (neg_pos_of_neg hr)] + b.oangle (r • x) y = b.oangle (-x) y := +by rw [←neg_neg r, neg_smul, ←smul_neg, b.oangle_smul_left_of_pos _ _ (neg_pos_of_neg hr)] /-- Multiplying the second vector passed to `oangle` by a negative real produces the same angle as negating that vector. -/ @[simp] lemma oangle_smul_right_of_neg (x y : V) {r : ℝ} (hr : r < 0) : - hb.oangle x (r • y) = hb.oangle x (-y) := -by rw [←neg_neg r, neg_smul, ←smul_neg, hb.oangle_smul_right_of_pos _ _ (neg_pos_of_neg hr)] + b.oangle x (r • y) = b.oangle x (-y) := +by rw [←neg_neg r, neg_smul, ←smul_neg, b.oangle_smul_right_of_pos _ _ (neg_pos_of_neg hr)] /-- The angle between a nonnegative multiple of a vector and that vector is 0. -/ @[simp] lemma oangle_smul_left_self_of_nonneg (x : V) {r : ℝ} (hr : 0 ≤ r) : - hb.oangle (r • x) x = 0 := + b.oangle (r • x) x = 0 := begin rcases hr.lt_or_eq with (h|h), { simp [h] }, @@ -217,7 +216,7 @@ end /-- The angle between a vector and a nonnegative multiple of that vector is 0. -/ @[simp] lemma oangle_smul_right_self_of_nonneg (x : V) {r : ℝ} (hr : 0 ≤ r) : - hb.oangle x (r • x) = 0 := + b.oangle x (r • x) = 0 := begin rcases hr.lt_or_eq with (h|h), { simp [h] }, @@ -226,7 +225,7 @@ end /-- The angle between two nonnegative multiples of the same vector is 0. -/ @[simp] lemma oangle_smul_smul_self_of_nonneg (x : V) {r₁ r₂ : ℝ} (hr₁ : 0 ≤ r₁) (hr₂ : 0 ≤ r₂) : - hb.oangle (r₁ • x) (r₂ • x) = 0 := + b.oangle (r₁ • x) (r₂ • x) = 0 := begin rcases hr₁.lt_or_eq with (h|h), { simp [h, hr₂] }, @@ -236,7 +235,7 @@ end /-- Multiplying the first vector passed to `oangle` by a nonzero real does not change twice the angle. -/ @[simp] lemma two_zsmul_oangle_smul_left_of_ne_zero (x y : V) {r : ℝ} (hr : r ≠ 0) : - (2 : ℤ) • hb.oangle (r • x) y = (2 : ℤ) • hb.oangle x y := + (2 : ℤ) • b.oangle (r • x) y = (2 : ℤ) • b.oangle x y := begin rcases hr.lt_or_lt with (h|h); simp [h] @@ -245,7 +244,7 @@ end /-- Multiplying the second vector passed to `oangle` by a nonzero real does not change twice the angle. -/ @[simp] lemma two_zsmul_oangle_smul_right_of_ne_zero (x y : V) {r : ℝ} (hr : r ≠ 0) : - (2 : ℤ) • hb.oangle x (r • y) = (2 : ℤ) • hb.oangle x y := + (2 : ℤ) • b.oangle x (r • y) = (2 : ℤ) • b.oangle x y := begin rcases hr.lt_or_lt with (h|h); simp [h] @@ -253,7 +252,7 @@ end /-- Twice the angle between a multiple of a vector and that vector is 0. -/ @[simp] lemma two_zsmul_oangle_smul_left_self (x : V) {r : ℝ} : - (2 : ℤ) • hb.oangle (r • x) x = 0 := + (2 : ℤ) • b.oangle (r • x) x = 0 := begin rcases lt_or_le r 0 with (h|h); simp [h] @@ -261,7 +260,7 @@ end /-- Twice the angle between a vector and a multiple of that vector is 0. -/ @[simp] lemma two_zsmul_oangle_smul_right_self (x : V) {r : ℝ} : - (2 : ℤ) • hb.oangle x (r • x) = 0 := + (2 : ℤ) • b.oangle x (r • x) = 0 := begin rcases lt_or_le r 0 with (h|h); simp [h] @@ -269,7 +268,7 @@ end /-- Twice the angle between two multiples of a vector is 0. -/ @[simp] lemma two_zsmul_oangle_smul_smul_self (x : V) {r₁ r₂ : ℝ} : - (2 : ℤ) • hb.oangle (r₁ • x) (r₂ • x) = 0 := + (2 : ℤ) • b.oangle (r₁ • x) (r₂ • x) = 0 := begin by_cases h : r₁ = 0; simp [h] @@ -277,47 +276,47 @@ end /-- The oriented angle between two vectors is zero if and only if the angle with the vectors swapped is zero. -/ -lemma oangle_eq_zero_iff_oangle_rev_eq_zero {x y : V} : hb.oangle x y = 0 ↔ hb.oangle y x = 0 := +lemma oangle_eq_zero_iff_oangle_rev_eq_zero {x y : V} : b.oangle x y = 0 ↔ b.oangle y x = 0 := by rw [oangle_rev, neg_eq_zero] /-- The oriented angle between two vectors is zero if and only if they are on the same ray. -/ -lemma oangle_eq_zero_iff_same_ray {x y : V} : hb.oangle x y = 0 ↔ same_ray ℝ x y := +lemma oangle_eq_zero_iff_same_ray {x y : V} : b.oangle x y = 0 ↔ same_ray ℝ x y := by rw [oangle, complex.arg_coe_angle_eq_iff_eq_to_real, real.angle.to_real_zero, ←complex.same_ray_iff_arg_div_eq_zero, ←linear_isometry_equiv.coe_to_linear_equiv, same_ray_map_iff, same_ray_comm] /-- The oriented angle between two vectors is `π` if and only if the angle with the vectors swapped is `π`. -/ -lemma oangle_eq_pi_iff_oangle_rev_eq_pi {x y : V} : hb.oangle x y = π ↔ hb.oangle y x = π := +lemma oangle_eq_pi_iff_oangle_rev_eq_pi {x y : V} : b.oangle x y = π ↔ b.oangle y x = π := by rw [oangle_rev, neg_eq_iff_neg_eq, eq_comm, real.angle.neg_coe_pi] /-- The oriented angle between two vectors is `π` if and only they are nonzero and the first is on the same ray as the negation of the second. -/ lemma oangle_eq_pi_iff_same_ray_neg {x y : V} : - hb.oangle x y = π ↔ x ≠ 0 ∧ y ≠ 0 ∧ same_ray ℝ x (-y) := + b.oangle x y = π ↔ x ≠ 0 ∧ y ≠ 0 ∧ same_ray ℝ x (-y) := begin - rw [←hb.oangle_eq_zero_iff_same_ray], + rw [←b.oangle_eq_zero_iff_same_ray], split, { intro h, by_cases hx : x = 0, { simpa [hx, real.angle.pi_ne_zero.symm] using h }, by_cases hy : y = 0, { simpa [hy, real.angle.pi_ne_zero.symm] using h }, refine ⟨hx, hy, _⟩, - rw [hb.oangle_neg_right hx hy, h, real.angle.coe_pi_add_coe_pi] }, + rw [b.oangle_neg_right hx hy, h, real.angle.coe_pi_add_coe_pi] }, { rintro ⟨hx, hy, h⟩, - rwa [hb.oangle_neg_right hx hy, ←real.angle.sub_coe_pi_eq_add_coe_pi, sub_eq_zero] at h } + rwa [b.oangle_neg_right hx hy, ←real.angle.sub_coe_pi_eq_add_coe_pi, sub_eq_zero] at h } end /-- The oriented angle between two vectors is zero or `π` if and only if those two vectors are not linearly independent. -/ lemma oangle_eq_zero_or_eq_pi_iff_not_linear_independent {x y : V} : - (hb.oangle x y = 0 ∨ hb.oangle x y = π) ↔ ¬ _root_.linear_independent ℝ ![x, y] := + (b.oangle x y = 0 ∨ b.oangle x y = π) ↔ ¬ linear_independent ℝ ![x, y] := by rw [oangle_eq_zero_iff_same_ray, oangle_eq_pi_iff_same_ray_neg, same_ray_or_ne_zero_and_same_ray_neg_iff_not_linear_independent] /-- The oriented angle between two vectors is zero or `π` if and only if the first vector is zero or the second is a multiple of the first. -/ lemma oangle_eq_zero_or_eq_pi_iff_right_eq_smul {x y : V} : - (hb.oangle x y = 0 ∨ hb.oangle x y = π) ↔ (x = 0 ∨ ∃ r : ℝ, y = r • x) := + (b.oangle x y = 0 ∨ b.oangle x y = π) ↔ (x = 0 ∨ ∃ r : ℝ, y = r • x) := begin rw [oangle_eq_zero_iff_same_ray, oangle_eq_pi_iff_same_ray_neg], refine ⟨λ h, _, λ h, _⟩, @@ -342,11 +341,11 @@ end /-- The oriented angle between two vectors is not zero or `π` if and only if those two vectors are linearly independent. -/ lemma oangle_ne_zero_and_ne_pi_iff_linear_independent {x y : V} : - (hb.oangle x y ≠ 0 ∧ hb.oangle x y ≠ π) ↔ _root_.linear_independent ℝ ![x, y] := + (b.oangle x y ≠ 0 ∧ b.oangle x y ≠ π) ↔ linear_independent ℝ ![x, y] := by rw [←not_or_distrib, ←not_iff_not, not_not, oangle_eq_zero_or_eq_pi_iff_not_linear_independent] /-- Two vectors are equal if and only if they have equal norms and zero angle between them. -/ -lemma eq_iff_norm_eq_and_oangle_eq_zero (x y : V) : x = y ↔ ∥x∥ = ∥y∥ ∧ hb.oangle x y = 0 := +lemma eq_iff_norm_eq_and_oangle_eq_zero (x y : V) : x = y ↔ ∥x∥ = ∥y∥ ∧ b.oangle x y = 0 := begin split, { intro h, @@ -356,9 +355,9 @@ begin by_cases hy0 : y = 0, { simpa [hy0] using hn }, { have hx0 : x ≠ 0 := norm_ne_zero_iff.1 (hn.symm ▸ norm_ne_zero_iff.2 hy0), - have hx0' : (complex.isometry_of_orthonormal hb).symm x ≠ 0, + have hx0' : (complex.isometry_of_orthonormal b).symm x ≠ 0, { simp [hx0] }, - have hy0' : (complex.isometry_of_orthonormal hb).symm y ≠ 0, + have hy0' : (complex.isometry_of_orthonormal b).symm y ≠ 0, { simp [hy0] }, rw [complex.arg_div_coe_angle hy0' hx0', sub_eq_zero, complex.arg_coe_angle_eq_iff, complex.arg_eq_arg_iff hy0' hx0', ←complex.norm_eq_abs, ←complex.norm_eq_abs, @@ -369,19 +368,19 @@ begin end /-- Two vectors with equal norms are equal if and only if they have zero angle between them. -/ -lemma eq_iff_oangle_eq_zero_of_norm_eq {x y : V} (h : ∥x∥ = ∥y∥) : x = y ↔ hb.oangle x y = 0 := -⟨λ he, ((hb.eq_iff_norm_eq_and_oangle_eq_zero x y).1 he).2, - λ ha, (hb.eq_iff_norm_eq_and_oangle_eq_zero x y).2 ⟨h, ha⟩⟩ +lemma eq_iff_oangle_eq_zero_of_norm_eq {x y : V} (h : ∥x∥ = ∥y∥) : x = y ↔ b.oangle x y = 0 := +⟨λ he, ((b.eq_iff_norm_eq_and_oangle_eq_zero x y).1 he).2, + λ ha, (b.eq_iff_norm_eq_and_oangle_eq_zero x y).2 ⟨h, ha⟩⟩ /-- Two vectors with zero angle between them are equal if and only if they have equal norms. -/ -lemma eq_iff_norm_eq_of_oangle_eq_zero {x y : V} (h : hb.oangle x y = 0) : x = y ↔ ∥x∥ = ∥y∥ := -⟨λ he, ((hb.eq_iff_norm_eq_and_oangle_eq_zero x y).1 he).1, - λ hn, (hb.eq_iff_norm_eq_and_oangle_eq_zero x y).2 ⟨hn, h⟩⟩ +lemma eq_iff_norm_eq_of_oangle_eq_zero {x y : V} (h : b.oangle x y = 0) : x = y ↔ ∥x∥ = ∥y∥ := +⟨λ he, ((b.eq_iff_norm_eq_and_oangle_eq_zero x y).1 he).1, + λ hn, (b.eq_iff_norm_eq_and_oangle_eq_zero x y).2 ⟨hn, h⟩⟩ /-- Given three nonzero vectors, the angle between the first and the second plus the angle between the second and the third equals the angle between the first and the third. -/ @[simp] lemma oangle_add {x y z : V} (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) : - hb.oangle x y + hb.oangle y z = hb.oangle x z := + b.oangle x y + b.oangle y z = b.oangle x z := begin simp_rw [oangle], rw ←complex.arg_mul_coe_angle, @@ -394,54 +393,54 @@ end /-- Given three nonzero vectors, the angle between the second and the third plus the angle between the first and the second equals the angle between the first and the third. -/ @[simp] lemma oangle_add_swap {x y z : V} (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) : - hb.oangle y z + hb.oangle x y = hb.oangle x z := -by rw [add_comm, hb.oangle_add hx hy hz] + b.oangle y z + b.oangle x y = b.oangle x z := +by rw [add_comm, b.oangle_add hx hy hz] /-- Given three nonzero vectors, the angle between the first and the third minus the angle between the first and the second equals the angle between the second and the third. -/ @[simp] lemma oangle_sub_left {x y z : V} (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) : - hb.oangle x z - hb.oangle x y = hb.oangle y z := -by rw [sub_eq_iff_eq_add, hb.oangle_add_swap hx hy hz] + b.oangle x z - b.oangle x y = b.oangle y z := +by rw [sub_eq_iff_eq_add, b.oangle_add_swap hx hy hz] /-- Given three nonzero vectors, the angle between the first and the third minus the angle between the second and the third equals the angle between the first and the second. -/ @[simp] lemma oangle_sub_right {x y z : V} (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) : - hb.oangle x z - hb.oangle y z = hb.oangle x y := -by rw [sub_eq_iff_eq_add, hb.oangle_add hx hy hz] + b.oangle x z - b.oangle y z = b.oangle x y := +by rw [sub_eq_iff_eq_add, b.oangle_add hx hy hz] /-- Given three nonzero vectors, adding the angles between them in cyclic order results in 0. -/ @[simp] lemma oangle_add_cyc3 {x y z : V} (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) : - hb.oangle x y + hb.oangle y z + hb.oangle z x = 0 := + b.oangle x y + b.oangle y z + b.oangle z x = 0 := by simp [hx, hy, hz] /-- Given three nonzero vectors, adding the angles between them in cyclic order, with the first vector in each angle negated, results in π. If the vectors add to 0, this is a version of the sum of the angles of a triangle. -/ @[simp] lemma oangle_add_cyc3_neg_left {x y z : V} (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) : - hb.oangle (-x) y + hb.oangle (-y) z + hb.oangle (-z) x = π := -by rw [hb.oangle_neg_left hx hy, hb.oangle_neg_left hy hz, hb.oangle_neg_left hz hx, - (show hb.oangle x y + π + (hb.oangle y z + π) + (hb.oangle z x + π) = - hb.oangle x y + hb.oangle y z + hb.oangle z x + (π + π + π : real.angle), by abel), - hb.oangle_add_cyc3 hx hy hz, real.angle.coe_pi_add_coe_pi, zero_add, zero_add] + b.oangle (-x) y + b.oangle (-y) z + b.oangle (-z) x = π := +by rw [b.oangle_neg_left hx hy, b.oangle_neg_left hy hz, b.oangle_neg_left hz hx, + (show b.oangle x y + π + (b.oangle y z + π) + (b.oangle z x + π) = + b.oangle x y + b.oangle y z + b.oangle z x + (π + π + π : real.angle), by abel), + b.oangle_add_cyc3 hx hy hz, real.angle.coe_pi_add_coe_pi, zero_add, zero_add] /-- Given three nonzero vectors, adding the angles between them in cyclic order, with the second vector in each angle negated, results in π. If the vectors add to 0, this is a version of the sum of the angles of a triangle. -/ @[simp] lemma oangle_add_cyc3_neg_right {x y z : V} (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) : - hb.oangle x (-y) + hb.oangle y (-z) + hb.oangle z (-x) = π := -by simp_rw [←oangle_neg_left_eq_neg_right, hb.oangle_add_cyc3_neg_left hx hy hz] + b.oangle x (-y) + b.oangle y (-z) + b.oangle z (-x) = π := +by simp_rw [←oangle_neg_left_eq_neg_right, b.oangle_add_cyc3_neg_left hx hy hz] /-- Pons asinorum, oriented vector angle form. -/ lemma oangle_sub_eq_oangle_sub_rev_of_norm_eq {x y : V} (h : ∥x∥ = ∥y∥) : - hb.oangle x (x - y) = hb.oangle (y - x) y := + b.oangle x (x - y) = b.oangle (y - x) y := begin by_cases hx : x = 0, { simp [hx] }, { have hy : y ≠ 0 := norm_ne_zero_iff.1 (h ▸ norm_ne_zero_iff.2 hx), - simp_rw [hb.oangle_rev y, oangle, linear_isometry_equiv.map_sub, + simp_rw [b.oangle_rev y, oangle, linear_isometry_equiv.map_sub, ←complex.arg_conj_coe_angle, sub_div, - div_self (((complex.isometry_of_orthonormal hb).symm.map_eq_zero_iff).not.2 hx), - div_self (((complex.isometry_of_orthonormal hb).symm.map_eq_zero_iff).not.2 hy), + div_self (((complex.isometry_of_orthonormal b).symm.map_eq_zero_iff).not.2 hx), + div_self (((complex.isometry_of_orthonormal b).symm.map_eq_zero_iff).not.2 hy), map_sub, map_one], rw ←inv_div, simp_rw [complex.inv_def, complex.norm_sq_div, ←complex.sq_abs, ←complex.norm_eq_abs, @@ -452,24 +451,24 @@ end /-- The angle at the apex of an isosceles triangle is `π` minus twice a base angle, oriented vector angle form. -/ lemma oangle_eq_pi_sub_two_zsmul_oangle_sub_of_norm_eq {x y : V} (hn : x ≠ y) (h : ∥x∥ = ∥y∥) : - hb.oangle y x = π - (2 : ℤ) • hb.oangle (y - x) y := + b.oangle y x = π - (2 : ℤ) • b.oangle (y - x) y := begin rw two_zsmul, - rw [←hb.oangle_sub_eq_oangle_sub_rev_of_norm_eq h] { occs := occurrences.pos [1] }, + rw [←b.oangle_sub_eq_oangle_sub_rev_of_norm_eq h] { occs := occurrences.pos [1] }, rw [eq_sub_iff_add_eq, ←oangle_neg_neg, ←add_assoc], have hy : y ≠ 0, { rintro rfl, rw [norm_zero, norm_eq_zero] at h, exact hn h }, have hx : x ≠ 0 := norm_ne_zero_iff.1 (h.symm ▸ norm_ne_zero_iff.2 hy), - convert hb.oangle_add_cyc3_neg_right (neg_ne_zero.2 hy) hx (sub_ne_zero_of_ne hn.symm); + convert b.oangle_add_cyc3_neg_right (neg_ne_zero.2 hy) hx (sub_ne_zero_of_ne hn.symm); simp end /-- Angle at center of a circle equals twice angle at circumference, oriented vector angle form. -/ lemma oangle_eq_two_zsmul_oangle_sub_of_norm_eq {x y z : V} (hxyne : x ≠ y) (hxzne : x ≠ z) - (hxy : ∥x∥ = ∥y∥) (hxz : ∥x∥ = ∥z∥) : hb.oangle y z = (2 : ℤ) • hb.oangle (y - x) (z - x) := + (hxy : ∥x∥ = ∥y∥) (hxz : ∥x∥ = ∥z∥) : b.oangle y z = (2 : ℤ) • b.oangle (y - x) (z - x) := begin have hy : y ≠ 0, { rintro rfl, @@ -477,15 +476,15 @@ begin exact hxyne hxy }, have hx : x ≠ 0 := norm_ne_zero_iff.1 (hxy.symm ▸ norm_ne_zero_iff.2 hy), have hz : z ≠ 0 := norm_ne_zero_iff.1 (hxz ▸ norm_ne_zero_iff.2 hx), - calc hb.oangle y z = hb.oangle x z - hb.oangle x y : (hb.oangle_sub_left hx hy hz).symm - ... = (π - (2 : ℤ) • hb.oangle (x - z) x) - - (π - (2 : ℤ) • hb.oangle (x - y) x) : - by rw [hb.oangle_eq_pi_sub_two_zsmul_oangle_sub_of_norm_eq hxzne.symm hxz.symm, - hb.oangle_eq_pi_sub_two_zsmul_oangle_sub_of_norm_eq hxyne.symm hxy.symm] - ... = (2 : ℤ) • (hb.oangle (x - y) x - hb.oangle (x - z) x) : by abel - ... = (2 : ℤ) • hb.oangle (x - y) (x - z) : - by rw hb.oangle_sub_right (sub_ne_zero_of_ne hxyne) (sub_ne_zero_of_ne hxzne) hx - ... = (2 : ℤ) • hb.oangle (y - x) (z - x) : + calc b.oangle y z = b.oangle x z - b.oangle x y : (b.oangle_sub_left hx hy hz).symm + ... = (π - (2 : ℤ) • b.oangle (x - z) x) - + (π - (2 : ℤ) • b.oangle (x - y) x) : + by rw [b.oangle_eq_pi_sub_two_zsmul_oangle_sub_of_norm_eq hxzne.symm hxz.symm, + b.oangle_eq_pi_sub_two_zsmul_oangle_sub_of_norm_eq hxyne.symm hxy.symm] + ... = (2 : ℤ) • (b.oangle (x - y) x - b.oangle (x - z) x) : by abel + ... = (2 : ℤ) • b.oangle (x - y) (x - z) : + by rw b.oangle_sub_right (sub_ne_zero_of_ne hxyne) (sub_ne_zero_of_ne hxzne) hx + ... = (2 : ℤ) • b.oangle (y - x) (z - x) : by rw [←oangle_neg_neg, neg_sub, neg_sub] end @@ -493,8 +492,8 @@ end form with radius specified. -/ lemma oangle_eq_two_zsmul_oangle_sub_of_norm_eq_real {x y z : V} (hxyne : x ≠ y) (hxzne : x ≠ z) {r : ℝ} (hx : ∥x∥ = r) (hy : ∥y∥ = r) (hz : ∥z∥ = r) : - hb.oangle y z = (2 : ℤ) • hb.oangle (y - x) (z - x) := -hb.oangle_eq_two_zsmul_oangle_sub_of_norm_eq hxyne hxzne (hy.symm ▸ hx) (hz.symm ▸ hx) + b.oangle y z = (2 : ℤ) • b.oangle (y - x) (z - x) := +b.oangle_eq_two_zsmul_oangle_sub_of_norm_eq hxyne hxzne (hy.symm ▸ hx) (hz.symm ▸ hx) /-- Oriented vector angle version of "angles in same segment are equal" and "opposite angles of a cyclic quadrilateral add to π", for oriented angles mod π (for which those are the same @@ -502,35 +501,35 @@ result), represented here as equality of twice the angles. -/ lemma two_zsmul_oangle_sub_eq_two_zsmul_oangle_sub_of_norm_eq {x₁ x₂ y z : V} (hx₁yne : x₁ ≠ y) (hx₁zne : x₁ ≠ z) (hx₂yne : x₂ ≠ y) (hx₂zne : x₂ ≠ z) {r : ℝ} (hx₁ : ∥x₁∥ = r) (hx₂ : ∥x₂∥ = r) (hy : ∥y∥ = r) (hz : ∥z∥ = r) : - (2 : ℤ) • hb.oangle (y - x₁) (z - x₁) = (2 : ℤ) • hb.oangle (y - x₂) (z - x₂) := -hb.oangle_eq_two_zsmul_oangle_sub_of_norm_eq_real hx₁yne hx₁zne hx₁ hy hz ▸ - hb.oangle_eq_two_zsmul_oangle_sub_of_norm_eq_real hx₂yne hx₂zne hx₂ hy hz + (2 : ℤ) • b.oangle (y - x₁) (z - x₁) = (2 : ℤ) • b.oangle (y - x₂) (z - x₂) := +b.oangle_eq_two_zsmul_oangle_sub_of_norm_eq_real hx₁yne hx₁zne hx₁ hy hz ▸ + b.oangle_eq_two_zsmul_oangle_sub_of_norm_eq_real hx₂yne hx₂zne hx₂ hy hz /-- A rotation by the oriented angle `θ`. -/ def rotation (θ : real.angle) : V ≃ₗᵢ[ℝ] V := -((complex.isometry_of_orthonormal hb).symm.trans (rotation (real.angle.exp_map_circle θ))).trans - (complex.isometry_of_orthonormal hb) +((complex.isometry_of_orthonormal b).symm.trans (rotation (real.angle.exp_map_circle θ))).trans + (complex.isometry_of_orthonormal b) /-- The determinant of `rotation` (as a linear map) is equal to `1`. -/ @[simp] lemma det_rotation (θ : real.angle) : - ((hb.rotation θ).to_linear_equiv : V →ₗ[ℝ] V).det = 1 := + ((b.rotation θ).to_linear_equiv : V →ₗ[ℝ] V).det = 1 := by simp [rotation, ←linear_isometry_equiv.to_linear_equiv_symm, ←linear_equiv.comp_coe] /-- The determinant of `rotation` (as a linear equiv) is equal to `1`. -/ @[simp] lemma linear_equiv_det_rotation (θ : real.angle) : - (hb.rotation θ).to_linear_equiv.det = 1 := + (b.rotation θ).to_linear_equiv.det = 1 := by simp [rotation, ←linear_isometry_equiv.to_linear_equiv_symm] /-- The inverse of `rotation` is rotation by the negation of the angle. -/ -@[simp] lemma rotation_symm (θ : real.angle) : (hb.rotation θ).symm = hb.rotation (-θ) := +@[simp] lemma rotation_symm (θ : real.angle) : (b.rotation θ).symm = b.rotation (-θ) := by simp [rotation, linear_isometry_equiv.trans_assoc] /-- Rotation by 0 is the identity. -/ -@[simp] lemma rotation_zero : hb.rotation 0 = linear_isometry_equiv.refl ℝ V := +@[simp] lemma rotation_zero : b.rotation 0 = linear_isometry_equiv.refl ℝ V := by simp [rotation] /-- Rotation by π is negation. -/ -lemma rotation_pi : hb.rotation π = linear_isometry_equiv.neg ℝ := +lemma rotation_pi : b.rotation π = linear_isometry_equiv.neg ℝ := begin ext x, simp [rotation] @@ -538,7 +537,7 @@ end /-- Rotating twice is equivalent to rotating by the sum of the angles. -/ @[simp] lemma rotation_trans (θ₁ θ₂ : real.angle) : - (hb.rotation θ₁).trans (hb.rotation θ₂) = hb.rotation (θ₂ + θ₁) := + (b.rotation θ₁).trans (b.rotation θ₂) = b.rotation (θ₂ + θ₁) := begin simp only [rotation, ←linear_isometry_equiv.trans_assoc], ext1 x, @@ -547,7 +546,7 @@ end /-- Rotating the first vector by `θ` subtracts `θ` from the angle between two vectors. -/ @[simp] lemma oangle_rotation_left {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : - hb.oangle (hb.rotation θ x) y = hb.oangle x y - θ := + b.oangle (b.rotation θ x) y = b.oangle x y - θ := begin simp [oangle, rotation, complex.arg_div_coe_angle, complex.arg_mul_coe_angle, hx, hy, ne_zero_of_mem_circle], @@ -556,7 +555,7 @@ end /-- Rotating the second vector by `θ` adds `θ` to the angle between two vectors. -/ @[simp] lemma oangle_rotation_right {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : - hb.oangle x (hb.rotation θ y) = hb.oangle x y + θ := + b.oangle x (b.rotation θ y) = b.oangle x y + θ := begin simp [oangle, rotation, complex.arg_div_coe_angle, complex.arg_mul_coe_angle, hx, hy, ne_zero_of_mem_circle], @@ -565,17 +564,17 @@ end /-- The rotation of a vector by `θ` has an angle of `-θ` from that vector. -/ @[simp] lemma oangle_rotation_self_left {x : V} (hx : x ≠ 0) (θ : real.angle) : - hb.oangle (hb.rotation θ x) x = -θ := + b.oangle (b.rotation θ x) x = -θ := by simp [hx] /-- A vector has an angle of `θ` from the rotation of that vector by `θ`. -/ @[simp] lemma oangle_rotation_self_right {x : V} (hx : x ≠ 0) (θ : real.angle) : - hb.oangle x (hb.rotation θ x) = θ := + b.oangle x (b.rotation θ x) = θ := by simp [hx] /-- Rotating the first vector by the angle between the two vectors results an an angle of 0. -/ @[simp] lemma oangle_rotation_oangle_left (x y : V) : - hb.oangle (hb.rotation (hb.oangle x y) x) y = 0 := + b.oangle (b.rotation (b.oangle x y) x) y = 0 := begin by_cases hx : x = 0, { simp [hx] }, @@ -587,7 +586,7 @@ end /-- Rotating the first vector by the angle between the two vectors and swapping the vectors results an an angle of 0. -/ @[simp] lemma oangle_rotation_oangle_right (x y : V) : - hb.oangle y (hb.rotation (hb.oangle x y) x) = 0 := + b.oangle y (b.rotation (b.oangle x y) x) = 0 := begin rw [oangle_rev], simp @@ -595,7 +594,7 @@ end /-- Rotating both vectors by the same angle does not change the angle between those vectors. -/ @[simp] lemma oangle_rotation (x y : V) (θ : real.angle) : - hb.oangle (hb.rotation θ x) (hb.rotation θ y) = hb.oangle x y := + b.oangle (b.rotation θ x) (b.rotation θ y) = b.oangle x y := begin by_cases hx : x = 0; by_cases hy : y = 0; simp [hx, hy] @@ -603,24 +602,24 @@ end /-- A rotation of a nonzero vector equals that vector if and only if the angle is zero. -/ @[simp] lemma rotation_eq_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : - hb.rotation θ x = x ↔ θ = 0 := + b.rotation θ x = x ↔ θ = 0 := begin split, { intro h, rw eq_comm, - simpa [hx, h] using hb.oangle_rotation_right hx hx θ }, + simpa [hx, h] using b.oangle_rotation_right hx hx θ }, { intro h, simp [h] } end /-- A nonzero vector equals a rotation of that vector if and only if the angle is zero. -/ @[simp] lemma eq_rotation_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : - x = hb.rotation θ x ↔ θ = 0 := -by rw [←hb.rotation_eq_self_iff_angle_eq_zero hx, eq_comm] + x = b.rotation θ x ↔ θ = 0 := +by rw [←b.rotation_eq_self_iff_angle_eq_zero hx, eq_comm] /-- A rotation of a vector equals that vector if and only if the vector or the angle is zero. -/ lemma rotation_eq_self_iff (x : V) (θ : real.angle) : - hb.rotation θ x = x ↔ x = 0 ∨ θ = 0 := + b.rotation θ x = x ↔ x = 0 ∨ θ = 0 := begin by_cases h : x = 0; simp [h] @@ -628,62 +627,62 @@ end /-- A vector equals a rotation of that vector if and only if the vector or the angle is zero. -/ lemma eq_rotation_self_iff (x : V) (θ : real.angle) : - x = hb.rotation θ x ↔ x = 0 ∨ θ = 0 := + x = b.rotation θ x ↔ x = 0 ∨ θ = 0 := by rw [←rotation_eq_self_iff, eq_comm] /-- Rotating a vector by the angle to another vector gives the second vector if and only if the norms are equal. -/ @[simp] lemma rotation_oangle_eq_iff_norm_eq (x y : V) : - hb.rotation (hb.oangle x y) x = y ↔ ∥x∥ = ∥y∥ := + b.rotation (b.oangle x y) x = y ↔ ∥x∥ = ∥y∥ := begin split, { intro h, rw [←h, linear_isometry_equiv.norm_map] }, { intro h, - rw hb.eq_iff_oangle_eq_zero_of_norm_eq; + rw b.eq_iff_oangle_eq_zero_of_norm_eq; simp [h] } end /-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first rotated by `θ` and scaled by the ratio of the norms. -/ lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) - (θ : real.angle) : hb.oangle x y = θ ↔ y = (∥y∥ / ∥x∥) • hb.rotation θ x := + (θ : real.angle) : b.oangle x y = θ ↔ y = (∥y∥ / ∥x∥) • b.rotation θ x := begin have hp := div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), split, { rintro rfl, - rw [←linear_isometry_equiv.map_smul, ←hb.oangle_smul_left_of_pos x y hp, + rw [←linear_isometry_equiv.map_smul, ←b.oangle_smul_left_of_pos x y hp, eq_comm, rotation_oangle_eq_iff_norm_eq, norm_smul, real.norm_of_nonneg hp.le, div_mul_cancel _ (norm_ne_zero_iff.2 hx)] }, { intro hye, - rw [hye, hb.oangle_smul_right_of_pos _ _ hp, hb.oangle_rotation_self_right hx] } + rw [hye, b.oangle_smul_right_of_pos _ _ hp, b.oangle_rotation_self_right hx] } end /-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first rotated by `θ` and scaled by a positive real. -/ lemma oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) - (θ : real.angle) : hb.oangle x y = θ ↔ ∃ r : ℝ, 0 < r ∧ y = r • hb.rotation θ x := + (θ : real.angle) : b.oangle x y = θ ↔ ∃ r : ℝ, 0 < r ∧ y = r • b.rotation θ x := begin split, { intro h, - rw hb.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy at h, + rw b.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy at h, exact ⟨∥y∥ / ∥x∥, div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), h⟩ }, { rintro ⟨r, hr, rfl⟩, - rw [hb.oangle_smul_right_of_pos _ _ hr, hb.oangle_rotation_self_right hx] } + rw [b.oangle_smul_right_of_pos _ _ hr, b.oangle_rotation_self_right hx] } end /-- The angle between two vectors is `θ` if and only if they are nonzero and the second vector is the first rotated by `θ` and scaled by the ratio of the norms, or `θ` and at least one of the vectors are zero. -/ lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : - hb.oangle x y = θ ↔ - (x ≠ 0 ∧ y ≠ 0 ∧ y = (∥y∥ / ∥x∥) • hb.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := + b.oangle x y = θ ↔ + (x ≠ 0 ∧ y ≠ 0 ∧ y = (∥y∥ / ∥x∥) • b.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := begin by_cases hx : x = 0, { simp [hx, eq_comm] }, { by_cases hy : y = 0, { simp [hy, eq_comm] }, - { rw hb.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy, + { rw b.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy, simp [hx, hy] } } end @@ -691,14 +690,14 @@ end is the first rotated by `θ` and scaled by a positive real, or `θ` and at least one of the vectors are zero. -/ lemma oangle_eq_iff_eq_pos_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : - hb.oangle x y = θ ↔ - (x ≠ 0 ∧ y ≠ 0 ∧ ∃ r : ℝ, 0 < r ∧ y = r • hb.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := + b.oangle x y = θ ↔ + (x ≠ 0 ∧ y ≠ 0 ∧ ∃ r : ℝ, 0 < r ∧ y = r • b.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := begin by_cases hx : x = 0, { simp [hx, eq_comm] }, { by_cases hy : y = 0, { simp [hy, eq_comm] }, - { rw hb.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero hx hy, + { rw b.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero hx hy, simp [hx, hy] } } end @@ -706,36 +705,36 @@ end depends on the choice of basis, not just on its orientation; for most geometrical purposes, the `reflection` definitions should be preferred instead. -/ def conj_lie : V ≃ₗᵢ[ℝ] V := -((complex.isometry_of_orthonormal hb).symm.trans complex.conj_lie).trans - (complex.isometry_of_orthonormal hb) +((complex.isometry_of_orthonormal b).symm.trans complex.conj_lie).trans + (complex.isometry_of_orthonormal b) /-- The determinant of `conj_lie` (as a linear map) is equal to `-1`. -/ -@[simp] lemma det_conj_lie : (hb.conj_lie.to_linear_equiv : V →ₗ[ℝ] V).det = -1 := +@[simp] lemma det_conj_lie : (b.conj_lie.to_linear_equiv : V →ₗ[ℝ] V).det = -1 := by simp [conj_lie, ←linear_isometry_equiv.to_linear_equiv_symm, ←linear_equiv.comp_coe] /-- The determinant of `conj_lie` (as a linear equiv) is equal to `-1`. -/ -@[simp] lemma linear_equiv_det_conj_lie : hb.conj_lie.to_linear_equiv.det = -1 := +@[simp] lemma linear_equiv_det_conj_lie : b.conj_lie.to_linear_equiv.det = -1 := by simp [conj_lie, ←linear_isometry_equiv.to_linear_equiv_symm] /-- `conj_lie` is its own inverse. -/ -@[simp] lemma conj_lie_symm : hb.conj_lie.symm = hb.conj_lie := +@[simp] lemma conj_lie_symm : b.conj_lie.symm = b.conj_lie := rfl /-- Applying `conj_lie` to both vectors negates the angle between those vectors. -/ @[simp] lemma oangle_conj_lie (x y : V) : - hb.oangle (hb.conj_lie x) (hb.conj_lie y) = -hb.oangle x y := -by simp only [orthonormal.conj_lie, linear_isometry_equiv.symm_apply_apply, orthonormal.oangle, - eq_self_iff_true, function.comp_app, complex.arg_coe_angle_eq_iff, + b.oangle (b.conj_lie x) (b.conj_lie y) = -b.oangle x y := +by simp only [orthonormal_basis.conj_lie, linear_isometry_equiv.symm_apply_apply, + orthonormal_basis.oangle, eq_self_iff_true, function.comp_app, complex.arg_coe_angle_eq_iff, linear_isometry_equiv.coe_trans, neg_inj, complex.conj_lie_apply, complex.arg_conj_coe_angle, ← map_div₀ (star_ring_end ℂ)] /-- Any linear isometric equivalence in `V` is `rotation` or `conj_lie` composed with `rotation`. -/ lemma exists_linear_isometry_equiv_eq (f : V ≃ₗᵢ[ℝ] V) : - ∃ θ : real.angle, f = hb.rotation θ ∨ f = hb.conj_lie.trans (hb.rotation θ) := + ∃ θ : real.angle, f = b.rotation θ ∨ f = b.conj_lie.trans (b.rotation θ) := begin - cases linear_isometry_complex (((complex.isometry_of_orthonormal hb).trans f).trans - (complex.isometry_of_orthonormal hb).symm) with a ha, + cases linear_isometry_complex (((complex.isometry_of_orthonormal b).trans f).trans + (complex.isometry_of_orthonormal b).symm) with a ha, use complex.arg a, rcases ha with (ha|ha), { left, @@ -756,9 +755,9 @@ end /-- Any linear isometric equivalence in `V` with positive determinant is `rotation`. -/ lemma exists_linear_isometry_equiv_eq_of_det_pos {f : V ≃ₗᵢ[ℝ] V} - (hd : 0 < (f.to_linear_equiv : V →ₗ[ℝ] V).det) : ∃ θ : real.angle, f = hb.rotation θ := + (hd : 0 < (f.to_linear_equiv : V →ₗ[ℝ] V).det) : ∃ θ : real.angle, f = b.rotation θ := begin - rcases hb.exists_linear_isometry_equiv_eq f with ⟨θ, (hf|hf)⟩, + rcases b.exists_linear_isometry_equiv_eq f with ⟨θ, (hf|hf)⟩, { exact ⟨θ, hf⟩ }, { simp [hf, ←linear_equiv.coe_det] at hd, norm_num at hd } @@ -768,36 +767,45 @@ end with `rotation`. -/ lemma exists_linear_isometry_equiv_eq_of_det_neg {f : V ≃ₗᵢ[ℝ] V} (hd : (f.to_linear_equiv : V →ₗ[ℝ] V).det < 0) : - ∃ θ : real.angle, f = hb.conj_lie.trans (hb.rotation θ) := + ∃ θ : real.angle, f = b.conj_lie.trans (b.rotation θ) := begin - rcases hb.exists_linear_isometry_equiv_eq f with ⟨θ, (hf|hf)⟩, + rcases b.exists_linear_isometry_equiv_eq f with ⟨θ, (hf|hf)⟩, { simp [hf, ←linear_equiv.coe_det] at hd, norm_num at hd }, { exact ⟨θ, hf⟩ } end /-- Two bases with the same orientation are related by a `rotation`. -/ -lemma exists_linear_isometry_equiv_map_eq_of_orientation_eq {b₂ : basis (fin 2) ℝ V} - (hb₂ : orthonormal ℝ b₂) (ho : b.orientation = b₂.orientation) : - ∃ θ : real.angle, b₂ = b.map (hb.rotation θ).to_linear_equiv := -begin - have h : b₂ = b.map (hb.equiv hb₂ (equiv.refl _)).to_linear_equiv, - { rw hb.map_equiv, simp }, - rw [eq_comm, h, b.orientation_comp_linear_equiv_eq_iff_det_pos] at ho, - cases hb.exists_linear_isometry_equiv_eq_of_det_pos ho with θ hθ, +lemma exists_linear_isometry_equiv_map_eq_of_orientation_eq (b₂ : orthonormal_basis (fin 2) ℝ V) + (ho : b.to_basis.orientation = b₂.to_basis.orientation) : + ∃ θ : real.angle, b₂ = b.map (b.rotation θ) := +begin + have h : b₂ = b.map (b.repr.trans b₂.repr.symm), + { cases b₂ with B₂, + cases b with B, + simp only [orthonormal_basis.map, linear_isometry_equiv.symm_trans, B.symm_trans_self, + linear_isometry_equiv.symm_symm, ← linear_isometry_equiv.trans_assoc B₂ B.symm B, + B₂.trans_refl] }, + rw [eq_comm, h, b.to_basis_map, b.to_basis.orientation_comp_linear_equiv_eq_iff_det_pos] at ho, + cases b.exists_linear_isometry_equiv_eq_of_det_pos ho with θ hθ, rw hθ at h, exact ⟨θ, h⟩ end /-- Two bases with opposite orientations are related by `conj_lie` composed with a `rotation`. -/ -lemma exists_linear_isometry_equiv_map_eq_of_orientation_eq_neg {b₂ : basis (fin 2) ℝ V} - (hb₂ : orthonormal ℝ b₂) (ho : b.orientation = -b₂.orientation) : - ∃ θ : real.angle, b₂ = b.map (hb.conj_lie.trans (hb.rotation θ)).to_linear_equiv := -begin - have h : b₂ = b.map (hb.equiv hb₂ (equiv.refl _)).to_linear_equiv, - { rw hb.map_equiv, simp }, - rw [eq_neg_iff_eq_neg, h, b.orientation_comp_linear_equiv_eq_neg_iff_det_neg] at ho, - cases hb.exists_linear_isometry_equiv_eq_of_det_neg ho with θ hθ, +lemma exists_linear_isometry_equiv_map_eq_of_orientation_eq_neg (b₂ : orthonormal_basis (fin 2) ℝ V) + (ho : b.to_basis.orientation = -b₂.to_basis.orientation) : + ∃ θ : real.angle, b₂ = b.map (b.conj_lie.trans (b.rotation θ)) := +begin + have h : b₂ = b.map (b.repr.trans b₂.repr.symm), + { cases b₂ with B₂, + cases b with B, + simp only [orthonormal_basis.map, linear_isometry_equiv.symm_trans, B.symm_trans_self, + linear_isometry_equiv.symm_symm, ← linear_isometry_equiv.trans_assoc B₂ B.symm B, + B₂.trans_refl] }, + rw [eq_neg_iff_eq_neg, h, b.to_basis_map, + b.to_basis.orientation_comp_linear_equiv_eq_neg_iff_det_neg] at ho, + cases b.exists_linear_isometry_equiv_eq_of_det_neg ho with θ hθ, rw hθ at h, exact ⟨θ, h⟩ end @@ -806,32 +814,35 @@ end isometric equivalence, equals the angle between those two vectors, transformed by the inverse of that equivalence, with respect to the original basis. -/ @[simp] lemma oangle_map (x y : V) (f : V ≃ₗᵢ[ℝ] V) : - (hb.map_linear_isometry_equiv f).oangle x y = hb.oangle (f.symm x) (f.symm y) := + (b.map f).oangle x y = b.oangle (f.symm x) (f.symm y) := by simp [oangle] /-- The value of `oangle` does not depend on the choice of basis for a given orientation. -/ -lemma oangle_eq_of_orientation_eq {b₂ : basis (fin 2) ℝ V} (hb₂ : orthonormal ℝ b₂) - (ho : b.orientation = b₂.orientation) (x y : V) : hb.oangle x y = hb₂.oangle x y := +lemma oangle_eq_of_orientation_eq (b₂ : orthonormal_basis (fin 2) ℝ V) + (ho : b.to_basis.orientation = b₂.to_basis.orientation) (x y : V) : + b.oangle x y = b₂.oangle x y := begin - obtain ⟨θ, rfl⟩ := hb.exists_linear_isometry_equiv_map_eq_of_orientation_eq hb₂ ho, - simp [hb] + obtain ⟨θ, rfl⟩ := b.exists_linear_isometry_equiv_map_eq_of_orientation_eq b₂ ho, + simp [b.orthonormal], end /-- Negating the orientation negates the value of `oangle`. -/ -lemma oangle_eq_neg_of_orientation_eq_neg {b₂ : basis (fin 2) ℝ V} (hb₂ : orthonormal ℝ b₂) - (ho : b.orientation = -b₂.orientation) (x y : V) : hb.oangle x y = -hb₂.oangle x y := +lemma oangle_eq_neg_of_orientation_eq_neg (b₂ : orthonormal_basis (fin 2) ℝ V) + (ho : b.to_basis.orientation = -b₂.to_basis.orientation) (x y : V) : + b.oangle x y = -b₂.oangle x y := begin - obtain ⟨θ, rfl⟩ := hb.exists_linear_isometry_equiv_map_eq_of_orientation_eq_neg hb₂ ho, - rw hb.oangle_map, - simp [hb] + obtain ⟨θ, rfl⟩ := b.exists_linear_isometry_equiv_map_eq_of_orientation_eq_neg b₂ ho, + rw b.oangle_map, + simp [b.orthonormal], end /-- `rotation` does not depend on the choice of basis for a given orientation. -/ -lemma rotation_eq_of_orientation_eq {b₂ : basis (fin 2) ℝ V} (hb₂ : orthonormal ℝ b₂) - (ho : b.orientation = b₂.orientation) (θ : real.angle) : hb.rotation θ = hb₂.rotation θ := +lemma rotation_eq_of_orientation_eq (b₂ : orthonormal_basis (fin 2) ℝ V) + (ho : b.to_basis.orientation = b₂.to_basis.orientation) (θ : real.angle) : + b.rotation θ = b₂.rotation θ := begin - obtain ⟨θ₂, rfl⟩ := hb.exists_linear_isometry_equiv_map_eq_of_orientation_eq hb₂ ho, - simp_rw [rotation, complex.map_isometry_of_orthonormal hb], + obtain ⟨θ₂, rfl⟩ := b.exists_linear_isometry_equiv_map_eq_of_orientation_eq b₂ ho, + simp_rw [rotation, complex.map_isometry_of_orthonormal b], simp only [linear_isometry_equiv.trans_assoc, linear_isometry_equiv.self_trans_symm, linear_isometry_equiv.refl_trans, linear_isometry_equiv.symm_trans], simp only [←linear_isometry_equiv.trans_assoc, _root_.rotation_symm, _root_.rotation_trans, @@ -839,12 +850,12 @@ begin end /-- Negating the orientation negates the angle in `rotation`. -/ -lemma rotation_eq_rotation_neg_of_orientation_eq_neg {b₂ : basis (fin 2) ℝ V} - (hb₂ : orthonormal ℝ b₂) (ho : b.orientation = -b₂.orientation) (θ : real.angle) : - hb.rotation θ = hb₂.rotation (-θ) := +lemma rotation_eq_rotation_neg_of_orientation_eq_neg (b₂ : orthonormal_basis (fin 2) ℝ V) + (ho : b.to_basis.orientation = -b₂.to_basis.orientation) (θ : real.angle) : + b.rotation θ = b₂.rotation (-θ) := begin - obtain ⟨θ₂, rfl⟩ := hb.exists_linear_isometry_equiv_map_eq_of_orientation_eq_neg hb₂ ho, - simp_rw [rotation, complex.map_isometry_of_orthonormal hb, conj_lie], + obtain ⟨θ₂, rfl⟩ := b.exists_linear_isometry_equiv_map_eq_of_orientation_eq_neg b₂ ho, + simp_rw [rotation, complex.map_isometry_of_orthonormal b, conj_lie], simp only [linear_isometry_equiv.trans_assoc, linear_isometry_equiv.self_trans_symm, linear_isometry_equiv.refl_trans, linear_isometry_equiv.symm_trans], congr' 1, @@ -864,7 +875,7 @@ end /-- The inner product of two vectors is the product of the norms and the cosine of the oriented angle between the vectors. -/ lemma inner_eq_norm_mul_norm_mul_cos_oangle (x y : V) : - ⟪x, y⟫ = ∥x∥ * ∥y∥ * real.angle.cos (hb.oangle x y) := + ⟪x, y⟫ = ∥x∥ * ∥y∥ * real.angle.cos (b.oangle x y) := begin by_cases hx : x = 0, { simp [hx] }, by_cases hy : y = 0, { simp [hy] }, @@ -875,20 +886,21 @@ begin is_R_or_C.I, complex.mul_I_re, complex.mul_I_im, complex.of_real_re, complex.of_real_im, basis.coord_apply, neg_zero, zero_add, add_zero], conv_lhs { rw [←b.sum_repr x, ←b.sum_repr y] }, - simp_rw [hb.inner_sum, (dec_trivial : (finset.univ : finset (fin 2)) = {0, 1}), + simp_rw [b.orthonormal.inner_sum, (dec_trivial : (finset.univ : finset (fin 2)) = {0, 1}), star_ring_end_apply, star_trivial], rw [finset.sum_insert (dec_trivial : (0 : fin 2) ∉ ({1} : finset (fin 2))), finset.sum_singleton], - field_simp only [norm_ne_zero_iff.2 hx, norm_ne_zero_iff.2 hy, ne.def, not_false_iff], + field_simp only [norm_ne_zero_iff.2 hx, norm_ne_zero_iff.2 hy, ne.def, not_false_iff, + orthonormal_basis.coe_to_basis_repr_apply], ring end /-- The cosine of the oriented angle between two nonzero vectors is the inner product divided by the product of the norms. -/ lemma cos_oangle_eq_inner_div_norm_mul_norm {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - real.angle.cos (hb.oangle x y) = ⟪x, y⟫ / (∥x∥ * ∥y∥) := + real.angle.cos (b.oangle x y) = ⟪x, y⟫ / (∥x∥ * ∥y∥) := begin - rw hb.inner_eq_norm_mul_norm_mul_cos_oangle, + rw b.inner_eq_norm_mul_norm_mul_cos_oangle, field_simp [norm_ne_zero_iff.2 hx, norm_ne_zero_iff.2 hy], ring end @@ -896,23 +908,23 @@ end /-- The cosine of the oriented angle between two nonzero vectors equals that of the unoriented angle. -/ lemma cos_oangle_eq_cos_angle {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - real.angle.cos (hb.oangle x y) = real.cos (inner_product_geometry.angle x y) := -by rw [hb.cos_oangle_eq_inner_div_norm_mul_norm hx hy, inner_product_geometry.cos_angle] + real.angle.cos (b.oangle x y) = real.cos (inner_product_geometry.angle x y) := +by rw [b.cos_oangle_eq_inner_div_norm_mul_norm hx hy, inner_product_geometry.cos_angle] /-- The oriented angle between two nonzero vectors is plus or minus the unoriented angle. -/ lemma oangle_eq_angle_or_eq_neg_angle {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - hb.oangle x y = inner_product_geometry.angle x y ∨ - hb.oangle x y = -inner_product_geometry.angle x y := -real.angle.cos_eq_real_cos_iff_eq_or_eq_neg.1 $ hb.cos_oangle_eq_cos_angle hx hy + b.oangle x y = inner_product_geometry.angle x y ∨ + b.oangle x y = -inner_product_geometry.angle x y := +real.angle.cos_eq_real_cos_iff_eq_or_eq_neg.1 $ b.cos_oangle_eq_cos_angle hx hy /-- The unoriented angle between two nonzero vectors is the absolute value of the oriented angle, converted to a real. -/ lemma angle_eq_abs_oangle_to_real {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - inner_product_geometry.angle x y = |(hb.oangle x y).to_real| := + inner_product_geometry.angle x y = |(b.oangle x y).to_real| := begin have h0 := inner_product_geometry.angle_nonneg x y, have hpi := inner_product_geometry.angle_le_pi x y, - rcases hb.oangle_eq_angle_or_eq_neg_angle hx hy with (h|h), + rcases b.oangle_eq_angle_or_eq_neg_angle hx hy with (h|h), { rw [h, eq_comm, real.angle.abs_to_real_coe_eq_self_iff], exact ⟨h0, hpi⟩ }, { rw [h, eq_comm, real.angle.abs_to_real_neg_coe_eq_self_iff], @@ -922,12 +934,12 @@ end /-- If the sign of the oriented angle between two vectors is zero, either one of the vectors is zero or the unoriented angle is 0 or π. -/ lemma eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero {x y : V} - (h : (hb.oangle x y).sign = 0) : + (h : (b.oangle x y).sign = 0) : x = 0 ∨ y = 0 ∨ inner_product_geometry.angle x y = 0 ∨ inner_product_geometry.angle x y = π := begin by_cases hx : x = 0, { simp [hx] }, by_cases hy : y = 0, { simp [hy] }, - rw hb.angle_eq_abs_oangle_to_real hx hy, + rw b.angle_eq_abs_oangle_to_real hx hy, rw real.angle.sign_eq_zero_iff at h, rcases h with h|h; simp [h, real.pi_pos.le] @@ -937,10 +949,10 @@ end equal, then the oriented angles are equal (even in degenerate cases). -/ lemma oangle_eq_of_angle_eq_of_sign_eq {w x y z : V} (h : inner_product_geometry.angle w x = inner_product_geometry.angle y z) - (hs : (hb.oangle w x).sign = (hb.oangle y z).sign) : hb.oangle w x = hb.oangle y z := + (hs : (b.oangle w x).sign = (b.oangle y z).sign) : b.oangle w x = b.oangle y z := begin by_cases h0 : (w = 0 ∨ x = 0) ∨ (y = 0 ∨ z = 0), - { have hs' : (hb.oangle w x).sign = 0 ∧ (hb.oangle y z).sign = 0, + { have hs' : (b.oangle w x).sign = 0 ∧ (b.oangle y z).sign = 0, { rcases h0 with (rfl|rfl)|rfl|rfl, { simpa using hs.symm }, { simpa using hs.symm }, @@ -960,80 +972,80 @@ begin { exact real.pi_pos.ne.symm hpi }, { exact two_ne_zero } }, have h0wx : (w = 0 ∨ x = 0), - { have h0' := hb.eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero hswx, + { have h0' := b.eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero hswx, simpa [hwx, real.pi_pos.ne.symm, hpi] using h0' }, have h0yz : (y = 0 ∨ z = 0), - { have h0' := hb.eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero hsyz, + { have h0' := b.eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero hsyz, simpa [hyz, real.pi_pos.ne.symm, hpi] using h0' }, rcases h0wx with h0wx|h0wx; rcases h0yz with h0yz|h0yz; simp [h0wx, h0yz] }, { push_neg at h0, rw real.angle.eq_iff_abs_to_real_eq_of_sign_eq hs, - rwa [hb.angle_eq_abs_oangle_to_real h0.1.1 h0.1.2, - hb.angle_eq_abs_oangle_to_real h0.2.1 h0.2.2] at h } + rwa [b.angle_eq_abs_oangle_to_real h0.1.1 h0.1.2, + b.angle_eq_abs_oangle_to_real h0.2.1 h0.2.2] at h } end /-- If the signs of two oriented angles between nonzero vectors are equal, the oriented angles are equal if and only if the unoriented angles are equal. -/ lemma oangle_eq_iff_angle_eq_of_sign_eq {w x y z : V} (hw : w ≠ 0) (hx : x ≠ 0) (hy : y ≠ 0) - (hz : z ≠ 0) (hs : (hb.oangle w x).sign = (hb.oangle y z).sign) : + (hz : z ≠ 0) (hs : (b.oangle w x).sign = (b.oangle y z).sign) : inner_product_geometry.angle w x = inner_product_geometry.angle y z ↔ - hb.oangle w x = hb.oangle y z := + b.oangle w x = b.oangle y z := begin - refine ⟨λ h, hb.oangle_eq_of_angle_eq_of_sign_eq h hs, λ h, _⟩, - rw [hb.angle_eq_abs_oangle_to_real hw hx, hb.angle_eq_abs_oangle_to_real hy hz, h] + refine ⟨λ h, b.oangle_eq_of_angle_eq_of_sign_eq h hs, λ h, _⟩, + rw [b.angle_eq_abs_oangle_to_real hw hx, b.angle_eq_abs_oangle_to_real hy hz, h] end /-- The oriented angle between two nonzero vectors is zero if and only if the unoriented angle is zero. -/ lemma oangle_eq_zero_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - hb.oangle x y = 0 ↔ inner_product_geometry.angle x y = 0 := + b.oangle x y = 0 ↔ inner_product_geometry.angle x y = 0 := begin refine ⟨λ h, _, λ h, _⟩, - { simpa [hb.angle_eq_abs_oangle_to_real hx hy] }, - { have ha := hb.oangle_eq_angle_or_eq_neg_angle hx hy, + { simpa [b.angle_eq_abs_oangle_to_real hx hy] }, + { have ha := b.oangle_eq_angle_or_eq_neg_angle hx hy, rw h at ha, simpa using ha } end /-- The oriented angle between two vectors is `π` if and only if the unoriented angle is `π`. -/ lemma oangle_eq_pi_iff_angle_eq_pi {x y : V} : - hb.oangle x y = π ↔ inner_product_geometry.angle x y = π := + b.oangle x y = π ↔ inner_product_geometry.angle x y = π := begin by_cases hx : x = 0, { simp [hx, real.angle.pi_ne_zero.symm, div_eq_mul_inv, mul_right_eq_self₀, not_or_distrib, real.pi_ne_zero], norm_num }, by_cases hy : y = 0, { simp [hy, real.angle.pi_ne_zero.symm, div_eq_mul_inv, mul_right_eq_self₀, not_or_distrib, real.pi_ne_zero], norm_num }, refine ⟨λ h, _, λ h, _⟩, - { rw [hb.angle_eq_abs_oangle_to_real hx hy, h], + { rw [b.angle_eq_abs_oangle_to_real hx hy, h], simp [real.pi_pos.le] }, - { have ha := hb.oangle_eq_angle_or_eq_neg_angle hx hy, + { have ha := b.oangle_eq_angle_or_eq_neg_angle hx hy, rw h at ha, simpa using ha } end /-- Negating the first vector passed to `oangle` negates the sign of the angle. -/ @[simp] lemma oangle_sign_neg_left (x y : V) : - (hb.oangle (-x) y).sign = -((hb.oangle x y).sign) := + (b.oangle (-x) y).sign = -((b.oangle x y).sign) := begin by_cases hx : x = 0, { simp [hx] }, by_cases hy : y = 0, { simp [hy] }, - rw [hb.oangle_neg_left hx hy, real.angle.sign_add_pi] + rw [b.oangle_neg_left hx hy, real.angle.sign_add_pi] end /-- Negating the second vector passed to `oangle` negates the sign of the angle. -/ @[simp] lemma oangle_sign_neg_right (x y : V) : - (hb.oangle x (-y)).sign = -((hb.oangle x y).sign) := + (b.oangle x (-y)).sign = -((b.oangle x y).sign) := begin by_cases hx : x = 0, { simp [hx] }, by_cases hy : y = 0, { simp [hy] }, - rw [hb.oangle_neg_right hx hy, real.angle.sign_add_pi] + rw [b.oangle_neg_right hx hy, real.angle.sign_add_pi] end /-- Multiplying the first vector passed to `oangle` by a real multiplies the sign of the angle by the sign of the real. -/ @[simp] lemma oangle_sign_smul_left (x y : V) (r : ℝ) : - (hb.oangle (r • x) y).sign = sign r * (hb.oangle x y).sign := + (b.oangle (r • x) y).sign = sign r * (b.oangle x y).sign := begin rcases lt_trichotomy r 0 with h|h|h; simp [h] @@ -1042,7 +1054,7 @@ end /-- Multiplying the second vector passed to `oangle` by a real multiplies the sign of the angle by the sign of the real. -/ @[simp] lemma oangle_sign_smul_right (x y : V) (r : ℝ) : - (hb.oangle x (r • y)).sign = sign r * (hb.oangle x y).sign := + (b.oangle x (r • y)).sign = sign r * (b.oangle x y).sign := begin rcases lt_trichotomy r 0 with h|h|h; simp [h] @@ -1051,8 +1063,8 @@ end /-- Auxiliary lemma for the proof of `oangle_sign_smul_add_right`; not intended to be used outside of that proof. -/ lemma oangle_smul_add_right_eq_zero_or_eq_pi_iff {x y : V} (r : ℝ) : - (hb.oangle x (r • x + y) = 0 ∨ hb.oangle x (r • x + y) = π) ↔ - (hb.oangle x y = 0 ∨ hb.oangle x y = π) := + (b.oangle x (r • x + y) = 0 ∨ b.oangle x (r • x + y) = π) ↔ + (b.oangle x y = 0 ∨ b.oangle x y = π) := begin simp_rw [oangle_eq_zero_or_eq_pi_iff_not_linear_independent, fintype.not_linear_independent_iff, fin.sum_univ_two, fin.exists_fin_two], @@ -1082,26 +1094,26 @@ end /-- Adding a multiple of the first vector passed to `oangle` to the second vector does not change the sign of the angle. -/ @[simp] lemma oangle_sign_smul_add_right (x y : V) (r : ℝ) : - (hb.oangle x (r • x + y)).sign = (hb.oangle x y).sign := + (b.oangle x (r • x + y)).sign = (b.oangle x y).sign := begin - by_cases h : hb.oangle x y = 0 ∨ hb.oangle x y = π, + by_cases h : b.oangle x y = 0 ∨ b.oangle x y = π, { rwa [real.angle.sign_eq_zero_iff.2 h, real.angle.sign_eq_zero_iff, oangle_smul_add_right_eq_zero_or_eq_pi_iff] }, - have h' : ∀ r' : ℝ, hb.oangle x (r' • x + y) ≠ 0 ∧ hb.oangle x (r' • x + y) ≠ π, + have h' : ∀ r' : ℝ, b.oangle x (r' • x + y) ≠ 0 ∧ b.oangle x (r' • x + y) ≠ π, { intro r', - rwa [←hb.oangle_smul_add_right_eq_zero_or_eq_pi_iff r', not_or_distrib] at h }, + rwa [←b.oangle_smul_add_right_eq_zero_or_eq_pi_iff r', not_or_distrib] at h }, let s : set (V × V) := (λ r' : ℝ, (x, r' • x + y)) '' set.univ, have hc : is_connected s := is_connected_univ.image _ ((continuous_const.prod_mk ((continuous_id.smul continuous_const).add continuous_const)).continuous_on), - have hf : continuous_on (λ z : V × V, hb.oangle z.1 z.2) s, - { refine continuous_at.continuous_on (λ z hz, hb.continuous_at_oangle _ _), + have hf : continuous_on (λ z : V × V, b.oangle z.1 z.2) s, + { refine continuous_at.continuous_on (λ z hz, b.continuous_at_oangle _ _), all_goals { simp_rw [s, set.mem_image] at hz, obtain ⟨r', -, rfl⟩ := hz, simp only [prod.fst, prod.snd], intro hz }, { simpa [hz] using (h' 0).1 }, { simpa [hz] using (h' r').1 } }, - have hs : ∀ z : V × V, z ∈ s → hb.oangle z.1 z.2 ≠ 0 ∧ hb.oangle z.1 z.2 ≠ π, + have hs : ∀ z : V × V, z ∈ s → b.oangle z.1 z.2 ≠ 0 ∧ b.oangle z.1 z.2 ≠ π, { intros z hz, simp_rw [s, set.mem_image] at hz, obtain ⟨r', -, rfl⟩ := hz, @@ -1116,28 +1128,28 @@ end /-- Adding a multiple of the second vector passed to `oangle` to the first vector does not change the sign of the angle. -/ @[simp] lemma oangle_sign_add_smul_left (x y : V) (r : ℝ) : - (hb.oangle (x + r • y) y).sign = (hb.oangle x y).sign := -by simp_rw [hb.oangle_rev y, real.angle.sign_neg, add_comm x, oangle_sign_smul_add_right] + (b.oangle (x + r • y) y).sign = (b.oangle x y).sign := +by simp_rw [b.oangle_rev y, real.angle.sign_neg, add_comm x, oangle_sign_smul_add_right] /-- Subtracting a multiple of the first vector passed to `oangle` from the second vector does not change the sign of the angle. -/ @[simp] lemma oangle_sign_sub_smul_right (x y : V) (r : ℝ) : - (hb.oangle x (y - r • x)).sign = (hb.oangle x y).sign := + (b.oangle x (y - r • x)).sign = (b.oangle x y).sign := by rw [sub_eq_add_neg, ←neg_smul, add_comm, oangle_sign_smul_add_right] /-- Subtracting a multiple of the second vector passed to `oangle` from the first vector does not change the sign of the angle. -/ @[simp] lemma oangle_sign_sub_smul_left (x y : V) (r : ℝ) : - (hb.oangle (x - r • y) y).sign = (hb.oangle x y).sign := + (b.oangle (x - r • y) y).sign = (b.oangle x y).sign := by rw [sub_eq_add_neg, ←neg_smul, oangle_sign_add_smul_left] /-- The sign of the angle between a vector, and a linear combination of that vector with a second vector, is the sign of the factor by which the second vector is multiplied in that combination multiplied by the sign of the angle between the two vectors. -/ @[simp] lemma oangle_sign_smul_add_smul_right (x y : V) (r₁ r₂ : ℝ) : - (hb.oangle x (r₁ • x + r₂ • y)).sign = sign r₂ * (hb.oangle x y).sign := + (b.oangle x (r₁ • x + r₂ • y)).sign = sign r₂ * (b.oangle x y).sign := begin - rw ←hb.oangle_sign_smul_add_right x (r₁ • x + r₂ • y) (-r₁), + rw ←b.oangle_sign_smul_add_right x (r₁ • x + r₂ • y) (-r₁), simp end @@ -1145,22 +1157,22 @@ end the sign of the factor by which the first vector is multiplied in that combination multiplied by the sign of the angle between the two vectors. -/ @[simp] lemma oangle_sign_smul_add_smul_left (x y : V) (r₁ r₂ : ℝ) : - (hb.oangle (r₁ • x + r₂ • y) y).sign = sign r₁ * (hb.oangle x y).sign := -by simp_rw [hb.oangle_rev y, real.angle.sign_neg, add_comm (r₁ • x), + (b.oangle (r₁ • x + r₂ • y) y).sign = sign r₁ * (b.oangle x y).sign := +by simp_rw [b.oangle_rev y, real.angle.sign_neg, add_comm (r₁ • x), oangle_sign_smul_add_smul_right, mul_neg] /-- The sign of the angle between two linear combinations of two vectors is the sign of the determinant of the factors in those combinations multiplied by the sign of the angle between the two vectors. -/ lemma oangle_sign_smul_add_smul_smul_add_smul (x y : V) (r₁ r₂ r₃ r₄ : ℝ) : - (hb.oangle (r₁ • x + r₂ • y) (r₃ • x + r₄ • y)).sign = - sign (r₁ * r₄ - r₂ * r₃) * (hb.oangle x y).sign := + (b.oangle (r₁ • x + r₂ • y) (r₃ • x + r₄ • y)).sign = + sign (r₁ * r₄ - r₂ * r₃) * (b.oangle x y).sign := begin by_cases hr₁ : r₁ = 0, { rw [hr₁, zero_smul, zero_mul, zero_add, zero_sub, left.sign_neg, oangle_sign_smul_left, add_comm, oangle_sign_smul_add_smul_right, oangle_rev, real.angle.sign_neg, sign_mul, mul_neg, mul_neg, neg_mul, mul_assoc] }, - { rw [←hb.oangle_sign_smul_add_right (r₁ • x + r₂ • y) (r₃ • x + r₄ • y) (-r₃ / r₁), + { rw [←b.oangle_sign_smul_add_right (r₁ • x + r₂ • y) (r₃ • x + r₄ • y) (-r₃ / r₁), smul_add, smul_smul, smul_smul, div_mul_cancel _ hr₁, neg_smul, ←add_assoc, add_comm (-(r₃ • x)), ←sub_eq_add_neg, sub_add_cancel, ←add_smul, oangle_sign_smul_right, oangle_sign_smul_add_smul_left, ←mul_assoc, ←sign_mul, @@ -1168,7 +1180,7 @@ begin neg_mul, ←sub_eq_add_neg, mul_comm r₄, mul_comm r₃] } end -end orthonormal +end orthonormal_basis namespace orientation @@ -1178,7 +1190,7 @@ variables {V : Type*} [inner_product_space ℝ V] variables [hd2 : fact (finrank ℝ V = 2)] (o : orientation ℝ V (fin 2)) include hd2 o -local notation `ob` := o.fin_orthonormal_basis_orthonormal dec_trivial hd2.out +local notation `ob` := o.fin_orthonormal_basis dec_trivial hd2.out /-- The oriented angle from `x` to `y`, modulo `2 * π`. If either vector is 0, this is 0. See `inner_product_geometry.angle` for the corresponding unoriented angle definition. -/ @@ -1592,17 +1604,18 @@ the inverse of that equivalence, with respect to the original orientation. -/ (orientation.map (fin 2) f.to_linear_equiv o).oangle x y = o.oangle (f.symm x) (f.symm y) := begin convert (ob).oangle_map x y f using 1, - refine orthonormal.oangle_eq_of_orientation_eq _ _ _ _ _, - simp_rw [basis.orientation_map, orientation.fin_orthonormal_basis_orientation] + refine orthonormal_basis.oangle_eq_of_orientation_eq _ _ _ _ _, + simp_rw [orthonormal_basis.to_basis_map, basis.orientation_map, + orientation.fin_orthonormal_basis_orientation], end /-- `orientation.oangle` equals `orthonormal.oangle` for any orthonormal basis with that orientation. -/ -lemma oangle_eq_basis_oangle {b : basis (fin 2) ℝ V} (hb : orthonormal ℝ b) - (h : b.orientation = o) (x y : V) : o.oangle x y = hb.oangle x y := +lemma oangle_eq_basis_oangle (b : orthonormal_basis (fin 2) ℝ V) + (h : b.to_basis.orientation = o) (x y : V) : o.oangle x y = b.oangle x y := begin rw oangle, - refine orthonormal.oangle_eq_of_orientation_eq _ _ _ _ _, + refine orthonormal_basis.oangle_eq_of_orientation_eq _ _ _ _ _, simp [h] end @@ -1610,17 +1623,17 @@ end lemma oangle_neg_orientation_eq_neg (x y : V) : (-o).oangle x y = -(o.oangle x y) := begin simp_rw oangle, - refine orthonormal.oangle_eq_neg_of_orientation_eq_neg _ _ _ _ _, + refine orthonormal_basis.oangle_eq_neg_of_orientation_eq_neg _ _ _ _ _, simp_rw orientation.fin_orthonormal_basis_orientation end /-- `orientation.rotation` equals `orthonormal.rotation` for any orthonormal basis with that orientation. -/ -lemma rotation_eq_basis_rotation {b : basis (fin 2) ℝ V} (hb : orthonormal ℝ b) - (h : b.orientation = o) (θ : ℝ) : o.rotation θ = hb.rotation θ := +lemma rotation_eq_basis_rotation (b : orthonormal_basis (fin 2) ℝ V) + (h : b.to_basis.orientation = o) (θ : ℝ) : o.rotation θ = b.rotation θ := begin rw rotation, - refine orthonormal.rotation_eq_of_orientation_eq _ _ _ _, + refine orthonormal_basis.rotation_eq_of_orientation_eq _ _ _ _, simp [h] end @@ -1629,7 +1642,7 @@ lemma rotation_neg_orientation_eq_neg (θ : real.angle) : (-o).rotation θ = o.rotation (-θ) := begin simp_rw rotation, - refine orthonormal.rotation_eq_rotation_neg_of_orientation_eq_neg _ _ _ _, + refine orthonormal_basis.rotation_eq_rotation_neg_of_orientation_eq_neg _ _ _ _, simp_rw orientation.fin_orthonormal_basis_orientation end From 866664b856176d3d9cff25e26432823c9ed47ea7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 30 Sep 2022 21:29:00 +0000 Subject: [PATCH 478/828] feat(data/polynomial/ring_division): add bUnion_roots_finite (#16670) Add three lemmas about polynomials needed for #15143 From flt-regular Co-authored-by: Damiano Testa --- src/data/polynomial/degree/definitions.lean | 14 ++++++ src/data/polynomial/ring_division.lean | 26 ++++++++++ src/topology/algebra/polynomial.lean | 55 ++++++++++++++------- 3 files changed, 77 insertions(+), 18 deletions(-) diff --git a/src/data/polynomial/degree/definitions.lean b/src/data/polynomial/degree/definitions.lean index 87273d4ba7f5f..d2d843cfeaa80 100644 --- a/src/data/polynomial/degree/definitions.lean +++ b/src/data/polynomial/degree/definitions.lean @@ -253,6 +253,20 @@ begin { rwa [degree_eq_nat_degree hp, with_bot.coe_lt_coe] } end +lemma ext_iff_nat_degree_le {p q : R[X]} {n : ℕ} (hp : p.nat_degree ≤ n) (hq : q.nat_degree ≤ n) : + p = q ↔ (∀ i ≤ n, p.coeff i = q.coeff i) := +begin + refine iff.trans polynomial.ext_iff _, + refine forall_congr (λ i, ⟨λ h _, h, λ h, _⟩), + refine (le_or_lt i n).elim h (λ k, _), + refine (coeff_eq_zero_of_nat_degree_lt (hp.trans_lt k)).trans + (coeff_eq_zero_of_nat_degree_lt (hq.trans_lt k)).symm, +end + +lemma ext_iff_degree_le {p q : R[X]} {n : ℕ} (hp : p.degree ≤ n) (hq : q.degree ≤ n) : + p = q ↔ (∀ i ≤ n, p.coeff i = q.coeff i) := +ext_iff_nat_degree_le (nat_degree_le_of_degree_le hp) (nat_degree_le_of_degree_le hq) + @[simp] lemma coeff_nat_degree_succ_eq_zero {p : R[X]} : p.coeff (p.nat_degree + 1) = 0 := coeff_eq_zero_of_nat_degree_lt (lt_add_one _) diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index 062ebffad0d13..074500c466758 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -665,6 +665,32 @@ lemma root_set_finite (p : T[X]) (S : Type*) [comm_ring S] [is_domain S] [algebra T S] : (p.root_set S).finite := set.to_finite _ +/-- The set of roots of all polynomials of bounded degree and having coefficients in a finite set +is finite. -/ +lemma bUnion_roots_finite {R S : Type*} [semiring R] [comm_ring S] [is_domain S] + (m : R →+* S) (d : ℕ) {U : set R} (h : U.finite) : + (⋃ (f : R[X]) (hf : f.nat_degree ≤ d ∧ ∀ i, (f.coeff i) ∈ U), + ((f.map m).roots.to_finset : set S)).finite := +begin + refine set.finite.bUnion _ _, + { -- We prove that the set of polynomials under consideration is finite because its + -- image by the injective map `π` is finite + let π : R[X] → finset.range (d+1) → R := λ f i, f.coeff i, + have h_inj : set.inj_on π {f : R[X] | f.nat_degree ≤ d ∧ ∀ (i : ℕ), f.coeff i ∈ U}, + { intros x hx y hy hxy, + rw ext_iff_nat_degree_le hx.1 hy.1, + exact_mod_cast λ i hi, congr_fun hxy ⟨i, finset.mem_range_succ_iff.mpr hi⟩, }, + have h_fin : (set.pi set.univ (λ e : finset.range (d+1), U)).finite := set.finite.pi (λ e, h), + refine set.finite.of_finite_image (set.finite.subset h_fin _) h_inj, + rw set.image_subset_iff, + intros f hf, + rw [set.mem_preimage, set.mem_univ_pi], + exact λ i, hf.2 i, }, + { intros i hi, + convert root_set_finite (i.map m) S, + simp only [algebra.id.map_eq_id, map_id], }, +end + theorem mem_root_set_iff' {p : T[X]} {S : Type*} [comm_ring S] [is_domain S] [algebra T S] (hp : p.map (algebra_map T S) ≠ 0) (a : S) : a ∈ p.root_set S ↔ (p.map (algebra_map T S)).eval a = 0 := diff --git a/src/topology/algebra/polynomial.lean b/src/topology/algebra/polynomial.lean index cc3aaef487808..bccc53ba4c7c8 100644 --- a/src/topology/algebra/polynomial.lean +++ b/src/topology/algebra/polynomial.lean @@ -140,12 +140,22 @@ variables {F K : Type*} [field F] [normed_field K] open multiset +lemma eq_one_of_roots_le {p : F[X]} {f : F →+* K} {B : ℝ} (hB : B < 0) + (h1 : p.monic) (h2 : splits f p) (h3 : ∀ z ∈ (map f p).roots, ∥z∥ ≤ B) : + p = 1 := +h1.nat_degree_eq_zero_iff_eq_one.mp begin + contrapose !hB, + rw nat_degree_eq_card_roots h2 at hB, + obtain ⟨z, hz⟩ := multiset.card_pos_iff_exists_mem.mp (zero_lt_iff.mpr hB), + exact le_trans (norm_nonneg _) (h3 z hz), +end + lemma coeff_le_of_roots_le {p : F[X]} {f : F →+* K} {B : ℝ} (i : ℕ) (h1 : p.monic) (h2 : splits f p) (h3 : ∀ z ∈ (map f p).roots, ∥z∥ ≤ B) : ∥ (map f p).coeff i ∥ ≤ B^(p.nat_degree - i) * p.nat_degree.choose i := begin have hcd : card (map f p).roots = p.nat_degree := (nat_degree_eq_card_roots h2).symm, - by_cases hB : 0 ≤ B, + obtain hB | hB := le_or_lt 0 B, { by_cases hi : i ≤ p.nat_degree, { rw [eq_prod_roots_of_splits h2, monic.def.mp h1, ring_hom.map_one, ring_hom.map_one, one_mul], rw prod_X_sub_C_coeff, @@ -184,23 +194,32 @@ begin rw_mod_cast mul_zero, { rwa monic.nat_degree_map h1, apply_instance, }}}, - { push_neg at hB, - have noroots : (map f p).roots = 0, - { contrapose! hB, - obtain ⟨z, hz⟩ := exists_mem_of_ne_zero hB, - exact le_trans (norm_nonneg z) (h3 z hz), }, - suffices : p.nat_degree = 0, - { by_cases hi : i = 0, - { rw [this, hi, (monic.nat_degree_eq_zero_iff_eq_one h1).mp this], - simp only [polynomial.map_one, coeff_one_zero, norm_one, pow_zero, nat.choose_self, - nat.cast_one, mul_one], }, - { replace hi := zero_lt_iff.mpr hi, - rw ←this at hi, - rw [nat.choose_eq_zero_of_lt hi, coeff_eq_zero_of_nat_degree_lt, norm_zero], - rw_mod_cast mul_zero, - { rwa monic.nat_degree_map h1, - apply_instance, }}}, - rw [←hcd, noroots, card_zero], }, + { rw [eq_one_of_roots_le hB h1 h2 h3, polynomial.map_one, nat_degree_one, zero_tsub, pow_zero, + one_mul, coeff_one], + split_ifs; norm_num [h], }, +end + +/-- The coefficients of the monic polynomials of bounded degree with bounded roots are +uniformely bounded. -/ +lemma coeff_bdd_of_roots_le {B : ℝ} {d : ℕ} (f : F →+* K) {p : F[X]} + (h1 : p.monic) (h2 : splits f p) (h3 : p.nat_degree ≤ d) (h4 : ∀ z ∈ (map f p).roots, ∥z∥ ≤ B) + (i : ℕ) : ∥(map f p).coeff i∥ ≤ (max B 1) ^ d * d.choose (d / 2) := +begin + obtain hB | hB := le_or_lt 0 B, + { apply (coeff_le_of_roots_le i h1 h2 h4).trans, + calc + _ ≤ (max B 1) ^ (p.nat_degree - i) * (p.nat_degree.choose i) + : mul_le_mul_of_nonneg_right (pow_le_pow_of_le_left hB (le_max_left _ _) _) _ + ... ≤ (max B 1) ^ d * (p.nat_degree.choose i) + : mul_le_mul_of_nonneg_right ((pow_mono (le_max_right _ _)) (le_trans (nat.sub_le _ _) h3)) _ + ... ≤ (max B 1) ^ d * d.choose (d / 2) + : mul_le_mul_of_nonneg_left (nat.cast_le.mpr ((i.choose_mono h3).trans + (i.choose_le_middle d))) _, + all_goals { positivity, }}, + { rw [eq_one_of_roots_le hB h1 h2 h4, polynomial.map_one, coeff_one], + refine trans _ (one_le_mul_of_one_le_of_one_le (one_le_pow_of_one_le (le_max_right B 1) d) _), + { split_ifs; norm_num, }, + { exact_mod_cast nat.succ_le_iff.mpr (nat.choose_pos (d.div_le_self 2)), }}, end end roots From aa701d8f25cc8868bfb9d5cc0e93f3a2dc200c64 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Fri, 30 Sep 2022 23:37:18 +0000 Subject: [PATCH 479/828] feat(algebra/module/localized_module): universal property of localized module (#15559) as $R$ module. `is_localized_module` also characterises localized module upto isomorphism as $R\_S$ module, this can be found in #16084 Co-authored-by: Andrew yang @erdOne --- src/algebra/algebra/basic.lean | 44 ++++ src/algebra/module/localized_module.lean | 301 ++++++++++++++++++++++- 2 files changed, 343 insertions(+), 2 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index 75def7eace672..c81a009eb39b1 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -476,6 +476,50 @@ lemma algebra_map_End_eq_smul_id (a : R) : ((algebra_map K (End K V)) a).ker = ⊥ := linear_map.ker_smul _ _ ha +section + +variables {R M} + +lemma End_is_unit_apply_inv_apply_of_is_unit {f : module.End R M} (h : is_unit f) (x : M) : + f (h.unit.inv x) = x := +show (f * h.unit.inv) x = x, by simp + +lemma End_is_unit_inv_apply_apply_of_is_unit {f : module.End R M} (h : is_unit f) (x : M) : + h.unit.inv (f x) = x := +(by simp : (h.unit.inv * f) x = x) + +lemma End_is_unit_iff (f : module.End R M) : + is_unit f ↔ function.bijective f := +⟨λ h, function.bijective_iff_has_inverse.mpr $ + ⟨h.unit.inv, ⟨End_is_unit_inv_apply_apply_of_is_unit h, + End_is_unit_apply_inv_apply_of_is_unit h⟩⟩, + λ H, let e : M ≃ₗ[R] M := { ..f, ..(equiv.of_bijective f H)} in + ⟨⟨_, e.symm, linear_map.ext e.right_inv, linear_map.ext e.left_inv⟩, rfl⟩⟩ + +lemma End_algebra_map_is_unit_inv_apply_eq_iff + {x : R} (h : is_unit (algebra_map R (module.End R M) x)) (m m' : M) : + h.unit⁻¹ m = m' ↔ m = x • m' := +{ mp := λ H, ((congr_arg h.unit H).symm.trans (End_is_unit_apply_inv_apply_of_is_unit h _)).symm, + mpr := λ H, H.symm ▸ + begin + apply_fun h.unit using ((module.End_is_unit_iff _).mp h).injective, + erw [End_is_unit_apply_inv_apply_of_is_unit], + refl, + end } + +lemma End_algebra_map_is_unit_inv_apply_eq_iff' + {x : R} (h : is_unit (algebra_map R (module.End R M) x)) (m m' : M) : + m' = h.unit⁻¹ m ↔ m = x • m' := +{ mp := λ H, ((congr_arg h.unit H).trans (End_is_unit_apply_inv_apply_of_is_unit h _)).symm, + mpr := λ H, H.symm ▸ + begin + apply_fun h.unit using ((module.End_is_unit_iff _).mp h).injective, + erw [End_is_unit_apply_inv_apply_of_is_unit], + refl, + end } + +end + end module namespace linear_map diff --git a/src/algebra/module/localized_module.lean b/src/algebra/module/localized_module.lean index 9a728c1798fc8..d0c6241b66ed1 100644 --- a/src/algebra/module/localized_module.lean +++ b/src/algebra/module/localized_module.lean @@ -356,8 +356,8 @@ section is_localized_module universes u v variables {R : Type u} [comm_ring R] (S : submonoid R) -variables {M M' : Type u} [add_comm_monoid M] [add_comm_monoid M'] -variables [module R M] [module R M'] (f : M →ₗ[R] M') +variables {M M' M'' : Type u} [add_comm_monoid M] [add_comm_monoid M'] [add_comm_monoid M''] +variables [module R M] [module R M'] [module R M''] (f : M →ₗ[R] M') (g : M →ₗ[R] M'') /-- The characteristic predicate for localized module. @@ -369,6 +369,118 @@ class is_localized_module : Prop := (surj [] : ∀ y : M', ∃ (x : M × S), x.2 • y = f x.1) (eq_iff_exists [] : ∀ {x₁ x₂}, f x₁ = f x₂ ↔ ∃ c : S, c • x₂ = c • x₁) +namespace localized_module + +/-- +If `g` is a linear map `M → M''` such that all scalar multiplication by `s : S` is invertible, then +there is a linear map `localized_module S M → M''`. +-/ +noncomputable def lift' (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) : + (localized_module S M) → M'' := +λ m, m.lift_on (λ p, (h $ p.2).unit⁻¹ $ g p.1) $ λ ⟨m, s⟩ ⟨m', s'⟩ ⟨c, eq1⟩, +begin + generalize_proofs h1 h2, + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, ←h2.unit⁻¹.1.map_smul], symmetry, + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff], dsimp, + have : c • s • g m' = c • s' • g m, + { erw [←g.map_smul, ←g.map_smul, ←g.map_smul, ←g.map_smul, eq1], refl, }, + have : function.injective (h c).unit.inv, + { rw function.injective_iff_has_left_inverse, refine ⟨(h c).unit, _⟩, + intros x, + change ((h c).unit.1 * (h c).unit.inv) x = x, + simp only [units.inv_eq_coe_inv, is_unit.mul_coe_inv, linear_map.one_apply], }, + apply_fun (h c).unit.inv, + erw [units.inv_eq_coe_inv, module.End_algebra_map_is_unit_inv_apply_eq_iff, + ←(h c).unit⁻¹.1.map_smul], symmetry, + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, + ←g.map_smul, ←g.map_smul, ←g.map_smul, ←g.map_smul, eq1], refl, +end + +lemma lift'_mk (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) (m : M) (s : S) : + localized_module.lift' S g h (localized_module.mk m s) = + (h s).unit⁻¹.1 (g m) := rfl + +lemma lift'_add (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) (x y) : + localized_module.lift' S g h (x + y) = + localized_module.lift' S g h x + localized_module.lift' S g h y := +localized_module.induction_on₂ begin + intros a a' b b', + erw [localized_module.lift'_mk, localized_module.lift'_mk, localized_module.lift'_mk], + dsimp, generalize_proofs h1 h2 h3, + erw [map_add, module.End_algebra_map_is_unit_inv_apply_eq_iff, + smul_add, ←h2.unit⁻¹.1.map_smul, ←h3.unit⁻¹.1.map_smul], + congr' 1; symmetry, + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, mul_smul, ←map_smul], refl, + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, mul_comm, mul_smul, ←map_smul], refl, +end x y + +lemma lift'_smul (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) + (r : R) (m) : + r • localized_module.lift' S g h m = localized_module.lift' S g h (r • m) := +m.induction_on begin + intros a b, + rw [localized_module.lift'_mk, localized_module.smul'_mk, localized_module.lift'_mk], + generalize_proofs h1 h2, + erw [←h1.unit⁻¹.1.map_smul, ←g.map_smul], +end + +/-- +If `g` is a linear map `M → M''` such that all scalar multiplication by `s : S` is invertible, then +there is a linear map `localized_module S M → M''`. +-/ +noncomputable def lift (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) : + (localized_module S M) →ₗ[R] M'' := +{ to_fun := localized_module.lift' S g h, + map_add' := localized_module.lift'_add S g h, + map_smul' := λ r x, by rw [localized_module.lift'_smul, ring_hom.id_apply] } + +/-- +If `g` is a linear map `M → M''` such that all scalar multiplication by `s : S` is invertible, then +`lift g m s = s⁻¹ • g m`. +-/ +lemma lift_mk (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) + (m : M) (s : S) : + localized_module.lift S g h (localized_module.mk m s) = (h s).unit⁻¹.1 (g m) := rfl + +/-- +If `g` is a linear map `M → M''` such that all scalar multiplication by `s : S` is invertible, then +there is a linear map `lift g ∘ mk_linear_map = g`. +-/ +lemma lift_comp (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) : + (lift S g h).comp (mk_linear_map S M) = g := +begin + ext x, dsimp, rw localized_module.lift_mk, + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, one_smul], +end + +/-- +If `g` is a linear map `M → M''` such that all scalar multiplication by `s : S` is invertible and +`l` is another linear map `localized_module S M ⟶ M''` such that `l ∘ mk_linear_map = g` then +`l = lift g` +-/ +lemma lift_unique (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) + (l : localized_module S M →ₗ[R] M'') + (hl : l.comp (localized_module.mk_linear_map S M) = g) : + localized_module.lift S g h = l := +begin + ext x, induction x using localized_module.induction_on with m s, + rw [localized_module.lift_mk], + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, ←hl, linear_map.coe_comp, function.comp_app, + localized_module.mk_linear_map_apply, ←l.map_smul, localized_module.smul'_mk], + congr' 1, rw localized_module.mk_eq, + refine ⟨1, _⟩, simp only [one_smul], refl, +end + +end localized_module + instance localized_module_is_localized_module : is_localized_module S (localized_module.mk_linear_map S M) := { map_units := λ s, ⟨⟨algebra_map R (module.End R (localized_module S M)) s, @@ -387,4 +499,189 @@ instance localized_module_is_localized_module : { mp := λ eq1, by simpa only [one_smul] using localized_module.mk_eq.mp eq1, mpr := λ ⟨c, eq1⟩, localized_module.mk_eq.mpr ⟨c, by simpa only [one_smul] using eq1⟩ } } +namespace is_localized_module + +variable [is_localized_module S f] + +/-- +If `(M', f : M ⟶ M')` satisfies universal property of localized module, there is a canonical map +`localized_module S M ⟶ M'`. +-/ +noncomputable def from_localized_module' : localized_module S M → M' := +λ p, p.lift_on (λ x, (is_localized_module.map_units f x.2).unit⁻¹ (f x.1)) +begin + rintros ⟨a, b⟩ ⟨a', b'⟩ ⟨c, eq1⟩, + dsimp, + generalize_proofs h1 h2, + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, ←h2.unit⁻¹.1.map_smul, + module.End_algebra_map_is_unit_inv_apply_eq_iff', ←linear_map.map_smul, ←linear_map.map_smul], + exact ((is_localized_module.eq_iff_exists S f).mpr ⟨c, eq1⟩).symm, +end + +@[simp] lemma from_localized_module'_mk (m : M) (s : S) : + from_localized_module' S f (localized_module.mk m s) = + (is_localized_module.map_units f s).unit⁻¹ (f m) := +rfl + +lemma from_localized_module'_add (x y : localized_module S M) : + from_localized_module' S f (x + y) = + from_localized_module' S f x + from_localized_module' S f y := +localized_module.induction_on₂ begin + intros a a' b b', + simp only [localized_module.mk_add_mk, from_localized_module'_mk], + generalize_proofs h1 h2 h3, + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, smul_add, ←h2.unit⁻¹.1.map_smul, + ←h3.unit⁻¹.1.map_smul, map_add], + congr' 1, + all_goals { erw [module.End_algebra_map_is_unit_inv_apply_eq_iff'] }, + { dsimp, erw [mul_smul, f.map_smul], refl, }, + { dsimp, erw [mul_comm, f.map_smul, mul_smul], refl, }, +end x y + +lemma from_localized_module'_smul (r : R) (x : localized_module S M) : + r • from_localized_module' S f x = from_localized_module' S f (r • x) := +localized_module.induction_on begin + intros a b, + rw [from_localized_module'_mk, localized_module.smul'_mk, from_localized_module'_mk], + generalize_proofs h1, erw [f.map_smul, h1.unit⁻¹.1.map_smul], refl, +end x + +/-- +If `(M', f : M ⟶ M')` satisfies universal property of localized module, there is a canonical map +`localized_module S M ⟶ M'`. +-/ +noncomputable def from_localized_module : localized_module S M →ₗ[R] M' := +{ to_fun := from_localized_module' S f, + map_add' := from_localized_module'_add S f, + map_smul' := λ r x, by rw [from_localized_module'_smul, ring_hom.id_apply] } + +lemma from_localized_module_mk (m : M) (s : S) : + from_localized_module S f (localized_module.mk m s) = + (is_localized_module.map_units f s).unit⁻¹ (f m) := +rfl + +lemma from_localized_module.inj : function.injective $ from_localized_module S f := +λ x y eq1, +begin + induction x using localized_module.induction_on with a b, + induction y using localized_module.induction_on with a' b', + simp only [from_localized_module_mk] at eq1, + generalize_proofs h1 h2 at eq1, + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff, ←linear_map.map_smul, + module.End_algebra_map_is_unit_inv_apply_eq_iff'] at eq1, + erw [localized_module.mk_eq, ←is_localized_module.eq_iff_exists S f, f.map_smul, f.map_smul, eq1], + refl, +end + +lemma from_localized_module.surj : function.surjective $ from_localized_module S f := +λ x, let ⟨⟨m, s⟩, eq1⟩ := is_localized_module.surj S f x in ⟨localized_module.mk m s, +by { rw [from_localized_module_mk, module.End_algebra_map_is_unit_inv_apply_eq_iff, ←eq1], refl }⟩ + +lemma from_localized_module.bij : function.bijective $ from_localized_module S f := +⟨from_localized_module.inj _ _, from_localized_module.surj _ _⟩ + +/-- +If `(M', f : M ⟶ M')` satisfies universal property of localized module, then `M'` is isomorphic to +`localized_module S M` as an `R`-module. +-/ +@[simps] noncomputable def iso : localized_module S M ≃ₗ[R] M' := +{ ..from_localized_module S f, + ..equiv.of_bijective (from_localized_module S f) $ from_localized_module.bij _ _} + +lemma iso_apply_mk (m : M) (s : S) : + iso S f (localized_module.mk m s) = (is_localized_module.map_units f s).unit⁻¹ (f m) := +rfl + +lemma iso_symm_apply_aux (m : M') : + (iso S f).symm m = localized_module.mk (is_localized_module.surj S f m).some.1 + (is_localized_module.surj S f m).some.2 := +begin + generalize_proofs _ h2, + apply_fun (iso S f) using linear_equiv.injective _, + rw [linear_equiv.apply_symm_apply], + simp only [iso_apply, linear_map.to_fun_eq_coe, from_localized_module_mk], + erw [module.End_algebra_map_is_unit_inv_apply_eq_iff', h2.some_spec], +end + +lemma iso_symm_apply' (m : M') (a : M) (b : S) (eq1 : b • m = f a) : + (iso S f).symm m = localized_module.mk a b := +(iso_symm_apply_aux S f m).trans $ localized_module.mk_eq.mpr $ +begin + generalize_proofs h1, + erw [←is_localized_module.eq_iff_exists S f, f.map_smul, f.map_smul, ←h1.some_spec, ←mul_smul, + mul_comm, mul_smul, eq1], +end + +lemma iso_symm_comp : (iso S f).symm.to_linear_map.comp f = localized_module.mk_linear_map S M := +begin + ext m, rw [linear_map.comp_apply, localized_module.mk_linear_map_apply], + change (iso S f).symm _ = _, rw [iso_symm_apply'], exact one_smul _ _, +end + +/-- +If `M'` is a localized module and `g` is a linear map `M' → M''` such that all scalar multiplication +by `s : S` is invertible, then there is a linear map `M' → M''`. +-/ +noncomputable def lift (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) : + M' →ₗ[R] M'' := +(localized_module.lift S g h).comp (iso S f).symm.to_linear_map + +lemma lift_comp (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) : + (lift S f g h).comp f = g := +begin + dunfold is_localized_module.lift, + rw [linear_map.comp_assoc], + convert localized_module.lift_comp S g h, + exact iso_symm_comp _ _, +end + +lemma lift_unique (g : M →ₗ[R] M'') + (h : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) + (l : M' →ₗ[R] M'') (hl : l.comp f = g) : + lift S f g h = l := +begin + dunfold is_localized_module.lift, + rw [localized_module.lift_unique S g h (l.comp (iso S f).to_linear_map), linear_map.comp_assoc, + show (iso S f).to_linear_map.comp (iso S f).symm.to_linear_map = linear_map.id, from _, + linear_map.comp_id], + { rw [linear_equiv.comp_to_linear_map_symm_eq, linear_map.id_comp], }, + { rw [linear_map.comp_assoc, ←hl], congr' 1, ext x, + erw [from_localized_module_mk, module.End_algebra_map_is_unit_inv_apply_eq_iff, one_smul], }, +end + +/-- +Universal property from localized module: +If `(M', f : M ⟶ M')` is a localized module then it satisfies the following universal property: +For every `R`-module `M''` which every `s : S`-scalar multiplication is invertible and for every +`R`-linear map `g : M ⟶ M''`, there is a unique `R`-linear map `l : M' ⟶ M''` such that +`l ∘ f = g`. +``` +M -----f----> M' +| / +|g / +| / l +v / +M'' +``` +-/ +lemma is_universal : + ∀ (g : M →ₗ[R] M'') (map_unit : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)), + ∃! (l : M' →ₗ[R] M''), l.comp f = g := +λ g h, ⟨lift S f g h, lift_comp S f g h, λ l hl, (lift_unique S f g h l hl).symm⟩ + +lemma ring_hom_ext (map_unit : ∀ (x : S), is_unit ((algebra_map R (module.End R M'')) x)) + ⦃j k : M' →ₗ[R] M''⦄ (h : j.comp f = k.comp f) : j = k := +by { rw [←lift_unique S f (k.comp f) map_unit j h, lift_unique], refl } + +/-- +If `(M', f)` and `(M'', g)` both satisfy universal property of localized module, then `M', M''` +are isomorphic as `R`-module +-/ +noncomputable def linear_equiv [is_localized_module S g] : M' ≃ₗ[R] M'' := +(iso S f).symm.trans (iso S g) + +end is_localized_module + end is_localized_module From e4ee4e30418efcb8cf304ba76ad653aeec04ba6e Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Fri, 30 Sep 2022 23:37:19 +0000 Subject: [PATCH 480/828] feat(category_theory/limits): colimits from finite colimits and filtered colimits (#16373) We will use this in the future to show that if C has finite colimits, then Ind(C) is cocomplete. --- src/category_theory/functor/flat.lean | 5 +- .../limits/constructions/filtered.lean | 88 +++++++++++++++++ src/category_theory/limits/filtered.lean | 47 +++++++++ src/category_theory/limits/opposites.lean | 99 +++++++++++++++---- 4 files changed, 218 insertions(+), 21 deletions(-) create mode 100644 src/category_theory/limits/constructions/filtered.lean create mode 100644 src/category_theory/limits/filtered.lean diff --git a/src/category_theory/functor/flat.lean b/src/category_theory/functor/flat.lean index 675ccdea65038..501fea2e5b457 100644 --- a/src/category_theory/functor/flat.lean +++ b/src/category_theory/functor/flat.lean @@ -169,8 +169,7 @@ variables {C : Type u₁} [category.{v₁} C] {D : Type u₂} [category.{v₁} D local attribute [instance] has_finite_limits_of_has_finite_limits_of_size -@[priority 100] -instance cofiltered_of_has_finite_limits [has_finite_limits C] : is_cofiltered C := +lemma cofiltered_of_has_finite_limits [has_finite_limits C] : is_cofiltered C := { cocone_objs := λ A B, ⟨limits.prod A B, limits.prod.fst, limits.prod.snd, trivial⟩, cocone_maps := λ A B f g, ⟨equalizer f g, equalizer.ι f g, equalizer.condition f g⟩, nonempty := ⟨⊤_ C⟩ } @@ -183,7 +182,7 @@ begin apply has_finite_limits_of_has_finite_limits_of_size.{v₁} (structured_arrow X F), intros J sJ fJ, resetI, constructor end, - apply_instance + exact cofiltered_of_has_finite_limits end⟩ namespace preserves_finite_limits_of_flat diff --git a/src/category_theory/limits/constructions/filtered.lean b/src/category_theory/limits/constructions/filtered.lean new file mode 100644 index 0000000000000..1836ccd9793c5 --- /dev/null +++ b/src/category_theory/limits/constructions/filtered.lean @@ -0,0 +1,88 @@ +/- +Copyright (c) 2022 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import category_theory.limits.constructions.limits_of_products_and_equalizers +import category_theory.limits.opposites + +/-! +# Constructing colimits from finite colimits and filtered colimits + +We construct colimits of size `w` from finite colimits and filtered colimits of size `w`. Since +`w`-sized colimits are constructured from coequalizers and `w`-sized coproducts, it suffices to +construct `w`-sized coproducts from finite coproducts and `w`-sized filtered colimits. + +The idea is simple: to construct coproducts of shape `α`, we take the colimit of the filtered +diagram of all coproducts of finite subsets of `α`. + +We also deduce the dual statement by invoking the original statement in `Cᵒᵖ`. +-/ + +universes w v u + +noncomputable theory + +open category_theory + +variables {C : Type u} [category.{v} C] {α : Type w} + +namespace category_theory.limits + +namespace coproducts_from_finite_filtered +local attribute [tidy] tactic.case_bash + +/-- If `C` has finite coproducts, a functor `discrete α ⥤ C` lifts to a functor + `finset (discrete α) ⥤ C` by taking coproducts. -/ +@[simps] +def lift_to_finset [has_finite_coproducts C] (F : discrete α ⥤ C) : finset (discrete α) ⥤ C := +{ obj := λ s, ∐ λ x : s, F.obj x, + map := λ s t h, sigma.desc (λ y, sigma.ι (λ x : t, F.obj x) ⟨y, h.down.down y.2⟩) } + +/-- If `C` has finite coproducts and filtered colimits, we can construct arbitrary coproducts by + taking the colimit of the diagram formed by the coproducts of finite sets over the indexing + type. -/ +@[simps] +def lift_to_finset_colimit_cocone [has_finite_coproducts C] [has_filtered_colimits_of_size.{w w} C] + [decidable_eq α] (F : discrete α ⥤ C) : colimit_cocone F := +{ cocone := + { X := colimit (lift_to_finset F), + ι := discrete.nat_trans $ λ j, + @sigma.ι _ _ _ (λ x : ({j} : finset (discrete α)), F.obj x) _ ⟨j, by simp⟩ ≫ + colimit.ι (lift_to_finset F) {j} }, + is_colimit := + { desc := λ s, colimit.desc (lift_to_finset F) + { X := s.X, + ι := { app := λ t, sigma.desc (λ x, s.ι.app x) } }, + uniq' := λ s m h, + begin + ext t ⟨⟨j, hj⟩⟩, + convert h j using 1, + { simp [← colimit.w (lift_to_finset F) ⟨⟨finset.singleton_subset_iff.2 hj⟩⟩], refl }, + { tidy } + end } } + +end coproducts_from_finite_filtered + +open coproducts_from_finite_filtered + +lemma has_coproducts_of_finite_and_filtered [has_finite_coproducts C] + [has_filtered_colimits_of_size.{w w} C] : has_coproducts.{w} C := +λ α, by { classical, exactI ⟨λ F, has_colimit.mk (lift_to_finset_colimit_cocone F)⟩ } + +lemma has_colimits_of_finite_and_filtered [has_finite_colimits C] + [has_filtered_colimits_of_size.{w w} C] : has_colimits_of_size.{w w} C := +have has_coproducts.{w} C, from has_coproducts_of_finite_and_filtered, +by exactI has_colimits_of_has_coequalizers_and_coproducts + +lemma has_products_of_finite_and_cofiltered [has_finite_products C] + [has_cofiltered_limits_of_size.{w w} C] : has_products.{w} C := +have has_coproducts.{w} Cᵒᵖ, from has_coproducts_of_finite_and_filtered, +by exactI has_products_of_opposite + +lemma has_limits_of_finite_and_cofiltered [has_finite_limits C] + [has_cofiltered_limits_of_size.{w w} C] : has_limits_of_size.{w w} C := +have has_products.{w} C, from has_products_of_finite_and_cofiltered, +by exactI has_limits_of_has_equalizers_and_products + +end category_theory.limits diff --git a/src/category_theory/limits/filtered.lean b/src/category_theory/limits/filtered.lean new file mode 100644 index 0000000000000..63a6b86d47597 --- /dev/null +++ b/src/category_theory/limits/filtered.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2022 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import category_theory.filtered +import category_theory.limits.has_limits + +/-! +# Possession of filtered colimits +-/ + +universes w' w v u + +noncomputable theory + +open category_theory + +variables {C : Type u} [category.{v} C] + +namespace category_theory.limits + +section +variables (C) + +/-- Class for having all cofiltered limits of a given size. -/ +class has_cofiltered_limits_of_size : Prop := +(has_limits_of_shape : Π (I : Type w) [category.{w'} I] [is_cofiltered I], has_limits_of_shape I C) + +/-- Class for having all filtered colimits of a given size. -/ +class has_filtered_colimits_of_size : Prop := +(has_colimits_of_shape : Π (I : Type w) [category.{w'} I] [is_filtered I], + has_colimits_of_shape I C) + +end + +@[priority 100] +instance has_limits_of_shape_of_has_cofiltered_limits [has_cofiltered_limits_of_size.{w' w} C] + (I : Type w) [category.{w'} I] [is_cofiltered I] : has_limits_of_shape I C := +has_cofiltered_limits_of_size.has_limits_of_shape _ + +@[priority 100] +instance has_colimits_of_shape_of_has_filtered_colimits [has_filtered_colimits_of_size.{w' w} C] + (I : Type w) [category.{w'} I] [is_filtered I] : has_colimits_of_shape I C := +has_filtered_colimits_of_size.has_colimits_of_shape _ + +end category_theory.limits diff --git a/src/category_theory/limits/opposites.lean b/src/category_theory/limits/opposites.lean index 0b0e88529db86..8a8fdd76fce99 100644 --- a/src/category_theory/limits/opposites.lean +++ b/src/category_theory/limits/opposites.lean @@ -3,6 +3,7 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Floris van Doorn -/ +import category_theory.limits.filtered import category_theory.limits.shapes.finite_products import category_theory.discrete_category import tactic.equiv_rw @@ -216,6 +217,11 @@ has_limit.mk { cone := cone_of_cocone_left_op (colimit.cocone F.left_op), is_limit := is_limit_cone_of_cocone_left_op _ (colimit.is_colimit _) } +lemma has_limit_of_has_colimit_op (F : J ⥤ C) [has_colimit F.op] : has_limit F := +has_limit.mk +{ cone := (colimit.cocone F.op).unop, + is_limit := is_limit_cocone_unop _ (colimit.is_colimit _) } + /-- If `C` has colimits of shape `Jᵒᵖ`, we can construct limits in `Cᵒᵖ` of shape `J`. -/ @@ -223,6 +229,10 @@ lemma has_limits_of_shape_op_of_has_colimits_of_shape [has_colimits_of_shape J has_limits_of_shape J Cᵒᵖ := { has_limit := λ F, has_limit_of_has_colimit_left_op F } +lemma has_limits_of_shape_of_has_colimits_of_shape_op [has_colimits_of_shape Jᵒᵖ Cᵒᵖ] : + has_limits_of_shape J C := +{ has_limit := λ F, has_limit_of_has_colimit_op F } + local attribute [instance] has_limits_of_shape_op_of_has_colimits_of_shape /-- @@ -230,6 +240,17 @@ If `C` has colimits, we can construct limits for `Cᵒᵖ`. -/ instance has_limits_op_of_has_colimits [has_colimits C] : has_limits Cᵒᵖ := ⟨infer_instance⟩ +lemma has_limits_of_has_colimits_op [has_colimits Cᵒᵖ] : has_limits C := +{ has_limits_of_shape := λ J hJ, by exactI has_limits_of_shape_of_has_colimits_of_shape_op } + +instance has_cofiltered_limits_op_of_has_filtered_colimits + [has_filtered_colimits_of_size.{v₂ u₂} C] : has_cofiltered_limits_of_size.{v₂ u₂} Cᵒᵖ := +{ has_limits_of_shape := λ I hI₁ hI₂, by exactI has_limits_of_shape_op_of_has_colimits_of_shape } + +lemma has_cofiltered_limits_of_has_filtered_colimits_op + [has_filtered_colimits_of_size.{v₂ u₂} Cᵒᵖ] : has_cofiltered_limits_of_size.{v₂ u₂} C := +{ has_limits_of_shape := λ I hI₂ hI₂, by exactI has_limits_of_shape_of_has_colimits_of_shape_op } + /-- If `F.left_op : Jᵒᵖ ⥤ C` has a limit, we can construct a colimit for `F : J ⥤ Cᵒᵖ`. -/ @@ -238,25 +259,43 @@ has_colimit.mk { cocone := cocone_of_cone_left_op (limit.cone F.left_op), is_colimit := is_colimit_cocone_of_cone_left_op _ (limit.is_limit _) } +lemma has_colimit_of_has_limit_op (F : J ⥤ C) [has_limit F.op] : has_colimit F := +has_colimit.mk +{ cocone := (limit.cone F.op).unop, + is_colimit := is_colimit_cone_unop _ (limit.is_limit _) } + /-- If `C` has colimits of shape `Jᵒᵖ`, we can construct limits in `Cᵒᵖ` of shape `J`. -/ -lemma has_colimits_of_shape_op_of_has_limits_of_shape [has_limits_of_shape Jᵒᵖ C] : +instance has_colimits_of_shape_op_of_has_limits_of_shape [has_limits_of_shape Jᵒᵖ C] : has_colimits_of_shape J Cᵒᵖ := { has_colimit := λ F, has_colimit_of_has_limit_left_op F } -local attribute [instance] has_colimits_of_shape_op_of_has_limits_of_shape +lemma has_colimits_of_shape_of_has_limits_of_shape_op [has_limits_of_shape Jᵒᵖ Cᵒᵖ] : + has_colimits_of_shape J C := +{ has_colimit := λ F, has_colimit_of_has_limit_op F } /-- If `C` has limits, we can construct colimits for `Cᵒᵖ`. -/ instance has_colimits_op_of_has_limits [has_limits C] : has_colimits Cᵒᵖ := ⟨infer_instance⟩ -variables (X : Type v₁) +lemma has_colimits_of_has_limits_op [has_limits Cᵒᵖ] : has_colimits C := +{ has_colimits_of_shape := λ J hJ, by exactI has_colimits_of_shape_of_has_limits_of_shape_op } + +instance has_filtered_colimits_op_of_has_cofiltered_limits + [has_cofiltered_limits_of_size.{v₂ u₂} C] : has_filtered_colimits_of_size.{v₂ u₂} Cᵒᵖ := +{ has_colimits_of_shape := λ I hI₁ hI₂, by exactI infer_instance } + +lemma has_filtered_colimits_of_has_cofiltered_limits_op + [has_cofiltered_limits_of_size.{v₂ u₂} Cᵒᵖ] : has_filtered_colimits_of_size.{v₂ u₂} C := +{ has_colimits_of_shape := λ I hI₁ hI₂, by exactI has_colimits_of_shape_of_has_limits_of_shape_op } + +variables (X : Type v₂) /-- If `C` has products indexed by `X`, then `Cᵒᵖ` has coproducts indexed by `X`. -/ -instance has_coproducts_opposite [has_products_of_shape X C] : +instance has_coproducts_of_shape_opposite [has_products_of_shape X C] : has_coproducts_of_shape X Cᵒᵖ := begin haveI : has_limits_of_shape (discrete X)ᵒᵖ C := @@ -264,10 +303,18 @@ begin apply_instance end +lemma has_coproducts_of_shape_of_opposite [has_products_of_shape X Cᵒᵖ] : + has_coproducts_of_shape X C := +begin + haveI : has_limits_of_shape (discrete X)ᵒᵖ Cᵒᵖ := + has_limits_of_shape_of_equivalence (discrete.opposite X).symm, + exact has_colimits_of_shape_of_has_limits_of_shape_op +end + /-- If `C` has coproducts indexed by `X`, then `Cᵒᵖ` has products indexed by `X`. -/ -instance has_products_opposite [has_coproducts_of_shape X C] : +instance has_products_of_shape_opposite [has_coproducts_of_shape X C] : has_products_of_shape X Cᵒᵖ := begin haveI : has_colimits_of_shape (discrete X)ᵒᵖ C := @@ -275,21 +322,37 @@ begin apply_instance end +lemma has_products_of_shape_of_opposite [has_coproducts_of_shape X Cᵒᵖ] : + has_products_of_shape X C := +begin + haveI : has_colimits_of_shape (discrete X)ᵒᵖ Cᵒᵖ := + has_colimits_of_shape_of_equivalence (discrete.opposite X).symm, + exact has_limits_of_shape_of_has_colimits_of_shape_op +end + +instance has_products_opposite [has_coproducts.{v₂} C] : has_products.{v₂} Cᵒᵖ := +λ X, infer_instance + +lemma has_products_of_opposite [has_coproducts.{v₂} Cᵒᵖ] : has_products.{v₂} C := +λ X, has_products_of_shape_of_opposite X + +instance has_coproducts_opposite [has_products.{v₂} C] : has_coproducts.{v₂} Cᵒᵖ := +λ X, infer_instance + +lemma has_coproducts_of_opposite [has_products.{v₂} Cᵒᵖ] : has_coproducts.{v₂} C := +λ X, has_coproducts_of_shape_of_opposite X + instance has_finite_coproducts_opposite [has_finite_products C] : has_finite_coproducts Cᵒᵖ := -{ out := λ J 𝒟, begin - resetI, - haveI : has_limits_of_shape (discrete J)ᵒᵖ C := - has_limits_of_shape_of_equivalence (discrete.opposite J).symm, - apply_instance, - end } +{ out := λ J _, by exactI infer_instance } + +lemma has_finite_coproducts_of_opposite [has_finite_products Cᵒᵖ] : has_finite_coproducts C := +{ out := λ J _, by exactI has_coproducts_of_shape_of_opposite J } instance has_finite_products_opposite [has_finite_coproducts C] : has_finite_products Cᵒᵖ := -{ out := λ J 𝒟, begin - resetI, - haveI : has_colimits_of_shape (discrete J)ᵒᵖ C := - has_colimits_of_shape_of_equivalence (discrete.opposite J).symm, - apply_instance, - end } +{ out := λ J _, by exactI infer_instance } + +lemma has_finite_products_of_opposite [has_finite_coproducts Cᵒᵖ] : has_finite_products C := +{ out := λ J _, by exactI has_products_of_shape_of_opposite J } instance has_equalizers_opposite [has_coequalizers C] : has_equalizers Cᵒᵖ := begin @@ -324,7 +387,7 @@ instance has_pushouts_opposite [has_pullbacks C] : has_pushouts Cᵒᵖ := begin haveI : has_limits_of_shape walking_spanᵒᵖ C := has_limits_of_shape_of_equivalence walking_span_op_equiv.symm, - apply has_colimits_of_shape_op_of_has_limits_of_shape, + apply_instance end /-- The canonical isomorphism relating `span f.op g.op` and `(cospan f g).op` -/ From c1a28cb2174d8efd91427c0eb69302fd6659e3c8 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Fri, 30 Sep 2022 23:37:20 +0000 Subject: [PATCH 481/828] feat(topology/semicontinuous): characterization by closed sets of semicontinuous maps to a linear order (#16442) --- src/topology/semicontinuous.lean | 40 +++++++++++++++++++++++++++----- 1 file changed, 34 insertions(+), 6 deletions(-) diff --git a/src/topology/semicontinuous.lean b/src/topology/semicontinuous.lean index bec1d2dd141c2..e6558228ef15a 100644 --- a/src/topology/semicontinuous.lean +++ b/src/topology/semicontinuous.lean @@ -226,16 +226,30 @@ end /-! #### Relationship with continuity -/ -theorem lower_semicontinuous_iff_is_open : +theorem lower_semicontinuous_iff_is_open_preimage : lower_semicontinuous f ↔ ∀ y, is_open (f ⁻¹' (Ioi y)) := ⟨λ H y, is_open_iff_mem_nhds.2 (λ x hx, H x y hx), λ H x y y_lt, is_open.mem_nhds (H y) y_lt⟩ lemma lower_semicontinuous.is_open_preimage (hf : lower_semicontinuous f) (y : β) : is_open (f ⁻¹' (Ioi y)) := -lower_semicontinuous_iff_is_open.1 hf y +lower_semicontinuous_iff_is_open_preimage.1 hf y section -variables {γ : Type*} [linear_order γ] [topological_space γ] [order_topology γ] +variables {γ : Type*} [linear_order γ] + +theorem lower_semicontinuous_iff_is_closed_preimage {f : α → γ} : + lower_semicontinuous f ↔ ∀ y, is_closed (f ⁻¹' (Iic y)) := +begin + rw lower_semicontinuous_iff_is_open_preimage, + congrm (∀ y, (_ : Prop)), + rw [← is_open_compl_iff, ← preimage_compl, compl_Iic] +end + +lemma lower_semicontinuous.is_closed_preimage {f : α → γ} (hf : lower_semicontinuous f) (y : γ) : + is_closed (f ⁻¹' (Iic y)) := +lower_semicontinuous_iff_is_closed_preimage.1 hf y + +variables [topological_space γ] [order_topology γ] lemma continuous_within_at.lower_semicontinuous_within_at {f : α → γ} (h : continuous_within_at f s x) : lower_semicontinuous_within_at f s x := @@ -665,16 +679,30 @@ end /-! #### Relationship with continuity -/ -theorem upper_semicontinuous_iff_is_open : +theorem upper_semicontinuous_iff_is_open_preimage : upper_semicontinuous f ↔ ∀ y, is_open (f ⁻¹' (Iio y)) := ⟨λ H y, is_open_iff_mem_nhds.2 (λ x hx, H x y hx), λ H x y y_lt, is_open.mem_nhds (H y) y_lt⟩ lemma upper_semicontinuous.is_open_preimage (hf : upper_semicontinuous f) (y : β) : is_open (f ⁻¹' (Iio y)) := -upper_semicontinuous_iff_is_open.1 hf y +upper_semicontinuous_iff_is_open_preimage.1 hf y section -variables {γ : Type*} [linear_order γ] [topological_space γ] [order_topology γ] +variables {γ : Type*} [linear_order γ] + +theorem upper_semicontinuous_iff_is_closed_preimage {f : α → γ} : + upper_semicontinuous f ↔ ∀ y, is_closed (f ⁻¹' (Ici y)) := +begin + rw upper_semicontinuous_iff_is_open_preimage, + congrm (∀ y, (_ : Prop)), + rw [← is_open_compl_iff, ← preimage_compl, compl_Ici] +end + +lemma upper_semicontinuous.is_closed_preimage {f : α → γ} (hf : upper_semicontinuous f) (y : γ) : + is_closed (f ⁻¹' (Ici y)) := +upper_semicontinuous_iff_is_closed_preimage.1 hf y + +variables [topological_space γ] [order_topology γ] lemma continuous_within_at.upper_semicontinuous_within_at {f : α → γ} (h : continuous_within_at f s x) : upper_semicontinuous_within_at f s x := From a753abcfb8123921ef18b59264ef74452ad18722 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 30 Sep 2022 23:37:21 +0000 Subject: [PATCH 482/828] feat(linear_algebra/{basis, determinant, orientation}): determinant of `adjust_to_orientation` of a basis (#16476) Performing the operation `adjust_to_orientation` on a basis either preserves the determinant with respect to that basis, or negates it. Formalized as part of the Sphere Eversion project. --- src/linear_algebra/basis.lean | 25 ++++++++++++++++++++++--- src/linear_algebra/determinant.lean | 17 +++++++++++++++-- src/linear_algebra/orientation.lean | 22 +++++++++++++++++++++- 3 files changed, 58 insertions(+), 6 deletions(-) diff --git a/src/linear_algebra/basis.lean b/src/linear_algebra/basis.lean index f32ebb284f918..9e2cbc63c2bae 100644 --- a/src/linear_algebra/basis.lean +++ b/src/linear_algebra/basis.lean @@ -68,7 +68,7 @@ universe u open function set submodule open_locale classical big_operators -variables {ι : Type*} {ι' : Type*} {R : Type*} {K : Type*} +variables {ι : Type*} {ι' : Type*} {R : Type*} {R₂ : Type*} {K : Type*} variables {M : Type*} {M' M'' : Type*} {V : Type u} {V' : Type*} section module @@ -859,8 +859,8 @@ section module open linear_map variables {v : ι → M} -variables [ring R] [add_comm_group M] [add_comm_group M'] [add_comm_group M''] -variables [module R M] [module R M'] [module R M''] +variables [ring R] [comm_ring R₂] [add_comm_group M] [add_comm_group M'] [add_comm_group M''] +variables [module R M] [module R₂ M] [module R M'] [module R M''] variables {c d : R} {x y : M} variables (b : basis ι R M) @@ -1020,6 +1020,25 @@ lemma units_smul_apply {v : basis ι R M} {w : ι → Rˣ} (i : ι) : mk_apply (v.linear_independent.units_smul w) (units_smul_span_eq_top v.span_eq).ge i +@[simp] lemma coord_units_smul (e : basis ι R₂ M) (w : ι → R₂ˣ) (i : ι) : + (e.units_smul w).coord i = (w i)⁻¹ • e.coord i := +begin + apply e.ext, + intros j, + transitivity ((e.units_smul w).coord i) ((w j)⁻¹ • (e.units_smul w) j), + { congr, + simp [basis.units_smul, ← mul_smul], }, + simp only [basis.coord_apply, linear_map.smul_apply, basis.repr_self, units.smul_def, + smul_hom_class.map_smul, finsupp.single_apply], + split_ifs with h h, + { simp [h] }, + { simp } +end + +@[simp] lemma repr_units_smul (e : basis ι R₂ M) (w : ι → R₂ˣ) (v : M) (i : ι) : + (e.units_smul w).repr v i = (w i)⁻¹ • e.repr v i := +congr_arg (λ f : M →ₗ[R₂] R₂, f v) (e.coord_units_smul w i) + /-- A version of `smul_of_units` that uses `is_unit`. -/ def is_unit_smul (v : basis ι R M) {w : ι → R} (hw : ∀ i, is_unit (w i)): basis ι R M := diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index 2d4b4870ca51d..3894ad659cc10 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -550,11 +550,24 @@ begin exact e.det.map_eq_zero_of_eq _ (by simp [hik, function.update_apply]) hik, }, end +/-- If a basis is multiplied columnwise by scalars `w : ι → Rˣ`, then the determinant with respect +to this basis is multiplied by the product of the inverse of these scalars. -/ +lemma basis.det_units_smul (e : basis ι R M) (w : ι → Rˣ) : + (e.units_smul w).det = (↑(∏ i, w i)⁻¹ : R) • e.det := +begin + ext f, + change matrix.det (λ i j, (e.units_smul w).repr (f j) i) + = (↑(∏ i, w i)⁻¹ : R) • matrix.det (λ i j, e.repr (f j) i), + simp only [e.repr_units_smul], + convert matrix.det_mul_column (λ i, (↑((w i)⁻¹) : R)) (λ i j, e.repr (f j) i), + simp [← finset.prod_inv_distrib] +end + /-- The determinant of a basis constructed by `units_smul` is the product of the given units. -/ -@[simp] lemma basis.det_units_smul (w : ι → Rˣ) : e.det (e.units_smul w) = ∏ i, w i := +@[simp] lemma basis.det_units_smul_self (w : ι → Rˣ) : e.det (e.units_smul w) = ∏ i, w i := by simp [basis.det_apply] /-- The determinant of a basis constructed by `is_unit_smul` is the product of the given units. -/ @[simp] lemma basis.det_is_unit_smul {w : ι → R} (hw : ∀ i, is_unit (w i)) : e.det (e.is_unit_smul hw) = ∏ i, w i := -e.det_units_smul _ +e.det_units_smul_self _ diff --git a/src/linear_algebra/orientation.lean b/src/linear_algebra/orientation.lean index bb3a10db19672..88990c4ec8947 100644 --- a/src/linear_algebra/orientation.lean +++ b/src/linear_algebra/orientation.lean @@ -108,7 +108,7 @@ lemma orientation_units_smul [nontrivial R] (e : basis ι R M) (w : ι → units (e.units_smul w).orientation = (∏ i, w i)⁻¹ • e.orientation := begin rw [basis.orientation, basis.orientation, smul_ray_of_ne_zero, ray_eq_iff, - e.det.eq_smul_basis_det (e.units_smul w), det_units_smul, units.smul_def, smul_smul], + e.det.eq_smul_basis_det (e.units_smul w), det_units_smul_self, units.smul_def, smul_smul], norm_cast, simp end @@ -207,6 +207,26 @@ begin simp [units_smul_apply, hi] } end +lemma det_adjust_to_orientation [nontrivial R] [nonempty ι] (e : basis ι R M) + (x : orientation R M ι) : + (e.adjust_to_orientation x).det = e.det ∨ (e.adjust_to_orientation x).det = - e.det := +begin + dsimp [basis.adjust_to_orientation], + split_ifs, + { left, + refl }, + { right, + simp [e.det_units_smul, ← units.coe_prod, finset.prod_update_of_mem] } +end + +@[simp] lemma abs_det_adjust_to_orientation [nontrivial R] [nonempty ι] (e : basis ι R M) + (x : orientation R M ι) (v : ι → M) : + |(e.adjust_to_orientation x).det v| = |e.det v| := +begin + cases e.det_adjust_to_orientation x with h h; + simp [h] +end + end basis end linear_ordered_comm_ring From 7aebb349fbbbf07c1b6e3867451b354730dc7799 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Fri, 30 Sep 2022 23:37:22 +0000 Subject: [PATCH 483/828] refactor(ring_theory/polynomial/tower): review APIs (#16678) Split from #16412. All of the following lemmas move from namespace `is_scalar_tower` to namespace `polynomial`. Rename `aeval_apply` to `aeval_map_algebra_map` and swap the sides of the equality, remove `aeval_map` which is just the same as `aeval_map_algebra_map`. Rename `algebra_map_aeval` to `aeval_algebra_map_apply` and swap the sides of the equality, rename original `aeval_algebra_map_apply` to `aeval_algebra_map_apply_eq_algebra_map_eval`. Make the new lemma a `simp` lemma and remove `simp` from the original one. Rename `aeval_eq_zero_of_aeval_algebra_map_eq_zero` to `aeval_algebra_map_eq_zero_iff_of_injective` and make it an iff lemma. Replace `aeval_eq_zero_of_aeval_algebra_map_eq_zero_field` by `aeval_algebra_map_eq_zero_iff`, it has weaker assumptions and is also an iff lemma. - [x] depends on: #16673 --- .../trigonometric/chebyshev.lean | 4 +- src/data/complex/is_R_or_C.lean | 2 +- src/data/polynomial/algebra_map.lean | 7 +- src/field_theory/abel_ruffini.lean | 2 +- src/field_theory/adjoin.lean | 2 +- src/field_theory/minpoly.lean | 17 ++--- src/field_theory/normal.lean | 7 +- src/field_theory/separable.lean | 3 +- src/field_theory/splitting_field.lean | 2 +- src/ring_theory/adjoin/power_basis.lean | 2 +- src/ring_theory/algebraic.lean | 4 +- src/ring_theory/integral_closure.lean | 7 +- src/ring_theory/localization/integral.lean | 6 +- src/ring_theory/polynomial/tower.lean | 72 +++++++++---------- 14 files changed, 64 insertions(+), 73 deletions(-) diff --git a/src/analysis/special_functions/trigonometric/chebyshev.lean b/src/analysis/special_functions/trigonometric/chebyshev.lean index 78ddae6ca6b6f..aa8a8e90c8672 100644 --- a/src/analysis/special_functions/trigonometric/chebyshev.lean +++ b/src/analysis/special_functions/trigonometric/chebyshev.lean @@ -28,11 +28,11 @@ by rw [aeval_def, eval₂_eq_eval_map, map_U] @[simp] lemma algebra_map_eval_T (x : R) (n : ℕ) : algebra_map R A ((T R n).eval x) = (T A n).eval (algebra_map R A x) := -by rw [←aeval_algebra_map_apply, aeval_T] +by rw [←aeval_algebra_map_apply_eq_algebra_map_eval, aeval_T] @[simp] lemma algebra_map_eval_U (x : R) (n : ℕ) : algebra_map R A ((U R n).eval x) = (U A n).eval (algebra_map R A x) := -by rw [←aeval_algebra_map_apply, aeval_U] +by rw [←aeval_algebra_map_apply_eq_algebra_map_eval, aeval_U] @[simp, norm_cast] lemma complex_of_real_eval_T : ∀ x n, ((T ℝ n).eval x : ℂ) = (T ℂ n).eval x := @algebra_map_eval_T ℝ ℂ _ _ _ diff --git a/src/data/complex/is_R_or_C.lean b/src/data/complex/is_R_or_C.lean index d98c8f09f741b..f8b57dac29b84 100644 --- a/src/data/complex/is_R_or_C.lean +++ b/src/data/complex/is_R_or_C.lean @@ -679,7 +679,7 @@ namespace polynomial open_locale polynomial lemma of_real_eval (p : ℝ[X]) (x : ℝ) : (p.eval x : K) = aeval ↑x p := -(@aeval_algebra_map_apply ℝ K _ _ _ x p).symm +(@aeval_algebra_map_apply_eq_algebra_map_eval ℝ K _ _ _ x p).symm end polynomial diff --git a/src/data/polynomial/algebra_map.lean b/src/data/polynomial/algebra_map.lean index fc08bdd5bb11f..aea595e742ab4 100644 --- a/src/data/polynomial/algebra_map.lean +++ b/src/data/polynomial/algebra_map.lean @@ -204,11 +204,6 @@ lemma aeval_comp {A : Type*} [comm_semiring A] [algebra R A] (x : A) : aeval x (p.comp q) = (aeval (aeval x q) p) := eval₂_comp (algebra_map R A) -@[simp] lemma aeval_map {A : Type*} [comm_semiring A] [algebra R A] [algebra A B] - [is_scalar_tower R A B] (b : B) (p : R[X]) : - aeval b (p.map (algebra_map R A)) = aeval b p := -by rw [aeval_def, eval₂_map, ←is_scalar_tower.algebra_map_eq, ←aeval_def] - theorem aeval_alg_hom (f : A →ₐ[R] B) (x : A) : aeval (f x) = f.comp (aeval x) := alg_hom_ext $ by simp only [aeval_X, alg_hom.comp_apply] @@ -233,7 +228,7 @@ theorem aeval_alg_equiv_apply (f : A ≃ₐ[R] B) (x : A) (p : R[X]) : aeval (f x) p = f (aeval x p) := aeval_alg_hom_apply (f : A →ₐ[R] B) x p -lemma aeval_algebra_map_apply (x : R) (p : R[X]) : +lemma aeval_algebra_map_apply_eq_algebra_map_eval (x : R) (p : R[X]) : aeval (algebra_map R A x) p = algebra_map R A (p.eval x) := aeval_alg_hom_apply (algebra.of_id R A) x p diff --git a/src/field_theory/abel_ruffini.lean b/src/field_theory/abel_ruffini.lean index 2ecab32a062ae..e93a32eeadef5 100644 --- a/src/field_theory/abel_ruffini.lean +++ b/src/field_theory/abel_ruffini.lean @@ -346,7 +346,7 @@ begin suffices : aeval (⟨γ, hγ⟩ : F ⟮α, β⟯) (minpoly F γ) = 0, { rw [aeval_alg_hom_apply, this, alg_hom.map_zero] }, apply (algebra_map F⟮α, β⟯ (solvable_by_rad F E)).injective, - rw [ring_hom.map_zero, is_scalar_tower.algebra_map_aeval], + rw [ring_hom.map_zero, ← aeval_algebra_map_apply], exact minpoly.aeval F γ, end (minpoly.monic (is_integral γ)), rw [P, key], diff --git a/src/field_theory/adjoin.lean b/src/field_theory/adjoin.lean index caaa9d554f777..453ae38d9b20f 100644 --- a/src/field_theory/adjoin.lean +++ b/src/field_theory/adjoin.lean @@ -619,7 +619,7 @@ begin ext, convert minpoly.aeval F α, conv in (aeval α) { rw [← adjoin_simple.algebra_map_gen F α] }, - exact is_scalar_tower.algebra_map_aeval F F⟮α⟯ E _ _ + exact (aeval_algebra_map_apply E (adjoin_simple.gen F α) _).symm end /-- algebra isomorphism between `adjoin_root` and `F⟮α⟯` -/ diff --git a/src/field_theory/minpoly.lean b/src/field_theory/minpoly.lean index 2cfa3279a3de6..0d9f3e85d189a 100644 --- a/src/field_theory/minpoly.lean +++ b/src/field_theory/minpoly.lean @@ -316,7 +316,7 @@ end lemma dvd_map_of_is_scalar_tower (A K : Type*) {R : Type*} [comm_ring A] [field K] [comm_ring R] [algebra A K] [algebra A R] [algebra K R] [is_scalar_tower A K R] (x : R) : minpoly K x ∣ (minpoly A x).map (algebra_map A K) := -by { refine minpoly.dvd K x _, rw [← is_scalar_tower.aeval_apply, minpoly.aeval] } +by { refine minpoly.dvd K x _, rw [aeval_map_algebra_map, minpoly.aeval] } /-- If `y` is a conjugate of `x` over a field `K`, then it is a conjugate over a subring `R`. -/ lemma aeval_of_is_scalar_tower (R : Type*) {K T U : Type*} [comm_ring R] [field K] [comm_ring T] @@ -324,9 +324,8 @@ lemma aeval_of_is_scalar_tower (R : Type*) {K T U : Type*} [comm_ring R] [field [comm_semiring U] [algebra K U] [algebra R U] [is_scalar_tower R K U] (x : T) (y : U) (hy : polynomial.aeval y (minpoly K x) = 0) : polynomial.aeval y (minpoly R x) = 0 := -by { rw is_scalar_tower.aeval_apply R K, - exact eval₂_eq_zero_of_dvd_of_eval₂_eq_zero (algebra_map K U) y - (minpoly.dvd_map_of_is_scalar_tower R K x) hy } +aeval_map_algebra_map K y (minpoly R x) ▸ eval₂_eq_zero_of_dvd_of_eval₂_eq_zero (algebra_map K U) + y (minpoly.dvd_map_of_is_scalar_tower R K x) hy variables {A x} @@ -361,9 +360,9 @@ lemma eq_of_algebra_map_eq {K S T : Type*} [field K] [comm_ring S] [comm_ring T] {x : S} {y : T} (hx : is_integral K x) (h : y = algebra_map S T x) : minpoly K x = minpoly K y := minpoly.unique _ _ (minpoly.monic hx) - (by rw [h, ← is_scalar_tower.algebra_map_aeval, minpoly.aeval, ring_hom.map_zero]) + (by rw [h, aeval_algebra_map_apply, minpoly.aeval, ring_hom.map_zero]) (λ q q_monic root_q, minpoly.min _ _ q_monic - (is_scalar_tower.aeval_eq_zero_of_aeval_algebra_map_eq_zero K S T hST + ((aeval_algebra_map_eq_zero_iff_of_injective hST).mp (h ▸ root_q : polynomial.aeval (algebra_map S T x) q = 0))) lemma add_algebra_map {B : Type*} [comm_ring B] [algebra A B] {x : B} @@ -442,8 +441,7 @@ begin refine (eq_of_irreducible_of_monic _ _ _).symm, { exact (polynomial.is_primitive.irreducible_iff_irreducible_map_fraction_map (polynomial.monic.is_primitive (monic hs))).1 (irreducible hs) }, - { rw [aeval_map, aeval_def, is_scalar_tower.algebra_map_eq R S L, ← eval₂_map, eval₂_at_apply, - eval_map, ← aeval_def, aeval, map_zero] }, + { rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval, map_zero] }, { exact (monic hs).map _ } end @@ -477,8 +475,7 @@ begin have hy : is_integral R y := hs.algebra_map, rw [← gcd_domain_eq_field_fractions K L hs], refine dvd _ _ _, - rw [aeval_map, aeval_def, is_scalar_tower.algebra_map_eq R S L, ← eval₂_map, eval₂_at_apply, - eval_map, ← aeval_def, aeval_prim_part_eq_zero hP hroot, map_zero] + rw [aeval_map_algebra_map, aeval_algebra_map_apply, aeval_prim_part_eq_zero hP hroot, map_zero] end /-- If an element `x` is a root of a nonzero polynomial `p`, then the degree of `p` is at least the diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index db49716d7fbc2..e84972970bbd9 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -100,8 +100,7 @@ lemma alg_hom.normal_bijective [h : normal F E] (ϕ : E →ₐ[F] K) : function. (minpoly.irreducible (is_integral_of_is_scalar_tower x ((is_integral_algebra_map_iff (algebra_map K E).injective).mp h1))) (minpoly.dvd E x ((algebra_map K E).injective (by - { rw [ring_hom.map_zero, aeval_map, ←is_scalar_tower.to_alg_hom_apply F K E, - ←alg_hom.comp_apply, ←aeval_alg_hom], + { rw [ring_hom.map_zero, aeval_map_algebra_map, ← aeval_algebra_map_apply], exact minpoly.aeval F (algebra_map K E x) })))) with y hy, exact ⟨y, hy⟩ }⟩ @@ -249,8 +248,8 @@ def alg_hom.restrict_normal_aux [h : normal F E] : exact or.resolve_left (h.splits z).def (minpoly.ne_zero (h.is_integral z)) (minpoly.irreducible $ is_integral_of_is_scalar_tower _ $ is_integral_alg_hom ϕ $ is_integral_alg_hom _ $ h.is_integral z) - (minpoly.dvd E _ $ by rw [aeval_map, aeval_alg_hom, aeval_alg_hom, alg_hom.comp_apply, - alg_hom.comp_apply, minpoly.aeval, alg_hom.map_zero, alg_hom.map_zero]) }⟩, + (minpoly.dvd E _ $ by rw [aeval_map_algebra_map, aeval_alg_hom_apply, aeval_alg_hom_apply, + minpoly.aeval, alg_hom.map_zero, alg_hom.map_zero]) }⟩, map_zero' := subtype.ext ϕ.map_zero, map_one' := subtype.ext ϕ.map_one, map_add' := λ x y, subtype.ext (ϕ.map_add x y), diff --git a/src/field_theory/separable.lean b/src/field_theory/separable.lean index 1ea975d6c5d14..1b889d1f75e7e 100644 --- a/src/field_theory/separable.lean +++ b/src/field_theory/separable.lean @@ -494,8 +494,7 @@ is_separable_iff.2 $ λ x, begin refine (is_separable_iff.1 h (algebra_map K E x)).imp is_integral_tower_bot_of_is_integral_field (λ hs, _), obtain ⟨q, hq⟩ := minpoly.dvd F x - (is_scalar_tower.aeval_eq_zero_of_aeval_algebra_map_eq_zero_field - (minpoly.aeval F ((algebra_map K E) x))), + ((aeval_algebra_map_eq_zero_iff _ _ _).mp (minpoly.aeval F ((algebra_map K E) x))), rw hq at hs, exact hs.of_mul_left end diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field.lean index 15420e98482cf..7babfe9e330f2 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field.lean @@ -476,7 +476,7 @@ begin ((polynomial.splits_map_iff _ _).2 _) (minpoly.dvd _ _ _), { rw ← is_scalar_tower.algebra_map_eq, exact H2 }, - { rw [← is_scalar_tower.aeval_apply, minpoly.aeval] } }, + { rw [polynomial.aeval_map_algebra_map, minpoly.aeval] } }, obtain ⟨y, hy⟩ := polynomial.exists_root_of_splits _ H6 (ne_of_lt (minpoly.degree_pos H5)).symm, refine ⟨subalgebra.of_restrict_scalars _ _ _⟩, refine (adjoin_root.lift_hom (minpoly (algebra.adjoin F (↑s : set K)) a) y hy).comp _, diff --git a/src/ring_theory/adjoin/power_basis.lean b/src/ring_theory/adjoin/power_basis.lean index a2b132111e4d0..e400d257dc74d 100644 --- a/src/ring_theory/adjoin/power_basis.lean +++ b/src/ring_theory/adjoin/power_basis.lean @@ -49,7 +49,7 @@ begin obtain ⟨f, rfl⟩ := (aeval x).mem_range.mp hy, use f, ext, - exact (is_scalar_tower.algebra_map_aeval K (adjoin K {x}) S ⟨x, _⟩ _).symm } } + exact aeval_algebra_map_apply S (⟨x, _⟩ : adjoin K {x}) _, } } end /-- The power basis `1, x, ..., x ^ (d - 1)` for `K[x]`, diff --git a/src/ring_theory/algebraic.lean b/src/ring_theory/algebraic.lean index fdb0c04e1ec27..46567de090a3e 100644 --- a/src/ring_theory/algebraic.lean +++ b/src/ring_theory/algebraic.lean @@ -112,7 +112,7 @@ open is_scalar_tower lemma is_algebraic_algebra_map_of_is_algebraic {a : S} : is_algebraic R a → is_algebraic R (algebra_map S A a) := -λ ⟨f, hf₁, hf₂⟩, ⟨f, hf₁, by rw [←algebra_map_aeval, hf₂, map_zero]⟩ +λ ⟨f, hf₁, hf₂⟩, ⟨f, hf₁, by rw [aeval_algebra_map_apply, hf₂, map_zero]⟩ /-- This is slightly more general than `is_algebraic_algebra_map_of_is_algebraic` in that it allows noncommutative intermediate rings `A`. -/ @@ -130,7 +130,7 @@ lemma _root_.alg_equiv.is_algebraic_iff {B} [ring B] [algebra R B] (e : A ≃ₐ lemma is_algebraic_algebra_map_iff {a : S} (h : function.injective (algebra_map S A)) : is_algebraic R (algebra_map S A a) ↔ is_algebraic R a := -⟨λ ⟨p, hp0, hp⟩, ⟨p, hp0, h (by rwa [map_zero, algebra_map_aeval])⟩, +⟨λ ⟨p, hp0, hp⟩, ⟨p, hp0, h (by rwa [map_zero, ← aeval_algebra_map_apply])⟩, is_algebraic_algebra_map_of_is_algebraic⟩ lemma is_algebraic_of_pow {r : A} {n : ℕ} (hn : 0 < n) (ht : is_algebraic R (r ^ n)) : diff --git a/src/ring_theory/integral_closure.lean b/src/ring_theory/integral_closure.lean index b02f24b384469..6b1dfa0366afd 100644 --- a/src/ring_theory/integral_closure.lean +++ b/src/ring_theory/integral_closure.lean @@ -126,7 +126,7 @@ theorem is_integral_of_is_scalar_tower [algebra A B] [is_scalar_tower R A B] (x : B) (hx : is_integral R x) : is_integral A x := let ⟨p, hp, hpx⟩ := hx in ⟨p.map $ algebra_map R A, hp.map _, - by rw [← aeval_def, ← is_scalar_tower.aeval_apply, aeval_def, hpx]⟩ + by rw [← aeval_def, aeval_map_algebra_map, aeval_def, hpx]⟩ theorem is_integral_of_subring {x : A} (T : subring R) (hx : is_integral T x) : is_integral R x := @@ -148,7 +148,7 @@ begin refine ⟨_, λ h, h.algebra_map⟩, rintros ⟨f, hf, hx⟩, use [f, hf], - exact is_scalar_tower.aeval_eq_zero_of_aeval_algebra_map_eq_zero R A B hAB hx, + exact (aeval_algebra_map_eq_zero_iff_of_injective hAB).mp hx, end theorem is_integral_iff_is_integral_closure_finite {r : A} : @@ -157,7 +157,8 @@ begin split; intro hr, { rcases hr with ⟨p, hmp, hpr⟩, refine ⟨_, finset.finite_to_set _, p.restriction, monic_restriction.2 hmp, _⟩, - erw [← aeval_def, is_scalar_tower.aeval_apply _ R, map_restriction, aeval_def, hpr] }, + rw [← aeval_def, ← aeval_map_algebra_map R r p.restriction, + map_restriction, aeval_def, hpr], }, rcases hr with ⟨s, hs, hsr⟩, exact is_integral_of_subring _ hsr end diff --git a/src/ring_theory/localization/integral.lean b/src/ring_theory/localization/integral.lean index 7f068d4d21281..b0c6ea47a692f 100644 --- a/src/ring_theory/localization/integral.lean +++ b/src/ring_theory/localization/integral.lean @@ -152,7 +152,7 @@ begin { refine ⟨p.map (algebra_map A K), λ h, hp (polynomial.ext (λ i, _)), _⟩, { have : algebra_map A K (p.coeff i) = 0 := trans (polynomial.coeff_map _ _).symm (by simp [h]), exact to_map_eq_zero_iff.mp this }, - { rwa is_scalar_tower.aeval_apply _ K at px } }, + { exact (polynomial.aeval_map_algebra_map K _ _).trans px, } }, { exact ⟨integer_normalization _ p, mt integer_normalization_eq_zero_iff.mp hp, integer_normalization_aeval_eq_zero _ p px⟩ }, @@ -386,11 +386,11 @@ begin (no_zero_smul_divisors.algebra_map_injective _ _) b h)))), rw [polynomial.aeval_def, ← inv_of_eq_inv, polynomial.eval₂_reverse_eq_zero_iff, polynomial.eval₂_map, ← is_scalar_tower.algebra_map_eq, ← polynomial.aeval_def, - ← is_scalar_tower.algebra_map_aeval, hf₂, ring_hom.map_zero] } } }, + polynomial.aeval_algebra_map_apply, hf₂, ring_hom.map_zero] } } }, { intros h x, obtain ⟨f, hf₁, hf₂⟩ := h (algebra_map S K x), use [f, hf₁], - rw [← is_scalar_tower.algebra_map_aeval] at hf₂, + rw [polynomial.aeval_algebra_map_apply] at hf₂, exact (injective_iff_map_eq_zero (algebra_map S K)).1 (no_zero_smul_divisors.algebra_map_injective _ _) _ hf₂ } end diff --git a/src/ring_theory/polynomial/tower.lean b/src/ring_theory/polynomial/tower.lean index 39c38f58e0d03..b1890b8312a18 100644 --- a/src/ring_theory/polynomial/tower.lean +++ b/src/ring_theory/polynomial/tower.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kenny Lau +Authors: Kenny Lau, Yuyang Zhao -/ import algebra.algebra.tower @@ -15,22 +15,22 @@ This file proves some basic results about the algebra tower structure for the ty This structure itself is provided elsewhere as `polynomial.is_scalar_tower` -/ -universes u v w u₁ open_locale polynomial -variables (R : Type u) (S : Type v) (A : Type w) (B : Type u₁) +variables (R A B : Type*) -namespace is_scalar_tower +namespace polynomial section semiring -variables [comm_semiring R] [comm_semiring S] [semiring A] [semiring B] -variables [algebra R S] [algebra S A] [algebra S B] [algebra R A] [algebra R B] -variables [is_scalar_tower R S A] [is_scalar_tower R S B] +variables [comm_semiring R] [comm_semiring A] [semiring B] +variables [algebra R A] [algebra A B] [algebra R B] +variables [is_scalar_tower R A B] + +variables {R B} -variables (R S A) {B} -theorem aeval_apply (x : A) (p : R[X]) : polynomial.aeval x p = - polynomial.aeval x (polynomial.map (algebra_map R S) p) := -by rw [polynomial.aeval_def, polynomial.aeval_def, polynomial.eval₂_map, algebra_map_eq R S A] +@[simp] theorem aeval_map_algebra_map (x : B) (p : R[X]) : + aeval x (map (algebra_map R A) p) = aeval x p := +by rw [aeval_def, aeval_def, eval₂_map, is_scalar_tower.algebra_map_eq R A B] end semiring @@ -38,41 +38,41 @@ section comm_semiring variables [comm_semiring R] [comm_semiring A] [semiring B] variables [algebra R A] [algebra A B] [algebra R B] [is_scalar_tower R A B] -lemma algebra_map_aeval (x : A) (p : R[X]) : - algebra_map A B (polynomial.aeval x p) = polynomial.aeval (algebra_map A B x) p := -by rw [polynomial.aeval_def, polynomial.aeval_def, polynomial.hom_eval₂, - ←is_scalar_tower.algebra_map_eq] - -lemma aeval_eq_zero_of_aeval_algebra_map_eq_zero {x : A} {p : R[X]} - (h : function.injective (algebra_map A B)) (hp : polynomial.aeval (algebra_map A B x) p = 0) : - polynomial.aeval x p = 0 := -begin - rw [← algebra_map_aeval, ← (algebra_map A B).map_zero] at hp, - exact h hp, -end - -lemma aeval_eq_zero_of_aeval_algebra_map_eq_zero_field {R A B : Type*} [comm_semiring R] [field A] - [comm_semiring B] [nontrivial B] [algebra R A] [algebra R B] [algebra A B] [is_scalar_tower R A B] - {x : A} {p : R[X]} (h : polynomial.aeval (algebra_map A B x) p = 0) : - polynomial.aeval x p = 0 := -aeval_eq_zero_of_aeval_algebra_map_eq_zero R A B (algebra_map A B).injective h +variables {R A} + +lemma aeval_algebra_map_apply (x : A) (p : R[X]) : + aeval (algebra_map A B x) p = algebra_map A B (aeval x p) := +by rw [aeval_def, aeval_def, hom_eval₂, ←is_scalar_tower.algebra_map_eq] + +@[simp] lemma aeval_algebra_map_eq_zero_iff [no_zero_smul_divisors A B] [nontrivial B] + (x : A) (p : R[X]) : + aeval (algebra_map A B x) p = 0 ↔ aeval x p = 0 := +by rw [aeval_algebra_map_apply, algebra.algebra_map_eq_smul_one, smul_eq_zero, + iff_false_intro (@one_ne_zero B _ _), or_false] + +variables {B} + +lemma aeval_algebra_map_eq_zero_iff_of_injective + {x : A} {p : R[X]} + (h : function.injective (algebra_map A B)) : + aeval (algebra_map A B x) p = 0 ↔ aeval x p = 0 := +by rw [aeval_algebra_map_apply, ← (algebra_map A B).map_zero, h.eq_iff] end comm_semiring -end is_scalar_tower +end polynomial namespace subalgebra -open is_scalar_tower +open polynomial section comm_semiring -variables (R) {S A} [comm_semiring R] [comm_semiring S] [comm_semiring A] -variables [algebra R S] [algebra S A] [algebra R A] [is_scalar_tower R S A] +variables {R A} [comm_semiring R] [comm_semiring A] [algebra R A] -@[simp] lemma aeval_coe {S : subalgebra R A} {x : S} {p : R[X]} : - polynomial.aeval (x : A) p = polynomial.aeval x p := -(algebra_map_aeval R S A x p).symm +@[simp] lemma aeval_coe (S : subalgebra R A) (x : S) (p : R[X]) : + aeval (x : A) p = aeval x p := +aeval_algebra_map_apply A x p end comm_semiring From 5fefdfa15e664c3606bcd6dd5c0a931301263a0b Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sat, 1 Oct 2022 02:05:34 +0000 Subject: [PATCH 484/828] feat(group_theory/schreier): The size of the commutator subgroup is bounded in terms of the index of the center and the number of commutators (#16679) This PR proves that the size of the commutator subgroup is bounded in terms of the index of the center and the number of commutators. The proof uses Schreier's lemma and the transfer homomorphism. I included lots of comments since the proof is rather technical. But please let me know if I went overboard. Ultimately, this is building up to a bound on the size of the commutator subgroup just in terms of the number of commutators. Co-authored-by: Thomas Browning --- src/group_theory/abelianization.lean | 15 ++++++++++- src/group_theory/schreier.lean | 39 ++++++++++++++++++++++++---- 2 files changed, 48 insertions(+), 6 deletions(-) diff --git a/src/group_theory/abelianization.lean b/src/group_theory/abelianization.lean index e7be75497aac9..4bb299c2123ac 100644 --- a/src/group_theory/abelianization.lean +++ b/src/group_theory/abelianization.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Michael Howes -/ import group_theory.commutator -import group_theory.quotient_group +import group_theory.finiteness /-! # The abelianization of a group @@ -46,6 +46,19 @@ by simp_rw [commutator, subgroup.commutator_def', subgroup.mem_top, exists_true_ instance commutator_characteristic : (commutator G).characteristic := subgroup.commutator_characteristic ⊤ ⊤ +instance [finite {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g}] : group.fg (commutator G) := +begin + rw commutator_eq_closure, + apply group.closure_finite_fg, +end + +lemma rank_commutator_le_card [finite {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g}] : + group.rank (commutator G) ≤ nat.card {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g} := +begin + rw subgroup.rank_congr (commutator_eq_closure G), + apply subgroup.rank_closure_finite_le_nat_card, +end + lemma commutator_centralizer_commutator_le_center : ⁅(commutator G).centralizer, (commutator G).centralizer⁆ ≤ subgroup.center G := begin diff --git a/src/group_theory/schreier.lean b/src/group_theory/schreier.lean index bf5b797a3e06a..6ea43693665e8 100644 --- a/src/group_theory/schreier.lean +++ b/src/group_theory/schreier.lean @@ -4,11 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ -import data.finset.pointwise -import group_theory.complement -import group_theory.finiteness -import group_theory.index -import tactic.group +import group_theory.abelianization +import group_theory.exponent +import group_theory.transfer /-! # Schreier's Lemma @@ -144,4 +142,35 @@ begin ... = H.index * group.rank G : congr_arg ((*) H.index) hS₀, end +variables (G) + +/-- If `G` has `n` commutators `[g₁, g₂]`, then `|G'| ∣ [G : Z(G)] ^ ([G : Z(G)] * n + 1)`, +where `G'` denotes the commutator of `G`. -/ +lemma card_commutator_dvd_index_center_pow [finite {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g}] : + nat.card (commutator G) ∣ + (center G).index ^ ((center G).index * nat.card {g | ∃ g₁ g₂ : G, ⁅g₁, g₂⁆ = g} + 1) := +begin + -- First handle the case when `Z(G)` has infinite index and `[G : Z(G)]` is defined to be `0` + by_cases hG : (center G).index = 0, + { simp_rw [hG, zero_mul, zero_add, pow_one, dvd_zero] }, + -- Rewrite as `|Z(G) ∩ G'| * [G' : Z(G) ∩ G'] ∣ [G : Z(G)] ^ ([G : Z(G)] * n) * [G : Z(G)]` + rw [←((center G).subgroup_of (commutator G)).card_mul_index, pow_succ'], + -- We have `h1 : [G' : Z(G) ∩ G'] ∣ [G : Z(G)]` + have h1 := relindex_dvd_index_of_normal (center G) (commutator G), + -- So we can reduce to proving `|Z(G) ∩ G'| ∣ [G : Z(G)] ^ ([G : Z(G)] * n)` + refine mul_dvd_mul _ h1, + -- We have `h2 : rank (Z(G) ∩ G') ≤ [G' : Z(G) ∩ G'] * rank G'` by Schreier's lemma + have h2 := rank_le_index_mul_rank (ne_zero_of_dvd_ne_zero hG h1), + -- We have `h3 : [G' : Z(G) ∩ G'] * rank G' ≤ [G : Z(G)] * n` by `h1` and `rank G' ≤ n` + have h3 := nat.mul_le_mul (nat.le_of_dvd (nat.pos_of_ne_zero hG) h1) (rank_commutator_le_card G), + -- So we can reduce to proving `|Z(G) ∩ G'| ∣ [G : Z(G)] ^ rank (Z(G) ∩ G')` + refine dvd_trans _ (pow_dvd_pow (center G).index (h2.trans h3)), + -- `Z(G) ∩ G'` is abelian, so it enough to prove that `g ^ [G : Z(G)] = 1` for `g ∈ Z(G) ∩ G'` + apply card_dvd_exponent_pow_rank' _ (λ g, _), + -- `Z(G)` is abelian, so `g ∈ Z(G) ∩ G' ≤ G' ≤ ker (transfer : G → Z(G))` + have := abelianization.commutator_subset_ker (monoid_hom.transfer_center_pow' hG) g.1.2, + -- `transfer g` is defeq to `g ^ [G : Z(G)]`, so we are done + simpa only [monoid_hom.mem_ker, subtype.ext_iff] using this, +end + end subgroup From 2a543b0da429133e8503d88787d9601e84dafe0c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 1 Oct 2022 02:05:35 +0000 Subject: [PATCH 485/828] perf(algebraic_topology/simplicial_object): Speedup `simplicial_cosimplicial_augmented_equiv` (#16729) Use `equivalence.mk` rather than `equivalence.mk'`. Restrict the `simps` to `functor` and `inverse`. This roughly speeds everything up by a factor of 2. --- src/algebraic_topology/simplicial_object.lean | 21 +++++++------------ 1 file changed, 7 insertions(+), 14 deletions(-) diff --git a/src/algebraic_topology/simplicial_object.lean b/src/algebraic_topology/simplicial_object.lean index b6fc69fe884cc..25b30aac57922 100644 --- a/src/algebraic_topology/simplicial_object.lean +++ b/src/algebraic_topology/simplicial_object.lean @@ -567,21 +567,14 @@ def cosimplicial_to_simplicial_augmented : /-- The contravariant categorical equivalence between augmented simplicial objects and augmented cosimplicial objects in the opposite category. -/ -@[simps] +@[simps functor inverse] def simplicial_cosimplicial_augmented_equiv : (simplicial_object.augmented C)ᵒᵖ ≌ cosimplicial_object.augmented Cᵒᵖ := -{ functor := simplicial_to_cosimplicial_augmented _, - inverse := cosimplicial_to_simplicial_augmented _, - unit_iso := nat_iso.of_components - (λ X, X.unop.right_op_left_op_iso.op) begin - intros X Y f, - dsimp, - rw (show f = f.unop.op, by simp), - simp_rw ← op_comp, - congr' 1, - tidy, - end, - counit_iso := nat_iso.of_components - (λ X, X.left_op_right_op_iso) (by tidy) } +equivalence.mk + (simplicial_to_cosimplicial_augmented _) + (cosimplicial_to_simplicial_augmented _) + (nat_iso.of_components (λ X, X.unop.right_op_left_op_iso.op) $ λ X Y f, + by { dsimp, rw ←f.op_unop, simp_rw ← op_comp, congr' 1, tidy }) + (nat_iso.of_components (λ X, X.left_op_right_op_iso) $ by tidy) end category_theory From 25c8333eca547c176ecfe8931038d1f54a3a2316 Mon Sep 17 00:00:00 2001 From: Jamie Reason <95385925+jamiereason@users.noreply.github.com> Date: Sat, 1 Oct 2022 07:15:36 +0000 Subject: [PATCH 486/828] chore(number_theory/bernoulli_polynomials): removed unnecessary cast_succ from sum_bernoulli (#16731) --- src/number_theory/bernoulli_polynomials.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/number_theory/bernoulli_polynomials.lean b/src/number_theory/bernoulli_polynomials.lean index 73b8263a76085..8b282f5befc8b 100644 --- a/src/number_theory/bernoulli_polynomials.lean +++ b/src/number_theory/bernoulli_polynomials.lean @@ -118,7 +118,7 @@ end begin simp_rw [bernoulli_def, finset.smul_sum, finset.range_eq_Ico, ←finset.sum_Ico_Ico_comm, finset.sum_Ico_eq_sum_range], - simp only [cast_succ, add_tsub_cancel_left, tsub_zero, zero_add, linear_map.map_add], + simp only [add_tsub_cancel_left, tsub_zero, zero_add, linear_map.map_add], simp_rw [smul_monomial, mul_comm (_root_.bernoulli _) _, smul_eq_mul, ←mul_assoc], conv_lhs { apply_congr, skip, conv { apply_congr, skip, @@ -127,7 +127,7 @@ begin mul_assoc, mul_comm, ←smul_eq_mul, ←smul_monomial] }, rw [←sum_smul], }, rw [sum_range_succ_comm], - simp only [add_right_eq_self, cast_succ, mul_one, cast_one, cast_add, add_tsub_cancel_left, + simp only [add_right_eq_self, mul_one, cast_one, cast_add, add_tsub_cancel_left, choose_succ_self_right, one_smul, _root_.bernoulli_zero, sum_singleton, zero_add, linear_map.map_add, range_one], apply sum_eq_zero (λ x hx, _), From 8f78eb9c07ad7a97657f64cfebb6bb60e6503afd Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 1 Oct 2022 08:30:23 +0000 Subject: [PATCH 487/828] feat(*): generalize to and add distrib_smul instances (#16132) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR builds on #16123 by adding `distrib_smul` (and `smul_zero_class`) instances for: * `finsupp` * `add_monoid_algebra` * `polynomial` * submodule quotients * scalar multiplication by `ℚ` It also generalizes some results by weakening `distrib_mul_action` to `distrib_smul`. The choice of instances and generalizations is based on which are necessary to solve instance diamonds in `splitting_field`, so I probably missed a lot of additional changes. Since I'm not so interested in generalizing everything at the moment, I'd rather leave missing instances to follow-up PRs. Co-authored-by: Eric Rodriguez --- src/algebra/monoid_algebra/basic.lean | 26 ++-- src/data/finsupp/basic.lean | 26 ++-- src/data/polynomial/basic.lean | 14 +- src/data/rat/cast.lean | 19 +++ src/field_theory/ratfunc.lean | 6 +- .../group_action/big_operators.lean | 10 +- src/group_theory/group_action/defs.lean | 129 ++++++++++++++++-- src/linear_algebra/basic.lean | 2 +- src/linear_algebra/quotient.lean | 17 +++ src/representation_theory/maschke.lean | 6 +- test/instance_diamonds.lean | 8 +- 11 files changed, 210 insertions(+), 53 deletions(-) diff --git a/src/algebra/monoid_algebra/basic.lean b/src/algebra/monoid_algebra/basic.lean index c045905d7fe03..612c868d3c783 100644 --- a/src/algebra/monoid_algebra/basic.lean +++ b/src/algebra/monoid_algebra/basic.lean @@ -268,9 +268,13 @@ instance [comm_ring k] [comm_monoid G] : comm_ring (monoid_algebra k G) := variables {S : Type*} -instance [monoid R] [semiring k] [distrib_mul_action R k] : - has_smul R (monoid_algebra k G) := -finsupp.has_smul +instance [semiring k] [smul_zero_class R k] : + smul_zero_class R (monoid_algebra k G) := +finsupp.smul_zero_class + +instance [semiring k] [distrib_smul R k] : + distrib_smul R (monoid_algebra k G) := +finsupp.distrib_smul _ _ instance [monoid R] [semiring k] [distrib_mul_action R k] : distrib_mul_action R (monoid_algebra k G) := @@ -451,7 +455,7 @@ end misc_theorems /-! #### Non-unital, non-associative algebra structure -/ section non_unital_non_assoc_algebra -variables (k) [monoid R] [semiring k] [distrib_mul_action R k] [has_mul G] +variables (k) [semiring k] [distrib_smul R k] [has_mul G] instance is_scalar_tower_self [is_scalar_tower R k k] : is_scalar_tower R (monoid_algebra k G) (monoid_algebra k G) := @@ -999,9 +1003,9 @@ end mul_one_class /-! #### Semiring structure -/ section semiring -instance {R : Type*} [monoid R] [semiring k] [distrib_mul_action R k] : - has_smul R (add_monoid_algebra k G) := -finsupp.has_smul +instance {R : Type*} [semiring k] [smul_zero_class R k] : + smul_zero_class R (add_monoid_algebra k G) := +finsupp.smul_zero_class variables [semiring k] [add_monoid G] @@ -1298,22 +1302,22 @@ variables {k G} section non_unital_non_assoc_algebra -variables (k) [monoid R] [semiring k] [distrib_mul_action R k] [has_add G] +variables (k) [semiring k] [distrib_smul R k] [has_add G] instance is_scalar_tower_self [is_scalar_tower R k k] : is_scalar_tower R (add_monoid_algebra k G) (add_monoid_algebra k G) := -@monoid_algebra.is_scalar_tower_self k (multiplicative G) R _ _ _ _ _ +@monoid_algebra.is_scalar_tower_self k (multiplicative G) R _ _ _ _ /-- Note that if `k` is a `comm_semiring` then we have `smul_comm_class k k k` and so we can take `R = k` in the below. In other words, if the coefficients are commutative amongst themselves, they also commute with the algebra multiplication. -/ instance smul_comm_class_self [smul_comm_class R k k] : smul_comm_class R (add_monoid_algebra k G) (add_monoid_algebra k G) := -@monoid_algebra.smul_comm_class_self k (multiplicative G) R _ _ _ _ _ +@monoid_algebra.smul_comm_class_self k (multiplicative G) R _ _ _ _ instance smul_comm_class_symm_self [smul_comm_class k R k] : smul_comm_class (add_monoid_algebra k G) R (add_monoid_algebra k G) := -@monoid_algebra.smul_comm_class_symm_self k (multiplicative G) R _ _ _ _ _ +@monoid_algebra.smul_comm_class_symm_self k (multiplicative G) R _ _ _ _ variables {A : Type u₃} [non_unital_non_assoc_semiring A] diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index 665c651c2fd76..e4b893f58a853 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -1296,36 +1296,42 @@ end end section -instance [monoid R] [add_monoid M] [distrib_mul_action R M] : has_smul R (α →₀ M) := -⟨λa v, v.map_range ((•) a) (smul_zero _)⟩ + +instance [has_zero M] [smul_zero_class R M] : smul_zero_class R (α →₀ M) := +{ smul := λ a v, v.map_range ((•) a) (smul_zero _), + smul_zero := λ a, by { ext, apply smul_zero } } /-! Throughout this section, some `monoid` and `semiring` arguments are specified with `{}` instead of `[]`. See note [implicit instance arguments]. -/ -@[simp] lemma coe_smul {_ : monoid R} [add_monoid M] [distrib_mul_action R M] +@[simp] lemma coe_smul [add_monoid M] [distrib_smul R M] (b : R) (v : α →₀ M) : ⇑(b • v) = b • v := rfl -lemma smul_apply {_ : monoid R} [add_monoid M] [distrib_mul_action R M] +lemma smul_apply [add_monoid M] [distrib_smul R M] (b : R) (v : α →₀ M) (a : α) : (b • v) a = b • (v a) := rfl -lemma _root_.is_smul_regular.finsupp {_ : monoid R} [add_monoid M] [distrib_mul_action R M] {k : R} +lemma _root_.is_smul_regular.finsupp [add_monoid M] [distrib_smul R M] {k : R} (hk : is_smul_regular M k) : is_smul_regular (α →₀ M) k := λ _ _ h, ext $ λ i, hk (congr_fun h i) -instance [monoid R] [nonempty α] [add_monoid M] [distrib_mul_action R M] [has_faithful_smul R M] : +instance [nonempty α] [add_monoid M] [distrib_smul R M] [has_faithful_smul R M] : has_faithful_smul R (α →₀ M) := { eq_of_smul_eq_smul := λ r₁ r₂ h, let ⟨a⟩ := ‹nonempty α› in eq_of_smul_eq_smul $ λ m : M, by simpa using congr_fun (h (single a m)) a } variables (α M) -instance [monoid R] [add_monoid M] [distrib_mul_action R M] : distrib_mul_action R (α →₀ M) := +instance [add_zero_class M] [distrib_smul R M] : distrib_smul R (α →₀ M) := { smul := (•), smul_add := λ a x y, ext $ λ _, smul_add _ _ _, + smul_zero := λ x, ext $ λ _, smul_zero _ } + +instance [monoid R] [add_monoid M] [distrib_mul_action R M] : distrib_mul_action R (α →₀ M) := +{ smul := (•), one_smul := λ x, ext $ λ _, one_smul _ _, mul_smul := λ r s x, ext $ λ _, mul_smul _ _ _, - smul_zero := λ x, ext $ λ _, smul_zero _ } + ..finsupp.distrib_smul _ _ } instance [monoid R] [monoid S] [add_monoid M] [distrib_mul_action R M] [distrib_mul_action S M] [has_smul R S] [is_scalar_tower R S M] : @@ -1416,14 +1422,14 @@ lemma sum_smul_index [semiring R] [add_comm_monoid M] {g : α →₀ R} {b : R} (h0 : ∀i, h i 0 = 0) : (b • g).sum h = g.sum (λi a, h i (b * a)) := finsupp.sum_map_range_index h0 -lemma sum_smul_index' [monoid R] [add_monoid M] [distrib_mul_action R M] [add_comm_monoid N] +lemma sum_smul_index' [add_monoid M] [distrib_smul R M] [add_comm_monoid N] {g : α →₀ M} {b : R} {h : α → M → N} (h0 : ∀i, h i 0 = 0) : (b • g).sum h = g.sum (λi c, h i (b • c)) := finsupp.sum_map_range_index h0 /-- A version of `finsupp.sum_smul_index'` for bundled additive maps. -/ lemma sum_smul_index_add_monoid_hom - [monoid R] [add_monoid M] [add_comm_monoid N] [distrib_mul_action R M] + [add_monoid M] [add_comm_monoid N] [distrib_smul R M] {g : α →₀ M} {b : R} {h : α → M →+ N} : (b • g).sum (λ a, h a) = g.sum (λ i c, h i (b • c)) := sum_map_range_index (λ i, (h i).map_zero) diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 3517ef332e485..545719846c0d8 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -97,8 +97,9 @@ instance : has_add R[X] := ⟨add⟩ instance {R : Type u} [ring R] : has_neg R[X] := ⟨neg⟩ instance {R : Type u} [ring R] : has_sub R[X] := ⟨λ a b, a + -b⟩ instance : has_mul R[X] := ⟨mul⟩ -instance {S : Type*} [monoid S] [distrib_mul_action S R] : has_smul S R[X] := -⟨λ r p, ⟨r • p.to_finsupp⟩⟩ +instance {S : Type*} [smul_zero_class S R] : smul_zero_class S R[X] := +{ smul := λ r p, ⟨r • p.to_finsupp⟩, + smul_zero := λ a, congr_arg of_finsupp (smul_zero a) } @[priority 1] -- to avoid a bug in the `ring` tactic instance has_pow : has_pow R[X] ℕ := { pow := λ p n, npow_rec n p } @@ -111,7 +112,7 @@ show _ = neg _, by rw neg @[simp] lemma of_finsupp_sub {R : Type u} [ring R] {a b} : (⟨a - b⟩ : R[X]) = ⟨a⟩ - ⟨b⟩ := by { rw [sub_eq_add_neg, of_finsupp_add, of_finsupp_neg], refl } @[simp] lemma of_finsupp_mul (a b) : (⟨a * b⟩ : R[X]) = ⟨a⟩ * ⟨b⟩ := show _ = mul _ _, by rw mul -@[simp] lemma of_finsupp_smul {S : Type*} [monoid S] [distrib_mul_action S R] (a : S) (b) : +@[simp] lemma of_finsupp_smul {S : Type*} [smul_zero_class S R] (a : S) (b) : (⟨a • b⟩ : R[X]) = (a • ⟨b⟩ : R[X]) := rfl @[simp] lemma of_finsupp_pow (a) (n : ℕ) : (⟨a ^ n⟩ : R[X]) = ⟨a⟩ ^ n := begin @@ -136,7 +137,7 @@ by { cases a, rw ←of_finsupp_neg } by { rw [sub_eq_add_neg, ←to_finsupp_neg, ←to_finsupp_add], refl } @[simp] lemma to_finsupp_mul (a b : R[X]) : (a * b).to_finsupp = a.to_finsupp * b.to_finsupp := by { cases a, cases b, rw ←of_finsupp_mul } -@[simp] lemma to_finsupp_smul {S : Type*} [monoid S] [distrib_mul_action S R] (a : S) (b : R[X]) : +@[simp] lemma to_finsupp_smul {S : Type*} [smul_zero_class S R] (a : S) (b : R[X]) : (a • b).to_finsupp = a • b.to_finsupp := rfl @[simp] lemma to_finsupp_pow (a : R[X]) (n : ℕ) : (a ^ n).to_finsupp = a.to_finsupp ^ n := by { cases a, rw ←of_finsupp_pow } @@ -196,6 +197,11 @@ instance {S₁ S₂} [has_smul S₁ S₂] [monoid S₁] [monoid S₂] [distrib_m [distrib_mul_action S₂ R] [is_scalar_tower S₁ S₂ R] : is_scalar_tower S₁ S₂ R[X] := ⟨by { rintros _ _ ⟨⟩, simp_rw [←of_finsupp_smul, smul_assoc] }⟩ +instance is_scalar_tower_right {α K : Type*} [semiring K] [distrib_smul α K] + [is_scalar_tower α K K] : is_scalar_tower α K[X] K[X] := +⟨by rintros _ ⟨⟩ ⟨⟩; + simp_rw [smul_eq_mul, ← of_finsupp_smul, ← of_finsupp_mul, ← of_finsupp_smul, smul_mul_assoc]⟩ + instance {S} [monoid S] [distrib_mul_action S R] [distrib_mul_action Sᵐᵒᵖ R] [is_central_scalar S R] : is_central_scalar S R[X] := ⟨by { rintros _ ⟨⟩, simp_rw [←of_finsupp_smul, op_smul_eq_smul] }⟩ diff --git a/src/data/rat/cast.lean b/src/data/rat/cast.lean index 780ae00f4e71f..920d46ed757a6 100644 --- a/src/data/rat/cast.lean +++ b/src/data/rat/cast.lean @@ -350,3 +350,22 @@ by rw [cast_def, div_eq_mul_inv, unop_mul, unop_inv, unop_nat_cast, unop_int_cas (commute.cast_int_right _ r.num).eq, cast_def, div_eq_mul_inv] end mul_opposite + +section smul + +namespace rat + +variables {K : Type*} [division_ring K] + +@[priority 100] +instance distrib_smul : distrib_smul ℚ K := +{ smul := (•), + smul_zero := λ a, by rw [smul_def, mul_zero], + smul_add := λ a x y, by simp only [smul_def, mul_add, cast_add] } + +instance is_scalar_tower_right : is_scalar_tower ℚ K K := +⟨λ a x y, by simp only [smul_def, smul_eq_mul, mul_assoc]⟩ + +end rat + +end smul diff --git a/src/field_theory/ratfunc.lean b/src/field_theory/ratfunc.lean index 234b2cfa9914f..87d5667ba7f53 100644 --- a/src/field_theory/ratfunc.lean +++ b/src/field_theory/ratfunc.lean @@ -140,7 +140,9 @@ of `∀ {p q a : K[X]} (hq : q ≠ 0) (ha : a ≠ 0), f (a * p) (a * q) = f p q) (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), p * q' = p' * q → f p q = f p' q') : P := -localization.lift_on (to_fraction_ring x) (λ p q, f p q) (λ p p' q q' h, H q.2 q'.2 +localization.lift_on + (by exact to_fraction_ring x) -- Fix timeout by manipulating elaboration order + (λ p q, f p q) (λ p p' q q' h, H q.2 q'.2 (let ⟨⟨c, hc⟩, mul_eq⟩ := (localization.r_iff_exists).mp h in mul_cancel_right_coe_non_zero_divisor.mp mul_eq)) @@ -148,7 +150,7 @@ lemma lift_on_of_fraction_ring_mk {P : Sort v} (n : K[X]) (d : K[X]⁰) (f : ∀ (p q : K[X]), P) (H : ∀ {p q p' q'} (hq : q ∈ K[X]⁰) (hq' : q' ∈ K[X]⁰), p * q' = p' * q → f p q = f p' q') : - ratfunc.lift_on (of_fraction_ring (localization.mk n d)) f @H = f n d := + ratfunc.lift_on (by exact of_fraction_ring (localization.mk n d)) f @H = f n d := begin unfold ratfunc.lift_on, exact localization.lift_on_mk _ _ _ _ diff --git a/src/group_theory/group_action/big_operators.lean b/src/group_theory/group_action/big_operators.lean index f1a2ec0b17ea8..af25d91dc1244 100644 --- a/src/group_theory/group_action/big_operators.lean +++ b/src/group_theory/group_action/big_operators.lean @@ -19,11 +19,11 @@ variables {α β γ : Type*} open_locale big_operators section -variables [monoid α] [add_monoid β] [distrib_mul_action α β] +variables [add_monoid β] [distrib_smul α β] lemma list.smul_sum {r : α} {l : list β} : r • l.sum = (l.map ((•) r)).sum := -(distrib_mul_action.to_add_monoid_hom β r).map_list_sum l +(distrib_smul.to_add_monoid_hom β r).map_list_sum l end @@ -37,15 +37,15 @@ lemma list.smul_prod {r : α} {l : list β} : end section -variables [monoid α] [add_comm_monoid β] [distrib_mul_action α β] +variables [add_comm_monoid β] [distrib_smul α β] lemma multiset.smul_sum {r : α} {s : multiset β} : r • s.sum = (s.map ((•) r)).sum := -(distrib_mul_action.to_add_monoid_hom β r).map_multiset_sum s +(distrib_smul.to_add_monoid_hom β r).map_multiset_sum s lemma finset.smul_sum {r : α} {f : γ → β} {s : finset γ} : r • ∑ x in s, f x = ∑ x in s, r • f x := -(distrib_mul_action.to_add_monoid_hom β r).map_sum f s +(distrib_smul.to_add_monoid_hom β r).map_sum f s end diff --git a/src/group_theory/group_action/defs.lean b/src/group_theory/group_action/defs.lean index 82865e043cd57..cec03a8941b80 100644 --- a/src/group_theory/group_action/defs.lean +++ b/src/group_theory/group_action/defs.lean @@ -494,9 +494,59 @@ end compatible_scalar class smul_zero_class (M A : Type*) [has_zero A] extends has_smul M A := (smul_zero : ∀ (a : M), a • (0 : A) = 0) -@[simp] theorem smul_zero [has_zero A] [smul_zero_class M A] (a : M) : a • (0 : A) = 0 := +section smul_zero + +variables [has_zero A] [smul_zero_class M A] + +@[simp] theorem smul_zero (a : M) : a • (0 : A) = 0 := smul_zero_class.smul_zero _ +/-- Pullback a zero-preserving scalar multiplication along an injective zero-preserving map. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.smul_zero_class [has_zero B] [has_smul M B] (f : zero_hom B A) + (hf : injective f) (smul : ∀ (c : M) x, f (c • x) = c • f x) : + smul_zero_class M B := +{ smul := (•), + smul_zero := λ c, hf $ by simp only [smul, map_zero, smul_zero] } + +/-- Pushforward a zero-preserving scalar multiplication along a zero-preserving map. +See note [reducible non-instances]. -/ +@[reducible] +protected def zero_hom.smul_zero_class [has_zero B] [has_smul M B] (f : zero_hom A B) + (smul : ∀ (c : M) x, f (c • x) = c • f x) : + smul_zero_class M B := +{ smul := (•), + smul_zero := λ c, by simp only [← map_zero f, ← smul, smul_zero] } + +/-- Push forward the multiplication of `R` on `M` along a compatible surjective map `f : R → S`. + +See also `function.surjective.distrib_mul_action_left`. +-/ +@[reducible] +def function.surjective.smul_zero_class_left {R S M : Type*} [has_zero M] [smul_zero_class R M] + [has_smul S M] (f : R → S) (hf : function.surjective f) (hsmul : ∀ c (x : M), f c • x = c • x) : + smul_zero_class S M := +{ smul := (•), + smul_zero := hf.forall.mpr $ λ c, by rw [hsmul, smul_zero] } + +variable (A) + +/-- Compose a `smul_zero_class` with a function, with scalar multiplication `f r' • m`. +See note [reducible non-instances]. -/ +@[reducible] def smul_zero_class.comp_fun (f : N → M) : + smul_zero_class N A := +{ smul := has_smul.comp.smul f, + smul_zero := λ x, smul_zero (f x) } + +/-- Each element of the scalars defines a zero-preserving map. -/ +@[simps] +def smul_zero_class.to_zero_hom (x : M) : zero_hom A A := +{ to_fun := (•) x, + map_zero' := smul_zero x } + +end smul_zero + /-- Typeclass for scalar multiplication that preserves `0` and `+` on the right. This is exactly `distrib_mul_action` without the `mul_action` part. @@ -505,10 +555,68 @@ This is exactly `distrib_mul_action` without the `mul_action` part. extends smul_zero_class M A := (smul_add : ∀ (a : M) (x y : A), a • (x + y) = a • x + a • y) -theorem smul_add [add_zero_class A] [distrib_smul M A] (a : M) (b₁ b₂ : A) : +section distrib_smul + +variables [add_zero_class A] [distrib_smul M A] + +theorem smul_add (a : M) (b₁ b₂ : A) : a • (b₁ + b₂) = a • b₁ + a • b₂ := distrib_smul.smul_add _ _ _ +/-- Pullback a distributive scalar multiplication along an injective additive monoid +homomorphism. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.injective.distrib_smul [add_zero_class B] [has_smul M B] (f : B →+ A) + (hf : injective f) (smul : ∀ (c : M) x, f (c • x) = c • f x) : + distrib_smul M B := +{ smul := (•), + smul_add := λ c x y, hf $ by simp only [smul, map_add, smul_add], + .. hf.smul_zero_class f.to_zero_hom smul } + +/-- Pushforward a distributive scalar multiplication along a surjective additive monoid +homomorphism. +See note [reducible non-instances]. -/ +@[reducible] +protected def function.surjective.distrib_smul [add_zero_class B] [has_smul M B] (f : A →+ B) + (hf : surjective f) (smul : ∀ (c : M) x, f (c • x) = c • f x) : + distrib_smul M B := +{ smul := (•), + smul_add := λ c x y, by { rcases hf x with ⟨x, rfl⟩, rcases hf y with ⟨y, rfl⟩, + simp only [smul_add, ← smul, ← map_add] }, + .. f.to_zero_hom.smul_zero_class smul } + +/-- Push forward the multiplication of `R` on `M` along a compatible surjective map `f : R → S`. + +See also `function.surjective.distrib_mul_action_left`. +-/ +@[reducible] +def function.surjective.distrib_smul_left {R S M : Type*} [add_zero_class M] [distrib_smul R M] + [has_smul S M] (f : R → S) (hf : function.surjective f) (hsmul : ∀ c (x : M), f c • x = c • x) : + distrib_smul S M := +{ smul := (•), + smul_add := hf.forall.mpr $ λ c x y, by simp only [hsmul, smul_add], + .. hf.smul_zero_class_left f hsmul } + +variable (A) + +/-- Compose a `distrib_smul` with a function, with scalar multiplication `f r' • m`. +See note [reducible non-instances]. -/ +@[reducible] def distrib_smul.comp_fun (f : N → M) : + distrib_smul N A := +{ smul := has_smul.comp.smul f, + smul_add := λ x, smul_add (f x), + .. smul_zero_class.comp_fun A f } + +/-- Each element of the scalars defines a additive monoid homomorphism. -/ +@[simps] +def distrib_smul.to_add_monoid_hom (x : M) : A →+ A := +{ to_fun := (•) x, + map_add' := smul_add x, + .. smul_zero_class.to_zero_hom A x } + +end distrib_smul + /-- Typeclass for multiplicative actions on additive structures. This generalizes group modules. -/ @[ext] class distrib_mul_action (M A : Type*) [monoid M] [add_monoid A] extends mul_action M A := @@ -537,8 +645,7 @@ protected def function.injective.distrib_mul_action [add_monoid B] [has_smul M B (hf : injective f) (smul : ∀ (c : M) x, f (c • x) = c • f x) : distrib_mul_action M B := { smul := (•), - smul_add := λ c x y, hf $ by simp only [smul, f.map_add, smul_add], - smul_zero := λ c, hf $ by simp only [smul, f.map_zero, smul_zero], + .. hf.distrib_smul f smul, .. hf.mul_action f smul } /-- Pushforward a distributive multiplicative action along a surjective additive monoid @@ -549,9 +656,7 @@ protected def function.surjective.distrib_mul_action [add_monoid B] [has_smul M (hf : surjective f) (smul : ∀ (c : M) x, f (c • x) = c • f x) : distrib_mul_action M B := { smul := (•), - smul_add := λ c x y, by { rcases hf x with ⟨x, rfl⟩, rcases hf y with ⟨y, rfl⟩, - simp only [smul_add, ← smul, ← f.map_add] }, - smul_zero := λ c, by simp only [← f.map_zero, ← smul, smul_zero], + .. hf.distrib_smul f smul, .. hf.mul_action f smul } /-- Push forward the action of `R` on `M` along a compatible surjective map `f : R →* S`. @@ -564,8 +669,7 @@ def function.surjective.distrib_mul_action_left {R S M : Type*} [monoid R] [add_ (f : R →* S) (hf : function.surjective f) (hsmul : ∀ c (x : M), f c • x = c • x) : distrib_mul_action S M := { smul := (•), - smul_zero := hf.forall.mpr $ λ c, by rw [hsmul, smul_zero], - smul_add := hf.forall.mpr $ λ c x y, by simp only [hsmul, smul_add], + .. hf.distrib_smul_left f hsmul, .. hf.mul_action_left f hsmul } variable (A) @@ -575,16 +679,13 @@ See note [reducible non-instances]. -/ @[reducible] def distrib_mul_action.comp_hom [monoid N] (f : N →* M) : distrib_mul_action N A := { smul := has_smul.comp.smul f, - smul_zero := λ x, smul_zero (f x), - smul_add := λ x, smul_add (f x), + .. distrib_smul.comp_fun A f, .. mul_action.comp_hom A f } /-- Each element of the monoid defines a additive monoid homomorphism. -/ @[simps] def distrib_mul_action.to_add_monoid_hom (x : M) : A →+ A := -{ to_fun := (•) x, - map_zero' := smul_zero x, - map_add' := smul_add x } +distrib_smul.to_add_monoid_hom A x variables (M) diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 55470ce78b889..9292567519370 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -69,7 +69,7 @@ variables {V : Type*} {V₂ : Type*} namespace finsupp lemma smul_sum {α : Type*} {β : Type*} {R : Type*} {M : Type*} - [has_zero β] [monoid R] [add_comm_monoid M] [distrib_mul_action R M] + [has_zero β] [add_comm_monoid M] [distrib_smul R M] {v : α →₀ β} {c : R} {h : α → β → M} : c • (v.sum h) = v.sum (λa b, c • h a b) := finset.smul_sum diff --git a/src/linear_algebra/quotient.lean b/src/linear_algebra/quotient.lean index 559dbecec9c26..6d0e5007241ac 100644 --- a/src/linear_algebra/quotient.lean +++ b/src/linear_algebra/quotient.lean @@ -109,6 +109,23 @@ function.surjective.mul_action mk (surjective_quot_mk _) P^.quotient.mk_smul instance mul_action (P : submodule R M) : mul_action R (M ⧸ P) := quotient.mul_action' P +instance smul_zero_class' [has_smul S R] [smul_zero_class S M] + [is_scalar_tower S R M] + (P : submodule R M) : smul_zero_class S (M ⧸ P) := +zero_hom.smul_zero_class ⟨mk, mk_zero _⟩ P^.quotient.mk_smul + +instance smul_zero_class (P : submodule R M) : smul_zero_class R (M ⧸ P) := +quotient.smul_zero_class' P + +instance distrib_smul' [has_smul S R] [distrib_smul S M] + [is_scalar_tower S R M] + (P : submodule R M) : distrib_smul S (M ⧸ P) := +function.surjective.distrib_smul + ⟨mk, rfl, λ _ _, rfl⟩ (surjective_quot_mk _) P^.quotient.mk_smul + +instance distrib_smul (P : submodule R M) : distrib_smul R (M ⧸ P) := +quotient.distrib_smul' P + instance distrib_mul_action' [monoid S] [has_smul S R] [distrib_mul_action S M] [is_scalar_tower S R M] (P : submodule R M) : distrib_mul_action S (M ⧸ P) := diff --git a/src/representation_theory/maschke.lean b/src/representation_theory/maschke.lean index 38d7d320ee236..afa617ccfe83b 100644 --- a/src/representation_theory/maschke.lean +++ b/src/representation_theory/maschke.lean @@ -104,8 +104,10 @@ In fact, the sum over `g : G` of the conjugate of `π` by `g` is a `k[G]`-linear def sum_of_conjugates_equivariant : W →ₗ[monoid_algebra k G] V := monoid_algebra.equivariant_of_linear_of_comm (π.sum_of_conjugates G) (λ g v, begin - dsimp [sum_of_conjugates], - simp only [linear_map.sum_apply, finset.smul_sum], + simp only [sum_of_conjugates, linear_map.sum_apply, + -- We have a `module (monoid_algebra k G)` instance but are working with `finsupp`s, + -- so help the elaborator unfold everything correctly. + @finset.smul_sum (monoid_algebra k G)], dsimp [conjugate], conv_lhs { rw [←finset.univ_map_embedding (mul_right_embedding g⁻¹)], diff --git a/test/instance_diamonds.lean b/test/instance_diamonds.lean index 4c6ed7631fd70..f6f3e1f2ba90a 100644 --- a/test/instance_diamonds.lean +++ b/test/instance_diamonds.lean @@ -150,9 +150,9 @@ end multiplicative section finsupp open finsupp -/-- `finsupp.comap_has_smul` can form a non-equal diamond with `finsupp.has_smul` -/ +/-- `finsupp.comap_has_smul` can form a non-equal diamond with `finsupp.smul_zero_class` -/ example {k : Type*} [semiring k] [nontrivial k] : - (finsupp.comap_has_smul : has_smul k (k →₀ k)) ≠ finsupp.has_smul := + (finsupp.comap_has_smul : has_smul k (k →₀ k)) ≠ finsupp.smul_zero_class.to_has_smul := begin obtain ⟨u : k, hu⟩ := exists_ne (1 : k), intro h, @@ -164,10 +164,10 @@ begin exact one_ne_zero h, end -/-- `finsupp.comap_has_smul` can form a non-equal diamond with `finsupp.has_smul` even when +/-- `finsupp.comap_has_smul` can form a non-equal diamond with `finsupp.smul_zero_class` even when the domain is a group. -/ example {k : Type*} [semiring k] [nontrivial kˣ] : - (finsupp.comap_has_smul : has_smul kˣ (kˣ →₀ k)) ≠ finsupp.has_smul := + (finsupp.comap_has_smul : has_smul kˣ (kˣ →₀ k)) ≠ finsupp.smul_zero_class.to_has_smul := begin obtain ⟨u : kˣ, hu⟩ := exists_ne (1 : kˣ), haveI : nontrivial k := ⟨⟨u, 1, units.ext.ne hu⟩⟩, From 563aed347eb59dc4181cb732cda0d124d736eaa3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 1 Oct 2022 11:17:06 +0000 Subject: [PATCH 488/828] feat(algebraic_topology/simplex_category): strong epi mono factorisations (#16276) In this PR, it is shown that there exists (unique) strong epi mono factorisations in `simplex_category`. --- src/algebraic_topology/simplex_category.lean | 97 ++++++++++--------- src/category_theory/functor/epi_mono.lean | 10 ++ src/category_theory/limits/shapes/images.lean | 25 +++++ src/order/category/NonemptyFinLinOrd.lean | 95 +++++++++++++++++- 4 files changed, 181 insertions(+), 46 deletions(-) diff --git a/src/algebraic_topology/simplex_category.lean b/src/algebraic_topology/simplex_category.lean index b1e224344c08e..fd05702bcada6 100644 --- a/src/algebraic_topology/simplex_category.lean +++ b/src/algebraic_topology/simplex_category.lean @@ -33,7 +33,7 @@ We provide the following functions to work with these objects: universe v -open category_theory +open category_theory category_theory.limits /-- The simplex category: * objects are natural numbers `n : ℕ` @@ -322,6 +322,10 @@ def skeletal_functor : simplex_category ⥤ NonemptyFinLinOrd.{v} := map_id' := λ a, by { ext, simp, }, map_comp' := λ a b c f g, by { ext, simp, }, } +lemma skeletal_functor.coe_map + {Δ₁ Δ₂ : simplex_category} (f : Δ₁ ⟶ Δ₂) : + coe_fn (skeletal_functor.{v}.map f) = ulift.up ∘ f.to_order_hom ∘ ulift.down := rfl + lemma skeletal : skeletal simplex_category := λ X Y ⟨I⟩, begin @@ -419,14 +423,11 @@ section epi_mono theorem mono_iff_injective {n m : simplex_category} {f : n ⟶ m} : mono f ↔ function.injective f.to_order_hom := begin - split, - { introsI m x y h, - have H : const n x ≫ f = const n y ≫ f, - { dsimp, rw h }, - change (n.const x).to_order_hom 0 = (n.const y).to_order_hom 0, - rw cancel_mono f at H, - rw H }, - { exact concrete_category.mono_of_injective f } + rw ← functor.mono_map_iff_mono skeletal_equivalence.functor.{0}, + dsimp only [skeletal_equivalence, functor.as_equivalence_functor], + rw [NonemptyFinLinOrd.mono_iff_injective, skeletal_functor.coe_map, + function.injective.of_comp_iff ulift.up_injective, + function.injective.of_comp_iff' _ ulift.down_bijective], end /-- A morphism in `simplex_category` is an epimorphism if and only if it is a surjective function @@ -434,39 +435,11 @@ end lemma epi_iff_surjective {n m : simplex_category} {f: n ⟶ m} : epi f ↔ function.surjective f.to_order_hom := begin - split, - { introsI hyp_f_epi x, - by_contra' h_ab, - -- The proof is by contradiction: assume f is not surjective, - -- then introduce two non-equal auxiliary functions equalizing f, and get a contradiction. - -- First we define the two auxiliary functions. - set chi_1 : m ⟶ [1] := hom.mk ⟨λ u, if u ≤ x then 0 else 1, begin - intros a b h, - dsimp only [], - split_ifs with h1 h2 h3, - any_goals { exact le_rfl }, - { exact bot_le }, - { exact false.elim (h1 (le_trans h h3)) } - end ⟩, - set chi_2 : m ⟶ [1] := hom.mk ⟨λ u, if u < x then 0 else 1, begin - intros a b h, - dsimp only [], - split_ifs with h1 h2 h3, - any_goals { exact le_rfl }, - { exact bot_le }, - { exact false.elim (h1 (lt_of_le_of_lt h h3)) } - end ⟩, - -- The two auxiliary functions equalize f - have f_comp_chi_i : f ≫ chi_1 = f ≫ chi_2, - { dsimp, - ext, - simp [le_iff_lt_or_eq, h_ab x_1] }, - -- We now just have to show the two auxiliary functions are not equal. - rw category_theory.cancel_epi f at f_comp_chi_i, rename f_comp_chi_i eq_chi_i, - apply_fun (λ e, e.to_order_hom x) at eq_chi_i, - suffices : (0 : fin 2) = 1, by exact bot_ne_top this, - simpa using eq_chi_i }, - { exact concrete_category.epi_of_surjective f } + rw ← functor.epi_map_iff_epi skeletal_equivalence.functor.{0}, + dsimp only [skeletal_equivalence, functor.as_equivalence_functor], + rw [NonemptyFinLinOrd.epi_iff_surjective, skeletal_functor.coe_map, + function.surjective.of_comp_iff' ulift.up_bijective, + function.surjective.of_comp_iff _ ulift.down_surjective], end /-- A monomorphism in `simplex_category` must increase lengths-/ @@ -571,7 +544,7 @@ begin refl, end -lemma eq_id_of_is_iso {x : simplex_category} {f : x ⟶ x} (hf : is_iso f) : f = 𝟙 _ := +lemma eq_id_of_is_iso {x : simplex_category} (f : x ⟶ x) [hf : is_iso f] : f = 𝟙 _ := congr_arg (λ (φ : _ ≅ _), φ.hom) (iso_eq_iso_refl (as_iso f)) lemma eq_σ_comp_of_not_injective' {n : ℕ} {Δ' : simplex_category} (θ : mk (n+1) ⟶ Δ') @@ -690,7 +663,8 @@ end lemma eq_id_of_mono {x : simplex_category} (i : x ⟶ x) [mono i] : i = 𝟙 _ := begin - apply eq_id_of_is_iso, + suffices : is_iso i, + { haveI := this, apply eq_id_of_is_iso, }, apply is_iso_of_bijective, dsimp, rw [fintype.bijective_iff_injective_and_card i.to_order_hom, ← mono_iff_injective, @@ -700,7 +674,8 @@ end lemma eq_id_of_epi {x : simplex_category} (i : x ⟶ x) [epi i] : i = 𝟙 _ := begin - apply eq_id_of_is_iso, + suffices : is_iso i, + { haveI := this, apply eq_id_of_is_iso, }, apply is_iso_of_bijective, dsimp, rw [fintype.bijective_iff_surjective_and_card i.to_order_hom, ← epi_iff_surjective, @@ -732,6 +707,38 @@ begin rw [h, eq_id_of_mono θ', category.id_comp], end +noncomputable instance : split_epi_category simplex_category := +skeletal_equivalence.{0}.inverse.split_epi_category_imp_of_is_equivalence + +instance : has_strong_epi_mono_factorisations simplex_category := +functor.has_strong_epi_mono_factorisations_imp_of_is_equivalence + simplex_category.skeletal_equivalence.{0}.inverse + +lemma image_eq {Δ Δ' Δ'' : simplex_category } {φ : Δ ⟶ Δ''} + {e : Δ ⟶ Δ'} [epi e] {i : Δ' ⟶ Δ''} [mono i] (fac : e ≫ i = φ) : + image φ = Δ' := +begin + haveI := strong_epi_of_epi e, + let e := image.iso_strong_epi_mono e i fac, + ext, + exact le_antisymm (len_le_of_epi (infer_instance : epi e.hom)) + (len_le_of_mono (infer_instance : mono e.hom)), +end + +lemma image_ι_eq {Δ Δ'' : simplex_category } {φ : Δ ⟶ Δ''} + {e : Δ ⟶ image φ} [epi e] {i : image φ ⟶ Δ''} [mono i] (fac : e ≫ i = φ) : + image.ι φ = i := +begin + haveI := strong_epi_of_epi e, + rw [← image.iso_strong_epi_mono_hom_comp_ι e i fac, + simplex_category.eq_id_of_is_iso (image.iso_strong_epi_mono e i fac).hom, category.id_comp], +end + +lemma factor_thru_image_eq {Δ Δ'' : simplex_category } {φ : Δ ⟶ Δ''} + {e : Δ ⟶ image φ} [epi e] {i : image φ ⟶ Δ''} [mono i] (fac : e ≫ i = φ) : + factor_thru_image φ = e := +by rw [← cancel_mono i, fac, ← image_ι_eq fac, image.fac] + end epi_mono /-- This functor `simplex_category ⥤ Cat` sends `[n]` (for `n : ℕ`) diff --git a/src/category_theory/functor/epi_mono.lean b/src/category_theory/functor/epi_mono.lean index 363bd3e4d8602..521a68d8ff733 100644 --- a/src/category_theory/functor/epi_mono.lean +++ b/src/category_theory/functor/epi_mono.lean @@ -249,6 +249,16 @@ begin exact F.map_mono f, }, end +/-- If `F : C ⥤ D` is an equivalence of categories and `C` is a `split_epi_category`, +then `D` also is. -/ +def split_epi_category_imp_of_is_equivalence [is_equivalence F] [split_epi_category C] : + split_epi_category D := +⟨λ X Y f, begin + introI, + rw ← F.inv.is_split_epi_iff f, + apply is_split_epi_of_epi, +end⟩ + end end category_theory.functor diff --git a/src/category_theory/limits/shapes/images.lean b/src/category_theory/limits/shapes/images.lean index 1457e0257f1e3..7ae28dbae9cab 100644 --- a/src/category_theory/limits/shapes/images.lean +++ b/src/category_theory/limits/shapes/images.lean @@ -824,3 +824,28 @@ lemma image.iso_strong_epi_mono_inv_comp_mono {I' : C} (e : X ⟶ I') (m : I' image.lift_fac _ end category_theory.limits + +namespace category_theory.functor + +open category_theory.limits + +variables {C D : Type*} [category C] [category D] + +lemma has_strong_epi_mono_factorisations_imp_of_is_equivalence (F : C ⥤ D) [is_equivalence F] + [h : has_strong_epi_mono_factorisations C] : + has_strong_epi_mono_factorisations D := +⟨λ X Y f, begin + let em : strong_epi_mono_factorisation (F.inv.map f) := + (has_strong_epi_mono_factorisations.has_fac (F.inv.map f)).some, + haveI : mono (F.map em.m ≫ F.as_equivalence.counit_iso.hom.app Y) := mono_comp _ _, + haveI : strong_epi (F.as_equivalence.counit_iso.inv.app X ≫ F.map em.e) := strong_epi_comp _ _, + exact nonempty.intro + { I := F.obj em.I, + e := F.as_equivalence.counit_iso.inv.app X ≫ F.map em.e, + m := F.map em.m ≫ F.as_equivalence.counit_iso.hom.app Y, + fac' := by simpa only [category.assoc, ← F.map_comp_assoc, em.fac', + is_equivalence.fun_inv_map, iso.inv_hom_id_app, iso.inv_hom_id_app_assoc] + using category.comp_id _, }, +end⟩ + +end category_theory.functor diff --git a/src/order/category/NonemptyFinLinOrd.lean b/src/order/category/NonemptyFinLinOrd.lean index b8a1b4a25f0f1..1a586aa688b4f 100644 --- a/src/order/category/NonemptyFinLinOrd.lean +++ b/src/order/category/NonemptyFinLinOrd.lean @@ -5,6 +5,8 @@ Authors: Johan Commelin -/ import data.fintype.order import order.category.LinearOrder +import category_theory.limits.shapes.images +import category_theory.limits.shapes.regular_mono /-! # Nonempty finite linear orders @@ -15,7 +17,7 @@ This is the index category for simplicial objects. universes u v -open category_theory +open category_theory category_theory.limits /-- A typeclass for nonempty finite linear orders. -/ class nonempty_fin_lin_ord (α : Type*) extends fintype α, linear_order α := @@ -80,6 +82,97 @@ equivalence.mk dual dual (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) (nat_iso.of_components (λ X, iso.mk $ order_iso.dual_dual X) $ λ X Y f, rfl) +lemma mono_iff_injective {A B : NonemptyFinLinOrd.{u}} (f : A ⟶ B) : + mono f ↔ function.injective f := +begin + refine ⟨_, concrete_category.mono_of_injective f⟩, + introI, + intros a₁ a₂ h, + let X : NonemptyFinLinOrd.{u} := ⟨ulift (fin 1)⟩, + let g₁ : X ⟶ A := ⟨λ x, a₁, λ x₁ x₂ h, by refl⟩, + let g₂ : X ⟶ A := ⟨λ x, a₂, λ x₁ x₂ h, by refl⟩, + change g₁ (ulift.up (0 : fin 1)) = g₂ (ulift.up (0 : fin 1)), + have eq : g₁ ≫ f = g₂ ≫ f := by { ext x, exact h, }, + rw cancel_mono at eq, + rw eq, +end + +lemma epi_iff_surjective {A B : NonemptyFinLinOrd.{u}} (f : A ⟶ B) : + epi f ↔ function.surjective f := +begin + split, + { introI, + by_contra' hf', + rcases hf' with ⟨m, hm⟩, + let Y : NonemptyFinLinOrd.{u} := ⟨ulift (fin 2)⟩, + let p₁ : B ⟶ Y := ⟨λ b, if b < m then ulift.up 0 else ulift.up 1, λ x₁ x₂ h, begin + simp only, + split_ifs with h₁ h₂ h₂, + any_goals { apply fin.zero_le, }, + { exfalso, + exact h₁ (lt_of_le_of_lt h h₂), }, + { refl, }, + end⟩, + let p₂ : B ⟶ Y := ⟨λ b, if b ≤ m then ulift.up 0 else ulift.up 1, λ x₁ x₂ h, begin + simp only, + split_ifs with h₁ h₂ h₂, + any_goals { apply fin.zero_le, }, + { exfalso, + exact h₁ (h.trans h₂), }, + { refl, }, + end⟩, + have h : p₁ m = p₂ m, + { congr, + rw ← cancel_epi f, + ext a : 2, + simp only [comp_apply, order_hom.coe_fun_mk], + split_ifs with h₁ h₂ h₂, + any_goals { refl, }, + { exfalso, exact h₂ (le_of_lt h₁), }, + { exfalso, exact hm a (eq_of_le_of_not_lt h₂ h₁), }, }, + simpa only [order_hom.coe_fun_mk, lt_self_iff_false, if_false, le_refl, if_true, + ulift.up_inj, fin.one_eq_zero_iff] using h, }, + { intro h, + exact concrete_category.epi_of_surjective f h, }, +end + +instance : split_epi_category NonemptyFinLinOrd.{u} := +⟨λ X Y f hf, begin + have H : ∀ (y : Y), nonempty (f⁻¹' { y }), + { rw epi_iff_surjective at hf, + intro y, + exact nonempty.intro ⟨(hf y).some, (hf y).some_spec⟩, }, + let φ : Y → X := λ y, (H y).some.1, + have hφ : ∀ (y : Y), f (φ y) = y := λ y, (H y).some.2, + refine is_split_epi.mk' ⟨⟨φ, _⟩, _⟩, swap, + { ext b, + apply hφ, }, + { intros a b, + contrapose, + intro h, + simp only [not_le] at h ⊢, + suffices : b ≤ a, + { apply lt_of_le_of_ne this, + intro h', + exfalso, + simpa only [h', lt_self_iff_false] using h, }, + simpa only [hφ] using f.monotone (le_of_lt h), }, +end⟩ + +instance : has_strong_epi_mono_factorisations NonemptyFinLinOrd.{u} := +⟨λ X Y f, begin + let I : NonemptyFinLinOrd.{u} := ⟨set.image (coe_fn f) ⊤, ⟨⟩⟩, + let e : X ⟶ I := ⟨λ x, ⟨f x, ⟨x, by tidy⟩⟩, λ x₁ x₂ h, f.monotone h⟩, + let m : I ⟶ Y := ⟨λ y, y, by tidy⟩, + haveI : epi e := by { rw epi_iff_surjective, tidy, }, + haveI : strong_epi e := strong_epi_of_epi e, + haveI : mono m := concrete_category.mono_of_injective _ (by tidy), + exact nonempty.intro + { I := I, + m := m, + e := e, }, +end⟩ + end NonemptyFinLinOrd lemma NonemptyFinLinOrd_dual_comp_forget_to_LinearOrder : From 9af20344b24ef1801b599d296aaed8b9fffdc360 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 1 Oct 2022 11:17:07 +0000 Subject: [PATCH 489/828] feat(algebraic_topology/dold_kan): decomposition of Q (#16357) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this PR, a decomposition of the endomorphisms `Q q` acting on `X _[n+1]` is obtained. In the case of abelian categories, it shows that the image of `Q q` is contained in a certain sum of images of degeneracy operators. Critical tools are also introduced in order to show (in a future PR) that the normalized Moore complex functor reflects isomorphisms. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- src/algebraic_topology/Moore_complex.lean | 4 +- .../dold_kan/decomposition.lean | 150 ++++++++++++++++++ src/algebraic_topology/simplicial_object.lean | 22 ++- src/data/fin/basic.lean | 19 +++ 4 files changed, 191 insertions(+), 4 deletions(-) create mode 100644 src/algebraic_topology/dold_kan/decomposition.lean diff --git a/src/algebraic_topology/Moore_complex.lean b/src/algebraic_topology/Moore_complex.lean index d9760eb05bb79..faf5988d0c504 100644 --- a/src/algebraic_topology/Moore_complex.lean +++ b/src/algebraic_topology/Moore_complex.lean @@ -131,8 +131,8 @@ chain_complex.of_hom _ _ _ _ _ _ end) (λ n, begin cases n; dsimp, - { ext, simp, erw f.naturality, refl, }, - { ext, simp, erw f.naturality, refl, }, + { ext, simp, }, + { ext, simp, }, end) end normalized_Moore_complex diff --git a/src/algebraic_topology/dold_kan/decomposition.lean b/src/algebraic_topology/dold_kan/decomposition.lean new file mode 100644 index 0000000000000..fa1471a071589 --- /dev/null +++ b/src/algebraic_topology/dold_kan/decomposition.lean @@ -0,0 +1,150 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.dold_kan.p_infty + +/-! + +# Decomposition of the Q endomorphisms + +In this file, we obtain a lemma `decomposition_Q` which expresses +explicitly the projection `(Q q).f (n+1) : X _[n+1] ⟶ X _[n+1]` +(`X : simplicial_object C` with `C` a preadditive category) as +a sum of terms which are postcompositions with degeneracies. + +(TODO @joelriou: when `C` is abelian, define the degenerate +subcomplex of the alternating face map complex of `X` and show +that it is a complement to the normalized Moore complex.) + +Then, we introduce an ad hoc structure `morph_components X n Z` which +can be used in order to define morphisms `X _[n+1] ⟶ Z` using the +decomposition provided by `decomposition_Q`. This shall play a critical +role in the proof that the functor +`N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ))` +reflects isomorphisms. + +-/ + +open category_theory category_theory.category category_theory.preadditive opposite +open_locale big_operators simplicial + +noncomputable theory + +namespace algebraic_topology + +namespace dold_kan + +variables {C : Type*} [category C] [preadditive C] {X X' : simplicial_object C} + +/-- In each positive degree, this lemma decomposes the idempotent endomorphism +`Q q` as a sum of morphisms which are postcompositions with suitable degeneracies. +As `Q q` is the complement projection to `P q`, this implies that in the case of +simplicial abelian groups, any $(n+1)$-simplex $x$ can be decomposed as +$x = x' + \sum (i=0}^{q-1} σ_{n-i}(y_i)$ where $x'$ is in the image of `P q` and +the $y_i$ are in degree $n$. -/ +lemma decomposition_Q (n q : ℕ) : + ((Q q).f (n+1) : X _[n+1] ⟶ X _[n+1]) = + ∑ (i : fin (n+1)) in finset.filter (λ i : fin(n+1), (i:ℕ) Date: Sat, 1 Oct 2022 11:17:08 +0000 Subject: [PATCH 490/828] feat(category_theory/morphism_property): miscellaneous basic results (#16444) This PR establishes miscellaneous results about `morphism_property` in category theory: how `stable_under_composition`, etc, behave with passing to the opposite category, the definition of the inverse image of a `morphism_property` by a functor, the definition of isomorphisms/monomorphisms/epimorphisms as morphism properties. --- .../limits/shapes/comm_sq.lean | 10 ++ src/category_theory/morphism_property.lean | 159 ++++++++++++++++++ 2 files changed, 169 insertions(+) diff --git a/src/category_theory/limits/shapes/comm_sq.lean b/src/category_theory/limits/shapes/comm_sq.lean index a58826108945c..63831ff1b8e2e 100644 --- a/src/category_theory/limits/shapes/comm_sq.lean +++ b/src/category_theory/limits/shapes/comm_sq.lean @@ -468,6 +468,16 @@ is_pullback.of_is_limit (is_limit.of_iso_limit (limits.pushout_cocone.is_colimit_equiv_is_limit_unop h.flip.cocone h.flip.is_colimit) h.to_comm_sq.flip.cocone_unop) +lemma of_horiz_is_iso [is_iso f] [is_iso inr] (sq : comm_sq f g inl inr) : + is_pushout f g inl inr := of_is_colimit' sq +begin + refine pushout_cocone.is_colimit.mk _ (λ s, inv inr ≫ s.inr) (λ s, _) (by tidy) (by tidy), + simp only [← cancel_epi f, s.condition, sq.w_assoc, is_iso.hom_inv_id_assoc], +end + +lemma of_vert_is_iso [is_iso g] [is_iso inl] (sq : comm_sq f g inl inr) : + is_pushout f g inl inr := (of_horiz_is_iso sq.flip).flip + end is_pushout namespace functor diff --git a/src/category_theory/morphism_property.lean b/src/category_theory/morphism_property.lean index c5da9502f0532..0a1f263e7c176 100644 --- a/src/category_theory/morphism_property.lean +++ b/src/category_theory/morphism_property.lean @@ -18,6 +18,8 @@ The following meta-properties are defined * `stable_under_composition`: `P` is stable under composition if `P f → P g → P (f ≫ g)`. * `stable_under_base_change`: `P` is stable under base change if in all pullback squares, the left map satisfies `P` if the right map satisfies it. +* `stable_under_cobase_change`: `P` is stable under cobase change if in all pushout + squares, the right map satisfies `P` if the left map satisfies it. -/ @@ -41,27 +43,69 @@ variable {C} namespace morphism_property +instance : has_subset (morphism_property C) := +⟨λ P₁ P₂, ∀ ⦃X Y : C⦄ (f : X ⟶ Y) (hf : P₁ f), P₂ f⟩ +instance : has_inter (morphism_property C) := +⟨λ P₁ P₂ X Y f, P₁ f ∧ P₂ f⟩ + +/-- The morphism property in `Cᵒᵖ` associated to a morphism property in `C` -/ +@[simp] def op (P : morphism_property C) : morphism_property Cᵒᵖ := λ X Y f, P f.unop + +/-- The morphism property in `C` associated to a morphism property in `Cᵒᵖ` -/ +@[simp] def unop (P : morphism_property Cᵒᵖ) : morphism_property C := λ X Y f, P f.op + +lemma unop_op (P : morphism_property C) : P.op.unop = P := rfl +lemma op_unop (P : morphism_property Cᵒᵖ) : P.unop.op = P := rfl + +/-- The inverse image of a `morphism_property D` by a functor `C ⥤ D` -/ +def inverse_image (P : morphism_property D) (F : C ⥤ D) : morphism_property C := +λ X Y f, P (F.map f) + /-- A morphism property `respects_iso` if it still holds when composed with an isomorphism -/ def respects_iso (P : morphism_property C) : Prop := (∀ {X Y Z} (e : X ≅ Y) (f : Y ⟶ Z), P f → P (e.hom ≫ f)) ∧ (∀ {X Y Z} (e : Y ≅ Z) (f : X ⟶ Y), P f → P (f ≫ e.hom)) +lemma respects_iso.op {P : morphism_property C} (h : respects_iso P) : respects_iso P.op := +⟨λ X Y Z e f hf, h.2 e.unop f.unop hf, λ X Y Z e f hf, h.1 e.unop f.unop hf⟩ + +lemma respects_iso.unop {P : morphism_property Cᵒᵖ} (h : respects_iso P) : respects_iso P.unop := +⟨λ X Y Z e f hf, h.2 e.op f.op hf, λ X Y Z e f hf, h.1 e.op f.op hf⟩ + /-- A morphism property is `stable_under_composition` if the composition of two such morphisms still falls in the class. -/ def stable_under_composition (P : morphism_property C) : Prop := ∀ ⦃X Y Z⦄ (f : X ⟶ Y) (g : Y ⟶ Z), P f → P g → P (f ≫ g) +lemma stable_under_composition.op {P : morphism_property C} (h : stable_under_composition P) : + stable_under_composition P.op := λ X Y Z f g hf hg, h g.unop f.unop hg hf + +lemma stable_under_composition.unop {P : morphism_property Cᵒᵖ} (h : stable_under_composition P) : + stable_under_composition P.unop := λ X Y Z f g hf hg, h g.op f.op hg hf + /-- A morphism property is `stable_under_inverse` if the inverse of a morphism satisfying the property still falls in the class. -/ def stable_under_inverse (P : morphism_property C) : Prop := ∀ ⦃X Y⦄ (e : X ≅ Y), P e.hom → P e.inv +lemma stable_under_inverse.op {P : morphism_property C} (h : stable_under_inverse P) : + stable_under_inverse P.op := λ X Y e he, h e.unop he + +lemma stable_under_inverse.unop {P : morphism_property Cᵒᵖ} (h : stable_under_inverse P) : + stable_under_inverse P.unop := λ X Y e he, h e.op he + /-- A morphism property is `stable_under_base_change` if the base change of such a morphism still falls in the class. -/ def stable_under_base_change (P : morphism_property C) : Prop := ∀ ⦃X Y Y' S : C⦄ ⦃f : X ⟶ S⦄ ⦃g : Y ⟶ S⦄ ⦃f' : Y' ⟶ Y⦄ ⦃g' : Y' ⟶ X⦄ (sq : is_pullback f' g' g f) (hg : P g), P g' +/-- A morphism property is `stable_under_cobase_change` if the cobase change of such a morphism +still falls in the class. -/ +def stable_under_cobase_change (P : morphism_property C) : Prop := +∀ ⦃A A' B B' : C⦄ ⦃f : A ⟶ A'⦄ ⦃g : A ⟶ B⦄ ⦃f' : B ⟶ B'⦄ ⦃g' : A' ⟶ B'⦄ + (sq : is_pushout g f f' g') (hf : P f), P f' + lemma stable_under_composition.respects_iso {P : morphism_property C} (hP : stable_under_composition P) (hP' : ∀ {X Y} (e : X ≅ Y), P e.hom) : respects_iso P := ⟨λ X Y Z e f hf, hP _ _ (hP' e) hf, λ X Y Z e f hf, hP _ _ hf (hP' e)⟩ @@ -164,6 +208,46 @@ begin hP.base_change_map _ (over.hom_mk _ e₁.symm : over.mk f ⟶ over.mk f') h₁], end +lemma stable_under_cobase_change.mk {P : morphism_property C} [has_pushouts C] + (hP₁ : respects_iso P) + (hP₂ : ∀ (A B A' : C) (f : A ⟶ A') (g : A ⟶ B) (hf : P f), P (pushout.inr : B ⟶ pushout f g)) : + stable_under_cobase_change P := λ A A' B B' f g f' g' sq hf, +begin + let e := sq.flip.iso_pushout, + rw [← hP₁.cancel_right_is_iso _ e.hom, sq.flip.inr_iso_pushout_hom], + exact hP₂ _ _ _ f g hf, +end + +lemma stable_under_cobase_change.respects_iso {P : morphism_property C} + (hP : stable_under_cobase_change P) : respects_iso P := +respects_iso.of_respects_arrow_iso _ (λ f g e, hP (is_pushout.of_horiz_is_iso (comm_sq.mk e.hom.w))) + +lemma stable_under_cobase_change.inl {P : morphism_property C} + (hP : stable_under_cobase_change P) {A B A' : C} (f : A ⟶ A') (g : A ⟶ B) [has_pushout f g] + (H : P g) : P (pushout.inl : A' ⟶ pushout f g) := +hP (is_pushout.of_has_pushout f g) H + +lemma stable_under_cobase_change.inr {P : morphism_property C} + (hP : stable_under_cobase_change P) {A B A' : C} (f : A ⟶ A') (g : A ⟶ B) [has_pushout f g] + (H : P f) : P (pushout.inr : B ⟶ pushout f g) := +hP (is_pushout.of_has_pushout f g).flip H + +lemma stable_under_cobase_change.op {P : morphism_property C} + (hP : stable_under_cobase_change P) : stable_under_base_change P.op := +λ X Y Y' S f g f' g' sq hg, hP sq.unop hg + +lemma stable_under_cobase_change.unop {P : morphism_property Cᵒᵖ} + (hP : stable_under_cobase_change P) : stable_under_base_change P.unop := +λ X Y Y' S f g f' g' sq hg, hP sq.op hg + +lemma stable_under_base_change.op {P : morphism_property C} + (hP : stable_under_base_change P) : stable_under_cobase_change P.op := +λ A A' B B' f g f' g' sq hf, hP sq.unop hf + +lemma stable_under_base_change.unop {P : morphism_property Cᵒᵖ} + (hP : stable_under_base_change P) : stable_under_cobase_change P.unop := +λ A A' B B' f g f' g' sq hf, hP sq.op hf + /-- If `P : morphism_property C` and `F : C ⥤ D`, then `P.is_inverted_by F` means that all morphisms in `P` are mapped by `F` to isomorphisms in `D`. -/ @@ -205,6 +289,81 @@ end end naturality_property +lemma respects_iso.inverse_image {P : morphism_property D} (h : respects_iso P) (F : C ⥤ D) : + respects_iso (P.inverse_image F) := +begin + split, + all_goals + { intros X Y Z e f hf, + dsimp [inverse_image], + rw F.map_comp, }, + exacts [h.1 (F.map_iso e) (F.map f) hf, h.2 (F.map_iso e) (F.map f) hf], +end + +lemma stable_under_composition.inverse_image {P : morphism_property D} + (h : stable_under_composition P) (F : C ⥤ D) : stable_under_composition (P.inverse_image F) := +λ X Y Z f g hf hg, by simpa only [← F.map_comp] using h (F.map f) (F.map g) hf hg + +variable (C) + +/-- The `morphism_property C` satisfied by isomorphisms in `C`. -/ +def isomorphisms : morphism_property C := λ X Y f, is_iso f + +/-- The `morphism_property C` satisfied by monomorphisms in `C`. -/ +def monomorphisms : morphism_property C := λ X Y f, mono f + +/-- The `morphism_property C` satisfied by epimorphisms in `C`. -/ +def epimorphisms : morphism_property C := λ X Y f, epi f + +section + +variables {C} {X Y : C} (f : X ⟶ Y) + +@[simp] lemma isomorphisms.iff : (isomorphisms C) f ↔ is_iso f := by refl +@[simp] lemma monomorphisms.iff : (monomorphisms C) f ↔ mono f := by refl +@[simp] lemma epimorphisms.iff : (epimorphisms C) f ↔ epi f := by refl + +lemma isomorphisms.infer_property [hf : is_iso f] : (isomorphisms C) f := hf +lemma monomorphisms.infer_property [hf : mono f] : (monomorphisms C) f := hf +lemma epimorphisms.infer_property [hf : epi f] : (epimorphisms C) f := hf + +end + +lemma respects_iso.monomorphisms : respects_iso (monomorphisms C) := +by { split; { intros X Y Z e f, simp only [monomorphisms.iff], introI, apply mono_comp, }, } + +lemma respects_iso.epimorphisms : respects_iso (epimorphisms C) := +by { split; { intros X Y Z e f, simp only [epimorphisms.iff], introI, apply epi_comp, }, } + +lemma respects_iso.isomorphisms : respects_iso (isomorphisms C) := +by { split; { intros X Y Z e f, simp only [isomorphisms.iff], introI, apply_instance, }, } + +lemma stable_under_composition.isomorphisms : stable_under_composition (isomorphisms C) := +λ X Y Z f g hf hg, begin + rw isomorphisms.iff at hf hg ⊢, + haveI := hf, + haveI := hg, + apply_instance, +end + +lemma stable_under_composition.monomorphisms : stable_under_composition (monomorphisms C) := +λ X Y Z f g hf hg, begin + rw monomorphisms.iff at hf hg ⊢, + haveI := hf, + haveI := hg, + apply mono_comp, +end + +lemma stable_under_composition.epimorphisms : stable_under_composition (epimorphisms C) := +λ X Y Z f g hf hg, begin + rw epimorphisms.iff at hf hg ⊢, + haveI := hf, + haveI := hg, + apply epi_comp, +end + +variable {C} + /-- The full subcategory of `C ⥤ D` consisting of functors inverting morphisms in `W` -/ @[derive category, nolint has_nonempty_instance] def functors_inverting (W : morphism_property C) (D : Type*) [category D] := From 7295c8118e292842d7f4771009453245c0e06c14 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sat, 1 Oct 2022 11:17:09 +0000 Subject: [PATCH 491/828] feat(topology/sheaves/skyscraper): skyscraper presheaf is a sheaf (#16649) - [x] depends on: #16694 --- src/topology/sheaves/skyscraper.lean | 24 ++++++++++++++++++++++-- 1 file changed, 22 insertions(+), 2 deletions(-) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index 6752445b8d277..2ab30db91089d 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -1,10 +1,10 @@ /- Copyright (c) 2022 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jujian Zhang +Authors: Jujian Zhang, Junyan Xu -/ import algebraic_geometry.sheafed_space -import topology.sheaves.sheaf_condition.opens_le_cover +import topology.sheaves.punit import topology.sheaves.stalks import category_theory.preadditive.injective @@ -64,6 +64,13 @@ point, then the skyscraper presheaf `𝓕` with value `A` is defined by `U ↦ A { rw [dif_neg hW], apply ((if_neg hW).symm.rec terminal_is_terminal).hom_ext } end } +lemma skyscraper_presheaf_eq_pushforward + [hd : Π (U : opens (Top.of punit.{u+1})), decidable (punit.star ∈ U)] : + skyscraper_presheaf p₀ A = + continuous_map.const (Top.of punit) p₀ _* skyscraper_presheaf punit.star A := +by convert_to @skyscraper_presheaf X p₀ + (λ U, hd $ (opens.map $ continuous_map.const _ p₀).obj U) C _ A _ = _; congr <|> refl + end section @@ -158,4 +165,17 @@ def skyscraper_presheaf_stalk_of_not_specializes_is_terminal [has_colimits C] {y : X} (h : ¬p₀ ⤳ y) : is_terminal ((skyscraper_presheaf p₀ A).stalk y) := is_terminal.of_iso terminal_is_terminal $ (skyscraper_presheaf_stalk_of_not_specializes _ _ h).symm +lemma skyscraper_presheaf_is_sheaf [has_products.{u} C] : (skyscraper_presheaf p₀ A).is_sheaf := +by classical; exact (presheaf.is_sheaf_iso_iff + (eq_to_iso $ skyscraper_presheaf_eq_pushforward p₀ A)).mpr + (sheaf.pushforward_sheaf_of_sheaf _ (presheaf.is_sheaf_on_punit_of_is_terminal _ + (by { dsimp, rw if_neg, exact terminal_is_terminal, exact set.not_mem_empty punit.star }))) + +/-- +The skyscraper presheaf supported at `p₀` with value `A` is the sheaf that assigns `A` to all opens +`U` that contain `p₀` and assigns `*` otherwise. +-/ +def skyscraper_sheaf [has_products.{u} C] : sheaf C X := +⟨skyscraper_presheaf p₀ A, skyscraper_presheaf_is_sheaf _ _⟩ + end From bf783181ae868c2f87e0d6af68fd26dc6e421943 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 1 Oct 2022 14:05:57 +0000 Subject: [PATCH 492/828] feat(number_theory/number_field): add mem_roots_of_unity_of_norm_eq_one (#15143) We prove that an algebraic integer whose conjugates are all of norm 1 is a root of unity. For that, we prove first that the set of algebraic integers (in a fixed number field) with bounded conjugates is finite. The counterpart of the result, that is roots of unity are of norm 1, is #16426 From flt-regular Co-authored-by: Alex J. Best [alex.j.best@gmail.com](mailto:alex.j.best@gmail.com) Co-authored-by: Riccardo Brasca --- src/field_theory/adjoin.lean | 8 +++ src/number_theory/number_field.lean | 81 +++++++++++++++++++++++++++-- 2 files changed, 85 insertions(+), 4 deletions(-) diff --git a/src/field_theory/adjoin.lean b/src/field_theory/adjoin.lean index 453ae38d9b20f..88ef68c1027d7 100644 --- a/src/field_theory/adjoin.lean +++ b/src/field_theory/adjoin.lean @@ -678,6 +678,14 @@ begin refl end +lemma _root_.minpoly.nat_degree_le {x : L} [finite_dimensional K L] (hx : is_integral K x) : + (minpoly K x).nat_degree ≤ finrank K L := +le_of_eq_of_le (intermediate_field.adjoin.finrank hx).symm K⟮x⟯.to_submodule.finrank_le + +lemma _root_.minpoly.degree_le {x : L} [finite_dimensional K L] (hx : is_integral K x) : + (minpoly K x).degree ≤ finrank K L := +degree_le_of_nat_degree_le (minpoly.nat_degree_le hx) + end power_basis /-- Algebra homomorphism `F⟮α⟯ →ₐ[F] K` are in bijection with the set of roots diff --git a/src/number_theory/number_field.lean b/src/number_theory/number_field.lean index 0317ff7dc52c9..acc6a8921d06a 100644 --- a/src/number_theory/number_field.lean +++ b/src/number_theory/number_field.lean @@ -6,6 +6,8 @@ Authors: Ashvni Narayanan, Anne Baanen import ring_theory.dedekind_domain.integral_closure import algebra.char_p.algebra +import analysis.normed_space.basic +import topology.algebra.polynomial /-! # Number fields @@ -199,7 +201,7 @@ open set polynomial /-- Let `A` an algebraically closed field and let `x ∈ K`, with `K` a number field. For `F`, subfield of `K`, the images of `x` by the `F`-algebra morphisms from `K` to `A` are exactly -the roots in `A` of the minimal polynomial of `x` over `F` -/ +the roots in `A` of the minimal polynomial of `x` over `F`. -/ lemma range_eq_roots (F K A : Type*) [field F] [number_field F] [field K] [number_field K] [field A] [is_alg_closed A] [algebra F K] [algebra F A] (x : K) : range (λ ψ : K →ₐ[F] A, ψ x) = (minpoly F x).root_set A := @@ -238,9 +240,8 @@ variables (K A : Type*) [field K] [number_field K] [field A] [char_zero A] [is_a /-- Let `A` be an algebraically closed field and let `x ∈ K`, with `K` a number field. The images of `x` by the embeddings of `K` in `A` are exactly the roots in `A` of -the minimal polynomial of `x` over `ℚ` -/ -lemma rat_range_eq_roots : -range (λ φ : K →+* A, φ x) = (minpoly ℚ x).root_set A := +the minimal polynomial of `x` over `ℚ`. -/ +lemma rat_range_eq_roots : range (λ φ : K →+* A, φ x) = (minpoly ℚ x).root_set A := begin convert range_eq_roots ℚ K A x using 1, ext a, @@ -249,4 +250,76 @@ end end roots +section bounded + +open finite_dimensional polynomial set + +variables {K : Type*} [field K] [number_field K] +variables {A : Type*} [normed_field A] [is_alg_closed A] [normed_algebra ℚ A] + +lemma coeff_bdd_of_norm_le {B : ℝ} {x : K} (h : ∀ φ : K →+* A, ∥φ x∥ ≤ B) (i : ℕ): + ∥(minpoly ℚ x).coeff i∥ ≤ (max B 1) ^ (finrank ℚ K) * (finrank ℚ K).choose ((finrank ℚ K) / 2) := +begin + have hx : is_integral ℚ x := is_separable.is_integral _ _, + rw (by rw [coeff_map, norm_algebra_map'] : + ∥(minpoly ℚ x).coeff i∥ = ∥(map (algebra_map ℚ A) (minpoly ℚ x)).coeff i∥), + refine coeff_bdd_of_roots_le _ (minpoly.monic hx) (is_alg_closed.splits_codomain _) + (minpoly.nat_degree_le hx) _ i, + intros z hz, + rsuffices ⟨φ, rfl⟩ : ∃ (φ : K →+* A), φ x = z, {exact h φ }, + letI : char_zero A := char_zero_of_injective_algebra_map (algebra_map ℚ _).injective, + rw [← set.mem_range, rat_range_eq_roots, mem_root_set_iff, aeval_def], + convert (mem_roots_map _).mp hz, + repeat { exact monic.ne_zero (minpoly.monic hx), }, +end + +variables (K A) + +/-- Let `B` be a real number. The set of algebraic integers in `K` whose conjugates are all +smaller in norm than `B` is finite. -/ +lemma finite_of_norm_le (B : ℝ) : + {x : K | is_integral ℤ x ∧ ∀ φ : K →+* A, ∥φ x∥ ≤ B}.finite := +begin + let C := nat.ceil ((max B 1) ^ (finrank ℚ K) * (finrank ℚ K).choose ((finrank ℚ K) / 2)), + have := bUnion_roots_finite (algebra_map ℤ K) (finrank ℚ K) (finite_Icc (-C : ℤ) C), + refine this.subset (λ x hx, _), + have h_map_rat_minpoly := minpoly.gcd_domain_eq_field_fractions' ℚ hx.1, + rw mem_Union, + use minpoly ℤ x, + rw [mem_Union, exists_prop], + refine ⟨⟨_, λ i, _⟩, _⟩, + { rw [← nat_degree_map_eq_of_injective (algebra_map ℤ ℚ).injective_int _, ← h_map_rat_minpoly], + exact minpoly.nat_degree_le (is_integral_of_is_scalar_tower _ hx.1), }, + { rw [mem_Icc, ← abs_le, ← @int.cast_le ℝ], + apply le_trans _ (nat.le_ceil _), + convert coeff_bdd_of_norm_le hx.2 i, + simp only [h_map_rat_minpoly, coeff_map, eq_int_cast, int.norm_cast_rat, int.norm_eq_abs, + int.cast_abs], }, + { rw [finset.mem_coe, multiset.mem_to_finset, mem_roots, is_root.def, ← eval₂_eq_eval_map, + ← aeval_def], + { exact minpoly.aeval ℤ x, }, + { exact monic.ne_zero (monic.map (algebra_map ℤ K) (minpoly.monic hx.1)), }}, +end + +/-- An algebraic integer whose conjugates are all of norm one is a root of unity. -/ +lemma pow_eq_one_of_norm_eq_one {x : K} + (hxi : is_integral ℤ x) (hx : ∀ φ : K →+* A, ∥φ x∥ = 1) : + ∃ (n : ℕ) (hn : 0 < n), x ^ n = 1 := +begin + obtain ⟨a, -, b, -, habne, h⟩ := @set.infinite.exists_ne_map_eq_of_maps_to _ _ _ _ + ((^) x : ℕ → K) set.infinite_univ _ (finite_of_norm_le K A (1:ℝ)), + { replace habne := habne.lt_or_lt, + wlog : a < b := habne using [a b], + refine ⟨b - a, tsub_pos_of_lt habne, _⟩, + have hxne : x ≠ 0, + { contrapose! hx, + simp only [hx, norm_zero, ring_hom.map_zero, ne.def, not_false_iff, zero_ne_one], + use (is_alg_closed.lift (number_field.is_algebraic K)).to_ring_hom, }, + { rw [pow_sub₀ _ hxne habne.le, h, mul_inv_cancel (pow_ne_zero b hxne)], }}, + { rw set.maps_univ_to, + exact λ a, ⟨hxi.pow a, λ φ, by simp [hx φ, norm_pow, one_pow]⟩, }, +end + +end bounded + end number_field.embeddings From 180f4f20d8f9833781682b52d27415bd5d4166e2 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Sat, 1 Oct 2022 14:05:58 +0000 Subject: [PATCH 493/828] feat(analysis/locally_convex): locally bounded implies continuous (#16550) We prove that locally bounded linear maps are continuous provided the domain is first countable. In the literature this is usually stated with pseudometrizable, but first countable is equivalent. --- src/analysis/locally_convex/bounded.lean | 121 +++++++++++++++++++++-- 1 file changed, 114 insertions(+), 7 deletions(-) diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index d14d5a29d2f6b..971e94e01dc22 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -4,10 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Moritz Doll -/ import analysis.locally_convex.basic +import analysis.locally_convex.balanced_core_hull +import analysis.normed_space.is_R_or_C import analysis.seminorm import topology.bornology.basic import topology.algebra.uniform_group -import analysis.locally_convex.balanced_core_hull +import topology.uniform_space.cauchy /-! # Von Neumann Boundedness @@ -24,6 +26,9 @@ absorbs `s`. * `bornology.is_vonN_bounded_of_topological_space_le`: A coarser topology admits more von Neumann-bounded sets. +* `bornology.is_vonN_bounded.image`: A continuous linear image of a bounded set is bounded. +* `linear_map.continuous_of_locally_bounded`: If `E` is first countable, then every +locally bounded linear map `E →ₛₗ[σ] F` is continuous. ## References @@ -31,7 +36,7 @@ von Neumann-bounded sets. -/ -variables {𝕜 E E' F ι : Type*} +variables {𝕜 𝕜' E E' F ι : Type*} open filter open_locale topological_space pointwise @@ -194,11 +199,13 @@ end uniform_add_group section continuous_linear_map -variables [nontrivially_normed_field 𝕜] -variables [add_comm_group E] [module 𝕜 E] -variables [uniform_space E] [uniform_add_group E] [has_continuous_smul 𝕜 E] -variables [add_comm_group F] [module 𝕜 F] -variables [uniform_space F] [uniform_add_group F] +variables [add_comm_group E] [uniform_space E] [uniform_add_group E] +variables [add_comm_group F] [uniform_space F] + +section nontrivially_normed_field + +variables [uniform_add_group F] +variables [nontrivially_normed_field 𝕜] [module 𝕜 E] [module 𝕜 F] [has_continuous_smul 𝕜 E] /-- Construct a continuous linear map from a linear map `f : E →ₗ[𝕜] F` and the existence of a neighborhood of zero that gets mapped into a bounded set in `F`. -/ @@ -237,6 +244,106 @@ lemma linear_map.clm_of_exists_bounded_image_coe {f : E →ₗ[𝕜] F} {h : ∃ (V : set E) (hV : V ∈ 𝓝 (0 : E)), bornology.is_vonN_bounded 𝕜 (f '' V)} {x : E} : f.clm_of_exists_bounded_image h x = f x := rfl +end nontrivially_normed_field + +section is_R_or_C + +open topological_space bornology + +variables [first_countable_topology E] +variables [is_R_or_C 𝕜] [module 𝕜 E] [has_continuous_smul 𝕜 E] +variables [is_R_or_C 𝕜'] [module 𝕜' F] [has_continuous_smul 𝕜' F] +variables {σ : 𝕜 →+* 𝕜'} + +lemma linear_map.continuous_at_zero_of_locally_bounded (f : E →ₛₗ[σ] F) + (hf : ∀ (s : set E) (hs : is_vonN_bounded 𝕜 s), is_vonN_bounded 𝕜' (f '' s)) : + continuous_at f 0 := +begin + -- Assume that f is not continuous at 0 + by_contradiction, + -- We use the a decreasing balanced basis for 0 : E and a balanced basis for 0 : F + -- and reformulate non-continuity in terms of these bases + rcases (nhds_basis_balanced 𝕜 E).exists_antitone_subbasis with ⟨b, bE1, bE⟩, + simp only [id.def] at bE, + have bE' : (𝓝 (0 : E)).has_basis (λ (x : ℕ), x ≠ 0) (λ n : ℕ, (n : 𝕜)⁻¹ • b n) := + begin + refine bE.1.to_has_basis _ _, + { intros n _, + use n+1, + simp only [ne.def, nat.succ_ne_zero, not_false_iff, nat.cast_add, nat.cast_one, true_and], + -- `b (n + 1) ⊆ b n` follows from `antitone`. + have h : b (n + 1) ⊆ b n := bE.2 (by simp), + refine subset_trans _ h, + rintros y ⟨x, hx, hy⟩, + -- Since `b (n + 1)` is balanced `(n+1)⁻¹ b (n + 1) ⊆ b (n + 1)` + rw ←hy, + refine (bE1 (n+1)).2.smul_mem _ hx, + have h' : 0 < (n : ℝ) + 1 := n.cast_add_one_pos, + rw [norm_inv, ←nat.cast_one, ←nat.cast_add, is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat, + nat.cast_add, nat.cast_one, inv_le h' zero_lt_one], + norm_cast, + simp, }, + intros n hn, + -- The converse direction follows from continuity of the scalar multiplication + have hcont : continuous_at (λ (x : E), (n : 𝕜) • x) 0 := + (continuous_const_smul (n : 𝕜)).continuous_at, + simp only [continuous_at, map_zero, smul_zero] at hcont, + rw bE.1.tendsto_left_iff at hcont, + rcases hcont (b n) (bE1 n).1 with ⟨i, _, hi⟩, + refine ⟨i, trivial, λ x hx, ⟨(n : 𝕜) • x, hi hx, _⟩⟩, + simp [←mul_smul, hn], + end, + rw [continuous_at, map_zero, bE'.tendsto_iff (nhds_basis_balanced 𝕜' F)] at h, + push_neg at h, + rcases h with ⟨V, ⟨hV, hV'⟩, h⟩, + simp only [id.def, forall_true_left] at h, + -- There exists `u : ℕ → E` such that for all `n : ℕ` we have `u n ∈ n⁻¹ • b n` and `f (u n) ∉ V` + choose! u hu hu' using h, + -- The sequence `(λ n, n • u n)` converges to `0` + have h_tendsto : tendsto (λ n : ℕ, (n : 𝕜) • u n) at_top (𝓝 (0 : E)) := + begin + apply bE.tendsto, + intros n, + by_cases h : n = 0, + { rw [h, nat.cast_zero, zero_smul], + refine mem_of_mem_nhds (bE.1.mem_of_mem $ by triv) }, + rcases hu n h with ⟨y, hy, hu1⟩, + convert hy, + rw [←hu1, ←mul_smul], + simp only [h, mul_inv_cancel, ne.def, nat.cast_eq_zero, not_false_iff, one_smul], + end, + -- The image `(λ n, n • u n)` is von Neumann bounded: + have h_bounded : is_vonN_bounded 𝕜 (set.range (λ n : ℕ, (n : 𝕜) • u n)) := + h_tendsto.cauchy_seq.totally_bounded_range.is_vonN_bounded 𝕜, + -- Since `range u` is bounded it absorbs `V` + rcases hf _ h_bounded hV with ⟨r, hr, h'⟩, + cases exists_nat_gt r with n hn, + -- We now find a contradiction between `f (u n) ∉ V` and the absorbing property + have h1 : r ≤ ∥(n : 𝕜')∥ := + by { rw [is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat], exact hn.le }, + have hn' : 0 < ∥(n : 𝕜')∥ := lt_of_lt_of_le hr h1, + rw [norm_pos_iff, ne.def, nat.cast_eq_zero] at hn', + have h'' : f (u n) ∈ V := + begin + simp only [set.image_subset_iff] at h', + specialize h' (n : 𝕜') h1 (set.mem_range_self n), + simp only [set.mem_preimage, linear_map.map_smulₛₗ, map_nat_cast] at h', + rcases h' with ⟨y, hy, h'⟩, + apply_fun (λ y : F, (n : 𝕜')⁻¹ • y) at h', + simp only [hn', inv_smul_smul₀, ne.def, nat.cast_eq_zero, not_false_iff] at h', + rwa ←h', + end, + exact hu' n hn' h'', +end + +/-- If `E` is first countable, then every locally bounded linear map `E →ₛₗ[σ] F` is continuous. -/ +lemma linear_map.continuous_of_locally_bounded [uniform_add_group F] (f : E →ₛₗ[σ] F) + (hf : ∀ (s : set E) (hs : is_vonN_bounded 𝕜 s), is_vonN_bounded 𝕜' (f '' s)) : + continuous f := +(uniform_continuous_of_continuous_at_zero f $ f.continuous_at_zero_of_locally_bounded hf).continuous + +end is_R_or_C + end continuous_linear_map section vonN_bornology_eq_metric From 21347ef22c3e7384ec305434f4dc810ebd746a5d Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Sat, 1 Oct 2022 14:05:59 +0000 Subject: [PATCH 494/828] feat(analysis/convolution): prove that parametric convolution tends to a specific value (#16704) * Prove `convolution_tendsto_right` which shows that the convolution tends to a value when all three relevant arguments vary (in the previous version only the first argument varied). * This lemma can be used to simplify a proof in the sphere eversion project --- src/analysis/calculus/specific_functions.lean | 14 +++ src/analysis/convolution.lean | 112 ++++++++++-------- src/topology/metric_space/basic.lean | 26 +++- 3 files changed, 104 insertions(+), 48 deletions(-) diff --git a/src/analysis/calculus/specific_functions.lean b/src/analysis/calculus/specific_functions.lean index d55ef2521e7f6..5756515b7b3ac 100644 --- a/src/analysis/calculus/specific_functions.lean +++ b/src/analysis/calculus/specific_functions.lean @@ -452,6 +452,20 @@ by simp_rw [tsupport, f.support_normed_eq, closure_ball _ f.R_pos.ne'] lemma has_compact_support_normed : has_compact_support (f.normed μ) := by simp_rw [has_compact_support, f.tsupport_normed_eq, is_compact_closed_ball] +lemma tendsto_support_normed_small_sets {ι} {φ : ι → cont_diff_bump_of_inner c} {l : filter ι} + (hφ : tendsto (λ i, (φ i).R) l (𝓝 0)) : + tendsto (λ i, support (λ x, (φ i).normed μ x)) l (𝓝 c).small_sets := +begin + simp_rw [normed_add_comm_group.tendsto_nhds_zero, real.norm_eq_abs, + abs_eq_self.mpr (φ _).R_pos.le] at hφ, + rw [tendsto_small_sets_iff], + intros t ht, + rcases metric.mem_nhds_iff.mp ht with ⟨ε, hε, ht⟩, + refine (hφ ε hε).mono (λ i hi, subset_trans _ ht), + simp_rw [(φ i).support_normed_eq], + exact ball_subset_ball hi.le +end + variable (μ) lemma integral_normed_smul (z : X) [complete_space X] : ∫ x, f.normed μ x • z ∂μ = z := by simp_rw [integral_smul_const, f.integral_normed, one_smul] diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index cbdf8be0e01fa..35d5e7879a533 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -78,7 +78,7 @@ The following notations are localized in the locale `convolution`: open set function filter measure_theory measure_theory.measure topological_space open continuous_linear_map metric -open_locale pointwise topological_space nnreal +open_locale pointwise topological_space nnreal filter variables {𝕜 G E E' E'' F F' F'' : Type*} variables [normed_add_comm_group E] [normed_add_comm_group E'] [normed_add_comm_group E''] @@ -605,26 +605,26 @@ variables [is_add_left_invariant μ] [sigma_finite μ] /-- Approximate `(f ⋆ g) x₀` if the support of the `f` is bounded within a ball, and `g` is near `g x₀` on a ball with the same radius around `x₀`. See `dist_convolution_le` for a special case. -We can simplify the second argument of `dist` further if we assume `f` is integrable, but also if -`L = (•)` or more generally if `L` has a `antilipschitz_with`-condition. -/ -lemma dist_convolution_le' {x₀ : G} {R ε : ℝ} +We can simplify the second argument of `dist` further if we add some extra type-classes on `E` +and `𝕜` or if `L` is scalar multiplication. -/ +lemma dist_convolution_le' {x₀ : G} {R ε : ℝ} {z₀ : E'} (hε : 0 ≤ ε) (hif : integrable f μ) (hf : support f ⊆ ball (0 : G) R) (hmg : ae_strongly_measurable g μ) - (hg : ∀ x ∈ ball x₀ R, dist (g x) (g x₀) ≤ ε) : - dist ((f ⋆[L, μ] g : G → F) x₀) (∫ t, L (f t) (g x₀) ∂μ) ≤ ∥L∥ * ∫ x, ∥f x∥ ∂μ * ε := + (hg : ∀ x ∈ ball x₀ R, dist (g x) z₀ ≤ ε) : + dist ((f ⋆[L, μ] g : G → F) x₀) (∫ t, L (f t) z₀ ∂μ) ≤ ∥L∥ * ∫ x, ∥f x∥ ∂μ * ε := begin have hfg : convolution_exists_at f g x₀ L μ, { refine bdd_above.convolution_exists_at L _ metric.is_open_ball.measurable_set (subset_trans _ hf) hif.integrable_on hif.ae_strongly_measurable hmg, swap, { refine λ t, mt (λ ht : f t = 0, _), simp_rw [ht, L.map_zero₂] }, rw [bdd_above_def], - refine ⟨∥g x₀∥ + ε, _⟩, + refine ⟨∥z₀∥ + ε, _⟩, rintro _ ⟨x, hx, rfl⟩, refine norm_le_norm_add_const_of_dist_le (hg x _), rwa [mem_ball_iff_norm, norm_sub_rev, ← mem_ball_zero_iff] }, - have h2 : ∀ t, dist (L (f t) (g (x₀ - t))) (L (f t) (g x₀)) ≤ ∥L (f t)∥ * ε, + have h2 : ∀ t, dist (L (f t) (g (x₀ - t))) (L (f t) z₀) ≤ ∥L (f t)∥ * ε, { intro t, by_cases ht : t ∈ support f, { have h2t := hf ht, rw [mem_ball_zero_iff] at h2t, @@ -636,7 +636,7 @@ begin simp_rw [ht, L.map_zero₂, L.map_zero, norm_zero, zero_mul, dist_self] } }, simp_rw [convolution_def], simp_rw [dist_eq_norm] at h2 ⊢, - rw [← integral_sub hfg.integrable], swap, { exact (L.flip (g x₀)).integrable_comp hif }, + rw [← integral_sub hfg.integrable], swap, { exact (L.flip z₀).integrable_comp hif }, refine (norm_integral_le_of_norm_le ((L.integrable_comp hif).norm.mul_const ε) (eventually_of_forall h2)).trans _, rw [integral_mul_right], @@ -653,14 +653,14 @@ on a ball with the same radius around `x₀`. This is a special case of `dist_convolution_le'` where `L` is `(•)`, `f` has integral 1 and `f` is nonnegative. -/ -lemma dist_convolution_le {f : G → ℝ} {x₀ : G} {R ε : ℝ} +lemma dist_convolution_le {f : G → ℝ} {x₀ : G} {R ε : ℝ} {z₀ : E'} (hε : 0 ≤ ε) (hf : support f ⊆ ball (0 : G) R) (hnf : ∀ x, 0 ≤ f x) (hintf : ∫ x, f x ∂μ = 1) (hmg : ae_strongly_measurable g μ) - (hg : ∀ x ∈ ball x₀ R, dist (g x) (g x₀) ≤ ε) : - dist ((f ⋆[lsmul ℝ ℝ, μ] g : G → E') x₀) (g x₀) ≤ ε := + (hg : ∀ x ∈ ball x₀ R, dist (g x) z₀ ≤ ε) : + dist ((f ⋆[lsmul ℝ ℝ, μ] g : G → E') x₀) z₀ ≤ ε := begin have hif : integrable f μ, { by_contra hif, exact zero_ne_one ((integral_undef hif).symm.trans hintf) }, @@ -670,25 +670,44 @@ begin exact (mul_le_mul_of_nonneg_right op_norm_lsmul_le hε).trans_eq (one_mul ε) } end -/-- `(φ i ⋆ g) x₀` tends to `g x₀` if `φ` is a sequence of nonnegative functions with integral 1 -whose support tends to small neighborhoods around `(0 : G)` and `g` is continuous at `x₀`. +/-- `(φ i ⋆ g i) (k i)` tends to `z₀` as `i` tends to some filter `l` if +* `φ` is a sequence of nonnegative functions with integral `1` as `i` tends to `l`; +* The support of `φ` tends to small neighborhoods around `(0 : G)` as `i` tends to `l`; +* `g i` is `mu`-a.e. strongly measurable as `i` tends to `l`; +* `g i x` tends to `z₀` as `(i, x)` tends to `l ×ᶠ 𝓝 x₀`; +* `k i` tends to `x₀`. -See also `cont_diff_bump_of_inner.convolution_tendsto_right'`. -/ -lemma convolution_tendsto_right {ι} {l : filter ι} {φ : ι → G → ℝ} - (hnφ : ∀ i x, 0 ≤ φ i x) - (hiφ : ∀ i, ∫ s, φ i s ∂μ = 1) +See also `cont_diff_bump_of_inner.convolution_tendsto_right`. +-/ +lemma convolution_tendsto_right + {ι} {g : ι → G → E'} {l : filter ι} {x₀ : G} {z₀ : E'} + {φ : ι → G → ℝ} {k : ι → G} + (hnφ : ∀ᶠ i in l, ∀ x, 0 ≤ φ i x) + (hiφ : ∀ᶠ i in l, ∫ x, φ i x ∂μ = 1) -- todo: we could weaken this to "the integral tends to 1" (hφ : tendsto (λ n, support (φ n)) l (𝓝 0).small_sets) - (hmg : ae_strongly_measurable g μ) {x₀ : G} (hcg : continuous_at g x₀) : - tendsto (λ i, (φ i ⋆[lsmul ℝ ℝ, μ] g : G → E') x₀) l (𝓝 (g x₀)) := + (hmg : ∀ᶠ i in l, ae_strongly_measurable (g i) μ) + (hcg : tendsto (uncurry g) (l ×ᶠ 𝓝 x₀) (𝓝 z₀)) + (hk : tendsto k l (𝓝 x₀)) : + tendsto (λ i : ι, (φ i ⋆[lsmul ℝ ℝ, μ] g i : G → E') (k i)) l (𝓝 z₀) := begin simp_rw [tendsto_small_sets_iff] at hφ, - rw [metric.continuous_at_iff] at hcg, - rw [metric.tendsto_nhds], + rw [metric.tendsto_nhds] at hcg ⊢, + simp_rw [metric.eventually_prod_nhds_iff] at hcg, intros ε hε, - rcases hcg (ε / 2) (half_pos hε) with ⟨δ, hδ, hgδ⟩, - refine (hφ (ball (0 : G) δ) $ ball_mem_nhds _ hδ).mono (λ i hi, _), - exact (dist_convolution_le (half_pos hε).le hi (hnφ i) (hiφ i) hmg (λ x hx, (hgδ hx.out).le)) - .trans_lt (half_lt_self hε) + have h2ε : 0 < ε / 3 := div_pos hε (by norm_num), + obtain ⟨p, hp, δ, hδ, hgδ⟩ := hcg _ h2ε, + dsimp only [uncurry] at hgδ, + have h2k := hk.eventually (ball_mem_nhds x₀ $ half_pos hδ), + have h2φ := (hφ (ball (0 : G) _) $ ball_mem_nhds _ (half_pos hδ)), + filter_upwards [hp, h2k, h2φ, hnφ, hiφ, hmg] with i hpi hki hφi hnφi hiφi hmgi, + have hgi : dist (g i (k i)) z₀ < ε / 3 := hgδ hpi (hki.trans $ half_lt_self hδ), + have h1 : ∀ x' ∈ ball (k i) (δ / 2), dist (g i x') (g i (k i)) ≤ ε / 3 + ε / 3, + { intros x' hx', + refine (dist_triangle_right _ _ _).trans (add_le_add (hgδ hpi _).le hgi.le), + exact ((dist_triangle _ _ _).trans_lt (add_lt_add hx'.out hki)).trans_eq (add_halves δ) }, + have := dist_convolution_le (add_pos h2ε h2ε).le hφi hnφi hiφi hmgi h1, + refine ((dist_triangle _ _ _).trans_lt (add_lt_add_of_le_of_lt this hgi)).trans_eq _, + field_simp, ring_nf end end normed_add_comm_group @@ -727,31 +746,30 @@ lemma dist_normed_convolution_le {x₀ : G} {ε : ℝ} dist_convolution_le (by simp_rw [← dist_self (g x₀), hg x₀ (mem_ball_self φ.R_pos)]) φ.support_normed_eq.subset φ.nonneg_normed φ.integral_normed hmg hg -/-- If `φ i` is a sequence of normed bump function, `(φ i ⋆ g) x₀` tends to `g x₀` if `(φ i).R` -tends to `0` and `g` is continuous at `x₀`. -/ -lemma convolution_tendsto_right' {ι} {φ : ι → cont_diff_bump_of_inner (0 : G)} - {l : filter ι} (hφ : tendsto (λ i, (φ i).R) l (𝓝 0)) - (hmg : ae_strongly_measurable g μ) {x₀ : G} (hcg : continuous_at g x₀) : - tendsto (λ i, ((λ x, (φ i).normed μ x) ⋆[lsmul ℝ ℝ, μ] g : G → E') x₀) l (𝓝 (g x₀)) := -begin - refine convolution_tendsto_right (λ i, (φ i).nonneg_normed) (λ i, (φ i).integral_normed) - _ hmg hcg, - rw [normed_add_comm_group.tendsto_nhds_zero] at hφ, - rw [tendsto_small_sets_iff], - intros t ht, - rcases metric.mem_nhds_iff.mp ht with ⟨ε, hε, ht⟩, - refine (hφ ε hε).mono (λ i hi, subset_trans _ ht), - simp_rw [(φ i).support_normed_eq], - rw [real.norm_eq_abs, abs_eq_self.mpr (φ i).R_pos.le] at hi, - exact ball_subset_ball hi.le -end - -/-- Special case of `cont_diff_bump_of_inner.convolution_tendsto_right'` where `g` is continuous. -/ +/-- `(φ i ⋆ g i) (k i)` tends to `z₀` as `i` tends to some filter `l` if +* `φ` is a sequence of normed bump functions such that `(φ i).R` tends to `0` as `i` tends to `l`; +* `g i` is `mu`-a.e. strongly measurable as `i` tends to `l`; +* `g i x` tends to `z₀` as `(i, x)` tends to `l ×ᶠ 𝓝 x₀`; +* `k i` tends to `x₀`. -/ lemma convolution_tendsto_right {ι} {φ : ι → cont_diff_bump_of_inner (0 : G)} + {g : ι → G → E'} {k : ι → G} {x₀ : G} {z₀ : E'} {l : filter ι} + (hφ : tendsto (λ i, (φ i).R) l (𝓝 0)) + (hig : ∀ᶠ i in l, ae_strongly_measurable (g i) μ) + (hcg : tendsto (uncurry g) (l ×ᶠ 𝓝 x₀) (𝓝 z₀)) + (hk : tendsto k l (𝓝 x₀)) : + tendsto (λ i, ((λ x, (φ i).normed μ x) ⋆[lsmul ℝ ℝ, μ] g i : G → E') (k i)) l (𝓝 z₀) := +convolution_tendsto_right (eventually_of_forall $ λ i, (φ i).nonneg_normed) + (eventually_of_forall $ λ i, (φ i).integral_normed) + (tendsto_support_normed_small_sets hφ) hig hcg hk + +/-- Special case of `cont_diff_bump_of_inner.convolution_tendsto_right` where `g` is continuous, + and the limit is taken only in the first function. -/ +lemma convolution_tendsto_right_of_continuous {ι} {φ : ι → cont_diff_bump_of_inner (0 : G)} {l : filter ι} (hφ : tendsto (λ i, (φ i).R) l (𝓝 0)) (hg : continuous g) (x₀ : G) : tendsto (λ i, ((λ x, (φ i).normed μ x) ⋆[lsmul ℝ ℝ, μ] g : G → E') x₀) l (𝓝 (g x₀)) := -convolution_tendsto_right' hφ hg.ae_strongly_measurable hg.continuous_at +convolution_tendsto_right hφ (eventually_of_forall $ λ _, hg.ae_strongly_measurable) + ((hg.tendsto x₀).comp tendsto_snd) tendsto_const_nhds end cont_diff_bump_of_inner diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index d480956f88b2a..ba64895084906 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -54,7 +54,7 @@ open set filter topological_space bornology open_locale uniformity topological_space big_operators filter nnreal ennreal universes u v w -variables {α : Type u} {β : Type v} {X : Type*} +variables {α : Type u} {β : Type v} {X ι : Type*} /-- Construct a uniform structure core from a distance function and metric space axioms. This is a technical construction that can be immediately used to construct a uniform structure @@ -867,6 +867,30 @@ lemma eventually_nhds_iff_ball {p : α → Prop} : (∀ᶠ y in 𝓝 x, p y) ↔ ∃ ε>0, ∀ y ∈ ball x ε, p y := mem_nhds_iff +/-- A version of `filter.eventually_prod_iff` where the second filter consists of neighborhoods +in a pseudo-metric space.-/ +lemma eventually_prod_nhds_iff {f : filter ι} {x₀ : α} {p : ι × α → Prop}: + (∀ᶠ x in f ×ᶠ 𝓝 x₀, p x) ↔ ∃ (pa : ι → Prop) (ha : ∀ᶠ i in f, pa i) (ε > 0), + ∀ {i}, pa i → ∀ {x}, dist x x₀ < ε → p (i, x) := +begin + simp_rw [eventually_prod_iff, metric.eventually_nhds_iff], + refine exists_congr (λ q, exists_congr $ λ hq, _), + split, + { rintro ⟨r, ⟨ε, hε, hεr⟩, hp⟩, exact ⟨ε, hε, λ i hi x hx, hp hi $ hεr hx⟩ }, + { rintro ⟨ε, hε, hp⟩, exact ⟨λ x, dist x x₀ < ε, ⟨ε, hε, λ y, id⟩, @hp⟩ } +end + +/-- A version of `filter.eventually_prod_iff` where the first filter consists of neighborhoods +in a pseudo-metric space.-/ +lemma eventually_nhds_prod_iff {ι α} [pseudo_metric_space α] {f : filter ι} {x₀ : α} + {p : α × ι → Prop}: + (∀ᶠ x in 𝓝 x₀ ×ᶠ f, p x) ↔ ∃ (ε > (0 : ℝ)) (pa : ι → Prop) (ha : ∀ᶠ i in f, pa i) , + ∀ {x}, dist x x₀ < ε → ∀ {i}, pa i → p (x, i) := +begin + rw [eventually_swap_iff, metric.eventually_prod_nhds_iff], + split; { rintro ⟨a1, a2, a3, a4, a5⟩, refine ⟨a3, a4, a1, a2, λ b1 b2 b3 b4, a5 b4 b2⟩ } +end + theorem nhds_basis_closed_ball : (𝓝 x).has_basis (λ ε:ℝ, 0 < ε) (closed_ball x) := nhds_basis_uniformity uniformity_basis_dist_le From c2d85cecf6a74eb88462df1270c02f5589fe53c2 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 1 Oct 2022 15:41:03 +0000 Subject: [PATCH 495/828] =?UTF-8?q?feat(data/list/basic):=20`=CE=BB=20a,?= =?UTF-8?q?=20[a]`=20is=20injective=20(#16716)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also add a `@[simp]` lemma `list.cons_eq_cons` --- src/data/list/basic.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 427146bdeb6fa..45d696485b69b 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -53,6 +53,13 @@ assume l₁ l₂, assume Pe, tail_eq_of_cons_eq Pe theorem cons_inj (a : α) {l l' : list α} : a::l = a::l' ↔ l = l' := cons_injective.eq_iff +theorem cons_eq_cons {a b : α} {l l' : list α} : a::l = b::l' ↔ a = b ∧ l = l' := +⟨list.cons.inj, λ h, h.1 ▸ h.2 ▸ rfl⟩ + +lemma singleton_injective : injective (λ a : α, [a]) := λ a b h, (cons_eq_cons.1 h).1 + +lemma singleton_inj {a b : α} : [a] = [b] ↔ a = b := singleton_injective.eq_iff + theorem exists_cons_of_ne_nil {l : list α} (h : l ≠ nil) : ∃ b L, l = b :: L := by { induction l with c l', contradiction, use [c,l'], } From 7f034c5203ebb12d7ea2a6319e6cee3bb8742841 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 1 Oct 2022 15:41:04 +0000 Subject: [PATCH 496/828] feat(topology/algebra/module/character_space): kernels of terms of the `character_space` (#16722) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This shows that the kernel of a element `φ` of `character_space 𝕜 A` is a maximal ideal. Moreover, `φ` and `ψ` are equal if their kernels coincide. In addition, we provide a missing `ext` lemma, and protect a lemma named `is_closed` in this namespace. --- src/analysis/normed_space/algebra.lean | 2 +- .../algebra/module/character_space.lean | 24 ++++++++++++++++++- 2 files changed, 24 insertions(+), 2 deletions(-) diff --git a/src/analysis/normed_space/algebra.lean b/src/analysis/normed_space/algebra.lean index dde5957d77518..c76180d6d6f24 100644 --- a/src/analysis/normed_space/algebra.lean +++ b/src/analysis/normed_space/algebra.lean @@ -51,7 +51,7 @@ begin { intros φ hφ, rw [set.mem_preimage, mem_closed_ball_zero_iff], exact (le_of_eq $ norm_one ⟨φ, ⟨hφ.1, hφ.2⟩⟩ : _), }, - exact compact_of_is_closed_subset (is_compact_closed_ball 𝕜 0 1) is_closed h, + exact compact_of_is_closed_subset (is_compact_closed_ball 𝕜 0 1) character_space.is_closed h, end end character_space diff --git a/src/topology/algebra/module/character_space.lean b/src/topology/algebra/module/character_space.lean index df3da5dab22e6..46744ed368b82 100644 --- a/src/topology/algebra/module/character_space.lean +++ b/src/topology/algebra/module/character_space.lean @@ -62,6 +62,8 @@ instance : continuous_linear_map_class (character_space 𝕜 A) 𝕜 A 𝕜 := map_add := λ φ, (φ : weak_dual 𝕜 A).map_add, map_continuous := λ φ, (φ : weak_dual 𝕜 A).cont } +@[ext] lemma ext {φ ψ : character_space 𝕜 A} (h : ∀ x, φ x = ψ x) : φ = ψ := fun_like.ext _ _ h + /-- An element of the character space, as a continuous linear map. -/ def to_clm (φ : character_space 𝕜 A) : A →L[𝕜] 𝕜 := (φ : weak_dual 𝕜 A) @@ -145,7 +147,7 @@ end /-- under suitable mild assumptions on `𝕜`, the character space is a closed set in `weak_dual 𝕜 A`. -/ -lemma is_closed [nontrivial 𝕜] [t2_space 𝕜] [has_continuous_mul 𝕜] : +protected lemma is_closed [nontrivial 𝕜] [t2_space 𝕜] [has_continuous_mul 𝕜] : is_closed (character_space 𝕜 A) := begin rw [eq_set_map_one_map_mul, set.set_of_and], @@ -163,10 +165,30 @@ variables [comm_ring 𝕜] [no_zero_divisors 𝕜] [topological_space 𝕜] [has lemma apply_mem_spectrum [nontrivial 𝕜] (φ : character_space 𝕜 A) (a : A) : φ a ∈ spectrum 𝕜 a := alg_hom.apply_mem_spectrum φ a +lemma ext_ker {φ ψ : character_space 𝕜 A} (h : ring_hom.ker φ = ring_hom.ker ψ) : φ = ψ := +begin + ext, + have : x - algebra_map 𝕜 A (ψ x) ∈ ring_hom.ker φ, + { simpa only [h, ring_hom.mem_ker, map_sub, alg_hom_class.commutes] using sub_self (ψ x) }, + { rwa [ring_hom.mem_ker, map_sub, alg_hom_class.commutes, sub_eq_zero] at this, } +end + end ring end character_space +section kernel + +variables [field 𝕜] [topological_space 𝕜] [has_continuous_add 𝕜] [has_continuous_const_smul 𝕜 𝕜] +variables [ring A] [topological_space A] [algebra 𝕜 A] + +/-- The `ring_hom.ker` of `φ : character_space 𝕜 A` is maximal. -/ +instance ker_is_maximal (φ : character_space 𝕜 A) : (ring_hom.ker φ).is_maximal := +ring_hom.ker_is_maximal_of_surjective φ $ λ z, ⟨algebra_map 𝕜 A z, + by simp only [alg_hom_class.commutes, algebra.id.map_eq_id, ring_hom.id_apply]⟩ + +end kernel + section gelfand_transform open continuous_map From ba197e48c03a7e31d4d11576ffae8cefa17a5377 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 1 Oct 2022 18:32:34 +0000 Subject: [PATCH 497/828] feat(order/filter/*,topology/noetherian_space): add lemmas about `cofinite` (#16498) * A filter is disjoint with `cofinite` iff there exists a finite set that belongs to this filter. * An ultrafilter is either less than or equal to `cofinite`, or is equal to `pure a` for some `a`. * Any type with cofinite topology is a Noetherian (hence, is a compact) space. --- src/order/filter/cofinite.lean | 17 ++++++++++----- src/order/filter/ultrafilter.lean | 33 ++++++++++++++++++++---------- src/topology/noetherian_space.lean | 11 ++++++++++ 3 files changed, 45 insertions(+), 16 deletions(-) diff --git a/src/order/filter/cofinite.lean b/src/order/filter/cofinite.lean index 6706656e87e21..a2f292575c2f5 100644 --- a/src/order/filter/cofinite.lean +++ b/src/order/filter/cofinite.lean @@ -23,7 +23,7 @@ Define filters for other cardinalities of the complement. open set function open_locale classical -variables {ι α β : Type*} +variables {ι α β : Type*} {l : filter α} namespace filter @@ -69,16 +69,14 @@ frequently_cofinite_iff_infinite.symm lemma eventually_cofinite_ne (x : α) : ∀ᶠ a in cofinite, a ≠ x := (set.finite_singleton x).eventually_cofinite_nmem -lemma le_cofinite_iff_compl_singleton_mem {l : filter α} : - l ≤ cofinite ↔ ∀ x, {x}ᶜ ∈ l := +lemma le_cofinite_iff_compl_singleton_mem : l ≤ cofinite ↔ ∀ x, {x}ᶜ ∈ l := begin refine ⟨λ h x, h (finite_singleton x).compl_mem_cofinite, λ h s (hs : sᶜ.finite), _⟩, rw [← compl_compl s, ← bUnion_of_singleton sᶜ, compl_Union₂,filter.bInter_mem hs], exact λ x _, h x end -lemma le_cofinite_iff_eventually_ne {l : filter α} : - l ≤ cofinite ↔ ∀ x, ∀ᶠ y in l, y ≠ x := +lemma le_cofinite_iff_eventually_ne : l ≤ cofinite ↔ ∀ x, ∀ᶠ y in l, y ≠ x := le_cofinite_iff_compl_singleton_mem /-- If `α` is a preorder with no maximal element, then `at_top ≤ cofinite`. -/ @@ -100,6 +98,15 @@ lemma Coprod_cofinite {α : ι → Type*} [finite ι] : filter.coext $ λ s, by simp only [compl_mem_Coprod, mem_cofinite, compl_compl, forall_finite_image_eval_iff] +@[simp] lemma disjoint_cofinite_left : disjoint cofinite l ↔ ∃ s ∈ l, set.finite s := +begin + simp only [has_basis_cofinite.disjoint_iff l.basis_sets, id, disjoint_compl_left_iff_subset], + exact ⟨λ ⟨s, hs, t, ht, hts⟩, ⟨t, ht, hs.subset hts⟩, λ ⟨s, hs, hsf⟩, ⟨s, hsf, s, hs, subset.rfl⟩⟩ +end + +@[simp] lemma disjoint_cofinite_right : disjoint l cofinite ↔ ∃ s ∈ l, set.finite s := +disjoint.comm.trans disjoint_cofinite_left + end filter open filter diff --git a/src/order/filter/ultrafilter.lean b/src/order/filter/ultrafilter.lean index b2f2b93c935aa..05c0fef9ebd00 100644 --- a/src/order/filter/ultrafilter.lean +++ b/src/order/filter/ultrafilter.lean @@ -67,6 +67,12 @@ le_of_inf_eq (f.unique inf_le_left hg) lemma le_of_inf_ne_bot' (f : ultrafilter α) {g : filter α} (hg : ne_bot (g ⊓ f)) : ↑f ≤ g := f.le_of_inf_ne_bot $ by rwa inf_comm +lemma inf_ne_bot_iff {f : ultrafilter α} {g : filter α} : ne_bot (↑f ⊓ g) ↔ ↑f ≤ g := +⟨le_of_inf_ne_bot f, λ h, (inf_of_le_left h).symm ▸ f.ne_bot⟩ + +lemma disjoint_iff_not_le {f : ultrafilter α} {g : filter α} : disjoint ↑f g ↔ ¬↑f ≤ g := +by rw [← inf_ne_bot_iff, ne_bot_iff, ne.def, not_not, disjoint_iff] + @[simp] lemma compl_not_mem_iff : sᶜ ∉ f ↔ s ∈ f := ⟨λ hsc, le_principal_iff.1 $ f.le_of_inf_ne_bot ⟨λ h, hsc $ mem_of_eq_bot$ by rwa compl_compl⟩, compl_not_mem⟩ @@ -92,6 +98,12 @@ lemma nonempty_of_mem (hs : s ∈ f) : s.nonempty := nonempty_of_mem hs lemma ne_empty_of_mem (hs : s ∈ f) : s ≠ ∅ := (nonempty_of_mem hs).ne_empty @[simp] lemma empty_not_mem : ∅ ∉ f := empty_not_mem f +@[simp] lemma le_sup_iff {u : ultrafilter α} {f g : filter α} : ↑u ≤ f ⊔ g ↔ ↑u ≤ f ∨ ↑u ≤ g := +not_iff_not.1 $ by simp only [← disjoint_iff_not_le, not_or_distrib, disjoint_sup_right] + +@[simp] lemma union_mem_iff : s ∪ t ∈ f ↔ s ∈ f ∨ t ∈ f := +by simp only [← mem_coe, ← le_principal_iff, ← sup_principal, le_sup_iff] + lemma mem_or_compl_mem (f : ultrafilter α) (s : set α) : s ∈ f ∨ sᶜ ∈ f := or_iff_not_imp_left.2 compl_mem_iff_not_mem.2 @@ -100,10 +112,7 @@ protected lemma em (f : ultrafilter α) (p : α → Prop) : f.mem_or_compl_mem {x | p x} lemma eventually_or : (∀ᶠ x in f, p x ∨ q x) ↔ (∀ᶠ x in f, p x) ∨ ∀ᶠ x in f, q x := -⟨λ H, (f.em p).imp_right $ λ hp, (H.and hp).mono $ λ x ⟨hx, hnx⟩, hx.resolve_left hnx, - λ H, H.elim (λ hp, hp.mono $ λ x, or.inl) (λ hp, hp.mono $ λ x, or.inr)⟩ - -lemma union_mem_iff : s ∪ t ∈ f ↔ s ∈ f ∨ t ∈ f := eventually_or +union_mem_iff lemma eventually_not : (∀ᶠ x in f, ¬p x) ↔ ¬∀ᶠ x in f, p x := compl_mem_iff_not_mem @@ -182,20 +191,22 @@ lemma pure_injective : injective (pure : α → ultrafilter α) := instance [inhabited α] : inhabited (ultrafilter α) := ⟨pure default⟩ instance [nonempty α] : nonempty (ultrafilter α) := nonempty.map pure infer_instance -lemma eq_pure_of_finite_mem (h : s.finite) (h' : s ∈ f) : ∃ x ∈ s, (f : filter α) = pure x := +lemma eq_pure_of_finite_mem (h : s.finite) (h' : s ∈ f) : ∃ x ∈ s, f = pure x := begin rw ← bUnion_of_singleton s at h', rcases (ultrafilter.finite_bUnion_mem_iff h).mp h' with ⟨a, has, haf⟩, - use [a, has], - change (f : filter α) = (pure a : ultrafilter α), - rw [ultrafilter.coe_inj, ← ultrafilter.coe_le_coe], - change (f : filter α) ≤ pure a, - rwa [← principal_singleton, le_principal_iff] + exact ⟨a, has, eq_of_le (filter.le_pure_iff.2 haf)⟩ end -lemma eq_pure_of_finite [finite α] (f : ultrafilter α) : ∃ a, (f : filter α) = pure a := +lemma eq_pure_of_finite [finite α] (f : ultrafilter α) : ∃ a, f = pure a := (eq_pure_of_finite_mem finite_univ univ_mem).imp $ λ a ⟨_, ha⟩, ha +lemma le_cofinite_or_eq_pure (f : ultrafilter α) : (f : filter α) ≤ cofinite ∨ ∃ a, f = pure a := +or_iff_not_imp_left.2 $ λ h, + let ⟨s, hs, hfin⟩ := filter.disjoint_cofinite_right.1 (disjoint_iff_not_le.2 h), + ⟨a, has, hf⟩ := eq_pure_of_finite_mem hfin hs + in ⟨a, hf⟩ + /-- Monadic bind for ultrafilters, coming from the one on filters defined in terms of map and join.-/ def bind (f : ultrafilter α) (m : α → ultrafilter β) : ultrafilter β := diff --git a/src/topology/noetherian_space.lean b/src/topology/noetherian_space.lean index 7806a0c6a7bb7..1cf5f3f65e33c 100644 --- a/src/topology/noetherian_space.lean +++ b/src/topology/noetherian_space.lean @@ -101,6 +101,17 @@ end variables {α β} +instance {α} : noetherian_space (cofinite_topology α) := +begin + simp only [noetherian_space_iff_opens, is_compact_iff_ultrafilter_le_nhds, + cofinite_topology.nhds_eq, ultrafilter.le_sup_iff], + intros s f hs, + rcases f.le_cofinite_or_eq_pure with hf|⟨a, rfl⟩, + { rcases filter.nonempty_of_mem (filter.le_principal_iff.1 hs) with ⟨a, ha⟩, + exact ⟨a, ha, or.inr hf⟩ }, + { exact ⟨a, filter.le_principal_iff.mp hs, or.inl le_rfl⟩ } +end + lemma noetherian_space.is_compact [h : noetherian_space α] (s : set α) : is_compact s := let H := (noetherian_space_tfae α).out 0 2 in H.mp h s From 567220c269ba1ca5895871a9356fb1239858d3b9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 1 Oct 2022 18:32:35 +0000 Subject: [PATCH 498/828] refactor(analysis/inner_product_space/pi_L2): Delete `orthonormal_basis_index` (#16698) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `std_orthonormal_basis` was indexed by `orthonormal_basis_index` while it can simply be indexed by `fin (finrank 𝕜 E)`. As a result, `fin_std_orthonormal_basis` becomes a trivial reindexation of `std_orthonormal_basis` so we delete it. --- .../inner_product_space/orientation.lean | 4 +-- src/analysis/inner_product_space/pi_L2.lean | 33 +++++++------------ 2 files changed, 13 insertions(+), 24 deletions(-) diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index cdc5569c7e729..f83291043780d 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -70,7 +70,7 @@ protected def orientation.fin_orthonormal_basis {n : ℕ} (hn : 0 < n) (h : finr begin haveI := fin.pos_iff_nonempty.1 hn, haveI := finite_dimensional_of_finrank (h.symm ▸ hn : 0 < finrank ℝ E), - exact (fin_std_orthonormal_basis h).adjust_to_orientation x + exact ((std_orthonormal_basis _ _).reindex $ fin_congr h).adjust_to_orientation x end /-- `orientation.fin_orthonormal_basis` gives a basis with the required orientation. -/ @@ -80,5 +80,5 @@ end begin haveI := fin.pos_iff_nonempty.1 hn, haveI := finite_dimensional_of_finrank (h.symm ▸ hn : 0 < finrank ℝ E), - exact (fin_std_orthonormal_basis h).orientation_adjust_to_orientation x + exact ((std_orthonormal_basis _ _).reindex $ fin_congr h).orientation_adjust_to_orientation x end diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 5cd0a253847e5..2ddd494dd1e6c 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -632,25 +632,16 @@ lemma _root_.exists_orthonormal_basis : let ⟨w, hw, hw', hw''⟩ := (orthonormal_empty 𝕜 E).exists_orthonormal_basis_extension in ⟨w, hw, hw''⟩ -/-- Index for an arbitrary orthonormal basis on a finite-dimensional `inner_product_space`. -/ -def orthonormal_basis_index : finset E := -classical.some (exists_orthonormal_basis 𝕜 E) - /-- A finite-dimensional `inner_product_space` has an orthonormal basis. -/ -def std_orthonormal_basis : orthonormal_basis (orthonormal_basis_index 𝕜 E) 𝕜 E := -classical.some (classical.some_spec (exists_orthonormal_basis 𝕜 E)) - -@[simp] lemma coe_std_orthonormal_basis : ⇑(std_orthonormal_basis 𝕜 E) = coe := -classical.some_spec (classical.some_spec (exists_orthonormal_basis 𝕜 E)) +def std_orthonormal_basis : orthonormal_basis (fin (finrank 𝕜 E)) 𝕜 E := +begin + let b := classical.some (classical.some_spec $ exists_orthonormal_basis 𝕜 E), + rw [finrank_eq_card_basis b.to_basis], + exact b.reindex (fintype.equiv_fin_of_card_eq rfl), +end variables {𝕜 E} -/-- An `n`-dimensional `inner_product_space` has an orthonormal basis indexed by `fin n`. -/ -def fin_std_orthonormal_basis {n : ℕ} (hn : finrank 𝕜 E = n) : orthonormal_basis (fin n) 𝕜 E := -have h : fintype.card (orthonormal_basis_index 𝕜 E) = n, -by rw [← finrank_eq_card_basis (std_orthonormal_basis 𝕜 E).to_basis, hn], -(std_orthonormal_basis 𝕜 E).reindex (fintype.equiv_fin_of_card_eq h) - section subordinate_orthonormal_basis open direct_sum variables {n : ℕ} (hn : finrank 𝕜 E = n) [decidable_eq ι] @@ -660,7 +651,7 @@ variables {n : ℕ} (hn : finrank 𝕜 E = n) [decidable_eq ι] inner product space `E`. This should not be accessed directly, but only via the subsequent API. -/ @[irreducible] def direct_sum.is_internal.sigma_orthonormal_basis_index_equiv (hV' : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : - (Σ i, orthonormal_basis_index 𝕜 (V i)) ≃ fin n := + (Σ i, fin (finrank 𝕜 (V i))) ≃ fin n := let b := hV.collected_orthonormal_basis hV' (λ i, (std_orthonormal_basis 𝕜 (V i))) in fintype.equiv_fin_of_card_eq $ (finite_dimensional.finrank_eq_card_basis b.to_basis).symm.trans hn @@ -704,7 +695,7 @@ space, there exists an isometry from the orthogonal complement of a nonzero sing def orthonormal_basis.from_orthogonal_span_singleton (n : ℕ) [fact (finrank 𝕜 E = n + 1)] {v : E} (hv : v ≠ 0) : orthonormal_basis (fin n) 𝕜 (𝕜 ∙ v)ᗮ := -(fin_std_orthonormal_basis (finrank_orthogonal_span_singleton hv)) +(std_orthonormal_basis _ _).reindex $ fin_congr $ finrank_orthogonal_span_singleton hv section linear_isometry @@ -731,11 +722,9 @@ begin ... = finrank 𝕜 V - finrank 𝕜 S : by simp only [linear_map.finrank_range_of_inj L.injective] ... = finrank 𝕜 Sᗮ : by simp only - [← S.finrank_add_finrank_orthogonal, add_tsub_cancel_left] - ... = d : dim_S_perp, - let BS := (fin_std_orthonormal_basis dim_S_perp), - let BLS := (fin_std_orthonormal_basis dim_LS_perp), - exact BS.repr.trans BLS.repr.symm }, + [← S.finrank_add_finrank_orthogonal, add_tsub_cancel_left], + exact (std_orthonormal_basis 𝕜 Sᗮ).repr.trans + ((std_orthonormal_basis 𝕜 LSᗮ).reindex $ fin_congr dim_LS_perp).repr.symm }, let L3 := (LS)ᗮ.subtypeₗᵢ.comp E.to_linear_isometry, -- Project onto S and Sᗮ haveI : complete_space S := finite_dimensional.complete 𝕜 S, From d638c7f7b812495fb60c946b20158d5b438b889a Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sat, 1 Oct 2022 18:32:37 +0000 Subject: [PATCH 499/828] =?UTF-8?q?feat(group=5Ftheory/subgroup/basic):=20?= =?UTF-8?q?If=20`H`=20is=20commutative,=20then=20`H=20=E2=89=A4=20H.centra?= =?UTF-8?q?lizer`=20(#16718)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `H` is commutative, then `H ≤ H.centralizer`. --- src/group_theory/subgroup/basic.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 721e922aa759a..d7af47900ef5f 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1859,6 +1859,12 @@ end⟩⟩ (H.subgroup_of K).is_commutative := H.comap_injective_is_commutative subtype.coe_injective +@[to_additive] lemma le_centralizer_iff_is_commutative : K ≤ K.centralizer ↔ K.is_commutative := +⟨λ h, ⟨⟨λ x y, subtype.ext (h y.2 x x.2)⟩⟩, λ h x hx y hy, congr_arg coe (h.1.1 ⟨y, hy⟩ ⟨x, hx⟩)⟩ + +@[to_additive] lemma le_centralizer [h : H.is_commutative] : H ≤ H.centralizer := +le_centralizer_iff_is_commutative.mpr h + end subgroup namespace group From 77db6f75252d3934c9f68800a8636d23ea5c2473 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Sat, 1 Oct 2022 21:30:39 +0000 Subject: [PATCH 500/828] feat(analysis/locally_convex/with_seminorms): split `continuous_from_bounded` to get an extra continuity criterion (#16710) This split should also help with proving that `continuous_from_bounded` is an `iff` under some assumptions, because the only remaining fact to prove is that a continuous seminorm is necessarily greater than a constant times the supremum of a finite number of generating seminorms. --- .../locally_convex/with_seminorms.lean | 56 ++++++++++++++----- 1 file changed, 41 insertions(+), 15 deletions(-) diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 709452696abbe..8f2a3de99edb4 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -302,6 +302,15 @@ begin exact add_group_filter_basis.nhds_zero_eq _, end +lemma with_seminorms.continuous_seminorm [module ℝ E] [normed_algebra ℝ 𝕜] [is_scalar_tower ℝ 𝕜 E] + [has_continuous_const_smul ℝ E] {p : seminorm_family 𝕜 E ι} (hp : with_seminorms p) + (i : ι) : continuous (p i) := +begin + refine seminorm.continuous _, + rw [p.with_seminorms_iff_nhds_eq_infi.mp hp, ball_zero_eq_preimage_ball], + exact filter.mem_infi_of_mem i (filter.preimage_mem_comap $ metric.ball_mem_nhds _ one_pos) +end + end topological_add_group section normed_space @@ -391,25 +400,42 @@ namespace seminorm variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [add_comm_group F] [module 𝕜 F] variables [nonempty ι] [nonempty ι'] +lemma continuous_of_continuous_comp {q : seminorm_family 𝕜 F ι'} + [topological_space E] [topological_add_group E] + [topological_space F] [topological_add_group F] (hq : with_seminorms q) + (f : E →ₗ[𝕜] F) (hf : ∀ i, continuous ((q i).comp f)) : continuous f := +begin + refine continuous_of_continuous_at_zero f _, + simp_rw [continuous_at, f.map_zero, q.with_seminorms_iff_nhds_eq_infi.mp hq, filter.tendsto_infi, + filter.tendsto_comap_iff], + intros i, + convert (hf i).continuous_at, + exact (map_zero _).symm +end + +lemma continuous_iff_continuous_comp [normed_algebra ℝ 𝕜] [module ℝ F] [is_scalar_tower ℝ 𝕜 F] + {q : seminorm_family 𝕜 F ι'} [topological_space E] [topological_add_group E] + [topological_space F] [topological_add_group F] [has_continuous_const_smul ℝ F] + (hq : with_seminorms q) (f : E →ₗ[𝕜] F) : + continuous f ↔ ∀ i, continuous ((q i).comp f) := +⟨λ h i, continuous.comp (hq.continuous_seminorm i) h, continuous_of_continuous_comp hq f⟩ + lemma continuous_from_bounded {p : seminorm_family 𝕜 E ι} {q : seminorm_family 𝕜 F ι'} - [uniform_space E] [uniform_add_group E] (hp : with_seminorms p) - [uniform_space F] [uniform_add_group F] (hq : with_seminorms q) + [topological_space E] [topological_add_group E] (hp : with_seminorms p) + [topological_space F] [topological_add_group F] (hq : with_seminorms q) (f : E →ₗ[𝕜] F) (hf : seminorm.is_bounded p q f) : continuous f := begin - refine continuous_of_continuous_at_zero f _, - rw [continuous_at_def, f.map_zero, hp.1], - intros U hU, - rw [hq.1, add_group_filter_basis.nhds_zero_eq, filter_basis.mem_filter_iff] at hU, - rcases hU with ⟨V, hV : V ∈ q.basis_sets, hU⟩, - rcases q.basis_sets_iff.mp hV with ⟨s₂, r, hr, hV⟩, - rw hV at hU, - rw [p.add_group_filter_basis.nhds_zero_eq, filter_basis.mem_filter_iff], - rcases (seminorm.is_bounded_sup hf s₂) with ⟨C, s₁, hC, hf⟩, - refine ⟨(s₁.sup p).ball 0 (r/C), p.basis_sets_mem _ (div_pos hr (nnreal.coe_pos.mpr hC)), _⟩, - refine subset.trans _ (preimage_mono hU), - simp_rw [←linear_map.map_zero f, ←ball_comp], + refine continuous_of_continuous_comp hq _ (λ i, seminorm.continuous_of_continuous_at_zero _), + rw [metric.continuous_at_iff', map_zero], + intros r hr, + rcases hf i with ⟨s₁, C, hC, hf⟩, + have hC' : 0 < C := hC.bot_lt, + rw hp.has_basis.eventually_iff, + refine ⟨(s₁.sup p).ball 0 (r/C), p.basis_sets_mem _ (by positivity), _⟩, + simp_rw [ ←metric.mem_ball, ←mem_preimage, ←ball_zero_eq_preimage_ball], refine subset.trans _ (ball_antitone hf), - rw ball_smul (s₁.sup p) hC, + rw ball_smul (s₁.sup p) hC', + refl end lemma cont_with_seminorms_normed_space (F) [seminormed_add_comm_group F] [normed_space 𝕜 F] From 852b23b5703f4c5bbc4eb500eb54d613636b5fca Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 1 Oct 2022 23:57:29 +0000 Subject: [PATCH 501/828] chore(topology,analysis): make the sup metric and norm computable (#16570) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit To make this work, we replace `max` with `sup` in the implementation (the two are defeq, but the former has stronger typeclass arguments), and add some missing shortcut instances for `ennreal` which are needed for the emetric structure. As a result, we can can now perform useless "computation"s like: ```lean #eval ∥((3 : ℤ), (4 : ℤ))∥ -- real.of_cauchy (sorry /- 4, 4, 4, 4, 4, 4, 4, 4, 4, 4, ... -/) ``` (this isn't totally useless; it's enough to inform the reader that they're not looking at a euclidean norm, which is a common source of confusion) --- src/analysis/normed/field/basic.lean | 17 ++++++++--------- src/analysis/normed/group/basic.lean | 11 +++++------ src/analysis/normed_space/star/basic.lean | 2 +- src/data/real/ennreal.lean | 1 + src/topology/metric_space/basic.lean | 15 ++++++++------- src/topology/metric_space/emetric_space.lean | 5 ++--- 6 files changed, 25 insertions(+), 26 deletions(-) diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index f3a458a9f9261..8ae17f85352d8 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -184,7 +184,7 @@ instance : non_unital_semi_normed_ring (ulift α) := /-- Non-unital seminormed ring structure on the product of two non-unital seminormed rings, using the sup norm. -/ -noncomputable instance prod.non_unital_semi_normed_ring [non_unital_semi_normed_ring β] : +instance prod.non_unital_semi_normed_ring [non_unital_semi_normed_ring β] : non_unital_semi_normed_ring (α × β) := { norm_mul := assume x y, calc @@ -201,7 +201,7 @@ noncomputable instance prod.non_unital_semi_normed_ring [non_unital_semi_normed_ /-- Non-unital seminormed ring structure on the product of finitely many non-unital seminormed rings, using the sup norm. -/ -noncomputable instance pi.non_unital_semi_normed_ring {π : ι → Type*} [fintype ι] +instance pi.non_unital_semi_normed_ring {π : ι → Type*} [fintype ι] [Π i, non_unital_semi_normed_ring (π i)] : non_unital_semi_normed_ring (Π i, π i) := { norm_mul := λ x y, nnreal.coe_mono $ @@ -314,14 +314,14 @@ instance : semi_normed_ring (ulift α) := /-- Seminormed ring structure on the product of two seminormed rings, using the sup norm. -/ -noncomputable instance prod.semi_normed_ring [semi_normed_ring β] : +instance prod.semi_normed_ring [semi_normed_ring β] : semi_normed_ring (α × β) := { ..prod.non_unital_semi_normed_ring, ..prod.seminormed_add_comm_group, } /-- Seminormed ring structure on the product of finitely many seminormed rings, using the sup norm. -/ -noncomputable instance pi.semi_normed_ring {π : ι → Type*} [fintype ι] +instance pi.semi_normed_ring {π : ι → Type*} [fintype ι] [Π i, semi_normed_ring (π i)] : semi_normed_ring (Π i, π i) := { ..pi.non_unital_semi_normed_ring, @@ -338,15 +338,14 @@ instance : non_unital_normed_ring (ulift α) := /-- Non-unital normed ring structure on the product of two non-unital normed rings, using the sup norm. -/ -noncomputable instance prod.non_unital_normed_ring [non_unital_normed_ring β] : +instance prod.non_unital_normed_ring [non_unital_normed_ring β] : non_unital_normed_ring (α × β) := { norm_mul := norm_mul_le, ..prod.seminormed_add_comm_group } /-- Normed ring structure on the product of finitely many non-unital normed rings, using the sup norm. -/ -noncomputable instance pi.non_unital_normed_ring {π : ι → Type*} [fintype ι] - [Π i, non_unital_normed_ring (π i)] : +instance pi.non_unital_normed_ring {π : ι → Type*} [fintype ι] [Π i, non_unital_normed_ring (π i)] : non_unital_normed_ring (Π i, π i) := { norm_mul := norm_mul_le, ..pi.normed_add_comm_group } @@ -368,12 +367,12 @@ instance : normed_ring (ulift α) := .. ulift.normed_add_comm_group } /-- Normed ring structure on the product of two normed rings, using the sup norm. -/ -noncomputable instance prod.normed_ring [normed_ring β] : normed_ring (α × β) := +instance prod.normed_ring [normed_ring β] : normed_ring (α × β) := { norm_mul := norm_mul_le, ..prod.normed_add_comm_group } /-- Normed ring structure on the product of finitely many normed rings, using the sup norm. -/ -noncomputable instance pi.normed_ring {π : ι → Type*} [fintype ι] [Π i, normed_ring (π i)] : +instance pi.normed_ring {π : ι → Type*} [fintype ι] [Π i, normed_ring (π i)] : normed_ring (Π i, π i) := { norm_mul := norm_mul_le, ..pi.normed_add_comm_group } diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 93adf28b3f9c6..305d481a27a85 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -856,8 +856,8 @@ lemma ulift.nnnorm_def (x : ulift E) : ∥x∥₊ = ∥x.down∥₊ := rfl @[simp] lemma ulift.nnnorm_up (x : E) : ∥ulift.up x∥₊ = ∥x∥₊ := rfl /-- seminormed group instance on the product of two seminormed groups, using the sup norm. -/ -noncomputable instance prod.seminormed_add_comm_group : seminormed_add_comm_group (E × F) := -{ norm := λx, max ∥x.1∥ ∥x.2∥, +instance prod.seminormed_add_comm_group : seminormed_add_comm_group (E × F) := +{ norm := λx, ∥x.1∥ ⊔ ∥x.2∥, dist_eq := assume (x y : E × F), show max (dist x.1 y.1) (dist x.2 y.2) = (max ∥(x - y).1∥ ∥(x - y).2∥), by simp [dist_eq_norm] } @@ -881,7 +881,7 @@ variables {π : ι → Type*} [fintype ι] [Π i, seminormed_add_comm_group (π /-- seminormed group instance on the product of finitely many seminormed groups, using the sup norm. -/ -noncomputable instance pi.seminormed_add_comm_group : seminormed_add_comm_group (Π i, π i) := +instance pi.seminormed_add_comm_group : seminormed_add_comm_group (Π i, π i) := { norm := λ f, ↑(finset.univ.sup (λ b, ∥f b∥₊)), dist_eq := assume x y, congr_arg (coe : ℝ≥0 → ℝ) $ congr_arg (finset.sup finset.univ) $ funext $ assume a, @@ -1208,12 +1208,11 @@ instance ulift.normed_add_comm_group : normed_add_comm_group (ulift E) := { ..ulift.seminormed_add_comm_group } /-- normed group instance on the product of two normed groups, using the sup norm. -/ -noncomputable instance prod.normed_add_comm_group : normed_add_comm_group (E × F) := +instance prod.normed_add_comm_group : normed_add_comm_group (E × F) := { ..prod.seminormed_add_comm_group } /-- normed group instance on the product of finitely many normed groups, using the sup norm. -/ -noncomputable instance pi.normed_add_comm_group {π : ι → Type*} [fintype ι] - [Π i, normed_add_comm_group (π i)] : +instance pi.normed_add_comm_group {π : ι → Type*} [fintype ι] [Π i, normed_add_comm_group (π i)] : normed_add_comm_group (Πi, π i) := { ..pi.seminormed_add_comm_group } lemma tendsto_norm_sub_self_punctured_nhds (a : E) : tendsto (λ x, ∥x - a∥) (𝓝[≠] a) (𝓝[>] 0) := diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index 7675e92c88d35..ba11121981588 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -149,7 +149,7 @@ instance _root_.prod.cstar_ring : cstar_ring (R₁ × R₂) := rw [sq_le_sq, abs_of_nonneg (norm_nonneg _)], exact (le_max_left _ _).trans (le_abs_self _), exact (le_max_right _ _).trans (le_abs_self _) }, - { rw le_max_iff, + { rw le_sup_iff, rcases le_total (∥x.fst∥) (∥x.snd∥) with (h | h); simp [h] } end } diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index e0de967a50f89..7e1f7dc313b3d 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -76,6 +76,7 @@ variables {α : Type*} {β : Type*} and is relevant as the codomain of a measure. -/ @[derive [ has_zero, add_comm_monoid_with_one, + semilattice_sup, distrib_lattice, order_bot, bounded_order, canonically_ordered_comm_semiring, complete_linear_order, densely_ordered, nontrivial, canonically_linear_ordered_add_monoid, has_sub, has_ordered_sub, linear_ordered_add_comm_monoid_with_top]] diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index ba64895084906..8c22175a37fc6 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1558,13 +1558,14 @@ end ulift section prod variables [pseudo_metric_space β] -noncomputable instance prod.pseudo_metric_space_max : +instance prod.pseudo_metric_space_max : pseudo_metric_space (α × β) := (pseudo_emetric_space.to_pseudo_metric_space_of_dist - (λ x y : α × β, max (dist x.1 y.1) (dist x.2 y.2)) + (λ x y : α × β, dist x.1 y.1 ⊔ dist x.2 y.2) (λ x y, (max_lt (edist_lt_top _ _) (edist_lt_top _ _)).ne) - (λ x y, by simp only [dist_edist, ← ennreal.to_real_max (edist_ne_top _ _) (edist_ne_top _ _), - prod.edist_eq])).replace_bornology $ + (λ x y, by simp only [sup_eq_max, dist_edist, + ← ennreal.to_real_max (edist_ne_top _ _) (edist_ne_top _ _), prod.edist_eq])) + .replace_bornology $ λ s, by { simp only [← is_bounded_image_fst_and_snd, is_bounded_iff_eventually, ball_image_iff, ← eventually_and, ← forall_and_distrib, ← max_le_iff], refl } @@ -1792,7 +1793,7 @@ open finset variables {π : β → Type*} [fintype β] [∀b, pseudo_metric_space (π b)] /-- A finite product of pseudometric spaces is a pseudometric space, with the sup distance. -/ -noncomputable instance pseudo_metric_space_pi : pseudo_metric_space (Πb, π b) := +instance pseudo_metric_space_pi : pseudo_metric_space (Πb, π b) := begin /- we construct the instance from the pseudoemetric space instance to avoid checking again that the uniformity is the same as the product uniformity, but we register nevertheless a nice formula @@ -2791,7 +2792,7 @@ metric_space.induced ulift.down ulift.down_injective ‹_› section prod -noncomputable instance prod.metric_space_max [metric_space β] : metric_space (γ × β) := +instance prod.metric_space_max [metric_space β] : metric_space (γ × β) := { eq_of_dist_eq_zero := λ x y h, begin cases max_le_iff.1 (le_of_eq h) with h₁ h₂, exact prod.ext_iff.2 ⟨dist_le_zero.1 h₁, dist_le_zero.1 h₂⟩ @@ -2805,7 +2806,7 @@ open finset variables {π : β → Type*} [fintype β] [∀b, metric_space (π b)] /-- A finite product of metric spaces is a metric space, with the sup distance. -/ -noncomputable instance metric_space_pi : metric_space (Πb, π b) := +instance metric_space_pi : metric_space (Πb, π b) := /- we construct the instance from the emetric space instance to avoid checking again that the uniformity is the same as the product uniformity, but we register nevertheless a nice formula for the distance -/ diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index 3a65a7af6808d..3f9fc8b553efa 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -28,7 +28,6 @@ to `emetric_space` at the end. -/ open set filter classical -noncomputable theory open_locale uniformity topological_space big_operators filter nnreal ennreal @@ -427,7 +426,7 @@ end ulift pseudometric spaces. We make sure that the uniform structure thus constructed is the one corresponding to the product of uniform spaces, to avoid diamond problems. -/ instance prod.pseudo_emetric_space_max [pseudo_emetric_space β] : pseudo_emetric_space (α × β) := -{ edist := λ x y, max (edist x.1 y.1) (edist x.2 y.2), +{ edist := λ x y, edist x.1 y.1 ⊔ edist x.2 y.2, edist_self := λ x, by simp, edist_comm := λ x y, by simp [edist_comm], edist_triangle := λ x y z, max_le @@ -746,7 +745,7 @@ end second_countable section diam /-- The diameter of a set in a pseudoemetric space, named `emetric.diam` -/ -def diam (s : set α) := ⨆ (x ∈ s) (y ∈ s), edist x y +noncomputable def diam (s : set α) := ⨆ (x ∈ s) (y ∈ s), edist x y lemma diam_le_iff {d : ℝ≥0∞} : diam s ≤ d ↔ ∀ (x ∈ s) (y ∈ s), edist x y ≤ d := From f063221aa6a669e3c5de030c56f42f5c5cd90c62 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sun, 2 Oct 2022 04:41:35 +0000 Subject: [PATCH 502/828] feat(topology/sheaves/skyscraper): skyscraper presheaf as a functor (#16685) --- src/topology/sheaves/skyscraper.lean | 61 +++++++++++++++++++++++++++- 1 file changed, 59 insertions(+), 2 deletions(-) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index 2ab30db91089d..3b8052b8dec06 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -42,7 +42,7 @@ variables {X : Top.{u}} (p₀ : X) [Π (U : opens X), decidable (p₀ ∈ U)] section -variables {C : Type v} [category.{w} C] (A : C) [has_terminal C] +variables {C : Type v} [category.{w} C] [has_terminal C] (A : C) /-- A skyscraper presheaf is a presheaf supported at a single point: if `p₀ ∈ X` is a specified @@ -69,7 +69,53 @@ lemma skyscraper_presheaf_eq_pushforward skyscraper_presheaf p₀ A = continuous_map.const (Top.of punit) p₀ _* skyscraper_presheaf punit.star A := by convert_to @skyscraper_presheaf X p₀ - (λ U, hd $ (opens.map $ continuous_map.const _ p₀).obj U) C _ A _ = _; congr <|> refl + (λ U, hd $ (opens.map $ continuous_map.const _ p₀).obj U) C _ _ A = _; congr <|> refl + +/-- +Taking skyscraper presheaf at a point is functorial: `c ↦ skyscraper p₀ c` defines a functor by +sending every `f : a ⟶ b` to the natural transformation `α` defined as: `α(U) = f : a ⟶ b` if +`p₀ ∈ U` and the unique morphism to a terminal object in `C` if `p₀ ∉ U`. +-/ +@[simps] def skyscraper_presheaf_functor.map' {a b : C} (f : a ⟶ b) : + skyscraper_presheaf p₀ a ⟶ skyscraper_presheaf p₀ b := +{ app := λ U, if h : p₀ ∈ U.unop + then eq_to_hom (if_pos h) ≫ f ≫ eq_to_hom (if_pos h).symm + else ((if_neg h).symm.rec terminal_is_terminal).from _, + naturality' := λ U V i, + begin + simp only [skyscraper_presheaf_map], by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom i.unop hV, split_ifs, + simpa only [eq_to_hom_trans_assoc, category.assoc, eq_to_hom_trans], }, + { apply ((if_neg hV).symm.rec terminal_is_terminal).hom_ext, }, + end } + +lemma skyscraper_presheaf_functor.map'_id {a : C} : + skyscraper_presheaf_functor.map' p₀ (𝟙 a) = 𝟙 _ := +begin + ext1, ext1, simp only [skyscraper_presheaf_functor.map'_app, nat_trans.id_app], split_ifs, + { simp only [category.id_comp, category.comp_id, eq_to_hom_trans, eq_to_hom_refl], }, + { apply ((if_neg h).symm.rec terminal_is_terminal).hom_ext, }, +end + +lemma skyscraper_presheaf_functor.map'_comp {a b c : C} (f : a ⟶ b) (g : b ⟶ c) : + skyscraper_presheaf_functor.map' p₀ (f ≫ g) = + skyscraper_presheaf_functor.map' p₀ f ≫ skyscraper_presheaf_functor.map' p₀ g := +begin + ext1, ext1, simp only [skyscraper_presheaf_functor.map'_app, nat_trans.comp_app], split_ifs, + { simp only [category.assoc, eq_to_hom_trans_assoc, eq_to_hom_refl, category.id_comp], }, + { apply ((if_neg h).symm.rec terminal_is_terminal).hom_ext, }, +end + +/-- +Taking skyscraper presheaf at a point is functorial: `c ↦ skyscraper p₀ c` defines a functor by +sending every `f : a ⟶ b` to the natural transformation `α` defined as: `α(U) = f : a ⟶ b` if +`p₀ ∈ U` and the unique morphism to a terminal object in `C` if `p₀ ∉ U`. +-/ +@[simps] def skyscraper_presheaf_functor : C ⥤ presheaf C X := +{ obj := skyscraper_presheaf p₀, + map := λ _ _, skyscraper_presheaf_functor.map' p₀, + map_id' := λ _, skyscraper_presheaf_functor.map'_id p₀, + map_comp' := λ _ _ _, skyscraper_presheaf_functor.map'_comp p₀ } end @@ -178,4 +224,15 @@ The skyscraper presheaf supported at `p₀` with value `A` is the sheaf that ass def skyscraper_sheaf [has_products.{u} C] : sheaf C X := ⟨skyscraper_presheaf p₀ A, skyscraper_presheaf_is_sheaf _ _⟩ +/-- +Taking skyscraper sheaf at a point is functorial: `c ↦ skyscraper p₀ c` defines a functor by +sending every `f : a ⟶ b` to the natural transformation `α` defined as: `α(U) = f : a ⟶ b` if +`p₀ ∈ U` and the unique morphism to a terminal object in `C` if `p₀ ∉ U`. +-/ +def skyscraper_sheaf_functor [has_products.{u} C] : C ⥤ sheaf C X := +{ obj := λ c, skyscraper_sheaf p₀ c, + map := λ a b f, Sheaf.hom.mk $ (skyscraper_presheaf_functor p₀).map f, + map_id' := λ c, Sheaf.hom.ext _ _ $ (skyscraper_presheaf_functor p₀).map_id _, + map_comp' := λ _ _ _ f g, Sheaf.hom.ext _ _ $ (skyscraper_presheaf_functor p₀).map_comp _ _ } + end From d79c3671549ace83a93d2a97f316d12ab42232ae Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sun, 2 Oct 2022 05:55:10 +0000 Subject: [PATCH 503/828] feat(algebraic_topology): extra degeneracy of augmented simplicial objects (#16411) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces the notion of extra degeneracy of augmented simplicial objects. In homotopy theory, this is a condition that is used to show that the connected components of simplicial sets are contractible. This notion is formalized for augmented simplicial objects in any category and it is shown that the standard `n`-simplex has an extra degeneracy. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- src/algebraic_topology/extra_degeneracy.lean | 188 ++++++++++++++++++ src/algebraic_topology/simplicial_object.lean | 8 + src/algebraic_topology/simplicial_set.lean | 22 +- 3 files changed, 217 insertions(+), 1 deletion(-) create mode 100644 src/algebraic_topology/extra_degeneracy.lean diff --git a/src/algebraic_topology/extra_degeneracy.lean b/src/algebraic_topology/extra_degeneracy.lean new file mode 100644 index 0000000000000..a853a23864d52 --- /dev/null +++ b/src/algebraic_topology/extra_degeneracy.lean @@ -0,0 +1,188 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.simplicial_set +import tactic.fin_cases + +/-! + +# Augmented simplicial objects with an extra degeneracy + +In simplicial homotopy theory, in order to prove that the connected components +of a simplicial set `X` are contractible, it suffices to construct an extra +degeneracy as it is defined in *Simplicial Homotopy Theory* by Goerss-Jardine p. 190. +It consists of a series of maps `π₀ X → X _[0]` and `X _[n] → X _[n+1]` which +behave formally like an extra degeneracy `σ (-1)`. It can be thought as a datum +associated to the augmented simplicial set `X → π₀ X`. + +In this file, we adapt this definition to the case of augmented +simplicial objects in any category. + +## Main definitions + +- the structure `extra_degeneracy X` for any `X : simplicial_object.augmented C` +- `extra_degeneracy.map`: extra degeneracies are preserved by the application of any +functor `C ⥤ D` +- `sSet.augmented.standard_simplex.extra_degeneracy`: the standard `n`-simplex has +an extra degeneracy + +TODO @joelriou: +1) when the category `C` is preadditive and has a zero object, and +`X : simplicial_object.augmented C` has an extra degeneracy, then the augmentation +on the alternating face map complex of `X` is a homotopy equivalence of chain +complexes. + +2) extra degeneracy for the Čech nerve of a split epi. In particular the +universal cover EG of the classifying space of a group G has an extra +degeneracy. + +## References +* [Paul G. Goerss, John F. Jardine, *Simplical Homotopy Theory*][goerss-jardine-2009] + +-/ + +open category_theory category_theory.category +open category_theory.simplicial_object.augmented +open opposite +open_locale simplicial + +universes u + +namespace simplicial_object + +namespace augmented + +variables {C : Type*} [category C] + +/-- The datum of an extra degeneracy is a technical condition on +augmented simplicial objects. The morphisms `s'` and `s n` of the +structure formally behave like extra degeneracies `σ (-1)`. -/ +@[ext] +structure extra_degeneracy (X : simplicial_object.augmented C) := +(s' : point.obj X ⟶ (drop.obj X) _[0]) +(s : Π (n : ℕ), (drop.obj X) _[n] ⟶ (drop.obj X) _[n+1]) +(s'_comp_ε' : s' ≫ X.hom.app (op [0]) = 𝟙 _) +(s₀_comp_δ₁' : s 0 ≫ (drop.obj X).δ 1 = X.hom.app (op [0]) ≫ s') +(s_comp_δ₀' : Π (n : ℕ), s n ≫ (drop.obj X).δ 0 = 𝟙 _) +(s_comp_δ' : Π (n : ℕ) (i : fin (n+2)), s (n+1) ≫ (drop.obj X).δ i.succ = + (drop.obj X).δ i ≫ s n) +(s_comp_σ' : Π (n : ℕ) (i : fin (n+1)), s n ≫ (drop.obj X).σ i.succ = + (drop.obj X).σ i ≫ s (n+1)) + +namespace extra_degeneracy + +restate_axiom s'_comp_ε' +restate_axiom s₀_comp_δ₁' +restate_axiom s_comp_δ₀' +restate_axiom s_comp_δ' +restate_axiom s_comp_σ' +attribute [reassoc] s'_comp_ε s₀_comp_δ₁ s_comp_δ₀ s_comp_δ s_comp_σ +attribute [simp] s'_comp_ε s_comp_δ₀ + +/-- If `ed` is an extra degeneracy for `X : simplicial_object.augmented C` and +`F : C ⥤ D` is a functor, then `ed.map F` is an extra degeneracy for the +augmented simplical object in `D` obtained by applying `F` to `X`. -/ +def map {D : Type*} [category D] + {X : simplicial_object.augmented C} (ed : extra_degeneracy X) (F : C ⥤ D) : + extra_degeneracy (((whiskering _ _).obj F).obj X) := +{ s' := F.map ed.s', + s := λ n, F.map (ed.s n), + s'_comp_ε' := by { dsimp, erw [comp_id, ← F.map_comp, ed.s'_comp_ε, F.map_id], }, + s₀_comp_δ₁' := by { dsimp, erw [comp_id, ← F.map_comp, ← F.map_comp, ed.s₀_comp_δ₁], }, + s_comp_δ₀' := λ n, by { dsimp, erw [← F.map_comp, ed.s_comp_δ₀, F.map_id], }, + s_comp_δ' := λ n i, by { dsimp, erw [← F.map_comp, ← F.map_comp, ed.s_comp_δ], refl, }, + s_comp_σ' := λ n i, by { dsimp, erw [← F.map_comp, ← F.map_comp, ed.s_comp_σ], refl, }, } + +end extra_degeneracy + +end augmented + +end simplicial_object + +namespace sSet + +namespace augmented + +namespace standard_simplex + +/-- When `[has_zero X]`, the shift of a map `f : fin n → X` +is a map `fin (n+1) → X` which sends `0` to `0` and `i.succ` to `f i`. -/ +def shift_fun {n : ℕ} {X : Type*} [has_zero X] (f : fin n → X) (i : fin (n+1)) : X := +dite (i = 0) (λ h, 0) (λ h, f (i.pred h)) + +@[simp] +lemma shift_fun_0 {n : ℕ} {X : Type*} [has_zero X] (f : fin n → X) : shift_fun f 0 = 0 := rfl + +@[simp] +lemma shift_fun_succ {n : ℕ} {X : Type*} [has_zero X] (f : fin n → X) + (i : fin n) : shift_fun f i.succ = f i := +begin + dsimp [shift_fun], + split_ifs, + { exfalso, + simpa only [fin.ext_iff, fin.coe_succ] using h, }, + { simp only [fin.pred_succ], }, +end + +/-- The shift of a morphism `f : [n] → Δ` in `simplex_category` corresponds to +the monotone map which sends `0` to `0` and `i.succ` to `f.to_order_hom i`. -/ +@[simp] +def shift {n : ℕ} {Δ : simplex_category} (f : [n] ⟶ Δ) : [n+1] ⟶ Δ := simplex_category.hom.mk +{ to_fun := shift_fun f.to_order_hom, + monotone' := λ i₁ i₂ hi, begin + by_cases h₁ : i₁ = 0, + { subst h₁, + simp only [shift_fun_0, fin.zero_le], }, + { have h₂ : i₂ ≠ 0 := by { intro h₂, subst h₂, exact h₁ (le_antisymm hi (fin.zero_le _)), }, + cases fin.eq_succ_of_ne_zero h₁ with j₁ hj₁, + cases fin.eq_succ_of_ne_zero h₂ with j₂ hj₂, + substs hj₁ hj₂, + simpa only [shift_fun_succ] using f.to_order_hom.monotone (fin.succ_le_succ_iff.mp hi), }, + end, } + +/-- The obvious extra degeneracy on the standard simplex. -/ +@[protected] +def extra_degeneracy (Δ : simplex_category) : + simplicial_object.augmented.extra_degeneracy (standard_simplex.obj Δ) := +{ s' := λ x, simplex_category.hom.mk (order_hom.const _ 0), + s := λ n f, shift f, + s'_comp_ε' := by { ext1 j, fin_cases j, }, + s₀_comp_δ₁' := by { ext x j, fin_cases j, refl, }, + s_comp_δ₀' := λ n, begin + ext φ i : 4, + dsimp [simplicial_object.δ, simplex_category.δ, sSet.standard_simplex], + simp only [shift_fun_succ], + end, + s_comp_δ' := λ n i, begin + ext φ j : 4, + dsimp [simplicial_object.δ, simplex_category.δ, sSet.standard_simplex], + by_cases j = 0, + { subst h, + simp only [fin.succ_succ_above_zero, shift_fun_0], }, + { cases fin.eq_succ_of_ne_zero h with k hk, + subst hk, + simp only [fin.succ_succ_above_succ, shift_fun_succ], }, + end, + s_comp_σ' := λ n i, begin + ext φ j : 4, + dsimp [simplicial_object.σ, simplex_category.σ, sSet.standard_simplex], + by_cases j = 0, + { subst h, + simpa only [shift_fun_0] using shift_fun_0 φ.to_order_hom, }, + { cases fin.eq_succ_of_ne_zero h with k hk, + subst hk, + simp only [fin.succ_pred_above_succ, shift_fun_succ], }, + end, } + +instance nonempty_extra_degeneracy_standard_simplex (Δ : simplex_category) : + nonempty (simplicial_object.augmented.extra_degeneracy (standard_simplex.obj Δ)) := +⟨standard_simplex.extra_degeneracy Δ⟩ + +end standard_simplex + +end augmented + +end sSet diff --git a/src/algebraic_topology/simplicial_object.lean b/src/algebraic_topology/simplicial_object.lean index 13aaae698beba..43e6a004e4e61 100644 --- a/src/algebraic_topology/simplicial_object.lean +++ b/src/algebraic_topology/simplicial_object.lean @@ -204,6 +204,14 @@ def to_arrow : augmented C ⥤ arrow C := refl, end } } +/-- The compatibility of a morphism with the augmentation, on 0-simplices -/ +@[reassoc] +lemma w₀ {X Y : augmented C} (f : X ⟶ Y) : + (augmented.drop.map f).app (op (simplex_category.mk 0)) ≫ + Y.hom.app (op (simplex_category.mk 0)) = + X.hom.app (op (simplex_category.mk 0)) ≫ augmented.point.map f := +by convert congr_app f.w (op (simplex_category.mk 0)) + variable (C) /-- Functor composition induces a functor on augmented simplicial objects. -/ diff --git a/src/algebraic_topology/simplicial_set.lean b/src/algebraic_topology/simplicial_set.lean index 9dd4523e29d54..6727591420e4e 100644 --- a/src/algebraic_topology/simplicial_set.lean +++ b/src/algebraic_topology/simplicial_set.lean @@ -32,7 +32,7 @@ a morphism `Δ[n] ⟶ ∂Δ[n]`. universes v u -open category_theory +open category_theory category_theory.limits open_locale simplicial @@ -120,6 +120,26 @@ def sk (n : ℕ) : sSet ⥤ sSet.truncated n := simplicial_object.sk n instance {n} : inhabited (sSet.truncated n) := ⟨(sk n).obj $ Δ[0]⟩ +/-- The category of augmented simplicial sets, as a particular case of +augmented simplicial objects. -/ +abbreviation augmented := simplicial_object.augmented (Type u) + +namespace augmented + +/-- The functor which sends `[n]` to the simplicial set `Δ[n]` equipped by +the obvious augmentation towards the terminal object of the category of sets. -/ +@[simps] +noncomputable def standard_simplex : simplex_category ⥤ sSet.augmented := +{ obj := λ Δ, + { left := sSet.standard_simplex.obj Δ, + right := terminal _, + hom := { app := λ Δ', terminal.from _, }, }, + map := λ Δ₁ Δ₂ θ, + { left := sSet.standard_simplex.map θ, + right := terminal.from _, }, } + +end augmented + end sSet /-- The functor associating the singular simplicial set to a topological space. -/ From df6a0b2e7b1bf8acffae0310f3ecc19153b33abd Mon Sep 17 00:00:00 2001 From: datokrat Date: Sun, 2 Oct 2022 06:53:32 +0000 Subject: [PATCH 504/828] feat(analysis/convex/body): define bodies and implement module instance (#16297) Defines the type `convex_body V` and endows it with a module structure over the nonnegative reals. This commit also introduces `set_like` and `inhabited` instances. Co-authored-by: Junyan Xu --- src/analysis/convex/body.lean | 115 +++++++++++++++++++++ src/topology/algebra/const_mul_action.lean | 5 + 2 files changed, 120 insertions(+) create mode 100644 src/analysis/convex/body.lean diff --git a/src/analysis/convex/body.lean b/src/analysis/convex/body.lean new file mode 100644 index 0000000000000..b3f1c32e2479e --- /dev/null +++ b/src/analysis/convex/body.lean @@ -0,0 +1,115 @@ +/- +Copyright (c) 2022 Paul A. Reichert. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Paul A. Reichert +-/ +import analysis.convex.basic +import analysis.normed_space.basic +import data.real.nnreal +import data.set.pointwise +import topology.subset_properties + +/-! +# convex bodies + +This file contains the definition of the type `convex_body V` +consisting of +convex, compact, nonempty subsets of a real normed space `V`. + +`convex_body V` is a module over the nonnegative reals (`nnreal`). + +TODOs: +- endow it with the Hausdorff metric +- define positive convex bodies, requiring the interior to be nonempty +- introduce support sets + +## Tags + +convex, convex body +-/ + +open_locale pointwise +open_locale nnreal + +variables (V : Type*) [seminormed_add_comm_group V] [normed_space ℝ V] + +/-- +Let `V` be a normed space. A subset of `V` is a convex body if and only if +it is convex, compact, and nonempty. +-/ +structure convex_body := +(carrier : set V) +(convex' : convex ℝ carrier) +(is_compact' : is_compact carrier) +(nonempty' : carrier.nonempty) + +namespace convex_body + +variables {V} + +instance : set_like (convex_body V) V := +{ coe := convex_body.carrier, + coe_injective' := λ K L h, by { cases K, cases L, congr' } } + +lemma convex (K : convex_body V) : convex ℝ (K : set V) := K.convex' +lemma is_compact (K : convex_body V) : is_compact (K : set V) := K.is_compact' +lemma nonempty (K : convex_body V) : (K : set V).nonempty := K.nonempty' + +@[ext] +protected lemma ext {K L : convex_body V} (h : (K : set V) = L) : K = L := set_like.ext' h + +@[simp] +lemma coe_mk (s : set V) (h₁ h₂ h₃) : (mk s h₁ h₂ h₃ : set V) = s := rfl + +instance : add_monoid (convex_body V) := +-- we cannot write K + L to avoid reducibility issues with the set.has_add instance +{ add := λ K L, ⟨set.image2 (+) K L, + K.convex.add L.convex, + K.is_compact.add L.is_compact, + K.nonempty.add L.nonempty⟩, + add_assoc := λ K L M, by { ext, simp only [coe_mk, set.image2_add, add_assoc] }, + zero := ⟨0, convex_singleton 0, is_compact_singleton, set.singleton_nonempty 0⟩, + zero_add := λ K, by { ext, simp only [coe_mk, set.image2_add, zero_add] }, + add_zero := λ K, by { ext, simp only [coe_mk, set.image2_add, add_zero] } } + +@[simp] +lemma coe_add (K L : convex_body V) : (↑(K + L) : set V) = (K : set V) + L := rfl + +@[simp] +lemma coe_zero : (↑(0 : convex_body V) : set V) = 0 := rfl + +instance : inhabited (convex_body V) := ⟨0⟩ + +instance : add_comm_monoid (convex_body V) := +{ add_comm := λ K L, by { ext, simp only [coe_add, add_comm] }, + .. convex_body.add_monoid } + +instance : has_smul ℝ (convex_body V) := +{ smul := λ c K, ⟨c • (K : set V), K.convex.smul _, K.is_compact.smul _, K.nonempty.smul_set⟩ } + +@[simp] +lemma coe_smul (c : ℝ) (K : convex_body V) : (↑(c • K) : set V) = c • (K : set V) := rfl + +instance : distrib_mul_action ℝ (convex_body V) := +{ to_has_smul := convex_body.has_smul, + one_smul := λ K, by { ext, simp only [coe_smul, one_smul] }, + mul_smul := λ c d K, by { ext, simp only [coe_smul, mul_smul] }, + smul_add := λ c K L, by { ext, simp only [coe_smul, coe_add, smul_add] }, + smul_zero := λ c, by { ext, simp only [coe_smul, coe_zero, smul_zero] } } + +@[simp] +lemma coe_smul' (c : ℝ≥0) (K : convex_body V) : (↑(c • K) : set V) = c • (K : set V) := rfl + +/-- +The convex bodies in a fixed space $V$ form a module over the nonnegative reals. +-/ +instance : module ℝ≥0 (convex_body V) := +{ add_smul := λ c d K, + begin + ext1, + simp only [coe_smul, coe_add], + exact convex.add_smul K.convex (nnreal.coe_nonneg _) (nnreal.coe_nonneg _), + end, + zero_smul := λ K, by { ext1, exact set.zero_smul_set K.nonempty } } + +end convex_body diff --git a/src/topology/algebra/const_mul_action.lean b/src/topology/algebra/const_mul_action.lean index 460e5dd64bdf6..f03025e3d0646 100644 --- a/src/topology/algebra/const_mul_action.lean +++ b/src/topology/algebra/const_mul_action.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Alex Kontorovich, Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex Kontorovich, Heather Macbeth -/ +import data.real.nnreal import topology.algebra.constructions import topology.homeomorph import group_theory.group_action.basic @@ -114,6 +115,10 @@ instance {ι : Type*} {γ : ι → Type*} [∀ i, topological_space (γ i)] [Π [∀ i, has_continuous_const_smul M (γ i)] : has_continuous_const_smul M (Π i, γ i) := ⟨λ _, continuous_pi $ λ i, (continuous_apply i).const_smul _⟩ +lemma is_compact.smul {α β} [has_smul α β] [topological_space β] + [has_continuous_const_smul α β] (a : α) {s : set β} + (hs : is_compact s) : is_compact (a • s) := hs.image (continuous_id'.const_smul a) + end has_smul section monoid From 73d05c46b08a3db2c1b14cce11121daaa7552ded Mon Sep 17 00:00:00 2001 From: mcdoll Date: Sun, 2 Oct 2022 12:39:50 +0000 Subject: [PATCH 505/828] feat(analysis/locally_convex): first countable topologies from countable families of seminorms (#16595) This PR proves that if the topology is induced by a countable family of seminorms, then it is first countable. --- .../locally_convex/with_seminorms.lean | 20 +++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 8f2a3de99edb4..065641be0c2bb 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -546,3 +546,23 @@ begin end end topological_constructions + +section topological_properties + +variables [nontrivially_normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [nonempty ι] [countable ι] +variables {p : seminorm_family 𝕜 E ι} +variables [uniform_space E] [uniform_add_group E] + +/-- If the topology of a space is induced by a countable family of seminorms, then the topology +is first countable. -/ +lemma with_seminorms.first_countable (hp : with_seminorms p) : + topological_space.first_countable_topology E := +begin + haveI : (𝓝 (0 : E)).is_countably_generated, + { rw p.with_seminorms_iff_nhds_eq_infi.mp hp, + exact filter.infi.is_countably_generated _ }, + haveI : (uniformity E).is_countably_generated := uniform_add_group.uniformity_countably_generated, + exact uniform_space.first_countable_topology E, +end + +end topological_properties From b2381e5a6452f45c1ed6232e8277d578597e5f65 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Sun, 2 Oct 2022 15:17:50 +0000 Subject: [PATCH 506/828] feat(analysis/locally_convex/with_seminorms): boundedness of images (#16674) This PR adds two lemmas that are very useful in proving that semilinear maps are locally bounded (and hence continuous). We also add some documentation for these lemmas and remove documentation for the ad-hoc lemmas. --- .../locally_convex/with_seminorms.lean | 28 +++++++++++++++++-- 1 file changed, 26 insertions(+), 2 deletions(-) diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 065641be0c2bb..4fc6e1004b13e 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -19,10 +19,24 @@ bounded by a finite number of seminorms in `E`. ## Main statements -* `continuous_from_bounded`: A bounded linear map `f : E →ₗ[𝕜] F` is continuous. * `seminorm_family.to_locally_convex_space`: A space equipped with a family of seminorms is locally convex. +## Continuity of semilinear maps + +If `E` and `F` are topological vector space with the topology induced by a family of seminorms, then +we have a direct method to prove that a linear map is continuous: +* `seminorm.continuous_from_bounded`: A bounded linear map `f : E →ₗ[𝕜] F` is continuous. + +If the topology of a space `E` is induced by a family of seminorms, then we can characterize von +Neumann boundedness in terms of that seminorm family. Together with +`linear_map.continuous_of_locally_bounded` this gives general criterion for continuity. + +* `with_seminorms.is_vonN_bounded_iff_finset_seminorm_bounded` +* `with_seminorms.is_vonN_bounded_iff_seminorm_bounded` +* `with_seminorms.image_is_vonN_bounded_iff_finset_seminorm_bounded` +* `with_seminorms.image_is_vonN_bounded_iff_seminorm_bounded` + ## Tags seminorm, locally convex @@ -370,7 +384,12 @@ begin exact (finset.sup I p).ball_zero_absorbs_ball_zero hr, end -lemma bornology.is_vonN_bounded_iff_seminorm_bounded {s : set E} (hp : with_seminorms p) : +lemma with_seminorms.image_is_vonN_bounded_iff_finset_seminorm_bounded (f : G → E) {s : set G} + (hp : with_seminorms p) : bornology.is_vonN_bounded 𝕜 (f '' s) ↔ + ∀ I : finset ι, ∃ r (hr : 0 < r), ∀ (x ∈ s), I.sup p (f x) < r := +by simp_rw [hp.is_vonN_bounded_iff_finset_seminorm_bounded, set.ball_image_iff] + +lemma with_seminorms.is_vonN_bounded_iff_seminorm_bounded {s : set E} (hp : with_seminorms p) : bornology.is_vonN_bounded 𝕜 s ↔ ∀ i : ι, ∃ r (hr : 0 < r), ∀ (x ∈ s), p i x < r := begin rw hp.is_vonN_bounded_iff_finset_seminorm_bounded, @@ -392,6 +411,11 @@ begin exact ⟨1, zero_lt_one, λ _ _, zero_lt_one⟩, end +lemma with_seminorms.image_is_vonN_bounded_iff_seminorm_bounded (f : G → E) {s : set G} + (hp : with_seminorms p) : + bornology.is_vonN_bounded 𝕜 (f '' s) ↔ ∀ i : ι, ∃ r (hr : 0 < r), ∀ (x ∈ s), p i (f x) < r := +by simp_rw [hp.is_vonN_bounded_iff_seminorm_bounded, set.ball_image_iff] + end nontrivially_normed_field section continuous_bounded From c4878d17d3c17dfe1cedb7f12e1c96ea82dea909 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Sun, 2 Oct 2022 16:23:54 +0000 Subject: [PATCH 507/828] feat(data/fintype): complement of univ (#16623) Co-authored-by: Bhavik Mehta --- src/data/fintype/basic.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index ff867a1a59911..3530f8ed67ccd 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -123,6 +123,8 @@ instance : order_top (finset α) := { top := univ, le_top := subset_univ } +@[simp] lemma top_eq_univ : (⊤ : finset α) = univ := rfl + section boolean_algebra variables [decidable_eq α] {a : α} @@ -141,6 +143,12 @@ set.ext $ λ x, mem_compl @[simp] lemma compl_empty : (∅ : finset α)ᶜ = univ := compl_bot +@[simp] lemma compl_univ : (univ : finset α)ᶜ = ∅ := compl_top + +@[simp] lemma compl_eq_empty_iff (s : finset α) : sᶜ = ∅ ↔ s = univ := compl_eq_bot + +@[simp] lemma compl_eq_univ_iff (s : finset α) : sᶜ = univ ↔ s = ∅ := compl_eq_top + @[simp] lemma union_compl (s : finset α) : s ∪ sᶜ = univ := sup_compl_eq_top @[simp] lemma inter_compl (s : finset α) : s ∩ sᶜ = ∅ := inf_compl_eq_bot From 2bb1bae8bae866933b1524747acbb38cc7334237 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Sun, 2 Oct 2022 19:22:08 +0000 Subject: [PATCH 508/828] chore(analysis/locally_convex/[continuous_of_]bounded): import `analysis.normed_space.is_R_or_C` later (#16758) We don't want to import `analysis/normed_space/is_R_or_C` in `analysis/locally_convex/bounded` because it will create a cycle when defining the strong topology on `continuous_linear_map`, because we will have to (transitively) import `analysis/locally_convex/bounded` in `analysis/normed_space/operator_norm`. So I moved the material of #16550 in a new file `analysis/locally_convex/continuous_of_bounded` --- src/analysis/locally_convex/bounded.lean | 152 --------------- .../locally_convex/continuous_of_bounded.lean | 174 ++++++++++++++++++ 2 files changed, 174 insertions(+), 152 deletions(-) create mode 100644 src/analysis/locally_convex/continuous_of_bounded.lean diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index 971e94e01dc22..cf0e7fa50f7f0 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -5,7 +5,6 @@ Authors: Moritz Doll -/ import analysis.locally_convex.basic import analysis.locally_convex.balanced_core_hull -import analysis.normed_space.is_R_or_C import analysis.seminorm import topology.bornology.basic import topology.algebra.uniform_group @@ -27,8 +26,6 @@ absorbs `s`. * `bornology.is_vonN_bounded_of_topological_space_le`: A coarser topology admits more von Neumann-bounded sets. * `bornology.is_vonN_bounded.image`: A continuous linear image of a bounded set is bounded. -* `linear_map.continuous_of_locally_bounded`: If `E` is first countable, then every -locally bounded linear map `E →ₛₗ[σ] F` is continuous. ## References @@ -197,155 +194,6 @@ end end uniform_add_group -section continuous_linear_map - -variables [add_comm_group E] [uniform_space E] [uniform_add_group E] -variables [add_comm_group F] [uniform_space F] - -section nontrivially_normed_field - -variables [uniform_add_group F] -variables [nontrivially_normed_field 𝕜] [module 𝕜 E] [module 𝕜 F] [has_continuous_smul 𝕜 E] - -/-- Construct a continuous linear map from a linear map `f : E →ₗ[𝕜] F` and the existence of a -neighborhood of zero that gets mapped into a bounded set in `F`. -/ -def linear_map.clm_of_exists_bounded_image (f : E →ₗ[𝕜] F) - (h : ∃ (V : set E) (hV : V ∈ 𝓝 (0 : E)), bornology.is_vonN_bounded 𝕜 (f '' V)) : E →L[𝕜] F := -⟨f, begin - -- It suffices to show that `f` is continuous at `0`. - refine continuous_of_continuous_at_zero f _, - rw [continuous_at_def, f.map_zero], - intros U hU, - -- Continuity means that `U ∈ 𝓝 0` implies that `f ⁻¹' U ∈ 𝓝 0`. - rcases h with ⟨V, hV, h⟩, - rcases h hU with ⟨r, hr, h⟩, - rcases normed_field.exists_lt_norm 𝕜 r with ⟨x, hx⟩, - specialize h x hx.le, - -- After unfolding all the definitions, we know that `f '' V ⊆ x • U`. We use this to show the - -- inclusion `x⁻¹ • V ⊆ f⁻¹' U`. - have x_ne := norm_pos_iff.mp (hr.trans hx), - have : x⁻¹ • V ⊆ f⁻¹' U := - calc x⁻¹ • V ⊆ x⁻¹ • (f⁻¹' (f '' V)) : set.smul_set_mono (set.subset_preimage_image ⇑f V) - ... ⊆ x⁻¹ • (f⁻¹' (x • U)) : set.smul_set_mono (set.preimage_mono h) - ... = f⁻¹' (x⁻¹ • (x • U)) : - by ext; simp only [set.mem_inv_smul_set_iff₀ x_ne, set.mem_preimage, linear_map.map_smul] - ... ⊆ f⁻¹' U : by rw inv_smul_smul₀ x_ne _, - -- Using this inclusion, it suffices to show that `x⁻¹ • V` is in `𝓝 0`, which is trivial. - refine mem_of_superset _ this, - convert set_smul_mem_nhds_smul hV (inv_ne_zero x_ne), - exact (smul_zero _).symm, -end⟩ - -lemma linear_map.clm_of_exists_bounded_image_coe {f : E →ₗ[𝕜] F} - {h : ∃ (V : set E) (hV : V ∈ 𝓝 (0 : E)), bornology.is_vonN_bounded 𝕜 (f '' V)} : - (f.clm_of_exists_bounded_image h : E →ₗ[𝕜] F) = f := rfl - -@[simp] lemma linear_map.clm_of_exists_bounded_image_apply {f : E →ₗ[𝕜] F} - {h : ∃ (V : set E) (hV : V ∈ 𝓝 (0 : E)), bornology.is_vonN_bounded 𝕜 (f '' V)} {x : E} : - f.clm_of_exists_bounded_image h x = f x := rfl - -end nontrivially_normed_field - -section is_R_or_C - -open topological_space bornology - -variables [first_countable_topology E] -variables [is_R_or_C 𝕜] [module 𝕜 E] [has_continuous_smul 𝕜 E] -variables [is_R_or_C 𝕜'] [module 𝕜' F] [has_continuous_smul 𝕜' F] -variables {σ : 𝕜 →+* 𝕜'} - -lemma linear_map.continuous_at_zero_of_locally_bounded (f : E →ₛₗ[σ] F) - (hf : ∀ (s : set E) (hs : is_vonN_bounded 𝕜 s), is_vonN_bounded 𝕜' (f '' s)) : - continuous_at f 0 := -begin - -- Assume that f is not continuous at 0 - by_contradiction, - -- We use the a decreasing balanced basis for 0 : E and a balanced basis for 0 : F - -- and reformulate non-continuity in terms of these bases - rcases (nhds_basis_balanced 𝕜 E).exists_antitone_subbasis with ⟨b, bE1, bE⟩, - simp only [id.def] at bE, - have bE' : (𝓝 (0 : E)).has_basis (λ (x : ℕ), x ≠ 0) (λ n : ℕ, (n : 𝕜)⁻¹ • b n) := - begin - refine bE.1.to_has_basis _ _, - { intros n _, - use n+1, - simp only [ne.def, nat.succ_ne_zero, not_false_iff, nat.cast_add, nat.cast_one, true_and], - -- `b (n + 1) ⊆ b n` follows from `antitone`. - have h : b (n + 1) ⊆ b n := bE.2 (by simp), - refine subset_trans _ h, - rintros y ⟨x, hx, hy⟩, - -- Since `b (n + 1)` is balanced `(n+1)⁻¹ b (n + 1) ⊆ b (n + 1)` - rw ←hy, - refine (bE1 (n+1)).2.smul_mem _ hx, - have h' : 0 < (n : ℝ) + 1 := n.cast_add_one_pos, - rw [norm_inv, ←nat.cast_one, ←nat.cast_add, is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat, - nat.cast_add, nat.cast_one, inv_le h' zero_lt_one], - norm_cast, - simp, }, - intros n hn, - -- The converse direction follows from continuity of the scalar multiplication - have hcont : continuous_at (λ (x : E), (n : 𝕜) • x) 0 := - (continuous_const_smul (n : 𝕜)).continuous_at, - simp only [continuous_at, map_zero, smul_zero] at hcont, - rw bE.1.tendsto_left_iff at hcont, - rcases hcont (b n) (bE1 n).1 with ⟨i, _, hi⟩, - refine ⟨i, trivial, λ x hx, ⟨(n : 𝕜) • x, hi hx, _⟩⟩, - simp [←mul_smul, hn], - end, - rw [continuous_at, map_zero, bE'.tendsto_iff (nhds_basis_balanced 𝕜' F)] at h, - push_neg at h, - rcases h with ⟨V, ⟨hV, hV'⟩, h⟩, - simp only [id.def, forall_true_left] at h, - -- There exists `u : ℕ → E` such that for all `n : ℕ` we have `u n ∈ n⁻¹ • b n` and `f (u n) ∉ V` - choose! u hu hu' using h, - -- The sequence `(λ n, n • u n)` converges to `0` - have h_tendsto : tendsto (λ n : ℕ, (n : 𝕜) • u n) at_top (𝓝 (0 : E)) := - begin - apply bE.tendsto, - intros n, - by_cases h : n = 0, - { rw [h, nat.cast_zero, zero_smul], - refine mem_of_mem_nhds (bE.1.mem_of_mem $ by triv) }, - rcases hu n h with ⟨y, hy, hu1⟩, - convert hy, - rw [←hu1, ←mul_smul], - simp only [h, mul_inv_cancel, ne.def, nat.cast_eq_zero, not_false_iff, one_smul], - end, - -- The image `(λ n, n • u n)` is von Neumann bounded: - have h_bounded : is_vonN_bounded 𝕜 (set.range (λ n : ℕ, (n : 𝕜) • u n)) := - h_tendsto.cauchy_seq.totally_bounded_range.is_vonN_bounded 𝕜, - -- Since `range u` is bounded it absorbs `V` - rcases hf _ h_bounded hV with ⟨r, hr, h'⟩, - cases exists_nat_gt r with n hn, - -- We now find a contradiction between `f (u n) ∉ V` and the absorbing property - have h1 : r ≤ ∥(n : 𝕜')∥ := - by { rw [is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat], exact hn.le }, - have hn' : 0 < ∥(n : 𝕜')∥ := lt_of_lt_of_le hr h1, - rw [norm_pos_iff, ne.def, nat.cast_eq_zero] at hn', - have h'' : f (u n) ∈ V := - begin - simp only [set.image_subset_iff] at h', - specialize h' (n : 𝕜') h1 (set.mem_range_self n), - simp only [set.mem_preimage, linear_map.map_smulₛₗ, map_nat_cast] at h', - rcases h' with ⟨y, hy, h'⟩, - apply_fun (λ y : F, (n : 𝕜')⁻¹ • y) at h', - simp only [hn', inv_smul_smul₀, ne.def, nat.cast_eq_zero, not_false_iff] at h', - rwa ←h', - end, - exact hu' n hn' h'', -end - -/-- If `E` is first countable, then every locally bounded linear map `E →ₛₗ[σ] F` is continuous. -/ -lemma linear_map.continuous_of_locally_bounded [uniform_add_group F] (f : E →ₛₗ[σ] F) - (hf : ∀ (s : set E) (hs : is_vonN_bounded 𝕜 s), is_vonN_bounded 𝕜' (f '' s)) : - continuous f := -(uniform_continuous_of_continuous_at_zero f $ f.continuous_at_zero_of_locally_bounded hf).continuous - -end is_R_or_C - -end continuous_linear_map - section vonN_bornology_eq_metric variables (𝕜 E) [nontrivially_normed_field 𝕜] [seminormed_add_comm_group E] [normed_space 𝕜 E] diff --git a/src/analysis/locally_convex/continuous_of_bounded.lean b/src/analysis/locally_convex/continuous_of_bounded.lean new file mode 100644 index 0000000000000..b183a05516361 --- /dev/null +++ b/src/analysis/locally_convex/continuous_of_bounded.lean @@ -0,0 +1,174 @@ +/- +Copyright (c) 2022 Anatole Dedecker. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Moritz Doll +-/ +import analysis.locally_convex.bounded +import analysis.normed_space.is_R_or_C + +/-! +# Continuity and Von Neumann boundedness + +This files proves that for `E` and `F` two topological vector spaces over `ℝ` or `ℂ`, +if `E` is first countable, then every locally bounded linear map `E →ₛₗ[σ] F` is continuous +(this is `linear_map.continuous_of_locally_bounded`). + +We keep this file separate from `analysis/locally_convex/bounded` in order not to import +`analysis/normed_space/is_R_or_C` there, because defining the strong topology on the space of +continuous linear maps will require importing `analysis/locally_convex/bounded` in +`analysis/normed_space/operator_norm`. + +## References + +* [Bourbaki, *Topological Vector Spaces*][bourbaki1987] + +-/ + +open topological_space bornology filter +open_locale topological_space pointwise + +variables {𝕜 𝕜' E F : Type*} +variables [add_comm_group E] [uniform_space E] [uniform_add_group E] +variables [add_comm_group F] [uniform_space F] + +section nontrivially_normed_field + +variables [uniform_add_group F] +variables [nontrivially_normed_field 𝕜] [module 𝕜 E] [module 𝕜 F] [has_continuous_smul 𝕜 E] + +/-- Construct a continuous linear map from a linear map `f : E →ₗ[𝕜] F` and the existence of a +neighborhood of zero that gets mapped into a bounded set in `F`. -/ +def linear_map.clm_of_exists_bounded_image (f : E →ₗ[𝕜] F) + (h : ∃ (V : set E) (hV : V ∈ 𝓝 (0 : E)), bornology.is_vonN_bounded 𝕜 (f '' V)) : E →L[𝕜] F := +⟨f, begin + -- It suffices to show that `f` is continuous at `0`. + refine continuous_of_continuous_at_zero f _, + rw [continuous_at_def, f.map_zero], + intros U hU, + -- Continuity means that `U ∈ 𝓝 0` implies that `f ⁻¹' U ∈ 𝓝 0`. + rcases h with ⟨V, hV, h⟩, + rcases h hU with ⟨r, hr, h⟩, + rcases normed_field.exists_lt_norm 𝕜 r with ⟨x, hx⟩, + specialize h x hx.le, + -- After unfolding all the definitions, we know that `f '' V ⊆ x • U`. We use this to show the + -- inclusion `x⁻¹ • V ⊆ f⁻¹' U`. + have x_ne := norm_pos_iff.mp (hr.trans hx), + have : x⁻¹ • V ⊆ f⁻¹' U := + calc x⁻¹ • V ⊆ x⁻¹ • (f⁻¹' (f '' V)) : set.smul_set_mono (set.subset_preimage_image ⇑f V) + ... ⊆ x⁻¹ • (f⁻¹' (x • U)) : set.smul_set_mono (set.preimage_mono h) + ... = f⁻¹' (x⁻¹ • (x • U)) : + by ext; simp only [set.mem_inv_smul_set_iff₀ x_ne, set.mem_preimage, linear_map.map_smul] + ... ⊆ f⁻¹' U : by rw inv_smul_smul₀ x_ne _, + -- Using this inclusion, it suffices to show that `x⁻¹ • V` is in `𝓝 0`, which is trivial. + refine mem_of_superset _ this, + convert set_smul_mem_nhds_smul hV (inv_ne_zero x_ne), + exact (smul_zero _).symm, +end⟩ + +lemma linear_map.clm_of_exists_bounded_image_coe {f : E →ₗ[𝕜] F} + {h : ∃ (V : set E) (hV : V ∈ 𝓝 (0 : E)), bornology.is_vonN_bounded 𝕜 (f '' V)} : + (f.clm_of_exists_bounded_image h : E →ₗ[𝕜] F) = f := rfl + +@[simp] lemma linear_map.clm_of_exists_bounded_image_apply {f : E →ₗ[𝕜] F} + {h : ∃ (V : set E) (hV : V ∈ 𝓝 (0 : E)), bornology.is_vonN_bounded 𝕜 (f '' V)} {x : E} : + f.clm_of_exists_bounded_image h x = f x := rfl + +end nontrivially_normed_field + +section is_R_or_C + +open topological_space bornology + +variables [first_countable_topology E] +variables [is_R_or_C 𝕜] [module 𝕜 E] [has_continuous_smul 𝕜 E] +variables [is_R_or_C 𝕜'] [module 𝕜' F] [has_continuous_smul 𝕜' F] +variables {σ : 𝕜 →+* 𝕜'} + +lemma linear_map.continuous_at_zero_of_locally_bounded (f : E →ₛₗ[σ] F) + (hf : ∀ (s : set E) (hs : is_vonN_bounded 𝕜 s), is_vonN_bounded 𝕜' (f '' s)) : + continuous_at f 0 := +begin + -- Assume that f is not continuous at 0 + by_contradiction, + -- We use a decreasing balanced basis for 0 : E and a balanced basis for 0 : F + -- and reformulate non-continuity in terms of these bases + rcases (nhds_basis_balanced 𝕜 E).exists_antitone_subbasis with ⟨b, bE1, bE⟩, + simp only [id.def] at bE, + have bE' : (𝓝 (0 : E)).has_basis (λ (x : ℕ), x ≠ 0) (λ n : ℕ, (n : 𝕜)⁻¹ • b n) := + begin + refine bE.1.to_has_basis _ _, + { intros n _, + use n+1, + simp only [ne.def, nat.succ_ne_zero, not_false_iff, nat.cast_add, nat.cast_one, true_and], + -- `b (n + 1) ⊆ b n` follows from `antitone`. + have h : b (n + 1) ⊆ b n := bE.2 (by simp), + refine subset_trans _ h, + rintros y ⟨x, hx, hy⟩, + -- Since `b (n + 1)` is balanced `(n+1)⁻¹ b (n + 1) ⊆ b (n + 1)` + rw ←hy, + refine (bE1 (n+1)).2.smul_mem _ hx, + have h' : 0 < (n : ℝ) + 1 := n.cast_add_one_pos, + rw [norm_inv, ←nat.cast_one, ←nat.cast_add, is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat, + nat.cast_add, nat.cast_one, inv_le h' zero_lt_one], + norm_cast, + simp, }, + intros n hn, + -- The converse direction follows from continuity of the scalar multiplication + have hcont : continuous_at (λ (x : E), (n : 𝕜) • x) 0 := + (continuous_const_smul (n : 𝕜)).continuous_at, + simp only [continuous_at, map_zero, smul_zero] at hcont, + rw bE.1.tendsto_left_iff at hcont, + rcases hcont (b n) (bE1 n).1 with ⟨i, _, hi⟩, + refine ⟨i, trivial, λ x hx, ⟨(n : 𝕜) • x, hi hx, _⟩⟩, + simp [←mul_smul, hn], + end, + rw [continuous_at, map_zero, bE'.tendsto_iff (nhds_basis_balanced 𝕜' F)] at h, + push_neg at h, + rcases h with ⟨V, ⟨hV, hV'⟩, h⟩, + simp only [id.def, forall_true_left] at h, + -- There exists `u : ℕ → E` such that for all `n : ℕ` we have `u n ∈ n⁻¹ • b n` and `f (u n) ∉ V` + choose! u hu hu' using h, + -- The sequence `(λ n, n • u n)` converges to `0` + have h_tendsto : tendsto (λ n : ℕ, (n : 𝕜) • u n) at_top (𝓝 (0 : E)) := + begin + apply bE.tendsto, + intros n, + by_cases h : n = 0, + { rw [h, nat.cast_zero, zero_smul], + refine mem_of_mem_nhds (bE.1.mem_of_mem $ by triv) }, + rcases hu n h with ⟨y, hy, hu1⟩, + convert hy, + rw [←hu1, ←mul_smul], + simp only [h, mul_inv_cancel, ne.def, nat.cast_eq_zero, not_false_iff, one_smul], + end, + -- The image `(λ n, n • u n)` is von Neumann bounded: + have h_bounded : is_vonN_bounded 𝕜 (set.range (λ n : ℕ, (n : 𝕜) • u n)) := + h_tendsto.cauchy_seq.totally_bounded_range.is_vonN_bounded 𝕜, + -- Since `range u` is bounded it absorbs `V` + rcases hf _ h_bounded hV with ⟨r, hr, h'⟩, + cases exists_nat_gt r with n hn, + -- We now find a contradiction between `f (u n) ∉ V` and the absorbing property + have h1 : r ≤ ∥(n : 𝕜')∥ := + by { rw [is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat], exact hn.le }, + have hn' : 0 < ∥(n : 𝕜')∥ := lt_of_lt_of_le hr h1, + rw [norm_pos_iff, ne.def, nat.cast_eq_zero] at hn', + have h'' : f (u n) ∈ V := + begin + simp only [set.image_subset_iff] at h', + specialize h' (n : 𝕜') h1 (set.mem_range_self n), + simp only [set.mem_preimage, linear_map.map_smulₛₗ, map_nat_cast] at h', + rcases h' with ⟨y, hy, h'⟩, + apply_fun (λ y : F, (n : 𝕜')⁻¹ • y) at h', + simp only [hn', inv_smul_smul₀, ne.def, nat.cast_eq_zero, not_false_iff] at h', + rwa ←h', + end, + exact hu' n hn' h'', +end + +/-- If `E` is first countable, then every locally bounded linear map `E →ₛₗ[σ] F` is continuous. -/ +lemma linear_map.continuous_of_locally_bounded [uniform_add_group F] (f : E →ₛₗ[σ] F) + (hf : ∀ (s : set E) (hs : is_vonN_bounded 𝕜 s), is_vonN_bounded 𝕜' (f '' s)) : + continuous f := +(uniform_continuous_of_continuous_at_zero f $ f.continuous_at_zero_of_locally_bounded hf).continuous + +end is_R_or_C From 51696a64f7a0ea235639a1f885b18f57a401b6f7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 2 Oct 2022 20:28:25 +0000 Subject: [PATCH 509/828] chore(data/analysis/*): Fix lint and docs (#16751) Write the missing docstrings and module docs for `data.analysis.filter` and `data.analysis.topology`. Satisfy the `has_nonempty_instance` and `unused_arguments` linters. --- src/data/analysis/filter.lean | 28 ++++++++++++++++-- src/data/analysis/topology.lean | 52 ++++++++++++++++++++++++++++----- 2 files changed, 69 insertions(+), 11 deletions(-) diff --git a/src/data/analysis/filter.lean b/src/data/analysis/filter.lean index 202b4fcd2eadc..85a8a7fa08cf7 100644 --- a/src/data/analysis/filter.lean +++ b/src/data/analysis/filter.lean @@ -2,10 +2,21 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro - -Computational realization of filters (experimental). -/ import order.filter.cofinite + +/-! +# Computational realization of filters (experimental) + +This file provides infrastructure to compute with filters. + +## Main declarations + +* `cfilter`: Realization of a filter base. Note that this is in the generality of filters on + lattices, while `filter` is filters of sets (so corresponding to `cfilter (set α) σ`). +* `filter.realizer`: Realization of a `filter`. `cfilter` that generates the given filter. +-/ + open set filter /-- A `cfilter α σ` is a realization of a filter (base) on `α`, @@ -20,6 +31,13 @@ structure cfilter (α σ : Type*) [partial_order α] := variables {α : Type*} {β : Type*} {σ : Type*} {τ : Type*} +instance [inhabited α] [semilattice_inf α] : inhabited (cfilter α α) := +⟨{ f := id, + pt := default, + inf := (⊓), + inf_le_left := λ _ _, inf_le_left, + inf_le_right := λ _ _, inf_le_right }⟩ + namespace cfilter section variables [partial_order α] (F : cfilter α σ) @@ -62,6 +80,7 @@ structure filter.realizer (f : filter α) := (F : cfilter (set α) σ) (eq : F.to_filter = f) +/-- A `cfilter` realizes the filter it generates. -/ protected def cfilter.to_realizer (F : cfilter (set α) σ) : F.to_filter.realizer := ⟨σ, F, rfl⟩ namespace filter.realizer @@ -69,7 +88,8 @@ namespace filter.realizer theorem mem_sets {f : filter α} (F : f.realizer) {a : set α} : a ∈ f ↔ ∃ b, F.F b ⊆ a := by cases F; subst f; simp --- Used because it has better definitional equalities than the eq.rec proof +/-- Transfer a realizer along an equality of filter. This has better definitional equalities than +the `eq.rec` proof. -/ def of_eq {f g : filter α} (e : f = g) (F : f.realizer) : g.realizer := ⟨F.σ, F.F, F.eq.trans e⟩ @@ -105,6 +125,8 @@ filter_eq $ set.ext $ λ x, @[simp] theorem principal_σ (s : set α) : (realizer.principal s).σ = unit := rfl @[simp] theorem principal_F (s : set α) (u : unit) : (realizer.principal s).F u = s := rfl +instance (s : set α) : inhabited (principal s).realizer := ⟨realizer.principal s⟩ + /-- `unit` is a realizer for the top filter -/ protected def top : (⊤ : filter α).realizer := (realizer.principal _).of_eq principal_univ diff --git a/src/data/analysis/topology.lean b/src/data/analysis/topology.lean index 2a08d83df31d9..36a161a35f71c 100644 --- a/src/data/analysis/topology.lean +++ b/src/data/analysis/topology.lean @@ -2,11 +2,23 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro - -Computational realization of topological spaces (experimental). -/ -import topology.bases import data.analysis.filter +import topology.bases + +/-! +# Computational realization of topological spaces (experimental) + +This file provides infrastructure to compute with topological spaces. + +## Main declarations + +* `ctop`: Realization of a topology basis. +* `ctop.realizer`: Realization of a topological space. `ctop` that generates the given topology. +* `locally_finite.realizer`: Realization of the local finiteness of an indexed family of sets. +* `compact.realizer`: Realization of the compactness of a set. +-/ + open set open filter (hiding realizer) open_locale topological_space @@ -24,6 +36,14 @@ structure ctop (α σ : Type*) := variables {α : Type*} {β : Type*} {σ : Type*} {τ : Type*} +instance : inhabited (ctop α (set α)) := +⟨{ f := id, + top := singleton, + top_mem := mem_singleton, + inter := λ s t _ _, s ∩ t, + inter_mem := λ s t a, id, + inter_sub := λ s t a ha, subset.rfl }⟩ + namespace ctop section variables (F : ctop α σ) @@ -74,9 +94,12 @@ structure ctop.realizer (α) [T : topological_space α] := (eq : F.to_topsp = T) open ctop +/-- A `ctop` realizes the topological space it generates. -/ protected def ctop.to_realizer (F : ctop α σ) : @ctop.realizer _ F.to_topsp := @ctop.realizer.mk _ F.to_topsp σ F rfl +instance (F : ctop α σ) : inhabited (@ctop.realizer _ F.to_topsp) := ⟨F.to_realizer⟩ + namespace ctop.realizer protected theorem is_basis [T : topological_space α] (F : realizer α) : @@ -122,6 +145,7 @@ ext' $ λ a s, ⟨H₂ a s, λ ⟨b, h₁, h₂⟩, mem_nhds_iff.2 ⟨_, h₂, H variable [topological_space α] +/-- The topological space realizer made of the open sets. -/ protected def id : realizer α := ⟨{x:set α // is_open x}, { f := subtype.val, top := λ _, ⟨univ, is_open_univ⟩, @@ -132,6 +156,7 @@ protected def id : realizer α := ⟨{x:set α // is_open x}, ext subtype.property $ λ x s h, let ⟨t, h, o, m⟩ := mem_nhds_iff.1 h in ⟨⟨t, o⟩, m, h⟩⟩ +/-- Replace the representation type of a `ctop` realizer. -/ def of_equiv (F : realizer α) (E : F.σ ≃ τ) : realizer α := ⟨τ, F.F.of_equiv E, ext' (λ a s, F.mem_nhds.trans $ ⟨λ ⟨s, h⟩, ⟨E s, by simpa using h⟩, λ ⟨t, h⟩, ⟨E.symm t, by simpa using h⟩⟩)⟩ @@ -140,6 +165,7 @@ def of_equiv (F : realizer α) (E : F.σ ≃ τ) : realizer α := @[simp] theorem of_equiv_F (F : realizer α) (E : F.σ ≃ τ) (s : τ) : (F.of_equiv E).F s = F.F (E.symm s) := by delta of_equiv; simp +/-- A realizer of the neighborhood of a point. -/ protected def nhds (F : realizer α) (a : α) : (𝓝 a).realizer := ⟨{s : F.σ // a ∈ F.F s}, { f := λ s, F.F s.1, @@ -151,10 +177,8 @@ filter_eq $ set.ext $ λ x, ⟨λ ⟨⟨s, as⟩, h⟩, mem_nhds_iff.2 ⟨_, h, F.is_open _, as⟩, λ h, let ⟨s, h, as⟩ := F.mem_nhds.1 h in ⟨⟨s, h⟩, as⟩⟩⟩ -@[simp] theorem nhds_σ (m : α → β) (F : realizer α) (a : α) : - (F.nhds a).σ = {s : F.σ // a ∈ F.F s} := rfl -@[simp] theorem nhds_F (m : α → β) (F : realizer α) (a : α) (s) : - (F.nhds a).F s = F.F s.1 := rfl +@[simp] lemma nhds_σ (F : realizer α) (a : α) : (F.nhds a).σ = {s : F.σ // a ∈ F.F s} := rfl +@[simp] lemma nhds_F (F : realizer α) (a : α) (s) : (F.nhds a).F s = F.F s.1 := rfl theorem tendsto_nhds_iff {m : β → α} {f : filter β} (F : f.realizer) (R : realizer α) {a : α} : tendsto m f (𝓝 a) ↔ ∀ t, a ∈ R.F t → ∃ s, ∀ x ∈ F.F s, m x ∈ R.F t := @@ -162,6 +186,9 @@ theorem tendsto_nhds_iff {m : β → α} {f : filter β} (F : f.realizer) (R : r end ctop.realizer +/-- A `locally_finite.realizer F f` is a realization that `f` is locally finite, namely it is a +choice of open sets from the basis of `F` such that they intersect only finitely many of the values +of `f`. -/ structure locally_finite.realizer [topological_space α] (F : realizer α) (f : β → set α) := (bas : ∀ a, {s // a ∈ F.F s}) (sets : ∀ x:α, fintype {i | (f i ∩ F.F (bas x)).nonempty}) @@ -183,6 +210,15 @@ theorem locally_finite_iff_exists_realizer [topological_space α] hi.mono (inter_subset_inter_right _ (h₂ x).2)⟩⟩, λ ⟨R⟩, R.to_locally_finite⟩ -def compact.realizer [topological_space α] (R : realizer α) (s : set α) := +instance [topological_space α] [finite β] (F : realizer α) (f : β → set α) : + nonempty (locally_finite.realizer F f) := +(locally_finite_iff_exists_realizer _).1 $ locally_finite_of_finite _ + +/-- A `compact.realizer s` is a realization that `s` is compact, namely it is a +choice of finite open covers for each set family covering `s`. -/ +def compact.realizer [topological_space α] (s : set α) := ∀ {f : filter α} (F : f.realizer) (x : F.σ), f ≠ ⊥ → F.F x ⊆ s → {a // a∈s ∧ 𝓝 a ⊓ f ≠ ⊥} + +instance [topological_space α] : inhabited (compact.realizer (∅ : set α)) := +⟨λ f F x h hF, by { cases h _, rw [←F.eq, eq_bot_iff], exact λ s _, ⟨x, hF.trans s.empty_subset⟩ }⟩ From 2ce380c2043cf93053a198c3f1501466d6e740e6 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Sun, 2 Oct 2022 20:28:26 +0000 Subject: [PATCH 510/828] docs(analysis/locally_convex/bounded): fix hyperlink (#16753) --- src/analysis/locally_convex/bounded.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/locally_convex/bounded.lean b/src/analysis/locally_convex/bounded.lean index cf0e7fa50f7f0..84c2c9eb289b3 100644 --- a/src/analysis/locally_convex/bounded.lean +++ b/src/analysis/locally_convex/bounded.lean @@ -23,7 +23,7 @@ absorbs `s`. ## Main results -* `bornology.is_vonN_bounded_of_topological_space_le`: A coarser topology admits more +* `bornology.is_vonN_bounded.of_topological_space_le`: A coarser topology admits more von Neumann-bounded sets. * `bornology.is_vonN_bounded.image`: A continuous linear image of a bounded set is bounded. From b2d70ba0ee5254e64757c42fdf12c54e27fd8a36 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Sun, 2 Oct 2022 21:39:12 +0000 Subject: [PATCH 511/828] feat(order/liminf_limsup): add limsup/liminf variants of some Limsup/Liminf lemmas (#16759) Also make sure that similar lemmas follow each other in the order Limsup, Liminf, limsup, liminf. --- src/order/liminf_limsup.lean | 63 +++++++++++++++++++++++++----------- 1 file changed, 45 insertions(+), 18 deletions(-) diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index e7b6f47ea369a..d6caaef0e8ee0 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -293,6 +293,16 @@ theorem le_Liminf_of_le {f : filter α} {a} (hf : f.is_cobounded (≥) . is_bounded_default) (h : ∀ᶠ n in f, a ≤ n) : a ≤ f.Liminf := le_cSup hf h +theorem limsup_le_of_le {f : filter β} {u : β → α} {a} + (hf : f.is_cobounded_under (≤) u . is_bounded_default) (h : ∀ᶠ n in f, u n ≤ a) : + f.limsup u ≤ a := +cInf_le hf h + +theorem le_liminf_of_le {f : filter β} {u : β → α} {a} + (hf : f.is_cobounded_under (≥) u . is_bounded_default) (h : ∀ᶠ n in f, a ≤ u n) : + a ≤ f.liminf u := +le_cSup hf h + theorem le_Limsup_of_le {f : filter α} {a} (hf : f.is_bounded (≤) . is_bounded_default) (h : ∀ b, (∀ᶠ n in f, n ≤ b) → a ≤ b) : a ≤ f.Limsup := @@ -303,31 +313,37 @@ theorem Liminf_le_of_le {f : filter α} {a} f.Liminf ≤ a := cSup_le hf h +theorem le_limsup_of_le {f : filter β} {u : β → α} {a} + (hf : f.is_bounded_under (≤) u . is_bounded_default) (h : ∀ b, (∀ᶠ n in f, u n ≤ b) → a ≤ b) : + a ≤ f.limsup u := +le_cInf hf h + +theorem liminf_le_of_le {f : filter β} {u : β → α} {a} + (hf : f.is_bounded_under (≥) u . is_bounded_default) (h : ∀ b, (∀ᶠ n in f, b ≤ u n) → b ≤ a) : + f.liminf u ≤ a := +cSup_le hf h + theorem Liminf_le_Limsup {f : filter α} [ne_bot f] (h₁ : f.is_bounded (≤) . is_bounded_default) (h₂ : f.is_bounded (≥) . is_bounded_default) : f.Liminf ≤ f.Limsup := Liminf_le_of_le h₂ $ assume a₀ ha₀, le_Limsup_of_le h₁ $ assume a₁ ha₁, show a₀ ≤ a₁, from let ⟨b, hb₀, hb₁⟩ := (ha₀.and ha₁).exists in le_trans hb₀ hb₁ -lemma Liminf_le_Liminf {f g : filter α} - (hf : f.is_bounded (≥) . is_bounded_default) (hg : g.is_cobounded (≥) . is_bounded_default) - (h : ∀ a, (∀ᶠ n in f, a ≤ n) → ∀ᶠ n in g, a ≤ n) : f.Liminf ≤ g.Liminf := -cSup_le_cSup hg hf h +lemma liminf_le_limsup {f : filter β} [ne_bot f] {u : β → α} + (h : f.is_bounded_under (≤) u . is_bounded_default) + (h' : f.is_bounded_under (≥) u . is_bounded_default) : + liminf f u ≤ limsup f u := +Liminf_le_Limsup h h' lemma Limsup_le_Limsup {f g : filter α} (hf : f.is_cobounded (≤) . is_bounded_default) (hg : g.is_bounded (≤) . is_bounded_default) (h : ∀ a, (∀ᶠ n in g, n ≤ a) → ∀ᶠ n in f, n ≤ a) : f.Limsup ≤ g.Limsup := cInf_le_cInf hf hg h -lemma Limsup_le_Limsup_of_le {f g : filter α} (h : f ≤ g) - (hf : f.is_cobounded (≤) . is_bounded_default) (hg : g.is_bounded (≤) . is_bounded_default) : - f.Limsup ≤ g.Limsup := -Limsup_le_Limsup hf hg (assume a ha, h ha) - -lemma Liminf_le_Liminf_of_le {f g : filter α} (h : g ≤ f) - (hf : f.is_bounded (≥) . is_bounded_default) (hg : g.is_cobounded (≥) . is_bounded_default) : - f.Liminf ≤ g.Liminf := -Liminf_le_Liminf hf hg (assume a ha, h ha) +lemma Liminf_le_Liminf {f g : filter α} + (hf : f.is_bounded (≥) . is_bounded_default) (hg : g.is_cobounded (≥) . is_bounded_default) + (h : ∀ a, (∀ᶠ n in f, a ≤ n) → ∀ᶠ n in g, a ≤ n) : f.Liminf ≤ g.Liminf := +cSup_le_cSup hg hf h lemma limsup_le_limsup {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} (h : u ≤ᶠ[f] v) @@ -343,6 +359,16 @@ lemma liminf_le_liminf {α : Type*} [conditionally_complete_lattice β] {f : fil f.liminf u ≤ f.liminf v := @limsup_le_limsup βᵒᵈ α _ _ _ _ h hv hu +lemma Limsup_le_Limsup_of_le {f g : filter α} (h : f ≤ g) + (hf : f.is_cobounded (≤) . is_bounded_default) (hg : g.is_bounded (≤) . is_bounded_default) : + f.Limsup ≤ g.Limsup := +Limsup_le_Limsup hf hg (assume a ha, h ha) + +lemma Liminf_le_Liminf_of_le {f g : filter α} (h : g ≤ f) + (hf : f.is_bounded (≥) . is_bounded_default) (hg : g.is_cobounded (≥) . is_bounded_default) : + f.Liminf ≤ g.Liminf := +Liminf_le_Liminf hf hg (assume a ha, h ha) + lemma limsup_le_limsup_of_le {α β} [conditionally_complete_lattice β] {f g : filter α} (h : f ≤ g) {u : α → β} (hf : f.is_cobounded_under (≤) u . is_bounded_default) (hg : g.is_bounded_under (≤) u . is_bounded_default) : @@ -383,11 +409,6 @@ lemma liminf_const {α : Type*} [conditionally_complete_lattice β] {f : filter (b : β) : liminf f (λ x, b) = b := @limsup_const βᵒᵈ α _ f _ b -lemma liminf_le_limsup {f : filter β} [ne_bot f] {u : β → α} - (h : f.is_bounded_under (≤) u . is_bounded_default) - (h' : f.is_bounded_under (≥) u . is_bounded_default) : - liminf f u ≤ limsup f u := -Liminf_le_Limsup h h' end conditionally_complete_lattice @@ -436,6 +457,12 @@ f.basis_sets.Limsup_eq_infi_Sup theorem Liminf_eq_supr_Inf {f : filter α} : f.Liminf = ⨆ s ∈ f, Inf s := @Limsup_eq_infi_Sup αᵒᵈ _ _ +theorem limsup_le_supr {f : filter β} {u : β → α} : f.limsup u ≤ ⨆ n, u n := +limsup_le_of_le (by is_bounded_default) (eventually_of_forall (le_supr u)) + +theorem infi_le_liminf {f : filter β} {u : β → α} : (⨅ n, u n) ≤ f.liminf u := +le_liminf_of_le (by is_bounded_default) (eventually_of_forall (infi_le u)) + /-- In a complete lattice, the limsup of a function is the infimum over sets `s` in the filter of the supremum of the function over `s` -/ theorem limsup_eq_infi_supr {f : filter β} {u : β → α} : f.limsup u = ⨅ s ∈ f, ⨆ a ∈ s, u a := From a4e0a656561ea93ed38048dd5aff3744997dcffa Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Matias=20Heikkil=C3=A4?= Date: Mon, 3 Oct 2022 01:13:30 +0000 Subject: [PATCH 512/828] feat(linear_algebra/matrix/hermitian): complex_hermitian_real_diagonal (#16380) The diagonal elements of a complex Hermitian matrix are real. --- src/linear_algebra/matrix/hermitian.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/linear_algebra/matrix/hermitian.lean b/src/linear_algebra/matrix/hermitian.lean index 15e98d01762ee..fbf01cba1c774 100644 --- a/src/linear_algebra/matrix/hermitian.lean +++ b/src/linear_algebra/matrix/hermitian.lean @@ -142,9 +142,20 @@ variables [ring α] [star_ring α] [ring β] [star_ring β] end ring section is_R_or_C +open is_R_or_C variables [is_R_or_C α] [is_R_or_C β] +/-- The diagonal elements of a complex hermitian matrix are real. -/ +lemma is_hermitian.coe_re_apply_self {A : matrix n n α} (h : A.is_hermitian) (i : n) : + (re (A i i) : α) = A i i := +by rw [←eq_conj_iff_re, ←star_def, ←conj_transpose_apply, h.eq] + +/-- The diagonal elements of a complex hermitian matrix are real. -/ +lemma is_hermitian.coe_re_diag {A : matrix n n α} (h : A.is_hermitian) : + (λ i, (re (A.diag i) : α)) = A.diag := +funext h.coe_re_apply_self + /-- A matrix is hermitian iff the corresponding linear map is self adjoint. -/ lemma is_hermitian_iff_is_symmetric [fintype n] [decidable_eq n] {A : matrix n n α} : is_hermitian A ↔ linear_map.is_symmetric From 3665254a1702986c6f6ccb1edca9f11a69d5af26 Mon Sep 17 00:00:00 2001 From: Apurva Nakade Date: Mon, 3 Oct 2022 02:06:51 +0000 Subject: [PATCH 513/828] feat(analysis/convex/cone/basic): add `has_add`, `add_zero_class`, and `add_comm_semigroup` instances to `convex_cone` (#16213) Adds `has_add`, `add_zero_class`, and `add_comm_semigroup` instance to `convex_cone`s. --- src/analysis/convex/cone/basic.lean | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/src/analysis/convex/cone/basic.lean b/src/analysis/convex/cone/basic.lean index 92aea1081b836..34135c4216f23 100644 --- a/src/analysis/convex/cone/basic.lean +++ b/src/analysis/convex/cone/basic.lean @@ -365,6 +365,32 @@ instance : has_zero (convex_cone 𝕜 E) := ⟨⟨0, λ _ _, by simp, λ _, by s lemma pointed_zero : (0 : convex_cone 𝕜 E).pointed := by rw [pointed, mem_zero] +instance : has_add (convex_cone 𝕜 E) := ⟨ λ K₁ K₂, +{ carrier := {z | ∃ (x y : E), x ∈ K₁ ∧ y ∈ K₂ ∧ x + y = z}, + smul_mem' := + begin + rintro c hc _ ⟨x, y, hx, hy, rfl⟩, + rw smul_add, + use [c • x, c • y, K₁.smul_mem hc hx, K₂.smul_mem hc hy], + end, + add_mem' := + begin + rintro _ ⟨x₁, x₂, hx₁, hx₂, rfl⟩ y ⟨y₁, y₂, hy₁, hy₂, rfl⟩, + use [x₁ + y₁, x₂ + y₂, K₁.add_mem hx₁ hy₁, K₂.add_mem hx₂ hy₂], + abel, + end } ⟩ + +@[simp] lemma mem_add {K₁ K₂ : convex_cone 𝕜 E} {a : E} : + a ∈ K₁ + K₂ ↔ ∃ (x y : E), x ∈ K₁ ∧ y ∈ K₂ ∧ x + y = a := iff.rfl + +instance : add_zero_class (convex_cone 𝕜 E) := +⟨0, has_add.add, λ _, by {ext, simp}, λ _, by {ext, simp}⟩ + +instance : add_comm_semigroup (convex_cone 𝕜 E) := +{ add := has_add.add, + add_assoc := λ _ _ _, set_like.coe_injective $ set.add_comm_semigroup.add_assoc _ _ _, + add_comm := λ _ _, set_like.coe_injective $ set.add_comm_semigroup.add_comm _ _ } + end module end ordered_semiring From 8818d82313df9bccc46d286581f0d2363d8c3cb2 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Mon, 3 Oct 2022 04:09:01 +0000 Subject: [PATCH 514/828] chore(scripts): update nolints.txt (#16773) I am happy to remove some nolints for you! --- scripts/nolints.txt | 18 ------------------ scripts/style-exceptions.txt | 2 -- 2 files changed, 20 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 639382f051ec5..53a65f47e0805 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -139,24 +139,6 @@ apply_nolint tactic.interactive.traverse_constructor unused_arguments apply_nolint tactic.interactive.traverse_field unused_arguments apply_nolint tactic.interactive.with_prefix doc_blame --- data/analysis/filter.lean -apply_nolint cfilter has_nonempty_instance -apply_nolint cfilter.to_realizer doc_blame -apply_nolint filter.realizer has_nonempty_instance -apply_nolint filter.realizer.of_eq doc_blame - --- data/analysis/topology.lean -apply_nolint compact.realizer has_nonempty_instance doc_blame unused_arguments -apply_nolint ctop has_nonempty_instance -apply_nolint ctop.realizer has_nonempty_instance -apply_nolint ctop.realizer.id doc_blame -apply_nolint ctop.realizer.nhds doc_blame -apply_nolint ctop.realizer.nhds_F unused_arguments -apply_nolint ctop.realizer.nhds_σ unused_arguments -apply_nolint ctop.realizer.of_equiv doc_blame -apply_nolint ctop.to_realizer doc_blame -apply_nolint locally_finite.realizer has_nonempty_instance doc_blame - -- data/finset/noncomm_prod.lean apply_nolint add_monoid_hom.pi_ext fintype_finite apply_nolint finset.noncomm_prod_union_of_disjoint to_additive_doc diff --git a/scripts/style-exceptions.txt b/scripts/style-exceptions.txt index cfeaa76b2fa8a..df01fa40c8658 100644 --- a/scripts/style-exceptions.txt +++ b/scripts/style-exceptions.txt @@ -2,8 +2,6 @@ src/control/basic.lean : line 9 : ERR_MOD : Module docstring missing, or too lat src/control/monad/cont.lean : line 13 : ERR_MOD : Module docstring missing, or too late src/control/monad/writer.lean : line 11 : ERR_MOD : Module docstring missing, or too late src/control/traversable/derive.lean : line 10 : ERR_MOD : Module docstring missing, or too late -src/data/analysis/filter.lean : line 9 : ERR_MOD : Module docstring missing, or too late -src/data/analysis/topology.lean : line 10 : ERR_MOD : Module docstring missing, or too late src/data/array/lemmas.lean : line 9 : ERR_MOD : Module docstring missing, or too late src/data/bitvec/basic.lean : line 11 : ERR_MOD : Module docstring missing, or too late src/data/buffer/basic.lean : line 12 : ERR_MOD : Module docstring missing, or too late From 08bb56fb507c72cafb7d310bf03fabc46087cb8d Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Mon, 3 Oct 2022 05:43:49 +0000 Subject: [PATCH 515/828] feat(algebra/module/projective): weaken assumptions in lifting_property (#16750) These `add_comm_group` structures are not necessary, nor is the universe restriction on `M`. --- src/algebra/module/projective.lean | 68 ++++++++++++++++-------------- 1 file changed, 37 insertions(+), 31 deletions(-) diff --git a/src/algebra/module/projective.lean b/src/algebra/module/projective.lean index ca4949996259c..f35d9b97c37ca 100644 --- a/src/algebra/module/projective.lean +++ b/src/algebra/module/projective.lean @@ -67,21 +67,20 @@ universes u v /-- An R-module is projective if it is a direct summand of a free module, or equivalently if maps from the module lift along surjections. There are several other equivalent definitions. -/ -class module.projective (R : Type u) [semiring R] (P : Type (max u v)) [add_comm_monoid P] +class module.projective (R : Type*) [semiring R] (P : Type*) [add_comm_monoid P] [module R P] : Prop := (out : ∃ s : P →ₗ[R] (P →₀ R), function.left_inverse (finsupp.total P P R id) s) namespace module -lemma projective_def {R : Type u} [semiring R] {P : Type (max u v)} [add_comm_monoid P] - [module R P] : projective R P ↔ - (∃ s : P →ₗ[R] (P →₀ R), function.left_inverse (finsupp.total P P R id) s) := -⟨λ h, h.1, λ h, ⟨h⟩⟩ - section semiring -variables {R : Type u} [semiring R] {P : Type (max u v)} [add_comm_monoid P] [module R P] - {M : Type (max u v)} [add_comm_group M] [module R M] {N : Type*} [add_comm_group N] [module R N] +variables {R : Type*} [semiring R] {P : Type*} [add_comm_monoid P] [module R P] + {M : Type*} [add_comm_monoid M] [module R M] {N : Type*} [add_comm_monoid N] [module R N] + +lemma projective_def : projective R P ↔ + (∃ s : P →ₗ[R] (P →₀ R), function.left_inverse (finsupp.total P P R id) s) := +⟨λ h, h.1, λ h, ⟨h⟩⟩ /-- A projective R-module has the property that maps from it lift along surjections. -/ theorem projective_lifting_property [h : projective R P] (f : M →ₗ[R] N) (g : P →ₗ[R] N) @@ -106,9 +105,37 @@ begin simp [φ, finsupp.total_apply, function.surj_inv_eq hf], end +end semiring + +section ring + +variables {R : Type*} [ring R] {P : Type*} [add_comm_group P] [module R P] + +/-- Free modules are projective. -/ +theorem projective_of_basis {ι : Type*} (b : basis ι R P) : projective R P := +begin + -- need P →ₗ (P →₀ R) for definition of projective. + -- get it from `ι → (P →₀ R)` coming from `b`. + use b.constr ℕ (λ i, finsupp.single (b i) (1 : R)), + intro m, + simp only [b.constr_apply, mul_one, id.def, finsupp.smul_single', finsupp.total_single, + linear_map.map_finsupp_sum], + exact b.total_repr m, +end + +@[priority 100] +instance projective_of_free [module.free R P] : module.projective R P := +projective_of_basis $ module.free.choose_basis R P + +end ring + +--This is in a different section because special universe restrictions are required. +section of_lifting_property + /-- A module which satisfies the universal property is projective. Note that the universe variables in `huniv` are somewhat restricted. -/ theorem projective_of_lifting_property' + {R : Type u} [semiring R] {P : Type (max u v)} [add_comm_monoid P] [module R P] -- If for all surjections of `R`-modules `M →ₗ N`, all maps `P →ₗ N` lift to `P →ₗ M`, (huniv : ∀ {M : Type (max v u)} {N : Type (max u v)} [add_comm_monoid M] [add_comm_monoid N], by exactI @@ -131,15 +158,10 @@ begin simp }, end -end semiring - -section ring - -variables {R : Type u} [ring R] {P : Type (max u v)} [add_comm_group P] [module R P] - /-- A variant of `of_lifting_property'` when we're working over a `[ring R]`, which only requires quantifying over modules with an `add_comm_group` instance. -/ theorem projective_of_lifting_property + {R : Type u} [ring R] {P : Type (max u v)} [add_comm_group P] [module R P] -- If for all surjections of `R`-modules `M →ₗ N`, all maps `P →ₗ N` lift to `P →ₗ M`, (huniv : ∀ {M : Type (max v u)} {N : Type (max u v)} [add_comm_group M] [add_comm_group N], by exactI @@ -165,22 +187,6 @@ begin simp }, end -/-- Free modules are projective. -/ -theorem projective_of_basis {ι : Type*} (b : basis ι R P) : projective R P := -begin - -- need P →ₗ (P →₀ R) for definition of projective. - -- get it from `ι → (P →₀ R)` coming from `b`. - use b.constr ℕ (λ i, finsupp.single (b i) (1 : R)), - intro m, - simp only [b.constr_apply, mul_one, id.def, finsupp.smul_single', finsupp.total_single, - linear_map.map_finsupp_sum], - exact b.total_repr m, -end - -@[priority 100] -instance projective_of_free [module.free R P] : module.projective R P := -projective_of_basis $ module.free.choose_basis R P - -end ring +end of_lifting_property end module From ebe6bda59dbe2f09a49e9e77caefbad7f19fc725 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Mon, 3 Oct 2022 06:36:39 +0000 Subject: [PATCH 516/828] feat(analysis/normed/group/add_circle): define the additive circle (#16201) Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> --- src/algebra/order/archimedean.lean | 3 +- src/algebra/order/to_interval_mod.lean | 10 +- src/analysis/normed/group/add_circle.lean | 109 ++++++++++++++ src/analysis/normed/group/quotient.lean | 4 + .../trigonometric/angle.lean | 10 +- src/group_theory/subgroup/basic.lean | 18 +++ src/topology/instances/add_circle.lean | 134 ++++++++++++++++++ 7 files changed, 278 insertions(+), 10 deletions(-) create mode 100644 src/analysis/normed/group/add_circle.lean create mode 100644 src/topology/instances/add_circle.lean diff --git a/src/algebra/order/archimedean.lean b/src/algebra/order/archimedean.lean index b04d316a1169d..c618b3a236034 100644 --- a/src/algebra/order/archimedean.lean +++ b/src/algebra/order/archimedean.lean @@ -330,7 +330,8 @@ floor_ring.of_floor α (λ a, classical.some (exists_floor a)) (λ z a, (classical.some_spec (exists_floor a) z).symm) /-- A linear ordered field that is a floor ring is archimedean. -/ -lemma floor_ring.archimedean (α) [linear_ordered_field α] [floor_ring α] : archimedean α := +@[priority 100] -- see Note [lower instance priority] +instance floor_ring.archimedean (α) [linear_ordered_field α] [floor_ring α] : archimedean α := begin rw archimedean_iff_int_le, exact λ x, ⟨⌈x⌉, int.le_ceil x⟩ diff --git a/src/algebra/order/to_interval_mod.lean b/src/algebra/order/to_interval_mod.lean index 064c523c4745d..1bc1ecf2e20bc 100644 --- a/src/algebra/order/to_interval_mod.lean +++ b/src/algebra/order/to_interval_mod.lean @@ -63,6 +63,10 @@ lemma to_Ico_mod_mem_Ico (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ico_mod a hb x ∈ set.Ico a (a + b) := add_to_Ico_div_zsmul_mem_Ico a hb x +lemma to_Ico_mod_mem_Ico' {b : α} (hb : 0 < b) (x : α) : + to_Ico_mod 0 hb x ∈ set.Ico 0 b := +by { convert to_Ico_mod_mem_Ico 0 hb x, exact (zero_add b).symm, } + lemma to_Ioc_mod_mem_Ioc (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ioc_mod a hb x ∈ set.Ioc a (a + b) := add_to_Ioc_div_zsmul_mem_Ioc a hb x @@ -426,8 +430,6 @@ section linear_ordered_field variables {α : Type*} [linear_ordered_field α] [floor_ring α] -local attribute [instance] floor_ring.archimedean - lemma to_Ico_div_eq_neg_floor (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ico_div a hb x = -⌊(x - a) / b⌋ := begin @@ -461,6 +463,10 @@ begin ring end +lemma to_Ico_mod_eq_fract_mul {b : α} (hb : 0 < b) (x : α) : + to_Ico_mod 0 hb x = int.fract (x / b) * b := +by simp [to_Ico_mod_eq_add_fract_mul] + lemma to_Ioc_mod_eq_sub_fract_mul (a : α) {b : α} (hb : 0 < b) (x : α) : to_Ioc_mod a hb x = a + b - int.fract ((a + b - x) / b) * b := begin diff --git a/src/analysis/normed/group/add_circle.lean b/src/analysis/normed/group/add_circle.lean new file mode 100644 index 0000000000000..9dcfec8067217 --- /dev/null +++ b/src/analysis/normed/group/add_circle.lean @@ -0,0 +1,109 @@ +/- +Copyright (c) 2022 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import analysis.normed.group.quotient +import topology.instances.add_circle + +/-! +# The additive circle as a normed group + +We define the normed group structure on `add_circle p`, for `p : ℝ`. For example if `p = 1` then: +`∥(x : add_circle 1)∥ = |x - round x|` for any `x : ℝ` (see `unit_add_circle.norm_eq`). + +## Main definitions: + + * `add_circle.norm_eq`: a characterisation of the norm on `add_circle p` + +## TODO + + * The fact `inner_product_geometry.angle (real.cos θ) (real.sin θ) = ∥(θ : real.angle)∥` + +-/ + +noncomputable theory + +open set int (hiding mem_zmultiples_iff) add_subgroup + +namespace add_circle + +variables (p : ℝ) + +instance : normed_add_comm_group (add_circle p) := add_subgroup.normed_add_comm_group_quotient _ + +@[simp] lemma norm_coe_mul (x : ℝ) (t : ℝ) : + ∥(↑(t * x) : add_circle (t * p))∥ = |t| * ∥(x : add_circle p)∥ := +begin + have aux : ∀ {a b c : ℝ}, a ∈ zmultiples b → c * a ∈ zmultiples (c * b) := λ a b c h, by + { simp only [mem_zmultiples_iff] at ⊢ h, + obtain ⟨n, rfl⟩ := h, + exact ⟨n, (mul_smul_comm n c b).symm⟩, }, + rcases eq_or_ne t 0 with rfl | ht, { simp, }, + have ht' : |t| ≠ 0 := (not_congr abs_eq_zero).mpr ht, + simp only [quotient_norm_eq, real.norm_eq_abs], + conv_rhs { rw [← smul_eq_mul, ← real.Inf_smul_of_nonneg (abs_nonneg t)], }, + simp only [quotient_add_group.mk'_apply, quotient_add_group.eq_iff_sub_mem], + congr' 1, + ext z, + rw mem_smul_set_iff_inv_smul_mem₀ ht', + show (∃ y, y - t * x ∈ zmultiples (t * p) ∧ |y| = z) ↔ ∃w, w - x ∈ zmultiples p ∧ |w| = |t|⁻¹ * z, + split, + { rintros ⟨y, hy, rfl⟩, + refine ⟨t⁻¹ * y, _, by rw [abs_mul, abs_inv]⟩, + rw [← inv_mul_cancel_left₀ ht x, ← inv_mul_cancel_left₀ ht p, ← mul_sub], + exact aux hy, }, + { rintros ⟨w, hw, hw'⟩, + refine ⟨t * w, _, by rw [← (eq_inv_mul_iff_mul_eq₀ ht').mp hw', abs_mul]⟩, + rw ← mul_sub, + exact aux hw, }, +end + +@[simp] lemma norm_eq_of_zero {x : ℝ} : ∥(x : add_circle (0 : ℝ))∥ = |x| := +begin + suffices : {y : ℝ | (y : add_circle (0 : ℝ)) = (x : add_circle (0 : ℝ)) } = { x }, + { rw [quotient_norm_eq, this, image_singleton, real.norm_eq_abs, cInf_singleton], }, + ext y, + simp [quotient_add_group.eq_iff_sub_mem, mem_zmultiples_iff, sub_eq_zero], +end + +lemma norm_eq {x : ℝ} : ∥(x : add_circle p)∥ = |x - round (p⁻¹ * x) * p| := +begin + suffices : ∀ (x : ℝ), ∥(x : add_circle (1 : ℝ))∥ = |x - round x|, + { rcases eq_or_ne p 0 with rfl | hp, { simp, }, + intros, + have hx := norm_coe_mul p x p⁻¹, + rw [abs_inv, eq_inv_mul_iff_mul_eq₀ ((not_congr abs_eq_zero).mpr hp)] at hx, + rw [← hx, inv_mul_cancel hp, this, ← abs_mul, mul_sub, mul_inv_cancel_left₀ hp, mul_comm p], }, + clear x p, + intros, + rw [quotient_norm_eq, abs_sub_round_eq_min], + have h₁ : bdd_below (abs '' {m : ℝ | (m : add_circle (1 : ℝ)) = x}) := + ⟨0, by simp [mem_lower_bounds]⟩, + have h₂ : (abs '' {m : ℝ | (m : add_circle (1 : ℝ)) = x}).nonempty := ⟨|x|, ⟨x, rfl, rfl⟩⟩, + apply le_antisymm, + { simp only [le_min_iff, real.norm_eq_abs, cInf_le_iff h₁ h₂], + intros b h, + refine ⟨mem_lower_bounds.1 h _ ⟨fract x, _, abs_fract⟩, + mem_lower_bounds.1 h _ ⟨fract x - 1, _, by rw [abs_sub_comm, abs_one_sub_fract]⟩⟩, + { simp only [mem_set_of_eq, fract, sub_eq_self, quotient_add_group.coe_sub, + quotient_add_group.eq_zero_iff, int_cast_mem_zmultiples_one], }, + { simp only [mem_set_of_eq, fract, sub_eq_self, quotient_add_group.coe_sub, + quotient_add_group.eq_zero_iff, int_cast_mem_zmultiples_one, sub_sub, + (by norm_cast : (⌊x⌋ : ℝ) + 1 = (↑(⌊x⌋ + 1) : ℝ))], }, }, + { simp only [quotient_add_group.mk'_apply, real.norm_eq_abs, le_cInf_iff h₁ h₂], + rintros b' ⟨b, hb, rfl⟩, + simp only [mem_set_of_eq, quotient_add_group.eq_iff_sub_mem, mem_zmultiples_iff, + smul_one_eq_coe] at hb, + obtain ⟨z, hz⟩ := hb, + rw [(by { rw hz, abel, } : x = b - z), fract_sub_int, ← abs_sub_round_eq_min], + exact abs_sub_round_le_abs_self _, }, +end + +end add_circle + +namespace unit_add_circle + +lemma norm_eq {x : ℝ} : ∥(x : unit_add_circle)∥ = |x - round x| := by simp [add_circle.norm_eq] + +end unit_add_circle diff --git a/src/analysis/normed/group/quotient.lean b/src/analysis/normed/group/quotient.lean index 2d01162f39f48..12e5cd58965d7 100644 --- a/src/analysis/normed/group/quotient.lean +++ b/src/analysis/normed/group/quotient.lean @@ -100,6 +100,10 @@ noncomputable instance norm_on_quotient (S : add_subgroup M) : has_norm (M ⧸ S) := { norm := λ x, Inf (norm '' {m | mk' S m = x}) } +lemma add_subgroup.quotient_norm_eq {S : add_subgroup M} (x : M ⧸ S) : + ∥x∥ = Inf (norm '' {m : M | (m : M ⧸ S) = x}) := +rfl + lemma image_norm_nonempty {S : add_subgroup M} : ∀ x : M ⧸ S, (norm '' {m | mk' S m = x}).nonempty := begin diff --git a/src/analysis/special_functions/trigonometric/angle.lean b/src/analysis/special_functions/trigonometric/angle.lean index 67809c5979110..88a2a9966c9c2 100644 --- a/src/analysis/special_functions/trigonometric/angle.lean +++ b/src/analysis/special_functions/trigonometric/angle.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Calle Sönne -/ import analysis.special_functions.trigonometric.basic +import analysis.normed.group.add_circle import algebra.char_zero.quotient import algebra.order.to_interval_mod import topology.instances.sign @@ -22,16 +23,11 @@ noncomputable theory namespace real /-- The type of angles -/ -@[derive [add_comm_group, topological_space, topological_add_group]] -def angle : Type := -ℝ ⧸ (add_subgroup.zmultiples (2 * π)) +@[derive [normed_add_comm_group, inhabited, has_coe_t ℝ]] +def angle : Type := add_circle (2 * π) namespace angle -instance : inhabited angle := ⟨0⟩ - -instance : has_coe ℝ angle := ⟨quotient_add_group.mk' _⟩ - @[continuity] lemma continuous_coe : continuous (coe : ℝ → angle) := continuous_quotient_mk diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index d7af47900ef5f..13597082341f8 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2652,6 +2652,9 @@ lemma mem_zpowers_iff {g h : G} : h ∈ zpowers g ↔ ∃ (k : ℤ), g ^ k = h := iff.rfl +@[simp] lemma zpow_mem_zpowers (g : G) (k : ℤ) : g^k ∈ zpowers g := +mem_zpowers_iff.mpr ⟨k, rfl⟩ + @[simp] lemma forall_zpowers {x : G} {p : zpowers x → Prop} : (∀ g, p g) ↔ ∀ m : ℤ, p ⟨x ^ m, m, rfl⟩ := set.forall_subtype_range_iff @@ -2684,11 +2687,26 @@ attribute [to_additive add_subgroup.zmultiples_eq_closure] subgroup.zpowers_eq_c attribute [to_additive add_subgroup.range_zmultiples_hom] subgroup.range_zpowers_hom attribute [to_additive add_subgroup.zmultiples_subset] subgroup.zpowers_subset attribute [to_additive add_subgroup.mem_zmultiples_iff] subgroup.mem_zpowers_iff +attribute [to_additive add_subgroup.zsmul_mem_zmultiples] subgroup.zpow_mem_zpowers attribute [to_additive add_subgroup.forall_zmultiples] subgroup.forall_zpowers attribute [to_additive add_subgroup.forall_mem_zmultiples] subgroup.forall_mem_zpowers attribute [to_additive add_subgroup.exists_zmultiples] subgroup.exists_zpowers attribute [to_additive add_subgroup.exists_mem_zmultiples] subgroup.exists_mem_zpowers +section ring + +variables {R : Type*} [ring R] (r : R) (k : ℤ) + +@[simp] lemma int_cast_mul_mem_zmultiples : + ↑(k : ℤ) * r ∈ zmultiples r := +by simpa only [← zsmul_eq_mul] using zsmul_mem_zmultiples r k + +@[simp] lemma int_cast_mem_zmultiples_one : + ↑(k : ℤ) ∈ zmultiples (1 : R) := +mem_zmultiples_iff.mp ⟨k, by simp⟩ + +end ring + end add_subgroup lemma int.mem_zmultiples_iff {a b : ℤ} : diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean new file mode 100644 index 0000000000000..cc7d4a0d9c703 --- /dev/null +++ b/src/topology/instances/add_circle.lean @@ -0,0 +1,134 @@ +/- +Copyright (c) 2022 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import algebra.order.floor +import algebra.order.to_interval_mod +import topology.instances.real + +/-! +# The additive circle + +We define the additive circle `add_circle p` as the quotient `𝕜 ⧸ (ℤ ∙ p)` for some period `p : 𝕜`. + +See also `circle` and `real.angle`. For the normed group structure on `add_circle`, see +`add_circle.normed_add_comm_group` in a later file. + +## Main definitions: + + * `add_circle`: the additive circle `𝕜 ⧸ (ℤ ∙ p)` for some period `p : 𝕜` + * `unit_add_circle`: the special case `ℝ ⧸ ℤ` + * `add_circle.equiv_add_circle`: the rescaling equivalence `add_circle p ≃+ add_circle q` + * `add_circle.equiv_Ico`: the natural equivalence `add_circle p ≃ Ico 0 p` + +## Implementation notes: + +Although the most important case is `𝕜 = ℝ` we wish to support other types of scalars, such as +the rational circle `add_circle (1 : ℚ)`, and so we set things up more generally. + +## TODO + + * Link with periodicity + * Measure space structure + * Lie group structure + * Exponential equivalence to `circle` + +-/ + +noncomputable theory + +open set int (hiding mem_zmultiples_iff) add_subgroup + +variables {𝕜 : Type*} + +/-- The "additive circle": `𝕜 ⧸ (ℤ ∙ p)`. See also `circle` and `real.angle`. -/ +@[derive [add_comm_group, topological_space, topological_add_group, inhabited, has_coe_t 𝕜], + nolint unused_arguments] +def add_circle [linear_ordered_add_comm_group 𝕜] [topological_space 𝕜] [order_topology 𝕜] (p : 𝕜) := +𝕜 ⧸ zmultiples p + +namespace add_circle + +section linear_ordered_field + +variables [linear_ordered_field 𝕜] [topological_space 𝕜] [order_topology 𝕜] (p q : 𝕜) + +@[continuity, nolint unused_arguments] protected lemma continuous_mk' : + continuous (quotient_add_group.mk' (zmultiples p) : 𝕜 → add_circle p) := +continuous_coinduced_rng + +/-- An auxiliary definition used only for constructing `add_circle.equiv_add_circle`. -/ +private def equiv_add_circle_aux (hp : p ≠ 0) : add_circle p →+ add_circle q := +quotient_add_group.lift _ + ((quotient_add_group.mk' (zmultiples q)).comp $ add_monoid_hom.mul_right (p⁻¹ * q)) + (λ x h, by obtain ⟨z, rfl⟩ := mem_zmultiples_iff.1 h; simp [hp, mul_assoc (z : 𝕜), ← mul_assoc p]) + +/-- The rescaling equivalence between additive circles with different periods. -/ +def equiv_add_circle (hp : p ≠ 0) (hq : q ≠ 0) : add_circle p ≃+ add_circle q := +{ to_fun := equiv_add_circle_aux p q hp, + inv_fun := equiv_add_circle_aux q p hq, + left_inv := by { rintros ⟨x⟩, show quotient_add_group.mk _ = _, congr, field_simp [hp, hq], }, + right_inv := by { rintros ⟨x⟩, show quotient_add_group.mk _ = _, congr, field_simp [hp, hq], }, + .. equiv_add_circle_aux p q hp } + +@[simp] lemma equiv_add_circle_apply_mk (hp : p ≠ 0) (hq : q ≠ 0) (x : 𝕜) : + equiv_add_circle p q hp hq (x : 𝕜) = (x * (p⁻¹ * q) : 𝕜) := +rfl + +@[simp] lemma equiv_add_circle_symm_apply_mk (hp : p ≠ 0) (hq : q ≠ 0) (x : 𝕜) : + (equiv_add_circle p q hp hq).symm (x : 𝕜) = (x * (q⁻¹ * p) : 𝕜) := +rfl + +variables [floor_ring 𝕜] + +/-- The natural equivalence between `add_circle p` and the half-open interval `[0, p)`. -/ +def equiv_Ico (hp : 0 < p) : add_circle p ≃ Ico 0 p := +{ inv_fun := quotient_add_group.mk' _ ∘ coe, + to_fun := λ x, ⟨(to_Ico_mod_periodic 0 hp).lift x, quot.induction_on x $ to_Ico_mod_mem_Ico' hp⟩, + right_inv := by { rintros ⟨x, hx⟩, ext, simp [to_Ico_mod_eq_self, hx.1, hx.2], }, + left_inv := + begin + rintros ⟨x⟩, + change quotient_add_group.mk (to_Ico_mod 0 hp x) = quotient_add_group.mk x, + rw [quotient_add_group.eq', neg_add_eq_sub, self_sub_to_Ico_mod, zsmul_eq_mul], + apply int_cast_mul_mem_zmultiples, + end } + +@[simp] lemma coe_equiv_Ico_mk_apply (hp : 0 < p) (x : 𝕜) : + (equiv_Ico p hp $ quotient_add_group.mk x : 𝕜) = fract (x / p) * p := +to_Ico_mod_eq_fract_mul hp x + +@[continuity] lemma continuous_equiv_Ico_symm (hp : 0 < p) : continuous (equiv_Ico p hp).symm := +continuous_coinduced_rng.comp continuous_induced_dom + +/-- The image of the closed interval `[0, p]` under the quotient map `𝕜 → add_circle p` is the +entire space. -/ +@[simp] lemma coe_image_Icc_eq (hp : 0 < p) : + (coe : 𝕜 → add_circle p) '' (Icc 0 p) = univ := +begin + refine eq_univ_iff_forall.mpr (λ x, _), + let y := equiv_Ico p hp x, + exact ⟨y, ⟨y.2.1, y.2.2.le⟩, (equiv_Ico p hp).symm_apply_apply x⟩, +end + +end linear_ordered_field + +variables (p : ℝ) + +lemma compact_space (hp : 0 < p) : compact_space $ add_circle p := +begin + rw [← is_compact_univ_iff, ← coe_image_Icc_eq p hp], + exact is_compact_Icc.image (add_circle.continuous_mk' p), +end + +end add_circle + +/-- The unit circle `ℝ ⧸ ℤ`. -/ +abbreviation unit_add_circle := add_circle (1 : ℝ) + +namespace unit_add_circle + +instance : compact_space unit_add_circle := add_circle.compact_space _ zero_lt_one + +end unit_add_circle From d2461773d08d81fb86cd06f44ac6f2cdd65e8da9 Mon Sep 17 00:00:00 2001 From: mcdoll Date: Mon, 3 Oct 2022 06:36:40 +0000 Subject: [PATCH 517/828] chore(analysis/locally_convex/with_seminorms): fix namespaces for dot-notation (#16771) Some lemmas where in the wrong namespaces. --- src/analysis/locally_convex/weak_dual.lean | 2 +- src/analysis/locally_convex/with_seminorms.lean | 6 +++--- src/analysis/schwartz_space.lean | 13 ++++++------- 3 files changed, 10 insertions(+), 11 deletions(-) diff --git a/src/analysis/locally_convex/weak_dual.lean b/src/analysis/locally_convex/weak_dual.lean index 6eb8b71ed36e2..257662dfbd58f 100644 --- a/src/analysis/locally_convex/weak_dual.lean +++ b/src/analysis/locally_convex/weak_dual.lean @@ -135,6 +135,6 @@ variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [add_comm_group variables [nonempty ι] [normed_space ℝ 𝕜] [module ℝ E] [is_scalar_tower ℝ 𝕜 E] instance {B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜} : locally_convex_space ℝ (weak_bilin B) := -seminorm_family.to_locally_convex_space (B.weak_bilin_with_seminorms) +(B.weak_bilin_with_seminorms).to_locally_convex_space end locally_convex diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 4fc6e1004b13e..6696bcebe14a3 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -271,7 +271,7 @@ variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [nonempty ι] structure with_seminorms (p : seminorm_family 𝕜 E ι) [t : topological_space E] : Prop := (topology_eq_with_seminorms : t = p.module_filter_basis.topology) -lemma seminorm_family.with_seminorms_eq {p : seminorm_family 𝕜 E ι} [t : topological_space E] +lemma with_seminorms.with_seminorms_eq {p : seminorm_family 𝕜 E ι} [t : topological_space E] (hp : with_seminorms p) : t = p.module_filter_basis.topology := hp.1 variables [topological_space E] @@ -493,7 +493,7 @@ variables [nonempty ι] [normed_field 𝕜] [normed_space ℝ 𝕜] [add_comm_group E] [module 𝕜 E] [module ℝ E] [is_scalar_tower ℝ 𝕜 E] [topological_space E] [topological_add_group E] -lemma seminorm_family.to_locally_convex_space {p : seminorm_family 𝕜 E ι} (hp : with_seminorms p) : +lemma with_seminorms.to_locally_convex_space {p : seminorm_family 𝕜 E ι} (hp : with_seminorms p) : locally_convex_space ℝ E := begin apply of_basis_zero ℝ E id (λ s, s ∈ p.basis_sets), @@ -516,7 +516,7 @@ variables (𝕜) [normed_field 𝕜] [normed_space ℝ 𝕜] [seminormed_add_com slightly weaker instance version. -/ lemma normed_space.to_locally_convex_space' [normed_space 𝕜 E] [module ℝ E] [is_scalar_tower ℝ 𝕜 E] : locally_convex_space ℝ E := -seminorm_family.to_locally_convex_space (norm_with_seminorms 𝕜 E) +(norm_with_seminorms 𝕜 E).to_locally_convex_space /-- See `normed_space.to_locally_convex_space'` for a slightly stronger version which is not an instance. -/ diff --git a/src/analysis/schwartz_space.lean b/src/analysis/schwartz_space.lean index 853527ca6dd2f..69e1458a09aae 100644 --- a/src/analysis/schwartz_space.lean +++ b/src/analysis/schwartz_space.lean @@ -22,7 +22,7 @@ natural numbers `k` and `n` we have uniform bounds `∥x∥^k * ∥iterated_fder This approach completely avoids using partial derivatives as well as polynomials. We construct the topology on the Schwartz space by a family of seminorms, which are the best constants in the above estimates, which is by abstract theory from -`seminorm_family.module_filter_basis` and `seminorm_family.to_locally_convex_space` turns the +`seminorm_family.module_filter_basis` and `with_seminorms.to_locally_convex_space` turns the Schwartz space into a locally convex topological vector space. ## Main definitions @@ -375,22 +375,21 @@ variables {𝕜 E F} instance : has_continuous_smul 𝕜 𝓢(E, F) := begin - rw seminorm_family.with_seminorms_eq (schwartz_with_seminorms 𝕜 E F), + rw (schwartz_with_seminorms 𝕜 E F).with_seminorms_eq, exact (schwartz_seminorm_family 𝕜 E F).module_filter_basis.has_continuous_smul, end instance : topological_add_group 𝓢(E, F) := -(schwartz_seminorm_family ℝ E F).module_filter_basis.to_add_group_filter_basis - .is_topological_add_group +(schwartz_seminorm_family ℝ E F).add_group_filter_basis.is_topological_add_group instance : uniform_space 𝓢(E, F) := -(schwartz_seminorm_family ℝ E F).module_filter_basis.to_add_group_filter_basis.uniform_space +(schwartz_seminorm_family ℝ E F).add_group_filter_basis.uniform_space instance : uniform_add_group 𝓢(E, F) := -(schwartz_seminorm_family ℝ E F).module_filter_basis.to_add_group_filter_basis.uniform_add_group +(schwartz_seminorm_family ℝ E F).add_group_filter_basis.uniform_add_group instance : locally_convex_space ℝ 𝓢(E, F) := -seminorm_family.to_locally_convex_space (schwartz_with_seminorms ℝ E F) +(schwartz_with_seminorms ℝ E F).to_locally_convex_space end topology From 3cf2547c8e46b61b0e5754c5fdc005de9f3195de Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 3 Oct 2022 08:59:05 +0000 Subject: [PATCH 518/828] feat(data/fintype/basic): Existence of a bijection with specified behaviour on a subset (#16479) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a function `f : α → β` which is injective on a set `s` in `α`, and a set `t` in `β` which has the same finite cardinality as the type `α` and which contains the image `f '' s`, extend/modify to a bijection between `α` and `t` which agrees on `s` with the original `f`. --- src/data/fintype/basic.lean | 43 +++++++++++++++++++++++++++++++++++++ src/data/set/function.lean | 5 +++++ 2 files changed, 48 insertions(+) diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 3530f8ed67ccd..ede6df4440323 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -1404,6 +1404,49 @@ lemma finset.univ_map_embedding {α : Type*} [fintype α] (e : α ↪ α) : univ.map e = univ := by rw [←e.equiv_of_fintype_self_embedding_to_embedding, univ_map_equiv_to_embedding] +/-- Any injection from a finset `s` in a fintype `α` to a finset `t` of the same cardinality as `α` +can be extended to a bijection between `α` and `t`. -/ +lemma finset.exists_equiv_extend_of_card_eq [fintype α] {t : finset β} + (hαt : fintype.card α = t.card) {s : finset α} {f : α → β} (hfst : s.image f ⊆ t) + (hfs : set.inj_on f s) : + ∃ g : α ≃ t, ∀ i ∈ s, (g i : β) = f i := +begin + classical, + induction s using finset.induction with a s has H generalizing f, + { obtain ⟨e⟩ : nonempty (α ≃ ↥t) := by rwa [← fintype.card_eq, fintype.card_coe], + use e, + simp }, + have hfst' : finset.image f s ⊆ t := (finset.image_mono _ (s.subset_insert a)).trans hfst, + have hfs' : set.inj_on f s := hfs.mono (s.subset_insert a), + obtain ⟨g', hg'⟩ := H hfst' hfs', + have hfat : f a ∈ t := hfst (mem_image_of_mem _ (s.mem_insert_self a)), + use g'.trans (equiv.swap (⟨f a, hfat⟩ : t) (g' a)), + simp_rw mem_insert, + rintro i (rfl | hi), + { simp }, + rw [equiv.trans_apply, equiv.swap_apply_of_ne_of_ne, hg' _ hi], + { exact ne_of_apply_ne subtype.val (ne_of_eq_of_ne (hg' _ hi) $ + hfs.ne (subset_insert _ _ hi) (mem_insert_self _ _) $ ne_of_mem_of_not_mem hi has) }, + { exact g'.injective.ne (ne_of_mem_of_not_mem hi has) }, +end + +/-- Any injection from a set `s` in a fintype `α` to a finset `t` of the same cardinality as `α` +can be extended to a bijection between `α` and `t`. -/ +lemma set.maps_to.exists_equiv_extend_of_card_eq [fintype α] {t : finset β} + (hαt : fintype.card α = t.card) {s : set α} {f : α → β} (hfst : s.maps_to f t) + (hfs : set.inj_on f s) : + ∃ g : α ≃ t, ∀ i ∈ s, (g i : β) = f i := +begin + classical, + let s' : finset α := s.to_finset, + have hfst' : s'.image f ⊆ t := by simpa [← finset.coe_subset] using hfst, + have hfs' : set.inj_on f s' := by simpa using hfs, + obtain ⟨g, hg⟩ := finset.exists_equiv_extend_of_card_eq hαt hfst' hfs', + refine ⟨g, λ i hi, _⟩, + apply hg, + simpa using hi, +end + namespace fintype /-- Given `fintype α`, `finset_equiv_set` is the equiv between `finset α` and `set α`. (All diff --git a/src/data/set/function.lean b/src/data/set/function.lean index 938aa1b4ce16b..f1cb1341f2ac1 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -413,6 +413,11 @@ theorem inj_on.eq_iff {x y} (h : inj_on f s) (hx : x ∈ s) (hy : y ∈ s) : f x = f y ↔ x = y := ⟨h hx hy, λ h, h ▸ rfl⟩ +lemma inj_on.ne_iff {x y} (h : inj_on f s) (hx : x ∈ s) (hy : y ∈ s) : f x ≠ f y ↔ x ≠ y := +(h.eq_iff hx hy).not + +alias inj_on.ne_iff ↔ _ inj_on.ne + theorem inj_on.congr (h₁ : inj_on f₁ s) (h : eq_on f₁ f₂ s) : inj_on f₂ s := λ x hx y hy, h hx ▸ h hy ▸ h₁ hx hy From 35bc69bc940ff78236b4fc602ec9c2fd9dd3f48f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 3 Oct 2022 08:59:07 +0000 Subject: [PATCH 519/828] feat(group_theory/subgroup): add several trivial lemmas (#16633) * add `subgroup.top_to_submonoid` and `subgroup.bot_to_submonoid`; * add `free_monoid.mrange_lift`. --- src/group_theory/subgroup/basic.lean | 9 +++++---- src/group_theory/submonoid/membership.lean | 8 ++++++-- 2 files changed, 11 insertions(+), 6 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 13597082341f8..a47d863fe907a 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -648,11 +648,12 @@ instance : inhabited (subgroup G) := ⟨⊥⟩ @[to_additive] instance : unique (⊥ : subgroup G) := ⟨⟨1⟩, λ g, subtype.ext g.2⟩ +@[simp, to_additive] lemma top_to_submonoid : (⊤ : subgroup G).to_submonoid = ⊤ := rfl + +@[simp, to_additive] lemma bot_to_submonoid : (⊥ : subgroup G).to_submonoid = ⊥ := rfl + @[to_additive] lemma eq_bot_iff_forall : H = ⊥ ↔ ∀ x ∈ H, x = (1 : G) := -begin - rw set_like.ext'_iff, - simp only [coe_bot, set.eq_singleton_iff_unique_mem, set_like.mem_coe, H.one_mem, true_and], -end +to_submonoid_injective.eq_iff.symm.trans $ submonoid.eq_bot_iff_forall _ @[to_additive] lemma eq_bot_of_subsingleton [subsingleton H] : H = ⊥ := begin diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index b90924f5ee5b8..fc77d6c08abb9 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -280,10 +280,14 @@ mem_closure_singleton.2 ⟨1, pow_one y⟩ lemma closure_singleton_one : closure ({1} : set M) = ⊥ := by simp [eq_bot_iff_forall, mem_closure_singleton] +@[to_additive] lemma _root_.free_monoid.mrange_lift {α} (f : α → M) : + (free_monoid.lift f).mrange = closure (set.range f) := +by rw [mrange_eq_map, ← free_monoid.closure_range_of, map_mclosure, ← set.range_comp, + free_monoid.lift_comp_of] + @[to_additive] lemma closure_eq_mrange (s : set M) : closure s = (free_monoid.lift (coe : s → M)).mrange := -by rw [mrange_eq_map, ← free_monoid.closure_range_of, map_mclosure, ← set.range_comp, - free_monoid.lift_comp_of, subtype.range_coe] +by rw [free_monoid.mrange_lift, subtype.range_coe] @[to_additive] lemma closure_eq_image_prod (s : set M) : (closure s : set M) = list.prod '' {l : list M | ∀ x ∈ l, x ∈ s} := From 2e433a50ea0c5d49d806f3eba6a84418681ee9c5 Mon Sep 17 00:00:00 2001 From: bottine Date: Mon, 3 Oct 2022 08:59:08 +0000 Subject: [PATCH 520/828] feat(category_theory/groupoid/free): free groupoid on a quiver (#16666) Define the free groupoid on a quiver and prove its universal property. Co-authored-by: Junyan Xu --- src/category_theory/functor/basic.lean | 10 + src/category_theory/groupoid.lean | 8 + .../groupoid/free_groupoid.lean | 190 ++++++++++++++++++ src/category_theory/path_category.lean | 48 +++++ src/category_theory/quotient.lean | 20 ++ src/combinatorics/quiver/basic.lean | 25 +++ .../quiver/connected_component.lean | 82 +++++++- src/combinatorics/quiver/path.lean | 1 + 8 files changed, 380 insertions(+), 4 deletions(-) create mode 100644 src/category_theory/groupoid/free_groupoid.lean diff --git a/src/category_theory/functor/basic.lean b/src/category_theory/functor/basic.lean index aaef887374d5f..76dc7ec26cec0 100644 --- a/src/category_theory/functor/basic.lean +++ b/src/category_theory/functor/basic.lean @@ -104,6 +104,16 @@ protected lemma id_comp (F : C ⥤ D) : (𝟭 C) ⋙ F = F := by cases F; refl F.map (if h : P then f h else g h) = if h : P then F.map (f h) else F.map (g h) := by { split_ifs; refl, } +@[simp] lemma to_prefunctor_obj (F : C ⥤ D) (X : C) : + F.to_prefunctor.obj X = F.obj X := rfl + +@[simp] lemma to_prefunctor_map (F : C ⥤ D) + {X Y : C} (f : X ⟶ Y) : F.to_prefunctor.map f = F.map f := rfl + +@[simp] lemma to_prefunctor_comp (F : C ⥤ D) (G : D ⥤ E) : + F.to_prefunctor.comp G.to_prefunctor = (F ⋙ G).to_prefunctor := rfl + + end end functor diff --git a/src/category_theory/groupoid.lean b/src/category_theory/groupoid.lean index 0140965cd0bdd..bd01812fbbaba 100644 --- a/src/category_theory/groupoid.lean +++ b/src/category_theory/groupoid.lean @@ -6,6 +6,9 @@ Authors: Reid Barton, Scott Morrison, David Wärn import category_theory.full_subcategory import category_theory.products.basic import category_theory.pi.basic +import category_theory.category.basic +import tactic.nth_rewrite +import combinatorics.quiver.connected_component /-! # Groupoids @@ -66,6 +69,11 @@ is_iso.eq_inv_of_hom_inv_id $ groupoid.comp_inv f @[simps] def groupoid.inv_equiv : (X ⟶ Y) ≃ (Y ⟶ X) := ⟨groupoid.inv, groupoid.inv, λ f, by simp, λ f, by simp⟩ +@[priority 100] +instance groupoid_has_involutive_reverse : quiver.has_involutive_reverse C := +{ reverse' := λ X Y f, groupoid.inv f +, inv' := λ X Y f, by { dsimp [quiver.reverse], simp, } } + variables (X Y) /-- In a groupoid, isomorphisms are equivalent to morphisms. -/ diff --git a/src/category_theory/groupoid/free_groupoid.lean b/src/category_theory/groupoid/free_groupoid.lean new file mode 100644 index 0000000000000..87c0f3d05bbdd --- /dev/null +++ b/src/category_theory/groupoid/free_groupoid.lean @@ -0,0 +1,190 @@ +/- +Copyright (c) 2022 Rémi Bottinelli. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémi Bottinelli +-/ +import category_theory.category.basic +import category_theory.functor.basic +import category_theory.groupoid +import combinatorics.quiver.basic +import combinatorics.quiver.connected_component +import logic.relation +import tactic.nth_rewrite +import category_theory.path_category +import category_theory.quotient + +/-! +# Free groupoid on a quiver + +This file defines the free groupoid on a quiver, the lifting of a prefunctor to its unique +extension as a functor from the free groupoid, and proves uniqueness of this extension. + +## Main results + +Given the type `V` and a quiver instance on `V`: + +- `free_groupoid V`: a type synonym for `V`. +- `free_groupoid_groupoid`: the `groupoid` instance on `free_groupoid V`. +- `lift`: the lifting of a prefunctor from `V` to `V'` where `V'` is a groupoid, to a functor. + `free_groupoid V ⥤ V'`. +- `lift_spec` and `lift_unique`: the proofs that, respectively, `lift` indeed is a lifting + and is the unique one. + +## Implementation notes + +The free groupoid is first defined by symmetrifying the quiver, taking the induced path category +and finally quotienting by the reducibility relation. + +-/ + + +open set classical function relation +local attribute [instance] prop_decidable + +namespace category_theory +namespace groupoid +namespace free + +universes u v u' v' + +variables {V : Type u} [quiver.{v+1} V] + +/-- Shorthand for the "forward" arrow corresponding to `f` in `symmetrify V` -/ +abbreviation quiver.hom.to_pos {X Y : V} (f : X ⟶ Y) : + (quiver.symmetrify_quiver V).hom X Y := sum.inl f + +/-- Shorthand for the "backward" arrow corresponding to `f` in `symmetrify V` -/ +abbreviation quiver.hom.to_neg {X Y : V} (f : X ⟶ Y) : + (quiver.symmetrify_quiver V).hom Y X := sum.inr f + +/-- Shorthand for the "forward" arrow corresponding to `f` in `paths $ symmetrify V` -/ +abbreviation quiver.hom.to_pos_path {X Y : V} (f : X ⟶ Y) : + ((category_theory.paths.category_paths $ quiver.symmetrify V).hom X Y) := f.to_pos.to_path + +/-- Shorthand for the "forward" arrow corresponding to `f` in `paths $ symmetrify V` -/ +abbreviation quiver.hom.to_neg_path {X Y : V} (f : X ⟶ Y) : + ((category_theory.paths.category_paths $ quiver.symmetrify V).hom Y X) := f.to_neg.to_path + +/-- The "reduction" relation -/ +inductive red_step : hom_rel (paths (quiver.symmetrify V)) +| step (X Z : quiver.symmetrify V) (f : X ⟶ Z) : + red_step (𝟙 X) (f.to_path ≫ (quiver.reverse f).to_path) + +/-- The underlying vertices of the free groupoid -/ +def free_groupoid (V) [Q : quiver.{v+1} V] := quotient (@red_step V Q) + +instance {V} [Q : quiver.{v+1} V] [h : nonempty V] : nonempty (free_groupoid V) := ⟨⟨h.some⟩⟩ + +lemma congr_reverse {X Y : paths $ quiver.symmetrify V} (p q : X ⟶ Y) : + quotient.comp_closure red_step p q → + quotient.comp_closure red_step (p.reverse) (q.reverse) := +begin + rintros ⟨U, W, XW, pp, qq, WY, ⟨_, Z, f⟩⟩, + have : quotient.comp_closure red_step (WY.reverse ≫ 𝟙 _ ≫ XW.reverse) + (WY.reverse ≫ (f.to_path ≫ (quiver.reverse f).to_path) ≫ XW.reverse), + { apply quotient.comp_closure.intro, + apply red_step.step, }, + simpa only [category_struct.comp, category_struct.id, quiver.path.reverse, quiver.path.nil_comp, + quiver.path.reverse_comp, quiver.reverse_reverse, quiver.path.reverse_to_path, + quiver.path.comp_assoc] using this, +end + +lemma congr_comp_reverse {X Y : paths $ quiver.symmetrify V} (p : X ⟶ Y) : + quot.mk (@quotient.comp_closure _ _ red_step _ _) (p ≫ p.reverse) = + quot.mk (@quotient.comp_closure _ _ red_step _ _) (𝟙 X) := +begin + apply quot.eqv_gen_sound, + induction p with _ _ q f ih, + { apply eqv_gen.refl, }, + { simp only [quiver.path.reverse], + fapply eqv_gen.trans, + { exact q ≫ q.reverse, }, + { apply eqv_gen.symm, apply eqv_gen.rel, + have : quotient.comp_closure + red_step (q ≫ (𝟙 _) ≫ q.reverse) + (q ≫ (f.to_path ≫ (quiver.reverse f).to_path) ≫ q.reverse), by + { apply quotient.comp_closure.intro, apply red_step.step, }, + have that : q.cons f = q.comp f.to_path, by refl, rw that, + simp only [category.assoc, category.id_comp] at this ⊢, + simp only [category_struct.comp, quiver.path.comp_assoc] at this ⊢, + exact this, }, + { exact ih }, }, +end + +lemma congr_reverse_comp {X Y : paths $ quiver.symmetrify V} (p : X ⟶ Y) : + quot.mk (@quotient.comp_closure _ _ red_step _ _) (p.reverse ≫ p) = + quot.mk (@quotient.comp_closure _ _ red_step _ _) (𝟙 Y) := +begin + nth_rewrite 1 ←quiver.path.reverse_reverse p, + apply congr_comp_reverse, +end + +instance : category (free_groupoid V) := quotient.category red_step + +/-- The inverse of an arrow in the free groupoid -/ +def quot_inv {X Y : free_groupoid V} (f : X ⟶ Y) : Y ⟶ X := +quot.lift_on f + (λ pp, quot.mk _ $ pp.reverse) + (λ pp qq con, quot.sound $ congr_reverse pp qq con) + +instance : groupoid (free_groupoid V) := +{ inv := λ X Y f, quot_inv f +, inv_comp' := λ X Y p, quot.induction_on p $ λ pp, congr_reverse_comp pp +, comp_inv' := λ X Y p, quot.induction_on p $ λ pp, congr_comp_reverse pp } + +/-- The inclusion of the quiver on `V` to the underlying quiver on `free_groupoid V`-/ +def of : prefunctor V (free_groupoid V) := +{ obj := λ X, ⟨X⟩ +, map := λ X Y f, quot.mk _ f.to_pos_path} + +lemma of_eq : of = + ((quiver.symmetrify.of).comp + paths.of).comp (quotient.functor $ @red_step V _).to_prefunctor := +begin + apply prefunctor.ext, rotate, + { rintro X, refl, }, + { rintro X Y f, refl, } +end + +section universal_property + +variables {V' : Type u'} [groupoid V'] (φ : prefunctor V V') + +/-- The lift of a prefunctor to a groupoid, to a functor from `free_groupoid V` -/ +def lift (φ : prefunctor V V') : free_groupoid V ⥤ V' := +quotient.lift _ + (paths.lift $ quiver.symmetrify.lift φ) + (by + { rintros _ _ _ _ ⟨X,Y,f⟩, + simp only [quiver.symmetrify.lift_reverse, paths.lift_nil, quiver.path.comp_nil, + paths.lift_cons, paths.lift_to_path], + symmetry, + apply groupoid.comp_inv, }) + +lemma lift_spec (φ : prefunctor V V') : of.comp (lift φ).to_prefunctor = φ := +begin + rw [of_eq, prefunctor.comp_assoc, prefunctor.comp_assoc, functor.to_prefunctor_comp], + dsimp [lift], + rw [quotient.lift_spec, paths.lift_spec, quiver.symmetrify.lift_spec], +end + +lemma lift_unique_spec (φ : prefunctor V V') (Φ : free_groupoid V ⥤ V') + (hΦ : of.comp Φ.to_prefunctor = φ) : Φ = (lift φ) := +begin + apply quotient.lift_spec_unique, + apply paths.lift_spec_unique, + apply quiver.symmetrify.lift_spec_unique, + { rw ←functor.to_prefunctor_comp, exact hΦ, }, + { rintros X Y f, + simp [←functor.to_prefunctor_comp,prefunctor.comp_map, paths.of_map, inv_eq_inv], + change Φ.map (inv ((quotient.functor red_step).to_prefunctor.map f.to_path)) = + inv (Φ.map ((quotient.functor red_step).to_prefunctor.map f.to_path)), + have := functor.map_inv Φ ((quotient.functor red_step).to_prefunctor.map f.to_path), + convert this; simp only [inv_eq_inv], }, +end + +end universal_property + +end free +end groupoid +end category_theory diff --git a/src/category_theory/path_category.lean b/src/category_theory/path_category.lean index 963973e54270a..60339f3cd66a5 100644 --- a/src/category_theory/path_category.lean +++ b/src/category_theory/path_category.lean @@ -52,6 +52,54 @@ def of : prefunctor V (paths V) := local attribute [ext] functor.ext +/-- Any prefunctor from `V` lifts to a functor from `paths V` -/ +def lift {C} [category C] (φ : prefunctor V C) : (paths V) ⥤ C := +{ obj := φ.obj +, map := λ X Y f, @quiver.path.rec V _ X (λ Y f, φ.obj X ⟶ φ.obj Y) (𝟙 $ φ.obj X) + (λ Y Z p f ihp, ihp ≫ (φ.map f)) Y f +, map_id' := λ X, by { refl, } +, map_comp' := λ X Y Z f g, by + { induction g with _ _ g' p ih _ _ _, + { rw category.comp_id, refl, }, + { have : f ≫ g'.cons p = (f ≫ g').cons p, by apply quiver.path.comp_cons, + rw this, simp only, rw [ih, category.assoc], }} } + +@[simp] lemma lift_nil {C} [category C] (φ : prefunctor V C) (X : V) : + (lift φ).map (quiver.path.nil) = 𝟙 (φ.obj X) := rfl + +@[simp] lemma lift_cons {C} [category C] (φ : prefunctor V C) {X Y Z: V} + (p : quiver.path X Y) (f : Y ⟶ Z) : + (lift φ).map (p.cons f) = (lift φ).map p ≫ (φ.map f) := rfl + +@[simp] lemma lift_to_path {C} [category C] (φ : prefunctor V C) {X Y : V} (f : X ⟶ Y) : + (lift φ).map f.to_path = φ.map f := by {dsimp [quiver.hom.to_path,lift], simp, } + +lemma lift_spec {C} [category C] (φ : prefunctor V C) : + of.comp (lift φ).to_prefunctor = φ := +begin + apply prefunctor.ext, rotate, + { rintro X, refl, }, + { rintro X Y f, rcases φ with ⟨φo,φm⟩, + dsimp [lift, quiver.hom.to_path], + simp only [category.id_comp], }, +end + +lemma lift_spec_unique {C} [category C] (φ : prefunctor V C) (Φ : paths V ⥤ C) + (hΦ : of.comp Φ.to_prefunctor = φ) : Φ = lift φ := +begin + subst_vars, + apply functor.ext, rotate, + { rintro X, refl, }, + { rintro X Y f, + dsimp [lift], + induction f with _ _ p f' ih, + { simp only [category.comp_id], apply functor.map_id, }, + { simp only [category.comp_id, category.id_comp] at ih ⊢, + have : Φ.map (p.cons f') = Φ.map p ≫ (Φ.map (f'.to_path)), by + { convert functor.map_comp Φ p (f'.to_path), }, + rw [this,ih], }, }, +end + /-- Two functors out of a path category are equal when they agree on singleton paths. -/ @[ext] lemma ext_functor {C} [category C] diff --git a/src/category_theory/quotient.lean b/src/category_theory/quotient.lean index 94b63d2b1c21d..1505a2c65360b 100644 --- a/src/category_theory/quotient.lean +++ b/src/category_theory/quotient.lean @@ -130,6 +130,26 @@ def lift : quotient r ⥤ D := map_id' := λ a, F.map_id a.as, map_comp' := by { rintros a b c ⟨f⟩ ⟨g⟩, exact F.map_comp f g, } } +lemma lift_spec : (functor r) ⋙ lift r F H = F := +begin + apply functor.ext, rotate, + { rintro X, refl, }, + { rintro X Y f, simp, }, +end + +lemma lift_spec_unique (Φ : quotient r ⥤ D) (hΦ : (functor r) ⋙ Φ = F) : Φ = lift r F H := +begin + subst_vars, + apply functor.hext, + { rintro X, dsimp [lift, functor], congr, ext, refl, }, + { rintro X Y f, + dsimp [lift, functor], + apply quot.induction_on f, + rintro ff, + simp only [quot.lift_on_mk, functor.comp_map], + congr; ext; refl, }, +end + /-- The original functor factors through the induced functor. -/ def lift.is_lift : (functor r) ⋙ lift r F H ≅ F := nat_iso.of_components (λ X, iso.refl _) (by tidy) diff --git a/src/combinatorics/quiver/basic.lean b/src/combinatorics/quiver/basic.lean index 37c9b5a9eb945..24b02fbed757e 100644 --- a/src/combinatorics/quiver/basic.lean +++ b/src/combinatorics/quiver/basic.lean @@ -53,6 +53,20 @@ structure prefunctor (V : Type u₁) [quiver.{v₁} V] (W : Type u₂) [quiver.{ namespace prefunctor +@[ext] +lemma ext {V : Type u} [quiver.{v₁} V] {W : Type u₂} [quiver.{v₂} W] + {F G : prefunctor V W} + (h_obj : ∀ X, F.obj X = G.obj X) + (h_map : ∀ (X Y : V) (f : X ⟶ Y), + F.map f = eq.rec_on (h_obj Y).symm (eq.rec_on (h_obj X).symm (G.map f))) : F = G := +begin + cases F with F_obj _, cases G with G_obj _, + obtain rfl : F_obj = G_obj, by { ext X, apply h_obj }, + congr, + funext X Y f, + simpa using h_map X Y f, +end + /-- The identity morphism between quivers. -/ @@ -72,6 +86,17 @@ def comp {U : Type*} [quiver U] {V : Type*} [quiver V] {W : Type*} [quiver W] { obj := λ X, G.obj (F.obj X), map := λ X Y f, G.map (F.map f), } +@[simp] +lemma comp_assoc + {U : Type*} [quiver U] {V : Type*} [quiver V] {W : Type*} [quiver W] {Z : Type*} [quiver Z] + (F : prefunctor U V) (G : prefunctor V W) (H : prefunctor W Z) : + (F.comp G).comp H = F.comp (G.comp H) := +begin + apply prefunctor.ext, rotate, + { rintro X, refl, }, + { rintro X Y Z, refl, } +end + end prefunctor namespace quiver diff --git a/src/combinatorics/quiver/connected_component.lean b/src/combinatorics/quiver/connected_component.lean index 6ce1276737e98..9d3404adf7159 100644 --- a/src/combinatorics/quiver/connected_component.lean +++ b/src/combinatorics/quiver/connected_component.lean @@ -36,18 +36,92 @@ variables (V : Type u) [quiver.{v+1} V] class has_reverse := (reverse' : Π {a b : V}, (a ⟶ b) → (b ⟶ a)) -instance : has_reverse (symmetrify V) := ⟨λ a b e, e.swap⟩ +/-- Reverse the direction of an arrow. -/ +def reverse {V} [quiver.{v+1} V] [has_reverse V] {a b : V} : (a ⟶ b) → (b ⟶ a) := + has_reverse.reverse' + +/-- A quiver `has_involutive_reverse` if reversing twice is the identity.`-/ +class has_involutive_reverse extends has_reverse V := +(inv' : Π {a b : V} (f : a ⟶ b), reverse (reverse f) = f) + +@[simp] lemma reverse_reverse {V} [quiver.{v+1} V] [h : has_involutive_reverse V] + {a b : V} (f : a ⟶ b) : reverse (reverse f) = f := by apply h.inv' variables {V} -/-- Reverse the direction of an arrow. -/ -def reverse [has_reverse V] {a b : V} : (a ⟶ b) → (b ⟶ a) := has_reverse.reverse' +instance : has_reverse (symmetrify V) := ⟨λ a b e, e.swap⟩ +instance : has_involutive_reverse (symmetrify V) := +{ to_has_reverse := ⟨λ a b e, e.swap⟩, + inv' := λ a b e, congr_fun sum.swap_swap_eq e } + /-- Reverse the direction of a path. -/ -def path.reverse [has_reverse V] {a : V} : Π {b}, path a b → path b a +@[simp] def path.reverse [has_reverse V] {a : V} : Π {b}, path a b → path b a | a path.nil := path.nil | b (path.cons p e) := (reverse e).to_path.comp p.reverse +@[simp] lemma path.reverse_to_path [has_reverse V] {a b : V} (f : a ⟶ b) : + f.to_path.reverse = (reverse f).to_path := rfl + +@[simp] lemma path.reverse_comp [has_reverse V] {a b c : V} (p : path a b) (q : path b c) : + (p.comp q).reverse = q.reverse.comp p.reverse := by +{ induction q, { simp, }, { simp [q_ih], }, } + +@[simp] lemma path.reverse_reverse [h : has_involutive_reverse V] {a b : V} (p : path a b) : + p.reverse.reverse = p := by +{ induction p, + { simp, }, + { simp only [path.reverse, path.reverse_comp, path.reverse_to_path, reverse_reverse, p_ih], + refl, }, } + +/-- The inclusion of a quiver in its symmetrification -/ +def symmetrify.of : prefunctor V (symmetrify V) := +{ obj := id +, map := λ X Y f, sum.inl f } + +/-- Given a quiver `V'` with reversible arrows, a prefunctor to `V'` can be lifted to one from + `symmetrify V` to `V'` -/ +def symmetrify.lift {V' : Type*} [quiver V'] [has_reverse V'] (φ : prefunctor V V') : + prefunctor (symmetrify V) V' := +{ obj := φ.obj +, map := λ X Y f, sum.rec (λ fwd, φ.map fwd) (λ bwd, reverse (φ.map bwd)) f } + +lemma symmetrify.lift_spec (V' : Type*) [quiver V'] [has_reverse V'] (φ : prefunctor V V') : + symmetrify.of.comp (symmetrify.lift φ) = φ := +begin + fapply prefunctor.ext, + { rintro X, refl, }, + { rintros X Y f, refl, }, +end + +lemma symmetrify.lift_reverse (V' : Type*) [quiver V'] [h : has_involutive_reverse V'] + (φ : prefunctor V V') + {X Y : symmetrify V} (f : X ⟶ Y) : + (symmetrify.lift φ).map (quiver.reverse f) = quiver.reverse ((symmetrify.lift φ).map f) := +begin + dsimp [symmetrify.lift], cases f, + { simp only, refl, }, + { simp only, rw h.inv', refl, } +end + +/-- `lift φ` is the only prefunctor extending `φ` and preserving reverses. -/ +lemma symmetrify.lift_spec_unique (V' : Type*) [quiver V'] [has_reverse V'] + (φ : prefunctor V V') + (Φ : prefunctor (symmetrify V) V') + (hΦ : symmetrify.of.comp Φ = φ) + (hΦinv : ∀ {X Y : V} (f : X ⟶ Y), Φ.map (reverse f) = reverse (Φ.map f)) : + Φ = symmetrify.lift φ := +begin + subst_vars, + fapply prefunctor.ext, + { rintro X, refl, }, + { rintros X Y f, + cases f, + { refl, }, + { dsimp [symmetrify.lift,symmetrify.of], + convert hΦinv (sum.inl f), }, }, +end + variables (V) /-- Two vertices are related in the zigzag setoid if there is a diff --git a/src/combinatorics/quiver/path.lean b/src/combinatorics/quiver/path.lean index ed87f1cafe5b7..3663aaf7f361a 100644 --- a/src/combinatorics/quiver/path.lean +++ b/src/combinatorics/quiver/path.lean @@ -126,3 +126,4 @@ def map_path {a : V} : lemma map_path_to_path {a b : V} (f : a ⟶ b) : F.map_path f.to_path = (F.map f).to_path := rfl end prefunctor + From 9dbe038ef150bf54cb30b53d20c4562abf30258d Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 3 Oct 2022 08:59:09 +0000 Subject: [PATCH 521/828] fix({ tactic/fin_cases + test/interval_cases }): add `focus1` and a test (#16752) Fix [this bug](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/interval_cases.20bug), reported by @hrmacbeth. The current solution is to move the `focus1` to `fin_cases_at`, instead of relying on `fin_cases` and `interval_cases` to call `focus1`. Having `interval_cases` wrapped in `focus1` is the reason for this PR. An older (compiling) solution was to just use `focus1` inside `interval_cases`, but calling it in `fin_cases_at` seems more stable. The link above is to a Zulip discussion with more details. --- src/tactic/fin_cases.lean | 8 ++++---- test/interval_cases.lean | 9 +++++++++ 2 files changed, 13 insertions(+), 4 deletions(-) diff --git a/src/tactic/fin_cases.lean b/src/tactic/fin_cases.lean index 6e0a575cd81c6..d0a7db95b2a9e 100644 --- a/src/tactic/fin_cases.lean +++ b/src/tactic/fin_cases.lean @@ -78,8 +78,8 @@ for example, to display nats as `n.succ` instead of `n+1`. These should be defeq to and in the same order as the terms in the enumeration of `α`. -/ meta def fin_cases_at (nm : option name) : Π (with_list : option pexpr) (e : expr), tactic unit -| with_list e := -do ty ← try_core $ guard_mem_fin e, +| with_list e := focus1 $ + do ty ← try_core $ guard_mem_fin e, match ty with | none := -- Deal with `x : A`, where `[fintype A]` is available: (do @@ -151,7 +151,7 @@ produces three goals with hypotheses -/ meta def fin_cases : parse hyp → parse (tk "with" *> texpr)? → parse (tk "using" *> ident)? → tactic unit -| none none nm := focus1 $ do +| none none nm := do ctx ← local_context, ctx.mfirst (fin_cases_at nm none) <|> fail ("No hypothesis of the forms `x ∈ A`, where " ++ @@ -160,7 +160,7 @@ meta def fin_cases : | (some n) with_list nm := do h ← get_local n, - focus1 $ fin_cases_at nm with_list h + fin_cases_at nm with_list h end interactive diff --git a/test/interval_cases.lean b/test/interval_cases.lean index 760ee945abfb6..7e9c8e528b0a1 100644 --- a/test/interval_cases.lean +++ b/test/interval_cases.lean @@ -127,6 +127,15 @@ begin { right, exact hrv } end +/- https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/interval_cases.20bug -/ +example {x : ℕ} (hx2 : x < 2) (h : false) : false := +begin + have : x ≤ 1, + interval_cases x, + exact zero_le_one, -- this solves the side-goal left by `interval_cases`, closing the `have` + exact h, +end + /- Sadly, this one doesn't work, reporting: `deep recursion was detected at 'expression equality test'` From c7aa2dfa8ab00da8ff6be6ecab56b41d0e2cfd1f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 3 Oct 2022 14:28:31 +0000 Subject: [PATCH 522/828] refactor(data/fin/basic): merge `fin.rev` and `order_iso.fin_equiv` (#16745) Also add some missing lemmas. --- src/data/fin/basic.lean | 80 ++++++++++++++++++++--------------------- 1 file changed, 40 insertions(+), 40 deletions(-) diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index 69732b5a93a31..304e3805f3460 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -94,10 +94,11 @@ lemma val_injective : function.injective (@fin.val n) := @fin.eq_of_veq n protected lemma prop (a : fin n) : a.val < n := a.2 -@[simp] lemma is_lt (a : fin n) : (a : nat) < n := a.2 +@[simp] lemma is_lt (a : fin n) : (a : ℕ) < n := a.2 -lemma pos_iff_nonempty {n : ℕ} : 0 < n ↔ nonempty (fin n) := -⟨λ h, ⟨⟨0, h⟩⟩, λ ⟨i⟩, lt_of_le_of_lt (nat.zero_le _) i.2⟩ +protected lemma pos (i : fin n) : 0 < n := lt_of_le_of_lt (nat.zero_le _) i.is_lt + +lemma pos_iff_nonempty {n : ℕ} : 0 < n ↔ nonempty (fin n) := ⟨λ h, ⟨⟨0, h⟩⟩, λ ⟨i⟩, i.pos⟩ /-- Equivalence between `fin n` and `{ i // i < n }`. -/ @[simps apply symm_apply] @@ -271,26 +272,46 @@ begin { right, exact ⟨⟨j, nat.lt_of_succ_lt_succ h⟩, rfl⟩, } end +lemma eq_succ_of_ne_zero {n : ℕ} {i : fin (n + 1)} (hi : i ≠ 0) : ∃ j : fin n, i = j.succ := +(eq_zero_or_eq_succ i).resolve_left hi + /-- The antitone involution `fin n → fin n` given by `i ↦ n-(i+1)`. -/ -def rev {n : ℕ} (i : fin n) : fin n := -⟨n-(i+1), begin - cases n, - { exfalso, - exact nat.not_lt_zero _ (i.is_lt), }, - { simpa only [nat.succ_sub_succ_eq_sub n i, nat.lt_succ_iff] using tsub_le_self, }, -end⟩ +def rev : equiv.perm (fin n) := +involutive.to_perm (λ i, ⟨n - (i + 1), tsub_lt_self i.pos (nat.succ_pos _)⟩) $ + λ i, ext $ by rw [coe_mk, coe_mk, ← tsub_tsub, + tsub_tsub_cancel_of_le (nat.add_one_le_iff.2 i.is_lt), add_tsub_cancel_right] + +@[simp] lemma coe_rev (i : fin n) : (i.rev : ℕ) = n - (i + 1) := rfl +lemma rev_involutive : involutive (@rev n) := involutive.to_perm_involutive _ +lemma rev_injective : injective (@rev n) := rev_involutive.injective +lemma rev_surjective : surjective (@rev n) := rev_involutive.surjective +lemma rev_bijective : bijective (@rev n) := rev_involutive.bijective +@[simp] lemma rev_inj {i j : fin n} : i.rev = j.rev ↔ i = j := rev_injective.eq_iff +@[simp] lemma rev_rev (i : fin n) : i.rev.rev = i := rev_involutive _ +@[simp] lemma rev_symm : (@rev n).symm = rev := rfl lemma rev_eq {n a : ℕ} (i : fin (n+1)) (h : n=a+i) : i.rev = ⟨a, nat.lt_succ_iff.mpr (nat.le.intro (h.symm))⟩ := begin ext, - dsimp [fin.rev], + dsimp, conv_lhs { congr, rw h, }, rw [add_assoc, add_tsub_cancel_right], end -lemma eq_succ_of_ne_zero {n : ℕ} {i : fin (n + 1)} (hi : i ≠ 0) : ∃ j : fin n, i = j.succ := -(eq_zero_or_eq_succ i).resolve_left hi +@[simp] lemma rev_le_rev {i j : fin n} : i.rev ≤ j.rev ↔ j ≤ i := +by simp only [le_iff_coe_le_coe, coe_rev, tsub_le_tsub_iff_left (nat.add_one_le_iff.2 j.is_lt), + add_le_add_iff_right] + +@[simp] lemma rev_lt_rev {i j : fin n} : i.rev < j.rev ↔ j < i := +lt_iff_lt_of_le_iff_le rev_le_rev + +/-- `fin.rev n` as an order-reversing isomorphism. -/ +@[simps apply to_equiv] def rev_order_iso {n} : (fin n)ᵒᵈ ≃o fin n := +⟨order_dual.of_dual.trans rev, λ i j, rev_le_rev⟩ + +@[simp] lemma rev_order_iso_symm_apply (i : fin n) : + rev_order_iso.symm i = order_dual.to_dual i.rev := rfl /-- The greatest value of `fin (n+1)` -/ def last (n : ℕ) : fin (n+1) := ⟨_, n.lt_succ_self⟩ @@ -1027,7 +1048,7 @@ def div_nat (i : fin (m * n)) : fin m := /-- Compute `i % n`, where `n` is a `nat` and inferred the type of `i`. -/ def mod_nat (i : fin (m * n)) : fin n := -⟨i % n, nat.mod_lt _ $ pos_of_mul_pos_right ((nat.zero_le i).trans_lt i.is_lt) m.zero_le⟩ +⟨i % n, nat.mod_lt _ $ pos_of_mul_pos_right i.pos m.zero_le⟩ @[simp] lemma coe_mod_nat (i : fin (m * n)) : (i.mod_nat : ℕ) = i % n := rfl @@ -1282,8 +1303,7 @@ section add_group open nat int /-- Negation on `fin n` -/ -instance (n : ℕ) : has_neg (fin n) := -⟨λ a, ⟨(n - a) % n, nat.mod_lt _ (lt_of_le_of_lt (nat.zero_le _) a.2)⟩⟩ +instance (n : ℕ) : has_neg (fin n) := ⟨λ a, ⟨(n - a) % n, nat.mod_lt _ a.pos⟩⟩ /-- Abelian group structure on `fin (n+1)`. -/ instance (n : ℕ) : add_comm_group (fin (n+1)) := @@ -1367,32 +1387,12 @@ begin simp end -lemma sub_one_lt_iff {n : ℕ} {k : fin (n + 1)} : +@[simp] lemma sub_one_lt_iff {n : ℕ} {k : fin (n + 1)} : k - 1 < k ↔ 0 < k := -begin - rw ←not_iff_not, - simp -end +not_iff_not.1 $ by simp only [not_lt, le_sub_one_iff, le_zero_iff] -/-- By sending `x` to `last n - x`, `fin n` is order-equivalent to its `order_dual`. -/ -def _root_.order_iso.fin_equiv : ∀ {n}, (fin n)ᵒᵈ ≃o fin n -| 0 := ⟨⟨elim0, elim0, elim0, elim0⟩, elim0⟩ -| (n+1) := order_iso.symm $ -{ to_fun := λ x, order_dual.to_dual (last n - x), - inv_fun := λ x, last n - x.of_dual, - left_inv := sub_sub_cancel _, - right_inv := sub_sub_cancel _, - map_rel_iff' := λ a b, - begin - simp only [equiv.coe_fn_mk, order_dual.to_dual_le_to_dual], - rw [le_iff_coe_le_coe, coe_sub_iff_le.mpr (le_last b), coe_sub_iff_le.mpr (le_last _), - tsub_le_tsub_iff_left, le_iff_coe_le_coe], - exact le_last _, - end } - -lemma _root_.order_iso.fin_equiv_apply (a) : order_iso.fin_equiv a = last n - a.of_dual := rfl -lemma _root_.order_iso.fin_equiv_symm_apply (a) : - order_iso.fin_equiv.symm a = order_dual.to_dual (last n - a) := rfl +lemma last_sub (i : fin (n + 1)) : last n - i = i.rev := +ext $ by rw [coe_sub_iff_le.2 i.le_last, coe_last, coe_rev, nat.succ_sub_succ_eq_sub] end add_group From 8a1b982fecc23f092f4bb5aa57e8d0dcd954dc04 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 3 Oct 2022 14:28:32 +0000 Subject: [PATCH 523/828] feat(data/{fin,}set/pointwise): add `{fin,}set.inv_insert` (#16748) --- src/data/finset/pointwise.lean | 3 +++ src/data/set/pointwise.lean | 3 +++ 2 files changed, 6 insertions(+) diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index a33a2ee6f9404..d9d754def5856 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -119,6 +119,9 @@ attribute [mono] neg_subset_neg @[simp, to_additive] lemma inv_singleton (a : α) : ({a} : finset α)⁻¹ = {a⁻¹} := image_singleton _ _ +@[simp, to_additive] +lemma inv_insert (a : α) (s : finset α) : (insert a s)⁻¹ = insert a⁻¹ s⁻¹ := image_insert _ _ _ + end has_inv open_locale pointwise diff --git a/src/data/set/pointwise.lean b/src/data/set/pointwise.lean index a7a5e028cf786..d8ea09db90344 100644 --- a/src/data/set/pointwise.lean +++ b/src/data/set/pointwise.lean @@ -166,6 +166,9 @@ lemma inv_subset_inv : s⁻¹ ⊆ t⁻¹ ↔ s ⊆ t := @[simp, to_additive] lemma inv_singleton (a : α) : ({a} : set α)⁻¹ = {a⁻¹} := by rw [←image_inv, image_singleton] +@[simp, to_additive] lemma inv_insert (a : α) (s : set α) : (insert a s)⁻¹ = insert a⁻¹ s⁻¹ := +by rw [insert_eq, union_inv, inv_singleton, insert_eq] + @[to_additive] lemma inv_range {ι : Sort*} {f : ι → α} : (range f)⁻¹ = range (λ i, (f i)⁻¹) := by { rw ←image_inv, exact (range_comp _ _).symm } From 81120aecec6dd7ad96704c3c08d8280e8ac21279 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Mon, 3 Oct 2022 14:28:33 +0000 Subject: [PATCH 524/828] refactor(number_theory): reorganize number field results into new subfolder (#16764) - Create a new dir `number_field` in `number_theory` - Move the current file `number_field.lean` to `number_theory/number_field/basic.lean` - Move the results about embeddings from this file to a new file `number_theory/number_field/embeddings.lean` - Move the file `number_theory/class_number/number_field.lean` to `number_theory/number_field/class_number.lean` --- src/number_theory/cyclotomic/basic.lean | 2 +- src/number_theory/number_field/basic.lean | 167 +++++++++++++++++ .../class_number.lean} | 2 +- .../embeddings.lean} | 175 ++---------------- src/ring_theory/discriminant.lean | 2 +- 5 files changed, 182 insertions(+), 166 deletions(-) create mode 100644 src/number_theory/number_field/basic.lean rename src/number_theory/{class_number/number_field.lean => number_field/class_number.lean} (97%) rename src/number_theory/{number_field.lean => number_field/embeddings.lean} (52%) diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index a01087be87cca..4e562176e7713 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -5,7 +5,7 @@ Authors: Riccardo Brasca -/ import ring_theory.polynomial.cyclotomic.basic -import number_theory.number_field +import number_theory.number_field.basic import algebra.char_p.algebra import field_theory.galois diff --git a/src/number_theory/number_field/basic.lean b/src/number_theory/number_field/basic.lean new file mode 100644 index 0000000000000..aaecadef620b5 --- /dev/null +++ b/src/number_theory/number_field/basic.lean @@ -0,0 +1,167 @@ +/- +Copyright (c) 2021 Ashvni Narayanan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Ashvni Narayanan, Anne Baanen +-/ + +import ring_theory.dedekind_domain.integral_closure +import algebra.char_p.algebra + +/-! +# Number fields +This file defines a number field and the ring of integers corresponding to it. + +## Main definitions + - `number_field` defines a number field as a field which has characteristic zero and is finite + dimensional over ℚ. + - `ring_of_integers` defines the ring of integers (or number ring) corresponding to a number field + as the integral closure of ℤ in the number field. + +## Implementation notes +The definitions that involve a field of fractions choose a canonical field of fractions, +but are independent of that choice. + +## References +* [D. Marcus, *Number Fields*][marcus1977number] +* [J.W.S. Cassels, A. Frölich, *Algebraic Number Theory*][cassels1967algebraic] +* [P. Samuel, *Algebraic Theory of Numbers*][samuel1970algebraic] + +## Tags +number field, ring of integers +-/ + +/-- A number field is a field which has characteristic zero and is finite +dimensional over ℚ. -/ +class number_field (K : Type*) [field K] : Prop := +[to_char_zero : char_zero K] +[to_finite_dimensional : finite_dimensional ℚ K] + +open function +open_locale classical big_operators + +/-- `ℤ` with its usual ring structure is not a field. -/ +lemma int.not_is_field : ¬ is_field ℤ := +λ h, int.not_even_one $ (h.mul_inv_cancel two_ne_zero).imp $ λ a, (by rw ← two_mul; exact eq.symm) + +namespace number_field + +variables (K L : Type*) [field K] [field L] [nf : number_field K] + +include nf + +-- See note [lower instance priority] +attribute [priority 100, instance] number_field.to_char_zero number_field.to_finite_dimensional + +protected lemma is_algebraic : algebra.is_algebraic ℚ K := algebra.is_algebraic_of_finite _ _ + +omit nf + +/-- The ring of integers (or number ring) corresponding to a number field +is the integral closure of ℤ in the number field. -/ +def ring_of_integers := integral_closure ℤ K + +localized "notation (name := ring_of_integers) + `𝓞` := number_field.ring_of_integers" in number_field + +lemma mem_ring_of_integers (x : K) : x ∈ 𝓞 K ↔ is_integral ℤ x := iff.rfl + +lemma is_integral_of_mem_ring_of_integers {K : Type*} [field K] {x : K} (hx : x ∈ 𝓞 K) : + is_integral ℤ (⟨x, hx⟩ : 𝓞 K) := +begin + obtain ⟨P, hPm, hP⟩ := hx, + refine ⟨P, hPm, _⟩, + rw [← polynomial.aeval_def, ← subalgebra.coe_eq_zero, polynomial.aeval_subalgebra_coe, + polynomial.aeval_def, subtype.coe_mk, hP] +end + +/-- Given an algebra between two fields, create an algebra between their two rings of integers. + +For now, this is not an instance by default as it creates an equal-but-not-defeq diamond with +`algebra.id` when `K = L`. This is caused by `x = ⟨x, x.prop⟩` not being defeq on subtypes. This +will likely change in Lean 4. -/ +def ring_of_integers_algebra [algebra K L] : algebra (𝓞 K) (𝓞 L) := ring_hom.to_algebra +{ to_fun := λ k, ⟨algebra_map K L k, is_integral.algebra_map k.2⟩, + map_zero' := subtype.ext $ by simp only [subtype.coe_mk, subalgebra.coe_zero, map_zero], + map_one' := subtype.ext $ by simp only [subtype.coe_mk, subalgebra.coe_one, map_one], + map_add' := λ x y, subtype.ext $ by simp only [map_add, subalgebra.coe_add, subtype.coe_mk], + map_mul' := λ x y, subtype.ext $ by simp only [subalgebra.coe_mul, map_mul, subtype.coe_mk] } + +namespace ring_of_integers + +variables {K} + +instance [number_field K] : is_fraction_ring (𝓞 K) K := +integral_closure.is_fraction_ring_of_finite_extension ℚ _ + +instance : is_integral_closure (𝓞 K) ℤ K := +integral_closure.is_integral_closure _ _ + +instance [number_field K] : is_integrally_closed (𝓞 K) := +integral_closure.is_integrally_closed_of_finite_extension ℚ + +lemma is_integral_coe (x : 𝓞 K) : is_integral ℤ (x : K) := +x.2 + +/-- The ring of integers of `K` are equivalent to any integral closure of `ℤ` in `K` -/ +protected noncomputable def equiv (R : Type*) [comm_ring R] [algebra R K] + [is_integral_closure R ℤ K] : 𝓞 K ≃+* R := +(is_integral_closure.equiv ℤ R K _).symm.to_ring_equiv + +variables (K) + +instance [number_field K] : char_zero (𝓞 K) := char_zero.of_module _ K + +instance [number_field K] : is_noetherian ℤ (𝓞 K) := is_integral_closure.is_noetherian _ ℚ K _ + +/-- The ring of integers of a number field is not a field. -/ +lemma not_is_field [number_field K] : ¬ is_field (𝓞 K) := +begin + have h_inj : function.injective ⇑(algebra_map ℤ (𝓞 K)), + { exact ring_hom.injective_int (algebra_map ℤ (𝓞 K)) }, + intro hf, + exact int.not_is_field + (((is_integral_closure.is_integral_algebra ℤ K).is_field_iff_is_field h_inj).mpr hf) +end + +instance [number_field K] : is_dedekind_domain (𝓞 K) := +is_integral_closure.is_dedekind_domain ℤ ℚ K _ + +end ring_of_integers + +end number_field + +namespace rat + +open number_field + +instance number_field : number_field ℚ := +{ to_char_zero := infer_instance, + to_finite_dimensional := + -- The vector space structure of `ℚ` over itself can arise in multiple ways: + -- all fields are vector spaces over themselves (used in `rat.finite_dimensional`) + -- all char 0 fields have a canonical embedding of `ℚ` (used in `number_field`). + -- Show that these coincide: + by convert (infer_instance : finite_dimensional ℚ ℚ), } + +/-- The ring of integers of `ℚ` as a number field is just `ℤ`. -/ +noncomputable def ring_of_integers_equiv : ring_of_integers ℚ ≃+* ℤ := +ring_of_integers.equiv ℤ + +end rat + +namespace adjoin_root + +section + +open_locale polynomial + +local attribute [-instance] algebra_rat + +/-- The quotient of `ℚ[X]` by the ideal generated by an irreducible polynomial of `ℚ[X]` +is a number field. -/ +instance {f : ℚ[X]} [hf : fact (irreducible f)] : number_field (adjoin_root f) := +{ to_char_zero := char_zero_of_injective_algebra_map (algebra_map ℚ _).injective, + to_finite_dimensional := by convert (adjoin_root.power_basis hf.out.ne_zero).finite_dimensional } +end + +end adjoin_root diff --git a/src/number_theory/class_number/number_field.lean b/src/number_theory/number_field/class_number.lean similarity index 97% rename from src/number_theory/class_number/number_field.lean rename to src/number_theory/number_field/class_number.lean index 19deef684a86e..4b8bc9a7b94df 100644 --- a/src/number_theory/class_number/number_field.lean +++ b/src/number_theory/number_field/class_number.lean @@ -5,7 +5,7 @@ Authors: Anne Baanen -/ import number_theory.class_number.admissible_abs import number_theory.class_number.finite -import number_theory.number_field +import number_theory.number_field.basic /-! # Class numbers of number fields diff --git a/src/number_theory/number_field.lean b/src/number_theory/number_field/embeddings.lean similarity index 52% rename from src/number_theory/number_field.lean rename to src/number_theory/number_field/embeddings.lean index acc6a8921d06a..543b924f44482 100644 --- a/src/number_theory/number_field.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -1,179 +1,28 @@ /- -Copyright (c) 2021 Ashvni Narayanan. All rights reserved. +Copyright (c) 2022 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Ashvni Narayanan, Anne Baanen +Authors: Alex J. Best, Xavier Roblot -/ -import ring_theory.dedekind_domain.integral_closure -import algebra.char_p.algebra import analysis.normed_space.basic +import number_theory.number_field.basic import topology.algebra.polynomial /-! -# Number fields -This file defines a number field, the ring of integers corresponding to it and includes some -basic facts about the embeddings into an algebraic closed field. +# Embeddings of number fields +This file defines the embeddings of a number field into an algebraic closed field. -## Main definitions - - `number_field` defines a number field as a field which has characteristic zero and is finite - dimensional over ℚ. - - `ring_of_integers` defines the ring of integers (or number ring) corresponding to a number field - as the integral closure of ℤ in the number field. - -## Main Result - - `eq_roots`: let `x ∈ K` with `K` number field and let `A` be an algebraic closed field of - char. 0, then the images of `x` by the embeddings of `K` in `A` are exactly the roots in - `A` of the minimal polynomial of `x` over `ℚ`. - -## Implementation notes -The definitions that involve a field of fractions choose a canonical field of fractions, -but are independent of that choice. - -## References -* [D. Marcus, *Number Fields*][marcus1977number] -* [J.W.S. Cassels, A. Frölich, *Algebraic Number Theory*][cassels1967algebraic] -* [P. Samuel, *Algebraic Theory of Numbers*][samuel1970algebraic] +## Main Results +* `number_field.embeddings.eq_roots`: let `x ∈ K` with `K` number field and let `A` be an algebraic + closed field of char. 0, then the images of `x` by the embeddings of `K` in `A` are exactly the + roots in `A` of the minimal polynomial of `x` over `ℚ`. +* `number_field.embeddings.pow_eq_one_of_norm_eq_one`: an algebraic integer whose conjugates are + all of norm one is a root of unity. ## Tags -number field, ring of integers +number field, embeddings -/ -/-- A number field is a field which has characteristic zero and is finite -dimensional over ℚ. -/ -class number_field (K : Type*) [field K] : Prop := -[to_char_zero : char_zero K] -[to_finite_dimensional : finite_dimensional ℚ K] - -open function -open_locale classical big_operators - -/-- `ℤ` with its usual ring structure is not a field. -/ -lemma int.not_is_field : ¬ is_field ℤ := -λ h, int.not_even_one $ (h.mul_inv_cancel two_ne_zero).imp $ λ a, (by rw ← two_mul; exact eq.symm) - -namespace number_field - -variables (K L : Type*) [field K] [field L] [nf : number_field K] - -include nf - --- See note [lower instance priority] -attribute [priority 100, instance] number_field.to_char_zero number_field.to_finite_dimensional - -protected lemma is_algebraic : algebra.is_algebraic ℚ K := algebra.is_algebraic_of_finite _ _ - -omit nf - -/-- The ring of integers (or number ring) corresponding to a number field -is the integral closure of ℤ in the number field. -/ -def ring_of_integers := integral_closure ℤ K - -localized "notation (name := ring_of_integers) - `𝓞` := number_field.ring_of_integers" in number_field - -lemma mem_ring_of_integers (x : K) : x ∈ 𝓞 K ↔ is_integral ℤ x := iff.rfl - -lemma is_integral_of_mem_ring_of_integers {K : Type*} [field K] {x : K} (hx : x ∈ 𝓞 K) : - is_integral ℤ (⟨x, hx⟩ : 𝓞 K) := -begin - obtain ⟨P, hPm, hP⟩ := hx, - refine ⟨P, hPm, _⟩, - rw [← polynomial.aeval_def, ← subalgebra.coe_eq_zero, polynomial.aeval_subalgebra_coe, - polynomial.aeval_def, subtype.coe_mk, hP] -end - -/-- Given an algebra between two fields, create an algebra between their two rings of integers. - -For now, this is not an instance by default as it creates an equal-but-not-defeq diamond with -`algebra.id` when `K = L`. This is caused by `x = ⟨x, x.prop⟩` not being defeq on subtypes. This -will likely change in Lean 4. -/ -def ring_of_integers_algebra [algebra K L] : algebra (𝓞 K) (𝓞 L) := ring_hom.to_algebra -{ to_fun := λ k, ⟨algebra_map K L k, is_integral.algebra_map k.2⟩, - map_zero' := subtype.ext $ by simp only [subtype.coe_mk, subalgebra.coe_zero, map_zero], - map_one' := subtype.ext $ by simp only [subtype.coe_mk, subalgebra.coe_one, map_one], - map_add' := λ x y, subtype.ext $ by simp only [map_add, subalgebra.coe_add, subtype.coe_mk], - map_mul' := λ x y, subtype.ext $ by simp only [subalgebra.coe_mul, map_mul, subtype.coe_mk] } - -namespace ring_of_integers - -variables {K} - -instance [number_field K] : is_fraction_ring (𝓞 K) K := -integral_closure.is_fraction_ring_of_finite_extension ℚ _ - -instance : is_integral_closure (𝓞 K) ℤ K := -integral_closure.is_integral_closure _ _ - -instance [number_field K] : is_integrally_closed (𝓞 K) := -integral_closure.is_integrally_closed_of_finite_extension ℚ - -lemma is_integral_coe (x : 𝓞 K) : is_integral ℤ (x : K) := -x.2 - -/-- The ring of integers of `K` are equivalent to any integral closure of `ℤ` in `K` -/ -protected noncomputable def equiv (R : Type*) [comm_ring R] [algebra R K] - [is_integral_closure R ℤ K] : 𝓞 K ≃+* R := -(is_integral_closure.equiv ℤ R K _).symm.to_ring_equiv - -variables (K) - -instance [number_field K] : char_zero (𝓞 K) := char_zero.of_module _ K - -instance [number_field K] : is_noetherian ℤ (𝓞 K) := is_integral_closure.is_noetherian _ ℚ K _ - -/-- The ring of integers of a number field is not a field. -/ -lemma not_is_field [number_field K] : ¬ is_field (𝓞 K) := -begin - have h_inj : function.injective ⇑(algebra_map ℤ (𝓞 K)), - { exact ring_hom.injective_int (algebra_map ℤ (𝓞 K)) }, - intro hf, - exact int.not_is_field - (((is_integral_closure.is_integral_algebra ℤ K).is_field_iff_is_field h_inj).mpr hf) -end - -instance [number_field K] : is_dedekind_domain (𝓞 K) := -is_integral_closure.is_dedekind_domain ℤ ℚ K _ - -end ring_of_integers - -end number_field - -namespace rat - -open number_field - -instance number_field : number_field ℚ := -{ to_char_zero := infer_instance, - to_finite_dimensional := - -- The vector space structure of `ℚ` over itself can arise in multiple ways: - -- all fields are vector spaces over themselves (used in `rat.finite_dimensional`) - -- all char 0 fields have a canonical embedding of `ℚ` (used in `number_field`). - -- Show that these coincide: - by convert (infer_instance : finite_dimensional ℚ ℚ), } - -/-- The ring of integers of `ℚ` as a number field is just `ℤ`. -/ -noncomputable def ring_of_integers_equiv : ring_of_integers ℚ ≃+* ℤ := -ring_of_integers.equiv ℤ - -end rat - -namespace adjoin_root - -section - -open_locale polynomial - -local attribute [-instance] algebra_rat - -/-- The quotient of `ℚ[X]` by the ideal generated by an irreducible polynomial of `ℚ[X]` -is a number field. -/ -instance {f : ℚ[X]} [hf : fact (irreducible f)] : number_field (adjoin_root f) := -{ to_char_zero := char_zero_of_injective_algebra_map (algebra_map ℚ _).injective, - to_finite_dimensional := by convert (adjoin_root.power_basis hf.out.ne_zero).finite_dimensional } -end - -end adjoin_root - namespace number_field.embeddings section fintype diff --git a/src/ring_theory/discriminant.lean b/src/ring_theory/discriminant.lean index 8e5e05041555a..1d851a10e5461 100644 --- a/src/ring_theory/discriminant.lean +++ b/src/ring_theory/discriminant.lean @@ -6,7 +6,7 @@ Authors: Riccardo Brasca import ring_theory.trace import ring_theory.norm -import number_theory.number_field +import number_theory.number_field.basic /-! # Discriminant of a family of vectors From 64b6d04ff460cf5cecb9a910227dbade57eea359 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 3 Oct 2022 14:28:34 +0000 Subject: [PATCH 525/828] feat(linear_algebra/affine_space/affine_subspace): `vector_span_insert_eq_vector_span` (#16765) Add a variant of `affine_span_insert_eq_affine_span` for the `vector_span`. --- src/linear_algebra/affine_space/affine_subspace.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index 062a94428f653..157374aa6a022 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -1161,6 +1161,14 @@ begin rw [←affine_span_insert_affine_span, set.insert_eq_of_mem h, affine_span_coe] end +variables {k} + +/-- If a point is in the affine span of a set, adding it to that set +does not change the vector span. -/ +lemma vector_span_insert_eq_vector_span {p : P} {ps : set P} (h : p ∈ affine_span k ps) : + vector_span k (insert p ps) = vector_span k ps := +by simp_rw [←direction_affine_span, affine_span_insert_eq_affine_span _ h] + end affine_space' namespace affine_subspace From 580e52cb7f3832b3ef5d1abe47cc29417e0a5e4f Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 3 Oct 2022 14:28:35 +0000 Subject: [PATCH 526/828] feat(analysis/convex/between): betweenness and `same_ray` (#16767) Add some lemmas about the relation between `wbtw` and `same_ray`. --- src/analysis/convex/between.lean | 53 ++++++++++++++++++++++++++++++++ 1 file changed, 53 insertions(+) diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index 4b03b63fea532..c0318ee8c8cff 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -5,6 +5,7 @@ Authors: Joseph Myers -/ import analysis.convex.segment import linear_algebra.affine_space.finite_dimensional +import tactic.field_simp /-! # Betweenness in affine spaces @@ -368,6 +369,40 @@ h₁.wbtw.trans_sbtw_right h₂ end ordered_ring +section ordered_comm_ring + +variables [ordered_comm_ring R] [add_comm_group V] [module R V] [add_torsor V P] + +include V + +variables {R} + +lemma wbtw.same_ray_vsub {x y z : P} (h : wbtw R x y z) : same_ray R (y -ᵥ x) (z -ᵥ y) := +begin + rcases h with ⟨t, ⟨ht0, ht1⟩, rfl⟩, + simp_rw line_map_apply, + rcases ht0.lt_or_eq with ht0' | rfl, swap, { simp }, + rcases ht1.lt_or_eq with ht1' | rfl, swap, { simp }, + refine or.inr (or.inr ⟨1 - t, t, sub_pos.2 ht1', ht0', _⟩), + simp [vsub_vadd_eq_vsub_sub, smul_sub, smul_smul, ←sub_smul], + ring_nf +end + +lemma wbtw.same_ray_vsub_left {x y z : P} (h : wbtw R x y z) : same_ray R (y -ᵥ x) (z -ᵥ x) := +begin + rcases h with ⟨t, ⟨ht0, ht1⟩, rfl⟩, + simpa [line_map_apply] using same_ray_nonneg_smul_left (z -ᵥ x) ht0 +end + +lemma wbtw.same_ray_vsub_right {x y z : P} (h : wbtw R x y z) : same_ray R (z -ᵥ x) (z -ᵥ y) := +begin + rcases h with ⟨t, ⟨ht0, ht1⟩, rfl⟩, + simpa [line_map_apply, vsub_vadd_eq_vsub_sub, sub_smul] using + same_ray_nonneg_smul_right (z -ᵥ x) (sub_nonneg.2 ht1) +end + +end ordered_comm_ring + section linear_ordered_field variables [linear_ordered_field R] [add_comm_group V] [module R V] [add_torsor V P] @@ -497,4 +532,22 @@ begin exact or.inl (wbtw_or_wbtw_smul_vadd_of_nonneg _ _ hy0.le hz0.le) } } end +lemma wbtw_iff_same_ray_vsub {x y z : P} : wbtw R x y z ↔ same_ray R (y -ᵥ x) (z -ᵥ y) := +begin + refine ⟨wbtw.same_ray_vsub, λ h, _⟩, + rcases h with h | h | ⟨r₁, r₂, hr₁, hr₂, h⟩, + { rw vsub_eq_zero_iff_eq at h, simp [h] }, + { rw vsub_eq_zero_iff_eq at h, simp [h] }, + { refine ⟨r₂ / (r₁ + r₂), + ⟨div_nonneg hr₂.le (add_nonneg hr₁.le hr₂.le), + div_le_one_of_le (le_add_of_nonneg_left hr₁.le) (add_nonneg hr₁.le hr₂.le)⟩, _⟩, + have h' : z = r₂⁻¹ • r₁ • (y -ᵥ x) +ᵥ y, { simp [h, hr₂.ne'] }, + rw eq_comm, + simp only [line_map_apply, h', vadd_vsub_assoc, smul_smul, ←add_smul, eq_vadd_iff_vsub_eq, + smul_add], + convert (one_smul _ _).symm, + field_simp [(add_pos hr₁ hr₂).ne', hr₂.ne'], + ring } +end + end linear_ordered_field From f1f9d66a4b37d4139671e20bac83387e9a5c565a Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 3 Oct 2022 14:28:37 +0000 Subject: [PATCH 527/828] feat(analysis/convex/between): `wbtw` and distinct end points (#16768) Add lemmas that `wbtw`, together with the middle point being distinct from one of the end points, implies the two end points are distinct, and use one of those lemmas in the proof of `sbtw.left_ne_right`. --- src/analysis/convex/between.lean | 16 +++++++++++++--- 1 file changed, 13 insertions(+), 3 deletions(-) diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index c0318ee8c8cff..86e7d20ebbcf0 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -242,13 +242,23 @@ end variables {R} -lemma sbtw.left_ne_right {x y z : P} (h : sbtw R x y z) : x ≠ z := +lemma wbtw.left_ne_right_of_ne_left {x y z : P} (h : wbtw R x y z) (hne : y ≠ x) : x ≠ z := begin rintro rfl, - rw [sbtw, wbtw_self_iff] at h, - exact h.2.1 h.1 + rw wbtw_self_iff at h, + exact hne h end +lemma wbtw.left_ne_right_of_ne_right {x y z : P} (h : wbtw R x y z) (hne : y ≠ z) : x ≠ z := +begin + rintro rfl, + rw wbtw_self_iff at h, + exact hne h +end + +lemma sbtw.left_ne_right {x y z : P} (h : sbtw R x y z) : x ≠ z := +h.wbtw.left_ne_right_of_ne_left h.2.1 + lemma sbtw_iff_mem_image_Ioo_and_ne [no_zero_smul_divisors R V] {x y z : P} : sbtw R x y z ↔ y ∈ line_map x z '' (set.Ioo (0 : R) 1) ∧ x ≠ z := begin From eb9738966c73ad83d7fcdb817ab74262f9e4e3df Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 3 Oct 2022 14:28:38 +0000 Subject: [PATCH 528/828] feat(linear_algebra/affine_space/affine_subspace): spans of two points (#16770) Add lemmas about the `vector_span` and `affine_span` of two points (analogous to `span_singleton` lemmas for `submodule`). --- .../affine_space/affine_subspace.lean | 92 +++++++++++++++++++ 1 file changed, 92 insertions(+) diff --git a/src/linear_algebra/affine_space/affine_subspace.lean b/src/linear_algebra/affine_space/affine_subspace.lean index 157374aa6a022..2eaf8026104f7 100644 --- a/src/linear_algebra/affine_space/affine_subspace.lean +++ b/src/linear_algebra/affine_space/affine_subspace.lean @@ -1140,6 +1140,98 @@ end variables (k) +/-- The `vector_span` of two points is the span of their difference. -/ +lemma vector_span_pair (p₁ p₂ : P) : vector_span k ({p₁, p₂} : set P) = k ∙ (p₁ -ᵥ p₂) := +by rw [vector_span_eq_span_vsub_set_left k (mem_insert p₁ _), image_pair, vsub_self, + submodule.span_insert_zero] + +/-- The `vector_span` of two points is the span of their difference (reversed). -/ +lemma vector_span_pair_rev (p₁ p₂ : P) : vector_span k ({p₁, p₂} : set P) = k ∙ (p₂ -ᵥ p₁) := +by rw [pair_comm, vector_span_pair] + +/-- The difference between two points lies in their `vector_span`. -/ +lemma vsub_mem_vector_span_pair (p₁ p₂ : P) : p₁ -ᵥ p₂ ∈ vector_span k ({p₁, p₂} : set P) := +vsub_mem_vector_span _ (set.mem_insert _ _) (set.mem_insert_of_mem _ (set.mem_singleton _)) + +/-- The difference between two points (reversed) lies in their `vector_span`. -/ +lemma vsub_rev_mem_vector_span_pair (p₁ p₂ : P) : p₂ -ᵥ p₁ ∈ vector_span k ({p₁, p₂} : set P) := +vsub_mem_vector_span _ (set.mem_insert_of_mem _ (set.mem_singleton _)) (set.mem_insert _ _) + +variables {k} + +/-- A multiple of the difference between two points lies in their `vector_span`. -/ +lemma smul_vsub_mem_vector_span_pair (r : k) (p₁ p₂ : P) : + r • (p₁ -ᵥ p₂) ∈ vector_span k ({p₁, p₂} : set P) := +submodule.smul_mem _ _ (vsub_mem_vector_span_pair k p₁ p₂) + +/-- A multiple of the difference between two points (reversed) lies in their `vector_span`. -/ +lemma smul_vsub_rev_mem_vector_span_pair (r : k) (p₁ p₂ : P) : + r • (p₂ -ᵥ p₁) ∈ vector_span k ({p₁, p₂} : set P) := +submodule.smul_mem _ _ (vsub_rev_mem_vector_span_pair k p₁ p₂) + +/-- A vector lies in the `vector_span` of two points if and only if it is a multiple of their +difference. -/ +lemma mem_vector_span_pair {p₁ p₂ : P} {v : V} : + v ∈ vector_span k ({p₁, p₂} : set P) ↔ ∃ r : k, r • (p₁ -ᵥ p₂) = v := +by rw [vector_span_pair, submodule.mem_span_singleton] + +/-- A vector lies in the `vector_span` of two points if and only if it is a multiple of their +difference (reversed). -/ +lemma mem_vector_span_pair_rev {p₁ p₂ : P} {v : V} : + v ∈ vector_span k ({p₁, p₂} : set P) ↔ ∃ r : k, r • (p₂ -ᵥ p₁) = v := +by rw [vector_span_pair_rev, submodule.mem_span_singleton] + +variables (k) + +/-- The first of two points lies in their affine span. -/ +lemma left_mem_affine_span_pair (p₁ p₂ : P) : p₁ ∈ affine_span k ({p₁, p₂} : set P) := +mem_affine_span _ (set.mem_insert _ _) + +/-- The second of two points lies in their affine span. -/ +lemma right_mem_affine_span_pair (p₁ p₂ : P) : p₂ ∈ affine_span k ({p₁, p₂} : set P) := +mem_affine_span _ (set.mem_insert_of_mem _ (set.mem_singleton _)) + +variables {k} + +/-- A combination of two points expressed with `line_map` lies in their affine span. -/ +lemma affine_map.line_map_mem_affine_span_pair (r : k) (p₁ p₂ : P) : + affine_map.line_map p₁ p₂ r ∈ affine_span k ({p₁, p₂} : set P) := +affine_map.line_map_mem _ (left_mem_affine_span_pair _ _ _) (right_mem_affine_span_pair _ _ _) + +/-- A combination of two points expressed with `line_map` (with the two points reversed) lies in +their affine span. -/ +lemma affine_map.line_map_rev_mem_affine_span_pair (r : k) (p₁ p₂ : P) : + affine_map.line_map p₂ p₁ r ∈ affine_span k ({p₁, p₂} : set P) := +affine_map.line_map_mem _ (right_mem_affine_span_pair _ _ _) (left_mem_affine_span_pair _ _ _) + +/-- A multiple of the difference of two points added to the first point lies in their affine +span. -/ +lemma smul_vsub_vadd_mem_affine_span_pair (r : k) (p₁ p₂ : P) : + r • (p₂ -ᵥ p₁) +ᵥ p₁ ∈ affine_span k ({p₁, p₂} : set P) := +affine_map.line_map_mem_affine_span_pair _ _ _ + +/-- A multiple of the difference of two points added to the second point lies in their affine +span. -/ +lemma smul_vsub_rev_vadd_mem_affine_span_pair (r : k) (p₁ p₂ : P) : + r • (p₁ -ᵥ p₂) +ᵥ p₂ ∈ affine_span k ({p₁, p₂} : set P) := +affine_map.line_map_rev_mem_affine_span_pair _ _ _ + +/-- A vector added to the first point lies in the affine span of two points if and only if it is +a multiple of their difference. -/ +lemma vadd_left_mem_affine_span_pair {p₁ p₂ : P} {v : V} : + v +ᵥ p₁ ∈ affine_span k ({p₁, p₂} : set P) ↔ ∃ r : k, r • (p₂ -ᵥ p₁) = v := +by rw [vadd_mem_iff_mem_direction _ (left_mem_affine_span_pair _ _ _), direction_affine_span, + mem_vector_span_pair_rev] + +/-- A vector added to the second point lies in the affine span of two points if and only if it is +a multiple of their difference. -/ +lemma vadd_right_mem_affine_span_pair {p₁ p₂ : P} {v : V} : + v +ᵥ p₂ ∈ affine_span k ({p₁, p₂} : set P) ↔ ∃ r : k, r • (p₁ -ᵥ p₂) = v := +by rw [vadd_mem_iff_mem_direction _ (right_mem_affine_span_pair _ _ _), direction_affine_span, + mem_vector_span_pair] + +variables (k) + /-- `affine_span` is monotone. -/ @[mono] lemma affine_span_mono {s₁ s₂ : set P} (h : s₁ ⊆ s₂) : affine_span k s₁ ≤ affine_span k s₂ := From ee02a30e209a2a77b93eac1254e8c66e76192f54 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 3 Oct 2022 17:40:43 +0000 Subject: [PATCH 529/828] feat(set_theory/surreal/basic): relabellings preserve being numeric (#14406) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- src/set_theory/surreal/basic.lean | 16 +++++++++++++++- 1 file changed, 15 insertions(+), 1 deletion(-) diff --git a/src/set_theory/surreal/basic.lean b/src/set_theory/surreal/basic.lean index 81cd1cd193329..13d1fb908e27b 100644 --- a/src/set_theory/surreal/basic.lean +++ b/src/set_theory/surreal/basic.lean @@ -54,7 +54,7 @@ The branch `surreal_mul` contains some progress on this proof. universes u -local infix ` ⧏ `:50 := pgame.lf +open_locale pgame namespace pgame @@ -95,6 +95,20 @@ theorem numeric_rec {C : pgame → Prop} | ⟨l, r, L, R⟩ ⟨h, hl, hr⟩ := H _ _ _ _ h hl hr (λ i, numeric_rec _ (hl i)) (λ i, numeric_rec _ (hr i)) +theorem relabelling.numeric_imp {x y : pgame} (r : x ≡r y) (ox : numeric x) : numeric y := +begin + induction x using pgame.move_rec_on with x IHl IHr generalizing y, + apply numeric.mk (λ i j, _) (λ i, _) (λ j, _), + { rw ←lt_congr (r.move_left_symm i).equiv (r.move_right_symm j).equiv, + apply ox.left_lt_right }, + { exact IHl _ (ox.move_left _) (r.move_left_symm i) }, + { exact IHr _ (ox.move_right _) (r.move_right_symm j) } +end + +/-- Relabellings preserve being numeric. -/ +theorem relabelling.numeric_congr {x y : pgame} (r : x ≡r y) : numeric x ↔ numeric y := +⟨r.numeric_imp, r.symm.numeric_imp⟩ + theorem lf_asymm {x y : pgame} (ox : numeric x) (oy : numeric y) : x ⧏ y → ¬ y ⧏ x := begin refine numeric_rec (λ xl xr xL xR hx oxl oxr IHxl IHxr, _) x ox y oy, From a347076985674932c0e91da09b9961ed0a79508c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 3 Oct 2022 17:40:45 +0000 Subject: [PATCH 530/828] feat(set_theory/game/birthday): birthday of addition equals natural addition of birthdays (#14549) --- src/order/lattice.lean | 10 +++++- src/set_theory/game/birthday.lean | 49 ++++++++++++++++++++------ src/set_theory/ordinal/arithmetic.lean | 20 +++++++++++ 3 files changed, 68 insertions(+), 11 deletions(-) diff --git a/src/order/lattice.lean b/src/order/lattice.lean index 66877bbc6e257..7a6ea3f91cb96 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -626,7 +626,7 @@ instance linear_order.to_lattice {α : Type u} [o : linear_order α] : ..o } section linear_order -variables [linear_order α] {a b c : α} +variables [linear_order α] {a b c d : α} lemma sup_eq_max : a ⊔ b = max a b := rfl lemma inf_eq_min : a ⊓ b = min a b := rfl @@ -655,6 +655,14 @@ lemma inf_ind (a b : α) {p : α → Prop} : p a → p b → p (a ⊓ b) := @sup @[simp] lemma inf_lt_iff : b ⊓ c < a ↔ b < a ∨ c < a := @lt_sup_iff αᵒᵈ _ _ _ _ @[simp] lemma lt_inf_iff : a < b ⊓ c ↔ a < b ∧ a < c := @sup_lt_iff αᵒᵈ _ _ _ _ +variables (a b c d) + +lemma max_max_max_comm : max (max a b) (max c d) = max (max a c) (max b d) := +sup_sup_sup_comm _ _ _ _ + +lemma min_min_min_comm : min (min a b) (min c d) = min (min a c) (min b d) := +inf_inf_inf_comm _ _ _ _ + end linear_order lemma sup_eq_max_default [semilattice_sup α] [decidable_rel ((≤) : α → α → Prop)] diff --git a/src/set_theory/game/birthday.lean b/src/set_theory/game/birthday.lean index d8d1345df6ff3..c9b2642f8e8b8 100644 --- a/src/set_theory/game/birthday.lean +++ b/src/set_theory/game/birthday.lean @@ -5,6 +5,7 @@ Authors: Violeta Hernández Palacios -/ import set_theory.game.ordinal +import set_theory.ordinal.natural_ops /-! # Birthdays of games @@ -27,7 +28,7 @@ universe u open ordinal -open_locale pgame +open_locale natural_ops pgame namespace pgame @@ -83,13 +84,7 @@ theorem relabelling.birthday_congr : ∀ {x y : pgame.{u}}, x ≡r y → birthda end using_well_founded { dec_tac := pgame_wf_tac } -@[simp] theorem birthday_add_zero (x : pgame) : birthday (x + 0) = birthday x := -(add_zero_relabelling x).birthday_congr - -@[simp] theorem birthday_zero_add (x : pgame) : birthday (0 + x) = birthday x := -(zero_add_relabelling x).birthday_congr - -@[simp] theorem birthday_eq_zero (x : pgame) : +@[simp] theorem birthday_eq_zero {x : pgame} : birthday x = 0 ↔ is_empty x.left_moves ∧ is_empty x.right_moves := by rw [birthday_def, max_eq_zero, lsub_eq_zero_iff, lsub_eq_zero_iff] @@ -123,7 +118,41 @@ theorem le_birthday : ∀ x : pgame, x ≤ x.birthday.to_pgame le_def.2 ⟨λ i, or.inl ⟨to_left_moves_to_pgame ⟨_, birthday_move_left_lt i⟩, by simp [le_birthday (xL i)]⟩, is_empty_elim⟩ -theorem neg_birthday_le (x : pgame) : -x.birthday.to_pgame ≤ x := -let h := le_birthday (-x) in by rwa [neg_birthday, neg_le_iff] at h +variables (a b x : pgame.{u}) + +theorem neg_birthday_le : -x.birthday.to_pgame ≤ x := +by simpa only [neg_birthday, ←neg_le_iff] using le_birthday (-x) + +@[simp] theorem birthday_add : ∀ x y : pgame.{u}, (x + y).birthday = x.birthday ♯ y.birthday +| ⟨xl, xr, xL, xR⟩ ⟨yl, yr, yL, yR⟩ := begin + rw [birthday_def, nadd_def], + simp only [birthday_add, lsub_sum, mk_add_move_left_inl, move_left_mk, mk_add_move_left_inr, + mk_add_move_right_inl, move_right_mk, mk_add_move_right_inr], + rw max_max_max_comm, + congr; apply le_antisymm, + any_goals + { exact max_le_iff.2 ⟨lsub_le_iff.2 (λ i, lt_blsub _ _ (birthday_move_left_lt i)), + lsub_le_iff.2 (λ i, lt_blsub _ _ (birthday_move_right_lt i))⟩ }, + all_goals + { apply blsub_le_iff.2 (λ i hi, _), + rcases lt_birthday_iff.1 hi with ⟨j, hj⟩ | ⟨j, hj⟩ }, + { exact lt_max_of_lt_left ((nadd_le_nadd_right hj _).trans_lt (lt_lsub _ _)) }, + { exact lt_max_of_lt_right ((nadd_le_nadd_right hj _).trans_lt (lt_lsub _ _)) }, + { exact lt_max_of_lt_left ((nadd_le_nadd_left hj _).trans_lt (lt_lsub _ _)) }, + { exact lt_max_of_lt_right ((nadd_le_nadd_left hj _).trans_lt (lt_lsub _ _)) } +end +using_well_founded { dec_tac := pgame_wf_tac } + +theorem birthday_add_zero : (a + 0).birthday = a.birthday := by simp +theorem birthday_zero_add : (0 + a).birthday = a.birthday := by simp +theorem birthday_add_one : (a + 1).birthday = order.succ a.birthday := by simp +theorem birthday_one_add : (1 + a).birthday = order.succ a.birthday := by simp + +@[simp] theorem birthday_nat_cast : ∀ n : ℕ, birthday n = n +| 0 := birthday_zero +| (n + 1) := by simp [birthday_nat_cast] + +theorem birthday_add_nat (n : ℕ) : (a + n).birthday = a.birthday + n := by simp +theorem birthday_nat_add (n : ℕ) : (↑n + a).birthday = a.birthday + n := by simp end pgame diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index cde5a7d017013..ebcc64e1c5e85 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -1086,6 +1086,19 @@ theorem sup_eq_of_range_eq {ι ι'} {f : ι → ordinal} {g : ι' → ordinal} (h : set.range f = set.range g) : sup.{u (max v w)} f = sup.{v (max u w)} g := (sup_le_of_range_subset h.le).antisymm (sup_le_of_range_subset.{v u w} h.ge) +@[simp] theorem sup_sum {α : Type u} {β : Type v} (f : α ⊕ β → ordinal) : sup.{(max u v) w} f = + max (sup.{u (max v w)} (λ a, f (sum.inl a))) (sup.{v (max u w)} (λ b, f (sum.inr b))) := +begin + apply (sup_le_iff.2 _).antisymm (max_le_iff.2 ⟨_, _⟩), + { rintro (i|i), + { exact le_max_of_le_left (le_sup _ i) }, + { exact le_max_of_le_right (le_sup _ i) } }, + all_goals + { apply sup_le_of_range_subset.{_ (max u v) w}, + rintros i ⟨a, rfl⟩, + apply mem_range_self } +end + lemma unbounded_range_of_sup_ge {α β : Type u} (r : α → α → Prop) [is_well_order α r] (f : β → α) (h : type r ≤ sup.{u u} (typein r ∘ f)) : unbounded r (range f) := (not_bounded_iff _).1 $ λ ⟨x, hx⟩, not_lt_of_le h $ lt_of_le_of_lt @@ -1333,6 +1346,10 @@ theorem lsub_eq_of_range_eq {ι ι'} {f : ι → ordinal} {g : ι' → ordinal} (h : set.range f = set.range g) : lsub.{u (max v w)} f = lsub.{v (max u w)} g := (lsub_le_of_range_subset h.le).antisymm (lsub_le_of_range_subset.{v u w} h.ge) +@[simp] theorem lsub_sum {α : Type u} {β : Type v} (f : α ⊕ β → ordinal) : lsub.{(max u v) w} f = + max (lsub.{u (max v w)} (λ a, f (sum.inl a))) (lsub.{v (max u w)} (λ b, f (sum.inr b))) := +sup_sum _ + theorem lsub_not_mem_range {ι} (f : ι → ordinal) : lsub f ∉ set.range f := λ ⟨i, h⟩, h.not_lt (lt_lsub f i) @@ -2164,6 +2181,9 @@ end /-! ### Casting naturals into ordinals, compatibility with operations -/ +@[simp] theorem one_add_nat_cast (m : ℕ) : 1 + (m : ordinal) = succ m := +by { rw [←nat.cast_one, ←nat.cast_add, add_comm], refl } + @[simp, norm_cast] theorem nat_cast_mul (m : ℕ) : ∀ n : ℕ, ((m * n : ℕ) : ordinal) = m * n | 0 := by simp | (n+1) := by rw [nat.mul_succ, nat.cast_add, nat_cast_mul, nat.cast_succ, mul_add_one] From 9aec1a288c6222ef5ccfd02a07155e9dd37b20de Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 3 Oct 2022 17:40:46 +0000 Subject: [PATCH 531/828] feat(ring_theory/dedekind_domain/selmer_group): add Selmer groups of Dedekind domains (#15405) --- .../dedekind_domain/selmer_group.lean | 223 ++++++++++++++++++ 1 file changed, 223 insertions(+) create mode 100644 src/ring_theory/dedekind_domain/selmer_group.lean diff --git a/src/ring_theory/dedekind_domain/selmer_group.lean b/src/ring_theory/dedekind_domain/selmer_group.lean new file mode 100644 index 0000000000000..dce7f6f6cf459 --- /dev/null +++ b/src/ring_theory/dedekind_domain/selmer_group.lean @@ -0,0 +1,223 @@ +/- +Copyright (c) 2022 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata +-/ + +import data.zmod.quotient +import ring_theory.dedekind_domain.adic_valuation + +/-! +# Selmer groups of fraction fields of Dedekind domains + +Let $K$ be the field of fractions of a Dedekind domain $R$. For any set $S$ of prime ideals in the +height one spectrum of $R$, and for any natural number $n$, the Selmer group $K(S, n)$ is defined to +be the subgroup of the unit group $K^\times$ modulo $n$-th powers where each element has $v$-adic +valuation divisible by $n$ for all prime ideals $v$ away from $S$. In other words, this is precisely +$$ K(S, n) := \{x(K^\times)^n \in K^\times / (K^\times)^n \ \mid \ + \forall v \notin S, \ \mathrm{ord}_v(x) \equiv 0 \pmod n\}. $$ + +There is a fundamental short exact sequence +$$ 1 \to R_S^\times / (R_S^\times)^n \to K(S, n) \to \mathrm{Cl}_S(R)[n] \to 0, $$ +where $R_S^\times$ is the $S$-unit group of $R$ and $\mathrm{Cl}_S(R)$ is the $S$-class group of +$R$. If the flanking groups are both finite, then $K(S, n)$ is finite by the first isomorphism +theorem. Such is the case when $R$ is the ring of integers of a number field $K$, $S$ is finite, and +$n$ is positive, in which case $R_S^\times$ is finitely generated by Dirichlet's unit theorem and +$\mathrm{Cl}_S(R)$ is finite by the class number theorem. + +This file defines the Selmer group $K(S, n)$ and some basic facts. + +## Main definitions + + * `is_dedekind_domain.selmer_group`: the Selmer group. + * TODO: maps in the sequence. + +## Main statements + + * TODO: proofs of exactness of the sequence. + * TODO: proofs of finiteness for global fields. + +## Notations + + * `K⟮S, n⟯`: the Selmer group with parameters `K`, `S`, and `n`. + +## Implementation notes + +The Selmer group is typically defined as a subgroup of the Galois cohomology group $H^1(K, \mu_n)$ +with certain local conditions defined by $v$-adic valuations, where $\mu_n$ is the group of $n$-th +roots of unity over a separable closure of $K$. Here $H^1(K, \mu_n)$ is identified with +$K^\times / (K^\times)^n$ by the long exact sequence from Kummer theory and Hilbert's theorem 90, +and the fundamental short exact sequence becomes an easy consequence of the snake lemma. This file +will define all the maps explicitly for computational purposes, but isomorphisms to the Galois +cohomological definition will be provided when possible. + +## References + +https://doc.sagemath.org/html/en/reference/number_fields/sage/rings/number_field/selmer_group.html + +## Tags + +class group, selmer group, unit group +-/ + +local notation (name := quot) K/n := Kˣ ⧸ (pow_monoid_hom n : Kˣ →* Kˣ).range + +namespace is_dedekind_domain + +noncomputable theory + +open_locale classical discrete_valuation non_zero_divisors + +universes u v + +variables {R : Type u} [comm_ring R] [is_domain R] [is_dedekind_domain R] {K : Type v} [field K] + [algebra R K] [is_fraction_ring R K] (v : height_one_spectrum R) + +/-! ### Valuations of non-zero elements -/ + +namespace height_one_spectrum + +/-- The multiplicative `v`-adic valuation on `Kˣ`. -/ +def valuation_of_ne_zero_to_fun (x : Kˣ) : multiplicative ℤ := +let hx := is_localization.sec R⁰ (x : K) in multiplicative.of_add $ + (-(associates.mk v.as_ideal).count (associates.mk $ ideal.span {hx.fst}).factors : ℤ) + - (-(associates.mk v.as_ideal).count (associates.mk $ ideal.span {(hx.snd : R)}).factors : ℤ) + +@[simp] lemma valuation_of_ne_zero_to_fun_eq (x : Kˣ) : + (v.valuation_of_ne_zero_to_fun x : ℤₘ₀) = v.valuation (x : K) := +begin + change _ = _ * _, + rw [units.coe_inv], + change _ = ite _ _ _ * (ite (coe _ = _) _ _)⁻¹, + rw [is_localization.to_localization_map_sec, + if_neg $ is_localization.sec_fst_ne_zero le_rfl x.ne_zero, + if_neg $ non_zero_divisors.coe_ne_zero _], + any_goals { exact is_domain.to_nontrivial R }, + refl +end + +/-- The multiplicative `v`-adic valuation on `Kˣ`. -/ +def valuation_of_ne_zero : Kˣ →* multiplicative ℤ := +{ to_fun := v.valuation_of_ne_zero_to_fun, + map_one' := by { rw [← with_zero.coe_inj, valuation_of_ne_zero_to_fun_eq], exact map_one _ }, + map_mul' := λ _ _, by { rw [← with_zero.coe_inj, with_zero.coe_mul], + simp only [valuation_of_ne_zero_to_fun_eq], exact map_mul _ _ _ } } + +@[simp] lemma valuation_of_ne_zero_eq (x : Kˣ) : + (v.valuation_of_ne_zero x : ℤₘ₀) = v.valuation (x : K) := +valuation_of_ne_zero_to_fun_eq v x + +@[simp] lemma valuation_of_unit_eq (x : Rˣ) : + v.valuation_of_ne_zero (units.map (algebra_map R K : R →* K) x) = 1 := +begin + rw [← with_zero.coe_inj, valuation_of_ne_zero_eq, units.coe_map, eq_iff_le_not_lt], + split, + { exact v.valuation_le_one x }, + { cases x with x _ hx _, + change ¬v.valuation (algebra_map R K x) < 1, + apply_fun v.int_valuation at hx, + rw [map_one, map_mul] at hx, + rw [not_lt, ← hx, ← mul_one $ v.valuation _, valuation_of_algebra_map, + mul_le_mul_left₀ $ left_ne_zero_of_mul_eq_one hx], + exact v.int_valuation_le_one _ } +end + +local attribute [semireducible] mul_opposite + +/-- The multiplicative `v`-adic valuation on `Kˣ` modulo `n`-th powers. -/ +def valuation_of_ne_zero_mod (n : ℕ) : K/n →* multiplicative (zmod n) := +(int.quotient_zmultiples_nat_equiv_zmod n).to_multiplicative.to_monoid_hom.comp $ + quotient_group.map (pow_monoid_hom n : Kˣ →* Kˣ).range + (add_subgroup.zmultiples (n : ℤ)).to_subgroup v.valuation_of_ne_zero +begin + rintro _ ⟨x, rfl⟩, + exact ⟨v.valuation_of_ne_zero x, by simpa only [pow_monoid_hom_apply, map_pow, int.to_add_pow]⟩ +end + +@[simp] lemma valuation_of_unit_mod_eq (n : ℕ) (x : Rˣ) : + v.valuation_of_ne_zero_mod n (units.map (algebra_map R K : R →* K) x : K/n) = 1 := +by rw [valuation_of_ne_zero_mod, monoid_hom.comp_apply, ← quotient_group.coe_mk', + quotient_group.map_mk', valuation_of_unit_eq, quotient_group.coe_one, map_one] + +end height_one_spectrum + +/-! ### Selmer groups -/ + +variables {S S' : set $ height_one_spectrum R} {n : ℕ} + +/-- The Selmer group `K⟮S, n⟯`. -/ +def selmer_group : subgroup $ K/n := +{ carrier := {x : K/n | ∀ v ∉ S, (v : height_one_spectrum R).valuation_of_ne_zero_mod n x = 1}, + one_mem' := λ _ _, by rw [map_one], + mul_mem' := λ _ _ hx hy v hv, by rw [map_mul, hx v hv, hy v hv, one_mul], + inv_mem' := λ _ hx v hv, by rw [map_inv, hx v hv, inv_one] } + +localized "notation K`⟮`S, n`⟯` := @selmer_group _ _ _ _ K _ _ _ S n" in selmer_group + +namespace selmer_group + +lemma monotone (hS : S ≤ S') : K⟮S, n⟯ ≤ (K⟮S', n⟯) := λ _ hx v, hx v ∘ mt (@hS v) + +/-- The multiplicative `v`-adic valuations on `K⟮S, n⟯` for all `v ∈ S`. -/ +def valuation : K⟮S, n⟯ →* S → multiplicative (zmod n) := +{ to_fun := λ x v, (v : height_one_spectrum R).valuation_of_ne_zero_mod n (x : K/n), + map_one' := funext $ λ v, map_one _, + map_mul' := λ x y, funext $ λ v, map_mul _ x y } + +lemma valuation_ker_eq : + valuation.ker = (K⟮(∅ : set $ height_one_spectrum R), n⟯).subgroup_of (K⟮S, n⟯) := +begin + ext ⟨_, hx⟩, + split, + { intros hx' v _, + by_cases hv : v ∈ S, + { exact congr_fun hx' ⟨v, hv⟩ }, + { exact hx v hv } }, + { exact λ hx', funext $ λ v, hx' v $ set.not_mem_empty v } +end + +/-- The natural homomorphism from `Rˣ` to `K⟮∅, n⟯`. -/ +def from_unit {n : ℕ} : Rˣ →* K⟮(∅ : set $ height_one_spectrum R), n⟯ := +{ to_fun := λ x, ⟨quotient_group.mk $ units.map (algebra_map R K).to_monoid_hom x, + λ v _, v.valuation_of_unit_mod_eq n x⟩, + map_one' := by simpa only [map_one], + map_mul' := λ _ _, by simpa only [map_mul] } + +lemma from_unit_ker [hn : fact $ 0 < n] : + (@from_unit R _ _ _ K _ _ _ n).ker = (pow_monoid_hom n : Rˣ →* Rˣ).range := +begin + ext ⟨_, _, _, _⟩, + split, + { intro hx, + rcases (quotient_group.eq_one_iff _).mp (subtype.mk.inj hx) with ⟨⟨v, i, vi, iv⟩, hx⟩, + have hv : ↑(_ ^ n : Kˣ) = algebra_map R K _ := congr_arg units.val hx, + have hi : ↑(_ ^ n : Kˣ)⁻¹ = algebra_map R K _ := congr_arg units.inv hx, + rw [units.coe_pow] at hv, + rw [← inv_pow, units.inv_mk, units.coe_pow] at hi, + rcases @is_integrally_closed.exists_algebra_map_eq_of_is_integral_pow R _ _ _ _ _ _ _ v _ + hn.out (hv.symm ▸ is_integral_algebra_map) with ⟨v', rfl⟩, + rcases @is_integrally_closed.exists_algebra_map_eq_of_is_integral_pow R _ _ _ _ _ _ _ i _ + hn.out (hi.symm ▸ is_integral_algebra_map) with ⟨i', rfl⟩, + rw [← map_mul, map_eq_one_iff _ $ no_zero_smul_divisors.algebra_map_injective R K] at vi, + rw [← map_mul, map_eq_one_iff _ $ no_zero_smul_divisors.algebra_map_injective R K] at iv, + rw [units.coe_mk, ← map_pow] at hv, + exact ⟨⟨v', i', vi, iv⟩, by simpa only [units.ext_iff, pow_monoid_hom_apply, units.coe_pow] + using no_zero_smul_divisors.algebra_map_injective R K hv⟩ }, + { rintro ⟨_, hx⟩, + rw [← hx], + exact subtype.mk_eq_mk.mpr + ((quotient_group.eq_one_iff _).mpr ⟨_, by simp only [pow_monoid_hom_apply, map_pow]⟩) } +end + +/-- The injection induced by the natural homomorphism from `Rˣ` to `K⟮∅, n⟯`. -/ +def from_unit_lift [fact $ 0 < n] : R/n →* K⟮(∅ : set $ height_one_spectrum R), n⟯ := +(quotient_group.ker_lift _).comp + (quotient_group.quotient_mul_equiv_of_eq from_unit_ker).symm.to_monoid_hom + +lemma from_unit_lift_injective [fact $ 0 < n] : + function.injective $ @from_unit_lift R _ _ _ K _ _ _ n _ := +function.injective.comp (quotient_group.ker_lift_injective _) (mul_equiv.injective _) + +end selmer_group + +end is_dedekind_domain From 77615d00fbf05cd7e55ba84176070a18bc097772 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 3 Oct 2022 17:40:47 +0000 Subject: [PATCH 532/828] =?UTF-8?q?chore(data/seq/*):=20`cases=5Fon`=20?= =?UTF-8?q?=E2=86=92=20`rec=5Fon`=20(#15843)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We rename recursion principles on `Sort` from `cases_on` to `rec_on`. Co-authored-by: Yaël Dillies --- scripts/nolints.txt | 2 +- src/data/seq/computation.lean | 25 +++++++++-------- src/data/seq/parallel.lean | 2 +- src/data/seq/seq.lean | 53 ++++++++++++++++++----------------- src/data/seq/wseq.lean | 48 ++++++++++++++++--------------- 5 files changed, 67 insertions(+), 63 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 53a65f47e0805..923931412d535 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -933,4 +933,4 @@ apply_nolint transfer.compute_transfer doc_blame apply_nolint tactic.wlog doc_blame -- topology/algebra/group.lean -apply_nolint topological_space.positive_compacts.locally_compact_space_of_group to_additive_doc \ No newline at end of file +apply_nolint topological_space.positive_compacts.locally_compact_space_of_group to_additive_doc diff --git a/src/data/seq/computation.lean b/src/data/seq/computation.lean index e1324f9536a29..8eed089c5e21e 100644 --- a/src/data/seq/computation.lean +++ b/src/data/seq/computation.lean @@ -128,7 +128,8 @@ by cases s with f al; apply subtype.eq; dsimp [tail, think]; rw [stream.tail_con theorem think_empty : empty α = think (empty α) := destruct_eq_think destruct_empty -def cases_on {C : computation α → Sort v} (s : computation α) +/-- Recursion principle for computations, compare with `list.rec_on`. -/ +def rec_on {C : computation α → Sort v} (s : computation α) (h1 : ∀ a, C (return a)) (h2 : ∀ s, C (think s)) : C s := begin induction H : destruct s with v v, { rw destruct_eq_ret H, apply h1 }, @@ -212,7 +213,7 @@ section bisim by cases s; refl, by cases s'; refl, r⟩) this, begin have := bisim r, revert r this, - apply cases_on s _ _; intros; apply cases_on s' _ _; intros; intros r this, + apply rec_on s _ _; intros; apply rec_on s' _ _; intros; intros r this, { constructor, dsimp at this, rw this, assumption }, { rw [destruct_ret, destruct_think] at this, exact false.elim this }, @@ -441,7 +442,7 @@ theorem eq_thinkN {s : computation α} {a n} (h : results s a n) : begin revert s, induction n with n IH; intro s; - apply cases_on s (λ a', _) (λ s, _); intro h, + apply rec_on s (λ a', _) (λ s, _); intro h, { rw ←eq_of_ret_mem h.mem, refl }, { cases of_results_think h with n h, cases h, contradiction }, { have := h.len_unique (results_ret _), contradiction }, @@ -506,7 +507,7 @@ def join (c : computation (computation α)) : computation α := c >>= id @[simp] theorem destruct_map (f : α → β) (s) : destruct (map f s) = lmap f (rmap (map f) (destruct s)) := -by apply s.cases_on; intro; simp +by apply s.rec_on; intro; simp @[simp] theorem map_id : ∀ (s : computation α), map id s = s | ⟨f, al⟩ := begin @@ -556,7 +557,7 @@ begin exact match c₁, c₂, h with | _, _, or.inl (eq.refl c) := begin cases destruct c with b cb; simp end | _, _, or.inr ⟨s, rfl, rfl⟩ := begin - apply cases_on s; intros s; simp, + apply rec_on s; intros s; simp, exact or.inr ⟨s, rfl, rfl⟩ end end }, { exact or.inr ⟨s, rfl, rfl⟩ } @@ -574,9 +575,9 @@ begin exact match c₁, c₂, h with | _, _, or.inl (eq.refl c) := by cases destruct c with b cb; simp | ._, ._, or.inr ⟨s, rfl, rfl⟩ := begin - apply cases_on s; intros s; simp, + apply rec_on s; intros s; simp, { generalize : f s = fs, - apply cases_on fs; intros t; simp, + apply rec_on fs; intros t; simp, { cases destruct (g t) with b cb; simp } }, { exact or.inr ⟨s, rfl, rfl⟩ } end end }, @@ -619,7 +620,7 @@ theorem of_results_bind {s : computation α} {f : α → computation β} {b k} : ∃ a m n, results s a m ∧ results (f a) b n ∧ k = n + m := begin induction k with n IH generalizing s; - apply cases_on s (λ a, _) (λ s', _); intro e, + apply rec_on s (λ a, _) (λ s', _); intro e, { simp [thinkN] at e, refine ⟨a, _, _, results_ret _, e, rfl⟩ }, { have := congr_arg head (eq_thinkN e), contradiction }, { simp at e, refine ⟨a, _, n+1, results_ret _, e, rfl⟩ }, @@ -705,7 +706,7 @@ destruct_eq_think $ by unfold has_orelse.orelse; simp [orelse] begin apply eq_of_bisim (λ c₁ c₂, (empty α <|> c₂) = c₁) _ rfl, intros s' s h, rw ←h, - apply cases_on s; intros s; rw think_empty; simp, + apply rec_on s; intros s; rw think_empty; simp, rw ←think_empty, end @@ -713,7 +714,7 @@ end begin apply eq_of_bisim (λ c₁ c₂, (c₂ <|> empty α) = c₁) _ rfl, intros s' s h, rw ←h, - apply cases_on s; intros s; rw think_empty; simp, + apply rec_on s; intros s; rw think_empty; simp, rw←think_empty, end @@ -920,7 +921,7 @@ attribute [simp] lift_rel_aux (C : computation α → computation β → Prop) (a cb) : lift_rel_aux R C (sum.inl a) (destruct cb) ↔ ∃ {b}, b ∈ cb ∧ R a b := begin - apply cb.cases_on (λ b, _) (λ cb, _), + apply cb.rec_on (λ b, _) (λ cb, _), { exact ⟨λ h, ⟨_, ret_mem _, h⟩, λ ⟨b', mb, h⟩, by rw [mem_unique (ret_mem _) mb]; exact h⟩ }, { rw [destruct_think], @@ -944,7 +945,7 @@ begin revert cb, refine mem_rec_on ha _ (λ ca' IH, _); intros cb Hc; have h := H Hc, { simp at h, simp [h] }, - { have h := H Hc, simp, revert h, apply cb.cases_on (λ b, _) (λ cb', _); + { have h := H Hc, simp, revert h, apply cb.rec_on (λ b, _) (λ cb', _); intro h; simp at h; simp [h], exact IH _ h } end diff --git a/src/data/seq/parallel.lean b/src/data/seq/parallel.lean index 32ecf2e561c6f..80f7dfc0832e7 100644 --- a/src/data/seq/parallel.lean +++ b/src/data/seq/parallel.lean @@ -187,7 +187,7 @@ begin cases list.foldr parallel.aux2._match_1 (sum.inr list.nil) l; simp [parallel.aux2], cases destruct c; simp }, simp [parallel.aux1], rw this, cases parallel.aux2 l with a l'; simp, - apply S.cases_on _ (λ c S, _) (λ S, _); simp; simp [parallel.aux1]; + apply S.rec_on _ (λ c S, _) (λ S, _); simp; simp [parallel.aux1]; exact ⟨_, _, rfl, rfl⟩ end end end diff --git a/src/data/seq/seq.lean b/src/data/seq/seq.lean index 74d600162366b..a56904f7f20b4 100644 --- a/src/data/seq/seq.lean +++ b/src/data/seq/seq.lean @@ -183,7 +183,8 @@ by cases s with f al; apply subtype.eq; dsimp [tail, cons]; rw [stream.tail_cons @[simp] theorem nth_tail (s : seq α) (n) : nth (tail s) n = nth s (n + 1) := rfl -def cases_on {C : seq α → Sort v} (s : seq α) +/-- Recursion principle for sequences, compare with `list.rec_on`. -/ +def rec_on {C : seq α → Sort v} (s : seq α) (h1 : C nil) (h2 : ∀ x s, C (cons x s)) : C s := begin induction H : destruct s with v v, { rw destruct_eq_nil H, apply h1 }, @@ -199,7 +200,7 @@ begin { apply destruct_eq_cons, unfold destruct nth functor.map, rw ←e, refl }, rw TH, apply h1 _ _ (or.inl rfl) }, - revert e, apply s.cases_on _ (λ b s', _); intro e, + revert e, apply s.rec_on _ (λ b s', _); intro e, { injection e }, { have h_eq : (cons b s').val (nat.succ k) = s'.val k, { cases s'; refl }, rw [h_eq] at e, @@ -270,7 +271,7 @@ section bisim by cases s; refl, by cases s'; refl, r⟩) this, begin have := bisim r, revert r this, - apply cases_on s _ _; intros; apply cases_on s' _ _; intros; intros r this, + apply rec_on s _ _; intros; apply rec_on s' _ _; intros; intros r this, { constructor, refl, assumption }, { rw [destruct_nil, destruct_cons] at this, exact false.elim this }, @@ -483,7 +484,7 @@ else sum.inr (to_stream s h) begin apply coinduction2, intro s, dsimp [append], rw [corec_eq], - dsimp [append], apply cases_on s _ _, + dsimp [append], apply rec_on s _ _, { trivial }, { intros x s, rw [destruct_cons], dsimp, @@ -500,7 +501,7 @@ end @[simp] theorem append_nil (s : seq α) : append s nil = s := begin apply coinduction2 s, intro s, - apply cases_on s _ _, + apply rec_on s _ _, { trivial }, { intros x s, rw [cons_append, destruct_cons, destruct_cons], dsimp, @@ -513,9 +514,9 @@ begin apply eq_of_bisim (λ s1 s2, ∃ s t u, s1 = append (append s t) u ∧ s2 = append s (append t u)), { intros s1 s2 h, exact match s1, s2, h with ._, ._, ⟨s, t, u, rfl, rfl⟩ := begin - apply cases_on s; simp, - { apply cases_on t; simp, - { apply cases_on u; simp, + apply rec_on s; simp, + { apply rec_on t; simp, + { apply rec_on u; simp, { intros x u, refine ⟨nil, nil, u, _, _⟩; simp } }, { intros x t, refine ⟨nil, t, u, _, _⟩; simp } }, { intros x s, exact ⟨s, t, u, rfl, rfl⟩ } @@ -550,8 +551,8 @@ begin apply eq_of_bisim (λ s1 s2, ∃ s t, s1 = map f (append s t) ∧ s2 = append (map f s) (map f t)) _ ⟨s, t, rfl, rfl⟩, intros s1 s2 h, exact match s1, s2, h with ._, ._, ⟨s, t, rfl, rfl⟩ := begin - apply cases_on s; simp, - { apply cases_on t; simp, + apply rec_on s; simp, + { apply rec_on t; simp, { intros x t, refine ⟨nil, t, _, _⟩; simp } }, { intros x s, refine ⟨s, t, rfl, rfl⟩ } end end @@ -584,11 +585,11 @@ begin intros s1 s2 h, exact match s1, s2, h with | _, _, (or.inl $ eq.refl s) := begin - apply cases_on s, { trivial }, + apply rec_on s, { trivial }, { intros x s, rw [destruct_cons], exact ⟨rfl, or.inl rfl⟩ } end | ._, ._, (or.inr ⟨a, s, S, rfl, rfl⟩) := begin - apply cases_on s, + apply rec_on s, { simp }, { intros x s, simp, refine or.inr ⟨x, s, S, rfl, rfl⟩ } end @@ -602,9 +603,9 @@ begin s1 = append s (join (append S T)) ∧ s2 = append s (append (join S) (join T))), { intros s1 s2 h, exact match s1, s2, h with ._, ._, ⟨s, S, T, rfl, rfl⟩ := begin - apply cases_on s; simp, - { apply cases_on S; simp, - { apply cases_on T, { simp }, + apply rec_on s; simp, + { apply rec_on S; simp, + { apply rec_on T, { simp }, { intros s T, cases s with a s; simp, refine ⟨s, nil, T, _, _⟩; simp } }, { intros s S, cases s with a s; simp, @@ -662,7 +663,7 @@ begin generalize e : append s₁ s₂ = ss, intro h, revert s₁, apply mem_rec_on h _, intros b s' o s₁, - apply s₁.cases_on _ (λ c t₁, _); intros m e; + apply s₁.rec_on _ (λ c t₁, _); intros m e; have := congr_arg destruct e, { apply or.inr, simpa using m }, { cases (show a = c ∨ a ∈ append t₁ s₂, by simpa using m) with e' m, @@ -721,7 +722,7 @@ def bind (s : seq1 α) (f : α → seq1 β) : seq1 β := join (map f s) @[simp] theorem join_map_ret (s : seq α) : seq.join (seq.map ret s) = s := -by apply coinduction2 s; intro s; apply cases_on s; simp [ret] +by apply coinduction2 s; intro s; apply rec_on s; simp [ret] @[simp] theorem bind_ret (f : α → β) : ∀ s, bind s (ret ∘ f) = map f s | ⟨a, s⟩ := begin @@ -733,7 +734,7 @@ end begin simp [ret, bind, map], cases f a with a s, - apply cases_on s; intros; simp + apply rec_on s; intros; simp end @[simp] theorem map_join' (f : α → β) (S) : @@ -743,8 +744,8 @@ begin ∃ s S, s1 = append s (seq.map f (seq.join S)) ∧ s2 = append s (seq.join (seq.map (map f) S))), { intros s1 s2 h, exact match s1, s2, h with ._, ._, ⟨s, S, rfl, rfl⟩ := begin - apply cases_on s; simp, - { apply cases_on S; simp, + apply rec_on s; simp, + { apply rec_on S; simp, { intros x S, cases x with a s; simp [map], exact ⟨_, _, rfl, rfl⟩ } }, { intros x s, refine ⟨s, S, rfl, rfl⟩ } @@ -753,7 +754,7 @@ begin end @[simp] theorem map_join (f : α → β) : ∀ S, map f (join S) = join (map (map f) S) -| ((a, s), S) := by apply cases_on s; intros; simp [map] +| ((a, s), S) := by apply rec_on s; intros; simp [map] @[simp] theorem join_join (SS : seq (seq1 (seq1 α))) : seq.join (seq.join SS) = seq.join (seq.map join SS) := @@ -762,10 +763,10 @@ begin ∃ s SS, s1 = seq.append s (seq.join (seq.join SS)) ∧ s2 = seq.append s (seq.join (seq.map join SS))), { intros s1 s2 h, exact match s1, s2, h with ._, ._, ⟨s, SS, rfl, rfl⟩ := begin - apply cases_on s; simp, - { apply cases_on SS; simp, + apply rec_on s; simp, + { apply rec_on SS; simp, { intros S SS, cases S with s S; cases s with x s; simp [map], - apply cases_on s; simp, + apply rec_on s; simp, { exact ⟨_, _, rfl, rfl⟩ }, { intros x s, refine ⟨cons x (append s (seq.join S)), SS, _, _⟩; simp } } }, @@ -784,8 +785,8 @@ begin rw [map_comp _ join], generalize : seq.map (map g ∘ f) s = SS, rcases map g (f a) with ⟨⟨a, s⟩, S⟩, - apply cases_on s; intros; apply cases_on S; intros; simp, - { cases x with x t, apply cases_on t; intros; simp }, + apply rec_on s; intros; apply rec_on S; intros; simp, + { cases x with x t, apply rec_on t; intros; simp }, { cases x_1 with y t; simp } end diff --git a/src/data/seq/wseq.lean b/src/data/seq/wseq.lean index 540787c0c45c5..0af202614f24a 100644 --- a/src/data/seq/wseq.lean +++ b/src/data/seq/wseq.lean @@ -66,9 +66,11 @@ computation.corec (λ s, match seq.destruct s with | some (some a, s') := sum.inl (some (a, s')) end) -def cases_on {C : wseq α → Sort v} (s : wseq α) (h1 : C nil) + +/-- Recursion principle for weak sequences, compare with `list.rec_on`. -/ +def rec_on {C : wseq α → Sort v} (s : wseq α) (h1 : C nil) (h2 : ∀ x s, C (cons x s)) (h3 : ∀ s, C (think s)) : C s := -seq.cases_on s h1 (λ o, option.cases_on o h3 h2) +seq.rec_on s h1 (λ o, option.rec_on o h3 h2) protected def mem (a : α) (s : wseq α) := seq.mem (some a) s @@ -500,7 +502,7 @@ begin intros c1 c2 h, exact match c1, c2, h with | _, _, (or.inl $ eq.refl c) := by cases c.destruct; simp | _, _, (or.inr ⟨c, rfl, rfl⟩) := begin - apply c.cases_on (λ a, _) (λ c', _); repeat {simp}, + apply c.rec_on (λ a, _) (λ c', _); repeat {simp}, { cases (destruct a).destruct; simp }, { exact or.inr ⟨c', rfl, rfl⟩ } end end @@ -677,7 +679,7 @@ theorem eq_or_mem_iff_mem {s : wseq α} {a a' s'} : begin generalize e : destruct s = c, intro h, revert s, apply computation.mem_rec_on h _ (λ c IH, _); intro s; - apply s.cases_on _ (λ x s, _) (λ s, _); intros m; + apply s.rec_on _ (λ x s, _) (λ s, _); intros m; have := congr_arg computation.destruct m; simp at this; cases this with i1 i2, { rw [i1, i2], @@ -704,7 +706,7 @@ theorem mem_cons (s : wseq α) (a) : a ∈ cons a s := theorem mem_of_mem_tail {s : wseq α} {a} : a ∈ tail s → a ∈ s := begin intro h, have := h, cases h with n e, revert s, simp [stream.nth], - induction n with n IH; intro s; apply s.cases_on _ (λ x s, _) (λ s, _); + induction n with n IH; intro s; apply s.rec_on _ (λ x s, _) (λ s, _); repeat{simp}; intros m e; injections, { exact or.inr m }, { exact or.inr m }, @@ -903,7 +905,7 @@ begin c2 = computation.map list.length (corec to_list._match_2 (l, s))) _ ⟨[], s, rfl, rfl⟩, intros s1 s2 h, rcases h with ⟨l, s, h⟩, rw [h.left, h.right], - apply s.cases_on _ (λ a s, _) (λ s, _); + apply s.rec_on _ (λ a s, _) (λ s, _); repeat {simp [to_list, nil, cons, think, length]}, { refine ⟨a::l, s, _, _⟩; simp }, { refine ⟨l, s, _, _⟩; simp } @@ -940,7 +942,7 @@ begin c2 = computation.map ((++) l.reverse) (corec to_list._match_2 (l', s))) _ ⟨[], s, rfl, rfl⟩, intros s1 s2 h, rcases h with ⟨l', s, h⟩, rw [h.left, h.right], - apply s.cases_on _ (λ a s, _) (λ s, _); + apply s.rec_on _ (λ a s, _) (λ s, _); repeat {simp [to_list, nil, cons, think, length]}, { refine ⟨a::l', s, _, _⟩; simp }, { refine ⟨l', s, _, _⟩; simp } @@ -971,7 +973,7 @@ by simp [head]; cases seq.head s; refl @[simp] theorem tail_of_seq (s : seq α) : tail (of_seq s) = of_seq s.tail := begin - simp [tail], apply s.cases_on _ (λ x s, _); simp [of_seq], {refl}, + simp [tail], apply s.rec_on _ (λ x s, _); simp [of_seq], {refl}, rw [seq.head_cons, seq.tail_cons], refl end @@ -1028,7 +1030,7 @@ suffices ∀ ss : wseq α, a ∈ ss → ∀ s S, append s (join S) = ss → (this _ h nil S (by simp) (by simp [h])).resolve_left (not_mem_nil _), begin intros ss h, apply mem_rec_on h (λ b ss o, _) (λ ss IH, _); intros s S, - { refine s.cases_on (S.cases_on _ (λ s S, _) (λ S, _)) (λ b' s, _) (λ s, _); + { refine s.rec_on (S.rec_on _ (λ s S, _) (λ S, _)) (λ b' s, _) (λ s, _); intros ej m; simp at ej; have := congr_arg seq.destruct ej; simp at this; try {cases this}; try {contradiction}, @@ -1037,7 +1039,7 @@ begin cases o with e IH, { simp [e] }, cases m with e m, { simp [e] }, exact or.imp_left or.inr (IH _ _ rfl m) }, - { refine s.cases_on (S.cases_on _ (λ s S, _) (λ S, _)) (λ b' s, _) (λ s, _); + { refine s.rec_on (S.rec_on _ (λ s S, _) (λ S, _)) (λ b' s, _) (λ s, _); intros ej m; simp at ej; have := congr_arg seq.destruct ej; simp at this; try { try {have := this.1}, contradiction }; subst ss, @@ -1063,7 +1065,7 @@ begin apply eq_of_bisim (λ c1 c2, ∃ s, c1 = destruct (map f s) ∧ c2 = computation.map (option.map (prod.map f (map f))) (destruct s)), { intros c1 c2 h, cases h with s h, rw [h.left, h.right], - apply s.cases_on _ (λ a s, _) (λ s, _); simp, + apply s.rec_on _ (λ a s, _) (λ s, _); simp, exact ⟨s, rfl, rfl⟩ }, { exact ⟨s, rfl, rfl⟩ } end @@ -1099,8 +1101,8 @@ begin apply eq_of_bisim (λ c1 c2, ∃ s t, c1 = destruct (append s t) ∧ c2 = (destruct s).bind (destruct_append.aux t)) _ ⟨s, t, rfl, rfl⟩, intros c1 c2 h, rcases h with ⟨s, t, h⟩, rw [h.left, h.right], - apply s.cases_on _ (λ a s, _) (λ s, _); simp, - { apply t.cases_on _ (λ b t, _) (λ t, _); simp, + apply s.rec_on _ (λ a s, _) (λ s, _); simp, + { apply t.rec_on _ (λ b t, _) (λ t, _); simp, { refine ⟨nil, t, _, _⟩; simp } }, { exact ⟨s, t, rfl, rfl⟩ } end @@ -1117,7 +1119,7 @@ begin intros c1 c2 h, exact match c1, c2, h with | _, _, (or.inl $ eq.refl c) := by cases c.destruct; simp | _, _, or.inr ⟨S, rfl, rfl⟩ := begin - apply S.cases_on _ (λ s S, _) (λ S, _); simp, + apply S.rec_on _ (λ s S, _) (λ S, _); simp, { refine or.inr ⟨S, rfl, rfl⟩ } end end end @@ -1243,7 +1245,7 @@ begin clear h _match, have : ∀ s, ∃ s' : wseq α, (map ret s).join.destruct = (map ret s').join.destruct ∧ destruct s = s'.destruct, from λ s, ⟨s, rfl, rfl⟩, - apply s.cases_on _ (λ a s, _) (λ s, _); simp [ret, ret_mem, this, option.exists] + apply s.rec_on _ (λ a s, _) (λ s, _); simp [ret, ret_mem, this, option.exists] end end }, { exact ⟨s, rfl, rfl⟩ } end @@ -1263,9 +1265,9 @@ begin intros c1 c2 h, exact match c1, c2, h with ._, ._, ⟨s, S, T, rfl, rfl⟩ := begin clear _match h h, - apply wseq.cases_on s _ (λ a s, _) (λ s, _); simp, - { apply wseq.cases_on S _ (λ s S, _) (λ S, _); simp, - { apply wseq.cases_on T _ (λ s T, _) (λ T, _); simp, + apply wseq.rec_on s _ (λ a s, _) (λ s, _); simp, + { apply wseq.rec_on S _ (λ s S, _) (λ S, _); simp, + { apply wseq.rec_on T _ (λ s T, _) (λ T, _); simp, { refine ⟨s, nil, T, _, _⟩; simp }, { refine ⟨nil, nil, T, _, _⟩; simp } }, { exact ⟨s, S, T, rfl, rfl⟩ }, @@ -1292,8 +1294,8 @@ begin s2 = append s (join (map (map f) S))), { intros s1 s2 h, exact match s1, s2, h with ._, ._, ⟨s, S, rfl, rfl⟩ := begin - apply wseq.cases_on s _ (λ a s, _) (λ s, _); simp, - { apply wseq.cases_on S _ (λ s S, _) (λ S, _); simp, + apply wseq.rec_on s _ (λ a s, _) (λ s, _); simp, + { apply wseq.rec_on S _ (λ s S, _) (λ S, _); simp, { exact ⟨map f s, S, rfl, rfl⟩ }, { refine ⟨nil, S, _, _⟩; simp } }, { exact ⟨_, _, rfl, rfl⟩ }, @@ -1318,9 +1320,9 @@ begin intros c1 c2 h, exact match c1, c2, h with ._, ._, ⟨s, S, SS, rfl, rfl⟩ := begin clear _match h h, - apply wseq.cases_on s _ (λ a s, _) (λ s, _); simp, - { apply wseq.cases_on S _ (λ s S, _) (λ S, _); simp, - { apply wseq.cases_on SS _ (λ S SS, _) (λ SS, _); simp, + apply wseq.rec_on s _ (λ a s, _) (λ s, _); simp, + { apply wseq.rec_on S _ (λ s S, _) (λ S, _); simp, + { apply wseq.rec_on SS _ (λ S SS, _) (λ SS, _); simp, { refine ⟨nil, S, SS, _, _⟩; simp }, { refine ⟨nil, nil, SS, _, _⟩; simp } }, { exact ⟨s, S, SS, rfl, rfl⟩ }, From 5cf34f8f1b20b989b0f8574519e65dbdbdb037d2 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 3 Oct 2022 17:40:48 +0000 Subject: [PATCH 533/828] feat(topology/algebra/continuous_monoid_hom): Continuity of composition (#16319) This PR adds some API regarding continuity of compositions of `continuous_monoid_hom`s. --- .../algebra/continuous_monoid_hom.lean | 109 ++++++++++++++++-- 1 file changed, 98 insertions(+), 11 deletions(-) diff --git a/src/topology/algebra/continuous_monoid_hom.lean b/src/topology/algebra/continuous_monoid_hom.lean index 4150f256e0c8b..e95e500b9095e 100644 --- a/src/topology/algebra/continuous_monoid_hom.lean +++ b/src/topology/algebra/continuous_monoid_hom.lean @@ -90,7 +90,7 @@ directly. -/ `fun_like.has_coe_to_fun` directly."] instance : has_coe_to_fun (continuous_monoid_hom A B) (λ _, A → B) := fun_like.has_coe_to_fun -@[to_additive] lemma ext {f g : continuous_monoid_hom A B} (h : ∀ x, f x = g x) : f = g := +@[to_additive, ext] lemma ext {f g : continuous_monoid_hom A B} (h : ∀ x, f x = g x) : f = g := fun_like.ext _ _ h /-- Reinterpret a `continuous_monoid_hom` as a `continuous_map`. -/ @@ -188,19 +188,21 @@ def coprod (f : continuous_monoid_hom A E) (g : continuous_monoid_hom B E) : inv := λ f, (inv E).comp f, mul_left_inv := λ f, ext (λ x, mul_left_inv (f x)) } -instance : topological_space (continuous_monoid_hom A B) := +@[to_additive] instance : topological_space (continuous_monoid_hom A B) := topological_space.induced to_continuous_map continuous_map.compact_open variables (A B C D E) -lemma is_inducing : inducing (to_continuous_map : continuous_monoid_hom A B → C(A, B)) := ⟨rfl⟩ +@[to_additive] lemma inducing_to_continuous_map : + inducing (to_continuous_map : continuous_monoid_hom A B → C(A, B)) := ⟨rfl⟩ -lemma is_embedding : embedding (to_continuous_map : continuous_monoid_hom A B → C(A, B)) := -⟨is_inducing A B, to_continuous_map_injective⟩ +@[to_additive] lemma embedding_to_continuous_map : + embedding (to_continuous_map : continuous_monoid_hom A B → C(A, B)) := +⟨inducing_to_continuous_map A B, to_continuous_map_injective⟩ -lemma is_closed_embedding [has_continuous_mul B] [t2_space B] : +@[to_additive] lemma closed_embedding_to_continuous_map [has_continuous_mul B] [t2_space B] : closed_embedding (to_continuous_map : continuous_monoid_hom A B → C(A, B)) := -⟨is_embedding A B, ⟨begin +⟨embedding_to_continuous_map A B, ⟨begin suffices : (set.range (to_continuous_map : continuous_monoid_hom A B → C(A, B))) = ({f | f '' {1} ⊆ {1}ᶜ} ∪ ⋃ (x y) (U V W) (hU : is_open U) (hV : is_open V) (hW : is_open W) (h : disjoint (U * V) W), {f | f '' {x} ⊆ U} ∩ {f | f '' {y} ⊆ V} ∩ {f | f '' {x * y} ⊆ W})ᶜ, @@ -231,16 +233,101 @@ end⟩⟩ variables {A B C D E} -instance [t2_space B] : t2_space (continuous_monoid_hom A B) := -(is_embedding A B).t2_space +@[to_additive] instance [t2_space B] : t2_space (continuous_monoid_hom A B) := +(embedding_to_continuous_map A B).t2_space -instance : topological_group (continuous_monoid_hom A E) := -let hi := is_inducing A E, hc := hi.continuous in +@[to_additive] instance : topological_group (continuous_monoid_hom A E) := +let hi := inducing_to_continuous_map A E, hc := hi.continuous in { continuous_mul := hi.continuous_iff.mpr (continuous_mul.comp (continuous.prod_map hc hc)), continuous_inv := hi.continuous_iff.mpr (continuous_inv.comp hc) } +@[to_additive] lemma continuous_of_continuous_uncurry {A : Type*} [topological_space A] + (f : A → continuous_monoid_hom B C) (h : continuous (function.uncurry (λ x y, f x y))) : + continuous f := +(inducing_to_continuous_map _ _).continuous_iff.mpr + (continuous_map.continuous_of_continuous_uncurry _ h) + +@[to_additive] lemma continuous_comp [locally_compact_space B] : + continuous (λ f : continuous_monoid_hom A B × continuous_monoid_hom B C, f.2.comp f.1) := +(inducing_to_continuous_map A C).continuous_iff.2 $ (continuous_map.continuous_comp'.comp + ((inducing_to_continuous_map A B).prod_mk (inducing_to_continuous_map B C)).continuous) + +@[to_additive] lemma continuous_comp_left (f : continuous_monoid_hom A B) : + continuous (λ g : continuous_monoid_hom B C, g.comp f) := +(inducing_to_continuous_map A C).continuous_iff.2 $ f.to_continuous_map.continuous_comp_left.comp + (inducing_to_continuous_map B C).continuous + +@[to_additive] lemma continuous_comp_right (f : continuous_monoid_hom B C) : + continuous (λ g : continuous_monoid_hom A B, f.comp g) := +(inducing_to_continuous_map A C).continuous_iff.2 $ f.to_continuous_map.continuous_comp.comp + (inducing_to_continuous_map A B).continuous + +variables (E) + +/-- `continuous_monoid_hom _ f` is a functor. -/ +@[to_additive "`continuous_add_monoid_hom _ f` is a functor."] +def comp_left (f : continuous_monoid_hom A B) : + continuous_monoid_hom (continuous_monoid_hom B E) (continuous_monoid_hom A E) := +{ to_fun := λ g, g.comp f, + map_one' := rfl, + map_mul' := λ g h, rfl, + continuous_to_fun := f.continuous_comp_left } + +variables (A) {E} + +/-- `continuous_monoid_hom f _` is a functor. -/ +@[to_additive "`continuous_add_monoid_hom f _` is a functor."] +def comp_right {B : Type*} [comm_group B] [topological_space B] + [topological_group B] (f : continuous_monoid_hom B E) : + continuous_monoid_hom (continuous_monoid_hom A B) (continuous_monoid_hom A E) := +{ to_fun := λ g, f.comp g, + map_one' := ext (λ a, map_one f), + map_mul' := λ g h, ext (λ a, map_mul f (g a) (h a)), + continuous_to_fun := f.continuous_comp_right } + end continuous_monoid_hom /-- The Pontryagin dual of `A` is the group of continuous homomorphism `A → circle`. -/ @[derive [topological_space, t2_space, comm_group, topological_group, inhabited]] def pontryagin_dual := continuous_monoid_hom A circle + +variables {A B C D E} + +namespace pontryagin_dual + +open continuous_monoid_hom + +noncomputable instance : continuous_monoid_hom_class (pontryagin_dual A) A circle := +continuous_monoid_hom.continuous_monoid_hom_class + +/-- `pontryagin_dual` is a functor. -/ +noncomputable def map (f : continuous_monoid_hom A B) : + continuous_monoid_hom (pontryagin_dual B) (pontryagin_dual A) := +f.comp_left circle + +@[simp] lemma map_apply (f : continuous_monoid_hom A B) (x : pontryagin_dual B) (y : A) : + map f x y = x (f y) := +rfl + +@[simp] lemma map_one : map (one A B) = one (pontryagin_dual B) (pontryagin_dual A) := +ext (λ x, ext (λ y, map_one x)) + +@[simp] lemma map_comp (g : continuous_monoid_hom B C) (f : continuous_monoid_hom A B) : + map (comp g f) = comp (map f) (map g) := +ext (λ x, ext (λ y, rfl)) + +@[simp] lemma map_mul (f g : continuous_monoid_hom A E) : map (f * g) = map f * map g := +ext (λ x, ext (λ y, map_mul x (f y) (g y))) + +variables (A B C D E) + +/-- `continuous_monoid_hom.dual` as a `continuous_monoid_hom`. -/ +noncomputable def map_hom [locally_compact_space E] : + continuous_monoid_hom (continuous_monoid_hom A E) + (continuous_monoid_hom (pontryagin_dual E) (pontryagin_dual A)) := +{ to_fun := map, + map_one' := map_one, + map_mul' := map_mul, + continuous_to_fun := continuous_of_continuous_uncurry _ continuous_comp } + +end pontryagin_dual From 486cb2f3bda4a67557c6285f5bd0c3348c1eea81 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Mon, 3 Oct 2022 17:40:49 +0000 Subject: [PATCH 534/828] feat(ring_theory/graded_algebra/homogeneous_localization): generalise homogeneous localization to any submonoid (#16653) homogeneous localization is generalised to work on any `submonoid` and introduced `homogneous_localization.prime` and `homogeneous_localization.away` --- .../projective_spectrum/structure_sheaf.lean | 13 +- .../homogeneous_localization.lean | 132 +++++++++++------- 2 files changed, 84 insertions(+), 61 deletions(-) diff --git a/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean b/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean index e68917a294a1c..d6d743bec77d1 100644 --- a/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean +++ b/src/algebraic_geometry/projective_spectrum/structure_sheaf.lean @@ -56,7 +56,7 @@ variables {R A: Type*} variables [comm_ring R] [comm_ring A] [algebra R A] variables (𝒜 : ℕ → submodule R A) [graded_algebra 𝒜] -local notation `at ` x := homogeneous_localization 𝒜 x.as_homogeneous_ideal.to_ideal +local notation `at ` x := homogeneous_localization.at_prime 𝒜 x.as_homogeneous_ideal.to_ideal namespace projective_spectrum.structure_sheaf @@ -120,7 +120,7 @@ begin obtain ⟨nin2, hy2⟩ := (wb (opens.inf_le_right Va Vb y)), dsimp only at hy1 hy2, erw [hy1, hy2], - simpa only [val_mk', add_mk, ← subtype.val_eq_coe, add_comm], } + simpa only [val_mk', add_mk, ← subtype.val_eq_coe, add_comm, mul_comm sa sb], } end lemma neg_mem' (U : (opens (projective_spectrum.Top 𝒜))ᵒᵖ) @@ -271,10 +271,7 @@ by { cases x, exact stalk_to_fiber_ring_hom_germ' 𝒜 U _ _ _ } lemma homogeneous_localization.mem_basic_open (x : projective_spectrum.Top 𝒜) (f : at x) : x ∈ projective_spectrum.basic_open 𝒜 f.denom := -begin - rw projective_spectrum.mem_basic_open, - exact homogeneous_localization.denom_not_mem _, -end +by { rw projective_spectrum.mem_basic_open, exact f.denom_mem } variable (𝒜) @@ -284,9 +281,9 @@ basic open set `D(f.denom)`-/ def section_in_basic_open (x : projective_spectrum.Top 𝒜) : Π (f : at x), (Proj.structure_sheaf 𝒜).1.obj (op (projective_spectrum.basic_open 𝒜 f.denom)) := -λ f, ⟨λ y, quotient.mk' ⟨f.deg, ⟨f.num, f.num_mem⟩, ⟨f.denom, f.denom_mem⟩, y.2⟩, +λ f, ⟨λ y, quotient.mk' ⟨f.deg, ⟨f.num, f.num_mem_deg⟩, ⟨f.denom, f.denom_mem_deg⟩, y.2⟩, λ y, ⟨projective_spectrum.basic_open 𝒜 f.denom, y.2, - ⟨𝟙 _, ⟨f.deg, ⟨⟨f.num, f.num_mem⟩, ⟨f.denom, f.denom_mem⟩, + ⟨𝟙 _, ⟨f.deg, ⟨⟨f.num, f.num_mem_deg⟩, ⟨f.denom, f.denom_mem_deg⟩, λ z, ⟨z.2, rfl⟩⟩⟩⟩⟩⟩ /--Given any point `x` and `f` in the homogeneous localization at `x`, there is an element in the diff --git a/src/ring_theory/graded_algebra/homogeneous_localization.lean b/src/ring_theory/graded_algebra/homogeneous_localization.lean index ee1f4e61e9b0e..714d3f8767d65 100644 --- a/src/ring_theory/graded_algebra/homogeneous_localization.lean +++ b/src/ring_theory/graded_algebra/homogeneous_localization.lean @@ -14,7 +14,7 @@ import ring_theory.graded_algebra.basic - `R` is a commutative semiring; - `A` is a commutative ring and an `R`-algebra; - `𝒜 : ι → submodule R A` is the grading of `A`; -- `x : ideal A` is a prime ideal. +- `x : submonoid A` is a submonoid ## Main definitions and results @@ -29,7 +29,7 @@ is a `num_denom_same_deg`, then generally, `c + (-c)` is not necessarily `0` for `0` is considered to have grade zero (see `deg_zero`) but `c + (-c)` has the same degree as `c`. To circumvent this, we quotient `num_denom_same_deg 𝒜 x` by the kernel of `c ↦ c.num / c.denom`. -* `homogeneous_localization.num_denom_same_deg.embedding` : for `x : prime ideal of A` and any +* `homogeneous_localization.num_denom_same_deg.embedding` : for `x : submonoid A` and any `c : num_denom_same_deg 𝒜 x`, or equivalent a numerator and a denominator of the same degree, we get an element `c.num / c.denom` of `Aₓ`. * `homogeneous_localization`: `num_denom_same_deg 𝒜 x` quotiented by kernel of `embedding 𝒜 x`. @@ -38,19 +38,20 @@ circumvent this, we quotient `num_denom_same_deg 𝒜 x` by the kernel of `c ↦ through `homogeneous_localization.val`. * `homogeneous_localization.num`: if `f : homogeneous_localization 𝒜 x`, then `f.num : A` is the numerator of `f`. -* `homogeneous_localization.num`: if `f : homogeneous_localization 𝒜 x`, then `f.denom : A` is the +* `homogeneous_localization.denom`: if `f : homogeneous_localization 𝒜 x`, then `f.denom : A` is the denominator of `f`. * `homogeneous_localization.deg`: if `f : homogeneous_localization 𝒜 x`, then `f.deg : ι` is the degree of `f` such that `f.num ∈ 𝒜 f.deg` and `f.denom ∈ 𝒜 f.deg` - (see `homogeneous_localization.num_mem` and `homogeneous_localization.denom_mem`). -* `homogeneous_localization.num_mem`: if `f : homogeneous_localization 𝒜 x`, then `f.num_mem` is a - proof that `f.num ∈ f.deg`. -* `homogeneous_localization.denom_mem`: if `f : homogeneous_localization 𝒜 x`, then `f.denom_mem` - is a proof that `f.denom ∈ f.deg`. + (see `homogeneous_localization.num_mem_deg` and `homogeneous_localization.denom_mem_deg`). +* `homogeneous_localization.num_mem_deg`: if `f : homogeneous_localization 𝒜 x`, then + `f.num_mem_deg` is a proof that `f.num ∈ 𝒜 f.deg`. +* `homogeneous_localization.denom_mem_deg`: if `f : homogeneous_localization 𝒜 x`, then + `f.denom_mem_deg` is a proof that `f.denom ∈ 𝒜 f.deg`. * `homogeneous_localization.eq_num_div_denom`: if `f : homogeneous_localization 𝒜 x`, then `f.val : Aₓ` is equal to `f.num / f.denom`. -* `homogeneous_localization.local_ring`: `homogeneous_localization 𝒜 x` is a local ring. +* `homogeneous_localization.local_ring`: `homogeneous_localization 𝒜 x` is a local ring when `x` is + the complement of some prime ideals. ## References @@ -68,22 +69,22 @@ variables {ι R A: Type*} variables [add_comm_monoid ι] [decidable_eq ι] variables [comm_ring R] [comm_ring A] [algebra R A] variables (𝒜 : ι → submodule R A) [graded_algebra 𝒜] -variables (x : ideal A) [ideal.is_prime x] +variables (x : submonoid A) -local notation `at ` x := localization.at_prime x +local notation `at ` x := localization x namespace homogeneous_localization section /-- -Let `x` be a prime ideal, then `num_denom_same_deg 𝒜 x` is a structure with a numerator and a -denominator with same grading such that the denominator is not contained in `x`. +Let `x` be a submonoid of `A`, then `num_denom_same_deg 𝒜 x` is a structure with a numerator and a +denominator with same grading such that the denominator is contained in `x`. -/ @[nolint has_nonempty_instance] structure num_denom_same_deg := (deg : ι) (num denom : 𝒜 deg) -(denom_not_mem : (denom : A) ∉ x) +(denom_mem : (denom : A) ∈ x) end @@ -99,8 +100,7 @@ begin rcases c1 with ⟨i1, ⟨n1, hn1⟩, ⟨d1, hd1⟩, h1⟩, rcases c2 with ⟨i2, ⟨n2, hn2⟩, ⟨d2, hd2⟩, h2⟩, dsimp only [subtype.coe_mk] at *, - simp only, - exact ⟨hdeg, by subst hdeg; subst hnum, by subst hdeg; subst hdenom⟩, + simp only, exact ⟨hdeg, by subst hdeg; subst hnum, by subst hdeg; subst hdenom⟩, end instance : has_one (num_denom_same_deg 𝒜 x) := @@ -108,14 +108,14 @@ instance : has_one (num_denom_same_deg 𝒜 x) := { deg := 0, num := ⟨1, one_mem⟩, denom := ⟨1, one_mem⟩, - denom_not_mem := λ r, (infer_instance : x.is_prime).ne_top $ x.eq_top_iff_one.mpr r } } + denom_mem := submonoid.one_mem _ } } @[simp] lemma deg_one : (1 : num_denom_same_deg 𝒜 x).deg = 0 := rfl @[simp] lemma num_one : ((1 : num_denom_same_deg 𝒜 x).num : A) = 1 := rfl @[simp] lemma denom_one : ((1 : num_denom_same_deg 𝒜 x).denom : A) = 1 := rfl instance : has_zero (num_denom_same_deg 𝒜 x) := -{ zero := ⟨0, 0, ⟨1, one_mem⟩, λ r, (infer_instance : x.is_prime).ne_top $ x.eq_top_iff_one.mpr r⟩ } +{ zero := ⟨0, 0, ⟨1, one_mem⟩, submonoid.one_mem _⟩ } @[simp] lemma deg_zero : (0 : num_denom_same_deg 𝒜 x).deg = 0 := rfl @[simp] lemma num_zero : (0 : num_denom_same_deg 𝒜 x).num = 0 := rfl @@ -126,8 +126,7 @@ instance : has_mul (num_denom_same_deg 𝒜 x) := { deg := p.deg + q.deg, num := ⟨p.num * q.num, mul_mem p.num.prop q.num.prop⟩, denom := ⟨p.denom * q.denom, mul_mem p.denom.prop q.denom.prop⟩, - denom_not_mem := λ r, or.elim - ((infer_instance : x.is_prime).mem_or_mem r) p.denom_not_mem q.denom_not_mem } } + denom_mem := submonoid.mul_mem _ p.denom_mem q.denom_mem } } @[simp] lemma deg_mul (c1 c2 : num_denom_same_deg 𝒜 x) : (c1 * c2).deg = c1.deg + c2.deg := rfl @[simp] lemma num_mul (c1 c2 : num_denom_same_deg 𝒜 x) : @@ -142,8 +141,7 @@ instance : has_add (num_denom_same_deg 𝒜 x) := add_mem (mul_mem c1.denom.2 c2.num.2) (add_comm c2.deg c1.deg ▸ mul_mem c2.denom.2 c1.num.2)⟩, denom := ⟨c1.denom * c2.denom, mul_mem c1.denom.2 c2.denom.2⟩, - denom_not_mem := λ r, or.elim - ((infer_instance : x.is_prime).mem_or_mem r) c1.denom_not_mem c2.denom_not_mem } } + denom_mem := submonoid.mul_mem _ c1.denom_mem c2.denom_mem } } @[simp] lemma deg_add (c1 c2 : num_denom_same_deg 𝒜 x) : (c1 + c2).deg = c1.deg + c2.deg := rfl @[simp] lemma num_add (c1 c2 : num_denom_same_deg 𝒜 x) : @@ -152,7 +150,7 @@ instance : has_add (num_denom_same_deg 𝒜 x) := ((c1 + c2).denom : A) = c1.denom * c2.denom := rfl instance : has_neg (num_denom_same_deg 𝒜 x) := -{ neg := λ c, ⟨c.deg, ⟨-c.num, neg_mem c.num.2⟩, c.denom, c.denom_not_mem⟩ } +{ neg := λ c, ⟨c.deg, ⟨-c.num, neg_mem c.num.2⟩, c.denom, c.denom_mem⟩ } @[simp] lemma deg_neg (c : num_denom_same_deg 𝒜 x) : (-c).deg = c.deg := rfl @[simp] lemma num_neg (c : num_denom_same_deg 𝒜 x) : ((-c).num : A) = -c.num := rfl @@ -171,11 +169,9 @@ instance : has_pow (num_denom_same_deg 𝒜 x) ℕ := @graded_monoid.gmonoid.gnpow _ (λ i, ↥(𝒜 i)) _ _ n _ c.num, @graded_monoid.gmonoid.gnpow _ (λ i, ↥(𝒜 i)) _ _ n _ c.denom, begin - cases n, - { simp only [graded_monoid.gmonoid.gnpow, subtype.coe_mk, pow_zero], - exact λ r, (infer_instance : x.is_prime).ne_top $ (ideal.eq_top_iff_one _).mpr r, }, - { exact λ r, c.denom_not_mem $ - ((infer_instance : x.is_prime).pow_mem_iff_mem n.succ (nat.zero_lt_succ _)).mp r } + induction n with n ih, + { simpa only [coe_gnpow, pow_zero] using submonoid.one_mem _ }, + { simpa only [pow_succ', coe_gnpow] using x.mul_mem ih c.denom_mem, }, end⟩ } @[simp] lemma deg_pow (c : num_denom_same_deg 𝒜 x) (n : ℕ) : (c ^ n).deg = n • c.deg := rfl @@ -187,7 +183,7 @@ section has_smul variables {α : Type*} [has_smul α R] [has_smul α A] [is_scalar_tower α R A] instance : has_smul α (num_denom_same_deg 𝒜 x) := -{ smul := λ m c, ⟨c.deg, m • c.num, c.denom, c.denom_not_mem⟩ } +{ smul := λ m c, ⟨c.deg, m • c.num, c.denom, c.denom_mem⟩ } @[simp] lemma deg_smul (c : num_denom_same_deg 𝒜 x) (m : α) : (m • c).deg = c.deg := rfl @[simp] lemma num_smul (c : num_denom_same_deg 𝒜 x) (m : α) : ((m • c).num : A) = m • c.num := rfl @@ -203,7 +199,7 @@ For `x : prime ideal of A` and any `p : num_denom_same_deg 𝒜 x`, or equivalen denominator of the same degree, we get an element `p.num / p.denom` of `Aₓ`. -/ def embedding (p : num_denom_same_deg 𝒜 x) : at x := -localization.mk p.num ⟨p.denom, p.denom_not_mem⟩ +localization.mk p.num ⟨p.denom, p.denom_mem⟩ end num_denom_same_deg @@ -231,12 +227,12 @@ def val (y : homogeneous_localization 𝒜 x) : at x := quotient.lift_on' y (num_denom_same_deg.embedding 𝒜 x) $ λ _ _, id @[simp] lemma val_mk' (i : num_denom_same_deg 𝒜 x) : - val (quotient.mk' i) = localization.mk i.num ⟨i.denom, i.denom_not_mem⟩ := + val (quotient.mk' i) = localization.mk i.num ⟨i.denom, i.denom_mem⟩ := rfl variable (x) lemma val_injective : - function.injective (@homogeneous_localization.val _ _ _ _ _ _ _ _ 𝒜 _ x _) := + function.injective (@homogeneous_localization.val _ _ _ _ _ _ _ _ 𝒜 _ x) := λ a b, quotient.rec_on_subsingleton₂' a b $ λ a b h, quotient.sound' h instance has_pow : has_pow (homogeneous_localization 𝒜 x) ℕ := @@ -390,10 +386,21 @@ show val (nat.unary_cast n) = _, by induction n; simp [nat.unary_cast, zero_val, @[simp] lemma int_cast_val (n : ℤ) : (n : homogeneous_localization 𝒜 x).val = n := show val (int.cast_def n) = _, by cases n; simp [int.cast_def, zero_val, one_val, *] -instance : comm_ring (homogeneous_localization 𝒜 x) := +instance homogenous_localization_comm_ring : comm_ring (homogeneous_localization 𝒜 x) := (homogeneous_localization.val_injective x).comm_ring _ zero_val one_val add_val mul_val neg_val sub_val (λ z n, smul_val x z n) (λ z n, smul_val x z n) pow_val nat_cast_val int_cast_val +instance homogeneous_localization_algebra : + algebra (homogeneous_localization 𝒜 x) (localization x) := +{ smul := λ p q, p.val * q, + to_fun := val, + map_one' := one_val, + map_mul' := mul_val, + map_zero' := zero_val, + map_add' := add_val, + commutes' := λ p q, mul_comm _ _, + smul_def' := λ p q, rfl } + end homogeneous_localization namespace homogeneous_localization @@ -415,18 +422,18 @@ def denom (f : homogeneous_localization 𝒜 x) : A := def deg (f : homogeneous_localization 𝒜 x) : ι := (quotient.out' f).deg -lemma denom_not_mem (f : homogeneous_localization 𝒜 x) : - f.denom ∉ x := -(quotient.out' f).denom_not_mem +lemma denom_mem (f : homogeneous_localization 𝒜 x) : + f.denom ∈ x := +(quotient.out' f).denom_mem -lemma num_mem (f : homogeneous_localization 𝒜 x) : f.num ∈ 𝒜 f.deg := +lemma num_mem_deg (f : homogeneous_localization 𝒜 x) : f.num ∈ 𝒜 f.deg := (quotient.out' f).num.2 -lemma denom_mem (f : homogeneous_localization 𝒜 x) : f.denom ∈ 𝒜 f.deg := +lemma denom_mem_deg (f : homogeneous_localization 𝒜 x) : f.denom ∈ 𝒜 f.deg := (quotient.out' f).denom.2 lemma eq_num_div_denom (f : homogeneous_localization 𝒜 x) : - f.val = localization.mk f.num ⟨f.denom, f.denom_not_mem⟩ := + f.val = localization.mk f.num ⟨f.denom, f.denom_mem⟩ := begin have := (quotient.out_eq' f), apply_fun homogeneous_localization.val at this, @@ -446,7 +453,14 @@ lemma ext_iff_val (f g : homogeneous_localization 𝒜 x) : f = g ↔ f.val = g. simpa only [quotient.lift_on'_mk] using h, end } -lemma is_unit_iff_is_unit_val (f : homogeneous_localization 𝒜 x) : +section + +variables (𝒜) (𝔭 : ideal A) [ideal.is_prime 𝔭] + +/--Localizing a ring homogeneously at a prime ideal-/ +abbreviation at_prime := homogeneous_localization 𝒜 𝔭.prime_compl + +lemma is_unit_iff_is_unit_val (f : homogeneous_localization.at_prime 𝒜 𝔭) : is_unit f.val ↔ is_unit f := ⟨λ h1, begin rcases h1 with ⟨⟨a, b, eq0, eq1⟩, (eq2 : a = f.val)⟩, @@ -455,23 +469,24 @@ lemma is_unit_iff_is_unit_val (f : homogeneous_localization 𝒜 x) : induction b using localization.induction_on with data, rcases data with ⟨a, ⟨b, hb⟩⟩, dsimp only at eq0 eq1, - have b_f_denom_not_mem : b * f.denom ∈ x.prime_compl := λ r, or.elim - (ideal.is_prime.mem_or_mem infer_instance r) (λ r2, hb r2) (λ r2, f.denom_not_mem r2), + have b_f_denom_not_mem : b * f.denom ∈ 𝔭.prime_compl := λ r, or.elim + (ideal.is_prime.mem_or_mem infer_instance r) (λ r2, hb r2) (λ r2, f.denom_mem r2), rw [f.eq_num_div_denom, localization.mk_mul, - show (⟨b, hb⟩ : x.prime_compl) * ⟨f.denom, _⟩ = ⟨b * f.denom, _⟩, from rfl, - show (1 : at x) = localization.mk 1 1, by erw localization.mk_self 1, + show (⟨b, hb⟩ : 𝔭.prime_compl) * ⟨f.denom, _⟩ = ⟨b * f.denom, _⟩, from rfl, + show (1 : localization.at_prime 𝔭) = localization.mk 1 1, by erw localization.mk_self 1, localization.mk_eq_mk', is_localization.eq] at eq1, rcases eq1 with ⟨⟨c, hc⟩, eq1⟩, simp only [← subtype.val_eq_coe] at eq1, change a * f.num * 1 * c = _ at eq1, simp only [one_mul, mul_one] at eq1, - have mem1 : a * f.num * c ∈ x.prime_compl := + have mem1 : a * f.num * c ∈ 𝔭.prime_compl := eq1.symm ▸ λ r, or.elim (ideal.is_prime.mem_or_mem infer_instance r) (by tauto)(by tauto), - have mem2 : f.num ∉ x, + have mem2 : f.num ∉ 𝔭, { contrapose! mem1, erw [not_not], exact ideal.mul_mem_right _ _ (ideal.mul_mem_left _ _ mem1), }, - refine ⟨⟨f, quotient.mk' ⟨f.deg, ⟨f.denom, f.denom_mem⟩, ⟨f.num, f.num_mem⟩, mem2⟩, _, _⟩, rfl⟩; + refine ⟨⟨f, quotient.mk' ⟨f.deg, ⟨f.denom, f.denom_mem_deg⟩, ⟨f.num, f.num_mem_deg⟩, mem2⟩, + _, _⟩, rfl⟩; simp only [ext_iff_val, mul_val, val_mk', ← subtype.val_eq_coe, f.eq_num_div_denom, localization.mk_mul, one_val]; convert localization.mk_self _; @@ -481,18 +496,18 @@ end, λ ⟨⟨_, b, eq1, eq2⟩, rfl⟩, begin exact ⟨⟨f.val, b.val, eq1, eq2⟩, rfl⟩ end⟩ -instance : nontrivial (homogeneous_localization 𝒜 x) := +instance : nontrivial (homogeneous_localization.at_prime 𝒜 𝔭) := ⟨⟨0, 1, λ r, by simpa [ext_iff_val, zero_val, one_val, zero_ne_one] using r⟩⟩ -instance : local_ring (homogeneous_localization 𝒜 x) := +instance : local_ring (homogeneous_localization.at_prime 𝒜 𝔭) := local_ring.of_is_unit_or_is_unit_one_sub_self $ λ a, begin simp only [← is_unit_iff_is_unit_val, sub_val, one_val], induction a using quotient.induction_on', simp only [homogeneous_localization.val_mk', ← subtype.val_eq_coe], - by_cases mem1 : a.num.1 ∈ x, + by_cases mem1 : a.num.1 ∈ 𝔭, { right, - have : a.denom.1 - a.num.1 ∈ x.prime_compl := λ h, a.denom_not_mem - ((sub_add_cancel a.denom.val a.num.val) ▸ ideal.add_mem _ h mem1 : a.denom.1 ∈ x), + have : a.denom.1 - a.num.1 ∈ 𝔭.prime_compl := λ h, a.denom_mem + ((sub_add_cancel a.denom.val a.num.val) ▸ ideal.add_mem _ h mem1 : a.denom.1 ∈ 𝔭), apply is_unit_of_mul_eq_one _ (localization.mk a.denom.1 ⟨a.denom.1 - a.num.1, this⟩), simp only [sub_mul, localization.mk_mul, one_mul, localization.sub_mk, ← subtype.val_eq_coe, submonoid.coe_mul], @@ -500,11 +515,22 @@ local_ring.of_is_unit_or_is_unit_one_sub_self $ λ a, begin simp only [← subtype.val_eq_coe, submonoid.coe_mul], ring, }, { left, - change _ ∈ x.prime_compl at mem1, + change _ ∈ 𝔭.prime_compl at mem1, apply is_unit_of_mul_eq_one _ (localization.mk a.denom.1 ⟨a.num.1, mem1⟩), rw [localization.mk_mul], convert localization.mk_self _, simpa only [mul_comm], }, end +end + +section + +variables (𝒜) (f : A) + +/--Localising away from powers of `f` homogeneously.-/ +abbreviation away := homogeneous_localization 𝒜 (submonoid.powers f) + +end + end homogeneous_localization From 5ca6e9d2bb25feeebdefe9bd9de163897bdecd05 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 3 Oct 2022 17:40:51 +0000 Subject: [PATCH 535/828] feat(logic/relation): characterization of the accessible predicate in terms of `well_founded_on` (#16720) + Show that `acc r a` (read: `a` is accessible under the relation `r`) if and only if `r` is well-founded on the subset of elements below `a` (transitively) under `r`, including (`refl_trans_gen`) or excluding (`trans_gen`) `a`. Stated as a `list.tfae`. + Move `relation.fibration` (used by the proof) to *logic/relation* to avoid importing *logic/hydra*. Add lemma `acc_trans_gen_iff` used in the proof. --- src/logic/hydra.lean | 37 +++------------------------------ src/logic/relation.lean | 36 +++++++++++++++++++++++++++++--- src/order/well_founded_set.lean | 23 ++++++++++++++++++++ 3 files changed, 59 insertions(+), 37 deletions(-) diff --git a/src/logic/hydra.lean b/src/logic/hydra.lean index 7a51cf7b3f246..b7ac0afeac559 100644 --- a/src/logic/hydra.lean +++ b/src/logic/hydra.lean @@ -5,7 +5,6 @@ Authors: Junyan Xu -/ import data.multiset.basic import order.game_add -import order.well_founded /-! # Termination of a hydra game @@ -24,9 +23,7 @@ valid "moves" of the game are modelled by the relation `cut_expand r` on `multis `cut_expand r s' s` is true iff `s'` is obtained by removing one head `a ∈ s` and adding back an arbitrary multiset `t` of heads such that all `a' ∈ t` satisfy `r a' a`. -To prove this theorem, we follow the proof by Peter LeFanu Lumsdaine at -https://mathoverflow.net/a/229084/3332, and along the way we introduce the notion of `fibration` -of relations. +We follow the proof by Peter LeFanu Lumsdaine at https://mathoverflow.net/a/229084/3332. TODO: formalize the relations corresponding to more powerful (e.g. Kirby–Paris and Buchholz) hydras, and prove their well-foundedness. @@ -34,37 +31,10 @@ hydras, and prove their well-foundedness. namespace relation -variables {α β : Type*} - -section fibration -variables (rα : α → α → Prop) (rβ : β → β → Prop) (f : α → β) - -/-- A function `f : α → β` is a fibration between the relation `rα` and `rβ` if for all - `a : α` and `b : β`, whenever `b : β` and `f a` are related by `rβ`, `b` is the image - of some `a' : α` under `f`, and `a'` and `a` are related by `rα`. -/ -def fibration := ∀ ⦃a b⦄, rβ b (f a) → ∃ a', rα a' a ∧ f a' = b - -variables {rα rβ} - -/-- If `f : α → β` is a fibration between relations `rα` and `rβ`, and `a : α` is - accessible under `rα`, then `f a` is accessible under `rβ`. -/ -lemma _root_.acc.of_fibration (fib : fibration rα rβ f) {a} (ha : acc rα a) : acc rβ (f a) := -begin - induction ha with a ha ih, - refine acc.intro (f a) (λ b hr, _), - obtain ⟨a', hr', rfl⟩ := fib hr, - exact ih a' hr', -end - -lemma _root_.acc.of_downward_closed (dc : ∀ {a b}, rβ b (f a) → b ∈ set.range f) - (a : α) (ha : acc (inv_image rβ f) a) : acc rβ (f a) := -ha.of_fibration f (λ a b h, let ⟨a', he⟩ := dc h in ⟨a', he.substr h, he⟩) - -end fibration - -section hydra open multiset prod +variables {α β : Type*} + /-- The relation that specifies valid moves in our hydra game. `cut_expand r s' s` means that `s'` is obtained by removing one head `a ∈ s` and adding back an arbitrary multiset `t` of heads such that all `a' ∈ t` satisfy `r a' a`. @@ -152,6 +122,5 @@ end /-- `cut_expand r` is well-founded when `r` is. -/ theorem _root_.well_founded.cut_expand (hr : well_founded r) : well_founded (cut_expand r) := ⟨by { letI h := hr.is_irrefl, exact λ s, acc_of_singleton $ λ a _, (hr.apply a).cut_expand }⟩ -end hydra end relation diff --git a/src/logic/relation.lean b/src/logic/relation.lean index 7a86397f89891..12abd7c2c4d82 100644 --- a/src/logic/relation.lean +++ b/src/logic/relation.lean @@ -147,6 +147,33 @@ end end comp +section fibration + +variables (rα : α → α → Prop) (rβ : β → β → Prop) (f : α → β) + +/-- A function `f : α → β` is a fibration between the relation `rα` and `rβ` if for all + `a : α` and `b : β`, whenever `b : β` and `f a` are related by `rβ`, `b` is the image + of some `a' : α` under `f`, and `a'` and `a` are related by `rα`. -/ +def fibration := ∀ ⦃a b⦄, rβ b (f a) → ∃ a', rα a' a ∧ f a' = b + +variables {rα rβ} + +/-- If `f : α → β` is a fibration between relations `rα` and `rβ`, and `a : α` is + accessible under `rα`, then `f a` is accessible under `rβ`. -/ +lemma _root_.acc.of_fibration (fib : fibration rα rβ f) {a} (ha : acc rα a) : acc rβ (f a) := +begin + induction ha with a ha ih, + refine acc.intro (f a) (λ b hr, _), + obtain ⟨a', hr', rfl⟩ := fib hr, + exact ih a' hr', +end + +lemma _root_.acc.of_downward_closed (dc : ∀ {a b}, rβ b (f a) → ∃ c, f c = b) + (a : α) (ha : acc (inv_image rβ f) a) : acc rβ (f a) := +ha.of_fibration f (λ a b h, let ⟨a', he⟩ := dc h in ⟨a', he.substr h, he⟩) + +end fibration + /-- The map of a relation `r` through a pair of functions pushes the relation to the codomains of the functions. The resulting relation is @@ -373,7 +400,7 @@ end end trans_gen -lemma _root_.acc.trans_gen {α} {r : α → α → Prop} {a : α} (h : acc r a) : acc (trans_gen r) a := +lemma _root_.acc.trans_gen (h : acc r a) : acc (trans_gen r) a := begin induction h with x _ H, refine acc.intro x (λ y hy, _), @@ -381,8 +408,11 @@ begin exacts [H y hyx, (H z hzx).inv hyz], end -lemma _root_.well_founded.trans_gen {α} {r : α → α → Prop} (h : well_founded r) : - well_founded (trans_gen r) := ⟨λ a, (h.apply a).trans_gen⟩ +lemma _root_.acc_trans_gen_iff : acc (trans_gen r) a ↔ acc r a := +⟨subrelation.accessible (λ _ _, trans_gen.single), acc.trans_gen⟩ + +lemma _root_.well_founded.trans_gen (h : well_founded r) : well_founded (trans_gen r) := +⟨λ a, (h.apply a).trans_gen⟩ section trans_gen diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index 5ae940977844e..47336543d5518 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -6,6 +6,7 @@ Authors: Aaron Anderson import order.antichain import order.order_iso_nat import order.well_founded +import tactic.tfae /-! # Well-founded sets @@ -96,6 +97,28 @@ end lemma subset (h : t.well_founded_on r) (hst : s ⊆ t) : s.well_founded_on r := h.mono le_rfl hst +open relation + +/-- `a` is accessible under the relation `r` iff `r` is well-founded on the downward transitive + closure of `a` under `r` (including `a` or not). -/ +lemma acc_iff_well_founded_on {α} {r : α → α → Prop} {a : α} : + [ acc r a, + {b | refl_trans_gen r b a}.well_founded_on r, + {b | trans_gen r b a}.well_founded_on r ].tfae := +begin + tfae_have : 1 → 2, + { refine λ h, ⟨λ b, _⟩, apply inv_image.accessible, + rw ← acc_trans_gen_iff at h ⊢, + obtain h'|h' := refl_trans_gen_iff_eq_or_trans_gen.1 b.2, + { rwa h' at h }, { exact h.inv h' } }, + tfae_have : 2 → 3, + { exact λ h, h.subset (λ _, trans_gen.to_refl) }, + tfae_have : 3 → 1, + { refine λ h, acc.intro _ (λ b hb, (h.apply ⟨b, trans_gen.single hb⟩).of_fibration subtype.val _), + exact λ ⟨c, hc⟩ d h, ⟨⟨d, trans_gen.head h hc⟩, h, rfl⟩ }, + tfae_finish, +end + end well_founded_on end any_rel From b0b0d5e3c424dce9695b03f2485d83a684e53e5a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 3 Oct 2022 17:40:52 +0000 Subject: [PATCH 536/828] feat(tactic/positivity): Extensions for `nat` constructions (#16728) Introduce the following `positivity` extensions: * `positivity_succ` * `positivity_factorial` * `positivity_asc_factorial` --- src/tactic/positivity.lean | 19 +++++++++++++++++++ test/positivity.lean | 24 ++++++++++++++---------- 2 files changed, 33 insertions(+), 10 deletions(-) diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index 2039db395f4f4..bfa185ef155cc 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -457,6 +457,25 @@ meta def positivity_coe : expr → tactic strictness end | _ := failed +/-- Extension for the `positivity` tactic: `nat.succ` is always positive. -/ +@[positivity] +meta def positivity_succ : expr → tactic strictness +| `(nat.succ %%a) := positive <$> mk_app `nat.succ_pos [a] +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `nat.succ n`" + +/-- Extension for the `positivity` tactic: `nat.factorial` is always positive. -/ +@[positivity] +meta def positivity_factorial : expr → tactic strictness +| `(nat.factorial %%a) := positive <$> mk_app ``nat.factorial_pos [a] +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `n!`" + +/-- Extension for the `positivity` tactic: `nat.asc_factorial` is always positive. -/ +@[positivity] +meta def positivity_asc_factorial : expr → tactic strictness +| `(nat.asc_factorial %%a %%b) := positive <$> mk_app ``nat.asc_factorial_pos [a, b] +| e := pp e >>= fail ∘ format.bracket "The expression `" + "` isn't of the form `nat.asc_factorial n k`" + private lemma card_univ_pos (α : Type*) [fintype α] [nonempty α] : 0 < (finset.univ : finset α).card := finset.univ_nonempty.card_pos diff --git a/test/positivity.lean b/test/positivity.lean index 1b41fa5b0343d..4cd0aab0a1c54 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -14,16 +14,7 @@ import tactic.positivity This tactic proves goals of the form `0 ≤ a` and `0 < a`. -/ -/- Test for instantiating meta-variables. Reported on -https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/New.20tactic.3A.20.60positivity.60/near/300639970 --/ -example : 0 ≤ 0 := -begin - apply le_trans _ le_rfl, - positivity, -end - -open_locale ennreal nnrat nnreal +open_locale ennreal nat nnrat nnreal universe u variables {α β : Type*} @@ -142,6 +133,10 @@ example : 0 ≤ max (-3 : ℤ) 5 := by positivity example [ordered_semiring α] [ordered_add_comm_monoid β] [smul_with_zero α β] [ordered_smul α β] {a : α} (ha : 0 < a) {b : β} (hb : 0 < b) : 0 ≤ a • b := by positivity +example (n : ℕ) : 0 < n.succ := by positivity +example (n : ℕ) : 0 < n! := by positivity +example (n k : ℕ) : 0 < n.asc_factorial k := by positivity + example {α : Type*} (s : finset α) (hs : s.nonempty) : 0 < s.card := by positivity example {α : Type*} [fintype α] [nonempty α] : 0 < fintype.card α := by positivity @@ -197,3 +192,12 @@ example {a : ℤ} (ha : a > 0) : 0 ≤ a := by positivity example {a : ℤ} (ha : 0 < a) : a ≥ 0 := by positivity example {a : ℤ} (ha : a > 0) : a ≥ 0 := by positivity + +/- +## Test for meta-variable instantiation + +Reported on +https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/New.20tactic.3A.20.60positivity.60/near/300639970 +-/ + +example : 0 ≤ 0 := by { apply le_trans _ le_rfl, positivity } From 6d5dda18b932258e8f39f24948c7759ebbd4c64c Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 3 Oct 2022 17:40:53 +0000 Subject: [PATCH 537/828] feat(geometry/euclidean/basic): angles and betweenness (#16734) Add some basic lemmas about the relation between (unoriented) angles and betweenness of points. --- src/geometry/euclidean/basic.lean | 135 ++++++++++++++++++++++++++++++ 1 file changed, 135 insertions(+) diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index fcfe5f33c0f0c..d7a4d6bbc8e23 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Manuel Candales -/ +import analysis.convex.between import analysis.inner_product_space.projection import analysis.special_functions.trigonometric.inverse import algebra.quadratic_discriminant @@ -578,6 +579,140 @@ lemma angle_right_midpoint_eq_pi_div_two_of_dist_eq {p1 p2 p3 : P} (h : dist p3 ∠ p3 (midpoint ℝ p1 p2) p2 = π / 2 := by rw [midpoint_comm p1 p2, angle_left_midpoint_eq_pi_div_two_of_dist_eq h.symm] +/-- If the second of three points is strictly between the other two, the angle at that point +is π. -/ +lemma _root_.sbtw.angle₁₂₃_eq_pi {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₁ p₂ p₃ = π := +begin + rw [angle, angle_eq_pi_iff], + rcases h with ⟨⟨r, ⟨hr0, hr1⟩, hp₂⟩, hp₂p₁, hp₂p₃⟩, + refine ⟨vsub_ne_zero.2 hp₂p₁.symm, -(1 - r) / r, _⟩, + have hr0' : r ≠ 0, + { rintro rfl, + rw ←hp₂ at hp₂p₁, + simpa using hp₂p₁ }, + have hr1' : r ≠ 1, + { rintro rfl, + rw ←hp₂ at hp₂p₃, + simpa using hp₂p₃ }, + replace hr0 := hr0.lt_of_ne hr0'.symm, + replace hr1 := hr1.lt_of_ne hr1', + refine ⟨div_neg_of_neg_of_pos (left.neg_neg_iff.2 (sub_pos.2 hr1)) hr0, _⟩, + rw [←hp₂, affine_map.line_map_apply, vsub_vadd_eq_vsub_sub, vsub_vadd_eq_vsub_sub, vsub_self, + zero_sub, smul_neg, smul_smul, div_mul_cancel _ hr0', neg_smul, neg_neg, sub_eq_iff_eq_add, + ←add_smul, sub_add_cancel, one_smul] +end + +/-- If the second of three points is strictly between the other two, the angle at that point +(reversed) is π. -/ +lemma _root_.sbtw.angle₃₂₁_eq_pi {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₃ p₂ p₁ = π := +by rw [←h.angle₁₂₃_eq_pi, angle_comm] + +/-- The angle between three points is π if and only if the second point is strictly between the +other two. -/ +lemma angle_eq_pi_iff_sbtw {p₁ p₂ p₃ : P} : ∠ p₁ p₂ p₃ = π ↔ sbtw ℝ p₁ p₂ p₃ := +begin + refine ⟨_, λ h, h.angle₁₂₃_eq_pi⟩, + rw [angle, angle_eq_pi_iff], + rintro ⟨hp₁p₂, r, hr, hp₃p₂⟩, + refine ⟨⟨1 / (1 - r), + ⟨div_nonneg zero_le_one (sub_nonneg.2 (hr.le.trans zero_le_one)), + (div_le_one (sub_pos.2 (hr.trans zero_lt_one))).2 ((le_sub_self_iff 1).2 hr.le)⟩, _⟩, + (vsub_ne_zero.1 hp₁p₂).symm, _⟩, + { rw ←eq_vadd_iff_vsub_eq at hp₃p₂, + rw [affine_map.line_map_apply, hp₃p₂, vadd_vsub_assoc, ←neg_vsub_eq_vsub_rev p₂ p₁, + smul_neg, ←neg_smul, smul_add, smul_smul, ←add_smul, eq_comm, eq_vadd_iff_vsub_eq], + convert (one_smul ℝ (p₂ -ᵥ p₁)).symm, + field_simp [(sub_pos.2 (hr.trans zero_lt_one)).ne.symm], + abel }, + { rw [ne_comm, ←@vsub_ne_zero V, hp₃p₂, smul_ne_zero_iff], + exact ⟨hr.ne, hp₁p₂⟩ } +end + +/-- If the second of three points is weakly between the other two, and not equal to the first, +the angle at the first point is zero. -/ +lemma _root_.wbtw.angle₂₁₃_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₁ : p₂ ≠ p₁) : + ∠ p₂ p₁ p₃ = 0 := +begin + rw [angle, angle_eq_zero_iff], + rcases h with ⟨r, ⟨hr0, hr1⟩, rfl⟩, + have hr0' : r ≠ 0, { rintro rfl, simpa using hp₂p₁ }, + replace hr0 := hr0.lt_of_ne hr0'.symm, + refine ⟨vsub_ne_zero.2 hp₂p₁, r⁻¹, inv_pos.2 hr0, _⟩, + rw [affine_map.line_map_apply, vadd_vsub_assoc, vsub_self, add_zero, smul_smul, + inv_mul_cancel hr0', one_smul] +end + +/-- If the second of three points is strictly between the other two, the angle at the first point +is zero. -/ +lemma _root_.sbtw.angle₂₁₃_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₂ p₁ p₃ = 0 := +h.wbtw.angle₂₁₃_eq_zero_of_ne h.ne_left + +/-- If the second of three points is weakly between the other two, and not equal to the first, +the angle at the first point (reversed) is zero. -/ +lemma _root_.wbtw.angle₃₁₂_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₁ : p₂ ≠ p₁) : + ∠ p₃ p₁ p₂ = 0 := +by rw [←h.angle₂₁₃_eq_zero_of_ne hp₂p₁, angle_comm] + +/-- If the second of three points is strictly between the other two, the angle at the first point +(reversed) is zero. -/ +lemma _root_.sbtw.angle₃₁₂_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₃ p₁ p₂ = 0 := +h.wbtw.angle₃₁₂_eq_zero_of_ne h.ne_left + +/-- If the second of three points is weakly between the other two, and not equal to the third, +the angle at the third point is zero. -/ +lemma _root_.wbtw.angle₂₃₁_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₃ : p₂ ≠ p₃) : + ∠ p₂ p₃ p₁ = 0 := +h.symm.angle₂₁₃_eq_zero_of_ne hp₂p₃ + +/-- If the second of three points is strictly between the other two, the angle at the third point +is zero. -/ +lemma _root_.sbtw.angle₂₃₁_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₂ p₃ p₁ = 0 := +h.wbtw.angle₂₃₁_eq_zero_of_ne h.ne_right + +/-- If the second of three points is weakly between the other two, and not equal to the third, +the angle at the third point (reversed) is zero. -/ +lemma _root_.wbtw.angle₁₃₂_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₃ : p₂ ≠ p₃) : + ∠ p₁ p₃ p₂ = 0 := +h.symm.angle₃₁₂_eq_zero_of_ne hp₂p₃ + +/-- If the second of three points is strictly between the other two, the angle at the third point +(reversed) is zero. -/ +lemma _root_.sbtw.angle₁₃₂_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₁ p₃ p₂ = 0 := +h.wbtw.angle₁₃₂_eq_zero_of_ne h.ne_right + +/-- The angle between three points is zero if and only if one of the first and third points is +weakly between the other two, and not equal to the second. -/ +lemma angle_eq_zero_iff_ne_and_wbtw {p₁ p₂ p₃ : P} : + ∠ p₁ p₂ p₃ = 0 ↔ (p₁ ≠ p₂ ∧ wbtw ℝ p₂ p₁ p₃) ∨ (p₃ ≠ p₂ ∧ wbtw ℝ p₂ p₃ p₁) := +begin + split, + { rw [angle, angle_eq_zero_iff], + rintro ⟨hp₁p₂, r, hr0, hp₃p₂⟩, + rcases le_or_lt 1 r with hr1 | hr1, + { refine or.inl ⟨vsub_ne_zero.1 hp₁p₂, r⁻¹, ⟨(inv_pos.2 hr0).le, inv_le_one hr1⟩, _⟩, + rw [affine_map.line_map_apply, hp₃p₂, smul_smul, inv_mul_cancel hr0.ne.symm, one_smul, + vsub_vadd] }, + { refine or.inr ⟨_, r, ⟨hr0.le, hr1.le⟩, _⟩, + { rw [←@vsub_ne_zero V, hp₃p₂, smul_ne_zero_iff], + exact ⟨hr0.ne.symm, hp₁p₂⟩ }, + { rw [affine_map.line_map_apply, ←hp₃p₂, vsub_vadd] } } }, + { rintro (⟨hp₁p₂, h⟩ | ⟨hp₃p₂, h⟩), + { exact h.angle₂₁₃_eq_zero_of_ne hp₁p₂ }, + { exact h.angle₃₁₂_eq_zero_of_ne hp₃p₂ } } +end + +/-- The angle between three points is zero if and only if one of the first and third points is +strictly between the other two, or those two points are equal but not equal to the second. -/ +lemma angle_eq_zero_iff_eq_and_ne_or_sbtw {p₁ p₂ p₃ : P} : + ∠ p₁ p₂ p₃ = 0 ↔ (p₁ = p₃ ∧ p₁ ≠ p₂) ∨ sbtw ℝ p₂ p₁ p₃ ∨ sbtw ℝ p₂ p₃ p₁ := +begin + rw angle_eq_zero_iff_ne_and_wbtw, + by_cases hp₁p₂ : p₁ = p₂, { simp [hp₁p₂] }, + by_cases hp₁p₃ : p₁ = p₃, { simp [hp₁p₃] }, + by_cases hp₃p₂ : p₃ = p₂, { simp [hp₃p₂] }, + simp [hp₁p₂, hp₁p₃, ne.symm hp₁p₃, sbtw, hp₃p₂] +end + /-- The inner product of two vectors given with `weighted_vsub`, in terms of the pairwise distances. -/ lemma inner_weighted_vsub {ι₁ : Type*} {s₁ : finset ι₁} {w₁ : ι₁ → ℝ} (p₁ : ι₁ → P) From 7802ba770c3ba6d9cc11bdb89fd852f7aff40598 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 3 Oct 2022 17:40:54 +0000 Subject: [PATCH 538/828] refactor(data/set/basic): review API of `set.nontrivial` (#16737) * make `z` an explicit argument in `set.nontrivial.exists_ne`; * rename `set.nontrivial.iff_exists_lt` to `set.nontrivial_iff_exists_lt`; * protect lemmas `set.nontrivial.nonempty` and `set.nontrivial.ne_empty`. --- src/data/set/basic.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 3a6274a4176a9..930f76a57c451 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1825,7 +1825,7 @@ lemma nontrivial_iff_pair_subset : s.nontrivial ↔ ∃ x y (hxy : x ≠ y), {x, lemma nontrivial_of_exists_ne {x} (hx : x ∈ s) (h : ∃ y ∈ s, y ≠ x) : s.nontrivial := let ⟨y, hy, hyx⟩ := h in ⟨y, hy, x, hx, hyx⟩ -lemma nontrivial.exists_ne {z} (hs : s.nontrivial) : ∃ x ∈ s, x ≠ z := +lemma nontrivial.exists_ne (hs : s.nontrivial) (z) : ∃ x ∈ s, x ≠ z := begin by_contra H, push_neg at H, rcases hs with ⟨x, hx, y, hy, hxy⟩, @@ -1834,7 +1834,7 @@ begin end lemma nontrivial_iff_exists_ne {x} (hx : x ∈ s) : s.nontrivial ↔ ∃ y ∈ s, y ≠ x := -⟨λ H, H.exists_ne, nontrivial_of_exists_ne hx⟩ +⟨λ H, H.exists_ne _, nontrivial_of_exists_ne hx⟩ lemma nontrivial_of_lt [preorder α] {x y} (hx : x ∈ s) (hy : y ∈ s) (hxy : x < y) : s.nontrivial := ⟨x, hx, y, hy, ne_of_lt hxy⟩ @@ -1846,12 +1846,13 @@ lemma nontrivial.exists_lt [linear_order α] (hs : s.nontrivial) : ∃ x y ∈ s let ⟨x, hx, y, hy, hxy⟩ := hs in or.elim (lt_or_gt_of_ne hxy) (λ H, ⟨x, hx, y, hy, H⟩) (λ H, ⟨y, hy, x, hx, H⟩) -lemma nontrivial.iff_exists_lt [linear_order α] : s.nontrivial ↔ ∃ x y ∈ s, x < y := +lemma nontrivial_iff_exists_lt [linear_order α] : s.nontrivial ↔ ∃ x y ∈ s, x < y := ⟨nontrivial.exists_lt, nontrivial_of_exists_lt⟩ -lemma nontrivial.nonempty (hs : s.nontrivial) : s.nonempty := let ⟨x, hx, _⟩ := hs in ⟨x, hx⟩ +protected lemma nontrivial.nonempty (hs : s.nontrivial) : s.nonempty := +let ⟨x, hx, _⟩ := hs in ⟨x, hx⟩ -lemma nontrivial.ne_empty (hs : s.nontrivial) : s ≠ ∅ := hs.nonempty.ne_empty +protected lemma nontrivial.ne_empty (hs : s.nontrivial) : s ≠ ∅ := hs.nonempty.ne_empty lemma nontrivial.not_subset_empty (hs : s.nontrivial) : ¬ s ⊆ ∅ := hs.nonempty.not_subset_empty From ce4a8b894ee9e8a0c4b415b463c31004e25da70a Mon Sep 17 00:00:00 2001 From: AlexKontorovich <58564076+AlexKontorovich@users.noreply.github.com> Date: Mon, 3 Oct 2022 17:40:55 +0000 Subject: [PATCH 539/828] feat(topology/algebra/group): quotient of a second countable group (#16738) Quotient of a second countable group is second countable Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> --- src/topology/algebra/const_mul_action.lean | 17 ++++++-- src/topology/algebra/group.lean | 7 +++ src/topology/algebra/monoid.lean | 5 ++- src/topology/bases.lean | 50 ++++++++++++++++++++++ 4 files changed, 75 insertions(+), 4 deletions(-) diff --git a/src/topology/algebra/const_mul_action.lean b/src/topology/algebra/const_mul_action.lean index f03025e3d0646..aa0a03f802a0d 100644 --- a/src/topology/algebra/const_mul_action.lean +++ b/src/topology/algebra/const_mul_action.lean @@ -7,6 +7,7 @@ import data.real.nnreal import topology.algebra.constructions import topology.homeomorph import group_theory.group_action.basic +import topology.bases /-! # Monoid actions continuous in the second variable @@ -37,7 +38,7 @@ Hausdorff, discrete group, properly discontinuous, quotient space open_locale topological_space pointwise -open filter set +open filter set topological_space local attribute [instance] mul_action.orbit_rel @@ -353,8 +354,10 @@ export properly_discontinuous_smul (finite_disjoint_inter_image) export properly_discontinuous_vadd (finite_disjoint_inter_image) -/-- The quotient map by a group action is open. -/ -@[to_additive "The quotient map by a group action is open."] +/-- The quotient map by a group action is open, i.e. the quotient by a group action is an open + quotient. -/ +@[to_additive "The quotient map by a group action is open, i.e. the quotient by a group +action is an open quotient. "] lemma is_open_map_quotient_mk_mul [has_continuous_const_smul Γ T] : is_open_map (quotient.mk : T → quotient (mul_action.orbit_rel Γ T)) := begin @@ -400,6 +403,14 @@ begin exact eq_empty_iff_forall_not_mem.mp H (γ • x) ⟨mem_image_of_mem _ x_in_K₀, h'⟩ }, end +/-- The quotient of a second countable space by a group action is second countable. -/ +@[to_additive "The quotient of a second countable space by an additive group action is second +countable."] +theorem has_continuous_const_smul.second_countable_topology [second_countable_topology T] + [has_continuous_const_smul Γ T] : + second_countable_topology (quotient (mul_action.orbit_rel Γ T)) := +topological_space.quotient.second_countable_topology is_open_map_quotient_mk_mul + section nhds section mul_action diff --git a/src/topology/algebra/group.lean b/src/topology/algebra/group.lean index 96f6b41799c14..2d89b6131302d 100644 --- a/src/topology/algebra/group.lean +++ b/src/topology/algebra/group.lean @@ -1210,6 +1210,13 @@ instance quotient_group.has_continuous_smul [locally_compact_space G] : exact quotient_map.continuous_lift_prod_right quotient_map_quotient_mk H, end } +/-- The quotient of a second countable topological group by a subgroup is second countable. -/ +@[to_additive "The quotient of a second countable additive topological group by a subgroup is second +countable."] +instance quotient_group.second_countable_topology [second_countable_topology G] : + second_countable_topology (G ⧸ Γ) := +has_continuous_const_smul.second_countable_topology + end quotient namespace units diff --git a/src/topology/algebra/monoid.lean b/src/topology/algebra/monoid.lean index 55b2fc8b7bcba..9f6fc51f299cb 100644 --- a/src/topology/algebra/monoid.lean +++ b/src/topology/algebra/monoid.lean @@ -456,7 +456,10 @@ instance is_scalar_tower.has_continuous_const_smul {R A : Type*} [monoid A] [has implies continuous scalar multiplication by constants. Notably, this instances applies when `R = Aᵐᵒᵖ` -/ -@[priority 100] +@[priority 100, to_additive "If the action of `R` on `A` commutes with left-addition, then +continuous addition implies continuous affine addition by constants. + +Notably, this instances applies when `R = Aᵃᵒᵖ`. "] instance smul_comm_class.has_continuous_const_smul {R A : Type*} [monoid A] [has_smul R A] [smul_comm_class R A A] [topological_space A] [has_continuous_mul A] : has_continuous_const_smul R A := diff --git a/src/topology/bases.lean b/src/topology/bases.lean index de9eec02d31e6..de84c646c4d37 100644 --- a/src/topology/bases.lean +++ b/src/topology/bases.lean @@ -809,6 +809,56 @@ end end sum +section quotient + +variables {X : Type*} [topological_space X] {Y : Type*} [topological_space Y] {π : X → Y} +omit t + +/-- The image of a topological basis under an open quotient map is a topological basis. -/ +lemma is_topological_basis.quotient_map {V : set (set X)} (hV : is_topological_basis V) + (h' : quotient_map π) (h : is_open_map π) : + is_topological_basis (set.image π '' V) := +begin + apply is_topological_basis_of_open_of_nhds, + { rintros - ⟨U, U_in_V, rfl⟩, + apply h U (hV.is_open U_in_V), }, + { intros y U y_in_U U_open, + obtain ⟨x, rfl⟩ := h'.surjective y, + let W := π ⁻¹' U, + have x_in_W : x ∈ W := y_in_U, + have W_open : is_open W := U_open.preimage h'.continuous, + obtain ⟨Z, Z_in_V, x_in_Z, Z_in_W⟩ := hV.exists_subset_of_mem_open x_in_W W_open, + have πZ_in_U : π '' Z ⊆ U := (set.image_subset _ Z_in_W).trans (image_preimage_subset π U), + exact ⟨π '' Z, ⟨Z, Z_in_V, rfl⟩, ⟨x, x_in_Z, rfl⟩, πZ_in_U⟩, }, +end + +/-- A second countable space is mapped by an open quotient map to a second countable space. -/ +lemma quotient_map.second_countable_topology [second_countable_topology X] (h' : quotient_map π) + (h : is_open_map π) : + second_countable_topology Y := +{ is_open_generated_countable := + begin + obtain ⟨V, V_countable, V_no_empty, V_generates⟩ := exists_countable_basis X, + exact ⟨set.image π '' V, V_countable.image (set.image π), + (V_generates.quotient_map h' h).eq_generate_from⟩, + end } + +variables {S : setoid X} + +/-- The image of a topological basis "downstairs" in an open quotient is a topological basis. -/ +lemma is_topological_basis.quotient {V : set (set X)} + (hV : is_topological_basis V) (h : is_open_map (quotient.mk : X → quotient S)) : + is_topological_basis (set.image (quotient.mk : X → quotient S) '' V) := +hV.quotient_map quotient_map_quotient_mk h + +/-- An open quotient of a second countable space is second countable. -/ +lemma quotient.second_countable_topology [second_countable_topology X] + (h : is_open_map (quotient.mk : X → quotient S)) : + second_countable_topology (quotient S) := +quotient_map_quotient_mk.second_countable_topology h + +end quotient + end topological_space open topological_space From b40ce7d04d8e40f54e1faeeac6e90492a697a0a2 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 3 Oct 2022 17:40:56 +0000 Subject: [PATCH 540/828] feat(measure_theory/integral/average): average of constant functions (#16776) Also make the API more coherent, by renaming `average_def'` to `average_eq` to match `set_average_eq`. And rename `lintegral_to_real_le_lintegral_nnnorm` to `lintegral_of_real_le_lintegral_nnnorm`. --- src/analysis/convex/integral.lean | 6 ++-- src/analysis/normed/group/basic.lean | 3 ++ src/measure_theory/function/l1_space.lean | 4 +++ src/measure_theory/integral/average.lean | 31 ++++++++++++++----- .../integral/integrable_on.lean | 2 +- src/measure_theory/integral/lebesgue.lean | 2 +- 6 files changed, 36 insertions(+), 12 deletions(-) diff --git a/src/analysis/convex/integral.lean b/src/analysis/convex/integral.lean index bcef36ed9e9f8..9c95f27474cd2 100644 --- a/src/analysis/convex/integral.lean +++ b/src/analysis/convex/integral.lean @@ -327,8 +327,8 @@ begin simp only [average_congr this, pi.zero_apply, average_zero], exact or.inl this }, by_cases hfi : integrable f μ, swap, - by simp [average_def', integral_undef hfi, hC0, ennreal.to_real_pos_iff], - cases (le_top : μ univ ≤ ∞).eq_or_lt with hμt hμt, { simp [average_def', hμt, hC0] }, + by simp [average_eq, integral_undef hfi, hC0, ennreal.to_real_pos_iff], + cases (le_top : μ univ ≤ ∞).eq_or_lt with hμt hμt, { simp [average_eq, hμt, hC0] }, haveI : is_finite_measure μ := ⟨hμt⟩, replace h_le : ∀ᵐ x ∂μ, f x ∈ closed_ball (0 : E) C, by simpa only [mem_closed_ball_zero_iff], simpa only [interior_closed_ball _ hC0.ne', mem_ball_zero_iff] @@ -347,7 +347,7 @@ begin have hμ : 0 < (μ univ).to_real, by simp [ennreal.to_real_pos_iff, pos_iff_ne_zero, h₀, measure_lt_top], refine (ae_eq_const_or_norm_average_lt_of_norm_le_const h_le).imp_right (λ H, _), - rwa [average_def', norm_smul, norm_inv, real.norm_eq_abs, abs_of_pos hμ, + rwa [average_eq, norm_smul, norm_inv, real.norm_eq_abs, abs_of_pos hμ, ← div_eq_inv_mul, div_lt_iff' hμ] at H end diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 305d481a27a85..62620f2883ef1 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -658,6 +658,9 @@ nnreal.coe_le_coe.1 $ norm_add_le g h @[simp] lemma nnnorm_neg (g : E) : ∥-g∥₊ = ∥g∥₊ := nnreal.eq $ norm_neg g +lemma nnnorm_sub_le (g h : E) : ∥g - h∥₊ ≤ ∥g∥₊ + ∥h∥₊ := +nnreal.coe_le_coe.1 $ norm_sub_le g h + lemma nndist_nnnorm_nnnorm_le (g h : E) : nndist ∥g∥₊ ∥h∥₊ ≤ ∥g - h∥₊ := nnreal.coe_le_coe.1 $ dist_norm_norm_le g h diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 552e0764bfc5a..4f6ac00fd7914 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -509,6 +509,10 @@ lemma integrable_smul_measure {f : α → β} {c : ℝ≥0∞} (h₁ : c ≠ 0) ⟨λ h, by simpa only [smul_smul, ennreal.inv_mul_cancel h₁ h₂, one_smul] using h.smul_measure (ennreal.inv_ne_top.2 h₁), λ h, h.smul_measure h₂⟩ +lemma integrable_inv_smul_measure {f : α → β} {c : ℝ≥0∞} (h₁ : c ≠ 0) (h₂ : c ≠ ∞) : + integrable f (c⁻¹ • μ) ↔ integrable f μ := +integrable_smul_measure (by simpa using h₂) (by simpa using h₁) + lemma integrable.to_average {f : α → β} (h : integrable f μ) : integrable f ((μ univ)⁻¹ • μ) := begin diff --git a/src/measure_theory/integral/average.lean b/src/measure_theory/integral/average.lean index 191635b0bd6cd..7ec0fab852c75 100644 --- a/src/measure_theory/integral/average.lean +++ b/src/measure_theory/integral/average.lean @@ -71,10 +71,10 @@ by rw [average, smul_zero, integral_zero_measure] @[simp] lemma average_neg (f : α → E) : ⨍ x, -f x ∂μ = -⨍ x, f x ∂μ := integral_neg f -lemma average_def (f : α → E) : ⨍ x, f x ∂μ = ∫ x, f x ∂((μ univ)⁻¹ • μ) := rfl +lemma average_eq' (f : α → E) : ⨍ x, f x ∂μ = ∫ x, f x ∂((μ univ)⁻¹ • μ) := rfl -lemma average_def' (f : α → E) : ⨍ x, f x ∂μ = (μ univ).to_real⁻¹ • ∫ x, f x ∂μ := -by rw [average_def, integral_smul_measure, ennreal.to_real_inv] +lemma average_eq (f : α → E) : ⨍ x, f x ∂μ = (μ univ).to_real⁻¹ • ∫ x, f x ∂μ := +by rw [average_eq', integral_smul_measure, ennreal.to_real_inv] lemma average_eq_integral [is_probability_measure μ] (f : α → E) : ⨍ x, f x ∂μ = ∫ x, f x ∂μ := @@ -85,19 +85,23 @@ by rw [average, measure_univ, inv_one, one_smul] begin cases eq_or_ne μ 0 with hμ hμ, { rw [hμ, integral_zero_measure, average_zero_measure, smul_zero] }, - { rw [average_def', smul_inv_smul₀], + { rw [average_eq, smul_inv_smul₀], refine (ennreal.to_real_pos _ $ measure_ne_top _ _).ne', rwa [ne.def, measure_univ_eq_zero] } end lemma set_average_eq (f : α → E) (s : set α) : ⨍ x in s, f x ∂μ = (μ s).to_real⁻¹ • ∫ x in s, f x ∂μ := -by rw [average_def', restrict_apply_univ] +by rw [average_eq, restrict_apply_univ] + +lemma set_average_eq' (f : α → E) (s : set α) : + ⨍ x in s, f x ∂μ = ∫ x, f x ∂((μ s)⁻¹ • μ.restrict s) := +by simp only [average_eq', restrict_apply_univ] variable {μ} lemma average_congr {f g : α → E} (h : f =ᵐ[μ] g) : ⨍ x, f x ∂μ = ⨍ x, g x ∂μ := -by simp only [average_def', integral_congr_ae h] +by simp only [average_eq, integral_congr_ae h] lemma average_add_measure [is_finite_measure μ] {ν : measure α} [is_finite_measure ν] {f : α → E} (hμ : integrable f μ) (hν : integrable f ν) : @@ -107,7 +111,7 @@ lemma average_add_measure [is_finite_measure μ] {ν : measure α} [is_finite_me begin simp only [div_eq_inv_mul, mul_smul, measure_smul_average, ← smul_add, ← integral_add_measure hμ hν, ← ennreal.to_real_add (measure_ne_top μ _) (measure_ne_top ν _)], - rw [average_def', measure.add_apply] + rw [average_eq, measure.add_apply] end lemma average_pair {f : α → E} {g : α → F} (hfi : integrable f μ) (hgi : integrable g μ) : @@ -162,4 +166,17 @@ by simpa only [union_compl_self, restrict_univ] using average_union_mem_open_segment ae_disjoint_compl_right hs.compl hs₀ hsc₀ (measure_ne_top _ _) (measure_ne_top _ _) hfi.integrable_on hfi.integrable_on +@[simp] lemma average_const [is_finite_measure μ] [h : μ.ae.ne_bot] (c : E) : + ⨍ x, c ∂μ = c := +by simp only [average_eq, integral_const, measure.restrict_apply, measurable_set.univ, one_smul, + univ_inter, smul_smul, ← ennreal.to_real_inv, ← ennreal.to_real_mul, ennreal.inv_mul_cancel, + measure_ne_top μ univ, ne.def, measure_univ_eq_zero, ae_ne_bot.1 h, not_false_iff, + ennreal.one_to_real] + +lemma set_average_const {s : set α} (hs₀ : μ s ≠ 0) (hs : μ s ≠ ∞) (c : E) : + ⨍ x in s, c ∂μ = c := +by simp only [set_average_eq, integral_const, measure.restrict_apply, measurable_set.univ, + univ_inter, smul_smul, ← ennreal.to_real_inv, ← ennreal.to_real_mul, + ennreal.inv_mul_cancel hs₀ hs, ennreal.one_to_real, one_smul] + end measure_theory diff --git a/src/measure_theory/integral/integrable_on.lean b/src/measure_theory/integral/integrable_on.lean index e97a455cab5b3..e4425cf536eba 100644 --- a/src/measure_theory/integral/integrable_on.lean +++ b/src/measure_theory/integral/integrable_on.lean @@ -249,7 +249,7 @@ end lemma integrable.lintegral_lt_top {f : α → ℝ} (hf : integrable f μ) : ∫⁻ x, ennreal.of_real (f x) ∂μ < ∞ := calc ∫⁻ x, ennreal.of_real (f x) ∂μ - ≤ ∫⁻ x, ↑∥f x∥₊ ∂μ : lintegral_to_real_le_lintegral_nnnorm f + ≤ ∫⁻ x, ↑∥f x∥₊ ∂μ : lintegral_of_real_le_lintegral_nnnorm f ... < ∞ : hf.2 lemma integrable_on.set_lintegral_lt_top {f : α → ℝ} {s : set α} (hf : integrable_on f s μ) : diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index ff220f69465d5..c60e524d6800d 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1227,7 +1227,7 @@ lemma set_lintegral_congr_fun {f g : α → ℝ≥0∞} {s : set α} (hs : measu ∫⁻ x in s, f x ∂μ = ∫⁻ x in s, g x ∂μ := by { rw lintegral_congr_ae, rw eventually_eq, rwa ae_restrict_iff' hs, } -lemma lintegral_to_real_le_lintegral_nnnorm (f : α → ℝ) : +lemma lintegral_of_real_le_lintegral_nnnorm (f : α → ℝ) : ∫⁻ x, ennreal.of_real (f x) ∂μ ≤ ∫⁻ x, ∥f x∥₊ ∂μ := begin simp_rw ← of_real_norm_eq_coe_nnnorm, From 075b3f7d19b9da85a0b54b3e33055a74fc388dec Mon Sep 17 00:00:00 2001 From: mcdoll Date: Mon, 3 Oct 2022 21:49:23 +0000 Subject: [PATCH 541/828] feat(linear_algebra/matrix): Add matrix properties for sesquilinear forms (#15906) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main goal of this PR is to move `linear_algebra/matrix/bilinear_form` to the curried bilinear map form. Since this file as quite few dependencies we copy it to a new file `linear_algebra/matrix/sesquilinear_form` and then move all the dependencies. The structure is taken literally from `linear_algebra/matrix/bilinear_form` with generalizations and easier proofs where possible. The new lemmas are named exactly as the old ones with the following changes: - the namespace changed from `bilin_form` to `linear_map` - `bilin` is changed to `linear_map₂`. In case there is the necessity for separate bilinear and semi-bilinear lemmas we use `linear_map₂` and `linear_mapₛₗ₂` - If `bilin` is not in the name of a lemma, `matrix` is changed to `matrix₂` to avoid nameclashes with lemmas for linear maps `M →ₛₗ[ρ₁₂] N` Moreover, the following changes were necessary: `linear_algebra/bilinear_map`: - Weakened some typeclass assumptions - Added bilinear version of `sum_repr_mul_repr_mul` and moved sesquilinear version to `sum_repr_mul_repr_mulₛₗ` `linear_algebra/matrix/bilinear_form`: - Moved `basis.equiv_fun_symm_std_basis` to `linear_algebra/std_basis` - `adjoint_pair` section: Removed a few definitions (they are now in `linear_algebra/matrix/sesquilinear_form`) and added a prime to the names of lemmas that have the same name as the version in `linear_algebra/matrix/sesquilinear_form` `linear_algebra/sesquilinear_form`: - Added a few missing lemmas about left-separating bilinear forms (note that `separating_left` was previously known as `nondegenerate`) `linear_algebra/std_basis`: - Lemma `std_basis_apply'` for calculating the application of `i' : ι` to `(std_basis R (λ (_x : ι), R) i)` `algebra/hom/ring/`: - Lemmas for calculating a ring homomorphism applied to `ite 0 1` and `ite 1 0` The last two additions are needed to get a reasonable proof for `matrix.to_linear_map₂'_aux_std_basis`. Co-authored-by: Moritz Doll --- src/algebra/hom/ring.lean | 8 + src/algebra/lie/skew_adjoint.lean | 2 +- src/linear_algebra/bilinear_map.lean | 59 +- src/linear_algebra/finsupp_vector_space.lean | 36 + src/linear_algebra/matrix/bilinear_form.lean | 58 +- .../matrix/sesquilinear_form.lean | 641 ++++++++++++++++++ src/linear_algebra/sesquilinear_form.lean | 42 +- src/linear_algebra/std_basis.lean | 7 + 8 files changed, 792 insertions(+), 61 deletions(-) create mode 100644 src/linear_algebra/matrix/sesquilinear_form.lean diff --git a/src/algebra/hom/ring.lean b/src/algebra/hom/ring.lean index b6dcd88cabb5b..1181c9e69f957 100644 --- a/src/algebra/hom/ring.lean +++ b/src/algebra/hom/ring.lean @@ -391,6 +391,14 @@ protected lemma map_bit0 (f : α →+* β) : ∀ a, f (bit0 a) = bit0 (f a) := m /-- Ring homomorphisms preserve `bit1`. -/ protected lemma map_bit1 (f : α →+* β) : ∀ a, f (bit1 a) = bit1 (f a) := map_bit1 f +@[simp] lemma map_ite_zero_one {F : Type*} [ring_hom_class F α β] (f : F) (p : Prop) [decidable p] : + f (ite p 0 1) = ite p 0 1 := +by { split_ifs; simp [h] } + +@[simp] lemma map_ite_one_zero {F : Type*} [ring_hom_class F α β] (f : F) (p : Prop) [decidable p] : + f (ite p 1 0) = ite p 1 0 := +by { split_ifs; simp [h] } + /-- `f : α →+* β` has a trivial codomain iff `f 1 = 0`. -/ lemma codomain_trivial_iff_map_one_eq_zero : (0 : β) = 1 ↔ f 1 = 0 := by rw [map_one, eq_comm] diff --git a/src/algebra/lie/skew_adjoint.lean b/src/algebra/lie/skew_adjoint.lean index 4c3ea50651cbe..4242eeeb37ec8 100644 --- a/src/algebra/lie/skew_adjoint.lean +++ b/src/algebra/lie/skew_adjoint.lean @@ -119,7 +119,7 @@ begin A ∈ skew_adjoint_matrices_submodule (Pᵀ ⬝ J ⬝ P), { simp only [lie_subalgebra.mem_coe, submodule.mem_map_equiv, lie_subalgebra.mem_map_submodule, coe_coe], exact this, }, - simp [matrix.is_skew_adjoint, J.is_adjoint_pair_equiv _ _ P (is_unit_of_invertible P)], + simp [matrix.is_skew_adjoint, J.is_adjoint_pair_equiv' _ _ P (is_unit_of_invertible P)], end lemma skew_adjoint_matrices_lie_subalgebra_equiv_apply diff --git a/src/linear_algebra/bilinear_map.lean b/src/linear_algebra/bilinear_map.lean index 5a4559067bb1e..2d9681ce2e4b9 100644 --- a/src/linear_algebra/bilinear_map.lean +++ b/src/linear_algebra/bilinear_map.lean @@ -256,6 +256,9 @@ include σ₄₃ f.compl₂ g m q = f m (g q) := rfl omit σ₄₃ +@[simp] theorem compl₂_id : f.compl₂ linear_map.id = f := +by { ext, rw [compl₂_apply, id_coe, id.def] } + /-- Composing linear maps `Q → M` and `Q' → N` with a bilinear map `M → N → P` to form a bilinear map `Q → Q' → P`. -/ def compl₁₂ (f : Mₗ →ₗ[R] Nₗ →ₗ[R] Pₗ) (g : Qₗ →ₗ[R] Mₗ) (g' : Qₗ' →ₗ[R] Nₗ) : @@ -265,6 +268,10 @@ def compl₁₂ (f : Mₗ →ₗ[R] Nₗ →ₗ[R] Pₗ) (g : Qₗ →ₗ[R] M @[simp] theorem compl₁₂_apply (f : Mₗ →ₗ[R] Nₗ →ₗ[R] Pₗ) (g : Qₗ →ₗ[R] Mₗ) (g' : Qₗ' →ₗ[R] Nₗ) (x : Qₗ) (y : Qₗ') : f.compl₁₂ g g' x y = f (g x) (g' y) := rfl +@[simp] theorem compl₁₂_id_id (f : Mₗ →ₗ[R] Nₗ →ₗ[R] Pₗ) : + f.compl₁₂ (linear_map.id) (linear_map.id) = f := +by { ext, simp_rw [compl₁₂_apply, id_coe, id.def] } + lemma compl₁₂_inj {f₁ f₂ : Mₗ →ₗ[R] Nₗ →ₗ[R] Pₗ} {g : Qₗ →ₗ[R] Mₗ} {g' : Qₗ' →ₗ[R] Nₗ} (hₗ : function.surjective g) (hᵣ : function.surjective g') : f₁.compl₁₂ g g' = f₂.compl₁₂ g g' ↔ f₁ = f₂ := @@ -301,20 +308,18 @@ end comm_semiring section comm_ring variables {R R₂ S S₂ M N P : Type*} +variables {Mₗ Nₗ Pₗ : Type*} variables [comm_ring R] [comm_ring S] [comm_ring R₂] [comm_ring S₂] -variables [add_comm_group M] [add_comm_group N] [add_comm_group P] + +section add_comm_monoid + +variables [add_comm_monoid M] [add_comm_monoid N] [add_comm_monoid P] +variables [add_comm_monoid Mₗ] [add_comm_monoid Nₗ] [add_comm_monoid Pₗ] variables [module R M] [module S N] [module R₂ P] [module S₂ P] +variables [module R Mₗ] [module R Nₗ] [module R Pₗ] variables [smul_comm_class S₂ R₂ P] variables {ρ₁₂ : R →+* R₂} {σ₁₂ : S →+* S₂} -variables (b₁ : basis ι₁ R M) (b₂ : basis ι₂ S N) - -lemma lsmul_injective [no_zero_smul_divisors R M] {x : R} (hx : x ≠ 0) : - function.injective (lsmul R M x) := -smul_right_injective _ hx - -lemma ker_lsmul [no_zero_smul_divisors R M] {a : R} (ha : a ≠ 0) : - (linear_map.lsmul R M a).ker = ⊥ := -linear_map.ker_eq_bot_of_injective (linear_map.lsmul_injective ha) +variables (b₁ : basis ι₁ R M) (b₂ : basis ι₂ S N) (b₁' : basis ι₁ R Mₗ) (b₂' : basis ι₂ R Nₗ) /-- Two bilinear maps are equal when they are equal on all basis vectors. -/ @@ -322,8 +327,10 @@ lemma ext_basis {B B' : M →ₛₗ[ρ₁₂] N →ₛₗ[σ₁₂] P} (h : ∀ i j, B (b₁ i) (b₂ j) = B' (b₁ i) (b₂ j)) : B = B' := b₁.ext $ λ i, b₂.ext $ λ j, h i j -/-- Write out `B x y` as a sum over `B (b i) (b j)` if `b` is a basis. -/ -lemma sum_repr_mul_repr_mul {B : M →ₛₗ[ρ₁₂] N →ₛₗ[σ₁₂] P} (x y) : +/-- Write out `B x y` as a sum over `B (b i) (b j)` if `b` is a basis. + +Version for semi-bilinear maps, see `sum_repr_mul_repr_mul` for the bilinear version. -/ +lemma sum_repr_mul_repr_mulₛₗ {B : M →ₛₗ[ρ₁₂] N →ₛₗ[σ₁₂] P} (x y) : (b₁.repr x).sum (λ i xi, (b₂.repr y).sum (λ j yj, (ρ₁₂ xi) • (σ₁₂ yj) • B (b₁ i) (b₂ j))) = B x y := begin @@ -332,6 +339,34 @@ begin linear_map.map_smulₛₗ₂, linear_map.map_smulₛₗ], end +/-- Write out `B x y` as a sum over `B (b i) (b j)` if `b` is a basis. + +Version for bilinear maps, see `sum_repr_mul_repr_mulₛₗ` for the semi-bilinear version. -/ +lemma sum_repr_mul_repr_mul {B : Mₗ →ₗ[R] Nₗ →ₗ[R] Pₗ} (x y) : + (b₁'.repr x).sum (λ i xi, (b₂'.repr y).sum (λ j yj, xi • yj • B (b₁' i) (b₂' j))) = + B x y := +begin + conv_rhs { rw [← b₁'.total_repr x, ← b₂'.total_repr y] }, + simp_rw [finsupp.total_apply, finsupp.sum, map_sum₂, map_sum, + linear_map.map_smul₂, linear_map.map_smul], +end + +end add_comm_monoid + +section add_comm_group + +variables [add_comm_group M] [add_comm_group N] [add_comm_group P] +variables [module R M] [module S N] [module R₂ P] [module S₂ P] + +lemma lsmul_injective [no_zero_smul_divisors R M] {x : R} (hx : x ≠ 0) : + function.injective (lsmul R M x) := +smul_right_injective _ hx + +lemma ker_lsmul [no_zero_smul_divisors R M] {a : R} (ha : a ≠ 0) : + (linear_map.lsmul R M a).ker = ⊥ := +linear_map.ker_eq_bot_of_injective (linear_map.lsmul_injective ha) + +end add_comm_group end comm_ring diff --git a/src/linear_algebra/finsupp_vector_space.lean b/src/linear_algebra/finsupp_vector_space.lean index b65a171d4ed5e..a73b6fea50d56 100644 --- a/src/linear_algebra/finsupp_vector_space.lean +++ b/src/linear_algebra/finsupp_vector_space.lean @@ -210,3 +210,39 @@ begin end end module + +namespace basis + +variables {R M n : Type*} +variables [decidable_eq n] [fintype n] +variables [semiring R] [add_comm_monoid M] [module R M] + +lemma _root_.finset.sum_single_ite (a : R) (i : n) : + finset.univ.sum (λ (x : n), finsupp.single x (ite (i = x) a 0)) = finsupp.single i a := +begin + rw finset.sum_congr_set {i} (λ (x : n), finsupp.single x (ite (i = x) a 0)) + (λ _, finsupp.single i a), + { simp }, + { intros x hx, + rw set.mem_singleton_iff at hx, + simp [hx] }, + intros x hx, + have hx' : ¬i = x := + begin + refine ne_comm.mp _, + rwa mem_singleton_iff at hx, + end, + simp [hx'], +end + +@[simp] lemma equiv_fun_symm_std_basis (b : basis n R M) (i : n) : + b.equiv_fun.symm (linear_map.std_basis R (λ _, R) i 1) = b i := +begin + have := equiv_like.injective b.repr, + apply_fun b.repr, + simp only [equiv_fun_symm_apply, std_basis_apply', linear_equiv.map_sum, + linear_equiv.map_smulₛₗ, ring_hom.id_apply, repr_self, finsupp.smul_single', boole_mul], + exact finset.sum_single_ite 1 i, +end + +end basis diff --git a/src/linear_algebra/matrix/bilinear_form.lean b/src/linear_algebra/matrix/bilinear_form.lean index 26506c84985c9..9e198c8f1b5b0 100644 --- a/src/linear_algebra/matrix/bilinear_form.lean +++ b/src/linear_algebra/matrix/bilinear_form.lean @@ -9,6 +9,7 @@ import linear_algebra.matrix.nondegenerate import linear_algebra.matrix.nonsingular_inverse import linear_algebra.matrix.to_linear_equiv import linear_algebra.bilinear_form +import linear_algebra.matrix.sesquilinear_form /-! # Bilinear form @@ -245,18 +246,6 @@ noncomputable def bilin_form.to_matrix : bilin_form R₂ M₂ ≃ₗ[R₂] matri noncomputable def matrix.to_bilin : matrix n n R₂ ≃ₗ[R₂] bilin_form R₂ M₂ := (bilin_form.to_matrix b).symm -@[simp] lemma basis.equiv_fun_symm_std_basis (i : n) : - b.equiv_fun.symm (std_basis R₂ (λ _, R₂) i 1) = b i := -begin - rw [b.equiv_fun_symm_apply, finset.sum_eq_single i], - { rw [std_basis_same, one_smul] }, - { rintros j - hj, - rw [std_basis_ne _ _ _ _ hj, zero_smul] }, - { intro, - have := mem_univ i, - contradiction } -end - @[simp] lemma bilin_form.to_matrix_apply (B : bilin_form R₂ M₂) (i j : n) : bilin_form.to_matrix b B i j = B (b i) (b j) := by rw [bilin_form.to_matrix, linear_equiv.trans_apply, bilin_form.to_matrix'_apply, congr_apply, @@ -373,18 +362,6 @@ variables {n : Type*} [fintype n] variables (b : basis n R₃ M₃) variables (J J₃ A A' : matrix n n R₃) -/-- The condition for the square matrices `A`, `A'` to be an adjoint pair with respect to the square -matrices `J`, `J₃`. -/ -def matrix.is_adjoint_pair := Aᵀ ⬝ J₃ = J ⬝ A' - -/-- The condition for a square matrix `A` to be self-adjoint with respect to the square matrix -`J`. -/ -def matrix.is_self_adjoint := matrix.is_adjoint_pair J J A A - -/-- The condition for a square matrix `A` to be skew-adjoint with respect to the square matrix -`J`. -/ -def matrix.is_skew_adjoint := matrix.is_adjoint_pair J J A (-A) - @[simp] lemma is_adjoint_pair_to_bilin' [decidable_eq n] : bilin_form.is_adjoint_pair (matrix.to_bilin' J) (matrix.to_bilin' J₃) (matrix.to_lin' A) (matrix.to_lin' A') ↔ @@ -421,7 +398,7 @@ begin refl, end -lemma matrix.is_adjoint_pair_equiv [decidable_eq n] (P : matrix n n R₃) (h : is_unit P) : +lemma matrix.is_adjoint_pair_equiv' [decidable_eq n] (P : matrix n n R₃) (h : is_unit P) : (Pᵀ ⬝ J ⬝ P).is_adjoint_pair (Pᵀ ⬝ J ⬝ P) A A' ↔ J.is_adjoint_pair J (P ⬝ A ⬝ P⁻¹) (P ⬝ A' ⬝ P⁻¹) := have h' : is_unit P.det := P.is_unit_iff_is_unit_det.mp h, @@ -444,45 +421,32 @@ variables [decidable_eq n] /-- The submodule of pair-self-adjoint matrices with respect to bilinear forms corresponding to given matrices `J`, `J₂`. -/ -def pair_self_adjoint_matrices_submodule : submodule R₃ (matrix n n R₃) := +def pair_self_adjoint_matrices_submodule' : submodule R₃ (matrix n n R₃) := (bilin_form.is_pair_self_adjoint_submodule (matrix.to_bilin' J) (matrix.to_bilin' J₃)).map ((linear_map.to_matrix' : ((n → R₃) →ₗ[R₃] (n → R₃)) ≃ₗ[R₃] matrix n n R₃) : ((n → R₃) →ₗ[R₃] (n → R₃)) →ₗ[R₃] matrix n n R₃) -@[simp] lemma mem_pair_self_adjoint_matrices_submodule : +lemma mem_pair_self_adjoint_matrices_submodule' : A ∈ (pair_self_adjoint_matrices_submodule J J₃) ↔ matrix.is_adjoint_pair J J₃ A A := -begin - simp only [pair_self_adjoint_matrices_submodule, linear_equiv.coe_coe, - linear_map.to_matrix'_apply, submodule.mem_map, bilin_form.mem_is_pair_self_adjoint_submodule], - split, - { rintros ⟨f, hf, hA⟩, - have hf' : f = A.to_lin' := by rw [←hA, matrix.to_lin'_to_matrix'], rw hf' at hf, - rw ← is_adjoint_pair_to_bilin', - exact hf, }, - { intros h, refine ⟨A.to_lin', _, linear_map.to_matrix'_to_lin' _⟩, - exact (is_adjoint_pair_to_bilin' _ _ _ _).mpr h, }, -end +by simp only [mem_pair_self_adjoint_matrices_submodule] /-- The submodule of self-adjoint matrices with respect to the bilinear form corresponding to the matrix `J`. -/ -def self_adjoint_matrices_submodule : submodule R₃ (matrix n n R₃) := +def self_adjoint_matrices_submodule' : submodule R₃ (matrix n n R₃) := pair_self_adjoint_matrices_submodule J J -@[simp] lemma mem_self_adjoint_matrices_submodule : +lemma mem_self_adjoint_matrices_submodule' : A ∈ self_adjoint_matrices_submodule J ↔ J.is_self_adjoint A := -by { erw mem_pair_self_adjoint_matrices_submodule, refl, } +by simp only [mem_self_adjoint_matrices_submodule] /-- The submodule of skew-adjoint matrices with respect to the bilinear form corresponding to the matrix `J`. -/ -def skew_adjoint_matrices_submodule : submodule R₃ (matrix n n R₃) := +def skew_adjoint_matrices_submodule' : submodule R₃ (matrix n n R₃) := pair_self_adjoint_matrices_submodule (-J) J -@[simp] lemma mem_skew_adjoint_matrices_submodule : +lemma mem_skew_adjoint_matrices_submodule' : A ∈ skew_adjoint_matrices_submodule J ↔ J.is_skew_adjoint A := -begin - erw mem_pair_self_adjoint_matrices_submodule, - simp [matrix.is_skew_adjoint, matrix.is_adjoint_pair], -end +by simp only [mem_skew_adjoint_matrices_submodule] end matrix_adjoints diff --git a/src/linear_algebra/matrix/sesquilinear_form.lean b/src/linear_algebra/matrix/sesquilinear_form.lean new file mode 100644 index 0000000000000..b2aa3ec2bcf12 --- /dev/null +++ b/src/linear_algebra/matrix/sesquilinear_form.lean @@ -0,0 +1,641 @@ +/- +Copyright (c) 2020 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen, Kexing Ying, Moritz Doll +-/ + +import linear_algebra.finsupp_vector_space +import linear_algebra.matrix.basis +import linear_algebra.matrix.nondegenerate +import linear_algebra.matrix.nonsingular_inverse +import linear_algebra.matrix.to_linear_equiv +import linear_algebra.sesquilinear_form + +/-! +# Sesquilinear form + +This file defines the conversion between sesquilinear forms and matrices. + +## Main definitions + + * `matrix.to_linear_map₂` given a basis define a bilinear form + * `matrix.to_linear_map₂'` define the bilinear form on `n → R` + * `linear_map.to_matrix₂`: calculate the matrix coefficients of a bilinear form + * `linear_map.to_matrix₂'`: calculate the matrix coefficients of a bilinear form on `n → R` + +## Todos + +At the moment this is quite a literal port from `matrix.bilinear_form`. Everything should be +generalized to fully semibilinear forms. + +## Tags + +sesquilinear_form, matrix, basis + +-/ + +variables {R R₁ R₂ M M₁ M₂ M₁' M₂' n m n' m' ι : Type*} + +open_locale big_operators +open finset linear_map matrix +open_locale matrix + +section aux_to_linear_map + +variables [comm_semiring R] [comm_semiring R₁] [comm_semiring R₂] +variables [fintype n] [fintype m] + +variables (σ₁ : R₁ →+* R) (σ₂ : R₂ →+* R) + + +/-- The map from `matrix n n R` to bilinear forms on `n → R`. + +This is an auxiliary definition for the equivalence `matrix.to_linear_map₂'`. -/ +def matrix.to_linear_map₂'_aux (f : matrix n m R) : + (n → R₁) →ₛₗ[σ₁] (m → R₂) →ₛₗ[σ₂] R := +mk₂'ₛₗ σ₁ σ₂ (λ (v : n → R₁) (w : m → R₂), ∑ i j, σ₁ (v i) * f i j * σ₂ (w j)) + (λ _ _ _, by simp only [pi.add_apply, map_add, add_mul, sum_add_distrib] ) + (λ _ _ _, by simp only [pi.smul_apply, smul_eq_mul, ring_hom.map_mul, mul_assoc, mul_sum] ) + (λ _ _ _, by simp only [pi.add_apply, map_add, mul_add, sum_add_distrib] ) + (λ _ _ _, by + simp only [pi.smul_apply, smul_eq_mul, ring_hom.map_mul, mul_assoc, mul_left_comm, mul_sum] ) + +variables [decidable_eq n] [decidable_eq m] + +lemma matrix.to_linear_map₂'_aux_std_basis (f : matrix n m R) (i : n) (j : m) : + f.to_linear_map₂'_aux σ₁ σ₂ (std_basis R₁ (λ _, R₁) i 1) (std_basis R₂ (λ _, R₂) j 1) = f i j := +begin + rw [matrix.to_linear_map₂'_aux, mk₂'ₛₗ_apply], + have : (∑ i' j', (if i = i' then 1 else 0) * f i' j' * (if j = j' then 1 else 0)) = f i j := + begin + simp_rw [mul_assoc, ←finset.mul_sum], + simp only [boole_mul, finset.sum_ite_eq, finset.mem_univ, if_true, mul_comm (f _ _)], + end, + rw ←this, + exact finset.sum_congr rfl (λ _ _, finset.sum_congr rfl (λ _ _, by simp)), +end + +end aux_to_linear_map + +section aux_to_matrix + +section comm_semiring + +variables [comm_semiring R] [comm_semiring R₁] [comm_semiring R₂] +variables [add_comm_monoid M₁] [module R₁ M₁] [add_comm_monoid M₂] [module R₂ M₂] + +variables {σ₁ : R₁ →+* R} {σ₂ : R₂ →+* R} + +/-- The linear map from sesquilinear forms to `matrix n m R` given an `n`-indexed basis for `M₁` +and an `m`-indexed basis for `M₂`. + +This is an auxiliary definition for the equivalence `matrix.to_linear_mapₛₗ₂'`. -/ +def linear_map.to_matrix₂_aux (b₁ : n → M₁) (b₂ : m → M₂) : + (M₁ →ₛₗ[σ₁] M₂ →ₛₗ[σ₂] R) →ₗ[R] matrix n m R := +{ to_fun := λ f, of $ λ i j, f (b₁ i) (b₂ j), + map_add' := λ f g, rfl, + map_smul' := λ f g, rfl } + +@[simp] lemma linear_map.to_matrix₂_aux_apply (f : M₁ →ₛₗ[σ₁] M₂ →ₛₗ[σ₂] R) + (b₁ : n → M₁) (b₂ : m → M₂) (i : n) (j : m) : + linear_map.to_matrix₂_aux b₁ b₂ f i j = f (b₁ i) (b₂ j) := rfl + +end comm_semiring + +section comm_ring + +variables [comm_ring R] [comm_ring R₁] [comm_ring R₂] +variables [add_comm_monoid M₁] [module R₁ M₁] [add_comm_monoid M₂] [module R₂ M₂] +variables [fintype n] [fintype m] +variables [decidable_eq n] [decidable_eq m] + +variables {σ₁ : R₁ →+* R} {σ₂ : R₂ →+* R} + +lemma linear_map.to_linear_map₂'_aux_to_matrix₂_aux (f : (n → R₁) →ₛₗ[σ₁] (m → R₂) →ₛₗ[σ₂] R) : + matrix.to_linear_map₂'_aux σ₁ σ₂ (linear_map.to_matrix₂_aux + (λ i, std_basis R₁ (λ _, R₁) i 1) (λ j, std_basis R₂ (λ _, R₂) j 1) f) = f := +begin + refine ext_basis (pi.basis_fun R₁ n) (pi.basis_fun R₂ m) (λ i j, _), + simp_rw [pi.basis_fun_apply, matrix.to_linear_map₂'_aux_std_basis, + linear_map.to_matrix₂_aux_apply], +end + +lemma matrix.to_matrix₂_aux_to_linear_map₂'_aux (f : matrix n m R) : + linear_map.to_matrix₂_aux (λ i, std_basis R₁ (λ _, R₁) i 1) (λ j, std_basis R₂ (λ _, R₂) j 1) + (f.to_linear_map₂'_aux σ₁ σ₂) = f := +by { ext i j, simp_rw [linear_map.to_matrix₂_aux_apply, matrix.to_linear_map₂'_aux_std_basis] } + +end comm_ring + +end aux_to_matrix + +section to_matrix' + +/-! ### Bilinear forms over `n → R` + +This section deals with the conversion between matrices and sesquilinear forms on `n → R`. +-/ + +variables [comm_ring R] [comm_ring R₁] [comm_ring R₂] +variables [fintype n] [fintype m] +variables [decidable_eq n] [decidable_eq m] + +variables {σ₁ : R₁ →+* R} {σ₂ : R₂ →+* R} + +/-- The linear equivalence between sesquilinear forms and `n × m` matrices -/ +def linear_map.to_matrixₛₗ₂' : ((n → R₁) →ₛₗ[σ₁] (m → R₂) →ₛₗ[σ₂] R) ≃ₗ[R] matrix n m R := +{ to_fun := linear_map.to_matrix₂_aux _ _, + inv_fun := matrix.to_linear_map₂'_aux σ₁ σ₂, + left_inv := linear_map.to_linear_map₂'_aux_to_matrix₂_aux, + right_inv := matrix.to_matrix₂_aux_to_linear_map₂'_aux, + ..linear_map.to_matrix₂_aux (λ i, std_basis R₁ (λ _, R₁) i 1) (λ j, std_basis R₂ (λ _, R₂) j 1) } + +/-- The linear equivalence between bilinear forms and `n × m` matrices -/ +def linear_map.to_matrix₂' : ((n → R) →ₗ[R] (m → R) →ₗ[R] R) ≃ₗ[R] matrix n m R := +linear_map.to_matrixₛₗ₂' + +variables (σ₁ σ₂) + +/-- The linear equivalence between `n × n` matrices and sesquilinear forms on `n → R` -/ +def matrix.to_linear_mapₛₗ₂' : matrix n m R ≃ₗ[R] ((n → R₁) →ₛₗ[σ₁] (m → R₂) →ₛₗ[σ₂] R) := +linear_map.to_matrixₛₗ₂'.symm + +/-- The linear equivalence between `n × n` matrices and bilinear forms on `n → R` -/ +def matrix.to_linear_map₂' : matrix n m R ≃ₗ[R] ((n → R) →ₗ[R] (m → R) →ₗ[R] R) := +linear_map.to_matrix₂'.symm + +lemma matrix.to_linear_mapₛₗ₂'_aux_eq (M : matrix n m R) : + matrix.to_linear_map₂'_aux σ₁ σ₂ M = matrix.to_linear_mapₛₗ₂' σ₁ σ₂ M := rfl + +lemma matrix.to_linear_mapₛₗ₂'_apply (M : matrix n m R) (x : n → R₁) (y : m → R₂) : + matrix.to_linear_mapₛₗ₂' σ₁ σ₂ M x y = ∑ i j, σ₁ (x i) * M i j * σ₂ (y j) := rfl + +lemma matrix.to_linear_map₂'_apply (M : matrix n m R) (x : n → R) (y : m → R) : + matrix.to_linear_map₂' M x y = ∑ i j, x i * M i j * y j := rfl + +lemma matrix.to_linear_map₂'_apply' (M : matrix n m R) (v : n → R) (w : m → R) : + matrix.to_linear_map₂' M v w = matrix.dot_product v (M.mul_vec w) := +begin + simp_rw [matrix.to_linear_map₂'_apply, matrix.dot_product, + matrix.mul_vec, matrix.dot_product], + refine finset.sum_congr rfl (λ _ _, _), + rw finset.mul_sum, + refine finset.sum_congr rfl (λ _ _, _), + rw ← mul_assoc, +end + +@[simp] lemma matrix.to_linear_mapₛₗ₂'_std_basis (M : matrix n m R) (i : n) (j : m) : + matrix.to_linear_mapₛₗ₂' σ₁ σ₂ M (std_basis R₁ (λ _, R₁) i 1) (std_basis R₂ (λ _, R₂) j 1) = + M i j := +matrix.to_linear_map₂'_aux_std_basis σ₁ σ₂ M i j + +@[simp] lemma matrix.to_linear_map₂'_std_basis (M : matrix n m R) (i : n) (j : m) : + matrix.to_linear_map₂' M (std_basis R (λ _, R) i 1) (std_basis R (λ _, R) j 1) = + M i j := +matrix.to_linear_map₂'_aux_std_basis _ _ M i j + +@[simp] lemma linear_map.to_matrixₛₗ₂'_symm : + (linear_map.to_matrixₛₗ₂'.symm : matrix n m R ≃ₗ[R] _) = matrix.to_linear_mapₛₗ₂' σ₁ σ₂ := +rfl + +@[simp] lemma matrix.to_linear_mapₛₗ₂'_symm : + ((matrix.to_linear_mapₛₗ₂' σ₁ σ₂).symm : _ ≃ₗ[R] matrix n m R) = linear_map.to_matrixₛₗ₂' := +linear_map.to_matrixₛₗ₂'.symm_symm + +@[simp] lemma matrix.to_linear_mapₛₗ₂'_to_matrix' (B : (n → R₁) →ₛₗ[σ₁] (m → R₂) →ₛₗ[σ₂] R) : + matrix.to_linear_mapₛₗ₂' σ₁ σ₂ (linear_map.to_matrixₛₗ₂' B) = B := +(matrix.to_linear_mapₛₗ₂' σ₁ σ₂).apply_symm_apply B + +@[simp] lemma matrix.to_linear_map₂'_to_matrix' (B : (n → R) →ₗ[R] (m → R) →ₗ[R] R) : + matrix.to_linear_map₂' (linear_map.to_matrix₂' B) = B := +matrix.to_linear_map₂'.apply_symm_apply B + +@[simp] lemma linear_map.to_matrix'_to_linear_mapₛₗ₂' (M : matrix n m R) : + linear_map.to_matrixₛₗ₂' (matrix.to_linear_mapₛₗ₂' σ₁ σ₂ M) = M := +linear_map.to_matrixₛₗ₂'.apply_symm_apply M + +@[simp] lemma linear_map.to_matrix'_to_linear_map₂' (M : matrix n m R) : + linear_map.to_matrix₂' (matrix.to_linear_map₂' M) = M := +linear_map.to_matrixₛₗ₂'.apply_symm_apply M + +@[simp] lemma linear_map.to_matrixₛₗ₂'_apply (B : (n → R₁) →ₛₗ[σ₁] (m → R₂) →ₛₗ[σ₂] R) (i : n) + (j : m): linear_map.to_matrixₛₗ₂' B i j = + B (std_basis R₁ (λ _, R₁) i 1) (std_basis R₂ (λ _, R₂) j 1) := +rfl + +@[simp] lemma linear_map.to_matrix₂'_apply (B : (n → R) →ₗ[R] (m → R) →ₗ[R] R) (i : n) (j : m): + linear_map.to_matrix₂' B i j = + B (std_basis R (λ _, R) i 1) (std_basis R (λ _, R) j 1) := +rfl + +variables [fintype n'] [fintype m'] +variables [decidable_eq n'] [decidable_eq m'] + +@[simp] lemma linear_map.to_matrix₂'_compl₁₂ (B : (n → R) →ₗ[R] (m → R) →ₗ[R] R) + (l : (n' → R) →ₗ[R] (n → R)) (r : (m' → R) →ₗ[R] (m → R)) : + (B.compl₁₂ l r).to_matrix₂' = l.to_matrix'ᵀ ⬝ B.to_matrix₂' ⬝ r.to_matrix' := +begin + ext i j, + simp only [linear_map.to_matrix₂'_apply, linear_map.compl₁₂_apply, transpose_apply, + matrix.mul_apply, linear_map.to_matrix', linear_equiv.coe_mk, sum_mul], + rw sum_comm, + conv_lhs { rw ←linear_map.sum_repr_mul_repr_mul (pi.basis_fun R n) (pi.basis_fun R m) + (l _) (r _) }, + rw finsupp.sum_fintype, + { apply sum_congr rfl, + rintros i' -, + rw finsupp.sum_fintype, + { apply sum_congr rfl, + rintros j' -, + simp only [smul_eq_mul, pi.basis_fun_repr, mul_assoc, mul_comm, mul_left_comm, + pi.basis_fun_apply, of_apply] }, + { intros, simp only [zero_smul, smul_zero] } }, + { intros, simp only [zero_smul, finsupp.sum_zero] } +end + +lemma linear_map.to_matrix₂'_comp (B : (n → R) →ₗ[R] (m → R) →ₗ[R] R) + (f : (n' → R) →ₗ[R] (n → R)) : (B.comp f).to_matrix₂' = f.to_matrix'ᵀ ⬝ B.to_matrix₂' := +by { rw [←linear_map.compl₂_id (B.comp f), ←linear_map.compl₁₂], simp } + +lemma linear_map.to_matrix₂'_compl₂ (B : (n → R) →ₗ[R] (m → R) →ₗ[R] R) + (f : (m' → R) →ₗ[R] (m → R)) : (B.compl₂ f).to_matrix₂' = B.to_matrix₂' ⬝ f.to_matrix' := +by { rw [←linear_map.comp_id B, ←linear_map.compl₁₂], simp } + +lemma linear_map.mul_to_matrix₂'_mul (B : (n → R) →ₗ[R] (m → R) →ₗ[R] R) + (M : matrix n' n R) (N : matrix m m' R) : + M ⬝ B.to_matrix₂' ⬝ N = (B.compl₁₂ Mᵀ.to_lin' N.to_lin').to_matrix₂' := +by simp + +lemma linear_map.mul_to_matrix' (B : (n → R) →ₗ[R] (m → R) →ₗ[R] R) (M : matrix n' n R) : + M ⬝ B.to_matrix₂' = (B.comp Mᵀ.to_lin').to_matrix₂' := +by simp only [B.to_matrix₂'_comp, transpose_transpose, to_matrix'_to_lin'] + +lemma linear_map.to_matrix₂'_mul (B : (n → R) →ₗ[R] (m → R) →ₗ[R] R) (M : matrix m m' R) : + B.to_matrix₂' ⬝ M = (B.compl₂ M.to_lin').to_matrix₂' := +by simp only [B.to_matrix₂'_compl₂, to_matrix'_to_lin'] + +lemma matrix.to_linear_map₂'_comp (M : matrix n m R) (P : matrix n n' R) (Q : matrix m m' R) : + M.to_linear_map₂'.compl₁₂ P.to_lin' Q.to_lin' = (Pᵀ ⬝ M ⬝ Q).to_linear_map₂' := +linear_map.to_matrix₂'.injective (by simp) + +end to_matrix' + +section to_matrix + +/-! ### Bilinear forms over arbitrary vector spaces + +This section deals with the conversion between matrices and bilinear forms on +a module with a fixed basis. +-/ + +variables [comm_ring R] +variables [add_comm_monoid M₁] [module R M₁] [add_comm_monoid M₂] [module R M₂] + +variables [decidable_eq n] [fintype n] +variables [decidable_eq m] [fintype m] +variables (b₁ : basis n R M₁) (b₂ : basis m R M₂) + +/-- `linear_map.to_matrix₂ b₁ b₂` is the equivalence between `R`-bilinear forms on `M` and +`n`-by-`n` matrices with entries in `R`, if `b` is an `R`-basis for `M`. -/ +noncomputable def linear_map.to_matrix₂ : (M₁ →ₗ[R] M₂ →ₗ[R] R) ≃ₗ[R] matrix n m R := +(b₁.equiv_fun.arrow_congr (b₂.equiv_fun.arrow_congr (linear_equiv.refl R R))).trans + linear_map.to_matrix₂' + +/-- `bilin_form.to_matrix b` is the equivalence between `R`-bilinear forms on `M` and +`n`-by-`n` matrices with entries in `R`, if `b` is an `R`-basis for `M`. -/ +noncomputable def matrix.to_linear_map₂ : matrix n m R ≃ₗ[R] (M₁ →ₗ[R] M₂ →ₗ[R] R) := +(linear_map.to_matrix₂ b₁ b₂).symm + +-- We make this and not `linear_map.to_matrix₂` a `simp` lemma to avoid timeouts +@[simp] lemma linear_map.to_matrix₂_apply (B : M₁ →ₗ[R] M₂ →ₗ[R] R) (i : n) (j : m) : + linear_map.to_matrix₂ b₁ b₂ B i j = B (b₁ i) (b₂ j) := +by simp only [linear_map.to_matrix₂, linear_equiv.trans_apply, linear_map.to_matrix₂'_apply, + linear_equiv.trans_apply, linear_map.to_matrix₂'_apply, linear_equiv.arrow_congr_apply, + basis.equiv_fun_symm_std_basis, linear_equiv.refl_apply] + +@[simp] lemma matrix.to_linear_map₂_apply (M : matrix n m R) (x : M₁) (y : M₂) : + matrix.to_linear_map₂ b₁ b₂ M x y = ∑ i j, b₁.repr x i * M i j * b₂.repr y j := +begin + rw [matrix.to_linear_map₂, linear_map.to_matrix₂, linear_equiv.symm_trans_apply, + ←matrix.to_linear_map₂'], + simp [matrix.to_linear_map₂'_apply], +end + +-- Not a `simp` lemma since `linear_map.to_matrix₂` needs an extra argument +lemma linear_map.to_matrix₂_aux_eq (B : M₁ →ₗ[R] M₂ →ₗ[R] R) : + linear_map.to_matrix₂_aux b₁ b₂ B = linear_map.to_matrix₂ b₁ b₂ B := +ext (λ i j, by rw [linear_map.to_matrix₂_apply, linear_map.to_matrix₂_aux_apply]) + +@[simp] lemma linear_map.to_matrix₂_symm : + (linear_map.to_matrix₂ b₁ b₂).symm = matrix.to_linear_map₂ b₁ b₂ := +rfl + +@[simp] lemma matrix.to_linear_map₂_symm : + (matrix.to_linear_map₂ b₁ b₂).symm = linear_map.to_matrix₂ b₁ b₂ := +(linear_map.to_matrix₂ b₁ b₂).symm_symm + +lemma matrix.to_linear_map₂_basis_fun : + matrix.to_linear_map₂ (pi.basis_fun R n) (pi.basis_fun R m) = matrix.to_linear_map₂' := +by { ext M, simp only [matrix.to_linear_map₂_apply, matrix.to_linear_map₂'_apply, pi.basis_fun_repr, + coe_comp, function.comp_app]} + +lemma linear_map.to_matrix₂_basis_fun : + linear_map.to_matrix₂ (pi.basis_fun R n) (pi.basis_fun R m) = linear_map.to_matrix₂' := +by { ext B, rw [linear_map.to_matrix₂_apply, linear_map.to_matrix₂'_apply, + pi.basis_fun_apply, pi.basis_fun_apply] } + +@[simp] lemma matrix.to_linear_map₂_to_matrix₂ (B : M₁ →ₗ[R] M₂ →ₗ[R] R) : + matrix.to_linear_map₂ b₁ b₂ (linear_map.to_matrix₂ b₁ b₂ B) = B := +(matrix.to_linear_map₂ b₁ b₂).apply_symm_apply B + +@[simp] lemma linear_map.to_matrix₂_to_linear_map₂ (M : matrix n m R) : + linear_map.to_matrix₂ b₁ b₂ (matrix.to_linear_map₂ b₁ b₂ M) = M := +(linear_map.to_matrix₂ b₁ b₂).apply_symm_apply M + +variables [add_comm_monoid M₁'] [module R M₁'] +variables [add_comm_monoid M₂'] [module R M₂'] +variables (b₁' : basis n' R M₁') +variables (b₂' : basis m' R M₂') +variables [fintype n'] [fintype m'] +variables [decidable_eq n'] [decidable_eq m'] + +-- Cannot be a `simp` lemma because `b₁` and `b₂` must be inferred. +lemma linear_map.to_matrix₂_compl₁₂ + (B : M₁ →ₗ[R] M₂ →ₗ[R] R) (l : M₁' →ₗ[R] M₁) (r : M₂' →ₗ[R] M₂) : + linear_map.to_matrix₂ b₁' b₂' (B.compl₁₂ l r) = + (to_matrix b₁' b₁ l)ᵀ ⬝ linear_map.to_matrix₂ b₁ b₂ B ⬝ to_matrix b₂' b₂ r := +begin + ext i j, + simp only [linear_map.to_matrix₂_apply, compl₁₂_apply, transpose_apply, matrix.mul_apply, + linear_map.to_matrix_apply, linear_equiv.coe_mk, sum_mul], + rw sum_comm, + conv_lhs { rw ←linear_map.sum_repr_mul_repr_mul b₁ b₂ }, + rw finsupp.sum_fintype, + { apply sum_congr rfl, + rintros i' -, + rw finsupp.sum_fintype, + { apply sum_congr rfl, + rintros j' -, + simp only [smul_eq_mul, linear_map.to_matrix_apply, + basis.equiv_fun_apply, mul_assoc, mul_comm, mul_left_comm] }, + { intros, simp only [zero_smul, smul_zero] } }, + { intros, simp only [zero_smul, finsupp.sum_zero] } +end + +lemma linear_map.to_matrix₂_comp (B : M₁ →ₗ[R] M₂ →ₗ[R] R) (f : M₁' →ₗ[R] M₁) : + linear_map.to_matrix₂ b₁' b₂ (B.comp f) = (to_matrix b₁' b₁ f)ᵀ ⬝ linear_map.to_matrix₂ b₁ b₂ B := +begin + rw [←linear_map.compl₂_id (B.comp f), ←linear_map.compl₁₂, linear_map.to_matrix₂_compl₁₂ b₁ b₂], + simp, +end + +lemma linear_map.to_matrix₂_compl₂ (B : M₁ →ₗ[R] M₂ →ₗ[R] R) (f : M₂' →ₗ[R] M₂) : + linear_map.to_matrix₂ b₁ b₂' (B.compl₂ f) = + linear_map.to_matrix₂ b₁ b₂ B ⬝ (to_matrix b₂' b₂ f) := +by { rw [←linear_map.comp_id B, ←linear_map.compl₁₂, linear_map.to_matrix₂_compl₁₂ b₁ b₂], simp } + +@[simp] +lemma linear_map.to_matrix₂_mul_basis_to_matrix (c₁ : basis n' R M₁) (c₂ : basis m' R M₂) + (B : M₁ →ₗ[R] M₂ →ₗ[R] R) : (b₁.to_matrix c₁)ᵀ ⬝ linear_map.to_matrix₂ b₁ b₂ B ⬝ b₂.to_matrix c₂ = + linear_map.to_matrix₂ c₁ c₂ B := +begin + simp_rw ←linear_map.to_matrix_id_eq_basis_to_matrix, + rw [←linear_map.to_matrix₂_compl₁₂, linear_map.compl₁₂_id_id], +end + +lemma linear_map.mul_to_matrix₂_mul (B : M₁ →ₗ[R] M₂ →ₗ[R] R) + (M : matrix n' n R) (N : matrix m m' R) : + M ⬝ linear_map.to_matrix₂ b₁ b₂ B ⬝ N = + linear_map.to_matrix₂ b₁' b₂' (B.compl₁₂ (to_lin b₁' b₁ Mᵀ) (to_lin b₂' b₂ N)) := +by simp_rw [linear_map.to_matrix₂_compl₁₂ b₁ b₂, to_matrix_to_lin, transpose_transpose] + +lemma linear_map.mul_to_matrix₂ (B : M₁ →ₗ[R] M₂ →ₗ[R] R) (M : matrix n' n R) : + M ⬝ linear_map.to_matrix₂ b₁ b₂ B = + linear_map.to_matrix₂ b₁' b₂ (B.comp (to_lin b₁' b₁ Mᵀ)) := +by rw [linear_map.to_matrix₂_comp b₁, to_matrix_to_lin, transpose_transpose] + +lemma linear_map.to_matrix₂_mul (B : M₁ →ₗ[R] M₂ →ₗ[R] R) (M : matrix m m' R) : + linear_map.to_matrix₂ b₁ b₂ B ⬝ M = + linear_map.to_matrix₂ b₁ b₂' (B.compl₂ (to_lin b₂' b₂ M)) := +by rw [linear_map.to_matrix₂_compl₂ b₁, to_matrix_to_lin] + +lemma matrix.to_linear_map₂_compl₁₂ (M : matrix n m R) (P : matrix n n' R) (Q : matrix m m' R) : + (matrix.to_linear_map₂ b₁ b₂ M).compl₁₂ (to_lin b₁' b₁ P) (to_lin b₂' b₂ Q) = + matrix.to_linear_map₂ b₁' b₂' (Pᵀ ⬝ M ⬝ Q) := +(linear_map.to_matrix₂ b₁' b₂').injective + (by simp only [linear_map.to_matrix₂_compl₁₂ b₁ b₂, linear_map.to_matrix₂_to_linear_map₂, + to_matrix_to_lin]) + +end to_matrix + +/-! ### Adjoint pairs-/ + +section matrix_adjoints +open_locale matrix + +variables [comm_ring R] +variables [add_comm_monoid M₁] [module R M₁] [add_comm_monoid M₂] [module R M₂] +variables [fintype n] [fintype n'] + +variables (b₁ : basis n R M₁) (b₂ : basis n' R M₂) +variables (J J₂ : matrix n n R) (J' : matrix n' n' R) +variables (A : matrix n' n R) (A' : matrix n n' R) +variables (A₁ : matrix n n R) + +/-- The condition for the matrices `A`, `A'` to be an adjoint pair with respect to the square +matrices `J`, `J₃`. -/ +def matrix.is_adjoint_pair := Aᵀ ⬝ J' = J ⬝ A' + +/-- The condition for a square matrix `A` to be self-adjoint with respect to the square matrix +`J`. -/ +def matrix.is_self_adjoint := matrix.is_adjoint_pair J J A₁ A₁ + +/-- The condition for a square matrix `A` to be skew-adjoint with respect to the square matrix +`J`. -/ +def matrix.is_skew_adjoint := matrix.is_adjoint_pair J J A₁ (-A₁) + +variables [decidable_eq n] [decidable_eq n'] + +@[simp] lemma is_adjoint_pair_to_linear_map₂' : + is_adjoint_pair (matrix.to_linear_map₂' J) (matrix.to_linear_map₂' J') + (matrix.to_lin' A) (matrix.to_lin' A') ↔ + matrix.is_adjoint_pair J J' A A' := +begin + rw is_adjoint_pair_iff_comp_eq_compl₂, + have h : ∀ (B B' : (n → R) →ₗ[R] (n' → R) →ₗ[R] R), B = B' ↔ + (linear_map.to_matrix₂' B) = (linear_map.to_matrix₂' B') := + begin + intros B B', + split; intros h, + { rw h }, + { exact linear_map.to_matrix₂'.injective h }, + end, + simp_rw [h, linear_map.to_matrix₂'_comp, linear_map.to_matrix₂'_compl₂, + linear_map.to_matrix'_to_lin', linear_map.to_matrix'_to_linear_map₂'], + refl, +end + +@[simp] lemma is_adjoint_pair_to_linear_map₂ : + is_adjoint_pair (matrix.to_linear_map₂ b₁ b₁ J) (matrix.to_linear_map₂ b₂ b₂ J') + (matrix.to_lin b₁ b₂ A) (matrix.to_lin b₂ b₁ A') ↔ + matrix.is_adjoint_pair J J' A A' := +begin + rw is_adjoint_pair_iff_comp_eq_compl₂, + have h : ∀ (B B' : M₁ →ₗ[R] M₂ →ₗ[R] R), B = B' ↔ + (linear_map.to_matrix₂ b₁ b₂ B) = (linear_map.to_matrix₂ b₁ b₂ B') := + begin + intros B B', + split; intros h, + { rw h }, + { exact (linear_map.to_matrix₂ b₁ b₂).injective h }, + end, + simp_rw [h, linear_map.to_matrix₂_comp b₂ b₂, linear_map.to_matrix₂_compl₂ b₁ b₁, + linear_map.to_matrix_to_lin, linear_map.to_matrix₂_to_linear_map₂], + refl, +end + +lemma matrix.is_adjoint_pair_equiv (P : matrix n n R) (h : is_unit P) : + (Pᵀ ⬝ J ⬝ P).is_adjoint_pair (Pᵀ ⬝ J ⬝ P) A₁ A₁ ↔ + J.is_adjoint_pair J (P ⬝ A₁ ⬝ P⁻¹) (P ⬝ A₁ ⬝ P⁻¹) := +have h' : is_unit P.det := P.is_unit_iff_is_unit_det.mp h, +begin + let u := P.nonsing_inv_unit h', + let v := Pᵀ.nonsing_inv_unit (P.is_unit_det_transpose h'), + let x := A₁ᵀ * Pᵀ * J, + let y := J * P * A₁, + suffices : x * ↑u = ↑v * y ↔ ↑v⁻¹ * x = y * ↑u⁻¹, + { dunfold matrix.is_adjoint_pair, + repeat { rw matrix.transpose_mul, }, + simp only [←matrix.mul_eq_mul, ←mul_assoc, P.transpose_nonsing_inv], + conv_lhs { to_rhs, rw [mul_assoc, mul_assoc], congr, skip, rw ←mul_assoc, }, + conv_rhs { rw [mul_assoc, mul_assoc], conv { to_lhs, congr, skip, rw ←mul_assoc }, }, + exact this, }, + rw units.eq_mul_inv_iff_mul_eq, conv_rhs { rw mul_assoc, }, rw v.inv_mul_eq_iff_eq_mul, +end + +/-- The submodule of pair-self-adjoint matrices with respect to bilinear forms corresponding to +given matrices `J`, `J₂`. -/ +def pair_self_adjoint_matrices_submodule : submodule R (matrix n n R) := +(is_pair_self_adjoint_submodule (matrix.to_linear_map₂' J) (matrix.to_linear_map₂' J₂)).map + ((linear_map.to_matrix' : ((n → R) →ₗ[R] (n → R)) ≃ₗ[R] matrix n n R) : + ((n → R) →ₗ[R] (n → R)) →ₗ[R] matrix n n R) + +@[simp] lemma mem_pair_self_adjoint_matrices_submodule : + A₁ ∈ (pair_self_adjoint_matrices_submodule J J₂) ↔ matrix.is_adjoint_pair J J₂ A₁ A₁ := +begin + simp only [pair_self_adjoint_matrices_submodule, linear_equiv.coe_coe, + linear_map.to_matrix'_apply, submodule.mem_map, mem_is_pair_self_adjoint_submodule], + split, + { rintros ⟨f, hf, hA⟩, + have hf' : f = A₁.to_lin' := by rw [←hA, matrix.to_lin'_to_matrix'], rw hf' at hf, + rw ← is_adjoint_pair_to_linear_map₂', + exact hf, }, + { intros h, refine ⟨A₁.to_lin', _, linear_map.to_matrix'_to_lin' _⟩, + exact (is_adjoint_pair_to_linear_map₂' _ _ _ _).mpr h, }, +end + +/-- The submodule of self-adjoint matrices with respect to the bilinear form corresponding to +the matrix `J`. -/ +def self_adjoint_matrices_submodule : submodule R (matrix n n R) := + pair_self_adjoint_matrices_submodule J J + +@[simp] lemma mem_self_adjoint_matrices_submodule : + A₁ ∈ self_adjoint_matrices_submodule J ↔ J.is_self_adjoint A₁ := +by { erw mem_pair_self_adjoint_matrices_submodule, refl, } + +/-- The submodule of skew-adjoint matrices with respect to the bilinear form corresponding to +the matrix `J`. -/ +def skew_adjoint_matrices_submodule : submodule R (matrix n n R) := + pair_self_adjoint_matrices_submodule (-J) J + +@[simp] lemma mem_skew_adjoint_matrices_submodule : + A₁ ∈ skew_adjoint_matrices_submodule J ↔ J.is_skew_adjoint A₁ := +begin + erw mem_pair_self_adjoint_matrices_submodule, + simp [matrix.is_skew_adjoint, matrix.is_adjoint_pair], +end + +end matrix_adjoints + +namespace linear_map + +/-! ### Nondegenerate bilinear forms-/ + +section det + +open matrix + +variables [comm_ring R₁] [add_comm_monoid M₁] [module R₁ M₁] +variables [decidable_eq ι] [fintype ι] + +lemma _root_.matrix.separating_left_to_linear_map₂'_iff_separating_left_to_linear_map₂ + {M : matrix ι ι R₁} (b : basis ι R₁ M₁) : M.to_linear_map₂'.separating_left ↔ + (matrix.to_linear_map₂ b b M).separating_left := +(separating_left_congr_iff b.equiv_fun.symm b.equiv_fun.symm).symm + +variables (B : M₁ →ₗ[R₁] M₁ →ₗ[R₁] R₁) + +-- Lemmas transferring nondegeneracy between a matrix and its associated bilinear form + +theorem _root_.matrix.nondegenerate.to_linear_map₂' {M : matrix ι ι R₁} (h : M.nondegenerate) : + M.to_linear_map₂'.separating_left := +λ x hx, h.eq_zero_of_ortho $ λ y, by simpa only [to_linear_map₂'_apply'] using hx y + +@[simp] lemma _root_.matrix.separating_left_to_linear_map₂'_iff {M : matrix ι ι R₁} : + M.to_linear_map₂'.separating_left ↔ M.nondegenerate := +⟨λ h v hv, h v $ λ w, (M.to_linear_map₂'_apply' _ _).trans $ hv w, + matrix.nondegenerate.to_linear_map₂'⟩ + +theorem _root_.matrix.nondegenerate.to_linear_map₂ {M : matrix ι ι R₁} (h : M.nondegenerate) + (b : basis ι R₁ M₁) : (to_linear_map₂ b b M).separating_left := +(matrix.separating_left_to_linear_map₂'_iff_separating_left_to_linear_map₂ b).mp h.to_linear_map₂' + +@[simp] lemma _root_.matrix.separating_left_to_linear_map₂_iff {M : matrix ι ι R₁} + (b : basis ι R₁ M₁) : (to_linear_map₂ b b M).separating_left ↔ M.nondegenerate := +by rw [←matrix.separating_left_to_linear_map₂'_iff_separating_left_to_linear_map₂, + matrix.separating_left_to_linear_map₂'_iff] + +-- Lemmas transferring nondegeneracy between a bilinear form and its associated matrix + +@[simp] theorem nondegenerate_to_matrix₂'_iff {B : (ι → R₁) →ₗ[R₁] (ι → R₁) →ₗ[R₁] R₁} : + B.to_matrix₂'.nondegenerate ↔ B.separating_left := +matrix.separating_left_to_linear_map₂'_iff.symm.trans $ + (matrix.to_linear_map₂'_to_matrix' B).symm ▸ iff.rfl + +theorem separating_left.to_matrix₂' {B : (ι → R₁) →ₗ[R₁] (ι → R₁) →ₗ[R₁] R₁} + (h : B.separating_left) : B.to_matrix₂'.nondegenerate := +nondegenerate_to_matrix₂'_iff.mpr h + +@[simp] theorem nondegenerate_to_matrix_iff {B : M₁ →ₗ[R₁] M₁ →ₗ[R₁] R₁} + (b : basis ι R₁ M₁) : (to_matrix₂ b b B).nondegenerate ↔ B.separating_left := +(matrix.separating_left_to_linear_map₂_iff b).symm.trans $ + (matrix.to_linear_map₂_to_matrix₂ b b B).symm ▸ iff.rfl + +theorem separating_left.to_matrix₂ {B : M₁ →ₗ[R₁] M₁ →ₗ[R₁] R₁} (h : B.separating_left) + (b : basis ι R₁ M₁) : (to_matrix₂ b b B).nondegenerate := +(nondegenerate_to_matrix_iff b).mpr h + +-- Some shorthands for combining the above with `matrix.nondegenerate_of_det_ne_zero` + +variables [is_domain R₁] + +lemma separating_left_to_linear_map₂'_iff_det_ne_zero {M : matrix ι ι R₁} : + M.to_linear_map₂'.separating_left ↔ M.det ≠ 0 := +by rw [matrix.separating_left_to_linear_map₂'_iff, matrix.nondegenerate_iff_det_ne_zero] + +theorem separating_left_to_linear_map₂'_of_det_ne_zero' (M : matrix ι ι R₁) (h : M.det ≠ 0) : + M.to_linear_map₂'.separating_left := +separating_left_to_linear_map₂'_iff_det_ne_zero.mpr h + +lemma separating_left_iff_det_ne_zero {B : M₁ →ₗ[R₁] M₁ →ₗ[R₁] R₁} + (b : basis ι R₁ M₁) : B.separating_left ↔ (to_matrix₂ b b B).det ≠ 0 := +by rw [←matrix.nondegenerate_iff_det_ne_zero, nondegenerate_to_matrix_iff] + +theorem separating_left_of_det_ne_zero {B : M₁ →ₗ[R₁] M₁ →ₗ[R₁] R₁} (b : basis ι R₁ M₁) + (h : (to_matrix₂ b b B).det ≠ 0) : + B.separating_left := +(separating_left_iff_det_ne_zero b).mpr h + +end det + +end linear_map diff --git a/src/linear_algebra/sesquilinear_form.lean b/src/linear_algebra/sesquilinear_form.lean index 90ed4b404fe95..b5132ac5ba9c3 100644 --- a/src/linear_algebra/sesquilinear_form.lean +++ b/src/linear_algebra/sesquilinear_form.lean @@ -37,7 +37,7 @@ Sesquilinear form, open_locale big_operators -variables {R R₁ R₂ R₃ M M₁ M₂ K K₁ K₂ V V₁ V₂ n: Type*} +variables {R R₁ R₂ R₃ M M₁ M₂ Mₗ₁ Mₗ₁' Mₗ₂ Mₗ₂' K K₁ K₂ V V₁ V₂ n : Type*} namespace linear_map @@ -549,6 +549,46 @@ for every nonzero `x` in `M₁`, there exists `y` in `M₂` with `B x y ≠ 0`.- def separating_left (B : M₁ →ₛₗ[I₁] M₂ →ₛₗ[I₂] R) : Prop := ∀ x : M₁, (∀ y : M₂, B x y = 0) → x = 0 +variables (M₁ M₂ I₁ I₂) + +/-- In a non-trivial module, zero is not non-degenerate. -/ +lemma not_separating_left_zero [nontrivial M₁] : ¬(0 : M₁ →ₛₗ[I₁] M₂ →ₛₗ[I₂] R).separating_left := +let ⟨m, hm⟩ := exists_ne (0 : M₁) in λ h, hm (h m $ λ n, rfl) + +variables {M₁ M₂ I₁ I₂} + +lemma separating_left.ne_zero [nontrivial M₁] {B : M₁ →ₛₗ[I₁] M₂ →ₛₗ[I₂] R} + (h : B.separating_left) : B ≠ 0 := +λ h0, not_separating_left_zero M₁ M₂ I₁ I₂ $ h0 ▸ h + +section linear + +variables [add_comm_monoid Mₗ₁] [add_comm_monoid Mₗ₂] [add_comm_monoid Mₗ₁'] [add_comm_monoid Mₗ₂'] +variables [module R Mₗ₁] [module R Mₗ₂] [module R Mₗ₁'] [module R Mₗ₂'] +variables {B : Mₗ₁ →ₗ[R] Mₗ₂ →ₗ[R] R} (e₁ : Mₗ₁ ≃ₗ[R] Mₗ₁') (e₂ : Mₗ₂ ≃ₗ[R] Mₗ₂') + +lemma separating_left.congr (h : B.separating_left) : + (e₁.arrow_congr (e₂.arrow_congr (linear_equiv.refl R R)) B).separating_left := +begin + intros x hx, + rw ←e₁.symm.map_eq_zero_iff, + refine h (e₁.symm x) (λ y, _), + specialize hx (e₂ y), + simp only [linear_equiv.arrow_congr_apply, linear_equiv.symm_apply_apply, + linear_equiv.map_eq_zero_iff] at hx, + exact hx, +end + +@[simp] lemma separating_left_congr_iff : + (e₁.arrow_congr (e₂.arrow_congr (linear_equiv.refl R R)) B).separating_left ↔ B.separating_left := +⟨λ h, begin + convert h.congr e₁.symm e₂.symm, + ext x y, + simp, +end, separating_left.congr e₁ e₂⟩ + +end linear + /-- A bilinear form is called right-separating if the only element that is right-orthogonal to every other element is `0`; i.e., for every nonzero `y` in `M₂`, there exists `x` in `M₁` with `B x y ≠ 0`.-/ diff --git a/src/linear_algebra/std_basis.lean b/src/linear_algebra/std_basis.lean index 90ee77f399777..7b332fc04061d 100644 --- a/src/linear_algebra/std_basis.lean +++ b/src/linear_algebra/std_basis.lean @@ -47,6 +47,13 @@ def std_basis : Π (i : ι), φ i →ₗ[R] (Πi, φ i) := single lemma std_basis_apply (i : ι) (b : φ i) : std_basis R φ i b = update 0 i b := rfl +@[simp] lemma std_basis_apply' (i i' : ι) : (std_basis R (λ (_x : ι), R) i) 1 i' = + ite (i = i') 1 0 := +begin + rw [linear_map.std_basis_apply, function.update_apply, pi.zero_apply], + congr' 1, rw [eq_iff_iff, eq_comm], +end + lemma coe_std_basis (i : ι) : ⇑(std_basis R φ i) = pi.single i := rfl From ec0b1144fc04ad74a06797f2746581947abb0649 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Mon, 3 Oct 2022 21:49:24 +0000 Subject: [PATCH 542/828] feat(measure_theory/group/prod.lean): use measure_preserving predicate (#16668) * Also use `quasi_measure_preserving` more in various files. In particular, `ae_[strongly_]measurable.comp_measurable'` is now `.comp_quasi_measure_preserving`. * Remove lemmas that are direct consequences of `[quasi_]measure_preserving` properties. * This simplifies various proofs. Co-authored-by: sgouezel --- src/analysis/convolution.lean | 8 +- src/measure_theory/constructions/prod.lean | 23 +-- .../function/strongly_measurable.lean | 8 +- src/measure_theory/group/measure.lean | 6 +- src/measure_theory/group/prod.lean | 136 ++++++++---------- src/measure_theory/measure/ae_measurable.lean | 6 +- src/measure_theory/measure/lebesgue.lean | 7 +- src/measure_theory/measure/measure_space.lean | 4 + 8 files changed, 96 insertions(+), 102 deletions(-) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 35d5e7879a533..64f55e61721a6 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -205,12 +205,14 @@ variables [sigma_finite μ] [is_add_left_invariant μ] lemma measure_theory.ae_strongly_measurable.convolution_integrand_snd (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (x : G) : ae_strongly_measurable (λ t, L (f t) (g (x - t))) μ := -hf.convolution_integrand_snd' L $ hg.mono' $ map_sub_left_absolutely_continuous μ x +hf.convolution_integrand_snd' L $ hg.mono' $ + (quasi_measure_preserving_sub_left μ x).absolutely_continuous lemma measure_theory.ae_strongly_measurable.convolution_integrand_swap_snd (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) (x : G) : ae_strongly_measurable (λ t, L (f (x - t)) (g t)) μ := -(hf.mono' (map_sub_left_absolutely_continuous μ x)).convolution_integrand_swap_snd' L hg +(hf.mono' (quasi_measure_preserving_sub_left μ x).absolutely_continuous) + .convolution_integrand_swap_snd' L hg end left @@ -309,7 +311,7 @@ lemma bdd_above.convolution_exists_at [sigma_finite μ] {x₀ : G} begin refine bdd_above.convolution_exists_at' L _ hs h2s hf hmf _, { simp_rw [← sub_eq_neg_add, hbg] }, - { exact hmg.mono' (map_sub_left_absolutely_continuous μ x₀) } + { exact hmg.mono' (quasi_measure_preserving_sub_left μ x₀).absolutely_continuous } end variables {L} [is_neg_invariant μ] diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod.lean index cff97ef13873a..a5cb2528f285e 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod.lean @@ -468,15 +468,15 @@ begin { simp_rw [Union_unpair_prod, hμ.spanning, hν.spanning, univ_prod_univ] } end -lemma prod_fst_absolutely_continuous : map prod.fst (μ.prod ν) ≪ μ := +lemma quasi_measure_preserving_fst : quasi_measure_preserving prod.fst (μ.prod ν) μ := begin - refine absolutely_continuous.mk (λ s hs h2s, _), + refine ⟨measurable_fst, absolutely_continuous.mk (λ s hs h2s, _)⟩, rw [map_apply measurable_fst hs, ← prod_univ, prod_prod, h2s, zero_mul], end -lemma prod_snd_absolutely_continuous : map prod.snd (μ.prod ν) ≪ ν := +lemma quasi_measure_preserving_snd : quasi_measure_preserving prod.snd (μ.prod ν) ν := begin - refine absolutely_continuous.mk (λ s hs h2s, _), + refine ⟨measurable_snd, absolutely_continuous.mk (λ s hs h2s, _)⟩, rw [map_apply measurable_snd hs, ← univ_prod, prod_prod, h2s, mul_zero] end @@ -515,6 +515,9 @@ begin simp_rw [map_apply measurable_swap (hs.prod ht), preimage_swap_prod, prod_prod, mul_comm] end +lemma measure_preserving_swap : measure_preserving prod.swap (μ.prod ν) (ν.prod μ) := +⟨measurable_swap, prod_swap⟩ + lemma prod_apply_symm {s : set (α × β)} (hs : measurable_set s) : μ.prod ν s = ∫⁻ y, μ ((λ x, (x, y)) ⁻¹' s) ∂ν := by { rw [← prod_swap, map_apply measurable_swap hs], @@ -697,21 +700,21 @@ by { rw ← prod_swap at hf, exact hf.comp_measurable measurable_swap } lemma ae_measurable.fst [sigma_finite ν] {f : α → γ} (hf : ae_measurable f μ) : ae_measurable (λ (z : α × β), f z.1) (μ.prod ν) := -hf.comp_measurable' measurable_fst prod_fst_absolutely_continuous +hf.comp_quasi_measure_preserving quasi_measure_preserving_fst lemma ae_measurable.snd [sigma_finite ν] {f : β → γ} (hf : ae_measurable f ν) : ae_measurable (λ (z : α × β), f z.2) (μ.prod ν) := -hf.comp_measurable' measurable_snd prod_snd_absolutely_continuous +hf.comp_quasi_measure_preserving quasi_measure_preserving_snd lemma measure_theory.ae_strongly_measurable.fst {γ} [topological_space γ] [sigma_finite ν] {f : α → γ} (hf : ae_strongly_measurable f μ) : ae_strongly_measurable (λ (z : α × β), f z.1) (μ.prod ν) := -hf.comp_measurable' measurable_fst prod_fst_absolutely_continuous +hf.comp_quasi_measure_preserving quasi_measure_preserving_fst lemma measure_theory.ae_strongly_measurable.snd {γ} [topological_space γ] [sigma_finite ν] {f : β → γ} (hf : ae_strongly_measurable f ν) : ae_strongly_measurable (λ (z : α × β), f z.2) (μ.prod ν) := -hf.comp_measurable' measurable_snd prod_snd_absolutely_continuous +hf.comp_quasi_measure_preserving quasi_measure_preserving_snd /-- The Bochner integral is a.e.-measurable. This shows that the integrand of (the right-hand-side of) Fubini's theorem is a.e.-measurable. -/ @@ -897,9 +900,7 @@ lemma integrable_prod_mul {f : α → ℝ} {g : β → ℝ} (hf : integrable f integrable (λ (z : α × β), f z.1 * g z.2) (μ.prod ν) := begin refine (integrable_prod_iff _).2 ⟨_, _⟩, - { apply ae_strongly_measurable.mul, - { exact (hf.1.mono' prod_fst_absolutely_continuous).comp_measurable measurable_fst }, - { exact (hg.1.mono' prod_snd_absolutely_continuous).comp_measurable measurable_snd } }, + { exact hf.1.fst.mul hg.1.snd }, { exact eventually_of_forall (λ x, hg.const_mul (f x)) }, { simpa only [norm_mul, integral_mul_left] using hf.norm.mul_const _ } end diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index 6011f1f1e1461..99974fc0ddb8a 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -1466,10 +1466,10 @@ lemma comp_measurable {γ : Type*} {mγ : measurable_space γ} {mα : measurable ae_strongly_measurable (g ∘ f) μ := hg.comp_ae_measurable hf.ae_measurable -lemma comp_measurable' {γ : Type*} {mγ : measurable_space γ} {mα : measurable_space α} {f : γ → α} - {μ : measure γ} {ν : measure α} (hg : ae_strongly_measurable g ν) (hf : measurable f) - (h : μ.map f ≪ ν) : ae_strongly_measurable (g ∘ f) μ := -(hg.mono' h).comp_measurable hf +lemma comp_quasi_measure_preserving {γ : Type*} {mγ : measurable_space γ} {mα : measurable_space α} + {f : γ → α} {μ : measure γ} {ν : measure α} (hg : ae_strongly_measurable g ν) + (hf : quasi_measure_preserving f μ ν) : ae_strongly_measurable (g ∘ f) μ := +(hg.mono' hf.absolutely_continuous).comp_measurable hf.measurable lemma is_separable_ae_range (hf : ae_strongly_measurable f μ) : ∃ (t : set β), is_separable t ∧ ∀ᵐ x ∂μ, f x ∈ t := diff --git a/src/measure_theory/group/measure.lean b/src/measure_theory/group/measure.lean index 2ade894b17f61..0c573bc97d2ea 100644 --- a/src/measure_theory/group/measure.lean +++ b/src/measure_theory/group/measure.lean @@ -159,9 +159,13 @@ lemma map_div_right_eq_self (μ : measure G) [is_mul_right_invariant μ] (g : G) map (/ g) μ = μ := by simp_rw [div_eq_mul_inv, map_mul_right_eq_self μ g⁻¹] - variables [has_measurable_mul G] +@[to_additive] +lemma measure_preserving_div_right (μ : measure G) [is_mul_right_invariant μ] + (g : G) : measure_preserving (/ g) μ μ := +by simp_rw [div_eq_mul_inv, measure_preserving_mul_right μ g⁻¹] + /-- We shorten this from `measure_preimage_mul_left`, since left invariant is the preferred option for measures in this formalization. -/ @[simp, to_additive "We shorten this from `measure_preimage_add_left`, since left invariant is the diff --git a/src/measure_theory/group/prod.lean b/src/measure_theory/group/prod.lean index c10d86f822fde..960ee25a60a95 100644 --- a/src/measure_theory/group/prod.lean +++ b/src/measure_theory/group/prod.lean @@ -54,25 +54,24 @@ namespace measure_theory open measure -/-- A shear mapping preserves the measure `μ.prod ν`. +/-- The multiplicative shear mapping `(x, y) ↦ (x, xy)` preserves the measure `μ × ν`. This condition is part of the definition of a measurable group in [Halmos, §59]. There, the map in this lemma is called `S`. -/ -@[to_additive map_prod_add_eq /-" An additive shear mapping preserves the measure `μ.prod ν`. "-/] -lemma map_prod_mul_eq [is_mul_left_invariant ν] : - map (λ z : G × G, (z.1, z.1 * z.2)) (μ.prod ν) = μ.prod ν := -((measure_preserving.id μ).skew_product measurable_mul - (filter.eventually_of_forall $ map_mul_left_eq_self ν)).map_eq - -/-- The function we are mapping along is `SR` in [Halmos, §59], - where `S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ -@[to_additive map_prod_add_eq_swap /-" "-/] -lemma map_prod_mul_eq_swap [is_mul_left_invariant μ] : - map (λ z : G × G, (z.2, z.2 * z.1)) (μ.prod ν) = ν.prod μ := -begin - rw [← prod_swap], - simp_rw [map_map (measurable_snd.prod_mk (measurable_snd.mul measurable_fst)) measurable_swap], - exact map_prod_mul_eq ν μ -end +@[to_additive measure_preserving_prod_add + /-" The shear mapping `(x, y) ↦ (x, x + y)` preserves the measure `μ.prod ν`. "-/] +lemma measure_preserving_prod_mul [is_mul_left_invariant ν] : + measure_preserving (λ z : G × G, (z.1, z.1 * z.2)) (μ.prod ν) (μ.prod ν) := +(measure_preserving.id μ).skew_product measurable_mul $ + filter.eventually_of_forall $ map_mul_left_eq_self ν + +/-- The map `(x, y) ↦ (y, yx)` sends the measure `μ × ν` to `ν × μ`. +This is the map `SR` in [Halmos, §59]. +`S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ +@[to_additive measure_preserving_prod_add_swap + /-" The map `(x, y) ↦ (y, y + x)` sends the measure `μ × ν` to `ν × μ`. "-/] +lemma measure_preserving_prod_mul_swap [is_mul_left_invariant μ] : + measure_preserving (λ z : G × G, (z.2, z.2 * z.1)) (μ.prod ν) (ν.prod μ) := +(measure_preserving_prod_mul ν μ).comp measure_preserving_swap @[to_additive] lemma measurable_measure_mul_right (hE : measurable_set E) : @@ -82,56 +81,51 @@ begin μ ((λ x, (x, y)) ⁻¹' ((λ z : G × G, ((1 : G), z.1 * z.2)) ⁻¹' (univ ×ˢ E)))), { convert this, ext1 x, congr' 1 with y : 1, simp }, apply measurable_measure_prod_mk_right, - exact measurable_const.prod_mk (measurable_fst.mul measurable_snd) (measurable_set.univ.prod hE) + exact measurable_const.prod_mk measurable_mul (measurable_set.univ.prod hE) end variables [has_measurable_inv G] -/-- The function we are mapping along is `S⁻¹` in [Halmos, §59], - where `S` is the map in `map_prod_mul_eq`. -/ -@[to_additive map_prod_neg_add_eq "The function we are mapping along is `-S` in [Halmos, §59], where -`S` is the map in `map_prod_add_eq`."] -lemma map_prod_inv_mul_eq [is_mul_left_invariant ν] : - map (λ z : G × G, (z.1, z.1⁻¹ * z.2)) (μ.prod ν) = μ.prod ν := -(measurable_equiv.shear_mul_right G).map_apply_eq_iff_map_symm_apply_eq.mp $ map_prod_mul_eq μ ν +/-- The map `(x, y) ↦ (x, x⁻¹y)` is measure-preserving. +This is the function `S⁻¹` in [Halmos, §59], +where `S` is the map in `measure_preserving_prod_mul`. -/ +@[to_additive measure_preserving_prod_neg_add + "The map `(x, y) ↦ (x, - x + y)` is measure-preserving."] +lemma measure_preserving_prod_inv_mul [is_mul_left_invariant ν] : + measure_preserving (λ z : G × G, (z.1, z.1⁻¹ * z.2)) (μ.prod ν) (μ.prod ν) := +(measure_preserving_prod_mul μ ν).symm $ measurable_equiv.shear_mul_right G @[to_additive] lemma quasi_measure_preserving_div [is_mul_right_invariant μ] : quasi_measure_preserving (λ (p : G × G), p.1 / p.2) (μ.prod μ) μ := begin - refine quasi_measure_preserving.prod_of_left measurable_div _, - simp_rw [div_eq_mul_inv], - apply eventually_of_forall, - refine λ y, ⟨measurable_mul_const y⁻¹, (map_mul_right_eq_self μ y⁻¹).absolutely_continuous⟩ + refine quasi_measure_preserving.prod_of_left measurable_div (eventually_of_forall $ λ y, _), + exact (measure_preserving_div_right μ y).quasi_measure_preserving end variables [is_mul_left_invariant μ] -/-- The function we are mapping along is `S⁻¹R` in [Halmos, §59], - where `S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ -@[to_additive map_prod_neg_add_eq_swap "The function we are mapping along is `-S + R` in -[Halmos, §59], where `S` is the map in `map_prod_add_eq` and `R` is `prod.swap`."] -lemma map_prod_inv_mul_eq_swap : map (λ z : G × G, (z.2, z.2⁻¹ * z.1)) (μ.prod ν) = ν.prod μ := +/-- The map `(x, y) ↦ (y, y⁻¹x)` sends `μ × ν` to `ν × μ`. +This is the function `S⁻¹R` in [Halmos, §59], +where `S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ +@[to_additive measure_preserving_prod_neg_add_swap + "The map `(x, y) ↦ (y, - y + x)` sends `μ × ν` to `ν × μ`."] +lemma measure_preserving_prod_inv_mul_swap : + measure_preserving (λ z : G × G, (z.2, z.2⁻¹ * z.1)) (μ.prod ν) (ν.prod μ) := +(measure_preserving_prod_inv_mul ν μ).comp measure_preserving_swap + +/-- The map `(x, y) ↦ (yx, x⁻¹)` is measure-preserving. +This is the function `S⁻¹RSR` in [Halmos, §59], +where `S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ +@[to_additive measure_preserving_add_prod_neg + "The map `(x, y) ↦ (y + x, - x)` is measure-preserving."] +lemma measure_preserving_mul_prod_inv [is_mul_left_invariant ν] : + measure_preserving (λ z : G × G, (z.2 * z.1, z.1⁻¹)) (μ.prod ν) (μ.prod ν) := begin - rw [← prod_swap], - simp_rw - [map_map (measurable_snd.prod_mk $ measurable_snd.inv.mul measurable_fst) measurable_swap], - exact map_prod_inv_mul_eq ν μ -end - -/-- The function we are mapping along is `S⁻¹RSR` in [Halmos, §59], - where `S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ -@[to_additive map_prod_add_neg_eq "The function we are mapping along is `-S + R + S + R ` in -[Halmos, §59], where `S` is the map in `map_prod_add_eq` and `R` is `prod.swap`."] -lemma map_prod_mul_inv_eq [is_mul_left_invariant ν] : - map (λ z : G × G, (z.2 * z.1, z.1⁻¹)) (μ.prod ν) = μ.prod ν := -begin - suffices : map ((λ z : G × G, (z.2, z.2⁻¹ * z.1)) ∘ (λ z : G × G, (z.2, z.2 * z.1))) (μ.prod ν) = - μ.prod ν, - { convert this, ext1 ⟨x, y⟩, simp }, - simp_rw [← map_map (measurable_snd.prod_mk (measurable_snd.inv.mul measurable_fst)) - (measurable_snd.prod_mk (measurable_snd.mul measurable_fst)), map_prod_mul_eq_swap μ ν, - map_prod_inv_mul_eq_swap ν μ] + convert (measure_preserving_prod_inv_mul_swap ν μ).comp + (measure_preserving_prod_mul_swap μ ν), + ext1 ⟨x, y⟩, + simp_rw [function.comp_apply, mul_inv_rev, inv_mul_cancel_right] end @[to_additive] lemma quasi_measure_preserving_inv : @@ -142,16 +136,13 @@ begin have hf : measurable (λ z : G × G, (z.2 * z.1, z.1⁻¹)) := (measurable_snd.mul measurable_fst).prod_mk measurable_fst.inv, suffices : map (λ z : G × G, (z.2 * z.1, z.1⁻¹)) (μ.prod μ) (s⁻¹ ×ˢ s⁻¹) = 0, - { simpa only [map_prod_mul_inv_eq μ μ, prod_prod, mul_eq_zero, or_self] using this }, + { simpa only [(measure_preserving_mul_prod_inv μ μ).map_eq, + prod_prod, mul_eq_zero, or_self] using this }, have hsm' : measurable_set (s⁻¹ ×ˢ s⁻¹) := hsm.inv.prod hsm.inv, simp_rw [map_apply hf hsm', prod_apply_symm (hf hsm'), preimage_preimage, mk_preimage_prod, inv_preimage, inv_inv, measure_mono_null (inter_subset_right _ _) hμs, lintegral_zero] end -@[to_additive] -lemma map_inv_absolutely_continuous : map has_inv.inv μ ≪ μ := -(quasi_measure_preserving_inv μ).absolutely_continuous - @[to_additive] lemma measure_inv_null : μ E⁻¹ = 0 ↔ μ E = 0 := begin @@ -174,12 +165,13 @@ lemma lintegral_lintegral_mul_inv [is_mul_left_invariant ν] begin have h : measurable (λ z : G × G, (z.2 * z.1, z.1⁻¹)) := (measurable_snd.mul measurable_fst).prod_mk measurable_fst.inv, - have h2f : ae_measurable (uncurry $ λ x y, f (y * x) x⁻¹) (μ.prod ν), - { apply hf.comp_measurable' h (map_prod_mul_inv_eq μ ν).absolutely_continuous }, + have h2f : ae_measurable (uncurry $ λ x y, f (y * x) x⁻¹) (μ.prod ν) := + hf.comp_quasi_measure_preserving (measure_preserving_mul_prod_inv μ ν).quasi_measure_preserving, simp_rw [lintegral_lintegral h2f, lintegral_lintegral hf], - conv_rhs { rw [← map_prod_mul_inv_eq μ ν] }, + conv_rhs { rw [← (measure_preserving_mul_prod_inv μ ν).map_eq] }, symmetry, - exact lintegral_map' (hf.mono' (map_prod_mul_inv_eq μ ν).absolutely_continuous) h.ae_measurable, + exact lintegral_map' (hf.mono' (measure_preserving_mul_prod_inv μ ν).map_eq.absolutely_continuous) + h.ae_measurable, end @[to_additive] @@ -194,17 +186,17 @@ lemma measure_mul_right_ne_zero (h2E : μ E ≠ 0) (y : G) : μ ((λ x, x * y) ⁻¹' E) ≠ 0 := (not_iff_not_of_iff (measure_mul_right_null μ y)).mpr h2E -@[to_additive] lemma quasi_measure_preserving_mul_right (g : G) : +/-- A *left*-invariant measure is quasi-preserved by *right*-multiplication. +This should not be confused with `(measure_preserving_mul_right μ g).quasi_measure_preserving`. -/ +@[to_additive /-"A *left*-invariant measure is quasi-preserved by *right*-addition. +This should not be confused with `(measure_preserving_add_right μ g).quasi_measure_preserving`. "-/] +lemma quasi_measure_preserving_mul_right (g : G) : quasi_measure_preserving (λ h : G, h * g) μ μ := begin refine ⟨measurable_mul_const g, absolutely_continuous.mk $ λ s hs, _⟩, rw [map_apply (measurable_mul_const g) hs, measure_mul_right_null], exact id, end -@[to_additive] -lemma map_mul_right_absolutely_continuous (g : G) : map (* g) μ ≪ μ := -(quasi_measure_preserving_mul_right μ g).absolutely_continuous - @[to_additive] lemma absolutely_continuous_map_mul_right (g : G) : μ ≪ map (* g) μ := begin @@ -215,17 +207,11 @@ end @[to_additive] lemma quasi_measure_preserving_div_left (g : G) : quasi_measure_preserving (λ h : G, g / h) μ μ := begin - refine ⟨measurable_const.div measurable_id, _⟩, simp_rw [div_eq_mul_inv], - rw [← map_map (measurable_const_mul g) measurable_inv], - refine ((map_inv_absolutely_continuous μ).map $ measurable_const_mul g).trans _, - rw [map_mul_left_eq_self], + exact (measure_preserving_mul_left μ g).quasi_measure_preserving.comp + (quasi_measure_preserving_inv μ) end -@[to_additive] -lemma map_div_left_absolutely_continuous (g : G) : map (λ h, g / h) μ ≪ μ := -(quasi_measure_preserving_div_left μ g).absolutely_continuous - @[to_additive] lemma absolutely_continuous_map_div_left (g : G) : μ ≪ map (λ h, g / h) μ := begin diff --git a/src/measure_theory/measure/ae_measurable.lean b/src/measure_theory/measure/ae_measurable.lean index 582015d62901a..5a97a8502d46f 100644 --- a/src/measure_theory/measure/ae_measurable.lean +++ b/src/measure_theory/measure/ae_measurable.lean @@ -129,9 +129,9 @@ lemma comp_measurable {f : α → δ} {g : δ → β} (hg : ae_measurable g (μ.map f)) (hf : measurable f) : ae_measurable (g ∘ f) μ := hg.comp_ae_measurable hf.ae_measurable -lemma comp_measurable' {ν : measure δ} {f : α → δ} {g : δ → β} (hg : ae_measurable g ν) - (hf : measurable f) (h : μ.map f ≪ ν) : ae_measurable (g ∘ f) μ := -(hg.mono' h).comp_measurable hf +lemma comp_quasi_measure_preserving {ν : measure δ} {f : α → δ} {g : δ → β} (hg : ae_measurable g ν) + (hf : quasi_measure_preserving f μ ν) : ae_measurable (g ∘ f) μ := +(hg.mono' hf.absolutely_continuous).comp_measurable hf.measurable lemma map_map_of_ae_measurable {g : β → γ} {f : α → β} (hg : ae_measurable g (measure.map f μ)) (hf : ae_measurable f μ) : diff --git a/src/measure_theory/measure/lebesgue.lean b/src/measure_theory/measure/lebesgue.lean index 9c643bbbf012a..71aca319c9d00 100644 --- a/src/measure_theory/measure/lebesgue.lean +++ b/src/measure_theory/measure/lebesgue.lean @@ -488,11 +488,8 @@ begin (μ.restrict s).prod volume (region_between (ae_measurable.mk f hf) (ae_measurable.mk g hg) s), { apply measure_congr, apply eventually_eq.rfl.inter, - exact - ((ae_eq_comp' measurable_fst.ae_measurable - hf.ae_eq_mk measure.prod_fst_absolutely_continuous).comp₂ _ eventually_eq.rfl).inter - (eventually_eq.rfl.comp₂ _ (ae_eq_comp' measurable_fst.ae_measurable - hg.ae_eq_mk measure.prod_fst_absolutely_continuous)) }, + exact ((quasi_measure_preserving_fst.ae_eq_comp hf.ae_eq_mk).comp₂ _ eventually_eq.rfl).inter + (eventually_eq.rfl.comp₂ _ $ quasi_measure_preserving_fst.ae_eq_comp hg.ae_eq_mk) }, rw [lintegral_congr_ae h₁, ← volume_region_between_eq_lintegral' hf.measurable_mk hg.measurable_mk hs], convert h₂ using 1, diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index ec1182d915da4..8ebb80a41982f 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -2109,6 +2109,10 @@ lemma ae_eq_comp' {ν : measure β} {f : α → β} {g g' : β → δ} (hf : ae_ (h : g =ᵐ[ν] g') (h2 : μ.map f ≪ ν) : g ∘ f =ᵐ[μ] g' ∘ f := (tendsto_ae_map hf).mono_right h2.ae_le h +lemma measure.quasi_measure_preserving.ae_eq_comp {ν : measure β} {f : α → β} {g g' : β → δ} + (hf : quasi_measure_preserving f μ ν) (h : g =ᵐ[ν] g') : g ∘ f =ᵐ[μ] g' ∘ f := +ae_eq_comp' hf.ae_measurable h hf.absolutely_continuous + lemma ae_eq_comp {f : α → β} {g g' : β → δ} (hf : ae_measurable f μ) (h : g =ᵐ[μ.map f] g') : g ∘ f =ᵐ[μ] g' ∘ f := ae_eq_comp' hf h absolutely_continuous.rfl From a1b96f80897671e6a73281ddb918a52f69bb48f1 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Tue, 4 Oct 2022 00:58:17 +0000 Subject: [PATCH 543/828] feat(group_theory/index): Index of finite intersection (#16393) This PR proves that the index of a finite intersection is at most the product of the indices. --- src/data/finite/card.lean | 41 +++++++++++++++++++++++++++++ src/group_theory/index.lean | 36 ++++++++++++++++++++----- src/set_theory/cardinal/finite.lean | 3 +++ 3 files changed, 73 insertions(+), 7 deletions(-) diff --git a/src/data/finite/card.lean b/src/data/finite/card.lean index 4adac16398fdd..20d2f09883f4c 100644 --- a/src/data/finite/card.lean +++ b/src/data/finite/card.lean @@ -92,6 +92,47 @@ by { haveI := fintype.of_finite α, haveI := fintype.of_surjective f hf, lemma card_eq_zero_iff [finite α] : nat.card α = 0 ↔ is_empty α := by { haveI := fintype.of_finite α, simp [fintype.card_eq_zero_iff] } +/-- If `f` is injective, then `nat.card α ≤ nat.card β`. We must also assume + `nat.card β = 0 → nat.card α = 0` since `nat.card` is defined to be `0` for infinite types. -/ +lemma card_le_of_injective' {f : α → β} (hf : function.injective f) + (h : nat.card β = 0 → nat.card α = 0) : nat.card α ≤ nat.card β := +(or_not_of_imp h).cases_on (λ h, le_of_eq_of_le h zero_le') + (λ h, @card_le_of_injective α β (nat.finite_of_card_ne_zero h) f hf) + +/-- If `f` is an embedding, then `nat.card α ≤ nat.card β`. We must also assume + `nat.card β = 0 → nat.card α = 0` since `nat.card` is defined to be `0` for infinite types. -/ +lemma card_le_of_embedding' (f : α ↪ β) (h : nat.card β = 0 → nat.card α = 0) : + nat.card α ≤ nat.card β := +card_le_of_injective' f.2 h + +/-- If `f` is surjective, then `nat.card β ≤ nat.card α`. We must also assume + `nat.card α = 0 → nat.card β = 0` since `nat.card` is defined to be `0` for infinite types. -/ +lemma card_le_of_surjective' {f : α → β} (hf : function.surjective f) + (h : nat.card α = 0 → nat.card β = 0) : nat.card β ≤ nat.card α := +(or_not_of_imp h).cases_on (λ h, le_of_eq_of_le h zero_le') + (λ h, @card_le_of_surjective α β (nat.finite_of_card_ne_zero h) f hf) + +/-- NB: `nat.card` is defined to be `0` for infinite types. -/ +lemma card_eq_zero_of_surjective {f : α → β} (hf : function.surjective f) (h : nat.card β = 0) : + nat.card α = 0 := +begin + casesI finite_or_infinite β, + { haveI := card_eq_zero_iff.mp h, + haveI := function.is_empty f, + exact nat.card_of_is_empty }, + { haveI := infinite.of_surjective f hf, + exact nat.card_eq_zero_of_infinite }, +end + +/-- NB: `nat.card` is defined to be `0` for infinite types. -/ +lemma card_eq_zero_of_injective [nonempty α] {f : α → β} (hf : function.injective f) + (h : nat.card α = 0) : nat.card β = 0 := +card_eq_zero_of_surjective (function.inv_fun_surjective hf) h + +/-- NB: `nat.card` is defined to be `0` for infinite types. -/ +lemma card_eq_zero_of_embedding [nonempty α] (f : α ↪ β) (h : nat.card α = 0) : nat.card β = 0 := +card_eq_zero_of_injective f.2 h + lemma card_sum [finite α] [finite β] : nat.card (α ⊕ β) = nat.card α + nat.card β := by { haveI := fintype.of_finite α, haveI := fintype.of_finite β, simp } diff --git a/src/group_theory/index.lean b/src/group_theory/index.lean index 52bcb63b48b14..76d67442cef44 100644 --- a/src/group_theory/index.lean +++ b/src/group_theory/index.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ +import data.finite.card import group_theory.quotient_group -import set_theory.cardinal.finite /-! # Index of a Subgroup @@ -33,7 +33,7 @@ Several theorems proved in this file are known as Lagrange's theorem. namespace subgroup -open_locale cardinal +open_locale big_operators cardinal variables {G : Type*} [group G] (H K L : subgroup G) @@ -228,9 +228,7 @@ eq_zero_of_zero_dvd (hKL ▸ (relindex_dvd_of_le_left L hHK)) @[to_additive] lemma relindex_eq_zero_of_le_right (hKL : K ≤ L) (hHK : H.relindex K = 0) : H.relindex L = 0 := -cardinal.to_nat_apply_of_aleph_0_le (le_trans (le_of_not_lt (λ h, cardinal.mk_ne_zero _ - ((cardinal.cast_to_nat_of_lt_aleph_0 h).symm.trans (cardinal.nat_cast_inj.mpr hHK)))) - (quotient_subgroup_of_embedding_of_le H hKL).cardinal_le) +finite.card_eq_zero_of_embedding (quotient_subgroup_of_embedding_of_le H hKL) hHK @[to_additive] lemma relindex_le_of_le_left (hHK : H ≤ K) (hHL : H.relindex L ≠ 0) : K.relindex L ≤ H.relindex L := @@ -238,8 +236,7 @@ nat.le_of_dvd (nat.pos_of_ne_zero hHL) (relindex_dvd_of_le_left L hHK) @[to_additive] lemma relindex_le_of_le_right (hKL : K ≤ L) (hHL : H.relindex L ≠ 0) : H.relindex K ≤ H.relindex L := -cardinal.to_nat_le_of_le_of_lt_aleph_0 (lt_of_not_ge (mt cardinal.to_nat_apply_of_aleph_0_le hHL)) - (cardinal.mk_le_of_injective (quotient_subgroup_of_embedding_of_le H hKL).2) +finite.card_le_of_embedding' (quotient_subgroup_of_embedding_of_le H hKL) (λ h, (hHL h).elim) @[to_additive] lemma relindex_ne_zero_trans (hHK : H.relindex K ≠ 0) (hKL : K.relindex L ≠ 0) : H.relindex L ≠ 0 := @@ -273,6 +270,31 @@ end @[to_additive] lemma index_inf_le : (H ⊓ K).index ≤ H.index * K.index := by simp_rw [←relindex_top_right, relindex_inf_le] +@[to_additive] lemma relindex_infi_ne_zero {ι : Type*} [hι : finite ι] {f : ι → subgroup G} + (hf : ∀ i, (f i).relindex L ≠ 0) : (⨅ i, f i).relindex L ≠ 0 := +begin + haveI := fintype.of_finite ι, + exact finset.prod_ne_zero_iff.mpr (λ i hi, hf i) ∘ nat.card_pi.symm.trans ∘ + finite.card_eq_zero_of_embedding (quotient_infi_embedding f L), +end + +@[to_additive] lemma relindex_infi_le {ι : Type*} [fintype ι] (f : ι → subgroup G) : + (⨅ i, f i).relindex L ≤ ∏ i, (f i).relindex L := +le_of_le_of_eq (finite.card_le_of_embedding' (quotient_infi_embedding f L) + (λ h, let ⟨i, hi, h⟩ := finset.prod_eq_zero_iff.mp (nat.card_pi.symm.trans h) in + relindex_eq_zero_of_le_left (infi_le f i) h)) nat.card_pi + +@[to_additive] lemma index_infi_ne_zero {ι : Type*} [finite ι] {f : ι → subgroup G} + (hf : ∀ i, (f i).index ≠ 0) : (⨅ i, f i).index ≠ 0 := +begin + simp_rw ← relindex_top_right at hf ⊢, + exact relindex_infi_ne_zero hf, +end + +@[to_additive] lemma index_infi_le {ι : Type*} [fintype ι] (f : ι → subgroup G) : + (⨅ i, f i).index ≤ ∏ i, (f i).index := +by simp_rw [←relindex_top_right, relindex_infi_le] + @[simp, to_additive index_eq_one] lemma index_eq_one : H.index = 1 ↔ H = ⊤ := ⟨λ h, quotient_group.subgroup_eq_top_of_subsingleton H (cardinal.to_nat_eq_one_iff_unique.mp h).1, λ h, (congr_arg index h).trans index_top⟩ diff --git a/src/set_theory/cardinal/finite.lean b/src/set_theory/cardinal/finite.lean index 7e1708e1a028e..253489173c994 100644 --- a/src/set_theory/cardinal/finite.lean +++ b/src/set_theory/cardinal/finite.lean @@ -35,6 +35,9 @@ lemma card_eq_fintype_card [fintype α] : nat.card α = fintype.card α := mk_to @[simp] lemma card_eq_zero_of_infinite [infinite α] : nat.card α = 0 := mk_to_nat_of_infinite +lemma finite_of_card_ne_zero (h : nat.card α ≠ 0) : finite α := +not_infinite_iff_finite.mp $ h ∘ @nat.card_eq_zero_of_infinite α + lemma card_congr (f : α ≃ β) : nat.card α = nat.card β := cardinal.to_nat_congr f From f3dd4ac6739eec04fba78657f23e517383f10bef Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Tue, 4 Oct 2022 00:58:18 +0000 Subject: [PATCH 544/828] feat(group_theory/sylow): Injectivity of `sylow.subtype` (#16717) `sylow.subtype` is injective if both Sylow subgroups are contained in the subgroup that you are restricting to. Co-authored-by: Thomas Browning --- src/group_theory/sylow.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index f88ccdd975903..f0d69dddb263a 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -106,6 +106,13 @@ P.comap_of_injective N.subtype subtype.coe_injective (by simp [h]) @[simp] lemma coe_subtype (h : ↑P ≤ N) : ↑(P.subtype h) = subgroup.comap N.subtype ↑P := rfl +lemma subtype_injective {P Q : sylow p G} {hP : ↑P ≤ N} {hQ : ↑Q ≤ N} + (h : P.subtype hP = Q.subtype hQ) : P = Q := +begin + rw set_like.ext_iff at h ⊢, + exact λ g, ⟨λ hg, (h ⟨g, hP hg⟩).mp hg, λ hg, (h ⟨g, hQ hg⟩).mpr hg⟩, +end + end sylow /-- A generalization of **Sylow's first theorem**. From 41a7254641a9379fb7bd1680f5c3db9537a52399 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 4 Oct 2022 03:15:12 +0000 Subject: [PATCH 545/828] feat(algebra/free_monoid): action of the `free_monoid` (#16746) --- src/algebra/free_monoid.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/src/algebra/free_monoid.lean b/src/algebra/free_monoid.lean index cbe4d5d79e439..53c3e796994ca 100644 --- a/src/algebra/free_monoid.lean +++ b/src/algebra/free_monoid.lean @@ -107,6 +107,17 @@ by { ext, simp } lemma hom_map_lift (g : M →* N) (f : α → M) (x : free_monoid α) : g (lift f x) = lift (g ∘ f) x := monoid_hom.ext_iff.1 (comp_lift g f) x +/-- Define a multiplicative action of `free_monoid α` on `β`. -/ +@[to_additive "Define an additive action of `free_add_monoid α` on `β`."] +def mk_mul_action (f : α → β → β) : mul_action (free_monoid α) β := +{ smul := λ l b, list.foldr f b l, + one_smul := λ x, rfl, + mul_smul := λ xs ys b, list.foldr_append _ _ _ _ } + +@[simp, to_additive] lemma of_smul (f : α → β → β) (x : α) (y : β) : + (by haveI := mk_mul_action f; exact of x • y) = f x y := +rfl + /-- The unique monoid homomorphism `free_monoid α →* free_monoid β` that sends each `of x` to `of (f x)`. -/ @[to_additive "The unique additive monoid homomorphism `free_add_monoid α →+ free_add_monoid β` @@ -127,4 +138,7 @@ hom_eq $ λ x, rfl lemma map_comp (g : β → γ) (f : α → β) : map (g ∘ f) = (map g).comp (map f) := hom_eq $ λ x, rfl +@[simp, to_additive] lemma map_id : map (@id α) = monoid_hom.id (free_monoid α) := +hom_eq $ λ x, rfl + end free_monoid From b5ce4aa22093fc4a15757007f0a739b49cff83b8 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 4 Oct 2022 05:52:14 +0000 Subject: [PATCH 546/828] feat(analysis/inner_product_space/pi_L2): extend an indexed orthonormal set to an orthonormal basis (#16478) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Given a finite-dimensional inner product space `E`, a fintype `ι` whose cardinality is the dimension of `E`, and a function `v : ι → E` whose restriction to some set `s` in `ι` is orthonormal, modify `v` to give an orthonormal basis for `E` indexed by `ι` which agrees on `s` with `v`. We already have this lemma for sets (`orthonormal.exists_orthonormal_basis_extension`); this is just some messing round to get the indexed version. Also rename `orthonormal.coe_range` to `orthonormal.to_subtype_range`, matching the naming convention for `linear_independent.to_subtype_range`. Formalized as part of the Sphere Eversion project. --- src/analysis/inner_product_space/basic.lean | 15 +++++++++++-- src/analysis/inner_product_space/pi_L2.lean | 24 ++++++++++++++++++++- 2 files changed, 36 insertions(+), 3 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 6311244e8f4a2..981c841163262 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -841,11 +841,22 @@ begin simp [hf.eq_iff] end +/-- An injective family `v : ι → E` is orthonormal if and only if `coe : (range v) → E` is +orthonormal. -/ +lemma orthonormal_subtype_range {v : ι → E} (hv : function.injective v) : + orthonormal 𝕜 (coe : set.range v → E) ↔ orthonormal 𝕜 v := +begin + let f : ι ≃ set.range v := equiv.of_injective v hv, + refine ⟨λ h, h.comp f f.injective, λ h, _⟩, + rw ← equiv.self_comp_of_injective_symm hv, + exact h.comp f.symm f.symm.injective, +end + /-- If `v : ι → E` is an orthonormal family, then `coe : (range v) → E` is an orthonormal family. -/ -lemma orthonormal.coe_range {v : ι → E} (hv : orthonormal 𝕜 v) : +lemma orthonormal.to_subtype_range {v : ι → E} (hv : orthonormal 𝕜 v) : orthonormal 𝕜 (coe : set.range v → E) := -by simpa using hv.comp _ (set.range_splitting_injective v) +(orthonormal_subtype_range hv.linear_independent.injective).2 hv /-- A linear combination of some subset of an orthonormal set is orthogonal to other members of the set. -/ diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 2ddd494dd1e6c..004096a224479 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -51,7 +51,7 @@ For consequences in infinite dimension (Hilbert bases, etc.), see the file -/ -open real set filter is_R_or_C submodule +open real set filter is_R_or_C submodule function open_locale big_operators uniformity topological_space nnreal ennreal complex_conjugate direct_sum noncomputable theory @@ -624,6 +624,28 @@ begin { simp }, end +lemma _root_.orthonormal.exists_orthonormal_basis_extension_of_card_eq + {ι : Type*} [fintype ι] (card_ι : finrank 𝕜 E = fintype.card ι) {v : ι → E} {s : set ι} + (hv : orthonormal 𝕜 (s.restrict v)) : + ∃ b : orthonormal_basis ι 𝕜 E, ∀ i ∈ s, b i = v i := +begin + have hsv : injective (s.restrict v) := hv.linear_independent.injective, + have hX : orthonormal 𝕜 (coe : set.range (s.restrict v) → E), + { rwa orthonormal_subtype_range hsv }, + obtain ⟨Y, b₀, hX, hb₀⟩ := hX.exists_orthonormal_basis_extension, + have hιY : fintype.card ι = Y.card, + { refine (card_ι.symm.trans _), + exact finite_dimensional.finrank_eq_card_finset_basis b₀.to_basis }, + have hvsY : s.maps_to v Y := (s.maps_to_image v).mono_right (by rwa ← range_restrict), + have hsv' : set.inj_on v s, + { rw set.inj_on_iff_injective, + exact hsv }, + obtain ⟨g, hg⟩ := hvsY.exists_equiv_extend_of_card_eq hιY hsv', + use b₀.reindex g.symm, + intros i hi, + { simp [hb₀, hg i hi] }, +end + variables (𝕜 E) /-- A finite-dimensional inner product space admits an orthonormal basis. -/ From 561202d0d491749f4c3fe9204405f4d2f575893d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 4 Oct 2022 05:52:15 +0000 Subject: [PATCH 547/828] refactor(data/list/count): review API, add lemmas (#16736) * add `list.countp_join`, `list.countp_map`, `list.countp_mono_left`, `list.countp_congr`, and `list.count_join`; * add `@[simp]` attrs to `list.countp_eq_zero`, `list.countp_eq_length`, `list.count_eq_zero`, and `list.count_eq_length`; * golf the proofs of `list.count_bind`, `list.count_map_of_injective`, and `list.count_le_count_map`. --- src/data/list/count.lean | 66 ++++++++++++++++++++++++---------------- 1 file changed, 40 insertions(+), 26 deletions(-) diff --git a/src/data/list/count.lean b/src/data/list/count.lean index b4479638e6b47..acff9e9fbff55 100644 --- a/src/data/list/count.lean +++ b/src/data/list/count.lean @@ -20,7 +20,7 @@ variables {α β : Type*} {l l₁ l₂ : list α} namespace list section countp -variables (p : α → Prop) [decidable_pred p] +variables (p q : α → Prop) [decidable_pred p] [decidable_pred q] @[simp] lemma countp_nil : countp p [] = 0 := rfl @@ -50,13 +50,17 @@ by simpa only [countp_eq_length_filter] using length_le_of_sublist (filter_subli @[simp] lemma countp_append (l₁ l₂) : countp p (l₁ ++ l₂) = countp p l₁ + countp p l₂ := by simp only [countp_eq_length_filter, filter_append, length_append] +lemma countp_join : ∀ l : list (list α), countp p l.join = (l.map (countp p)).sum +| [] := rfl +| (a :: l) := by rw [join, countp_append, map_cons, sum_cons, countp_join] + lemma countp_pos {l} : 0 < countp p l ↔ ∃ a ∈ l, p a := by simp only [countp_eq_length_filter, length_pos_iff_exists_mem, mem_filter, exists_prop] -theorem countp_eq_zero {l} : countp p l = 0 ↔ ∀ a ∈ l, ¬ p a := +@[simp] theorem countp_eq_zero {l} : countp p l = 0 ↔ ∀ a ∈ l, ¬ p a := by { rw [← not_iff_not, ← ne.def, ← pos_iff_ne_zero, countp_pos], simp } -lemma countp_eq_length {l} : countp p l = l.length ↔ ∀ a ∈ l, p a := +@[simp] lemma countp_eq_length {l} : countp p l = l.length ↔ ∀ a ∈ l, p a := by rw [countp_eq_length_filter, filter_length_eq_length] lemma length_filter_lt_length_iff_exists (l) : length (filter p l) < length l ↔ ∃ x ∈ l, ¬p x := @@ -65,15 +69,32 @@ by rw [length_eq_countp_add_countp p l, ← countp_pos, countp_eq_length_filter, lemma sublist.countp_le (s : l₁ <+ l₂) : countp p l₁ ≤ countp p l₂ := by simpa only [countp_eq_length_filter] using length_le_of_sublist (s.filter p) -@[simp] lemma countp_filter {q} [decidable_pred q] (l : list α) : - countp p (filter q l) = countp (λ a, p a ∧ q a) l := +@[simp] lemma countp_filter (l : list α) : countp p (filter q l) = countp (λ a, p a ∧ q a) l := by simp only [countp_eq_length_filter, filter_filter] -@[simp] lemma countp_true : l.countp (λ _, true) = l.length := -by simp [countp_eq_length_filter] +@[simp] lemma countp_true : l.countp (λ _, true) = l.length := by simp + +@[simp] lemma countp_false : l.countp (λ _, false) = 0 := by simp + +@[simp] lemma countp_map (p : β → Prop) [decidable_pred p] (f : α → β) : + ∀ l, countp p (map f l) = countp (p ∘ f) l +| [] := rfl +| (a::l) := by rw [map_cons, countp_cons, countp_cons, countp_map] + +variables {p q} + +lemma countp_mono_left (h : ∀ x ∈ l, p x → q x) : countp p l ≤ countp q l := +begin + induction l with a l ihl, { refl }, + rw [forall_mem_cons] at h, cases h with ha hl, + rw [countp_cons, countp_cons], + refine add_le_add (ihl hl) _, + split_ifs; try { simp only [le_rfl, zero_le] }, + exact absurd (ha ‹_›) ‹_› +end -@[simp] lemma countp_false : l.countp (λ _, false) = 0 := -by simp [countp_eq_length_filter] +lemma countp_congr (h : ∀ x ∈ l, p x ↔ q x) : countp p l = countp q l := +le_antisymm (countp_mono_left $ λ x hx, (h x hx).1) (countp_mono_left $ λ x hx, (h x hx).2) end countp @@ -115,6 +136,9 @@ lemma count_singleton' (a b : α) : count a [b] = ite (a = b) 1 0 := rfl @[simp] lemma count_append (a : α) : ∀ l₁ l₂, count a (l₁ ++ l₂) = count a l₁ + count a l₂ := countp_append _ +lemma count_join (l : list (list α)) (a : α) : l.join.count a = (l.map (count a)).sum := +countp_join _ _ + lemma count_concat (a : α) (l : list α) : count a (concat l a) = succ (count a l) := by simp [-add_comm] @@ -131,11 +155,11 @@ decidable.by_contradiction $ λ h', h $ count_pos.1 (nat.pos_of_ne_zero h') lemma not_mem_of_count_eq_zero {a : α} {l : list α} (h : count a l = 0) : a ∉ l := λ h', (count_pos.2 h').ne' h -lemma count_eq_zero {a : α} {l} : count a l = 0 ↔ a ∉ l := +@[simp] lemma count_eq_zero {a : α} {l} : count a l = 0 ↔ a ∉ l := ⟨not_mem_of_count_eq_zero, count_eq_zero_of_not_mem⟩ -lemma count_eq_length {a : α} {l} : count a l = l.length ↔ ∀ b ∈ l, a = b := -by rw [count, countp_eq_length] +@[simp] lemma count_eq_length {a : α} {l} : count a l = l.length ↔ ∀ b ∈ l, a = b := +countp_eq_length _ @[simp] lemma count_repeat (a : α) (n : ℕ) : count a (repeat a n) = n := by rw [count, countp_eq_length_filter, filter_eq_self.2, length_repeat]; @@ -160,28 +184,18 @@ by simp only [count, countp_filter, show (λ b, a = b ∧ p b) = eq a, by { ext lemma count_bind {α β} [decidable_eq β] (l : list α) (f : α → list β) (x : β) : count x (l.bind f) = sum (map (count x ∘ f) l) := -begin - induction l with hd tl IH, - { simp }, - { simpa } -end +by rw [list.bind, count_join, map_map] @[simp] lemma count_map_of_injective {α β} [decidable_eq α] [decidable_eq β] (l : list α) (f : α → β) (hf : function.injective f) (x : α) : count (f x) (map f l) = count x l := -begin - induction l with y l IH generalizing x, - { simp }, - { simp [map_cons, count_cons', IH, hf.eq_iff] } -end +by simp only [count, countp_map, (∘), hf.eq_iff] lemma count_le_count_map [decidable_eq β] (l : list α) (f : α → β) (x : α) : count x l ≤ count (f x) (map f l) := begin - induction l with a as IH, { simp }, - rcases eq_or_ne x a with rfl | hxa, - { simp [succ_le_succ IH] }, - { simp [hxa, le_add_right IH, count_cons'] } + rw [count, count, countp_map], + exact countp_mono_left (λ y hyl, congr_arg f), end @[simp] lemma count_erase_self (a : α) : From 522ace495538604270b08649e9aee5a226f20eae Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 4 Oct 2022 05:52:16 +0000 Subject: [PATCH 548/828] feat(order/antisymmetrization): (<) is well-founded for a preorder iff well-founded on its antisymmetrization (#16741) + Show that the antisymmetrization of a well-founded preorder is a well-founded partial order. Useful for reducing well-foundedness of preorders to partial orders. + fix reference to renamed lemma `well_founded.prod_game_add` in docstring. --- src/order/antisymmetrization.lean | 17 +++++++++++++++++ src/order/game_add.lean | 4 ++-- 2 files changed, 19 insertions(+), 2 deletions(-) diff --git a/src/order/antisymmetrization.lean b/src/order/antisymmetrization.lean index 4e8098fb5a811..ca73c3d0b8a09 100644 --- a/src/order/antisymmetrization.lean +++ b/src/order/antisymmetrization.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import order.hom.basic +import logic.relation /-! # Turning a preorder into a partial order @@ -108,6 +109,22 @@ instance : partial_order (antisymmetrization α (≤)) := lt_iff_le_not_le := λ a b, quotient.induction_on₂' a b $ λ a b, lt_iff_le_not_le, le_antisymm := λ a b, quotient.induction_on₂' a b $ λ a b hab hba, quotient.sound' ⟨hab, hba⟩ } +lemma antisymmetrization_fibration : + relation.fibration (<) (<) (@to_antisymmetrization α (≤) _) := +by { rintro a ⟨b⟩ h, exact ⟨b, h, rfl⟩ } + +lemma acc_antisymmetrization_iff : acc (<) (to_antisymmetrization (≤) a) ↔ acc (<) a := +⟨λ h, by { have := inv_image.accessible _ h, exact this }, + acc.of_fibration _ antisymmetrization_fibration⟩ + +lemma well_founded_antisymmetrization_iff : + well_founded (@has_lt.lt (antisymmetrization α (≤)) _) ↔ well_founded (@has_lt.lt α _) := +⟨λ h, ⟨λ a, acc_antisymmetrization_iff.1 $ h.apply _⟩, + λ h, ⟨by { rintro ⟨a⟩, exact acc_antisymmetrization_iff.2 (h.apply a) }⟩⟩ + +instance [well_founded_lt α] : well_founded_lt (antisymmetrization α (≤)) := +⟨well_founded_antisymmetrization_iff.2 is_well_founded.wf⟩ + instance [@decidable_rel α (≤)] [@decidable_rel α (<)] [is_total α (≤)] : linear_order (antisymmetrization α (≤)) := { le_total := λ a b, quotient.induction_on₂' a b $ total_of (≤), diff --git a/src/order/game_add.lean b/src/order/game_add.lean index c72282c7729f9..8bc7d3cc51ee8 100644 --- a/src/order/game_add.lean +++ b/src/order/game_add.lean @@ -18,8 +18,8 @@ subsequency relation on the addition of combinatorial games. ## Main definitions and results - `prod.game_add`: the game addition relation on ordered pairs. -- `well_founded.game_add`: formalizes induction on ordered pairs, where exactly one entry decreases - at a time. +- `well_founded.prod_game_add`: formalizes induction on ordered pairs, where exactly one entry + decreases at a time. ## Todo From 8230606cba155eb0d016600606b377c58130c385 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Tue, 4 Oct 2022 05:52:17 +0000 Subject: [PATCH 549/828] feat(ring_theory/roots_of_unity): add geom_sum_eq_zero (#16755) We add `geom_sum_eq_zero`: the geometric sum of a primitive root of unity is `0`. From flt-regular --- src/ring_theory/roots_of_unity.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity.lean index 453fd90d7e9b9..c14112380a9fc 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity.lean @@ -319,6 +319,9 @@ end lemma pow_ne_one_of_pos_of_lt (h0 : 0 < l) (hl : l < k) : ζ ^ l ≠ 1 := mt (nat.le_of_dvd h0 ∘ h.dvd_of_pow_eq_one _) $ not_le_of_lt hl +lemma ne_one {ζ : M} (h : is_primitive_root ζ k) (hk : 1 < k) : ζ ≠ 1 := +h.pow_ne_one_of_pos_of_lt zero_lt_one hk ∘ (pow_one ζ).trans + lemma pow_inj (h : is_primitive_root ζ k) ⦃i j : ℕ⦄ (hi : i < k) (hj : j < k) (H : ζ ^ i = ζ ^ j) : i = j := begin @@ -599,6 +602,20 @@ begin rwa ring_char.eq_iff.mpr h end +/-- If `1 < k` then `(∑ i in range k, ζ ^ i) = 0`. -/ +lemma geom_sum_eq_zero [is_domain R] {ζ : R} (hζ : is_primitive_root ζ k) (hk : 1 < k) : + (∑ i in range k, ζ ^ i) = 0 := +begin + refine eq_zero_of_ne_zero_of_mul_left_eq_zero (sub_ne_zero_of_ne (hζ.ne_one hk).symm) _, + rw [mul_neg_geom_sum, hζ.pow_eq_one, sub_self] +end + +/-- If `1 < k`, then `ζ ^ k.pred = -(∑ i in range k.pred, ζ ^ i)`. -/ +lemma pow_sub_one_eq [is_domain R] {ζ : R} (hζ : is_primitive_root ζ k) (hk : 1 < k) : + ζ ^ k.pred = -(∑ i in range k.pred, ζ ^ i) := +by rw [eq_neg_iff_add_eq_zero, add_comm, ←sum_range_succ, ←nat.succ_eq_add_one, + nat.succ_pred_eq_of_pos (pos_of_gt hk), hζ.geom_sum_eq_zero hk] + /-- The (additive) monoid equivalence between `zmod k` and the powers of a primitive root of unity `ζ`. -/ def zmod_equiv_zpowers (h : is_primitive_root ζ k) : zmod k ≃+ additive (subgroup.zpowers ζ) := From 3f6d1de30ca0a351c03ab6df1d59a2439a238d33 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 4 Oct 2022 05:52:18 +0000 Subject: [PATCH 550/828] feat(overview.yaml): add divergence theorem (#16781) --- docs/overview.yaml | 1 + 1 file changed, 1 insertion(+) diff --git a/docs/overview.yaml b/docs/overview.yaml index b59896809a824..fae0637b299f0 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -333,6 +333,7 @@ Analysis: approximation by convolution: 'cont_diff_bump_of_inner.convolution_tendsto_right' regularization by convolution: 'has_compact_support.cont_diff_convolution_left' change of variables formula: 'measure_theory.integral_image_eq_integral_abs_det_fderiv_smul' + divergence theorem: 'measure_theory.integral_divergence_of_has_fderiv_within_at_off_countable' Complex analysis: Cauchy integral formula: 'diff_cont_on_cl.circle_integral_sub_inv_smul' From efb178db9dbb0902369361e1099ab4edaad140fd Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Tue, 4 Oct 2022 05:52:19 +0000 Subject: [PATCH 551/828] chore(scripts): update nolints.txt (#16793) I am happy to remove some nolints for you! --- scripts/nolints.txt | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 923931412d535..aade010d56cee 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -331,7 +331,6 @@ apply_nolint rbnode.is_bad_red_black doc_blame apply_nolint computation.bind.F doc_blame apply_nolint computation.bind.G doc_blame apply_nolint computation.bisim_o doc_blame -apply_nolint computation.cases_on doc_blame apply_nolint computation.corec.F doc_blame apply_nolint computation.is_bisimulation doc_blame apply_nolint computation.lift_rel_aux doc_blame @@ -347,14 +346,12 @@ apply_nolint computation.parallel_rec doc_blame -- data/seq/seq.lean apply_nolint seq.bisim_o doc_blame -apply_nolint seq.cases_on doc_blame apply_nolint seq.corec.F doc_blame apply_nolint seq.is_bisimulation doc_blame apply_nolint seq.mem doc_blame -- data/seq/wseq.lean apply_nolint wseq.bisim_o doc_blame -apply_nolint wseq.cases_on doc_blame apply_nolint wseq.destruct_append.aux doc_blame apply_nolint wseq.destruct_join.aux doc_blame apply_nolint wseq.drop.aux doc_blame @@ -933,4 +930,4 @@ apply_nolint transfer.compute_transfer doc_blame apply_nolint tactic.wlog doc_blame -- topology/algebra/group.lean -apply_nolint topological_space.positive_compacts.locally_compact_space_of_group to_additive_doc +apply_nolint topological_space.positive_compacts.locally_compact_space_of_group to_additive_doc \ No newline at end of file From 7aaa53f772917468e4774f7da5aa8c4205fc2b18 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 4 Oct 2022 09:14:55 +0000 Subject: [PATCH 552/828] refactor(order/succ_pred/limit): redefine successor/predecessor limits (#15655) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I realized too late that there is a much simpler definition for a successor/predecessor limit (defined in #15001) in terms of the covering relation. We redefine these predicates and port over the existing API. Note that this new definition is only equivalent to the old one in partial orders. The API has remained exactly the same with two caveats: - A few theorems need their assumptions strengthened from preorders to partial orders. - Since `is_succ_limit` is now defined outside of the `order` namespace, all its theorems are moved outside of it too. Co-authored-by: Yaël Dillies --- src/order/succ_pred/limit.lean | 257 ++++++++++++++++++--------------- 1 file changed, 141 insertions(+), 116 deletions(-) diff --git a/src/order/succ_pred/limit.lean b/src/order/succ_pred/limit.lean index f3033b25fb13d..cdd69c0596b33 100644 --- a/src/order/succ_pred/limit.lean +++ b/src/order/succ_pred/limit.lean @@ -9,8 +9,9 @@ import order.succ_pred.basic /-! # Successor and predecessor limits -We define the predicate `order.is_succ_limit` for "successor limits", values that aren't successors, -except possibly of themselves. We define `order.is_pred_limit` analogously, and prove basic results. +We define the predicate `order.is_succ_limit` for "successor limits", values that don't cover any +others. They are so named since they can't be the successors of anything smaller. We define +`order.is_pred_limit` analogously, and prove basic results. ## Todo @@ -18,80 +19,55 @@ The plan is to eventually replace `ordinal.is_limit` and `cardinal.is_limit` wit predicate `order.is_succ_limit`. -/ -variables {α : Type*} {a b : α} +variables {α : Type*} -open function set +namespace order +open function set order_dual /-! ### Successor limits -/ -namespace order -section preorder -variables [preorder α] [succ_order α] +section has_lt +variables [has_lt α] -/-- A successor limit is a value that isn't a successor, except possibly of itself. -/ -def is_succ_limit (a : α) : Prop := ∀ b, succ b = a → is_max b +/-- A successor limit is a value that doesn't cover any other. -protected lemma is_succ_limit.is_max (h : is_succ_limit (succ a)) : is_max a := h a rfl +It's so named because in a successor order, a successor limit can't be the successor of anything +smaller. -/ +def is_succ_limit (a : α) : Prop := ∀ b, ¬ b ⋖ a -lemma not_is_succ_limit_succ_of_not_is_max (ha : ¬ is_max a) : ¬ is_succ_limit (succ a) := -λ h, ha (h a rfl) +lemma not_is_succ_limit_iff_exists_covby (a : α) : ¬ is_succ_limit a ↔ ∃ b, b ⋖ a := +by simp [is_succ_limit] -protected lemma _root_.is_min.is_succ_limit (h : is_min a) : is_succ_limit a := -by { rintros b rfl, exact max_of_succ_le (h $ le_succ b) } +@[simp] lemma is_succ_limit_of_dense [densely_ordered α] (a : α) : is_succ_limit a := λ b, not_covby -lemma is_succ_limit_of_succ_ne (h : ∀ b, succ b ≠ a) : is_succ_limit a := λ b hb, (h _ hb).elim +end has_lt -lemma not_is_succ_limit_iff : ¬ is_succ_limit a ↔ ∃ b, ¬ is_max b ∧ succ b = a := -by simp [is_succ_limit, and_comm] +section preorder +variables [preorder α] {a : α} -/-- See `order.not_is_succ_limit_iff` for a version that states that `a` is a successor of a value -other than itself. -/ -lemma mem_range_succ_of_not_is_succ_limit (h : ¬ is_succ_limit a) : a ∈ range (@succ α _ _) := -by { cases not_is_succ_limit_iff.1 h with b hb, exact ⟨b, hb.2⟩ } +protected lemma _root_.is_min.is_succ_limit : is_min a → is_succ_limit a := +λ h b hab, not_is_min_of_lt hab.lt h -lemma is_succ_limit_of_succ_lt (H : ∀ a < b, succ a < b) : is_succ_limit b := -by { rintros a rfl, by_contra ha, exact (H a $ lt_succ_of_not_is_max ha).false } +lemma is_succ_limit_bot [order_bot α] : is_succ_limit (⊥ : α) := is_min_bot.is_succ_limit -/-- A value can be built by building it on successors and successor limits. +variables [succ_order α] -Note that you need a partial order for data built using this to behave nicely on successors. -/ -@[elab_as_eliminator] noncomputable def is_succ_limit_rec_on {C : α → Sort*} (b) - (hs : Π a, ¬ is_max a → C (succ a)) (hl : Π a, is_succ_limit a → C a) : C b := -begin - by_cases hb : is_succ_limit b, - { exact hl b hb }, - { have H := classical.some_spec (not_is_succ_limit_iff.1 hb), - rw ←H.2, - exact hs _ H.1 } -end +protected lemma is_succ_limit.is_max (h : is_succ_limit (succ a)) : is_max a := +by { by_contra H, exact h a (covby_succ_of_not_is_max H) } -lemma is_succ_limit_rec_on_limit {C : α → Sort*} (hs : Π a, ¬ is_max a → C (succ a)) - (hl : Π a, is_succ_limit a → C a) (hb : is_succ_limit b) : - @is_succ_limit_rec_on α _ _ C b hs hl = hl b hb := -by { classical, exact dif_pos hb } +lemma not_is_succ_limit_succ_of_not_is_max (ha : ¬ is_max a) : ¬ is_succ_limit (succ a) := +by { contrapose! ha, exact ha.is_max } section no_max_order variables [no_max_order α] -lemma is_succ_limit.succ_ne (h : is_succ_limit a) (b) : succ b ≠ a := λ hb, not_is_max b (h b hb) +lemma is_succ_limit.succ_ne (h : is_succ_limit a) (b : α) : succ b ≠ a := +by { rintro rfl, exact not_is_max _ h.is_max } @[simp] lemma not_is_succ_limit_succ (a : α) : ¬ is_succ_limit (succ a) := λ h, h.succ_ne _ rfl -lemma is_succ_limit_iff_succ_ne : is_succ_limit a ↔ ∀ b, succ b ≠ a := -⟨is_succ_limit.succ_ne, is_succ_limit_of_succ_ne⟩ - -lemma not_is_succ_limit_iff' : ¬ is_succ_limit a ↔ a ∈ range (@succ α _ _) := -by { simp_rw [is_succ_limit_iff_succ_ne, not_forall, not_ne_iff], refl } - end no_max_order -section order_bot -variable [order_bot α] - -lemma is_succ_limit_bot : is_succ_limit (⊥ : α) := is_min_bot.is_succ_limit - -end order_bot - section is_succ_archimedean variable [is_succ_archimedean α] @@ -112,7 +88,25 @@ end is_succ_archimedean end preorder section partial_order -variables [partial_order α] [succ_order α] +variables [partial_order α] [succ_order α] {a b : α} {C : α → Sort*} + +lemma is_succ_limit_of_succ_ne (h : ∀ b, succ b ≠ a) : is_succ_limit a := λ b hba, h b hba.succ_eq + +lemma not_is_succ_limit_iff : ¬ is_succ_limit a ↔ ∃ b, ¬ is_max b ∧ succ b = a := +begin + rw not_is_succ_limit_iff_exists_covby, + refine exists_congr (λ b, ⟨λ hba, ⟨hba.lt.not_is_max, hba.succ_eq⟩, _⟩), + rintro ⟨h, rfl⟩, + exact covby_succ_of_not_is_max h +end + +/-- See `not_is_succ_limit_iff` for a version that states that `a` is a successor of a value other +than itself. -/ +lemma mem_range_succ_of_not_is_succ_limit (h : ¬ is_succ_limit a) : a ∈ range (@succ α _ _) := +by { cases not_is_succ_limit_iff.1 h with b hb, exact ⟨b, hb.2⟩ } + +lemma is_succ_limit_of_succ_lt (H : ∀ a < b, succ a < b) : is_succ_limit b := +λ a hab, (H a hab.lt).ne hab.succ_eq lemma is_succ_limit.succ_lt (hb : is_succ_limit b) (ha : a < b) : succ a < b := begin @@ -130,7 +124,23 @@ lemma is_succ_limit.succ_lt_iff (hb : is_succ_limit b) : succ a < b ↔ a < b := lemma is_succ_limit_iff_succ_lt : is_succ_limit b ↔ ∀ a < b, succ a < b := ⟨λ hb a, hb.succ_lt, is_succ_limit_of_succ_lt⟩ -lemma is_succ_limit_rec_on_succ' {C : α → Sort*} (hs : Π a, ¬ is_max a → C (succ a)) +/-- A value can be built by building it on successors and successor limits. -/ +@[elab_as_eliminator] noncomputable def is_succ_limit_rec_on (b : α) + (hs : Π a, ¬ is_max a → C (succ a)) (hl : Π a, is_succ_limit a → C a) : C b := +begin + by_cases hb : is_succ_limit b, + { exact hl b hb }, + { have H := classical.some_spec (not_is_succ_limit_iff.1 hb), + rw ←H.2, + exact hs _ H.1 } +end + +lemma is_succ_limit_rec_on_limit (hs : Π a, ¬ is_max a → C (succ a)) + (hl : Π a, is_succ_limit a → C a) (hb : is_succ_limit b) : + @is_succ_limit_rec_on α _ _ C b hs hl = hl b hb := +by { classical, exact dif_pos hb } + +lemma is_succ_limit_rec_on_succ' (hs : Π a, ¬ is_max a → C (succ a)) (hl : Π a, is_succ_limit a → C a) {b : α} (hb : ¬ is_max b) : @is_succ_limit_rec_on α _ _ C (succ b) hs hl = hs b hb := begin @@ -146,11 +156,17 @@ end section no_max_order variables [no_max_order α] -@[simp] lemma is_succ_limit_rec_on_succ {C : α → Sort*} (hs : Π a, ¬ is_max a → C (succ a)) +@[simp] lemma is_succ_limit_rec_on_succ (hs : Π a, ¬ is_max a → C (succ a)) (hl : Π a, is_succ_limit a → C a) (b : α) : @is_succ_limit_rec_on α _ _ C (succ b) hs hl = hs b (not_is_max b) := is_succ_limit_rec_on_succ' _ _ _ +lemma is_succ_limit_iff_succ_ne : is_succ_limit a ↔ ∀ b, succ b ≠ a := +⟨is_succ_limit.succ_ne, is_succ_limit_of_succ_ne⟩ + +lemma not_is_succ_limit_iff' : ¬ is_succ_limit a ↔ a ∈ range (@succ α _ _) := +by { simp_rw [is_succ_limit_iff_succ_ne, not_forall, not_ne_iff], refl } + end no_max_order section is_succ_archimedean @@ -175,75 +191,65 @@ end partial_order /-! ### Predecessor limits -/ -section preorder -variables [preorder α] [pred_order α] +section has_lt +variables [has_lt α] {a : α} -/-- A predecessor limit is a value that isn't a predecessor, except possibly of itself. -/ -def is_pred_limit (a : α) : Prop := ∀ b, pred b = a → is_min b +/-- A predecessor limit is a value that isn't covered by any other. -protected lemma is_pred_limit.is_min (h : is_pred_limit (pred a)) : is_min a := h a rfl +It's so named because in a predecessor order, a predecessor limit can't be the predecessor of +anything greater. -/ +def is_pred_limit (a : α) : Prop := ∀ b, ¬ a ⋖ b -lemma not_is_pred_limit_pred_of_not_is_min (ha : ¬ is_min a) : ¬ is_pred_limit (pred a) := -λ h, ha (h a rfl) +lemma not_is_pred_limit_iff_exists_covby (a : α) : ¬ is_pred_limit a ↔ ∃ b, a ⋖ b := +by simp [is_pred_limit] -protected lemma _root_.is_max.is_pred_limit : is_max a → is_pred_limit a := -@is_min.is_succ_limit αᵒᵈ a _ _ +lemma is_pred_limit_of_dense [densely_ordered α] (a : α) : is_pred_limit a := λ b, not_covby -lemma is_pred_limit_of_pred_ne (h : ∀ b, pred b ≠ a) : is_pred_limit a := λ b hb, (h _ hb).elim +@[simp] lemma is_succ_limit_to_dual_iff : is_succ_limit (to_dual a) ↔ is_pred_limit a := +by simp [is_succ_limit, is_pred_limit] -lemma not_is_pred_limit_iff : ¬ is_pred_limit a ↔ ∃ b, ¬ is_min b ∧ pred b = a := -@not_is_succ_limit_iff αᵒᵈ _ _ _ +@[simp] lemma is_pred_limit_to_dual_iff : is_pred_limit (to_dual a) ↔ is_succ_limit a := +by simp [is_succ_limit, is_pred_limit] -/-- See `order.not_is_pred_limit_iff` for a version that states that `a` is a predecessor of a value -other than itself. -/ -lemma mem_range_pred_of_not_is_pred_limit : ¬ is_pred_limit a → a ∈ range (@pred α _ _) := -@mem_range_succ_of_not_is_succ_limit αᵒᵈ _ _ _ +alias is_succ_limit_to_dual_iff ↔ _ is_pred_limit.dual +alias is_pred_limit_to_dual_iff ↔ _ is_succ_limit.dual -lemma is_pred_limit_of_lt_pred : (∀ b > a, a < pred b) → is_pred_limit a := -@is_succ_limit_of_succ_lt αᵒᵈ a _ _ +end has_lt -/-- A value can be built by building it on predecessors and predecessor limits. +section preorder +variables [preorder α] {a : α} -Note that you need a partial order for data built using this to behave nicely on successors. -/ -@[elab_as_eliminator] noncomputable def is_pred_limit_rec_on : Π {C : α → Sort*} (b) - (hs : Π a, ¬ is_min a → C (pred a)) (hl : Π a, is_pred_limit a → C a), C b := -@is_succ_limit_rec_on αᵒᵈ _ _ +protected lemma _root_.is_max.is_pred_limit : is_max a → is_pred_limit a := +λ h b hab, not_is_max_of_lt hab.lt h -theorem is_pred_limit_rec_on_limit : Π {C : α → Sort*} (hs : Π a, ¬ is_min a → C (pred a)) - (hl : Π a, is_pred_limit a → C a) (hb : is_pred_limit b), - @is_pred_limit_rec_on α _ _ C b hs hl = hl b hb := -@is_succ_limit_rec_on_limit αᵒᵈ b _ _ +lemma is_pred_limit_top [order_top α] : is_pred_limit (⊤ : α) := is_max_top.is_pred_limit + +variables [pred_order α] + +protected lemma is_pred_limit.is_min (h : is_pred_limit (pred a)) : is_min a := +by { by_contra H, exact h a (pred_covby_of_not_is_min H) } + +lemma not_is_pred_limit_pred_of_not_is_min (ha : ¬ is_min a) : ¬ is_pred_limit (pred a) := +by { contrapose! ha, exact ha.is_min } section no_min_order variables [no_min_order α] -lemma is_pred_limit.pred_ne (h : is_pred_limit a) (b) : pred b ≠ a := λ hb, not_is_min b (h b hb) +lemma is_pred_limit.pred_ne (h : is_pred_limit a) (b : α) : pred b ≠ a := +by { rintro rfl, exact not_is_min _ h.is_min } @[simp] lemma not_is_pred_limit_pred (a : α) : ¬ is_pred_limit (pred a) := λ h, h.pred_ne _ rfl -lemma is_pred_limit_iff_pred_ne : is_pred_limit a ↔ ∀ b, pred b ≠ a := -@is_succ_limit_iff_succ_ne αᵒᵈ a _ _ _ - -lemma not_is_pred_limit_iff' : ¬ is_pred_limit a ↔ a ∈ range (@pred α _ _) := -@not_is_succ_limit_iff' αᵒᵈ a _ _ _ - end no_min_order -section order_top -variable [order_top α] - -lemma is_pred_limit_top : is_pred_limit (⊤ : α) := is_max_top.is_pred_limit - -end order_top - section is_pred_archimedean -variable [is_pred_archimedean α] +variables [is_pred_archimedean α] -protected lemma is_pred_limit.is_max_of_no_min [no_min_order α] : is_pred_limit a → is_max a := -@is_succ_limit.is_min_of_no_max αᵒᵈ a _ _ _ _ +protected lemma is_pred_limit.is_max_of_no_min [no_min_order α] (h : is_pred_limit a) : is_max a := +h.dual.is_min_of_no_max @[simp] lemma is_pred_limit_iff_of_no_min [no_min_order α] : is_pred_limit a ↔ is_max a := -@is_succ_limit_iff_of_no_max αᵒᵈ a _ _ _ _ +is_succ_limit_to_dual_iff.symm.trans is_succ_limit_iff_of_no_max lemma not_is_pred_limit_of_no_min [no_min_order α] [no_max_order α] : ¬ is_pred_limit a := by simp @@ -252,40 +258,59 @@ end is_pred_archimedean end preorder section partial_order -variables [partial_order α] [pred_order α] +variables [partial_order α] [pred_order α] {a b : α} {C : α → Sort*} + +lemma is_pred_limit_of_pred_ne (h : ∀ b, pred b ≠ a) : is_pred_limit a := λ b hba, h b hba.pred_eq + +lemma not_is_pred_limit_iff : ¬ is_pred_limit a ↔ ∃ b, ¬ is_min b ∧ pred b = a := +by { rw ←is_succ_limit_to_dual_iff, exact not_is_succ_limit_iff } -lemma is_pred_limit.lt_pred : ∀ (ha : is_pred_limit a), a < b → a < pred b := -@is_succ_limit.succ_lt αᵒᵈ b a _ _ +/-- See `not_is_pred_limit_iff` for a version that states that `a` is a successor of a value other +than itself. -/ +lemma mem_range_pred_of_not_is_pred_limit (h : ¬ is_pred_limit a) : a ∈ range (@pred α _ _) := +by { cases not_is_pred_limit_iff.1 h with b hb, exact ⟨b, hb.2⟩ } -lemma is_pred_limit.lt_pred_iff : ∀ (ha : is_pred_limit a), a < pred b ↔ a < b := -@is_succ_limit.succ_lt_iff αᵒᵈ b a _ _ +lemma is_pred_limit_of_pred_lt (H : ∀ a > b, pred a < b) : is_pred_limit b := +λ a hab, (H a hab.lt).ne hab.pred_eq -lemma is_pred_limit_iff_lt_pred : is_pred_limit a ↔ ∀ b > a, a < pred b := -@is_succ_limit_iff_succ_lt αᵒᵈ a _ _ +lemma is_pred_limit.lt_pred (h : is_pred_limit a) : a < b → a < pred b := h.dual.succ_lt +lemma is_pred_limit.lt_pred_iff (h : is_pred_limit a) : a < pred b ↔ a < b := h.dual.succ_lt_iff -lemma is_pred_limit_rec_on_pred' : ∀ {C : α → Sort*} (hs : Π a, ¬ is_min a → C (pred a)) - (hl : Π a, is_pred_limit a → C a) {b : α} (hb : ¬ is_min b), +lemma is_pred_limit_iff_lt_pred : is_pred_limit a ↔ ∀ ⦃b⦄, a < b → a < pred b := +is_succ_limit_to_dual_iff.symm.trans is_succ_limit_iff_succ_lt + +/-- A value can be built by building it on predecessors and predecessor limits. -/ +@[elab_as_eliminator] noncomputable def is_pred_limit_rec_on (b : α) + (hs : Π a, ¬ is_min a → C (pred a)) (hl : Π a, is_pred_limit a → C a) : C b := +@is_succ_limit_rec_on αᵒᵈ _ _ _ _ hs (λ a ha, hl _ ha.dual) + +lemma is_pred_limit_rec_on_limit (hs : Π a, ¬ is_min a → C (pred a)) + (hl : Π a, is_pred_limit a → C a) (hb : is_pred_limit b) : + @is_pred_limit_rec_on α _ _ C b hs hl = hl b hb := +is_succ_limit_rec_on_limit _ _ hb.dual + +lemma is_pred_limit_rec_on_pred' (hs : Π a, ¬ is_min a → C (pred a)) + (hl : Π a, is_pred_limit a → C a) {b : α} (hb : ¬ is_min b) : @is_pred_limit_rec_on α _ _ C (pred b) hs hl = hs b hb := -@is_succ_limit_rec_on_succ' αᵒᵈ _ _ +is_succ_limit_rec_on_succ' _ _ _ section no_min_order variables [no_min_order α] -@[simp] theorem is_pred_limit_rec_on_pred : ∀ {C : α → Sort*} (hs : Π a, ¬ is_min a → C (pred a)) - (hl : Π a, is_pred_limit a → C a) (b : α), +@[simp] theorem is_pred_limit_rec_on_pred (hs : Π a, ¬ is_min a → C (pred a)) + (hl : Π a, is_pred_limit a → C a) (b : α) : @is_pred_limit_rec_on α _ _ C (pred b) hs hl = hs b (not_is_min b) := -@is_succ_limit_rec_on_succ αᵒᵈ _ _ _ +is_succ_limit_rec_on_succ _ _ _ end no_min_order section is_pred_archimedean variable [is_pred_archimedean α] -protected lemma is_pred_limit.is_max : is_pred_limit a → is_max a := -@is_succ_limit.is_min αᵒᵈ a _ _ _ +protected lemma is_pred_limit.is_max (h : is_pred_limit a) : is_max a := h.dual.is_min @[simp] lemma is_pred_limit_iff : is_pred_limit a ↔ is_max a := -@is_succ_limit_iff αᵒᵈ a _ _ _ +is_succ_limit_to_dual_iff.symm.trans is_succ_limit_iff lemma not_is_pred_limit [no_max_order α] : ¬ is_pred_limit a := by simp From 933547832736be61a5de6576e22db351c6c2fbfd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 4 Oct 2022 10:22:24 +0000 Subject: [PATCH 553/828] refactor(algebra/symmetrized): Bundle `sym` and `unsym` as equivalences (#14334) Coercions to and from a type synonym should be bundled as (decorated) equivalences to be maximally useful. This PR bundles `sym_alg.sym` and `sym_alg.unsym` that way. As a result, `sym_alg.sym_equiv` is now redundant. --- src/algebra/symmetrized.lean | 25 +++++++++++-------------- 1 file changed, 11 insertions(+), 14 deletions(-) diff --git a/src/algebra/symmetrized.lean b/src/algebra/symmetrized.lean index f1016159ca8b7..c335e5b9864c5 100644 --- a/src/algebra/symmetrized.lean +++ b/src/algebra/symmetrized.lean @@ -42,28 +42,25 @@ variables {α : Type*} /-- The element of `sym_alg α` that represents `a : α`. -/ @[pattern,pp_nodot] -def sym : α → αˢʸᵐ := id +def sym : α ≃ αˢʸᵐ := equiv.refl _ /-- The element of `α` represented by `x : αˢʸᵐ`. -/ @[pp_nodot] -def unsym : αˢʸᵐ → α := id +def unsym : αˢʸᵐ ≃ α := equiv.refl _ @[simp] lemma unsym_sym (a : α) : unsym (sym a) = a := rfl @[simp] lemma sym_unsym (a : α) : sym (unsym a) = a := rfl - @[simp] lemma sym_comp_unsym : (sym : α → αˢʸᵐ) ∘ unsym = id := rfl @[simp] lemma unsym_comp_sym : (unsym : αˢʸᵐ → α) ∘ sym = id := rfl - -/-- The canonical bijection between `α` and `αˢʸᵐ`. -/ -@[simps apply symm_apply { fully_applied := ff }] -def sym_equiv : α ≃ αˢʸᵐ := ⟨sym, unsym, unsym_sym, sym_unsym⟩ - -lemma sym_bijective : bijective (sym : α → αˢʸᵐ) := sym_equiv.bijective -lemma unsym_bijective : bijective (unsym : αˢʸᵐ → α) := sym_equiv.symm.bijective -lemma sym_injective : injective (sym : α → αˢʸᵐ) := sym_bijective.injective -lemma sym_surjective : surjective (sym : α → αˢʸᵐ) := sym_bijective.surjective -lemma unsym_injective : injective (unsym : αˢʸᵐ → α) := unsym_bijective.injective -lemma unsym_surjective : surjective (unsym : αˢʸᵐ → α) := unsym_bijective.surjective +@[simp] lemma sym_symm : (@sym α).symm = unsym := rfl +@[simp] lemma unsym_symm : (@unsym α).symm = sym := rfl + +lemma sym_bijective : bijective (sym : α → αˢʸᵐ) := sym.bijective +lemma unsym_bijective : bijective (unsym : αˢʸᵐ → α) := unsym.symm.bijective +lemma sym_injective : injective (sym : α → αˢʸᵐ) := sym.injective +lemma sym_surjective : surjective (sym : α → αˢʸᵐ) := sym.surjective +lemma unsym_injective : injective (unsym : αˢʸᵐ → α) := unsym.injective +lemma unsym_surjective : surjective (unsym : αˢʸᵐ → α) := unsym.surjective @[simp] lemma sym_inj {a b : α} : sym a = sym b ↔ a = b := sym_injective.eq_iff @[simp] lemma unsym_inj {a b : αˢʸᵐ} : unsym a = unsym b ↔ a = b := unsym_injective.eq_iff From 16749fc4661828cba18cd0f4e3c5eb66a8e80598 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 4 Oct 2022 10:22:25 +0000 Subject: [PATCH 554/828] feat(algebra/monoid_algebra/support): lemmas about support and multiplying with single (#16326) The new lemmas use weaker assumptions and can be used to shorten some proofs. --- src/algebra/monoid_algebra/support.lean | 84 ++++++++++++++++++------- 1 file changed, 62 insertions(+), 22 deletions(-) diff --git a/src/algebra/monoid_algebra/support.lean b/src/algebra/monoid_algebra/support.lean index 5ed1a6baba907..cb6909ab07316 100644 --- a/src/algebra/monoid_algebra/support.lean +++ b/src/algebra/monoid_algebra/support.lean @@ -13,7 +13,55 @@ universes u₁ u₂ u₃ namespace monoid_algebra open finset finsupp -variables {k : Type u₁} {G : Type u₂} {R : Type u₃} [semiring k] +variables {k : Type u₁} {G : Type u₂} [semiring k] + +lemma support_single_mul_subset [decidable_eq G] [has_mul G] + (f : monoid_algebra k G) (r : k) (a : G) : + (single a r * f : monoid_algebra k G).support ⊆ finset.image ((*) a) f.support := +begin + intros x hx, + contrapose hx, + have : ∀ y, a * y = x → f y = 0, + { simpa only [not_and', mem_image, mem_support_iff, exists_prop, not_exists, not_not] using hx }, + simp only [mem_support_iff, mul_apply, sum_single_index, zero_mul, if_t_t, sum_zero, not_not], + exact finset.sum_eq_zero (by simp only [this, mem_support_iff, mul_zero, ne.def, + ite_eq_right_iff, eq_self_iff_true, implies_true_iff] {contextual := tt}), +end + +lemma support_mul_single_subset [decidable_eq G] [has_mul G] + (f : monoid_algebra k G) (r : k) (a : G) : + (f * single a r).support ⊆ finset.image (* a) f.support := +begin + intros x hx, + contrapose hx, + have : ∀ y, y * a = x → f y = 0, + { simpa only [not_and', mem_image, mem_support_iff, exists_prop, not_exists, not_not] using hx }, + simp only [mem_support_iff, mul_apply, sum_single_index, zero_mul, if_t_t, sum_zero, not_not], + exact finset.sum_eq_zero (by simp only [this, sum_single_index, ite_eq_right_iff, + eq_self_iff_true, implies_true_iff, zero_mul] {contextual := tt}), +end + +lemma support_single_mul_eq_image [decidable_eq G] [has_mul G] + (f : monoid_algebra k G) {r : k} (hr : ∀ y, r * y = 0 ↔ y = 0) {x : G} (lx : is_left_regular x) : + (single x r * f : monoid_algebra k G).support = finset.image ((*) x) f.support := +begin + refine subset_antisymm (support_single_mul_subset f _ _) (λ y hy, _), + obtain ⟨y, yf, rfl⟩ : ∃ (a : G), a ∈ f.support ∧ x * a = y, + { simpa only [finset.mem_image, exists_prop] using hy }, + simp only [mul_apply, mem_support_iff.mp yf, hr, mem_support_iff, sum_single_index, + finsupp.sum_ite_eq', ne.def, not_false_iff, if_true, zero_mul, if_t_t, sum_zero, lx.eq_iff] +end + +lemma support_mul_single_eq_image [decidable_eq G] [has_mul G] + (f : monoid_algebra k G) {r : k} (hr : ∀ y, y * r = 0 ↔ y = 0) {x : G} (rx : is_right_regular x) : + (f * single x r).support = finset.image (* x) f.support := +begin + refine subset_antisymm (support_mul_single_subset f _ _) (λ y hy, _), + obtain ⟨y, yf, rfl⟩ : ∃ (a : G), a ∈ f.support ∧ a * x = y, + { simpa only [finset.mem_image, exists_prop] using hy }, + simp only [mul_apply, mem_support_iff.mp yf, hr, mem_support_iff, sum_single_index, + finsupp.sum_ite_eq', ne.def, not_false_iff, if_true, mul_zero, if_t_t, sum_zero, rx.eq_iff] +end lemma support_mul [has_mul G] [decidable_eq G] (a b : monoid_algebra k G) : (a * b).support ⊆ a.support.bUnion (λa₁, b.support.bUnion $ λa₂, {a₁ * a₂}) := @@ -24,47 +72,39 @@ lemma support_mul_single [right_cancel_semigroup G] (f : monoid_algebra k G) (r : k) (hr : ∀ y, y * r = 0 ↔ y = 0) (x : G) : (f * single x r).support = f.support.map (mul_right_embedding x) := begin - ext y, simp only [mem_support_iff, mem_map, exists_prop, mul_right_embedding_apply], - by_cases H : ∃ a, a * x = y, - { rcases H with ⟨a, rfl⟩, - rw [mul_single_apply_aux f (λ _, mul_left_inj x)], - simp [hr] }, - { push_neg at H, - classical, - simp [mul_apply, H] } + classical, + ext, + simp only [support_mul_single_eq_image f hr (is_right_regular_of_right_cancel_semigroup x), + mem_image, mem_map, mul_right_embedding_apply], end lemma support_single_mul [left_cancel_semigroup G] (f : monoid_algebra k G) (r : k) (hr : ∀ y, r * y = 0 ↔ y = 0) (x : G) : (single x r * f : monoid_algebra k G).support = f.support.map (mul_left_embedding x) := begin - ext y, simp only [mem_support_iff, mem_map, exists_prop, mul_left_embedding_apply], - by_cases H : ∃ a, x * a = y, - { rcases H with ⟨a, rfl⟩, - rw [single_mul_apply_aux f (λ _, mul_right_inj x)], - simp [hr] }, - { push_neg at H, - classical, - simp [mul_apply, H] } + classical, + ext, + simp only [support_single_mul_eq_image f hr (is_left_regular_of_left_cancel_semigroup x), + mem_image, mem_map, mul_left_embedding_apply], end + section span variables [mul_one_class G] -/-- An element of `monoid_algebra R M` is in the subalgebra generated by its support. -/ +/-- An element of `monoid_algebra k G` is in the subalgebra generated by its support. -/ lemma mem_span_support (f : monoid_algebra k G) : f ∈ submodule.span k (of k G '' (f.support : set G)) := by rw [of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single, finsupp.mem_supported] end span - end monoid_algebra namespace add_monoid_algebra open finset finsupp mul_opposite -variables {k : Type u₁} {G : Type u₂} {R : Type u₃} [semiring k] +variables {k : Type u₁} {G : Type u₂} [semiring k] lemma support_mul [decidable_eq G] [has_add G] (a b : add_monoid_algebra k G) : (a * b).support ⊆ a.support.bUnion (λa₁, b.support.bUnion $ λa₂, {a₁ + a₂}) := @@ -82,12 +122,12 @@ lemma support_single_mul [add_left_cancel_semigroup G] section span -/-- An element of `add_monoid_algebra R M` is in the submodule generated by its support. -/ +/-- An element of `add_monoid_algebra k G` is in the submodule generated by its support. -/ lemma mem_span_support [add_zero_class G] (f : add_monoid_algebra k G) : f ∈ submodule.span k (of k G '' (f.support : set G)) := by rw [of, monoid_hom.coe_mk, ← finsupp.supported_eq_span_single, finsupp.mem_supported] -/-- An element of `add_monoid_algebra R M` is in the subalgebra generated by its support, using +/-- An element of `add_monoid_algebra k G` is in the subalgebra generated by its support, using unbundled inclusion. -/ lemma mem_span_support' (f : add_monoid_algebra k G) : f ∈ submodule.span k (of' k G '' (f.support : set G)) := From 91df768e6138ea91f937fe896a3276ab9bbf34a1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 4 Oct 2022 10:22:27 +0000 Subject: [PATCH 555/828] feat(tactic/positivity): Extension for `int.nat_abs` (#16787) Add `positivity_nat_abs`, a positivity extension for `int.nat_abs`. --- src/tactic/positivity.lean | 15 +++++++++++++++ test/positivity.lean | 2 ++ 2 files changed, 17 insertions(+) diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index bfa185ef155cc..92b8d0586e829 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -413,6 +413,21 @@ meta def positivity_abs : expr → tactic strictness nonnegative <$> mk_app ``abs_nonneg [a] -- else report nonnegativity | _ := failed +private lemma int_nat_abs_pos {n : ℤ} (hn : 0 < n) : 0 < n.nat_abs := +int.nat_abs_pos_of_ne_zero hn.ne' + +/-- Extension for the `positivity` tactic: `int.nat_abs` is positive when its input is. + +Since the output type of `int.nat_abs` is `ℕ`, the nonnegative case is handled by the default +`positivity` tactic. +-/ +@[positivity] +meta def positivity_nat_abs : expr → tactic strictness +| `(int.nat_abs %%a) := do + positive p ← core a, + positive <$> mk_app ``int_nat_abs_pos [p] +| _ := failed + private lemma nat_cast_pos [ordered_semiring α] [nontrivial α] {n : ℕ} : 0 < n → 0 < (n : α) := nat.cast_pos.2 diff --git a/test/positivity.lean b/test/positivity.lean index 4cd0aab0a1c54..fdf579374298c 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -105,6 +105,8 @@ example {a : ℤ} : 0 ≤ |a| := by positivity example {a : ℤ} : 0 < |a| + 3 := by positivity +example {n : ℤ} (hn : 0 < n) : 0 < n.nat_abs := by positivity + example {a : ℤ} (ha : 1 < a) : 0 < |(3:ℤ) + a| := by positivity example {a : ℝ} (ha : 0 ≤ a) : 0 ≤ real.sqrt a := by positivity From 9fe85cd025655fc6285758c29b6a392c0fa96743 Mon Sep 17 00:00:00 2001 From: BillyMiao Date: Tue, 4 Oct 2022 12:55:47 +0000 Subject: [PATCH 556/828] feat(tactic/qify) : `qify` tactic (#16313) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This file mainly introduces `qify` which casts goals or hypotheses about natural numbers or integers to the ones about rational numbers. Note that this file is following from `zify`. This is motivated by thinking about division in natural numbers or integers. Division in natural numbers or integers is not always working fine (e.g. (5 : ℕ) / 2 = 2), so it's easier to work in `ℚ`, where division and subtraction are well behaved. Co-authored-by: Alex J Best --- src/tactic/qify.lean | 160 +++++++++++++++++++++++++++++++++++++++++++ test/qify.lean | 23 +++++++ 2 files changed, 183 insertions(+) create mode 100644 src/tactic/qify.lean create mode 100644 test/qify.lean diff --git a/src/tactic/qify.lean b/src/tactic/qify.lean new file mode 100644 index 0000000000000..2f7b8be0dbd17 --- /dev/null +++ b/src/tactic/qify.lean @@ -0,0 +1,160 @@ +/- +Copyright (c) 2022 Jiale Miao. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jiale Miao +-/ + +import tactic.norm_cast +import data.rat.cast + +/-! +# A tactic to shift `ℕ` or `ℤ` goals to `ℚ` + +Note that this file is following from `zify`. + +Division in `ℕ` and `ℤ` is not always working fine (e.g. (5 : ℕ) / 2 = 2), so it's easier +to work in `ℚ`, where division and subtraction are well behaved. `qify` can be used to cast goals +and hypotheses about natural numbers or integers to rational numbers. It makes use of `push_cast`, +part of the `norm_cast` family, to simplify these goals. + +## Implementation notes + +`qify` is extensible, using the attribute `@[qify]` to label lemmas used for moving propositions +from `ℕ` or `ℤ` to `ℚ`. +`qify` lemmas should have the form `∀ a₁ ... aₙ : ℕ, Pq (a₁ : ℚ) ... (aₙ : ℚ) ↔ Pn a₁ ... aₙ`. +For example, `rat.coe_nat_le_coe_nat_iff : ∀ (m n : ℕ), ↑m ≤ ↑n ↔ m ≤ n` is a `qify` lemma. + +`qify` is very nearly just `simp only with qify push_cast`. There are a few minor differences: +* `qify` lemmas are used in the opposite order of the standard simp form. + E.g. we will rewrite with `rat.coe_nat_le_coe_nat_iff` from right to left. +* `qify` should fail if no `qify` lemma applies (i.e. it was unable to shift any proposition to ℚ). + However, once this succeeds, it does not necessarily need to rewrite with any `push_cast` rules. +-/ + +open tactic + +namespace qify + +/-- +The `qify` attribute is used by the `qify` tactic. It applies to lemmas that shift propositions +from `nat` or `int` to `rat`. + +`qify` lemmas should have the form `∀ a₁ ... aₙ : ℕ, Pq (a₁ : ℚ) ... (aₙ : ℚ) ↔ Pn a₁ ... aₙ` or + `∀ a₁ ... aₙ : ℤ, Pq (a₁ : ℚ) ... (aₙ : ℚ) ↔ Pz a₁ ... aₙ`. +For example, `rat.coe_nat_le_coe_nat_iff : ∀ (m n : ℕ), ↑m ≤ ↑n ↔ m ≤ n` is a `qify` lemma. +-/ +@[user_attribute] +meta def qify_attr : user_attribute simp_lemmas unit := +{ name := `qify, + descr := "Used to tag lemmas for use in the `qify` tactic", + cache_cfg := + { mk_cache := + λ ns, mmap (λ n, do c ← mk_const n, return (c, tt)) ns >>= simp_lemmas.mk.append_with_symm, + dependencies := [] } } + +/-- +Given an expression `e`, `lift_to_q e` looks for subterms of `e` that are propositions "about" +natural numbers or integers and change them to propositions about rational numbers. + +Returns an expression `e'` and a proof that `e = e'`. + +Includes `ge_iff_le` and `gt_iff_lt` in the simp set. These can't be tagged with `qify` as we +want to use them in the "forward", not "backward", direction. +-/ +meta def lift_to_q (e : expr) : tactic (expr × expr) := +do sl ← qify_attr.get_cache, + sl ← sl.add_simp `ge_iff_le, sl ← sl.add_simp `gt_iff_lt, + (e', prf, _) ← simplify sl [] e, + return (e', prf) + +@[qify] lemma rat.coe_nat_le_coe_nat_iff (a b : ℕ) : (a : ℚ) ≤ b ↔ a ≤ b := by simp +@[qify] lemma rat.coe_nat_lt_coe_nat_iff (a b : ℕ) : (a : ℚ) < b ↔ a < b := by simp +@[qify] lemma rat.coe_nat_eq_coe_nat_iff (a b : ℕ) : (a : ℚ) = b ↔ a = b := by simp +@[qify] lemma rat.coe_nat_ne_coe_nat_iff (a b : ℕ) : (a : ℚ) ≠ b ↔ a ≠ b := by simp + +@[qify] lemma rat.coe_int_le_coe_int_iff (a b : ℤ) : (a : ℚ) ≤ b ↔ a ≤ b := by simp +@[qify] lemma rat.coe_int_lt_coe_int_iff (a b : ℤ) : (a : ℚ) < b ↔ a < b := by simp +@[qify] lemma rat.coe_int_eq_coe_int_iff (a b : ℤ) : (a : ℚ) = b ↔ a = b := by simp +@[qify] lemma rat.coe_int_ne_coe_int_iff (a b : ℤ) : (a : ℚ) ≠ b ↔ a ≠ b := by simp + +end qify + +/-- +`qify extra_lems e` is used to shift propositions in `e` from `ℕ` or `ℤ` to `ℚ`. +This is often useful since `ℚ` has well-behaved division and subtraction. + +The list of extra lemmas is used in the `push_cast` step. + +Returns an expression `e'` and a proof that `e = e'`.-/ +meta def tactic.qify (extra_lems : list simp_arg_type) : expr → tactic (expr × expr) := λ q, +do (q1, p1) ← qify.lift_to_q q <|> fail "failed to find an applicable qify lemma", + (q2, p2) ← norm_cast.derive_push_cast extra_lems q1, + prod.mk q2 <$> mk_eq_trans p1 p2 + +/-- +A variant of `tactic.qify` that takes `h`, a proof of a proposition about natural numbers +or integers, and returns a proof of the qified version of that propositon. +-/ +meta def tactic.qify_proof (extra_lems : list simp_arg_type) (h : expr) : tactic expr := +do (_, pf) ← infer_type h >>= tactic.qify extra_lems, + mk_eq_mp pf h + +section + +setup_tactic_parser + +/-- +The `qify` tactic is used to shift propositions from `ℕ` or `ℤ` to `ℚ`. +This is often useful since `ℚ` has well-behaved division and subtraction. + +```lean +example (a b c : ℕ) (x y z : ℤ) (h : ¬ x*y*z < 0) : c < a + 3*b := +begin + qify, + qify at h, + /- + h : ¬↑x * ↑y * ↑z < 0 + ⊢ ↑c < ↑a + 3 * ↑b + -/ +end +``` + +`qify` can be given extra lemmas to use in simplification. This is especially useful in the +presence of subtraction and division: passing `≤` or `∣` arguments will allow `push_cast` +to do more work. +``` +example (a b c : ℕ) (h : a - b < c) (hab : b ≤ a) : false := +begin + qify [hab] at h, + /- h : ↑a - ↑b < ↑c -/ +end +``` + +``` +example (a b c : ℕ) (h : a / b = c) (hab : b ∣ a) : false := +begin + qify [hab] at h, + /- h : ↑a / ↑b = ↑c -/ +end +``` + +`qify` makes use of the `@[qify]` attribute to move propositions, +and the `push_cast` tactic to simplify the `ℚ`-valued expressions. +-/ +meta def tactic.interactive.qify (sl : parse simp_arg_list) (l : parse location) : tactic unit := +do locs ← l.get_locals, +replace_at (tactic.qify sl) locs l.include_goal >>= guardb + +end + +add_tactic_doc +{ name := "qify", + category := doc_category.attr, + decl_names := [`qify.qify_attr], + tags := ["coercions", "transport"] } + +add_tactic_doc +{ name := "qify", + category := doc_category.tactic, + decl_names := [`tactic.interactive.qify], + tags := ["coercions", "transport"] } diff --git a/test/qify.lean b/test/qify.lean new file mode 100644 index 0000000000000..733cacf73dcb7 --- /dev/null +++ b/test/qify.lean @@ -0,0 +1,23 @@ +import tactic.qify + +example (a b : ℕ) : (a : ℚ) ≤ b ↔ a ≤ b := by qify ; refl +example (a b : ℕ) : (a : ℚ) < b ↔ a < b := by qify ; refl +example (a b : ℕ) : (a : ℚ) = b ↔ a = b := by qify ; refl +example (a b : ℕ) : (a : ℚ) ≠ b ↔ a ≠ b := by qify ; refl + +example (a b : ℤ) : (a : ℚ) ≤ b ↔ a ≤ b := by qify ; refl +example (a b : ℤ) : (a : ℚ) < b ↔ a < b := by qify ; refl +example (a b : ℤ) : (a : ℚ) = b ↔ a = b := by qify ; refl +example (a b : ℤ) : (a : ℚ) ≠ b ↔ a ≠ b := by qify ; refl + +example (a b c : ℕ) (h : a - b = c) (hab : b ≤ a) : a = c + b := +begin + qify [hab] at h ⊢, -- `zify` does the same thing here. + exact sub_eq_iff_eq_add.1 h, +end + +example (a b c : ℤ) (h : a / b = c) (hab : b ∣ a) (hb : b ≠ 0) : a = c * b := +begin + qify [hab] at h hb ⊢, + exact (div_eq_iff hb).1 h, +end From b2e818bb0daee5ed718825330a7a1c59603b0fb4 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 4 Oct 2022 12:55:48 +0000 Subject: [PATCH 557/828] chore(data/complex/basic): golf 2 proofs (#16788) --- src/data/complex/basic.lean | 10 ++-------- 1 file changed, 2 insertions(+), 8 deletions(-) diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index a6ee5d7cbc964..58860cb1368cc 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -554,16 +554,10 @@ map_pow abs z n map_zpow₀ abs z n lemma abs_re_le_abs (z : ℂ) : |z.re| ≤ abs z := -begin - rw [mul_self_le_mul_self_iff (abs_nonneg z.re) (abs.nonneg _), - abs_mul_abs_self, mul_self_abs], - apply re_sq_le_norm_sq -end +real.abs_le_sqrt $ by { rw [norm_sq_apply, ← sq], exact le_add_of_nonneg_right (mul_self_nonneg _) } lemma abs_im_le_abs (z : ℂ) : |z.im| ≤ abs z := -by rw [mul_self_le_mul_self_iff (abs_nonneg z.im) (abs.nonneg _), - abs_mul_abs_self, mul_self_abs]; - apply im_sq_le_norm_sq +real.abs_le_sqrt $ by { rw [norm_sq_apply, ← sq, ← sq], exact le_add_of_nonneg_left (sq_nonneg _) } lemma re_le_abs (z : ℂ) : z.re ≤ abs z := (abs_le.1 (abs_re_le_abs _)).2 From 6858e178a7f5a8bc0480046a3f9381d75c55b5da Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Tue, 4 Oct 2022 18:24:25 +0000 Subject: [PATCH 558/828] feat(number_theory/padic_numbers): bundle padic_norm_e (#16650) We turn `padic_norm_e` into a docs#absolute_value, removing the need for many specific methods. Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- src/number_theory/padics/padic_numbers.lean | 92 ++++++++------------- 1 file changed, 33 insertions(+), 59 deletions(-) diff --git a/src/number_theory/padics/padic_numbers.lean b/src/number_theory/padics/padic_numbers.lean index 139d6412590f8..1d39e59165e9e 100644 --- a/src/number_theory/padics/padic_numbers.lean +++ b/src/number_theory/padics/padic_numbers.lean @@ -446,6 +446,8 @@ def mk : padic_seq p → ℚ_[p] := quotient.mk variables (p) +lemma zero_def : (0 : ℚ_[p]) = ⟦0⟧ := rfl + lemma mk_eq {f g : padic_seq p} : mk f = mk g ↔ f ≈ g := quotient.eq lemma const_equiv {q r : ℚ} : const (padic_norm p) q ≈ const (padic_norm p) r ↔ q = r := @@ -469,8 +471,19 @@ end padic /-- The rational-valued `p`-adic norm on `ℚ_[p]` is lifted from the norm on Cauchy sequences. The canonical form of this function is the normed space instance, with notation `∥ ∥`. -/ -def padic_norm_e {p : ℕ} [hp : fact p.prime] : ℚ_[p] → ℚ := -quotient.lift padic_seq.norm $ @padic_seq.norm_equiv _ _ +def padic_norm_e {p : ℕ} [hp : fact p.prime] : absolute_value ℚ_[p] ℚ := +{ to_fun := quotient.lift padic_seq.norm $ @padic_seq.norm_equiv _ _, + map_mul' := λ q r, quotient.induction_on₂ q r $ padic_seq.norm_mul, + nonneg' := λ q, quotient.induction_on q $ padic_seq.norm_nonneg, + eq_zero' := λ q, quotient.induction_on q $ + by simpa only [padic.zero_def, quotient.eq] using padic_seq.norm_zero_iff, + add_le' := λ q r, begin + transitivity max ((quotient.lift padic_seq.norm $ @padic_seq.norm_equiv _ _) q) + ((quotient.lift padic_seq.norm $ @padic_seq.norm_equiv _ _) r), + exact (quotient.induction_on₂ q r $ padic_seq.norm_nonarchimedean), + refine max_le_add_of_nonneg (quotient.induction_on q $ padic_seq.norm_nonneg) _, + exact (quotient.induction_on r $ padic_seq.norm_nonneg) + end } namespace padic_norm_e section embedding @@ -479,6 +492,7 @@ variables {p : ℕ} [fact p.prime] lemma defn (f : padic_seq p) {ε : ℚ} (hε : 0 < ε) : ∃ N, ∀ i ≥ N, padic_norm_e (⟦f⟧ - f i) < ε := begin + dsimp [padic_norm_e], change ∃ N, ∀ i ≥ N, (f - const _ (f i)).norm < ε, by_contra' h, cases cauchy₂ f hε with N hN, @@ -494,22 +508,6 @@ begin exact hN _ le_rfl _ hi } end -protected lemma nonneg (q : ℚ_[p]) : 0 ≤ padic_norm_e q := quotient.induction_on q norm_nonneg - -lemma zero_def : (0 : ℚ_[p]) = ⟦0⟧ := rfl - -lemma zero_iff (q : ℚ_[p]) : padic_norm_e q = 0 ↔ q = 0 := -quotient.induction_on q $ by simpa only [zero_def, quotient.eq] using norm_zero_iff - -@[simp] protected lemma zero : padic_norm_e (0 : ℚ_[p]) = 0 := (zero_iff _).2 rfl - -/-- Theorems about `padic_norm_e` are named with a `'` so the names do not conflict with the -equivalent theorems about `norm` (`∥ ∥`). -/ -@[simp] protected lemma one' : padic_norm_e (1 : ℚ_[p]) = 1 := norm_one - -@[simp] protected lemma neg (q : ℚ_[p]) : padic_norm_e (-q) = padic_norm_e q := -quotient.induction_on q $ norm_neg - /-- Theorems about `padic_norm_e` are named with a `'` so the names do not conflict with the equivalent theorems about `norm` (`∥ ∥`). -/ theorem nonarchimedean' (q r : ℚ_[p]) : @@ -522,28 +520,6 @@ theorem add_eq_max_of_ne' {q r : ℚ_[p]} : padic_norm_e q ≠ padic_norm_e r → padic_norm_e (q + r) = max (padic_norm_e q) (padic_norm_e r) := quotient.induction_on₂ q r $ λ _ _, padic_seq.add_eq_max_of_ne -lemma triangle_ineq (x y z : ℚ_[p]) : - padic_norm_e (x - z) ≤ padic_norm_e (x - y) + padic_norm_e (y - z) := -calc padic_norm_e (x - z) = padic_norm_e ((x - y) + (y - z)) : by rw sub_add_sub_cancel - ... ≤ max (padic_norm_e (x - y)) (padic_norm_e (y - z)) : padic_norm_e.nonarchimedean' _ _ - ... ≤ padic_norm_e (x - y) + padic_norm_e (y - z) : - max_le_add_of_nonneg (padic_norm_e.nonneg _) (padic_norm_e.nonneg _) - -protected lemma add (q r : ℚ_[p]) : padic_norm_e (q + r) ≤ (padic_norm_e q) + (padic_norm_e r) := -calc - padic_norm_e (q + r) ≤ max (padic_norm_e q) (padic_norm_e r) : nonarchimedean' _ _ - ... ≤ (padic_norm_e q) + (padic_norm_e r) : - max_le_add_of_nonneg (padic_norm_e.nonneg _) (padic_norm_e.nonneg _) - -protected lemma mul' (q r : ℚ_[p]) : padic_norm_e (q * r) = (padic_norm_e q) * (padic_norm_e r) := -quotient.induction_on₂ q r $ norm_mul - -instance : is_absolute_value (@padic_norm_e p _) := -{ abv_nonneg := padic_norm_e.nonneg, - abv_eq_zero := zero_iff, - abv_add := padic_norm_e.add, - abv_mul := padic_norm_e.mul' } - @[simp] lemma eq_padic_norm' (q : ℚ) : padic_norm_e (q : ℚ_[p]) = padic_norm p q := norm_const _ @@ -552,9 +528,6 @@ quotient.induction_on q $ λ f hf, have ¬ f ≈ 0, from (ne_zero_iff_nequiv_zero f).1 hf, norm_values_discrete f this -lemma sub_rev (q r : ℚ_[p]) : padic_norm_e (q - r) = padic_norm_e (r - q) := -by rw ← padic_norm_e.neg; simp - end embedding end padic_norm_e @@ -572,6 +545,7 @@ quotient.induction_on q $ λ q', let ⟨N, hN⟩ := this in ⟨q' N, begin + dsimp [padic_norm_e], change padic_seq.norm (q' - const _ (q' N)) < ε, cases decidable.em ((q' - const (padic_norm p) (q' N)) ≈ 0) with heq hne', { simpa only [heq, padic_seq.norm, dif_pos] }, @@ -620,7 +594,7 @@ begin rw [← padic_norm_e.eq_padic_norm'], exact_mod_cast this }, { apply lt_of_le_of_lt, - { apply padic_norm_e.add }, + { apply padic_norm_e.add_le }, { have : (3 : ℚ) ≠ 0, by norm_num, have : ε = ε / 3 + ε / 3 + ε / 3, { field_simp [this], simp only [bit0, bit1, mul_add, mul_one] }, @@ -629,9 +603,9 @@ begin { suffices : padic_norm_e ((lim_seq f j - f j) + (f j - f (max N N2))) < ε / 3 + ε / 3, by simpa only [sub_add_sub_cancel], apply lt_of_le_of_lt, - { apply padic_norm_e.add }, + { apply padic_norm_e.add_le }, { apply add_lt_add, - { rw [padic_norm_e.sub_rev], + { rw [padic_norm_e.map_sub], apply_mod_cast hN, exact le_of_max_le_left hj }, { exact hN2 _ (le_of_max_le_right hj) _ (le_max_right _ _) } } }, @@ -654,12 +628,12 @@ theorem complete' : ∃ q : ℚ_[p], ∀ ε > 0, ∃ N, ∀ i ≥ N, padic_norm_ suffices : padic_norm_e ((lim f - lim' f i) + (lim' f i - f i)) < ε, { ring_nf at this; exact this }, { apply lt_of_le_of_lt, - { apply padic_norm_e.add }, + { apply padic_norm_e.add_le }, { have : ε = ε / 2 + ε / 2, by rw ← add_self_div_two ε; simp, rw this, apply add_lt_add, { apply hN2, exact le_of_max_le_right hi }, - { rw_mod_cast [padic_norm_e.sub_rev], + { rw_mod_cast [padic_norm_e.map_sub], apply hN, exact le_of_max_le_left hi }}} end ⟩ @@ -674,17 +648,17 @@ instance : has_dist ℚ_[p] := ⟨λ x y, padic_norm_e (x - y)⟩ instance : metric_space ℚ_[p] := { dist_self := by simp [dist], dist := dist, - dist_comm := λ x y, by unfold dist; rw ← padic_norm_e.neg (x - y); simp, - dist_triangle := - begin - intros, unfold dist, - exact_mod_cast padic_norm_e.triangle_ineq _ _ _, - end, + dist_comm := λ x y, by simp [dist, ←padic_norm_e.map_neg (x - y)], + dist_triangle := λ x y z, + begin + unfold dist, + exact_mod_cast padic_norm_e.sub_le _ _ _, + end, eq_of_dist_eq_zero := begin unfold dist, intros _ _ h, apply eq_of_sub_eq_zero, - apply (padic_norm_e.zero_iff _).1, + apply padic_norm_e.eq_zero.1, exact_mod_cast h end } @@ -692,14 +666,14 @@ instance : has_norm ℚ_[p] := ⟨λ x, padic_norm_e x⟩ instance : normed_field ℚ_[p] := { dist_eq := λ _ _, rfl, - norm_mul' := by simp [has_norm.norm, padic_norm_e.mul'], + norm_mul' := by simp [has_norm.norm, map_mul], norm := norm, .. padic.field, .. padic.metric_space p } instance is_absolute_value : is_absolute_value (λ a : ℚ_[p], ∥a∥) := { abv_nonneg := norm_nonneg, abv_eq_zero := λ _, norm_eq_zero, abv_add := norm_add_le, - abv_mul := by simp [has_norm.norm, padic_norm_e.mul'] } + abv_mul := by simp [has_norm.norm, map_mul] } theorem rat_dense (q : ℚ_[p]) {ε : ℝ} (hε : 0 < ε) : ∃ r : ℚ, ∥q - r∥ < ε := let ⟨ε', hε'l, hε'r⟩ := exists_rat_btwn hε, @@ -715,7 +689,7 @@ variables {p : ℕ} [hp : fact p.prime] include hp @[simp] protected lemma mul (q r : ℚ_[p]) : ∥q * r∥ = ∥q∥ * ∥r∥ := -by simp [has_norm.norm, padic_norm_e.mul'] +by simp [has_norm.norm, map_mul] protected lemma is_norm (q : ℚ_[p]) : ↑(padic_norm_e q) = ∥q∥ := rfl @@ -881,7 +855,7 @@ begin cases hq ε' hε'.1 with N hN, existsi N, intros i hi, let h := hN i hi, unfold norm, - rw_mod_cast [cau_seq.sub_apply, padic_norm_e.sub_rev], + rw_mod_cast [cau_seq.sub_apply, padic_norm_e.map_sub], refine lt_trans _ hε'.2, exact_mod_cast hN i hi end From 5c88020dbba7f4ccc22562b7bd4a2b84fd2e5057 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 4 Oct 2022 18:24:26 +0000 Subject: [PATCH 559/828] feat(data/{finsupp,pi}/lex): generalize from linear_order to partial_order (#16740) Included in #16777. Should be closed if #16777 is merged first. Co-authored-by: Junyan Xu --- src/data/finsupp/lex.lean | 11 ++++++++++- src/data/pi/lex.lean | 16 +++++++++------- 2 files changed, 19 insertions(+), 8 deletions(-) diff --git a/src/data/finsupp/lex.lean b/src/data/finsupp/lex.lean index 9178940be7cae..b5dd06e4545ca 100644 --- a/src/data/finsupp/lex.lean +++ b/src/data/finsupp/lex.lean @@ -51,6 +51,8 @@ See `finsupp.lex.linear_order` for a proof that this partial order is in fact li instance lex.partial_order [partial_order N] : partial_order (lex (α →₀ N)) := partial_order.lift (λ x, to_lex ⇑(of_lex x)) finsupp.coe_fn_injective--fun_like.coe_injective +section linear_order + variable [linear_order N] /-- Auxiliary helper to case split computably. There is no need for this to be public, as it @@ -96,8 +98,15 @@ instance lex.linear_order : linear_order (lex (α →₀ N)) := decidable_eq := by apply_instance, ..lex.partial_order } +end linear_order + +variable [partial_order N] + lemma lex.le_of_forall_le {a b : lex (α →₀ N)} (h : ∀ i, of_lex a i ≤ of_lex b i) : a ≤ b := -le_of_not_lt (λ ⟨i, hi⟩, (h i).not_lt hi.2) +le_of_lt_or_eq $ or_iff_not_imp_right.2 $ λ hne, by classical; exact + ⟨finset.min' _ (nonempty_ne_locus_iff.2 hne), + λ j hj, not_mem_ne_locus.1 (λ h, (finset.min'_le _ _ h).not_lt hj), + (h _).lt_of_ne (mem_ne_locus.1 $ finset.min'_mem _ _)⟩ lemma lex.le_of_of_lex_le {a b : lex (α →₀ N)} (h : of_lex a ≤ of_lex b) : a ≤ b := lex.le_of_forall_le h diff --git a/src/data/pi/lex.lean b/src/data/pi/lex.lean index ebdda5aaa60c3..46b225975a289 100644 --- a/src/data/pi/lex.lean +++ b/src/data/pi/lex.lean @@ -86,29 +86,31 @@ noncomputable instance [linear_order ι] [is_well_order ι (<)] [∀ a, linear_o @linear_order_of_STO (Πₗ i, β i) (<) { to_is_trichotomous := is_trichotomous_lex _ _ is_well_founded.wf } (classical.dec_rel _) -lemma lex.le_of_forall_le [linear_order ι] [is_well_order ι (<)] [Π a, linear_order (β a)] +lemma lex.le_of_forall_le [linear_order ι] [hwf : is_well_order ι (<)] [Π a, partial_order (β a)] {a b : lex (Π i, β i)} (h : ∀ i, a i ≤ b i) : a ≤ b := -le_of_not_lt (λ ⟨i, hi⟩, (h i).not_lt hi.2) +or_iff_not_imp_left.2 $ λ hne, + let ⟨i, hi, hl⟩ := hwf.to_is_well_founded.wf.has_min {i | a i ≠ b i} (function.ne_iff.1 hne) in + ⟨i, λ j hj, by { contrapose! hl, exact ⟨j, hl, hj⟩ }, (h i).lt_of_ne hi⟩ -lemma lex.le_of_of_lex_le [linear_order ι] [is_well_order ι (<)] [Π a, linear_order (β a)] +lemma lex.le_of_of_lex_le [linear_order ι] [is_well_order ι (<)] [Π a, partial_order (β a)] {a b : lex (Π i, β i)} (h : of_lex a ≤ of_lex b) : a ≤ b := lex.le_of_forall_le h -lemma to_lex_monotone [linear_order ι] [is_well_order ι (<)] [Π a, linear_order (β a)] : +lemma to_lex_monotone [linear_order ι] [is_well_order ι (<)] [Π a, partial_order (β a)] : monotone (@to_lex (Π i, β i)) := λ _ _, lex.le_of_forall_le -instance [linear_order ι] [is_well_order ι (<)] [Π a, linear_order (β a)] +instance [linear_order ι] [is_well_order ι (<)] [Π a, partial_order (β a)] [Π a, order_bot (β a)] : order_bot (lex (Π a, β a)) := { bot := to_lex ⊥, bot_le := λ f, lex.le_of_of_lex_le bot_le } -instance [linear_order ι] [is_well_order ι (<)] [Π a, linear_order (β a)] +instance [linear_order ι] [is_well_order ι (<)] [Π a, partial_order (β a)] [Π a, order_top (β a)] : order_top (lex (Π a, β a)) := { top := to_lex ⊤, le_top := λ f, lex.le_of_of_lex_le le_top } -instance [linear_order ι] [is_well_order ι (<)] [Π a, linear_order (β a)] +instance [linear_order ι] [is_well_order ι (<)] [Π a, partial_order (β a)] [Π a, bounded_order (β a)] : bounded_order (lex (Π a, β a)) := { .. pi.lex.order_bot, .. pi.lex.order_top } From 8f40883edaf337ff63a9d8da80f463af4aa42a94 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 4 Oct 2022 18:24:27 +0000 Subject: [PATCH 560/828] chore(topology/algebra/polynomial): golf `polynomial.coeff_le_of_roots_le` (#16789) ... and generalize from [field F] to [comm_ring F]. --- src/topology/algebra/polynomial.lean | 73 ++++++++++------------------ 1 file changed, 25 insertions(+), 48 deletions(-) diff --git a/src/topology/algebra/polynomial.lean b/src/topology/algebra/polynomial.lean index bccc53ba4c7c8..f6683a87c15d5 100644 --- a/src/topology/algebra/polynomial.lean +++ b/src/topology/algebra/polynomial.lean @@ -133,10 +133,9 @@ else ⟨p.coeff 0, by rw [eq_C_of_degree_le_zero (le_of_not_gt hp0)]; simp⟩ section roots -open_locale polynomial -open_locale nnreal +open_locale polynomial nnreal -variables {F K : Type*} [field F] [normed_field K] +variables {F K : Type*} [comm_ring F] [normed_field K] open multiset @@ -145,8 +144,8 @@ lemma eq_one_of_roots_le {p : F[X]} {f : F →+* K} {B : ℝ} (hB : B < 0) p = 1 := h1.nat_degree_eq_zero_iff_eq_one.mp begin contrapose !hB, - rw nat_degree_eq_card_roots h2 at hB, - obtain ⟨z, hz⟩ := multiset.card_pos_iff_exists_mem.mp (zero_lt_iff.mpr hB), + rw [← h1.nat_degree_map f, nat_degree_eq_card_roots' h2] at hB, + obtain ⟨z, hz⟩ := card_pos_iff_exists_mem.mp (zero_lt_iff.mpr hB), exact le_trans (norm_nonneg _) (h3 z hz), end @@ -154,49 +153,27 @@ lemma coeff_le_of_roots_le {p : F[X]} {f : F →+* K} {B : ℝ} (i : ℕ) (h1 : p.monic) (h2 : splits f p) (h3 : ∀ z ∈ (map f p).roots, ∥z∥ ≤ B) : ∥ (map f p).coeff i ∥ ≤ B^(p.nat_degree - i) * p.nat_degree.choose i := begin - have hcd : card (map f p).roots = p.nat_degree := (nat_degree_eq_card_roots h2).symm, - obtain hB | hB := le_or_lt 0 B, - { by_cases hi : i ≤ p.nat_degree, - { rw [eq_prod_roots_of_splits h2, monic.def.mp h1, ring_hom.map_one, ring_hom.map_one, one_mul], - rw prod_X_sub_C_coeff, - swap, rwa hcd, - rw [norm_mul, (by norm_num : ∥(-1 : K) ^ (card (map f p).roots - i)∥= 1), one_mul], - apply le_trans (le_sum_of_subadditive norm _ _ _ ), - rotate, exact norm_zero, exact norm_add_le, - rw multiset.map_map, - suffices : ∀ r ∈ multiset.map (norm_hom ∘ prod) - (powerset_len (card (map f p).roots - i) (map f p).roots), r ≤ B^(p.nat_degree - i), - { convert sum_le_sum_map _ this, - simp only [hi, hcd, multiset.map_const, card_map, card_powerset_len, nat.choose_symm, - sum_repeat, nsmul_eq_mul, mul_comm], }, - { intros r hr, - obtain ⟨t, ht⟩ := multiset.mem_map.mp hr, - have hbounds : ∀ x ∈ (multiset.map norm_hom t), 0 ≤ x ∧ x ≤ B, - { intros x hx, - obtain ⟨z, hz⟩ := multiset.mem_map.mp hx, - rw ←hz.right, - exact ⟨norm_nonneg z, - h3 z (mem_of_le (mem_powerset_len.mp ht.left).left hz.left)⟩, }, - lift B to ℝ≥0 using hB, - lift (multiset.map norm_hom t) to (multiset ℝ≥0) using (λ x hx, (hbounds x hx).left) - with normt hn, - rw (by rw_mod_cast [←ht.right, function.comp_apply, ←prod_hom t norm_hom, ←hn] : - r = normt.prod), - convert multiset.prod_le_pow_card normt _ _, - { rw (_ : card normt = card (multiset.map coe normt)), - rwa [hn, ←hcd, card_map, (mem_powerset_len.mp ht.left).right.symm], - rwa card_map, }, - { intros x hx, - have xmem : (x : ℝ) ∈ multiset.map coe normt := mem_map_of_mem _ hx, - exact (hbounds x xmem).right, }}}, - { push_neg at hi, - rw [nat.choose_eq_zero_of_lt hi, coeff_eq_zero_of_nat_degree_lt, norm_zero], - rw_mod_cast mul_zero, - { rwa monic.nat_degree_map h1, - apply_instance, }}}, - { rw [eq_one_of_roots_le hB h1 h2 h3, polynomial.map_one, nat_degree_one, zero_tsub, pow_zero, - one_mul, coeff_one], - split_ifs; norm_num [h], }, + obtain hB | hB := lt_or_le B 0, + { rw [eq_one_of_roots_le hB h1 h2 h3, polynomial.map_one, + nat_degree_one, zero_tsub, pow_zero, one_mul, coeff_one], + split_ifs; norm_num [h] }, + rw ← h1.nat_degree_map f, + obtain hi | hi := lt_or_le (map f p).nat_degree i, + { rw [coeff_eq_zero_of_nat_degree_lt hi, norm_zero], positivity }, + rw [coeff_eq_esymm_roots_of_splits ((splits_id_iff_splits f).2 h2) hi, + (h1.map _).leading_coeff, one_mul, norm_mul, norm_pow, norm_neg, norm_one, one_pow, one_mul], + apply ((norm_multiset_sum_le _).trans $ sum_le_card_nsmul _ _ $ λ r hr, _).trans, + { rw [multiset.map_map, card_map, card_powerset_len, + ←nat_degree_eq_card_roots' h2, nat.choose_symm hi, mul_comm, nsmul_eq_mul] }, + simp_rw multiset.mem_map at hr, + obtain ⟨_, ⟨s, hs, rfl⟩, rfl⟩ := hr, + rw mem_powerset_len at hs, + lift B to ℝ≥0 using hB, + rw [←coe_nnnorm, ←nnreal.coe_pow, nnreal.coe_le_coe, + ←nnnorm_hom_apply, ←monoid_hom.coe_coe, monoid_hom.map_multiset_prod], + refine (prod_le_pow_card _ B $ λ x hx, _).trans_eq (by rw [card_map, hs.2]), + obtain ⟨z, hz, rfl⟩ := multiset.mem_map.1 hx, + exact h3 z (mem_of_le hs.1 hz), end /-- The coefficients of the monic polynomials of bounded degree with bounded roots are From 3946a60230b888dab6b586f882b54319cdee28bc Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Tue, 4 Oct 2022 21:21:05 +0000 Subject: [PATCH 561/828] =?UTF-8?q?fix(docs/references.bib):=20fix=20Fr?= =?UTF-8?q?=C3=B6hlich's=20name=20in=20the=20references=20(#16802)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Correct the BibTeX entry for the book by Cassels and Frohlich, copy-pasting it from MathSciNet. --- docs/references.bib | 14 ++++++++++---- 1 file changed, 10 insertions(+), 4 deletions(-) diff --git a/docs/references.bib b/docs/references.bib index fffe6800f3a4e..a99aa70dbb716 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -429,11 +429,17 @@ @InProceedings{ carneiro2019 } @Book{ cassels1967algebraic, - title = {Algebraic number theory: proceedings of an instructional - conference}, - author = {Cassels, John William Scott and Fr{\"o}lich, Albrecht}, + title = {Algebraic number theory}, + booktitle = {Proceedings of an instructional conference organized by + the {L}ondon {M}athematical {S}ociety (a {NATO} {A}dvanced + {S}tudy {I}nstitute) with the support of the + {I}nternational {M}athematical {U}nion}, + editor = {Cassels, John William Scott and Fr\"{o}hlich, Albrecht}, + publisher = {Academic Press, London; Thompson Book Co., Inc., + Washington, D.C.}, year = {1967}, - publisher = {Academic Pr} + pages = {xviii+366}, + mrclass = {00.04 (10.00)} } @InProceedings{ Chou1994, From e36ae183078af8f644dac9e55448f5d23bdf5ee4 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Tue, 4 Oct 2022 21:21:06 +0000 Subject: [PATCH 562/828] feat(data/set/prod): `set.off_diag` (#16803) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the off-diagonal of a set, which is the set of pairs of points whose components are distinct. Co-authored-by: Yaël Dillies Co-authored-by: Kyle Miller --- src/data/set/prod.lean | 54 ++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 54 insertions(+) diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index 06acc5f6fbb53..6e57bbce93d5a 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -16,6 +16,7 @@ type. * `set.prod`: Binary product of sets. For `s : set α`, `t : set β`, we have `s.prod t : set (α × β)`. * `set.diagonal`: Diagonal of a type. `set.diagonal α = {(x, x) | x : α}`. +* `set.off_diag`: Off-diagonal. `s ×ˢ s` without the diagonal. * `set.pi`: Arbitrary product of sets. -/ @@ -332,6 +333,59 @@ lemma diag_preimage_prod_self (s : set α) : (λ x, (x, x)) ⁻¹' (s ×ˢ s) = end diagonal +section off_diag +variables {α : Type*} {s t : set α} {x : α × α} {a : α} + +/-- The off-diagonal of a set `s` is the set of pairs `(a, b)` with `a, b ∈ s` and `a ≠ b`. -/ +def off_diag (s : set α) : set (α × α) := {x | x.1 ∈ s ∧ x.2 ∈ s ∧ x.1 ≠ x.2} + +@[simp] lemma mem_off_diag : x ∈ s.off_diag ↔ x.1 ∈ s ∧ x.2 ∈ s ∧ x.1 ≠ x.2 := iff.rfl + +lemma off_diag_mono : monotone (off_diag : set α → set (α × α)) := +λ s t h x, and.imp (@h _) $ and.imp_left $ @h _ + +@[simp] lemma off_diag_nonempty : s.off_diag.nonempty ↔ s.nontrivial := +by simp [off_diag, set.nonempty, set.nontrivial] + +@[simp] lemma off_diag_eq_empty : s.off_diag = ∅ ↔ s.subsingleton := +by rw [←not_nonempty_iff_eq_empty, ←not_nontrivial_iff, off_diag_nonempty.not] + +alias off_diag_nonempty ↔ _ nontrivial.off_diag_nonempty +alias off_diag_nonempty ↔ _ subsingleton.off_diag_eq_empty + +variables (s t) + +lemma off_diag_subset_prod : s.off_diag ⊆ s ×ˢ s := λ x hx, ⟨hx.1, hx.2.1⟩ +lemma off_diag_eq_sep_prod : s.off_diag = {x ∈ s ×ˢ s | x.1 ≠ x.2} := ext $ λ _, and.assoc.symm + +@[simp] lemma off_diag_empty : (∅ : set α).off_diag = ∅ := by simp +@[simp] lemma off_diag_singleton (a : α) : ({a} : set α).off_diag = ∅ := by simp +@[simp] lemma off_diag_univ : (univ : set α).off_diag = (diagonal α)ᶜ := ext $ by simp + +@[simp] lemma prod_sdiff_diagonal : s ×ˢ s \ diagonal α = s.off_diag := ext $ λ _, and.assoc +@[simp] lemma disjoint_diagonal_off_diag : disjoint (diagonal α) s.off_diag := λ x hx, hx.2.2.2 hx.1 + +lemma off_diag_inter : (s ∩ t).off_diag = s.off_diag ∩ t.off_diag := +ext $ λ x, by { simp only [mem_off_diag, mem_inter_iff], tauto } + +variables {s t} + +lemma off_diag_union (h : disjoint s t) : + (s ∪ t).off_diag = s.off_diag ∪ t.off_diag ∪ s ×ˢ t ∪ t ×ˢ s := +begin + rw [off_diag_eq_sep_prod, union_prod, prod_union, prod_union, union_comm _ (t ×ˢ t), union_assoc, + union_left_comm (s ×ˢ t), ←union_assoc, sep_union, sep_union, ←off_diag_eq_sep_prod, + ←off_diag_eq_sep_prod, sep_eq_self_iff_mem_true.2, ←union_assoc], + simp only [mem_union, mem_prod, ne.def, prod.forall], + rintro i j (⟨hi, hj⟩ | ⟨hi, hj⟩) rfl; exact h ⟨‹_›, ‹_›⟩, +end + +lemma off_diag_insert (ha : a ∉ s) : (insert a s).off_diag = s.off_diag ∪ {a} ×ˢ s ∪ s ×ˢ {a} := +by { rw [insert_eq, union_comm, off_diag_union, off_diag_singleton, union_empty, union_right_comm], + rintro b ⟨hb, (rfl : b = a)⟩, exact ha hb } + +end off_diag + /-! ### Cartesian set-indexed product of sets -/ section pi From 8f83c26d7bef24e056fa7b60199a88fc1924116c Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 5 Oct 2022 01:42:24 +0000 Subject: [PATCH 563/828] feat(ring_theory/ideal/minimal_prime): Minimal prime ideals over an ideal (#16136) --- src/ring_theory/ideal/basic.lean | 14 ++ src/ring_theory/ideal/minimal_prime.lean | 217 +++++++++++++++++++++++ 2 files changed, 231 insertions(+) create mode 100644 src/ring_theory/ideal/minimal_prime.lean diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index 31cc795eb9cf9..af7b52c4b160c 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -306,6 +306,20 @@ lemma mem_pi (x : ι → α) : x ∈ I.pi ι ↔ ∀ i, x i ∈ I := iff.rfl end pi +lemma Inf_is_prime_of_is_chain {s : set (ideal α)} (hs : s.nonempty) (hs' : is_chain (≤) s) + (H : ∀ p ∈ s, ideal.is_prime p) : + (Inf s).is_prime := +⟨λ e, let ⟨x, hx⟩ := hs in (H x hx).ne_top (eq_top_iff.mpr (e.symm.trans_le (Inf_le hx))), + λ x y e, or_iff_not_imp_left.mpr $ λ hx, begin + rw ideal.mem_Inf at hx ⊢ e, + push_neg at hx, + obtain ⟨I, hI, hI'⟩ := hx, + intros J hJ, + cases hs'.total hI hJ, + { exact h (((H I hI).mem_or_mem (e hI)).resolve_left hI') }, + { exact ((H J hJ).mem_or_mem (e hJ)).resolve_left (λ x, hI' $ h x) }, + end⟩ + end ideal end semiring diff --git a/src/ring_theory/ideal/minimal_prime.lean b/src/ring_theory/ideal/minimal_prime.lean new file mode 100644 index 0000000000000..397f2e12c22b3 --- /dev/null +++ b/src/ring_theory/ideal/minimal_prime.lean @@ -0,0 +1,217 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ + +import ring_theory.localization.integral +import ring_theory.localization.at_prime +import order.minimal + +/-! + +# Minimal primes + +We provide various results concerning the minimal primes above an ideal + +## Main results +- `ideal.minimal_primes`: `I.minimal_primes` is the set of ideals that are minimal primes over `I`. +- `minimal_primes`: `minimal_primes R` is the set of minimal primes of `R`. +- `ideal.exists_minimal_primes_le`: Every prime ideal over `I` contains a minimal prime over `I`. +- `ideal.radical_minimal_primes`: The minimal primes over `I.radical` are precisely + the minimal primes over `I`. +- `ideal.Inf_minimal_primes`: The intersection of minimal primes over `I` is `I.radical`. +- `ideal.exists_minimal_primes_comap_eq` If `p` is a minimal prime over `f ⁻¹ I`, then it is the + preimage of some minimal prime over `I`. +- `ideal.minimal_primes_eq_comap`: The minimal primes over `I` are precisely the preimages of + minimal primes of `R ⧸ I`. + + +-/ + + +section + +variables {R S : Type*} [comm_ring R] [comm_ring S] (I J : ideal R) + +/-- `I.minimal_primes` is the set of ideals that are minimal primes over `I`. -/ +def ideal.minimal_primes : set (ideal R) := +minimals (≤) { p | p.is_prime ∧ I ≤ p } + +/-- `minimal_primes R` is the set of minimal primes of `R`. +This is defined as `ideal.minimal_primes ⊥`. -/ +def minimal_primes (R : Type*) [comm_ring R] : set (ideal R) := ideal.minimal_primes ⊥ + +variables {I J} + +lemma ideal.exists_minimal_primes_le [J.is_prime] (e : I ≤ J) : + ∃ p ∈ I.minimal_primes, p ≤ J := +begin + suffices : ∃ m ∈ { p : (ideal R)ᵒᵈ | ideal.is_prime p ∧ I ≤ order_dual.of_dual p }, + (order_dual.to_dual J) ≤ m ∧ + ∀ z ∈ { p : (ideal R)ᵒᵈ | ideal.is_prime p ∧ I ≤ p }, m ≤ z → z = m, + { obtain ⟨p, h₁, h₂, h₃⟩ := this, + simp_rw ← @eq_comm _ p at h₃, + exact ⟨p, ⟨h₁, λ a b c, (h₃ a b c).le⟩, h₂⟩ }, + apply zorn_nonempty_partial_order₀, + swap, { refine ⟨show J.is_prime, by apply_instance, e⟩ }, + rintros (c : set (ideal R)) hc hc' J' hJ', + refine ⟨order_dual.to_dual (Inf c), + ⟨ideal.Inf_is_prime_of_is_chain ⟨J', hJ'⟩ hc'.symm (λ x hx, (hc hx).1), _⟩, _⟩, + { rw order_dual.of_dual_to_dual, convert le_Inf _, intros x hx, exact (hc hx).2 }, + { rintro z hz, + rw order_dual.le_to_dual, + exact Inf_le hz } +end + +@[simp] +lemma ideal.radical_minimal_primes : I.radical.minimal_primes = I.minimal_primes := +begin + rw [ideal.minimal_primes, ideal.minimal_primes], + congr, + ext p, + exact ⟨λ ⟨a, b⟩, ⟨a, ideal.le_radical.trans b⟩, λ ⟨a, b⟩, ⟨a, a.radical_le_iff.mpr b⟩⟩, +end + +@[simp] +lemma ideal.Inf_minimal_primes : + Inf I.minimal_primes = I.radical := +begin + rw I.radical_eq_Inf, + apply le_antisymm, + { intros x hx, + rw ideal.mem_Inf at hx ⊢, + rintros J ⟨e, hJ⟩, + resetI, + obtain ⟨p, hp, hp'⟩ := ideal.exists_minimal_primes_le e, + exact hp' (hx hp) }, + { apply Inf_le_Inf _, + intros I hI, + exact hI.1.symm }, +end + +lemma ideal.exists_comap_eq_of_mem_minimal_primes_of_injective {f : R →+* S} + (hf : function.injective f) (p ∈ minimal_primes R) : + ∃ p' : ideal S, p'.is_prime ∧ p'.comap f = p := +begin + haveI := H.1.1, + haveI : nontrivial (localization (submonoid.map f p.prime_compl)), + { refine ⟨⟨1, 0, _⟩⟩, + convert (is_localization.map_injective_of_injective p.prime_compl (localization.at_prime p) + (localization $ p.prime_compl.map f) hf).ne one_ne_zero, + { rw map_one }, { rw map_zero } }, + obtain ⟨M, hM⟩ := ideal.exists_maximal (localization (submonoid.map f p.prime_compl)), + resetI, + refine ⟨M.comap (algebra_map S $ localization (submonoid.map f p.prime_compl)), + infer_instance, _⟩, + rw [ideal.comap_comap, ← @@is_localization.map_comp _ _ _ _ localization.is_localization _ + p.prime_compl.le_comap_map _ localization.is_localization, ← ideal.comap_comap], + suffices : _ ≤ p, + { exact this.antisymm (H.2 ⟨infer_instance, bot_le⟩ this) }, + intros x hx, + by_contra h, + apply hM.ne_top, + apply M.eq_top_of_is_unit_mem hx, + apply is_unit.map, + apply is_localization.map_units _ (show p.prime_compl, from ⟨x, h⟩), + apply_instance +end + +lemma ideal.exists_comap_eq_of_mem_minimal_primes {I : ideal S} + (f : R →+* S) (p ∈ (I.comap f).minimal_primes) : + ∃ p' : ideal S, p'.is_prime ∧ I ≤ p' ∧ p'.comap f = p := +begin + haveI := H.1.1, + let f' := I^.quotient.mk^.comp f, + have e : (I^.quotient.mk^.comp f).ker = I.comap f, + { ext1, exact (submodule.quotient.mk_eq_zero _) }, + have : (I^.quotient.mk^.comp f).ker^.quotient.mk^.ker ≤ p, + { rw [ideal.mk_ker, e], exact H.1.2 }, + obtain ⟨p', hp₁, hp₂⟩ := ideal.exists_comap_eq_of_mem_minimal_primes_of_injective + (I^.quotient.mk^.comp f).ker_lift_injective (p.map (I^.quotient.mk^.comp f).ker^.quotient.mk) _, + { resetI, + refine ⟨p'.comap I^.quotient.mk, ideal.is_prime.comap _, _, _⟩, + { exact ideal.mk_ker.symm.trans_le (ideal.comap_mono bot_le) }, + convert congr_arg (ideal.comap (I^.quotient.mk^.comp f).ker^.quotient.mk) hp₂, + rwa [ideal.comap_map_of_surjective (I^.quotient.mk^.comp f).ker^.quotient.mk + ideal.quotient.mk_surjective, eq_comm, sup_eq_left] }, + refine ⟨⟨_, bot_le⟩, _⟩, + { apply ideal.map_is_prime_of_surjective _ this, exact ideal.quotient.mk_surjective }, + { rintro q ⟨hq, -⟩ hq', + rw ← ideal.map_comap_of_surjective (I^.quotient.mk^.comp f).ker^.quotient.mk + ideal.quotient.mk_surjective q, + apply ideal.map_mono, + resetI, + apply H.2, + { refine ⟨infer_instance, (ideal.mk_ker.trans e).symm.trans_le (ideal.comap_mono bot_le)⟩ }, + { refine (ideal.comap_mono hq').trans _, rw ideal.comap_map_of_surjective, + exacts [sup_le rfl.le this, ideal.quotient.mk_surjective] } } +end + +lemma ideal.exists_minimal_primes_comap_eq {I : ideal S} + (f : R →+* S) (p ∈ (I.comap f).minimal_primes) : + ∃ p' ∈ I.minimal_primes, ideal.comap f p' = p := +begin + obtain ⟨p', h₁, h₂, h₃⟩ := ideal.exists_comap_eq_of_mem_minimal_primes f p H, + resetI, + obtain ⟨q, hq, hq'⟩ := ideal.exists_minimal_primes_le h₂, + refine ⟨q, hq, eq.symm _⟩, + haveI := hq.1.1, + have := (ideal.comap_mono hq').trans_eq h₃, + exact (H.2 ⟨infer_instance, ideal.comap_mono hq.1.2⟩ this).antisymm this +end + +lemma ideal.mimimal_primes_comap_of_surjective {f : R →+* S} (hf : function.surjective f) + {I J : ideal S} (h : J ∈ I.minimal_primes) : + J.comap f ∈ (I.comap f).minimal_primes := +begin + haveI := h.1.1, + refine ⟨⟨infer_instance, ideal.comap_mono h.1.2⟩, _⟩, + rintros K ⟨hK, e₁⟩ e₂, + have : f.ker ≤ K := (ideal.comap_mono bot_le).trans e₁, + rw [← sup_eq_left.mpr this, ring_hom.ker_eq_comap_bot, ← ideal.comap_map_of_surjective f hf], + apply ideal.comap_mono _, + apply h.2 _ _, + { exactI ⟨ideal.map_is_prime_of_surjective hf this, + ideal.le_map_of_comap_le_of_surjective f hf e₁⟩ }, + { exact ideal.map_le_of_le_comap e₂ } +end + +lemma ideal.comap_minimal_primes_eq_of_surjective {f : R →+* S} (hf : function.surjective f) + (I : ideal S) : + (I.comap f).minimal_primes = ideal.comap f '' I.minimal_primes := +begin + ext J, + split, + { intro H, obtain ⟨p, h, rfl⟩ := ideal.exists_minimal_primes_comap_eq f J H, exact ⟨p, h, rfl⟩ }, + { rintros ⟨J, hJ, rfl⟩, exact ideal.mimimal_primes_comap_of_surjective hf hJ } +end + +lemma ideal.minimal_primes_eq_comap : + I.minimal_primes = ideal.comap I^.quotient.mk '' minimal_primes (R ⧸ I) := +begin + rw [minimal_primes, ← ideal.comap_minimal_primes_eq_of_surjective ideal.quotient.mk_surjective, + ← ring_hom.ker_eq_comap_bot, ideal.mk_ker], +end + +lemma ideal.minimal_primes_eq_subsingleton (hI : I.is_primary) : + I.minimal_primes = {I.radical} := +begin + ext J, + split, + { exact λ H, let e := H.1.1.radical_le_iff.mpr H.1.2 in + (H.2 ⟨ideal.is_prime_radical hI, ideal.le_radical⟩ e).antisymm e }, + { rintro (rfl : J = I.radical), + exact ⟨⟨ideal.is_prime_radical hI, ideal.le_radical⟩, λ _ H _, H.1.radical_le_iff.mpr H.2⟩ } +end + +lemma ideal.minimal_primes_eq_subsingleton_self [I.is_prime] : + I.minimal_primes = {I} := +begin + ext J, + split, + { exact λ H, (H.2 ⟨infer_instance, rfl.le⟩ H.1.2).antisymm H.1.2 }, + { unfreezingI { rintro (rfl : J = I) }, refine ⟨⟨infer_instance, rfl.le⟩, λ _ h _, h.2⟩ }, +end + +end From ae17d05fbf2da5a6829f2d8776be108f55cc8eaa Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 5 Oct 2022 04:10:07 +0000 Subject: [PATCH 564/828] chore(algebra/hom/non_unital_alg): add missing `non_unital_ring_hom_class` instance (#16783) This instance was missing because `non_unital_alg_hom` was defined before `non_unital_ring_hom`. --- src/algebra/hom/non_unital_alg.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/algebra/hom/non_unital_alg.lean b/src/algebra/hom/non_unital_alg.lean index 8d9221d46a2d8..15136e3e05e8f 100644 --- a/src/algebra/hom/non_unital_alg.lean +++ b/src/algebra/hom/non_unital_alg.lean @@ -73,6 +73,14 @@ attribute [nolint dangerous_instance] non_unital_alg_hom_class.to_mul_hom_class namespace non_unital_alg_hom_class +-- `R` becomes a metavariable but that's fine because it's an `out_param` +@[priority 100, nolint dangerous_instance] -- See note [lower instance priority] +instance non_unital_alg_hom_class.to_non_unital_ring_hom_class {F R A B : Type*} [monoid R] + [non_unital_non_assoc_semiring A] [distrib_mul_action R A] + [non_unital_non_assoc_semiring B] [distrib_mul_action R B] + [non_unital_alg_hom_class F R A B] : non_unital_ring_hom_class F A B := +{ coe := coe_fn, ..‹non_unital_alg_hom_class F R A B› } + variables [semiring R] [non_unital_non_assoc_semiring A] [module R A] [non_unital_non_assoc_semiring B] [module R B] From 52a270e2ea4e342c2587c106f8be904524214a4b Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 5 Oct 2022 06:38:07 +0000 Subject: [PATCH 565/828] feat(category_theory/limits): bicartesian squares (#14375) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ``` X ⊞ Y --fst--> X | | snd 0 | | v v Y Co-authored-by: Scott Morrison --- .../limits/constructions/zero_objects.lean | 156 +++++++++ .../limits/shapes/binary_products.lean | 30 ++ .../limits/shapes/comm_sq.lean | 301 +++++++++++++++++- 3 files changed, 473 insertions(+), 14 deletions(-) create mode 100644 src/category_theory/limits/constructions/zero_objects.lean diff --git a/src/category_theory/limits/constructions/zero_objects.lean b/src/category_theory/limits/constructions/zero_objects.lean new file mode 100644 index 0000000000000..c33665cd16695 --- /dev/null +++ b/src/category_theory/limits/constructions/zero_objects.lean @@ -0,0 +1,156 @@ +/- +Copyright (c) 2022 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ +import category_theory.limits.shapes.pullbacks +import category_theory.limits.shapes.zero_morphisms +import category_theory.limits.constructions.binary_products + +/-! +# Limits involving zero objects + +Binary products and coproducts with a zero object always exist, +and pullbacks/pushouts over a zero object are products/coproducts. +-/ + +noncomputable theory + +open category_theory + +variables {C : Type*} [category C] + +namespace category_theory.limits + +variables [has_zero_object C] [has_zero_morphisms C] +open_locale zero_object + +/-- The limit cone for the product with a zero object. -/ +def binary_fan_zero_left (X : C) : binary_fan (0 : C) X := +binary_fan.mk 0 (𝟙 X) + +/-- The limit cone for the product with a zero object is limiting. -/ +def binary_fan_zero_left_is_limit (X : C) : is_limit (binary_fan_zero_left X) := +binary_fan.is_limit_mk (λ s, binary_fan.snd s) (by tidy) (by tidy) (by tidy) + +instance has_binary_product_zero_left (X : C) : has_binary_product (0 : C) X := +has_limit.mk ⟨_, binary_fan_zero_left_is_limit X⟩ + +/-- A zero object is a left unit for categorical product. -/ +def zero_prod_iso (X : C) : (0 : C) ⨯ X ≅ X := +limit.iso_limit_cone ⟨_, binary_fan_zero_left_is_limit X⟩ + +@[simp] lemma zero_prod_iso_hom (X : C) : (zero_prod_iso X).hom = prod.snd := +rfl +@[simp] lemma zero_prod_iso_inv_snd (X : C) : (zero_prod_iso X).inv ≫ prod.snd = 𝟙 X := +by { dsimp [zero_prod_iso, binary_fan_zero_left], simp, } + +/-- The limit cone for the product with a zero object. -/ +def binary_fan_zero_right (X : C) : binary_fan X (0 : C) := +binary_fan.mk (𝟙 X) 0 + +/-- The limit cone for the product with a zero object is limiting. -/ +def binary_fan_zero_right_is_limit (X : C) : is_limit (binary_fan_zero_right X) := +binary_fan.is_limit_mk (λ s, binary_fan.fst s) (by tidy) (by tidy) (by tidy) + +instance has_binary_product_zero_right (X : C) : has_binary_product X (0 : C) := +has_limit.mk ⟨_, binary_fan_zero_right_is_limit X⟩ + +/-- A zero object is a right unit for categorical product. -/ +def prod_zero_iso (X : C) : X ⨯ (0 : C) ≅ X := +limit.iso_limit_cone ⟨_, binary_fan_zero_right_is_limit X⟩ + +@[simp] lemma prod_zero_iso_hom (X : C) : (prod_zero_iso X).hom = prod.fst := +rfl +@[simp] lemma prod_zero_iso_iso_inv_snd (X : C) : (prod_zero_iso X).inv ≫ prod.fst = 𝟙 X := +by { dsimp [prod_zero_iso, binary_fan_zero_right], simp, } + +/-- The colimit cocone for the coproduct with a zero object. -/ +def binary_cofan_zero_left (X : C) : binary_cofan (0 : C) X := +binary_cofan.mk 0 (𝟙 X) + +/-- The colimit cocone for the coproduct with a zero object is colimiting. -/ +def binary_cofan_zero_left_is_colimit (X : C) : is_colimit (binary_cofan_zero_left X) := +binary_cofan.is_colimit_mk (λ s, binary_cofan.inr s) (by tidy) (by tidy) (by tidy) + +instance has_binary_coproduct_zero_left (X : C) : has_binary_coproduct (0 : C) X := +has_colimit.mk ⟨_, binary_cofan_zero_left_is_colimit X⟩ + +/-- A zero object is a left unit for categorical coproduct. -/ +def zero_coprod_iso (X : C) : (0 : C) ⨿ X ≅ X := +colimit.iso_colimit_cocone ⟨_, binary_cofan_zero_left_is_colimit X⟩ + +@[simp] lemma inr_zero_coprod_iso_hom (X : C) : coprod.inr ≫ (zero_coprod_iso X).hom = 𝟙 X := +by { dsimp [zero_coprod_iso, binary_cofan_zero_left], simp, } +@[simp] lemma zero_coprod_iso_inv (X : C) : (zero_coprod_iso X).inv = coprod.inr := +rfl + +/-- The colimit cocone for the coproduct with a zero object. -/ +def binary_cofan_zero_right (X : C) : binary_cofan X (0 : C) := +binary_cofan.mk (𝟙 X) 0 + +/-- The colimit cocone for the coproduct with a zero object is colimiting. -/ +def binary_cofan_zero_right_is_colimit (X : C) : is_colimit (binary_cofan_zero_right X) := +binary_cofan.is_colimit_mk (λ s, binary_cofan.inl s) (by tidy) (by tidy) (by tidy) + +instance has_binary_coproduct_zero_right (X : C) : has_binary_coproduct X (0 : C) := +has_colimit.mk ⟨_, binary_cofan_zero_right_is_colimit X⟩ + +/-- A zero object is a right unit for categorical coproduct. -/ +def coprod_zero_iso (X : C) : X ⨿ (0 : C) ≅ X := +colimit.iso_colimit_cocone ⟨_, binary_cofan_zero_right_is_colimit X⟩ + +@[simp] lemma inr_coprod_zeroiso_hom (X : C) : coprod.inl ≫ (coprod_zero_iso X).hom = 𝟙 X := +by { dsimp [coprod_zero_iso, binary_cofan_zero_right], simp, } +@[simp] lemma coprod_zero_iso_inv (X : C) : (coprod_zero_iso X).inv = coprod.inl := +rfl + +instance has_pullback_over_zero + (X Y : C) [has_binary_product X Y] : has_pullback (0 : X ⟶ 0) (0 : Y ⟶ 0) := +has_limit.mk ⟨_, is_pullback_of_is_terminal_is_product _ _ _ _ + has_zero_object.zero_is_terminal (prod_is_prod X Y)⟩ + +/-- The pullback over the zeron object is the product. -/ +def pullback_zero_zero_iso (X Y : C) [has_binary_product X Y] : + pullback (0 : X ⟶ 0) (0 : Y ⟶ 0) ≅ X ⨯ Y := +limit.iso_limit_cone ⟨_, is_pullback_of_is_terminal_is_product _ _ _ _ + has_zero_object.zero_is_terminal (prod_is_prod X Y)⟩ + +@[simp] lemma pullback_zero_zero_iso_inv_fst (X Y : C) [has_binary_product X Y] : + (pullback_zero_zero_iso X Y).inv ≫ pullback.fst = prod.fst := +by { dsimp [pullback_zero_zero_iso], simp, } +@[simp] lemma pullback_zero_zero_iso_inv_snd (X Y : C) [has_binary_product X Y] : + (pullback_zero_zero_iso X Y).inv ≫ pullback.snd = prod.snd := +by { dsimp [pullback_zero_zero_iso], simp, } +@[simp] lemma pullback_zero_zero_iso_hom_fst (X Y : C) [has_binary_product X Y] : + (pullback_zero_zero_iso X Y).hom ≫ prod.fst = pullback.fst := +by { simp [←iso.eq_inv_comp], } +@[simp] lemma pullback_zero_zero_iso_hom_snd (X Y : C) [has_binary_product X Y] : + (pullback_zero_zero_iso X Y).hom ≫ prod.snd = pullback.snd := +by { simp [←iso.eq_inv_comp], } + +instance has_pushout_over_zero + (X Y : C) [has_binary_coproduct X Y] : has_pushout (0 : 0 ⟶ X) (0 : 0 ⟶ Y) := +has_colimit.mk ⟨_, is_pushout_of_is_initial_is_coproduct _ _ _ _ + has_zero_object.zero_is_initial (coprod_is_coprod X Y)⟩ + +/-- The pushout over the zero object is the coproduct. -/ +def pushout_zero_zero_iso + (X Y : C) [has_binary_coproduct X Y] : pushout (0 : 0 ⟶ X) (0 : 0 ⟶ Y) ≅ X ⨿ Y := +colimit.iso_colimit_cocone ⟨_, is_pushout_of_is_initial_is_coproduct _ _ _ _ + has_zero_object.zero_is_initial (coprod_is_coprod X Y)⟩ + +@[simp] lemma inl_pushout_zero_zero_iso_hom (X Y : C) [has_binary_coproduct X Y] : + pushout.inl ≫ (pushout_zero_zero_iso X Y).hom = coprod.inl := +by { dsimp [pushout_zero_zero_iso], simp, } +@[simp] lemma inr_pushout_zero_zero_iso_hom (X Y : C) [has_binary_coproduct X Y] : + pushout.inr ≫ (pushout_zero_zero_iso X Y).hom = coprod.inr := +by { dsimp [pushout_zero_zero_iso], simp, } +@[simp] lemma inl_pushout_zero_zero_iso_inv (X Y : C) [has_binary_coproduct X Y] : + coprod.inl ≫ (pushout_zero_zero_iso X Y).inv = pushout.inl := +by { simp [iso.comp_inv_eq], } +@[simp] lemma inr_pushout_zero_zero_iso_inv (X Y : C) [has_binary_coproduct X Y] : + coprod.inr ≫ (pushout_zero_zero_iso X Y).inv = pushout.inr := +by { simp [iso.comp_inv_eq], } + +end category_theory.limits diff --git a/src/category_theory/limits/shapes/binary_products.lean b/src/category_theory/limits/shapes/binary_products.lean index 76698c5a4a482..e88b5d31b797f 100644 --- a/src/category_theory/limits/shapes/binary_products.lean +++ b/src/category_theory/limits/shapes/binary_products.lean @@ -219,6 +219,36 @@ cones.ext (iso.refl _) (λ j, by discrete_cases; cases j; tidy) def iso_binary_cofan_mk {X Y : C} (c : binary_cofan X Y) : c ≅ binary_cofan.mk c.inl c.inr := cocones.ext (iso.refl _) (λ j, by discrete_cases; cases j; tidy) +/-- +This is a more convenient formulation to show that a `binary_fan` constructed using +`binary_fan.mk` is a limit cone. +-/ +def binary_fan.is_limit_mk {W : C} {fst : W ⟶ X} {snd : W ⟶ Y} + (lift : Π (s : binary_fan X Y), s.X ⟶ W) + (fac_left : ∀ (s : binary_fan X Y), lift s ≫ fst = s.fst) + (fac_right : ∀ (s : binary_fan X Y), lift s ≫ snd = s.snd) + (uniq : ∀ (s : binary_fan X Y) (m : s.X ⟶ W) + (w_fst : m ≫ fst = s.fst) (w_snd : m ≫ snd = s.snd), m = lift s) : + is_limit (binary_fan.mk fst snd) := +{ lift := lift, + fac' := λ s j, by { rcases j with ⟨⟨⟩⟩, exacts [fac_left s, fac_right s], }, + uniq' := λ s m w, uniq s m (w ⟨walking_pair.left⟩) (w ⟨walking_pair.right⟩) } + +/-- +This is a more convenient formulation to show that a `binary_cofan` constructed using +`binary_cofan.mk` is a colimit cocone. +-/ +def binary_cofan.is_colimit_mk {W : C} {inl : X ⟶ W} {inr : Y ⟶ W} + (desc : Π (s : binary_cofan X Y), W ⟶ s.X) + (fac_left : ∀ (s : binary_cofan X Y), inl ≫ desc s = s.inl) + (fac_right : ∀ (s : binary_cofan X Y), inr ≫ desc s = s.inr) + (uniq : ∀ (s : binary_cofan X Y) (m : W ⟶ s.X) + (w_inl : inl ≫ m = s.inl) (w_inr : inr ≫ m = s.inr), m = desc s) : + is_colimit (binary_cofan.mk inl inr) := +{ desc := desc, + fac' := λ s j, by { rcases j with ⟨⟨⟩⟩, exacts [fac_left s, fac_right s], }, + uniq' := λ s m w, uniq s m (w ⟨walking_pair.left⟩) (w ⟨walking_pair.right⟩) } + /-- If `s` is a limit binary fan over `X` and `Y`, then every pair of morphisms `f : W ⟶ X` and `g : W ⟶ Y` induces a morphism `l : W ⟶ s.X` satisfying `l ≫ s.fst = f` and `l ≫ s.snd = g`. -/ diff --git a/src/category_theory/limits/shapes/comm_sq.lean b/src/category_theory/limits/shapes/comm_sq.lean index 63831ff1b8e2e..1484f31ad2389 100644 --- a/src/category_theory/limits/shapes/comm_sq.lean +++ b/src/category_theory/limits/shapes/comm_sq.lean @@ -4,13 +4,15 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Joël Riou -/ import category_theory.comm_sq +import category_theory.limits.opposites +import category_theory.limits.shapes.biproducts import category_theory.limits.preserves.shapes.pullbacks import category_theory.limits.shapes.zero_morphisms import category_theory.limits.constructions.binary_products -import category_theory.limits.opposites +import category_theory.limits.constructions.zero_objects /-! -# Pullback and pushout squares +# Pullback and pushout squares, and bicartesian squares We provide another API for pullbacks and pushouts. @@ -35,8 +37,7 @@ for the usual `pullback f g` provided by the `has_limit` API. We don't attempt to restate everything we know about pullbacks in this language, but do restate the pasting lemmas. -## Future work -Bicartesian squares, and +We define bicartesian squares, and show that the pullback and pushout squares for a biproduct are bicartesian. -/ @@ -51,6 +52,8 @@ namespace category_theory variables {C : Type u₁} [category.{v₁} C] +attribute [simp] comm_sq.mk + namespace comm_sq variables {W X Y Z : C} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} @@ -105,10 +108,9 @@ end comm_sq Y ---g---> Z ``` -is a pullback square. +is a pullback square. (Also known as a fibered product or cartesian square.) -/ -structure is_pullback - {P X Y Z : C} (fst : P ⟶ X) (snd : P ⟶ Y) (f : X ⟶ Z) (g : Y ⟶ Z) +structure is_pullback {P X Y Z : C} (fst : P ⟶ X) (snd : P ⟶ Y) (f : X ⟶ Z) (g : Y ⟶ Z) extends comm_sq fst snd f g : Prop := (is_limit' : nonempty (is_limit (pullback_cone.mk _ _ w))) @@ -122,13 +124,36 @@ structure is_pullback Y --inr--> P ``` -is a pushout square. +is a pushout square. (Also known as a fiber coproduct or cocartesian square.) -/ -structure is_pushout - {Z X Y P : C} (f : Z ⟶ X) (g : Z ⟶ Y) (inl : X ⟶ P) (inr : Y ⟶ P) +structure is_pushout {Z X Y P : C} (f : Z ⟶ X) (g : Z ⟶ Y) (inl : X ⟶ P) (inr : Y ⟶ P) extends comm_sq f g inl inr : Prop := (is_colimit' : nonempty (is_colimit (pushout_cocone.mk _ _ w))) + +section +set_option old_structure_cmd true + +/-- A *bicartesian* square is a commutative square +``` + W ---f---> X + | | + g h + | | + v v + Y ---i---> Z + +``` +that is both a pullback square and a pushout square. +-/ +structure bicartesian_sq {W X Y Z : C} (f : W ⟶ X) (g : W ⟶ Y) (h : X ⟶ Z) (i : Y ⟶ Z) + extends is_pullback f g h i, is_pushout f g h i : Prop + +-- Lean should make these parent projections as `lemma`, not `def`. +attribute [nolint def_lemma doc_blame] bicartesian_sq.to_is_pullback bicartesian_sq.to_is_pushout + +end + /-! We begin by providing some glue between `is_pullback` and the `is_limit` and `has_limit` APIs. (And similarly for `is_pushout`.) @@ -175,6 +200,11 @@ lemma of_is_product {c : binary_fan X Y} (h : limits.is_limit c) (t : is_termina of_is_limit (is_pullback_of_is_terminal_is_product _ _ _ _ t (is_limit.of_iso_limit h (limits.cones.ext (iso.refl c.X) (by rintro ⟨⟨⟩⟩; { dsimp, simp, })))) +/-- A variant of `of_is_product` that is more useful with `apply`. -/ +lemma of_is_product' (h : limits.is_limit (binary_fan.mk fst snd)) (t : is_terminal Z) : + is_pullback fst snd (t.from _) (t.from _) := +of_is_product h t + variables (X Y) lemma of_has_binary_product' [has_binary_product X Y] [has_terminal C] : @@ -264,6 +294,11 @@ of_is_colimit (is_pushout_of_is_initial_is_coproduct _ _ _ _ t (is_colimit.of_iso_colimit h (limits.cocones.ext (iso.refl c.X) (by rintro ⟨⟨⟩⟩; { dsimp, simp, })))) +/-- A variant of `of_is_coproduct` that is more useful with `apply`. -/ +lemma of_is_coproduct' (h : limits.is_colimit (binary_cofan.mk inl inr)) (t : is_initial Z) : + is_pushout (t.to _) (t.to _) inl inr := +of_is_coproduct h t + variables (X Y) lemma of_has_binary_coproduct' [has_binary_coproduct X Y] [has_initial C] : @@ -318,7 +353,7 @@ variables [has_zero_object C] [has_zero_morphisms C] open_locale zero_object /-- The square with `0 : 0 ⟶ 0` on the left and `𝟙 X` on the right is a pullback square. -/ -lemma zero_left (X : C) : is_pullback (0 : 0 ⟶ X) (0 : 0 ⟶ 0) (𝟙 X) (0 : 0 ⟶ X) := +@[simp] lemma zero_left (X : C) : is_pullback (0 : 0 ⟶ X) (0 : 0 ⟶ 0) (𝟙 X) (0 : 0 ⟶ X) := { w := by simp, is_limit' := ⟨{ lift := λ s, 0, @@ -326,9 +361,18 @@ lemma zero_left (X : C) : is_pullback (0 : 0 ⟶ X) (0 : 0 ⟶ 0) (𝟙 X) (0 : (by simpa using (pullback_cone.condition s).symm), }⟩ } /-- The square with `0 : 0 ⟶ 0` on the top and `𝟙 X` on the bottom is a pullback square. -/ -lemma zero_top (X : C) : is_pullback (0 : 0 ⟶ 0) (0 : 0 ⟶ X) (0 : 0 ⟶ X) (𝟙 X) := +@[simp] lemma zero_top (X : C) : is_pullback (0 : 0 ⟶ 0) (0 : 0 ⟶ X) (0 : 0 ⟶ X) (𝟙 X) := (zero_left X).flip +/-- The square with `0 : 0 ⟶ 0` on the right and `𝟙 X` on the left is a pullback square. -/ +@[simp] lemma zero_right (X : C) : is_pullback (0 : X ⟶ 0) (𝟙 X) (0 : 0 ⟶ 0) (0 : X ⟶ 0) := +of_iso_pullback (by simp) ((zero_prod_iso X).symm ≪≫ (pullback_zero_zero_iso _ _).symm) + (by simp) (by simp) + +/-- The square with `0 : 0 ⟶ 0` on the bottom and `𝟙 X` on the top is a pullback square. -/ +@[simp] lemma zero_bot (X : C) : is_pullback (𝟙 X) (0 : X ⟶ 0) (0 : X ⟶ 0) (0 : 0 ⟶ 0) := +(zero_right X).flip + end /-- Paste two pullback squares "vertically" to obtain another pullback square. -/ @@ -371,6 +415,78 @@ lemma of_right {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} is_pullback h₁₁ v₁₁ v₁₂ h₂₁ := (of_bot s.flip p.symm t.flip).flip +section + +variables [has_zero_object C] [has_zero_morphisms C] +open_locale zero_object + +lemma of_is_bilimit {b : binary_bicone X Y} (h : b.is_bilimit) : + is_pullback b.fst b.snd (0 : X ⟶ 0) (0 : Y ⟶ 0) := +by convert is_pullback.of_is_product' h.is_limit has_zero_object.zero_is_terminal + +@[simp] lemma of_has_biproduct (X Y : C) [has_binary_biproduct X Y] : + is_pullback biprod.fst biprod.snd (0 : X ⟶ 0) (0 : Y ⟶ 0) := +of_is_bilimit (binary_biproduct.is_bilimit X Y) + +lemma inl_snd' {b : binary_bicone X Y} (h : b.is_bilimit) : + is_pullback b.inl (0 : X ⟶ 0) b.snd (0 : 0 ⟶ Y) := +by { refine of_right _ (by simp) (of_is_bilimit h), simp, } + +/-- +The square +``` + X --inl--> X ⊞ Y + | | + 0 snd + | | + v v + 0 ---0-----> Y +``` +is a pullback square. +-/ +@[simp] lemma inl_snd (X Y : C) [has_binary_biproduct X Y] : + is_pullback biprod.inl (0 : X ⟶ 0) biprod.snd (0 : 0 ⟶ Y) := +inl_snd' (binary_biproduct.is_bilimit X Y) + +lemma inr_fst' {b : binary_bicone X Y} (h : b.is_bilimit) : + is_pullback b.inr (0 : Y ⟶ 0) b.fst (0 : 0 ⟶ X) := +by { apply flip, refine of_bot _ (by simp) (of_is_bilimit h), simp, } + +/-- +The square +``` + Y --inr--> X ⊞ Y + | | + 0 fst + | | + v v + 0 ---0-----> X +``` +is a pullback square. +-/ +@[simp] lemma inr_fst (X Y : C) [has_binary_biproduct X Y] : + is_pullback biprod.inr (0 : Y ⟶ 0) biprod.fst (0 : 0 ⟶ X) := +inr_fst' (binary_biproduct.is_bilimit X Y) + +lemma of_is_bilimit' {b : binary_bicone X Y} (h : b.is_bilimit) : + is_pullback (0 : 0 ⟶ X) (0 : 0 ⟶ Y) b.inl b.inr := +by { refine is_pullback.of_right _ (by simp) (is_pullback.inl_snd' h).flip, simp, } + +lemma of_has_binary_biproduct (X Y : C) [has_binary_biproduct X Y] : + is_pullback (0 : 0 ⟶ X) (0 : 0 ⟶ Y) biprod.inl biprod.inr := +of_is_bilimit' (binary_biproduct.is_bilimit X Y) + +instance has_pullback_biprod_fst_biprod_snd [has_binary_biproduct X Y] : + has_pullback (biprod.inl : X ⟶ _) (biprod.inr : Y ⟶ _) := +has_limit.mk ⟨_, (of_has_binary_biproduct X Y).is_limit⟩ + +/-- The pullback of `biprod.inl` and `biprod.inr` is the zero object. -/ +def pullback_biprod_inl_biprod_inr [has_binary_biproduct X Y] : + pullback (biprod.inl : X ⟶ _) (biprod.inr : Y ⟶ _) ≅ 0 := +limit.iso_limit_cone ⟨_, (of_has_binary_biproduct X Y).is_limit⟩ + +end + lemma op (h : is_pullback fst snd f g) : is_pushout g.op f.op snd.op fst.op := is_pushout.of_is_colimit (is_colimit.of_iso_colimit (limits.pullback_cone.is_limit_equiv_is_colimit_op h.flip.cone h.flip.is_limit) @@ -400,7 +516,7 @@ variables [has_zero_object C] [has_zero_morphisms C] open_locale zero_object /-- The square with `0 : 0 ⟶ 0` on the right and `𝟙 X` on the left is a pushout square. -/ -lemma zero_right (X : C) : is_pushout (0 : X ⟶ 0) (𝟙 X) (0 : 0 ⟶ 0) (0 : X ⟶ 0) := +@[simp] lemma zero_right (X : C) : is_pushout (0 : X ⟶ 0) (𝟙 X) (0 : 0 ⟶ 0) (0 : X ⟶ 0) := { w := by simp, is_colimit' := ⟨{ desc := λ s, 0, @@ -412,9 +528,18 @@ lemma zero_right (X : C) : is_pushout (0 : X ⟶ 0) (𝟙 X) (0 : 0 ⟶ 0) (0 : end }⟩ } /-- The square with `0 : 0 ⟶ 0` on the bottom and `𝟙 X` on the top is a pushout square. -/ -lemma zero_bot (X : C) : is_pushout (𝟙 X) (0 : X ⟶ 0) (0 : X ⟶ 0) (0 : 0 ⟶ 0) := +@[simp] lemma zero_bot (X : C) : is_pushout (𝟙 X) (0 : X ⟶ 0) (0 : X ⟶ 0) (0 : 0 ⟶ 0) := (zero_right X).flip +/-- The square with `0 : 0 ⟶ 0` on the right left `𝟙 X` on the right is a pushout square. -/ +@[simp] lemma zero_left (X : C) : is_pushout (0 : 0 ⟶ X) (0 : 0 ⟶ 0) (𝟙 X) (0 : 0 ⟶ X) := +of_iso_pushout (by simp) ((coprod_zero_iso X).symm ≪≫ (pushout_zero_zero_iso _ _).symm) + (by simp) (by simp) + +/-- The square with `0 : 0 ⟶ 0` on the top and `𝟙 X` on the bottom is a pushout square. -/ +@[simp] lemma zero_top (X : C) : is_pushout (0 : 0 ⟶ 0) (0 : 0 ⟶ X) (0 : 0 ⟶ X) (𝟙 X) := +(zero_left X).flip + end /-- Paste two pushout squares "vertically" to obtain another pushout square. -/ @@ -457,6 +582,78 @@ lemma of_right {X₁₁ X₁₂ X₁₃ X₂₁ X₂₂ X₂₃ : C} is_pushout h₁₂ v₁₂ v₁₃ h₂₂ := (of_bot s.flip p.symm t.flip).flip +section + +variables [has_zero_object C] [has_zero_morphisms C] +open_locale zero_object + +lemma of_is_bilimit {b : binary_bicone X Y} (h : b.is_bilimit) : + is_pushout (0 : 0 ⟶ X) (0 : 0 ⟶ Y) b.inl b.inr := +by convert is_pushout.of_is_coproduct' h.is_colimit has_zero_object.zero_is_initial + +@[simp] lemma of_has_biproduct (X Y : C) [has_binary_biproduct X Y] : + is_pushout (0 : 0 ⟶ X) (0 : 0 ⟶ Y) biprod.inl biprod.inr := +of_is_bilimit (binary_biproduct.is_bilimit X Y) + +lemma inl_snd' {b : binary_bicone X Y} (h : b.is_bilimit) : + is_pushout b.inl (0 : X ⟶ 0) b.snd (0 : 0 ⟶ Y) := +by { apply flip, refine of_right _ (by simp) (of_is_bilimit h), simp, } + +/-- +The square +``` + X --inl--> X ⊞ Y + | | + 0 snd + | | + v v + 0 ---0-----> Y +``` +is a pushout square. +-/ +lemma inl_snd (X Y : C) [has_binary_biproduct X Y] : + is_pushout biprod.inl (0 : X ⟶ 0) biprod.snd (0 : 0 ⟶ Y) := +inl_snd' (binary_biproduct.is_bilimit X Y) + +lemma inr_fst' {b : binary_bicone X Y} (h : b.is_bilimit) : + is_pushout b.inr (0 : Y ⟶ 0) b.fst (0 : 0 ⟶ X) := +by { refine of_bot _ (by simp) (of_is_bilimit h), simp, } + +/-- +The square +``` + Y --inr--> X ⊞ Y + | | + 0 fst + | | + v v + 0 ---0-----> X +``` +is a pushout square. +-/ +lemma inr_fst (X Y : C) [has_binary_biproduct X Y] : + is_pushout biprod.inr (0 : Y ⟶ 0) biprod.fst (0 : 0 ⟶ X) := +inr_fst' (binary_biproduct.is_bilimit X Y) + +lemma of_is_bilimit' {b : binary_bicone X Y} (h : b.is_bilimit) : + is_pushout b.fst b.snd (0 : X ⟶ 0) (0 : Y ⟶ 0) := +by { refine is_pushout.of_right _ (by simp) (is_pushout.inl_snd' h), simp, } + +lemma of_has_binary_biproduct (X Y : C) [has_binary_biproduct X Y] : + is_pushout biprod.fst biprod.snd (0 : X ⟶ 0) (0 : Y ⟶ 0) := +of_is_bilimit' (binary_biproduct.is_bilimit X Y) + +instance has_pushout_biprod_fst_biprod_snd [has_binary_biproduct X Y] : + has_pushout (biprod.fst : _ ⟶ X) (biprod.snd : _ ⟶ Y) := +has_colimit.mk ⟨_, (of_has_binary_biproduct X Y).is_colimit⟩ + +/-- The pushout of `biprod.fst` and `biprod.snd` is the zero object. -/ +def pushout_biprod_fst_biprod_snd [has_binary_biproduct X Y] : + pushout (biprod.fst : _ ⟶ X) (biprod.snd : _ ⟶ Y) ≅ 0 := +colimit.iso_colimit_cocone ⟨_, (of_has_binary_biproduct X Y).is_colimit⟩ + +end + lemma op (h : is_pushout f g inl inr) : is_pullback inr.op inl.op g.op f.op := is_pullback.of_is_limit (is_limit.of_iso_limit (limits.pushout_cocone.is_colimit_equiv_is_limit_op h.flip.cocone h.flip.is_colimit) @@ -480,6 +677,82 @@ lemma of_vert_is_iso [is_iso g] [is_iso inl] (sq : comm_sq f g inl inr) : end is_pushout +namespace bicartesian_sq + +variables {W X Y Z : C} {f : W ⟶ X} {g : W ⟶ Y} {h : X ⟶ Z} {i : Y ⟶ Z} + +lemma of_is_pullback_is_pushout (p₁ : is_pullback f g h i) (p₂ : is_pushout f g h i) : + bicartesian_sq f g h i := +bicartesian_sq.mk p₁.to_comm_sq ⟨p₁.is_limit⟩ ⟨p₂.is_colimit⟩ + +lemma flip (p : bicartesian_sq f g h i) : bicartesian_sq g f i h := +of_is_pullback_is_pushout p.to_is_pullback.flip p.to_is_pushout.flip + +variables [has_zero_object C] [has_zero_morphisms C] +open_locale zero_object + +/-- +``` + X ⊞ Y --fst--> X + | | + snd 0 + | | + v v + Y -----0---> 0 +``` +is a bicartesian square. +-/ +lemma of_is_biproduct₁ {b : binary_bicone X Y} (h : b.is_bilimit) : + bicartesian_sq b.fst b.snd (0 : X ⟶ 0) (0 : Y ⟶ 0) := +of_is_pullback_is_pushout (is_pullback.of_is_bilimit h) (is_pushout.of_is_bilimit' h) + +/-- +``` + 0 -----0---> X + | | + 0 inl + | | + v v + Y --inr--> X ⊞ Y +``` +is a bicartesian square. +-/ +lemma of_is_biproduct₂ {b : binary_bicone X Y} (h : b.is_bilimit) : + bicartesian_sq (0 : 0 ⟶ X) (0 : 0 ⟶ Y) b.inl b.inr := +of_is_pullback_is_pushout (is_pullback.of_is_bilimit' h) (is_pushout.of_is_bilimit h) + +/-- +``` + X ⊞ Y --fst--> X + | | + snd 0 + | | + v v + Y -----0---> 0 +``` +is a bicartesian square. +-/ +@[simp] lemma of_has_biproduct₁ [has_binary_biproduct X Y] : + bicartesian_sq biprod.fst biprod.snd (0 : X ⟶ 0) (0 : Y ⟶ 0) := +by convert of_is_biproduct₁ (binary_biproduct.is_bilimit X Y) + +/-- +``` + 0 -----0---> X + | | + 0 inl + | | + v v + Y --inr--> X ⊞ Y +``` +is a bicartesian square. +-/ +@[simp] lemma of_has_biproduct₂ [has_binary_biproduct X Y] : + bicartesian_sq (0 : 0 ⟶ X) (0 : 0 ⟶ Y) biprod.inl biprod.inr := +by convert of_is_biproduct₂ (binary_biproduct.is_bilimit X Y) + +end bicartesian_sq + namespace functor variables {D : Type u₂} [category.{v₂} D] From 8a8470cc4d27ead7252b37f604f845096101a643 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Wed, 5 Oct 2022 06:38:09 +0000 Subject: [PATCH 566/828] feat(data/finset/prod): `finset.diag` and `finset.off_diag` lemmas (#16808) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Additional lemmas for `diag` and `off_diag` including `finset.coe_off_diag` to relate `finset.off_diag` to `set.off_diag`. Co-authored-by: Yaël Dillies Co-authored-by: Kyle Miller --- src/data/finset/prod.lean | 27 +++++++++++++++++---------- 1 file changed, 17 insertions(+), 10 deletions(-) diff --git a/src/data/finset/prod.lean b/src/data/finset/prod.lean index 5e8f04e86d7cb..58e887e9e010b 100644 --- a/src/data/finset/prod.lean +++ b/src/data/finset/prod.lean @@ -142,7 +142,7 @@ by { ext ⟨x, y⟩, simp only [and_or_distrib_left, mem_union, mem_product] } end prod section diag -variables (s t : finset α) [decidable_eq α] +variables [decidable_eq α] (s t : finset α) /-- Given a finite set `s`, the diagonal, `s.diag` is the set of pairs of the form `(a, a)` for `a ∈ s`. -/ @@ -160,6 +160,9 @@ by { simp only [diag, mem_filter, mem_product], split; intros h; by { simp only [off_diag, mem_filter, mem_product], split; intros h; simp only [h, ne.def, not_false_iff, and_self] } +@[simp, norm_cast] lemma coe_off_diag : (s.off_diag : set (α × α)) = (s : set α).off_diag := +set.ext $ mem_off_diag _ + @[simp] lemma diag_card : (diag s).card = s.card := begin suffices : diag s = s.image (λ a, (a, a)), @@ -177,6 +180,12 @@ begin apply filter_card_add_filter_neg_card_eq_card, end +@[mono] lemma diag_mono : monotone (diag : finset α → finset (α × α)) := +λ s t h x hx, (mem_diag _ _).2 $ and.imp_left (@h _) $ (mem_diag _ _).1 hx + +@[mono] lemma off_diag_mono : monotone (off_diag : finset α → finset (α × α)) := +λ s t h x hx, (mem_off_diag _ _).2 $ and.imp (@h _) (and.imp_left $ @h _) $ (mem_off_diag _ _).1 hx + @[simp] lemma diag_empty : (∅ : finset α).diag = ∅ := rfl @[simp] lemma off_diag_empty : (∅ : finset α).off_diag = ∅ := rfl @@ -193,6 +202,12 @@ by rw [←diag_union_off_diag, union_comm, union_sdiff_self, lemma product_sdiff_off_diag : s ×ˢ s \ s.off_diag = s.diag := by rw [←diag_union_off_diag, union_sdiff_self, sdiff_eq_self_of_disjoint (disjoint_diag_off_diag _)] +lemma diag_inter : (s ∩ t).diag = s.diag ∩ t.diag := +ext $ λ x, by simpa only [mem_diag, mem_inter] using and_and_distrib_right _ _ _ + +lemma off_diag_inter : (s ∩ t).off_diag = s.off_diag ∩ t.off_diag := +coe_injective $ by { push_cast, exact set.off_diag_inter _ _ } + lemma diag_union : (s ∪ t).diag = s.diag ∪ t.diag := by { ext ⟨i, j⟩, simp only [mem_diag, mem_union, or_and_distrib_right] } @@ -200,15 +215,7 @@ variables {s t} lemma off_diag_union (h : disjoint s t) : (s ∪ t).off_diag = s.off_diag ∪ t.off_diag ∪ s ×ˢ t ∪ t ×ˢ s := -begin - rw [off_diag, union_product, product_union, product_union, union_comm _ (t ×ˢ t), - union_assoc, union_left_comm (s ×ˢ t), ←union_assoc, filter_union, filter_union, ←off_diag, - ←off_diag, filter_true_of_mem, ←union_assoc], - simp only [mem_union, mem_product, ne.def, prod.forall], - rintro i j (⟨hi, hj⟩ | ⟨hi, hj⟩), - { exact h.forall_ne_finset hi hj }, - { exact h.symm.forall_ne_finset hi hj }, -end +coe_injective $ by { push_cast, exact set.off_diag_union (disjoint_coe.2 h) } variables (a : α) From 96782a2d6dcded92116d8ac9ae48efb41d46a27c Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Wed, 5 Oct 2022 09:32:55 +0000 Subject: [PATCH 567/828] feat(algebra/char_p/mixed_char_zero): define mixed/equal characteristic zero (#14672) API for equal/mixed characteristic zero with main theorem to split propositions by characteristics. Co-authored-by: Jon Co-authored-by: Jon Eugster --- src/algebra/char_p/algebra.lean | 30 ++ src/algebra/char_p/mixed_char_zero.lean | 404 ++++++++++++++++++++++++ 2 files changed, 434 insertions(+) create mode 100644 src/algebra/char_p/mixed_char_zero.lean diff --git a/src/algebra/char_p/algebra.lean b/src/algebra/char_p/algebra.lean index feebf7234a385..0748eeb869bcb 100644 --- a/src/algebra/char_p/algebra.lean +++ b/src/algebra/char_p/algebra.lean @@ -58,6 +58,36 @@ lemma char_zero_of_injective_algebra_map {R A : Type*} [comm_semiring R] [semiri -- `char_p.char_p_to_char_zero A _ (char_p_of_injective_algebra_map h 0)` does not work -- here as it would require `ring A`. +/-! +As an application, a `ℚ`-algebra has characteristic zero. +-/ +section Q_algebra + +variables (R : Type*) [nontrivial R] + +/-- A nontrivial `ℚ`-algebra has `char_p` equal to zero. + +This cannot be a (local) instance because it would immediately form a loop with the +instance `algebra_rat`. It's probably easier to go the other way: prove `char_zero R` and +automatically receive an `algebra ℚ R` instance. +-/ +lemma algebra_rat.char_p_zero [semiring R] [algebra ℚ R] : char_p R 0 := +char_p_of_injective_algebra_map (algebra_map ℚ R).injective 0 + +/-- A nontrivial `ℚ`-algebra has characteristic zero. + +This cannot be a (local) instance because it would immediately form a loop with the +instance `algebra_rat`. It's probably easier to go the other way: prove `char_zero R` and +automatically receive an `algebra ℚ R` instance. +-/ +lemma algebra_rat.char_zero [ring R] [algebra ℚ R] : char_zero R := +@char_p.char_p_to_char_zero R _ (algebra_rat.char_p_zero R) + +end Q_algebra + +/-! +An algebra over a field has the same characteristic as the field. +-/ section variables (K L : Type*) [field K] [comm_semiring L] [nontrivial L] [algebra K L] diff --git a/src/algebra/char_p/mixed_char_zero.lean b/src/algebra/char_p/mixed_char_zero.lean new file mode 100644 index 0000000000000..3b60ebd4fd69f --- /dev/null +++ b/src/algebra/char_p/mixed_char_zero.lean @@ -0,0 +1,404 @@ +/- +Copyright (c) 2022 Jon Eugster. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jon Eugster +-/ +import algebra.char_p.algebra +import algebra.char_p.local_ring +import data.pnat.basic +import ring_theory.ideal.quotient +import tactic.field_simp + +/-! +# Equal and mixed characteristic + +In commutative algebra, some statments are simpler when working over a `ℚ`-algebra `R`, in which +case one also says that the ring has "equal characteristic zero". A ring that is not a +`ℚ`-algebra has either positive characteristic or there exists a prime ideal `I ⊂ R` such that +the quotient `R ⧸ I` has positive characteristic `p > 0`. In this case one speaks of +"mixed characteristic `(0, p)`", where `p` is only unique if `R` is local. + +Examples of mixed characteristic rings are `ℤ` or the `p`-adic integers/numbers. + +This file provides the main theorem `split_by_characteristic` that splits any proposition `P` into +the following three cases: + +1) Positive characteristic: `char_p R p` (where `p ≠ 0`) +2) Equal characteristic zero: `algebra ℚ R` +3) Mixed characteristic: `mixed_char_zero R p` (where `p` is prime) + +## Main definitions + +- `mixed_char_zero` : A ring has mixed characteristic `(0, p)` if it has characteristic zero + and there exists an ideal such that the quotient `R ⧸ I` has characteristic `p`. + +## Main results + +- `split_equal_mixed_char` : Split a statement into equal/mixed characteristic zero. + +This main theorem has the following three corollaries which include the positive +characteristic case for convenience: + +- `split_by_characteristic` : Generally consider positive char `p ≠ 0`. +- `split_by_characteristic_domain` : In a domain we can assume that `p` is prime. +- `split_by_characteristic_local_ring` : In a local ring we can assume that `p` is a prime power. + +## TODO + +- Relate mixed characteristic in a local ring to p-adic numbers [number_theory.padics]. + +-/ + +variables (R : Type*) [comm_ring R] + +/-! +### Mixed characteristic +-/ + +/-- +A ring of characteristic zero is of "mixed characteristic `(0, p)`" if there exists an ideal +such that the quotient `R ⧸ I` has caracteristic `p`. + +**Remark:** For `p = 0`, `mixed_char R 0` is a meaningless definition as `R ⧸ ⊥ ≅ R` has by +definition always characteristic zero. +One could require `(I ≠ ⊥)` in the definition, but then `mixed_char R 0` would mean something +like `ℤ`-algebra of extension degree `≥ 1` and would be completely independent from +whether something is a `ℚ`-algebra or not (e.g. `ℚ[X]` would satisfy it but `ℚ` wouldn't). +-/ +class mixed_char_zero (p : ℕ) : Prop := +[to_char_zero : char_zero R] +(char_p_quotient : ∃ (I : ideal R), (I ≠ ⊤) ∧ char_p (R ⧸ I) p) + +namespace mixed_char_zero + +/-- +Reduction to `p` prime: When proving any statement `P` about mixed characteristic rings we +can always assume that `p` is prime. +-/ +lemma reduce_to_p_prime {P : Prop} : + (∀ p > 0, mixed_char_zero R p → P) ↔ + (∀ (p : ℕ), p.prime → mixed_char_zero R p → P) := +begin + split, + { intros h q q_prime q_mixed_char, + exact h q (nat.prime.pos q_prime) q_mixed_char }, + { intros h q q_pos q_mixed_char, + rcases q_mixed_char.char_p_quotient with ⟨I, hI_ne_top, hI_char⟩, + + -- Krull's Thm: There exists a prime ideal `P` such that `I ≤ P` + rcases ideal.exists_le_maximal I hI_ne_top with ⟨M, hM_max, h_IM⟩, + resetI, -- make `hI_char : char_p (R ⧸ I) q` an instance. + + let r := ring_char (R ⧸ M), + have r_pos : r ≠ 0, + { have q_zero := congr_arg (ideal.quotient.factor I M h_IM) (char_p.cast_eq_zero (R ⧸ I) q), + simp only [map_nat_cast, map_zero] at q_zero, + apply ne_zero_of_dvd_ne_zero (ne_of_gt q_pos), + exact (char_p.cast_eq_zero_iff (R ⧸ M) r q).mp q_zero }, + have r_prime : nat.prime r := + or_iff_not_imp_right.1 (char_p.char_is_prime_or_zero (R ⧸ M) r) r_pos, + apply h r r_prime, + haveI : char_zero R := q_mixed_char.to_char_zero, + exact ⟨⟨M, hM_max.ne_top, ring_char.of_eq rfl⟩⟩ } +end + +/-- +Reduction to `I` prime ideal: When proving statements about mixed characteristic rings, +after we reduced to `p` prime, we can assume that the ideal `I` in the definition is maximal. +-/ +lemma reduce_to_maximal_ideal {p : ℕ} (hp : nat.prime p) : + (∃ (I : ideal R), (I ≠ ⊤) ∧ char_p (R ⧸ I) p) ↔ + (∃ (I : ideal R), (I.is_maximal) ∧ char_p (R ⧸ I) p) := +begin + split, + { intro g, + rcases g with ⟨I, ⟨hI_not_top, hI⟩⟩, + + -- Krull's Thm: There exists a prime ideal `M` such that `I ≤ M`. + rcases ideal.exists_le_maximal I hI_not_top with ⟨M, ⟨hM_max, hM⟩⟩, + + use M, + split, + exact hM_max, + { cases char_p.exists (R ⧸ M) with r hr, + convert hr, + resetI, -- make `hr : char_p (R ⧸ M) r` an instance. + + have r_dvd_p : r ∣ p, + { rw ←char_p.cast_eq_zero_iff (R ⧸ M) r p, + convert congr_arg (ideal.quotient.factor I M hM) (char_p.cast_eq_zero (R ⧸ I) p) }, + symmetry, + apply (nat.prime.eq_one_or_self_of_dvd hp r r_dvd_p).resolve_left, + exact char_p.char_ne_one (R ⧸ M) r }}, + { rintro ⟨I, hI_max, hI⟩, + use I, + exact ⟨ideal.is_maximal.ne_top hI_max, hI⟩ } +end + +end mixed_char_zero + +/-! +### Equal characteristic zero + +A commutative ring `R` has "equal characteristic zero" if it satisfies one of the following +equivalent properties: + +1) `R` is a `ℚ`-algebra. +2) The quotient `R ⧸ I` has characteristic zero for any proper ideal `I ⊂ R`. +3) `R` has characteristic zero and does not have mixed characteristic for any prime `p`. + +We show `(1) ↔ (2) ↔ (3)`, and most of the following is concerned with constructing +an explicit algebra map `ℚ →+* R` (given by `x ↦ (x.num : R) /ₚ ↑x.pnat_denom`) +for the direction `(1) ← (2)`. + +Note: Property `(2)` is denoted as `equal_char_zero` in the statement names below. +-/ +section equal_char_zero + +/-- +`ℚ`-algebra implies equal characteristic. +-/ +@[nolint unused_arguments] -- argument `[nontrivial R]` is used in the first line of the proof. +lemma Q_algebra_to_equal_char_zero [nontrivial R] [algebra ℚ R] : + ∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I) := +begin + haveI : char_zero R := algebra_rat.char_zero R, + intros I hI, + constructor, + intros a b h_ab, + contrapose! hI, + -- `↑a - ↑b` is a unit contained in `I`, which contradicts `I ≠ ⊤`. + refine I.eq_top_of_is_unit_mem _ (is_unit.map (algebra_map ℚ R) (is_unit.mk0 (a - b : ℚ) _)), + { simpa only [← ideal.quotient.eq_zero_iff_mem, map_sub, sub_eq_zero, map_nat_cast] }, + simpa only [ne.def, sub_eq_zero] using (@nat.cast_injective ℚ _ _).ne hI +end + +section construction_of_Q_algebra + +/-- Internal: Not intended to be used outside this local construction. -/ +lemma equal_char_zero.pnat_coe_is_unit [h : fact (∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I))] + (n : ℕ+) : is_unit (n : R) := +begin + -- `n : R` is a unit iff `(n)` is not a proper ideal in `R`. + rw ← ideal.span_singleton_eq_top, + -- So by contrapositive, we should show the quotient does not have characteristic zero. + apply not_imp_comm.mp (h.elim (ideal.span {n})), + unfreezingI { intro h_char_zero }, + -- In particular, the image of `n` in the quotient should be nonzero. + apply (h_char_zero.cast_injective).ne n.ne_zero, + -- But `n` generates the ideal, so its image is clearly zero. + rw [←map_nat_cast (ideal.quotient.mk _), nat.cast_zero, ideal.quotient.eq_zero_iff_mem], + exact ideal.subset_span (set.mem_singleton _) +end + +/-- Internal: Not intended to be used outside this local construction. -/ +noncomputable instance equal_char_zero.pnat_has_coe_units + [fact (∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I))] : has_coe_t ℕ+ Rˣ := +⟨λn, (equal_char_zero.pnat_coe_is_unit R n).unit⟩ + +/-- Internal: Not intended to be used outside this local construction. -/ +lemma equal_char_zero.pnat_coe_units_eq_one [fact (∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I))] : + ((1 : ℕ+) : Rˣ) = 1 := +begin + apply units.ext, + rw units.coe_one, + change ((equal_char_zero.pnat_coe_is_unit R 1).unit : R) = 1, + rw is_unit.unit_spec (equal_char_zero.pnat_coe_is_unit R 1), + rw [coe_coe, pnat.one_coe, nat.cast_one], +end + +/-- Internal: Not intended to be used outside this local construction. -/ +lemma equal_char_zero.pnat_coe_units_coe_eq_coe + [fact (∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I))] (n : ℕ+) : + ((n : Rˣ) : R) = ↑n := +begin + change ((equal_char_zero.pnat_coe_is_unit R n).unit : R) = ↑n, + simp only [is_unit.unit_spec], +end + +/-- +Equal characteristic implies `ℚ`-algebra. +-/ +noncomputable def equal_char_zero_to_Q_algebra (h : ∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I)) : + algebra ℚ R := +by haveI : fact (∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I)) := ⟨h⟩; exact +ring_hom.to_algebra + { to_fun := λ x, x.num /ₚ ↑(x.pnat_denom), + map_zero' := by simp [divp], + map_one' := by simp [equal_char_zero.pnat_coe_units_eq_one], + map_mul' := + begin + intros a b, + field_simp, + repeat { rw equal_char_zero.pnat_coe_units_coe_eq_coe R }, + transitivity (↑((a * b).num * (a.denom) * (b.denom)) : R), + { simp_rw [int.cast_mul, int.cast_coe_nat, coe_coe, rat.coe_pnat_denom], + ring }, + rw rat.mul_num_denom' a b, + simp + end, + map_add' := + begin + intros a b, + field_simp, + repeat { rw equal_char_zero.pnat_coe_units_coe_eq_coe R }, + transitivity (↑((a + b).num * a.denom * b.denom) : R), + { simp_rw [int.cast_mul, int.cast_coe_nat, coe_coe, rat.coe_pnat_denom], + ring }, + rw rat.add_num_denom' a b, + simp + end } + +end construction_of_Q_algebra + +end equal_char_zero + +/-- +Not mixed characteristic implies equal characteristic. +-/ +lemma not_mixed_char_to_equal_char_zero [char_zero R] (h : ∀ p > 0, ¬mixed_char_zero R p) : + ∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I) := +begin + intros I hI_ne_top, + apply char_p.char_p_to_char_zero _, + cases char_p.exists (R ⧸ I) with p hp, + cases p, + { exact hp }, + { have h_mixed : mixed_char_zero R p.succ := ⟨⟨I, ⟨hI_ne_top, hp⟩⟩⟩, + exact absurd h_mixed (h p.succ p.succ_pos) } +end + +/-- +Equal characteristic implies not mixed characteristic. +-/ +lemma equal_char_zero_to_not_mixed_char (h : ∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I)) : + ∀ p > 0, ¬mixed_char_zero R p := +begin + intros p p_pos, + by_contradiction hp_mixed_char, + rcases hp_mixed_char.char_p_quotient with ⟨I, hI_ne_top, hI_p⟩, + replace hI_zero : char_p (R ⧸ I) 0 := @char_p.of_char_zero _ _ (h I hI_ne_top), + exact absurd (char_p.eq (R ⧸ I) hI_p hI_zero) (ne_of_gt p_pos), +end + +/-- +A ring of characteristic zero has equal characteristic iff it does not +have mixed characteristic for any `p`. +-/ +lemma equal_char_zero_iff_not_mixed_char [char_zero R] : + (∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I)) ↔ (∀ p > 0, ¬mixed_char_zero R p) := +⟨equal_char_zero_to_not_mixed_char R, not_mixed_char_to_equal_char_zero R⟩ + +/-- +A ring is a `ℚ`-algebra iff it has equal characteristic zero. +-/ +theorem Q_algebra_iff_equal_char_zero [nontrivial R] : + nonempty (algebra ℚ R) ↔ ∀ (I : ideal R), I ≠ ⊤ → char_zero (R ⧸ I) := +begin + split, + { intro h_alg, + haveI h_alg' : algebra ℚ R := h_alg.some, + apply Q_algebra_to_equal_char_zero }, + { intro h, + apply nonempty.intro, + exact equal_char_zero_to_Q_algebra R h } +end + +/-- +A ring of characteristic zero is not a `ℚ`-algebra iff it has mixed characteristic for some `p`. +-/ +theorem not_Q_algebra_iff_not_equal_char_zero [char_zero R] : + is_empty (algebra ℚ R) ↔ (∃ p > 0, mixed_char_zero R p) := +begin + rw ←not_iff_not, + push_neg, + rw [not_is_empty_iff, ←equal_char_zero_iff_not_mixed_char], + apply Q_algebra_iff_equal_char_zero, +end + +/-! +# Splitting statements into different characteristic + +Statements to split a proof by characteristic. There are 3 theorems here that are very +similar. They only differ in the assumptions we can make on the positive characteristic +case: +Generally we need to consider all `p ≠ 0`, but if `R` is a local ring, we can assume +that `p` is a prime power. And if `R` is a domain, we can even assume that `p` is prime. +-/ +section main_statements + +variable {P : Prop} + +/-- +Split a `Prop` in characteristic zero into equal and mixed characteristic. +-/ +theorem split_equal_mixed_char [char_zero R] + (h_equal : algebra ℚ R → P) + (h_mixed : ∀ (p : ℕ), (nat.prime p → mixed_char_zero R p → P)) : P := +begin + by_cases h : ∃ p > 0, mixed_char_zero R p, + { rcases h with ⟨p, ⟨H, hp⟩⟩, + rw ←mixed_char_zero.reduce_to_p_prime at h_mixed, + exact h_mixed p H hp }, + { apply h_equal, + rw [←not_Q_algebra_iff_not_equal_char_zero, not_is_empty_iff] at h, + exact h.some }, +end + +example (n : ℕ) (h : n ≠ 0) :0 < n := zero_lt_iff.mpr h + + +/-- Split any `Prop` over `R` into the three cases: +- positive characteristic. +- equal characteristic zero. +- mixed characteristic `(0, p)`. +-/ +theorem split_by_characteristic + (h_pos : ∀ (p : ℕ), (p ≠ 0 → char_p R p → P)) + (h_equal : algebra ℚ R → P) + (h_mixed : ∀ (p : ℕ), (nat.prime p → mixed_char_zero R p → P)) : P := +begin + cases char_p.exists R with p p_char, + by_cases p = 0, + { rw h at p_char, + resetI, -- make `p_char : char_p R 0` an instance. + haveI h0 : char_zero R := char_p.char_p_to_char_zero R, + exact split_equal_mixed_char R h_equal h_mixed }, + exact h_pos p h p_char, +end + +/-- In a `is_domain R`, split any `Prop` over `R` into the three cases: +- *prime* characteristic. +- equal characteristic zero. +- mixed characteristic `(0, p)`. +-/ +theorem split_by_characteristic_domain [is_domain R] + (h_pos : ∀ (p : ℕ), (nat.prime p → char_p R p → P)) + (h_equal : algebra ℚ R → P) + (h_mixed : ∀ (p : ℕ), (nat.prime p → mixed_char_zero R p → P)) : P := +begin + refine split_by_characteristic R _ h_equal h_mixed, + introsI p p_pos p_char, + have p_prime : nat.prime p := + or_iff_not_imp_right.mp (char_p.char_is_prime_or_zero R p) p_pos, + exact h_pos p p_prime p_char, +end + +/-- In a `local_ring R`, split any `Prop` over `R` into the three cases: +- *prime power* characteristic. +- equal characteristic zero. +- mixed characteristic `(0, p)`. +-/ +theorem split_by_characteristic_local_ring [local_ring R] + (h_pos : ∀ (p : ℕ), (is_prime_pow p → char_p R p → P)) + (h_equal : algebra ℚ R → P) + (h_mixed : ∀ (p : ℕ), (nat.prime p → mixed_char_zero R p → P)) : P := +begin + refine split_by_characteristic R _ h_equal h_mixed, + introsI p p_pos p_char, + have p_ppow : is_prime_pow (p : ℕ) := + or_iff_not_imp_left.mp (char_p_zero_or_prime_power R p) p_pos, + exact h_pos p p_ppow p_char, +end + +end main_statements From e24808b301c8ad0e27f2c63d38309ea4c5d94fb7 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 5 Oct 2022 09:32:56 +0000 Subject: [PATCH 568/828] feat(ring_theory/filtration): Artin-Rees lemma and Krull's intersection theorems (#16000) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/ring_theory/filtration.lean | 252 +++++++++++++++++++++++++++++++- src/ring_theory/noetherian.lean | 17 +++ 2 files changed, 268 insertions(+), 1 deletion(-) diff --git a/src/ring_theory/filtration.lean b/src/ring_theory/filtration.lean index 6a117b916757d..6d8bced3a3ac9 100644 --- a/src/ring_theory/filtration.lean +++ b/src/ring_theory/filtration.lean @@ -7,6 +7,8 @@ Authors: Andrew Yang import ring_theory.noetherian import ring_theory.rees_algebra import ring_theory.finiteness +import data.polynomial.module +import order.hom.complete_lattice /-! @@ -14,6 +16,25 @@ import ring_theory.finiteness This file contains the definitions and basic results around (stable) `I`-filtrations of modules. +## Main results + +- `ideal.filtration`: An `I`-filtration on the module `M` is a sequence of decreasing submodules + `N i` such that `I • N ≤ I (i + 1)`. Note that we do not require the filtration to start from `⊤`. +- `ideal.filtration.stable`: An `I`-filtration is stable if `I • (N i) = N (i + 1)` for large + enough `i`. +- `ideal.filtration.submodule`: The associated module `⨁ Nᵢ` of a filtration, implemented as a + submodule of `M[X]`. +- `ideal.filtration.submodule_fg_iff_stable`: If `F.N i` are all finitely generated, then + `F.stable` iff `F.submodule.fg`. +- `ideal.filtration.stable.of_le`: In a finite module over a noetherian ring, + if `F' ≤ F`, then `F.stable → F'.stable`. +- `ideal.exists_pow_inf_eq_pow_smul`: **Artin-Rees lemma**. + given `N ≤ M`, there exists a `k` such that `IⁿM ⊓ N = Iⁿ⁻ᵏ(IᵏM ⊓ N)` for all `n ≥ k`. +- `ideal.infi_pow_eq_bot_of_local_ring`: + **Krull's intersection theorem** (`⨅ i, I ^ i = ⊥`) for noetherian local rings. +- `ideal.infi_pow_eq_bot_of_is_domain`: + **Krull's intersection theorem** (`⨅ i, I ^ i = ⊥`) for noetherian domains. + -/ @@ -25,7 +46,7 @@ open polynomial open_locale polynomial big_operators /-- An `I`-filtration on the module `M` is a sequence of decreasing submodules `N i` such that -`I • N ≤ I (i + 1)`. Note that we do not require the filtration to start from `⊤`. -/ +`I • (N i) ≤ N (i + 1)`. Note that we do not require the filtration to start from `⊤`. -/ @[ext] structure ideal.filtration (M : Type u) [add_comm_group M] [module R M] := (N : ℕ → submodule R M) @@ -196,4 +217,233 @@ begin refine ⟨(F.antitone _).trans (h₁ n), (F'.antitone _).trans (h₂ n)⟩; simp end +open polynomial_module + +variables (F F') + +/-- The `R[IX]`-submodule of `M[X]` associated with an `I`-filtration. -/ +protected +def submodule : submodule (rees_algebra I) (polynomial_module R M) := +{ carrier := { f | ∀ i, f i ∈ F.N i }, + add_mem' := λ f g hf hg i, submodule.add_mem _ (hf i) (hg i), + zero_mem' := λ i, submodule.zero_mem _, + smul_mem' := λ r f hf i, begin + rw [subalgebra.smul_def, polynomial_module.smul_apply], + apply submodule.sum_mem, + rintro ⟨j, k⟩ e, + rw finset.nat.mem_antidiagonal at e, + subst e, + exact F.pow_smul_le j k (submodule.smul_mem_smul (r.2 j) (hf k)) + end } + +@[simp] +lemma mem_submodule (f : polynomial_module R M) : f ∈ F.submodule ↔ ∀ i, f i ∈ F.N i := iff.rfl + +lemma inf_submodule : (F ⊓ F').submodule = F.submodule ⊓ F'.submodule := +by { ext, exact forall_and_distrib } + +variables (I M) + +/-- `ideal.filtration.submodule` as an `inf_hom` -/ +def submodule_inf_hom : + inf_hom (I.filtration M) (submodule (rees_algebra I) (polynomial_module R M)) := +{ to_fun := ideal.filtration.submodule, map_inf' := inf_submodule } + +variables {I M} + +lemma submodule_closure_single : + add_submonoid.closure (⋃ i, single R i '' (F.N i : set M)) = F.submodule.to_add_submonoid := +begin + apply le_antisymm, + { rw [add_submonoid.closure_le, set.Union_subset_iff], + rintro i _ ⟨m, hm, rfl⟩ j, + rw single_apply, + split_ifs, + { rwa ← h }, + { exact (F.N j).zero_mem } }, + { intros f hf, + rw [← f.sum_single], + apply add_submonoid.sum_mem _ _, + rintros c -, + exact add_submonoid.subset_closure (set.subset_Union _ c $ set.mem_image_of_mem _ (hf c)) } +end + +lemma submodule_span_single : + submodule.span (rees_algebra I) (⋃ i, single R i '' (F.N i : set M)) = F.submodule := +begin + rw [← submodule.span_closure, submodule_closure_single], + simp, +end + +lemma submodule_eq_span_le_iff_stable_ge (n₀ : ℕ) : + F.submodule = submodule.span _ (⋃ i ≤ n₀, single R i '' (F.N i : set M)) ↔ + ∀ n ≥ n₀, I • F.N n = F.N (n + 1) := +begin + rw [← submodule_span_single, ← has_le.le.le_iff_eq, submodule.span_le, + set.Union_subset_iff], + swap, { exact submodule.span_mono (set.Union₂_subset_Union _ _) }, + split, + { intros H n hn, + refine (F.smul_le n).antisymm _, + intros x hx, + obtain ⟨l, hl⟩ := (finsupp.mem_span_iff_total _ _ _).mp (H _ ⟨x, hx, rfl⟩), + replace hl := congr_arg (λ f : ℕ →₀ M, f (n + 1)) hl, + dsimp only at hl, + erw finsupp.single_eq_same at hl, + rw [← hl, finsupp.total_apply, finsupp.sum_apply], + apply submodule.sum_mem _ _, + rintros ⟨_, _, ⟨n', rfl⟩, _, ⟨hn', rfl⟩, m, hm, rfl⟩ -, + dsimp only [subtype.coe_mk], + rw [subalgebra.smul_def, smul_single_apply, if_pos (show n' ≤ n + 1, by linarith)], + have e : n' ≤ n := by linarith, + have := F.pow_smul_le_pow_smul (n - n') n' 1, + rw [tsub_add_cancel_of_le e, pow_one, add_comm _ 1, ← add_tsub_assoc_of_le e, add_comm] at this, + exact this (submodule.smul_mem_smul ((l _).2 $ n + 1 - n') hm) }, + { let F' := submodule.span (rees_algebra I) (⋃ i ≤ n₀, single R i '' (F.N i : set M)), + intros hF i, + have : ∀ i ≤ n₀, single R i '' (F.N i : set M) ⊆ F' := + λ i hi, set.subset.trans (set.subset_Union₂ i hi) submodule.subset_span, + induction i with j hj, + { exact this _ (zero_le _) }, + by_cases hj' : j.succ ≤ n₀, + { exact this _ hj' }, + simp only [not_le, nat.lt_succ_iff] at hj', + rw [nat.succ_eq_add_one, ← hF _ hj'], + rintro _ ⟨m, hm, rfl⟩, + apply submodule.smul_induction_on hm, + { intros r hr m' hm', + rw [add_comm, ← monomial_smul_single], + exact F'.smul_mem ⟨_, rees_algebra.monomial_mem.mpr (by rwa pow_one)⟩ + (hj $ set.mem_image_of_mem _ hm') }, + { intros x y hx hy, rw map_add, exact F'.add_mem hx hy } } +end + +/-- If the components of a filtration are finitely generated, then the filtration is stable iff +its associated submodule of is finitely generated. -/ +lemma submodule_fg_iff_stable (hF' : ∀ i, (F.N i).fg) : + F.submodule.fg ↔ F.stable := +begin + classical, + delta ideal.filtration.stable, + simp_rw ← F.submodule_eq_span_le_iff_stable_ge, + split, + { rintro H, + apply H.stablizes_of_supr_eq + ⟨λ n₀, submodule.span _ (⋃ (i : ℕ) (H : i ≤ n₀), single R i '' ↑(F.N i)), _⟩, + { dsimp, + rw [← submodule.span_Union, ← submodule_span_single], + congr' 1, + ext, + simp only [set.mem_Union, set.mem_image, set_like.mem_coe, exists_prop], + split, + { rintro ⟨-, i, -, e⟩, exact ⟨i, e⟩ }, + { rintro ⟨i, e⟩, exact ⟨i, i, le_refl i, e⟩ } }, + { intros n m e, + rw [submodule.span_le, set.Union₂_subset_iff], + intros i hi, + refine (set.subset.trans _ (set.subset_Union₂ i (hi.trans e : _))).trans + submodule.subset_span, + refl } }, + { rintros ⟨n, hn⟩, + rw hn, + simp_rw [submodule.span_Union₂, ← finset.mem_range_succ_iff, supr_subtype'], + apply submodule.fg_supr, + rintro ⟨i, hi⟩, + obtain ⟨s, hs⟩ := hF' i, + have : submodule.span (rees_algebra I) (s.image (lsingle R i) : set (polynomial_module R M)) + = submodule.span _ (single R i '' (F.N i : set M)), + { rw [finset.coe_image, ← submodule.span_span_of_tower R, ← submodule.map_span, hs], refl }, + rw [subtype.coe_mk, ← this], + exact ⟨_, rfl⟩ } +end +. +variables {F} + +lemma stable.of_le [is_noetherian_ring R] [h : module.finite R M] (hF : F.stable) + {F' : I.filtration M} (hf : F' ≤ F) : F'.stable := +begin + haveI := is_noetherian_of_fg_of_noetherian' h.1, + rw ← submodule_fg_iff_stable at hF ⊢, + any_goals { intro i, exact is_noetherian.noetherian _ }, + have := is_noetherian_of_fg_of_noetherian _ hF, + rw is_noetherian_submodule at this, + exact this _ (order_hom_class.mono (submodule_inf_hom M I) hf), +end + +lemma stable.inter_right [is_noetherian_ring R] [h : module.finite R M] (hF : F.stable) : + (F ⊓ F').stable := +hF.of_le inf_le_left + +lemma stable.inter_left [is_noetherian_ring R] [h : module.finite R M] (hF : F.stable) : + (F' ⊓ F).stable := +hF.of_le inf_le_right + end ideal.filtration + +variable (I) + +/-- **Artin-Rees lemma** -/ +lemma ideal.exists_pow_inf_eq_pow_smul [is_noetherian_ring R] [h : module.finite R M] + (N : submodule R M) : ∃ k : ℕ, ∀ n ≥ k, I ^ n • ⊤ ⊓ N = I ^ (n - k) • (I ^ k • ⊤ ⊓ N) := +((I.stable_filtration_stable ⊤).inter_right (I.trivial_filtration N)).exists_pow_smul_eq_of_ge + +lemma ideal.mem_infi_smul_pow_eq_bot_iff [is_noetherian_ring R] [hM : module.finite R M] (x : M) : + x ∈ (⨅ i : ℕ, I ^ i • ⊤ : submodule R M) ↔ ∃ r : I, (r : R) • x = x := +begin + let N := (⨅ i : ℕ, I ^ i • ⊤ : submodule R M), + have hN : ∀ k, (I.stable_filtration ⊤ ⊓ I.trivial_filtration N).N k = N, + { intro k, exact inf_eq_right.mpr ((infi_le _ k).trans $ le_of_eq $ by simp) }, + split, + { haveI := is_noetherian_of_fg_of_noetherian' hM.out, + obtain ⟨r, hr₁, hr₂⟩ := submodule.exists_mem_and_smul_eq_self_of_fg_of_le_smul I N + (is_noetherian.noetherian N) _, + { intro H, exact ⟨⟨r, hr₁⟩, hr₂ _ H⟩ }, + obtain ⟨k, hk⟩ := (I.stable_filtration_stable ⊤).inter_right (I.trivial_filtration N), + have := hk k (le_refl _), + rw [hN, hN] at this, + exact le_of_eq this.symm }, + { rintro ⟨r, eq⟩, + rw submodule.mem_infi, + intro i, + induction i with i hi, + { simp }, + { rw [nat.succ_eq_one_add, pow_add, ← smul_smul, pow_one, ← eq], + exact submodule.smul_mem_smul r.prop hi } } +end + +lemma ideal.infi_pow_smul_eq_bot_of_local_ring [is_noetherian_ring R] [local_ring R] + [module.finite R M] (h : I ≠ ⊤) : + (⨅ i : ℕ, I ^ i • ⊤ : submodule R M) = ⊥ := +begin + rw eq_bot_iff, + intros x hx, + obtain ⟨r, hr⟩ := (I.mem_infi_smul_pow_eq_bot_iff x).mp hx, + have := local_ring.is_unit_one_sub_self_of_mem_nonunits _ (local_ring.le_maximal_ideal h r.prop), + apply this.smul_left_cancel.mp, + swap, { apply_instance }, + simp [sub_smul, hr], +end + +/-- **Krull's intersection theorem** for noetherian local rings. -/ +lemma ideal.infi_pow_eq_bot_of_local_ring [is_noetherian_ring R] [local_ring R] (h : I ≠ ⊤) : + (⨅ i : ℕ, I ^ i) = ⊥ := +begin + convert I.infi_pow_smul_eq_bot_of_local_ring h, + ext i, + rw [smul_eq_mul, ← ideal.one_eq_top, mul_one], + apply_instance, +end + +/-- **Krull's intersection theorem** for noetherian domains. -/ +lemma ideal.infi_pow_eq_bot_of_is_domain [is_noetherian_ring R] [is_domain R] (h : I ≠ ⊤) : + (⨅ i : ℕ, I ^ i) = ⊥ := +begin + rw eq_bot_iff, + intros x hx, + by_contra hx', + have := ideal.mem_infi_smul_pow_eq_bot_iff I x, + simp_rw [smul_eq_mul, ← ideal.one_eq_top, mul_one] at this, + obtain ⟨r, hr⟩ := this.mp hx, + have := mul_right_cancel₀ hx' (hr.trans (one_mul x).symm), + exact I.eq_top_iff_one.not.mp h (this ▸ r.prop), +end diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index 9066a23fb6f53..d5f7472f8141b 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -314,6 +314,23 @@ begin exact submodule.span_eq_restrict_scalars R S M X h end +lemma fg.stablizes_of_supr_eq {M' : submodule R M} (hM' : M'.fg) + (N : ℕ →o submodule R M) (H : supr N = M') : ∃ n, M' = N n := +begin + obtain ⟨S, hS⟩ := hM', + have : ∀ s : S, ∃ n, (s : M) ∈ N n := + λ s, (submodule.mem_supr_of_chain N s).mp + (by { rw [H, ← hS], exact submodule.subset_span s.2 }), + choose f hf, + use S.attach.sup f, + apply le_antisymm, + { conv_lhs { rw ← hS }, + rw submodule.span_le, + intros s hs, + exact N.2 (finset.le_sup $ S.mem_attach ⟨s, hs⟩) (hf _) }, + { rw ← H, exact le_supr _ _ } +end + /-- Finitely generated submodules are precisely compact elements in the submodule lattice. -/ theorem fg_iff_compact (s : submodule R M) : s.fg ↔ complete_lattice.is_compact_element s := begin From a4d242389fcdfcb93f0a43601fe7f9fb23f4d2b0 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 5 Oct 2022 09:32:57 +0000 Subject: [PATCH 569/828] feat(geometry/euclidean/basic): bundled spheres (#16184) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define a bundled `sphere` structure, for various uses in Euclidean geometry where it's convenient to pass around `center` and `radius` together. Some basic API is set up for this structure, including `sphere` versions of a few existing lemmas that could naturally be expressed in that way. The construction of `circumcenter` and `circumradius` is also changed to pass around a `sphere` instead of using `P × ℝ`. It is likely that other existing lemmas can usefully have bundled sphere versions added in followups. Certainly there are plenty of other definitions and results about spheres that can usefully be built up on top of this basic API. Notes: * As with `cospherical`, the definition and some of the most basic lemmas don't actually need anything more than the metric space structure. `sphere` is defined alongside `cospherical`, but it would also be reasonable to define both in some metric space file. In that case, the name of `sphere` would need to change to avoid conflicts with the existing `metric.sphere` (which is part of a family of unbundled definitions with `metric.ball` and `metric.closed_ball`, so should probably remain as-is). * The definition doesn't include any non-degeneracy conditions, so avoiding the need for users to prove such conditions when constructing a `sphere` for a lemma that doesn't need them. Note that the base case for the induction constructing the circumsphere uses a radius of zero. * I haven't forgotten the discussion in #4088 of simplifying the proof of `eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two` using bundled spheres, a definition of the radical subspace and a proof that a one-dimensional sphere has at most two points (so ending up proving the unbundled lemma using the bundled one rather than vice versa), but I think quite a lot more API about power of a point and radical subspaces would still need adding before that could be done. --- src/geometry/euclidean/basic.lean | 108 ++++++++++++++++++++++ src/geometry/euclidean/circumcenter.lean | 113 ++++++++++++----------- src/geometry/euclidean/inversion.lean | 2 +- 3 files changed, 166 insertions(+), 57 deletions(-) diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index d7a4d6bbc8e23..bbc3264657584 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -1315,6 +1315,69 @@ reflection_orthogonal_vadd hp₁ omit V +variables (P) + +/-- A `sphere P` bundles a `center` and `radius`. This definition does not require the radius to +be positive; that should be given as a hypothesis to lemmas that require it. -/ +@[ext] structure sphere := +(center : P) +(radius : ℝ) + +variables {P} + +instance [nonempty P] : nonempty (sphere P) := ⟨⟨classical.arbitrary P, 0⟩⟩ + +instance : has_coe (sphere P) (set P) := ⟨λ s, metric.sphere s.center s.radius⟩ +instance : has_mem P (sphere P) := ⟨λ p s, p ∈ (s : set P)⟩ + +lemma sphere.mk_center (c : P) (r : ℝ) : (⟨c, r⟩ : sphere P).center = c := rfl + +lemma sphere.mk_radius (c : P) (r : ℝ) : (⟨c, r⟩ : sphere P).radius = r := rfl + +@[simp] lemma sphere.mk_center_radius (s : sphere P) : (⟨s.center, s.radius⟩ : sphere P) = s := +by ext; refl + +lemma sphere.coe_def (s : sphere P) : (s : set P) = metric.sphere s.center s.radius := rfl + +@[simp] lemma sphere.coe_mk (c : P) (r : ℝ) : ↑(⟨c, r⟩ : sphere P) = metric.sphere c r := rfl + +@[simp] lemma sphere.mem_coe {p : P} {s : sphere P} : p ∈ (s : set P) ↔ p ∈ s := iff.rfl + +lemma mem_sphere {p : P} {s : sphere P} : p ∈ s ↔ dist p s.center = s.radius := iff.rfl + +lemma mem_sphere' {p : P} {s : sphere P} : p ∈ s ↔ dist s.center p = s.radius := +metric.mem_sphere' + +lemma subset_sphere {ps : set P} {s : sphere P} : ps ⊆ s ↔ ∀ p ∈ ps, p ∈ s := iff.rfl + +lemma dist_of_mem_subset_sphere {p : P} {ps : set P} {s : sphere P} (hp : p ∈ ps) + (hps : ps ⊆ (s : set P)) : dist p s.center = s.radius := +mem_sphere.1 (sphere.mem_coe.1 (set.mem_of_mem_of_subset hp hps)) + +lemma dist_of_mem_subset_mk_sphere {p c : P} {ps : set P} {r : ℝ} (hp : p ∈ ps) + (hps : ps ⊆ ↑(⟨c, r⟩ : sphere P)) : dist p c = r := +dist_of_mem_subset_sphere hp hps + +lemma sphere.ne_iff {s₁ s₂ : sphere P} : + s₁ ≠ s₂ ↔ s₁.center ≠ s₂.center ∨ s₁.radius ≠ s₂.radius := +by rw [←not_and_distrib, ←sphere.ext_iff] + +lemma sphere.center_eq_iff_eq_of_mem {s₁ s₂ : sphere P} {p : P} (hs₁ : p ∈ s₁) (hs₂ : p ∈ s₂) : + s₁.center = s₂.center ↔ s₁ = s₂ := +begin + refine ⟨λ h, sphere.ext _ _ h _, λ h, h ▸ rfl⟩, + rw mem_sphere at hs₁ hs₂, + rw [←hs₁, ←hs₂, h] +end + +lemma sphere.center_ne_iff_ne_of_mem {s₁ s₂ : sphere P} {p : P} (hs₁ : p ∈ s₁) (hs₂ : p ∈ s₂) : + s₁.center ≠ s₂.center ↔ s₁ ≠ s₂ := +(sphere.center_eq_iff_eq_of_mem hs₁ hs₂).not + +lemma dist_center_eq_dist_center_of_mem_sphere {p₁ p₂ : P} {s : sphere P} (hp₁ : p₁ ∈ s) + (hp₂ : p₂ ∈ s) : dist p₁ s.center = dist p₂ s.center := +by rw [mem_sphere.1 hp₁, mem_sphere.1 hp₂] + /-- A set of points is cospherical if they are equidistant from some point. In two dimensions, this is the same thing as being concyclic. -/ @@ -1326,6 +1389,21 @@ lemma cospherical_def (ps : set P) : cospherical ps ↔ ∃ (center : P) (radius : ℝ), ∀ p ∈ ps, dist p center = radius := iff.rfl +/-- A set of points is cospherical if and only if they lie in some sphere. -/ +lemma cospherical_iff_exists_sphere {ps : set P} : + cospherical ps ↔ ∃ s : sphere P, ps ⊆ (s : set P) := +begin + refine ⟨λ h, _, λ h, _⟩, + { rcases h with ⟨c, r, h⟩, + exact ⟨⟨c, r⟩, h⟩ }, + { rcases h with ⟨s, h⟩, + exact ⟨s.center, s.radius, h⟩ } +end + +/-- The set of points in a sphere is cospherical. -/ +lemma sphere.cospherical (s : sphere P) : cospherical (s : set P) := +cospherical_iff_exists_sphere.2 ⟨s, set.subset.rfl⟩ + /-- A subset of a cospherical set is cospherical. -/ lemma cospherical_subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (hc : cospherical ps₂) : cospherical ps₁ := @@ -1410,4 +1488,34 @@ begin exact (dec_trivial : (1 : fin 3) ≠ 2) (hfi hf12) end +/-- Suppose that `p₁` and `p₂` lie in spheres `s₁` and `s₂`. Then the vector between the centers +of those spheres is orthogonal to that between `p₁` and `p₂`; this is a version of +`inner_vsub_vsub_of_dist_eq_of_dist_eq` for bundled spheres. (In two dimensions, this says that +the diagonals of a kite are orthogonal.) -/ +lemma inner_vsub_vsub_of_mem_sphere_of_mem_sphere {p₁ p₂ : P} {s₁ s₂ : sphere P} + (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) : + ⟪s₂.center -ᵥ s₁.center, p₂ -ᵥ p₁⟫ = 0 := +inner_vsub_vsub_of_dist_eq_of_dist_eq (dist_center_eq_dist_center_of_mem_sphere hp₁s₁ hp₂s₁) + (dist_center_eq_dist_center_of_mem_sphere hp₁s₂ hp₂s₂) + +/-- Two spheres intersect in at most two points in a two-dimensional subspace containing their +centers; this is a version of `eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two` for bundled +spheres. -/ +lemma eq_of_mem_sphere_of_mem_sphere_of_mem_of_finrank_eq_two {s : affine_subspace ℝ P} + [finite_dimensional ℝ s.direction] (hd : finrank ℝ s.direction = 2) {s₁ s₂ : sphere P} + {p₁ p₂ p : P} (hs₁ : s₁.center ∈ s) (hs₂ : s₂.center ∈ s) (hp₁s : p₁ ∈ s) (hp₂s : p₂ ∈ s) + (hps : p ∈ s) (hs : s₁ ≠ s₂) (hp : p₁ ≠ p₂) (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hps₁ : p ∈ s₁) + (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) (hps₂ : p ∈ s₂) : p = p₁ ∨ p = p₂ := +eq_of_dist_eq_of_dist_eq_of_mem_of_finrank_eq_two hd hs₁ hs₂ hp₁s hp₂s hps + ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ + +/-- Two spheres intersect in at most two points in two-dimensional space; this is a version of +`eq_of_dist_eq_of_dist_eq_of_finrank_eq_two` for bundled spheres. -/ +lemma eq_of_mem_sphere_of_mem_sphere_of_finrank_eq_two [finite_dimensional ℝ V] + (hd : finrank ℝ V = 2) {s₁ s₂ : sphere P} {p₁ p₂ p : P} (hs : s₁ ≠ s₂) (hp : p₁ ≠ p₂) + (hp₁s₁ : p₁ ∈ s₁) (hp₂s₁ : p₂ ∈ s₁) (hps₁ : p ∈ s₁) (hp₁s₂ : p₁ ∈ s₂) (hp₂s₂ : p₂ ∈ s₂) + (hps₂ : p ∈ s₂) : p = p₁ ∨ p = p₂ := +eq_of_dist_eq_of_dist_eq_of_finrank_eq_two hd ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) + hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ + end euclidean_geometry diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index 69ada84e2f8bf..9d1f9f40241c2 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -95,21 +95,21 @@ subspace with `p` added. -/ lemma exists_unique_dist_eq_of_insert {s : affine_subspace ℝ P} [complete_space s.direction] {ps : set P} (hnps : ps.nonempty) {p : P} (hps : ps ⊆ s) (hp : p ∉ s) - (hu : ∃! cccr : (P × ℝ), cccr.fst ∈ s ∧ ∀ p1 ∈ ps, dist p1 cccr.fst = cccr.snd) : - ∃! cccr₂ : (P × ℝ), cccr₂.fst ∈ affine_span ℝ (insert p (s : set P)) ∧ - ∀ p1 ∈ insert p ps, dist p1 cccr₂.fst = cccr₂.snd := + (hu : ∃! cs : sphere P, cs.center ∈ s ∧ ps ⊆ (cs : set P)) : + ∃! cs₂ : sphere P, cs₂.center ∈ affine_span ℝ (insert p (s : set P)) ∧ + (insert p ps) ⊆ (cs₂ : set P) := begin haveI : nonempty s := set.nonempty.to_subtype (hnps.mono hps), rcases hu with ⟨⟨cc, cr⟩, ⟨hcc, hcr⟩, hcccru⟩, - simp only [prod.fst, prod.snd] at hcc hcr hcccru, + simp only at hcc hcr hcccru, let x := dist cc (orthogonal_projection s p), let y := dist p (orthogonal_projection s p), have hy0 : y ≠ 0 := dist_orthogonal_projection_ne_zero_of_not_mem hp, let ycc₂ := (x * x + y * y - cr * cr) / (2 * y), let cc₂ := (ycc₂ / y) • (p -ᵥ orthogonal_projection s p : V) +ᵥ cc, let cr₂ := real.sqrt (cr * cr + ycc₂ * ycc₂), - use (cc₂, cr₂), - simp only [prod.fst, prod.snd], + use ⟨cc₂, cr₂⟩, + simp only, have hpo : p = (1 : ℝ) • (p -ᵥ orthogonal_projection s p : V) +ᵥ orthogonal_projection s p, { simp }, split, @@ -120,7 +120,7 @@ begin (vsub_mem_vector_span ℝ (set.mem_insert _ _) (set.mem_insert_of_mem _ (orthogonal_projection_mem _))) }, { intros p1 hp1, - rw [←mul_self_inj_of_nonneg dist_nonneg (real.sqrt_nonneg _), + rw [sphere.mem_coe, mem_sphere, ←mul_self_inj_of_nonneg dist_nonneg (real.sqrt_nonneg _), real.mul_self_sqrt (add_nonneg (mul_self_nonneg _) (mul_self_nonneg _))], cases hp1, { rw hp1, @@ -134,39 +134,39 @@ begin { rw [dist_sq_eq_dist_orthogonal_projection_sq_add_dist_orthogonal_projection_sq _ (hps hp1), orthogonal_projection_vadd_smul_vsub_orthogonal_projection _ _ hcc, subtype.coe_mk, - hcr _ hp1, dist_eq_norm_vsub V cc₂ cc, vadd_vsub, norm_smul, ←dist_eq_norm_vsub V, + dist_of_mem_subset_mk_sphere hp1 hcr, + dist_eq_norm_vsub V cc₂ cc, vadd_vsub, norm_smul, ←dist_eq_norm_vsub V, real.norm_eq_abs, abs_div, abs_of_nonneg dist_nonneg, div_mul_cancel _ hy0, abs_mul_abs_self] } } }, { rintros ⟨cc₃, cr₃⟩ ⟨hcc₃, hcr₃⟩, - simp only [prod.fst, prod.snd] at hcc₃ hcr₃, + simp only at hcc₃ hcr₃, obtain ⟨t₃, cc₃', hcc₃', hcc₃''⟩ : ∃ (r : ℝ) (p0 : P) (hp0 : p0 ∈ s), cc₃ = r • (p -ᵥ ↑((orthogonal_projection s) p)) +ᵥ p0, { rwa mem_affine_span_insert_iff (orthogonal_projection_mem p) at hcc₃ }, have hcr₃' : ∃ r, ∀ p1 ∈ ps, dist p1 cc₃ = r := - ⟨cr₃, λ p1 hp1, hcr₃ p1 (set.mem_insert_of_mem _ hp1)⟩, + ⟨cr₃, λ p1 hp1, dist_of_mem_subset_mk_sphere (set.mem_insert_of_mem _ hp1) hcr₃⟩, rw [exists_dist_eq_iff_exists_dist_orthogonal_projection_eq hps cc₃, hcc₃'', orthogonal_projection_vadd_smul_vsub_orthogonal_projection _ _ hcc₃'] at hcr₃', cases hcr₃' with cr₃' hcr₃', - have hu := hcccru (cc₃', cr₃'), - simp only [prod.fst, prod.snd] at hu, + have hu := hcccru ⟨cc₃', cr₃'⟩, + simp only at hu, replace hu := hu ⟨hcc₃', hcr₃'⟩, - rw prod.ext_iff at hu, - simp only [prod.fst, prod.snd] at hu, cases hu with hucc hucr, substs hucc hucr, have hcr₃val : cr₃ = real.sqrt (cr₃' * cr₃' + (t₃ * y) * (t₃ * y)), { cases hnps with p0 hp0, have h' : ↑(⟨cc₃', hcc₃'⟩ : s) = cc₃' := rfl, - rw [←hcr₃ p0 (set.mem_insert_of_mem _ hp0), hcc₃'', + rw [←dist_of_mem_subset_mk_sphere (set.mem_insert_of_mem _ hp0) hcr₃, hcc₃'', ←mul_self_inj_of_nonneg dist_nonneg (real.sqrt_nonneg _), real.mul_self_sqrt (add_nonneg (mul_self_nonneg _) (mul_self_nonneg _)), dist_sq_eq_dist_orthogonal_projection_sq_add_dist_orthogonal_projection_sq _ (hps hp0), - orthogonal_projection_vadd_smul_vsub_orthogonal_projection _ _ hcc₃', h', hcr p0 hp0, + orthogonal_projection_vadd_smul_vsub_orthogonal_projection _ _ hcc₃', h', + dist_of_mem_subset_mk_sphere hp0 hcr, dist_eq_norm_vsub V _ cc₃', vadd_vsub, norm_smul, ←dist_eq_norm_vsub V p, real.norm_eq_abs, ←mul_assoc, mul_comm _ (|t₃|), ←mul_assoc, abs_mul_abs_self], ring }, - replace hcr₃ := hcr₃ p (set.mem_insert _ _), + replace hcr₃ := dist_of_mem_subset_mk_sphere (set.mem_insert _ _) hcr₃, rw [hpo, hcc₃'', hcr₃val, ←mul_self_inj_of_nonneg dist_nonneg (real.sqrt_nonneg _), dist_sq_smul_orthogonal_vadd_smul_orthogonal_vadd (orthogonal_projection_mem p) hcc₃' _ _ @@ -193,8 +193,7 @@ there is a unique (circumcenter, circumradius) pair for those points in the affine subspace they span. -/ lemma _root_.affine_independent.exists_unique_dist_eq {ι : Type*} [hne : nonempty ι] [fintype ι] {p : ι → P} (ha : affine_independent ℝ p) : - ∃! cccr : (P × ℝ), cccr.fst ∈ affine_span ℝ (set.range p) ∧ - ∀ i, dist (p i) cccr.fst = cccr.snd := + ∃! cs : sphere P, cs.center ∈ affine_span ℝ (set.range p) ∧ set.range p ⊆ (cs : set P) := begin unfreezingI { induction hn : fintype.card ι with m hm generalizing ι }, { exfalso, @@ -205,21 +204,17 @@ begin { rw fintype.card_eq_one_iff at hn, cases hn with i hi, haveI : unique ι := ⟨⟨i⟩, hi⟩, - use (p i, 0), - simp only [prod.fst, prod.snd, set.range_unique, affine_subspace.mem_affine_span_singleton], + use ⟨p i, 0⟩, + simp only [set.range_unique, affine_subspace.mem_affine_span_singleton], split, - { simp_rw [hi default], - use rfl, - intro i1, - rw hi i1, - exact dist_self _ }, + { simp_rw [hi default, set.singleton_subset_iff, sphere.mem_coe, mem_sphere, dist_self], + exact ⟨rfl, rfl⟩ }, { rintros ⟨cc, cr⟩, - simp only [prod.fst, prod.snd], + simp only, rintros ⟨rfl, hdist⟩, - rw hi default, - congr', - rw ←hdist default, - exact dist_self _ } }, + simp_rw [set.singleton_subset_iff, sphere.mem_coe, mem_sphere, dist_self] at hdist, + rw [hi default, hdist], + exact ⟨rfl, rfl⟩ } }, { have i := hne.some, let ι2 := {x // x ≠ i}, have hc : fintype.card ι2 = m + 1, @@ -239,13 +234,7 @@ begin rw [←set.image_eq_range, ←set.image_univ, ←set.image_insert_eq], congr' with j, simp [classical.em] }, - change ∃! (cccr : P × ℝ), (_ ∧ ∀ i2, (λ q, dist q cccr.fst = cccr.snd) (p i2)), - conv { congr, funext, conv { congr, skip, rw ←set.forall_range_iff } }, - dsimp only, - rw hr, - change ∃! (cccr : P × ℝ), (_ ∧ ∀ (i2 : ι2), (λ q, dist q cccr.fst = cccr.snd) (p i2)) at hm, - conv at hm { congr, funext, conv { congr, skip, rw ←set.forall_range_iff } }, - rw ←affine_span_insert_affine_span, + rw [hr, ←affine_span_insert_affine_span], refine exists_unique_dist_eq_of_insert (set.range_nonempty _) (subset_span_points ℝ _) @@ -269,36 +258,46 @@ variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V -/-- The pair (circumcenter, circumradius) of a simplex. -/ -def circumcenter_circumradius {n : ℕ} (s : simplex ℝ P n) : (P × ℝ) := +/-- The circumsphere of a simplex. -/ +def circumsphere {n : ℕ} (s : simplex ℝ P n) : sphere P := s.independent.exists_unique_dist_eq.some -/-- The property satisfied by the (circumcenter, circumradius) pair. -/ -lemma circumcenter_circumradius_unique_dist_eq {n : ℕ} (s : simplex ℝ P n) : - (s.circumcenter_circumradius.fst ∈ affine_span ℝ (set.range s.points) ∧ - ∀ i, dist (s.points i) s.circumcenter_circumradius.fst = s.circumcenter_circumradius.snd) ∧ - (∀ cccr : (P × ℝ), (cccr.fst ∈ affine_span ℝ (set.range s.points) ∧ - ∀ i, dist (s.points i) cccr.fst = cccr.snd) → cccr = s.circumcenter_circumradius) := +/-- The property satisfied by the circumsphere. -/ +lemma circumsphere_unique_dist_eq {n : ℕ} (s : simplex ℝ P n) : + (s.circumsphere.center ∈ affine_span ℝ (set.range s.points) ∧ + set.range s.points ⊆ s.circumsphere) ∧ + (∀ cs : sphere P, (cs.center ∈ affine_span ℝ (set.range s.points) ∧ + set.range s.points ⊆ cs → cs = s.circumsphere)) := s.independent.exists_unique_dist_eq.some_spec /-- The circumcenter of a simplex. -/ def circumcenter {n : ℕ} (s : simplex ℝ P n) : P := -s.circumcenter_circumradius.fst +s.circumsphere.center /-- The circumradius of a simplex. -/ def circumradius {n : ℕ} (s : simplex ℝ P n) : ℝ := -s.circumcenter_circumradius.snd +s.circumsphere.radius + +/-- The center of the circumsphere is the circumcenter. -/ +@[simp] lemma circumsphere_center {n : ℕ} (s : simplex ℝ P n) : + s.circumsphere.center = s.circumcenter := +rfl + +/-- The radius of the circumsphere is the circumradius. -/ +@[simp] lemma circumsphere_radius {n : ℕ} (s : simplex ℝ P n) : + s.circumsphere.radius = s.circumradius := +rfl /-- The circumcenter lies in the affine span. -/ lemma circumcenter_mem_affine_span {n : ℕ} (s : simplex ℝ P n) : s.circumcenter ∈ affine_span ℝ (set.range s.points) := -s.circumcenter_circumradius_unique_dist_eq.1.1 +s.circumsphere_unique_dist_eq.1.1 /-- All points have distance from the circumcenter equal to the circumradius. -/ -@[simp] lemma dist_circumcenter_eq_circumradius {n : ℕ} (s : simplex ℝ P n) : - ∀ i, dist (s.points i) s.circumcenter = s.circumradius := -s.circumcenter_circumradius_unique_dist_eq.1.2 +@[simp] lemma dist_circumcenter_eq_circumradius {n : ℕ} (s : simplex ℝ P n) (i : fin (n + 1)) : + dist (s.points i) s.circumcenter = s.circumradius := +dist_of_mem_subset_sphere (set.mem_range_self _) s.circumsphere_unique_dist_eq.1.2 /-- All points have distance to the circumcenter equal to the circumradius. -/ @@ -316,8 +315,9 @@ lemma eq_circumcenter_of_dist_eq {n : ℕ} (s : simplex ℝ P n) {p : P} (hp : p ∈ affine_span ℝ (set.range s.points)) {r : ℝ} (hr : ∀ i, dist (s.points i) p = r) : p = s.circumcenter := begin - have h := s.circumcenter_circumradius_unique_dist_eq.2 (p, r), - simp only [hp, hr, forall_const, eq_self_iff_true, and_self, prod.ext_iff] at h, + have h := s.circumsphere_unique_dist_eq.2 ⟨p, r⟩, + simp only [hp, hr, forall_const, eq_self_iff_true, subset_sphere, sphere.ext_iff, + set.forall_range_iff, mem_sphere, true_and] at h, exact h.1 end @@ -327,8 +327,9 @@ lemma eq_circumradius_of_dist_eq {n : ℕ} (s : simplex ℝ P n) {p : P} (hp : p ∈ affine_span ℝ (set.range s.points)) {r : ℝ} (hr : ∀ i, dist (s.points i) p = r) : r = s.circumradius := begin - have h := s.circumcenter_circumradius_unique_dist_eq.2 (p, r), - simp only [hp, hr, forall_const, eq_self_iff_true, and_self, prod.ext_iff] at h, + have h := s.circumsphere_unique_dist_eq.2 ⟨p, r⟩, + simp only [hp, hr, forall_const, eq_self_iff_true, subset_sphere, sphere.ext_iff, + set.forall_range_iff, mem_sphere, true_and] at h, exact h.2 end diff --git a/src/geometry/euclidean/inversion.lean b/src/geometry/euclidean/inversion.lean index 0ad5c99bb4357..e9dd02b11c4b0 100644 --- a/src/geometry/euclidean/inversion.lean +++ b/src/geometry/euclidean/inversion.lean @@ -49,7 +49,7 @@ begin rwa [dist_ne_zero] } end -lemma inversion_of_mem_sphere (h : x ∈ sphere c R) : inversion c R x = x := +lemma inversion_of_mem_sphere (h : x ∈ metric.sphere c R) : inversion c R x = x := h.out ▸ inversion_dist_center c x /-- Distance from the image of a point under inversion to the center. This formula accidentally From 5abbfc939f22d5c07569a0b2382e6e6622a1922a Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 5 Oct 2022 09:32:58 +0000 Subject: [PATCH 570/828] feat(logic/equiv/fin): `succ_above` as an equivalence (#16278) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add a version of `succ_above` that's a bundled order isomorphism between `fin n` and `{x : fin (n + 1) // x ≠ p}`. --- src/logic/equiv/fin.lean | 55 ++++++++++++++++++++++++++++++++++++++++ 1 file changed, 55 insertions(+) diff --git a/src/logic/equiv/fin.lean b/src/logic/equiv/fin.lean index 596142c38943b..6beacf7bd6b2b 100644 --- a/src/logic/equiv/fin.lean +++ b/src/logic/equiv/fin.lean @@ -176,6 +176,61 @@ fin_succ_equiv_symm_some m lemma fin_succ_equiv'_zero {n : ℕ} : fin_succ_equiv' (0 : fin (n + 1)) = fin_succ_equiv n := rfl +lemma fin_succ_equiv'_last_apply {n : ℕ} {i : fin (n + 1)} (h : i ≠ fin.last n) : + fin_succ_equiv' (fin.last n) i = + fin.cast_lt i (lt_of_le_of_ne (fin.le_last _) (fin.coe_injective.ne_iff.2 h) : ↑i < n) := +begin + have h' : ↑i < n := lt_of_le_of_ne (fin.le_last _) (fin.coe_injective.ne_iff.2 h), + conv_lhs { rw ←fin.cast_succ_cast_lt i h' }, + convert fin_succ_equiv'_below _, + rw fin.cast_succ_cast_lt i h', + exact h' +end + +lemma fin_succ_equiv'_ne_last_apply {i j : fin (n + 1)} (hi : i ≠ fin.last n) + (hj : j ≠ i) : + fin_succ_equiv' i j = (i.cast_lt (lt_of_le_of_ne (fin.le_last _) + (fin.coe_injective.ne_iff.2 hi) : ↑i < n)).pred_above j := +begin + rw [fin.pred_above], + have hi' : ↑i < n := lt_of_le_of_ne (fin.le_last _) (fin.coe_injective.ne_iff.2 hi), + rcases hj.lt_or_lt with hij | hij, + { simp only [hij.not_lt, fin.cast_succ_cast_lt, not_false_iff, dif_neg], + convert fin_succ_equiv'_below _, + { simp }, + { exact hij } }, + { simp only [hij, fin.cast_succ_cast_lt, dif_pos], + convert fin_succ_equiv'_above _, + { simp }, + { simp [fin.le_cast_succ_iff, hij] } } +end + +/-- `succ_above` as an order isomorphism between `fin n` and `{x : fin (n + 1) // x ≠ p}`. -/ +def fin_succ_above_equiv (p : fin (n + 1)) : fin n ≃o {x : fin (n + 1) // x ≠ p} := +{ map_rel_iff' := λ _ _, p.succ_above.map_rel_iff', + ..equiv.option_subtype p ⟨(fin_succ_equiv' p).symm, rfl⟩ } + +lemma fin_succ_above_equiv_apply (p : fin (n + 1)) (i : fin n) : + fin_succ_above_equiv p i = ⟨p.succ_above i, p.succ_above_ne i⟩ := +rfl + +lemma fin_succ_above_equiv_symm_apply_last (x : {x : fin (n + 1) // x ≠ fin.last n}) : + (fin_succ_above_equiv (fin.last n)).symm x = + fin.cast_lt (x : fin (n + 1)) + (lt_of_le_of_ne (fin.le_last _) (fin.coe_injective.ne_iff.2 x.property)) := +begin + rw [←option.some_inj, ←option.coe_def], + simpa [fin_succ_above_equiv, order_iso.symm] using fin_succ_equiv'_last_apply x.property, +end + +lemma fin_succ_above_equiv_symm_apply_ne_last {p : fin (n + 1)} (h : p ≠ fin.last n) + (x : {x : fin (n + 1) // x ≠ p}) : (fin_succ_above_equiv p).symm x = + (p.cast_lt (lt_of_le_of_ne (fin.le_last _) (fin.coe_injective.ne_iff.2 h))).pred_above x := +begin + rw [←option.some_inj, ←option.coe_def], + simpa [fin_succ_above_equiv, order_iso.symm] using fin_succ_equiv'_ne_last_apply h x.property +end + /-- `equiv` between `fin (n + 1)` and `option (fin n)` sending `fin.last n` to `none` -/ def fin_succ_equiv_last {n : ℕ} : fin (n + 1) ≃ option (fin n) := fin_succ_equiv' (fin.last n) From 2eac7392000df99bb96aea05231e7425d6c66c73 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 5 Oct 2022 09:32:59 +0000 Subject: [PATCH 571/828] feat(linear_algebra/affine_space/finite_dimensional): `collinear` lemmas, implicit arguments (#16332) Add more lemmas about `collinear`, and change existing `iff` lemmas about `collinear` to follow usual mathlib conventions about implicit arguments for such lemmas. --- src/analysis/convex/between.lean | 2 +- src/geometry/euclidean/basic.lean | 2 +- .../affine_space/finite_dimensional.lean | 35 ++++++++++++++++--- 3 files changed, 32 insertions(+), 7 deletions(-) diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index 86e7d20ebbcf0..638f58f508156 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -519,7 +519,7 @@ end lemma collinear.wbtw_or_wbtw_or_wbtw {x y z : P} (h : collinear R ({x, y, z} : set P)) : wbtw R x y z ∨ wbtw R y z x ∨ wbtw R z x y := begin - rw collinear_iff_of_mem R (set.mem_insert _ _) at h, + rw collinear_iff_of_mem (set.mem_insert _ _) at h, rcases h with ⟨v, h⟩, simp_rw [set.mem_insert_iff, set.mem_singleton_iff] at h, have hy := h y (or.inr (or.inl rfl)), diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index bbc3264657584..7b478e10299b6 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -1455,7 +1455,7 @@ lemma cospherical.affine_independent {s : set P} (hs : cospherical s) {p : fin 3 begin rw affine_independent_iff_not_collinear, intro hc, - rw collinear_iff_of_mem ℝ (set.mem_range_self (0 : fin 3)) at hc, + rw collinear_iff_of_mem (set.mem_range_self (0 : fin 3)) at hc, rcases hc with ⟨v, hv⟩, rw set.forall_range_iff at hv, have hv0 : v ≠ 0, diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index ce6230126093a..35208c1d243d4 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -309,10 +309,12 @@ def collinear (s : set P) : Prop := module.rank k (vector_span k s) ≤ 1 lemma collinear_iff_dim_le_one (s : set P) : collinear k s ↔ module.rank k (vector_span k s) ≤ 1 := iff.rfl +variables {k} + /-- A set of points, whose `vector_span` is finite-dimensional, is collinear if and only if their `vector_span` has dimension at most `1`. -/ -lemma collinear_iff_finrank_le_one (s : set P) [finite_dimensional k (vector_span k s)] : +lemma collinear_iff_finrank_le_one {s : set P} [finite_dimensional k (vector_span k s)] : collinear k s ↔ finrank k (vector_span k s) ≤ 1 := begin have h := collinear_iff_dim_le_one k s, @@ -320,7 +322,24 @@ begin exact_mod_cast h end -variables (P) +alias collinear_iff_finrank_le_one ↔ collinear.finrank_le_one _ + +/-- A subset of a collinear set is collinear. -/ +lemma collinear.subset {s₁ s₂ : set P} (hs : s₁ ⊆ s₂) (h : collinear k s₂) : collinear k s₁ := +(dim_le_of_submodule (vector_span k s₁) (vector_span k s₂) (vector_span_mono k hs)).trans h + +/-- The `vector_span` of collinear points is finite-dimensional. -/ +lemma collinear.finite_dimensional_vector_span {s : set P} (h : collinear k s) : + finite_dimensional k (vector_span k s) := +is_noetherian.iff_fg.1 + (is_noetherian.iff_dim_lt_aleph_0.2 (lt_of_le_of_lt h cardinal.one_lt_aleph_0)) + +/-- The direction of the affine span of collinear points is finite-dimensional. -/ +lemma collinear.finite_dimensional_direction_affine_span {s : set P} (h : collinear k s) : + finite_dimensional k (affine_span k s).direction := +(direction_affine_span k s).symm ▸ h.finite_dimensional_vector_span + +variables (k P) /-- The empty set is collinear. -/ lemma collinear_empty : collinear k (∅ : set P) := @@ -338,6 +357,8 @@ begin simp end +variables {k} + /-- Given a point `p₀` in a set of points, that set is collinear if and only if the points can all be expressed as multiples of the same vector, added to `p₀`. -/ @@ -377,7 +398,7 @@ lemma collinear_iff_exists_forall_eq_smul_vadd (s : set P) : begin rcases set.eq_empty_or_nonempty s with rfl | ⟨⟨p₁, hp₁⟩⟩, { simp [collinear_empty] }, - { rw collinear_iff_of_mem k hp₁, + { rw collinear_iff_of_mem hp₁, split, { exact λ h, ⟨p₁, h⟩ }, { rintros ⟨p, v, hv⟩, @@ -389,6 +410,8 @@ begin simp [vadd_vadd, ←add_smul] } } end +variables (k) + /-- Two points are collinear. -/ lemma collinear_pair (p₁ p₂ : P) : collinear k ({p₁, p₂} : set P) := begin @@ -403,16 +426,18 @@ begin simp [hp] } end +variables {k} + /-- Three points are affinely independent if and only if they are not collinear. -/ -lemma affine_independent_iff_not_collinear (p : fin 3 → P) : +lemma affine_independent_iff_not_collinear {p : fin 3 → P} : affine_independent k p ↔ ¬ collinear k (set.range p) := by rw [collinear_iff_finrank_le_one, affine_independent_iff_not_finrank_vector_span_le k p (fintype.card_fin 3)] /-- Three points are collinear if and only if they are not affinely independent. -/ -lemma collinear_iff_not_affine_independent (p : fin 3 → P) : +lemma collinear_iff_not_affine_independent {p : fin 3 → P} : collinear k (set.range p) ↔ ¬ affine_independent k p := by rw [collinear_iff_finrank_le_one, finrank_vector_span_le_iff_not_affine_independent k p (fintype.card_fin 3)] From 609b9346d68b839487bdfe305f68f31daddecca5 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 5 Oct 2022 09:33:02 +0000 Subject: [PATCH 572/828] feat(ring_theory/mv_polynomial/tower): add analogs of existing lemmas for `mv_polynomial` (#16412) --- src/ring_theory/mv_polynomial/tower.lean | 80 ++++++++++++++++++++++++ src/ring_theory/polynomial/tower.lean | 3 + 2 files changed, 83 insertions(+) create mode 100644 src/ring_theory/mv_polynomial/tower.lean diff --git a/src/ring_theory/mv_polynomial/tower.lean b/src/ring_theory/mv_polynomial/tower.lean new file mode 100644 index 0000000000000..0145d9f7a24b2 --- /dev/null +++ b/src/ring_theory/mv_polynomial/tower.lean @@ -0,0 +1,80 @@ +/- +Copyright (c) 2022 Yuyang Zhao. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yuyang Zhao +-/ + +import algebra.algebra.tower +import data.mv_polynomial.basic + +/-! +# Algebra towers for multivariate polynomial + +This file proves some basic results about the algebra tower structure for the type +`mv_polynomial σ R`. + +This structure itself is provided elsewhere as `mv_polynomial.is_scalar_tower` + +When you update this file, you can also try to make a corresponding update in +`ring_theory.polynomial.tower`. +-/ + +variables (R A B : Type*) {σ : Type*} + +namespace mv_polynomial + +section semiring +variables [comm_semiring R] [comm_semiring A] [comm_semiring B] +variables [algebra R A] [algebra A B] [algebra R B] +variables [is_scalar_tower R A B] + +variables {R B} + +theorem aeval_map_algebra_map (x : σ → B) (p : mv_polynomial σ R) : + aeval x (map (algebra_map R A) p) = aeval x p := +by rw [aeval_def, aeval_def, eval₂_map, is_scalar_tower.algebra_map_eq R A B] + +end semiring + +section comm_semiring +variables [comm_semiring R] [comm_semiring A] [comm_semiring B] +variables [algebra R A] [algebra A B] [algebra R B] [is_scalar_tower R A B] + +variables {R A} + +lemma aeval_algebra_map_apply (x : σ → A) (p : mv_polynomial σ R) : + aeval (algebra_map A B ∘ x) p = algebra_map A B (mv_polynomial.aeval x p) := +by rw [aeval_def, aeval_def, ← coe_eval₂_hom, ← coe_eval₂_hom, map_eval₂_hom, + ←is_scalar_tower.algebra_map_eq] + +lemma aeval_algebra_map_eq_zero_iff [no_zero_smul_divisors A B] [nontrivial B] + (x : σ → A) (p : mv_polynomial σ R) : + aeval (algebra_map A B ∘ x) p = 0 ↔ aeval x p = 0 := +by rw [aeval_algebra_map_apply, algebra.algebra_map_eq_smul_one, smul_eq_zero, + iff_false_intro (@one_ne_zero B _ _), or_false] + +lemma aeval_algebra_map_eq_zero_iff_of_injective + {x : σ → A} {p : mv_polynomial σ R} + (h : function.injective (algebra_map A B)) : + aeval (algebra_map A B ∘ x) p = 0 ↔ aeval x p = 0 := +by rw [aeval_algebra_map_apply, ← (algebra_map A B).map_zero, h.eq_iff] + +end comm_semiring + +end mv_polynomial + +namespace subalgebra + +open mv_polynomial + +section comm_semiring + +variables {R A} [comm_semiring R] [comm_semiring A] [algebra R A] + +@[simp] lemma mv_polynomial_aeval_coe (S : subalgebra R A) (x : σ → S) (p : mv_polynomial σ R) : + aeval (λ i, (x i : A)) p = aeval x p := +by convert aeval_algebra_map_apply A x p + +end comm_semiring + +end subalgebra diff --git a/src/ring_theory/polynomial/tower.lean b/src/ring_theory/polynomial/tower.lean index b1890b8312a18..56c6aabf810bf 100644 --- a/src/ring_theory/polynomial/tower.lean +++ b/src/ring_theory/polynomial/tower.lean @@ -13,6 +13,9 @@ import data.polynomial.algebra_map This file proves some basic results about the algebra tower structure for the type `polynomial R`. This structure itself is provided elsewhere as `polynomial.is_scalar_tower` + +When you update this file, you can also try to make a corresponding update in +`ring_theory.mv_polynomial.tower`. -/ open_locale polynomial From caf8c55f3442b4344f610c52e54b7babbc168d88 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 5 Oct 2022 09:33:03 +0000 Subject: [PATCH 573/828] =?UTF-8?q?feat(tactic/positivity):=20Handle=20`a?= =?UTF-8?q?=20=E2=89=A0=200`=20assumptions=20(#16529)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Make `positivity` handle `a ≠ 0` assumptions. This involves * adding a third constructor to `positivity.strictness`, that I've called `nonzero`. It is meant to hold proofs of `a ≠ 0`, not `0 ≠ a` * modifying the existing extensions to handle that new case. * changing `positivity.orelse'` to account for the fact that if we can prove `0 ≤ a` and `a ≠ 0` then we can prove `0 < a`. * making `compare_hyp` handle `eq` and `ne` hypotheses. ## Other changes * Introduce `tac1 ≤|≥ tac2` as notation for `positivity.orelse' tac1 tac2`. This has the same precedence as the `<|>` notation for `orelse`. * Make `positivity` extensions fail with more informative error messages. This is useless when running `positivity` alone but: * It clearly marks within the code what each failure means * The errors can show up when calling an extension directly. Not sure why you would do that, but you could. * Add a `has_to_format strictness` instance. This is mostly useful for debugging. --- src/algebra/order/floor.lean | 2 + src/algebra/order/smul.lean | 26 +- src/analysis/special_functions/pow.lean | 1 + src/data/real/ereal.lean | 37 +- src/data/real/hyperreal.lean | 9 +- src/data/real/nnreal.lean | 2 +- src/order/bounded_order.lean | 12 +- src/ring_theory/dedekind_domain/ideal.lean | 4 +- src/tactic/positivity.lean | 473 +++++++++++++++------ test/positivity.lean | 87 +++- 10 files changed, 494 insertions(+), 159 deletions(-) diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index 9b56e2ef2613c..53bd2d11f81a1 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -949,6 +949,7 @@ meta def positivity_floor : expr → tactic strictness match strictness_a with | positive p := nonnegative <$> mk_app ``int_floor_nonneg_of_pos [p] | nonnegative p := nonnegative <$> mk_app ``int_floor_nonneg [p] + | _ := failed end | e := pp e >>= fail ∘ format.bracket "The expression `" "` is not of the form `⌊a⌋`" @@ -969,6 +970,7 @@ meta def positivity_ceil : expr → tactic strictness match strictness_a with | positive p := positive <$> mk_app ``int_ceil_pos [p] | nonnegative p := nonnegative <$> mk_app ``int.ceil_nonneg [p] + | _ := failed end | e := pp e >>= fail ∘ format.bracket "The expression `" "` is not of the form `⌈a⌉₊` nor `⌈a⌉`" diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index b47859245c111..64806d2156e7b 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -256,6 +256,7 @@ variables {M} end linear_ordered_semifield namespace tactic +section ordered_smul variables [ordered_semiring R] [ordered_add_comm_monoid M] [smul_with_zero R M] [ordered_smul R M] {a : R} {b : M} @@ -265,13 +266,26 @@ smul_nonneg ha.le hb private lemma smul_nonneg_of_nonneg_of_pos (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ a • b := smul_nonneg ha hb.le +end ordered_smul + +section no_zero_smul_divisors +variables [has_zero R] [has_zero M] [has_smul R M] [no_zero_smul_divisors R M] {a : R} {b : M} + +private lemma smul_ne_zero_of_pos_of_ne_zero [preorder R] (ha : 0 < a) (hb : b ≠ 0) : a • b ≠ 0 := +smul_ne_zero ha.ne' hb + +private lemma smul_ne_zero_of_ne_zero_of_pos [preorder M] (ha : a ≠ 0) (hb : 0 < b) : a • b ≠ 0 := +smul_ne_zero ha hb.ne' + +end no_zero_smul_divisors + open positivity -/-- Extension for the `positivity` tactic: scalar multiplication is nonnegative if both sides are -nonnegative, and strictly positive if both sides are. -/ +/-- Extension for the `positivity` tactic: scalar multiplication is nonnegative/positive/nonzero if +both sides are. -/ @[positivity] meta def positivity_smul : expr → tactic strictness -| `(%%a • %%b) := do +| e@`(%%a • %%b) := do strictness_a ← core a, strictness_b ← core b, match strictness_a, strictness_b with @@ -279,7 +293,11 @@ meta def positivity_smul : expr → tactic strictness | positive pa, nonnegative pb := nonnegative <$> mk_app ``smul_nonneg_of_pos_of_nonneg [pa, pb] | nonnegative pa, positive pb := nonnegative <$> mk_app ``smul_nonneg_of_nonneg_of_pos [pa, pb] | nonnegative pa, nonnegative pb := nonnegative <$> mk_app ``smul_nonneg [pa, pb] + | positive pa, nonzero pb := nonzero <$> to_expr ``(smul_ne_zero_of_pos_of_ne_zero %%pa %%pb) + | nonzero pa, positive pb := nonzero <$> to_expr ``(smul_ne_zero_of_ne_zero_of_pos %%pa %%pb) + | nonzero pa, nonzero pb := nonzero <$> to_expr ``(smul_ne_zero %%pa %%pb) + | sa@_, sb@ _ := positivity_fail e a b sa sb end -| _ := failed +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `a • b`" end tactic diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index c87976de0f018..300519f9ebd3b 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -2159,6 +2159,7 @@ do match strictness_a with | nonnegative p := nonnegative <$> mk_app ``real.rpow_nonneg_of_nonneg [p, b] | positive p := positive <$> mk_app ``real.rpow_pos_of_pos [p, b] + | _ := failed end private lemma nnrpow_pos {a : ℝ≥0} (ha : 0 < a) (b : ℝ) : 0 < a ^ b := nnreal.rpow_pos ha diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index dffc98ea3e875..a29a8597cdebe 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -49,10 +49,11 @@ also do some limits stuff (liminf/limsup etc). See https://isabelle.in.tum.de/dist/library/HOL/HOL-Library/Extended_Real.html -/ +open function open_locale ennreal nnreal /-- ereal : The type `[-∞, ∞]` -/ -@[derive [has_top, comm_monoid_with_zero, +@[derive [has_top, comm_monoid_with_zero, nontrivial, has_Sup, has_Inf, complete_linear_order, linear_ordered_add_comm_monoid_with_top]] def ereal := with_top (with_bot ℝ) @@ -69,12 +70,19 @@ instance : has_bot ereal := ⟨some ⊥⟩ @[simp] lemma bot_ne_top : (⊥ : ereal) ≠ ⊤ := bot_lt_top.ne instance : has_coe ℝ ereal := ⟨real.to_ereal⟩ + +lemma coe_strict_mono : strict_mono (coe : ℝ → ereal) := +with_top.coe_strict_mono.comp with_bot.coe_strict_mono + +lemma coe_injective : injective (coe : ℝ → ereal) := coe_strict_mono.injective + @[simp, norm_cast] protected lemma coe_le_coe_iff {x y : ℝ} : (x : ereal) ≤ (y : ereal) ↔ x ≤ y := -by { unfold_coes, simp [real.to_ereal] } +coe_strict_mono.le_iff_le @[simp, norm_cast] protected lemma coe_lt_coe_iff {x y : ℝ} : (x : ereal) < (y : ereal) ↔ x < y := -by { unfold_coes, simp [real.to_ereal] } +coe_strict_mono.lt_iff_lt @[simp, norm_cast] protected lemma coe_eq_coe_iff {x y : ℝ} : (x : ereal) = (y : ereal) ↔ x = y := -by { unfold_coes, simp [real.to_ereal, option.some_inj] } +coe_injective.eq_iff +protected lemma coe_ne_coe_iff {x y : ℝ} : (x : ereal) ≠ (y : ereal) ↔ x ≠ y := coe_injective.ne_iff /-- The canonical map from nonnegative extended reals to extended reals -/ def _root_.ennreal.to_ereal : ℝ≥0∞ → ereal @@ -154,6 +162,11 @@ rfl @[simp, norm_cast] lemma coe_zero : ((0 : ℝ) : ereal) = 0 := rfl +@[simp, norm_cast] lemma coe_eq_zero {x : ℝ} : (x : ereal) = 0 ↔ x = 0 := ereal.coe_eq_coe_iff +@[simp, norm_cast] lemma coe_eq_one {x : ℝ} : (x : ereal) = 1 ↔ x = 1 := ereal.coe_eq_coe_iff +lemma coe_ne_zero {x : ℝ} : (x : ereal) ≠ 0 ↔ x ≠ 0 := ereal.coe_ne_coe_iff +lemma coe_ne_one {x : ℝ} : (x : ereal) ≠ 1 ↔ x ≠ 1 := ereal.coe_ne_coe_iff + @[simp, norm_cast] protected lemma coe_nonneg {x : ℝ} : (0 : ereal) ≤ x ↔ 0 ≤ x := ereal.coe_le_coe_iff @@ -225,6 +238,7 @@ end lemma coe_nnreal_eq_coe_real (x : ℝ≥0) : ((x : ℝ≥0∞) : ereal) = (x : ℝ) := rfl @[simp, norm_cast] lemma coe_ennreal_zero : ((0 : ℝ≥0∞) : ereal) = 0 := rfl +@[simp, norm_cast] lemma coe_ennreal_one : ((1 : ℝ≥0∞) : ereal) = 1 := rfl @[simp, norm_cast] lemma coe_ennreal_top : ((⊤ : ℝ≥0∞) : ereal) = ⊤ := rfl @[simp] lemma coe_ennreal_eq_top_iff : ∀ {x : ℝ≥0∞}, (x : ereal) = ⊤ ↔ x = ⊤ @@ -255,6 +269,17 @@ lemma coe_nnreal_ne_top (x : ℝ≥0) : ((x : ℝ≥0∞) : ereal) ≠ ⊤ := de | ⊤ (some y) := by simp [(coe_nnreal_lt_top y).ne'] | (some x) (some y) := by simp [coe_nnreal_eq_coe_real] +@[simp, norm_cast] lemma coe_ennreal_eq_zero {x : ℝ≥0∞} : (x : ereal) = 0 ↔ x = 0 := +by rw [←coe_ennreal_eq_coe_ennreal_iff, coe_ennreal_zero] + +@[simp, norm_cast] lemma coe_ennreal_eq_one {x : ℝ≥0∞} : (x : ereal) = 1 ↔ x = 1 := +by rw [←coe_ennreal_eq_coe_ennreal_iff, coe_ennreal_one] + +@[norm_cast] lemma coe_ennreal_ne_zero {x : ℝ≥0∞} : (x : ereal) ≠ 0 ↔ x ≠ 0 := +coe_ennreal_eq_zero.not + +@[norm_cast] lemma coe_ennreal_ne_one {x : ℝ≥0∞} : (x : ereal) ≠ 1 ↔ x ≠ 1 := coe_ennreal_eq_one.not + lemma coe_ennreal_nonneg (x : ℝ≥0∞) : (0 : ereal) ≤ x := coe_ennreal_le_coe_ennreal_iff.2 (zero_le x) @@ -531,6 +556,7 @@ end ereal namespace tactic open positivity +private lemma ereal_coe_ne_zero {r : ℝ} : r ≠ 0 → (r : ereal) ≠ 0 := ereal.coe_ne_zero.2 private lemma ereal_coe_nonneg {r : ℝ} : 0 ≤ r → 0 ≤ (r : ereal) := ereal.coe_nonneg.2 private lemma ereal_coe_pos {r : ℝ} : 0 < r → 0 < (r : ereal) := ereal.coe_pos.2 private lemma ereal_coe_ennreal_pos {r : ℝ≥0∞} : 0 < r → 0 < (r : ereal) := ereal.coe_ennreal_pos.2 @@ -544,6 +570,7 @@ meta def positivity_coe_real_ereal : expr → tactic strictness match strictness_a with | positive p := positive <$> mk_app ``ereal_coe_pos [p] | nonnegative p := nonnegative <$> mk_mapp ``ereal_coe_nonneg [a, p] + | nonzero p := nonzero <$> mk_mapp ``ereal_coe_ne_zero [a, p] end | e := pp e >>= fail ∘ format.bracket "The expression " " is not of the form `(r : ereal)` for `r : ℝ`" @@ -556,7 +583,7 @@ meta def positivity_coe_ennreal_ereal : expr → tactic strictness strictness_a ← core a, match strictness_a with | positive p := positive <$> mk_app ``ereal_coe_ennreal_pos [p] - | nonnegative _ := nonnegative <$> mk_mapp `ereal.coe_ennreal_nonneg [a] + | _ := nonnegative <$> mk_mapp `ereal.coe_ennreal_nonneg [a] end | e := pp e >>= fail ∘ format.bracket "The expression " " is not of the form `(r : ereal)` for `r : ℝ≥0∞`" diff --git a/src/data/real/hyperreal.lean b/src/data/real/hyperreal.lean index 5a6f648ed4a7d..9c8b7679f2ebb 100644 --- a/src/data/real/hyperreal.lean +++ b/src/data/real/hyperreal.lean @@ -23,12 +23,13 @@ notation `ℝ*` := hyperreal noncomputable instance : has_coe_t ℝ ℝ* := ⟨λ x, (↑x : germ _ _)⟩ -@[simp, norm_cast] -lemma coe_eq_coe {x y : ℝ} : (x : ℝ*) = y ↔ x = y := -germ.const_inj +@[simp, norm_cast] lemma coe_eq_coe {x y : ℝ} : (x : ℝ*) = y ↔ x = y := germ.const_inj +lemma coe_ne_coe {x y : ℝ} : (x : ℝ*) ≠ y ↔ x ≠ y := coe_eq_coe.not @[simp, norm_cast] lemma coe_eq_zero {x : ℝ} : (x : ℝ*) = 0 ↔ x = 0 := coe_eq_coe @[simp, norm_cast] lemma coe_eq_one {x : ℝ} : (x : ℝ*) = 1 ↔ x = 1 := coe_eq_coe +@[norm_cast] lemma coe_ne_zero {x : ℝ} : (x : ℝ*) ≠ 0 ↔ x ≠ 0 := coe_ne_coe +@[norm_cast] lemma coe_ne_one {x : ℝ} : (x : ℝ*) ≠ 1 ↔ x ≠ 1 := coe_ne_coe @[simp, norm_cast] lemma coe_one : ↑(1 : ℝ) = (1 : ℝ*) := rfl @[simp, norm_cast] lemma coe_zero : ↑(0 : ℝ) = (0 : ℝ*) := rfl @@ -793,6 +794,7 @@ end hyperreal namespace tactic open positivity +private lemma hyperreal_coe_ne_zero {r : ℝ} : r ≠ 0 → (r : ℝ*) ≠ 0 := hyperreal.coe_ne_zero.2 private lemma hyperreal_coe_nonneg {r : ℝ} : 0 ≤ r → 0 ≤ (r : ℝ*) := hyperreal.coe_nonneg.2 private lemma hyperreal_coe_pos {r : ℝ} : 0 < r → 0 < (r : ℝ*) := hyperreal.coe_pos.2 @@ -805,6 +807,7 @@ meta def positivity_coe_real_hyperreal : expr → tactic strictness match strictness_a with | positive p := positive <$> mk_app ``hyperreal_coe_pos [p] | nonnegative p := nonnegative <$> mk_app ``hyperreal_coe_nonneg [p] + | nonzero p := nonzero <$> mk_app ``hyperreal_coe_ne_zero [p] end | e := pp e >>= fail ∘ format.bracket "The expression " " is not of the form `(r : ℝ*)` for `r : ℝ`" diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 75ef1fe775d56..f94cdbe95aeb7 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -852,7 +852,7 @@ meta def positivity_coe_nnreal_real : expr → tactic strictness strictness_a ← core a, match strictness_a with | positive p := positive <$> mk_app ``nnreal_coe_pos [p] - | nonnegative _ := nonnegative <$> mk_app ``nnreal.coe_nonneg [a] + | _ := nonnegative <$> mk_app ``nnreal.coe_nonneg [a] end | e := pp e >>= fail ∘ format.bracket "The expression " " is not of the form `(r : ℝ)` for `r : ℝ≥0`" diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index cf60b2f747c44..6bb58fcc28940 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -688,6 +688,8 @@ instance [partial_order α] : partial_order (with_bot α) := end, .. with_bot.preorder } +lemma coe_strict_mono [preorder α] : strict_mono (coe : α → with_bot α) := λ a b, some_lt_some.2 + lemma map_le_iff [preorder α] [preorder β] (f : α → β) (mono_iff : ∀ {a b}, f a ≤ f b ↔ a ≤ b) : ∀ (a b : with_bot α), a.map f ≤ b.map f ↔ a ≤ b | ⊥ _ := by simp only [map_bot, bot_le] @@ -1155,6 +1157,8 @@ instance [partial_order α] : partial_order (with_top α) := { le_antisymm := λ _ _, by { simp_rw [←to_dual_le_to_dual_iff], exact function.swap le_antisymm }, .. with_top.preorder } +lemma coe_strict_mono [preorder α] : strict_mono (coe : α → with_top α) := λ a b, some_lt_some.2 + lemma map_le_iff [preorder α] [preorder β] (f : α → β) (a b : with_top α) (mono_iff : ∀ {a b}, f a ≤ f b ↔ a ≤ b) : a.map f ≤ b.map f ↔ a ≤ b := @@ -1777,11 +1781,9 @@ section nontrivial variables [partial_order α] [bounded_order α] [nontrivial α] -lemma bot_ne_top : (⊥ : α) ≠ ⊤ := -λ H, not_nontrivial_iff_subsingleton.mpr (subsingleton_of_bot_eq_top H) ‹_› - -lemma top_ne_bot : (⊤ : α) ≠ ⊥ := bot_ne_top.symm -lemma bot_lt_top : (⊥ : α) < ⊤ := lt_top_iff_ne_top.2 bot_ne_top +@[simp] lemma bot_ne_top : (⊥ : α) ≠ ⊤ := λ h, not_subsingleton _ $ subsingleton_of_bot_eq_top h +@[simp] lemma top_ne_bot : (⊤ : α) ≠ ⊥ := bot_ne_top.symm +@[simp] lemma bot_lt_top : (⊥ : α) < ⊤ := lt_top_iff_ne_top.2 bot_ne_top end nontrivial diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 2ebf8d471d1f6..ebad5844b9405 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -585,10 +585,8 @@ noncomputable instance fractional_ideal.semifield : inv_zero := inv_zero' _, div := (/), div_eq_mul_inv := fractional_ideal.div_eq_mul_inv, - exists_pair_ne := ⟨0, 1, (coe_to_fractional_ideal_injective le_rfl).ne - (by simpa using @zero_ne_one (ideal A) _ _)⟩, mul_inv_cancel := λ I, fractional_ideal.mul_inv_cancel, - .. fractional_ideal.comm_semiring } + .. fractional_ideal.comm_semiring, ..(coe_to_fractional_ideal_injective le_rfl).nontrivial } /-- Fractional ideals have cancellative multiplication in a Dedekind domain. diff --git a/src/tactic/positivity.lean b/src/tactic/positivity.lean index 92b8d0586e829..50b65868112c2 100644 --- a/src/tactic/positivity.lean +++ b/src/tactic/positivity.lean @@ -1,15 +1,15 @@ /- Copyright (c) 2022 Mario Carneiro, Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Heather Macbeth +Authors: Mario Carneiro, Heather Macbeth, Yaël Dillies -/ import tactic.norm_num /-! # `positivity` tactic - -The `positivity` tactic in this file solves goals of the form `0 ≤ x` and `0 < x`. The tactic works -recursively according to the syntax of the expression `x`. For example, a goal of the form -`0 ≤ 3 * a ^ 2 + b * c` can be solved either +αᵒᵈ βᵒᵈ +The `positivity` tactic in this file solves goals of the form `0 ≤ x`, `0 < x` and `x ≠ 0`. The +tactic works recursively according to the syntax of the expression `x`. For example, a goal of the +form `0 ≤ 3 * a ^ 2 + b * c` can be solved either * by a hypothesis such as `5 ≤ 3 * a ^ 2 + b * c` which directly implies the nonegativity of `3 * a ^ 2 + b * c`; or, * by the application of the lemma `add_nonneg` and the success of the `positivity` tactic on the two @@ -44,8 +44,11 @@ introduce these operations. `tactic.positivity.strictness`, containing an `expr` which is a proof of the strict-positivity/nonnegativity of `e` as well as an indication of whether what could be proved was strict-positivity or nonnegativity +* `tactic.positivity.order_rel` is a custom inductive type recording whether the goal is + `0 ≤ e`/`e ≥ 0`, `0 < e`/`e > 0`, `e ≠ 0` or `0 ≠ e`. * `tactic.interactive.positivity` is the user-facing tactic. It parses the goal and, if it is of - one of the forms `0 ≤ e`, `0 < e`, `e > 0`, `e ≥ 0`, it sends `e` to `tactic.positivity.core`. + one of the forms `0 ≤ e`, `0 < e`, `e > 0`, `e ≥ 0`, `e ≠ 0`, `0 ≠ e`, it sends `e` to + `tactic.positivity.core`. ## TODO @@ -60,14 +63,24 @@ namespace tactic meta inductive positivity.strictness : Type | positive : expr → positivity.strictness | nonnegative : expr → positivity.strictness +| nonzero : expr → positivity.strictness + +export positivity.strictness (positive nonnegative nonzero) + +meta instance : has_to_string strictness := +⟨λ s, match s with + | positive p := "strictness.positive (" ++ to_string p ++ ")" + | nonnegative p := "strictness.nonnegative (" ++ to_string p ++ ")" + | nonzero p := "strictness.nonzero (" ++ to_string p ++ ")" + end⟩ -export positivity.strictness (positive nonnegative) +meta instance : has_to_format strictness := ⟨λ s, to_string s⟩ -private lemma lt_of_lt_of_eq'' {α} [preorder α] {b b' a : α} : b = b' → a < b' → a < b := +private lemma lt_of_eq_of_lt'' {α} [preorder α] {b b' a : α} : b = b' → a < b' → a < b := λ h1 h2, lt_of_lt_of_eq h2 h1.symm /-- First base case of the `positivity` tactic. We try `norm_num` to prove directly that an -expression `e` is positive or nonnegative. -/ +expression `e` is positive, nonnegative or non-zero. -/ meta def norm_num.positivity (e : expr) : tactic strictness := do (e', p) ← norm_num.derive e <|> refl_conv e, e'' ← e'.to_rat, @@ -75,12 +88,12 @@ meta def norm_num.positivity (e : expr) : tactic strictness := do ic ← mk_instance_cache typ, if e'' > 0 then do (ic, p₁) ← norm_num.prove_pos ic e', - p ← mk_app ``lt_of_lt_of_eq'' [p, p₁], - pure (positive p) - else if e'' = 0 then do - p' ← mk_app ``ge_of_eq [p], - pure (nonnegative p') - else failed + positive <$> mk_app ``lt_of_eq_of_lt'' [p, p₁] + else if e'' = 0 then + nonnegative <$> mk_app ``ge_of_eq [p] + else do + (ic, p₁) ← norm_num.prove_ne_zero' ic e', + nonzero <$> to_expr ``(ne_of_eq_of_ne %%p %%p₁) /-- Second base case of the `positivity` tactic: Any element of a canonically ordered additive monoid is nonnegative. -/ @@ -89,47 +102,153 @@ meta def positivity_canon : expr → tactic strictness namespace positivity +/-- Inductive type recording whether the goal `positivity` is called on is nonnegativity, positivity +or different from `0`. -/ +@[derive inhabited] +inductive order_rel : Type +| le : order_rel -- `0 ≤ a` +| lt : order_rel -- `0 < a` +| ne : order_rel -- `a ≠ 0` +| ne' : order_rel -- `0 ≠ a` + +meta instance : has_to_format order_rel := +⟨λ r, match r with + | order_rel.le := "order_rel.le" + | order_rel.lt := "order_rel.lt" + | order_rel.ne := "order_rel.ne" + | order_rel.ne' := "order_rel.ne'" + end⟩ + /-- Given two tactics whose result is `strictness`, report a `strictness`: - if at least one gives `positive`, report `positive` and one of the expressions giving a proof of positivity -- if neither gives `pos` but at least one gives `nonnegative`, report `nonnegative` and one of the +- if one reports `nonnegative` and the other reports `nonzero`, report `positive` +- else if at least one reports `nonnegative`, report `nonnegative` and one of the expressions giving a proof of nonnegativity +- else if at least one reports `nonzero`, report `nonzero` and one of the expressions giving a proof + of nonzeroness - if both fail, fail -/ -meta def orelse' (tac1 tac2 : tactic strictness) : tactic strictness := do - res ← try_core tac1, - match res with +protected meta def orelse (tac1 tac2 : tactic strictness) : tactic strictness := do + res1 ← try_core tac1, + match res1 with | none := tac2 - | some res@(nonnegative e) := tac2 <|> pure res - | some res@(positive _) := pure res + | some p1@(positive _) := pure p1 + | some (nonnegative e1) := do + res2 ← try_core tac2, + match res2 with + | some p2@(positive _) := pure p2 + | some (nonzero e2) := positive <$> mk_app ``lt_of_le_of_ne' [e1, e2] + | _ := pure (nonnegative e1) + end + | some (nonzero e1) := do + res2 ← try_core tac2, + match res2 with + | some p2@(positive _) := pure p2 + | some (nonnegative e2) := positive <$> mk_app ``lt_of_le_of_ne' [e2, e1] + | _ := pure (nonzero e1) + end end +localized "infixr ` ≤|≥ `:2 := tactic.positivity.orelse" in positivity + +/-- This tactic fails with a message saying that `positivity` couldn't prove anything about `e` +if we only know that `a` and `b` are positive/nonnegative/nonzero (according to `pa` and `pb`). -/ +meta def positivity_fail {α : Type*} (e a b : expr) (pa pb : strictness) : tactic α := +do + e' ← pp e, + a' ← pp a, + b' ← pp b, + let f : strictness → format → format := λ p c, match p with + | positive _ := "0 < " ++ c + | nonnegative _ := "0 ≤ " ++ c + | nonzero _ := c ++ " ≠ 0" + end, + fail (↑"`positivity` can't say anything about `" ++ e' ++ "` knowing only `" ++ f pa a' ++ + "` and `" ++ f pb b' ++ "`") + /-! ### Core logic of the `positivity` tactic -/ -/-- Third base case of the `positivity` tactic. Prove an expression `e` is positive/nonnegative by -finding a hypothesis of the form `a < e` or `a ≤ e` in which `a` can be proved positive/nonnegative -by `norm_num`. -/ +private lemma ne_of_ne_of_eq' {α : Type*} {a b c : α} (ha : a ≠ c) (h : a = b) : b ≠ c := by rwa ←h + +/-- Calls `norm_num` on `a` to prove positivity/nonnegativity of `e` assuming `b` is defeq to `e` +and `p₂ : a ≤ b`. -/ +meta def compare_hyp_le (e a b p₂ : expr) : tactic strictness := do + is_def_eq e b, + strict_a ← norm_num.positivity a, + match strict_a with + | positive p₁ := positive <$> mk_app ``lt_of_lt_of_le [p₁, p₂] + | nonnegative p₁ := nonnegative <$> mk_app ``le_trans [p₁, p₂] + | _ := do + e' ← pp e, + p₂' ← pp p₂, + fail (↑"`norm_num` can't prove nonnegativity of " ++ e' ++ " using " ++ p₂') + end + +/-- Calls `norm_num` on `a` to prove positivity/nonnegativity of `e` assuming `b` is defeq to `e` +and `p₂ : a < b`. -/ +meta def compare_hyp_lt (e a b p₂ : expr) : tactic strictness := do + is_def_eq e b, + strict_a ← norm_num.positivity a, + match strict_a with + | positive p₁ := positive <$> mk_app ``lt_trans [p₁, p₂] + | nonnegative p₁ := positive <$> mk_app ``lt_of_le_of_lt [p₁, p₂] + | _ := do + e' ← pp e, + p₂' ← pp p₂, + fail (↑"`norm_num` can't prove positivity of " ++ e' ++ " using " ++ p₂') + end + +/-- Calls `norm_num` on `a` to prove positivity/nonnegativity/nonzeroness of `e` assuming `b` is +defeq to `e` and `p₂ : a = b`. -/ +meta def compare_hyp_eq (e a b p₂ : expr) : tactic strictness := do + is_def_eq e b, + strict_a ← norm_num.positivity a, + match strict_a with + | positive p₁ := positive <$> mk_app ``lt_of_lt_of_eq [p₁, p₂] + | nonnegative p₁ := nonnegative <$> mk_app ``le_of_le_of_eq [p₁, p₂] + | nonzero p₁ := nonzero <$> to_expr ``(ne_of_ne_of_eq' %%p₁ %%p₂) + end + +/-- Calls `norm_num` on `a` to prove nonzeroness of `e` assuming `b` is defeq to `e` and +`p₂ : b ≠ a`. -/ +meta def compare_hyp_ne (e a b p₂ : expr) : tactic strictness := do + is_def_eq e b, + (a', p₁) ← norm_num.derive a <|> refl_conv a, + a'' ← a'.to_rat, + if a'' = 0 then + nonzero <$> mk_mapp ``ne_of_ne_of_eq [none, none, none, none, p₂, p₁] + else do + e' ← pp e, + p₂' ← pp p₂, + a' ← pp a, + fail (↑"`norm_num` can't prove non-zeroness of " ++ e' ++ " using " ++ p₂' ++ " because " + ++ a' ++ " is non-zero") + +/-- Third base case of the `positivity` tactic. Prove an expression `e` is +positive/nonnegative/nonzero by finding a hypothesis of the form `a < e`, `a ≤ e` or `a = e` in +which `a` can be proved positive/nonnegative/nonzero by `norm_num`. -/ meta def compare_hyp (e p₂ : expr) : tactic strictness := do p_typ ← infer_type p₂, - (lo, hi, strict₂) ← match p_typ with -- TODO also handle equality hypotheses - | `(%%lo ≤ %%hi) := pure (lo, hi, ff) - | `(%%hi ≥ %%lo) := pure (lo, hi, ff) - | `(%%lo < %%hi) := pure (lo, hi, tt) - | `(%%hi > %%lo) := pure (lo, hi, tt) - | _ := failed - end, - is_def_eq e hi, - strictness₁ ← norm_num.positivity lo, - match strictness₁, strict₂ with - | (positive p₁), tt := positive <$> mk_app ``lt_trans [p₁, p₂] - | (positive p₁), ff := positive <$> mk_app `lt_of_lt_of_le [p₁, p₂] - | (nonnegative p₁), tt := positive <$> mk_app `lt_of_le_of_lt [p₁, p₂] - | (nonnegative p₁), ff := nonnegative <$> mk_app `le_trans [p₁, p₂] + match p_typ with + | `(%%lo ≤ %%hi) := compare_hyp_le e lo hi p₂ + | `(%%hi ≥ %%lo) := compare_hyp_le e lo hi p₂ + | `(%%lo < %%hi) := compare_hyp_lt e lo hi p₂ + | `(%%hi > %%lo) := compare_hyp_lt e lo hi p₂ + | `(%%lo = %%hi) := compare_hyp_eq e lo hi p₂ <|> do + p₂' ← mk_app ``eq.symm [p₂], + compare_hyp_eq e hi lo p₂' + | `(%%hi ≠ %%lo) := compare_hyp_ne e lo hi p₂ <|> do + p₂' ← mk_mapp ``ne.symm [none, none, none, p₂], + compare_hyp_ne e hi lo p₂' + | e := do + p₂' ← pp p₂, + fail (p₂' ++ "is not of the form `a ≤ b`, `a < b`, `a = b` or `a ≠ b`") end -/-- Attribute allowing a user to tag a tactic as an extension for `tactic.positivity`. The main -(recursive) step of this tactic is to try successively all the extensions tagged with this attribute -on the expression at hand, and also to try the two "base case" tactics `tactic.norm_num.positivity`, -`tactic.positivity.compare_hyp` on the expression at hand. -/ +/-- Attribute allowing a user to tag a tactic as an extension for `tactic.interactive.positivity`. +The main (recursive) step of this tactic is to try successively all the extensions tagged with this +attribute on the expression at hand, and also to try the two "base case" tactics +`tactic.norm_num.positivity`, `tactic.positivity.compare_hyp` on the expression at hand. -/ @[user_attribute] meta def attr : user_attribute (expr → tactic strictness) unit := { name := `positivity, @@ -139,35 +258,37 @@ meta def attr : user_attribute (expr → tactic strictness) unit := { t ← ns.mfoldl (λ (t : expr → tactic strictness) n, do t' ← eval_expr (expr → tactic strictness) (expr.const n []), - pure (λ e, orelse' (t' e) (t e))) + pure (λ e, t' e ≤|≥ t e)) (λ _, failed), - pure $ λ e, orelse' - (t e) $ orelse' -- run all the extensions on `e` - (norm_num.positivity e) $ orelse' -- directly try `norm_num` on `e` - (positivity_canon e) $ -- try showing nonnegativity from canonicity of the order + pure $ λ e, + t e -- run all the extensions on `e` + ≤|≥ norm_num.positivity e -- directly try `norm_num` on `e` + ≤|≥ positivity_canon e -- try showing nonnegativity from canonicity of the order -- loop over hypotheses and try to compare with `e` - local_context >>= list.foldl (λ tac h, orelse' tac (compare_hyp e h)) failed }, + ≤|≥ local_context >>= list.foldl (λ tac h, tac ≤|≥ compare_hyp e h) + (fail "no applicable positivity extension found") }, dependencies := [] } } /-- Look for a proof of positivity/nonnegativity of an expression `e`; if found, return the proof together with a `strictness` stating whether the proof found was for strict positivity -(`positive p`) or only for nonnegativity (`nonnegative p`). -/ +(`positive p`), nonnegativity (`nonnegative p`), or nonzeroness (`nonzero p`). -/ meta def core (e : expr) : tactic strictness := do f ← attr.get_cache, - f e <|> fail "failed to prove positivity/nonnegativity" + f e <|> fail "failed to prove positivity/nonnegativity/nonzeroness" end positivity open positivity +open_locale positivity namespace interactive setup_tactic_parser -/-- Tactic solving goals of the form `0 ≤ x` and `0 < x`. The tactic works recursively according to -the syntax of the expression `x`, if the atoms composing the expression all have numeric lower -bounds which can be proved positive/nonnegative by `norm_num`. This tactic either closes the goal -or fails. +/-- Tactic solving goals of the form `0 ≤ x`, `0 < x` and `x ≠ 0`. The tactic works recursively +according to the syntax of the expression `x`, if the atoms composing the expression all have +numeric lower bounds which can be proved positive/nonnegative/nonzero by `norm_num`. This tactic +either closes the goal or fails. Examples: ``` @@ -181,19 +302,39 @@ example {b : ℤ} : 0 ≤ max (-3) (b ^ 2) := by positivity meta def positivity : tactic unit := focus1 $ do t ← target >>= instantiate_mvars, (rel_desired, a) ← match t with - | `(0 ≤ %%e₂) := pure (ff, e₂) - | `(%%e₂ ≥ 0) := pure (ff, e₂) - | `(0 < %%e₂) := pure (tt, e₂) - | `(%%e₂ > 0) := pure (tt, e₂) - | _ := fail "not a positivity/nonnegativity goal" + | `(0 ≤ %%e) := pure (order_rel.le, e) + | `(%%e ≥ 0) := pure (order_rel.le, e) + | `(0 < %%e) := pure (order_rel.lt, e) + | `(%%e > 0) := pure (order_rel.lt, e) + | `(%%e₁ ≠ %%e₂) := do + match e₂ with + | `(has_zero.zero) := pure (order_rel.ne, e₁) + | _ := match e₁ with + | `(has_zero.zero) := pure (order_rel.ne', e₂) + | _ := fail "not a positivity/nonnegativity/nonzeroness goal" + end + end + | _ := fail "not a positivity/nonnegativity/nonzeroness goal" end, strictness_proved ← tactic.positivity.core a, match rel_desired, strictness_proved with - | tt, (positive p) := pure p - | tt, (nonnegative _) := fail ("failed to prove strict positivity, but it would be possible to " - ++ "prove nonnegativity if desired") - | ff, (positive p) := mk_app ``le_of_lt [p] - | ff, (nonnegative p) := pure p + | order_rel.lt, positive p := pure p + | order_rel.lt, nonnegative _ := fail ("failed to prove strict positivity, but it would be " ++ + "possible to prove nonnegativity if desired") + | order_rel.lt, nonzero _ := fail ("failed to prove strict positivity, but it would be " ++ + "possible to prove nonzeroness if desired") + | order_rel.le, positive p := mk_app ``le_of_lt [p] + | order_rel.le, nonnegative p := pure p + | order_rel.le, nonzero _ := fail ("failed to prove nonnegativity, but it would be " ++ + "possible to prove nonzeroness if desired") + | order_rel.ne, positive p := to_expr ``(ne_of_gt %%p) + | order_rel.ne, nonnegative _ := fail ("failed to prove nonzeroness, but it would be " ++ + "possible to prove nonnegativity if desired") + | order_rel.ne, nonzero p := pure p + | order_rel.ne', positive p := to_expr ``(ne_of_lt %%p) + | order_rel.ne', nonnegative _ := fail ("failed to prove nonzeroness, but it would be " ++ + "possible to prove nonnegativity if desired") + | order_rel.ne', nonzero p := to_expr ``(ne.symm %%p) end >>= tactic.exact add_tactic_doc @@ -208,19 +349,26 @@ variables {α R : Type*} /-! ### `positivity` extensions for particular arithmetic operations -/ -private lemma le_min_of_lt_of_le [linear_order R] (a b c : R) (ha : a < b) (hb : a ≤ c) : - a ≤ min b c := -le_min ha.le hb +section linear_order +variables [linear_order R] {a b c : R} + +private lemma le_min_of_lt_of_le (ha : a < b) (hb : a ≤ c) : a ≤ min b c := le_min ha.le hb +private lemma le_min_of_le_of_lt (ha : a ≤ b) (hb : a < c) : a ≤ min b c := le_min ha hb.le +private lemma min_ne (ha : a ≠ c) (hb : b ≠ c) : min a b ≠ c := +by { rw min_def, split_ifs; assumption } +private lemma min_ne_of_ne_of_lt (ha : a ≠ c) (hb : c < b) : min a b ≠ c := min_ne ha hb.ne' +private lemma min_ne_of_lt_of_ne (ha : c < a) (hb : b ≠ c) : min a b ≠ c := min_ne ha.ne' hb + +private lemma max_ne (ha : a ≠ c) (hb : b ≠ c) : max a b ≠ c := +by { rw max_def, split_ifs; assumption } -private lemma le_min_of_le_of_lt [linear_order R] (a b c : R) (ha : a ≤ b) (hb : a < c) : - a ≤ min b c := -le_min ha hb.le +end linear_order /-- Extension for the `positivity` tactic: the `min` of two numbers is nonnegative if both are nonnegative, and strictly positive if both are. -/ @[positivity] meta def positivity_min : expr → tactic strictness -| `(min %%a %%b) := do +| e@`(min %%a %%b) := do strictness_a ← core a, strictness_b ← core b, match strictness_a, strictness_b with @@ -228,34 +376,46 @@ meta def positivity_min : expr → tactic strictness | (positive pa), (nonnegative pb) := nonnegative <$> mk_app ``le_min_of_lt_of_le [pa, pb] | (nonnegative pa), (positive pb) := nonnegative <$> mk_app ``le_min_of_le_of_lt [pa, pb] | (nonnegative pa), (nonnegative pb) := nonnegative <$> mk_app ``le_min [pa, pb] + | positive pa, nonzero pb := nonzero <$> to_expr ``(min_ne_of_lt_of_ne %%pa %%pb) + | nonzero pa, positive pb := nonzero <$> to_expr ``(min_ne_of_ne_of_lt %%pa %%pb) + | nonzero pa, nonzero pb := nonzero <$> to_expr ``(min_ne %%pa %%pb) + | sa@_, sb@ _ := positivity_fail e a b sa sb end -| _ := failed +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `min a b`" /-- Extension for the `positivity` tactic: the `max` of two numbers is nonnegative if at least one -is nonnegative, and strictly positive if at least one is positive. -/ +is nonnegative, strictly positive if at least one is positive, and nonzero if both are nonzero. -/ @[positivity] meta def positivity_max : expr → tactic strictness -| `(max %%a %%b) := tactic.positivity.orelse' (do - strictness_a ← core a, +| `(max %%a %%b) := do + strictness_a ← try_core (core a), + (do match strictness_a with - | (positive pa) := positive <$> mk_mapp ``lt_max_of_lt_left [none, none, none, a, b, pa] - | (nonnegative pa) := + | some (positive pa) := positive <$> mk_mapp ``lt_max_of_lt_left [none, none, none, a, b, pa] + | some (nonnegative pa) := nonnegative <$> mk_mapp ``le_max_of_le_left [none, none, none, a, b, pa] - end) + | _ := failed + -- If `a ≠ 0`, we might prove `max a b ≠ 0` if `b ≠ 0` but we don't want to evaluate + -- `b` before having ruled out `0 < a`, for performance. So we do that in the second branch + -- of the `orelse'`. + end) ≤|≥ (do strictness_b ← core b, match strictness_b with | (positive pb) := positive <$> mk_mapp ``lt_max_of_lt_right [none, none, none, a, b, pb] | (nonnegative pb) := nonnegative <$> mk_mapp ``le_max_of_le_right [none, none, none, a, b, pb] + | nonzero pb := do + nonzero pa ← strictness_a, + nonzero <$> to_expr ``(max_ne %%pa %%pb) end) -| _ := failed +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `max a b`" /-- Extension for the `positivity` tactic: addition is nonnegative if both summands are nonnegative, and strictly positive if at least one summand is. -/ @[positivity] meta def positivity_add : expr → tactic strictness -| `(%%a + %%b) := do +| e@`(%%a + %%b) := do strictness_a ← core a, strictness_b ← core b, match strictness_a, strictness_b with @@ -263,24 +423,34 @@ meta def positivity_add : expr → tactic strictness | (positive pa), (nonnegative pb) := positive <$> mk_app ``lt_add_of_pos_of_le [pa, pb] | (nonnegative pa), (positive pb) := positive <$> mk_app ``lt_add_of_le_of_pos [pa, pb] | (nonnegative pa), (nonnegative pb) := nonnegative <$> mk_app ``add_nonneg [pa, pb] + | sa@_, sb@ _ := positivity_fail e a b sa sb end -| _ := failed +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `a + b`" + +section ordered_semiring +variables [ordered_semiring R] {a b : R} -private lemma mul_nonneg_of_pos_of_nonneg [linear_ordered_semiring R] (a b : R) (ha : 0 < a) - (hb : 0 ≤ b) : - 0 ≤ a * b := +private lemma mul_nonneg_of_pos_of_nonneg (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ a * b := mul_nonneg ha.le hb -private lemma mul_nonneg_of_nonneg_of_pos [linear_ordered_semiring R] (a b : R) (ha : 0 ≤ a) - (hb : 0 < b) : - 0 ≤ a * b := +private lemma mul_nonneg_of_nonneg_of_pos (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ a * b := mul_nonneg ha hb.le -/-- Extension for the `positivity` tactic: multiplication is nonnegative if both multiplicands are -nonnegative, and strictly positive if both multiplicands are. -/ +private lemma mul_ne_zero_of_pos_of_ne_zero [no_zero_divisors R] (ha : 0 < a) (hb : b ≠ 0) : + a * b ≠ 0 := +mul_ne_zero ha.ne' hb + +private lemma mul_ne_zero_of_ne_zero_of_pos [no_zero_divisors R] (ha : a ≠ 0) (hb : 0 < b) : + a * b ≠ 0 := +mul_ne_zero ha hb.ne' + +end ordered_semiring + +/-- Extension for the `positivity` tactic: multiplication is nonnegative/positive/nonzero if both +multiplicands are. -/ @[positivity] meta def positivity_mul : expr → tactic strictness -| `(%%a * %%b) := do +| e@`(%%a * %%b) := do strictness_a ← core a, strictness_b ← core b, match strictness_a, strictness_b with @@ -288,19 +458,30 @@ meta def positivity_mul : expr → tactic strictness | (positive pa), (nonnegative pb) := nonnegative <$> mk_app ``mul_nonneg_of_pos_of_nonneg [pa, pb] | (nonnegative pa), (positive pb) := nonnegative <$> mk_app ``mul_nonneg_of_nonneg_of_pos [pa, pb] | (nonnegative pa), (nonnegative pb) := nonnegative <$> mk_app ``mul_nonneg [pa, pb] + | positive pa, nonzero pb := nonzero <$> to_expr ``(mul_ne_zero_of_pos_of_ne_zero %%pa %%pb) + | nonzero pa, positive pb := nonzero <$> to_expr ``(mul_ne_zero_of_ne_zero_of_pos %%pa %%pb) + | nonzero pa, nonzero pb := nonzero <$> to_expr ``(mul_ne_zero %%pa %%pb) + | sa@_, sb@ _ := positivity_fail e a b sa sb end -| _ := failed +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `a * b`" -private lemma div_nonneg_of_pos_of_nonneg [linear_ordered_field R] {a b : R} (ha : 0 < a) - (hb : 0 ≤ b) : - 0 ≤ a / b := +section linear_ordered_semifield +variables [linear_ordered_semifield R] {a b : R} + +private lemma div_nonneg_of_pos_of_nonneg (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ a / b := div_nonneg ha.le hb -private lemma div_nonneg_of_nonneg_of_pos [linear_ordered_field R] {a b : R} (ha : 0 ≤ a) - (hb : 0 < b) : - 0 ≤ a / b := +private lemma div_nonneg_of_nonneg_of_pos (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ a / b := div_nonneg ha hb.le +private lemma div_ne_zero_of_pos_of_ne_zero (ha : 0 < a) (hb : b ≠ 0) : a / b ≠ 0 := +div_ne_zero ha.ne' hb + +private lemma div_ne_zero_of_ne_zero_of_pos (ha : a ≠ 0) (hb : 0 < b) : a / b ≠ 0 := +div_ne_zero ha hb.ne' + +end linear_ordered_semifield + private lemma int_div_self_pos {a : ℤ} (ha : 0 < a) : 0 < a / a := by { rw int.div_self ha.ne', exact zero_lt_one } @@ -317,7 +498,7 @@ int.div_nonneg ha.le hb.le are nonnegative, and strictly positive if both numerator and denominator are. -/ @[positivity] meta def positivity_div : expr → tactic strictness -| `(@has_div.div int _ %%a %%b) := do +| e@`(@has_div.div ℤ _ %%a %%b) := do strictness_a ← core a, strictness_b ← core b, match strictness_a, strictness_b with @@ -331,8 +512,9 @@ meta def positivity_div : expr → tactic strictness | nonnegative pa, positive pb := nonnegative <$> mk_app ``int_div_nonneg_of_nonneg_of_pos [pa, pb] | nonnegative pa, nonnegative pb := nonnegative <$> mk_app ``int.div_nonneg [pa, pb] + | sa@_, sb@ _ := positivity_fail e a b sa sb end -| `(%%a / %%b) := do +| e@`(%%a / %%b) := do strictness_a ← core a, strictness_b ← core b, match strictness_a, strictness_b with @@ -340,8 +522,12 @@ meta def positivity_div : expr → tactic strictness | positive pa, nonnegative pb := nonnegative <$> mk_app ``div_nonneg_of_pos_of_nonneg [pa, pb] | nonnegative pa, positive pb := nonnegative <$> mk_app ``div_nonneg_of_nonneg_of_pos [pa, pb] | nonnegative pa, nonnegative pb := nonnegative <$> mk_app ``div_nonneg [pa, pb] + | positive pa, nonzero pb := nonzero <$> to_expr ``(div_ne_zero_of_pos_of_ne_zero %%pa %%pb) + | nonzero pa, positive pb := nonzero <$> to_expr ``(div_ne_zero_of_ne_zero_of_pos %%pa %%pb) + | nonzero pa, nonzero pb := nonzero <$> to_expr ``(div_ne_zero %%pa %%pb) + | sa@_, sb@ _ := positivity_fail e a b sa sb end -| _ := failed +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `a / b`" /-- Extension for the `positivity` tactic: an inverse of a positive number is positive, an inverse of a nonnegative number is nonnegative. -/ @@ -352,8 +538,9 @@ meta def positivity_inv : expr → tactic strictness match strictness_a with | (positive pa) := positive <$> mk_app ``inv_pos_of_pos [pa] | (nonnegative pa) := nonnegative <$> mk_app ``inv_nonneg_of_nonneg [pa] + | nonzero pa := nonzero <$> to_expr ``(inv_ne_zero %%pa) end -| _ := failed +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `a⁻¹`" private lemma pow_zero_pos [ordered_semiring R] [nontrivial R] (a : R) : 0 < a ^ 0 := zero_lt_one.trans_le (pow_zero a).ge @@ -366,52 +553,66 @@ positive if `n = 0` (since `a ^ 0 = 1`) or if `0 < a`, and is nonnegative if `n` are nonnegative) or if `0 ≤ a`. -/ @[positivity] meta def positivity_pow : expr → tactic strictness -| `(%%a ^ %%n) := do +| e@`(%%a ^ %%n) := do typ ← infer_type n, (do unify typ `(ℕ), if n = `(0) then positive <$> mk_app ``pow_zero_pos [a] - else positivity.orelse' - (do -- even powers are nonnegative + else + do -- even powers are nonnegative + -- Note this is automatically strengthened to `0 < a ^ n` when `a ≠ 0` thanks to the `orelse'` match n with -- TODO: Decision procedure for parity | `(bit0 %% n) := nonnegative <$> mk_app ``pow_bit0_nonneg [a, n] - | _ := failed - end) $ + | _ := do + e' ← pp e, + fail (e' ++ "is not an even power so positivity can't prove it's nonnegative") + end ≤|≥ do -- `a ^ n` is positive if `a` is, and nonnegative if `a` is strictness_a ← core a, match strictness_a with | positive p := positive <$> mk_app ``pow_pos [p, n] | nonnegative p := nonnegative <$> mk_app `pow_nonneg [p, n] + | nonzero p := nonzero <$> to_expr ``(pow_ne_zero %%n %%p) end) <|> - (do - unify typ `(ℤ), - if n = `(0 : ℤ) then - positive <$> mk_app ``zpow_zero_pos [a] - else positivity.orelse' - (do -- even powers are nonnegative - match n with -- TODO: Decision procedure for parity - | `(bit0 %% n) := nonnegative <$> mk_app ``zpow_bit0_nonneg [a, n] - | _ := failed - end) $ - do -- `a ^ n` is positive if `a` is, and nonnegative if `a` is - strictness_a ← core a, - match strictness_a with - | positive p := positive <$> mk_app ``zpow_pos_of_pos [p, n] - | nonnegative p := nonnegative <$> mk_app ``zpow_nonneg [p, n] - end) -| _ := failed + (do + unify typ `(ℤ), + if n = `(0 : ℤ) then + positive <$> mk_app ``zpow_zero_pos [a] + else + do -- even powers are nonnegative + -- Note this is automatically strengthened to `0 < a ^ n` when `a ≠ 0` thanks to the `orelse'` + match n with -- TODO: Decision procedure for parity + | `(bit0 %%n) := nonnegative <$> mk_app ``zpow_bit0_nonneg [a, n] + | _ := do + e' ← pp e, + fail (e' ++ "is not an even power so positivity can't prove it's nonnegative") + end ≤|≥ + do -- `a ^ n` is positive if `a` is, and nonnegative if `a` is + strictness_a ← core a, + match strictness_a with + | positive p := positive <$> mk_app ``zpow_pos_of_pos [p, n] + | nonnegative p := nonnegative <$> mk_app ``zpow_nonneg [p, n] + | nonzero p := nonzero <$> to_expr ``(zpow_ne_zero %%n %%p) + end) +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `a ^ n`" + +private alias abs_pos ↔ _ abs_pos_of_ne_zero /-- Extension for the `positivity` tactic: an absolute value is nonnegative, and is strictly -positive if its input is. -/ +positive if its input is nonzero. -/ @[positivity] meta def positivity_abs : expr → tactic strictness | `(|%%a|) := do - (do -- if can prove `0 < a`, report positivity - positive pa ← core a, - positive <$> mk_app ``abs_pos_of_pos [pa]) <|> + (do -- if can prove `0 < a` or `a ≠ 0`, report positivity + strict_a ← core a, + match strict_a with + | positive p := positive <$> mk_app ``abs_pos_of_pos [p] + | nonzero p := positive <$> mk_app ``abs_pos_of_ne_zero [p] + | _ := failed + end) <|> nonnegative <$> mk_app ``abs_nonneg [a] -- else report nonnegativity -| _ := failed +| e := pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `|a|`" private lemma int_nat_abs_pos {n : ℤ} (hn : 0 < n) : 0 < n.nat_abs := int.nat_abs_pos_of_ne_zero hn.ne' @@ -424,8 +625,12 @@ Since the output type of `int.nat_abs` is `ℕ`, the nonnegative case is handled @[positivity] meta def positivity_nat_abs : expr → tactic strictness | `(int.nat_abs %%a) := do - positive p ← core a, - positive <$> mk_app ``int_nat_abs_pos [p] + strict_a ← core a, + match strict_a with + | positive p := positive <$> mk_app ``int_nat_abs_pos [p] + | nonzero p := positive <$> mk_app ``int.nat_abs_pos_of_ne_zero [p] + | _ := failed + end | _ := failed private lemma nat_cast_pos [ordered_semiring α] [nontrivial α] {n : ℕ} : 0 < n → 0 < (n : α) := @@ -434,11 +639,15 @@ nat.cast_pos.2 private lemma int_coe_nat_nonneg (n : ℕ) : 0 ≤ (n : ℤ) := n.cast_nonneg private lemma int_coe_nat_pos {n : ℕ} : 0 < n → 0 < (n : ℤ) := nat.cast_pos.2 +private lemma int_cast_ne_zero [add_group_with_one α] [char_zero α] {n : ℤ} : n ≠ 0 → (n : α) ≠ 0 := +int.cast_ne_zero.2 private lemma int_cast_nonneg [ordered_ring α] {n : ℤ} (hn : 0 ≤ n) : 0 ≤ (n : α) := by { rw ←int.cast_zero, exact int.cast_mono hn } private lemma int_cast_pos [ordered_ring α] [nontrivial α] {n : ℤ} : 0 < n → 0 < (n : α) := int.cast_pos.2 +private lemma rat_cast_ne_zero [division_ring α] [char_zero α] {q : ℚ} : q ≠ 0 → (q : α) ≠ 0 := +rat.cast_ne_zero.2 private lemma rat_cast_nonneg [linear_ordered_field α] {q : ℚ} : 0 ≤ q → 0 ≤ (q : α) := rat.cast_nonneg.2 private lemma rat_cast_pos [linear_ordered_field α] {q : ℚ} : 0 < q → 0 < (q : α) := rat.cast_pos.2 @@ -461,9 +670,13 @@ meta def positivity_coe : expr → tactic strictness | `(int.cast_coe), positive p := positive <$> mk_mapp ``int_cast_pos [typ, none, none, none, p] | `(int.cast_coe), nonnegative p := nonnegative <$> mk_mapp ``int_cast_nonneg [typ, none, none, p] + | `(int.cast_coe), nonzero p := nonzero <$> + mk_mapp ``int_cast_ne_zero [typ, none, none, none, p] | `(rat.cast_coe), positive p := positive <$> mk_mapp ``rat_cast_pos [typ, none, none, p] | `(rat.cast_coe), nonnegative p := nonnegative <$> mk_mapp ``rat_cast_nonneg [typ, none, none, p] + | `(rat.cast_coe), nonzero p := nonzero <$> + mk_mapp ``rat_cast_ne_zero [typ, none, none, none, p] | `(@coe_base _ _ int.has_coe), positive p := positive <$> mk_app ``int_coe_nat_pos [p] | `(@coe_base _ _ int.has_coe), _ := nonnegative <$> mk_app ``int_coe_nat_nonneg [a] | _, _ := failed diff --git a/test/positivity.lean b/test/positivity.lean index fdf579374298c..05bfb619c2efe 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -29,22 +29,92 @@ example : 0 < 3 := by positivity /- ## Goals working directly from a hypothesis -/ -example {a : ℤ} (ha : 0 ≤ a) : 0 ≤ a := by positivity - +example {a : ℤ} (ha : 0 < a) : 0 < a := by positivity example {a : ℤ} (ha : 0 < a) : 0 ≤ a := by positivity +example {a : ℤ} (ha : 0 < a) : a ≠ 0 := by positivity +example {a : ℤ} (ha : 0 ≤ a) : 0 ≤ a := by positivity +example {a : ℤ} (ha : a ≠ 0) : a ≠ 0 := by positivity +example {a : ℤ} (ha : a = 0) : 0 ≤ a := by positivity -example {a : ℤ} (ha : 0 < a) : 0 < a := by positivity +/- ### Reversing hypotheses -/ +example {a : ℤ} (ha : a > 0) : 0 < a := by positivity +example {a : ℤ} (ha : a > 0) : 0 ≤ a := by positivity +example {a : ℤ} (ha : a > 0) : a ≠ 0 := by positivity +example {a : ℤ} (ha : a ≥ 0) : 0 ≤ a := by positivity +example {a : ℤ} (ha : 0 ≠ a) : a ≠ 0 := by positivity +example {a : ℤ} (ha : 0 < a) : a > 0 := by positivity +example {a : ℤ} (ha : 0 < a) : a ≥ 0 := by positivity +example {a : ℤ} (ha : 0 < a) : 0 ≠ a := by positivity +example {a : ℤ} (ha : 0 ≤ a) : a ≥ 0 := by positivity +example {a : ℤ} (ha : a ≠ 0) : 0 ≠ a := by positivity +example {a : ℤ} (ha : a = 0) : a ≥ 0 := by positivity +example {a : ℤ} (ha : 0 = a) : 0 ≤ a := by positivity +example {a : ℤ} (ha : 0 = a) : a ≥ 0 := by positivity + +/- ### Calling `norm_num` -/ + +example {a : ℤ} (ha : 3 = a) : 0 ≤ a := by positivity +example {a : ℤ} (ha : 3 = a) : a ≠ 0 := by positivity +example {a : ℤ} (ha : 3 = a) : 0 < a := by positivity +example {a : ℤ} (ha : a = -1) : a ≠ 0 := by positivity + +example {a : ℤ} (ha : 3 ≤ a) : 0 ≤ a := by positivity +example {a : ℤ} (ha : 3 ≤ a) : a ≠ 0 := by positivity example {a : ℤ} (ha : 3 ≤ a) : 0 < a := by positivity example {a : ℤ} (ha : 3 < a) : 0 ≤ a := by positivity - +example {a : ℤ} (ha : 3 < a) : a ≠ 0 := by positivity example {a : ℤ} (ha : 3 < a) : 0 < a := by positivity example {a b : ℤ} (h : 0 ≤ a + b) : 0 ≤ a + b := by positivity +example {a : ℤ} (hlt : 0 ≤ a) (hne : a ≠ 0) : 0 < a := by positivity + /- ## Tests of the @[positivity] plugin tactics (addition, multiplication, division) -/ +example {a b : ℚ} (ha : 0 < a) (hb : 0 < b) : 0 < min a b := by positivity +example {a b : ℚ} (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ min a b := by positivity +example {a b : ℚ} (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ min a b := by positivity +example {a b : ℚ} (ha : 0 < a) (hb : b ≠ 0) : min a b ≠ 0 := by positivity +example {a b : ℚ} (ha : a ≠ 0) (hb : 0 < b) : min a b ≠ 0 := by positivity +example {a b : ℚ} (ha : a ≠ 0) (hb : b ≠ 0) : min a b ≠ 0 := by positivity + +example {a b : ℚ} (ha : 0 < a) (hb : 0 < b) : 0 < a * b := by positivity +example {a b : ℚ} (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ a * b := by positivity +example {a b : ℚ} (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ a * b := by positivity +example {a b : ℚ} (ha : 0 < a) (hb : b ≠ 0) : a * b ≠ 0 := by positivity +example {a b : ℚ} (ha : a ≠ 0) (hb : 0 < b) : a * b ≠ 0 := by positivity +example {a b : ℚ} (ha : a ≠ 0) (hb : b ≠ 0) : a * b ≠ 0 := by positivity + +example {a b : ℚ} (ha : 0 < a) (hb : 0 < b) : 0 < a / b := by positivity +example {a b : ℚ} (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ a / b := by positivity +example {a b : ℚ} (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ a / b := by positivity +example {a b : ℚ} (ha : 0 < a) (hb : b ≠ 0) : a / b ≠ 0 := by positivity +example {a b : ℚ} (ha : a ≠ 0) (hb : 0 < b) : a / b ≠ 0 := by positivity +example {a b : ℚ} (ha : a ≠ 0) (hb : b ≠ 0) : a / b ≠ 0 := by positivity + +example {a : ℚ} (ha : 0 < a) : 0 < a⁻¹ := by positivity +example {a : ℚ} (ha : 0 ≤ a) : 0 ≤ a⁻¹ := by positivity +example {a : ℚ} (ha : a ≠ 0) : a⁻¹ ≠ 0 := by positivity + +example {a : ℚ} (n : ℕ) (ha : 0 < a) : 0 < a ^ n := by positivity +example {a : ℚ} (n : ℕ) (ha : 0 ≤ a) : 0 ≤ a ^ n := by positivity +example {a : ℚ} (n : ℕ) (ha : a ≠ 0) : a ^ n ≠ 0 := by positivity +example {a : ℚ} (n : ℕ) : 0 ≤ a ^ bit0 n := by positivity +example {a : ℚ} (n : ℕ) (ha : a ≠ 0) : 0 < a ^ bit0 n := by positivity + +example {a : ℚ} (ha : 0 < a) : 0 < |a| := by positivity +example {a : ℚ} (ha : a ≠ 0) : 0 < |a| := by positivity +example (a : ℚ) : 0 ≤ |a| := by positivity + +example {a : ℤ} {b : ℚ} (ha : 0 < a) (hb : 0 < b) : 0 < a • b := by positivity +example {a : ℤ} {b : ℚ} (ha : 0 < a) (hb : 0 ≤ b) : 0 ≤ a • b := by positivity +example {a : ℤ} {b : ℚ} (ha : 0 ≤ a) (hb : 0 < b) : 0 ≤ a • b := by positivity +example {a : ℤ} {b : ℚ} (ha : 0 < a) (hb : b ≠ 0) : a • b ≠ 0 := by positivity +example {a : ℤ} {b : ℚ} (ha : a ≠ 0) (hb : 0 < b) : a • b ≠ 0 := by positivity +example {a : ℤ} {b : ℚ} (ha : a ≠ 0) (hb : b ≠ 0) : a • b ≠ 0 := by positivity + example {a : ℤ} (ha : 3 < a) : 0 ≤ a + a := by positivity example {a b : ℤ} (ha : 3 < a) (hb : 4 ≤ b) : 0 ≤ 3 + a + b + b + 14 := by positivity @@ -64,12 +134,12 @@ example {a : ℤ} (ha : 0 < a) : 0 < a / a := by positivity /-! ### Exponentiation -/ example [ordered_semiring α] [nontrivial α] (a : α) : 0 < a ^ 0 := by positivity -example [linear_ordered_ring α] (a : α) (n : ℕ) : 0 ≤ a ^ (bit0 n) := by positivity +example [linear_ordered_ring α] (a : α) (n : ℕ) : 0 ≤ a ^ bit0 n := by positivity example [ordered_semiring α] {a : α} {n : ℕ} (ha : 0 ≤ a) : 0 ≤ a ^ n := by positivity example [ordered_semiring α] {a : α} {n : ℕ} (ha : 0 < a) : 0 < a ^ n := by positivity example [linear_ordered_semifield α] (a : α) : 0 < a ^ (0 : ℤ) := by positivity -example [linear_ordered_field α] (a : α) (n : ℤ) : 0 ≤ a ^ (bit0 n) := by positivity +example [linear_ordered_field α] (a : α) (n : ℤ) : 0 ≤ a ^ bit0 n := by positivity example [linear_ordered_semifield α] {a : α} {n : ℤ} (ha : 0 ≤ a) : 0 ≤ a ^ n := by positivity example [linear_ordered_semifield α] {a : α} {n : ℤ} (ha : 0 < a) : 0 < a ^ n := by positivity @@ -99,13 +169,12 @@ example {a b : ℤ} (ha : 3 < a) (hb : b ≥ 4) : 0 ≤ 3 * a ^ 2 * b + b * 7 + example {a b : ℤ} (ha : 3 < a) (hb : b ≥ 4) : 0 < 3 * a ^ 2 * b + b * 7 + 14 := by positivity -example {x : ℚ} (hx : 0 ≤ x) : 0 ≤ x⁻¹ := by positivity - example {a : ℤ} : 0 ≤ |a| := by positivity example {a : ℤ} : 0 < |a| + 3 := by positivity example {n : ℤ} (hn : 0 < n) : 0 < n.nat_abs := by positivity +example {n : ℤ} (hn : n ≠ 0) : 0 < n.nat_abs := by positivity example {a : ℤ} (ha : 1 < a) : 0 < |(3:ℤ) + a| := by positivity @@ -167,8 +236,10 @@ example {a : ℝ≥0∞} : 0 ≤ a := by positivity example {a : ℕ} : (0 : ℤ) ≤ a := by positivity example {a : ℕ} (ha : 0 < a) : (0 : ℤ) < a := by positivity +example {a : ℤ} (ha : a ≠ 0) : (a : ℚ) ≠ 0 := by positivity example {a : ℤ} (ha : 0 ≤ a) : (0 : ℚ) ≤ a := by positivity example {a : ℤ} (ha : 0 < a) : (0 : ℚ) < a := by positivity +example {a : ℚ} (ha : a ≠ 0) : (a : ℝ) ≠ 0 := by positivity example {a : ℚ} (ha : 0 ≤ a) : (0 : ℝ) ≤ a := by positivity example {a : ℚ} (ha : 0 < a) : (0 : ℝ) < a := by positivity example {r : ℝ≥0} : (0 : ℝ) ≤ r := by positivity From b8e7a8d9ce8d8fb73d6e81392620262fd22cf343 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 5 Oct 2022 09:33:04 +0000 Subject: [PATCH 574/828] feat(linear_algebra/affine_space/combination): `sdiff`, `subtype`, `filter` lemmas (#16644) Add various lemmas about affine combinations in relation to subsets of the index type extracted with `sdiff`, `subtype` or `filter` (largely analogous to such lemmas present for `finset.sum`). --- .../affine_space/combination.lean | 95 +++++++++++++++++++ 1 file changed, 95 insertions(+) diff --git a/src/linear_algebra/affine_space/combination.lean b/src/linear_algebra/affine_space/combination.lean index cbfaee749a4c1..2e114b79778f1 100644 --- a/src/linear_algebra/affine_space/combination.lean +++ b/src/linear_algebra/affine_space/combination.lean @@ -179,6 +179,38 @@ lemma sum_smul_const_vsub_eq_sub_weighted_vsub_of_point (w : ι → k) (p₂ : ∑ i in s, w i • (p₁ -ᵥ p₂ i) = (∑ i in s, w i) • (p₁ -ᵥ b) - s.weighted_vsub_of_point p₂ b w := by rw [sum_smul_vsub_eq_weighted_vsub_of_point_sub, weighted_vsub_of_point_apply_const] +/-- A weighted sum may be split into such sums over two subsets. -/ +lemma weighted_vsub_of_point_sdiff [decidable_eq ι] {s₂ : finset ι} (h : s₂ ⊆ s) (w : ι → k) + (p : ι → P) (b : P) : (s \ s₂).weighted_vsub_of_point p b w + s₂.weighted_vsub_of_point p b w = + s.weighted_vsub_of_point p b w := +by simp_rw [weighted_vsub_of_point_apply, sum_sdiff h] + +/-- A weighted sum may be split into a subtraction of such sums over two subsets. -/ +lemma weighted_vsub_of_point_sdiff_sub [decidable_eq ι] {s₂ : finset ι} (h : s₂ ⊆ s) (w : ι → k) + (p : ι → P) (b : P) : (s \ s₂).weighted_vsub_of_point p b w - s₂.weighted_vsub_of_point p b (-w) = + s.weighted_vsub_of_point p b w := +by rw [map_neg, sub_neg_eq_add, s.weighted_vsub_of_point_sdiff h] + +/-- A weighted sum over `s.subtype pred` equals one over `s.filter pred`. -/ +lemma weighted_vsub_of_point_subtype_eq_filter (w : ι → k) (p : ι → P) (b : P) + (pred : ι → Prop) [decidable_pred pred] : + (s.subtype pred).weighted_vsub_of_point (λ i, p i) b (λ i, w i) = + (s.filter pred).weighted_vsub_of_point p b w := +by rw [weighted_vsub_of_point_apply, weighted_vsub_of_point_apply, ←sum_subtype_eq_sum_filter] + +/-- A weighted sum over `s.filter pred` equals one over `s` if all the weights at indices in `s` +not satisfying `pred` are zero. -/ +lemma weighted_vsub_of_point_filter_of_ne (w : ι → k) (p : ι → P) (b : P) {pred : ι → Prop} + [decidable_pred pred] (h : ∀ i ∈ s, w i ≠ 0 → pred i) : + (s.filter pred).weighted_vsub_of_point p b w = s.weighted_vsub_of_point p b w := +begin + rw [weighted_vsub_of_point_apply, weighted_vsub_of_point_apply, sum_filter_of_ne], + intros i hi hne, + refine h i hi _, + intro hw, + simpa [hw] using hne, +end + /-- A weighted sum of the results of subtracting a default base point from the given points, as a linear map on the weights. This is intended to be used when the sum of the weights is 0; that condition @@ -246,6 +278,29 @@ lemma sum_smul_const_vsub_eq_neg_weighted_vsub (w : ι → k) (p₂ : ι → P) ∑ i in s, w i • (p₁ -ᵥ p₂ i) = -s.weighted_vsub p₂ w := by rw [sum_smul_vsub_eq_weighted_vsub_sub, s.weighted_vsub_apply_const _ _ h, zero_sub] +/-- A weighted sum may be split into such sums over two subsets. -/ +lemma weighted_vsub_sdiff [decidable_eq ι] {s₂ : finset ι} (h : s₂ ⊆ s) (w : ι → k) + (p : ι → P) : (s \ s₂).weighted_vsub p w + s₂.weighted_vsub p w = s.weighted_vsub p w := +s.weighted_vsub_of_point_sdiff h _ _ _ + +/-- A weighted sum may be split into a subtraction of such sums over two subsets. -/ +lemma weighted_vsub_sdiff_sub [decidable_eq ι] {s₂ : finset ι} (h : s₂ ⊆ s) (w : ι → k) + (p : ι → P) : (s \ s₂).weighted_vsub p w - s₂.weighted_vsub p (-w) = s.weighted_vsub p w := +s.weighted_vsub_of_point_sdiff_sub h _ _ _ + +/-- A weighted sum over `s.subtype pred` equals one over `s.filter pred`. -/ +lemma weighted_vsub_subtype_eq_filter (w : ι → k) (p : ι → P) (pred : ι → Prop) + [decidable_pred pred] : + (s.subtype pred).weighted_vsub (λ i, p i) (λ i, w i) = (s.filter pred).weighted_vsub p w := +s.weighted_vsub_of_point_subtype_eq_filter _ _ _ _ + +/-- A weighted sum over `s.filter pred` equals one over `s` if all the weights at indices in `s` +not satisfying `pred` are zero. -/ +lemma weighted_vsub_filter_of_ne (w : ι → k) (p : ι → P) {pred : ι → Prop} + [decidable_pred pred] (h : ∀ i ∈ s, w i ≠ 0 → pred i) : + (s.filter pred).weighted_vsub p w = s.weighted_vsub p w := +s.weighted_vsub_of_point_filter_of_ne _ _ _ h + /-- A weighted sum of the results of subtracting a default base point from the given points, added to that base point, as an affine map on the weights. This is intended to be used when the sum of the weights @@ -390,6 +445,46 @@ lemma sum_smul_const_vsub_eq_vsub_affine_combination (w : ι → k) (p₂ : ι ∑ i in s, w i • (p₁ -ᵥ p₂ i) = p₁ -ᵥ s.affine_combination p₂ w := by rw [sum_smul_vsub_eq_affine_combination_vsub, affine_combination_apply_const _ _ _ h] +/-- A weighted sum may be split into a subtraction of affine combinations over two subsets. -/ +lemma affine_combination_sdiff_sub [decidable_eq ι] {s₂ : finset ι} (h : s₂ ⊆ s) (w : ι → k) + (p : ι → P) : + (s \ s₂).affine_combination p w -ᵥ s₂.affine_combination p (-w) = s.weighted_vsub p w := +begin + simp_rw [affine_combination_apply, vadd_vsub_vadd_cancel_right], + exact s.weighted_vsub_sdiff_sub h _ _ +end + +/-- If a weighted sum is zero and one of the weights is `-1`, the corresponding point is +the affine combination of the other points with the given weights. -/ +lemma affine_combination_eq_of_weighted_vsub_eq_zero_of_eq_neg_one {w : ι → k} {p : ι → P} + (hw : s.weighted_vsub p w = (0 : V)) {i : ι} [decidable_pred (≠ i)] (his : i ∈ s) + (hwi : w i = -1) : (s.filter (≠ i)).affine_combination p w = p i := +begin + classical, + rw [←@vsub_eq_zero_iff_eq V, ←hw, + ←s.affine_combination_sdiff_sub (singleton_subset_iff.2 his), sdiff_singleton_eq_erase, + ←filter_ne'], + congr, + refine (affine_combination_of_eq_one_of_eq_zero _ _ _ (mem_singleton_self _) _ _).symm, + { simp [hwi] }, + { simp } +end + +/-- An affine combination over `s.subtype pred` equals one over `s.filter pred`. -/ +lemma affine_combination_subtype_eq_filter (w : ι → k) (p : ι → P) (pred : ι → Prop) + [decidable_pred pred] : + (s.subtype pred).affine_combination (λ i, p i) (λ i, w i) = + (s.filter pred).affine_combination p w := +by rw [affine_combination_apply, affine_combination_apply, weighted_vsub_of_point_subtype_eq_filter] + +/-- An affine combination over `s.filter pred` equals one over `s` if all the weights at indices +in `s` not satisfying `pred` are zero. -/ +lemma affine_combination_filter_of_ne (w : ι → k) (p : ι → P) {pred : ι → Prop} + [decidable_pred pred] (h : ∀ i ∈ s, w i ≠ 0 → pred i) : + (s.filter pred).affine_combination p w = s.affine_combination p w := +by rw [affine_combination_apply, affine_combination_apply, + s.weighted_vsub_of_point_filter_of_ne _ _ _ h] + variables {V} /-- Suppose an indexed family of points is given, along with a subset From 095a77bb04499791977270ffc881b331422b0667 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 5 Oct 2022 09:33:05 +0000 Subject: [PATCH 575/828] feat(linear_algebra/tensor_product): add variant of ext_fourfold (#16724) This came up in Oliver Nash's Xena workshop project. --- src/linear_algebra/tensor_product.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/linear_algebra/tensor_product.lean b/src/linear_algebra/tensor_product.lean index c5f504961e7b4..e4bc0279f7e81 100644 --- a/src/linear_algebra/tensor_product.lean +++ b/src/linear_algebra/tensor_product.lean @@ -526,6 +526,15 @@ begin exact H w x y z, end +/-- Two linear maps (M ⊗ N) ⊗ (P ⊗ Q) → S which agree on all elements of the +form (m ⊗ₜ n) ⊗ₜ (p ⊗ₜ q) are equal. -/ +theorem ext_fourfold' {φ ψ : (M ⊗[R] N) ⊗[R] (P ⊗[R] Q) →ₗ[R] S} + (H : ∀ w x y z, φ ((w ⊗ₜ x) ⊗ₜ (y ⊗ₜ z)) = ψ ((w ⊗ₜ x) ⊗ₜ (y ⊗ₜ z))) : φ = ψ := +begin + ext m n p q, + exact H m n p q, +end + end UMP variables {M N} From 7c1d0d8e6ebee9b0324003f5520b12780f0ba501 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Wed, 5 Oct 2022 09:33:06 +0000 Subject: [PATCH 576/828] misc(linear_algebra, ring_theory): more simple facts about `is_simple_module` and order on submodules (#16786) The main result is `is_coatom_ker_of_surjective`, but I added various miscellaneous lemmas along the way --- src/algebra/module/equiv.lean | 10 ++++++---- src/linear_algebra/basic.lean | 25 +++++++++++++++++++++++++ src/ring_theory/ideal/basic.lean | 4 ++++ src/ring_theory/simple_module.lean | 19 +++++++++++++++++++ 4 files changed, 54 insertions(+), 4 deletions(-) diff --git a/src/algebra/module/equiv.lean b/src/algebra/module/equiv.lean index 7d35aad669c66..e5fb497d8656f 100644 --- a/src/algebra/module/equiv.lean +++ b/src/algebra/module/equiv.lean @@ -340,13 +340,15 @@ omit module_M₃ @[simp] lemma refl_symm [module R M] : (refl R M).symm = linear_equiv.refl R M := rfl -@[simp] lemma self_trans_symm [module R M] [module R M₂] (f : M ≃ₗ[R] M₂) : - f.trans f.symm = linear_equiv.refl R M := +include re₁₂ re₂₁ module_M₁ module_M₂ +@[simp] lemma self_trans_symm (f : M₁ ≃ₛₗ[σ₁₂] M₂) : + f.trans f.symm = linear_equiv.refl R₁ M₁ := by { ext x, simp } -@[simp] lemma symm_trans_self [module R M] [module R M₂] (f : M ≃ₗ[R] M₂) : - f.symm.trans f = linear_equiv.refl R M₂ := +@[simp] lemma symm_trans_self (f : M₁ ≃ₛₗ[σ₁₂] M₂) : + f.symm.trans f = linear_equiv.refl R₂ M₂ := by { ext x, simp } +omit re₁₂ re₂₁ module_M₁ module_M₂ @[simp, norm_cast] lemma refl_to_linear_map [module R M] : (linear_equiv.refl R M : M →ₗ[R] M) = linear_map.id := diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 9292567519370..9369f5c9cd7fe 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -808,6 +808,21 @@ lemma map_strict_mono_of_injective : strict_mono (map f) := end galois_coinsertion +section order_iso +omit sc +include σ₁₂ σ₂₁ +variables [semilinear_equiv_class F σ₁₂ M M₂] + +/-- A linear isomorphism induces an order isomorphism of submodules. -/ +@[simps symm_apply apply] def order_iso_map_comap (f : F) : submodule R M ≃o submodule R₂ M₂ := +{ to_fun := map f, + inv_fun := comap f, + left_inv := comap_map_eq_of_injective $ equiv_like.injective f, + right_inv := map_comap_eq_of_surjective $ equiv_like.surjective f, + map_rel_iff' := map_le_map_iff_of_injective $ equiv_like.injective f } + +end order_iso + --TODO(Mario): is there a way to prove this from order properties? lemma map_inf_eq_map_inf_comap [ring_hom_surjective σ₁₂] {f : F} {p : submodule R M} {p' : submodule R₂ M₂} : @@ -1963,6 +1978,16 @@ lemma comap_equiv_eq_map_symm (e : M ≃ₛₗ[τ₁₂] M₂) (K : submodule R K.comap (e : M →ₛₗ[τ₁₂] M₂) = K.map (e.symm : M₂ →ₛₗ[τ₂₁] M) := (map_equiv_eq_comap_symm e.symm K).symm +include τ₂₁ +lemma order_iso_map_comap_apply' (e : M ≃ₛₗ[τ₁₂] M₂) (p : submodule R M) : + order_iso_map_comap e p = comap e.symm p := +p.map_equiv_eq_comap_symm _ + +lemma order_iso_map_comap_symm_apply' (e : M ≃ₛₗ[τ₁₂] M₂) (p : submodule R₂ M₂) : + (order_iso_map_comap e).symm p = map e.symm p := +p.comap_equiv_eq_map_symm _ +omit τ₂₁ + lemma comap_le_comap_smul (fₗ : N →ₗ[R] N₂) (c : R) : comap fₗ qₗ ≤ comap (c • fₗ) qₗ := begin diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index af7b52c4b160c..b51c3ad89cbe2 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -519,6 +519,10 @@ begin simpa [H, h1] using I.mul_mem_left r⁻¹ hr, end +/-- Ideals of a `division_ring` are a simple order. Thanks to the way abbreviations work, this +automatically gives a `is_simple_module K` instance. -/ +instance : is_simple_order (ideal K) := ⟨eq_bot_or_top⟩ + lemma eq_bot_of_prime [h : I.is_prime] : I = ⊥ := or_iff_not_imp_right.mp I.eq_bot_or_top h.1 diff --git a/src/ring_theory/simple_module.lean b/src/ring_theory/simple_module.lean index 9036f849b7885..a737e013f6ba8 100644 --- a/src/ring_theory/simple_module.lean +++ b/src/ring_theory/simple_module.lean @@ -6,6 +6,7 @@ Authors: Aaron Anderson import linear_algebra.span import order.atoms +import linear_algebra.isomorphisms /-! # Simple Modules @@ -47,6 +48,9 @@ end⟩⟩ variables {R} {M} {m : submodule R M} {N : Type*} [add_comm_group N] [module R N] +lemma is_simple_module.congr (l : M ≃ₗ[R] N) [is_simple_module R N] : is_simple_module R M := +(submodule.order_iso_map_comap l).is_simple_order + theorem is_simple_module_iff_is_atom : is_simple_module R m ↔ is_atom m := begin @@ -55,6 +59,14 @@ begin exact submodule.map_subtype.rel_iso m, end +theorem is_simple_module_iff_is_coatom : + is_simple_module R (M ⧸ m) ↔ is_coatom m := +begin + rw ← set.is_simple_order_Ici_iff_is_coatom, + apply order_iso.is_simple_order_iff, + exact submodule.comap_mkq.rel_iso m, +end + namespace is_simple_module variable [hm : is_simple_module R m] @@ -131,6 +143,13 @@ theorem bijective_of_ne_zero [is_simple_module R M] [is_simple_module R N] function.bijective f := f.bijective_or_eq_zero.resolve_right h +theorem is_coatom_ker_of_surjective [is_simple_module R N] {f : M →ₗ[R] N} + (hf : function.surjective f) : is_coatom f.ker := +begin + rw ←is_simple_module_iff_is_coatom, + exact is_simple_module.congr (f.quot_ker_equiv_of_surjective hf) +end + /-- Schur's Lemma makes the endomorphism ring of a simple module a division ring. -/ noncomputable instance _root_.module.End.division_ring [decidable_eq (module.End R M)] [is_simple_module R M] : From 0830c456112bdb856fd779bb3910c666a831f8bf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 5 Oct 2022 09:33:08 +0000 Subject: [PATCH 577/828] feat(algebra/gcd_monoid/finset): Generalize `finset.gcd_image` (#16795) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `finset.gcd_image` and `finset.gcd_eq_gcd_image` were unnecessarily assuming `[is_idempotent α gcd]`. Remove that assumption and add the corresponding `finset.lcm` lemmas. --- src/algebra/gcd_monoid/finset.lean | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/src/algebra/gcd_monoid/finset.lean b/src/algebra/gcd_monoid/finset.lean index 496290e35e67b..40a2dd8ec9beb 100644 --- a/src/algebra/gcd_monoid/finset.lean +++ b/src/algebra/gcd_monoid/finset.lean @@ -87,6 +87,11 @@ lcm_dvd (λ b hb, (h b hb).trans (dvd_lcm hb)) lemma lcm_mono (h : s₁ ⊆ s₂) : s₁.lcm f ∣ s₂.lcm f := lcm_dvd $ assume b hb, dvd_lcm (h hb) +lemma lcm_image [decidable_eq β] {g : γ → β} (s : finset γ) : (s.image g).lcm f = s.lcm (f ∘ g) := +by { classical, induction s using finset.induction with c s hc ih; simp [*] } + +lemma lcm_eq_lcm_image [decidable_eq α] : s.lcm f = (s.image f).lcm id := eq.symm $ lcm_image _ + theorem lcm_eq_zero_iff [nontrivial α] : s.lcm f = 0 ↔ 0 ∈ f '' s := by simp only [multiset.mem_map, lcm_def, multiset.lcm_eq_zero_iff, set.mem_image, mem_coe, ← finset.mem_def] @@ -147,11 +152,10 @@ dvd_gcd (λ b hb, (gcd_dvd hb).trans (h b hb)) lemma gcd_mono (h : s₁ ⊆ s₂) : s₂.gcd f ∣ s₁.gcd f := dvd_gcd $ assume b hb, gcd_dvd (h hb) -theorem gcd_image {g : γ → β} (s: finset γ) [decidable_eq β] [is_idempotent α gcd_monoid.gcd] : - (s.image g).gcd f = s.gcd (f ∘ g) := by simp [gcd, fold_image_idem] +lemma gcd_image [decidable_eq β] {g : γ → β} (s : finset γ) : (s.image g).gcd f = s.gcd (f ∘ g) := +by { classical, induction s using finset.induction with c s hc ih; simp [*] } -theorem gcd_eq_gcd_image [decidable_eq α] [is_idempotent α gcd_monoid.gcd] : - s.gcd f = (s.image f).gcd id := (@gcd_image _ _ _ _ _ id _ _ _ _).symm +lemma gcd_eq_gcd_image [decidable_eq α] : s.gcd f = (s.image f).gcd id := eq.symm $ gcd_image _ theorem gcd_eq_zero_iff : s.gcd f = 0 ↔ ∀ (x : β), x ∈ s → f x = 0 := begin From 1b2d1a4a544576f8cdaf72d192e3c03ab0ec99c1 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Wed, 5 Oct 2022 09:33:09 +0000 Subject: [PATCH 578/828] feat(group_theory/sylow): Technical fusion lemma for Burnside's transfer theorem (#16800) Burnside's transfer theorem requires a technical lemma obtained by applying Sylow's second theorem to the centralizer of an element. This technical lemmas states that if `P` is abelian, then `G`-conjugate elements of `P` are `N(P)`-conjugate. The fancy way of saying this is that "`N(P)` controls fusion in `P`. --- src/group_theory/sylow.lean | 24 ++++++++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index f0d69dddb263a..5d5c5449025c6 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -296,6 +296,30 @@ lemma sylow.stabilizer_eq_normalizer (P : sylow p G) : stabilizer G P = (P : subgroup G).normalizer := ext (λ g, sylow.smul_eq_iff_mem_normalizer) +lemma sylow.conj_eq_normalizer_conj_of_mem_centralizer + [fact p.prime] [finite (sylow p G)] (P : sylow p G) (x g : G) + (hx : x ∈ (P : subgroup G).centralizer) (hy : g⁻¹ * x * g ∈ (P : subgroup G).centralizer) : + ∃ n ∈ (P : subgroup G).normalizer, g⁻¹ * x * g = n⁻¹ * x * n := +begin + have h1 : ↑P ≤ (zpowers x).centralizer, + { rwa [le_centralizer_iff, zpowers_le] }, + have h2 : ↑(g • P) ≤ (zpowers x).centralizer, + { rw [le_centralizer_iff, zpowers_le], + rintros - ⟨z, hz, rfl⟩, + specialize hy z hz, + rwa [←mul_assoc, ←eq_mul_inv_iff_mul_eq, mul_assoc, mul_assoc, mul_assoc, ←mul_assoc, + eq_inv_mul_iff_mul_eq, ←mul_assoc, ←mul_assoc] at hy }, + obtain ⟨h, hh⟩ := exists_smul_eq (zpowers x).centralizer ((g • P).subtype h2) (P.subtype h1), + simp_rw [sylow.smul_subtype, smul_def, smul_smul] at hh, + refine ⟨h * g, sylow.smul_eq_iff_mem_normalizer.mp (sylow.subtype_injective hh), _⟩, + rw [←mul_assoc, commute.right_comm (h.prop x (mem_zpowers x)), mul_inv_rev, inv_mul_cancel_right] +end + +lemma sylow.conj_eq_normalizer_conj_of_mem [fact p.prime] [finite (sylow p G)] (P : sylow p G) + [hP : (P : subgroup G).is_commutative] (x g : G) (hx : x ∈ P) (hy : g⁻¹ * x * g ∈ P) : + ∃ n ∈ (P : subgroup G).normalizer, g⁻¹ * x * g = n⁻¹ * x * n := +P.conj_eq_normalizer_conj_of_mem_centralizer x g (le_centralizer P hx) (le_centralizer P hy) + /-- Sylow `p`-subgroups are in bijection with cosets of the normalizer of a Sylow `p`-subgroup -/ noncomputable def sylow.equiv_quotient_normalizer [fact p.prime] [fintype (sylow p G)] (P : sylow p G) : sylow p G ≃ G ⧸ (P : subgroup G).normalizer := From fd4806c1e951d64122af3cee5d8ec482e00d7e55 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 5 Oct 2022 09:33:10 +0000 Subject: [PATCH 579/828] feat(topology/algebra/module/weak_dual): add a `t2_space` instance for the `weak_dual` (#16801) --- src/topology/algebra/module/weak_dual.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/topology/algebra/module/weak_dual.lean b/src/topology/algebra/module/weak_dual.lean index 627aa29f3062c..e942a19eef6b2 100644 --- a/src/topology/algebra/module/weak_dual.lean +++ b/src/topology/algebra/module/weak_dual.lean @@ -238,6 +238,14 @@ lemma continuous_of_continuous_eval [topological_space α] {g : α → weak_dual (h : ∀ y, continuous (λ a, (g a) y)) : continuous g := continuous_induced_rng.2 (continuous_pi_iff.mpr h) +instance [t2_space 𝕜] : t2_space (weak_dual 𝕜 E) := +begin + refine t2_iff_is_closed_diagonal.mpr _, + simp_rw [set.diagonal, fun_like.ext_iff, set.set_of_forall], + exact is_closed_Inter (λ x, is_closed_eq ((weak_dual.eval_continuous x).comp continuous_fst) $ + (weak_dual.eval_continuous x).comp continuous_snd), +end + end weak_dual /-- The weak topology is the topology coarsest topology on `E` such that all From 7316286ff2942aa14e540add9058c6b0aa1c8070 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 5 Oct 2022 12:41:17 +0000 Subject: [PATCH 580/828] feat(algebraic_geometry/morphisms/basic): Morphism properties on the diagonal morphism. (#16111) --- src/algebraic_geometry/morphisms/basic.lean | 124 ++++++++++++++++++ src/algebraic_geometry/pullbacks.lean | 43 +++++- .../limits/shapes/pullbacks.lean | 26 ++++ src/category_theory/morphism_property.lean | 49 ++++++- 4 files changed, 239 insertions(+), 3 deletions(-) diff --git a/src/algebraic_geometry/morphisms/basic.lean b/src/algebraic_geometry/morphisms/basic.lean index 31d5de9befe66..1369fbf22f0a0 100644 --- a/src/algebraic_geometry/morphisms/basic.lean +++ b/src/algebraic_geometry/morphisms/basic.lean @@ -436,4 +436,128 @@ end end affine_target_morphism_property +/-- +The `affine_target_morphism_property` associated to `(target_affine_locally P).diagonal`. +See `diagonal_target_affine_locally_eq_target_affine_locally`. +-/ +def affine_target_morphism_property.diagonal (P : affine_target_morphism_property) : + affine_target_morphism_property := +λ X Y f hf, ∀ {U₁ U₂ : Scheme} (f₁ : U₁ ⟶ X) (f₂ : U₂ ⟶ X) [is_affine U₁] [is_affine U₂] + [is_open_immersion f₁] [is_open_immersion f₂], + by exactI P (pullback.map_desc f₁ f₂ f) + +lemma affine_target_morphism_property.diagonal_respects_iso (P : affine_target_morphism_property) + (hP : P.to_property.respects_iso) : + P.diagonal.to_property.respects_iso := +begin + delta affine_target_morphism_property.diagonal, + apply affine_target_morphism_property.respects_iso_mk, + { introv H _ _, + resetI, + rw [pullback.map_desc_comp, affine_cancel_left_is_iso hP, affine_cancel_right_is_iso hP], + apply H }, + { introv H _ _, + resetI, + rw [pullback.map_desc_comp, affine_cancel_right_is_iso hP], + apply H } +end + +lemma diagonal_target_affine_locally_of_open_cover (P : affine_target_morphism_property) + (hP : P.is_local) + {X Y : Scheme.{u}} (f : X ⟶ Y) + (𝒰 : Scheme.open_cover.{u} Y) + [∀ i, is_affine (𝒰.obj i)] (𝒰' : Π i, Scheme.open_cover.{u} (pullback f (𝒰.map i))) + [∀ i j, is_affine ((𝒰' i).obj j)] + (h𝒰' : ∀ i j k, P (pullback.map_desc ((𝒰' i).map j) ((𝒰' i).map k) pullback.snd)) : + (target_affine_locally P).diagonal f := +begin + refine (hP.affine_open_cover_iff _ _).mpr _, + { exact ((Scheme.pullback.open_cover_of_base 𝒰 f f).bind (λ i, + Scheme.pullback.open_cover_of_left_right.{u u} (𝒰' i) (𝒰' i) pullback.snd pullback.snd)) }, + { intro i, + dsimp at *, + apply_instance }, + { rintro ⟨i, j, k⟩, + dsimp, + convert (affine_cancel_left_is_iso hP.1 + (pullback_diagonal_map_iso _ _ ((𝒰' i).map j) ((𝒰' i).map k)).inv pullback.snd).mp _, + swap 3, + { convert h𝒰' i j k, apply pullback.hom_ext; simp, }, + all_goals + { apply pullback.hom_ext; simp only [category.assoc, pullback.lift_fst, pullback.lift_snd, + pullback.lift_fst_assoc, pullback.lift_snd_assoc] } } +end + +lemma affine_target_morphism_property.diagonal_of_target_affine_locally + (P : affine_target_morphism_property) + (hP : P.is_local) {X Y U : Scheme.{u}} (f : X ⟶ Y) (g : U ⟶ Y) + [is_affine U] [is_open_immersion g] (H : (target_affine_locally P).diagonal f) : + P.diagonal (pullback.snd : pullback f g ⟶ _) := +begin + rintros U V f₁ f₂ _ _ _ _, + resetI, + replace H := ((hP.affine_open_cover_tfae (pullback.diagonal f)).out 0 3).mp H, + let g₁ := pullback.map (f₁ ≫ pullback.snd) + (f₂ ≫ pullback.snd) f f + (f₁ ≫ pullback.fst) + (f₂ ≫ pullback.fst) g + (by rw [category.assoc, category.assoc, pullback.condition]) + (by rw [category.assoc, category.assoc, pullback.condition]), + let g₂ : pullback f₁ f₂ ⟶ pullback f g := pullback.fst ≫ f₁, + specialize H g₁, + rw ← affine_cancel_left_is_iso hP.1 (pullback_diagonal_map_iso f _ f₁ f₂).hom, + convert H, + { apply pullback.hom_ext; simp only [category.assoc, pullback.lift_fst, pullback.lift_snd, + pullback.lift_fst_assoc, pullback.lift_snd_assoc, category.comp_id, + pullback_diagonal_map_iso_hom_fst, pullback_diagonal_map_iso_hom_snd], } +end + +lemma affine_target_morphism_property.is_local.diagonal_affine_open_cover_tfae + {P : affine_target_morphism_property} + (hP : P.is_local) {X Y : Scheme.{u}} (f : X ⟶ Y) : + tfae [(target_affine_locally P).diagonal f, + ∃ (𝒰 : Scheme.open_cover.{u} Y) [∀ i, is_affine (𝒰.obj i)], by exactI + ∀ (i : 𝒰.J), P.diagonal (pullback.snd : pullback f (𝒰.map i) ⟶ _), + ∀ (𝒰 : Scheme.open_cover.{u} Y) [∀ i, is_affine (𝒰.obj i)] (i : 𝒰.J), by exactI + P.diagonal (pullback.snd : pullback f (𝒰.map i) ⟶ _), + ∀ {U : Scheme} (g : U ⟶ Y) [is_affine U] [is_open_immersion g], by exactI + P.diagonal (pullback.snd : pullback f g ⟶ _), + ∃ (𝒰 : Scheme.open_cover.{u} Y) [∀ i, is_affine (𝒰.obj i)] + (𝒰' : Π i, Scheme.open_cover.{u} (pullback f (𝒰.map i))) [∀ i j, is_affine ((𝒰' i).obj j)], + by exactI ∀ i j k, P (pullback.map_desc ((𝒰' i).map j) ((𝒰' i).map k) pullback.snd)] := +begin + tfae_have : 1 → 4, + { introv H hU hg _ _, resetI, apply P.diagonal_of_target_affine_locally; assumption }, + tfae_have : 4 → 3, + { introv H h𝒰, resetI, apply H }, + tfae_have : 3 → 2, + { exact λ H, ⟨Y.affine_cover, infer_instance, H Y.affine_cover⟩ }, + tfae_have : 2 → 5, + { rintro ⟨𝒰, h𝒰, H⟩, + resetI, + refine ⟨𝒰, infer_instance, λ _, Scheme.affine_cover _, infer_instance, _⟩, + intros i j k, + apply H }, + tfae_have : 5 → 1, + { rintro ⟨𝒰, _, 𝒰', _, H⟩, + exactI diagonal_target_affine_locally_of_open_cover P hP f 𝒰 𝒰' H, }, + tfae_finish +end + +lemma affine_target_morphism_property.is_local.diagonal {P : affine_target_morphism_property} + (hP : P.is_local) : P.diagonal.is_local := +affine_target_morphism_property.is_local_of_open_cover_imply + P.diagonal + (P.diagonal_respects_iso hP.1) + (λ _ _ f, ((hP.diagonal_affine_open_cover_tfae f).out 1 3).mp) + +lemma diagonal_target_affine_locally_eq_target_affine_locally (P : affine_target_morphism_property) + (hP : P.is_local) : + (target_affine_locally P).diagonal = target_affine_locally P.diagonal := +begin + ext _ _ f, + exact ((hP.diagonal_affine_open_cover_tfae f).out 0 1).trans + ((hP.diagonal.affine_open_cover_tfae f).out 1 0), +end + end algebraic_geometry diff --git a/src/algebraic_geometry/pullbacks.lean b/src/algebraic_geometry/pullbacks.lean index d3fbcd6d6fb5a..726ffd94eed13 100644 --- a/src/algebraic_geometry/pullbacks.lean +++ b/src/algebraic_geometry/pullbacks.lean @@ -5,7 +5,8 @@ Authors: Andrew Yang -/ import algebraic_geometry.gluing import category_theory.limits.opposites -import algebraic_geometry.Gamma_Spec_adjunction +import algebraic_geometry.AffineScheme +import category_theory.limits.shapes.diagonal /-! # Fibred products of schemes @@ -245,7 +246,8 @@ def gluing : Scheme.glue_data.{u} := t_fac := λ i j k, begin apply pullback.hom_ext, apply pullback.hom_ext, - all_goals { simp } + all_goals { simp only [t'_snd_fst_fst, t'_snd_fst_snd, t'_snd_snd, + t_fst_fst, t_fst_snd, t_snd, category.assoc] } end, cocycle := λ i j k, cocycle 𝒰 f g i j k } @@ -558,6 +560,13 @@ has_pullback_of_cover (Z.affine_cover.pullback_cover f) f g instance : has_pullbacks Scheme := has_pullbacks_of_has_limit_cospan _ +instance {X Y Z : Scheme} (f : X ⟶ Z) (g : Y ⟶ Z) [is_affine X] [is_affine Y] [is_affine Z] : + is_affine (pullback f g) := +is_affine_of_iso (pullback.map f g (Spec.map (Γ.map f.op).op) (Spec.map (Γ.map g.op).op) + (Γ_Spec.adjunction.unit.app X) (Γ_Spec.adjunction.unit.app Y) (Γ_Spec.adjunction.unit.app Z) + (Γ_Spec.adjunction.unit.naturality f) (Γ_Spec.adjunction.unit.naturality g) ≫ + (preserves_pullback.iso Spec _ _).inv) + /-- Given an open cover `{ Xᵢ }` of `X`, then `X ×[Z] Y` is covered by `Xᵢ ×[Z] Y`. -/ @[simps J obj map] def open_cover_of_left (𝒰 : open_cover X) (f : X ⟶ Z) (g : Y ⟶ Z) : open_cover (pullback f g) := @@ -592,6 +601,23 @@ begin apply pullback.hom_ext; simp, end +/-- Given an open cover `{ Xᵢ }` of `X` and an open cover `{ Yⱼ }` of `Y`, then +`X ×[Z] Y` is covered by `Xᵢ ×[Z] Yⱼ`. -/ +@[simps J obj map] +def open_cover_of_left_right (𝒰X : X.open_cover) (𝒰Y : Y.open_cover) + (f : X ⟶ Z) (g : Y ⟶ Z) : (pullback f g).open_cover := +begin + fapply ((open_cover_of_left 𝒰X f g).bind (λ x, open_cover_of_right 𝒰Y (𝒰X.map x ≫ f) g)).copy + (𝒰X.J × 𝒰Y.J) + (λ ij, pullback (𝒰X.map ij.1 ≫ f) (𝒰Y.map ij.2 ≫ g)) + (λ ij, pullback.map _ _ _ _ (𝒰X.map ij.1) (𝒰Y.map ij.2) (𝟙 _) + (category.comp_id _) (category.comp_id _)) + (equiv.sigma_equiv_prod _ _).symm + (λ _, iso.refl _), + rintro ⟨i, j⟩, + apply pullback.hom_ext; simpa, +end + /-- (Implementation). Use `open_cover_of_base` instead. -/ def open_cover_of_base' (𝒰 : open_cover Z) (f : X ⟶ Z) (g : Y ⟶ Z) : open_cover (pullback f g) := begin @@ -639,3 +665,16 @@ end end pullback end algebraic_geometry.Scheme + +namespace algebraic_geometry + +instance {X Y S X' Y' S' : Scheme} (f : X ⟶ S) (g : Y ⟶ S) (f' : X' ⟶ S') + (g' : Y' ⟶ S') (i₁ : X ⟶ X') (i₂ : Y ⟶ Y') (i₃ : S ⟶ S') (e₁ : f ≫ i₃ = i₁ ≫ f') + (e₂ : g ≫ i₃ = i₂ ≫ g') [is_open_immersion i₁] [is_open_immersion i₂] [mono i₃] : + is_open_immersion (pullback.map f g f' g' i₁ i₂ i₃ e₁ e₂) := +begin + rw pullback_map_eq_pullback_fst_fst_iso_inv, + apply_instance +end + +end algebraic_geometry diff --git a/src/category_theory/limits/shapes/pullbacks.lean b/src/category_theory/limits/shapes/pullbacks.lean index dfd1089f8fbf1..f88e38a7a8145 100644 --- a/src/category_theory/limits/shapes/pullbacks.lean +++ b/src/category_theory/limits/shapes/pullbacks.lean @@ -946,6 +946,12 @@ abbreviation pullback.map {W X Y Z S T : C} (f₁ : W ⟶ S) (f₂ : X ⟶ S) [h pullback.lift (pullback.fst ≫ i₁) (pullback.snd ≫ i₂) (by simp [← eq₁, ← eq₂, pullback.condition_assoc]) +/-- The canonical map `X ×ₛ Y ⟶ X ×ₜ Y` given `S ⟶ T`. -/ +abbreviation pullback.map_desc {X Y S T : C} (f : X ⟶ S) (g : Y ⟶ S) (i : S ⟶ T) + [has_pullback f g] [has_pullback (f ≫ i) (g ≫ i)] : + pullback f g ⟶ pullback (f ≫ i) (g ≫ i) := +pullback.map f g (f ≫ i) (g ≫ i) (𝟙 _) (𝟙 _) i (category.id_comp _).symm (category.id_comp _).symm + /-- Given such a diagram, then there is a natural morphism `W ⨿ₛ X ⟶ Y ⨿ₜ Z`. @@ -963,6 +969,11 @@ abbreviation pushout.map {W X Y Z S T : C} (f₁ : S ⟶ W) (f₂ : S ⟶ X) [ha pushout.desc (i₁ ≫ pushout.inl) (i₂ ≫ pushout.inr) (by { simp only [← category.assoc, eq₁, eq₂], simp [pushout.condition] }) +/-- The canonical map `X ⨿ₛ Y ⟶ X ⨿ₜ Y` given `S ⟶ T`. -/ +abbreviation pushout.map_lift {X Y S T : C} (f : T ⟶ X) (g : T ⟶ Y) (i : S ⟶ T) + [has_pushout f g] [has_pushout (i ≫ f) (i ≫ g)] : + pushout (i ≫ f) (i ≫ g) ⟶ pushout f g := +pushout.map (i ≫ f) (i ≫ g) f g (𝟙 _) (𝟙 _) i (category.comp_id _) (category.comp_id _) /-- Two morphisms into a pullback are equal if their compositions with the pullback morphisms are equal -/ @@ -1077,6 +1088,13 @@ begin tidy end +lemma pullback.map_desc_comp {X Y S T S' : C} (f : X ⟶ T) (g : Y ⟶ T) (i : T ⟶ S) + (i' : S ⟶ S') [has_pullback f g] [has_pullback (f ≫ i) (g ≫ i)] + [has_pullback (f ≫ i ≫ i') (g ≫ i ≫ i')] [has_pullback ((f ≫ i) ≫ i') ((g ≫ i) ≫ i')] : + pullback.map_desc f g (i ≫ i') = pullback.map_desc f g i ≫ pullback.map_desc _ _ i' ≫ + (pullback.congr_hom (category.assoc _ _ _) (category.assoc _ _ _)).hom := +by { ext; simp } + /-- If `f₁ = f₂` and `g₁ = g₂`, we may construct a canonical isomorphism `pushout f₁ g₁ ≅ pullback f₂ g₂` -/ @[simps hom] @@ -1102,6 +1120,14 @@ begin rw category.id_comp } end +lemma pushout.map_lift_comp {X Y S T S' : C} (f : T ⟶ X) (g : T ⟶ Y) (i : S ⟶ T) + (i' : S' ⟶ S) [has_pushout f g] [has_pushout (i ≫ f) (i ≫ g)] + [has_pushout (i' ≫ i ≫ f) (i' ≫ i ≫ g)] [has_pushout ((i' ≫ i) ≫ f) ((i' ≫ i) ≫ g)] : + pushout.map_lift f g (i' ≫ i) = + (pushout.congr_hom (category.assoc _ _ _) (category.assoc _ _ _)).hom ≫ + pushout.map_lift _ _ i' ≫ pushout.map_lift f g i := +by { ext; simp } + section variables (G : C ⥤ D) diff --git a/src/category_theory/morphism_property.lean b/src/category_theory/morphism_property.lean index 0a1f263e7c176..deaa7b2209268 100644 --- a/src/category_theory/morphism_property.lean +++ b/src/category_theory/morphism_property.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import category_theory.limits.shapes.pullbacks +import category_theory.limits.shapes.diagonal import category_theory.arrow import category_theory.limits.shapes.comm_sq @@ -373,6 +373,53 @@ full_subcategory (λ (F : C ⥤ D), W.is_inverted_by F) def functors_inverting.mk {W : morphism_property C} {D : Type*} [category D] (F : C ⥤ D) (hF : W.is_inverted_by F) : W.functors_inverting D := ⟨F, hF⟩ +section diagonal + +variables [has_pullbacks C] {P : morphism_property C} + +/-- For `P : morphism_property C`, `P.diagonal` is a morphism property that holds for `f : X ⟶ Y` +whenever `P` holds for `X ⟶ Y xₓ Y`. -/ +def diagonal (P : morphism_property C) : morphism_property C := +λ X Y f, P (pullback.diagonal f) + +lemma diagonal_iff {X Y : C} {f : X ⟶ Y} : P.diagonal f ↔ P (pullback.diagonal f) := iff.rfl + +lemma respects_iso.diagonal (hP : P.respects_iso) : P.diagonal.respects_iso := +begin + split, + { introv H, + rwa [diagonal_iff, pullback.diagonal_comp, hP.cancel_left_is_iso, hP.cancel_left_is_iso, + ← hP.cancel_right_is_iso _ _, ← pullback.condition, hP.cancel_left_is_iso], + apply_instance }, + { introv H, + delta diagonal, + rwa [pullback.diagonal_comp, hP.cancel_right_is_iso] } +end + +lemma stable_under_composition.diagonal + (hP : stable_under_composition P) (hP' : respects_iso P) (hP'' : stable_under_base_change P) : + P.diagonal.stable_under_composition := +begin + introv X h₁ h₂, + rw [diagonal_iff, pullback.diagonal_comp], + apply hP, { assumption }, + rw hP'.cancel_left_is_iso, + apply hP''.snd, + assumption +end + +lemma stable_under_base_change.diagonal + (hP : stable_under_base_change P) (hP' : respects_iso P) : + P.diagonal.stable_under_base_change := +stable_under_base_change.mk hP'.diagonal +begin + introv h, + rw [diagonal_iff, diagonal_pullback_fst, hP'.cancel_left_is_iso, hP'.cancel_right_is_iso], + convert hP.base_change_map f _ _; simp; assumption +end + +end diagonal + end morphism_property end category_theory From a17aefd343f32f8b9877303cb6a52f8c29bc2e93 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Wed, 5 Oct 2022 12:41:20 +0000 Subject: [PATCH 581/828] feat(measure_theory/covering/density_theorem): add a version of Lebesgue's density theorem (#16762) --- src/algebra/group_power/basic.lean | 4 + src/algebra/group_with_zero/basic.lean | 6 + src/data/set/intervals/image_preimage.lean | 6 + src/measure_theory/covering/besicovitch.lean | 3 +- .../covering/density_theorem.lean | 157 ++++++++++++++++++ .../covering/vitali_family.lean | 19 +++ src/order/filter/basic.lean | 4 + src/topology/algebra/order/basic.lean | 50 ++++++ src/topology/continuous_on.lean | 14 ++ 9 files changed, 262 insertions(+), 1 deletion(-) create mode 100644 src/measure_theory/covering/density_theorem.lean diff --git a/src/algebra/group_power/basic.lean b/src/algebra/group_power/basic.lean index 6448fec7ca7dc..93e14beb11826 100644 --- a/src/algebra/group_power/basic.lean +++ b/src/algebra/group_power/basic.lean @@ -68,6 +68,10 @@ by rw [pow_succ, pow_one] alias pow_two ← sq +theorem pow_three' (a : M) : a^3 = a * a * a := by rw [pow_succ', pow_two] + +theorem pow_three (a : M) : a^3 = a * (a * a) := by rw [pow_succ, pow_two] + @[to_additive] theorem pow_mul_comm' (a : M) (n : ℕ) : a^n * a = a * a^n := commute.pow_self a n diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index 228f125178567..e642256f50ef4 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -843,6 +843,12 @@ classical.by_cases (assume ha, ha) (assume ha, ((one_div_ne_zero ha) h).elim) +lemma mul_left_surjective₀ {a : G₀} (h : a ≠ 0) : surjective (λ g, a * g) := +λ g, ⟨a⁻¹ * g, by simp [← mul_assoc, mul_inv_cancel h]⟩ + +lemma mul_right_surjective₀ {a : G₀} (h : a ≠ 0) : surjective (λ g, g * a) := +λ g, ⟨g * a⁻¹, by simp [mul_assoc, inv_mul_cancel h]⟩ + end group_with_zero section comm_group_with_zero -- comm diff --git a/src/data/set/intervals/image_preimage.lean b/src/data/set/intervals/image_preimage.lean index 414413825f40a..1439d43c32867 100644 --- a/src/data/set/intervals/image_preimage.lean +++ b/src/data/set/intervals/image_preimage.lean @@ -531,6 +531,12 @@ end lemma inv_Ioi {a : k} (ha : 0 < a) : (Ioi a)⁻¹ = Ioo 0 a⁻¹ := by rw [inv_eq_iff_inv_eq, inv_Ioo_0_left (inv_pos.2 ha), inv_inv] +lemma image_const_mul_Ioi_zero {k : Type*} [linear_ordered_field k] + {x : k} (hx : 0 < x) : + (λ y, x * y) '' Ioi (0 : k) = Ioi 0 := +by erw [(units.mk0 x hx.ne').mul_left.image_eq_preimage, preimage_const_mul_Ioi 0 (inv_pos.mpr hx), + zero_div] + /-! ### Images under `x ↦ a * x + b` -/ diff --git a/src/measure_theory/covering/besicovitch.lean b/src/measure_theory/covering/besicovitch.lean index 34435f1851700..f3ce826a3ee72 100644 --- a/src/measure_theory/covering/besicovitch.lean +++ b/src/measure_theory/covering/besicovitch.lean @@ -1139,7 +1139,8 @@ end to `1` when `r` tends to `0`, for almost every `x` in `s`. This shows that almost every point of `s` is a Lebesgue density point for `s`. A stronger version holds for measurable sets, see `ae_tendsto_measure_inter_div_of_measurable_set`. --/ + +See also `is_doubling_measure.ae_tendsto_measure_inter_div`. -/ lemma ae_tendsto_measure_inter_div (μ : measure β) [is_locally_finite_measure μ] (s : set β) : ∀ᵐ x ∂(μ.restrict s), tendsto (λ r, μ (s ∩ (closed_ball x r)) / μ (closed_ball x r)) (𝓝[>] 0) (𝓝 1) := diff --git a/src/measure_theory/covering/density_theorem.lean b/src/measure_theory/covering/density_theorem.lean new file mode 100644 index 0000000000000..7c88d649c5f7b --- /dev/null +++ b/src/measure_theory/covering/density_theorem.lean @@ -0,0 +1,157 @@ +/- +Copyright (c) 2022 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import measure_theory.covering.vitali +import measure_theory.covering.differentiation +import analysis.special_functions.log.base + +/-! +# Doubling measures and Lebesgue's density theorem + +A doubling measure `μ` on a metric space is a measure for which there exists a constant `C` such +that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius +`2 * ε` is bounded by `C` times the measure of the concentric ball of radius `ε`. + +Lebesgue's density theorem states that given a set `S` in a proper metric space with locally-finite +doubling measure `μ` then for almost all points `x` in `S`, for any sequence of closed balls +`B₀, B₁, B₂, ...` containing `x`, the limit `μ (S ∩ Bⱼ) / μ (Bⱼ) → 1` as `j → ∞`. + +In this file we combine general results about existence of Vitali families for doubling measures +with results about differentiation along a Vitali family to obtain an explicit form of Lebesgue's +density theorem. + +## Main results + + * `is_doubling_measure`: the definition of a doubling measure (as a typeclass). + * `is_doubling_measure.doubling_constant`: a function yielding the doubling constant `C` appearing + in the definition of a doubling measure. + * `is_doubling_measure.ae_tendsto_measure_inter_div`: a version of Lebesgue's density theorem for + sequences of balls converging on a point but whose centres are not required to be fixed. + +-/ + +noncomputable theory + +open set filter metric measure_theory +open_locale nnreal topological_space + +/-- A measure `μ` is said to be a doubling measure if there exists a constant `C` such that for +all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius `2 * ε` is +bounded by `C` times the measure of the concentric ball of radius `ε`. + +Note: it is important that this definition makes a demand only for sufficiently small `ε`. For +example we want hyperbolic space to carry the instance `is_doubling_measure volume` but volumes grow +exponentially in hyperbolic space. To be really explicit, consider the hyperbolic plane of +curvature -1, the area of a disc of radius `ε` is `A(ε) = 2π(cosh(ε) - 1)` so `A(2ε)/A(ε) ~ exp(ε)`. +-/ +class is_doubling_measure {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) := +(exists_measure_closed_ball_le_mul [] : + ∃ (C : ℝ≥0), ∀ᶠ ε in 𝓝[>] 0, ∀ x, μ (closed_ball x (2 * ε)) ≤ C * μ (closed_ball x ε)) + +namespace is_doubling_measure + +variables {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) [is_doubling_measure μ] + +/-- A doubling constant for a doubling measure. + +See also `is_doubling_measure.scaling_constant_of`. -/ +def doubling_constant : ℝ≥0 := classical.some $ exists_measure_closed_ball_le_mul μ + +lemma exists_measure_closed_ball_le_mul' : + ∀ᶠ ε in 𝓝[>] 0, ∀ x, μ (closed_ball x (2 * ε)) ≤ doubling_constant μ * μ (closed_ball x ε) := +classical.some_spec $ exists_measure_closed_ball_le_mul μ + +lemma exists_eventually_forall_measure_closed_ball_le_mul (K : ℝ) : + ∃ (C : ℝ≥0), ∀ᶠ ε in 𝓝[>] 0, ∀ x t (ht : t ≤ K), + μ (closed_ball x (t * ε)) ≤ C * μ (closed_ball x ε) := +begin + let C := doubling_constant μ, + have hμ : ∀ (n : ℕ), ∀ᶠ ε in 𝓝[>] 0, ∀ x, + μ (closed_ball x (2^n * ε)) ≤ ↑(C^n) * μ (closed_ball x ε), + { intros n, + induction n with n ih, { simp, }, + replace ih := eventually_nhds_within_pos_mul_left (two_pos : 0 < (2 : ℝ)) ih, + refine (ih.and (exists_measure_closed_ball_le_mul' μ)).mono (λ ε hε x, _), + calc μ (closed_ball x (2^(n + 1) * ε)) + = μ (closed_ball x (2^n * (2 * ε))) : by rw [pow_succ', mul_assoc] + ... ≤ ↑(C^n) * μ (closed_ball x (2 * ε)) : hε.1 x + ... ≤ ↑(C^n) * (C * μ (closed_ball x ε)) : ennreal.mul_left_mono (hε.2 x) + ... = ↑(C^(n + 1)) * μ (closed_ball x ε) : by rw [← mul_assoc, pow_succ', ennreal.coe_mul], }, + rcases lt_or_le K 1 with hK | hK, + { refine ⟨1, _⟩, + simp only [ennreal.coe_one, one_mul], + exact eventually_mem_nhds_within.mono (λ ε hε x t ht, + measure_mono $ closed_ball_subset_closed_ball (by nlinarith [mem_Ioi.mp hε])), }, + { refine ⟨C^⌈real.logb 2 K⌉₊, ((hμ ⌈real.logb 2 K⌉₊).and eventually_mem_nhds_within).mono + (λ ε hε x t ht, le_trans (measure_mono $ closed_ball_subset_closed_ball _) (hε.1 x))⟩, + refine mul_le_mul_of_nonneg_right (ht.trans _) (mem_Ioi.mp hε.2).le, + conv_lhs { rw ← real.rpow_logb two_pos (by norm_num) (by linarith : 0 < K), }, + rw ← real.rpow_nat_cast, + exact real.rpow_le_rpow_of_exponent_le one_le_two (nat.le_ceil (real.logb 2 K)), }, +end + +/-- A variant of `is_doubling_measure.doubling_constant` which allows for scaling the radius by +values other than `2`. -/ +def scaling_constant_of (K : ℝ) : ℝ≥0 := +classical.some $ exists_eventually_forall_measure_closed_ball_le_mul μ K + +lemma eventually_scaling_constant_of (K : ℝ) : + ∀ᶠ ε in 𝓝[>] 0, ∀ x t (ht : t ≤ K), + μ (closed_ball x (t * ε)) ≤ (scaling_constant_of μ K) * μ (closed_ball x ε) := +classical.some_spec $ exists_eventually_forall_measure_closed_ball_le_mul μ K + +variables [proper_space α] [borel_space α] [is_locally_finite_measure μ] + +/-- A Vitali family in space with doubling measure with a covering proportion controlled by `K`. -/ +def vitali_family (K : ℝ) (hK : 6 ≤ K) : vitali_family μ := +vitali.vitali_family μ (scaling_constant_of μ K) $ λ x ε hε, +begin + have h := eventually_scaling_constant_of μ K, + replace h := forall_eventually_of_eventually_forall (forall_eventually_of_eventually_forall h x), + replace h := eventually_imp_distrib_left.mp (h 6) hK, + simpa only [exists_prop] using ((eventually_nhds_within_pos_mem_Ioc hε).and h).exists, +end + +/-- A version of *Lebesgue's density theorem* for a sequence of closed balls whose centres are +not required to be fixed. + +See also `besicovitch.ae_tendsto_measure_inter_div`. -/ +lemma ae_tendsto_measure_inter_div (S : set α) (K : ℝ) (hK : K ∈ unit_interval) : + ∀ᵐ x ∂μ.restrict S, ∀ {ι : Type*} {l : filter ι} (w : ι → α) (δ : ι → ℝ) + (δlim : tendsto δ l (𝓝[>] 0)) + (xmem : ∀ᶠ j in l, x ∈ closed_ball (w j) (K * δ j)), + tendsto (λ j, μ (S ∩ closed_ball (w j) (δ j)) / μ (closed_ball (w j) (δ j))) l (𝓝 1) := +begin + let v := is_doubling_measure.vitali_family μ 7 (by norm_num), + filter_upwards [v.ae_tendsto_measure_inter_div S] with x hx ι l w δ δlim xmem, + suffices : tendsto (λ j, closed_ball (w j) (δ j)) l (v.filter_at x), { exact hx.comp this, }, + refine v.tendsto_filter_at_iff.mpr ⟨_, λ ε hε, _⟩, + { simp only [v, vitali.vitali_family], + have δpos : ∀ᶠ j in l, 0 < δ j := eventually_mem_of_tendsto_nhds_within δlim, + replace xmem : ∀ᶠ (j : ι) in l, x ∈ closed_ball (w j) (δ j) := (δpos.and xmem).mono + (λ j hj, closed_ball_subset_closed_ball (by nlinarith [hj.1, hK.2]) hj.2), + apply ((δlim.eventually (eventually_scaling_constant_of μ 7)).and (xmem.and δpos)).mono, + rintros j ⟨hjC, hjx, hjδ⟩, + have hdiam : 3 * diam (closed_ball (w j) (δ j)) ≤ 6 * δ j, + { linarith [@diam_closed_ball _ _ (w j) _ hjδ.le], }, + refine ⟨hjx, is_closed_ball, (nonempty_ball.mpr hjδ).mono ball_subset_interior_closed_ball, + (measure_mono (closed_ball_subset_closed_ball hdiam)).trans _⟩, + suffices : closed_ball x (6 * δ j) ⊆ closed_ball (w j) (7 * δ j), + { exact (measure_mono this).trans ((hjC (w j) 7 (by norm_num)).trans $ le_refl _), }, + intros y hy, + simp only [mem_closed_ball, dist_comm x (w j)] at hjx hy ⊢, + linarith [dist_triangle_right y (w j) x], }, + { have δpos := eventually_mem_of_tendsto_nhds_within δlim, + replace δlim := tendsto_nhds_of_tendsto_nhds_within δlim, + replace hK : 0 < K + 1 := by linarith [hK.1], + apply (((metric.tendsto_nhds.mp δlim _ (div_pos hε hK)).and δpos).and xmem).mono, + rintros j ⟨⟨hjε, hj₀ : 0 < δ j⟩, hx⟩ y hy, + replace hjε : (K + 1) * δ j < ε := + by simpa [abs_eq_self.mpr hj₀.le] using (lt_div_iff' hK).mp hjε, + simp only [mem_closed_ball] at hx hy ⊢, + linarith [dist_triangle_right y x (w j)], }, +end + +end is_doubling_measure diff --git a/src/measure_theory/covering/vitali_family.lean b/src/measure_theory/covering/vitali_family.lean index b8d1c2c9a38be..852e5353ae45f 100644 --- a/src/measure_theory/covering/vitali_family.lean +++ b/src/measure_theory/covering/vitali_family.lean @@ -203,6 +203,25 @@ begin exact ⟨1, zero_lt_one⟩ end +lemma eventually_filter_at_subset_closed_ball (x : α) {ε : ℝ} (hε : 0 < ε) : + ∀ᶠ (a : set α) in v.filter_at x, a ⊆ closed_ball x ε := +begin + simp only [v.eventually_filter_at_iff], + exact ⟨ε, hε, λ a ha ha', ha'⟩, +end + +lemma tendsto_filter_at_iff {ι : Type*} {l : filter ι} {f : ι → set α} {x : α} : + tendsto f l (v.filter_at x) ↔ + (∀ᶠ i in l, f i ∈ v.sets_at x) ∧ (∀ (ε > (0 : ℝ)), ∀ᶠ i in l, f i ⊆ closed_ball x ε) := +begin + refine ⟨λ H, + ⟨H.eventually $ v.eventually_filter_at_mem_sets x, + λ ε hε, H.eventually $ v.eventually_filter_at_subset_closed_ball x hε⟩, + λ H s hs, (_ : ∀ᶠ i in l, f i ∈ s)⟩, + obtain ⟨ε, εpos, hε⟩ := v.mem_filter_at_iff.mp hs, + filter_upwards [H.1, H.2 ε εpos] with i hi hiε using hε _ hi hiε, +end + lemma eventually_filter_at_measurable_set (x : α) : ∀ᶠ a in v.filter_at x, measurable_set a := by { filter_upwards [v.eventually_filter_at_mem_sets x] with _ ha using v.measurable_set' _ _ ha } diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 70763cb535b72..fcf4c0a0c7d52 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -969,6 +969,10 @@ lemma eventually_of_forall {p : α → Prop} {f : filter α} (hp : ∀ x, p x) : ∀ᶠ x in f, p x := univ_mem' hp +lemma forall_eventually_of_eventually_forall {f : filter α} {p : α → β → Prop} + (h : ∀ᶠ x in f, ∀ y, p x y) : ∀ y, ∀ᶠ x in f, p x y := +by { intros y, filter_upwards [h], tauto, } + @[simp] lemma eventually_false_iff_eq_bot {f : filter α} : (∀ᶠ x in f, false) ↔ f = ⊥ := empty_mem_iff_bot diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index b845c9e7603c8..4ca8225e2a0d2 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -2952,4 +2952,54 @@ lemma monotone.tendsto_nhds_within_Ioi {α β : Type*} end conditionally_complete_linear_order +section nhds_with_pos + +section linear_ordered_add_comm_group + +variables [linear_ordered_add_comm_group α] [topological_space α] [order_topology α] + +lemma eventually_nhds_within_pos_mem_Ioo {ε : α} (h : 0 < ε) : + ∀ᶠ x in 𝓝[>] 0, x ∈ Ioo 0 ε := +begin + rw [eventually_iff, mem_nhds_within], + exact ⟨Ioo (-ε) ε, is_open_Ioo, by simp [h], λ x hx, ⟨hx.2, hx.1.2⟩⟩, +end + +lemma eventually_nhds_within_pos_mem_Ioc {ε : α} (h : 0 < ε) : + ∀ᶠ x in 𝓝[>] 0, x ∈ Ioc 0 ε := +(eventually_nhds_within_pos_mem_Ioo h).mono Ioo_subset_Ioc_self + +end linear_ordered_add_comm_group + +section linear_ordered_field + +variables [linear_ordered_field α] [topological_space α] [order_topology α] + +lemma nhds_within_pos_comap_mul_left {x : α} (hx : 0 < x) : + comap (λ ε, x * ε) (𝓝[>] 0) = 𝓝[>] 0 := +begin + suffices : ∀ {x : α} (hx : 0 < x), 𝓝[>] 0 ≤ comap (λ ε, x * ε) (𝓝[>] 0), + { refine le_antisymm _ (this hx), + have hr : 𝓝[>] (0 : α) = ((𝓝[>] (0 : α)).comap (λ ε, x⁻¹ * ε)).comap (λ ε, x * ε), + { simp [comap_comap, inv_mul_cancel hx.ne.symm, comap_id, one_mul_eq_id], }, + conv_rhs { rw hr, }, + rw comap_le_comap_iff (by convert univ_mem; exact (mul_left_surjective₀ hx.ne.symm).range_eq), + exact this (inv_pos.mpr hx), }, + intros x hx, + convert nhds_within_le_comap (continuous_mul_left x).continuous_within_at, + { exact (mul_zero _).symm, }, + { rw image_const_mul_Ioi_zero hx, }, +end + +lemma eventually_nhds_within_pos_mul_left {x : α} (hx : 0 < x) + {p : α → Prop} (h : ∀ᶠ ε in 𝓝[>] 0, p ε) : ∀ᶠ ε in 𝓝[>] 0, p (x * ε) := +begin + convert h.comap (λ ε, x * ε), + exact (nhds_within_pos_comap_mul_left hx).symm, +end + +end linear_ordered_field + +end nhds_with_pos + end order_topology diff --git a/src/topology/continuous_on.lean b/src/topology/continuous_on.lean index e4c16d4191367..7ab9203159cd7 100644 --- a/src/topology/continuous_on.lean +++ b/src/topology/continuous_on.lean @@ -338,6 +338,20 @@ theorem tendsto_nhds_within_of_tendsto_nhds {f : α → β} {a : α} tendsto f (𝓝[s] a) l := h.mono_left inf_le_left +lemma eventually_mem_of_tendsto_nhds_within {f : β → α} {a : α} + {s : set α} {l : filter β} (h : tendsto f l (𝓝[s] a)) : + ∀ᶠ i in l, f i ∈ s := +begin + simp_rw [nhds_within_eq, tendsto_infi, mem_set_of_eq, tendsto_principal, mem_inter_iff, + eventually_and] at h, + exact (h univ ⟨mem_univ a, is_open_univ⟩).2, +end + +lemma tendsto_nhds_of_tendsto_nhds_within {f : β → α} {a : α} + {s : set α} {l : filter β} (h : tendsto f l (𝓝[s] a)) : + tendsto f l (𝓝 a) := +h.mono_right nhds_within_le_nhds + theorem principal_subtype {α : Type*} (s : set α) (t : set {x // x ∈ s}) : 𝓟 t = comap coe (𝓟 ((coe : s → α) '' t)) := by rw [comap_principal, set.preimage_image_eq _ subtype.coe_injective] From 074ba985b0784b47e67f37add1b8252f0e656360 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 5 Oct 2022 16:33:06 +0000 Subject: [PATCH 582/828] feat(geometry/euclidean/oriented_angle): oriented angles between points (#16279) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the oriented angle between three points, notation `∡`, in terms of that between two vectors, and set up some corresponding basic API (including relating it to unoriented angles between three points). For angles between points, the choice of orientation is implicit, via use of the `module.oriented` typeclass to specify a preferred orientation. --- src/geometry/euclidean/oriented_angle.lean | 198 ++++++++++++++++++++- 1 file changed, 196 insertions(+), 2 deletions(-) diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 9d4937d23c5a4..eb7bbee14a857 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -12,7 +12,7 @@ import geometry.euclidean.basic /-! # Oriented angles. -This file defines oriented angles in real inner product spaces. +This file defines oriented angles in real inner product spaces and Euclidean affine spaces. ## Main definitions @@ -20,6 +20,9 @@ This file defines oriented angles in real inner product spaces. * `orientation.rotation` is the rotation by an oriented angle with respect to an orientation. +* `euclidean_geometry.oangle`, with notation `∡`, is the oriented angle determined by three + points. + ## Implementation notes The definitions here use the `real.angle` type, angles modulo `2 * π`. For some purposes, @@ -30,7 +33,10 @@ modulo `2 * π` as equalities of `(2 : ℤ) • θ`. Definitions and results in the `orthonormal` namespace, with respect to a particular choice of orthonormal basis, are mainly for use in setting up the API and proving that certain definitions do not depend on the choice of basis for a given orientation. Applications should -generally use the definitions and results in the `orientation` namespace instead. +generally use the definitions and results in the `orientation` namespace instead, when working +with vectors in a real inner product space, or those in the `euclidean_geometry` namespace, +when working with points in a Euclidean affine space (of which a choice of orientation has +been fixed with `module.oriented`). ## References @@ -40,6 +46,7 @@ generally use the definitions and results in the `orientation` namespace instead noncomputable theory +open_locale euclidean_geometry open_locale real open_locale real_inner_product_space @@ -1778,3 +1785,190 @@ lemma oangle_sign_smul_add_smul_smul_add_smul (x y : V) (r₁ r₂ r₃ r₄ : (ob).oangle_sign_smul_add_smul_smul_add_smul x y r₁ r₂ r₃ r₄ end orientation + +namespace euclidean_geometry + +open finite_dimensional + +variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] +variables [normed_add_torsor V P] [hd2 : fact (finrank ℝ V = 2)] [module.oriented ℝ V (fin 2)] +include hd2 + +local notation `o` := module.oriented.positive_orientation + +/-- The oriented angle at `p₂` between the line segments to `p₁` and `p₃`, modulo `2 * π`. If +either of those points equals `p₂`, this is 0. See `euclidean_geometry.angle` for the +corresponding unoriented angle definition. -/ +def oangle (p₁ p₂ p₃ : P) : real.angle := (o).oangle (p₁ -ᵥ p₂) (p₃ -ᵥ p₂) + +localized "notation (name := oangle) `∡` := euclidean_geometry.oangle" in euclidean_geometry + +/-- Oriented angles are continuous when neither end point equals the middle point. -/ +lemma continuous_at_oangle {x : P × P × P} (hx12 : x.1 ≠ x.2.1) (hx32 : x.2.2 ≠ x.2.1) : + continuous_at (λ y : P × P × P, ∡ y.1 y.2.1 y.2.2) x := +begin + let f : P × P × P → V × V := λ y, (y.1 -ᵥ y.2.1, y.2.2 -ᵥ y.2.1), + have hf1 : (f x).1 ≠ 0, by simp [hx12], + have hf2 : (f x).2 ≠ 0, by simp [hx32], + exact ((o).continuous_at_oangle hf1 hf2).comp + ((continuous_fst.vsub continuous_snd.fst).prod_mk + (continuous_snd.snd.vsub continuous_snd.fst)).continuous_at +end + +/-- The angle ∡AAB at a point. -/ +@[simp] lemma oangle_self_left (p₁ p₂ : P) : ∡ p₁ p₁ p₂ = 0 := +by simp [oangle] + +/-- The angle ∡ABB at a point. -/ +@[simp] lemma oangle_self_right (p₁ p₂ : P) : ∡ p₁ p₂ p₂ = 0 := +by simp [oangle] + +/-- The angle ∡ABA at a point. -/ +@[simp] lemma oangle_self_left_right (p₁ p₂ : P) : ∡ p₁ p₂ p₁ = 0 := +(o).oangle_self _ + +/-- Reversing the order of the points passed to `oangle` negates the angle. -/ +lemma oangle_rev (p₁ p₂ p₃ : P) : ∡ p₃ p₂ p₁ = -∡ p₁ p₂ p₃ := +(o).oangle_rev _ _ + +/-- Adding an angle to that with the order of the points reversed results in 0. -/ +@[simp] lemma oangle_add_oangle_rev (p₁ p₂ p₃ : P) : ∡ p₁ p₂ p₃ + ∡ p₃ p₂ p₁ = 0 := +(o).oangle_add_oangle_rev _ _ + +/-- An oriented angle is zero if and only if the angle with the order of the points reversed is +zero. -/ +lemma oangle_eq_zero_iff_oangle_rev_eq_zero {p₁ p₂ p₃ : P} : ∡ p₁ p₂ p₃ = 0 ↔ ∡ p₃ p₂ p₁ = 0 := +(o).oangle_eq_zero_iff_oangle_rev_eq_zero + +/-- An oriented angle is `π` if and only if the angle with the order of the points reversed is +`π`. -/ +lemma oangle_eq_pi_iff_oangle_rev_eq_pi {p₁ p₂ p₃ : P} : ∡ p₁ p₂ p₃ = π ↔ ∡ p₃ p₂ p₁ = π := +(o).oangle_eq_pi_iff_oangle_rev_eq_pi + +/-- Given three points not equal to `p`, the angle between the first and the second at `p` plus +the angle between the second and the third equals the angle between the first and the third. -/ +@[simp] lemma oangle_add {p p₁ p₂ p₃ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) (hp₃ : p₃ ≠ p) : + ∡ p₁ p p₂ + ∡ p₂ p p₃ = ∡ p₁ p p₃ := +(o).oangle_add (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₂) (vsub_ne_zero.2 hp₃) + +/-- Given three points not equal to `p`, the angle between the second and the third at `p` plus +the angle between the first and the second equals the angle between the first and the third. -/ +@[simp] lemma oangle_add_swap {p p₁ p₂ p₃ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) (hp₃ : p₃ ≠ p) : + ∡ p₂ p p₃ + ∡ p₁ p p₂ = ∡ p₁ p p₃ := +(o).oangle_add_swap (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₂) (vsub_ne_zero.2 hp₃) + +/-- Given three points not equal to `p`, the angle between the first and the third at `p` minus +the angle between the first and the second equals the angle between the second and the third. -/ +@[simp] lemma oangle_sub_left {p p₁ p₂ p₃ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) (hp₃ : p₃ ≠ p) : + ∡ p₁ p p₃ - ∡ p₁ p p₂ = ∡ p₂ p p₃ := +(o).oangle_sub_left (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₂) (vsub_ne_zero.2 hp₃) + +/-- Given three points not equal to `p`, the angle between the first and the third at `p` minus +the angle between the second and the third equals the angle between the first and the second. -/ +@[simp] lemma oangle_sub_right {p p₁ p₂ p₃ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) (hp₃ : p₃ ≠ p) : + ∡ p₁ p p₃ - ∡ p₂ p p₃ = ∡ p₁ p p₂ := +(o).oangle_sub_right (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₂) (vsub_ne_zero.2 hp₃) + +/-- Given three points not equal to `p`, adding the angles between them at `p` in cyclic order +results in 0. -/ +@[simp] lemma oangle_add_cyc3 {p p₁ p₂ p₃ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) (hp₃ : p₃ ≠ p) : + ∡ p₁ p p₂ + ∡ p₂ p p₃ + ∡ p₃ p p₁ = 0 := +(o).oangle_add_cyc3 (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₂) (vsub_ne_zero.2 hp₃) + +/-- Pons asinorum, oriented angle-at-point form. -/ +lemma oangle_eq_oangle_of_dist_eq {p₁ p₂ p₃ : P} (h : dist p₁ p₂ = dist p₁ p₃) : + ∡ p₁ p₂ p₃ = ∡ p₂ p₃ p₁ := +begin + simp_rw dist_eq_norm_vsub at h, + rw [oangle, oangle, ←vsub_sub_vsub_cancel_left p₃ p₂ p₁, ←vsub_sub_vsub_cancel_left p₂ p₃ p₁, + (o).oangle_sub_eq_oangle_sub_rev_of_norm_eq h] +end + +/-- The angle at the apex of an isosceles triangle is `π` minus twice a base angle, oriented +angle-at-point form. -/ +lemma oangle_eq_pi_sub_two_zsmul_oangle_of_dist_eq {p₁ p₂ p₃ : P} (hn : p₂ ≠ p₃) + (h : dist p₁ p₂ = dist p₁ p₃) : ∡ p₃ p₁ p₂ = π - (2 : ℤ) • ∡ p₁ p₂ p₃ := +begin + simp_rw dist_eq_norm_vsub at h, + rw [oangle, oangle], + convert (o).oangle_eq_pi_sub_two_zsmul_oangle_sub_of_norm_eq _ h using 1, + { rw [←neg_vsub_eq_vsub_rev p₁ p₃, ←neg_vsub_eq_vsub_rev p₁ p₂, (o).oangle_neg_neg] }, + { rw [←(o).oangle_sub_eq_oangle_sub_rev_of_norm_eq h], simp }, + { simpa using hn } +end + +/-- The cosine of the oriented angle at `p` between two points not equal to `p` equals that of the +unoriented angle. -/ +lemma cos_oangle_eq_cos_angle {p p₁ p₂ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) : + real.angle.cos (∡ p₁ p p₂) = real.cos (∠ p₁ p p₂) := +(o).cos_oangle_eq_cos_angle (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₂) + +/-- The oriented angle at `p` between two points not equal to `p` is plus or minus the unoriented +angle. -/ +lemma oangle_eq_angle_or_eq_neg_angle {p p₁ p₂ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) : + ∡ p₁ p p₂ = ∠ p₁ p p₂ ∨ ∡ p₁ p p₂ = -∠ p₁ p p₂ := +(o).oangle_eq_angle_or_eq_neg_angle (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₂) + +/-- The unoriented angle at `p` between two points not equal to `p` is the absolute value of the +oriented angle. -/ +lemma angle_eq_abs_oangle_to_real {p p₁ p₂ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) : + ∠ p₁ p p₂ = |(∡ p₁ p p₂).to_real| := +(o).angle_eq_abs_oangle_to_real (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₂) + +/-- If the sign of the oriented angle at `p` between two points is zero, either one of the points +equals `p` or the unoriented angle is 0 or π. -/ +lemma eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero {p p₁ p₂ : P} + (h : (∡ p₁ p p₂).sign = 0) : p₁ = p ∨ p₂ = p ∨ ∠ p₁ p p₂ = 0 ∨ ∠ p₁ p p₂ = π := +begin + convert (o).eq_zero_or_angle_eq_zero_or_pi_of_sign_oangle_eq_zero h; + simp +end + +/-- If two unoriented angles are equal, and the signs of the corresponding oriented angles are +equal, then the oriented angles are equal (even in degenerate cases). -/ +lemma oangle_eq_of_angle_eq_of_sign_eq {p₁ p₂ p₃ p₄ p₅ p₆ : P} (h : ∠ p₁ p₂ p₃ = ∠ p₄ p₅ p₆) + (hs : (∡ p₁ p₂ p₃).sign = (∡ p₄ p₅ p₆).sign) : ∡ p₁ p₂ p₃ = ∡ p₄ p₅ p₆ := +(o).oangle_eq_of_angle_eq_of_sign_eq h hs + +/-- If the signs of two nondegenerate oriented angles between points are equal, the oriented +angles are equal if and only if the unoriented angles are equal. -/ +lemma oangle_eq_iff_angle_eq_of_sign_eq {p₁ p₂ p₃ p₄ p₅ p₆ : P} (hp₁ : p₁ ≠ p₂) (hp₃ : p₃ ≠ p₂) + (hp₄ : p₄ ≠ p₅) (hp₆ : p₆ ≠ p₅) (hs : (∡ p₁ p₂ p₃).sign = (∡ p₄ p₅ p₆).sign) : + ∠ p₁ p₂ p₃ = ∠ p₄ p₅ p₆ ↔ ∡ p₁ p₂ p₃ = ∡ p₄ p₅ p₆ := +(o).oangle_eq_iff_angle_eq_of_sign_eq (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₃) + (vsub_ne_zero.2 hp₄) (vsub_ne_zero.2 hp₆) hs + +/-- The unoriented angle at `p` between two points not equal to `p` is zero if and only if the +unoriented angle is zero. -/ +lemma oangle_eq_zero_iff_angle_eq_zero {p p₁ p₂ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) : + ∡ p₁ p p₂ = 0 ↔ ∠ p₁ p p₂ = 0 := +(o).oangle_eq_zero_iff_angle_eq_zero (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₂) + +/-- The oriented angle between three points is `π` if and only if the unoriented angle is `π`. -/ +lemma oangle_eq_pi_iff_angle_eq_pi {p₁ p₂ p₃ : P} : ∡ p₁ p₂ p₃ = π ↔ ∠ p₁ p₂ p₃ = π := +(o).oangle_eq_pi_iff_angle_eq_pi + +/-- Swapping the first and second points in an oriented angle negates the sign of that angle. -/ +lemma oangle_swap₁₂_sign (p₁ p₂ p₃ : P) : -(∡ p₁ p₂ p₃).sign = (∡ p₂ p₁ p₃).sign := +begin + rw [eq_comm, oangle, oangle, ←(o).oangle_neg_neg, neg_vsub_eq_vsub_rev, neg_vsub_eq_vsub_rev, + ←vsub_sub_vsub_cancel_left p₁ p₃ p₂, ←neg_vsub_eq_vsub_rev p₃ p₂, sub_eq_add_neg, + neg_vsub_eq_vsub_rev p₂ p₁, add_comm, ←@neg_one_smul ℝ], + nth_rewrite 1 [←one_smul ℝ (p₁ -ᵥ p₂)], + rw (o).oangle_sign_smul_add_smul_right, + simp +end + +/-- Swapping the first and third points in an oriented angle negates the sign of that angle. -/ +lemma oangle_swap₁₃_sign (p₁ p₂ p₃ : P) : -(∡ p₁ p₂ p₃).sign = (∡ p₃ p₂ p₁).sign := +by rw [oangle_rev, real.angle.sign_neg, neg_neg] + +/-- Swapping the second and third points in an oriented angle negates the sign of that angle. -/ +lemma oangle_swap₂₃_sign (p₁ p₂ p₃ : P) : -(∡ p₁ p₂ p₃).sign = (∡ p₁ p₃ p₂).sign := +by rw [oangle_swap₁₃_sign, ←oangle_swap₁₂_sign, oangle_swap₁₃_sign] + +/-- Rotating the points in an oriented angle does not change the sign of that angle. -/ +lemma oangle_rotate_sign (p₁ p₂ p₃ : P) : (∡ p₂ p₃ p₁).sign = (∡ p₁ p₂ p₃).sign := +by rw [←oangle_swap₁₂_sign, oangle_swap₁₃_sign] + +end euclidean_geometry From 735ce9d9456b4d892242bdfbcdb4b6944555ed11 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 5 Oct 2022 18:04:11 +0000 Subject: [PATCH 583/828] refactor(algebraic_geometry/projective_spectrum/scheme): use homogenous localization API (#16693) --- .../projective_spectrum/scheme.lean | 335 ++++++++---------- 1 file changed, 145 insertions(+), 190 deletions(-) diff --git a/src/algebraic_geometry/projective_spectrum/scheme.lean b/src/algebraic_geometry/projective_spectrum/scheme.lean index 19c57bcbd7859..1ae0d6c22164a 100644 --- a/src/algebraic_geometry/projective_spectrum/scheme.lean +++ b/src/algebraic_geometry/projective_spectrum/scheme.lean @@ -103,82 +103,7 @@ local notation `Spec ` ring := Spec.LocallyRingedSpace_obj (CommRing.of ring) local notation `Spec.T ` ring := (Spec.LocallyRingedSpace_obj (CommRing.of ring)).to_SheafedSpace.to_PresheafedSpace.1 -- the underlying topological space of `Spec` - -section -variable {𝒜} -/-- -The degree zero part of the localized ring `Aₓ` is the subring of elements of the form `a/x^n` such -that `a` and `x^n` have the same degree. --/ -def degree_zero_part {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) : subring (away f) := -{ carrier := { y | ∃ (n : ℕ) (a : 𝒜 (m * n)), y = mk a ⟨f^n, ⟨n, rfl⟩⟩ }, - mul_mem' := λ _ _ ⟨n, ⟨a, h⟩⟩ ⟨n', ⟨b, h'⟩⟩, h.symm ▸ h'.symm ▸ - ⟨n+n', ⟨⟨a.1 * b.1, (mul_add m n n').symm ▸ mul_mem a.2 b.2⟩, - by {rw mk_mul, congr' 1, simp only [pow_add], refl }⟩⟩, - one_mem' := ⟨0, ⟨1, (mul_zero m).symm ▸ one_mem⟩, - by { symmetry, rw subtype.coe_mk, convert ← mk_self 1, simp only [pow_zero], refl, }⟩, - add_mem' := λ _ _ ⟨n, ⟨a, h⟩⟩ ⟨n', ⟨b, h'⟩⟩, h.symm ▸ h'.symm ▸ - ⟨n+n', ⟨⟨f ^ n * b.1 + f ^ n' * a.1, (mul_add m n n').symm ▸ - add_mem (mul_mem (by { rw mul_comm, exact set_like.pow_mem_graded n f_deg }) b.2) - begin - rw add_comm, - refine mul_mem _ a.2, - rw mul_comm, - exact set_like.pow_mem_graded _ f_deg - end⟩, begin - rw add_mk, - congr' 1, - simp only [pow_add], - refl, - end⟩⟩, - zero_mem' := ⟨0, ⟨0, (mk_zero _).symm⟩⟩, - neg_mem' := λ x ⟨n, ⟨a, h⟩⟩, h.symm ▸ ⟨n, ⟨-a, neg_mk _ _⟩⟩ } - -end - -local notation `A⁰_ ` f_deg := degree_zero_part f_deg - -section - -variable {𝒜} - -instance (f : A) {m : ℕ} (f_deg : f ∈ 𝒜 m) : comm_ring (A⁰_ f_deg) := -(degree_zero_part f_deg).to_comm_ring - -/-- -Every element in the degree zero part of `Aₓ` can be written as `a/x^n` for some `a` and `n : ℕ`, -`degree_zero_part.deg` picks this natural number `n` --/ -def degree_zero_part.deg {f : A} {m : ℕ} {f_deg : f ∈ 𝒜 m} (x : A⁰_ f_deg) : ℕ := -x.2.some - -/-- -Every element in the degree zero part of `Aₓ` can be written as `a/x^n` for some `a` and `n : ℕ`, -`degree_zero_part.deg` picks the numerator `a` --/ -def degree_zero_part.num {f : A} {m : ℕ} {f_deg : f ∈ 𝒜 m} (x : A⁰_ f_deg) : A := -x.2.some_spec.some.1 - -lemma degree_zero_part.num_mem {f : A} {m : ℕ} {f_deg : f ∈ 𝒜 m} (x : A⁰_ f_deg) : - degree_zero_part.num x ∈ 𝒜 (m * degree_zero_part.deg x) := -x.2.some_spec.some.2 - -lemma degree_zero_part.eq {f : A} {m : ℕ} {f_deg : f ∈ 𝒜 m} (x : A⁰_ f_deg) : - (x : away f) = mk (degree_zero_part.num x) ⟨f^(degree_zero_part.deg x), ⟨_, rfl⟩⟩ := -x.2.some_spec.some_spec - -lemma degree_zero_part.coe_mul {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) (x y : A⁰_ f_deg) : - (↑(x * y) : away f) = x * y := rfl - -lemma degree_zero_part.coe_one {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) : - (↑(1 : A⁰_ f_deg) : away f) = 1 := rfl - -lemma degree_zero_part.coe_sum {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) {ι : Type*} - (s : finset ι) (g : ι → A⁰_ f_deg) : - (↑(∑ i in s, g i) : away f) = ∑ i in s, (g i : away f) := -by { classical, induction s using finset.induction_on with i s hi ih; simp } - -end +local notation `A⁰_ ` f := homogeneous_localization.away 𝒜 f namespace Proj_iso_Spec_Top_component @@ -200,25 +125,24 @@ variables {𝒜} {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) (x : Proj| (pbo f)) /--For any `x` in `Proj| (pbo f)`, the corresponding ideal in `Spec A⁰_f`. This fact that this ideal is prime is proven in `Top_component.forward.to_fun`-/ -def carrier : ideal (A⁰_ f_deg) := -ideal.comap (algebra_map (A⁰_ f_deg) (away f)) - (ideal.span $ algebra_map A (away f) '' x.1.as_homogeneous_ideal) +def carrier : ideal (A⁰_ f) := +ideal.comap (algebra_map (A⁰_ f) (away f)) + (ideal.span $ algebra_map A (away f) '' x.val.as_homogeneous_ideal) -lemma mem_carrier_iff (z : A⁰_ f_deg) : - z ∈ carrier f_deg x ↔ - ↑z ∈ ideal.span (algebra_map A (away f) '' x.1.as_homogeneous_ideal) := +lemma mem_carrier_iff (z : A⁰_ f) : + z ∈ carrier 𝒜 x ↔ + z.val ∈ ideal.span (algebra_map A (away f) '' x.1.as_homogeneous_ideal) := iff.rfl -lemma mem_carrier.clear_denominator [decidable_eq (away f)] - {z : A⁰_ f_deg} (hz : z ∈ carrier f_deg x) : +lemma mem_carrier.clear_denominator' [decidable_eq (away f)] + {z : localization.away f} + (hz : z ∈ span ((algebra_map A (away f)) '' x.val.as_homogeneous_ideal)) : ∃ (c : algebra_map A (away f) '' x.1.as_homogeneous_ideal →₀ away f) - (N : ℕ) - (acd : Π y ∈ c.support.image c, A), - f ^ N • ↑z = - algebra_map A (away f) (∑ i in c.support.attach, - acd (c i) (finset.mem_image.mpr ⟨i, ⟨i.2, rfl⟩⟩) * classical.some i.1.2) := + (N : ℕ) (acd : Π y ∈ c.support.image c, A), + f ^ N • z = algebra_map A (away f) + (∑ i in c.support.attach, acd (c i) (finset.mem_image.mpr ⟨i, ⟨i.2, rfl⟩⟩) * i.1.2.some) := begin - rw [mem_carrier_iff, ←submodule_span_eq, finsupp.span_eq_range_total, linear_map.mem_range] at hz, + rw [←submodule_span_eq, finsupp.span_eq_range_total, linear_map.mem_range] at hz, rcases hz with ⟨c, eq1⟩, rw [finsupp.total_apply, finsupp.sum] at eq1, obtain ⟨⟨_, N, rfl⟩, hN⟩ := is_localization.exist_integer_multiples_of_finset (submonoid.powers f) @@ -233,6 +157,14 @@ begin refl end +lemma mem_carrier.clear_denominator [decidable_eq (away f)] + {z : A⁰_ f} (hz : z ∈ carrier 𝒜 x) : + ∃ (c : algebra_map A (away f) '' x.1.as_homogeneous_ideal →₀ away f) + (N : ℕ) (acd : Π y ∈ c.support.image c, A), + f ^ N • z.val = algebra_map A (away f) + (∑ i in c.support.attach, acd (c i) (finset.mem_image.mpr ⟨i, ⟨i.2, rfl⟩⟩) * i.1.2.some) := +mem_carrier.clear_denominator' x $ (mem_carrier_iff 𝒜 x z).mpr hz + lemma disjoint : (disjoint (x.1.as_homogeneous_ideal.to_ideal : set A) (submonoid.powers f : set A)) := begin @@ -249,13 +181,13 @@ begin end lemma carrier_ne_top : - carrier f_deg x ≠ ⊤ := + carrier 𝒜 x ≠ ⊤ := begin have eq_top := disjoint x, classical, contrapose! eq_top, obtain ⟨c, N, acd, eq1⟩ := mem_carrier.clear_denominator _ x ((ideal.eq_top_iff_one _).mp eq_top), - rw [algebra.smul_def, subring.coe_one, mul_one] at eq1, + rw [algebra.smul_def, homogeneous_localization.one_val, mul_one] at eq1, change localization.mk (f ^ N) 1 = mk (∑ _, _) 1 at eq1, simp only [mk_eq_mk', is_localization.eq] at eq1, rcases eq1 with ⟨⟨_, ⟨M, rfl⟩⟩, eq1⟩, @@ -268,18 +200,22 @@ begin exact (classical.some_spec h).1, end +variable (f) /--The function between the basic open set `D(f)` in `Proj` to the corresponding basic open set in `Spec A⁰_f`. This is bundled into a continuous map in `Top_component.forward`. -/ -def to_fun (x : Proj.T| (pbo f)) : (Spec.T (A⁰_ f_deg)) := -⟨carrier f_deg x, carrier_ne_top f_deg x, λ x1 x2 hx12, begin - classical, - rcases ⟨x1, x2⟩ with ⟨⟨x1, hx1⟩, ⟨x2, hx2⟩⟩, +def to_fun (x : Proj.T| (pbo f)) : (Spec.T (A⁰_ f)) := +⟨carrier 𝒜 x, carrier_ne_top x, λ x1 x2 hx12, begin + classical, simp only [mem_carrier_iff] at hx12 ⊢, + let J := span (⇑(algebra_map A (away f)) '' x.val.as_homogeneous_ideal), + suffices h : ∀ (x y : localization.away f), x * y ∈ J → x ∈ J ∨ y ∈ J, + { rw [homogeneous_localization.mul_val] at hx12, exact h x1.val x2.val hx12, }, + clear' x1 x2 hx12, intros x1 x2 hx12, induction x1 using localization.induction_on with data_x1, induction x2 using localization.induction_on with data_x2, rcases ⟨data_x1, data_x2⟩ with ⟨⟨a1, _, ⟨n1, rfl⟩⟩, ⟨a2, _, ⟨n2, rfl⟩⟩⟩, - rcases mem_carrier.clear_denominator f_deg x hx12 with ⟨c, N, acd, eq1⟩, - simp only [degree_zero_part.coe_mul, algebra.smul_def] at eq1, + rcases mem_carrier.clear_denominator' x hx12 with ⟨c, N, acd, eq1⟩, + simp only [algebra.smul_def] at eq1, change localization.mk (f ^ N) 1 * (mk _ _ * mk _ _) = mk (∑ _, _) _ at eq1, simp only [localization.mk_mul, one_mul] at eq1, simp only [mk_eq_mk', is_localization.eq] at eq1, @@ -290,22 +226,17 @@ def to_fun (x : Proj.T| (pbo f)) : (Spec.T (A⁰_ f_deg)) := rcases x.1.is_prime.mem_or_mem (show a1 * a2 * f ^ N * f ^ M ∈ _, from _) with h1|rid2, rcases x.1.is_prime.mem_or_mem h1 with h1|rid1, rcases x.1.is_prime.mem_or_mem h1 with h1|h2, - { left, - rw mem_carrier_iff, - simp only [show (mk a1 ⟨f ^ n1, _⟩ : away f) = mk a1 1 * mk 1 ⟨f^n1, ⟨n1, rfl⟩⟩, + { left, simp only [show (mk a1 ⟨f ^ n1, _⟩ : away f) = mk a1 1 * mk 1 ⟨f^n1, ⟨n1, rfl⟩⟩, by rw [localization.mk_mul, mul_one, one_mul]], exact ideal.mul_mem_right _ _ (ideal.subset_span ⟨_, h1, rfl⟩), }, - { right, - rw mem_carrier_iff, - simp only [show (mk a2 ⟨f ^ n2, _⟩ : away f) = mk a2 1 * mk 1 ⟨f^n2, ⟨n2, rfl⟩⟩, + { right, simp only [show (mk a2 ⟨f ^ n2, _⟩ : away f) = mk a2 1 * mk 1 ⟨f^n2, ⟨n2, rfl⟩⟩, by rw [localization.mk_mul, mul_one, one_mul]], exact ideal.mul_mem_right _ _ (ideal.subset_span ⟨_, h2, rfl⟩), }, { exact false.elim (x.2 (x.1.is_prime.mem_of_pow_mem N rid1)), }, { exact false.elim (x.2 (x.1.is_prime.mem_of_pow_mem M rid2)), }, { rw [mul_comm _ (f^N), eq1], refine mul_mem_right _ _ (mul_mem_right _ _ (sum_mem _ (λ i hi, mul_mem_left _ _ _))), - generalize_proofs h, - exact (classical.some_spec h).1 }, + generalize_proofs h, exact (classical.some_spec h).1 }, end⟩ /- @@ -313,11 +244,11 @@ The preimage of basic open set `D(a/f^n)` in `Spec A⁰_f` under the forward map `Spec A⁰_f` is the basic open set `D(a) ∩ D(f)` in `Proj A`. This lemma is used to prove that the forward map is continuous. -/ -lemma preimage_eq (a : A) (n : ℕ) - (a_mem_degree_zero : (mk a ⟨f ^ n, ⟨n, rfl⟩⟩ : away f) ∈ A⁰_ f_deg) : - to_fun 𝒜 f_deg ⁻¹' - ((sbo (⟨mk a ⟨f ^ n, ⟨_, rfl⟩⟩, a_mem_degree_zero⟩ : A⁰_ f_deg)) : - set (prime_spectrum {x // x ∈ A⁰_ f_deg})) +lemma preimage_eq (a b : A) (k : ℕ) (a_mem : a ∈ 𝒜 k) (b_mem1 : b ∈ 𝒜 k) + (b_mem2 : b ∈ submonoid.powers f) : to_fun 𝒜 f ⁻¹' + ((@prime_spectrum.basic_open (A⁰_ f) _ + (quotient.mk' ⟨k, ⟨a, a_mem⟩, ⟨b, b_mem1⟩, b_mem2⟩)) : + set (prime_spectrum (homogeneous_localization.away 𝒜 f))) = {x | x.1 ∈ (pbo f) ⊓ (pbo a)} := begin classical, @@ -327,16 +258,17 @@ begin rw projective_spectrum.mem_coe_basic_open, intro a_mem_y, apply hy, - rw [to_fun, mem_carrier_iff], - simp only [show (mk a ⟨f^n, ⟨_, rfl⟩⟩ : away f) = mk 1 ⟨f^n, ⟨_, rfl⟩⟩ * mk a 1, - by rw [mk_mul, one_mul, mul_one]], + rw [to_fun, mem_carrier_iff, homogeneous_localization.val_mk', subtype.coe_mk], + dsimp, rcases b_mem2 with ⟨k, hk⟩, + simp only [show (mk a ⟨b, ⟨k, hk⟩⟩ : away f) = mk 1 ⟨f^k, ⟨_, rfl⟩⟩ * mk a 1, + by { rw [mk_mul, one_mul, mul_one], congr, rw hk }], exact ideal.mul_mem_left _ _ (ideal.subset_span ⟨_, a_mem_y, rfl⟩), }, { change y.1 ∈ _ at hy, rcases hy with ⟨hy1, hy2⟩, rw projective_spectrum.mem_coe_basic_open at hy1 hy2, rw [set.mem_preimage, to_fun, opens.mem_coe, prime_spectrum.mem_basic_open], - intro rid, - rcases mem_carrier.clear_denominator f_deg _ rid with ⟨c, N, acd, eq1⟩, + intro rid, dsimp at rid, + rcases mem_carrier.clear_denominator 𝒜 _ rid with ⟨c, N, acd, eq1⟩, rw [algebra.smul_def] at eq1, change localization.mk (f^N) 1 * mk _ _ = mk (∑ _, _) _ at eq1, rw [mk_mul, one_mul, mk_eq_mk', is_localization.eq] at eq1, @@ -351,8 +283,7 @@ begin { exact y.2 (y.1.is_prime.mem_of_pow_mem M H3), }, { rw [mul_comm _ (f^N), eq1], refine mul_mem_right _ _ (mul_mem_right _ _ (sum_mem _ (λ i hi, mul_mem_left _ _ _))), - generalize_proofs h, - exact (classical.some_spec h).1, }, }, + generalize_proofs h, exact (classical.some_spec h).1, }, }, end end to_Spec @@ -364,16 +295,12 @@ variable {𝒜} /--The continuous function between the basic open set `D(f)` in `Proj` to the corresponding basic open set in `Spec A⁰_f`. -/ -def to_Spec {f : A} (m : ℕ) (f_deg : f ∈ 𝒜 m) : - (Proj.T| (pbo f)) ⟶ (Spec.T (A⁰_ f_deg)) := -{ to_fun := to_Spec.to_fun 𝒜 f_deg, +def to_Spec {f : A} : (Proj.T| (pbo f)) ⟶ (Spec.T (A⁰_ f)) := +{ to_fun := to_Spec.to_fun 𝒜 f, continuous_to_fun := begin apply is_topological_basis.continuous (prime_spectrum.is_topological_basis_basic_opens), - rintros _ ⟨⟨g, hg⟩, rfl⟩, - induction g using localization.induction_on with data, - obtain ⟨a, ⟨_, ⟨n, rfl⟩⟩⟩ := data, - - erw to_Spec.preimage_eq, + rintros _ ⟨⟨k, ⟨a, ha⟩, ⟨b, hb1⟩, ⟨k', hb2⟩⟩, rfl⟩, dsimp, + erw to_Spec.preimage_eq f a b k ha hb1 ⟨k', hb2⟩, refine is_open_induced_iff.mpr ⟨(pbo f).1 ⊓ (pbo a).1, is_open.inter (pbo f).2 (pbo a).2, _⟩, ext z, split; intros hz; simpa [set.mem_preimage], end } @@ -383,14 +310,17 @@ end namespace from_Spec open graded_algebra set_like finset (hiding mk_zero) +open _root_.homogeneous_localization -variables {𝒜} {f : A} {m : ℕ} {f_deg : f ∈ 𝒜 m} +variables {𝒜} {f : A} {m : ℕ} (f_deg : f ∈ 𝒜 m) private meta def mem_tac : tactic unit := let b : tactic unit := - `[exact pow_mem_graded _ (submodule.coe_mem _) <|> exact nat_cast_mem_graded _ _] in + `[exact pow_mem_graded _ (submodule.coe_mem _) <|> exact nat_cast_mem_graded _ _ <|> + exact pow_mem_graded _ f_deg] in b <|> `[by repeat { all_goals { apply graded_monoid.mul_mem } }; b] +include f_deg /--The function from `Spec A⁰_f` to `Proj|D(f)` is defined by `q ↦ {a | aᵢᵐ/fⁱ ∈ q}`, i.e. sending `q` a prime ideal in `A⁰_f` to the homogeneous prime relevant ideal containing only and all the elements `a : A` such that for every `i`, the degree 0 element formed by dividing the `m`-th power @@ -403,113 +333,137 @@ The set `{a | aᵢᵐ/fⁱ ∈ q}` * is prime, as proved in `carrier.as_ideal.prime`; * is relevant, as proved in `carrier.relevant`. -/ -def carrier (q : Spec.T (A⁰_ f_deg)) : set A := -{a | ∀ i, (⟨mk (proj 𝒜 i a ^ m) ⟨_, _, rfl⟩, i, ⟨_, by mem_tac⟩, rfl⟩ : A⁰_ f_deg) ∈ q.1} - -lemma mem_carrier_iff (q : Spec.T (A⁰_ f_deg)) (a : A) : - a ∈ carrier q ↔ - ∀ i, (⟨mk (proj 𝒜 i a ^ m) ⟨_, _, rfl⟩, i, ⟨_, by mem_tac⟩, rfl⟩ : A⁰_ f_deg) ∈ q.1 := +def carrier (q : Spec.T (A⁰_ f)) : set A := +{a | ∀ i, (quotient.mk' ⟨m * i, ⟨proj 𝒜 i a ^ m, by mem_tac⟩, + ⟨f^i, by rw mul_comm; mem_tac⟩, ⟨_, rfl⟩⟩ : A⁰_ f) ∈ q.1} + +lemma mem_carrier_iff (q : Spec.T (A⁰_ f)) (a : A) : + a ∈ carrier f_deg q ↔ + ∀ i, (quotient.mk' ⟨m * i, ⟨proj 𝒜 i a ^ m, by mem_tac⟩, ⟨f^i, by rw mul_comm; mem_tac⟩, ⟨_, rfl⟩⟩ + : A⁰_ f) ∈ q.1 := iff.rfl -lemma carrier.add_mem (q : Spec.T (A⁰_ f_deg)) {a b : A} (ha : a ∈ carrier q) (hb : b ∈ carrier q) : - a + b ∈ carrier q := +lemma mem_carrier_iff' (q : Spec.T (A⁰_ f)) (a : A) : + a ∈ carrier f_deg q ↔ + ∀ i, (localization.mk (proj 𝒜 i a ^ m) ⟨f^i, ⟨i, rfl⟩⟩ : localization.away f) ∈ + (algebra_map (homogeneous_localization.away 𝒜 f) (localization.away f)) '' q.1.1 := +(mem_carrier_iff f_deg q a).trans begin + split; intros h i; specialize h i, + { rw set.mem_image, refine ⟨_, h, rfl⟩, }, + { rw set.mem_image at h, rcases h with ⟨x, h, hx⟩, + convert h, rw [ext_iff_val, val_mk'], dsimp only [subtype.coe_mk], rw ←hx, refl, }, +end + +lemma carrier.add_mem (q : Spec.T (A⁰_ f)) {a b : A} (ha : a ∈ carrier f_deg q) + (hb : b ∈ carrier f_deg q) : + a + b ∈ carrier f_deg q := begin refine λ i, (q.2.mem_or_mem _).elim id id, - change subtype.mk (localization.mk _ _ * mk _ _) _ ∈ q.1, - simp_rw [mk_mul, ← pow_add, map_add, add_pow, mk_sum, mul_comm, ← nsmul_eq_mul, ← smul_mk], - let g : ℕ → A⁰_ f_deg := λ j, (m + m).choose j • if h2 : m + m < j then 0 else if h1 : j ≤ m - then ⟨mk (proj 𝒜 i a ^ j * proj 𝒜 i b ^ (m - j)) ⟨_, i, rfl⟩, i, ⟨_, _⟩, rfl⟩ * - ⟨mk (proj 𝒜 i b ^ m) ⟨_, i, rfl⟩, i, ⟨_, by mem_tac⟩, rfl⟩ - else ⟨mk (proj 𝒜 i a ^ m) ⟨_, i, rfl⟩, i, ⟨_, by mem_tac⟩, rfl⟩ * - ⟨mk (proj 𝒜 i a ^ (j - m) * proj 𝒜 i b ^ (m + m - j)) ⟨_, i, rfl⟩, i, ⟨_, _⟩, rfl⟩, + change (quotient.mk' ⟨_, _, _, _⟩ : A⁰_ f) ∈ q.1, dsimp only [subtype.coe_mk], + simp_rw [←pow_add, map_add, add_pow, mul_comm, ← nsmul_eq_mul], + let g : ℕ → A⁰_ f := λ j, (m + m).choose j • if h2 : m + m < j then 0 else if h1 : j ≤ m + then quotient.mk' ⟨m * i, ⟨proj 𝒜 i a^j * proj 𝒜 i b ^ (m - j), _⟩, + ⟨_, by rw mul_comm; mem_tac⟩, ⟨i, rfl⟩⟩ * + quotient.mk' ⟨m * i, ⟨proj 𝒜 i b ^ m, by mem_tac⟩, ⟨_, by rw mul_comm; mem_tac⟩, ⟨i, rfl⟩⟩ + else quotient.mk' ⟨m * i, ⟨proj 𝒜 i a ^ m, by mem_tac⟩, + ⟨_, by rw mul_comm; mem_tac⟩, ⟨i, rfl⟩⟩ * quotient.mk' ⟨m * i, ⟨proj 𝒜 i a ^ (j - m) * + proj 𝒜 i b ^ (m + m - j), _⟩, ⟨_, by rw mul_comm; mem_tac⟩, ⟨i, rfl⟩⟩, rotate, - { rw (_ : m * i = _), mem_tac, rw [← add_smul, nat.add_sub_of_le h1], refl }, - { rw (_ : m * i = _), mem_tac, rw ← add_smul, congr, - zify [le_of_not_lt h2, le_of_not_le h1], abel }, + { rw (_ : m*i = _), mem_tac, rw [← add_smul, nat.add_sub_of_le h1], refl }, + { rw (_ : m*i = _), mem_tac, rw ←add_smul, congr, zify [le_of_not_lt h2, le_of_not_le h1], abel }, convert_to ∑ i in range (m + m + 1), g i ∈ q.1, swap, { refine q.1.sum_mem (λ j hj, nsmul_mem _ _), split_ifs, exacts [q.1.zero_mem, q.1.mul_mem_left _ (hb i), q.1.mul_mem_right _ (ha i)] }, - apply subtype.ext, - rw [degree_zero_part.coe_sum, subtype.coe_mk], + rw [ext_iff_val, val_mk'], + change _ = (algebra_map (homogeneous_localization.away 𝒜 f) (localization.away f)) _, + dsimp only [subtype.coe_mk], rw [map_sum, mk_sum], apply finset.sum_congr rfl (λ j hj, _), - congr' 1, split_ifs with h2 h1, + change _ = homogeneous_localization.val _, + rw [homogeneous_localization.smul_val], + split_ifs with h2 h1, { exact ((finset.mem_range.1 hj).not_le h2).elim }, - all_goals { simp only [subtype.val_eq_coe, degree_zero_part.coe_mul, subtype.coe_mk, mk_mul] }, - { rw [mul_assoc, ← pow_add, add_comm (m - j), nat.add_sub_assoc h1] }, - { rw [← mul_assoc, ← pow_add, nat.add_sub_of_le (le_of_not_le h1)] }, + all_goals { simp only [mul_val, zero_val, val_mk', subtype.coe_mk, mk_mul, ←smul_mk], congr' 2 }, + { rw [mul_assoc, ←pow_add, add_comm (m-j), nat.add_sub_assoc h1] }, { simp_rw [pow_add], refl }, + { rw [← mul_assoc, ←pow_add, nat.add_sub_of_le (le_of_not_le h1)] }, { simp_rw [pow_add], refl }, end -variables (hm : 0 < m) (q : Spec.T (A⁰_ f_deg)) +variables (hm : 0 < m) (q : Spec.T (A⁰_ f)) include hm -lemma carrier.zero_mem : (0 : A) ∈ carrier q := -λ i, by simpa only [linear_map.map_zero, zero_pow hm, mk_zero] using submodule.zero_mem _ +lemma carrier.zero_mem : (0 : A) ∈ carrier f_deg q := λ i, begin + convert submodule.zero_mem q.1 using 1, + rw [ext_iff_val, val_mk', zero_val], simp_rw [map_zero, zero_pow hm], + convert localization.mk_zero _ using 1, +end -lemma carrier.smul_mem (c x : A) (hx : x ∈ carrier q) : c • x ∈ carrier q := +lemma carrier.smul_mem (c x : A) (hx : x ∈ carrier f_deg q) : c • x ∈ carrier f_deg q := begin revert c, refine direct_sum.decomposition.induction_on 𝒜 _ _ _, - { rw zero_smul, exact carrier.zero_mem hm _ }, + { rw zero_smul, exact carrier.zero_mem f_deg hm _ }, { rintros n ⟨a, ha⟩ i, simp_rw [subtype.coe_mk, proj_apply, smul_eq_mul, coe_decompose_mul_of_left_mem 𝒜 i ha], split_ifs, - { convert_to (⟨mk _ ⟨_, n, rfl⟩, n, ⟨_, pow_mem_graded m ha⟩, rfl⟩ : A⁰_ f_deg) * - ⟨mk _ ⟨_, i - n, rfl⟩, _, ⟨proj 𝒜 (i - n) x ^ m, by mem_tac⟩, rfl⟩ ∈ q.1, - { erw [subtype.ext_iff, subring.coe_mul, mk_mul, subtype.coe_mk, mul_pow], + { convert_to (quotient.mk' ⟨_, ⟨a^m, pow_mem_graded m ha⟩, ⟨_, _⟩, ⟨n, rfl⟩⟩ * quotient.mk' + ⟨_, ⟨proj 𝒜 (i - n) x ^ m, by mem_tac⟩, ⟨_, _⟩, ⟨i - n, rfl⟩⟩ : A⁰_ f) ∈ q.1, + { erw [ext_iff_val, val_mk', mul_val, val_mk', val_mk', subtype.coe_mk], + simp_rw [mul_pow, subtype.coe_mk], rw [localization.mk_mul], congr, erw [← pow_add, nat.add_sub_of_le h] }, - { exact ideal.mul_mem_left _ _ (hx _) } }, - { simp_rw [zero_pow hm, mk_zero], exact q.1.zero_mem } }, - { simp_rw add_smul, exact λ _ _, carrier.add_mem q }, + { exact ideal.mul_mem_left _ _ (hx _), rw [smul_eq_mul, mul_comm], mem_tac, } }, + { simp_rw [zero_pow hm], convert carrier.zero_mem f_deg hm q i, rw [map_zero, zero_pow hm] } }, + { simp_rw add_smul, exact λ _ _, carrier.add_mem f_deg q }, end /-- For a prime ideal `q` in `A⁰_f`, the set `{a | aᵢᵐ/fⁱ ∈ q}` as an ideal. -/ def carrier.as_ideal : ideal A := -{ carrier := carrier q, - zero_mem' := carrier.zero_mem hm q, - add_mem' := λ a b, carrier.add_mem q, - smul_mem' := carrier.smul_mem hm q } +{ carrier := carrier f_deg q, + zero_mem' := carrier.zero_mem f_deg hm q, + add_mem' := λ a b, carrier.add_mem f_deg q, + smul_mem' := carrier.smul_mem f_deg hm q } -lemma carrier.as_ideal.homogeneous : (carrier.as_ideal hm q).is_homogeneous 𝒜 := +lemma carrier.as_ideal.homogeneous : (carrier.as_ideal f_deg hm q).is_homogeneous 𝒜 := λ i a ha j, (em (i = j)).elim (λ h, h ▸ by simpa only [proj_apply, decompose_coe, of_eq_same] using ha _) - (λ h, by simpa only [proj_apply, decompose_of_mem_ne 𝒜 (submodule.coe_mem (decompose 𝒜 a i)) h, - zero_pow hm, mk_zero] using submodule.zero_mem _) + (λ h, begin + simp only [proj_apply, decompose_of_mem_ne 𝒜 (submodule.coe_mem (decompose 𝒜 a i)) h, + zero_pow hm], convert carrier.zero_mem f_deg hm q j, rw [map_zero, zero_pow hm], + end) /-- For a prime ideal `q` in `A⁰_f`, the set `{a | aᵢᵐ/fⁱ ∈ q}` as a homogeneous ideal. -/ def carrier.as_homogeneous_ideal : homogeneous_ideal 𝒜 := -⟨carrier.as_ideal hm q, carrier.as_ideal.homogeneous hm q⟩ +⟨carrier.as_ideal f_deg hm q, carrier.as_ideal.homogeneous f_deg hm q⟩ -lemma carrier.denom_not_mem : f ∉ carrier.as_ideal hm q := +lemma carrier.denom_not_mem : f ∉ carrier.as_ideal f_deg hm q := λ rid, q.is_prime.ne_top $ (ideal.eq_top_iff_one _).mpr begin convert rid m, - simpa only [subtype.ext_iff, degree_zero_part.coe_one, subtype.coe_mk, proj_apply, - decompose_of_mem_same _ f_deg] using (mk_self (⟨_, m, rfl⟩ : submonoid.powers f)).symm, + simpa only [ext_iff_val, one_val, proj_apply, decompose_of_mem_same _ f_deg, val_mk'] using + (mk_self (⟨_, m, rfl⟩ : submonoid.powers f)).symm, end -lemma carrier.relevant : ¬ homogeneous_ideal.irrelevant 𝒜 ≤ carrier.as_homogeneous_ideal hm q := -λ rid, carrier.denom_not_mem hm q $ rid $ direct_sum.decompose_of_mem_ne 𝒜 f_deg hm.ne' +lemma carrier.relevant : + ¬homogeneous_ideal.irrelevant 𝒜 ≤ carrier.as_homogeneous_ideal f_deg hm q := +λ rid, carrier.denom_not_mem f_deg hm q $ rid $ direct_sum.decompose_of_mem_ne 𝒜 f_deg hm.ne' -lemma carrier.as_ideal.ne_top : (carrier.as_ideal hm q) ≠ ⊤ := -λ rid, carrier.denom_not_mem hm q (rid.symm ▸ submodule.mem_top) +lemma carrier.as_ideal.ne_top : (carrier.as_ideal f_deg hm q) ≠ ⊤ := +λ rid, carrier.denom_not_mem f_deg hm q (rid.symm ▸ submodule.mem_top) -lemma carrier.as_ideal.prime : (carrier.as_ideal hm q).is_prime := -(carrier.as_ideal.homogeneous hm q).is_prime_of_homogeneous_mem_or_mem - (carrier.as_ideal.ne_top hm q) $ λ x y ⟨nx, hnx⟩ ⟨ny, hny⟩ hxy, show (∀ i, _ ∈ _) ∨ ∀ i, _ ∈ _, -begin +lemma carrier.as_ideal.prime : (carrier.as_ideal f_deg hm q).is_prime := +(carrier.as_ideal.homogeneous f_deg hm q).is_prime_of_homogeneous_mem_or_mem + (carrier.as_ideal.ne_top f_deg hm q) $ λ x y ⟨nx, hnx⟩ ⟨ny, hny⟩ hxy, +show (∀ i, _ ∈ _) ∨ ∀ i, _ ∈ _, begin rw [← and_forall_ne nx, and_iff_left, ← and_forall_ne ny, and_iff_left], { apply q.2.mem_or_mem, convert hxy (nx + ny) using 1, simp_rw [proj_apply, decompose_of_mem_same 𝒜 hnx, decompose_of_mem_same 𝒜 hny, decompose_of_mem_same 𝒜 (mul_mem hnx hny), mul_pow, pow_add], - exact subtype.ext (mk_mul _ _ _ _) }, - all_goals { intros n hn, - convert q.1.zero_mem using 2, - rw [proj_apply, decompose_of_mem_ne 𝒜 _ hn.symm], - { rw [zero_pow hm, mk_zero] }, + simpa only [ext_iff_val, val_mk', mul_val, mk_mul], }, + all_goals { intros n hn, convert q.1.zero_mem using 1, + rw [ext_iff_val, val_mk', zero_val], simp_rw [proj_apply, subtype.coe_mk], + convert mk_zero _, rw [decompose_of_mem_ne 𝒜 _ hn.symm, zero_pow hm], { exact hnx <|> exact hny } }, end @@ -517,9 +471,10 @@ variable (f_deg) /-- The function `Spec A⁰_f → Proj|D(f)` by sending `q` to `{a | aᵢᵐ/fⁱ ∈ q}`. -/ -def to_fun : (Spec.T (A⁰_ f_deg)) → (Proj.T| (pbo f)) := -λ q, ⟨⟨carrier.as_homogeneous_ideal hm q, carrier.as_ideal.prime hm q, carrier.relevant hm q⟩, - (projective_spectrum.mem_basic_open _ f _).mp $ carrier.denom_not_mem hm q⟩ +def to_fun : (Spec.T (A⁰_ f)) → (Proj.T| (pbo f)) := +λ q, ⟨⟨carrier.as_homogeneous_ideal f_deg hm q, carrier.as_ideal.prime f_deg hm q, + carrier.relevant f_deg hm q⟩, + (projective_spectrum.mem_basic_open _ f _).mp $ carrier.denom_not_mem f_deg hm q⟩ end from_Spec From cb7da1b4d0f6aea6959add981b6ac9272cf7c7f2 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 5 Oct 2022 18:04:12 +0000 Subject: [PATCH 584/828] refactor(algebra/group/with_one): `inv_one_class` and related instances (#16697) Add instances for `neg_zero_class`, `inv_one_class` and `div_inv_one_monoid` for `with_zero`, and `inv_one_class` for `with_one`, under appropriate conditions. --- src/algebra/group/with_one.lean | 17 ++++++++++++++--- 1 file changed, 14 insertions(+), 3 deletions(-) diff --git a/src/algebra/group/with_one.lean b/src/algebra/group/with_one.lean index c568d8b9acd00..0b06f4b89f4bc 100644 --- a/src/algebra/group/with_one.lean +++ b/src/algebra/group/with_one.lean @@ -51,6 +51,11 @@ instance [has_mul α] : has_mul (with_one α) := ⟨option.lift_or_get (*)⟩ { inv_inv := λ a, (option.map_map _ _ _).trans $ by simp_rw [inv_comp_inv, option.map_id, id], ..with_one.has_inv } +@[to_additive] instance [has_inv α] : inv_one_class (with_one α) := +{ inv_one := rfl, + ..with_one.has_one, + ..with_one.has_inv } + @[to_additive] instance : inhabited (with_one α) := ⟨1⟩ @@ -326,6 +331,11 @@ instance [has_involutive_inv α] : has_involutive_inv (with_zero α) := { inv_inv := λ a, (option.map_map _ _ _).trans $ by simp_rw [inv_comp_inv, option.map_id, id], ..with_zero.has_inv } +instance [inv_one_class α] : inv_one_class (with_zero α) := +{ inv_one := show ((1⁻¹ : α) : with_zero α) = 1, by simp, + ..with_zero.has_one, + ..with_zero.has_inv } + instance [has_div α] : has_div (with_zero α) := ⟨λ o₁ o₂, o₁.bind (λ a, option.map (λ b, a / b) o₂)⟩ @@ -365,6 +375,10 @@ instance [div_inv_monoid α] : div_inv_monoid (with_zero α) := .. with_zero.has_inv, .. with_zero.monoid_with_zero, } +instance [div_inv_one_monoid α] : div_inv_one_monoid (with_zero α) := +{ ..with_zero.div_inv_monoid, + ..with_zero.inv_one_class } + instance [division_monoid α] : division_monoid (with_zero α) := { mul_inv_rev := λ a b, match a, b with | none, none := rfl @@ -386,9 +400,6 @@ instance [division_comm_monoid α] : division_comm_monoid (with_zero α) := section group variables [group α] -@[simp] lemma inv_one : (1 : with_zero α)⁻¹ = 1 := -show ((1⁻¹ : α) : with_zero α) = 1, by simp - /-- if `G` is a group then `with_zero G` is a group with zero. -/ instance : group_with_zero (with_zero α) := { inv_zero := inv_zero, From 83bf40c1fcc0310a8ddc5f62e7da4b9bc2f441e1 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Wed, 5 Oct 2022 18:04:13 +0000 Subject: [PATCH 585/828] refactor(ring_theory/roots_of_unity): Remove redundant variables (#16798) Followup to #16755 --- src/ring_theory/roots_of_unity.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/ring_theory/roots_of_unity.lean b/src/ring_theory/roots_of_unity.lean index c14112380a9fc..c5ebf12cbdb37 100644 --- a/src/ring_theory/roots_of_unity.lean +++ b/src/ring_theory/roots_of_unity.lean @@ -319,7 +319,7 @@ end lemma pow_ne_one_of_pos_of_lt (h0 : 0 < l) (hl : l < k) : ζ ^ l ≠ 1 := mt (nat.le_of_dvd h0 ∘ h.dvd_of_pow_eq_one _) $ not_le_of_lt hl -lemma ne_one {ζ : M} (h : is_primitive_root ζ k) (hk : 1 < k) : ζ ≠ 1 := +lemma ne_one (hk : 1 < k) : ζ ≠ 1 := h.pow_ne_one_of_pos_of_lt zero_lt_one hk ∘ (pow_one ζ).trans lemma pow_inj (h : is_primitive_root ζ k) ⦃i j : ℕ⦄ (hi : i < k) (hj : j < k) (H : ζ ^ i = ζ ^ j) : From daa660ca996766ca2b4dddb292793997f9bfa8c5 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Wed, 5 Oct 2022 18:04:14 +0000 Subject: [PATCH 586/828] feat(set_theory/cardinal/finite): Add `nat.card_fun` (#16799) This PR adds `nat.card_fun`, a quick consequence of `nat.card_pi`. --- src/set_theory/cardinal/finite.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/set_theory/cardinal/finite.lean b/src/set_theory/cardinal/finite.lean index 253489173c994..0aa0b280c13cd 100644 --- a/src/set_theory/cardinal/finite.lean +++ b/src/set_theory/cardinal/finite.lean @@ -84,6 +84,12 @@ card_congr equiv.plift lemma card_pi {β : α → Type*} [fintype α] : nat.card (Π a, β a) = ∏ a, nat.card (β a) := by simp_rw [nat.card, mk_pi, prod_eq_of_fintype, to_nat_lift, to_nat_finset_prod] +lemma card_fun [finite α] : nat.card (α → β) = nat.card β ^ nat.card α := +begin + haveI := fintype.of_finite α, + rw [nat.card_pi, finset.prod_const, finset.card_univ, ←nat.card_eq_fintype_card], +end + @[simp] lemma card_zmod (n : ℕ) : nat.card (zmod n) = n := begin cases n, From 7c0ff89beed56d8dff40e4bfee8a3ea286f646e3 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 5 Oct 2022 18:04:16 +0000 Subject: [PATCH 587/828] perf(analysis/calculus/specific_functions): speed up `exp_neg_inv_glue.f_aux_deriv` (#16812) This caused some [bors failures](https://github.com/leanprover-community/mathlib/actions/runs/3182965873/jobs/5189596416) yesterday. The new proof features a 8x time reduction from >20s to <2.5s on my machine. --- src/analysis/calculus/specific_functions.lean | 13 ++++++++----- 1 file changed, 8 insertions(+), 5 deletions(-) diff --git a/src/analysis/calculus/specific_functions.lean b/src/analysis/calculus/specific_functions.lean index 5756515b7b3ac..866d58ee774ee 100644 --- a/src/analysis/calculus/specific_functions.lean +++ b/src/analysis/calculus/specific_functions.lean @@ -64,7 +64,7 @@ derivatives for `x > 0`. The `n`-th derivative is of the form `P_aux n (x) exp(- where `P_aux n` is computed inductively. -/ noncomputable def P_aux : ℕ → polynomial ℝ | 0 := 1 -| (n+1) := X^2 * (P_aux n).derivative + (1 - C ↑(2 * n) * X) * (P_aux n) +| (n+1) := X^2 * (P_aux n).derivative + (1 - C ↑(2 * n) * X) * (P_aux n) /-- Formula for the `n`-th derivative of `exp_neg_inv_glue`, as an auxiliary function `f_aux`. -/ def f_aux (n : ℕ) (x : ℝ) : ℝ := @@ -86,13 +86,16 @@ lemma f_aux_deriv (n : ℕ) (x : ℝ) (hx : x ≠ 0) : has_deriv_at (λx, (P_aux n).eval x * exp (-x⁻¹) / x^(2 * n)) ((P_aux (n+1)).eval x * exp (-x⁻¹) / x^(2 * (n + 1))) x := begin - have A : ∀ k : ℕ, 2 * (k + 1) - 1 = 2 * k + 1 := λ k, rfl, + simp only [P_aux, eval_add, eval_sub, eval_mul, eval_pow, eval_X, eval_C, eval_one], convert (((P_aux n).has_deriv_at x).mul (((has_deriv_at_exp _).comp x (has_deriv_at_inv hx).neg))).div (has_deriv_at_pow (2 * n) x) (pow_ne_zero _ hx) using 1, - field_simp [hx, P_aux], - -- `ring_exp` can't solve `p ∨ q` goal generated by `mul_eq_mul_right_iff` - cases n; simp [nat.succ_eq_add_one, A, -mul_eq_mul_right_iff]; ring_exp + rw div_eq_div_iff, + { have := pow_ne_zero 2 hx, field_simp only, + cases n, + { simp only [mul_zero, nat.cast_zero, mul_one], ring }, + { rw (id rfl : 2 * n.succ - 1 = 2 * n + 1), ring_exp } }, + all_goals { apply_rules [pow_ne_zero] }, end /-- For positive values, the derivative of the `n`-th auxiliary function `f_aux n` From 7952bc6d522d4e16997826a9d9e17ad08ede11f7 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 5 Oct 2022 21:07:33 +0000 Subject: [PATCH 588/828] feat(topology/locally_finite): add lemmas about `option _` and `sum _ _` (#15923) Co-authored-by: Jireh Loreaux --- src/data/set/basic.lean | 5 +++ src/data/set/finite.lean | 5 +++ src/topology/locally_finite.lean | 58 +++++++++++++++++++++----------- 3 files changed, 48 insertions(+), 20 deletions(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 930f76a57c451..29ff5f21a2380 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -2158,6 +2158,11 @@ is_compl.compl_eq is_compl_range_inl_range_inr @[simp] lemma compl_range_inr : (range (sum.inr : β → α ⊕ β))ᶜ = range (sum.inl : α → α ⊕ β) := is_compl.compl_eq is_compl_range_inl_range_inr.symm +theorem image_preimage_inl_union_image_preimage_inr (s : set (α ⊕ β)) : + sum.inl '' (sum.inl ⁻¹' s) ∪ sum.inr '' (sum.inr ⁻¹' s) = s := +by rw [image_preimage_eq_inter_range, image_preimage_eq_inter_range, ← inter_distrib_left, + range_inl_union_range_inr, inter_univ] + @[simp] theorem range_quot_mk (r : α → α → Prop) : range (quot.mk r) = univ := (surjective_quot_mk r).range_eq diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 91ab26af7a1e3..a96ebf110fa03 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -597,6 +597,11 @@ theorem finite_mem_finset (s : finset α) : {a | a ∈ s}.finite := to_finite _ lemma subsingleton.finite {s : set α} (h : s.subsingleton) : s.finite := h.induction_on finite_empty finite_singleton +lemma finite_preimage_inl_and_inr {s : set (α ⊕ β)} : + (sum.inl ⁻¹' s).finite ∧ (sum.inr ⁻¹' s).finite ↔ s.finite := +⟨λ h, image_preimage_inl_union_image_preimage_inr s ▸ (h.1.image _).union (h.2.image _), + λ h, ⟨h.preimage (sum.inl_injective.inj_on _), h.preimage (sum.inr_injective.inj_on _)⟩⟩ + theorem exists_finite_iff_finset {p : set α → Prop} : (∃ s : set α, s.finite ∧ p s) ↔ ∃ s : finset α, p ↑s := ⟨λ ⟨s, hs, hps⟩, ⟨hs.to_finset, hs.coe_to_finset.symm ▸ hps⟩, diff --git a/src/topology/locally_finite.lean b/src/topology/locally_finite.lean index ded3c28916569..ef1b37e5b8318 100644 --- a/src/topology/locally_finite.lean +++ b/src/topology/locally_finite.lean @@ -19,7 +19,8 @@ In this file we give the definition and prove basic properties of locally finite open set function filter open_locale topological_space filter -variables {ι ι' α X Y : Type*} [topological_space X] [topological_space Y] +universe u +variables {ι : Type u} {ι' α X Y : Type*} [topological_space X] [topological_space Y] {f g : ι → set X} /-- A family of sets in `set X` is locally finite if at every point `x : X`, @@ -45,34 +46,25 @@ lemma comp_inj_on {g : ι' → ι} (hf : locally_finite f) λ x, let ⟨t, htx, htf⟩ := hf x in ⟨t, htx, htf.preimage $ hg.mono $ λ i hi, hi.out.mono $ inter_subset_left _ _⟩ -lemma comp_injective {g : ι' → ι} (hf : locally_finite f) - (hg : function.injective g) : locally_finite (f ∘ g) := +lemma comp_injective {g : ι' → ι} (hf : locally_finite f) (hg : injective g) : + locally_finite (f ∘ g) := hf.comp_inj_on (hg.inj_on _) -lemma eventually_finite (hf : locally_finite f) (x : X) : +lemma _root_.locally_finite_iff_small_sets : + locally_finite f ↔ ∀ x, ∀ᶠ s in (𝓝 x).small_sets, {i | (f i ∩ s).nonempty}.finite := +forall_congr $ λ x, iff.symm $ eventually_small_sets' $ λ s t hst ht, ht.subset $ + λ i hi, hi.mono $ inter_subset_inter_right _ hst + +protected lemma eventually_small_sets (hf : locally_finite f) (x : X) : ∀ᶠ s in (𝓝 x).small_sets, {i | (f i ∩ s).nonempty}.finite := -eventually_small_sets.2 $ let ⟨s, hsx, hs⟩ := hf x in - ⟨s, hsx, λ t hts, hs.subset $ λ i hi, hi.out.mono $ inter_subset_inter_right _ hts⟩ +locally_finite_iff_small_sets.mp hf x lemma exists_mem_basis {ι' : Sort*} (hf : locally_finite f) {p : ι' → Prop} {s : ι' → set X} {x : X} (hb : (𝓝 x).has_basis p s) : ∃ i (hi : p i), {j | (f j ∩ s i).nonempty}.finite := -let ⟨i, hpi, hi⟩ := hb.small_sets.eventually_iff.mp (hf.eventually_finite x) +let ⟨i, hpi, hi⟩ := hb.small_sets.eventually_iff.mp (hf.eventually_small_sets x) in ⟨i, hpi, hi subset.rfl⟩ -lemma sum_elim {g : ι' → set X} (hf : locally_finite f) (hg : locally_finite g) : - locally_finite (sum.elim f g) := -begin - intro x, - obtain ⟨s, hsx, hsf, hsg⟩ : - ∃ s, s ∈ 𝓝 x ∧ {i | (f i ∩ s).nonempty}.finite ∧ {j | (g j ∩ s).nonempty}.finite, - from ((𝓝 x).frequently_small_sets_mem.and_eventually - ((hf.eventually_finite x).and (hg.eventually_finite x))).exists, - refine ⟨s, hsx, _⟩, - convert (hsf.image sum.inl).union (hsg.image sum.inr) using 1, - ext (i|j); simp -end - protected lemma closure (hf : locally_finite f) : locally_finite (λ i, closure (f i)) := begin intro x, @@ -162,3 +154,29 @@ lemma preimage_continuous {g : Y → X} (hf : locally_finite f) (hg : continuous in ⟨g ⁻¹' s, hg.continuous_at hsx, hs.subset $ λ i ⟨y, hy⟩, ⟨g y, hy⟩⟩ end locally_finite + +@[simp] lemma equiv.locally_finite_comp_iff (e : ι' ≃ ι) : + locally_finite (f ∘ e) ↔ locally_finite f := +⟨λ h, by simpa only [(∘), e.apply_symm_apply] using h.comp_injective e.symm.injective, + λ h, h.comp_injective e.injective⟩ + +lemma locally_finite_sum {f : ι ⊕ ι' → set X} : + locally_finite f ↔ locally_finite (f ∘ sum.inl) ∧ locally_finite (f ∘ sum.inr) := +by simp only [locally_finite_iff_small_sets, ← forall_and_distrib, ← finite_preimage_inl_and_inr, + preimage_set_of_eq, (∘), eventually_and] + +lemma locally_finite.sum_elim {g : ι' → set X} (hf : locally_finite f) (hg : locally_finite g) : + locally_finite (sum.elim f g) := +locally_finite_sum.mpr ⟨hf, hg⟩ + +lemma locally_finite_option {f : option ι → set X} : + locally_finite f ↔ locally_finite (f ∘ some) := +begin + simp only [← (equiv.option_equiv_sum_punit.{u} ι).symm.locally_finite_comp_iff, + locally_finite_sum, locally_finite_of_finite, and_true], + refl +end + +lemma locally_finite.option_elim (hf : locally_finite f) (s : set X) : + locally_finite (option.elim s f) := +locally_finite_option.2 hf From 96055b68cc87d46e7d2691c0152034780e7596d7 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 5 Oct 2022 21:07:34 +0000 Subject: [PATCH 589/828] feat(analysis/inner_product_space/gram_schmidt_ortho): generalize orthonormal basis construction (#16477) Previously, the construction `gram_schmidt_orthonormal_basis` performed Gram-Schmidt orthogonalization on an existing basis to produce an orthonormal basis. This changes it so that it outputs an orthonormal basis if the input is any family of vectors with the right sized index set. It augments the naive Gram-Schmidt output (an orthonormal-or-zero family of vectors, not necessarily spanning) arbitrarily to make an orthonormal basis. Formalized as part of the Sphere Eversion project. --- .../gram_schmidt_ortho.lean | 201 +++++++++++++++--- 1 file changed, 174 insertions(+), 27 deletions(-) diff --git a/src/analysis/inner_product_space/gram_schmidt_ortho.lean b/src/analysis/inner_product_space/gram_schmidt_ortho.lean index 12eff1b6066b1..a02e9265910bc 100644 --- a/src/analysis/inner_product_space/gram_schmidt_ortho.lean +++ b/src/analysis/inner_product_space/gram_schmidt_ortho.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jiale Miao, Kevin Buzzard, Alexander Bentkamp -/ -import analysis.inner_product_space.projection -import order.well_founded_set import analysis.inner_product_space.pi_L2 +import order.well_founded_set +import linear_algebra.matrix.block /-! # Gram-Schmidt Orthogonalization and Orthonormalization @@ -32,13 +32,12 @@ and outputs a set of orthogonal vectors which have the same span. the normalized `gram_schmidt` (i.e each vector in `gram_schmidt_normed` has unit length.) - `gram_schmidt_orthornormal` : `gram_schmidt_normed` produces an orthornormal system of vectors. - -## TODO - Construct a version with an orthonormal basis from Gram-Schmidt process. +- `gram_schmidt_orthonormal_basis`: orthonormal basis constructed by the Gram-Schmidt process from + an indexed set of vectors of the right size -/ open_locale big_operators -open finset +open finset submodule finite_dimensional variables (𝕜 : Type*) {E : Type*} [is_R_or_C 𝕜] [inner_product_space 𝕜 E] variables {ι : Type*} [linear_order ι] [locally_finite_order_bot ι] [is_well_order ι (<)] @@ -64,6 +63,15 @@ lemma gram_schmidt_def' (f : ι → E) (n : ι): orthogonal_projection (𝕜 ∙ gram_schmidt 𝕜 f i) (f n) := by rw [gram_schmidt_def, sub_add_cancel] +lemma gram_schmidt_def'' (f : ι → E) (n : ι): + f n = gram_schmidt 𝕜 f n + + ∑ i in Iio n, (⟪gram_schmidt 𝕜 f i, f n⟫ / ∥gram_schmidt 𝕜 f i∥ ^ 2) • gram_schmidt 𝕜 f i := +begin + convert gram_schmidt_def' 𝕜 f n, + ext i, + rw orthogonal_projection_singleton, +end + @[simp] lemma gram_schmidt_zero {ι : Type*} [linear_order ι] [locally_finite_order ι] [order_bot ι] [is_well_order ι (<)] (f : ι → E) : gram_schmidt 𝕜 f ⊥ = f ⊥ := by rw [gram_schmidt_def, Iio_eq_Ico, finset.Ico_self, finset.sum_empty, sub_zero] @@ -104,6 +112,21 @@ theorem gram_schmidt_pairwise_orthogonal (f : ι → E) : pairwise (λ a b, ⟪gram_schmidt 𝕜 f a, gram_schmidt 𝕜 f b⟫ = 0) := λ a b, gram_schmidt_orthogonal 𝕜 f +lemma gram_schmidt_inv_triangular (v : ι → E) {i j : ι} (hij : i < j) : + ⟪gram_schmidt 𝕜 v j, v i⟫ = 0 := +begin + rw gram_schmidt_def'' 𝕜 v, + simp only [inner_add_right, inner_sum, inner_smul_right], + set b : ι → E := gram_schmidt 𝕜 v, + convert zero_add (0:𝕜), + { exact gram_schmidt_orthogonal 𝕜 v hij.ne' }, + apply finset.sum_eq_zero, + rintros k hki', + have hki : k < i := by simpa using hki', + have : ⟪b j, b k⟫ = 0 := gram_schmidt_orthogonal 𝕜 v (hki.trans hij).ne', + simp [this], +end + open submodule set order lemma mem_span_gram_schmidt (f : ι → E) {i j : ι} (hij : i ≤ j) : @@ -149,8 +172,34 @@ span_eq_span (range_subset_iff.2 $ λ i, span_mono (image_subset_range _ _) $ gram_schmidt_mem_span _ _ le_rfl) $ range_subset_iff.2 $ λ i, span_mono (image_subset_range _ _) $ mem_span_gram_schmidt _ _ le_rfl +lemma gram_schmidt_of_orthogonal {f : ι → E} (hf : pairwise (λ i j, ⟪f i, f j⟫ = 0)) : + gram_schmidt 𝕜 f = f := +begin + ext i, + rw gram_schmidt_def, + transitivity f i - 0, + { congr, + apply finset.sum_eq_zero, + intros j hj, + rw coe_eq_zero, + suffices : span 𝕜 (f '' set.Iic j) ≤ (𝕜 ∙ f i)ᗮ, + { apply orthogonal_projection_mem_subspace_orthogonal_complement_eq_zero, + apply mem_orthogonal_singleton_of_inner_left, + apply inner_right_of_mem_orthogonal_singleton, + exact this (gram_schmidt_mem_span 𝕜 f (le_refl j)) }, + rw span_le, + rintros - ⟨k, hk, rfl⟩, + apply mem_orthogonal_singleton_of_inner_left, + apply hf, + refine (lt_of_le_of_lt hk _).ne, + simpa using hj }, + { simp }, +end + +variables {𝕜} + lemma gram_schmidt_ne_zero_coe - (f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 (f ∘ (coe : set.Iic n → ι))) : + {f : ι → E} (n : ι) (h₀ : linear_independent 𝕜 (f ∘ (coe : set.Iic n → ι))) : gram_schmidt 𝕜 f n ≠ 0 := begin by_contra h, @@ -174,9 +223,9 @@ end /-- If the input vectors of `gram_schmidt` are linearly independent, then the output vectors are non-zero. -/ -lemma gram_schmidt_ne_zero (f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 f) : +lemma gram_schmidt_ne_zero {f : ι → E} (n : ι) (h₀ : linear_independent 𝕜 f) : gram_schmidt 𝕜 f n ≠ 0 := -gram_schmidt_ne_zero_coe _ _ _ (linear_independent.comp h₀ _ subtype.coe_injective) +gram_schmidt_ne_zero_coe _ (linear_independent.comp h₀ _ subtype.coe_injective) /-- `gram_schmidt` produces a triangular matrix of vectors when given a basis. -/ lemma gram_schmidt_triangular {i j : ι} (hij : i < j) (b : basis ι 𝕜 E) : @@ -193,38 +242,51 @@ begin end /-- `gram_schmidt` produces linearly independent vectors when given linearly independent vectors. -/ -lemma gram_schmidt_linear_independent (f : ι → E) (h₀ : linear_independent 𝕜 f) : +lemma gram_schmidt_linear_independent {f : ι → E} (h₀ : linear_independent 𝕜 f) : linear_independent 𝕜 (gram_schmidt 𝕜 f) := linear_independent_of_ne_zero_of_inner_eq_zero - (λ i, gram_schmidt_ne_zero _ _ _ h₀) (λ i j, gram_schmidt_orthogonal 𝕜 f) + (λ i, gram_schmidt_ne_zero _ h₀) (λ i j, gram_schmidt_orthogonal 𝕜 f) /-- When given a basis, `gram_schmidt` produces a basis. -/ noncomputable def gram_schmidt_basis (b : basis ι 𝕜 E) : basis ι 𝕜 E := basis.mk - (gram_schmidt_linear_independent 𝕜 b b.linear_independent) + (gram_schmidt_linear_independent b.linear_independent) ((span_gram_schmidt 𝕜 b).trans b.span_eq).ge lemma coe_gram_schmidt_basis (b : basis ι 𝕜 E) : - (gram_schmidt_basis 𝕜 b : ι → E) = gram_schmidt 𝕜 b := basis.coe_mk _ _ + (gram_schmidt_basis b : ι → E) = gram_schmidt 𝕜 b := basis.coe_mk _ _ + +variables (𝕜) /-- the normalized `gram_schmidt` (i.e each vector in `gram_schmidt_normed` has unit length.) -/ noncomputable def gram_schmidt_normed (f : ι → E) (n : ι) : E := (∥gram_schmidt 𝕜 f n∥ : 𝕜)⁻¹ • (gram_schmidt 𝕜 f n) +variables {𝕜} + lemma gram_schmidt_normed_unit_length_coe - (f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 (f ∘ (coe : set.Iic n → ι))) : + {f : ι → E} (n : ι) (h₀ : linear_independent 𝕜 (f ∘ (coe : set.Iic n → ι))) : ∥gram_schmidt_normed 𝕜 f n∥ = 1 := -by simp only [gram_schmidt_ne_zero_coe 𝕜 f n h₀, +by simp only [gram_schmidt_ne_zero_coe n h₀, gram_schmidt_normed, norm_smul_inv_norm, ne.def, not_false_iff] -lemma gram_schmidt_normed_unit_length (f : ι → E) (n : ι) (h₀ : linear_independent 𝕜 f) : +lemma gram_schmidt_normed_unit_length {f : ι → E} (n : ι) (h₀ : linear_independent 𝕜 f) : ∥gram_schmidt_normed 𝕜 f n∥ = 1 := -gram_schmidt_normed_unit_length_coe _ _ _ (linear_independent.comp h₀ _ subtype.coe_injective) +gram_schmidt_normed_unit_length_coe _ (linear_independent.comp h₀ _ subtype.coe_injective) + +lemma gram_schmidt_normed_unit_length' {f : ι → E} {n : ι} (hn : gram_schmidt_normed 𝕜 f n ≠ 0) : + ∥gram_schmidt_normed 𝕜 f n∥ = 1 := +begin + rw gram_schmidt_normed at *, + rw [norm_smul_inv_norm], + simpa using hn, +end /-- **Gram-Schmidt Orthonormalization**: -`gram_schmidt_normed` produces an orthornormal system of vectors. -/ -theorem gram_schmidt_orthonormal (f : ι → E) (h₀ : linear_independent 𝕜 f) : +`gram_schmidt_normed` applied to a linearly independent set of vectors produces an orthornormal +system of vectors. -/ +theorem gram_schmidt_orthonormal {f : ι → E} (h₀ : linear_independent 𝕜 f) : orthonormal 𝕜 (gram_schmidt_normed 𝕜 f) := begin unfold orthonormal, @@ -237,6 +299,18 @@ begin exact gram_schmidt_orthogonal 𝕜 f hij } end +/-- **Gram-Schmidt Orthonormalization**: +`gram_schmidt_normed` produces an orthornormal system of vectors after removing the vectors which +become zero in the process. -/ +lemma gram_schmidt_orthonormal' (f : ι → E) : + orthonormal 𝕜 (λ i : {i | gram_schmidt_normed 𝕜 f i ≠ 0}, gram_schmidt_normed 𝕜 f i) := +begin + refine ⟨λ i, gram_schmidt_normed_unit_length' i.prop, _⟩, + rintros i j (hij : ¬ _), + rw subtype.ext_iff at hij, + simp [gram_schmidt_normed, inner_smul_left, inner_smul_right, gram_schmidt_orthogonal 𝕜 f hij], +end + lemma span_gram_schmidt_normed (f : ι → E) (s : set ι) : span 𝕜 (gram_schmidt_normed 𝕜 f '' s) = span 𝕜 (gram_schmidt 𝕜 f '' s) := begin @@ -252,11 +326,84 @@ end lemma span_gram_schmidt_normed_range (f : ι → E) : span 𝕜 (range (gram_schmidt_normed 𝕜 f)) = span 𝕜 (range (gram_schmidt 𝕜 f)) := -by simpa only [image_univ.symm] using span_gram_schmidt_normed 𝕜 f univ - -/-- When given a basis, `gram_schmidt_normed` produces an orthonormal basis. -/ -noncomputable def gram_schmidt_orthonormal_basis [fintype ι] (b : basis ι 𝕜 E) : - orthonormal_basis ι 𝕜 E := -orthonormal_basis.mk - (gram_schmidt_orthonormal 𝕜 b b.linear_independent) - (((span_gram_schmidt_normed_range 𝕜 b).trans (span_gram_schmidt 𝕜 b)).trans b.span_eq).ge +by simpa only [image_univ.symm] using span_gram_schmidt_normed f univ + +section orthonormal_basis +variables [fintype ι] [finite_dimensional 𝕜 E] (h : finrank 𝕜 E = fintype.card ι) (f : ι → E) +include h + +/-- Given an indexed family `f : ι → E` of vectors in an inner product space `E`, for which the +size of the index set is the dimension of `E`, produce an orthonormal basis for `E` which agrees +with the orthonormal set produced by the Gram-Schmidt orthonormalization process on the elements of +`ι` for which this process gives a nonzero number. -/ +noncomputable def gram_schmidt_orthonormal_basis : orthonormal_basis ι 𝕜 E := +((gram_schmidt_orthonormal' f).exists_orthonormal_basis_extension_of_card_eq h).some + +lemma gram_schmidt_orthonormal_basis_apply {f : ι → E} {i : ι} + (hi : gram_schmidt_normed 𝕜 f i ≠ 0) : + gram_schmidt_orthonormal_basis h f i = gram_schmidt_normed 𝕜 f i := +((gram_schmidt_orthonormal' f).exists_orthonormal_basis_extension_of_card_eq h).some_spec i hi + +lemma gram_schmidt_orthonormal_basis_apply_of_orthogonal {f : ι → E} + (hf : pairwise (λ i j, ⟪f i, f j⟫ = 0)) {i : ι} (hi : f i ≠ 0) : + gram_schmidt_orthonormal_basis h f i = (∥f i∥⁻¹ : 𝕜) • f i := +begin + have H : gram_schmidt_normed 𝕜 f i = (∥f i∥⁻¹ : 𝕜) • f i, + { rw [gram_schmidt_normed, gram_schmidt_of_orthogonal 𝕜 hf] }, + rw [gram_schmidt_orthonormal_basis_apply h, H], + simpa [H] using hi, +end + +lemma inner_gram_schmidt_orthonormal_basis_eq_zero {f : ι → E} {i : ι} + (hi : gram_schmidt_normed 𝕜 f i = 0) (j : ι) : + ⟪gram_schmidt_orthonormal_basis h f i, f j⟫ = 0 := +begin + apply inner_right_of_mem_orthogonal_singleton, + suffices : span 𝕜 (gram_schmidt_normed 𝕜 f '' Iic j) + ≤ (𝕜 ∙ gram_schmidt_orthonormal_basis h f i)ᗮ, + { apply this, + rw span_gram_schmidt_normed, + simpa using mem_span_gram_schmidt 𝕜 f (le_refl j) }, + rw span_le, + rintros - ⟨k, -, rfl⟩, + apply mem_orthogonal_singleton_of_inner_left, + by_cases hk : gram_schmidt_normed 𝕜 f k = 0, + { simp [hk] }, + rw ← gram_schmidt_orthonormal_basis_apply h hk, + have : k ≠ i, + { rintros rfl, + exact hk hi }, + exact (gram_schmidt_orthonormal_basis h f).orthonormal.2 this, +end + +lemma gram_schmidt_orthonormal_basis_inv_triangular {i j : ι} (hij : i < j) : + ⟪gram_schmidt_orthonormal_basis h f j, f i⟫ = 0 := +begin + by_cases hi : gram_schmidt_normed 𝕜 f j = 0, + { rw inner_gram_schmidt_orthonormal_basis_eq_zero h hi }, + { simp [gram_schmidt_orthonormal_basis_apply h hi, gram_schmidt_normed, inner_smul_left, + gram_schmidt_inv_triangular 𝕜 f hij] } +end + +lemma gram_schmidt_orthonormal_basis_inv_triangular' {i j : ι} (hij : i < j) : + (gram_schmidt_orthonormal_basis h f).repr (f i) j = 0 := +by simpa [orthonormal_basis.repr_apply_apply] + using gram_schmidt_orthonormal_basis_inv_triangular h f hij + +/-- Given an indexed family `f : ι → E` of vectors in an inner product space `E`, for which the +size of the index set is the dimension of `E`, the matrix of coefficients of `f` with respect to the +orthonormal basis `gram_schmidt_orthonormal_basis` constructed from `f` is upper-triangular. -/ +lemma gram_schmidt_orthonormal_basis_inv_block_triangular : + ((gram_schmidt_orthonormal_basis h f).to_basis.to_matrix f).block_triangular id := +λ i j, gram_schmidt_orthonormal_basis_inv_triangular' h f + +lemma gram_schmidt_orthonormal_basis_det : + (gram_schmidt_orthonormal_basis h f).to_basis.det f = + ∏ i, ⟪gram_schmidt_orthonormal_basis h f i, f i⟫ := +begin + convert matrix.det_of_upper_triangular (gram_schmidt_orthonormal_basis_inv_block_triangular h f), + ext i, + exact ((gram_schmidt_orthonormal_basis h f).repr_apply_apply (f i) i).symm, +end + +end orthonormal_basis From 248bafa20fe59782de114dcda581f8d7d6354645 Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Wed, 5 Oct 2022 21:07:35 +0000 Subject: [PATCH 590/828] feat(algebra/module/projective): projective modules are closed under product (#16735) --- src/algebra/module/projective.lean | 25 +++++++++++++++++++++++++ src/linear_algebra/finsupp.lean | 11 +++++++++++ 2 files changed, 36 insertions(+) diff --git a/src/algebra/module/projective.lean b/src/algebra/module/projective.lean index f35d9b97c37ca..ece899c19dd22 100644 --- a/src/algebra/module/projective.lean +++ b/src/algebra/module/projective.lean @@ -62,6 +62,8 @@ projective module universes u v +open linear_map finsupp + /- The actual implementation we choose: `P` is projective if the natural surjection from the free `R`-module on `P` to `P` splits. -/ /-- An R-module is projective if it is a direct summand of a free module, or equivalently @@ -82,6 +84,10 @@ lemma projective_def : projective R P ↔ (∃ s : P →ₗ[R] (P →₀ R), function.left_inverse (finsupp.total P P R id) s) := ⟨λ h, h.1, λ h, ⟨h⟩⟩ +theorem projective_def' : projective R P ↔ + (∃ s : P →ₗ[R] (P →₀ R), (finsupp.total P P R id) ∘ₗ s = id) := +by simp_rw [projective_def, fun_like.ext_iff, function.left_inverse, coe_comp, id_coe, id.def] + /-- A projective R-module has the property that maps from it lift along surjections. -/ theorem projective_lifting_property [h : projective R P] (f : M →ₗ[R] N) (g : P →ₗ[R] N) (hf : function.surjective f) : ∃ (h : P →ₗ[R] M), f.comp h = g := @@ -105,6 +111,25 @@ begin simp [φ, finsupp.total_apply, function.surj_inv_eq hf], end +variables {Q : Type*} [add_comm_monoid Q] [module R Q] + +instance [hP : projective R P] [hQ : projective R Q] : projective R (P × Q) := +begin + rw module.projective_def', + cases hP.out with sP hsP, + cases hQ.out with sQ hsQ, + use coprod (lmap_domain R R (inl R P Q)) (lmap_domain R R (inr R P Q)) ∘ₗ sP.prod_map sQ, + ext; simp only [coe_inl, coe_inr, coe_comp, function.comp_app, prod_map_apply, map_zero, + coprod_apply, lmap_domain_apply, map_domain_zero, add_zero, zero_add, id_comp, + total_map_domain R _ (prod.mk.inj_right (0 : Q)), + total_map_domain R _ (prod.mk.inj_left (0 : P))], + + { rw [←fst_apply _, apply_total R], exact hsP x, }, + { rw [←snd_apply _, apply_total R], exact finsupp.total_zero_apply _ (sP x), }, + { rw [←fst_apply _, apply_total R], exact finsupp.total_zero_apply _ (sQ x), }, + { rw [←snd_apply _, apply_total R], exact hsQ x, }, +end + end semiring section ring diff --git a/src/linear_algebra/finsupp.lean b/src/linear_algebra/finsupp.lean index 533caf150664d..07fd734dfa757 100644 --- a/src/linear_algebra/finsupp.lean +++ b/src/linear_algebra/finsupp.lean @@ -434,6 +434,17 @@ finset.sum_subset hs $ λ x _ hxg, show l x • v x = 0, by rw [not_mem_support_ finsupp.total α M R v (single a c) = c • (v a) := by simp [total_apply, sum_single_index] +lemma total_zero_apply (x : α →₀ R) : + (finsupp.total α M R 0) x = 0 := by simp [finsupp.total_apply] + +variables (α M) + +@[simp] lemma total_zero : + finsupp.total α M R 0 = 0 := +linear_map.ext (total_zero_apply R) + +variables {α M} + theorem apply_total (f : M →ₗ[R] M') (v) (l : α →₀ R) : f (finsupp.total α M R v l) = finsupp.total α M' R (f ∘ v) l := by apply finsupp.induction_linear l; simp { contextual := tt, } From fe37df4fe5bbb3664a0dc1dc3a8e4cd2907b262c Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 5 Oct 2022 21:07:36 +0000 Subject: [PATCH 591/828] feat(topology/sheaves/skyscraper): stalk functor is adjoint to skyscraper functor (#16790) --- src/topology/sheaves/skyscraper.lean | 148 +++++++++++++++++++++++++++ 1 file changed, 148 insertions(+) diff --git a/src/topology/sheaves/skyscraper.lean b/src/topology/sheaves/skyscraper.lean index 3b8052b8dec06..7d745751dfc97 100644 --- a/src/topology/sheaves/skyscraper.lean +++ b/src/topology/sheaves/skyscraper.lean @@ -235,4 +235,152 @@ def skyscraper_sheaf_functor [has_products.{u} C] : C ⥤ sheaf C X := map_id' := λ c, Sheaf.hom.ext _ _ $ (skyscraper_presheaf_functor p₀).map_id _, map_comp' := λ _ _ _ f g, Sheaf.hom.ext _ _ $ (skyscraper_presheaf_functor p₀).map_comp _ _ } +namespace stalk_skyscraper_presheaf_adjunction_auxs + +variables [has_colimits C] + +/-- +If `f : 𝓕.stalk p₀ ⟶ c`, then a natural transformation `𝓕 ⟶ skyscraper_presheaf p₀ c` can be +defined by: `𝓕.germ p₀ ≫ f : 𝓕(U) ⟶ c` if `p₀ ∈ U` and the unique morphism to a terminal object +if `p₀ ∉ U`. +-/ +@[simps] def to_skyscraper_presheaf {𝓕 : presheaf C X} {c : C} (f : 𝓕.stalk p₀ ⟶ c) : + 𝓕 ⟶ skyscraper_presheaf p₀ c := +{ app := λ U, if h : p₀ ∈ U.unop + then 𝓕.germ ⟨p₀, h⟩ ≫ f ≫ eq_to_hom (if_pos h).symm + else ((if_neg h).symm.rec terminal_is_terminal).from _, + naturality' := λ U V inc, + begin + dsimp, by_cases hV : p₀ ∈ V.unop, + { have hU : p₀ ∈ U.unop := le_of_hom inc.unop hV, split_ifs, + erw [←category.assoc, 𝓕.germ_res inc.unop, category.assoc, category.assoc, eq_to_hom_trans], + refl, }, + { split_ifs, apply ((if_neg hV).symm.rec terminal_is_terminal).hom_ext }, + end } + +/-- +If `f : 𝓕 ⟶ skyscraper_presheaf p₀ c` is a natural transformation, then there is a morphism +`𝓕.stalk p₀ ⟶ c` defined as the morphism from colimit to cocone at `c`. +-/ +def from_stalk {𝓕 : presheaf C X} {c : C} (f : 𝓕 ⟶ skyscraper_presheaf p₀ c) : + 𝓕.stalk p₀ ⟶ c := +let χ : cocone ((open_nhds.inclusion p₀).op ⋙ 𝓕) := cocone.mk c $ +{ app := λ U, f.app (op U.unop.1) ≫ eq_to_hom (if_pos U.unop.2), + naturality' := λ U V inc, + begin + dsimp, erw [category.comp_id, ←category.assoc, comp_eq_to_hom_iff, category.assoc, + eq_to_hom_trans, f.naturality, skyscraper_presheaf_map], + have hV : p₀ ∈ (open_nhds.inclusion p₀).obj V.unop := V.unop.2, split_ifs, + simpa only [comp_eq_to_hom_iff, category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.comp_id], + end } in colimit.desc _ χ + +lemma to_skyscraper_from_stalk {𝓕 : presheaf C X} {c : C} (f : 𝓕 ⟶ skyscraper_presheaf p₀ c) : + to_skyscraper_presheaf p₀ (from_stalk _ f) = f := +nat_trans.ext _ _ $ funext $ λ U, (em (p₀ ∈ U.unop)).elim +(λ h, by { dsimp, split_ifs, erw [←category.assoc, colimit.ι_desc, category.assoc, + eq_to_hom_trans, eq_to_hom_refl, category.comp_id], refl }) $ +λ h, by { dsimp, split_ifs, apply ((if_neg h).symm.rec terminal_is_terminal).hom_ext } + +lemma from_stalk_to_skyscraper {𝓕 : presheaf C X} {c : C} (f : 𝓕.stalk p₀ ⟶ c) : + from_stalk p₀ (to_skyscraper_presheaf _ f) = f := +colimit.hom_ext $ λ U, by { erw [colimit.ι_desc], dsimp, rw dif_pos U.unop.2, rw [category.assoc, + category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id, presheaf.germ], congr' 3, + apply_fun opposite.unop using unop_injective, rw [unop_op], ext, refl } + +/-- +The unit in `presheaf.stalk ⊣ skyscraper_presheaf_functor` +-/ +@[simps] protected def unit : + 𝟭 (presheaf C X) ⟶ presheaf.stalk_functor C p₀ ⋙ skyscraper_presheaf_functor p₀ := +{ app := λ 𝓕, to_skyscraper_presheaf _ $ 𝟙 _, + naturality' := λ 𝓕 𝓖 f, + begin + ext U, dsimp, split_ifs, + { simp only [category.id_comp, ←category.assoc], rw [comp_eq_to_hom_iff], + simp only [category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id], + erw [colimit.ι_map], refl, }, + { apply ((if_neg h).symm.rec terminal_is_terminal).hom_ext, }, + end } + +/-- +The counit in `presheaf.stalk ⊣ skyscraper_presheaf_functor` +-/ +@[simps] protected def counit : + (skyscraper_presheaf_functor p₀ ⋙ (presheaf.stalk_functor C p₀ : presheaf C X ⥤ C)) ⟶ 𝟭 C := +{ app := λ c, (skyscraper_presheaf_stalk_of_specializes p₀ c specializes_rfl).hom, + naturality' := λ x y f, colimit.hom_ext $ λ U, + begin + erw [←category.assoc, colimit.ι_map, colimit.iso_colimit_cocone_ι_hom_assoc, + skyscraper_presheaf_cocone_of_specializes_ι_app, category.assoc, colimit.ι_desc, + whiskering_left_obj_map, whisker_left_app, skyscraper_presheaf_functor.map'_app, + dif_pos U.unop.2, skyscraper_presheaf_cocone_of_specializes_ι_app, comp_eq_to_hom_iff, + category.assoc, eq_to_hom_comp_iff, ←category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.id_comp, comp_eq_to_hom_iff, category.assoc, eq_to_hom_trans, eq_to_hom_refl, + category.comp_id, category_theory.functor.id_map], + end } + +end stalk_skyscraper_presheaf_adjunction_auxs + +section + +open stalk_skyscraper_presheaf_adjunction_auxs + +/-- +`skyscraper_presheaf_functor` is the right adjoint of `presheaf.stalk_functor` +-/ +def skyscraper_presheaf_stalk_adjunction [has_colimits C] : + (presheaf.stalk_functor C p₀ : presheaf C X ⥤ C) ⊣ skyscraper_presheaf_functor p₀ := +{ hom_equiv := λ c 𝓕, + { to_fun := to_skyscraper_presheaf _, + inv_fun := from_stalk _, + left_inv := from_stalk_to_skyscraper _, + right_inv := to_skyscraper_from_stalk _ }, + unit := stalk_skyscraper_presheaf_adjunction_auxs.unit _, + counit := stalk_skyscraper_presheaf_adjunction_auxs.counit _, + hom_equiv_unit' := λ 𝓕 c α, + begin + ext U, simp only [equiv.coe_fn_mk, to_skyscraper_presheaf_app, nat_trans.comp_app, + skyscraper_presheaf_functor.map'_app, skyscraper_presheaf_functor_map, unit_app], split_ifs, + { erw [category.id_comp, ←category.assoc, comp_eq_to_hom_iff, category.assoc, category.assoc, + category.assoc, category.assoc, eq_to_hom_trans, eq_to_hom_refl, category.comp_id, + ←category.assoc _ _ α, eq_to_hom_trans, eq_to_hom_refl, category.id_comp], }, + { apply ((if_neg h).symm.rec terminal_is_terminal).hom_ext } + end, + hom_equiv_counit' := λ 𝓕 c α, + begin + ext U, simp only [equiv.coe_fn_symm_mk, counit_app], + erw [colimit.ι_desc, ←category.assoc, colimit.ι_map, whisker_left_app, category.assoc, + colimit.ι_desc], refl, + end } + +instance [has_colimits C] : is_right_adjoint (skyscraper_presheaf_functor p₀ : C ⥤ presheaf C X) := +⟨_, skyscraper_presheaf_stalk_adjunction _⟩ + +instance [has_colimits C] : is_left_adjoint (presheaf.stalk_functor C p₀) := +⟨_, skyscraper_presheaf_stalk_adjunction _⟩ + +/-- +Taking stalks of a sheaf is the left adjoint functor to `skyscraper_sheaf_functor` +-/ +def stalk_skyscraper_sheaf_adjunction [has_colimits C] [has_products.{u} C] : + sheaf.forget C X ⋙ presheaf.stalk_functor _ p₀ ⊣ skyscraper_sheaf_functor p₀ := +{ hom_equiv := λ 𝓕 c, + ⟨λ f, ⟨to_skyscraper_presheaf p₀ f⟩, λ g, from_stalk p₀ g.1, from_stalk_to_skyscraper p₀, + λ g, by { ext1, apply to_skyscraper_from_stalk }⟩, + unit := + { app := λ 𝓕, ⟨(stalk_skyscraper_presheaf_adjunction_auxs.unit p₀).app 𝓕.1⟩, + naturality' := λ 𝓐 𝓑 ⟨f⟩, + by { ext1, apply (stalk_skyscraper_presheaf_adjunction_auxs.unit p₀).naturality } }, + counit := stalk_skyscraper_presheaf_adjunction_auxs.counit p₀, + hom_equiv_unit' := λ 𝓐 c f, + by { ext1, exact (skyscraper_presheaf_stalk_adjunction p₀).hom_equiv_unit }, + hom_equiv_counit' := λ 𝓐 c f, (skyscraper_presheaf_stalk_adjunction p₀).hom_equiv_counit } + +instance [has_colimits C] [has_products.{u} C] : + is_right_adjoint (skyscraper_sheaf_functor p₀ : C ⥤ sheaf C X) := +⟨_, stalk_skyscraper_sheaf_adjunction _⟩ + +end + end From 224ee5dbb599f878a2b690fc517e9ca3d9b71257 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Wed, 5 Oct 2022 21:07:37 +0000 Subject: [PATCH 592/828] fix(docs/references.bib): re-introduce authors in Cassels-Frohlich (#16813) authors were missing in the reference to Cassels-Frohlich --- docs/references.bib | 1 + 1 file changed, 1 insertion(+) diff --git a/docs/references.bib b/docs/references.bib index a99aa70dbb716..94c9ed1995f68 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -430,6 +430,7 @@ @InProceedings{ carneiro2019 @Book{ cassels1967algebraic, title = {Algebraic number theory}, + author = {Cassels, John William Scott and Fr{\"o}hlich, Albrecht}, booktitle = {Proceedings of an instructional conference organized by the {L}ondon {M}athematical {S}ociety (a {NATO} {A}dvanced {S}tudy {I}nstitute) with the support of the From d1e687517fa8c7bf39ff1771e2e7be17d0494cb0 Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Wed, 5 Oct 2022 21:07:41 +0000 Subject: [PATCH 593/828] chore(algebra/category/Group/injective): remove redundant empty line (#16814) --- src/algebra/category/Group/injective.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/src/algebra/category/Group/injective.lean b/src/algebra/category/Group/injective.lean index aa98c46120ec5..a07d2230fbee7 100644 --- a/src/algebra/category/Group/injective.lean +++ b/src/algebra/category/Group/injective.lean @@ -19,7 +19,6 @@ groups. -/ - open category_theory open_locale pointwise From 02625770565bab3fa4633946fc5aa08d784d3207 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 5 Oct 2022 21:07:42 +0000 Subject: [PATCH 594/828] perf(number_theory/padics/hensel): fix timeout in `calc_eval_z' (#16823) Simply replace the term-mode def by a tactic-mode one appears to fix the timeout observed in multiple branches. Co-authored-by: Junyan Xu --- src/number_theory/padics/hensel.lean | 33 ++++++++++++++-------------- 1 file changed, 16 insertions(+), 17 deletions(-) diff --git a/src/number_theory/padics/hensel.lean b/src/number_theory/padics/hensel.lean index 499605aa26613..2651cfcbe2baa 100644 --- a/src/number_theory/padics/hensel.lean +++ b/src/number_theory/padics/hensel.lean @@ -154,23 +154,22 @@ calc private def calc_eval_z' {z z' z1 : ℤ_[p]} (hz' : z' = z - z1) {n} (hz : ih n z) (h1 : ∥(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)∥ ≤ 1) (hzeq : z1 = ⟨_, h1⟩) : {q : ℤ_[p] // F.eval z' = q * z1^2} := -have hdzne' : (↑(F.derivative.eval z) : ℚ_[p]) ≠ 0, from - have hdzne : F.derivative.eval z ≠ 0, - from mt norm_eq_zero.2 (by rw hz.1; apply deriv_norm_ne_zero; assumption), - λ h, hdzne $ subtype.ext_iff_val.2 h, -let ⟨q, hq⟩ := F.binom_expansion z (-z1) in -have ∥(↑(F.derivative.eval z) * (↑(F.eval z) / ↑(F.derivative.eval z)) : ℚ_[p])∥ ≤ 1, - by { rw padic_norm_e.mul, exact mul_le_one (padic_int.norm_le_one _) (norm_nonneg _) h1 }, -have F.derivative.eval z * (-z1) = -F.eval z, from calc - F.derivative.eval z * (-z1) - = (F.derivative.eval z) * -⟨↑(F.eval z) / ↑(F.derivative.eval z), h1⟩ : by rw [hzeq] -... = -((F.derivative.eval z) * ⟨↑(F.eval z) / ↑(F.derivative.eval z), h1⟩) : mul_neg _ _ -... = -(⟨↑(F.derivative.eval z) * (↑(F.eval z) / ↑(F.derivative.eval z)), this⟩) : - subtype.ext $ by simp only [padic_int.coe_neg, padic_int.coe_mul, subtype.coe_mk] -... = -(F.eval z) : by simp only [mul_div_cancel' _ hdzne', subtype.coe_eta], -have heq : F.eval z' = q * z1^2, by simpa only [sub_eq_add_neg, this, hz', add_right_neg, neg_sq, - zero_add] using hq, -⟨q, heq⟩ +begin + have hdzne : F.derivative.eval z ≠ 0 := + mt norm_eq_zero.2 (by rw hz.1; apply deriv_norm_ne_zero; assumption), + have hdzne' : (↑(F.derivative.eval z) : ℚ_[p]) ≠ 0 := λ h, hdzne (subtype.ext_iff_val.2 h), + obtain ⟨q, hq⟩ := F.binom_expansion z (-z1), + have : ∥(↑(F.derivative.eval z) * (↑(F.eval z) / ↑(F.derivative.eval z)) : ℚ_[p])∥ ≤ 1, + { rw padic_norm_e.mul, exact mul_le_one (padic_int.norm_le_one _) (norm_nonneg _) h1 }, + have : F.derivative.eval z * (-z1) = -F.eval z, + { calc F.derivative.eval z * (-z1) + = (F.derivative.eval z) * -⟨↑(F.eval z) / ↑(F.derivative.eval z), h1⟩ : by rw [hzeq] + ... = -((F.derivative.eval z) * ⟨↑(F.eval z) / ↑(F.derivative.eval z), h1⟩) : mul_neg _ _ + ... = -(⟨↑(F.derivative.eval z) * (↑(F.eval z) / ↑(F.derivative.eval z)), this⟩) : + subtype.ext $ by simp only [padic_int.coe_neg, padic_int.coe_mul, subtype.coe_mk] + ... = -(F.eval z) : by simp only [mul_div_cancel' _ hdzne', subtype.coe_eta] }, + exact ⟨q, by simpa only [sub_eq_add_neg, this, hz', add_right_neg, neg_sq, zero_add] using hq⟩, +end private def calc_eval_z'_norm {z z' z1 : ℤ_[p]} {n} (hz : ih n z) {q} (heq : F.eval z' = q * z1^2) (h1 : ∥(↑(F.eval z) : ℚ_[p]) / ↑(F.derivative.eval z)∥ ≤ 1) From 07cdf30bd0e5335f83f83dd55b1a4f165eaedfd6 Mon Sep 17 00:00:00 2001 From: jakelev Date: Thu, 6 Oct 2022 00:32:16 +0000 Subject: [PATCH 595/828] feat(combinatorics/ssyt): add definition and basic API for semistandard Young tableaux (#16195) Add basic definition and initial API for semistandard Young tableaux (SSYTs). Co-authored-by: Kyle Miller --- .../young_tableaux/semistandard.lean | 114 ++++++++++++++++++ 1 file changed, 114 insertions(+) create mode 100644 src/combinatorics/young_tableaux/semistandard.lean diff --git a/src/combinatorics/young_tableaux/semistandard.lean b/src/combinatorics/young_tableaux/semistandard.lean new file mode 100644 index 0000000000000..e0784ecf00138 --- /dev/null +++ b/src/combinatorics/young_tableaux/semistandard.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2022 Jake Levinson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jake Levinson +-/ +import combinatorics.young_diagram + +/-! +# Semistandard Young tableaux + +A semistandard Young tableau is a filling of a Young diagram by natural numbers, such that +the entries are weakly increasing left-to-right along rows (i.e. for fixed `i`), and +strictly-increasing top-to-bottom along columns (i.e. for fixed `j`). + +An example of an SSYT of shape `μ = [4, 2, 1]` is: + +```text +0 0 0 2 +1 1 +2 +``` + +We represent an SSYT as a function `ℕ → ℕ → ℕ`, which is required to be zero for all pairs +`(i, j) ∉ μ` and to satisfy the row-weak and column-strict conditions on `μ`. + + +## Main definitions + +- `ssyt (μ : young_diagram)` : semistandard Young tableaux of shape `μ`. There is + a `has_coe_to_fun` instance such that `T i j` is value of the `(i, j)` entry of the SSYT `T`. +- `ssyt.highest_weight (μ : young_diagram)`: the semistandard Young tableau whose `i`th row + consists entirely of `i`s, for each `i`. + +## Tags + +Semistandard Young tableau + +## References + + + +-/ + +/-- A semistandard Young tableau (SSYT) is a filling of the cells of a Young diagram by natural +numbers, such that the entries in each row are weakly increasing (left to right), and the entries +in each column are strictly increasing (top to bottom). + +Here, an SSYT is represented as an unrestricted function `ℕ → ℕ → ℕ` that, for reasons +of extensionality, is required to vanish outside `μ`. --/ +structure ssyt (μ : young_diagram) := +(entry : ℕ → ℕ → ℕ) +(row_weak' : ∀ {i j1 j2 : ℕ}, j1 < j2 → (i, j2) ∈ μ → entry i j1 ≤ entry i j2) +(col_strict' : ∀ {i1 i2 j : ℕ}, i1 < i2 → (i2, j) ∈ μ → entry i1 j < entry i2 j) +(zeros' : ∀ {i j}, (i, j) ∉ μ → entry i j = 0) + +namespace ssyt + +instance fun_like {μ : young_diagram} : fun_like (ssyt μ) ℕ (λ _, ℕ → ℕ) := +{ coe := ssyt.entry, + coe_injective' := λ T T' h, by { cases T, cases T', congr' } } + +/-- Helper instance for when there's too many metavariables to apply +`fun_like.has_coe_to_fun` directly. -/ +instance {μ : young_diagram} : has_coe_to_fun (ssyt μ) (λ _, ℕ → ℕ → ℕ) := +fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe {μ : young_diagram} {T : ssyt μ} : T.entry = (T : ℕ → ℕ → ℕ) := rfl + +@[ext] theorem ext {μ : young_diagram} {T T' : ssyt μ} (h : ∀ i j, T i j = T' i j) : T = T' := +fun_like.ext T T' (λ x, by { funext, apply h }) + +/-- Copy of an `ssyt μ` with a new `entry` equal to the old one. Useful to fix definitional +equalities. -/ +protected def copy {μ : young_diagram} {T : ssyt μ} (entry' : ℕ → ℕ → ℕ) (h : entry' = T) : +ssyt μ := +{ entry := entry', + row_weak' := λ _ _ _, h.symm ▸ T.row_weak', + col_strict' := λ _ _ _, h.symm ▸ T.col_strict', + zeros' := λ _ _, h.symm ▸ T.zeros' } + +lemma row_weak {μ : young_diagram} (T : ssyt μ) {i j1 j2 : ℕ} + (hj : j1 < j2) (hcell : (i, j2) ∈ μ) : T i j1 ≤ T i j2 := +T.row_weak' hj hcell + +lemma col_strict {μ : young_diagram} (T : ssyt μ) {i1 i2 j : ℕ} + (hi : i1 < i2) (hcell : (i2, j) ∈ μ) : T i1 j < T i2 j := +T.col_strict' hi hcell + +lemma zeros {μ : young_diagram} (T : ssyt μ) + {i j : ℕ} (not_cell : (i, j) ∉ μ) : T i j = 0 := T.zeros' not_cell + +lemma row_weak_of_le {μ : young_diagram} (T : ssyt μ) {i j1 j2 : ℕ} + (hj : j1 ≤ j2) (cell : (i, j2) ∈ μ) : T i j1 ≤ T i j2 := +by { cases eq_or_lt_of_le hj, subst h, exact T.row_weak h cell } + +lemma col_weak {μ : young_diagram} (T : ssyt μ) {i1 i2 j : ℕ} + (hi : i1 ≤ i2) (cell : (i2, j) ∈ μ) : T i1 j ≤ T i2 j := +by { cases eq_or_lt_of_le hi, subst h, exact le_of_lt (T.col_strict h cell) } + +/-- The "highest weight" SSYT of a given shape is has all i's in row i, for each i. --/ +def highest_weight (μ : young_diagram) : ssyt μ := +{ entry := λ i j, if (i, j) ∈ μ then i else 0, + row_weak' := λ i j1 j2 hj hcell, + by rw [if_pos hcell, if_pos (μ.up_left_mem (by refl) (le_of_lt hj) hcell)], + col_strict' := λ i1 i2 j hi hcell, + by rwa [if_pos hcell, if_pos (μ.up_left_mem (le_of_lt hi) (by refl) hcell)], + zeros' := λ i j not_cell, if_neg not_cell } + +@[simp] lemma highest_weight_apply {μ : young_diagram} {i j : ℕ} : + highest_weight μ i j = if (i, j) ∈ μ then i else 0 := rfl + +instance {μ : young_diagram} : inhabited (ssyt μ) := ⟨ssyt.highest_weight μ⟩ + +end ssyt From 881d122b0c86b2332476cbf5150f2db06dadf9c0 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 6 Oct 2022 01:40:10 +0000 Subject: [PATCH 596/828] feat(algebra/indicator_function): the inverse image under a constant indicator function has four possibilities (#16041) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The possibilities are `set.univ, U, Uᶜ, ∅`, This lemma is from LTE. Co-authored-by: Yury G. Kudryashov --- src/algebra/group/pi.lean | 5 +++++ src/algebra/indicator_function.lean | 28 ++++++++++++++++++++++++++++ 2 files changed, 33 insertions(+) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index d1685c1e91563..64403ce5224ff 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -20,6 +20,11 @@ variable {I : Type u} -- The indexing type variable {f : I → Type v} -- The family of types already equipped with instances variables (x y : Π i, f i) (i : I) +@[to_additive] +lemma set.preimage_one {α β : Type*} [has_one β] (s : set β) [decidable ((1 : β) ∈ s)] : + (1 : α → β) ⁻¹' s = if (1 : β) ∈ s then set.univ else ∅ := +set.preimage_const 1 s + namespace pi @[to_additive] diff --git a/src/algebra/indicator_function.lean b/src/algebra/indicator_function.lean index 1e783b1b83f72..09e8caa86a4d0 100644 --- a/src/algebra/indicator_function.lean +++ b/src/algebra/indicator_function.lean @@ -183,6 +183,34 @@ end (mul_indicator s f)⁻¹' B = s.ite (f ⁻¹' B) (1 ⁻¹' B) := by letI := classical.dec_pred (∈ s); exact piecewise_preimage s f 1 B +@[to_additive] lemma mul_indicator_one_preimage (s : set M) : + t.mul_indicator 1 ⁻¹' s ∈ ({set.univ, ∅} : set (set α)) := +begin + classical, + rw [mul_indicator_one', preimage_one], + split_ifs; simp +end + +@[to_additive] lemma mul_indicator_const_preimage_eq_union (U : set α) (s : set M) (a : M) + [decidable (a ∈ s)] [decidable ((1 : M) ∈ s)] : + U.mul_indicator (λ x, a) ⁻¹' s = (if a ∈ s then U else ∅) ∪ (if (1 : M) ∈ s then Uᶜ else ∅) := +begin + rw [mul_indicator_preimage, preimage_one, preimage_const], + split_ifs; simp [← compl_eq_univ_diff] +end + +@[to_additive] lemma mul_indicator_const_preimage (U : set α) (s : set M) (a : M) : + U.mul_indicator (λ x, a) ⁻¹' s ∈ ({set.univ, U, Uᶜ, ∅} : set (set α)) := +begin + classical, + rw [mul_indicator_const_preimage_eq_union], + split_ifs; simp +end + +lemma indicator_one_preimage [has_zero M] (U : set α) (s : set M) : + U.indicator 1 ⁻¹' s ∈ ({set.univ, U, Uᶜ, ∅} : set (set α)) := +indicator_const_preimage _ _ 1 + @[to_additive] lemma mul_indicator_preimage_of_not_mem (s : set α) (f : α → M) {t : set M} (ht : (1:M) ∉ t) : (mul_indicator s f)⁻¹' t = f ⁻¹' t ∩ s := From a4dddb594422128756d8f79fb92f8cdecbbf980d Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 6 Oct 2022 04:39:25 +0000 Subject: [PATCH 597/828] chore(topology/algebra/module/weak_dual): golf `weak_dual.t2_space` (#16816) --- src/topology/algebra/module/weak_dual.lean | 8 ++------ 1 file changed, 2 insertions(+), 6 deletions(-) diff --git a/src/topology/algebra/module/weak_dual.lean b/src/topology/algebra/module/weak_dual.lean index e942a19eef6b2..b69a3329ea8c2 100644 --- a/src/topology/algebra/module/weak_dual.lean +++ b/src/topology/algebra/module/weak_dual.lean @@ -239,12 +239,8 @@ lemma continuous_of_continuous_eval [topological_space α] {g : α → weak_dual continuous_induced_rng.2 (continuous_pi_iff.mpr h) instance [t2_space 𝕜] : t2_space (weak_dual 𝕜 E) := -begin - refine t2_iff_is_closed_diagonal.mpr _, - simp_rw [set.diagonal, fun_like.ext_iff, set.set_of_forall], - exact is_closed_Inter (λ x, is_closed_eq ((weak_dual.eval_continuous x).comp continuous_fst) $ - (weak_dual.eval_continuous x).comp continuous_snd), -end +embedding.t2_space $ weak_bilin.embedding $ + show function.injective (top_dual_pairing 𝕜 E), from continuous_linear_map.coe_injective end weak_dual From 78dddf53c1be72bdef02e4d2553b0b940497fe43 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Thu, 6 Oct 2022 05:52:08 +0000 Subject: [PATCH 598/828] =?UTF-8?q?feat(algebraic=5Ftopology/dold=5Fkan):?= =?UTF-8?q?=20N=E2=82=81=20reflects=20isomorphisms=20(#16778)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this PR, it is shown that `N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ)` reflects isomorphisms for any preadditive category `C`. --- .../dold_kan/n_reflects_iso.lean | 74 +++++++++++++++++++ .../idempotents/homological_complex.lean | 52 +++++++++++++ 2 files changed, 126 insertions(+) create mode 100644 src/algebraic_topology/dold_kan/n_reflects_iso.lean create mode 100644 src/category_theory/idempotents/homological_complex.lean diff --git a/src/algebraic_topology/dold_kan/n_reflects_iso.lean b/src/algebraic_topology/dold_kan/n_reflects_iso.lean new file mode 100644 index 0000000000000..48fc279d6716e --- /dev/null +++ b/src/algebraic_topology/dold_kan/n_reflects_iso.lean @@ -0,0 +1,74 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.dold_kan.functor_n +import algebraic_topology.dold_kan.decomposition +import category_theory.idempotents.homological_complex + +/-! + +# N₁ and N₂ reflects isomorphisms + +In this file, it is shown that the functor +`N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ)` +reflects isomorphisms for any preadditive category `C`. + +TODO @joelriou: deduce that `N₂ : karoubi (simplicial_object C) ⥤ karoubi (chain_complex C ℕ)` +also reflects isomorphisms. + +-/ + +open category_theory +open category_theory.category +open category_theory.idempotents +open opposite +open_locale simplicial + +namespace algebraic_topology + +namespace dold_kan + +variables {C : Type*} [category C] [preadditive C] + +open morph_components + +instance : reflects_isomorphisms (N₁ : simplicial_object C ⥤ karoubi (chain_complex C ℕ)) := +⟨λ X Y f, begin + introI, + /- restating the result in a way that allows induction on the degree n -/ + suffices : ∀ (n : ℕ), is_iso (f.app (op [n])), + { haveI : ∀ (Δ : simplex_categoryᵒᵖ), is_iso (f.app Δ) := λ Δ, this Δ.unop.len, + apply nat_iso.is_iso_of_is_iso_app, }, + /- restating the assumption in a more practical form -/ + have h₁ := homological_complex.congr_hom (karoubi.hom_ext.mp (is_iso.hom_inv_id (N₁.map f))), + have h₂ := homological_complex.congr_hom (karoubi.hom_ext.mp (is_iso.inv_hom_id (N₁.map f))), + have h₃ := λ n, karoubi.homological_complex.p_comm_f_assoc (inv (N₁.map f)) (n) (f.app (op [n])), + simp only [N₁_map_f, karoubi.comp, homological_complex.comp_f, + alternating_face_map_complex.map_f, N₁_obj_p, karoubi.id_eq, assoc] at h₁ h₂ h₃, + /- we have to construct an inverse to f in degree n, by induction on n -/ + intro n, + induction n with n hn, + /- degree 0 -/ + { use (inv (N₁.map f)).f.f 0, + have h₁₀ := h₁ 0, + have h₂₀ := h₂ 0, + dsimp at h₁₀ h₂₀, + simp only [id_comp, comp_id] at h₁₀ h₂₀, + tauto, }, + /- induction step -/ + { haveI := hn, + use φ + { a := P_infty.f (n+1) ≫ (inv (N₁.map f)).f.f (n+1), + b := λ i, inv (f.app (op [n])) ≫ X.σ i, }, + simp only [morph_components.id, ← id_φ, ← pre_comp_φ, pre_comp, ← post_comp_φ, + post_comp, P_infty_f_naturality_assoc, is_iso.hom_inv_id_assoc, assoc, + is_iso.inv_hom_id_assoc, simplicial_object.σ_naturality, h₁, h₂, h₃], + tauto, }, +end⟩ + +end dold_kan + +end algebraic_topology diff --git a/src/category_theory/idempotents/homological_complex.lean b/src/category_theory/idempotents/homological_complex.lean new file mode 100644 index 0000000000000..9f1409bbc5c68 --- /dev/null +++ b/src/category_theory/idempotents/homological_complex.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebra.homology.homological_complex +import category_theory.idempotents.karoubi + +/-! +# Idempotent completeness and homological complexes + +This file contains simplifications lemmas for categories +`karoubi (homological_complex C c)`. + +TODO @joelriou : Get an equivalence of categories +`karoubi (homological_complex C c) ≌ homological_complex (karoubi C) c` +for all preadditive categories `C` and complex shape `c`. + +-/ + +namespace category_theory + +variables {C : Type*} [category C] [preadditive C] {ι : Type*} {c : complex_shape ι} + +namespace idempotents + +namespace karoubi + +namespace homological_complex + +variables {P Q : karoubi (homological_complex C c)} (f : P ⟶ Q) (n : ι) + +@[simp, reassoc] +lemma p_comp_d : P.p.f n ≫ f.f.f n = f.f.f n := +homological_complex.congr_hom (p_comp f) n + +@[simp, reassoc] +lemma comp_p_d : f.f.f n ≫ Q.p.f n = f.f.f n := +homological_complex.congr_hom (comp_p f) n + +@[reassoc] +lemma p_comm_f : P.p.f n ≫ f.f.f n = f.f.f n ≫ Q.p.f n := +homological_complex.congr_hom (p_comm f) n + +end homological_complex + +end karoubi + +end idempotents + +end category_theory From 54619e055bd8c9395572508c20fb2640f15b86dc Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 6 Oct 2022 07:37:58 +0000 Subject: [PATCH 599/828] feat(algebra/{hom,ring}): extra coercion lemmas for `{mul,add,ring}_equiv` (#16725) This PR adds more lemmas for the coercion of `refl` and `trans` of `{mul,add,ring}_equiv` to other types of maps. In particular, it ensures these types come with: * `coe_{type}_refl` and `coe_{type}_trans` where `type` ranges over the types of bundled maps that the equivs inherit from * `self_trans_symm` and `symm_trans_self` * `coe_trans` Of course, it would be great if we figured out some generic way of stating all these results so we wouldn't have to go through and add all these lemmas. Co-authored-by: Anne Baanen --- src/algebra/hom/equiv.lean | 16 ++++++++++++++++ src/algebra/ring/equiv.lean | 28 +++++++++++++++++++++++++++- 2 files changed, 43 insertions(+), 1 deletion(-) diff --git a/src/algebra/hom/equiv.lean b/src/algebra/hom/equiv.lean index 4664e0d76730d..d97a8cacf05eb 100644 --- a/src/algebra/hom/equiv.lean +++ b/src/algebra/hom/equiv.lean @@ -297,6 +297,22 @@ e.to_equiv.eq_symm_apply @[to_additive] lemma symm_comp_eq {α : Type*} (e : M ≃* N) (f : α → M) (g : α → N) : e.symm ∘ g = f ↔ g = e ∘ f := e.to_equiv.symm_comp_eq f g +@[simp, to_additive] +theorem symm_trans_self (e : M ≃* N) : e.symm.trans e = refl N := +fun_like.ext _ _ e.apply_symm_apply + +@[simp, to_additive] +theorem self_trans_symm (e : M ≃* N) : e.trans e.symm = refl M := +fun_like.ext _ _ e.symm_apply_apply + +@[to_additive, simp] lemma coe_monoid_hom_refl {M} [mul_one_class M] : + (refl M : M →* M) = monoid_hom.id M := rfl + +@[to_additive, simp] lemma coe_monoid_hom_trans {M N P} + [mul_one_class M] [mul_one_class N] [mul_one_class P] (e₁ : M ≃* N) (e₂ : N ≃* P) : + (e₁.trans e₂ : M →* P) = (e₂ : N →* P).comp ↑e₁ := +rfl + /-- Two multiplicative isomorphisms agree if they are defined by the same underlying function. -/ @[ext, to_additive diff --git a/src/algebra/ring/equiv.lean b/src/algebra/ring/equiv.lean index 207189ee280a7..cb682e18ebac1 100644 --- a/src/algebra/ring/equiv.lean +++ b/src/algebra/ring/equiv.lean @@ -207,13 +207,19 @@ symm_bijective.injective $ ext $ λ x, rfl @[trans] protected def trans (e₁ : R ≃+* S) (e₂ : S ≃+* S') : R ≃+* S' := { .. (e₁.to_mul_equiv.trans e₂.to_mul_equiv), .. (e₁.to_add_equiv.trans e₂.to_add_equiv) } -@[simp] lemma trans_apply (e₁ : R ≃+* S) (e₂ : S ≃+* S') (a : R) : +lemma trans_apply (e₁ : R ≃+* S) (e₂ : S ≃+* S') (a : R) : e₁.trans e₂ a = e₂ (e₁ a) := rfl +@[simp] lemma coe_trans (e₁ : R ≃+* S) (e₂ : S ≃+* S') : + (e₁.trans e₂ : R → S') = e₂ ∘ e₁ := rfl + @[simp] lemma symm_trans_apply (e₁ : R ≃+* S) (e₂ : S ≃+* S') (a : S') : (e₁.trans e₂).symm a = e₁.symm (e₂.symm a) := rfl +lemma symm_trans (e₁ : R ≃+* S) (e₂ : S ≃+* S') : + (e₁.trans e₂).symm = e₂.symm.trans (e₁.symm) := rfl + protected lemma bijective (e : R ≃+* S) : function.bijective e := equiv_like.bijective e protected lemma injective (e : R ≃+* S) : function.injective e := equiv_like.injective e protected lemma surjective (e : R ≃+* S) : function.surjective e := equiv_like.surjective e @@ -224,6 +230,11 @@ protected lemma surjective (e : R ≃+* S) : function.surjective e := equiv_like lemma image_eq_preimage (e : R ≃+* S) (s : set R) : e '' s = e.symm ⁻¹' s := e.to_equiv.image_eq_preimage s +@[simp] lemma coe_mul_equiv_trans (e₁ : R ≃+* S) (e₂ : S ≃+* S') : + (e₁.trans e₂ : R ≃* S') = (e₁ : R ≃* S).trans ↑e₂:= rfl +@[simp] lemma coe_add_equiv_trans (e₁ : R ≃+* S) (e₂ : S ≃+* S') : + (e₁.trans e₂ : R ≃+ S') = (e₁ : R ≃+ S).trans ↑e₂:= rfl + end basic section opposite @@ -335,6 +346,21 @@ protected lemma map_eq_one_iff : f x = 1 ↔ x = 1 := mul_equiv_class.map_eq_one lemma map_ne_one_iff : f x ≠ 1 ↔ x ≠ 1 := mul_equiv_class.map_ne_one_iff f +lemma coe_monoid_hom_refl : (ring_equiv.refl R : R →* R) = monoid_hom.id R := rfl +@[simp] lemma coe_add_monoid_hom_refl : (ring_equiv.refl R : R →+ R) = add_monoid_hom.id R := rfl +/-! `ring_equiv.coe_mul_equiv_refl` and `ring_equiv.coe_add_equiv_refl` are proved above +in higher generality -/ +@[simp] lemma coe_ring_hom_refl : (ring_equiv.refl R : R →* R) = ring_hom.id R := rfl + +@[simp] lemma coe_monoid_hom_trans [non_assoc_semiring S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') : + (e₁.trans e₂ : R →* S') = (e₂ : S →* S').comp ↑e₁ := rfl +@[simp] lemma coe_add_monoid_hom_trans [non_assoc_semiring S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') : + (e₁.trans e₂ : R →+ S') = (e₂ : S →+ S').comp ↑e₁ := rfl +/-! `ring_equiv.coe_mul_equiv_trans` and `ring_equiv.coe_add_equiv_trans` are proved above +in higher generality -/ +@[simp] lemma coe_ring_hom_trans [non_assoc_semiring S'] (e₁ : R ≃+* S) (e₂ : S ≃+* S') : + (e₁.trans e₂ : R →+* S') = (e₂ : S →+* S').comp ↑e₁ := rfl + end semiring section non_unital_ring From 28e5032b3adc497848c263a93059884339d7bfb4 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 6 Oct 2022 10:46:52 +0000 Subject: [PATCH 600/828] =?UTF-8?q?feat(ring=5Ftheory/derivation):=20Endom?= =?UTF-8?q?orphisms=20of=20the=20K=C3=A4hler=20differential=20module=20(#1?= =?UTF-8?q?5854)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/algebra/module/submodule/basic.lean | 6 ++ .../group_action/sub_mul_action.lean | 5 + src/ring_theory/derivation.lean | 91 ++++++++++++++++++- src/ring_theory/ideal/cotangent.lean | 8 ++ src/ring_theory/ideal/operations.lean | 6 ++ src/ring_theory/ideal/quotient.lean | 2 + 6 files changed, 116 insertions(+), 2 deletions(-) diff --git a/src/algebra/module/submodule/basic.lean b/src/algebra/module/submodule/basic.lean index 7c376204bdf8f..ccdd822920f23 100644 --- a/src/algebra/module/submodule/basic.lean +++ b/src/algebra/module/submodule/basic.lean @@ -167,6 +167,12 @@ instance [has_smul S R] [has_smul S M] [is_scalar_tower S R M] : instance [has_smul S R] [has_smul S M] [is_scalar_tower S R M] : is_scalar_tower S R p := p.to_sub_mul_action.is_scalar_tower +instance is_scalar_tower' {S' : Type*} + [has_smul S R] [has_smul S M] [has_smul S' R] [has_smul S' M] [has_smul S S'] + [is_scalar_tower S' R M] [is_scalar_tower S S' M] [is_scalar_tower S R M] : + is_scalar_tower S S' p := +p.to_sub_mul_action.is_scalar_tower' + instance [has_smul S R] [has_smul S M] [is_scalar_tower S R M] [has_smul Sᵐᵒᵖ R] [has_smul Sᵐᵒᵖ M] [is_scalar_tower Sᵐᵒᵖ R M] diff --git a/src/group_theory/group_action/sub_mul_action.lean b/src/group_theory/group_action/sub_mul_action.lean index 5d03221e6e8cf..ae2bc2753c281 100644 --- a/src/group_theory/group_action/sub_mul_action.lean +++ b/src/group_theory/group_action/sub_mul_action.lean @@ -117,6 +117,11 @@ instance has_smul' : has_smul S p := instance : is_scalar_tower S R p := { smul_assoc := λ s r x, subtype.ext $ smul_assoc s r ↑x } +instance is_scalar_tower' {S' : Type*} [has_smul S' R] [has_smul S' S] + [has_smul S' M] [is_scalar_tower S' R M] [is_scalar_tower S' S M] : + is_scalar_tower S' S p := +{ smul_assoc := λ s r x, subtype.ext $ smul_assoc s r ↑x } + @[simp, norm_cast] lemma coe_smul_of_tower (s : S) (x : p) : ((s • x : p) : M) = s • ↑x := rfl @[simp] lemma smul_mem_iff' {G} [group G] [has_smul G R] [mul_action G M] diff --git a/src/ring_theory/derivation.lean b/src/ring_theory/derivation.lean index 903b95af667f5..94400d2a334f1 100644 --- a/src/ring_theory/derivation.lean +++ b/src/ring_theory/derivation.lean @@ -6,6 +6,7 @@ Authors: Nicolò Cavalleri, Andrew Yang import ring_theory.adjoin.basic import algebra.lie.of_associative +import ring_theory.ideal.cotangent import ring_theory.tensor_product import ring_theory.ideal.cotangent @@ -211,7 +212,7 @@ section push_forward variables {N : Type*} [add_comm_monoid N] [module A N] [module R N] [is_scalar_tower R A M] [is_scalar_tower R A N] -variables (f : M →ₗ[A] N) +variables (f : M →ₗ[A] N) (e : M ≃ₗ[A] N) /-- We can push forward derivations using linear maps, i.e., the composition of a derivation with a linear map is a derivation. Furthermore, this operation is linear on the spaces of derivations. -/ @@ -233,11 +234,20 @@ rfl rfl /-- The composition of a derivation with a linear map as a bilinear map -/ +@[simps] def llcomp : (M →ₗ[A] N) →ₗ[A] derivation R A M →ₗ[R] derivation R A N := { to_fun := λ f, f.comp_der, map_add' := λ f₁ f₂, by { ext, refl }, map_smul' := λ r D, by { ext, refl } } +/-- Pushing a derivation foward through a linear equivalence is an equivalence. -/ +def _root_.linear_equiv.comp_der : derivation R A M ≃ₗ[R] derivation R A N := +{ inv_fun := e.symm.to_linear_map.comp_der, + left_inv := λ D, by { ext a, exact e.symm_apply_apply (D a) }, + right_inv := λ D, by { ext a, exact e.apply_symm_apply (D a) }, + ..e.to_linear_map.comp_der } + + end push_forward end @@ -356,7 +366,7 @@ section to_square_zero universes u v w -variables {R : Type u} {A : Type u} {B : Type w} [comm_semiring R] [comm_semiring A] [comm_ring B] +variables {R : Type u} {A : Type v} {B : Type w} [comm_semiring R] [comm_semiring A] [comm_ring B] variables [algebra R A] [algebra R B] (I : ideal B) (hI : I ^ 2 = ⊥) /-- If `f₁ f₂ : A →ₐ[R] B` are two lifts of the same `A →ₐ[R] B ⧸ I`, @@ -721,4 +731,81 @@ def kaehler_differential.linear_map_equiv_derivation : (Ω[S⁄R] →ₗ[S] M) right_inv := derivation.lift_kaehler_differential_comp, ..(derivation.llcomp.flip $ kaehler_differential.D R S) } +/-- The quotient ring of `S ⊗ S ⧸ J ^ 2` by `Ω[S⁄R]` is isomorphic to `S`. -/ +def kaehler_differential.quotient_cotangent_ideal_ring_equiv : + (S ⊗ S ⧸ kaehler_differential.ideal R S ^ 2) ⧸ + (kaehler_differential.ideal R S).cotangent_ideal ≃+* S := +begin + have : function.right_inverse tensor_product.include_left + (↑(tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S) : S ⊗[R] S →+* S), + { intro x, rw [alg_hom.coe_to_ring_hom, ← alg_hom.comp_apply, + tensor_product.lmul'_comp_include_left], refl }, + refine (ideal.quot_cotangent _).trans _, + refine (ideal.quot_equiv_of_eq _).trans (ring_hom.quotient_ker_equiv_of_right_inverse this), + ext, refl, +end + +/-- The quotient ring of `S ⊗ S ⧸ J ^ 2` by `Ω[S⁄R]` is isomorphic to `S` as an `S`-algebra. -/ +def kaehler_differential.quotient_cotangent_ideal : + ((S ⊗ S ⧸ kaehler_differential.ideal R S ^ 2) ⧸ + (kaehler_differential.ideal R S).cotangent_ideal) ≃ₐ[S] S := +{ commutes' := (kaehler_differential.quotient_cotangent_ideal_ring_equiv R S).apply_symm_apply, + ..kaehler_differential.quotient_cotangent_ideal_ring_equiv R S } + +lemma kaehler_differential.End_equiv_aux (f : S →ₐ[R] S ⊗ S ⧸ kaehler_differential.ideal R S ^ 2) : + (ideal.quotient.mkₐ R (kaehler_differential.ideal R S).cotangent_ideal).comp f = + is_scalar_tower.to_alg_hom R S _ ↔ + (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift.comp f = alg_hom.id R S := +begin + rw [alg_hom.ext_iff, alg_hom.ext_iff], + apply forall_congr, + intro x, + have e₁ : (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift (f x) = + kaehler_differential.quotient_cotangent_ideal_ring_equiv R S + (ideal.quotient.mk (kaehler_differential.ideal R S).cotangent_ideal $ f x), + { generalize : f x = y, obtain ⟨y, rfl⟩ := ideal.quotient.mk_surjective y, refl }, + have e₂ : x = kaehler_differential.quotient_cotangent_ideal_ring_equiv + R S (is_scalar_tower.to_alg_hom R S _ x), + { exact ((tensor_product.lmul'_apply_tmul x 1).trans (mul_one x)).symm }, + split, + { intro e, + exact (e₁.trans (@ring_equiv.congr_arg _ _ _ _ _ _ + (kaehler_differential.quotient_cotangent_ideal_ring_equiv R S) _ _ e)).trans e₂.symm }, + { intro e, apply (kaehler_differential.quotient_cotangent_ideal_ring_equiv R S).injective, + exact e₁.symm.trans (e.trans e₂) } +end + +/-- Derivations into `Ω[S⁄R]` is equivalent to derivations +into `(kaehler_differential.ideal R S).cotangent_ideal` -/ +-- This has type +-- `derivation R S Ω[ S / R ] ≃ₗ[R] derivation R S (kaehler_differential.ideal R S).cotangent_ideal` +-- But lean times-out if this is given explicitly. +noncomputable +def kaehler_differential.End_equiv_derivation' := +@linear_equiv.comp_der R _ _ _ _ Ω[S⁄R] _ _ _ _ _ _ _ _ _ + ((kaehler_differential.ideal R S).cotangent_equiv_ideal.restrict_scalars S) + +/-- (Implementation) An `equiv` version of `kaehler_differential.End_equiv_aux`. +Used in `kaehler_differential.End_equiv`. -/ +def kaehler_differential.End_equiv_aux_equiv : + {f // (ideal.quotient.mkₐ R (kaehler_differential.ideal R S).cotangent_ideal).comp f = + is_scalar_tower.to_alg_hom R S _ } ≃ + { f // (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift.comp f = alg_hom.id R S } := +(equiv.refl _).subtype_equiv (kaehler_differential.End_equiv_aux R S) + +/-- +The endomorphisms of `Ω[S⁄R]` corresponds to sections of the surjection `S ⊗[R] S ⧸ J ^ 2 →ₐ[R] S`, +with `J` being the kernel of the multiplication map `S ⊗[R] S →ₐ[R] S`. +-/ +noncomputable +def kaehler_differential.End_equiv : + module.End S Ω[S⁄R] ≃ + { f // (tensor_product.lmul' R : S ⊗[R] S →ₐ[R] S).ker_square_lift.comp f = alg_hom.id R S } := +(kaehler_differential.linear_map_equiv_derivation R S).to_equiv.trans $ + (kaehler_differential.End_equiv_derivation' R S).to_equiv.trans $ + (derivation_to_square_zero_equiv_lift + (kaehler_differential.ideal R S).cotangent_ideal + (kaehler_differential.ideal R S).cotangent_ideal_square).trans $ + kaehler_differential.End_equiv_aux_equiv R S + end kaehler_differential diff --git a/src/ring_theory/ideal/cotangent.lean b/src/ring_theory/ideal/cotangent.lean index d54cc285765d2..4b4efd7ded916 100644 --- a/src/ring_theory/ideal/cotangent.lean +++ b/src/ring_theory/ideal/cotangent.lean @@ -179,6 +179,14 @@ begin { rintros _ ⟨x, hx, rfl⟩, exact hx } end +/-- The quotient ring of `I ⧸ I ^ 2` is `R ⧸ I`. -/ +def quot_cotangent : ((R ⧸ I ^ 2) ⧸ I.cotangent_ideal) ≃+* R ⧸ I := +begin + refine (ideal.quot_equiv_of_eq (ideal.map_eq_submodule_map _ _).symm).trans _, + refine (double_quot.quot_quot_equiv_quot_sup _ _).trans _, + exact (ideal.quot_equiv_of_eq (sup_eq_right.mpr $ ideal.pow_le_self two_ne_zero)), +end + end ideal namespace local_ring diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index bc0a4c1cce383..5a111557b1ab6 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -1206,6 +1206,12 @@ lemma mem_map_iff_of_surjective {I : ideal R} {y} : lemma le_map_of_comap_le_of_surjective : comap f K ≤ I → K ≤ map f I := λ h, (map_comap_of_surjective f hf K) ▸ map_mono h +omit hf + +lemma map_eq_submodule_map (f : R →+* S) [h : ring_hom_surjective f] (I : ideal R) : + I.map f = submodule.map f.to_semilinear_map I := +submodule.ext (λ x, mem_map_iff_of_surjective f h.1) + end surjective section injective diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index 41952cd53bd88..8bff80f2e77d5 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -121,6 +121,8 @@ instance : unique (R ⧸ (⊤ : ideal R)) := lemma mk_surjective : function.surjective (mk I) := λ y, quotient.induction_on' y (λ x, exists.intro x rfl) +instance : ring_hom_surjective (mk I) := ⟨mk_surjective⟩ + /-- If `I` is an ideal of a commutative ring `R`, if `q : R → R/I` is the quotient map, and if `s ⊆ R` is a subset, then `q⁻¹(q(s)) = ⋃ᵢ(i + s)`, the union running over all `i ∈ I`. -/ lemma quotient_ring_saturate (I : ideal R) (s : set R) : From a87fad7b1594d57b7c879afadd9f081063431984 Mon Sep 17 00:00:00 2001 From: damiano Date: Thu, 6 Oct 2022 10:46:53 +0000 Subject: [PATCH 601/828] feat(counterexamples/zero_divisors_in_add_monoid_algebras + data/finsupp/lex): add a counterexample (#16236) This PR describes a counterexample to a possible weakening of lemmas `finsupp.lex.covariant_class_le_left` and its analogue `finsupp.lex.covariant_class_le_right`: assuming only monotonicity of addition, instead of strict monotonicity would not be strong enough to prove monotonicity of addition for `finsupp`s. See also [this Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/144837-PR-reviews/topic/.2316157.20no-zero-divisors.20in.20add_monoid_algebras) for related material. --- .../zero_divisors_in_add_monoid_algebras.lean | 211 ++++++++++++++++++ src/data/finsupp/lex.lean | 4 +- 2 files changed, 214 insertions(+), 1 deletion(-) create mode 100644 counterexamples/zero_divisors_in_add_monoid_algebras.lean diff --git a/counterexamples/zero_divisors_in_add_monoid_algebras.lean b/counterexamples/zero_divisors_in_add_monoid_algebras.lean new file mode 100644 index 0000000000000..e828dbfb07c58 --- /dev/null +++ b/counterexamples/zero_divisors_in_add_monoid_algebras.lean @@ -0,0 +1,211 @@ +/- +Copyright (c) 2022 Damiano Testa. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Damiano Testa +-/ +import algebra.char_p.basic +import algebra.geom_sum +import algebra.monoid_algebra.basic +import data.finsupp.lex +import data.zmod.basic +import group_theory.order_of_element + +/-! +# Examples of zero-divisors in `add_monoid_algebra`s + +This file contains an easy source of zero-divisors in an `add_monoid_algebra`. +If `k` is a field and `G` is an additive group containing a non-zero torsion element, then +`add_monoid_algebra k G` contains non-zero zero-divisors: this is lemma `zero_divisors_of_torsion`. + +There is also a version for periodic elements of an additive monoid: `zero_divisors_of_periodic`. + +The converse of this statement is +[Kaplansky's zero divisor conjecture](https://en.wikipedia.org/wiki/Kaplansky%27s_conjectures). + +The formalized example generalizes in trivial ways the assumptions: the field `k` can be any +nontrivial ring `R` and the additive group `G` with a torsion element can be any additive monoid +`A` with a non-zero periodic element. + +Besides this example, we also address a comment in `data.finsupp.lex` to the effect that the proof +that addition is monotone on `α →₀ N` uses that it is *strictly* monotone on `N`. + +The specific statement is about `finsupp.lex.covariant_class_le_left` and its analogue +`finsupp.lex.covariant_class_le_right`. We do not need two separate counterexamples, since the +operation is commutative. + +The example is very simple. Let `F = {0, 1}` with order determined by `0 < 1` and absorbing +addition (which is the same as `max` in this case). We denote a function `f : F → F` (which is +automatically finitely supported!) by `[f 0, f 1]`, listing its values. Recall that the order on +finitely supported function is lexicographic, matching the list notation. The inequality +`[0, 1] ≤ [1, 0]` holds. However, adding `[1, 0]` to both sides yields the *reversed* inequality +`[1, 1] > [1, 0]`. +-/ +open finsupp add_monoid_algebra + +/-- This is a simple example showing that if `R` is a non-trivial ring and `A` is an additive +monoid with an element `a` satisfying `n • a = a` and `(n - 1) • a ≠ a`, for some `2 ≤ n`, +then `add_monoid_algebra R A` contains non-zero zero-divisors. The elements are easy to write down: +`[a]` and `[a] ^ (n - 1) - 1` are non-zero elements of `add_monoid_algebra R A` whose product +is zero. + +Observe that such an element `a` *cannot* be invertible. In particular, this lemma never applies +if `A` is a group. -/ +lemma zero_divisors_of_periodic {R A} [nontrivial R] [ring R] [add_monoid A] {n : ℕ} {a : A} + (n2 : 2 ≤ n) (na : n • a = a) (na1 : (n - 1) • a ≠ 0) : + ∃ f g : add_monoid_algebra R A, f ≠ 0 ∧ g ≠ 0 ∧ f * g = 0 := +begin + refine ⟨single a 1, single ((n - 1) • a) 1 - single 0 1, by simp, _, _⟩, + { exact sub_ne_zero.mpr (by simpa [single_eq_single_iff]) }, + { rw [mul_sub, add_monoid_algebra.single_mul_single, add_monoid_algebra.single_mul_single, + sub_eq_zero, add_zero, ← succ_nsmul, nat.sub_add_cancel (one_le_two.trans n2), na] }, +end + +lemma single_zero_one {R A} [semiring R] [has_zero A] : + single (0 : A) (1 : R) = (1 : add_monoid_algebra R A) := rfl + +/-- This is a simple example showing that if `R` is a non-trivial ring and `A` is an additive +monoid with a non-zero element `a` of finite order `oa`, then `add_monoid_algebra R A` contains +non-zero zero-divisors. The elements are easy to write down: +`∑ i in finset.range oa, [a] ^ i` and `[a] - 1` are non-zero elements of `add_monoid_algebra R A` +whose product is zero. + +In particular, this applies whenever the additive monoid `A` is an additive group with a non-zero +torsion element. -/ +lemma zero_divisors_of_torsion {R A} [nontrivial R] [ring R] [add_monoid A] (a : A) + (o2 : 2 ≤ add_order_of a) : + ∃ f g : add_monoid_algebra R A, f ≠ 0 ∧ g ≠ 0 ∧ f * g = 0 := +begin + refine ⟨(finset.range (add_order_of a)).sum (λ (i : ℕ), (single a 1) ^ i), + single a 1 - single 0 1, _, _, _⟩, + { apply_fun (λ x : add_monoid_algebra R A, x 0), + refine ne_of_eq_of_ne (_ : (_ : R) = 1) one_ne_zero, + simp_rw finset.sum_apply', + refine (finset.sum_eq_single 0 _ _).trans _, + { intros b hb b0, + rw [single_pow, one_pow, single_eq_of_ne], + exact nsmul_ne_zero_of_lt_add_order_of' b0 (finset.mem_range.mp hb) }, + { simp only [(zero_lt_two.trans_le o2).ne', finset.mem_range, not_lt, le_zero_iff, + false_implies_iff] }, + { rw [single_pow, one_pow, zero_smul, single_eq_same] } }, + { apply_fun (λ x : add_monoid_algebra R A, x 0), + refine sub_ne_zero.mpr (ne_of_eq_of_ne (_ : (_ : R) = 0) _), + { have a0 : a ≠ 0 := ne_of_eq_of_ne (one_nsmul a).symm + (nsmul_ne_zero_of_lt_add_order_of' one_ne_zero (nat.succ_le_iff.mp o2)), + simp only [a0, single_eq_of_ne, ne.def, not_false_iff] }, + { simpa only [single_eq_same] using zero_ne_one, } }, + { convert commute.geom_sum₂_mul _ (add_order_of a), + { ext, rw [single_zero_one, one_pow, mul_one] }, + { rw [single_pow, one_pow, add_order_of_nsmul_eq_zero, single_zero_one, one_pow, sub_self] }, + { simp only [single_zero_one, commute.one_right] } }, +end + +example {R} [ring R] [nontrivial R] (n : ℕ) (n0 : 2 ≤ n) : + ∃ f g : add_monoid_algebra R (zmod n), f ≠ 0 ∧ g ≠ 0 ∧ f * g = 0 := +zero_divisors_of_torsion (1 : zmod n) (n0.trans_eq (zmod.add_order_of_one _).symm) + +/-- `F` is the type with two elements `zero` and `one`. We define the "obvious" linear order and +absorbing addition on it to generate our counterexample. -/ +@[derive [decidable_eq, inhabited]] inductive F | zero | one + +/-- The same as `list.get_rest`, except that we take the "rest" from the first match, rather than +from the beginning, returning `[]` if there is no match. For instance, +```lean +#eval [1,2].drop_until [3,1,2,4,1,2] -- [4, 1, 2] +``` +-/ +def list.drop_until {α} [decidable_eq α] : list α → list α → list α +| l [] := [] +| l (a::as) := ((a::as).get_rest l).get_or_else (l.drop_until as) + +/-- `guard_decl_in_file na loc` makes sure that the declaration with name `na` is in the file with +relative path `"src/" ++ "/".intercalate loc ++ ".lean"`. +```lean +#eval guard_decl_in_file `nat.nontrivial ["data", "nat", "basic"] -- does nothing + +#eval guard_decl_in_file `nat.nontrivial ["not", "in", "here"] +-- fails giving the location 'data/nat/basic.lean' +``` + +This test makes sure that the comment referring to this example is in the file claimed in the +doc-module to this counterexample. -/ +meta def guard_decl_in_file (na : name) (loc : list string) : tactic unit := +do env ← tactic.get_env, + some fil ← pure $ env.decl_olean na | fail!"the instance `{na}` is not imported!", + let path : string := ⟨list.drop_until "/src/".to_list fil.to_list⟩, + let locdot : string := ".".intercalate loc, + guard (fil.ends_with ("src/" ++ "/".intercalate loc ++ ".lean")) <|> + fail!("instance `{na}` is no longer in `{locdot}`.\n\n" ++ + "Please, update the doc-module and this check with the correct location:\n\n'{path}'\n") + +#eval guard_decl_in_file `finsupp.lex.covariant_class_le_left ["data", "finsupp", "lex"] +#eval guard_decl_in_file `finsupp.lex.covariant_class_le_right ["data", "finsupp", "lex"] + +namespace F + +instance : has_zero F := ⟨F.zero⟩ + +/-- `1` is not really needed, but it is nice to use the notation. -/ +instance : has_one F := ⟨F.one⟩ + +/-- A tactic to prove trivial goals by enumeration. -/ +meta def boom : tactic unit := +`[ repeat { rintro ⟨⟩ }; dec_trivial ] + +/-- `val` maps `0 1 : F` to their counterparts in `ℕ`. +We use it to lift the linear order on `ℕ`. -/ +def val : F → ℕ +| 0 := 0 +| 1 := 1 + +instance : linear_order F := linear_order.lift' val (by boom) + +@[simp] lemma z01 : (0 : F) < 1 := by boom + +/-- `F` would be a `comm_semiring`, using `min` as multiplication. Again, we do not need this. -/ +instance : add_comm_monoid F := +{ add := max, + add_assoc := by boom, + zero := 0, + zero_add := by boom, + add_zero := by boom, + add_comm := by boom } + +/-- The `covariant_class`es asserting monotonicity of addition hold for `F`. -/ +instance covariant_class_add_le : covariant_class F F (+) (≤) := ⟨by boom⟩ +example : covariant_class F F (function.swap (+)) (≤) := by apply_instance + +/-- The following examples show that `F` has all the typeclasses used by +`finsupp.lex.covariant_class_le_left`... -/ +example : linear_order F := by apply_instance +example : add_monoid F := by apply_instance + +/-- ... except for the strict monotonicity of addition, the crux of the matter. -/ +example : ¬ covariant_class F F (+) (<) := λ h, lt_irrefl 1 $ (h.elim : covariant F F (+) (<)) 1 z01 + +/-- A few `simp`-lemmas to take care of trivialities in the proof of the example below. -/ +@[simp] lemma f1 : ∀ (a : F), 1 + a = 1 := by boom +@[simp] lemma f011 : of_lex (single (0 : F) (1 : F)) 1 = 0 := single_apply_eq_zero.mpr (λ h, h) +@[simp] lemma f010 : of_lex (single (0 : F) (1 : F)) 0 = 1 := single_eq_same +@[simp] lemma f111 : of_lex (single (1 : F) (1 : F)) 1 = 1 := single_eq_same +@[simp] lemma f110 : of_lex (single (1 : F) (1 : F)) 0 = 0 := single_apply_eq_zero.mpr (λ h, h.symm) + +/-- Here we see that (not-necessarily strict) monotonicity of addition on `lex (F →₀ F)` is not +a consequence of monotonicity of addition on `F`. Strict monotonicity of addition on `F` is +enough and is the content of `finsupp.lex.covariant_class_le_left`. -/ +example : ¬ covariant_class (lex (F →₀ F)) (lex (F →₀ F)) (+) (≤) := +begin + rintro ⟨h⟩, + refine not_lt.mpr (h (single (0 : F) (1 : F)) (_ : single 1 1 ≤ single 0 1)) ⟨1, _⟩, + { exact or.inr ⟨0, by simp [(by boom : ∀ j : F, j < 0 ↔ false)]⟩ }, + { simp only [(by boom : ∀ j : F, j < 1 ↔ j = 0), of_lex_add, coe_add, pi.to_lex_apply, + pi.add_apply, forall_eq, f010, f1, eq_self_iff_true, f011, f111, zero_add, and_self] }, +end + +example {α} [ring α] [nontrivial α] : + ∃ f g : add_monoid_algebra α F, f ≠ 0 ∧ g ≠ 0 ∧ f * g = 0 := +zero_divisors_of_periodic le_rfl ((two_smul _ _).trans (by refl)) z01.ne' + +example {α} [has_zero α] : 2 • (single 0 1 : α →₀ F) = single 0 1 ∧ (single 0 1 : α →₀ F) ≠ 0 := +⟨smul_single _ _ _, by simpa only [ne.def, single_eq_zero] using z01.ne⟩ + +end F diff --git a/src/data/finsupp/lex.lean b/src/data/finsupp/lex.lean index b5dd06e4545ca..aa4e4066bf532 100644 --- a/src/data/finsupp/lex.lean +++ b/src/data/finsupp/lex.lean @@ -126,7 +126,9 @@ variables [linear_order α] [add_monoid N] [linear_order N] /-! We are about to sneak in a hypothesis that might appear to be too strong. We assume `covariant_class` with *strict* inequality `<` also when proving the one with the *weak* inequality `≤`. This is actually necessary: addition on `lex (α →₀ N)` may fail to be -monotone, when it is "just" monotone on `N`. -/ +monotone, when it is "just" monotone on `N`. + +See `counterexamples.zero_divisors_in_add_monoid_algebras` for a counterexample. -/ section left variables [covariant_class N N (+) (<)] From 9cf494e0c4e277137e102249e0d098274ee23081 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 6 Oct 2022 10:46:54 +0000 Subject: [PATCH 602/828] fix(geometry/euclidean/oriented_angle): correct lemma name (#16821) `oangle_eq_iff_angle_eq_of_sign_eq` should be called `angle_eq_iff_oangle_eq_of_sign_eq` to match the actual order of the `iff` arguments in the statement. --- src/geometry/euclidean/oriented_angle.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index eb7bbee14a857..f98d366ec9e62 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -994,7 +994,7 @@ end /-- If the signs of two oriented angles between nonzero vectors are equal, the oriented angles are equal if and only if the unoriented angles are equal. -/ -lemma oangle_eq_iff_angle_eq_of_sign_eq {w x y z : V} (hw : w ≠ 0) (hx : x ≠ 0) (hy : y ≠ 0) +lemma angle_eq_iff_oangle_eq_of_sign_eq {w x y z : V} (hw : w ≠ 0) (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) (hs : (b.oangle w x).sign = (b.oangle y z).sign) : inner_product_geometry.angle w x = inner_product_geometry.angle y z ↔ b.oangle w x = b.oangle y z := @@ -1699,11 +1699,11 @@ lemma oangle_eq_of_angle_eq_of_sign_eq {w x y z : V} /-- If the signs of two oriented angles between nonzero vectors are equal, the oriented angles are equal if and only if the unoriented angles are equal. -/ -lemma oangle_eq_iff_angle_eq_of_sign_eq {w x y z : V} (hw : w ≠ 0) (hx : x ≠ 0) (hy : y ≠ 0) +lemma angle_eq_iff_oangle_eq_of_sign_eq {w x y z : V} (hw : w ≠ 0) (hx : x ≠ 0) (hy : y ≠ 0) (hz : z ≠ 0) (hs : (o.oangle w x).sign = (o.oangle y z).sign) : inner_product_geometry.angle w x = inner_product_geometry.angle y z ↔ o.oangle w x = o.oangle y z := -(ob).oangle_eq_iff_angle_eq_of_sign_eq hw hx hy hz hs +(ob).angle_eq_iff_oangle_eq_of_sign_eq hw hx hy hz hs /-- The oriented angle between two nonzero vectors is zero if and only if the unoriented angle is zero. -/ @@ -1932,10 +1932,10 @@ lemma oangle_eq_of_angle_eq_of_sign_eq {p₁ p₂ p₃ p₄ p₅ p₆ : P} (h : /-- If the signs of two nondegenerate oriented angles between points are equal, the oriented angles are equal if and only if the unoriented angles are equal. -/ -lemma oangle_eq_iff_angle_eq_of_sign_eq {p₁ p₂ p₃ p₄ p₅ p₆ : P} (hp₁ : p₁ ≠ p₂) (hp₃ : p₃ ≠ p₂) +lemma angle_eq_iff_oangle_eq_of_sign_eq {p₁ p₂ p₃ p₄ p₅ p₆ : P} (hp₁ : p₁ ≠ p₂) (hp₃ : p₃ ≠ p₂) (hp₄ : p₄ ≠ p₅) (hp₆ : p₆ ≠ p₅) (hs : (∡ p₁ p₂ p₃).sign = (∡ p₄ p₅ p₆).sign) : ∠ p₁ p₂ p₃ = ∠ p₄ p₅ p₆ ↔ ∡ p₁ p₂ p₃ = ∡ p₄ p₅ p₆ := -(o).oangle_eq_iff_angle_eq_of_sign_eq (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₃) +(o).angle_eq_iff_oangle_eq_of_sign_eq (vsub_ne_zero.2 hp₁) (vsub_ne_zero.2 hp₃) (vsub_ne_zero.2 hp₄) (vsub_ne_zero.2 hp₆) hs /-- The unoriented angle at `p` between two points not equal to `p` is zero if and only if the From eaf5cf6bf9e92e51264fa262b008d15b9328ab4c Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Thu, 6 Oct 2022 14:33:11 +0000 Subject: [PATCH 603/828] chore(number_theory/number_field/embeddings): golf (#16804) Also golfs `bUnion_roots_finite` in *data/polynomial/ring_division*. --- src/data/polynomial/ring_division.lean | 28 ++--- src/field_theory/is_alg_closed/basic.lean | 23 ++++ .../number_field/embeddings.lean | 107 +++++------------- 3 files changed, 61 insertions(+), 97 deletions(-) diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index 074500c466758..608c51c1090f9 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -671,25 +671,15 @@ lemma bUnion_roots_finite {R S : Type*} [semiring R] [comm_ring S] [is_domain S] (m : R →+* S) (d : ℕ) {U : set R} (h : U.finite) : (⋃ (f : R[X]) (hf : f.nat_degree ≤ d ∧ ∀ i, (f.coeff i) ∈ U), ((f.map m).roots.to_finset : set S)).finite := -begin - refine set.finite.bUnion _ _, - { -- We prove that the set of polynomials under consideration is finite because its - -- image by the injective map `π` is finite - let π : R[X] → finset.range (d+1) → R := λ f i, f.coeff i, - have h_inj : set.inj_on π {f : R[X] | f.nat_degree ≤ d ∧ ∀ (i : ℕ), f.coeff i ∈ U}, - { intros x hx y hy hxy, - rw ext_iff_nat_degree_le hx.1 hy.1, - exact_mod_cast λ i hi, congr_fun hxy ⟨i, finset.mem_range_succ_iff.mpr hi⟩, }, - have h_fin : (set.pi set.univ (λ e : finset.range (d+1), U)).finite := set.finite.pi (λ e, h), - refine set.finite.of_finite_image (set.finite.subset h_fin _) h_inj, - rw set.image_subset_iff, - intros f hf, - rw [set.mem_preimage, set.mem_univ_pi], - exact λ i, hf.2 i, }, - { intros i hi, - convert root_set_finite (i.map m) S, - simp only [algebra.id.map_eq_id, map_id], }, -end +set.finite.bUnion begin + -- We prove that the set of polynomials under consideration is finite because its + -- image by the injective map `π` is finite + let π : R[X] → fin (d+1) → R := λ f i, f.coeff i, + refine ((set.finite.pi $ λ e, h).subset $ _).of_finite_image (_ : set.inj_on π _), + { exact set.image_subset_iff.2 (λ f hf i _, hf.2 i) }, + { refine λ x hx y hy hxy, (ext_iff_nat_degree_le hx.1 hy.1).2 (λ i hi, _), + exact id congr_fun hxy ⟨i, nat.lt_succ_of_le hi⟩ }, +end $ λ i hi, finset.finite_to_set _ theorem mem_root_set_iff' {p : T[X]} {S : Type*} [comm_ring S] [is_domain S] [algebra T S] (hp : p.map (algebra_map T S) ≠ 0) (a : S) : diff --git a/src/field_theory/is_alg_closed/basic.lean b/src/field_theory/is_alg_closed/basic.lean index 5f7495206061a..5a2282e04aba4 100644 --- a/src/field_theory/is_alg_closed/basic.lean +++ b/src/field_theory/is_alg_closed/basic.lean @@ -486,3 +486,26 @@ ring_hom.ext_iff.2 (equiv_of_equiv_symm_algebra_map L M hSR) end equiv_of_equiv end is_alg_closure + +/-- Let `A` be an algebraically closed field and let `x ∈ K`, with `K/F` an algebraic extension + of fields. Then the images of `x` by the `F`-algebra morphisms from `K` to `A` are exactly + the roots in `A` of the minimal polynomial of `x` over `F`. -/ +lemma algebra.is_algebraic.range_eval_eq_root_set_minpoly {F K} (A) [field F] [field K] [field A] + [is_alg_closed A] [algebra F K] (hK : algebra.is_algebraic F K) [algebra F A] (x : K) : + set.range (λ ψ : K →ₐ[F] A, ψ x) = (minpoly F x).root_set A := +begin + have := algebra.is_algebraic_iff_is_integral.1 hK, + ext a, rw mem_root_set_iff (minpoly.ne_zero $ this x) a, + refine ⟨_, λ ha, _⟩, + { rintro ⟨ψ, rfl⟩, rw [aeval_alg_hom_apply ψ x, minpoly.aeval, map_zero] }, + let Fx := adjoin_root (minpoly F x), + have hx : aeval x (minpoly F x) = 0 := minpoly.aeval F x, + letI : algebra Fx A := (adjoin_root.lift (algebra_map F A) a ha).to_algebra, + letI : algebra Fx K := (adjoin_root.lift (algebra_map F K) x hx).to_algebra, + haveI : is_scalar_tower F Fx A := is_scalar_tower.of_ring_hom (adjoin_root.lift_hom _ a ha), + haveI : is_scalar_tower F Fx K := is_scalar_tower.of_ring_hom (adjoin_root.lift_hom _ x hx), + haveI : fact (irreducible $ minpoly F x) := ⟨minpoly.irreducible $ this x⟩, + let ψ₀ : K →ₐ[Fx] A := is_alg_closed.lift (algebra.is_algebraic_of_larger_base F Fx hK), + exact ⟨ψ₀.restrict_scalars F, (congr_arg ψ₀ (adjoin_root.lift_root hx).symm).trans $ + (ψ₀.commutes _).trans $ adjoin_root.lift_root ha⟩, +end diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index 543b924f44482..53e4582bd1cce 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -48,51 +48,15 @@ section roots open set polynomial -/-- Let `A` an algebraically closed field and let `x ∈ K`, with `K` a number field. For `F`, -subfield of `K`, the images of `x` by the `F`-algebra morphisms from `K` to `A` are exactly -the roots in `A` of the minimal polynomial of `x` over `F`. -/ -lemma range_eq_roots (F K A : Type*) [field F] [number_field F] [field K] [number_field K] - [field A] [is_alg_closed A] [algebra F K] [algebra F A] (x : K) : - range (λ ψ : K →ₐ[F] A, ψ x) = (minpoly F x).root_set A := -begin - haveI : finite_dimensional F K := finite_dimensional.right ℚ _ _ , - have hx : is_integral F x := is_separable.is_integral F x, - ext a, split, - { rintro ⟨ψ, hψ⟩, - rw [mem_root_set_iff, ←hψ], - { rw aeval_alg_hom_apply ψ x (minpoly F x), - simp only [minpoly.aeval, map_zero], }, - exact minpoly.ne_zero hx, }, - { intro ha, - let Fx := adjoin_root (minpoly F x), - haveI : fact (irreducible $ minpoly F x) := ⟨minpoly.irreducible hx⟩, - have hK : (aeval x) (minpoly F x) = 0 := minpoly.aeval _ _, - have hA : (aeval a) (minpoly F x) = 0, - { rwa [aeval_def, ←eval_map, ←mem_root_set_iff'], - exact monic.ne_zero (monic.map (algebra_map F A) (minpoly.monic hx)), }, - letI : algebra Fx A := ring_hom.to_algebra (by convert adjoin_root.lift (algebra_map F A) a hA), - letI : algebra Fx K := ring_hom.to_algebra (by convert adjoin_root.lift (algebra_map F K) x hK), - haveI : finite_dimensional Fx K := finite_dimensional.right ℚ _ _ , - let ψ₀ : K →ₐ[Fx] A := is_alg_closed.lift (algebra.is_algebraic_of_finite _ _), - haveI : is_scalar_tower F Fx K := is_scalar_tower.of_ring_hom (adjoin_root.lift_hom _ _ hK), - haveI : is_scalar_tower F Fx A := is_scalar_tower.of_ring_hom (adjoin_root.lift_hom _ _ hA), - let ψ : K →ₐ[F] A := alg_hom.restrict_scalars F ψ₀, - refine ⟨ψ, _⟩, - rw (_ : x = (algebra_map Fx K) (adjoin_root.root (minpoly F x))), - rw (_ : a = (algebra_map Fx A) (adjoin_root.root (minpoly F x))), - exact alg_hom.commutes _ _, - exact (adjoin_root.lift_root hA).symm, - exact (adjoin_root.lift_root hK).symm, }, -end - -variables (K A : Type*) [field K] [number_field K] [field A] [char_zero A] [is_alg_closed A] (x : K) +variables (K A : Type*) [field K] [number_field K] + [field A] [algebra ℚ A] [is_alg_closed A] (x : K) /-- Let `A` be an algebraically closed field and let `x ∈ K`, with `K` a number field. The images of `x` by the embeddings of `K` in `A` are exactly the roots in `A` of the minimal polynomial of `x` over `ℚ`. -/ -lemma rat_range_eq_roots : range (λ φ : K →+* A, φ x) = (minpoly ℚ x).root_set A := +lemma range_eval_eq_root_set_minpoly : range (λ φ : K →+* A, φ x) = (minpoly ℚ x).root_set A := begin - convert range_eq_roots ℚ K A x using 1, + convert (number_field.is_algebraic K).range_eval_eq_root_set_minpoly A x using 1, ext a, exact ⟨λ ⟨φ, hφ⟩, ⟨φ.to_rat_alg_hom, hφ⟩, λ ⟨φ, hφ⟩, ⟨φ.to_ring_hom, hφ⟩⟩, end @@ -106,20 +70,16 @@ open finite_dimensional polynomial set variables {K : Type*} [field K] [number_field K] variables {A : Type*} [normed_field A] [is_alg_closed A] [normed_algebra ℚ A] -lemma coeff_bdd_of_norm_le {B : ℝ} {x : K} (h : ∀ φ : K →+* A, ∥φ x∥ ≤ B) (i : ℕ): +lemma coeff_bdd_of_norm_le {B : ℝ} {x : K} (h : ∀ φ : K →+* A, ∥φ x∥ ≤ B) (i : ℕ) : ∥(minpoly ℚ x).coeff i∥ ≤ (max B 1) ^ (finrank ℚ K) * (finrank ℚ K).choose ((finrank ℚ K) / 2) := begin - have hx : is_integral ℚ x := is_separable.is_integral _ _, - rw (by rw [coeff_map, norm_algebra_map'] : - ∥(minpoly ℚ x).coeff i∥ = ∥(map (algebra_map ℚ A) (minpoly ℚ x)).coeff i∥), + have hx := is_separable.is_integral ℚ x, + rw [← norm_algebra_map' A, ← coeff_map (algebra_map ℚ A)], refine coeff_bdd_of_roots_le _ (minpoly.monic hx) (is_alg_closed.splits_codomain _) - (minpoly.nat_degree_le hx) _ i, - intros z hz, - rsuffices ⟨φ, rfl⟩ : ∃ (φ : K →+* A), φ x = z, {exact h φ }, - letI : char_zero A := char_zero_of_injective_algebra_map (algebra_map ℚ _).injective, - rw [← set.mem_range, rat_range_eq_roots, mem_root_set_iff, aeval_def], - convert (mem_roots_map _).mp hz, - repeat { exact monic.ne_zero (minpoly.monic hx), }, + (minpoly.nat_degree_le hx) (λ z hz, _) i, + classical, rw ← multiset.mem_to_finset at hz, + obtain ⟨φ, rfl⟩ := (range_eval_eq_root_set_minpoly K A x).symm.subset hz, + exact h φ, end variables (K A) @@ -131,42 +91,33 @@ lemma finite_of_norm_le (B : ℝ) : begin let C := nat.ceil ((max B 1) ^ (finrank ℚ K) * (finrank ℚ K).choose ((finrank ℚ K) / 2)), have := bUnion_roots_finite (algebra_map ℤ K) (finrank ℚ K) (finite_Icc (-C : ℤ) C), - refine this.subset (λ x hx, _), - have h_map_rat_minpoly := minpoly.gcd_domain_eq_field_fractions' ℚ hx.1, - rw mem_Union, - use minpoly ℤ x, - rw [mem_Union, exists_prop], - refine ⟨⟨_, λ i, _⟩, _⟩, - { rw [← nat_degree_map_eq_of_injective (algebra_map ℤ ℚ).injective_int _, ← h_map_rat_minpoly], - exact minpoly.nat_degree_le (is_integral_of_is_scalar_tower _ hx.1), }, - { rw [mem_Icc, ← abs_le, ← @int.cast_le ℝ], - apply le_trans _ (nat.le_ceil _), - convert coeff_bdd_of_norm_le hx.2 i, - simp only [h_map_rat_minpoly, coeff_map, eq_int_cast, int.norm_cast_rat, int.norm_eq_abs, - int.cast_abs], }, - { rw [finset.mem_coe, multiset.mem_to_finset, mem_roots, is_root.def, ← eval₂_eq_eval_map, - ← aeval_def], - { exact minpoly.aeval ℤ x, }, - { exact monic.ne_zero (monic.map (algebra_map ℤ K) (minpoly.monic hx.1)), }}, + refine this.subset (λ x hx, _), simp_rw mem_Union, + have h_map_ℚ_minpoly := minpoly.gcd_domain_eq_field_fractions' ℚ hx.1, + refine ⟨_, ⟨_, λ i, _⟩, (mem_root_set_iff (minpoly.ne_zero hx.1) x).2 (minpoly.aeval ℤ x)⟩, + { rw [← (minpoly.monic hx.1).nat_degree_map (algebra_map ℤ ℚ), ← h_map_ℚ_minpoly], + exact minpoly.nat_degree_le (is_integral_of_is_scalar_tower x hx.1), }, + rw [mem_Icc, ← abs_le, ← @int.cast_le ℝ], + refine (eq.trans_le _ $ coeff_bdd_of_norm_le hx.2 i).trans (nat.le_ceil _), + rw [h_map_ℚ_minpoly, coeff_map, eq_int_cast, int.norm_cast_rat, int.norm_eq_abs, int.cast_abs], end /-- An algebraic integer whose conjugates are all of norm one is a root of unity. -/ lemma pow_eq_one_of_norm_eq_one {x : K} - (hxi : is_integral ℤ x) (hx : ∀ φ : K →+* A, ∥φ x∥ = 1) : + (hxi : is_integral ℤ x) (hx : ∀ φ : K →+* A, ∥φ x∥ = 1) : ∃ (n : ℕ) (hn : 0 < n), x ^ n = 1 := begin obtain ⟨a, -, b, -, habne, h⟩ := @set.infinite.exists_ne_map_eq_of_maps_to _ _ _ _ ((^) x : ℕ → K) set.infinite_univ _ (finite_of_norm_le K A (1:ℝ)), { replace habne := habne.lt_or_lt, - wlog : a < b := habne using [a b], - refine ⟨b - a, tsub_pos_of_lt habne, _⟩, - have hxne : x ≠ 0, - { contrapose! hx, - simp only [hx, norm_zero, ring_hom.map_zero, ne.def, not_false_iff, zero_ne_one], - use (is_alg_closed.lift (number_field.is_algebraic K)).to_ring_hom, }, - { rw [pow_sub₀ _ hxne habne.le, h, mul_inv_cancel (pow_ne_zero b hxne)], }}, - { rw set.maps_univ_to, - exact λ a, ⟨hxi.pow a, λ φ, by simp [hx φ, norm_pow, one_pow]⟩, }, + have : _, swap, cases habne, swap, + { revert a b, exact this }, + { exact this b a h.symm habne }, + refine λ a b h hlt, ⟨a - b, tsub_pos_of_lt hlt, _⟩, + rw [← nat.sub_add_cancel hlt.le, pow_add, mul_left_eq_self₀] at h, + refine h.resolve_right (λ hp, _), + specialize hx (is_alg_closed.lift (number_field.is_algebraic K)).to_ring_hom, + rw [pow_eq_zero hp, map_zero, norm_zero] at hx, norm_num at hx }, + { exact λ a _, ⟨hxi.pow a, λ φ, by simp only [hx φ, norm_pow, one_pow, map_pow]⟩ }, end end bounded From 5b0dd140dc3b854e9eb25e5124d5d5476e9c3d1c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 6 Oct 2022 17:50:20 +0000 Subject: [PATCH 604/828] feat(analysis/normed/group/basic): Multiplicative, noncommutative normed groups (#15705) Define * `seminormed_group` * `seminormed_add_group` * `seminormed_comm_group` * `normed_group` * `normed_add_group` * `normed_comm_group` Multiplicativize all lemmas in `analysis.normed.group.basic`. To disambiguate names, I add one or two `'` to the multiplicative version where needed. --- src/algebra/order/hom/basic.lean | 11 +- src/analysis/asymptotics/asymptotics.lean | 15 +- .../calculus/uniform_limits_deriv.lean | 11 +- src/analysis/complex/schwarz.lean | 4 +- src/analysis/convolution.lean | 4 +- src/analysis/inner_product_space/basic.lean | 12 +- src/analysis/inner_product_space/pi_L2.lean | 3 +- .../locally_convex/with_seminorms.lean | 4 +- src/analysis/matrix.lean | 1 + src/analysis/normed/group/add_torsor.lean | 5 + src/analysis/normed/group/basic.lean | 2141 ++++++++++------- .../normed_space/bounded_linear_maps.lean | 3 +- .../normed_space/continuous_affine_map.lean | 2 +- .../normed_space/lattice_ordered_group.lean | 10 +- src/analysis/specific_limits/normed.lean | 6 +- src/topology/continuous_function/bounded.lean | 2 +- 16 files changed, 1322 insertions(+), 912 deletions(-) diff --git a/src/algebra/order/hom/basic.lean b/src/algebra/order/hom/basic.lean index c21ad5adfbbb5..eaeb26fa3d6f7 100644 --- a/src/algebra/order/hom/basic.lean +++ b/src/algebra/order/hom/basic.lean @@ -60,14 +60,19 @@ attribute [simp] map_nonneg [submultiplicative_hom_class F α β] (f : F) (a b : α) : f a ≤ f b * f (a / b) := by simpa only [mul_comm, div_mul_cancel'] using map_mul_le_mul f (a / b) b +@[to_additive] lemma le_map_add_map_div [group α] [add_comm_semigroup β] [has_le β] + [mul_le_add_hom_class F α β] (f : F) (a b : α) : f a ≤ f b + f (a / b) := +by simpa only [add_comm, div_mul_cancel'] using map_mul_le_add f (a / b) b + @[to_additive] lemma le_map_div_mul_map_div [group α] [comm_semigroup β] [has_le β] [submultiplicative_hom_class F α β] (f : F) (a b c: α) : f (a / c) ≤ f (a / b) * f (b / c) := by simpa only [div_mul_div_cancel'] using map_mul_le_mul f (a / b) (b / c) -@[to_additive] lemma le_map_add_map_div [group α] [add_comm_semigroup β] [has_le β] - [mul_le_add_hom_class F α β] (f : F) (a b : α) : f a ≤ f b + f (a / b) := -by simpa only [add_comm, div_mul_cancel'] using map_mul_le_add f (a / b) b +@[to_additive] +lemma le_map_div_add_map_div [group α] [add_comm_semigroup β] [has_le β] + [mul_le_add_hom_class F α β] (f : F) (a b c: α) : f (a / c) ≤ f (a / b) + f (b / c) := +by simpa only [div_mul_div_cancel'] using map_mul_le_add f (a / b) (b / c) namespace tactic open positivity diff --git a/src/analysis/asymptotics/asymptotics.lean b/src/analysis/asymptotics/asymptotics.lean index a452a1bc2d16a..be4c66e886f89 100644 --- a/src/analysis/asymptotics/asymptotics.lean +++ b/src/analysis/asymptotics/asymptotics.lean @@ -655,11 +655,8 @@ lemma is_O_with_prod_left : lemma is_O.prod_left (hf : f' =O[l] k') (hg : g' =O[l] k') : (λ x, (f' x, g' x)) =O[l] k' := let ⟨c, hf⟩ := hf.is_O_with, ⟨c', hg⟩ := hg.is_O_with in (hf.prod_left hg).is_O -lemma is_O.prod_left_fst (h : (λ x, (f' x, g' x)) =O[l] k') : f' =O[l] k' := -is_O_fst_prod.trans h - -lemma is_O.prod_left_snd (h : (λ x, (f' x, g' x)) =O[l] k') : g' =O[l] k' := -is_O_snd_prod.trans h +lemma is_O.prod_left_fst : (λ x, (f' x, g' x)) =O[l] k' → f' =O[l] k' := is_O.trans is_O_fst_prod +lemma is_O.prod_left_snd : (λ x, (f' x, g' x)) =O[l] k' → g' =O[l] k' := is_O.trans is_O_snd_prod @[simp] lemma is_O_prod_left : (λ x, (f' x, g' x)) =O[l] k' ↔ f' =O[l] k' ∧ g' =O[l] k' := ⟨λ h, ⟨h.prod_left_fst, h.prod_left_snd⟩, λ h, h.1.prod_left h.2⟩ @@ -667,11 +664,11 @@ is_O_snd_prod.trans h lemma is_o.prod_left (hf : f' =o[l] k') (hg : g' =o[l] k') : (λ x, (f' x, g' x)) =o[l] k' := is_o.of_is_O_with $ λ c hc, (hf.forall_is_O_with hc).prod_left_same (hg.forall_is_O_with hc) -lemma is_o.prod_left_fst (h : (λ x, (f' x, g' x)) =o[l] k') : f' =o[l] k' := -is_O_fst_prod.trans_is_o h +lemma is_o.prod_left_fst : (λ x, (f' x, g' x)) =o[l] k' → f' =o[l] k' := +is_O.trans_is_o is_O_fst_prod -lemma is_o.prod_left_snd (h : (λ x, (f' x, g' x)) =o[l] k') : g' =o[l] k' := -is_O_snd_prod.trans_is_o h +lemma is_o.prod_left_snd : (λ x, (f' x, g' x)) =o[l] k' → g' =o[l] k' := +is_O.trans_is_o is_O_snd_prod @[simp] lemma is_o_prod_left : (λ x, (f' x, g' x)) =o[l] k' ↔ f' =o[l] k' ∧ g' =o[l] k' := ⟨λ h, ⟨h.prod_left_fst, h.prod_left_snd⟩, λ h, h.1.prod_left h.2⟩ diff --git a/src/analysis/calculus/uniform_limits_deriv.lean b/src/analysis/calculus/uniform_limits_deriv.lean index 4416629d8836d..b549459a49fcf 100644 --- a/src/analysis/calculus/uniform_limits_deriv.lean +++ b/src/analysis/calculus/uniform_limits_deriv.lean @@ -114,7 +114,7 @@ lemma uniform_cauchy_seq_on_filter_of_tendsto_uniformly_on_filter_fderiv (hfg : tendsto (λ n, f n x) l (𝓝 (g x))) : uniform_cauchy_seq_on_filter f l (𝓝 x) := begin - rw normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero at + rw seminormed_add_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero at hf' ⊢, suffices : tendsto_uniformly_on_filter @@ -180,7 +180,7 @@ lemma uniform_cauchy_seq_on_ball_of_tendsto_uniformly_on_ball_fderiv (hfg : tendsto (λ n, f n x) l (𝓝 (g x))) : uniform_cauchy_seq_on f l (metric.ball x r) := begin - rw normed_add_comm_group.uniform_cauchy_seq_on_iff_tendsto_uniformly_on_zero at hf' ⊢, + rw seminormed_add_group.uniform_cauchy_seq_on_iff_tendsto_uniformly_on_zero at hf' ⊢, suffices : tendsto_uniformly_on (λ (n : ι × ι) (z : E), f n.1 z - f n.2 z - (f n.1 x - f n.2 x)) 0 @@ -235,12 +235,11 @@ lemma difference_quotients_converge_uniformly begin refine uniform_cauchy_seq_on_filter.tendsto_uniformly_on_filter_of_tendsto _ ((hfg.and (eventually_const.mpr hfg.self_of_nhds)).mono (λ y hy, (hy.1.sub hy.2).const_smul _)), - rw normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero, + rw seminormed_add_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero, rw metric.tendsto_uniformly_on_filter_iff, have hfg' := hf'.uniform_cauchy_seq_on_filter, - rw normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero at - hfg', + rw seminormed_add_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero at hfg', rw metric.tendsto_uniformly_on_filter_iff at hfg', intros ε hε, obtain ⟨q, hqpos, hqε⟩ := exists_pos_rat_lt hε, @@ -431,7 +430,7 @@ begin -- metrics are written in terms of `<`. So we need to shrink `ε` utilizing the archimedean -- property of `ℝ` - rw [normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero, + rw [seminormed_add_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero, metric.tendsto_uniformly_on_filter_iff] at hf' ⊢, intros ε hε, obtain ⟨q, hq, hq'⟩ := exists_between hε.lt, diff --git a/src/analysis/complex/schwarz.lean b/src/analysis/complex/schwarz.lean index 83e5125bb8ded..469986d422a44 100644 --- a/src/analysis/complex/schwarz.lean +++ b/src/analysis/complex/schwarz.lean @@ -77,8 +77,8 @@ begin { rw frontier_ball c hr₀.ne', intros z hz, have hz' : z ≠ c, from ne_of_mem_sphere hz hr₀.ne', - rw [dslope_of_ne _ hz', slope_def_module, norm_smul, norm_inv, - (mem_sphere_iff_norm _ _ _).1 hz, ← div_eq_inv_mul, div_le_div_right hr₀, ← dist_eq_norm], + rw [dslope_of_ne _ hz', slope_def_module, norm_smul, norm_inv, mem_sphere_iff_norm.1 hz, + ← div_eq_inv_mul, div_le_div_right hr₀, ← dist_eq_norm], exact le_of_lt (h_maps (mem_ball.2 (by { rw mem_sphere.1 hz, exact hr.2 }))) }, { rw [closure_ball c hr₀.ne', mem_closed_ball], exact hr.1.le } diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 64f55e61721a6..88cdb7da1a1ea 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -643,7 +643,9 @@ begin (eventually_of_forall h2)).trans _, rw [integral_mul_right], refine mul_le_mul_of_nonneg_right _ hε, - have h3 : ∀ t, ∥L (f t)∥ ≤ ∥L∥ * ∥f t∥ := λ t, L.le_op_norm (f t), + have h3 : ∀ t, ∥L (f t)∥ ≤ ∥L∥ * ∥f t∥, + { intros t, + exact L.le_op_norm (f t) }, refine (integral_mono (L.integrable_comp hif).norm (hif.norm.const_mul _) h3).trans_eq _, rw [integral_mul_left] end diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 981c841163262..316fb4b786889 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -1821,7 +1821,8 @@ end bessels_inequality /-- A field `𝕜` satisfying `is_R_or_C` is itself a `𝕜`-inner product space. -/ instance is_R_or_C.inner_product_space : inner_product_space 𝕜 𝕜 := -{ inner := (λ x y, (conj x) * y), +{ to_normed_add_comm_group := non_unital_normed_ring.to_normed_add_comm_group, + inner := λ x y, conj x * y, norm_sq_eq_inner := λ x, by { unfold inner, rw [mul_comm, mul_conj, of_real_re, norm_sq_eq_def'] }, conj_sym := λ x y, by simp [mul_comm], @@ -1834,7 +1835,8 @@ instance is_R_or_C.inner_product_space : inner_product_space 𝕜 𝕜 := /-- Induced inner product on a submodule. -/ instance submodule.inner_product_space (W : submodule 𝕜 E) : inner_product_space 𝕜 W := -{ inner := λ x y, ⟪(x:E), (y:E)⟫, +{ to_normed_add_comm_group := submodule.normed_add_comm_group _, + inner := λ x y, ⟪(x:E), (y:E)⟫, conj_sym := λ _ _, inner_conj_sym _ _ , norm_sq_eq_inner := λ _, norm_sq_eq_inner _, add_left := λ _ _ _ , inner_add_left, @@ -2088,7 +2090,8 @@ registered as an instance since it creates problems with the case `𝕜 = ℝ`, proof to obtain a real inner product space structure from a given `𝕜`-inner product space structure. -/ def inner_product_space.is_R_or_C_to_real : inner_product_space ℝ E := -{ norm_sq_eq_inner := norm_sq_eq_inner, +{ to_normed_add_comm_group := inner_product_space.to_normed_add_comm_group 𝕜, + norm_sq_eq_inner := norm_sq_eq_inner, conj_sym := λ x y, inner_re_symm, add_left := λ x y z, by { change re ⟪x + y, z⟫ = re ⟪x, z⟫ + re ⟪y, z⟫, @@ -2376,7 +2379,8 @@ protected lemma continuous.inner {α : Type*} [topological_space α] uniform_space.completion.continuous_inner.comp (hf.prod_mk hg : _) instance : inner_product_space 𝕜 (completion E) := -{ norm_sq_eq_inner := λ x, completion.induction_on x +{ to_normed_add_comm_group := infer_instance, + norm_sq_eq_inner := λ x, completion.induction_on x (is_closed_eq (continuous_norm.pow 2) (continuous_re.comp (continuous.inner continuous_id' continuous_id'))) diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index 004096a224479..b711d17671bc9 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -70,7 +70,8 @@ we use instead `pi_Lp 2 f` for the product space, which is endowed with the `L^2 -/ instance pi_Lp.inner_product_space {ι : Type*} [fintype ι] (f : ι → Type*) [Π i, inner_product_space 𝕜 (f i)] : inner_product_space 𝕜 (pi_Lp 2 f) := -{ inner := λ x y, ∑ i, inner (x i) (y i), +{ to_normed_add_comm_group := infer_instance, + inner := λ x y, ∑ i, inner (x i) (y i), norm_sq_eq_inner := λ x, by simp only [pi_Lp.norm_sq_eq_of_L2, add_monoid_hom.map_sum, ← norm_sq_eq_inner, one_div], conj_sym := diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 6696bcebe14a3..d295bea227674 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -334,8 +334,8 @@ lemma norm_with_seminorms (𝕜 E) [normed_field 𝕜] [seminormed_add_comm_grou with_seminorms (λ (_ : fin 1), norm_seminorm 𝕜 E) := begin let p : seminorm_family 𝕜 E (fin 1) := λ _, norm_seminorm 𝕜 E, - refine ⟨topological_add_group.ext normed_top_group - (p.add_group_filter_basis.is_topological_add_group) _⟩, + refine ⟨seminormed_add_comm_group.to_topological_add_group.ext + p.add_group_filter_basis.is_topological_add_group _⟩, refine filter.has_basis.eq_of_same_basis metric.nhds_basis_ball _, rw ←ball_norm_seminorm 𝕜 E, refine filter.has_basis.to_has_basis p.add_group_filter_basis.nhds_zero_has_basis _ diff --git a/src/analysis/matrix.lean b/src/analysis/matrix.lean index a4a043dd3870a..471c73f1c2de7 100644 --- a/src/analysis/matrix.lean +++ b/src/analysis/matrix.lean @@ -132,6 +132,7 @@ end congr_arg coe $ nnnorm_diagonal v /-- Note this is safe as an instance as it carries no data. -/ +@[nolint fails_quickly] instance [nonempty n] [decidable_eq n] [has_one α] [norm_one_class α] : norm_one_class (matrix n n α) := ⟨(norm_diagonal _).trans $ norm_one⟩ diff --git a/src/analysis/normed/group/add_torsor.lean b/src/analysis/normed/group/add_torsor.lean index fcd36ce4130b0..fc06853e8b08b 100644 --- a/src/analysis/normed/group/add_torsor.lean +++ b/src/analysis/normed/group/add_torsor.lean @@ -30,6 +30,11 @@ class normed_add_torsor (V : out_param $ Type*) (P : Type*) extends add_torsor V P := (dist_eq_norm' : ∀ (x y : P), dist x y = ∥(x -ᵥ y : V)∥) +/-- Shortcut instance to help typeclass inference out. -/ +@[priority 100] +instance normed_add_torsor.to_add_torsor' {V P : Type*} [normed_add_comm_group V] [metric_space P] + [normed_add_torsor V P] : add_torsor V P := normed_add_torsor.to_add_torsor + variables {α V P W Q : Type*} [seminormed_add_comm_group V] [pseudo_metric_space P] [normed_add_torsor V P] [normed_add_comm_group W] [metric_space Q] [normed_add_torsor W Q] diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 62620f2883ef1..50a2579e204c1 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -1,7 +1,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 +Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies -/ import algebra.module.ulift import analysis.normed.group.seminorm @@ -14,85 +14,264 @@ import topology.sequences /-! # Normed (semi)groups -In this file we define four classes: +In this file we define 10 classes: * `has_norm`, `has_nnnorm`: auxiliary classes endowing a type `α` with a function `norm : α → ℝ` (notation: `∥x∥`) and `nnnorm : α → ℝ≥0` (notation: `∥x∥₊`), respectively; -* `seminormed_add_comm_group`: a seminormed group is an additive group with a norm and a - pseudo-metric space structures that agree with each other: `∀ x y, dist x y = ∥x - y∥`; -* `normed_add_comm_group`: a normed group is an additive group with a norm and a metric space - structures that agree with each other: `∀ x y, dist x y = ∥x - y∥`. +* `seminormed_..._group`: A seminormed (additive) (commutative) group is an (additive) (commutative) + group with a norm and a compatible pseudometric space structure: + `∀ x y, dist x y = ∥x / y∥` or `∀ x y, dist x y = ∥x - y∥`, depending on the group operation. +* `normed_..._group`: A normed (additive) (commutative) group is an (additive) (commutative) group + with a norm and a compatible metric space structure. We also prove basic properties of (semi)normed groups and provide some instances. +## Notes + +The current convention `dist x y = ∥x - y∥` means that the distance is invariant under right +addition, but actions in mathlib are usually from the left. This means we might want to change it to +`dist x y = ∥-x + y∥`. + +The normed group hierarchy would lend itself well to a mixin design (that is, having +`seminormed_group` and `seminormed_add_group` not extend `group` and `add_group`), but we choose not +to for performance concerns. + ## Tags normed group -/ -variables {α ι E F G : Type*} +variables {𝓕 𝕜 α ι κ E F G : Type*} -open filter metric -open_locale topological_space big_operators nnreal ennreal uniformity pointwise +open filter function metric +open_locale big_operators ennreal filter nnreal uniformity pointwise topological_space /-- Auxiliary class, endowing a type `E` with a function `norm : E → ℝ` with notation `∥x∥`. This class is designed to be extended in more interesting classes specifying the properties of the norm. -/ -class has_norm (E : Type*) := (norm : E → ℝ) +@[notation_class] class has_norm (E : Type*) := (norm : E → ℝ) + +/-- Auxiliary class, endowing a type `α` with a function `nnnorm : α → ℝ≥0` with notation `∥x∥₊`. -/ +@[notation_class] class has_nnnorm (E : Type*) := (nnnorm : E → ℝ≥0) export has_norm (norm) +export has_nnnorm (nnnorm) notation `∥` e `∥` := norm e +notation `∥` e `∥₊` := nnnorm e + +/-- A seminormed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` +defines a pseudometric space structure. -/ +class seminormed_add_group (E : Type*) extends has_norm E, add_group E, pseudo_metric_space E := +(dist := λ x y, ∥x - y∥) +(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) + +/-- A seminormed group is a group endowed with a norm for which `dist x y = ∥x / y∥` defines a +pseudometric space structure. -/ +@[to_additive] +class seminormed_group (E : Type*) extends has_norm E, group E, pseudo_metric_space E := +(dist := λ x y, ∥x / y∥) +(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) + +/-- A normed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` defines a +metric space structure. -/ +class normed_add_group (E : Type*) extends has_norm E, add_group E, metric_space E := +(dist := λ x y, ∥x - y∥) +(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) + +/-- A normed group is a group endowed with a norm for which `dist x y = ∥x / y∥` defines a metric +space structure. -/ +@[to_additive] +class normed_group (E : Type*) extends has_norm E, group E, metric_space E := +(dist := λ x y, ∥x / y∥) +(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) /-- A seminormed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` defines a pseudometric space structure. -/ class seminormed_add_comm_group (E : Type*) extends has_norm E, add_comm_group E, pseudo_metric_space E := -(dist_eq : ∀ x y : E, dist x y = norm (x - y)) +(dist := λ x y, ∥x - y∥) +(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) -/-- A normed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` defines -a metric space structure. -/ +/-- A seminormed group is a group endowed with a norm for which `dist x y = ∥x / y∥` +defines a pseudometric space structure. -/ +@[to_additive] +class seminormed_comm_group (E : Type*) + extends has_norm E, comm_group E, pseudo_metric_space E := +(dist := λ x y, ∥x / y∥) +(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) + +/-- A normed group is an additive group endowed with a norm for which `dist x y = ∥x - y∥` defines a +metric space structure. -/ class normed_add_comm_group (E : Type*) extends has_norm E, add_comm_group E, metric_space E := -(dist_eq : ∀ x y : E, dist x y = norm (x - y)) - -/-- A normed group is a seminormed group. -/ -@[priority 100] -- see Note [lower instance priority] -instance normed_add_comm_group.to_seminormed_add_comm_group [h : normed_add_comm_group E] : - seminormed_add_comm_group E := -{ .. h } - -/-- Construct a seminormed group from a translation invariant pseudodistance. -/ -def seminormed_add_comm_group.of_add_dist [has_norm E] [add_comm_group E] [pseudo_metric_space E] - (H1 : ∀ x : E, ∥x∥ = dist x 0) - (H2 : ∀ x y z : E, dist x y ≤ dist (x + z) (y + z)) : seminormed_add_comm_group E := +(dist := λ x y, ∥x - y∥) +(dist_eq : ∀ x y, dist x y = ∥x - y∥ . obviously) + +/-- A normed group is a group endowed with a norm for which `dist x y = ∥x / y∥` defines a metric +space structure. -/ +@[to_additive] +class normed_comm_group (E : Type*) extends has_norm E, comm_group E, metric_space E := +(dist := λ x y, ∥x / y∥) +(dist_eq : ∀ x y, dist x y = ∥x / y∥ . obviously) + +@[priority 100, to_additive] -- See note [lower instance priority] +instance normed_group.to_seminormed_group [normed_group E] : seminormed_group E := +{ ..‹normed_group E› } + +@[priority 100, to_additive] -- See note [lower instance priority] +instance normed_comm_group.to_seminormed_comm_group [normed_comm_group E] : + seminormed_comm_group E := +{ ..‹normed_comm_group E› } + +@[priority 100, to_additive] -- See note [lower instance priority] +instance seminormed_comm_group.to_seminormed_group [seminormed_comm_group E] : seminormed_group E := +{ ..‹seminormed_comm_group E› } + +@[priority 100, to_additive] -- See note [lower instance priority] +instance normed_comm_group.to_normed_group [normed_comm_group E] : normed_group E := +{ ..‹normed_comm_group E› } + +/-- Construct a `normed_group` from a `seminormed_group` satisfying `∀ x, ∥x∥ = 0 → x = 1`. This +avoids having to go back to the `(pseudo_)metric_space` level when declaring a `normed_group` +instance as a special case of a more general `seminormed_group` instance. -/ +@[to_additive "Construct a `normed_add_group` from a `seminormed_add_group` satisfying +`∀ x, ∥x∥ = 0 → x = 0`. This avoids having to go back to the `(pseudo_)metric_space` level when +declaring a `normed_add_group` instance as a special case of a more general `seminormed_add_group` +instance.", reducible] -- See note [reducible non-instances] +def normed_group.of_separation [seminormed_group E] (h : ∀ x : E, ∥x∥ = 0 → x = 1) : + normed_group E := +{ to_metric_space := + { eq_of_dist_eq_zero := λ x y hxy, div_eq_one.1 $ h _ $ by rwa ←‹seminormed_group E›.dist_eq }, + ..‹seminormed_group E› } + +/-- Construct a `normed_comm_group` from a `seminormed_comm_group` satisfying +`∀ x, ∥x∥ = 0 → x = 1`. This avoids having to go back to the `(pseudo_)metric_space` level when +declaring a `normed_comm_group` instance as a special case of a more general `seminormed_comm_group` +instance. -/ +@[to_additive "Construct a `normed_add_comm_group` from a `seminormed_add_comm_group` satisfying +`∀ x, ∥x∥ = 0 → x = 0`. This avoids having to go back to the `(pseudo_)metric_space` level when +declaring a `normed_add_comm_group` instance as a special case of a more general +`seminormed_add_comm_group` instance.", reducible] -- See note [reducible non-instances] +def normed_comm_group.of_separation [seminormed_comm_group E] (h : ∀ x : E, ∥x∥ = 0 → x = 1) : + normed_comm_group E := +{ ..‹seminormed_comm_group E›, ..normed_group.of_separation h } + +/-- Construct a seminormed group from a multiplication-invariant distance. -/ +@[to_additive "Construct a seminormed group from a translation-invariant distance."] +def seminormed_group.of_mul_dist [has_norm E] [group E] [pseudo_metric_space E] + (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : + seminormed_group E := { dist_eq := λ x y, begin - rw H1, apply le_antisymm, - { rw [sub_eq_add_neg, ← add_right_neg y], apply H2 }, - { have := H2 (x - y) 0 y, rwa [sub_add_cancel, zero_add] at this } + rw h₁, apply le_antisymm, + { simpa only [div_eq_mul_inv, ← mul_right_inv y] using h₂ _ _ _ }, + { simpa only [div_mul_cancel', one_mul] using h₂ (x/y) 1 y } end } -/-- Construct a seminormed group from a translation invariant pseudodistance -/ -def seminormed_add_comm_group.of_add_dist' [has_norm E] [add_comm_group E] [pseudo_metric_space E] - (H1 : ∀ x : E, ∥x∥ = dist x 0) - (H2 : ∀ x y z : E, dist (x + z) (y + z) ≤ dist x y) : seminormed_add_comm_group E := +/-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ +@[to_additive "Construct a seminormed group from a translation-invariant pseudodistance."] +def seminormed_group.of_mul_dist' [has_norm E] [group E] [pseudo_metric_space E] + (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : + seminormed_group E := { dist_eq := λ x y, begin - rw H1, apply le_antisymm, - { have := H2 (x - y) 0 y, rwa [sub_add_cancel, zero_add] at this }, - { rw [sub_eq_add_neg, ← add_right_neg y], apply H2 } + rw h₁, apply le_antisymm, + { simpa only [div_mul_cancel', one_mul] using h₂ (x/y) 1 y }, + { simpa only [div_eq_mul_inv, ← mul_right_inv y] using h₂ _ _ _ } end } +/-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ +@[to_additive "Construct a seminormed group from a translation-invariant pseudodistance."] +def seminormed_comm_group.of_mul_dist [has_norm E] [comm_group E] [pseudo_metric_space E] + (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : + seminormed_comm_group E := +{ ..seminormed_group.of_mul_dist h₁ h₂ } + +/-- Construct a seminormed group from a multiplication-invariant pseudodistance. -/ +@[to_additive "Construct a seminormed group from a translation-invariant pseudodistance."] +def seminormed_comm_group.of_mul_dist' [has_norm E] [comm_group E] [pseudo_metric_space E] + (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : + seminormed_comm_group E := +{ ..seminormed_group.of_mul_dist' h₁ h₂ } + +/-- Construct a normed group from a multiplication-invariant distance. -/ +@[to_additive "Construct a normed group from a translation-invariant distance."] +def normed_group.of_mul_dist [has_norm E] [group E] [metric_space E] + (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : + normed_group E := +{ ..seminormed_group.of_mul_dist h₁ h₂ } + +/-- Construct a normed group from a multiplication-invariant pseudodistance. -/ +@[to_additive "Construct a normed group from a translation-invariant pseudodistance."] +def normed_group.of_mul_dist' [has_norm E] [group E] [metric_space E] + (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : + normed_group E := +{ ..seminormed_group.of_mul_dist' h₁ h₂ } + +/-- Construct a normed group from a multiplication-invariant pseudodistance. -/ +@[to_additive "Construct a normed group from a translation-invariant pseudodistance."] +def normed_comm_group.of_mul_dist [has_norm E] [comm_group E] [metric_space E] + (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist x y ≤ dist (x * z) (y * z)) : + normed_comm_group E := +{ ..normed_group.of_mul_dist h₁ h₂ } + +/-- Construct a normed group from a multiplication-invariant pseudodistance. -/ +@[to_additive "Construct a normed group from a translation-invariant pseudodistance."] +def normed_comm_group.of_mul_dist' [has_norm E] [comm_group E] [metric_space E] + (h₁ : ∀ x : E, ∥x∥ = dist x 1) (h₂ : ∀ x y z : E, dist (x * z) (y * z) ≤ dist x y) : + normed_comm_group E := +{ ..normed_group.of_mul_dist' h₁ h₂ } + +set_option old_structure_cmd true + /-- Construct a seminormed group from a seminorm, i.e., registering the pseudodistance and the pseudometric space structure from the seminorm properties. Note that in most cases this instance creates bad definitional equalities (e.g., it does not take into account a possibly existing `uniform_space` instance on `E`). -/ -def add_group_seminorm.to_seminormed_add_comm_group [add_comm_group E] (f : add_group_seminorm E) : - seminormed_add_comm_group E := -{ dist := λ x y, f (x - y), +@[to_additive "Construct a seminormed group from a seminorm, i.e., registering the pseudodistance* +and the pseudometric space structure from the seminorm properties. Note that in most cases this +instance creates bad definitional equalities (e.g., it does not take into account a possibly +existing `uniform_space` instance on `E`)."] +def group_seminorm.to_seminormed_group [group E] (f : group_seminorm E) : seminormed_group E := +{ dist := λ x y, f (x / y), norm := f, dist_eq := λ x y, rfl, - dist_self := λ x, by simp only [sub_self, map_zero], - dist_triangle := le_map_sub_add_map_sub f, - dist_comm := map_sub_rev f } + dist_self := λ x, by simp only [div_self', map_one_eq_zero], + dist_triangle := le_map_div_add_map_div f, + dist_comm := map_div_rev f } + +/-- Construct a seminormed group from a seminorm, i.e., registering the pseudodistance and the +pseudometric space structure from the seminorm properties. Note that in most cases this instance +creates bad definitional equalities (e.g., it does not take into account a possibly existing +`uniform_space` instance on `E`). -/ +@[to_additive "Construct a seminormed group from a seminorm, i.e., registering the pseudodistance* +and the pseudometric space structure from the seminorm properties. Note that in most cases this +instance creates bad definitional equalities (e.g., it does not take into account a possibly +existing `uniform_space` instance on `E`)."] +def group_seminorm.to_seminormed_comm_group [comm_group E] (f : group_seminorm E) : + seminormed_comm_group E := +{ ..f.to_seminormed_group } + +/-- Construct a normed group from a norm, i.e., registering the distance and the metric space +structure from the norm properties. Note that in most cases this instance creates bad definitional +equalities (e.g., it does not take into account a possibly existing `uniform_space` instance on +`E`). -/ +@[to_additive "Construct a normed group from a norm, i.e., registering the distance and the metric +space structure from the norm properties. Note that in most cases this instance creates bad +definitional equalities (e.g., it does not take into account a possibly existing `uniform_space` +instance on `E`)."] +def group_norm.to_normed_group [group E] (f : group_norm E) : normed_group E := +{ eq_of_dist_eq_zero := λ x y h, div_eq_one.1 $ eq_one_of_map_eq_zero f h, + ..f.to_group_seminorm.to_seminormed_group } + +/-- Construct a normed group from a norm, i.e., registering the distance and the metric space +structure from the norm properties. Note that in most cases this instance creates bad definitional +equalities (e.g., it does not take into account a possibly existing `uniform_space` instance on +`E`). -/ +@[to_additive "Construct a normed group from a norm, i.e., registering the distance and the metric +space structure from the norm properties. Note that in most cases this instance creates bad +definitional equalities (e.g., it does not take into account a possibly existing `uniform_space` +instance on `E`)."] +def group_norm.to_normed_comm_group [comm_group E] (f : group_norm E) : normed_comm_group E := +{ ..f.to_normed_group } instance : normed_add_comm_group punit := { norm := function.const _ 0, @@ -104,1138 +283,1362 @@ instance : has_norm ℝ := { norm := λ x, |x| } @[simp] lemma real.norm_eq_abs (r : ℝ) : ∥r∥ = |r| := rfl -instance : normed_add_comm_group ℝ := -{ dist_eq := assume x y, rfl } +instance : normed_add_comm_group ℝ := ⟨λ x y, rfl⟩ -section seminormed_add_comm_group -variables [seminormed_add_comm_group E] [seminormed_add_comm_group F] [seminormed_add_comm_group G] +section seminormed_group +variables [seminormed_group E] [seminormed_group F] [seminormed_group G] {s : set E} + {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} -lemma dist_eq_norm (g h : E) : dist g h = ∥g - h∥ := -seminormed_add_comm_group.dist_eq _ _ +@[to_additive] +lemma dist_eq_norm_div (a b : E) : dist a b = ∥a / b∥ := seminormed_group.dist_eq _ _ -lemma dist_eq_norm' (g h : E) : dist g h = ∥h - g∥ := -by rw [dist_comm, dist_eq_norm] +@[to_additive] +lemma dist_eq_norm_div' (a b : E) : dist a b = ∥b / a∥ := by rw [dist_comm, dist_eq_norm_div] -@[simp] lemma dist_zero_right (g : E) : dist g 0 = ∥g∥ := -by rw [dist_eq_norm, sub_zero] +alias dist_eq_norm_sub ← dist_eq_norm +alias dist_eq_norm_sub' ← dist_eq_norm' -@[simp] lemma dist_zero_left : dist (0 : E) = norm := -funext $ λ g, by rw [dist_comm, dist_zero_right] +@[simp, to_additive] lemma dist_one_right (a : E) : dist a 1 = ∥a∥ := +by rw [dist_eq_norm_div, div_one] -lemma isometry.norm_map_of_map_zero {f : E → F} (hi : isometry f) (h0 : f 0 = 0) (x : E) : +@[simp, to_additive] lemma dist_one_left : dist (1 : E) = norm := +funext $ λ a, by rw [dist_comm, dist_one_right] + +@[to_additive] +lemma isometry.norm_map_of_map_one {f : E → F} (hi : isometry f) (h₁ : f 1 = 1) (x : E) : ∥f x∥ = ∥x∥ := -by rw [←dist_zero_right, ←h0, hi.dist_eq, dist_zero_right] +by rw [←dist_one_right, ←h₁, hi.dist_eq, dist_one_right] -lemma tendsto_norm_cocompact_at_top [proper_space E] : - tendsto norm (cocompact E) at_top := -by simpa only [dist_zero_right] using tendsto_dist_right_cocompact_at_top (0 : E) +@[to_additive tendsto_norm_cocompact_at_top] +lemma tendsto_norm_cocompact_at_top' [proper_space E] : tendsto norm (cocompact E) at_top := +by simpa only [dist_one_right] using tendsto_dist_right_cocompact_at_top (1 : E) -lemma norm_sub_rev (g h : E) : ∥g - h∥ = ∥h - g∥ := -by simpa only [dist_eq_norm] using dist_comm g h +@[to_additive] lemma norm_div_rev (a b : E) : ∥a / b∥ = ∥b / a∥ := +by simpa only [dist_eq_norm_div] using dist_comm a b -@[simp] lemma norm_neg (g : E) : ∥-g∥ = ∥g∥ := -by simpa using norm_sub_rev 0 g +@[simp, to_additive norm_neg] +lemma norm_inv' (a : E) : ∥a⁻¹∥ = ∥a∥ := by simpa using norm_div_rev 1 a -@[simp] lemma dist_add_left (g h₁ h₂ : E) : dist (g + h₁) (g + h₂) = dist h₁ h₂ := -by simp [dist_eq_norm] +@[simp, to_additive] lemma dist_mul_right (a₁ a₂ b : E) : dist (a₁ * b) (a₂ * b) = dist a₁ a₂ := +by simp [dist_eq_norm_div] -@[simp] lemma dist_add_right (g₁ g₂ h : E) : dist (g₁ + h) (g₂ + h) = dist g₁ g₂ := -by simp [dist_eq_norm] +@[to_additive] lemma dist_div_right (a₁ a₂ b : E) : dist (a₁ / b) (a₂ / b) = dist a₁ a₂ := +by simpa only [div_eq_mul_inv] using dist_mul_right _ _ _ -lemma dist_neg (x y : E) : dist (-x) y = dist x (-y) := -by simp_rw [dist_eq_norm, ←norm_neg (-x - y), neg_sub, sub_neg_eq_add, add_comm] +@[simp, to_additive] lemma dist_div_eq_dist_mul_left (a b c : E) : + dist (a / b) c = dist a (c * b) := +by rw [←dist_mul_right _ _ b, div_mul_cancel'] -@[simp] lemma dist_neg_neg (g h : E) : dist (-g) (-h) = dist g h := by rw [dist_neg, neg_neg] +@[simp, to_additive] lemma dist_div_eq_dist_mul_right (a b c : E) : + dist a (b / c) = dist (a * c) b := +by rw [←dist_mul_right _ _ c, div_mul_cancel'] -@[simp] lemma dist_sub_left (g h₁ h₂ : E) : dist (g - h₁) (g - h₂) = dist h₁ h₂ := -by simp only [sub_eq_add_neg, dist_add_left, dist_neg_neg] +/-- In a (semi)normed group, inversion `x ↦ x⁻¹` tends to infinity at infinity. TODO: use +`bornology.cobounded` instead of `filter.comap has_norm.norm filter.at_top`. -/ +@[to_additive "In a (semi)normed group, negation `x ↦ -x` tends to infinity at infinity. TODO: use +`bornology.cobounded` instead of `filter.comap has_norm.norm filter.at_top`."] +lemma filter.tendsto_inv_cobounded : + tendsto (has_inv.inv : E → E) (comap norm at_top) (comap norm at_top) := +by simpa only [norm_inv', tendsto_comap_iff, (∘)] using tendsto_comap -@[simp] lemma dist_sub_right (g₁ g₂ h : E) : dist (g₁ - h) (g₂ - h) = dist g₁ g₂ := -by simpa only [sub_eq_add_neg] using dist_add_right _ _ _ +/-- **Triangle inequality** for the norm. -/ +@[to_additive norm_add_le "**Triangle inequality** for the norm."] +lemma norm_mul_le' (a b : E) : ∥a * b∥ ≤ ∥a∥ + ∥b∥ := +by simpa [dist_eq_norm_div] using dist_triangle a 1 b⁻¹ -@[simp] theorem dist_self_add_right (g h : E) : dist g (g + h) = ∥h∥ := -by rw [← dist_zero_left, ← dist_add_left g 0 h, add_zero] +@[to_additive] lemma norm_mul_le_of_le (h₁ : ∥a₁∥ ≤ r₁) (h₂ : ∥a₂∥ ≤ r₂) : ∥a₁ * a₂∥ ≤ r₁ + r₂ := +(norm_mul_le' a₁ a₂).trans $ add_le_add h₁ h₂ -@[simp] theorem dist_self_add_left (g h : E) : dist (g + h) g = ∥h∥ := -by rw [dist_comm, dist_self_add_right] +@[to_additive norm_add₃_le] lemma norm_mul₃_le (a b c : E) : ∥a * b * c∥ ≤ ∥a∥ + ∥b∥ + ∥c∥ := +norm_mul_le_of_le (norm_mul_le' _ _) le_rfl -@[simp] theorem dist_self_sub_right (g h : E) : dist g (g - h) = ∥h∥ := -by rw [sub_eq_add_neg, dist_self_add_right, norm_neg] +@[simp, to_additive norm_nonneg] lemma norm_nonneg' (a : E) : 0 ≤ ∥a∥ := +by { rw [←dist_one_right], exact dist_nonneg } -@[simp] theorem dist_self_sub_left (g h : E) : dist (g - h) g = ∥h∥ := -by rw [dist_comm, dist_self_sub_right] +section +open tactic tactic.positivity -/-- In a (semi)normed group, negation `x ↦ -x` tends to infinity at infinity. TODO: use -`bornology.cobounded` instead of `filter.comap has_norm.norm filter.at_top`. -/ -lemma filter.tendsto_neg_cobounded : - tendsto (has_neg.neg : E → E) (comap norm at_top) (comap norm at_top) := -by simpa only [norm_neg, tendsto_comap_iff, (∘)] using tendsto_comap +/-- Extension for the `positivity` tactic: norms are nonnegative. -/ +@[positivity] +meta def _root_.tactic.positivity_norm : expr → tactic strictness +| `(∥%%a∥) := nonnegative <$> mk_app ``norm_nonneg [a] <|> nonnegative <$> mk_app ``norm_nonneg' [a] +| _ := failed -/-- **Triangle inequality** for the norm. -/ -lemma norm_add_le (g h : E) : ∥g + h∥ ≤ ∥g∥ + ∥h∥ := -by simpa [dist_eq_norm] using dist_triangle g 0 (-h) +end -lemma norm_add_le_of_le {g₁ g₂ : E} {n₁ n₂ : ℝ} (H₁ : ∥g₁∥ ≤ n₁) (H₂ : ∥g₂∥ ≤ n₂) : - ∥g₁ + g₂∥ ≤ n₁ + n₂ := -le_trans (norm_add_le g₁ g₂) (add_le_add H₁ H₂) +@[simp, to_additive norm_zero] lemma norm_one' : ∥(1 : E)∥ = 0 := by rw [←dist_one_right, dist_self] -lemma norm_add₃_le (x y z : E) : ∥x + y + z∥ ≤ ∥x∥ + ∥y∥ + ∥z∥ := -norm_add_le_of_le (norm_add_le _ _) le_rfl +@[to_additive] lemma ne_one_of_norm_ne_zero : ∥a∥ ≠ 0 → a ≠ 1 := +mt $ by { rintro rfl, exact norm_one' } -lemma dist_add_add_le (g₁ g₂ h₁ h₂ : E) : - dist (g₁ + g₂) (h₁ + h₂) ≤ dist g₁ h₁ + dist g₂ h₂ := -by simpa only [dist_add_left, dist_add_right] using dist_triangle (g₁ + g₂) (h₁ + g₂) (h₁ + h₂) +@[nontriviality, to_additive norm_of_subsingleton] +lemma norm_of_subsingleton' [subsingleton E] (a : E) : ∥a∥ = 0 := +by rw [subsingleton.elim a 1, norm_one'] -lemma dist_add_add_le_of_le {g₁ g₂ h₁ h₂ : E} {d₁ d₂ : ℝ} - (H₁ : dist g₁ h₁ ≤ d₁) (H₂ : dist g₂ h₂ ≤ d₂) : - dist (g₁ + g₂) (h₁ + h₂) ≤ d₁ + d₂ := -le_trans (dist_add_add_le g₁ g₂ h₁ h₂) (add_le_add H₁ H₂) +attribute [nontriviality] norm_of_subsingleton -lemma dist_sub_sub_le (g₁ g₂ h₁ h₂ : E) : - dist (g₁ - g₂) (h₁ - h₂) ≤ dist g₁ h₁ + dist g₂ h₂ := -by simpa only [sub_eq_add_neg, dist_neg_neg] using dist_add_add_le g₁ (-g₂) h₁ (-h₂) +@[to_additive] lemma norm_div_le (a b : E) : ∥a / b∥ ≤ ∥a∥ + ∥b∥ := +by simpa [dist_eq_norm_div] using dist_triangle a 1 b -lemma dist_sub_sub_le_of_le {g₁ g₂ h₁ h₂ : E} {d₁ d₂ : ℝ} - (H₁ : dist g₁ h₁ ≤ d₁) (H₂ : dist g₂ h₂ ≤ d₂) : - dist (g₁ - g₂) (h₁ - h₂) ≤ d₁ + d₂ := -le_trans (dist_sub_sub_le g₁ g₂ h₁ h₂) (add_le_add H₁ H₂) +@[to_additive] lemma norm_div_le_of_le {r₁ r₂ : ℝ} (H₁ : ∥a₁∥ ≤ r₁) (H₂ : ∥a₂∥ ≤ r₂) : + ∥a₁ / a₂∥ ≤ r₁ + r₂ := +(norm_div_le a₁ a₂).trans $ add_le_add H₁ H₂ -lemma abs_dist_sub_le_dist_add_add (g₁ g₂ h₁ h₂ : E) : - |dist g₁ h₁ - dist g₂ h₂| ≤ dist (g₁ + g₂) (h₁ + h₂) := -by simpa only [dist_add_left, dist_add_right, dist_comm h₂] - using abs_dist_sub_le (g₁ + g₂) (h₁ + h₂) (h₁ + g₂) +@[to_additive] lemma dist_le_norm_mul_norm (a b : E) : dist a b ≤ ∥a∥ + ∥b∥ := +by { rw dist_eq_norm_div, apply norm_div_le } -@[simp] lemma norm_nonneg (g : E) : 0 ≤ ∥g∥ := -by { rw[←dist_zero_right], exact dist_nonneg } +@[to_additive abs_norm_sub_norm_le] lemma abs_norm_sub_norm_le' (a b : E) : |∥a∥ - ∥b∥| ≤ ∥a / b∥ := +by simpa [dist_eq_norm_div] using abs_dist_sub_le a b 1 -section -open tactic tactic.positivity +@[to_additive norm_sub_norm_le] lemma norm_sub_norm_le' (a b : E) : ∥a∥ - ∥b∥ ≤ ∥a / b∥ := +(le_abs_self _).trans (abs_norm_sub_norm_le' a b) -/-- Extension for the `positivity` tactic: norms are nonnegative. -/ -@[positivity] -meta def _root_.tactic.positivity_norm : expr → tactic strictness -| `(∥%%a∥) := nonnegative <$> mk_app ``norm_nonneg [a] -| _ := failed +@[to_additive dist_norm_norm_le] lemma dist_norm_norm_le' (a b : E) : dist ∥a∥ ∥b∥ ≤ ∥a / b∥ := +abs_norm_sub_norm_le' a b -end +@[to_additive] lemma norm_le_norm_add_norm_div' (u v : E) : ∥u∥ ≤ ∥v∥ + ∥u / v∥ := +by { rw add_comm, refine (norm_mul_le' _ _).trans_eq' _, rw div_mul_cancel' } -@[simp] lemma norm_zero : ∥(0 : E)∥ = 0 := by rw [← dist_zero_right, dist_self] +@[to_additive] lemma norm_le_norm_add_norm_div (u v : E) : ∥v∥ ≤ ∥u∥ + ∥u / v∥ := +by { rw norm_div_rev, exact norm_le_norm_add_norm_div' v u } -lemma ne_zero_of_norm_ne_zero {g : E} : ∥g∥ ≠ 0 → g ≠ 0 := mt $ by { rintro rfl, exact norm_zero } +alias norm_le_norm_add_norm_sub' ← norm_le_insert' +alias norm_le_norm_add_norm_sub ← norm_le_insert -@[nontriviality] lemma norm_of_subsingleton [subsingleton E] (x : E) : ∥x∥ = 0 := -by rw [subsingleton.elim x 0, norm_zero] +@[to_additive] lemma norm_le_mul_norm_add (u v : E) : ∥u∥ ≤ ∥u * v∥ + ∥v∥ := +calc ∥u∥ = ∥u * v / v∥ : by rw mul_div_cancel'' +... ≤ ∥u * v∥ + ∥v∥ : norm_div_le _ _ -lemma norm_multiset_sum_le (m : multiset E) : - ∥m.sum∥ ≤ (m.map (λ x, ∥x∥)).sum := -m.le_sum_of_subadditive norm norm_zero norm_add_le +@[to_additive ball_eq] lemma ball_eq' (y : E) (ε : ℝ) : ball y ε = {x | ∥x / y∥ < ε} := +set.ext $ λ a, by simp [dist_eq_norm_div] -lemma norm_sum_le (s : finset ι) (f : ι → E) : ∥∑ i in s, f i∥ ≤ ∑ i in s, ∥f i∥ := -s.le_sum_of_subadditive norm norm_zero norm_add_le f +@[to_additive] lemma ball_one_eq (r : ℝ) : ball (1 : E) r = {x | ∥x∥ < r} := +set.ext $ assume a, by simp -lemma norm_sum_le_of_le (s : finset ι) {f : ι → E} {n : ι → ℝ} (h : ∀ b ∈ s, ∥f b∥ ≤ n b) : - ∥∑ b in s, f b∥ ≤ ∑ b in s, n b := -le_trans (norm_sum_le s f) (finset.sum_le_sum h) +@[to_additive mem_ball_iff_norm] lemma mem_ball_iff_norm'' : b ∈ ball a r ↔ ∥b / a∥ < r := +by rw [mem_ball, dist_eq_norm_div] -lemma dist_sum_sum_le_of_le (s : finset ι) {f g : ι → E} {d : ι → ℝ} - (h : ∀ b ∈ s, dist (f b) (g b) ≤ d b) : - dist (∑ b in s, f b) (∑ b in s, g b) ≤ ∑ b in s, d b := -begin - simp only [dist_eq_norm, ← finset.sum_sub_distrib] at *, - exact norm_sum_le_of_le s h -end +@[to_additive mem_ball_iff_norm'] lemma mem_ball_iff_norm''' : b ∈ ball a r ↔ ∥a / b∥ < r := +by rw [mem_ball', dist_eq_norm_div] -lemma dist_sum_sum_le (s : finset ι) (f g : ι → E) : - dist (∑ b in s, f b) (∑ b in s, g b) ≤ ∑ b in s, dist (f b) (g b) := -dist_sum_sum_le_of_le s (λ _ _, le_rfl) +@[simp, to_additive] lemma mem_ball_one_iff : a ∈ ball (1 : E) r ↔ ∥a∥ < r := +by rw [mem_ball, dist_one_right] -lemma norm_sub_le (g h : E) : ∥g - h∥ ≤ ∥g∥ + ∥h∥ := -by simpa [dist_eq_norm] using dist_triangle g 0 h +@[to_additive mem_closed_ball_iff_norm] +lemma mem_closed_ball_iff_norm'' : b ∈ closed_ball a r ↔ ∥b / a∥ ≤ r := +by rw [mem_closed_ball, dist_eq_norm_div] -lemma norm_sub_le_of_le {g₁ g₂ : E} {n₁ n₂ : ℝ} (H₁ : ∥g₁∥ ≤ n₁) (H₂ : ∥g₂∥ ≤ n₂) : - ∥g₁ - g₂∥ ≤ n₁ + n₂ := -le_trans (norm_sub_le g₁ g₂) (add_le_add H₁ H₂) +@[simp, to_additive] lemma mem_closed_ball_one_iff : a ∈ closed_ball (1 : E) r ↔ ∥a∥ ≤ r := +by rw [mem_closed_ball, dist_one_right] -lemma dist_le_norm_add_norm (g h : E) : dist g h ≤ ∥g∥ + ∥h∥ := -by { rw dist_eq_norm, apply norm_sub_le } +@[to_additive mem_closed_ball_iff_norm'] +lemma mem_closed_ball_iff_norm''' : b ∈ closed_ball a r ↔ ∥a / b∥ ≤ r := +by rw [mem_closed_ball', dist_eq_norm_div] -lemma abs_norm_sub_norm_le (g h : E) : |∥g∥ - ∥h∥| ≤ ∥g - h∥ := -by simpa [dist_eq_norm] using abs_dist_sub_le g h 0 +@[to_additive norm_le_of_mem_closed_ball] +lemma norm_le_of_mem_closed_ball' (h : b ∈ closed_ball a r) : ∥b∥ ≤ ∥a∥ + r := +(norm_le_norm_add_norm_div' _ _).trans $ add_le_add_left (by rwa ←dist_eq_norm_div) _ -lemma norm_sub_norm_le (g h : E) : ∥g∥ - ∥h∥ ≤ ∥g - h∥ := -le_trans (le_abs_self _) (abs_norm_sub_norm_le g h) +@[to_additive norm_le_norm_add_const_of_dist_le] +lemma norm_le_norm_add_const_of_dist_le' : dist a b ≤ r → ∥a∥ ≤ ∥b∥ + r := +norm_le_of_mem_closed_ball' -lemma dist_norm_norm_le (g h : E) : dist ∥g∥ ∥h∥ ≤ ∥g - h∥ := -abs_norm_sub_norm_le g h +@[to_additive norm_lt_of_mem_ball] +lemma norm_lt_of_mem_ball' (h : b ∈ ball a r) : ∥b∥ < ∥a∥ + r := +(norm_le_norm_add_norm_div' _ _).trans_lt $ add_lt_add_left (by rwa ←dist_eq_norm_div) _ -/-- The direct path from `0` to `v` is shorter than the path with `u` inserted in between. -/ -lemma norm_le_insert (u v : E) : ∥v∥ ≤ ∥u∥ + ∥u - v∥ := -calc ∥v∥ = ∥u - (u - v)∥ : by abel -... ≤ ∥u∥ + ∥u - v∥ : norm_sub_le u _ +@[to_additive bounded_iff_forall_norm_le] +lemma bounded_iff_forall_norm_le' : bounded s ↔ ∃ C, ∀ x ∈ s, ∥x∥ ≤ C := +by simpa only [set.subset_def, mem_closed_ball_one_iff] using bounded_iff_subset_ball (1 : E) -lemma norm_le_insert' (u v : E) : ∥u∥ ≤ ∥v∥ + ∥u - v∥ := -by { rw norm_sub_rev, exact norm_le_insert v u } +@[simp, to_additive mem_sphere_iff_norm] +lemma mem_sphere_iff_norm' : b ∈ sphere a r ↔ ∥b / a∥ = r := by simp [dist_eq_norm_div] -lemma norm_le_add_norm_add (u v : E) : - ∥u∥ ≤ ∥u + v∥ + ∥v∥ := -calc ∥u∥ = ∥u + v - v∥ : by rw add_sub_cancel -... ≤ ∥u + v∥ + ∥v∥ : norm_sub_le _ _ +@[simp, to_additive] lemma mem_sphere_one_iff_norm : a ∈ sphere (1 : E) r ↔ ∥a∥ = r := +by simp [dist_eq_norm_div] -lemma ball_eq (y : E) (ε : ℝ) : metric.ball y ε = { x | ∥x - y∥ < ε} := -by { ext, simp [dist_eq_norm], } +@[simp, to_additive norm_eq_of_mem_sphere] +lemma norm_eq_of_mem_sphere' (x : sphere (1:E) r) : ∥(x : E)∥ = r := mem_sphere_one_iff_norm.mp x.2 -lemma ball_zero_eq (ε : ℝ) : ball (0 : E) ε = {x | ∥x∥ < ε} := -set.ext $ assume a, by simp +@[to_additive] lemma ne_one_of_mem_sphere (hr : r ≠ 0) (x : sphere (1 : E) r) : (x : E) ≠ 1 := +ne_one_of_norm_ne_zero $ by rwa norm_eq_of_mem_sphere' x -lemma mem_ball_iff_norm {g h : E} {r : ℝ} : - h ∈ ball g r ↔ ∥h - g∥ < r := -by rw [mem_ball, dist_eq_norm] +@[to_additive ne_zero_of_mem_unit_sphere] +lemma ne_one_of_mem_unit_sphere (x : sphere (1 : E) 1) : (x:E) ≠ 1 := +ne_one_of_mem_sphere one_ne_zero _ -lemma add_mem_ball_iff_norm {g h : E} {r : ℝ} : - g + h ∈ ball g r ↔ ∥h∥ < r := -by rw [mem_ball_iff_norm, add_sub_cancel'] +variables (E) -lemma mem_ball_iff_norm' {g h : E} {r : ℝ} : - h ∈ ball g r ↔ ∥g - h∥ < r := -by rw [mem_ball', dist_eq_norm] +/-- The norm of a seminormed group as a group seminorm. -/ +@[to_additive "The norm of a seminormed group as an additive group seminorm."] +def norm_group_seminorm : group_seminorm E := ⟨norm, norm_one', norm_mul_le', norm_inv'⟩ -@[simp] lemma mem_ball_zero_iff {ε : ℝ} {x : E} : x ∈ ball (0 : E) ε ↔ ∥x∥ < ε := -by rw [mem_ball, dist_zero_right] +@[simp, to_additive] lemma coe_norm_group_seminorm : ⇑(norm_group_seminorm E) = norm := rfl -lemma mem_closed_ball_iff_norm {g h : E} {r : ℝ} : - h ∈ closed_ball g r ↔ ∥h - g∥ ≤ r := -by rw [mem_closed_ball, dist_eq_norm] +variables {E} -@[simp] lemma mem_closed_ball_zero_iff {ε : ℝ} {x : E} : x ∈ closed_ball (0 : E) ε ↔ ∥x∥ ≤ ε := -by rw [mem_closed_ball, dist_zero_right] +namespace isometric +-- TODO This material is superseded by similar constructions such as +-- `affine_isometry_equiv.const_vadd`; deduplicate -lemma add_mem_closed_ball_iff_norm {g h : E} {r : ℝ} : - g + h ∈ closed_ball g r ↔ ∥h∥ ≤ r := -by rw [mem_closed_ball_iff_norm, add_sub_cancel'] +/-- Multiplication `y ↦ y * x` as an `isometry`. -/ +@[to_additive "Addition `y ↦ y + x` as an `isometry`"] +protected def mul_right (x : E) : E ≃ᵢ E := +{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_mul_right _ _ _, + .. equiv.mul_right x } -lemma mem_closed_ball_iff_norm' {g h : E} {r : ℝ} : - h ∈ closed_ball g r ↔ ∥g - h∥ ≤ r := -by rw [mem_closed_ball', dist_eq_norm] +@[simp, to_additive] +lemma mul_right_to_equiv (x : E) : (isometric.mul_right x).to_equiv = equiv.mul_right x := rfl -lemma norm_le_of_mem_closed_ball {g h : E} {r : ℝ} (H : h ∈ closed_ball g r) : - ∥h∥ ≤ ∥g∥ + r := -calc - ∥h∥ = ∥g + (h - g)∥ : by rw [add_sub_cancel'_right] - ... ≤ ∥g∥ + ∥h - g∥ : norm_add_le _ _ - ... ≤ ∥g∥ + r : by { apply add_le_add_left, rw ← dist_eq_norm, exact H } +@[simp, to_additive] +lemma coe_mul_right (x : E) : (isometric.mul_right x : E → E) = λ y, y * x := rfl -lemma norm_le_norm_add_const_of_dist_le {a b : E} {c : ℝ} (h : dist a b ≤ c) : - ∥a∥ ≤ ∥b∥ + c := -norm_le_of_mem_closed_ball h +@[to_additive] lemma mul_right_apply (x y : E) : (isometric.mul_right x : E → E) y = y * x := rfl -lemma norm_lt_of_mem_ball {g h : E} {r : ℝ} (H : h ∈ ball g r) : - ∥h∥ < ∥g∥ + r := -calc - ∥h∥ = ∥g + (h - g)∥ : by rw [add_sub_cancel'_right] - ... ≤ ∥g∥ + ∥h - g∥ : norm_add_le _ _ - ... < ∥g∥ + r : by { apply add_lt_add_left, rw ← dist_eq_norm, exact H } +@[simp, to_additive] +lemma mul_right_symm (x : E) : (isometric.mul_right x).symm = isometric.mul_right x⁻¹ := +ext $ λ y, rfl -lemma norm_lt_norm_add_const_of_dist_lt {a b : E} {c : ℝ} (h : dist a b < c) : - ∥a∥ < ∥b∥ + c := -norm_lt_of_mem_ball h +end isometric -lemma bounded_iff_forall_norm_le {s : set E} : bounded s ↔ ∃ C, ∀ x ∈ s, ∥x∥ ≤ C := -by simpa only [set.subset_def, mem_closed_ball_iff_norm, sub_zero] - using bounded_iff_subset_ball (0 : E) +@[to_additive] lemma normed_comm_group.tendsto_nhds_one {f : α → E} {l : filter α} : + tendsto f l (𝓝 1) ↔ ∀ ε > 0, ∀ᶠ x in l, ∥ f x ∥ < ε := +metric.tendsto_nhds.trans $ by simp only [dist_one_right] -@[simp] lemma preimage_add_ball (x y : E) (r : ℝ) : ((+) y) ⁻¹' (ball x r) = ball (x - y) r := -begin - ext z, - simp only [dist_eq_norm, set.mem_preimage, mem_ball], - abel -end +@[to_additive] lemma normed_comm_group.tendsto_nhds_nhds {f : E → F} {x : E} {y : F} : + tendsto f (𝓝 x) (𝓝 y) ↔ ∀ ε > 0, ∃ δ > 0, ∀ x', ∥x' / x∥ < δ → ∥f x' / y∥ < ε := +by simp_rw [metric.tendsto_nhds_nhds, dist_eq_norm_div] -@[simp] lemma preimage_add_closed_ball (x y : E) (r : ℝ) : - ((+) y) ⁻¹' (closed_ball x r) = closed_ball (x - y) r := -begin - ext z, - simp only [dist_eq_norm, set.mem_preimage, mem_closed_ball], - abel -end +@[to_additive] lemma normed_comm_group.cauchy_seq_iff [nonempty α] [semilattice_sup α] {u : α → E} : + cauchy_seq u ↔ ∀ ε > 0, ∃ N, ∀ m, N ≤ m → ∀ n, N ≤ n → ∥u m / u n∥ < ε := +by simp [metric.cauchy_seq_iff, dist_eq_norm_div] + +@[to_additive] lemma normed_comm_group.nhds_basis_norm_lt (x : E) : + (𝓝 x).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {y | ∥y / x∥ < ε}) := +by { simp_rw ← ball_eq', exact metric.nhds_basis_ball } + +@[to_additive] lemma normed_comm_group.nhds_one_basis_norm_lt : + (𝓝 (1 : E)).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {y | ∥y∥ < ε}) := +by { convert normed_comm_group.nhds_basis_norm_lt (1 : E), simp } + +@[to_additive] lemma normed_comm_group.uniformity_basis_dist : + (𝓤 E).has_basis (λ ε : ℝ, 0 < ε) (λ ε, {p : E × E | ∥p.fst / p.snd∥ < ε}) := +by { convert metric.uniformity_basis_dist, simp [dist_eq_norm_div] } + +open finset + +/-- A homomorphism `f` of seminormed groups is Lipschitz, if there exists a constant `C` such that +for all `x`, one has `∥f x∥ ≤ C * ∥x∥`. The analogous condition for a linear map of +(semi)normed spaces is in `normed_space.operator_norm`. -/ +@[to_additive "A homomorphism `f` of seminormed groups is Lipschitz, if there exists a constant `C` +such that for all `x`, one has `∥f x∥ ≤ C * ∥x∥`. The analogous condition for a linear map of +(semi)normed spaces is in `normed_space.operator_norm`."] +lemma monoid_hom_class.lipschitz_of_bound [monoid_hom_class 𝓕 E F] (f : 𝓕) (C : ℝ) + (h : ∀ x, ∥f x∥ ≤ C * ∥x∥) : lipschitz_with (real.to_nnreal C) f := +lipschitz_with.of_dist_le' $ λ x y, by simpa only [dist_eq_norm_div, map_div] using h (x / y) -@[simp] lemma mem_sphere_iff_norm (v w : E) (r : ℝ) : w ∈ sphere v r ↔ ∥w - v∥ = r := -by simp [dist_eq_norm] +@[to_additive] lemma lipschitz_on_with_iff_norm_div_le {f : E → F} {C : ℝ≥0} : + lipschitz_on_with C f s ↔ ∀ ⦃x⦄, x ∈ s → ∀ ⦃y⦄, y ∈ s → ∥f x / f y∥ ≤ C * ∥x / y∥ := +by simp only [lipschitz_on_with_iff_dist_le_mul, dist_eq_norm_div] -@[simp] lemma mem_sphere_zero_iff_norm {w : E} {r : ℝ} : w ∈ sphere (0:E) r ↔ ∥w∥ = r := -by simp [dist_eq_norm] +alias lipschitz_on_with_iff_norm_div_le ↔ lipschitz_on_with.norm_div_le _ -@[simp] lemma norm_eq_of_mem_sphere {r : ℝ} (x : sphere (0:E) r) : ∥(x:E)∥ = r := -mem_sphere_zero_iff_norm.mp x.2 +attribute [to_additive] lipschitz_on_with.norm_div_le + +@[to_additive] lemma lipschitz_on_with.norm_div_le_of_le {f : E → F} {C : ℝ≥0} + (h : lipschitz_on_with C f s) (ha : a ∈ s) (hb : b ∈ s) (hr : ∥a / b∥ ≤ r) : + ∥f a / f b∥ ≤ C * r := +(h.norm_div_le ha hb).trans $ mul_le_mul_of_nonneg_left hr C.2 + +@[to_additive] lemma lipschitz_with_iff_norm_div_le {f : E → F} {C : ℝ≥0} : + lipschitz_with C f ↔ ∀ x y, ∥f x / f y∥ ≤ C * ∥x / y∥ := +by simp only [lipschitz_with_iff_dist_le_mul, dist_eq_norm_div] + +alias lipschitz_with_iff_norm_div_le ↔ lipschitz_with.norm_div_le _ + +attribute [to_additive] lipschitz_with.norm_div_le + +@[to_additive] lemma lipschitz_with.norm_div_le_of_le {f : E → F} {C : ℝ≥0} (h : lipschitz_with C f) + (hr : ∥a / b∥ ≤ r) : ∥f a / f b∥ ≤ C * r := +(h.norm_div_le _ _).trans $ mul_le_mul_of_nonneg_left hr C.2 + +/-- A homomorphism `f` of seminormed groups is continuous, if there exists a constant `C` such that +for all `x`, one has `∥f x∥ ≤ C * ∥x∥`. -/ +@[to_additive "A homomorphism `f` of seminormed groups is continuous, if there exists a constant `C` +such that for all `x`, one has `∥f x∥ ≤ C * ∥x∥`"] +lemma monoid_hom_class.continuous_of_bound [monoid_hom_class 𝓕 E F] (f : 𝓕) (C : ℝ) + (h : ∀ x, ∥f x∥ ≤ C * ∥x∥) : continuous f := +(monoid_hom_class.lipschitz_of_bound f C h).continuous + +@[to_additive] lemma monoid_hom_class.uniform_continuous_of_bound [monoid_hom_class 𝓕 E F] + (f : 𝓕) (C : ℝ) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : uniform_continuous f := +(monoid_hom_class.lipschitz_of_bound f C h).uniform_continuous + +@[to_additive is_compact.exists_bound_of_continuous_on] +lemma is_compact.exists_bound_of_continuous_on' [topological_space α] {s : set α} + (hs : is_compact s) {f : α → E} (hf : continuous_on f s) : + ∃ C, ∀ x ∈ s, ∥f x∥ ≤ C := +(bounded_iff_forall_norm_le'.1 (hs.image_of_continuous_on hf).bounded).imp $ λ C hC x hx, + hC _ $ set.mem_image_of_mem _ hx -lemma preimage_add_sphere (x y : E) (r : ℝ) : - ((+) y) ⁻¹' (sphere x r) = sphere (x - y) r := +@[to_additive] lemma monoid_hom_class.isometry_iff_norm [monoid_hom_class 𝓕 E F] (f : 𝓕) : + isometry f ↔ ∀ x, ∥f x∥ = ∥x∥ := begin - ext z, - simp only [set.mem_preimage, mem_sphere_iff_norm], - abel + simp only [isometry_iff_dist_eq, dist_eq_norm_div, ←map_div], + refine ⟨λ h x, _, λ h x y, h _⟩, + simpa using h x 1, end -lemma ne_zero_of_mem_sphere {r : ℝ} (hr : r ≠ 0) (x : sphere (0 : E) r) : (x : E) ≠ 0 := -ne_zero_of_norm_ne_zero $ by rwa norm_eq_of_mem_sphere x +alias monoid_hom_class.isometry_iff_norm ↔ _ monoid_hom_class.isometry_of_norm -lemma ne_zero_of_mem_unit_sphere (x : sphere (0:E) 1) : (x:E) ≠ 0 := -ne_zero_of_mem_sphere one_ne_zero _ +attribute [to_additive] monoid_hom_class.isometry_of_norm -variables (E) +section nnnorm -/-- The norm of a seminormed group as an additive group seminorm. -/ -def norm_add_group_seminorm : add_group_seminorm E := ⟨norm, norm_zero, norm_add_le, norm_neg⟩ +@[priority 100, to_additive] -- See note [lower instance priority] +instance seminormed_group.to_has_nnnorm : has_nnnorm E := ⟨λ a, ⟨∥a∥, norm_nonneg' a⟩⟩ -@[simp] lemma coe_norm_add_group_seminorm : ⇑(norm_add_group_seminorm E) = norm := rfl +@[simp, norm_cast, to_additive coe_nnnorm] lemma coe_nnnorm' (a : E) : (∥a∥₊ : ℝ) = ∥a∥ := rfl -variables {E} +@[simp, to_additive coe_comp_nnnorm] +lemma coe_comp_nnnorm' : (coe : ℝ≥0 → ℝ) ∘ (nnnorm : E → ℝ≥0) = norm := rfl -namespace isometric --- TODO This material is superseded by similar constructions such as --- `affine_isometry_equiv.const_vadd`; deduplicate +@[to_additive norm_to_nnreal] +lemma norm_to_nnreal' : ∥a∥.to_nnreal = ∥a∥₊ := @real.to_nnreal_coe ∥a∥₊ -/-- Addition `y ↦ y + x` as an `isometry`. -/ -protected def add_right (x : E) : E ≃ᵢ E := -{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_add_right _ _ _, - .. equiv.add_right x } +@[to_additive] +lemma nndist_eq_nnnorm_div (a b : E) : nndist a b = ∥a / b∥₊ := nnreal.eq $ dist_eq_norm_div _ _ -@[simp] lemma add_right_to_equiv (x : E) : - (isometric.add_right x).to_equiv = equiv.add_right x := rfl +alias nndist_eq_nnnorm_sub ← nndist_eq_nnnorm -@[simp] lemma coe_add_right (x : E) : (isometric.add_right x : E → E) = λ y, y + x := rfl +@[simp, to_additive nnnorm_zero] lemma nnnorm_one' : ∥(1 : E)∥₊ = 0 := nnreal.eq norm_one' -lemma add_right_apply (x y : E) : (isometric.add_right x : E → E) y = y + x := rfl +@[to_additive] lemma ne_one_of_nnnorm_ne_zero {a : E} : ∥a∥₊ ≠ 0 → a ≠ 1 := +mt $ by { rintro rfl, exact nnnorm_one' } -@[simp] lemma add_right_symm (x : E) : - (isometric.add_right x).symm = isometric.add_right (-x) := -ext $ λ y, rfl +@[to_additive nnnorm_add_le] +lemma nnnorm_mul_le' (a b : E) : ∥a * b∥₊ ≤ ∥a∥₊ + ∥b∥₊ := nnreal.coe_le_coe.1 $ norm_mul_le' a b -/-- Addition `y ↦ x + y` as an `isometry`. -/ -protected def add_left (x : E) : E ≃ᵢ E := -{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_add_left _ _ _, - to_equiv := equiv.add_left x } +@[simp, to_additive nnnorm_neg] lemma nnnorm_inv' (a : E) : ∥a⁻¹∥₊ = ∥a∥₊ := nnreal.eq $ norm_inv' a -@[simp] lemma add_left_to_equiv (x : E) : - (isometric.add_left x).to_equiv = equiv.add_left x := rfl +@[to_additive] lemma nnnorm_div_le (a b : E) : ∥a / b∥₊ ≤ ∥a∥₊ + ∥b∥₊ := +nnreal.coe_le_coe.1 $ norm_div_le _ _ -@[simp] lemma coe_add_left (x : E) : ⇑(isometric.add_left x) = (+) x := rfl +@[to_additive nndist_nnnorm_nnnorm_le] +lemma nndist_nnnorm_nnnorm_le' (a b : E) : nndist ∥a∥₊ ∥b∥₊ ≤ ∥a / b∥₊ := +nnreal.coe_le_coe.1 $ dist_norm_norm_le' a b -@[simp] lemma add_left_symm (x : E) : - (isometric.add_left x).symm = isometric.add_left (-x) := -ext $ λ y, rfl +@[to_additive] lemma nnnorm_le_nnnorm_add_nnnorm_div (a b : E) : ∥b∥₊ ≤ ∥a∥₊ + ∥a / b∥₊ := +norm_le_norm_add_norm_div _ _ -variable (E) +@[to_additive] lemma nnnorm_le_nnnorm_add_nnnorm_div' (a b : E) : ∥a∥₊ ≤ ∥b∥₊ + ∥a / b∥₊ := +norm_le_norm_add_norm_div' _ _ -/-- Negation `x ↦ -x` as an `isometry`. -/ -protected def neg : E ≃ᵢ E := -{ isometry_to_fun := isometry.of_dist_eq $ λ x y, dist_neg_neg _ _, - to_equiv := equiv.neg E } +alias nnnorm_le_nnnorm_add_nnnorm_sub' ← nnnorm_le_insert' +alias nnnorm_le_nnnorm_add_nnnorm_sub ← nnnorm_le_insert -variable {E} +@[to_additive] +lemma nnnorm_le_mul_nnnorm_add (a b : E) : ∥a∥₊ ≤ ∥a * b∥₊ + ∥b∥₊ := norm_le_mul_norm_add _ _ -@[simp] lemma neg_symm : (isometric.neg E).symm = isometric.neg E := rfl +@[to_additive of_real_norm_eq_coe_nnnorm] +lemma of_real_norm_eq_coe_nnnorm' (a : E) : ennreal.of_real ∥a∥ = ∥a∥₊ := +ennreal.of_real_eq_coe_nnreal _ -@[simp] lemma neg_to_equiv : (isometric.neg E).to_equiv = equiv.neg E := rfl +@[to_additive] lemma edist_eq_coe_nnnorm_div (a b : E) : edist a b = ∥a / b∥₊ := +by rw [edist_dist, dist_eq_norm_div, of_real_norm_eq_coe_nnnorm'] -@[simp] lemma coe_neg : ⇑(isometric.neg E) = has_neg.neg := rfl +@[to_additive edist_eq_coe_nnnorm] lemma edist_eq_coe_nnnorm' (x : E) : edist x 1 = (∥x∥₊ : ℝ≥0∞) := +by rw [edist_eq_coe_nnnorm_div, div_one] -end isometric +@[to_additive] +lemma mem_emetric_ball_one_iff {r : ℝ≥0∞} : a ∈ emetric.ball (1 : E) r ↔ ↑∥a∥₊ < r := +by rw [emetric.mem_ball, edist_eq_coe_nnnorm'] + +@[simp, to_additive] lemma edist_mul_right (a₁ a₂ b : E) : edist (a₁ * b) (a₂ * b) = edist a₁ a₂ := +by simp [edist_dist] + +@[simp, to_additive] lemma edist_div_right (a₁ a₂ b : E) : edist (a₁ / b) (a₂ / b) = edist a₁ a₂ := +by simpa only [div_eq_mul_inv] using edist_mul_right _ _ _ + +@[to_additive] lemma monoid_hom_class.lipschitz_of_bound_nnnorm [monoid_hom_class 𝓕 E F] (f : 𝓕) + (C : ℝ≥0) (h : ∀ x, ∥f x∥₊ ≤ C * ∥x∥₊) : lipschitz_with C f := +@real.to_nnreal_coe C ▸ monoid_hom_class.lipschitz_of_bound f C h + +@[to_additive] lemma monoid_hom_class.antilipschitz_of_bound [monoid_hom_class 𝓕 E F] (f : 𝓕) + {K : ℝ≥0} (h : ∀ x, ∥x∥ ≤ K * ∥f x∥) : + antilipschitz_with K f := +antilipschitz_with.of_le_mul_dist $ λ x y, by simpa only [dist_eq_norm_div, map_div] using h (x / y) + +@[to_additive] lemma monoid_hom_class.bound_of_antilipschitz [monoid_hom_class 𝓕 E F] (f : 𝓕) + {K : ℝ≥0} (h : antilipschitz_with K f) (x) : ∥x∥ ≤ K * ∥f x∥ := +by simpa only [dist_one_right, map_one] using h.le_mul_dist x 1 + +end nnnorm -theorem normed_add_comm_group.tendsto_nhds_zero {f : α → E} {l : filter α} : - tendsto f l (𝓝 0) ↔ ∀ ε > 0, ∀ᶠ x in l, ∥ f x ∥ < ε := -metric.tendsto_nhds.trans $ by simp only [dist_zero_right] +@[to_additive] lemma tendsto_iff_norm_tendsto_one {f : α → E} {a : filter α} {b : E} : + tendsto f a (𝓝 b) ↔ tendsto (λ e, ∥f e / b∥) a (𝓝 0) := +by { convert tendsto_iff_dist_tendsto_zero, simp [dist_eq_norm_div] } -lemma normed_add_comm_group.tendsto_nhds_nhds {f : E → F} {x : E} {y : F} : - tendsto f (𝓝 x) (𝓝 y) ↔ ∀ ε > 0, ∃ δ > 0, ∀ x', ∥x' - x∥ < δ → ∥f x' - y∥ < ε := -by simp_rw [metric.tendsto_nhds_nhds, dist_eq_norm] +@[to_additive] lemma tendsto_one_iff_norm_tendsto_one {f : α → E} {a : filter α} : + tendsto f a (𝓝 1) ↔ tendsto (λ e, ∥f e∥) a (𝓝 0) := +by { rw tendsto_iff_norm_tendsto_one, simp only [div_one] } -lemma normed_add_comm_group.cauchy_seq_iff [nonempty α] [semilattice_sup α] {u : α → E} : - cauchy_seq u ↔ ∀ ε > 0, ∃ N, ∀ m, N ≤ m → ∀ n, N ≤ n → ∥u m - u n∥ < ε := -by simp [metric.cauchy_seq_iff, dist_eq_norm] +@[to_additive] lemma comap_norm_nhds_one : comap norm (𝓝 0) = 𝓝 (1 : E) := +by simpa only [dist_one_right] using nhds_comap_dist (1 : E) -lemma normed_add_comm_group.nhds_basis_norm_lt (x : E) : - (𝓝 x).has_basis (λ (ε : ℝ), 0 < ε) (λ (ε : ℝ), { y | ∥y - x∥ < ε }) := +/-- Special case of the sandwich theorem: if the norm of `f` is eventually bounded by a real +function `a` which tends to `0`, then `f` tends to `1`. In this pair of lemmas (`squeeze_one_norm'` +and `squeeze_one_norm`), following a convention of similar lemmas in `topology.metric_space.basic` +and `topology.algebra.order`, the `'` version is phrased using "eventually" and the non-`'` version +is phrased absolutely. -/ +@[to_additive "Special case of the sandwich theorem: if the norm of `f` is eventually bounded by a +real function `a` which tends to `0`, then `f` tends to `1`. In this pair of lemmas +(`squeeze_zero_norm'` and `squeeze_zero_norm`), following a convention of similar lemmas in +`topology.metric_space.basic` and `topology.algebra.order`, the `'` version is phrased using +\"eventually\" and the non-`'` version is phrased absolutely."] +lemma squeeze_one_norm' {f : α → E} {a : α → ℝ} {t₀ : filter α} (h : ∀ᶠ n in t₀, ∥f n∥ ≤ a n) + (h' : tendsto a t₀ (𝓝 0)) : tendsto f t₀ (𝓝 1) := +tendsto_one_iff_norm_tendsto_one.2 $ squeeze_zero' (eventually_of_forall $ λ n, norm_nonneg' _) h h' + +/-- Special case of the sandwich theorem: if the norm of `f` is bounded by a real function `a` which +tends to `0`, then `f` tends to `1`. -/ +@[to_additive "Special case of the sandwich theorem: if the norm of `f` is bounded by a real +function `a` which tends to `0`, then `f` tends to `0`."] +lemma squeeze_one_norm {f : α → E} {a : α → ℝ} {t₀ : filter α} (h : ∀ n, ∥f n∥ ≤ a n) : + tendsto a t₀ (𝓝 0) → tendsto f t₀ (𝓝 1) := +squeeze_one_norm' $ eventually_of_forall h + +@[to_additive] lemma tendsto_norm_div_self (x : E) : tendsto (λ a, ∥a / x∥) (𝓝 x) (𝓝 0) := +by simpa [dist_eq_norm_div] using + tendsto_id.dist (tendsto_const_nhds : tendsto (λ a, (x:E)) (𝓝 x) _) + +@[to_additive tendsto_norm]lemma tendsto_norm' {x : E} : tendsto (λ a, ∥a∥) (𝓝 x) (𝓝 ∥x∥) := +by simpa using tendsto_id.dist (tendsto_const_nhds : tendsto (λ a, (1:E)) _ _) + +@[to_additive] lemma tendsto_norm_one : tendsto (λ a : E, ∥a∥) (𝓝 1) (𝓝 0) := +by simpa using tendsto_norm_div_self (1:E) + +@[continuity, to_additive continuous_norm] lemma continuous_norm' : continuous (λ a : E, ∥a∥) := +by simpa using continuous_id.dist (continuous_const : continuous (λ a, (1:E))) + +@[continuity, to_additive continuous_nnnorm] +lemma continuous_nnnorm' : continuous (λ a : E, ∥a∥₊) := continuous_norm'.subtype_mk _ + +@[to_additive lipschitz_with_one_norm] lemma lipschitz_with_one_norm' : + lipschitz_with 1 (norm : E → ℝ) := +by simpa only [dist_one_left] using lipschitz_with.dist_right (1 : E) + +@[to_additive lipschitz_with_one_nnnorm] lemma lipschitz_with_one_nnnorm' : + lipschitz_with 1 (has_nnnorm.nnnorm : E → ℝ≥0) := +lipschitz_with_one_norm' + +@[to_additive uniform_continuous_norm] +lemma uniform_continuous_norm' : uniform_continuous (norm : E → ℝ) := +lipschitz_with_one_norm'.uniform_continuous + +@[to_additive uniform_continuous_nnnorm] +lemma uniform_continuous_nnnorm' : uniform_continuous (λ (a : E), ∥a∥₊) := +uniform_continuous_norm'.subtype_mk _ + +/-- A helper lemma used to prove that the (scalar or usual) product of a function that tends to one +and a bounded function tends to one. This lemma is formulated for any binary operation +`op : E → F → G` with an estimate `∥op x y∥ ≤ A * ∥x∥ * ∥y∥` for some constant A instead of +multiplication so that it can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`. -/ +@[to_additive "A helper lemma used to prove that the (scalar or usual) product of a function that +tends to zero and a bounded function tends to zero. This lemma is formulated for any binary +operation `op : E → F → G` with an estimate `∥op x y∥ ≤ A * ∥x∥ * ∥y∥` for some constant A instead +of multiplication so that it can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`."] +lemma filter.tendsto.op_one_is_bounded_under_le' {f : α → E} {g : α → F} {l : filter α} + (hf : tendsto f l (𝓝 1)) (hg : is_bounded_under (≤) l (norm ∘ g)) (op : E → F → G) + (h_op : ∃ A, ∀ x y, ∥op x y∥ ≤ A * ∥x∥ * ∥y∥) : + tendsto (λ x, op (f x) (g x)) l (𝓝 1) := begin - simp_rw ← ball_eq, - exact metric.nhds_basis_ball, + cases h_op with A h_op, + rcases hg with ⟨C, hC⟩, rw eventually_map at hC, + rw normed_comm_group.tendsto_nhds_one at hf ⊢, + intros ε ε₀, + rcases exists_pos_mul_lt ε₀ (A * C) with ⟨δ, δ₀, hδ⟩, + filter_upwards [hf δ δ₀, hC] with i hf hg, + refine (h_op _ _).trans_lt _, + cases le_total A 0 with hA hA, + { exact (mul_nonpos_of_nonpos_of_nonneg (mul_nonpos_of_nonpos_of_nonneg hA $ norm_nonneg' _) $ + norm_nonneg' _).trans_lt ε₀ }, + calc A * ∥f i∥ * ∥g i∥ ≤ A * δ * C : + mul_le_mul (mul_le_mul_of_nonneg_left hf.le hA) hg (norm_nonneg' _) (mul_nonneg hA δ₀.le) + ... = A * C * δ : mul_right_comm _ _ _ + ... < ε : hδ, end -lemma normed_add_comm_group.nhds_zero_basis_norm_lt : - (𝓝 (0 : E)).has_basis (λ (ε : ℝ), 0 < ε) (λ (ε : ℝ), { y | ∥y∥ < ε }) := -begin - convert normed_add_comm_group.nhds_basis_norm_lt (0 : E), - simp, +/-- A helper lemma used to prove that the (scalar or usual) product of a function that tends to one +and a bounded function tends to one. This lemma is formulated for any binary operation +`op : E → F → G` with an estimate `∥op x y∥ ≤ ∥x∥ * ∥y∥` instead of multiplication so that it +can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`. -/ +@[to_additive "A helper lemma used to prove that the (scalar or usual) product of a function that +tends to zero and a bounded function tends to zero. This lemma is formulated for any binary +operation `op : E → F → G` with an estimate `∥op x y∥ ≤ ∥x∥ * ∥y∥` instead of multiplication so that +it can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`."] +lemma filter.tendsto.op_one_is_bounded_under_le {f : α → E} {g : α → F} {l : filter α} + (hf : tendsto f l (𝓝 1)) (hg : is_bounded_under (≤) l (norm ∘ g)) (op : E → F → G) + (h_op : ∀ x y, ∥op x y∥ ≤ ∥x∥ * ∥y∥) : + tendsto (λ x, op (f x) (g x)) l (𝓝 1) := +hf.op_one_is_bounded_under_le' hg op ⟨1, λ x y, (one_mul (∥x∥)).symm ▸ h_op x y⟩ + +section +variables {l : filter α} {f : α → E} + +@[to_additive filter.tendsto.norm] lemma filter.tendsto.norm' (h : tendsto f l (𝓝 a)) : + tendsto (λ x, ∥f x∥) l (𝓝 ∥a∥) := +tendsto_norm'.comp h + +@[to_additive filter.tendsto.nnnorm] lemma filter.tendsto.nnnorm' (h : tendsto f l (𝓝 a)) : + tendsto (λ x, ∥f x∥₊) l (𝓝 (∥a∥₊)) := +tendsto.comp continuous_nnnorm'.continuous_at h + end -lemma normed_add_comm_group.uniformity_basis_dist : - (𝓤 E).has_basis (λ (ε : ℝ), 0 < ε) (λ ε, {p : E × E | ∥p.fst - p.snd∥ < ε}) := +section +variables [topological_space α] {f : α → E} + +@[to_additive continuous.norm] lemma continuous.norm' : continuous f → continuous (λ x, ∥f x∥) := +continuous_norm'.comp + +@[to_additive continuous.nnnorm] +lemma continuous.nnnorm' : continuous f → continuous (λ x, ∥f x∥₊) := continuous_nnnorm'.comp + +@[to_additive continuous_at.norm] +lemma continuous_at.norm' {a : α} (h : continuous_at f a) : continuous_at (λ x, ∥f x∥) a := h.norm' + +@[to_additive continuous_at.nnnorm] +lemma continuous_at.nnnorm' {a : α} (h : continuous_at f a) : continuous_at (λ x, ∥f x∥₊) a := +h.nnnorm' + +@[to_additive continuous_within_at.norm] +lemma continuous_within_at.norm' {s : set α} {a : α} (h : continuous_within_at f s a) : + continuous_within_at (λ x, ∥f x∥) s a := +h.norm' + +@[to_additive continuous_within_at.nnnorm] +lemma continuous_within_at.nnnorm' {s : set α} {a : α} (h : continuous_within_at f s a) : + continuous_within_at (λ x, ∥f x∥₊) s a := +h.nnnorm' + +@[to_additive continuous_on.norm] +lemma continuous_on.norm' {s : set α} (h : continuous_on f s) : continuous_on (λ x, ∥f x∥) s := +λ x hx, (h x hx).norm' + +@[to_additive continuous_on.nnnorm] +lemma continuous_on.nnnorm' {s : set α} (h : continuous_on f s) : continuous_on (λ x, ∥f x∥₊) s := +λ x hx, (h x hx).nnnorm' + +end + +/-- If `∥y∥ → ∞`, then we can assume `y ≠ x` for any fixed `x`. -/ +@[to_additive eventually_ne_of_tendsto_norm_at_top "If `∥y∥→∞`, then we can assume `y≠x` for any +fixed `x`"] +lemma eventually_ne_of_tendsto_norm_at_top' {l : filter α} {f : α → E} + (h : tendsto (λ y, ∥f y∥) l at_top) (x : E) : + ∀ᶠ y in l, f y ≠ x := +(h.eventually_ne_at_top _).mono $ λ x, ne_of_apply_ne norm + +@[to_additive] lemma seminormed_comm_group.mem_closure_iff : + a ∈ closure s ↔ ∀ ε, 0 < ε → ∃ b ∈ s, ∥a / b∥ < ε := +by simp [metric.mem_closure_iff, dist_eq_norm_div] + +@[to_additive norm_le_zero_iff'] lemma norm_le_zero_iff''' [t0_space E] {a : E} : ∥a∥ ≤ 0 ↔ a = 1 := begin - convert metric.uniformity_basis_dist, - simp [dist_eq_norm] + letI : normed_group E := + { to_metric_space := metric.of_t0_pseudo_metric_space E, ..‹seminormed_group E› }, + rw [←dist_one_right, dist_le_zero], end -section tendsto_uniformly -/-- The results in this section do not require `E'` to any particular structure -/ -variables {E' : Type*} {f : ι → E' → G} {s : set E'} {l : filter ι} +@[to_additive norm_eq_zero'] lemma norm_eq_zero''' [t0_space E] {a : E} : ∥a∥ = 0 ↔ a = 1 := +(norm_nonneg' a).le_iff_eq.symm.trans norm_le_zero_iff''' + +@[to_additive norm_pos_iff'] lemma norm_pos_iff''' [t0_space E] {a : E} : 0 < ∥a∥ ↔ a ≠ 1 := +by rw [← not_le, norm_le_zero_iff'''] -lemma normed_add_comm_group.tendsto_uniformly_on_zero : - tendsto_uniformly_on f 0 l s ↔ ∀ ε > 0, ∀ᶠ (N : ι) in l, ∀ x : E', x ∈ s → ∥f N x∥ < ε := -by simp_rw [tendsto_uniformly_on_iff, pi.zero_apply, dist_zero_left] +@[to_additive] +lemma seminormed_group.tendsto_uniformly_on_one {f : ι → κ → G} {s : set κ} {l : filter ι} : + tendsto_uniformly_on f 1 l s ↔ ∀ ε > 0, ∀ᶠ i in l, ∀ x ∈ s, ∥f i x∥ < ε := +by simp_rw [tendsto_uniformly_on_iff, pi.one_apply, dist_one_left] -lemma normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero - {l' : filter E'} : uniform_cauchy_seq_on_filter f l l' ↔ - tendsto_uniformly_on_filter (λ n : ι × ι, λ z : E', f n.fst z - f n.snd z) 0 (l.prod l) l' := +@[to_additive] +lemma seminormed_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_one + {f : ι → κ → G} {l : filter ι} {l' : filter κ} : uniform_cauchy_seq_on_filter f l l' ↔ + tendsto_uniformly_on_filter (λ n : ι × ι, λ z, f n.fst z / f n.snd z) 1 (l ×ᶠ l) l' := begin - split, - { intros hf u hu, - obtain ⟨ε, hε, H⟩ := uniformity_basis_dist.mem_uniformity_iff.mp hu, + refine ⟨λ hf u hu, _, λ hf u hu, _⟩, + { obtain ⟨ε, hε, H⟩ := uniformity_basis_dist.mem_uniformity_iff.mp hu, refine (hf {p : G × G | dist p.fst p.snd < ε} $ dist_mem_uniformity hε).mono (λ x hx, - H 0 (f x.fst.fst x.snd - f x.fst.snd x.snd) _), - simpa [dist_eq_norm, norm_sub_rev] using hx, }, - { intros hf u hu, - obtain ⟨ε, hε, H⟩ := uniformity_basis_dist.mem_uniformity_iff.mp hu, + H 1 (f x.fst.fst x.snd / f x.fst.snd x.snd) _), + simpa [dist_eq_norm_div, norm_div_rev] using hx }, + { obtain ⟨ε, hε, H⟩ := uniformity_basis_dist.mem_uniformity_iff.mp hu, refine (hf {p : G × G | dist p.fst p.snd < ε} $ dist_mem_uniformity hε).mono (λ x hx, H (f x.fst.fst x.snd) (f x.fst.snd x.snd) _), - simpa [dist_eq_norm, norm_sub_rev] using hx, }, + simpa [dist_eq_norm_div, norm_div_rev] using hx } end -lemma normed_add_comm_group.uniform_cauchy_seq_on_iff_tendsto_uniformly_on_zero : +@[to_additive] +lemma seminormed_group.uniform_cauchy_seq_on_iff_tendsto_uniformly_on_one + {f : ι → κ → G} {s : set κ} {l : filter ι} : uniform_cauchy_seq_on f l s ↔ - tendsto_uniformly_on (λ n : ι × ι, λ z : E', f n.fst z - f n.snd z) 0 (l.prod l) s := + tendsto_uniformly_on (λ n : ι × ι, λ z, f n.fst z / f n.snd z) 1 (l ×ᶠ l) s := +by rw [tendsto_uniformly_on_iff_tendsto_uniformly_on_filter, + uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter, + seminormed_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_one] + +end seminormed_group + +section induced + +/-- A group homomorphism from a `group` to a `seminormed_group` induces a `seminormed_group` +structure on the domain. -/ +@[reducible, -- See note [reducible non-instances] +to_additive "A group homomorphism from an `add_group` to a `seminormed_add_group` induces a +`seminormed_add_group` structure on the domain."] +def seminormed_group.induced [group E] [seminormed_group F] (f : E →* F) : seminormed_group E := +{ norm := λ x, ∥f x∥, + dist_eq := λ x y, by simpa only [monoid_hom.map_div, ←dist_eq_norm_div], + ..pseudo_metric_space.induced f _ } + +/-- A group homomorphism from a `comm_group` to a `seminormed_group` induces a +`seminormed_comm_group` structure on the domain. -/ +@[reducible, -- See note [reducible non-instances] +to_additive "A group homomorphism from an `add_comm_group` to a `seminormed_add_group` induces a +`seminormed_add_comm_group` structure on the domain."] +def seminormed_comm_group.induced [comm_group E] [seminormed_group F] (f : E →* F) : + seminormed_comm_group E := +{ ..seminormed_group.induced f } + +/-- An injective group homomorphism from a `group` to a `normed_group` induces a `normed_group` +structure on the domain. -/ +@[reducible, -- See note [reducible non-instances]. +to_additive "An injective group homomorphism from an `add_group` to a `normed_add_group` induces a +`normed_add_group` structure on the domain."] +def normed_group.induced [group E] [normed_group F] (f : E →* F) (h : injective f) : + normed_group E := +{ ..seminormed_group.induced f, ..metric_space.induced f h _ } + +/-- An injective group homomorphism from an `comm_group` to a `normed_group` induces a +`normed_comm_group` structure on the domain. -/ +@[reducible, -- See note [reducible non-instances]. +to_additive "An injective group homomorphism from an `comm_group` to a `normed_comm_group` induces a +`normed_comm_group` structure on the domain."] +def normed_comm_group.induced [comm_group E] [normed_group F] (f : E →* F) (h : injective f) : + normed_comm_group E := +{ ..seminormed_group.induced f, ..metric_space.induced f h _ } + +end induced + +section seminormed_comm_group +variables [seminormed_comm_group E] [seminormed_comm_group F] {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} + +@[simp, to_additive] lemma dist_mul_left (a b₁ b₂ : E) : dist (a * b₁) (a * b₂) = dist b₁ b₂ := +by simp [dist_eq_norm_div] + +@[to_additive] lemma dist_inv (x y : E) : dist x⁻¹ y = dist x y⁻¹ := +by simp_rw [dist_eq_norm_div, ←norm_inv' (x⁻¹ / y), inv_div, div_inv_eq_mul, mul_comm] + +@[simp, to_additive] lemma dist_inv_inv (a b : E) : dist a⁻¹ b⁻¹ = dist a b := +by rw [dist_inv, inv_inv] + +@[simp, to_additive] lemma dist_div_left (a b₁ b₂ : E) : dist (a / b₁) (a / b₂) = dist b₁ b₂ := +by simp only [div_eq_mul_inv, dist_mul_left, dist_inv_inv] + +@[simp, to_additive] lemma dist_self_mul_right (a b : E) : dist a (a * b) = ∥b∥ := +by rw [←dist_one_left, ←dist_mul_left a 1 b, mul_one] + +@[simp, to_additive] lemma dist_self_mul_left (a b : E) : dist (a * b) a = ∥b∥ := +by rw [dist_comm, dist_self_mul_right] + +@[simp, to_additive] lemma dist_self_div_right (a b : E) : dist a (a / b) = ∥b∥ := +by rw [div_eq_mul_inv, dist_self_mul_right, norm_inv'] + +@[simp, to_additive] lemma dist_self_div_left (a b : E) : dist (a / b) a = ∥b∥ := +by rw [dist_comm, dist_self_div_right] + +@[to_additive] lemma dist_mul_mul_le (a₁ a₂ b₁ b₂ : E) : + dist (a₁ * a₂) (b₁ * b₂) ≤ dist a₁ b₁ + dist a₂ b₂ := +by simpa only [dist_mul_left, dist_mul_right] using dist_triangle (a₁ * a₂) (b₁ * a₂) (b₁ * b₂) + +@[to_additive] lemma dist_mul_mul_le_of_le (h₁ : dist a₁ b₁ ≤ r₁) (h₂ : dist a₂ b₂ ≤ r₂) : + dist (a₁ * a₂) (b₁ * b₂) ≤ r₁ + r₂ := +(dist_mul_mul_le a₁ a₂ b₁ b₂).trans $ add_le_add h₁ h₂ + +@[to_additive] lemma dist_div_div_le (a₁ a₂ b₁ b₂ : E) : + dist (a₁ / a₂) (b₁ / b₂) ≤ dist a₁ b₁ + dist a₂ b₂ := +by simpa only [div_eq_mul_inv, dist_inv_inv] using dist_mul_mul_le a₁ a₂⁻¹ b₁ b₂⁻¹ + +@[to_additive] lemma dist_div_div_le_of_le (h₁ : dist a₁ b₁ ≤ r₁) (h₂ : dist a₂ b₂ ≤ r₂) : + dist (a₁ / a₂) (b₁ / b₂) ≤ r₁ + r₂ := +(dist_div_div_le a₁ a₂ b₁ b₂).trans $ add_le_add h₁ h₂ + +@[to_additive] lemma abs_dist_sub_le_dist_mul_mul (a₁ a₂ b₁ b₂ : E) : + |dist a₁ b₁ - dist a₂ b₂| ≤ dist (a₁ * a₂) (b₁ * b₂) := +by simpa only [dist_mul_left, dist_mul_right, dist_comm b₂] + using abs_dist_sub_le (a₁ * a₂) (b₁ * b₂) (b₁ * a₂) + +lemma norm_multiset_sum_le {E} [seminormed_add_comm_group E] (m : multiset E) : + ∥m.sum∥ ≤ (m.map (λ x, ∥x∥)).sum := +m.le_sum_of_subadditive norm norm_zero norm_add_le + +@[to_additive] +lemma norm_multiset_prod_le (m : multiset E) : ∥m.prod∥ ≤ (m.map $ λ x, ∥x∥).sum := begin - rw tendsto_uniformly_on_iff_tendsto_uniformly_on_filter, - rw uniform_cauchy_seq_on_iff_uniform_cauchy_seq_on_filter, - exact normed_add_comm_group.uniform_cauchy_seq_on_filter_iff_tendsto_uniformly_on_filter_zero, + rw [←multiplicative.of_add_le, of_add_multiset_prod, multiset.map_map], + refine multiset.le_prod_of_submultiplicative (multiplicative.of_add ∘ norm) _ (λ x y, _) _, + { simp only [comp_app, norm_one', of_add_zero] }, + { exact norm_mul_le' _ _ } end -end tendsto_uniformly +lemma norm_sum_le {E} [seminormed_add_comm_group E] (s : finset ι) (f : ι → E) : + ∥∑ i in s, f i∥ ≤ ∑ i in s, ∥f i∥ := +s.le_sum_of_subadditive norm norm_zero norm_add_le f -open finset +@[to_additive] lemma norm_prod_le (s : finset ι) (f : ι → E) : ∥∏ i in s, f i∥ ≤ ∑ i in s, ∥f i∥ := +begin + rw [←multiplicative.of_add_le, of_add_sum], + refine finset.le_prod_of_submultiplicative (multiplicative.of_add ∘ norm) _ (λ x y, _) _ _, + { simp only [comp_app, norm_one', of_add_zero] }, + { exact norm_mul_le' _ _ } +end -/-- A homomorphism `f` of seminormed groups is Lipschitz, if there exists a constant `C` such that -for all `x`, one has `∥f x∥ ≤ C * ∥x∥`. The analogous condition for a linear map of -(semi)normed spaces is in `normed_space.operator_norm`. -/ -lemma add_monoid_hom_class.lipschitz_of_bound {𝓕 : Type*} [add_monoid_hom_class 𝓕 E F] (f : 𝓕) - (C : ℝ) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : lipschitz_with (real.to_nnreal C) f := -lipschitz_with.of_dist_le' $ λ x y, by simpa only [dist_eq_norm, map_sub] using h (x - y) +@[to_additive] +lemma norm_prod_le_of_le (s : finset ι) {f : ι → E} {n : ι → ℝ} (h : ∀ b ∈ s, ∥f b∥ ≤ n b) : + ∥∏ b in s, f b∥ ≤ ∑ b in s, n b := +(norm_prod_le s f).trans $ finset.sum_le_sum h -lemma lipschitz_on_with_iff_norm_sub_le {f : E → F} {C : ℝ≥0} {s : set E} : - lipschitz_on_with C f s ↔ ∀ (x ∈ s) (y ∈ s), ∥f x - f y∥ ≤ C * ∥x - y∥ := -by simp only [lipschitz_on_with_iff_dist_le_mul, dist_eq_norm] +@[to_additive] lemma dist_prod_prod_le_of_le (s : finset ι) {f a : ι → E} {d : ι → ℝ} + (h : ∀ b ∈ s, dist (f b) (a b) ≤ d b) : + dist (∏ b in s, f b) (∏ b in s, a b) ≤ ∑ b in s, d b := +by { simp only [dist_eq_norm_div, ← finset.prod_div_distrib] at *, exact norm_prod_le_of_le s h } -lemma lipschitz_on_with.norm_sub_le {f : E → F} {C : ℝ≥0} {s : set E} (h : lipschitz_on_with C f s) - {x y : E} (x_in : x ∈ s) (y_in : y ∈ s) : ∥f x - f y∥ ≤ C * ∥x - y∥ := -lipschitz_on_with_iff_norm_sub_le.mp h x x_in y y_in +@[to_additive] lemma dist_prod_prod_le (s : finset ι) (f a : ι → E) : + dist (∏ b in s, f b) (∏ b in s, a b) ≤ ∑ b in s, dist (f b) (a b) := +dist_prod_prod_le_of_le s $ λ _ _, le_rfl -lemma lipschitz_on_with.norm_sub_le_of_le {f : E → F} {C : ℝ≥0} {s : set E} - (h : lipschitz_on_with C f s){x y : E} (x_in : x ∈ s) (y_in : y ∈ s) {d : ℝ} (hd : ∥x - y∥ ≤ d) : - ∥f x - f y∥ ≤ C * d := -(h.norm_sub_le x_in y_in).trans $ mul_le_mul_of_nonneg_left hd C.2 +@[to_additive] lemma mul_mem_ball_iff_norm : a * b ∈ ball a r ↔ ∥b∥ < r := +by rw [mem_ball_iff_norm'', mul_div_cancel'''] -lemma lipschitz_with_iff_norm_sub_le {f : E → F} {C : ℝ≥0} : - lipschitz_with C f ↔ ∀ x y, ∥f x - f y∥ ≤ C * ∥x - y∥ := -by simp only [lipschitz_with_iff_dist_le_mul, dist_eq_norm] +@[to_additive] lemma mul_mem_closed_ball_iff_norm : a * b ∈ closed_ball a r ↔ ∥b∥ ≤ r := +by rw [mem_closed_ball_iff_norm'', mul_div_cancel'''] -alias lipschitz_with_iff_norm_sub_le ↔ lipschitz_with.norm_sub_le _ +@[simp, to_additive] lemma preimage_mul_ball (a b : E) (r : ℝ) : + ((*) b) ⁻¹' ball a r = ball (a / b) r := +by { ext c, simp only [dist_eq_norm_div, set.mem_preimage, mem_ball, div_div_eq_mul_div, mul_comm] } -lemma lipschitz_with.norm_sub_le_of_le {f : E → F} {C : ℝ≥0} (h : lipschitz_with C f) - {x y : E} {d : ℝ} (hd : ∥x - y∥ ≤ d) : - ∥f x - f y∥ ≤ C * d := -(h.norm_sub_le x y).trans $ mul_le_mul_of_nonneg_left hd C.2 +@[simp, to_additive] lemma preimage_mul_closed_ball (a b : E) (r : ℝ) : + ((*) b) ⁻¹' (closed_ball a r) = closed_ball (a / b) r := +by { ext c, + simp only [dist_eq_norm_div, set.mem_preimage, mem_closed_ball, div_div_eq_mul_div, mul_comm] } -/-- A homomorphism `f` of seminormed groups is continuous, if there exists a constant `C` such that -for all `x`, one has `∥f x∥ ≤ C * ∥x∥`. -/ -lemma add_monoid_hom_class.continuous_of_bound {𝓕 : Type*} [add_monoid_hom_class 𝓕 E F] (f : 𝓕) - (C : ℝ) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : continuous f := -(add_monoid_hom_class.lipschitz_of_bound f C h).continuous +@[simp, to_additive] lemma preimage_mul_sphere (a b : E) (r : ℝ) : + ((*) b) ⁻¹' sphere a r = sphere (a / b) r := +by { ext c, simp only [set.mem_preimage, mem_sphere_iff_norm', div_div_eq_mul_div, mul_comm] } -lemma add_monoid_hom_class.uniform_continuous_of_bound {𝓕 : Type*} [add_monoid_hom_class 𝓕 E F] - (f : 𝓕) (C : ℝ) (h : ∀x, ∥f x∥ ≤ C * ∥x∥) : uniform_continuous f := -(add_monoid_hom_class.lipschitz_of_bound f C h).uniform_continuous +namespace isometric -lemma is_compact.exists_bound_of_continuous_on [topological_space α] - {s : set α} (hs : is_compact s) {f : α → E} (hf : continuous_on f s) : - ∃ C, ∀ x ∈ s, ∥f x∥ ≤ C := -begin - have : bounded (f '' s) := (hs.image_of_continuous_on hf).bounded, - rcases bounded_iff_forall_norm_le.1 this with ⟨C, hC⟩, - exact ⟨C, λ x hx, hC _ (set.mem_image_of_mem _ hx)⟩, -end +/-- Multiplication `y ↦ x * y` as an `isometry`. -/ +@[to_additive "Addition `y ↦ x + y` as an `isometry`"] +protected def mul_left (x : E) : E ≃ᵢ E := +{ isometry_to_fun := isometry.of_dist_eq $ λ y z, dist_mul_left _ _ _, + to_equiv := equiv.mul_left x } -lemma add_monoid_hom_class.isometry_iff_norm {𝓕 : Type*} [add_monoid_hom_class 𝓕 E F] - (f : 𝓕) : isometry f ↔ ∀ x, ∥f x∥ = ∥x∥ := -begin - simp only [isometry_iff_dist_eq, dist_eq_norm, ←map_sub], - refine ⟨λ h x, _, λ h x y, h _⟩, - simpa using h x 0 -end +@[simp, to_additive] lemma mul_left_to_equiv (x : E) : + (isometric.mul_left x).to_equiv = equiv.mul_left x := rfl + +@[simp, to_additive] lemma coe_mul_left (x : E) : ⇑(isometric.mul_left x) = (*) x := rfl + +@[simp, to_additive] lemma mul_left_symm (x : E) : + (isometric.mul_left x).symm = isometric.mul_left x⁻¹ := +ext $ λ y, rfl + +variables (E) -lemma add_monoid_hom_class.isometry_of_norm {𝓕 : Type*} [add_monoid_hom_class 𝓕 E F] - (f : 𝓕) (hf : ∀ x, ∥f x∥ = ∥x∥) : isometry f := -(add_monoid_hom_class.isometry_iff_norm f).2 hf +/-- Inversion `x ↦ x⁻¹` as an `isometry`. -/ +@[to_additive "Negation `x ↦ -x` as an `isometry`."] protected def inv : E ≃ᵢ E := +{ isometry_to_fun := isometry.of_dist_eq $ λ x y, dist_inv_inv _ _, + to_equiv := equiv.inv E } + +variables {E} + +@[simp, to_additive] lemma inv_symm : (isometric.inv E).symm = isometric.inv E := rfl +@[simp, to_additive] lemma inv_to_equiv : (isometric.inv E).to_equiv = equiv.inv E := rfl +@[simp, to_additive] lemma coe_inv : ⇑(isometric.inv E) = has_inv.inv := rfl + +end isometric -lemma controlled_sum_of_mem_closure {s : add_subgroup E} {g : E} - (hg : g ∈ closure (s : set E)) {b : ℕ → ℝ} (b_pos : ∀ n, 0 < b n) : +open finset + +@[to_additive] lemma controlled_prod_of_mem_closure {s : subgroup E} (hg : a ∈ closure (s : set E)) + {b : ℕ → ℝ} (b_pos : ∀ n, 0 < b n) : ∃ v : ℕ → E, - tendsto (λ n, ∑ i in range (n+1), v i) at_top (𝓝 g) ∧ + tendsto (λ n, ∏ i in range (n+1), v i) at_top (𝓝 a) ∧ (∀ n, v n ∈ s) ∧ - ∥v 0 - g∥ < b 0 ∧ - ∀ n > 0, ∥v n∥ < b n := + ∥v 0 / a∥ < b 0 ∧ + ∀ n, 0 < n → ∥v n∥ < b n := begin - obtain ⟨u : ℕ → E, u_in : ∀ n, u n ∈ s, lim_u : tendsto u at_top (𝓝 g)⟩ := + obtain ⟨u : ℕ → E, u_in : ∀ n, u n ∈ s, lim_u : tendsto u at_top (𝓝 a)⟩ := mem_closure_iff_seq_limit.mp hg, - obtain ⟨n₀, hn₀⟩ : ∃ n₀, ∀ n ≥ n₀, ∥u n - g∥ < b 0, - { have : {x | ∥x - g∥ < b 0} ∈ 𝓝 g, - { simp_rw ← dist_eq_norm, + obtain ⟨n₀, hn₀⟩ : ∃ n₀, ∀ n ≥ n₀, ∥u n / a∥ < b 0, + { have : {x | ∥x / a∥ < b 0} ∈ 𝓝 a, + { simp_rw ← dist_eq_norm_div, exact metric.ball_mem_nhds _ (b_pos _) }, exact filter.tendsto_at_top'.mp lim_u _ this }, set z : ℕ → E := λ n, u (n + n₀), - have lim_z : tendsto z at_top (𝓝 g) := lim_u.comp (tendsto_add_at_top_nat n₀), - have mem_𝓤 : ∀ n, {p : E × E | ∥p.1 - p.2∥ < b (n + 1)} ∈ 𝓤 E := - λ n, by simpa [← dist_eq_norm] using metric.dist_mem_uniformity (b_pos $ n+1), + have lim_z : tendsto z at_top (𝓝 a) := lim_u.comp (tendsto_add_at_top_nat n₀), + have mem_𝓤 : ∀ n, {p : E × E | ∥p.1 / p.2∥ < b (n + 1)} ∈ 𝓤 E := + λ n, by simpa [← dist_eq_norm_div] using metric.dist_mem_uniformity (b_pos $ n+1), obtain ⟨φ : ℕ → ℕ, φ_extr : strict_mono φ, - hφ : ∀ n, ∥z (φ $ n + 1) - z (φ n)∥ < b (n + 1)⟩ := + hφ : ∀ n, ∥z (φ $ n + 1) / z (φ n)∥ < b (n + 1)⟩ := lim_z.cauchy_seq.subseq_mem mem_𝓤, set w : ℕ → E := z ∘ φ, - have hw : tendsto w at_top (𝓝 g), + have hw : tendsto w at_top (𝓝 a), from lim_z.comp φ_extr.tendsto_at_top, - set v : ℕ → E := λ i, if i = 0 then w 0 else w i - w (i - 1), - refine ⟨v, tendsto.congr (finset.eq_sum_range_sub' w) hw , _, + set v : ℕ → E := λ i, if i = 0 then w 0 else w i / w (i - 1), + refine ⟨v, tendsto.congr (finset.eq_prod_range_div' w) hw , _, hn₀ _ (n₀.le_add_left _), _⟩, { rintro ⟨⟩, { change w 0 ∈ s, apply u_in }, - { apply s.sub_mem ; apply u_in }, }, + { apply s.div_mem ; apply u_in }, }, { intros l hl, - obtain ⟨k, rfl⟩ : ∃ k, l = k+1, exact nat.exists_eq_succ_of_ne_zero (ne_of_gt hl), - apply hφ }, + obtain ⟨k, rfl⟩ : ∃ k, l = k+1, exact nat.exists_eq_succ_of_ne_zero hl.ne', + apply hφ } end -lemma controlled_sum_of_mem_closure_range {j : E →+ F} {h : F} - (Hh : h ∈ (closure $ (j.range : set F))) {b : ℕ → ℝ} (b_pos : ∀ n, 0 < b n) : - ∃ g : ℕ → E, - tendsto (λ n, ∑ i in range (n+1), j (g i)) at_top (𝓝 h) ∧ - ∥j (g 0) - h∥ < b 0 ∧ - ∀ n > 0, ∥j (g n)∥ < b n := +@[to_additive] lemma controlled_prod_of_mem_closure_range {j : E →* F} {b : F} + (hb : b ∈ closure (j.range : set F)) {f : ℕ → ℝ} (b_pos : ∀ n, 0 < f n) : + ∃ a : ℕ → E, + tendsto (λ n, ∏ i in range (n + 1), j (a i)) at_top (𝓝 b) ∧ + ∥j (a 0) / b∥ < f 0 ∧ + ∀ n, 0 < n → ∥j (a n)∥ < f n := begin - rcases controlled_sum_of_mem_closure Hh b_pos with ⟨v, sum_v, v_in, hv₀, hv_pos⟩, + obtain ⟨v, sum_v, v_in, hv₀, hv_pos⟩ := controlled_prod_of_mem_closure hb b_pos, choose g hg using v_in, - change ∀ (n : ℕ), j (g n) = v n at hg, refine ⟨g, by simpa [← hg] using sum_v, by simpa [hg 0] using hv₀, λ n hn, - by simpa [hg] using hv_pos n hn⟩ + by simpa [hg] using hv_pos n hn⟩, end -section nnnorm +@[to_additive] lemma nndist_mul_mul_le (a₁ a₂ b₁ b₂ : E) : + nndist (a₁ * a₂) (b₁ * b₂) ≤ nndist a₁ b₁ + nndist a₂ b₂ := +nnreal.coe_le_coe.1 $ dist_mul_mul_le a₁ a₂ b₁ b₂ -/-- Auxiliary class, endowing a type `α` with a function `nnnorm : α → ℝ≥0` with notation `∥x∥₊`. -/ -class has_nnnorm (E : Type*) := (nnnorm : E → ℝ≥0) +@[to_additive] +lemma edist_mul_mul_le (a₁ a₂ b₁ b₂ : E) : edist (a₁ * a₂) (b₁ * b₂) ≤ edist a₁ b₁ + edist a₂ b₂ := +by { simp only [edist_nndist], norm_cast, apply nndist_mul_mul_le } -export has_nnnorm (nnnorm) +@[simp, to_additive] lemma edist_mul_left (a b₁ b₂ : E) : edist (a * b₁) (a * b₂) = edist b₁ b₂ := +by simp [edist_dist] -notation `∥`e`∥₊` := nnnorm e +@[to_additive] +lemma edist_inv (a b : E) : edist a⁻¹ b = edist a b⁻¹ := by simp_rw [edist_dist, dist_inv] -@[priority 100] -- see Note [lower instance priority] -instance seminormed_add_comm_group.to_has_nnnorm : has_nnnorm E := ⟨λ a, ⟨norm a, norm_nonneg a⟩⟩ +@[simp, to_additive] lemma edist_inv_inv (x y : E) : edist x⁻¹ y⁻¹ = edist x y := +by rw [edist_inv, inv_inv] -@[simp, norm_cast] lemma coe_nnnorm (a : E) : (∥a∥₊ : ℝ) = norm a := rfl +@[simp, to_additive] lemma edist_div_left (a b₁ b₂ : E) : edist (a / b₁) (a / b₂) = edist b₁ b₂ := +by simp only [div_eq_mul_inv, edist_mul_left, edist_inv_inv] -@[simp] lemma coe_comp_nnnorm : (coe : ℝ≥0 → ℝ) ∘ (nnnorm : E → ℝ≥0) = norm := rfl +@[to_additive] +lemma nnnorm_multiset_prod_le (m : multiset E) : ∥m.prod∥₊ ≤ (m.map (λ x, ∥x∥₊)).sum := +nnreal.coe_le_coe.1 $ by { push_cast, rw multiset.map_map, exact norm_multiset_prod_le _ } -lemma norm_to_nnreal {a : E} : ∥a∥.to_nnreal = ∥a∥₊ := -@real.to_nnreal_coe ∥a∥₊ +@[to_additive] lemma nnnorm_prod_le (s : finset ι) (f : ι → E) : + ∥∏ a in s, f a∥₊ ≤ ∑ a in s, ∥f a∥₊ := +nnreal.coe_le_coe.1 $ by { push_cast, exact norm_prod_le _ _ } -lemma nndist_eq_nnnorm (a b : E) : nndist a b = ∥a - b∥₊ := nnreal.eq $ dist_eq_norm _ _ +@[to_additive] +lemma nnnorm_prod_le_of_le (s : finset ι) {f : ι → E} {n : ι → ℝ≥0} (h : ∀ b ∈ s, ∥f b∥₊ ≤ n b) : + ∥∏ b in s, f b∥₊ ≤ ∑ b in s, n b := +(norm_prod_le_of_le s h).trans_eq nnreal.coe_sum.symm -@[simp] lemma nnnorm_zero : ∥(0 : E)∥₊ = 0 := -nnreal.eq norm_zero +namespace lipschitz_with +variables [pseudo_emetric_space α] {K Kf Kg : ℝ≥0} {f g : α → E} -lemma ne_zero_of_nnnorm_ne_zero {g : E} : ∥g∥₊ ≠ 0 → g ≠ 0 := -mt $ by { rintro rfl, exact nnnorm_zero } +@[to_additive] lemma inv (hf : lipschitz_with K f) : lipschitz_with K (λ x, (f x)⁻¹) := +λ x y, (edist_inv_inv _ _).trans_le $ hf x y -lemma nnnorm_add_le (g h : E) : ∥g + h∥₊ ≤ ∥g∥₊ + ∥h∥₊ := -nnreal.coe_le_coe.1 $ norm_add_le g h +@[to_additive add] lemma mul' (hf : lipschitz_with Kf f) (hg : lipschitz_with Kg g) : + lipschitz_with (Kf + Kg) (λ x, f x * g x) := +λ x y, calc + edist (f x * g x) (f y * g y) ≤ edist (f x) (f y) + edist (g x) (g y) : edist_mul_mul_le _ _ _ _ +... ≤ Kf * edist x y + Kg * edist x y : add_le_add (hf x y) (hg x y) +... = (Kf + Kg) * edist x y : (add_mul _ _ _).symm -@[simp] lemma nnnorm_neg (g : E) : ∥-g∥₊ = ∥g∥₊ := -nnreal.eq $ norm_neg g +@[to_additive] lemma div (hf : lipschitz_with Kf f) (hg : lipschitz_with Kg g) : + lipschitz_with (Kf + Kg) (λ x, f x / g x) := +by simpa only [div_eq_mul_inv] using hf.mul' hg.inv -lemma nnnorm_sub_le (g h : E) : ∥g - h∥₊ ≤ ∥g∥₊ + ∥h∥₊ := -nnreal.coe_le_coe.1 $ norm_sub_le g h +end lipschitz_with -lemma nndist_nnnorm_nnnorm_le (g h : E) : nndist ∥g∥₊ ∥h∥₊ ≤ ∥g - h∥₊ := -nnreal.coe_le_coe.1 $ dist_norm_norm_le g h +namespace antilipschitz_with +variables [pseudo_emetric_space α] {K Kf Kg : ℝ≥0} {f g : α → E} -/-- The direct path from `0` to `v` is shorter than the path with `u` inserted in between. -/ -lemma nnnorm_le_insert (u v : E) : ∥v∥₊ ≤ ∥u∥₊ + ∥u - v∥₊ := norm_le_insert u v +@[to_additive] lemma mul_lipschitz_with (hf : antilipschitz_with Kf f) (hg : lipschitz_with Kg g) + (hK : Kg < Kf⁻¹) : antilipschitz_with (Kf⁻¹ - Kg)⁻¹ (λ x, f x * g x) := +begin + letI : pseudo_metric_space α := pseudo_emetric_space.to_pseudo_metric_space hf.edist_ne_top, + refine antilipschitz_with.of_le_mul_dist (λ x y, _), + rw [nnreal.coe_inv, ← div_eq_inv_mul], + rw le_div_iff (nnreal.coe_pos.2 $ tsub_pos_iff_lt.2 hK), + rw [mul_comm, nnreal.coe_sub hK.le, sub_mul], + calc ↑Kf⁻¹ * dist x y - Kg * dist x y ≤ dist (f x) (f y) - dist (g x) (g y) : + sub_le_sub (hf.mul_le_dist x y) (hg.dist_le_mul x y) + ... ≤ _ : le_trans (le_abs_self _) (abs_dist_sub_le_dist_mul_mul _ _ _ _), +end -lemma nnnorm_le_insert' (u v : E) : ∥u∥₊ ≤ ∥v∥₊ + ∥u - v∥₊ := norm_le_insert' u v +@[to_additive] lemma mul_div_lipschitz_with (hf : antilipschitz_with Kf f) + (hg : lipschitz_with Kg (g / f)) (hK : Kg < Kf⁻¹) : antilipschitz_with (Kf⁻¹ - Kg)⁻¹ g := +by simpa only [pi.div_apply, mul_div_cancel'_right] using hf.mul_lipschitz_with hg hK -lemma nnnorm_le_add_nnnorm_add (u v : E) : ∥u∥₊ ≤ ∥u + v∥₊ + ∥v∥₊ := norm_le_add_norm_add u v +@[to_additive] lemma le_mul_norm_div {f : E → F} (hf : antilipschitz_with K f) (x y : E) : + ∥x / y∥ ≤ K * ∥f x / f y∥ := +by simp [← dist_eq_norm_div, hf.le_mul_dist x y] -lemma of_real_norm_eq_coe_nnnorm (x : E) : ennreal.of_real ∥x∥ = (∥x∥₊ : ℝ≥0∞) := -ennreal.of_real_eq_coe_nnreal _ +end antilipschitz_with -lemma edist_eq_coe_nnnorm_sub (x y : E) : edist x y = (∥x - y∥₊ : ℝ≥0∞) := -by rw [edist_dist, dist_eq_norm, of_real_norm_eq_coe_nnnorm] +@[priority 100, to_additive] -- See note [lower instance priority] +instance seminormed_comm_group.to_has_lipschitz_mul : has_lipschitz_mul E := +⟨⟨1 + 1, lipschitz_with.prod_fst.mul' lipschitz_with.prod_snd⟩⟩ -lemma edist_eq_coe_nnnorm (x : E) : edist x 0 = (∥x∥₊ : ℝ≥0∞) := -by rw [edist_eq_coe_nnnorm_sub, _root_.sub_zero] +/-- A seminormed group is a uniform group, i.e., multiplication and division are uniformly +continuous. -/ +@[priority 100, to_additive "A seminormed group is a uniform additive group, i.e., addition and +subtraction are uniformly continuous."] -- See note [lower instance priority] +instance seminormed_comm_group.to_uniform_group : uniform_group E := +⟨(lipschitz_with.prod_fst.div lipschitz_with.prod_snd).uniform_continuous⟩ + + -- short-circuit type class inference +@[priority 100, to_additive] -- See note [lower instance priority] +instance seminormed_comm_group.to_topological_group : topological_group E := infer_instance + +@[to_additive] lemma cauchy_seq_prod_of_eventually_eq {u v : ℕ → E} {N : ℕ} + (huv : ∀ n ≥ N, u n = v n) (hv : cauchy_seq (λ n, ∏ k in range (n+1), v k)) : + cauchy_seq (λ n, ∏ k in range (n + 1), u k) := +begin + let d : ℕ → E := λ n, ∏ k in range (n + 1), (u k / v k), + rw show (λ n, ∏ k in range (n + 1), u k) = d * (λ n, ∏ k in range (n + 1), v k), + by { ext n, simp [d] }, + suffices : ∀ n ≥ N, d n = d N, + { exact (tendsto_at_top_of_eventually_const this).cauchy_seq.mul hv }, + intros n hn, + dsimp [d], + rw eventually_constant_prod _ hn, + intros m hm, + simp [huv m hm], +end -lemma mem_emetric_ball_zero_iff {x : E} {r : ℝ≥0∞} : x ∈ emetric.ball (0 : E) r ↔ ↑∥x∥₊ < r := -by rw [emetric.mem_ball, edist_eq_coe_nnnorm] +end seminormed_comm_group -lemma nndist_add_add_le (g₁ g₂ h₁ h₂ : E) : - nndist (g₁ + g₂) (h₁ + h₂) ≤ nndist g₁ h₁ + nndist g₂ h₂ := -nnreal.coe_le_coe.1 $ dist_add_add_le g₁ g₂ h₁ h₂ +section normed_group +variables [normed_group E] [normed_group F] {a b : E} -lemma edist_add_add_le (g₁ g₂ h₁ h₂ : E) : - edist (g₁ + g₂) (h₁ + h₂) ≤ edist g₁ h₁ + edist g₂ h₂ := -by { simp only [edist_nndist], norm_cast, apply nndist_add_add_le } +@[simp, to_additive norm_eq_zero] lemma norm_eq_zero'' : ∥a∥ = 0 ↔ a = 1 := norm_eq_zero''' -@[simp] lemma edist_add_left (g h₁ h₂ : E) : edist (g + h₁) (g + h₂) = edist h₁ h₂ := -by simp [edist_dist] +@[to_additive norm_ne_zero_iff] lemma norm_ne_zero_iff' : ∥a∥ ≠ 0 ↔ a ≠ 1 := norm_eq_zero''.not -@[simp] lemma edist_add_right (g₁ g₂ h : E) : edist (g₁ + h) (g₂ + h) = edist g₁ g₂ := -by simp [edist_dist] +@[simp, to_additive norm_pos_iff] lemma norm_pos_iff'' : 0 < ∥a∥ ↔ a ≠ 1 := norm_pos_iff''' -lemma edist_neg (x y : E) : edist (-x) y = edist x (-y) := by simp_rw [edist_dist, dist_neg] -@[simp] lemma edist_neg_neg (x y : E) : edist (-x) (-y) = edist x y := by rw [edist_neg, neg_neg] +@[simp, to_additive norm_le_zero_iff] +lemma norm_le_zero_iff'' : ∥a∥ ≤ 0 ↔ a = 1 := norm_le_zero_iff''' -@[simp] lemma edist_sub_left (g h₁ h₂ : E) : edist (g - h₁) (g - h₂) = edist h₁ h₂ := -by simp only [sub_eq_add_neg, edist_add_left, edist_neg_neg] +@[to_additive] +lemma norm_div_eq_zero_iff : ∥a / b∥ = 0 ↔ a = b := by rw [norm_eq_zero'', div_eq_one] -@[simp] lemma edist_sub_right (g₁ g₂ h : E) : edist (g₁ - h) (g₂ - h) = edist g₁ g₂ := -by simpa only [sub_eq_add_neg] using edist_add_right _ _ _ +@[to_additive] lemma norm_div_pos_iff : 0 < ∥a / b∥ ↔ a ≠ b := +by { rw [(norm_nonneg' _).lt_iff_ne, ne_comm], exact norm_div_eq_zero_iff.not } -lemma nnnorm_multiset_sum_le (m : multiset E) : - ∥m.sum∥₊ ≤ (m.map (λ x, ∥x∥₊)).sum := -m.le_sum_of_subadditive nnnorm nnnorm_zero nnnorm_add_le +@[to_additive] lemma eq_of_norm_div_le_zero (h : ∥a / b∥ ≤ 0) : a = b := +by rwa [←div_eq_one, ← norm_le_zero_iff''] -lemma nnnorm_sum_le (s : finset ι) (f : ι → E) : - ∥∑ a in s, f a∥₊ ≤ ∑ a in s, ∥f a∥₊ := -s.le_sum_of_subadditive nnnorm nnnorm_zero nnnorm_add_le f +alias norm_div_eq_zero_iff ↔ eq_of_norm_div_eq_zero _ -lemma nnnorm_sum_le_of_le (s : finset ι) {f : ι → E} {n : ι → ℝ≥0} (h : ∀ b ∈ s, ∥f b∥₊ ≤ n b) : - ∥∑ b in s, f b∥₊ ≤ ∑ b in s, n b := -(norm_sum_le_of_le s h).trans_eq nnreal.coe_sum.symm +attribute [to_additive] eq_of_norm_div_eq_zero -lemma add_monoid_hom_class.lipschitz_of_bound_nnnorm {𝓕 : Type*} [add_monoid_hom_class 𝓕 E F] - (f : 𝓕) (C : ℝ≥0) (h : ∀ x, ∥f x∥₊ ≤ C * ∥x∥₊) : lipschitz_with C f := -@real.to_nnreal_coe C ▸ add_monoid_hom_class.lipschitz_of_bound f C h +@[simp, to_additive nnnorm_eq_zero] lemma nnnorm_eq_zero' : ∥a∥₊ = 0 ↔ a = 1 := +by rw [← nnreal.coe_eq_zero, coe_nnnorm', norm_eq_zero''] -lemma add_monoid_hom_class.antilipschitz_of_bound {𝓕 : Type*} - [add_monoid_hom_class 𝓕 E F] (f : 𝓕) {K : ℝ≥0} (h : ∀ x, ∥x∥ ≤ K * ∥f x∥) : - antilipschitz_with K f := -antilipschitz_with.of_le_mul_dist $ -λ x y, by simpa only [dist_eq_norm, map_sub] using h (x - y) +@[to_additive nnnorm_ne_zero_iff] +lemma nnnorm_ne_zero_iff' : ∥a∥₊ ≠ 0 ↔ a ≠ 1 := nnnorm_eq_zero'.not -lemma add_monoid_hom_class.bound_of_antilipschitz {𝓕 : Type*} [add_monoid_hom_class 𝓕 E F] (f : 𝓕) - {K : ℝ≥0} (h : antilipschitz_with K f) (x) : ∥x∥ ≤ K * ∥f x∥ := -by simpa only [dist_zero_right, map_zero] using h.le_mul_dist x 0 +@[to_additive] +lemma tendsto_norm_div_self_punctured_nhds (a : E) : tendsto (λ x, ∥x / a∥) (𝓝[≠] a) (𝓝[>] 0) := +(tendsto_norm_div_self a).inf $ tendsto_principal_principal.2 $ λ x hx, norm_pos_iff''.2 $ + div_ne_one.2 hx +@[to_additive] lemma tendsto_norm_nhds_within_one : tendsto (norm : E → ℝ) (𝓝[≠] 1) (𝓝[>] 0) := +tendsto_norm_one.inf $ tendsto_principal_principal.2 $ λ x, norm_pos_iff''.2 -end nnnorm +variables (E) -namespace lipschitz_with +/-- The norm of a normed group as a group norm. -/ +@[to_additive "The norm of a normed group as an additive group norm."] +def norm_group_norm : group_norm E := +{ eq_one_of_map_eq_zero' := λ _, norm_eq_zero''.1, ..norm_group_seminorm _ } -variables [pseudo_emetric_space α] {K Kf Kg : ℝ≥0} {f g : α → E} +@[simp] lemma coe_norm_group_norm : ⇑(norm_group_norm E) = norm := rfl -lemma neg (hf : lipschitz_with K f) : lipschitz_with K (λ x, -f x) := -λ x y, (edist_neg_neg _ _).trans_le $ hf x y +end normed_group -lemma add (hf : lipschitz_with Kf f) (hg : lipschitz_with Kg g) : - lipschitz_with (Kf + Kg) (λ x, f x + g x) := -λ x y, -calc edist (f x + g x) (f y + g y) ≤ edist (f x) (f y) + edist (g x) (g y) : - edist_add_add_le _ _ _ _ -... ≤ Kf * edist x y + Kg * edist x y : - add_le_add (hf x y) (hg x y) -... = (Kf + Kg) * edist x y : - (add_mul _ _ _).symm +section normed_add_group +variables [normed_add_group E] [topological_space α] {f : α → E} -lemma sub (hf : lipschitz_with Kf f) (hg : lipschitz_with Kg g) : - lipschitz_with (Kf + Kg) (λ x, f x - g x) := -by simpa only [sub_eq_add_neg] using hf.add hg.neg +/-! Some relations with `has_compact_support` -/ -end lipschitz_with +lemma has_compact_support_norm_iff : has_compact_support (λ x, ∥f x∥) ↔ has_compact_support f := +has_compact_support_comp_left $ λ x, norm_eq_zero -namespace antilipschitz_with +alias has_compact_support_norm_iff ↔ _ has_compact_support.norm -variables [pseudo_emetric_space α] {K Kf Kg : ℝ≥0} {f g : α → E} +lemma continuous.bounded_above_of_compact_support (hf : continuous f) (h : has_compact_support f) : + ∃ C, ∀ x, ∥f x∥ ≤ C := +by simpa [bdd_above_def] using hf.norm.bdd_above_range_of_has_compact_support h.norm -lemma add_lipschitz_with (hf : antilipschitz_with Kf f) (hg : lipschitz_with Kg g) - (hK : Kg < Kf⁻¹) : antilipschitz_with (Kf⁻¹ - Kg)⁻¹ (λ x, f x + g x) := -begin - letI : pseudo_metric_space α := pseudo_emetric_space.to_pseudo_metric_space hf.edist_ne_top, - refine antilipschitz_with.of_le_mul_dist (λ x y, _), - rw [nnreal.coe_inv, ← div_eq_inv_mul], - rw le_div_iff (nnreal.coe_pos.2 $ tsub_pos_iff_lt.2 hK), - rw [mul_comm, nnreal.coe_sub hK.le, sub_mul], - calc ↑Kf⁻¹ * dist x y - Kg * dist x y ≤ dist (f x) (f y) - dist (g x) (g y) : - sub_le_sub (hf.mul_le_dist x y) (hg.dist_le_mul x y) - ... ≤ _ : le_trans (le_abs_self _) (abs_dist_sub_le_dist_add_add _ _ _ _) -end +end normed_add_group -lemma add_sub_lipschitz_with (hf : antilipschitz_with Kf f) (hg : lipschitz_with Kg (g - f)) - (hK : Kg < Kf⁻¹) : antilipschitz_with (Kf⁻¹ - Kg)⁻¹ g := -by simpa only [pi.sub_apply, add_sub_cancel'_right] using hf.add_lipschitz_with hg hK +/-! ### `ulift` -/ -lemma le_mul_norm_sub {f : E → F} (hf : antilipschitz_with K f) (x y : E) : - ∥x - y∥ ≤ K * ∥f x - f y∥ := -by simp [← dist_eq_norm, hf.le_mul_dist x y] +namespace ulift +section has_norm +variables [has_norm E] -end antilipschitz_with +instance : has_norm (ulift E) := ⟨λ x, ∥x.down∥⟩ -/-- A group homomorphism from an `add_comm_group` to a `seminormed_add_comm_group` induces a -`seminormed_add_comm_group` structure on the domain. +lemma norm_def (x : ulift E) : ∥x∥ = ∥x.down∥ := rfl +@[simp] lemma norm_up (x : E) : ∥ulift.up x∥ = ∥x∥ := rfl +@[simp] lemma norm_down (x : ulift E) : ∥x.down∥ = ∥x∥ := rfl -See note [reducible non-instances] -/ -@[reducible] -def seminormed_add_comm_group.induced {E} [add_comm_group E] (f : E →+ F) : - seminormed_add_comm_group E := -{ norm := λ x, ∥f x∥, - dist_eq := λ x y, by simpa only [add_monoid_hom.map_sub, ← dist_eq_norm], - .. pseudo_metric_space.induced f seminormed_add_comm_group.to_pseudo_metric_space, } +end has_norm -/-- A subgroup of a seminormed group is also a seminormed group, -with the restriction of the norm. -/ -instance add_subgroup.seminormed_add_comm_group (s : add_subgroup E) : - seminormed_add_comm_group s := -seminormed_add_comm_group.induced s.subtype +section has_nnnorm +variables [has_nnnorm E] -/-- If `x` is an element of a subgroup `s` of a seminormed group `E`, its norm in `s` is equal to -its norm in `E`. -/ -@[simp] lemma add_subgroup.coe_norm {E : Type*} [seminormed_add_comm_group E] - {s : add_subgroup E} (x : s) : - ∥(x : s)∥ = ∥(x:E)∥ := -rfl +instance : has_nnnorm (ulift E) := ⟨λ x, ∥x.down∥₊⟩ -/-- If `x` is an element of a subgroup `s` of a seminormed group `E`, its norm in `s` is equal to -its norm in `E`. +lemma nnnorm_def (x : ulift E) : ∥x∥₊ = ∥x.down∥₊ := rfl +@[simp] lemma nnnorm_up (x : E) : ∥ulift.up x∥₊ = ∥x∥₊ := rfl +@[simp] lemma nnnorm_down (x : ulift E) : ∥x.down∥₊ = ∥x∥₊ := rfl -This is a reversed version of the `simp` lemma `add_subgroup.coe_norm` for use by `norm_cast`. --/ +end has_nnnorm -@[norm_cast] lemma add_subgroup.norm_coe {E : Type*} [seminormed_add_comm_group E] - {s : add_subgroup E} (x : s) : - ∥(x : E)∥ = ∥(x : s)∥ := -rfl +@[to_additive] instance seminormed_group [seminormed_group E] : seminormed_group (ulift E) := +seminormed_group.induced ⟨ulift.down, rfl, λ _ _, rfl⟩ -/-- A submodule of a seminormed group is also a seminormed group, with the restriction of the norm. +@[to_additive] +instance seminormed_comm_group [seminormed_comm_group E] : seminormed_comm_group (ulift E) := +seminormed_comm_group.induced ⟨ulift.down, rfl, λ _ _, rfl⟩ -See note [implicit instance arguments]. -/ -instance submodule.seminormed_add_comm_group {𝕜 E : Type*} {_ : ring 𝕜} - [seminormed_add_comm_group E] {_ : module 𝕜 E} (s : submodule 𝕜 E) : - seminormed_add_comm_group s := -{ norm := λx, norm (x : E), - dist_eq := λx y, dist_eq_norm (x : E) (y : E) } +@[to_additive] instance normed_group [normed_group E] : normed_group (ulift E) := +normed_group.induced ⟨ulift.down, rfl, λ _ _, rfl⟩ down_injective -/-- If `x` is an element of a submodule `s` of a normed group `E`, its norm in `s` is equal to its -norm in `E`. +@[to_additive] +instance normed_comm_group [normed_comm_group E] : normed_comm_group (ulift E) := +normed_comm_group.induced ⟨ulift.down, rfl, λ _ _, rfl⟩ down_injective -See note [implicit instance arguments]. -/ -@[simp] lemma submodule.coe_norm {𝕜 : Type*} {_ : ring 𝕜} - {E : Type*} [seminormed_add_comm_group E] {_ : module 𝕜 E} {s : submodule 𝕜 E} (x : s) : - ∥(x : s)∥ = ∥(x : E)∥ := -rfl +end ulift -/-- If `x` is an element of a submodule `s` of a normed group `E`, its norm in `E` is equal to its -norm in `s`. +/-! ### `additive`, `multiplicative` -/ -This is a reversed version of the `simp` lemma `submodule.coe_norm` for use by `norm_cast`. +section additive_multiplicative -See note [implicit instance arguments]. -/ -@[norm_cast] lemma submodule.norm_coe {𝕜 : Type*} {_ : ring 𝕜} - {E : Type*} [seminormed_add_comm_group E] {_ : module 𝕜 E} {s : submodule 𝕜 E} (x : s) : - ∥(x : E)∥ = ∥(x : s)∥ := -rfl +open additive multiplicative -instance ulift.seminormed_add_comm_group : seminormed_add_comm_group (ulift E) := -seminormed_add_comm_group.induced ⟨ulift.down, rfl, λ _ _, rfl⟩ +section has_norm +variables [has_norm E] -lemma ulift.norm_def (x : ulift E) : ∥x∥ = ∥x.down∥ := rfl -lemma ulift.nnnorm_def (x : ulift E) : ∥x∥₊ = ∥x.down∥₊ := rfl +instance : has_norm (additive E) := ‹has_norm E› +instance : has_norm (multiplicative E) := ‹has_norm E› -@[simp] lemma ulift.norm_up (x : E) : ∥ulift.up x∥ = ∥x∥ := rfl -@[simp] lemma ulift.nnnorm_up (x : E) : ∥ulift.up x∥₊ = ∥x∥₊ := rfl +@[simp] lemma norm_to_mul (x) : ∥(to_mul x : E)∥ = ∥x∥ := rfl +@[simp] lemma norm_of_mul (x : E) : ∥of_mul x∥ = ∥x∥ := rfl +@[simp] lemma norm_to_add (x) : ∥(to_add x : E)∥ = ∥x∥ := rfl +@[simp] lemma norm_of_add (x : E) : ∥of_add x∥ = ∥x∥ := rfl -/-- seminormed group instance on the product of two seminormed groups, using the sup norm. -/ -instance prod.seminormed_add_comm_group : seminormed_add_comm_group (E × F) := -{ norm := λx, ∥x.1∥ ⊔ ∥x.2∥, - dist_eq := assume (x y : E × F), - show max (dist x.1 y.1) (dist x.2 y.2) = (max ∥(x - y).1∥ ∥(x - y).2∥), by simp [dist_eq_norm] } +end has_norm -lemma prod.norm_def (x : E × F) : ∥x∥ = (max ∥x.1∥ ∥x.2∥) := rfl +section has_nnnorm +variables [has_nnnorm E] -lemma prod.nnnorm_def (x : E × F) : ∥x∥₊ = max (∥x.1∥₊) (∥x.2∥₊) := -by { have := x.norm_def, simp only [← coe_nnnorm] at this, exact_mod_cast this } +instance : has_nnnorm (additive E) := ‹has_nnnorm E› +instance : has_nnnorm (multiplicative E) := ‹has_nnnorm E› -lemma norm_fst_le (x : E × F) : ∥x.1∥ ≤ ∥x∥ := -le_max_left _ _ +@[simp] lemma nnnorm_to_mul (x) : ∥(to_mul x : E)∥₊ = ∥x∥₊ := rfl +@[simp] lemma nnnorm_of_mul (x : E) : ∥of_mul x∥₊ = ∥x∥₊ := rfl +@[simp] lemma nnnorm_to_add (x) : ∥(to_add x : E)∥₊ = ∥x∥₊ := rfl +@[simp] lemma nnnorm_of_add (x : E) : ∥of_add x∥₊ = ∥x∥₊ := rfl -lemma norm_snd_le (x : E × F) : ∥x.2∥ ≤ ∥x∥ := -le_max_right _ _ +end has_nnnorm -lemma norm_prod_le_iff {x : E × F} {r : ℝ} : - ∥x∥ ≤ r ↔ ∥x.1∥ ≤ r ∧ ∥x.2∥ ≤ r := -max_le_iff +instance [seminormed_group E] : seminormed_add_group (additive E) := +{ dist_eq := dist_eq_norm_div } -section pi -variables {π : ι → Type*} [fintype ι] [Π i, seminormed_add_comm_group (π i)] (f : Π i, π i) +instance [seminormed_add_group E] : seminormed_group (multiplicative E) := +{ dist_eq := dist_eq_norm_sub } -/-- seminormed group instance on the product of finitely many seminormed groups, -using the sup norm. -/ -instance pi.seminormed_add_comm_group : seminormed_add_comm_group (Π i, π i) := -{ norm := λ f, ↑(finset.univ.sup (λ b, ∥f b∥₊)), - dist_eq := assume x y, - congr_arg (coe : ℝ≥0 → ℝ) $ congr_arg (finset.sup finset.univ) $ funext $ assume a, - show nndist (x a) (y a) = ∥x a - y a∥₊, from nndist_eq_nnnorm _ _ } +instance [seminormed_comm_group E] : seminormed_add_comm_group (additive E) := +{ ..additive.seminormed_add_group } -lemma pi.norm_def : ∥f∥ = ↑(finset.univ.sup (λ b, ∥f b∥₊)) := rfl -lemma pi.nnnorm_def : ∥f∥₊ = finset.univ.sup (λ b, ∥f b∥₊) := subtype.eta _ _ +instance [seminormed_add_comm_group E] : seminormed_comm_group (multiplicative E) := +{ ..multiplicative.seminormed_group } -/-- The seminorm of an element in a product space is `≤ r` if and only if the norm of each -component is. -/ -lemma pi_norm_le_iff {r : ℝ} (hr : 0 ≤ r) {x : Π i, π i} : ∥x∥ ≤ r ↔ ∀ i, ∥x i∥ ≤ r := -by simp only [← dist_zero_right, dist_pi_le_iff hr, pi.zero_apply] +instance [normed_group E] : normed_add_group (additive E) := +{ ..additive.seminormed_add_group } -lemma pi_nnnorm_le_iff {r : ℝ≥0} {x : Π i, π i} : ∥x∥₊ ≤ r ↔ ∀i, ∥x i∥₊ ≤ r := -pi_norm_le_iff r.coe_nonneg +instance [normed_add_group E] : normed_group (multiplicative E) := +{ ..multiplicative.seminormed_group } -/-- The seminorm of an element in a product space is `< r` if and only if the norm of each -component is. -/ -lemma pi_norm_lt_iff {r : ℝ} (hr : 0 < r) {x : Π i, π i} : ∥x∥ < r ↔ ∀ i, ∥x i∥ < r := -by simp only [← dist_zero_right, dist_pi_lt_iff hr, pi.zero_apply] +instance [normed_comm_group E] : normed_add_comm_group (additive E) := +{ ..additive.seminormed_add_group } -lemma pi_nnnorm_lt_iff {r : ℝ≥0} (hr : 0 < r) {x : Π i, π i} : ∥x∥₊ < r ↔ ∀ i, ∥x i∥₊ < r := -pi_norm_lt_iff hr +instance [normed_add_comm_group E] : normed_comm_group (multiplicative E) := +{ ..multiplicative.seminormed_group } -lemma norm_le_pi_norm (i : ι) : ∥f i∥ ≤ ∥f∥ := (pi_norm_le_iff $ norm_nonneg _).1 le_rfl i +end additive_multiplicative -lemma nnnorm_le_pi_nnnorm (i : ι) : ∥f i∥₊ ≤ ∥f∥₊ := norm_le_pi_norm _ i +/-! ### Order dual -/ -@[simp] lemma pi_norm_const [nonempty ι] (a : E) : ∥(λ i : ι, a)∥ = ∥a∥ := -by simpa only [← dist_zero_right] using dist_pi_const a 0 +section order_dual -@[simp] lemma pi_nnnorm_const [nonempty ι] (a : E) : ∥(λ i : ι, a)∥₊ = ∥a∥₊ := -nnreal.eq $ pi_norm_const a +open order_dual -/-- The $L^1$ norm is less than the $L^\infty$ norm scaled by the cardinality. -/ -lemma pi.sum_norm_apply_le_norm : ∑ i, ∥f i∥ ≤ fintype.card ι • ∥f∥ := -finset.sum_le_card_nsmul _ _ _ $ λ i hi, norm_le_pi_norm _ i +section has_norm +variables [has_norm E] -/-- The $L^1$ norm is less than the $L^\infty$ norm scaled by the cardinality. -/ -lemma pi.sum_nnnorm_apply_le_nnnorm : ∑ i, ∥f i∥₊ ≤ fintype.card ι • ∥f∥₊ := -nnreal.coe_sum.trans_le $ pi.sum_norm_apply_le_norm _ +instance : has_norm Eᵒᵈ := ‹has_norm E› -end pi - -lemma tendsto_iff_norm_tendsto_zero {f : α → E} {a : filter α} {b : E} : - tendsto f a (𝓝 b) ↔ tendsto (λ e, ∥f e - b∥) a (𝓝 0) := -by { convert tendsto_iff_dist_tendsto_zero, simp [dist_eq_norm] } +@[simp] lemma norm_to_dual (x : E) : ∥to_dual x∥ = ∥x∥ := rfl +@[simp] lemma norm_of_dual (x : Eᵒᵈ) : ∥of_dual x∥ = ∥x∥ := rfl -lemma tendsto_zero_iff_norm_tendsto_zero {f : α → E} {a : filter α} : - tendsto f a (𝓝 0) ↔ tendsto (λ e, ∥f e∥) a (𝓝 0) := -by { rw [tendsto_iff_norm_tendsto_zero], simp only [sub_zero] } +end has_norm -lemma comap_norm_nhds_zero : comap norm (𝓝 0) = 𝓝 (0 : E) := -by simpa only [dist_zero_right] using nhds_comap_dist (0 : E) +section has_nnnorm +variables [has_nnnorm E] -/-- Special case of the sandwich theorem: if the norm of `f` is eventually bounded by a real -function `g` which tends to `0`, then `f` tends to `0`. -In this pair of lemmas (`squeeze_zero_norm'` and `squeeze_zero_norm`), following a convention of -similar lemmas in `topology.metric_space.basic` and `topology.algebra.order`, the `'` version is -phrased using "eventually" and the non-`'` version is phrased absolutely. -/ -lemma squeeze_zero_norm' {f : α → E} {g : α → ℝ} {t₀ : filter α} - (h : ∀ᶠ n in t₀, ∥f n∥ ≤ g n) - (h' : tendsto g t₀ (𝓝 0)) : tendsto f t₀ (𝓝 0) := -tendsto_zero_iff_norm_tendsto_zero.mpr - (squeeze_zero' (eventually_of_forall (λ n, norm_nonneg _)) h h') +instance : has_nnnorm Eᵒᵈ := ‹has_nnnorm E› -/-- Special case of the sandwich theorem: if the norm of `f` is bounded by a real function `g` which -tends to `0`, then `f` tends to `0`. -/ -lemma squeeze_zero_norm {f : α → E} {g : α → ℝ} {t₀ : filter α} - (h : ∀ n, ∥f n∥ ≤ g n) (h' : tendsto g t₀ (𝓝 0)) : - tendsto f t₀ (𝓝 0) := -squeeze_zero_norm' (eventually_of_forall h) h' +@[simp] lemma nnnorm_to_dual (x : E) : ∥to_dual x∥₊ = ∥x∥₊ := rfl +@[simp] lemma nnnorm_of_dual (x : Eᵒᵈ) : ∥of_dual x∥₊ = ∥x∥₊ := rfl -lemma tendsto_norm_sub_self (x : E) : tendsto (λ g : E, ∥g - x∥) (𝓝 x) (𝓝 0) := -by simpa [dist_eq_norm] using tendsto_id.dist (tendsto_const_nhds : tendsto (λ g, (x:E)) (𝓝 x) _) +end has_nnnorm -lemma lipschitz_with_one_norm : lipschitz_with 1 (norm : E → ℝ) := -by simpa only [dist_zero_left] using lipschitz_with.dist_right (0 : E) +@[priority 100, to_additive] -- See note [lower instance priority] +instance [seminormed_group E] : seminormed_group Eᵒᵈ := ‹seminormed_group E› -lemma lipschitz_with_one_nnnorm : lipschitz_with 1 (has_nnnorm.nnnorm : E → ℝ≥0) := -lipschitz_with_one_norm +@[priority 100, to_additive] -- See note [lower instance priority] +instance [seminormed_comm_group E] : seminormed_comm_group Eᵒᵈ := ‹seminormed_comm_group E› -lemma uniform_continuous_norm : uniform_continuous (norm : E → ℝ) := -lipschitz_with_one_norm.uniform_continuous +@[priority 100, to_additive] -- See note [lower instance priority] +instance [normed_group E] : normed_group Eᵒᵈ := ‹normed_group E› -lemma uniform_continuous_nnnorm : uniform_continuous (λ (a : E), ∥a∥₊) := -uniform_continuous_norm.subtype_mk _ +@[priority 100, to_additive] -- See note [lower instance priority] +instance [normed_comm_group E] : normed_comm_group Eᵒᵈ := ‹normed_comm_group E› -@[continuity] lemma continuous_norm : continuous (λg:E, ∥g∥) := -uniform_continuous_norm.continuous +end order_dual -@[continuity] -lemma continuous_nnnorm : continuous (λ (a : E), ∥a∥₊) := -uniform_continuous_nnnorm.continuous +/-! ### Binary product of normed groups -/ -lemma tendsto_norm {x : E} : tendsto (λg : E, ∥g∥) (𝓝 x) (𝓝 ∥x∥) := -continuous_norm.continuous_at +section has_norm +variables [has_norm E] [has_norm F] {x : E × F} {r : ℝ} -lemma tendsto_norm_zero : tendsto (λg : E, ∥g∥) (𝓝 0) (𝓝 0) := -continuous_norm.tendsto' 0 0 norm_zero +instance : has_norm (E × F) := ⟨λ x, ∥x.1∥ ⊔ ∥x.2∥⟩ -/-- A helper lemma used to prove that the (scalar or usual) product of a function that tends to zero -and a bounded function tends to zero. This lemma is formulated for any binary operation -`op : E → F → G` with an estimate `∥op x y∥ ≤ A * ∥x∥ * ∥y∥` for some constant A instead of -multiplication so that it can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`. -/ -lemma filter.tendsto.op_zero_is_bounded_under_le' {f : α → E} {g : α → F} {l : filter α} - (hf : tendsto f l (𝓝 0)) (hg : is_bounded_under (≤) l (norm ∘ g)) (op : E → F → G) - (h_op : ∃ A, ∀ x y, ∥op x y∥ ≤ A * ∥x∥ * ∥y∥) : - tendsto (λ x, op (f x) (g x)) l (𝓝 0) := -begin - cases h_op with A h_op, - rcases hg with ⟨C, hC⟩, rw eventually_map at hC, - rw normed_add_comm_group.tendsto_nhds_zero at hf ⊢, - intros ε ε₀, - rcases exists_pos_mul_lt ε₀ (A * C) with ⟨δ, δ₀, hδ⟩, - filter_upwards [hf δ δ₀, hC] with i hf hg, - refine (h_op _ _).trans_lt _, - cases le_total A 0 with hA hA, - { exact (mul_nonpos_of_nonpos_of_nonneg (mul_nonpos_of_nonpos_of_nonneg hA (norm_nonneg _)) - (norm_nonneg _)).trans_lt ε₀ }, - calc A * ∥f i∥ * ∥g i∥ ≤ A * δ * C : - mul_le_mul (mul_le_mul_of_nonneg_left hf.le hA) hg (norm_nonneg _) (mul_nonneg hA δ₀.le) - ... = A * C * δ : mul_right_comm _ _ _ - ... < ε : hδ -end - -/-- A helper lemma used to prove that the (scalar or usual) product of a function that tends to zero -and a bounded function tends to zero. This lemma is formulated for any binary operation -`op : E → F → G` with an estimate `∥op x y∥ ≤ ∥x∥ * ∥y∥` instead of multiplication so that it -can be applied to `(*)`, `flip (*)`, `(•)`, and `flip (•)`. -/ -lemma filter.tendsto.op_zero_is_bounded_under_le {f : α → E} {g : α → F} {l : filter α} - (hf : tendsto f l (𝓝 0)) (hg : is_bounded_under (≤) l (norm ∘ g)) (op : E → F → G) - (h_op : ∀ x y, ∥op x y∥ ≤ ∥x∥ * ∥y∥) : - tendsto (λ x, op (f x) (g x)) l (𝓝 0) := -hf.op_zero_is_bounded_under_le' hg op ⟨1, λ x y, (one_mul (∥x∥)).symm ▸ h_op x y⟩ +lemma prod.norm_def (x : E × F) : ∥x∥ = (max ∥x.1∥ ∥x.2∥) := rfl +lemma norm_fst_le (x : E × F) : ∥x.1∥ ≤ ∥x∥ := le_max_left _ _ +lemma norm_snd_le (x : E × F) : ∥x.2∥ ≤ ∥x∥ := le_max_right _ _ -section +lemma norm_prod_le_iff : ∥x∥ ≤ r ↔ ∥x.1∥ ≤ r ∧ ∥x.2∥ ≤ r := max_le_iff -variables {l : filter α} {f : α → E} {a : E} +end has_norm -lemma filter.tendsto.norm (h : tendsto f l (𝓝 a)) : tendsto (λ x, ∥f x∥) l (𝓝 ∥a∥) := -tendsto_norm.comp h +section seminormed_group +variables [seminormed_group E] [seminormed_group F] -lemma filter.tendsto.nnnorm (h : tendsto f l (𝓝 a)) : - tendsto (λ x, ∥f x∥₊) l (𝓝 (∥a∥₊)) := -tendsto.comp continuous_nnnorm.continuous_at h - -end +/-- Product of seminormed groups, using the sup norm. -/ +@[to_additive "Product of seminormed groups, using the sup norm."] +instance : seminormed_group (E × F) := +⟨λ x y, by simp only [prod.norm_def, prod.dist_eq, dist_eq_norm_div, prod.fst_div, prod.snd_div]⟩ -section +@[to_additive prod.nnnorm_def'] +lemma prod.nnorm_def (x : E × F) : ∥x∥₊ = (max ∥x.1∥₊ ∥x.2∥₊) := rfl -variables [topological_space α] {f : α → E} {s : set α} {a : α} {b : E} +end seminormed_group -lemma continuous.norm (h : continuous f) : continuous (λ x, ∥f x∥) := continuous_norm.comp h +/-- Product of seminormed groups, using the sup norm. -/ +@[to_additive "Product of seminormed groups, using the sup norm."] +instance [seminormed_comm_group E] [seminormed_comm_group F] : seminormed_comm_group (E × F) := +{ ..prod.seminormed_group } -lemma continuous.nnnorm (h : continuous f) : continuous (λ x, ∥f x∥₊) := -continuous_nnnorm.comp h +/-- Product of normed groups, using the sup norm. -/ +@[to_additive "Product of normed groups, using the sup norm."] +instance [normed_group E] [normed_group F] : normed_group (E × F) := { ..prod.seminormed_group } -lemma continuous_at.norm (h : continuous_at f a) : continuous_at (λ x, ∥f x∥) a := h.norm +/-- Product of normed groups, using the sup norm. -/ +@[to_additive "Product of normed groups, using the sup norm."] +instance [normed_comm_group E] [normed_comm_group F] : normed_comm_group (E × F) := +{ ..prod.seminormed_group } -lemma continuous_at.nnnorm (h : continuous_at f a) : continuous_at (λ x, ∥f x∥₊) a := h.nnnorm -lemma continuous_within_at.norm (h : continuous_within_at f s a) : - continuous_within_at (λ x, ∥f x∥) s a := -h.norm +/-! ### Finite product of normed groups -/ -lemma continuous_within_at.nnnorm (h : continuous_within_at f s a) : - continuous_within_at (λ x, ∥f x∥₊) s a := -h.nnnorm +section pi +variables {π : ι → Type*} [fintype ι] -lemma continuous_on.norm (h : continuous_on f s) : continuous_on (λ x, ∥f x∥) s := -λ x hx, (h x hx).norm +section seminormed_group +variables [Π i, seminormed_group (π i)] [seminormed_group E] (f : Π i, π i) {x : Π i, π i} {r : ℝ} -lemma continuous_on.nnnorm (h : continuous_on f s) : continuous_on (λ x, ∥f x∥₊) s := -λ x hx, (h x hx).nnnorm +/-- Finite product of seminormed groups, using the sup norm. -/ +@[to_additive "Finite product of seminormed groups, using the sup norm."] +instance : seminormed_group (Π i, π i) := +{ norm := λ f, ↑(finset.univ.sup (λ b, ∥f b∥₊)), + dist_eq := λ x y, + congr_arg (coe : ℝ≥0 → ℝ) $ congr_arg (finset.sup finset.univ) $ funext $ λ a, + show nndist (x a) (y a) = ∥x a / y a∥₊, from nndist_eq_nnnorm_div (x a) (y a) } -end +@[to_additive pi.norm_def] lemma pi.norm_def' : ∥f∥ = ↑(finset.univ.sup (λ b, ∥f b∥₊)) := rfl +@[to_additive pi.nnnorm_def] lemma pi.nnnorm_def' : ∥f∥₊ = finset.univ.sup (λ b, ∥f b∥₊) := +subtype.eta _ _ -/-- If `∥y∥→∞`, then we can assume `y≠x` for any fixed `x`. -/ -lemma eventually_ne_of_tendsto_norm_at_top {l : filter α} {f : α → E} - (h : tendsto (λ y, ∥f y∥) l at_top) (x : E) : - ∀ᶠ y in l, f y ≠ x := -(h.eventually_ne_at_top _).mono $ λ x, ne_of_apply_ne norm +/-- The seminorm of an element in a product space is `≤ r` if and only if the norm of each +component is. -/ +@[to_additive pi_norm_le_iff "The seminorm of an element in a product space is `≤ r` if and only if +the norm of each component is."] +lemma pi_norm_le_iff' (hr : 0 ≤ r) : ∥x∥ ≤ r ↔ ∀ i, ∥x i∥ ≤ r := +by simp only [←dist_one_right, dist_pi_le_iff hr, pi.one_apply] -@[priority 100] -- see Note [lower instance priority] -instance seminormed_add_comm_group.has_lipschitz_add : has_lipschitz_add E := -{ lipschitz_add := ⟨2, lipschitz_with.prod_fst.add lipschitz_with.prod_snd⟩ } +@[to_additive pi_nnnorm_le_iff] +lemma pi_nnnorm_le_iff' {r : ℝ≥0} : ∥x∥₊ ≤ r ↔ ∀ i, ∥x i∥₊ ≤ r := pi_norm_le_iff' r.coe_nonneg -/-- A seminormed group is a uniform additive group, i.e., addition and subtraction are uniformly -continuous. -/ -@[priority 100] -- see Note [lower instance priority] -instance normed_uniform_group : uniform_add_group E := -⟨(lipschitz_with.prod_fst.sub lipschitz_with.prod_snd).uniform_continuous⟩ +/-- The seminorm of an element in a product space is `< r` if and only if the norm of each +component is. -/ +@[to_additive pi_norm_lt_iff "The seminorm of an element in a product space is `< r` if and only if +the norm of each component is."] +lemma pi_norm_lt_iff' (hr : 0 < r) : ∥x∥ < r ↔ ∀ i, ∥x i∥ < r := +by simp only [←dist_one_right, dist_pi_lt_iff hr, pi.one_apply] -@[priority 100] -- see Note [lower instance priority] -instance normed_top_group : topological_add_group E := -by apply_instance -- short-circuit type class inference +@[to_additive pi_nnnorm_lt_iff] +lemma pi_nnnorm_lt_iff' {r : ℝ≥0} (hr : 0 < r) : ∥x∥₊ < r ↔ ∀ i, ∥x i∥₊ < r := pi_norm_lt_iff' hr -lemma seminormed_add_comm_group.mem_closure_iff {s : set E} {x : E} : - x ∈ closure s ↔ ∀ ε > 0, ∃ y ∈ s, ∥x - y∥ < ε := -by simp [metric.mem_closure_iff, dist_eq_norm] +@[to_additive norm_le_pi_norm] +lemma norm_le_pi_norm' (i : ι) : ∥f i∥ ≤ ∥f∥ := (pi_norm_le_iff' $ norm_nonneg' _).1 le_rfl i -lemma norm_le_zero_iff' [t0_space E] {g : E} : - ∥g∥ ≤ 0 ↔ g = 0 := -begin - letI : normed_add_comm_group E := { to_metric_space := metric.of_t0_pseudo_metric_space E, - .. ‹seminormed_add_comm_group E› }, - rw [← dist_zero_right], exact dist_le_zero -end +@[to_additive nnnorm_le_pi_nnnorm] +lemma nnnorm_le_pi_nnnorm' (i : ι) : ∥f i∥₊ ≤ ∥f∥₊ := norm_le_pi_norm' _ i -lemma norm_eq_zero_iff' [t0_space E] {g : E} : ∥g∥ = 0 ↔ g = 0 := -(norm_nonneg g).le_iff_eq.symm.trans norm_le_zero_iff' +@[simp, to_additive pi_norm_const] +lemma pi_norm_const' [nonempty ι] (a : E) : ∥(λ i : ι, a)∥ = ∥a∥ := +by simpa only [←dist_one_right] using dist_pi_const a 1 -lemma norm_pos_iff' [t0_space E] {g : E} : 0 < ∥g∥ ↔ g ≠ 0 := -by rw [← not_le, norm_le_zero_iff'] +@[simp, to_additive pi_nnnorm_const] +lemma pi_nnnorm_const' [nonempty ι] (a : E) : ∥(λ i : ι, a)∥₊ = ∥a∥₊ := nnreal.eq $ pi_norm_const' a -lemma cauchy_seq_sum_of_eventually_eq {u v : ℕ → E} {N : ℕ} (huv : ∀ n ≥ N, u n = v n) - (hv : cauchy_seq (λ n, ∑ k in range (n+1), v k)) : cauchy_seq (λ n, ∑ k in range (n + 1), u k) := -begin - let d : ℕ → E := λ n, ∑ k in range (n + 1), (u k - v k), - rw show (λ n, ∑ k in range (n + 1), u k) = d + (λ n, ∑ k in range (n + 1), v k), - by { ext n, simp [d] }, - have : ∀ n ≥ N, d n = d N, - { intros n hn, - dsimp [d], - rw eventually_constant_sum _ hn, - intros m hm, - simp [huv m hm] }, - exact (tendsto_at_top_of_eventually_const this).cauchy_seq.add hv -end +/-- The $L^1$ norm is less than the $L^\infty$ norm scaled by the cardinality. -/ +@[to_additive pi.sum_norm_apply_le_norm "The $L^1$ norm is less than the $L^\\infty$ norm scaled by +the cardinality."] +lemma pi.sum_norm_apply_le_norm' : ∑ i, ∥f i∥ ≤ fintype.card ι • ∥f∥ := +finset.sum_le_card_nsmul _ _ _ $ λ i hi, norm_le_pi_norm' _ i -end seminormed_add_comm_group +/-- The $L^1$ norm is less than the $L^\infty$ norm scaled by the cardinality. -/ +@[to_additive pi.sum_nnnorm_apply_le_nnnorm "The $L^1$ norm is less than the $L^\\infty$ norm scaled +by the cardinality."] +lemma pi.sum_nnnorm_apply_le_nnnorm' : ∑ i, ∥f i∥₊ ≤ fintype.card ι • ∥f∥₊ := +nnreal.coe_sum.trans_le $ pi.sum_norm_apply_le_norm' _ -section normed_add_comm_group +end seminormed_group -/-- Construct a `normed_add_comm_group` from a `seminormed_add_comm_group` satisfying -`∀ x, ∥x∥ = 0 → x = 0`. This avoids having to go back to the `(pseudo_)metric_space` level -when declaring a `normed_add_comm_group` instance as a special case of a more general -`seminormed_add_comm_group` instance. -/ -def normed_add_comm_group.of_separation [h₁ : seminormed_add_comm_group E] - (h₂ : ∀ x : E, ∥x∥ = 0 → x = 0) : normed_add_comm_group E := -{ to_metric_space := - { eq_of_dist_eq_zero := λ x y hxy, by rw h₁.dist_eq at hxy; rw ← sub_eq_zero; exact h₂ _ hxy }, - ..h₁ } +/-- Finite product of seminormed groups, using the sup norm. -/ +@[to_additive "Finite product of seminormed groups, using the sup norm."] +instance pi.seminormed_comm_group [Π i, seminormed_comm_group (π i)] : + seminormed_comm_group (Π i, π i) := +{ ..pi.seminormed_group } -/-- Construct a normed group from a translation invariant distance -/ -def normed_add_comm_group.of_add_dist [has_norm E] [add_comm_group E] [metric_space E] - (H1 : ∀ x : E, ∥x∥ = dist x 0) - (H2 : ∀ x y z : E, dist x y ≤ dist (x + z) (y + z)) : normed_add_comm_group E := -{ dist_eq := λ x y, begin - rw H1, apply le_antisymm, - { rw [sub_eq_add_neg, ← add_right_neg y], apply H2 }, - { have := H2 (x-y) 0 y, rwa [sub_add_cancel, zero_add] at this } - end } +/-- Finite product of normed groups, using the sup norm. -/ +@[to_additive "Finite product of seminormed groups, using the sup norm."] +instance pi.normed_group [Π i, normed_group (π i)] : normed_group (Π i, π i) := +{ ..pi.seminormed_group } -/-- Construct a normed group from a norm, i.e., registering the distance and the metric space -structure from the norm properties. Note that in most cases this instance creates bad definitional -equalities (e.g., it does not take into account a possibly existing `uniform_space` instance on -`E`). -/ -def add_group_norm.to_normed_add_comm_group [add_comm_group E] (f : add_group_norm E) : - normed_add_comm_group E := -{ eq_of_dist_eq_zero := λ x y h, sub_eq_zero.1 $ eq_zero_of_map_eq_zero f h, - ..f.to_add_group_seminorm.to_seminormed_add_comm_group } +/-- Finite product of normed groups, using the sup norm. -/ +@[to_additive "Finite product of seminormed groups, using the sup norm."] +instance pi.normed_comm_group [Π i, normed_comm_group (π i)] : normed_comm_group (Π i, π i) := +{ ..pi.seminormed_group } -variables [normed_add_comm_group E] [normed_add_comm_group F] {x y : E} +end pi -@[simp] lemma norm_eq_zero {g : E} : ∥g∥ = 0 ↔ g = 0 := norm_eq_zero_iff' +/-! ### Subgroups of normed groups -/ -lemma norm_ne_zero_iff {g : E} : ∥g∥ ≠ 0 ↔ g ≠ 0 := not_congr norm_eq_zero +namespace subgroup +section seminormed_group +variables [seminormed_group E] {s : subgroup E} -@[simp] lemma norm_pos_iff {g : E} : 0 < ∥ g ∥ ↔ g ≠ 0 := norm_pos_iff' +/-- A subgroup of a seminormed group is also a seminormed group, +with the restriction of the norm. -/ +@[to_additive "A subgroup of a seminormed group is also a seminormed group, +with the restriction of the norm."] +instance seminormed_group : seminormed_group s := seminormed_group.induced s.subtype -@[simp] lemma norm_le_zero_iff {g : E} : ∥g∥ ≤ 0 ↔ g = 0 := norm_le_zero_iff' +/-- If `x` is an element of a subgroup `s` of a seminormed group `E`, its norm in `s` is equal to +its norm in `E`. -/ +@[simp, to_additive "If `x` is an element of a subgroup `s` of a seminormed group `E`, its norm in +`s` is equal to its norm in `E`."] +lemma coe_norm (x : s) : ∥x∥ = ∥(x : E)∥ := rfl -lemma norm_sub_eq_zero_iff {u v : E} : ∥u - v∥ = 0 ↔ u = v := -by rw [norm_eq_zero, sub_eq_zero] +/-- If `x` is an element of a subgroup `s` of a seminormed group `E`, its norm in `s` is equal to +its norm in `E`. -lemma norm_sub_pos_iff : 0 < ∥x - y∥ ↔ x ≠ y := -by { rw [(norm_nonneg _).lt_iff_ne, ne_comm], exact norm_sub_eq_zero_iff.not } +This is a reversed version of the `simp` lemma `subgroup.coe_norm` for use by `norm_cast`. -/ +@[norm_cast, to_additive "If `x` is an element of a subgroup `s` of a seminormed group `E`, its norm +in `s` is equal to its norm in `E`. -lemma eq_of_norm_sub_le_zero {g h : E} (a : ∥g - h∥ ≤ 0) : g = h := -by rwa [← sub_eq_zero, ← norm_le_zero_iff] +This is a reversed version of the `simp` lemma `add_subgroup.coe_norm` for use by `norm_cast`."] +lemma norm_coe {s : subgroup E} (x : s) : ∥(x : E)∥ = ∥x∥ := rfl -lemma eq_of_norm_sub_eq_zero {u v : E} (h : ∥u - v∥ = 0) : u = v := -norm_sub_eq_zero_iff.1 h +end seminormed_group -@[simp] lemma nnnorm_eq_zero {a : E} : ∥a∥₊ = 0 ↔ a = 0 := -by rw [← nnreal.coe_eq_zero, coe_nnnorm, norm_eq_zero] +@[to_additive] instance seminormed_comm_group [seminormed_comm_group E] {s : subgroup E} : + seminormed_comm_group s := +seminormed_comm_group.induced s.subtype -lemma nnnorm_ne_zero_iff {g : E} : ∥g∥₊ ≠ 0 ↔ g ≠ 0 := not_congr nnnorm_eq_zero +@[to_additive] instance normed_group [normed_group E] {s : subgroup E} : normed_group s := +normed_group.induced s.subtype subtype.coe_injective -variables (E) +@[to_additive] +instance normed_comm_group [normed_comm_group E] {s : subgroup E} : normed_comm_group s := +normed_comm_group.induced s.subtype subtype.coe_injective -/-- The norm of a normed group as an additive group norm. -/ -def norm_add_group_norm : add_group_norm E := ⟨norm, norm_zero, norm_add_le, norm_neg, - λ {g : E}, norm_eq_zero.mp⟩ +end subgroup -@[simp] lemma coe_norm_add_group_norm : ⇑(norm_add_group_norm E) = norm := rfl +/-! ### Submodules of normed groups -/ -variables {E} +namespace submodule -/-- An injective group homomorphism from an `add_comm_group` to a `normed_add_comm_group` induces a -`normed_add_comm_group` structure on the domain. +/-- A submodule of a seminormed group is also a seminormed group, with the restriction of the norm. +-/ +-- See note [implicit instance arguments] +instance seminormed_add_comm_group {_ : ring 𝕜} [seminormed_add_comm_group E] {_ : module 𝕜 E} + (s : submodule 𝕜 E) : + seminormed_add_comm_group s := +seminormed_add_comm_group.induced s.subtype.to_add_monoid_hom -See note [reducible non-instances]. -/ -@[reducible] -def normed_add_comm_group.induced {E} [add_comm_group E] - (f : E →+ F) (h : function.injective f) : normed_add_comm_group E := -{ .. seminormed_add_comm_group.induced f, - .. metric_space.induced f h normed_add_comm_group.to_metric_space, } +/-- If `x` is an element of a submodule `s` of a normed group `E`, its norm in `s` is equal to its +norm in `E`. -/ +-- See note [implicit instance arguments]. +@[simp] lemma coe_norm {_ : ring 𝕜} [seminormed_add_comm_group E] {_ : module 𝕜 E} + {s : submodule 𝕜 E} (x : s) : + ∥x∥ = ∥(x : E)∥ := rfl -/-- A subgroup of a normed group is also a normed group, with the restriction of the norm. -/ -instance add_subgroup.normed_add_comm_group (s : add_subgroup E) : normed_add_comm_group s := -normed_add_comm_group.induced s.subtype subtype.coe_injective +/-- If `x` is an element of a submodule `s` of a normed group `E`, its norm in `E` is equal to its +norm in `s`. -/-- A submodule of a normed group is also a normed group, with the restriction of the norm. +This is a reversed version of the `simp` lemma `submodule.coe_norm` for use by `norm_cast`. -/ +-- See note [implicit instance arguments]. +@[norm_cast] lemma norm_coe {_ : ring 𝕜} [seminormed_add_comm_group E] {_ : module 𝕜 E} + {s : submodule 𝕜 E} (x : s) : + ∥(x : E)∥ = ∥x∥ := rfl -See note [implicit instance arguments]. -/ -instance submodule.normed_add_comm_group {𝕜 E : Type*} {_ : ring 𝕜} [normed_add_comm_group E] - {_ : module 𝕜 E} (s : submodule 𝕜 E) : normed_add_comm_group s := +/-- A submodule of a normed group is also a normed group, with the restriction of the norm. -/ +-- See note [implicit instance arguments]. +instance {_ : ring 𝕜} [normed_add_comm_group E] {_ : module 𝕜 E} (s : submodule 𝕜 E) : + normed_add_comm_group s := { ..submodule.seminormed_add_comm_group s } -instance ulift.normed_add_comm_group : normed_add_comm_group (ulift E) := -{ ..ulift.seminormed_add_comm_group } - -/-- normed group instance on the product of two normed groups, using the sup norm. -/ -instance prod.normed_add_comm_group : normed_add_comm_group (E × F) := -{ ..prod.seminormed_add_comm_group } - -/-- normed group instance on the product of finitely many normed groups, using the sup norm. -/ -instance pi.normed_add_comm_group {π : ι → Type*} [fintype ι] [Π i, normed_add_comm_group (π i)] : - normed_add_comm_group (Πi, π i) := { ..pi.seminormed_add_comm_group } - -lemma tendsto_norm_sub_self_punctured_nhds (a : E) : tendsto (λ x, ∥x - a∥) (𝓝[≠] a) (𝓝[>] 0) := -(tendsto_norm_sub_self a).inf $ tendsto_principal_principal.2 $ λ x hx, - norm_pos_iff.2 $ sub_ne_zero.2 hx - -lemma tendsto_norm_nhds_within_zero : tendsto (norm : E → ℝ) (𝓝[≠] 0) (𝓝[>] 0) := -tendsto_norm_zero.inf $ tendsto_principal_principal.2 $ λ x, norm_pos_iff.2 - -/-! Some relations with `has_compact_support` -/ - -lemma has_compact_support_norm_iff [topological_space α] {f : α → E} : - has_compact_support (λ x, ∥ f x ∥) ↔ has_compact_support f := -has_compact_support_comp_left $ λ x, norm_eq_zero - -alias has_compact_support_norm_iff ↔ _ has_compact_support.norm - -lemma continuous.bounded_above_of_compact_support [topological_space α] {f : α → E} - (hf : continuous f) (hsupp : has_compact_support f) : ∃ C, ∀ x, ∥f x∥ ≤ C := -by simpa [bdd_above_def] using hf.norm.bdd_above_range_of_has_compact_support hsupp.norm - - -end normed_add_comm_group +end submodule diff --git a/src/analysis/normed_space/bounded_linear_maps.lean b/src/analysis/normed_space/bounded_linear_maps.lean index 67d91463d48ba..3e53c007d2d19 100644 --- a/src/analysis/normed_space/bounded_linear_maps.lean +++ b/src/analysis/normed_space/bounded_linear_maps.lean @@ -315,7 +315,8 @@ H.is_O.comp_tendsto le_top protected lemma is_bounded_bilinear_map.is_O' (h : is_bounded_bilinear_map 𝕜 f) : f =O[⊤] (λ p : E × F, ∥p∥ * ∥p∥) := -h.is_O.trans (asymptotics.is_O_fst_prod'.norm_norm.mul asymptotics.is_O_snd_prod'.norm_norm) +h.is_O.trans $ (@asymptotics.is_O_fst_prod' _ E F _ _ _ _).norm_norm.mul + (@asymptotics.is_O_snd_prod' _ E F _ _ _ _).norm_norm lemma is_bounded_bilinear_map.map_sub_left (h : is_bounded_bilinear_map 𝕜 f) {x y : E} {z : F} : f (x - y, z) = f (x, z) - f(y, z) := diff --git a/src/analysis/normed_space/continuous_affine_map.lean b/src/analysis/normed_space/continuous_affine_map.lean index 9ab473114ca9d..aa33495ccfded 100644 --- a/src/analysis/normed_space/continuous_affine_map.lean +++ b/src/analysis/normed_space/continuous_affine_map.lean @@ -176,7 +176,7 @@ add_group_norm.to_normed_add_comm_group simp only [function.const_apply, coe_const, norm_eq_zero] at h₁, rw h₁, refl, }, - { rw [norm_eq_zero_iff', cont_linear_eq_zero_iff_exists_const] at h₁, + { rw [norm_eq_zero', cont_linear_eq_zero_iff_exists_const] at h₁, obtain ⟨q, rfl⟩ := h₁, simp only [function.const_apply, coe_const, norm_le_zero_iff] at h₂, rw h₂, diff --git a/src/analysis/normed_space/lattice_ordered_group.lean b/src/analysis/normed_space/lattice_ordered_group.lean index 594b54a8bf7de..8f873933c3b34 100644 --- a/src/analysis/normed_space/lattice_ordered_group.lean +++ b/src/analysis/normed_space/lattice_ordered_group.lean @@ -60,12 +60,6 @@ A normed lattice ordered group is an ordered additive commutative group instance normed_lattice_add_comm_group_to_ordered_add_comm_group {α : Type*} [h : normed_lattice_add_comm_group α] : ordered_add_comm_group α := { ..h } -/-- -Let `α` be a normed group with a partial order. Then the order dual is also a normed group. --/ -@[priority 100] -- see Note [lower instance priority] -instance {α : Type*} : Π [normed_add_comm_group α], normed_add_comm_group αᵒᵈ := id - variables {α : Type*} [normed_lattice_add_comm_group α] open lattice_ordered_comm_group @@ -86,8 +80,8 @@ normed lattice ordered group. -/ @[priority 100] -- see Note [lower instance priority] instance : normed_lattice_add_comm_group αᵒᵈ := -{ add_le_add_left := λ a b, add_le_add_left, - solid := dual_solid } +{ solid := dual_solid, + ..order_dual.ordered_add_comm_group, ..order_dual.normed_add_comm_group } lemma norm_abs_eq_norm (a : α) : ∥|a|∥ = ∥a∥ := (solid (abs_abs a).le).antisymm (solid (abs_abs a).symm.le) diff --git a/src/analysis/specific_limits/normed.lean b/src/analysis/specific_limits/normed.lean index 6ae1ad1fc27b0..07aa0f37e9d40 100644 --- a/src/analysis/specific_limits/normed.lean +++ b/src/analysis/specific_limits/normed.lean @@ -558,14 +558,12 @@ begin nat.succ_sub_succ_eq_sub, tsub_zero], apply (normed_field.tendsto_zero_smul_of_tendsto_zero_of_bounded hf0 ⟨b, eventually_map.mpr $ eventually_of_forall $ λ n, hgb $ n+1⟩).cauchy_seq.add, - apply (cauchy_seq_range_of_norm_bounded _ _ (_ : ∀ n, _ ≤ b * |f(n+1) - f(n)|)).neg, - { exact normed_uniform_group }, + refine (cauchy_seq_range_of_norm_bounded _ _ (λ n, _ : ∀ n, _ ≤ b * |f(n+1) - f(n)|)).neg, { simp_rw [abs_of_nonneg (sub_nonneg_of_le (hfa (nat.le_succ _))), ← mul_sum], apply real.uniform_continuous_const_mul.comp_cauchy_seq, simp_rw [sum_range_sub, sub_eq_add_neg], exact (tendsto.cauchy_seq hf0).add_const }, - { intro n, - rw [norm_smul, mul_comm], + { rw [norm_smul, mul_comm], exact mul_le_mul_of_nonneg_right (hgb _) (abs_nonneg _) }, end diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 4462c792638b4..30eaf3bf12450 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -1341,7 +1341,7 @@ instance : normed_lattice_add_comm_group (α →ᵇ β) := rw norm_le (norm_nonneg _), exact λ t, (i1 t).trans (norm_coe_le_norm g t), end, - ..bounded_continuous_function.lattice, } + ..bounded_continuous_function.lattice, ..bounded_continuous_function.seminormed_add_comm_group } end normed_lattice_ordered_group From 9cb7e93708e2c66ea7049dd2450ead6b9897f20c Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 6 Oct 2022 22:17:01 +0000 Subject: [PATCH 605/828] =?UTF-8?q?feat(tactic/norm=5Fnum):=20permit=20dis?= =?UTF-8?q?abling=20the=20calculation=20of=20`/`,=20`%`,=20`=E2=88=A3`=20(?= =?UTF-8?q?#16588)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For teaching, I would like to be able to turn off `norm_num`'s power to calculate `/`, `%`, `∣`. This can be achieved by moving `norm_num.eval_nat_succ_ext` from the core tactic to an extension, so that it can be locally disabled with ```lean local attribute [-norm_num] norm_num.eval_nat_int_ext ``` (Interested users: note that this is only useful if you also make `int.has_div`, `nat.has_dvd`, etc irreducible with ```lean local attribute [irreducible] int.has_div nat.has_dvd ... ``` since otherwise many of these goals can be solved by `refl`.) Zulip: https://leanprover.zulipchat.com/#narrow/stream/239415-metaprogramming-.2F-tactics/topic/Selectively.20weaken.20norm_num --- src/tactic/norm_num.lean | 176 +++++++++++++++++++++------------------ test/norm_num.lean | 26 ++++++ 2 files changed, 120 insertions(+), 82 deletions(-) diff --git a/src/tactic/norm_num.lean b/src/tactic/norm_num.lean index 0d70e6fa9d0b2..38083d87ff708 100644 --- a/src/tactic/norm_num.lean +++ b/src/tactic/norm_num.lean @@ -1177,69 +1177,6 @@ meta def prove_nat_succ (ic : instance_cache) : expr → tactic (instance_cache p ← mk_eq_refl e, return (ic, n, e, p) -lemma nat_div (a b q r m : ℕ) (hm : q * b = m) (h : r + m = a) (h₂ : r < b) : a / b = q := -by rw [← h, ← hm, nat.add_mul_div_right _ _ (lt_of_le_of_lt (nat.zero_le _) h₂), - nat.div_eq_of_lt h₂, zero_add] - -lemma int_div (a b q r m : ℤ) (hm : q * b = m) (h : r + m = a) (h₁ : 0 ≤ r) (h₂ : r < b) : - a / b = q := -by rw [← h, ← hm, int.add_mul_div_right _ _ (ne_of_gt (lt_of_le_of_lt h₁ h₂)), - int.div_eq_zero_of_lt h₁ h₂, zero_add] - -lemma nat_mod (a b q r m : ℕ) (hm : q * b = m) (h : r + m = a) (h₂ : r < b) : a % b = r := -by rw [← h, ← hm, nat.add_mul_mod_self_right, nat.mod_eq_of_lt h₂] - -lemma int_mod (a b q r m : ℤ) (hm : q * b = m) (h : r + m = a) (h₁ : 0 ≤ r) (h₂ : r < b) : - a % b = r := -by rw [← h, ← hm, int.add_mul_mod_self, int.mod_eq_of_lt h₁ h₂] - -lemma int_div_neg (a b c' c : ℤ) (h : a / b = c') (h₂ : -c' = c) : a / -b = c := -h₂ ▸ h ▸ int.div_neg _ _ - -lemma int_mod_neg (a b c : ℤ) (h : a % b = c) : a % -b = c := -(int.mod_neg _ _).trans h - -/-- Given `a`,`b` numerals in `nat` or `int`, - * `prove_div_mod ic a b ff` returns `(c, ⊢ a / b = c)` - * `prove_div_mod ic a b tt` returns `(c, ⊢ a % b = c)` --/ -meta def prove_div_mod (ic : instance_cache) : - expr → expr → bool → tactic (instance_cache × expr × expr) -| a b mod := - match match_neg b with - | some b := do - (ic, c', p) ← prove_div_mod a b mod, - if mod then - return (ic, c', `(int_mod_neg).mk_app [a, b, c', p]) - else do - (ic, c, p₂) ← prove_neg ic c', - return (ic, c, `(int_div_neg).mk_app [a, b, c', c, p, p₂]) - | none := do - nb ← b.to_nat, - na ← a.to_int, - let nq := na / nb, - let nr := na % nb, - let nm := nq * nr, - (ic, q) ← ic.of_int nq, - (ic, r) ← ic.of_int nr, - (ic, m, pm) ← prove_mul_rat ic q b (rat.of_int nq) (rat.of_int nb), - (ic, p) ← prove_add_rat ic r m a (rat.of_int nr) (rat.of_int nm) (rat.of_int na), - (ic, p') ← prove_lt_nat ic r b, - if ic.α = `(nat) then - if mod then return (ic, r, `(nat_mod).mk_app [a, b, q, r, m, pm, p, p']) - else return (ic, q, `(nat_div).mk_app [a, b, q, r, m, pm, p, p']) - else if ic.α = `(int) then do - (ic, p₀) ← prove_nonneg ic r, - if mod then return (ic, r, `(int_mod).mk_app [a, b, q, r, m, pm, p, p₀, p']) - else return (ic, q, `(int_div).mk_app [a, b, q, r, m, pm, p, p₀, p']) - else failed - end - -theorem dvd_eq_nat (a b c : ℕ) (p) (h₁ : b % a = c) (h₂ : (c = 0) = p) : (a ∣ b) = p := -(propext $ by rw [← h₁, nat.dvd_iff_mod_eq_zero]).trans h₂ -theorem dvd_eq_int (a b c : ℤ) (p) (h₁ : b % a = c) (h₂ : (c = 0) = p) : (a ∣ b) = p := -(propext $ by rw [← h₁, int.dvd_iff_mod_eq_zero]).trans h₂ - theorem int_to_nat_pos (a : ℤ) (b : ℕ) (h : (by haveI := @nat.cast_coe ℤ; exact b : ℤ) = a) : a.to_nat = b := by rw ← h; simp theorem int_to_nat_neg (a : ℤ) (h : 0 < a) : (-a).to_nat = 0 := @@ -1254,28 +1191,12 @@ theorem neg_succ_of_nat (a b : ℕ) (c : ℤ) (h₁ : a + 1 = b) (h₂ : (by haveI := @nat.cast_coe ℤ; exact b : ℤ) = c) : -[1+ a] = -c := by rw [← h₂, ← h₁]; refl -/-- Evaluates some extra numeric operations on `nat` and `int`, specifically -`nat.succ`, `/` and `%`, and `∣` (divisibility). -/ -meta def eval_nat_int_ext : expr → tactic (expr × expr) +/-- Evaluates `nat.succ`, `int.to_nat`, `int.nat_abs`, `int.neg_succ_of_nat`. -/ +meta def eval_nat_int : expr → tactic (expr × expr) | e@`(nat.succ _) := do ic ← mk_instance_cache `(ℕ), (_, _, ep) ← prove_nat_succ ic e, return ep -| `(%%a / %%b) := do - c ← infer_type a >>= mk_instance_cache, - prod.snd <$> prove_div_mod c a b ff -| `(%%a % %%b) := do - c ← infer_type a >>= mk_instance_cache, - prod.snd <$> prove_div_mod c a b tt -| `(%%a ∣ %%b) := do - α ← infer_type a, - ic ← mk_instance_cache α, - th ← if α = `(nat) then return (`(dvd_eq_nat):expr) else - if α = `(int) then return `(dvd_eq_int) else failed, - (ic, c, p₁) ← prove_div_mod ic b a tt, - (ic, z) ← ic.mk_app ``has_zero.zero [], - (e', p₂) ← mk_app ``eq [c, z] >>= eval_ineq, - return (e', th.mk_app [a, b, c, e', p₁, p₂]) | `(int.to_nat %%a) := do n ← a.to_int, ic ← mk_instance_cache `(ℤ), @@ -1353,7 +1274,7 @@ meta def eval_cast : expr → tactic (expr × expr) /-- This version of `derive` does not fail when the input is already a numeral -/ meta def derive.step (e : expr) : tactic (expr × expr) := -eval_field e <|> eval_pow e <|> eval_ineq e <|> eval_cast e <|> eval_nat_int_ext e +eval_field e <|> eval_pow e <|> eval_ineq e <|> eval_cast e <|> eval_nat_int e /-- An attribute for adding additional extensions to `norm_num`. To use this attribute, put `@[norm_num]` on a tactic of type `expr → tactic (expr × expr)`; the tactic will be called on @@ -1626,3 +1547,94 @@ add_tactic_doc tags := ["simplification", "arithmetic", "decision procedure"] } end tactic + +namespace norm_num +section elementary_number_theory + +open tactic + +lemma nat_div (a b q r m : ℕ) (hm : q * b = m) (h : r + m = a) (h₂ : r < b) : a / b = q := +by rw [← h, ← hm, nat.add_mul_div_right _ _ (lt_of_le_of_lt (nat.zero_le _) h₂), + nat.div_eq_of_lt h₂, zero_add] + +lemma int_div (a b q r m : ℤ) (hm : q * b = m) (h : r + m = a) (h₁ : 0 ≤ r) (h₂ : r < b) : + a / b = q := +by rw [← h, ← hm, int.add_mul_div_right _ _ (ne_of_gt (lt_of_le_of_lt h₁ h₂)), + int.div_eq_zero_of_lt h₁ h₂, zero_add] + +lemma nat_mod (a b q r m : ℕ) (hm : q * b = m) (h : r + m = a) (h₂ : r < b) : a % b = r := +by rw [← h, ← hm, nat.add_mul_mod_self_right, nat.mod_eq_of_lt h₂] + +lemma int_mod (a b q r m : ℤ) (hm : q * b = m) (h : r + m = a) (h₁ : 0 ≤ r) (h₂ : r < b) : + a % b = r := +by rw [← h, ← hm, int.add_mul_mod_self, int.mod_eq_of_lt h₁ h₂] + +lemma int_div_neg (a b c' c : ℤ) (h : a / b = c') (h₂ : -c' = c) : a / -b = c := +h₂ ▸ h ▸ int.div_neg _ _ + +lemma int_mod_neg (a b c : ℤ) (h : a % b = c) : a % -b = c := +(int.mod_neg _ _).trans h + +/-- Given `a`,`b` numerals in `nat` or `int`, + * `prove_div_mod ic a b ff` returns `(c, ⊢ a / b = c)` + * `prove_div_mod ic a b tt` returns `(c, ⊢ a % b = c)` +-/ +meta def prove_div_mod (ic : instance_cache) : + expr → expr → bool → tactic (instance_cache × expr × expr) +| a b mod := + match match_neg b with + | some b := do + (ic, c', p) ← prove_div_mod a b mod, + if mod then + return (ic, c', `(int_mod_neg).mk_app [a, b, c', p]) + else do + (ic, c, p₂) ← prove_neg ic c', + return (ic, c, `(int_div_neg).mk_app [a, b, c', c, p, p₂]) + | none := do + nb ← b.to_nat, + na ← a.to_int, + let nq := na / nb, + let nr := na % nb, + let nm := nq * nr, + (ic, q) ← ic.of_int nq, + (ic, r) ← ic.of_int nr, + (ic, m, pm) ← prove_mul_rat ic q b (rat.of_int nq) (rat.of_int nb), + (ic, p) ← prove_add_rat ic r m a (rat.of_int nr) (rat.of_int nm) (rat.of_int na), + (ic, p') ← prove_lt_nat ic r b, + if ic.α = `(nat) then + if mod then return (ic, r, `(nat_mod).mk_app [a, b, q, r, m, pm, p, p']) + else return (ic, q, `(nat_div).mk_app [a, b, q, r, m, pm, p, p']) + else if ic.α = `(int) then do + (ic, p₀) ← prove_nonneg ic r, + if mod then return (ic, r, `(int_mod).mk_app [a, b, q, r, m, pm, p, p₀, p']) + else return (ic, q, `(int_div).mk_app [a, b, q, r, m, pm, p, p₀, p']) + else failed + end + +theorem dvd_eq_nat (a b c : ℕ) (p) (h₁ : b % a = c) (h₂ : (c = 0) = p) : (a ∣ b) = p := +(propext $ by rw [← h₁, nat.dvd_iff_mod_eq_zero]).trans h₂ +theorem dvd_eq_int (a b c : ℤ) (p) (h₁ : b % a = c) (h₂ : (c = 0) = p) : (a ∣ b) = p := +(propext $ by rw [← h₁, int.dvd_iff_mod_eq_zero]).trans h₂ + +/-- Evaluates some extra numeric operations on `nat` and `int`, specifically +`/` and `%`, and `∣` (divisibility). -/ +@[norm_num] meta def eval_nat_int_ext : expr → tactic (expr × expr) +| `(%%a / %%b) := do + c ← infer_type a >>= mk_instance_cache, + prod.snd <$> prove_div_mod c a b ff +| `(%%a % %%b) := do + c ← infer_type a >>= mk_instance_cache, + prod.snd <$> prove_div_mod c a b tt +| `(%%a ∣ %%b) := do + α ← infer_type a, + ic ← mk_instance_cache α, + th ← if α = `(nat) then return (`(dvd_eq_nat):expr) else + if α = `(int) then return `(dvd_eq_int) else failed, + (ic, c, p₁) ← prove_div_mod ic b a tt, + (ic, z) ← ic.mk_app ``has_zero.zero [], + (e', p₂) ← mk_app ``eq [c, z] >>= eval_ineq, + return (e', th.mk_app [a, b, c, e', p₁, p₂]) +| _ := failed + +end elementary_number_theory +end norm_num diff --git a/test/norm_num.lean b/test/norm_num.lean index 849cdc70793bd..0353099d98096 100644 --- a/test/norm_num.lean +++ b/test/norm_num.lean @@ -300,3 +300,29 @@ example : (- ((- (((66 - 86) - 36) / 94) - 3) / - - (77 / (56 - - - 79))) + 87) (312254/3619 : α) := by norm_num example : 2 ^ 13 - 1 = int.of_nat 8191 := by norm_num + +/-! Test the behaviour of removing one `norm_num` extension tactic. -/ +section remove_extension + +-- turn off the `norm_num` extension which deals with `/`, `%`, `∣` +local attribute [-norm_num] norm_num.eval_nat_int_ext + +example : (5 / 2:ℕ) = 2 := by success_if_fail { solve1 { norm_num } }; refl + +example : 10 = (-1 : ℤ) % 11 := by success_if_fail { solve1 { norm_num } }; refl + +example (h : (5 : ℤ) ∣ 2) : false := +begin + success_if_fail { norm_num at h }, + have : (2:ℤ) ≠ 0 := by norm_num, + exact this (int.mod_eq_zero_of_dvd h), +end + +example : 2^4-1 ∣ 2^16-1 := +begin + success_if_fail { solve1 { norm_num } }, + use 4369, + norm_num, +end + +end remove_extension From e36abd0e2fd9bf352cd0ffacc04ef7abedec5578 Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Fri, 7 Oct 2022 09:18:17 +0000 Subject: [PATCH 606/828] feat(topology): Completable field lemmas (#16805) Subfields of completable fields are completable and complete fields are completable. --- src/topology/algebra/uniform_field.lean | 31 +++++++++++++++++++ src/topology/uniform_space/separation.lean | 3 ++ .../uniform_space/uniform_embedding.lean | 4 +++ 3 files changed, 38 insertions(+) diff --git a/src/topology/algebra/uniform_field.lean b/src/topology/algebra/uniform_field.lean index d751bca617980..7539f91be14b0 100644 --- a/src/topology/algebra/uniform_field.lean +++ b/src/topology/algebra/uniform_field.lean @@ -5,6 +5,7 @@ Authors: Patrick Massot -/ import topology.algebra.uniform_ring import topology.algebra.field +import field_theory.subfield /-! # Completion of topological fields @@ -173,3 +174,33 @@ instance : topological_division_ring (hat K) := end completion end uniform_space + +variables (L : Type*) [field L] [uniform_space L] [completable_top_field L] + +instance subfield.completable_top_field (K : subfield L) : completable_top_field K := +{ nice := begin + intros F F_cau inf_F, + let i : K →+* L := K.subtype, + have hi : uniform_inducing i, from uniform_embedding_subtype_coe.to_uniform_inducing, + rw ← hi.cauchy_map_iff at F_cau ⊢, + rw [map_comm (show (i ∘ λ x, x⁻¹) = (λ x, x⁻¹) ∘ i, by {ext, refl})], + apply completable_top_field.nice _ F_cau, + rw [← filter.push_pull', ← map_zero i, ← hi.inducing.nhds_eq_comap, inf_F, filter.map_bot] + end, + ..subtype.separated_space (K : set L) } + +@[priority 100] +instance completable_top_field_of_complete (L : Type*) [field L] + [uniform_space L] [topological_division_ring L] [separated_space L] [complete_space L] : + completable_top_field L := +{ nice := λ F cau_F hF, begin + haveI : ne_bot F := cau_F.1, + rcases complete_space.complete cau_F with ⟨x, hx⟩, + have hx' : x ≠ 0, + { rintro rfl, + rw inf_eq_right.mpr hx at hF, + exact cau_F.1.ne hF }, + exact filter.tendsto.cauchy_map (calc map (λ x, x⁻¹) F ≤ map (λ x, x⁻¹) (𝓝 x) : map_mono hx + ... ≤ 𝓝 (x⁻¹) : continuous_at_inv₀ hx') + end, + ..‹separated_space L›} diff --git a/src/topology/uniform_space/separation.lean b/src/topology/uniform_space/separation.lean index a9e4be4296b81..7064aaa7c91d3 100644 --- a/src/topology/uniform_space/separation.lean +++ b/src/topology/uniform_space/separation.lean @@ -198,6 +198,9 @@ end instance separated_t3 [separated_space α] : t3_space α := by { haveI := separated_iff_t2.mp ‹_›, exact ⟨⟩ } +instance subtype.separated_space [separated_space α] (s : set α) : separated_space s := +separated_iff_t2.mpr subtype.t2_space + lemma is_closed_of_spaced_out [separated_space α] {V₀ : set (α × α)} (V₀_in : V₀ ∈ 𝓤 α) {s : set α} (hs : s.pairwise (λ x y, (x, y) ∉ V₀)) : is_closed s := begin diff --git a/src/topology/uniform_space/uniform_embedding.lean b/src/topology/uniform_space/uniform_embedding.lean index 76480ca943ace..a24dd8b853142 100644 --- a/src/topology/uniform_space/uniform_embedding.lean +++ b/src/topology/uniform_space/uniform_embedding.lean @@ -45,6 +45,10 @@ lemma uniform_inducing.basis_uniformity {f : α → β} (hf : uniform_inducing f (𝓤 α).has_basis p (λ i, prod.map f f ⁻¹' s i) := hf.1 ▸ H.comap _ +lemma uniform_inducing.cauchy_map_iff {f : α → β} (hf : uniform_inducing f) {F : filter α} : + cauchy (map f F) ↔ cauchy F := +by simp only [cauchy, map_ne_bot_iff, prod_map_map_eq, map_le_iff_le_comap, ← hf.comap_uniformity] + lemma uniform_inducing_of_compose {f : α → β} {g : β → γ} (hf : uniform_continuous f) (hg : uniform_continuous g) (hgf : uniform_inducing (g ∘ f)) : uniform_inducing f := begin From c965e54c4a6ffe899035a17cc1c1b6d73e875bb7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Fri, 7 Oct 2022 11:25:29 +0000 Subject: [PATCH 607/828] =?UTF-8?q?feat(algebraic=5Ftopology):=20the=20?= =?UTF-8?q?=C4=8Cech=20nerve=20of=20a=20split=20epi=20has=20an=20extra=20d?= =?UTF-8?q?egeneracy=20(#16779)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In this PR, it is shown that the Čech nerve of a split epimorphism has an extra degeneracy. It is also shown that if two augmented simplicial objects are isomorphic, then an extra degeneracy for one of these transports as an extra degeneracy for the other. --- src/algebraic_topology/extra_degeneracy.lean | 163 ++++++++++++++++++- 1 file changed, 157 insertions(+), 6 deletions(-) diff --git a/src/algebraic_topology/extra_degeneracy.lean b/src/algebraic_topology/extra_degeneracy.lean index a853a23864d52..414c52dbb93e4 100644 --- a/src/algebraic_topology/extra_degeneracy.lean +++ b/src/algebraic_topology/extra_degeneracy.lean @@ -5,6 +5,7 @@ Authors: Joël Riou -/ import algebraic_topology.simplicial_set +import algebraic_topology.cech_nerve import tactic.fin_cases /-! @@ -28,6 +29,8 @@ simplicial objects in any category. functor `C ⥤ D` - `sSet.augmented.standard_simplex.extra_degeneracy`: the standard `n`-simplex has an extra degeneracy +- `arrow.augmented_cech_nerve.extra_degeneracy`: the Čech nerve of a split +epimorphism has an extra degeneracy TODO @joelriou: 1) when the category `C` is preadditive and has a zero object, and @@ -35,10 +38,6 @@ TODO @joelriou: on the alternating face map complex of `X` is a homotopy equivalence of chain complexes. -2) extra degeneracy for the Čech nerve of a split epi. In particular the -universal cover EG of the classifying space of a group G has an extra -degeneracy. - ## References * [Paul G. Goerss, John F. Jardine, *Simplical Homotopy Theory*][goerss-jardine-2009] @@ -49,8 +48,6 @@ open category_theory.simplicial_object.augmented open opposite open_locale simplicial -universes u - namespace simplicial_object namespace augmented @@ -96,6 +93,38 @@ def map {D : Type*} [category D] s_comp_δ' := λ n i, by { dsimp, erw [← F.map_comp, ← F.map_comp, ed.s_comp_δ], refl, }, s_comp_σ' := λ n i, by { dsimp, erw [← F.map_comp, ← F.map_comp, ed.s_comp_σ], refl, }, } +/-- If `X` and `Y` are isomorphic augmented simplicial objects, then an extra +degeneracy for `X` gives also an extra degeneracy for `Y` -/ +def of_iso {X Y : simplicial_object.augmented C} (e : X ≅ Y) (ed : extra_degeneracy X) : + extra_degeneracy Y := +{ s' := (point.map_iso e).inv ≫ ed.s' ≫ (drop.map_iso e).hom.app (op [0]), + s := λ n, (drop.map_iso e).inv.app (op [n]) ≫ ed.s n ≫ (drop.map_iso e).hom.app (op [n+1]), + s'_comp_ε' := by simpa only [functor.map_iso, assoc, w₀, ed.s'_comp_ε_assoc] + using (point.map_iso e).inv_hom_id, + s₀_comp_δ₁' := begin + have h := w₀ e.inv, + dsimp at h ⊢, + simp only [assoc, ← simplicial_object.δ_naturality, ed.s₀_comp_δ₁_assoc, reassoc_of h], + end, + s_comp_δ₀' := λ n, begin + have h := ed.s_comp_δ₀', + dsimp at ⊢ h, + simpa only [assoc, ← simplicial_object.δ_naturality, reassoc_of h] + using congr_app (drop.map_iso e).inv_hom_id (op [n]), + end, + s_comp_δ' := λ n i, begin + have h := ed.s_comp_δ' n i, + dsimp at ⊢ h, + simp only [assoc, ← simplicial_object.δ_naturality, reassoc_of h, + ← simplicial_object.δ_naturality_assoc], + end, + s_comp_σ' := λ n i, begin + have h := ed.s_comp_σ' n i, + dsimp at ⊢ h, + simp only [assoc, ← simplicial_object.σ_naturality, reassoc_of h, + ← simplicial_object.σ_naturality_assoc], + end,} + end extra_degeneracy end augmented @@ -186,3 +215,125 @@ end standard_simplex end augmented end sSet + +namespace category_theory + +open limits + +namespace arrow + +namespace augmented_cech_nerve + +variables {C : Type*} [category C] (f : arrow C) + [∀ n : ℕ, has_wide_pullback f.right (λ i : fin (n+1), f.left) (λ i, f.hom)] + (S : split_epi f.hom) + +include S + +/-- The extra degeneracy map on the Čech nerve of a split epi. It is +given on the `0`-projection by the given section of the split epi, +and by shifting the indices on the other projections. -/ +noncomputable def extra_degeneracy.s (n : ℕ) : + f.cech_nerve.obj (op [n]) ⟶ f.cech_nerve.obj (op [n + 1]) := +wide_pullback.lift (wide_pullback.base _) + (λ i, dite (i = 0) (λ h, wide_pullback.base _ ≫ S.section_) + (λ h, wide_pullback.π _ (i.pred h))) + (λ i, begin + split_ifs, + { subst h, + simp only [assoc, split_epi.id, comp_id], }, + { simp only [wide_pullback.π_arrow], }, + end) + +@[simp] +lemma extra_degeneracy.s_comp_π_0 (n : ℕ) : + extra_degeneracy.s f S n ≫ wide_pullback.π _ 0 = wide_pullback.base _ ≫ S.section_ := +by { dsimp [extra_degeneracy.s], simpa only [wide_pullback.lift_π], } + +@[simp] +lemma extra_degeneracy.s_comp_π_succ (n : ℕ) (i : fin (n+1)) : + extra_degeneracy.s f S n ≫ wide_pullback.π _ i.succ = wide_pullback.π _ i := +begin + dsimp [extra_degeneracy.s], + simp only [wide_pullback.lift_π], + split_ifs, + { exfalso, + simpa only [fin.ext_iff, fin.coe_succ, fin.coe_zero, nat.succ_ne_zero] using h, }, + { congr, + apply fin.pred_succ, }, +end + +@[simp] +lemma extra_degeneracy.s_comp_base (n : ℕ) : + extra_degeneracy.s f S n ≫ wide_pullback.base _ = wide_pullback.base _ := +by apply wide_pullback.lift_base + +/-- The augmented Čech nerve associated to a split epimorphism has an extra degeneracy. -/ +noncomputable def extra_degeneracy : + simplicial_object.augmented.extra_degeneracy f.augmented_cech_nerve := +{ s' := S.section_ ≫ wide_pullback.lift f.hom (λ i, 𝟙 _) (λ i, by rw id_comp), + s := λ n, extra_degeneracy.s f S n, + s'_comp_ε' := + by simp only [augmented_cech_nerve_hom_app, assoc, wide_pullback.lift_base, split_epi.id], + s₀_comp_δ₁' := begin + dsimp [cech_nerve, simplicial_object.δ, simplex_category.δ], + ext j, + { fin_cases j, + simpa only [assoc, wide_pullback.lift_π, comp_id] using extra_degeneracy.s_comp_π_0 f S 0, }, + { simpa only [assoc, wide_pullback.lift_base, split_epi.id, comp_id] + using extra_degeneracy.s_comp_base f S 0, }, + end, + s_comp_δ₀' := λ n, begin + dsimp [cech_nerve, simplicial_object.δ, simplex_category.δ], + ext j, + { simpa only [assoc, wide_pullback.lift_π, id_comp] + using extra_degeneracy.s_comp_π_succ f S n j, }, + { simpa only [assoc, wide_pullback.lift_base, id_comp] + using extra_degeneracy.s_comp_base f S n, }, + end, + s_comp_δ' := λ n i, begin + dsimp [cech_nerve, simplicial_object.δ, simplex_category.δ], + ext j, + { simp only [assoc, wide_pullback.lift_π], + by_cases j = 0, + { subst h, + erw [fin.succ_succ_above_zero, extra_degeneracy.s_comp_π_0, extra_degeneracy.s_comp_π_0], + dsimp, + simp only [wide_pullback.lift_base_assoc], }, + { cases fin.eq_succ_of_ne_zero h with k hk, + subst hk, + erw [fin.succ_succ_above_succ, extra_degeneracy.s_comp_π_succ, + extra_degeneracy.s_comp_π_succ], + dsimp, + simp only [wide_pullback.lift_π], }, }, + { simp only [assoc, wide_pullback.lift_base], + erw [extra_degeneracy.s_comp_base, extra_degeneracy.s_comp_base], + dsimp, + simp only [wide_pullback.lift_base], }, + end, + s_comp_σ' := λ n i, begin + dsimp [cech_nerve, simplicial_object.σ, simplex_category.σ], + ext j, + { simp only [assoc, wide_pullback.lift_π], + by_cases j = 0, + { subst h, + erw [extra_degeneracy.s_comp_π_0, extra_degeneracy.s_comp_π_0], + dsimp, + simp only [wide_pullback.lift_base_assoc], }, + { cases fin.eq_succ_of_ne_zero h with k hk, + subst hk, + erw [fin.succ_pred_above_succ, extra_degeneracy.s_comp_π_succ, + extra_degeneracy.s_comp_π_succ], + dsimp, + simp only [wide_pullback.lift_π], }, }, + { simp only [assoc, wide_pullback.lift_base], + erw [extra_degeneracy.s_comp_base, extra_degeneracy.s_comp_base], + dsimp, + simp only [wide_pullback.lift_base], }, + end, } + +end augmented_cech_nerve + +end arrow + +end category_theory From 04e93ae30146a23686884192805c61f6483470c4 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 7 Oct 2022 13:20:34 +0000 Subject: [PATCH 608/828] perf(topology/continuous_function/stone_weierstrass): fix timeout (#16844) Squeeze simps and replace a slow `convert` by `eq.subst` with explicit motive (maybe `convert` was unfolding the instances?). From >20s to 4s on my machine. Co-authored-by: sgouezel --- src/measure_theory/integral/circle_transform.lean | 7 +++++-- .../continuous_function/stone_weierstrass.lean | 10 +++++----- 2 files changed, 10 insertions(+), 7 deletions(-) diff --git a/src/measure_theory/integral/circle_transform.lean b/src/measure_theory/integral/circle_transform.lean index 7086f0aef089c..d5beca6060b0c 100644 --- a/src/measure_theory/integral/circle_transform.lean +++ b/src/measure_theory/integral/circle_transform.lean @@ -132,8 +132,11 @@ begin have cts := continuous_on_abs_circle_transform_bounding_function hr z, have comp : is_compact (closed_ball z r ×ˢ [0, 2 * π]), { apply_rules [is_compact.prod, proper_space.is_compact_closed_ball z r, is_compact_interval], }, - have none := (nonempty_closed_ball.2 hr').prod nonempty_interval, - simpa using is_compact.exists_forall_ge comp none (cts.mono (by { intro z, simp, tauto })), + have none : (closed_ball z r ×ˢ [0, 2 * π]).nonempty := + (nonempty_closed_ball.2 hr').prod nonempty_interval, + have := is_compact.exists_forall_ge comp none (cts.mono + (by { intro z, simp only [mem_prod, mem_closed_ball, mem_univ, and_true, and_imp], tauto })), + simpa only [set_coe.forall, subtype.coe_mk, set_coe.exists], end /-- The derivative of a `circle_transform` is locally bounded. -/ diff --git a/src/topology/continuous_function/stone_weierstrass.lean b/src/topology/continuous_function/stone_weierstrass.lean index 6ba08ff989c62..3d6fa35af48fc 100644 --- a/src/topology/continuous_function/stone_weierstrass.lean +++ b/src/topology/continuous_function/stone_weierstrass.lean @@ -393,10 +393,9 @@ begin let F : C(X, 𝕜) := f - const _ (f x₂), -- Subtract the constant `f x₂` from `f`; this is still an element of the subalgebra have hFA : F ∈ A, - { refine A.sub_mem hfA _, - convert A.smul_mem A.one_mem (f x₂), - ext1, - simp }, + { refine A.sub_mem hfA (@eq.subst _ (∈ A) _ _ _ $ A.smul_mem A.one_mem $ f x₂), + ext1, simp only [coe_smul, coe_one, pi.smul_apply, + pi.one_apply, algebra.id.smul_eq_mul, mul_one, const_apply] }, -- Consider now the function `λ x, |f x - f x₂| ^ 2` refine ⟨_, ⟨(⟨is_R_or_C.norm_sq, continuous_norm_sq⟩ : C(𝕜, ℝ)).comp F, _, rfl⟩, _⟩, { -- This is also an element of the subalgebra, and takes only real values @@ -407,7 +406,8 @@ begin exact (is_R_or_C.mul_conj _).symm }, { -- And it also separates the points `x₁`, `x₂` have : f x₁ - f x₂ ≠ 0 := sub_ne_zero.mpr hf, - simpa using this }, + simpa only [comp_apply, coe_sub, coe_const, pi.sub_apply, + coe_mk, sub_self, map_zero, ne.def, norm_sq_eq_zero] using this }, end variables [compact_space X] From 07c3cf2d851866ff7198219ed3fedf42e901f25c Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 7 Oct 2022 16:10:13 +0000 Subject: [PATCH 609/828] fix(algebra/star/free): speedup (#16851) --- src/algebra/star/free.lean | 10 +++++++--- 1 file changed, 7 insertions(+), 3 deletions(-) diff --git a/src/algebra/star/free.lean b/src/algebra/star/free.lean index 9e8ef68b90740..670cce3b058be 100644 --- a/src/algebra/star/free.lean +++ b/src/algebra/star/free.lean @@ -42,9 +42,13 @@ instance : star_ring (free_algebra R X) := star_involutive := λ x, by { unfold has_star.star, simp only [function.comp_apply], - refine free_algebra.induction R X _ _ _ _ x; intros; simp [*] }, - star_mul := λ a b, by simp, - star_add := λ a b, by simp } + refine free_algebra.induction R X _ _ _ _ x, + { intros, simp only [alg_hom.commutes, mul_opposite.algebra_map_apply, mul_opposite.unop_op] }, + { intros, simp only [lift_ι_apply, mul_opposite.unop_op] }, + { intros, simp only [*, map_mul, mul_opposite.unop_mul] }, + { intros, simp only [*, map_add, mul_opposite.unop_add] } }, + star_mul := λ a b, by simp only [function.comp_app, map_mul, mul_opposite.unop_mul], + star_add := λ a b, by simp only [function.comp_app, map_add, mul_opposite.unop_add]} @[simp] lemma star_ι (x : X) : star (ι R x) = (ι R x) := From 68054185e2290e6f1ff4a2321e63ceb94d6f307a Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 7 Oct 2022 17:33:40 +0000 Subject: [PATCH 610/828] chore(data/rat,data/nat): modularize and reduce imports (#16810) While trying to understand the import tree leading to `data.rat.basic`, a small amount of modularization and import reduction. Co-authored-by: Scott Morrison --- src/algebra/euclidean_domain.lean | 2 -- src/algebra/field/basic.lean | 2 +- src/data/nat/fib.lean | 2 +- src/data/nat/{gcd.lean => gcd/basic.lean} | 19 -------------- src/data/nat/gcd/big_operators.lean | 30 ++++++++++++++++++++++ src/data/nat/prime.lean | 2 +- src/data/num/lemmas.lean | 2 +- src/data/rat/basic.lean | 4 --- src/data/rat/cast.lean | 1 + src/data/rat/defs.lean | 11 +++----- src/data/rat/encodable.lean | 22 ++++++++++++++++ src/group_theory/noncomm_pi_coprod.lean | 1 + src/number_theory/arithmetic_function.lean | 1 + src/topology/instances/ereal.lean | 1 + src/topology/instances/irrational.lean | 1 + 15 files changed, 65 insertions(+), 36 deletions(-) rename src/data/nat/{gcd.lean => gcd/basic.lean} (96%) create mode 100644 src/data/nat/gcd/big_operators.lean create mode 100644 src/data/rat/encodable.lean diff --git a/src/algebra/euclidean_domain.lean b/src/algebra/euclidean_domain.lean index a4709703046c1..f749980d318ee 100644 --- a/src/algebra/euclidean_domain.lean +++ b/src/algebra/euclidean_domain.lean @@ -3,8 +3,6 @@ Copyright (c) 2018 Louis Carlin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Louis Carlin, Mario Carneiro -/ - -import data.int.basic import algebra.field.basic /-! diff --git a/src/algebra/field/basic.lean b/src/algebra/field/basic.lean index 3ba5aa0e83b4e..1a0a0e24bda88 100644 --- a/src/algebra/field/basic.lean +++ b/src/algebra/field/basic.lean @@ -3,8 +3,8 @@ Copyright (c) 2014 Robert Lewis. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert Lewis, Leonardo de Moura, Johannes Hölzl, Mario Carneiro -/ -import algebra.hom.ring import data.rat.defs +import algebra.group_with_zero.power /-! # Division (semi)rings and (semi)fields diff --git a/src/data/nat/fib.lean b/src/data/nat/fib.lean index acf6505001d9e..71650b4c2502f 100644 --- a/src/data/nat/fib.lean +++ b/src/data/nat/fib.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Kevin Kappelmann. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kevin Kappelmann, Kyle Miller, Mario Carneiro -/ -import data.nat.gcd +import data.nat.gcd.basic import logic.function.iterate import data.finset.nat_antidiagonal import algebra.big_operators.basic diff --git a/src/data/nat/gcd.lean b/src/data/nat/gcd/basic.lean similarity index 96% rename from src/data/nat/gcd.lean rename to src/data/nat/gcd/basic.lean index c046bce81afd5..8e58b46a37551 100644 --- a/src/data/nat/gcd.lean +++ b/src/data/nat/gcd/basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ import algebra.group_power.order -import algebra.big_operators.basic /-! # Definitions and properties of `gcd`, `lcm`, and `coprime` @@ -478,24 +477,6 @@ begin exact coprime.coprime_mul_right_right hac, end -section big_operators - -open_locale big_operators - -/-- See `is_coprime.prod_left` for the corresponding lemma about `is_coprime` -/ -lemma coprime_prod_left - {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : - (∀ (i : ι), i ∈ t → coprime (s i) x) → coprime (∏ (i : ι) in t, s i) x := -finset.prod_induction s (λ y, y.coprime x) (λ a b, coprime.mul) (by simp) - -/-- See `is_coprime.prod_right` for the corresponding lemma about `is_coprime` -/ -lemma coprime_prod_right - {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : - (∀ (i : ι), i ∈ t → coprime x (s i)) → coprime x (∏ (i : ι) in t, s i) := -finset.prod_induction s (λ y, x.coprime y) (λ a b, coprime.mul_right) (by simp) - -end big_operators - lemma coprime.eq_of_mul_eq_zero {m n : ℕ} (h : m.coprime n) (hmn : m * n = 0) : m = 0 ∧ n = 1 ∨ m = 1 ∧ n = 0 := (nat.eq_zero_of_mul_eq_zero hmn).imp diff --git a/src/data/nat/gcd/big_operators.lean b/src/data/nat/gcd/big_operators.lean new file mode 100644 index 0000000000000..98c199b463738 --- /dev/null +++ b/src/data/nat/gcd/big_operators.lean @@ -0,0 +1,30 @@ +/- +Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura +-/ +import data.nat.gcd.basic +import algebra.big_operators.basic + +/-! # Lemmas about coprimality with big products. + +These lemmas are kept separate from `data.nat.gcd.basic` in order to minimize imports. +-/ + +namespace nat + +open_locale big_operators + +/-- See `is_coprime.prod_left` for the corresponding lemma about `is_coprime` -/ +lemma coprime_prod_left + {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : + (∀ (i : ι), i ∈ t → coprime (s i) x) → coprime (∏ (i : ι) in t, s i) x := +finset.prod_induction s (λ y, y.coprime x) (λ a b, coprime.mul) (by simp) + +/-- See `is_coprime.prod_right` for the corresponding lemma about `is_coprime` -/ +lemma coprime_prod_right + {ι : Type*} {x : ℕ} {s : ι → ℕ} {t : finset ι} : + (∀ (i : ι), i ∈ t → coprime x (s i)) → coprime x (∏ (i : ι) in t, s i) := +finset.prod_induction s (λ y, x.coprime y) (λ a b, coprime.mul_right) (by simp) + +end nat diff --git a/src/data/nat/prime.lean b/src/data/nat/prime.lean index dab7a6f614ecd..6baff00b813e8 100644 --- a/src/data/nat/prime.lean +++ b/src/data/nat/prime.lean @@ -5,7 +5,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import data.list.prime import data.list.sort -import data.nat.gcd +import data.nat.gcd.basic import data.nat.sqrt_norm_num import data.set.finite import tactic.wlog diff --git a/src/data/num/lemmas.lean b/src/data/num/lemmas.lean index ca40d688db3e9..a3b46d4c1e3c3 100644 --- a/src/data/num/lemmas.lean +++ b/src/data/num/lemmas.lean @@ -5,7 +5,7 @@ Authors: Mario Carneiro -/ import data.num.bitwise import data.int.char_zero -import data.nat.gcd +import data.nat.gcd.basic import data.nat.psub /-! diff --git a/src/data/rat/basic.lean b/src/data/rat/basic.lean index a3bb6366f1be5..089d227ffda4a 100644 --- a/src/data/rat/basic.lean +++ b/src/data/rat/basic.lean @@ -4,10 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ import algebra.euclidean_domain -import data.int.cast -import data.nat.gcd -import data.rat.defs -import logic.encodable.basic /-! # Field Structure on the Rational Numbers diff --git a/src/data/rat/cast.lean b/src/data/rat/cast.lean index 920d46ed757a6..7c317f67a2d11 100644 --- a/src/data/rat/cast.lean +++ b/src/data/rat/cast.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro import data.rat.order import data.int.char_zero import algebra.field.opposite +import algebra.big_operators.basic /-! # Casts for Rational Numbers diff --git a/src/data/rat/defs.lean b/src/data/rat/defs.lean index 67ffc9802c7d9..76fda20a3f398 100644 --- a/src/data/rat/defs.lean +++ b/src/data/rat/defs.lean @@ -3,9 +3,10 @@ Copyright (c) 2019 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro -/ -import data.int.basic -import data.nat.gcd -import logic.encodable.basic +import data.int.cast +import data.nat.gcd.basic +import data.pnat.basic +import tactic.nth_rewrite /-! # Basics for the Rational Numbers @@ -60,10 +61,6 @@ instance : has_repr ℚ := ⟨rat.repr⟩ instance : has_to_string ℚ := ⟨rat.repr⟩ meta instance : has_to_format ℚ := ⟨coe ∘ rat.repr⟩ -instance : encodable ℚ := encodable.of_equiv (Σ n : ℤ, {d : ℕ // 0 < d ∧ n.nat_abs.coprime d}) - ⟨λ ⟨a, b, c, d⟩, ⟨a, b, c, d⟩, λ⟨a, b, c, d⟩, ⟨a, b, c, d⟩, - λ ⟨a, b, c, d⟩, rfl, λ⟨a, b, c, d⟩, rfl⟩ - /-- Embed an integer as a rational number -/ def of_int (n : ℤ) : ℚ := ⟨n, 1, nat.one_pos, nat.coprime_one_right _⟩ diff --git a/src/data/rat/encodable.lean b/src/data/rat/encodable.lean new file mode 100644 index 0000000000000..57d6c081c6f6c --- /dev/null +++ b/src/data/rat/encodable.lean @@ -0,0 +1,22 @@ +/- +Copyright (c) 2019 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Mario Carneiro +-/ +import data.rat.defs +import logic.encodable.basic + +/-! # The rationals are `encodable`. + +As a consequence we also get the instance `countable ℚ`. + +This is kept separate from `data.rat.defs` in order to minimize imports. +-/ + +namespace rat + +instance : encodable ℚ := encodable.of_equiv (Σ n : ℤ, {d : ℕ // 0 < d ∧ n.nat_abs.coprime d}) + ⟨λ ⟨a, b, c, d⟩, ⟨a, b, c, d⟩, λ⟨a, b, c, d⟩, ⟨a, b, c, d⟩, + λ ⟨a, b, c, d⟩, rfl, λ⟨a, b, c, d⟩, rfl⟩ + +end rat diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index fdac562bc6549..3ed0b04fc1417 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -6,6 +6,7 @@ Authors: Joachim Breitner import group_theory.order_of_element import data.finset.noncomm_prod import data.fintype.card +import data.nat.gcd.big_operators /-! # Canonical homomorphism from a finite family of monoids diff --git a/src/number_theory/arithmetic_function.lean b/src/number_theory/arithmetic_function.lean index 80539f7af0fc3..8cf8cf809b847 100644 --- a/src/number_theory/arithmetic_function.lean +++ b/src/number_theory/arithmetic_function.lean @@ -6,6 +6,7 @@ Authors: Aaron Anderson import algebra.big_operators.ring import number_theory.divisors import data.nat.squarefree +import data.nat.gcd.big_operators import algebra.invertible import data.nat.factorization.basic diff --git a/src/topology/instances/ereal.lean b/src/topology/instances/ereal.lean index ba32823a51d9b..504077e238b47 100644 --- a/src/topology/instances/ereal.lean +++ b/src/topology/instances/ereal.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ +import data.rat.encodable import data.real.ereal import topology.algebra.order.monotone_continuity import topology.instances.ennreal diff --git a/src/topology/instances/irrational.lean b/src/topology/instances/irrational.lean index d2206a83d768d..1fe1ed9145659 100644 --- a/src/topology/instances/irrational.lean +++ b/src/topology/instances/irrational.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ import data.real.irrational +import data.rat.encodable import topology.metric_space.baire /-! From 64fc7238fb41b1a4f12ff05e3d5edfa360dd768c Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Fri, 7 Oct 2022 20:10:10 +0000 Subject: [PATCH 611/828] feat(algebraic_geometry/morphisms/finite_type): Morphisms locally of finite type. (#16124) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/algebraic_geometry/Scheme.lean | 7 ++ .../morphisms/finite_type.lean | 101 ++++++++++++++++++ .../morphisms/ring_hom_properties.lean | 32 +++++- src/ring_theory/ring_hom/finite_type.lean | 5 +- 4 files changed, 143 insertions(+), 2 deletions(-) create mode 100644 src/algebraic_geometry/morphisms/finite_type.lean diff --git a/src/algebraic_geometry/Scheme.lean b/src/algebraic_geometry/Scheme.lean index d3ba561ccede5..7ebac7a659aff 100644 --- a/src/algebraic_geometry/Scheme.lean +++ b/src/algebraic_geometry/Scheme.lean @@ -119,6 +119,13 @@ begin refl end +/-- Given a morphism of schemes `f : X ⟶ Y`, and open sets `U ⊆ Y`, `V ⊆ f ⁻¹' U`, +this is the induced map `Γ(Y, U) ⟶ Γ(X, V)`. -/ +abbreviation hom.app_le {X Y : Scheme} + (f : X ⟶ Y) {V : opens X.carrier} {U : opens Y.carrier} (e : V ≤ (opens.map f.1.base).obj U) : + Y.presheaf.obj (op U) ⟶ X.presheaf.obj (op V) := +f.1.c.app (op U) ≫ X.presheaf.map (hom_of_le e).op + /-- The spectrum of a commutative ring, as a scheme. -/ diff --git a/src/algebraic_geometry/morphisms/finite_type.lean b/src/algebraic_geometry/morphisms/finite_type.lean new file mode 100644 index 0000000000000..472b69021a5b6 --- /dev/null +++ b/src/algebraic_geometry/morphisms/finite_type.lean @@ -0,0 +1,101 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import algebraic_geometry.morphisms.ring_hom_properties +import algebraic_geometry.morphisms.quasi_compact +import ring_theory.ring_hom.finite_type + +/-! +# Morphisms of finite type + +A morphism of schemes `f : X ⟶ Y` is locally of finite type if for each affine `U ⊆ Y` and +`V ⊆ f ⁻¹' U`, The induced map `Γ(Y, U) ⟶ Γ(X, V)` is of finite type. + +A morphism of schemes is of finite type if it is both locally of finite type and quasi-compact. + +We show that these properties are local, and are stable under compositions. + +-/ + +noncomputable theory + +open category_theory category_theory.limits opposite topological_space + +universes v u + +namespace algebraic_geometry + +variables {X Y : Scheme.{u}} (f : X ⟶ Y) + +/-- +A morphism of schemes `f : X ⟶ Y` is locally of finite type if for each affine `U ⊆ Y` and +`V ⊆ f ⁻¹' U`, The induced map `Γ(Y, U) ⟶ Γ(X, V)` is of finite type. +-/ +@[mk_iff] +class locally_of_finite_type (f : X ⟶ Y) : Prop := +(finite_type_of_affine_subset : + ∀ (U : Y.affine_opens) (V : X.affine_opens) (e : V.1 ≤ (opens.map f.1.base).obj U.1), + (f.app_le e).finite_type) + +lemma locally_of_finite_type_eq : + @locally_of_finite_type = affine_locally @ring_hom.finite_type := +begin + ext X Y f, + rw [locally_of_finite_type_iff, affine_locally_iff_affine_opens_le], + exact ring_hom.finite_type_respects_iso +end + +@[priority 900] +instance locally_of_finite_type_of_is_open_immersion {X Y : Scheme} (f : X ⟶ Y) + [is_open_immersion f] : locally_of_finite_type f := +locally_of_finite_type_eq.symm ▸ ring_hom.finite_type_is_local.affine_locally_of_is_open_immersion f + +lemma locally_of_finite_type_stable_under_composition : + morphism_property.stable_under_composition @locally_of_finite_type := +locally_of_finite_type_eq.symm ▸ +ring_hom.finite_type_is_local.affine_locally_stable_under_composition + +instance locally_of_finite_type_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) + [hf : locally_of_finite_type f] [hg : locally_of_finite_type g] : + locally_of_finite_type (f ≫ g) := +locally_of_finite_type_stable_under_composition f g hf hg + +lemma locally_of_finite_type_of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) + [hf : locally_of_finite_type (f ≫ g)] : + locally_of_finite_type f := +begin + unfreezingI { revert hf }, + rw [locally_of_finite_type_eq], + apply ring_hom.finite_type_is_local.affine_locally_of_comp, + introv H, + exactI ring_hom.finite_type.of_comp_finite_type H, +end + +lemma locally_of_finite_type.affine_open_cover_iff {X Y : Scheme.{u}} (f : X ⟶ Y) + (𝒰 : Scheme.open_cover.{u} Y) [∀ i, is_affine (𝒰.obj i)] + (𝒰' : ∀ i, Scheme.open_cover.{u} ((𝒰.pullback_cover f).obj i)) + [∀ i j, is_affine ((𝒰' i).obj j)] : + locally_of_finite_type f ↔ + (∀ i j, (Scheme.Γ.map ((𝒰' i).map j ≫ pullback.snd).op).finite_type) := +locally_of_finite_type_eq.symm ▸ ring_hom.finite_type_is_local.affine_open_cover_iff f 𝒰 𝒰' + +lemma locally_of_finite_type.source_open_cover_iff {X Y : Scheme.{u}} (f : X ⟶ Y) + (𝒰 : Scheme.open_cover.{u} X) : + locally_of_finite_type f ↔ (∀ i, locally_of_finite_type (𝒰.map i ≫ f)) := +locally_of_finite_type_eq.symm ▸ ring_hom.finite_type_is_local.source_open_cover_iff f 𝒰 + +lemma locally_of_finite_type.open_cover_iff {X Y : Scheme.{u}} (f : X ⟶ Y) + (𝒰 : Scheme.open_cover.{u} Y) : + locally_of_finite_type f ↔ + (∀ i, locally_of_finite_type (pullback.snd : pullback f (𝒰.map i) ⟶ _)) := +locally_of_finite_type_eq.symm ▸ + ring_hom.finite_type_is_local.is_local_affine_locally.open_cover_iff f 𝒰 + +lemma locally_of_finite_type_respects_iso : + morphism_property.respects_iso @locally_of_finite_type := +locally_of_finite_type_eq.symm ▸ target_affine_locally_respects_iso + (source_affine_locally_respects_iso ring_hom.finite_type_respects_iso) + +end algebraic_geometry diff --git a/src/algebraic_geometry/morphisms/ring_hom_properties.lean b/src/algebraic_geometry/morphisms/ring_hom_properties.lean index 89deac727eda6..154ec3f3d4527 100644 --- a/src/algebraic_geometry/morphisms/ring_hom_properties.lean +++ b/src/algebraic_geometry/morphisms/ring_hom_properties.lean @@ -175,7 +175,7 @@ lemma affine_locally_iff_affine_opens_le (hP : ring_hom.respects_iso @P) {X Y : Scheme} (f : X ⟶ Y) : affine_locally @P f ↔ (∀ (U : Y.affine_opens) (V : X.affine_opens) (e : V.1 ≤ (opens.map f.1.base).obj U.1), - P (f.1.c.app (op U) ≫ X.presheaf.map (hom_of_le e).op)) := + P (f.app_le e)) := begin apply forall_congr, intro U, @@ -491,6 +491,36 @@ begin { intro i, exact H } end +lemma affine_locally_of_comp + (H : ∀ {R S T : Type.{u}} [comm_ring R] [comm_ring S] [comm_ring T], by exactI + ∀ (f : R →+* S) (g : S →+* T), P (g.comp f) → P g) + {X Y Z : Scheme} {f : X ⟶ Y} {g : Y ⟶ Z} (h : affine_locally @P (f ≫ g)) : + affine_locally @P f := +begin + let 𝒰 : ∀ i, ((Z.affine_cover.pullback_cover (f ≫ g)).obj i).open_cover, + { intro i, + refine Scheme.open_cover.bind _ (λ i, Scheme.affine_cover _), + apply Scheme.open_cover.pushforward_iso _ + (pullback_right_pullback_fst_iso g (Z.affine_cover.map i) f).hom, + apply Scheme.pullback.open_cover_of_right, + exact (pullback g (Z.affine_cover.map i)).affine_cover }, + haveI h𝒰 : ∀ i j, is_affine ((𝒰 i).obj j), by { dsimp, apply_instance }, + let 𝒰' := (Z.affine_cover.pullback_cover g).bind (λ i, Scheme.affine_cover _), + haveI h𝒰' : ∀ i, is_affine (𝒰'.obj i), by { dsimp, apply_instance }, + rw hP.affine_open_cover_iff f 𝒰' (λ i, Scheme.affine_cover _), + rw hP.affine_open_cover_iff (f ≫ g) Z.affine_cover 𝒰 at h, + rintros ⟨i, j⟩ k, + dsimp at i j k, + specialize h i ⟨j, k⟩, + dsimp only [Scheme.open_cover.bind_map, Scheme.open_cover.pushforward_iso_obj, + Scheme.pullback.open_cover_of_right_obj, Scheme.open_cover.pushforward_iso_map, + Scheme.pullback.open_cover_of_right_map, Scheme.open_cover.bind_obj, + Scheme.open_cover.pullback_cover_obj, Scheme.open_cover.pullback_cover_map] at h ⊢, + rw [category.assoc, category.assoc, pullback_right_pullback_fst_iso_hom_snd, + pullback.lift_snd_assoc, category.assoc, ← category.assoc, op_comp, functor.map_comp] at h, + exact H _ _ h, +end + lemma affine_locally_stable_under_composition : (affine_locally @P).stable_under_composition := begin diff --git a/src/ring_theory/ring_hom/finite_type.lean b/src/ring_theory/ring_hom/finite_type.lean index ecce446a133de..6f6f5ab0c978a 100644 --- a/src/ring_theory/ring_hom/finite_type.lean +++ b/src/ring_theory/ring_hom/finite_type.lean @@ -84,9 +84,12 @@ begin { rw ht, trivial } } end -lemma finite_is_local : +lemma finite_type_is_local : property_is_local @finite_type := ⟨localization_finite_type, finite_type_of_localization_span_target, finite_type_stable_under_composition, finite_type_holds_for_localization_away⟩ +lemma finite_type_respects_iso : ring_hom.respects_iso @ring_hom.finite_type := +ring_hom.finite_type_is_local.respects_iso + end ring_hom From 84066cbfea48e8dfc9bb12e8be078054229626da Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Fri, 7 Oct 2022 21:24:35 +0000 Subject: [PATCH 612/828] feat(logic/equiv, topology/homeomorph): split a product at a coordinate (#16210) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit + Add `equiv.pi_split_at` and `homeomorph.pi_split_at`: for every "coordinate" `i : α`, the bijection/homeomorphism between a product `Π j : α, β j` of types/topological spaces and the binary product of the type/space `β i` at that coordinate and the product `Π j : {j // j ≠ i}, β j` over all other coordinates. + Specialize to get the non-dependent versions, which concerns a product of copies of the same type/topological space, in which case non-dependent versions are friendlier to unification. Moreover, separate definitions allow the use of `@[simps]` to generate lemmas in which type casts are automatically removed. Prerequisite of #15681, where we deal with cubes, which are products of copies of the unit interval, for the purpose of showing commutativity of homotopy groups. --- src/logic/equiv/basic.lean | 14 ++++++++++++++ src/topology/homeomorph.lean | 32 ++++++++++++++++++++++++++++++++ 2 files changed, 46 insertions(+) diff --git a/src/logic/equiv/basic.lean b/src/logic/equiv/basic.lean index 7cc44e6d615ff..207261d9339ea 100644 --- a/src/logic/equiv/basic.lean +++ b/src/logic/equiv/basic.lean @@ -1600,6 +1600,20 @@ depending on whether they satisfy a predicate `p` or not. -/ refl }, end } +/-- A product of types can be split as the binary product of one of the types and the product + of all the remaining types. -/ +@[simps] def pi_split_at {α : Type*} [decidable_eq α] (i : α) (β : α → Type*) : + (Π j, β j) ≃ β i × Π j : {j // j ≠ i}, β j := +{ to_fun := λ f, ⟨f i, λ j, f j⟩, + inv_fun := λ f j, if h : j = i then h.symm.rec f.1 else f.2 ⟨j, h⟩, + right_inv := λ f, by { ext, exacts [dif_pos rfl, (dif_neg x.2).trans (by cases x; refl)] }, + left_inv := λ f, by { ext, dsimp only, split_ifs, { subst h }, { refl } } } + +/-- A product of copies of a type can be split as the binary product of one copy and the product + of all the remaining copies. -/ +@[simps] def fun_split_at {α : Type*} [decidable_eq α] (i : α) (β : Type*) : + (α → β) ≃ β × ({j // j ≠ i} → β) := pi_split_at i _ + end section subtype_equiv_codomain diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index dbaf081171be3..06860a720f17f 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -470,6 +470,38 @@ def set.univ (α : Type*) [topological_space α] : (univ : set α) ≃ₜ α := continuous_inv_fun := (continuous_subtype_coe.fst'.prod_mk continuous_subtype_coe.snd').subtype_mk _ } +section + +variable {ι : Type*} + +/-- The topological space `Π i, β i` can be split as a product by separating the indices in ι + depending on whether they satisfy a predicate p or not.-/ +@[simps] def pi_equiv_pi_subtype_prod (p : ι → Prop) (β : ι → Type*) [Π i, topological_space (β i)] + [decidable_pred p] : (Π i, β i) ≃ₜ (Π i : {x // p x}, β i) × Π i : {x // ¬p x}, β i := +{ to_equiv := equiv.pi_equiv_pi_subtype_prod p β, + continuous_to_fun := by apply continuous.prod_mk; exact continuous_pi (λ j, continuous_apply j), + continuous_inv_fun := continuous_pi $ λ j, begin + dsimp only [equiv.pi_equiv_pi_subtype_prod], split_ifs, + exacts [(continuous_apply _).comp continuous_fst, (continuous_apply _).comp continuous_snd], + end } + +variables [decidable_eq ι] (i : ι) + +/-- A product of topological spaces can be split as the binary product of one of the spaces and + the product of all the remaining spaces. -/ +@[simps] def pi_split_at (β : ι → Type*) [Π j, topological_space (β j)] : + (Π j, β j) ≃ₜ β i × Π j : {j // j ≠ i}, β j := +{ to_equiv := equiv.pi_split_at i β, + continuous_to_fun := (continuous_apply i).prod_mk (continuous_pi $ λ j, continuous_apply j), + continuous_inv_fun := continuous_pi $ λ j, by { dsimp only [equiv.pi_split_at], + split_ifs, subst h, exacts [continuous_fst, (continuous_apply _).comp continuous_snd] } } + +/-- A product of copies of a topological space can be split as the binary product of one copy and + the product of all the remaining copies. -/ +@[simps] def fun_split_at : (ι → β) ≃ₜ β × ({j // j ≠ i} → β) := pi_split_at i _ + +end + end homeomorph /-- An inducing equiv between topological spaces is a homeomorphism. -/ From 54f74fc806aab56cdabf8f964643d80bc31b72af Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 8 Oct 2022 00:38:14 +0000 Subject: [PATCH 613/828] =?UTF-8?q?feat(topology/continuous=5Ffunction/ide?= =?UTF-8?q?als):=20construct=20the=20Galois=20correspondence=20between=20c?= =?UTF-8?q?losed=20ideals=20in=20`C(X,=20=F0=9D=95=9C)`=20and=20open=20set?= =?UTF-8?q?s=20in=20`X`=20(#16677)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For a topological ring `R` and a topological space `X` there is a Galois connection between `ideal C(X, R)` and `set X` given by sending each `I : ideal C(X, R)` to `{x : X | ∀ f ∈ I, f x = 0}ᶜ` and mapping `s : set X` to the ideal with carrier `{f : C(X, R) | ∀ x ∈ sᶜ, f x = 0}`, and we call these maps `continuous_map.set_of_ideal` and `continuous_map.ideal_of_set`. As long as `R` is Hausdorff, `continuous_map.set_of_ideal I` is open, and if, in addition, `X` is locally compact, then `continuous_map.set_of_ideal s` is closed. --- src/topology/continuous_function/ideals.lean | 303 +++++++++++++++++++ 1 file changed, 303 insertions(+) create mode 100644 src/topology/continuous_function/ideals.lean diff --git a/src/topology/continuous_function/ideals.lean b/src/topology/continuous_function/ideals.lean new file mode 100644 index 0000000000000..ed779750f0aad --- /dev/null +++ b/src/topology/continuous_function/ideals.lean @@ -0,0 +1,303 @@ +/- +Copyright (c) 2022 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ + +import topology.algebra.algebra +import topology.continuous_function.compact +import topology.urysohns_lemma +import data.complex.is_R_or_C + +/-! +# Ideals of continuous functions + +For a topological ring `R` and a topological space `X` there is a Galois connection between +`ideal C(X, R)` and `set X` given by sending each `I : ideal C(X, R)` to +`{x : X | ∀ f ∈ I, f x = 0}ᶜ` and mapping `s : set X` to the ideal with carrier +`{f : C(X, R) | ∀ x ∈ sᶜ, f x = 0}`, and we call these maps `continuous_map.set_of_ideal` and +`continuous_map.ideal_of_set`. As long as `R` is Hausdorff, `continuous_map.set_of_ideal I` is open, +and if, in addition, `X` is locally compact, then `continuous_map.set_of_ideal s` is closed. + +When `R = 𝕜` with `is_R_or_C 𝕜` and `X` is compact Hausdorff, then this Galois connection can be +improved to a true Galois correspondence (i.e., order isomorphism) between the type `opens X` and +the subtype of closed ideals of `C(X, R)`. + +## Main definitions + +* `continuous_map.ideal_of_set`: ideal of functions which vanish on the complement of a set. +* `continuous_map.set_of_ideal`: complement of the set on which all functions in the ideal vanish. +* `continuous_map.opens_of_ideal`: `continuous_map.set_of_ideal` as a term of `opens X`. +* `continuous_map.ideal_opens_gi`: The Galois insertion `continuous_map.opens_of_ideal` and + `λ s, continuous_map.ideal_of_set ↑s`. + +## Main statements + +* `ideal_of_set_of_ideal_eq_closure`: when `X` is compact Hausdorff and `is_R_or_C 𝕜`, + `ideal_of_set 𝕜 (set_of_ideal I) = I.closure` for any ideal `I : ideal C(X, 𝕜)`. +* `set_of_ideal_of_set_eq_interior`: when `X` is compact Hausdorff and `is_R_or_C 𝕜`, + `set_of_ideal (ideal_of_set 𝕜 s) = interior s` for any `s : set X`. + +## Implementation details + +Because there does not currently exist a bundled type of closed ideals, we don't provide the actual +order isomorphism described above, and instead we only consider the Galois insertion +`continuous_map.ideal_opens_gi`. + +## TODO + +* Show that maximal ideals in `C(X, 𝕜)` correspond to (complements of) singletons. + +## Tags + +ideal, continuous function, compact, Hausdorff +-/ + + +open_locale nnreal + +namespace continuous_map + +open topological_space + +section topological_ring + +variables {X R : Type*} [topological_space X] [ring R] [topological_space R] [topological_ring R] + +variable (R) + +/-- Given a topological ring `R` and `s : set X`, construct the ideal in `C(X, R)` of functions +which vanish on the complement of `s`. -/ +def ideal_of_set (s : set X) : ideal C(X, R) := +{ carrier := {f : C(X, R) | ∀ x ∈ sᶜ, f x = 0}, + add_mem' := λ f g hf hg x hx, by simp only [hf x hx, hg x hx, coe_add, pi.add_apply, add_zero], + zero_mem' := λ _ _, rfl, + smul_mem' := λ c f hf x hx, mul_zero (c x) ▸ congr_arg (λ y, c x * y) (hf x hx), } + +lemma ideal_of_set_closed [locally_compact_space X] [t2_space R] (s : set X) : + is_closed (ideal_of_set R s : set C(X, R) ) := +begin + simp only [ideal_of_set, submodule.coe_set_mk, set.set_of_forall], + exact is_closed_Inter (λ x, is_closed_Inter $ + λ hx, is_closed_eq (continuous_eval_const' x) continuous_const), +end + +variable {R} + +lemma mem_ideal_of_set {s : set X} {f : C(X, R)} : + f ∈ ideal_of_set R s ↔ ∀ ⦃x : X⦄, x ∈ sᶜ → f x = 0 := iff.rfl + +lemma not_mem_ideal_of_set {s : set X} {f : C(X, R)} : + f ∉ ideal_of_set R s ↔ ∃ x ∈ sᶜ, f x ≠ 0 := +by { simp_rw [mem_ideal_of_set, exists_prop], push_neg } + +/-- Given an ideal `I` of `C(X, R)`, construct the set of points for which every function in the +ideal vanishes on the complement. -/ +def set_of_ideal (I : ideal C(X, R)) : set X := +{x : X | ∀ f ∈ I, (f : C(X, R)) x = 0}ᶜ + +lemma not_mem_set_of_ideal {I : ideal C(X, R)} {x : X} : + x ∉ set_of_ideal I ↔ ∀ ⦃f : C(X, R)⦄, f ∈ I → f x = 0 := +by rw [←set.mem_compl_iff, set_of_ideal, compl_compl, set.mem_set_of] + +lemma mem_set_of_ideal {I : ideal C(X, R)} {x : X} : + x ∈ set_of_ideal I ↔ ∃ f ∈ I, (f : C(X, R)) x ≠ 0 := +by { simp_rw [set_of_ideal, set.mem_compl_iff, set.mem_set_of, exists_prop], push_neg } + +lemma set_of_ideal_open [t2_space R] (I : ideal C(X, R)) : is_open (set_of_ideal I) := +begin + simp only [set_of_ideal, set.set_of_forall, is_open_compl_iff], + exact is_closed_Inter (λ f, is_closed_Inter $ + λ hf, is_closed_eq (map_continuous f) continuous_const) +end + +/-- The open set `set_of_ideal I` realized as a term of `opens X`. -/ +@[simps] def opens_of_ideal [t2_space R] (I : ideal C(X, R)) : opens X := +⟨set_of_ideal I, set_of_ideal_open I⟩ + +@[simp] lemma set_of_top_eq_univ [nontrivial R] : (set_of_ideal (⊤ : ideal C(X, R))) = set.univ := +set.univ_subset_iff.mp $ λ x hx, mem_set_of_ideal.mpr ⟨1, submodule.mem_top, one_ne_zero⟩ + +@[simp] lemma ideal_of_empty_eq_bot : (ideal_of_set R (∅ : set X)) = ⊥ := +ideal.ext (λ f, by simpa only [mem_ideal_of_set, set.compl_empty, set.mem_univ, forall_true_left, + ideal.mem_bot, fun_like.ext_iff] using iff.rfl) + +variables (X R) +lemma ideal_gc : galois_connection (set_of_ideal : ideal C(X, R) → set X) (ideal_of_set R) := +begin + refine λ I s, ⟨λ h f hf, _, λ h x hx, _⟩, + { by_contra h', + rcases not_mem_ideal_of_set.mp h' with ⟨x, hx, hfx⟩, + exact hfx (not_mem_set_of_ideal.mp (mt (@h x) hx) hf) }, + { obtain ⟨f, hf, hfx⟩ := mem_set_of_ideal.mp hx, + by_contra hx', + exact not_mem_ideal_of_set.mpr ⟨x, hx', hfx⟩ (h hf) }, +end + +end topological_ring + +section is_R_or_C +open is_R_or_C + +variables {X 𝕜 : Type*} [is_R_or_C 𝕜] [topological_space X] + +/-- An auxiliary lemma used in the proof of `ideal_of_set_of_ideal_eq_closure` which may be useful +on its own. -/ +lemma exists_mul_le_one_eq_on_ge (f : C(X, ℝ≥0)) {c : ℝ≥0} (hc : 0 < c) : + ∃ g : C(X, ℝ≥0), (∀ x : X, (g * f) x ≤ 1) ∧ {x : X | c ≤ f x}.eq_on (g * f) 1 := +⟨{ to_fun := (f ⊔ (const X c))⁻¹, + continuous_to_fun := ((map_continuous f).sup $ map_continuous _).inv₀ + (λ _, (hc.trans_le le_sup_right).ne')}, + λ x, (inv_mul_le_iff (hc.trans_le le_sup_right)).mpr ((mul_one (f x ⊔ c)).symm ▸ le_sup_left), + λ x hx, by simpa only [coe_const, coe_mk, pi.mul_apply, pi.inv_apply, pi.sup_apply, + function.const_apply, pi.one_apply, sup_eq_left.mpr (set.mem_set_of.mp hx)] + using inv_mul_cancel (hc.trans_le hx).ne'⟩ + +@[simp] lemma ideal_of_set_of_ideal_eq_closure [compact_space X] [t2_space X] (I : ideal C(X, 𝕜)) : + ideal_of_set 𝕜 (set_of_ideal I) = I.closure := +begin + /- Since `ideal_of_set 𝕜 (set_of_ideal I)` is closed and contains `I`, it contains `I.closure`. + For the reverse inclusion, given `f ∈ ideal_of_set 𝕜 (set_of_ideal I)` and `(ε : ℝ≥0) > 0` it + suffices to show that `f` is within `ε` of `I`.-/ + refine le_antisymm (λ f hf, metric.mem_closure_iff.mpr (λ ε hε, _)) + ((ideal_of_set_closed 𝕜 $ set_of_ideal I).closure_subset_iff.mpr $ + λ f hf x hx, not_mem_set_of_ideal.mp hx hf), + lift ε to ℝ≥0 using hε.lt.le, + replace hε := (show (0 : ℝ≥0) < ε, from hε), + simp_rw dist_nndist, + norm_cast, + -- Let `t := {x : X | ε / 2 ≤ ∥f x∥₊}}` which is closed and disjoint from `set_of_ideal I`. + set t := {x : X | ε / 2 ≤ ∥f x∥₊}, + have ht : is_closed t := is_closed_le continuous_const (map_continuous f).nnnorm, + have htI : disjoint t (set_of_ideal I)ᶜ, + { refine set.subset_compl_iff_disjoint_left.mp (λ x hx, _), + simpa only [t, set.mem_set_of, set.mem_compl_iff, not_le] + using (nnnorm_eq_zero.mpr (mem_ideal_of_set.mp hf hx)).trans_lt (half_pos hε), }, + /- It suffices to produce `g : C(X, ℝ≥0)` which takes values in `[0,1]` and is constantly `1` on + `t` such that when composed with the natural embedding of `ℝ≥0` into `𝕜` lies in the ideal `I`. + Indeed, then `∥f - f * ↑g∥ ≤ ∥f * (1 - ↑g)∥ ≤ ⨆ ∥f * (1 - ↑g) x∥`. When `x ∉ t`, `∥f x∥ < ε / 2` + and `∥(1 - ↑g) x∥ ≤ 1`, and when `x ∈ t`, `(1 - ↑g) x = 0`, and clearly `f * ↑g ∈ I`. -/ + suffices : ∃ g : C(X, ℝ≥0), + (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g ∈ I ∧ (∀ x, g x ≤ 1) ∧ t.eq_on g 1, + { obtain ⟨g, hgI, hg, hgt⟩ := this, + refine ⟨f * (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g, I.mul_mem_left f hgI, _⟩, + rw nndist_eq_nnnorm, + refine (nnnorm_lt_iff _ hε).2 (λ x, _), + simp only [coe_sub, coe_mul, pi.sub_apply, pi.mul_apply], + by_cases hx : x ∈ t, + { simpa only [hgt hx, comp_apply, pi.one_apply, continuous_map.coe_coe, algebra_map_clm_apply, + map_one, mul_one, sub_self, nnnorm_zero] using hε, }, + { refine lt_of_le_of_lt _ (half_lt_self hε), + have := calc ∥((1 - (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) x : 𝕜)∥₊ + = ∥1 - algebra_map ℝ≥0 𝕜 (g x)∥₊ + : by simp only [coe_sub, coe_one, coe_comp, continuous_map.coe_coe, pi.sub_apply, + pi.one_apply, function.comp_app, algebra_map_clm_apply] + ... = ∥algebra_map ℝ≥0 𝕜 (1 - g x)∥₊ + : by simp only [algebra.algebra_map_eq_smul_one, nnreal.smul_def, nnreal.coe_sub (hg x), + sub_smul, nonneg.coe_one, one_smul] + ... ≤ 1 : (nnnorm_algebra_map_nnreal 𝕜 (1 - g x)).trans_le tsub_le_self, + calc ∥f x - f x * (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g x∥₊ + = ∥f x * (1 - (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) x∥₊ + : by simp only [mul_sub, coe_sub, coe_one, pi.sub_apply, pi.one_apply, mul_one] + ... ≤ (ε / 2) * ∥(1 - (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) x∥₊ + : (nnnorm_mul_le _ _).trans (mul_le_mul_right' + (not_le.mp $ show ¬ ε / 2 ≤ ∥f x∥₊, from hx).le _) + ... ≤ ε / 2 : by simpa only [mul_one] using mul_le_mul_left' this _, } }, + /- There is some `g' : C(X, ℝ≥0)` which is strictly positive on `t` such that the composition + `↑g` with the natural embedding of `ℝ≥0` into `𝕜` lies in `I`. This follows from compactness of + `t` and that we can do it in any neighborhood of a point `x ∈ t`. Indeed, since `x ∈ t`, then + `fₓ x ≠ 0` for some `fₓ ∈ I` and so `λ y, ∥(star fₓ * fₓ) y∥₊` is strictly posiive in a + neighborhood of `y`. Moreover, `(∥(star fₓ * fₓ) y∥₊ : 𝕜) = (star fₓ * fₓ) y`, so composition of + this map with the natural embedding is just `star fₓ * fₓ ∈ I`. -/ + have : ∃ g' : C(X, ℝ≥0), (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g' ∈ I ∧ (∀ x ∈ t, 0 < g' x), + { refine @is_compact.induction_on _ _ _ ht.is_compact (λ s, ∃ g' : C(X, ℝ≥0), + (algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g' ∈ I ∧ (∀ x ∈ s, 0 < g' x)) _ _ _ _, + { refine ⟨0, _, λ x hx, false.elim hx⟩, + convert I.zero_mem, + ext, + simp only [coe_zero, pi.zero_apply, continuous_map.coe_coe, continuous_map.coe_comp, + map_zero, pi.comp_zero] }, + { rintro s₁ s₂ hs ⟨g, hI, hgt⟩, exact ⟨g, hI, λ x hx, hgt x (hs hx)⟩, }, + { rintro s₁ s₂ ⟨g₁, hI₁, hgt₁⟩ ⟨g₂, hI₂, hgt₂⟩, + refine ⟨g₁ + g₂, _, λ x hx, _⟩, + { convert I.add_mem hI₁ hI₂, + ext y, + simp only [coe_add, pi.add_apply, map_add, coe_comp, function.comp_app, + continuous_map.coe_coe]}, + { rcases hx with (hx | hx), + simpa only [zero_add] using add_lt_add_of_lt_of_le (hgt₁ x hx) zero_le', + simpa only [zero_add] using add_lt_add_of_le_of_lt zero_le' (hgt₂ x hx), } }, + { intros x hx, + replace hx := htI.subset_compl_right hx, + rw [compl_compl, mem_set_of_ideal] at hx, + obtain ⟨g, hI, hgx⟩ := hx, + have := (map_continuous g).continuous_at.eventually_ne hgx, + refine ⟨{y : X | g y ≠ 0} ∩ t, mem_nhds_within_iff_exists_mem_nhds_inter.mpr + ⟨_, this, set.subset.rfl⟩, ⟨⟨λ x, ∥g x∥₊ ^ 2, (map_continuous g).nnnorm.pow 2⟩, _, + λ x hx, pow_pos (norm_pos_iff.mpr hx.1) 2⟩⟩, + convert I.mul_mem_left (star g) hI, + ext, + simp only [comp_apply, coe_mk, algebra_map_clm_coe, map_pow, coe_mul, coe_star, + pi.mul_apply, pi.star_apply, star_def, continuous_map.coe_coe], + simpa only [norm_sq_eq_def', conj_mul_eq_norm_sq_left, of_real_pow], }, }, + /- Get the function `g'` which is guaranteed to exist above. By the extreme value theorem and + compactness of `t`, there is some `0 < c` such that `c ≤ g' x` for all `x ∈ t`. Then by + `main_lemma_aux` there is some `g` for which `g * g'` is the desired function. -/ + obtain ⟨g', hI', hgt'⟩ := this, + obtain (⟨c, hc, hgc'⟩ : ∃ c (hc : 0 < c), ∀ y : X, y ∈ t → c ≤ g' y) := + t.eq_empty_or_nonempty.elim (λ ht', ⟨1, zero_lt_one, λ y hy, false.elim (by rwa ht' at hy)⟩) + (λ ht', let ⟨x, hx, hx'⟩ := ht.is_compact.exists_forall_le ht' (map_continuous g').continuous_on + in ⟨g' x, hgt' x hx, hx'⟩), + obtain ⟨g, hg, hgc⟩ := exists_mul_le_one_eq_on_ge g' hc, + refine ⟨g * g', _, hg, hgc.mono hgc'⟩, + convert I.mul_mem_left ((algebra_map_clm ℝ≥0 𝕜 : C(ℝ≥0, 𝕜)).comp g) hI', + ext, + simp only [algebra_map_clm_coe, continuous_map.coe_coe, comp_apply, coe_mul, pi.mul_apply, + map_mul], +end + +lemma ideal_of_set_of_ideal_is_closed [compact_space X] [t2_space X] {I : ideal C(X, 𝕜)} + (hI : is_closed (I : set C(X, 𝕜))) : ideal_of_set 𝕜 (set_of_ideal I) = I := +(ideal_of_set_of_ideal_eq_closure I).trans (ideal.ext $ set.ext_iff.mp hI.closure_eq) + +variable (𝕜) + +@[simp] lemma set_of_ideal_of_set_eq_interior [compact_space X] [t2_space X] (s : set X) : + set_of_ideal (ideal_of_set 𝕜 s) = interior s:= +begin + refine set.subset.antisymm ((set_of_ideal_open (ideal_of_set 𝕜 s)).subset_interior_iff.mpr + (λ x hx, let ⟨f, hf, hfx⟩ := mem_set_of_ideal.mp hx + in set.not_mem_compl_iff.mp (mt (@hf x) hfx))) (λ x hx, _), + /- If `x ∉ closure sᶜ`, we must produce `f : C(X, 𝕜)` which is zero on `sᶜ` and `f x ≠ 0`. -/ + rw [←compl_compl (interior s), ←closure_compl] at hx, + simp_rw [mem_set_of_ideal, mem_ideal_of_set], + haveI : normal_space X := normal_of_compact_t2, + /- Apply Urysohn's lemma to get `g : C(X, ℝ)` which is zero on `sᶜ` and `g x ≠ 0`, then compose + with the natural embedding `ℝ ↪ 𝕜` to produce the desired `f`. -/ + obtain ⟨g, hgs, (hgx : set.eq_on g 1 {x}), -⟩ := exists_continuous_zero_one_of_closed + is_closed_closure is_closed_singleton (set.disjoint_singleton_right.mpr hx), + exact ⟨⟨λ x, g x, continuous_of_real.comp (map_continuous g)⟩, + by simpa only [coe_mk, of_real_eq_zero] using λ x hx, hgs (subset_closure hx), + by simpa only [coe_mk, hgx (set.mem_singleton x), pi.one_apply, is_R_or_C.of_real_one] + using one_ne_zero⟩, +end + +lemma set_of_ideal_of_set_of_is_open [compact_space X] [t2_space X] {s : set X} + (hs : is_open s) : set_of_ideal (ideal_of_set 𝕜 s) = s := +(set_of_ideal_of_set_eq_interior 𝕜 s).trans hs.interior_eq + +variable (X) + +/-- The Galois insertion `continuous_map.opens_of_ideal : ideal C(X, 𝕜) → opens X` and +`λ s, continuous_map.ideal_of_set ↑s`. -/ +@[simps] def ideal_opens_gi [compact_space X] [t2_space X] : + galois_insertion (opens_of_ideal : ideal C(X, 𝕜) → opens X) (λ s, ideal_of_set 𝕜 s) := +{ choice := λ I hI, opens_of_ideal I.closure, + gc := λ I s, ideal_gc X 𝕜 I s, + le_l_u := λ s, (set_of_ideal_of_set_of_is_open 𝕜 s.prop).ge, + choice_eq := λ I hI, congr_arg _ $ ideal.ext (set.ext_iff.mp (is_closed_of_closure_subset $ + (ideal_of_set_of_ideal_eq_closure I ▸ hI : I.closure ≤ I)).closure_eq) } + +end is_R_or_C + +end continuous_map From 40e377e572f9b12453aa4a82de98900fbb713130 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Sat, 8 Oct 2022 00:38:15 +0000 Subject: [PATCH 614/828] feat(data/list/indexes): add map_with_index_append (#16785) --- src/data/list/indexes.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/data/list/indexes.lean b/src/data/list/indexes.lean index f8aa55881c805..bc47dbbec6dcf 100644 --- a/src/data/list/indexes.lean +++ b/src/data/list/indexes.lean @@ -49,6 +49,14 @@ end by simp [map_with_index_eq_enum_map, enum_eq_zip_range, map_uncurry_zip_eq_zip_with, range_succ_eq_map, zip_with_map_left] +lemma map_with_index_append {α} (K L : list α) (f : ℕ → α → β) : + (K ++ L).map_with_index f = K.map_with_index f ++ L.map_with_index (λ i a, f (i + K.length) a) := +begin + induction K with a J IH generalizing f, + { simp }, + { simp [IH (λ i, f (i+1)), add_assoc], } +end + @[simp] lemma length_map_with_index {α β} (l : list α) (f : ℕ → α → β) : (l.map_with_index f).length = l.length := begin From cf43320e83fda9c6263175c4cbc0af800acaeac1 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sat, 8 Oct 2022 00:38:17 +0000 Subject: [PATCH 615/828] feat(group_theory/transfer): Homomorphism for Burnside's transfer theorem (#16818) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR constructs the homomorphism `G →* P` for Burnside's transfer theorem and proves that it is the `[G : P]`th power map on elements of `P`. Burnside's transfer theorem will follow reasonably quickly from this (the homomorphism is an isomorphism on `P`, so the kernel is disjoint from `P`, so the kernel is disjoint from every Sylow by normality, so the kernel cannot have order divisible by p, so the homomorphism is surjective and the kernel is a normal p-complement). --- src/group_theory/transfer.lean | 32 ++++++++++++++++++++++++++++++-- 1 file changed, 30 insertions(+), 2 deletions(-) diff --git a/src/group_theory/transfer.lean b/src/group_theory/transfer.lean index 431458421f488..3d8f0070003ea 100644 --- a/src/group_theory/transfer.lean +++ b/src/group_theory/transfer.lean @@ -5,8 +5,7 @@ Authors: Thomas Browning -/ import group_theory.complement -import group_theory.group_action.basic -import group_theory.index +import group_theory.sylow /-! # The Transfer Homomorphism @@ -162,4 +161,33 @@ noncomputable def transfer_center_pow' (h : (center G).index ≠ 0) : G →* cen ↑(transfer_center_pow' h g) = g ^ (center G).index := rfl +section burnside_transfer + +variables {p : ℕ} (P : sylow p G) (hP : (P : subgroup G).normalizer ≤ (P : subgroup G).centralizer) + +include hP + +/-- The homomorphism `G →* P` in Burnside's transfer theorem. -/ +noncomputable def transfer_sylow [fintype (G ⧸ (P : subgroup G))] : G →* (P : subgroup G) := +@transfer G _ P P (@subgroup.is_commutative.comm_group G _ P + ⟨⟨λ a b, subtype.ext (hP (le_normalizer b.2) a a.2)⟩⟩) (monoid_hom.id P) _ + +/-- Auxillary lemma in order to state `transfer_sylow_eq_pow`. -/ +lemma transfer_sylow_eq_pow_aux [fact p.prime] [finite (sylow p G)] (g : G) (hg : g ∈ P) (k : ℕ) + (g₀ : G) (h : g₀⁻¹ * g ^ k * g₀ ∈ P) : g₀⁻¹ * g ^ k * g₀ = g ^ k := +begin + haveI : (P : subgroup G).is_commutative := ⟨⟨λ a b, subtype.ext (hP (le_normalizer b.2) a a.2)⟩⟩, + replace hg := (P : subgroup G).pow_mem hg k, + obtain ⟨n, hn, h⟩ := P.conj_eq_normalizer_conj_of_mem (g ^ k) g₀ hg h, + exact h.trans (commute.inv_mul_cancel (hP hn (g ^ k) hg).symm), +end + +lemma transfer_sylow_eq_pow [fact p.prime] [finite (sylow p G)] [fintype (G ⧸ (P : subgroup G))] + (hP : P.1.normalizer ≤ P.1.centralizer) (g : G) (hg : g ∈ P) : + transfer_sylow P hP g = ⟨g ^ (P : subgroup G).index, + transfer_eq_pow_aux g (transfer_sylow_eq_pow_aux P hP g hg)⟩ := +by apply transfer_eq_pow + +end burnside_transfer + end monoid_hom From 706345b3b4565493a73b5ec0ea28d65d0a2cc49c Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 8 Oct 2022 00:38:18 +0000 Subject: [PATCH 616/828] feat(algebra/free_monoid/count): new file (#16829) * add `algebra.free_monoid.count`; * move `algebra.free_monoid` to `algebra.free_monoid.basic`. --- src/algebra/category/Mon/adjunctions.lean | 2 +- .../basic.lean} | 0 src/algebra/free_monoid/count.lean | 71 +++++++++++++++++++ src/control/fold.lean | 2 +- src/group_theory/free_product.lean | 2 +- src/group_theory/submonoid/membership.lean | 2 +- 6 files changed, 75 insertions(+), 4 deletions(-) rename src/algebra/{free_monoid.lean => free_monoid/basic.lean} (100%) create mode 100644 src/algebra/free_monoid/count.lean diff --git a/src/algebra/category/Mon/adjunctions.lean b/src/algebra/category/Mon/adjunctions.lean index a5444e80fbeee..2b766727588f3 100644 --- a/src/algebra/category/Mon/adjunctions.lean +++ b/src/algebra/category/Mon/adjunctions.lean @@ -6,7 +6,7 @@ Authors: Julian Kuelshammer import algebra.category.Mon.basic import algebra.category.Semigroup.basic import algebra.group.with_one -import algebra.free_monoid +import algebra.free_monoid.basic /-! # Adjunctions regarding the category of monoids diff --git a/src/algebra/free_monoid.lean b/src/algebra/free_monoid/basic.lean similarity index 100% rename from src/algebra/free_monoid.lean rename to src/algebra/free_monoid/basic.lean diff --git a/src/algebra/free_monoid/count.lean b/src/algebra/free_monoid/count.lean new file mode 100644 index 0000000000000..f526f01c6b861 --- /dev/null +++ b/src/algebra/free_monoid/count.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2022 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import algebra.free_monoid.basic +import data.list.count + +/-! +# `list.count` as a bundled homomorphism + +In this file we define `free_monoid.countp`, `free_monoid.count`, `free_add_monoid.countp`, and +`free_add_monoid.count`. These are `list.countp` and `list.count` bundled as multiplicative and +additive homomorphisms from `free_monoid` and `free_add_monoid`. + +We do not use `to_additive` because it can't map `multiplicative ℕ` to `ℕ`. +-/ + +variables {α : Type*} (p : α → Prop) [decidable_pred p] + +namespace free_add_monoid + +/-- `list.countp` as a bundled additive monoid homomorphism. -/ +def countp (p : α → Prop) [decidable_pred p] : free_add_monoid α →+ ℕ := +⟨list.countp p, list.countp_nil p, list.countp_append _⟩ + +lemma countp_of (x : α) : countp p (of x) = if p x then 1 else 0 := rfl + +lemma countp_apply (l : free_add_monoid α) : countp p l = list.countp p l := rfl + +/-- `list.count` as a bundled additive monoid homomorphism. -/ +def count [decidable_eq α] (x : α) : free_add_monoid α →+ ℕ := countp (eq x) + +lemma count_of [decidable_eq α] (x y : α) : count x (of y) = pi.single x 1 y := +by simp only [count, countp_of, pi.single_apply, eq_comm] + +lemma count_apply [decidable_eq α] (x : α) (l : free_add_monoid α) : + count x l = list.count x l := +rfl + +end free_add_monoid + +namespace free_monoid + +/-- `list.countp` as a bundled multiplicative monoid homomorphism. -/ +def countp (p : α → Prop) [decidable_pred p] : free_monoid α →* multiplicative ℕ := +(free_add_monoid.countp p).to_multiplicative + +lemma countp_of' (x : α) : + countp p (of x) = if p x then multiplicative.of_add 1 else multiplicative.of_add 0 := +rfl + +lemma countp_of (x : α) : countp p (of x) = if p x then multiplicative.of_add 1 else 1 := +by rw [countp_of', of_add_zero] -- `rfl` is not transitive + +lemma countp_apply (l : free_add_monoid α) : + countp p l = multiplicative.of_add (list.countp p l) := +rfl + +/-- `list.count` as a bundled additive monoid homomorphism. -/ +def count [decidable_eq α] (x : α) : free_monoid α →* multiplicative ℕ := countp (eq x) + +lemma count_apply [decidable_eq α] (x : α) (l : free_add_monoid α) : + count x l = multiplicative.of_add (list.count x l) := +rfl + +lemma count_of [decidable_eq α] (x y : α) : + count x (of y) = @pi.mul_single α (λ _, multiplicative ℕ) _ _ x (multiplicative.of_add 1) y := +by simp only [count, countp_of, pi.mul_single_apply, eq_comm] + +end free_monoid diff --git a/src/control/fold.lean b/src/control/fold.lean index efb64c4d35c80..2944d754f29cf 100644 --- a/src/control/fold.lean +++ b/src/control/fold.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Sean Leather -/ -import algebra.free_monoid +import algebra.free_monoid.basic import algebra.opposites import control.traversable.instances import control.traversable.lemmas diff --git a/src/group_theory/free_product.lean b/src/group_theory/free_product.lean index 89c1f09895630..f6147c59b5231 100644 --- a/src/group_theory/free_product.lean +++ b/src/group_theory/free_product.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 David Wärn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Wärn, Joachim Breitner -/ -import algebra.free_monoid +import algebra.free_monoid.basic import group_theory.congruence import group_theory.is_free_group import group_theory.subgroup.pointwise diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index fc77d6c08abb9..d70d06f8b1aa5 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -6,7 +6,7 @@ Amelia Livingston, Yury Kudryashov -/ import group_theory.submonoid.operations import algebra.big_operators.basic -import algebra.free_monoid +import algebra.free_monoid.basic import data.finset.noncomm_prod /-! From 827f0b99066e44c65a0abb240f132dec23b474e9 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Sat, 8 Oct 2022 00:38:19 +0000 Subject: [PATCH 617/828] chore(data/pfunctor/univariate): remove use of `h_generalize` (#16839) to help with the port we remove all uses of the tactic `h_generalize`, there is only one in fact --- src/data/pfunctor/univariate/M.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/data/pfunctor/univariate/M.lean b/src/data/pfunctor/univariate/M.lean index 0d7980a33d14c..b2d6f9fceef29 100644 --- a/src/data/pfunctor/univariate/M.lean +++ b/src/data/pfunctor/univariate/M.lean @@ -295,7 +295,9 @@ begin { rw [← head_succ' n,h,head'], apply x.consistent }, revert ch, rw h', intros, congr, { ext a, dsimp only [children], - h_generalize! hh : a == a'', + generalize hh : cast _ a = a'', + rw cast_eq_iff_heq at hh, + revert a'', rw h, intros, cases hh, refl }, end From 90fab282f1c4f8e56a2592b91990fda0a67fcd38 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Sat, 8 Oct 2022 00:38:20 +0000 Subject: [PATCH 618/828] feat(topology/continuous_on): generalise lemma slightly (#16840) Generalise `frequently_nhds_within_iff` to allow `within` any set, not just complement of singleton Co-authored-by: Bhavik Mehta --- src/topology/continuous_on.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/topology/continuous_on.lean b/src/topology/continuous_on.lean index 7ab9203159cd7..e58361997501d 100644 --- a/src/topology/continuous_on.lean +++ b/src/topology/continuous_on.lean @@ -44,9 +44,9 @@ lemma eventually_nhds_within_iff {a : α} {s : set α} {p : α → Prop} : (∀ᶠ x in 𝓝[s] a, p x) ↔ ∀ᶠ x in 𝓝 a, x ∈ s → p x := eventually_inf_principal -lemma frequently_nhds_within_iff {z : α} {p : α → Prop} : - (∃ᶠ x in 𝓝[≠] z, p x) ↔ (∃ᶠ x in 𝓝 z, p x ∧ x ≠ z) := -iff.not (by simp [eventually_nhds_within_iff, not_imp_not]) +lemma frequently_nhds_within_iff {z : α} {s : set α} {p : α → Prop} : + (∃ᶠ x in 𝓝[s] z, p x) ↔ (∃ᶠ x in 𝓝 z, p x ∧ x ∈ s) := +iff.not (by simp [eventually_nhds_within_iff, not_and']) lemma mem_closure_ne_iff_frequently_within {z : α} {s : set α} : z ∈ closure (s \ {z}) ↔ ∃ᶠ x in 𝓝[≠] z, x ∈ s := From 46146be043b2a840a385432efaff4c8291a4488f Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 8 Oct 2022 03:03:10 +0000 Subject: [PATCH 619/828] feat(linear_algebra/matrix/charpoly) Generalize Cayley-Hamilton theorem. (#15105) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- docs/references.bib | 13 + .../matrix/charpoly/linear_map.lean | 245 ++++++++++++++++++ 2 files changed, 258 insertions(+) create mode 100644 src/linear_algebra/matrix/charpoly/linear_map.lean diff --git a/docs/references.bib b/docs/references.bib index 94c9ed1995f68..135bb7eddf5fe 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -609,6 +609,19 @@ @Book{ EinsiedlerWard2017 doi = {10.1007/978-3-319-58540-6} } +@Book{ Eisenbud1995, + title = "Commutative algebra", + author = "Eisenbud, David", + publisher = "Springer", + series = "Graduate Texts in Mathematics", + month = mar, + year = 1995, + address = "New York, NY", + language = "en", + isbn = {978-0-387-94268-1}, + doi = {10.1007/978-1-4612-5350-1} +} + @Book{ Elephant, title = {Sketches of an Elephant – A Topos Theory Compendium}, author = {Peter Johnstone}, diff --git a/src/linear_algebra/matrix/charpoly/linear_map.lean b/src/linear_algebra/matrix/charpoly/linear_map.lean new file mode 100644 index 0000000000000..34d47b0fcf5dc --- /dev/null +++ b/src/linear_algebra/matrix/charpoly/linear_map.lean @@ -0,0 +1,245 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import linear_algebra.matrix.charpoly.coeff +import linear_algebra.matrix.to_lin + +/-! + +# Calyley-Hamilton theorem for f.g. modules. + +Given a fixed finite spanning set `b : ι → M` of a `R`-module `M`, we say that a matrix `M` +represents an endomorphism `f : M →ₗ[R] M` if the matrix as an endomorphism of `ι → R` commutes +with `f` via the projection `(ι → R) →ₗ[R] M` given by `b`. + +We show that every endomorphism has a matrix representation, and if `f.range ≤ I • ⊤` for some +ideal `I`, we may furthermore obtain a matrix representation whose entries fall in `I`. + +This is used to conclude the Cayley-Hamilton theorem for f.g. modules over arbitrary rings. +-/ + +variables {ι : Type*} [fintype ι] +variables {M : Type*} [add_comm_group M] (R : Type*) [comm_ring R] [module R M] (I : ideal R) +variables (b : ι → M) (hb : submodule.span R (set.range b) = ⊤) + +open_locale big_operators + +/-- The composition of a matrix (as an endomporphism of `ι → R`) with the projection +`(ι → R) →ₗ[R] M`. -/ +def pi_to_module.from_matrix [decidable_eq ι] : matrix ι ι R →ₗ[R] (ι → R) →ₗ[R] M := +(linear_map.llcomp R _ _ _ (fintype.total R R b)).comp alg_equiv_matrix'.symm.to_linear_map + +lemma pi_to_module.from_matrix_apply [decidable_eq ι] (A : matrix ι ι R) (w : ι → R) : + pi_to_module.from_matrix R b A w = fintype.total R R b (A.mul_vec w) := rfl + +lemma pi_to_module.from_matrix_apply_single_one [decidable_eq ι] (A : matrix ι ι R) (j : ι) : + pi_to_module.from_matrix R b A (pi.single j 1) = ∑ (i : ι), A i j • b i := +begin + rw [pi_to_module.from_matrix_apply, fintype.total_apply, matrix.mul_vec_single], + simp_rw [mul_one] +end + +/-- The endomorphisms of `M` acts on `(ι → R) →ₗ[R] M`, and takes the projection +to a `(ι → R) →ₗ[R] M`. -/ +def pi_to_module.from_End : (module.End R M) →ₗ[R] (ι → R) →ₗ[R] M := +linear_map.lcomp _ _ (fintype.total R R b) + +lemma pi_to_module.from_End_apply (f : module.End R M) (w : ι → R) : + pi_to_module.from_End R b f w = f (fintype.total R R b w) := rfl + +lemma pi_to_module.from_End_apply_single_one [decidable_eq ι] (f : module.End R M) (i : ι) : + pi_to_module.from_End R b f (pi.single i 1) = f (b i) := +begin + rw pi_to_module.from_End_apply, + congr, + convert fintype.total_apply_single R b i 1, + rw one_smul, +end + +lemma pi_to_module.from_End_injective (hb : submodule.span R (set.range b) = ⊤) : + function.injective (pi_to_module.from_End R b) := +begin + intros x y e, + ext m, + obtain ⟨m, rfl⟩ : m ∈ (fintype.total R R b).range, + { rw (fintype.range_total R b).trans hb, trivial }, + exact (linear_map.congr_fun e m : _) +end + +section + +variables {R} [decidable_eq ι] + +/-- We say that a matrix represents an endomorphism of `M` if the matrix acting on `ι → R` is +equal to `f` via the projection `(ι → R) →ₗ[R] M` given by a fixed (spanning) set. -/ +def matrix.represents (A : matrix ι ι R) (f : module.End R M) : Prop := +pi_to_module.from_matrix R b A = pi_to_module.from_End R b f + +variables {b} + +lemma matrix.represents.congr_fun {A : matrix ι ι R} {f : module.End R M} + (h : A.represents b f) (x) : + fintype.total R R b (A.mul_vec x) = f (fintype.total R R b x) := +linear_map.congr_fun h x + +lemma matrix.represents_iff {A : matrix ι ι R} {f : module.End R M} : + A.represents b f ↔ + ∀ x, fintype.total R R b (A.mul_vec x) = f (fintype.total R R b x) := +⟨λ e x, e.congr_fun x, λ H, linear_map.ext $ λ x, H x⟩ + +lemma matrix.represents_iff' {A : matrix ι ι R} {f : module.End R M} : + A.represents b f ↔ ∀ j, ∑ (i : ι), A i j • b i = f (b j) := +begin + split, + { intros h i, + have := linear_map.congr_fun h (pi.single i 1), + rwa [pi_to_module.from_End_apply_single_one, + pi_to_module.from_matrix_apply_single_one] at this }, + { intros h, + ext, + simp_rw [linear_map.comp_apply, linear_map.coe_single, pi_to_module.from_End_apply_single_one, + pi_to_module.from_matrix_apply_single_one], + apply h } +end + +lemma matrix.represents.mul {A A' : matrix ι ι R} {f f' : module.End R M} + (h : A.represents b f) (h' : matrix.represents b A' f') : + (A * A').represents b (f * f') := +begin + delta matrix.represents pi_to_module.from_matrix at ⊢, + rw [linear_map.comp_apply, alg_equiv.to_linear_map_apply, map_mul], + ext, + dsimp [pi_to_module.from_End], + rw [← h'.congr_fun, ← h.congr_fun], + refl, +end + +lemma matrix.represents.one : (1 : matrix ι ι R).represents b 1 := +begin + delta matrix.represents pi_to_module.from_matrix, + rw [linear_map.comp_apply, alg_equiv.to_linear_map_apply, map_one], + ext, + refl +end + +lemma matrix.represents.add {A A' : matrix ι ι R} {f f' : module.End R M} + (h : A.represents b f) (h' : matrix.represents b A' f') : + (A + A').represents b (f + f') := +begin + delta matrix.represents at ⊢ h h', rw [map_add, map_add, h, h'], +end + +lemma matrix.represents.zero : + (0 : matrix ι ι R).represents b 0 := +begin + delta matrix.represents, rw [map_zero, map_zero], +end + +lemma matrix.represents.smul {A : matrix ι ι R} {f : module.End R M} + (h : A.represents b f) (r : R) : + (r • A).represents b (r • f) := +begin + delta matrix.represents at ⊢ h, rw [map_smul, map_smul, h], +end + +lemma matrix.represents.eq {A : matrix ι ι R} {f f' : module.End R M} + (h : A.represents b f) (h' : A.represents b f') : f = f' := +pi_to_module.from_End_injective R b hb (h.symm.trans h') + +variables (b R) + +/-- The subalgebra of `matrix ι ι R` that consists of matrices that actually represent +endomorphisms on `M`. -/ +def matrix.is_representation : subalgebra R (matrix ι ι R) := +{ carrier := { A | ∃ f : module.End R M, A.represents b f }, + mul_mem' := λ A₁ A₂ ⟨f₁, e₁⟩ ⟨f₂, e₂⟩, ⟨f₁ * f₂, e₁.mul e₂⟩, + one_mem' := ⟨1, matrix.represents.one⟩, + add_mem' := λ A₁ A₂ ⟨f₁, e₁⟩ ⟨f₂, e₂⟩, ⟨f₁ + f₂, e₁.add e₂⟩, + zero_mem' := ⟨0, matrix.represents.zero⟩, + algebra_map_mem' := λ r, ⟨r • 1, matrix.represents.one.smul r⟩ } + +/-- The map sending a matrix to the endomorphism it represents. This is an `R`-algebra morphism. -/ +noncomputable +def matrix.is_representation.to_End : matrix.is_representation R b →ₐ[R] module.End R M := +{ to_fun := λ A, A.2.some, + map_one' := (1 : matrix.is_representation R b).2.some_spec.eq hb matrix.represents.one, + map_mul' := λ A₁ A₂, (A₁ * A₂).2.some_spec.eq hb (A₁.2.some_spec.mul A₂.2.some_spec), + map_zero' := (0 : matrix.is_representation R b).2.some_spec.eq hb matrix.represents.zero, + map_add' := λ A₁ A₂, (A₁ + A₂).2.some_spec.eq hb (A₁.2.some_spec.add A₂.2.some_spec), + commutes' := λ r, (r • 1 : matrix.is_representation R b).2.some_spec.eq + hb (matrix.represents.one.smul r) } + +lemma matrix.is_representation.to_End_represents (A : matrix.is_representation R b) : + (A : matrix ι ι R).represents b (matrix.is_representation.to_End R b hb A) := +A.2.some_spec + +lemma matrix.is_representation.eq_to_End_of_represents (A : matrix.is_representation R b) + {f : module.End R M} (h : (A : matrix ι ι R).represents b f) : + matrix.is_representation.to_End R b hb A = f := +A.2.some_spec.eq hb h + +lemma matrix.is_representation.to_End_exists_mem_ideal + (f : module.End R M) (I : ideal R) (hI : f.range ≤ I • ⊤) : + ∃ M, matrix.is_representation.to_End R b hb M = f ∧ ∀ i j, M.1 i j ∈ I := +begin + have : ∀ x, f x ∈ (ideal.finsupp_total ι M I b).range, + { rw [ideal.range_finsupp_total, hb], exact λ x, hI (f.mem_range_self x) }, + choose bM' hbM', + let A : matrix ι ι R := λ i j, bM' (b j) i, + have : A.represents b f, + { rw matrix.represents_iff', + dsimp [A], + intro j, + specialize hbM' (b j), + rwa ideal.finsupp_total_apply_eq_of_fintype at hbM' }, + exact ⟨⟨A, f, this⟩, matrix.is_representation.eq_to_End_of_represents R b hb ⟨A, f, this⟩ this, + λ i j, (bM' (b j) i).prop⟩, +end + +lemma matrix.is_representation.to_End_surjective : + function.surjective (matrix.is_representation.to_End R b hb) := +begin + intro f, + obtain ⟨M, e, -⟩ := matrix.is_representation.to_End_exists_mem_ideal R b hb f ⊤ _, + exact ⟨M, e⟩, + simp, +end + +end + +/-- +The **Cayley-Hamilton Theorem** for f.g. modules over arbitrary rings states that for each +`R`-endomorphism `φ` of an `R`-module `M` such that `φ(M) ≤ I • M` for some ideal `I`, there +exists some `n` and some `aᵢ ∈ Iⁱ` such that `φⁿ + a₁ φⁿ⁻¹ + ⋯ + aₙ = 0`. + +This is the version found in Eisenbud 4.3, which is slightly weaker than Matsumura 2.1 +(this lacks the constraint on `n`), and is slightly stronger than Atiyah-Macdonald 2.4. +-/ +lemma linear_map.exists_monic_and_coeff_mem_pow_and_aeval_eq_zero_of_range_le_smul + [module.finite R M] (f : module.End R M) (I : ideal R) (hI : f.range ≤ I • ⊤) : + ∃ p : polynomial R, + p.monic ∧ (∀ k, p.coeff k ∈ I ^ (p.nat_degree - k)) ∧ polynomial.aeval f p = 0 := +begin + classical, + casesI subsingleton_or_nontrivial R, + { exactI ⟨0, polynomial.monic_of_subsingleton _, by simp⟩ }, + obtain ⟨s : finset M, hs : submodule.span R (s : set M) = ⊤⟩ := module.finite.out, + obtain ⟨A, rfl, h⟩ := matrix.is_representation.to_End_exists_mem_ideal R (coe : s → M) + (by rw [subtype.range_coe_subtype, finset.set_of_mem, hs]) f I hI, + refine ⟨A.1.charpoly, A.1.charpoly_monic, _, _⟩, + { rw A.1.charpoly_nat_degree_eq_dim, + exact coeff_charpoly_mem_ideal_pow h }, + { rw [polynomial.aeval_alg_hom_apply, ← map_zero (matrix.is_representation.to_End R coe _)], + congr' 1, + ext1, + rw [polynomial.aeval_subalgebra_coe, subtype.val_eq_coe, matrix.aeval_self_charpoly, + subalgebra.coe_zero] }, + { apply_instance } +end + +lemma linear_map.exists_monic_and_aeval_eq_zero [module.finite R M] + (f : module.End R M) : ∃ p : polynomial R, p.monic ∧ polynomial.aeval f p = 0 := +(linear_map.exists_monic_and_coeff_mem_pow_and_aeval_eq_zero_of_range_le_smul R f ⊤ (by simp)).imp + (λ p h, h.imp_right and.elim_right) From 51cbe88849c5bcf9eaf8e38f2bbdf1a44bbabe0c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 8 Oct 2022 03:03:11 +0000 Subject: [PATCH 620/828] feat(order/category/HeytAlg): The category of Heyting algebras (#16250) Define `HeytAlg`, the category of Heyting algebras with Heyting morphisms. --- src/order/category/BoolAlg.lean | 14 +++++++-- src/order/category/HeytAlg.lean | 55 +++++++++++++++++++++++++++++++++ 2 files changed, 66 insertions(+), 3 deletions(-) create mode 100644 src/order/category/HeytAlg.lean diff --git a/src/order/category/BoolAlg.lean b/src/order/category/BoolAlg.lean index 7301e685995a9..6aaaa3a2fab5f 100644 --- a/src/order/category/BoolAlg.lean +++ b/src/order/category/BoolAlg.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ -import order.category.BoundedDistribLattice +import order.category.HeytAlg /-! # The category of boolean algebras @@ -43,8 +43,16 @@ instance : concrete_category BoolAlg := induced_category.concrete_category to_Bo instance has_forget_to_BoundedDistribLattice : has_forget₂ BoolAlg BoundedDistribLattice := induced_category.has_forget₂ to_BoundedDistribLattice -/-- Constructs an equivalence between boolean algebras from an order isomorphism -between them. -/ +section + +local attribute [instance] bounded_lattice_hom_class.to_biheyting_hom_class + +@[simps] instance has_forget_to_HeytAlg : has_forget₂ BoolAlg HeytAlg := +{ forget₂ := { obj := λ X, ⟨X⟩, map := λ X Y f, show bounded_lattice_hom X Y, from f } } + +end + +/-- Constructs an equivalence between Boolean algebras from an order isomorphism between them. -/ @[simps] def iso.mk {α β : BoolAlg.{u}} (e : α ≃o β) : α ≅ β := { hom := (e : bounded_lattice_hom α β), inv := (e.symm : bounded_lattice_hom β α), diff --git a/src/order/category/HeytAlg.lean b/src/order/category/HeytAlg.lean new file mode 100644 index 0000000000000..5e3c7af5b84b4 --- /dev/null +++ b/src/order/category/HeytAlg.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import order.category.BoundedDistribLattice +import order.heyting.hom + +/-! +# The category of Heyting algebras + +This file defines `HeytAlg`, the category of Heyting algebras. +-/ + +universes u + +open category_theory opposite order + +/-- The category of Heyting algebras. -/ +def HeytAlg := bundled heyting_algebra + +namespace HeytAlg + +instance : has_coe_to_sort HeytAlg Type* := bundled.has_coe_to_sort +instance (X : HeytAlg) : heyting_algebra X := X.str + +/-- Construct a bundled `HeytAlg` from a `heyting_algebra`. -/ +def of (α : Type*) [heyting_algebra α] : HeytAlg := bundled.of α + +@[simp] lemma coe_of (α : Type*) [heyting_algebra α] : ↥(of α) = α := rfl + +instance : inhabited HeytAlg := ⟨of punit⟩ + +instance bundled_hom : bundled_hom heyting_hom := +{ to_fun := λ α β [heyting_algebra α] [heyting_algebra β], + by exactI (coe_fn : heyting_hom α β → α → β), + id := heyting_hom.id, + comp := @heyting_hom.comp, + hom_ext := λ α β [heyting_algebra α] [heyting_algebra β], by exactI fun_like.coe_injective } + +attribute [derive [large_category, concrete_category]] HeytAlg + +@[simps] +instance has_forget_to_Lattice : has_forget₂ HeytAlg BoundedDistribLattice := +{ forget₂ := { obj := λ X, BoundedDistribLattice.of X, + map := λ X Y f, (f : bounded_lattice_hom X Y) } } + +/-- Constructs an isomorphism of Heyting algebras from an order isomorphism between them. -/ +@[simps] def iso.mk {α β : HeytAlg.{u}} (e : α ≃o β) : α ≅ β := +{ hom := e, + inv := e.symm, + hom_inv_id' := by { ext, exact e.symm_apply_apply _ }, + inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } + +end HeytAlg From 052b6cbe18941597628794e179225ff585cb93bf Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sat, 8 Oct 2022 03:03:12 +0000 Subject: [PATCH 621/828] feat(topology/algebra/order/left_right): left and right limits (#16551) We define left and right limits of a function, and prove that a monotone function is continuous at a point if and only if its left and right limits coincide. We also refactor the file on Stieltjes functions using this general definition instead of the ad hoc definition that is currently used. --- src/measure_theory/measure/stieltjes.lean | 112 +++--- src/topology/algebra/order/left_right.lean | 16 + .../algebra/order/left_right_lim.lean | 364 ++++++++++++++++++ 3 files changed, 441 insertions(+), 51 deletions(-) create mode 100644 src/topology/algebra/order/left_right_lim.lean diff --git a/src/measure_theory/measure/stieltjes.lean b/src/measure_theory/measure/stieltjes.lean index 77600cdcf4d72..da738a8787086 100644 --- a/src/measure_theory/measure/stieltjes.lean +++ b/src/measure_theory/measure/stieltjes.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov, Sébastien Gouëzel -/ import measure_theory.constructions.borel_space +import topology.algebra.order.left_right_lim /-! # Stieltjes measures on the real line @@ -16,14 +17,13 @@ corrresponding measure, giving mass `f b - f a` to the interval `(a, b]`. * `stieltjes_function` is a structure containing a function from `ℝ → ℝ`, together with the assertions that it is monotone and right-continuous. To `f : stieltjes_function`, one associates a Borel measure `f.measure`. -* `f.left_lim x` is the limit of `f` to the left of `x`. * `f.measure_Ioc` asserts that `f.measure (Ioc a b) = of_real (f b - f a)` -* `f.measure_Ioo` asserts that `f.measure (Ioo a b) = of_real (f.left_lim b - f a)`. +* `f.measure_Ioo` asserts that `f.measure (Ioo a b) = of_real (left_lim f b - f a)`. * `f.measure_Icc` and `f.measure_Ico` are analogous. -/ noncomputable theory -open classical set filter +open classical set filter function open ennreal (of_real) open_locale big_operators ennreal nnreal topological_space measure_theory @@ -47,46 +47,54 @@ lemma mono : monotone f := f.mono' lemma right_continuous (x : ℝ) : continuous_within_at f (Ici x) x := f.right_continuous' x -/-- The limit of a Stieltjes function to the left of `x` (it exists by monotonicity). The fact that -it is indeed a left limit is asserted in `tendsto_left_lim` -/ -@[irreducible] def left_lim (x : ℝ) := Sup (f '' (Iio x)) - -lemma tendsto_left_lim (x : ℝ) : tendsto f (𝓝[<] x) (𝓝 (f.left_lim x)) := -by { rw left_lim, exact f.mono.tendsto_nhds_within_Iio x } - -lemma left_lim_le {x y : ℝ} (h : x ≤ y) : f.left_lim x ≤ f y := -begin - apply le_of_tendsto (f.tendsto_left_lim x), - filter_upwards [self_mem_nhds_within] with _ hz using (f.mono (le_of_lt hz)).trans (f.mono h), -end - -lemma le_left_lim {x y : ℝ} (h : x < y) : f x ≤ f.left_lim y := -begin - apply ge_of_tendsto (f.tendsto_left_lim y), - apply mem_nhds_within_Iio_iff_exists_Ioo_subset.2 ⟨x, h, _⟩, - assume z hz, - exact f.mono hz.1.le, -end - -lemma left_lim_le_left_lim {x y : ℝ} (h : x ≤ y) : f.left_lim x ≤ f.left_lim y := -begin - rcases eq_or_lt_of_le h with rfl|hxy, - { exact le_rfl }, - { exact (f.left_lim_le le_rfl).trans (f.le_left_lim hxy) } -end - /-- The identity of `ℝ` as a Stieltjes function, used to construct Lebesgue measure. -/ @[simps] protected def id : stieltjes_function := { to_fun := id, mono' := λ x y, id, right_continuous' := λ x, continuous_within_at_id } -@[simp] lemma id_left_lim (x : ℝ) : stieltjes_function.id.left_lim x = x := -tendsto_nhds_unique (stieltjes_function.id.tendsto_left_lim x) $ +@[simp] lemma id_left_lim (x : ℝ) : left_lim stieltjes_function.id x = x := +tendsto_nhds_unique (stieltjes_function.id.mono.tendsto_left_lim x) $ (continuous_at_id).tendsto.mono_left nhds_within_le_nhds instance : inhabited stieltjes_function := ⟨stieltjes_function.id⟩ +/-- If a function `f : ℝ → ℝ` is monotone, then the function mapping `x` to the right limit of `f` +at `x` is a Stieltjes function, i.e., it is monotone and right-continuous. -/ +noncomputable def _root_.monotone.stieltjes_function {f : ℝ → ℝ} (hf : monotone f) : + stieltjes_function := +{ to_fun := right_lim f, + mono' := λ x y hxy, hf.right_lim hxy, + right_continuous' := + begin + assume x s hs, + obtain ⟨l, u, hlu, lus⟩ : ∃ (l u : ℝ), right_lim f x ∈ Ioo l u ∧ Ioo l u ⊆ s := + mem_nhds_iff_exists_Ioo_subset.1 hs, + obtain ⟨y, xy, h'y⟩ : ∃ (y : ℝ) (H : x < y), Ioc x y ⊆ f ⁻¹' (Ioo l u) := + mem_nhds_within_Ioi_iff_exists_Ioc_subset.1 + (hf.tendsto_right_lim x (Ioo_mem_nhds hlu.1 hlu.2)), + change ∀ᶠ y in 𝓝[≥] x, right_lim f y ∈ s, + filter_upwards [Ico_mem_nhds_within_Ici ⟨le_refl x, xy⟩] with z hz, + apply lus, + refine ⟨hlu.1.trans_le (hf.right_lim hz.1), _⟩, + obtain ⟨a, za, ay⟩ : ∃ (a : ℝ), z < a ∧ a < y := exists_between hz.2, + calc right_lim f z ≤ f a : hf.right_lim_le za + ... < u : (h'y ⟨hz.1.trans_lt za, ay.le⟩).2, + end } + +lemma _root_.monotone.stieltjes_function_eq {f : ℝ → ℝ} (hf : monotone f) (x : ℝ) : + hf.stieltjes_function x = right_lim f x := rfl + +lemma countable_left_lim_ne (f : stieltjes_function) : + set.countable {x | left_lim f x ≠ f x} := +begin + apply countable.mono _ (f.mono.countable_not_continuous_at), + assume x hx h'x, + apply hx, + exact tendsto_nhds_unique (f.mono.tendsto_left_lim x) (h'x.tendsto.mono_left nhds_within_le_nhds), +end + + /-! ### The outer measure associated to a Stieltjes function -/ /-- Length of an interval. This is the largest monotone function which correctly measures all @@ -281,7 +289,7 @@ interval `(a, b]`. -/ @[simp] lemma measure_Ioc (a b : ℝ) : f.measure (Ioc a b) = of_real (f b - f a) := by { rw stieltjes_function.measure, exact f.outer_Ioc a b } -@[simp] lemma measure_singleton (a : ℝ) : f.measure {a} = of_real (f a - f.left_lim a) := +@[simp] lemma measure_singleton (a : ℝ) : f.measure {a} = of_real (f a - left_lim f a) := begin obtain ⟨u, u_mono, u_lt_a, u_lim⟩ : ∃ (u : ℕ → ℝ), strict_mono u ∧ (∀ (n : ℕ), u n < a) ∧ tendsto u at_top (𝓝 a) := exists_seq_strict_mono_tendsto a, @@ -295,54 +303,56 @@ begin refine tendsto_measure_Inter (λ n, measurable_set_Ioc) (λ m n hmn, _) _, { exact Ioc_subset_Ioc (u_mono.monotone hmn) le_rfl }, { exact ⟨0, by simpa only [measure_Ioc] using ennreal.of_real_ne_top⟩ } }, - have L2 : tendsto (λ n, f.measure (Ioc (u n) a)) at_top (𝓝 (of_real (f a - f.left_lim a))), + have L2 : tendsto (λ n, f.measure (Ioc (u n) a)) at_top (𝓝 (of_real (f a - left_lim f a))), { simp only [measure_Ioc], - have : tendsto (λ n, f (u n)) at_top (𝓝 (f.left_lim a)), - { apply (f.tendsto_left_lim a).comp, + have : tendsto (λ n, f (u n)) at_top (𝓝 (left_lim f a)), + { apply (f.mono.tendsto_left_lim a).comp, exact tendsto_nhds_within_of_tendsto_nhds_of_eventually_within _ u_lim (eventually_of_forall (λ n, u_lt_a n)) }, exact ennreal.continuous_of_real.continuous_at.tendsto.comp (tendsto_const_nhds.sub this) }, exact tendsto_nhds_unique L1 L2 end -@[simp] lemma measure_Icc (a b : ℝ) : f.measure (Icc a b) = of_real (f b - f.left_lim a) := +@[simp] lemma measure_Icc (a b : ℝ) : f.measure (Icc a b) = of_real (f b - left_lim f a) := begin rcases le_or_lt a b with hab|hab, { have A : disjoint {a} (Ioc a b), by simp, - simp [← Icc_union_Ioc_eq_Icc le_rfl hab, -singleton_union, ← ennreal.of_real_add, f.left_lim_le, - measure_union A measurable_set_Ioc, f.mono hab] }, + simp [← Icc_union_Ioc_eq_Icc le_rfl hab, -singleton_union, ← ennreal.of_real_add, + f.mono.left_lim_le, measure_union A measurable_set_Ioc, f.mono hab] }, { simp only [hab, measure_empty, Icc_eq_empty, not_le], symmetry, - simp [ennreal.of_real_eq_zero, f.le_left_lim hab] } + simp [ennreal.of_real_eq_zero, f.mono.le_left_lim hab] } end -@[simp] lemma measure_Ioo {a b : ℝ} : f.measure (Ioo a b) = of_real (f.left_lim b - f a) := +@[simp] lemma measure_Ioo {a b : ℝ} : f.measure (Ioo a b) = of_real (left_lim f b - f a) := begin rcases le_or_lt b a with hab|hab, { simp only [hab, measure_empty, Ioo_eq_empty, not_lt], symmetry, - simp [ennreal.of_real_eq_zero, f.left_lim_le hab] }, + simp [ennreal.of_real_eq_zero, f.mono.left_lim_le hab] }, { have A : disjoint (Ioo a b) {b}, by simp, - have D : f b - f a = (f b - f.left_lim b) + (f.left_lim b - f a), by abel, + have D : f b - f a = (f b - left_lim f b) + (left_lim f b - f a), by abel, have := f.measure_Ioc a b, simp only [←Ioo_union_Icc_eq_Ioc hab le_rfl, measure_singleton, measure_union A (measurable_set_singleton b), Icc_self] at this, rw [D, ennreal.of_real_add, add_comm] at this, { simpa only [ennreal.add_right_inj ennreal.of_real_ne_top] }, - { simp only [f.left_lim_le, sub_nonneg] }, - { simp only [f.le_left_lim hab, sub_nonneg] } }, + { simp only [f.mono.left_lim_le, sub_nonneg] }, + { simp only [f.mono.le_left_lim hab, sub_nonneg] } }, end -@[simp] lemma measure_Ico (a b : ℝ) : f.measure (Ico a b) = of_real (f.left_lim b - f.left_lim a) := +@[simp] lemma measure_Ico (a b : ℝ) : f.measure (Ico a b) = of_real (left_lim f b - left_lim f a) := begin rcases le_or_lt b a with hab|hab, { simp only [hab, measure_empty, Ico_eq_empty, not_lt], symmetry, - simp [ennreal.of_real_eq_zero, f.left_lim_le_left_lim hab] }, + simp [ennreal.of_real_eq_zero, f.mono.left_lim hab] }, { have A : disjoint {a} (Ioo a b) := by simp, - simp [← Icc_union_Ioo_eq_Ico le_rfl hab, -singleton_union, hab.ne, f.left_lim_le, - measure_union A measurable_set_Ioo, f.le_left_lim hab, - ← ennreal.of_real_add] } + simp [← Icc_union_Ioo_eq_Ico le_rfl hab, -singleton_union, hab.ne, f.mono.left_lim_le, + measure_union A measurable_set_Ioo, f.mono.le_left_lim hab, ← ennreal.of_real_add] } end +instance : is_locally_finite_measure f.measure := +⟨λ x, ⟨Ioo (x-1) (x+1), Ioo_mem_nhds (by linarith) (by linarith), by simp⟩⟩ + end stieltjes_function diff --git a/src/topology/algebra/order/left_right.lean b/src/topology/algebra/order/left_right.lean index 21e4190f327e4..12a4d6582dd20 100644 --- a/src/topology/algebra/order/left_right.lean +++ b/src/topology/algebra/order/left_right.lean @@ -37,8 +37,18 @@ lemma continuous_within_at_Iio_iff_Iic {a : α} {f : α → β} : continuous_within_at f (Iio a) a ↔ continuous_within_at f (Iic a) a := @continuous_within_at_Ioi_iff_Ici αᵒᵈ _ ‹topological_space α› _ _ _ f +lemma nhds_left'_le_nhds_ne (a : α) : + 𝓝[<] a ≤ 𝓝[≠] a := +nhds_within_mono a (λ y hy, ne_of_lt hy) + +lemma nhds_right'_le_nhds_ne (a : α) : + 𝓝[>] a ≤ 𝓝[≠] a := +nhds_within_mono a (λ y hy, ne_of_gt hy) + end partial_order +section topological_space + variables {α β : Type*} [topological_space α] [linear_order α] [topological_space β] lemma nhds_left_sup_nhds_right (a : α) : @@ -53,6 +63,10 @@ lemma nhds_left_sup_nhds_right' (a : α) : 𝓝[≤] a ⊔ 𝓝[>] a = 𝓝 a := by rw [← nhds_within_union, Iic_union_Ioi, nhds_within_univ] +lemma nhds_left'_sup_nhds_right' (a : α) : + 𝓝[<] a ⊔ 𝓝[>] a = 𝓝[≠] a := +by rw [← nhds_within_union, Iio_union_Ioi] + lemma continuous_at_iff_continuous_left_right {a : α} {f : α → β} : continuous_at f a ↔ continuous_within_at f (Iic a) a ∧ continuous_within_at f (Ici a) a := by simp only [continuous_within_at, continuous_at, ← tendsto_sup, nhds_left_sup_nhds_right] @@ -61,3 +75,5 @@ lemma continuous_at_iff_continuous_left'_right' {a : α} {f : α → β} : continuous_at f a ↔ continuous_within_at f (Iio a) a ∧ continuous_within_at f (Ioi a) a := by rw [continuous_within_at_Ioi_iff_Ici, continuous_within_at_Iio_iff_Iic, continuous_at_iff_continuous_left_right] + +end topological_space diff --git a/src/topology/algebra/order/left_right_lim.lean b/src/topology/algebra/order/left_right_lim.lean new file mode 100644 index 0000000000000..0e3a0d1ca0e8d --- /dev/null +++ b/src/topology/algebra/order/left_right_lim.lean @@ -0,0 +1,364 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import topology.algebra.order.basic +import topology.algebra.order.left_right + +/-! +# Left and right limits + +We define the (strict) left and right limits of a function. + +* `left_lim f x` is the strict left limit of `f` at `x` (using `f x` as a garbage value if `x` + is isolated to its left). +* `right_lim f x` is the strict right limit of `f` at `x` (using `f x` as a garbage value if `x` + is isolated to its right). + +We develop a comprehensive API for monotone functions. Notably, + +* `monotone.continuous_at_iff_left_lim_eq_right_lim` states that a monotone function is continuous + at a point if and only if its left and right limits coincide. +* `monotone.countable_not_continuous_at` asserts that a monotone function taking values in a + second-countable space has at most countably many discontinuity points. + +We also port the API to antitone functions. + +## TODO + +Prove corresponding stronger results for strict_mono and strict_anti functions. +-/ + +open set filter +open_locale topological_space + +section + +variables {α β : Type*} [linear_order α] [topological_space β] + +/-- Let `f : α → β` be a function from a linear order `α` to a topological_space `β`, and +let `a : α`. The limit strictly to the left of `f` at `a`, denoted with `left_lim f a`, is defined +by using the order topology on `α`. If `a` is isolated to its left or the function has no left +limit, we use `f a` instead to guarantee a good behavior in most cases. -/ +@[irreducible] noncomputable def function.left_lim (f : α → β) (a : α) : β := +begin + classical, + haveI : nonempty β := ⟨f a⟩, + letI : topological_space α := preorder.topology α, + exact if (𝓝[<] a = ⊥) ∨ ¬(∃ y, tendsto f (𝓝[<] a) (𝓝 y)) then f a + else lim (𝓝[<] a) f +end + +/-- Let `f : α → β` be a function from a linear order `α` to a topological_space `β`, and +let `a : α`. The limit strictly to the right of `f` at `a`, denoted with `right_lim f a`, is defined +by using the order topology on `α`. If `a` is isolated to its right or the function has no right +limit, , we use `f a` instead to guarantee a good behavior in most cases. -/ +noncomputable def function.right_lim (f : α → β) (a : α) : β := +@function.left_lim αᵒᵈ β _ _ f a + +open function + +lemma left_lim_eq_of_tendsto + [hα : topological_space α] [h'α : order_topology α] [t2_space β] + {f : α → β} {a : α} {y : β} (h : 𝓝[<] a ≠ ⊥) (h' : tendsto f (𝓝[<] a) (𝓝 y)) : + left_lim f a = y := +begin + have h'' : ∃ y, tendsto f (𝓝[<] a) (𝓝 y) := ⟨y, h'⟩, + rw [h'α.topology_eq_generate_intervals] at h h' h'', + simp only [left_lim, h, h'', not_true, or_self, if_false], + haveI := ne_bot_iff.2 h, + exact h'.lim_eq, +end + +lemma left_lim_eq_of_eq_bot [hα : topological_space α] [h'α : order_topology α] + (f : α → β) {a : α} (h : 𝓝[<] a = ⊥) : + left_lim f a = f a := +begin + rw [h'α.topology_eq_generate_intervals] at h, + simp [left_lim, ite_eq_left_iff, h], +end + +end + +open function + +namespace monotone + +variables {α β : Type*} [linear_order α] [conditionally_complete_linear_order β] +[topological_space β] [order_topology β] {f : α → β} (hf : monotone f) {x y : α} +include hf + +lemma left_lim_eq_Sup [topological_space α] [order_topology α] (h : 𝓝[<] x ≠ ⊥) : + left_lim f x = Sup (f '' (Iio x)) := +left_lim_eq_of_tendsto h (hf.tendsto_nhds_within_Iio x) + +lemma left_lim_le (h : x ≤ y) : left_lim f x ≤ f y := +begin + letI : topological_space α := preorder.topology α, + haveI : order_topology α := ⟨rfl⟩, + rcases eq_or_ne (𝓝[<] x) ⊥ with h'|h', + { simpa [left_lim, h'] using hf h }, + haveI A : ne_bot (𝓝[<] x) := ne_bot_iff.2 h', + rw left_lim_eq_Sup hf h', + refine cSup_le _ _, + { simp only [nonempty_image_iff], + exact (forall_mem_nonempty_iff_ne_bot.2 A) _ self_mem_nhds_within }, + { simp only [mem_image, mem_Iio, forall_exists_index, and_imp, forall_apply_eq_imp_iff₂], + assume z hz, + exact hf (hz.le.trans h) }, +end + +lemma le_left_lim (h : x < y) : f x ≤ left_lim f y := +begin + letI : topological_space α := preorder.topology α, + haveI : order_topology α := ⟨rfl⟩, + rcases eq_or_ne (𝓝[<] y) ⊥ with h'|h', + { rw left_lim_eq_of_eq_bot _ h', exact hf h.le }, + rw left_lim_eq_Sup hf h', + refine le_cSup ⟨f y, _⟩ (mem_image_of_mem _ h), + simp only [upper_bounds, mem_image, mem_Iio, forall_exists_index, and_imp, + forall_apply_eq_imp_iff₂, mem_set_of_eq], + assume z hz, + exact hf hz.le +end + +@[mono] protected lemma left_lim : monotone (left_lim f) := +begin + assume x y h, + rcases eq_or_lt_of_le h with rfl|hxy, + { exact le_rfl }, + { exact (hf.left_lim_le le_rfl).trans (hf.le_left_lim hxy) } +end + +lemma le_right_lim (h : x ≤ y) : f x ≤ right_lim f y := +hf.dual.left_lim_le h + +lemma right_lim_le (h : x < y) : right_lim f x ≤ f y := +hf.dual.le_left_lim h + +@[mono] protected lemma right_lim : monotone (right_lim f) := +λ x y h, hf.dual.left_lim h + +lemma left_lim_le_right_lim (h : x ≤ y) : left_lim f x ≤ right_lim f y := +(hf.left_lim_le le_rfl).trans (hf.le_right_lim h) + +lemma right_lim_le_left_lim (h : x < y) : right_lim f x ≤ left_lim f y := +begin + letI : topological_space α := preorder.topology α, + haveI : order_topology α := ⟨rfl⟩, + rcases eq_or_ne (𝓝[<] y) ⊥ with h'|h', + { simp [left_lim, h'], + exact right_lim_le hf h }, + obtain ⟨a, ⟨xa, ay⟩⟩ : (Ioo x y).nonempty := + forall_mem_nonempty_iff_ne_bot.2 (ne_bot_iff.2 h') (Ioo x y) + (Ioo_mem_nhds_within_Iio ⟨h, le_refl _⟩), + calc right_lim f x ≤ f a : hf.right_lim_le xa + ... ≤ left_lim f y : hf.le_left_lim ay +end + +variables [topological_space α] [order_topology α] + +lemma tendsto_left_lim (x : α) : tendsto f (𝓝[<] x) (𝓝 (left_lim f x)) := +begin + rcases eq_or_ne (𝓝[<] x) ⊥ with h'|h', + { simp [h'] }, + rw left_lim_eq_Sup hf h', + exact hf.tendsto_nhds_within_Iio x +end + +lemma tendsto_left_lim_within (x : α) : tendsto f (𝓝[<] x) (𝓝[≤] (left_lim f x)) := +begin + apply tendsto_nhds_within_of_tendsto_nhds_of_eventually_within f (hf.tendsto_left_lim x), + filter_upwards [self_mem_nhds_within] with y hy using hf.le_left_lim hy, +end + +lemma tendsto_right_lim (x : α) : + tendsto f (𝓝[>] x) (𝓝 (right_lim f x)) := +hf.dual.tendsto_left_lim x + +lemma tendsto_right_lim_within (x : α) : + tendsto f (𝓝[>] x) (𝓝[≥] (right_lim f x)) := +hf.dual.tendsto_left_lim_within x + +/-- A monotone function is continuous to the left at a point if and only if its left limit +coincides with the value of the function. -/ +lemma continuous_within_at_Iio_iff_left_lim_eq : + continuous_within_at f (Iio x) x ↔ left_lim f x = f x := +begin + rcases eq_or_ne (𝓝[<] x) ⊥ with h'|h', + { simp [left_lim_eq_of_eq_bot f h', continuous_within_at, h'] }, + haveI : (𝓝[Iio x] x).ne_bot := ne_bot_iff.2 h', + refine ⟨λ h, tendsto_nhds_unique (hf.tendsto_left_lim x) h.tendsto, λ h, _⟩, + have := hf.tendsto_left_lim x, + rwa h at this, +end + +/-- A monotone function is continuous to the right at a point if and only if its right limit +coincides with the value of the function. -/ +lemma continuous_within_at_Ioi_iff_right_lim_eq : + continuous_within_at f (Ioi x) x ↔ right_lim f x = f x := +hf.dual.continuous_within_at_Iio_iff_left_lim_eq + +/-- A monotone function is continuous at a point if and only if its left and right limits +coincide. -/ +lemma continuous_at_iff_left_lim_eq_right_lim : + continuous_at f x ↔ left_lim f x = right_lim f x := +begin + refine ⟨λ h, _, λ h, _⟩, + { have A : left_lim f x = f x, + from (hf.continuous_within_at_Iio_iff_left_lim_eq).1 h.continuous_within_at, + have B : right_lim f x = f x, + from (hf.continuous_within_at_Ioi_iff_right_lim_eq).1 h.continuous_within_at, + exact A.trans B.symm }, + { have h' : left_lim f x = f x, + { apply le_antisymm (left_lim_le hf (le_refl _)), + rw h, + exact le_right_lim hf (le_refl _) }, + refine continuous_at_iff_continuous_left'_right'.2 ⟨_, _⟩, + { exact hf.continuous_within_at_Iio_iff_left_lim_eq.2 h' }, + { rw h at h', + exact hf.continuous_within_at_Ioi_iff_right_lim_eq.2 h' } }, +end + +/-- In a second countable space, the set of points where a monotone function is not right-continuous +is at most countable. Superseded by `countable_not_continuous_at` which gives the two-sided +version. -/ +lemma countable_not_continuous_within_at_Ioi [topological_space.second_countable_topology β] : + set.countable {x | ¬(continuous_within_at f (Ioi x) x)} := +begin + /- If `f` is not continuous on the right at `x`, there is an interval `(f x, z x)` which is not + reached by `f`. This gives a family of disjoint open intervals in `β`. Such a family can only + be countable as `β` is second-countable. -/ + nontriviality α, + let s := {x | ¬(continuous_within_at f (Ioi x) x)}, + have : ∀ x, x ∈ s → ∃ z, f x < z ∧ ∀ y, x < y → z ≤ f y, + { rintros x (hx : ¬(continuous_within_at f (Ioi x) x)), + contrapose! hx, + refine tendsto_order.2 ⟨λ m hm, _, λ u hu, _⟩, + { filter_upwards [self_mem_nhds_within] with y hy using hm.trans_le (hf (le_of_lt hy)) }, + rcases hx u hu with ⟨v, xv, fvu⟩, + have : Ioo x v ∈ 𝓝[>] x, from Ioo_mem_nhds_within_Ioi ⟨le_refl _, xv⟩, + filter_upwards [this] with y hy, + apply (hf hy.2.le).trans_lt fvu }, + -- choose `z x` such that `f` does not take the values in `(f x, z x)`. + choose! z hz using this, + have I : inj_on f s, + { apply strict_mono_on.inj_on, + assume x hx y hy hxy, + calc f x < z x : (hz x hx).1 + ... ≤ f y : (hz x hx).2 y hxy }, + -- show that `f s` is countable by arguing that a disjoint family of disjoint open intervals + -- (the intervals `(f x, z x)`) is at most countable. + have fs_count : (f '' s).countable, + { have A : (f '' s).pairwise_disjoint (λ x, Ioo x (z (inv_fun_on f s x))), + { rintros _ ⟨u, us, rfl⟩ _ ⟨v, vs, rfl⟩ huv, + wlog h'uv : u ≤ v := le_total u v using [u v, v u] tactic.skip, + { rcases eq_or_lt_of_le h'uv with rfl|h''uv, + { exact (huv rfl).elim }, + apply disjoint_iff_forall_ne.2, + rintros a ha b hb rfl, + simp [I.left_inv_on_inv_fun_on us, I.left_inv_on_inv_fun_on vs] at ha hb, + exact lt_irrefl _ ((ha.2.trans_le ((hz u us).2 v h''uv)).trans hb.1) }, + { assume hu hv h'uv, + exact (this hv hu h'uv.symm).symm } }, + apply set.pairwise_disjoint.countable_of_Ioo A, + rintros _ ⟨y, ys, rfl⟩, + simpa only [I.left_inv_on_inv_fun_on ys] using (hz y ys).1 }, + exact maps_to.countable_of_inj_on (maps_to_image f s) I fs_count, +end + +/-- In a second countable space, the set of points where a monotone function is not left-continuous +is at most countable. Superseded by `countable_not_continuous_at` which gives the two-sided +version. -/ +lemma countable_not_continuous_within_at_Iio [topological_space.second_countable_topology β] : + set.countable {x | ¬(continuous_within_at f (Iio x) x)} := +hf.dual.countable_not_continuous_within_at_Ioi + +/-- In a second countable space, the set of points where a monotone function is not continuous +is at most countable. -/ +lemma countable_not_continuous_at [topological_space.second_countable_topology β] : + set.countable {x | ¬(continuous_at f x)} := +begin + apply (hf.countable_not_continuous_within_at_Ioi.union + hf.countable_not_continuous_within_at_Iio).mono _, + refine compl_subset_compl.1 _, + simp only [compl_union], + rintros x ⟨hx, h'x⟩, + simp only [mem_set_of_eq, not_not, mem_compl_iff] at hx h'x ⊢, + exact continuous_at_iff_continuous_left'_right'.2 ⟨h'x, hx⟩ +end + +end monotone + +namespace antitone + +variables {α β : Type*} [linear_order α] [conditionally_complete_linear_order β] +[topological_space β] [order_topology β] {f : α → β} (hf : antitone f) {x y : α} +include hf + +lemma le_left_lim (h : x ≤ y) : f y ≤ left_lim f x := +hf.dual_right.left_lim_le h + +lemma left_lim_le (h : x < y) : left_lim f y ≤ f x := +hf.dual_right.le_left_lim h + +@[mono] protected lemma left_lim : antitone (left_lim f) := +hf.dual_right.left_lim + +lemma right_lim_le (h : x ≤ y) : right_lim f y ≤ f x := +hf.dual_right.le_right_lim h + +lemma le_right_lim (h : x < y) : f y ≤ right_lim f x := +hf.dual_right.right_lim_le h + +@[mono] protected lemma right_lim : antitone (right_lim f) := +hf.dual_right.right_lim + +lemma right_lim_le_left_lim (h : x ≤ y) : right_lim f y ≤ left_lim f x := +hf.dual_right.left_lim_le_right_lim h + +lemma left_lim_le_right_lim (h : x < y) : left_lim f y ≤ right_lim f x := +hf.dual_right.right_lim_le_left_lim h + +variables [topological_space α] [order_topology α] + +lemma tendsto_left_lim (x : α) : tendsto f (𝓝[<] x) (𝓝 (left_lim f x)) := +hf.dual_right.tendsto_left_lim x + +lemma tendsto_left_lim_within (x : α) : tendsto f (𝓝[<] x) (𝓝[≥] (left_lim f x)) := +hf.dual_right.tendsto_left_lim_within x + +lemma tendsto_right_lim (x : α) : + tendsto f (𝓝[>] x) (𝓝 (right_lim f x)) := +hf.dual_right.tendsto_right_lim x + +lemma tendsto_right_lim_within (x : α) : + tendsto f (𝓝[>] x) (𝓝[≤] (right_lim f x)) := +hf.dual_right.tendsto_right_lim_within x + +/-- An antitone function is continuous to the left at a point if and only if its left limit +coincides with the value of the function. -/ +lemma continuous_within_at_Iio_iff_left_lim_eq : + continuous_within_at f (Iio x) x ↔ left_lim f x = f x := +hf.dual_right.continuous_within_at_Iio_iff_left_lim_eq + +/-- An antitone function is continuous to the right at a point if and only if its right limit +coincides with the value of the function. -/ +lemma continuous_within_at_Ioi_iff_right_lim_eq : + continuous_within_at f (Ioi x) x ↔ right_lim f x = f x := +hf.dual_right.continuous_within_at_Ioi_iff_right_lim_eq + +/-- An antitone function is continuous at a point if and only if its left and right limits +coincide. -/ +lemma continuous_at_iff_left_lim_eq_right_lim : + continuous_at f x ↔ left_lim f x = right_lim f x := +hf.dual_right.continuous_at_iff_left_lim_eq_right_lim + +/-- In a second countable space, the set of points where an antitone function is not continuous +is at most countable. -/ +lemma countable_not_continuous_at [topological_space.second_countable_topology β] : + set.countable {x | ¬(continuous_at f x)} := +hf.dual_right.countable_not_continuous_at + +end antitone From 5d25c8ab4001350d885c46e7cf56c546365e2768 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 8 Oct 2022 03:03:15 +0000 Subject: [PATCH 622/828] feat(order/atoms): Galois (co)insertions and (co)atoms (#16663) --- src/order/atoms.lean | 125 +++++++++++++++++++++++++++++++++++-------- 1 file changed, 102 insertions(+), 23 deletions(-) diff --git a/src/order/atoms.lean b/src/order/atoms.lean index 5acd4f34dacb0..8e4b4bfa05ad1 100644 --- a/src/order/atoms.lean +++ b/src/order/atoms.lean @@ -50,7 +50,7 @@ which are lattices with only two elements, and related ideas. -/ -variable {α : Type*} +variables {α β : Type*} section atoms @@ -565,40 +565,120 @@ is_simple_order_iff_is_coatom_bot.trans $ and_congr (not_congr subtype.mk_eq_mk) end set +namespace order_embedding + +variables [partial_order α] [partial_order β] + +lemma is_atom_of_map_bot_of_image [order_bot α] [order_bot β] (f : β ↪o α) (hbot : f ⊥ = ⊥) {b : β} + (hb : is_atom (f b)) : is_atom b := +by { simp only [←bot_covby_iff] at hb ⊢, exact covby.of_image f (hbot.symm ▸ hb) } + +lemma is_coatom_of_map_top_of_image [order_top α] [order_top β] (f : β ↪o α) (htop : f ⊤ = ⊤) + {b : β} (hb : is_coatom (f b)) : is_coatom b := +f.dual.is_atom_of_map_bot_of_image htop hb + +end order_embedding + +namespace galois_insertion + +variables [partial_order α] [partial_order β] + +lemma is_atom_of_u_bot [order_bot α] [order_bot β] {l : α → β} {u : β → α} + (gi : galois_insertion l u) (hbot : u ⊥ = ⊥) {b : β} (hb : is_atom (u b)) : is_atom b := +order_embedding.is_atom_of_map_bot_of_image + ⟨⟨u, gi.u_injective⟩, @galois_insertion.u_le_u_iff _ _ _ _ _ _ gi⟩ hbot hb + +lemma is_atom_iff [order_bot α] [is_atomic α] [order_bot β] {l : α → β} {u : β → α} + (gi : galois_insertion l u) (hbot : u ⊥ = ⊥) (h_atom : ∀ a, is_atom a → u (l a) = a) (a : α) : + is_atom (l a) ↔ is_atom a := +begin + refine ⟨λ hla, _, λ ha, gi.is_atom_of_u_bot hbot ((h_atom a ha).symm ▸ ha)⟩, + obtain ⟨a', ha', hab'⟩ := (eq_bot_or_exists_atom_le (u (l a))).resolve_left + (hbot ▸ λ h, hla.1 (gi.u_injective h)), + have := (hla.le_iff.mp $ (gi.l_u_eq (l a) ▸ gi.gc.monotone_l hab' : l a' ≤ l a)).resolve_left + (λ h, ha'.1 (hbot ▸ (h_atom a' ha') ▸ congr_arg u h)), + have haa' : a = a' := (ha'.le_iff.mp $ + (gi.gc.le_u_l a).trans_eq (h_atom a' ha' ▸ congr_arg u this.symm)).resolve_left + (mt (congr_arg l) (gi.gc.l_bot.symm ▸ hla.1)), + exact haa'.symm ▸ ha' +end + +lemma is_atom_iff' [order_bot α] [is_atomic α] [order_bot β] {l : α → β} {u : β → α} + (gi : galois_insertion l u) (hbot : u ⊥ = ⊥) (h_atom : ∀ a, is_atom a → u (l a) = a) (b : β) : + is_atom (u b) ↔ is_atom b := +by rw [←gi.is_atom_iff hbot h_atom, gi.l_u_eq] + +lemma is_coatom_of_image [order_top α] [order_top β] {l : α → β} {u : β → α} + (gi : galois_insertion l u) {b : β} (hb : is_coatom (u b)) : is_coatom b := +order_embedding.is_coatom_of_map_top_of_image + ⟨⟨u, gi.u_injective⟩, @galois_insertion.u_le_u_iff _ _ _ _ _ _ gi⟩ gi.gc.u_top hb + +lemma is_coatom_iff [order_top α] [is_coatomic α] [order_top β] {l : α → β} {u : β → α} + (gi : galois_insertion l u) (h_coatom : ∀ a : α, is_coatom a → u (l a) = a) (b : β) : + is_coatom (u b) ↔ is_coatom b := +begin + refine ⟨λ hb, gi.is_coatom_of_image hb, λ hb, _⟩, + obtain ⟨a, ha, hab⟩ := (eq_top_or_exists_le_coatom (u b)).resolve_left + (λ h, hb.1 $ (gi.gc.u_top ▸ gi.l_u_eq ⊤ : l ⊤ = ⊤) ▸ gi.l_u_eq b ▸ congr_arg l h), + have : l a = b := (hb.le_iff.mp ((gi.l_u_eq b ▸ gi.gc.monotone_l hab) : b ≤ l a)).resolve_left + (λ hla, ha.1 (gi.gc.u_top ▸ h_coatom a ha ▸ congr_arg u hla)), + exact this ▸ (h_coatom a ha).symm ▸ ha, +end + +end galois_insertion + +namespace galois_coinsertion + +variables [partial_order α] [partial_order β] + +lemma is_coatom_of_l_top [order_top α] [order_top β] {l : α → β} {u : β → α} + (gi : galois_coinsertion l u) (hbot : l ⊤ = ⊤) {a : α} (hb : is_coatom (l a)) : is_coatom a := +gi.dual.is_atom_of_u_bot hbot hb.dual + +lemma is_coatom_iff [order_top α] [order_top β] [is_coatomic β] {l : α → β} {u : β → α} + (gi : galois_coinsertion l u) (htop : l ⊤ = ⊤) (h_coatom : ∀ b, is_coatom b → l (u b) = b) + (b : β) : is_coatom (u b) ↔ is_coatom b := +gi.dual.is_atom_iff htop h_coatom b + +lemma is_coatom_iff' [order_top α] [order_top β] [is_coatomic β] {l : α → β} {u : β → α} + (gi : galois_coinsertion l u) (htop : l ⊤ = ⊤) (h_coatom : ∀ b, is_coatom b → l (u b) = b) + (a : α) : is_coatom (l a) ↔ is_coatom a := +gi.dual.is_atom_iff' htop h_coatom a + +lemma is_atom_of_image [order_bot α] [order_bot β] {l : α → β} {u : β → α} + (gi : galois_coinsertion l u) {a : α} (hb : is_atom (l a)) : is_atom a := +gi.dual.is_coatom_of_image hb.dual + +lemma is_atom_iff [order_bot α] [order_bot β] [is_atomic β] {l : α → β} {u : β → α} + (gi : galois_coinsertion l u) (h_atom : ∀ b, is_atom b → l (u b) = b) (a : α) : + is_atom (l a) ↔ is_atom a := +gi.dual.is_coatom_iff h_atom a + +end galois_coinsertion + namespace order_iso -variables {β : Type*} +variables [partial_order α] [partial_order β] -@[simp] lemma is_atom_iff [partial_order α] [order_bot α] [partial_order β] [order_bot β] - (f : α ≃o β) (a : α) : +@[simp] lemma is_atom_iff [order_bot α] [order_bot β] (f : α ≃o β) (a : α) : is_atom (f a) ↔ is_atom a := -and_congr (not_congr ⟨λ h, f.injective (f.map_bot.symm ▸ h), λ h, f.map_bot ▸ (congr rfl h)⟩) - ⟨λ h b hb, f.injective ((h (f b) ((f : α ↪o β).lt_iff_lt.2 hb)).trans f.map_bot.symm), - λ h b hb, f.symm.injective begin - rw f.symm.map_bot, - apply h, - rw [← f.symm_apply_apply a], - exact (f.symm : β ↪o α).lt_iff_lt.2 hb, - end⟩ - -@[simp] lemma is_coatom_iff [partial_order α] [order_top α] [partial_order β] [order_top β] - (f : α ≃o β) (a : α) : +⟨f.to_galois_coinsertion.is_atom_of_image, + λ ha, f.to_galois_insertion.is_atom_of_u_bot (map_bot f.symm) $ (f.symm_apply_apply a).symm ▸ ha⟩ + +@[simp] lemma is_coatom_iff [order_top α] [order_top β] (f : α ≃o β) (a : α) : is_coatom (f a) ↔ is_coatom a := f.dual.is_atom_iff a -lemma is_simple_order_iff [partial_order α] [bounded_order α] [partial_order β] [bounded_order β] - (f : α ≃o β) : +lemma is_simple_order_iff [bounded_order α] [bounded_order β] (f : α ≃o β) : is_simple_order α ↔ is_simple_order β := by rw [is_simple_order_iff_is_atom_top, is_simple_order_iff_is_atom_top, ← f.is_atom_iff ⊤, f.map_top] -lemma is_simple_order [partial_order α] [bounded_order α] [partial_order β] [bounded_order β] - [h : is_simple_order β] (f : α ≃o β) : +lemma is_simple_order [bounded_order α] [bounded_order β] [h : is_simple_order β] (f : α ≃o β) : is_simple_order α := f.is_simple_order_iff.mpr h -lemma is_atomic_iff [partial_order α] [order_bot α] [partial_order β] [order_bot β] (f : α ≃o β) : - is_atomic α ↔ is_atomic β := +lemma is_atomic_iff [order_bot α] [order_bot β] (f : α ≃o β) : is_atomic α ↔ is_atomic β := begin suffices : (∀ b : α, b = ⊥ ∨ ∃ (a : α), is_atom a ∧ a ≤ b) ↔ (∀ b : β, b = ⊥ ∨ ∃ (a : β), is_atom a ∧ a ≤ b), @@ -614,8 +694,7 @@ begin rwa [←f.le_iff_le, f.apply_symm_apply], }, }, end -lemma is_coatomic_iff [partial_order α] [order_top α] [partial_order β] [order_top β] (f : α ≃o β) : - is_coatomic α ↔ is_coatomic β := +lemma is_coatomic_iff [order_top α] [order_top β] (f : α ≃o β) : is_coatomic α ↔ is_coatomic β := by { rw [←is_atomic_dual_iff_is_coatomic, ←is_atomic_dual_iff_is_coatomic], exact f.dual.is_atomic_iff } From ffd1e5cb54129bb18ebbe0f6778db0f6fdc87913 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 8 Oct 2022 03:03:16 +0000 Subject: [PATCH 623/828] feat(group_theory/quotient_group): congr for lifting an isomorphism to group quotients (#16726) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If `G` and `H` are isomorphic, and the isomorphism maps the subgroup `G'` to `H'`, then the quotients `G ⧸ G'` and `H ⧸ H'` are also isomorphic. Useful for showing the class group doesn't depend on a choice of field of fractions. --- src/group_theory/quotient_group.lean | 42 ++++++++++++++++++++++++++++ src/group_theory/subgroup/basic.lean | 11 ++++++++ 2 files changed, 53 insertions(+) diff --git a/src/group_theory/quotient_group.lean b/src/group_theory/quotient_group.lean index b45d3a0feec7e..43af3522e1aed 100644 --- a/src/group_theory/quotient_group.lean +++ b/src/group_theory/quotient_group.lean @@ -220,6 +220,48 @@ lemma map_comp_map {I : Type*} [group I] (M : subgroup H) (O : subgroup I) monoid_hom.ext (map_map N M O f g hf hg hgf) omit nN + +section congr + +variables (G' : subgroup G) (H' : subgroup H) [subgroup.normal G'] [subgroup.normal H'] + +/-- `quotient_group.congr` lifts the isomorphism `e : G ≃ H` to `G ⧸ G' ≃ H ⧸ H'`, +given that `e` maps `G` to `H`. -/ +@[to_additive "`quotient_add_group.congr` lifts the isomorphism `e : G ≃ H` to `G ⧸ G' ≃ H ⧸ H'`, +given that `e` maps `G` to `H`."] +def congr (e : G ≃* H) (he : G'.map ↑e = H') : G ⧸ G' ≃* H ⧸ H' := +{ to_fun := map G' H' ↑e (he ▸ G'.le_comap_map e), + inv_fun := map H' G' ↑e.symm (he ▸ (G'.map_equiv_eq_comap_symm e).le), + left_inv := λ x, by rw map_map; -- `simp` doesn't like this lemma... + simp only [map_map, ← mul_equiv.coe_monoid_hom_trans, mul_equiv.self_trans_symm, + mul_equiv.coe_monoid_hom_refl, map_id_apply], + right_inv := λ x, by rw map_map; -- `simp` doesn't like this lemma... + simp only [← mul_equiv.coe_monoid_hom_trans, mul_equiv.symm_trans_self, + mul_equiv.coe_monoid_hom_refl, map_id_apply], + .. map G' H' ↑e (he ▸ G'.le_comap_map e) } + +@[simp] lemma congr_mk (e : G ≃* H) (he : G'.map ↑e = H') + (x) : congr G' H' e he (mk x) = e x := +map_mk' G' _ _ (he ▸ G'.le_comap_map e) _ + +lemma congr_mk' (e : G ≃* H) (he : G'.map ↑e = H') + (x) : congr G' H' e he (mk' G' x) = mk' H' (e x) := +map_mk' G' _ _ (he ▸ G'.le_comap_map e) _ + +@[simp] lemma congr_apply (e : G ≃* H) (he : G'.map ↑e = H') + (x : G) : congr G' H' e he x = mk' H' (e x) := +map_mk' G' _ _ (he ▸ G'.le_comap_map e) _ + +@[simp] lemma congr_refl (he : G'.map (mul_equiv.refl G : G →* G) = G' := subgroup.map_id G') : + congr G' G' (mul_equiv.refl G) he = mul_equiv.refl (G ⧸ G') := +by ext x; refine induction_on' x (λ x', _); simp + +@[simp] lemma congr_symm (e : G ≃* H) (he : G'.map ↑e = H') : + (congr G' H' e he).symm = congr H' G' e.symm ((subgroup.map_symm_eq_iff_map_eq _).mpr he) := +rfl + +end congr + variables (φ : G →* H) open function monoid_hom diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index a47d863fe907a..dcbc5dc1bb42a 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -1200,6 +1200,17 @@ lemma comap_equiv_eq_map_symm (f : N ≃* G) (K : subgroup G) : K.comap f.to_monoid_hom = K.map f.symm.to_monoid_hom := (map_equiv_eq_comap_symm f.symm K).symm +@[to_additive] +lemma map_symm_eq_iff_map_eq {H : subgroup N} {e : G ≃* N} : + H.map ↑e.symm = K ↔ K.map ↑e = H := +begin + split; rintro rfl, + { rw [map_map, ← mul_equiv.coe_monoid_hom_trans, mul_equiv.symm_trans_self, + mul_equiv.coe_monoid_hom_refl, map_id] }, + { rw [map_map, ← mul_equiv.coe_monoid_hom_trans, mul_equiv.self_trans_symm, + mul_equiv.coe_monoid_hom_refl, map_id] }, +end + @[to_additive] lemma map_le_iff_le_comap {f : G →* N} {K : subgroup G} {H : subgroup N} : K.map f ≤ H ↔ K ≤ H.comap f := From 85453a2a14be8da64caf15ca50930cf4c6e5d8de Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Sat, 8 Oct 2022 03:03:17 +0000 Subject: [PATCH 624/828] =?UTF-8?q?feat(probability/variance):=20introduce?= =?UTF-8?q?=20`=E2=84=9D=E2=89=A50=E2=88=9E`-valued=20variance=20(#16730)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces `ℝ≥0∞`-valued variance and define the real-valued variance in terms of it. --- src/analysis/normed/field/basic.lean | 8 + src/analysis/normed/group/basic.lean | 7 + src/data/real/nnreal.lean | 3 + src/probability/ident_distrib.lean | 10 +- src/probability/moments.lean | 4 +- src/probability/strong_law.lean | 2 +- src/probability/variance.lean | 310 +++++++++++++++++++++------ 7 files changed, 269 insertions(+), 75 deletions(-) diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 8ae17f85352d8..fd035c2cbe88e 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -660,6 +660,14 @@ begin exact bot_le } end +lemma to_nnreal_mul_nnnorm {x : ℝ} (y : ℝ) (hr : 0 ≤ x) : x.to_nnreal * ∥y∥₊ = ∥x * y∥₊ := +begin + rw real.to_nnreal_of_nonneg hr, + simp only [nnnorm_mul, mul_eq_mul_right_iff], + refine or.inl (nnreal.eq _), + simp only [subtype.coe_mk, coe_nnnorm, real.norm_eq_abs, abs_of_nonneg hr] +end + /-- If `E` is a nontrivial topological module over `ℝ`, then `E` has no isolated points. This is a particular case of `module.punctured_nhds_ne_bot`. -/ instance punctured_nhds_module_ne_bot diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 50a2579e204c1..2ec4b28c52b73 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -1153,6 +1153,13 @@ lemma nnnorm_prod_le_of_le (s : finset ι) {f : ι → E} {n : ι → ℝ≥0} ( ∥∏ b in s, f b∥₊ ≤ ∑ b in s, n b := (norm_prod_le_of_le s h).trans_eq nnreal.coe_sum.symm +lemma real.to_nnreal_eq_nnnorm_of_nonneg {r : ℝ} (hr : 0 ≤ r) : r.to_nnreal = ∥r∥₊ := +begin + rw real.to_nnreal_of_nonneg hr, + congr, + rw [real.norm_eq_abs, abs_of_nonneg hr], +end + namespace lipschitz_with variables [pseudo_emetric_space α] {K Kf Kg : ℝ≥0} {f g : α → E} diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index f94cdbe95aeb7..a8e2aa4935eef 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -89,6 +89,9 @@ noncomputable def _root_.real.to_nnreal (r : ℝ) : ℝ≥0 := ⟨max r 0, le_ma lemma _root_.real.coe_to_nnreal (r : ℝ) (hr : 0 ≤ r) : (real.to_nnreal r : ℝ) = r := max_eq_left hr +lemma _root_.real.to_nnreal_of_nonneg {r : ℝ} (hr : 0 ≤ r) : r.to_nnreal = ⟨r, hr⟩ := +by simp_rw [real.to_nnreal, max_eq_left hr] + lemma _root_.real.le_coe_to_nnreal (r : ℝ) : r ≤ real.to_nnreal r := le_max_left r 0 diff --git a/src/probability/ident_distrib.lean b/src/probability/ident_distrib.lean index e397568c6c6f4..89b563ddf294d 100644 --- a/src/probability/ident_distrib.lean +++ b/src/probability/ident_distrib.lean @@ -280,14 +280,18 @@ lemma const_div [has_div γ] [has_measurable_div γ] (h : ident_distrib f g μ ident_distrib (λ x, c / f x) (λ x, c / g x) μ ν := h.comp (has_measurable_div.measurable_const_div c) -lemma variance_eq {f : α → ℝ} {g : β → ℝ} (h : ident_distrib f g μ ν) : - variance f μ = variance g ν := +lemma evariance_eq {f : α → ℝ} {g : β → ℝ} (h : ident_distrib f g μ ν) : + evariance f μ = evariance g ν := begin - convert (h.sub_const (∫ x, f x ∂μ)).sq.integral_eq, + convert (h.sub_const (∫ x, f x ∂μ)).nnnorm.coe_nnreal_ennreal.sq.lintegral_eq, rw h.integral_eq, refl end +lemma variance_eq {f : α → ℝ} {g : β → ℝ} (h : ident_distrib f g μ ν) : + variance f μ = variance g ν := +by { rw [variance, h.evariance_eq], refl, } + end ident_distrib section uniform_integrable diff --git a/src/probability/moments.lean b/src/probability/moments.lean index c1ef19a82846c..392a55356ec10 100644 --- a/src/probability/moments.lean +++ b/src/probability/moments.lean @@ -84,7 +84,9 @@ begin rw integral_undef this, }, end -@[simp] lemma central_moment_two_eq_variance : central_moment X 2 μ = variance X μ := rfl +lemma central_moment_two_eq_variance [is_finite_measure μ] (hX : mem_ℒp X 2 μ) : + central_moment X 2 μ = variance X μ := +by { rw hX.variance_eq, refl, } section moment_generating_function diff --git a/src/probability/strong_law.lean b/src/probability/strong_law.lean index ecdad27ce878c..9b96133760e04 100644 --- a/src/probability/strong_law.lean +++ b/src/probability/strong_law.lean @@ -481,7 +481,7 @@ begin apply sum_le_sum (λ j hj, _), refine mul_le_mul_of_nonneg_left _ (inv_nonneg.2 (sq_nonneg _)), rw (hident j).truncation.variance_eq, - exact variance_le_expectation_sq, + exact variance_le_expectation_sq (hX 0).truncation, end ... ≤ 2 * 𝔼[X 0] : sum_variance_truncation_le hint (hnonneg 0) K }, let C := (c ^ 5 * (c - 1) ⁻¹ ^ 3) * (2 * 𝔼[X 0]), diff --git a/src/probability/variance.lean b/src/probability/variance.lean index cb59ffa0df42e..189e780273c24 100644 --- a/src/probability/variance.lean +++ b/src/probability/variance.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Kexing Ying -/ import probability.notation import probability.integration @@ -12,14 +12,23 @@ import probability.integration We define the variance of a real-valued random variable as `Var[X] = 𝔼[(X - 𝔼[X])^2]` (in the `probability_theory` locale). -We prove the basic properties of the variance: -* `variance_le_expectation_sq`: the inequality `Var[X] ≤ 𝔼[X^2]`. -* `meas_ge_le_variance_div_sq`: Chebyshev's inequality, i.e., +## Main definitions + +* `probability_theory.evariance`: the variance of a real-valued random variable as a extended + non-negative real. +* `probability_theory.variance`: the variance of a real-valued random variable as a real number. + +## Main results + +* `probability_theory.variance_le_expectation_sq`: the inequality `Var[X] ≤ 𝔼[X^2]`. +* `probability_theory.meas_ge_le_variance_div_sq`: Chebyshev's inequality, i.e., `ℙ {ω | c ≤ |X ω - 𝔼[X]|} ≤ ennreal.of_real (Var[X] / c ^ 2)`. -* `indep_fun.variance_add`: the variance of the sum of two independent random variables is the sum - of the variances. -* `indep_fun.variance_sum`: the variance of a finite sum of pairwise independent random variables is - the sum of the variances. +* `probability_theory.meas_ge_le_evariance_div_sq`: Chebyshev's inequality formulated with + `evariance` without requiring the random variables to be L². +* `probability_theory.indep_fun.variance_add`: the variance of the sum of two independent + random variables is the sum of the variances. +* `probability_theory.indep_fun.variance_sum`: the variance of a finite sum of pairwise + independent random variables is the sum of the variances. -/ open measure_theory filter finset @@ -30,38 +39,165 @@ open_locale big_operators measure_theory probability_theory ennreal nnreal namespace probability_theory -/-- The variance of a random variable is `𝔼[X^2] - 𝔼[X]^2` or, equivalently, `𝔼[(X - 𝔼[X])^2]`. We -use the latter as the definition, to ensure better behavior even in garbage situations. -/ -def variance {Ω : Type*} {m : measurable_space Ω} (f : Ω → ℝ) (μ : measure Ω) : ℝ := -μ[(f - (λ ω, μ[f])) ^ 2] +/-- The `ℝ≥0∞`-valued variance of a real-valued random variable defined as the Lebesgue integral of +`(X - 𝔼[X])^2`. -/ +def evariance {Ω : Type*} {m : measurable_space Ω} (X : Ω → ℝ) (μ : measure Ω) : ℝ≥0∞ := +∫⁻ ω, ∥X ω - μ[X]∥₊^2 ∂μ -@[simp] lemma variance_zero {Ω : Type*} {m : measurable_space Ω} (μ : measure Ω) : - variance 0 μ = 0 := -by simp [variance] +/-- The `ℝ`-valued variance of a real-valued random variable defined by applying `ennreal.to_real` +to `evariance`. -/ +def variance {Ω : Type*} {m : measurable_space Ω} (X : Ω → ℝ) (μ : measure Ω) : ℝ := +(evariance X μ).to_real -lemma variance_nonneg {Ω : Type*} {m : measurable_space Ω} (f : Ω → ℝ) (μ : measure Ω) : - 0 ≤ variance f μ := -integral_nonneg (λ ω, sq_nonneg _) +variables {Ω : Type*} {m : measurable_space Ω} {X : Ω → ℝ} {μ : measure Ω} -lemma variance_mul {Ω : Type*} {m : measurable_space Ω} (c : ℝ) (f : Ω → ℝ) (μ : measure Ω) : - variance (λ ω, c * f ω) μ = c^2 * variance f μ := -calc -variance (λ ω, c * f ω) μ - = ∫ x, (c * f x - ∫ y, c * f y ∂μ) ^ 2 ∂μ : rfl -... = ∫ x, (c * (f x - ∫ y, f y ∂μ)) ^ 2 ∂μ : - by { congr' 1 with x, simp_rw [integral_mul_left, mul_sub] } -... = c^2 * variance f μ : - by { simp_rw [mul_pow, integral_mul_left], refl } +lemma _root_.measure_theory.mem_ℒp.evariance_lt_top [is_finite_measure μ] (hX : mem_ℒp X 2 μ) : + evariance X μ < ∞ := +begin + have := ennreal.pow_lt_top (hX.sub $ mem_ℒp_const $ μ[X]).2 2, + rw [snorm_eq_lintegral_rpow_nnnorm ennreal.two_ne_zero ennreal.two_ne_top, + ← ennreal.rpow_two] at this, + simp only [pi.sub_apply, ennreal.to_real_bit0, ennreal.one_to_real, one_div] at this, + rw [← ennreal.rpow_mul, inv_mul_cancel (two_ne_zero : (2 : ℝ) ≠ 0), ennreal.rpow_one] at this, + simp_rw ennreal.rpow_two at this, + exact this, +end + +lemma evariance_eq_top [is_finite_measure μ] + (hXm : ae_strongly_measurable X μ) (hX : ¬ mem_ℒp X 2 μ) : + evariance X μ = ∞ := +begin + by_contra h, + rw [← ne.def, ← lt_top_iff_ne_top] at h, + have : mem_ℒp (λ ω, X ω - μ[X]) 2 μ, + { refine ⟨hXm.sub ae_strongly_measurable_const, _⟩, + rw snorm_eq_lintegral_rpow_nnnorm ennreal.two_ne_zero ennreal.two_ne_top, + simp only [ennreal.to_real_bit0, ennreal.one_to_real, ennreal.rpow_two, ne.def], + exact ennreal.rpow_lt_top_of_nonneg (by simp) h.ne }, + refine hX _, + convert this.add (mem_ℒp_const $ μ[X]), + ext ω, + rw [pi.add_apply, sub_add_cancel], +end + +lemma evariance_lt_top_iff_mem_ℒp [is_finite_measure μ] + (hX : ae_strongly_measurable X μ) : + evariance X μ < ∞ ↔ mem_ℒp X 2 μ := +begin + refine ⟨_, measure_theory.mem_ℒp.evariance_lt_top⟩, + contrapose, + rw [not_lt, top_le_iff], + exact evariance_eq_top hX +end + +lemma _root_.measure_theory.mem_ℒp.of_real_variance_eq [is_finite_measure μ] + (hX : mem_ℒp X 2 μ) : + ennreal.of_real (variance X μ) = evariance X μ := +by { rw [variance, ennreal.of_real_to_real], exact hX.evariance_lt_top.ne, } + +include m + +lemma evariance_eq_lintegral_of_real (X : Ω → ℝ) (μ : measure Ω) : + evariance X μ = ∫⁻ ω, ennreal.of_real ((X ω - μ[X])^2) ∂μ := +begin + rw evariance, + congr, + ext1 ω, + rw [pow_two, ← ennreal.coe_mul, ← nnnorm_mul, ← pow_two], + congr, + exact (real.to_nnreal_eq_nnnorm_of_nonneg $ sq_nonneg _).symm, +end + +lemma _root_.measure_theory.mem_ℒp.variance_eq_of_integral_eq_zero + (hX : mem_ℒp X 2 μ) (hXint : μ[X] = 0) : + variance X μ = μ[X^2] := +begin + rw [variance, evariance_eq_lintegral_of_real, ← of_real_integral_eq_lintegral_of_real, + ennreal.to_real_of_real]; + simp_rw [hXint, sub_zero], + { refl }, + { exact integral_nonneg (λ ω, pow_two_nonneg _) }, + { convert hX.integrable_norm_rpow ennreal.two_ne_zero ennreal.two_ne_top, + ext ω, + simp only [pi.sub_apply, real.norm_eq_abs, ennreal.to_real_bit0, ennreal.one_to_real, + real.rpow_two, pow_bit0_abs] }, + { exact ae_of_all _ (λ ω, pow_two_nonneg _) } +end + +lemma _root_.measure_theory.mem_ℒp.variance_eq [is_finite_measure μ] + (hX : mem_ℒp X 2 μ) : + variance X μ = μ[(X - (λ ω, μ[X]))^2] := +begin + rw [variance, evariance_eq_lintegral_of_real, ← of_real_integral_eq_lintegral_of_real, + ennreal.to_real_of_real], + { refl }, + { exact integral_nonneg (λ ω, pow_two_nonneg _) }, + { convert (hX.sub $ mem_ℒp_const (μ[X])).integrable_norm_rpow + ennreal.two_ne_zero ennreal.two_ne_top, + ext ω, + simp only [pi.sub_apply, real.norm_eq_abs, ennreal.to_real_bit0, ennreal.one_to_real, + real.rpow_two, pow_bit0_abs] }, + { exact ae_of_all _ (λ ω, pow_two_nonneg _) } +end + +@[simp] lemma evariance_zero : evariance 0 μ = 0 := +by simp [evariance] + +lemma evariance_eq_zero_iff (hX : ae_measurable X μ) : + evariance X μ = 0 ↔ X =ᵐ[μ] λ ω, μ[X] := +begin + rw [evariance, lintegral_eq_zero_iff'], + split; intro hX; filter_upwards [hX] with ω hω, + { simp only [pi.zero_apply, pow_eq_zero_iff, nat.succ_pos', ennreal.coe_eq_zero, + nnnorm_eq_zero, sub_eq_zero] at hω, + exact hω }, + { rw hω, + simp }, + { measurability } +end + +lemma evariance_mul (c : ℝ) (X : Ω → ℝ) (μ : measure Ω) : + evariance (λ ω, c * X ω) μ = ennreal.of_real (c^2) * evariance X μ := +begin + rw [evariance, evariance, ← lintegral_const_mul' _ _ ennreal.of_real_lt_top.ne], + congr, + ext1 ω, + rw [ennreal.of_real, ← ennreal.coe_pow, ← ennreal.coe_pow, ← ennreal.coe_mul], + congr, + rw [← sq_abs, ← real.rpow_two, real.to_nnreal_rpow_of_nonneg (abs_nonneg _), nnreal.rpow_two, + ← mul_pow, real.to_nnreal_mul_nnnorm _ (abs_nonneg _)], + conv_rhs { rw [← nnnorm_norm, norm_mul, norm_abs_eq_norm, ← norm_mul, nnnorm_norm, mul_sub] }, + congr, + rw mul_comm, + simp_rw [← smul_eq_mul, ← integral_smul_const, smul_eq_mul, mul_comm], +end + +localized "notation (name := probability_theory.evariance) `eVar[` X `]` := + probability_theory.evariance X measure_theory.measure_space.volume" in probability_theory + +@[simp] lemma variance_zero (μ : measure Ω) : variance 0 μ = 0 := +by simp only [variance, evariance_zero, ennreal.zero_to_real] + +lemma variance_nonneg (X : Ω → ℝ) (μ : measure Ω) : + 0 ≤ variance X μ := +ennreal.to_real_nonneg + +lemma variance_mul (c : ℝ) (X : Ω → ℝ) (μ : measure Ω) : + variance (λ ω, c * X ω) μ = c^2 * variance X μ := +begin + rw [variance, evariance_mul, ennreal.to_real_mul, ennreal.to_real_of_real (sq_nonneg _)], + refl, +end -lemma variance_smul {Ω : Type*} {m : measurable_space Ω} (c : ℝ) (f : Ω → ℝ) (μ : measure Ω) : - variance (c • f) μ = c^2 * variance f μ := -variance_mul c f μ +lemma variance_smul (c : ℝ) (X : Ω → ℝ) (μ : measure Ω) : + variance (c • X) μ = c^2 * variance X μ := +variance_mul c X μ lemma variance_smul' {A : Type*} [comm_semiring A] [algebra A ℝ] - {Ω : Type*} {m : measurable_space Ω} (c : A) (f : Ω → ℝ) (μ : measure Ω) : - variance (c • f) μ = c^2 • variance f μ := + (c : A) (X : Ω → ℝ) (μ : measure Ω) : + variance (c • X) μ = c^2 • variance X μ := begin - convert variance_smul (algebra_map A ℝ c) f μ, + convert variance_smul (algebra_map A ℝ c) X μ, { ext1 x, simp only [algebra_map_smul], }, { simp only [algebra.smul_def, map_pow], } end @@ -69,12 +205,15 @@ end localized "notation (name := probability_theory.variance) `Var[` X `]` := probability_theory.variance X measure_theory.measure_space.volume" in probability_theory -variables {Ω : Type*} [measure_space Ω] [is_probability_measure (volume : measure Ω)] +omit m -lemma variance_def' {X : Ω → ℝ} (hX : mem_ℒp X 2) : +variables [measure_space Ω] + +lemma variance_def' [is_probability_measure (ℙ : measure Ω)] + {X : Ω → ℝ} (hX : mem_ℒp X 2) : Var[X] = 𝔼[X^2] - 𝔼[X]^2 := begin - rw [variance, sub_sq', integral_sub', integral_add'], rotate, + rw [hX.variance_eq, sub_sq', integral_sub', integral_add'], rotate, { exact hX.integrable_sq }, { convert integrable_const (𝔼[X] ^ 2), apply_instance }, @@ -88,58 +227,88 @@ begin ring, end -lemma variance_le_expectation_sq {X : Ω → ℝ} : +lemma variance_le_expectation_sq [is_probability_measure (ℙ : measure Ω)] + {X : Ω → ℝ} (hm : ae_strongly_measurable X ℙ) : Var[X] ≤ 𝔼[X^2] := begin - by_cases h_int : integrable X, swap, - { simp only [variance, integral_undef h_int, pi.pow_apply, pi.sub_apply, sub_zero] }, by_cases hX : mem_ℒp X 2, { rw variance_def' hX, simp only [sq_nonneg, sub_le_self_iff] }, - { rw [variance, integral_undef], + rw [variance, evariance_eq_lintegral_of_real, ← integral_eq_lintegral_of_nonneg_ae], + by_cases hint : integrable X, swap, + { simp only [integral_undef hint, pi.pow_apply, pi.sub_apply, sub_zero] }, + { rw integral_undef, { exact integral_nonneg (λ a, sq_nonneg _) }, - { assume h, + { intro h, have A : mem_ℒp (X - λ (ω : Ω), 𝔼[X]) 2 ℙ := (mem_ℒp_two_iff_integrable_sq - (h_int.ae_strongly_measurable.sub ae_strongly_measurable_const)).2 h, + (hint.ae_strongly_measurable.sub ae_strongly_measurable_const)).2 h, have B : mem_ℒp (λ (ω : Ω), 𝔼[X]) 2 ℙ := mem_ℒp_const _, apply hX, convert A.add B, - simp } } + simp } }, + { exact ae_of_all _ (λ x, sq_nonneg _) }, + { exact (ae_measurable.pow_const (hm.ae_measurable.sub_const _) _).ae_strongly_measurable }, end -/-- *Chebyshev's inequality* : one can control the deviation probability of a real random variable -from its expectation in terms of the variance. -/ -theorem meas_ge_le_variance_div_sq {X : Ω → ℝ} (hX : mem_ℒp X 2) {c : ℝ} (hc : 0 < c) : - ℙ {ω | c ≤ |X ω - 𝔼[X]|} ≤ ennreal.of_real (Var[X] / c ^ 2) := +lemma evariance_def' [is_probability_measure (ℙ : measure Ω)] + {X : Ω → ℝ} (hX : ae_strongly_measurable X ℙ) : + eVar[X] = (∫⁻ ω, ∥X ω∥₊^2) - ennreal.of_real (𝔼[X]^2) := +begin + by_cases hℒ : mem_ℒp X 2, + { rw [← hℒ.of_real_variance_eq, variance_def' hℒ, ennreal.of_real_sub _ (sq_nonneg _)], + congr, + simp_rw ← ennreal.coe_pow, + rw lintegral_coe_eq_integral, + { congr' 2 with ω, + simp only [pi.pow_apply, nnreal.coe_pow, coe_nnnorm, real.norm_eq_abs, pow_bit0_abs] }, + { exact hℒ.abs.integrable_sq } }, + { symmetry, + rw [evariance_eq_top hX hℒ, ennreal.sub_eq_top_iff], + refine ⟨_, ennreal.of_real_ne_top⟩, + rw [mem_ℒp, not_and] at hℒ, + specialize hℒ hX, + simp only [snorm_eq_lintegral_rpow_nnnorm ennreal.two_ne_zero ennreal.two_ne_top, not_lt, + top_le_iff, ennreal.to_real_bit0, ennreal.one_to_real, ennreal.rpow_two, one_div, + ennreal.rpow_eq_top_iff, inv_lt_zero, inv_pos, zero_lt_bit0, zero_lt_one, and_true, + or_iff_not_imp_left, not_and_distrib] at hℒ, + exact hℒ (λ _, zero_le_two) } +end + +/-- *Chebyshev's inequality* for `ℝ≥0∞`-valued variance. -/ +theorem meas_ge_le_evariance_div_sq {X : Ω → ℝ} + (hX : ae_strongly_measurable X ℙ) {c : ℝ≥0} (hc : c ≠ 0) : + ℙ {ω | ↑c ≤ |X ω - 𝔼[X]|} ≤ eVar[X] / c ^ 2 := begin - have A : (ennreal.of_real c : ℝ≥0∞) ≠ 0, - by simp only [hc, ne.def, ennreal.of_real_eq_zero, not_le], + have A : (c : ℝ≥0∞) ≠ 0, { rwa [ne.def, ennreal.coe_eq_zero] }, have B : ae_strongly_measurable (λ (ω : Ω), 𝔼[X]) ℙ := ae_strongly_measurable_const, - convert meas_ge_le_mul_pow_snorm ℙ ennreal.two_ne_zero ennreal.two_ne_top - (hX.ae_strongly_measurable.sub B) A, + convert meas_ge_le_mul_pow_snorm ℙ ennreal.two_ne_zero ennreal.two_ne_top (hX.sub B) A, { ext ω, - set d : ℝ≥0 := ⟨c, hc.le⟩ with hd, - have cd : c = d, by simp only [subtype.coe_mk], simp only [pi.sub_apply, ennreal.coe_le_coe, ← real.norm_eq_abs, ← coe_nnnorm, - nnreal.coe_le_coe, cd, ennreal.of_real_coe_nnreal] }, - { rw (hX.sub (mem_ℒp_const _)).snorm_eq_integral_rpow_norm - ennreal.two_ne_zero ennreal.two_ne_top, - simp only [pi.sub_apply, ennreal.to_real_bit0, ennreal.one_to_real], - rw ennreal.of_real_rpow_of_nonneg _ zero_le_two, rotate, - { apply real.rpow_nonneg_of_nonneg, - exact integral_nonneg (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _) }, - rw [variance, ← real.rpow_mul, inv_mul_cancel], rotate, - { exact two_ne_zero }, - { exact integral_nonneg (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _) }, - simp only [pi.pow_apply, pi.sub_apply, real.rpow_two, real.rpow_one, real.norm_eq_abs, - pow_bit0_abs, ennreal.of_real_inv_of_pos hc, ennreal.rpow_two], - rw [← ennreal.of_real_pow (inv_nonneg.2 hc.le), ← ennreal.of_real_mul (sq_nonneg _), - div_eq_inv_mul, inv_pow] } + nnreal.coe_le_coe, ennreal.of_real_coe_nnreal] }, + { rw snorm_eq_lintegral_rpow_nnnorm ennreal.two_ne_zero ennreal.two_ne_top, + simp only [ennreal.to_real_bit0, ennreal.one_to_real, pi.sub_apply, one_div], + rw [div_eq_mul_inv, ennreal.inv_pow, mul_comm, ennreal.rpow_two], + congr, + simp_rw [← ennreal.rpow_mul, inv_mul_cancel (two_ne_zero : (2 : ℝ) ≠ 0), ennreal.rpow_two, + ennreal.rpow_one, evariance] } +end + +/-- *Chebyshev's inequality* : one can control the deviation probability of a real random variable +from its expectation in terms of the variance. -/ +theorem meas_ge_le_variance_div_sq [is_finite_measure (ℙ : measure Ω)] + {X : Ω → ℝ} (hX : mem_ℒp X 2) {c : ℝ} (hc : 0 < c) : + ℙ {ω | c ≤ |X ω - 𝔼[X]|} ≤ ennreal.of_real (Var[X] / c ^ 2) := +begin + rw [ennreal.of_real_div_of_pos (sq_pos_of_ne_zero _ hc.ne.symm), hX.of_real_variance_eq], + convert @meas_ge_le_evariance_div_sq _ _ _ hX.1 (c.to_nnreal) (by simp [hc]), + { simp only [real.coe_to_nnreal', max_le_iff, abs_nonneg, and_true] }, + { rw ennreal.of_real_pow hc.le, + refl } end /-- The variance of the sum of two independent random variables is the sum of the variances. -/ -theorem indep_fun.variance_add {X Y : Ω → ℝ} - (hX : mem_ℒp X 2) (hY : mem_ℒp Y 2) (h : indep_fun X Y) : +theorem indep_fun.variance_add [is_probability_measure (ℙ : measure Ω)] + {X Y : Ω → ℝ} (hX : mem_ℒp X 2) (hY : mem_ℒp Y 2) (h : indep_fun X Y) : Var[X + Y] = Var[X] + Var[Y] := calc Var[X + Y] = 𝔼[λ a, (X a)^2 + (Y a)^2 + 2 * X a * Y a] - 𝔼[X+Y]^2 : @@ -167,7 +336,8 @@ end /-- The variance of a finite sum of pairwise independent random variables is the sum of the variances. -/ -theorem indep_fun.variance_sum {ι : Type*} {X : ι → Ω → ℝ} {s : finset ι} +theorem indep_fun.variance_sum [is_probability_measure (ℙ : measure Ω)] + {ι : Type*} {X : ι → Ω → ℝ} {s : finset ι} (hs : ∀ i ∈ s, mem_ℒp (X i) 2) (h : set.pairwise ↑s (λ i j, indep_fun (X i) (X j))) : Var[∑ i in s, X i] = ∑ i in s, Var[X i] := begin From 2a528d11813a4452202e79a3c2d82fd7db2bc6e8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 8 Oct 2022 03:03:18 +0000 Subject: [PATCH 625/828] feat(set_theory/cardinal/ordinal): Beth numbers form a normal family (#16853) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove `is_normal (λ o, (beth o).ord)` --- src/set_theory/cardinal/ordinal.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/set_theory/cardinal/ordinal.lean b/src/set_theory/cardinal/ordinal.lean index 792112ce3328a..03545ce188f25 100644 --- a/src/set_theory/cardinal/ordinal.lean +++ b/src/set_theory/cardinal/ordinal.lean @@ -364,6 +364,8 @@ begin exact le_csupr (bdd_above_of_small _) (⟨_, hb.succ_lt h⟩ : Iio b) } end +lemma beth_mono : monotone beth := beth_strict_mono.monotone + @[simp] theorem beth_lt {o₁ o₂ : ordinal} : beth o₁ < beth o₂ ↔ o₁ < o₂ := beth_strict_mono.lt_iff_lt @@ -391,6 +393,10 @@ aleph_0_pos.trans_le $ aleph_0_le_beth o theorem beth_ne_zero (o : ordinal) : beth o ≠ 0 := (beth_pos o).ne' +lemma beth_normal : is_normal.{u} (λ o, (beth o).ord) := +(is_normal_iff_strict_mono_limit _).2 ⟨ord_strict_mono.comp beth_strict_mono, λ o ho a ha, + by { rw [beth_limit ho, ord_le], exact csupr_le' (λ b, ord_le.1 (ha _ b.2)) }⟩ + /-! ### Properties of `mul` -/ /-- If `α` is an infinite type, then `α × α` and `α` have the same cardinality. -/ From 666d8787a472bc85b9b8883d8056846fbd560ef6 Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Sat, 8 Oct 2022 05:41:09 +0000 Subject: [PATCH 626/828] feat(algebra/module/projective): projective modules are closed under direct sums (#16743) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit I'm unsure what the universe should be for the index type `ι` in algebra/module/projective (`{ι : Type*}` doesn't work). - [x] depends on: #16735 Co-authored-by: antoinelab01 <66086247+antoinelab01@users.noreply.github.com> --- src/algebra/module/projective.lean | 37 +++++++++++++++++++++++++++++- src/linear_algebra/dfinsupp.lean | 19 +++++++++++++++ 2 files changed, 55 insertions(+), 1 deletion(-) diff --git a/src/algebra/module/projective.lean b/src/algebra/module/projective.lean index ece899c19dd22..f85a04b29f43e 100644 --- a/src/algebra/module/projective.lean +++ b/src/algebra/module/projective.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kevin Buzzard +Authors: Kevin Buzzard, Antoine Labelle -/ import algebra.module.basic @@ -130,6 +130,41 @@ begin { rw [←snd_apply _, apply_total R], exact hsQ x, }, end +variables {ι : Type*} (A : ι → Type*) [Π (i : ι), add_comm_monoid (A i)] + [Π (i : ι), module R (A i)] + +instance [h : Π (i : ι), projective R (A i)] : projective R (Π₀ i, A i) := +begin + classical, + rw module.projective_def', + simp_rw projective_def at h, choose s hs using h, + + letI : Π (i : ι), add_comm_monoid (A i →₀ R) := λ i, by apply_instance, + letI : Π (i : ι), module R (A i →₀ R) := λ i, by apply_instance, + letI : add_comm_monoid (Π₀ (i : ι), A i →₀ R) := @dfinsupp.add_comm_monoid ι (λ i, A i →₀ R) _, + letI : module R (Π₀ (i : ι), A i →₀ R) := @dfinsupp.module ι R (λ i, A i →₀ R) _ _ _, + + let f := λ i, lmap_domain R R (dfinsupp.single i : A i → Π₀ i, A i), + use dfinsupp.coprod_map f ∘ₗ dfinsupp.map_range.linear_map s, + + ext i x j, + simp only [dfinsupp.coprod_map, direct_sum.lof, total_map_domain R _ dfinsupp.single_injective, + coe_comp, coe_lsum, id_coe, linear_equiv.coe_to_linear_map, finsupp_lequiv_dfinsupp_symm_apply, + function.comp_app, dfinsupp.lsingle_apply, dfinsupp.map_range.linear_map_apply, + dfinsupp.map_range_single, lmap_domain_apply, dfinsupp.to_finsupp_single, + finsupp.sum_single_index, id.def, function.comp.left_id, dfinsupp.single_apply], + rw [←dfinsupp.lapply_apply j, apply_total R], + + obtain rfl | hij := eq_or_ne i j, + + { convert (hs i) x, + { ext, simp }, + { simp } }, + { convert finsupp.total_zero_apply _ ((s i) x), + { ext, simp [hij] }, + { simp [hij] } } +end + end semiring section ring diff --git a/src/linear_algebra/dfinsupp.lean b/src/linear_algebra/dfinsupp.lean index 225d7c9c70756..2aaa4f83e9b08 100644 --- a/src/linear_algebra/dfinsupp.lean +++ b/src/linear_algebra/dfinsupp.lean @@ -206,6 +206,25 @@ lemma map_range.linear_equiv_symm (e : Π i, β₁ i ≃ₗ[R] β₂ i) : end map_range +section coprod_map + +variables [decidable_eq ι] [Π (x : N), decidable (x ≠ 0)] + +/-- Given a family of linear maps `f i : M i →ₗ[R] N`, we can form a linear map +`(Π₀ i, M i) →ₗ[R] N` which sends `x : Π₀ i, M i` to the sum over `i` of `f i` applied to `x i`. +This is the map coming from the universal property of `Π₀ i, M i` as the coproduct of the `M i`. +See also `linear_map.coprod` for the binary product version. -/ +noncomputable def coprod_map (f : Π (i : ι), M i →ₗ[R] N) : (Π₀ i, M i) →ₗ[R] N := +finsupp.lsum ℕ (λ i : ι, linear_map.id) ∘ₗ +(@finsupp_lequiv_dfinsupp ι R N _ _ _ _ _).symm.to_linear_map ∘ₗ +(dfinsupp.map_range.linear_map f) + +lemma coprod_map_apply (f : Π (i : ι), M i →ₗ[R] N) (x : Π₀ i, M i) : + coprod_map f x = + finsupp.sum (map_range (λ i, f i) (λ i, linear_map.map_zero _) x).to_finsupp (λ i, id) := rfl + +end coprod_map + section basis /-- The direct sum of free modules is free. From f585481bced7e7cc7ce108a9ed2c723ec398db69 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 8 Oct 2022 05:41:10 +0000 Subject: [PATCH 627/828] feat(data/set/prod): add lemmas about `set.pi` (#16828) --- src/data/set/basic.lean | 5 +++++ src/data/set/lattice.lean | 5 ----- src/data/set/prod.lean | 25 ++++++++++++++++++------- 3 files changed, 23 insertions(+), 12 deletions(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 29ff5f21a2380..a42e241734b0b 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -888,6 +888,11 @@ ssubset_singleton_iff.1 hs /-! ### Disjointness -/ +protected theorem disjoint_iff : disjoint s t ↔ s ∩ t ⊆ ∅ := iff.rfl + +theorem disjoint_iff_inter_eq_empty : disjoint s t ↔ s ∩ t = ∅ := +disjoint_iff + lemma _root_.disjoint.inter_eq : disjoint s t → s ∩ t = ∅ := disjoint.eq_bot lemma disjoint_left : disjoint s t ↔ ∀ ⦃a⦄, a ∈ s → a ∉ t := forall_congr $ λ _, not_and diff --git a/src/data/set/lattice.lean b/src/data/set/lattice.lean index dd78286c6fb12..60165659b97db 100644 --- a/src/data/set/lattice.lean +++ b/src/data/set/lattice.lean @@ -1579,11 +1579,6 @@ end disjoint namespace set -protected theorem disjoint_iff : disjoint s t ↔ s ∩ t ⊆ ∅ := iff.rfl - -theorem disjoint_iff_inter_eq_empty : disjoint s t ↔ s ∩ t = ∅ := -disjoint_iff - lemma not_disjoint_iff : ¬disjoint s t ↔ ∃ x, x ∈ s ∧ x ∈ t := not_forall.trans $ exists_congr $ λ x, not_not diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index 6e57bbce93d5a..ca10b1cf8880b 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -430,20 +430,19 @@ lemma pi_eq_empty_iff : s.pi t = ∅ ↔ ∃ i, is_empty (α i) ∨ i ∈ s ∧ begin rw [← not_nonempty_iff_eq_empty, pi_nonempty_iff], push_neg, - refine exists_congr (λ i, ⟨λ h, (is_empty_or_nonempty (α i)).imp_right _, _⟩), - { rintro ⟨x⟩, - exact ⟨(h x).1, by simp [eq_empty_iff_forall_not_mem, h]⟩ }, - { rintro (h | h) x, - { exact h.elim' x }, - { simp [h] } } + refine exists_congr (λ i, _), + casesI is_empty_or_nonempty (α i); simp [*, forall_and_distrib, eq_empty_iff_forall_not_mem], end -lemma univ_pi_eq_empty_iff : pi univ t = ∅ ↔ ∃ i, t i = ∅ := +@[simp] lemma univ_pi_eq_empty_iff : pi univ t = ∅ ↔ ∃ i, t i = ∅ := by simp [← not_nonempty_iff_eq_empty, univ_pi_nonempty_iff] @[simp] lemma univ_pi_empty [h : nonempty ι] : pi univ (λ i, ∅ : Π i, set (α i)) = ∅ := univ_pi_eq_empty_iff.2 $ h.elim $ λ x, ⟨x, rfl⟩ +@[simp] lemma disjoint_univ_pi : disjoint (pi univ t₁) (pi univ t₂) ↔ ∃ i, disjoint (t₁ i) (t₂ i) := +by simp only [disjoint_iff_inter_eq_empty, ← pi_inter_distrib, univ_pi_eq_empty_iff] + @[simp] lemma range_dcomp (f : Π i, α i → β i) : range (λ (g : Π i, α i), (λ i, f i (g i))) = pi univ (λ i, range (f i)) := begin @@ -525,6 +524,18 @@ lemma eval_image_pi (hs : i ∈ s) (ht : (s.pi t).nonempty) : eval i '' s.pi t = (λ f : Π i, α i, f i) '' pi univ t = t i := eval_image_pi (mem_univ i) ht +lemma pi_subset_pi_iff : pi s t₁ ⊆ pi s t₂ ↔ (∀ i ∈ s, t₁ i ⊆ t₂ i) ∨ pi s t₁ = ∅ := +begin + refine ⟨λ h, or_iff_not_imp_right.2 _, λ h, h.elim pi_mono (λ h', h'.symm ▸ empty_subset _)⟩, + rw [← ne.def, ne_empty_iff_nonempty], + intros hne i hi, + simpa only [eval_image_pi hi hne, eval_image_pi hi (hne.mono h)] + using image_subset (λ f : Π i, α i, f i) h +end + +lemma univ_pi_subset_univ_pi_iff : pi univ t₁ ⊆ pi univ t₂ ↔ (∀ i, t₁ i ⊆ t₂ i) ∨ ∃ i, t₁ i = ∅ := +by simp [pi_subset_pi_iff] + lemma eval_preimage [decidable_eq ι] {s : set (α i)} : eval i ⁻¹' s = pi univ (update (λ i, univ) i s) := by { ext x, simp [@forall_update_iff _ (λ i, set (α i)) _ _ _ _ (λ i' y, x i' ∈ y)] } From a67fbae61f4021038cad203c89af722c3aca992a Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sat, 8 Oct 2022 05:41:11 +0000 Subject: [PATCH 628/828] feat(measure_theory/covering): improve Vitali families and Lebesgue density theorem (#16830) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit #16762 has shown weaknesses of our current implementation of Vitali families. Notably, it enforces that the sets based at `x` all contain `x`, which is not natural for some applications. We refactor Vitali families to solve this issue. Here are the main changes: * in the definition of Vitali families, in the covering property it is now allowed to use several sets based at the same point (which means that the covering is not indexed by `α` but by `α × set α`) * We modify the Vitali covering theorem to deal with general indexed families, to fit in this framework. * This makes it possible to define better Vitali families for doubling measures. In particular, for any `K`, we define a Vitali family such that the sets based at `x` contain all balls `closed_ball y r` when `dist x y ≤ K * r`. * This gives a better Lebesgue density theorem, solving the issue pointed out in #16762 Co-authored-by: sgouezel --- src/measure_theory/covering/besicovitch.lean | 8 +- .../covering/density_theorem.lean | 183 +++++--- src/measure_theory/covering/vitali.lean | 407 ++++++++---------- .../covering/vitali_family.lean | 87 ++-- .../metric_space/hausdorff_distance.lean | 19 + 5 files changed, 389 insertions(+), 315 deletions(-) diff --git a/src/measure_theory/covering/besicovitch.lean b/src/measure_theory/covering/besicovitch.lean index f3ce826a3ee72..f0576fecbea21 100644 --- a/src/measure_theory/covering/besicovitch.lean +++ b/src/measure_theory/covering/besicovitch.lean @@ -1081,7 +1081,13 @@ protected def vitali_family (μ : measure α) [sigma_finite μ] : ∧ μ (s \ (⋃ (x ∈ t), closed_ball x (r x))) = 0 ∧ t.pairwise_disjoint (λ x, closed_ball x (r x)) := exists_disjoint_closed_ball_covering_ae μ g s A (λ _, 1) (λ _ _, zero_lt_one), - exact ⟨t, λ x, closed_ball x (r x), ts, tdisj, λ x xt, (tg x xt).1.2, μt⟩, + let F : α → α × set α := λ x, (x, closed_ball x (r x)), + refine ⟨F '' t, _, _, _, _⟩, + { rintros - ⟨x, hx, rfl⟩, exact ts hx }, + { rintros p ⟨x, hx, rfl⟩ q ⟨y, hy, rfl⟩ hxy, + exact tdisj hx hy (ne_of_apply_ne F hxy) }, + { rintros - ⟨x, hx, rfl⟩, exact (tg x hx).1.2 }, + { rwa bUnion_image } end } /-- The main feature of the Besicovitch Vitali family is that its filter at a point `x` corresponds diff --git a/src/measure_theory/covering/density_theorem.lean b/src/measure_theory/covering/density_theorem.lean index 7c88d649c5f7b..fe98c8be7ad7a 100644 --- a/src/measure_theory/covering/density_theorem.lean +++ b/src/measure_theory/covering/density_theorem.lean @@ -14,9 +14,9 @@ A doubling measure `μ` on a metric space is a measure for which there exists a that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius `2 * ε` is bounded by `C` times the measure of the concentric ball of radius `ε`. -Lebesgue's density theorem states that given a set `S` in a proper metric space with locally-finite -doubling measure `μ` then for almost all points `x` in `S`, for any sequence of closed balls -`B₀, B₁, B₂, ...` containing `x`, the limit `μ (S ∩ Bⱼ) / μ (Bⱼ) → 1` as `j → ∞`. +Lebesgue's density theorem states that given a set `S` in a sigma compact metric space with +locally-finite doubling measure `μ` then for almost all points `x` in `S`, for any sequence of +closed balls `B₀, B₁, B₂, ...` containing `x`, the limit `μ (S ∩ Bⱼ) / μ (Bⱼ) → 1` as `j → ∞`. In this file we combine general results about existence of Vitali families for doubling measures with results about differentiation along a Vitali family to obtain an explicit form of Lebesgue's @@ -29,14 +29,15 @@ density theorem. in the definition of a doubling measure. * `is_doubling_measure.ae_tendsto_measure_inter_div`: a version of Lebesgue's density theorem for sequences of balls converging on a point but whose centres are not required to be fixed. - -/ noncomputable theory -open set filter metric measure_theory +open set filter metric measure_theory topological_space open_locale nnreal topological_space +local attribute [instance] emetric.second_countable_of_sigma_compact + /-- A measure `μ` is said to be a doubling measure if there exists a constant `C` such that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius `2 * ε` is bounded by `C` times the measure of the concentric ball of radius `ε`. @@ -95,63 +96,149 @@ end /-- A variant of `is_doubling_measure.doubling_constant` which allows for scaling the radius by values other than `2`. -/ def scaling_constant_of (K : ℝ) : ℝ≥0 := -classical.some $ exists_eventually_forall_measure_closed_ball_le_mul μ K +max (classical.some $ exists_eventually_forall_measure_closed_ball_le_mul μ K) 1 + +lemma eventually_measure_mul_le_scaling_constant_of_mul (K : ℝ) : + ∃ (R : ℝ), 0 < R ∧ ∀ x t r (ht : t ∈ Ioc 0 K) (hr : r ≤ R), + μ (closed_ball x (t * r)) ≤ scaling_constant_of μ K * μ (closed_ball x r) := +begin + have h := classical.some_spec (exists_eventually_forall_measure_closed_ball_le_mul μ K), + rcases mem_nhds_within_Ioi_iff_exists_Ioc_subset.1 h with ⟨R, Rpos, hR⟩, + refine ⟨R, Rpos, λ x t r ht hr, _⟩, + rcases lt_trichotomy r 0 with rneg|rfl|rpos, + { have : t * r < 0, from mul_neg_of_pos_of_neg ht.1 rneg, + simp only [closed_ball_eq_empty.2 this, measure_empty, zero_le'] }, + { simp only [mul_zero, closed_ball_zero], + refine le_mul_of_one_le_of_le _ le_rfl, + apply ennreal.one_le_coe_iff.2 (le_max_right _ _) }, + { apply (hR ⟨rpos, hr⟩ x t ht.2).trans _, + exact ennreal.mul_le_mul (ennreal.coe_le_coe.2 (le_max_left _ _)) le_rfl } +end + +/-- A scale below which the doubling measure `μ` satisfies good rescaling properties when one +multiplies the radius of balls by at most `K`, as stated +in `measure_mul_le_scaling_constant_of_mul`. -/ +def scaling_scale_of (K : ℝ) : ℝ := +(eventually_measure_mul_le_scaling_constant_of_mul μ K).some + +lemma scaling_scale_of_pos (K : ℝ) : 0 < scaling_scale_of μ K := +(eventually_measure_mul_le_scaling_constant_of_mul μ K).some_spec.1 -lemma eventually_scaling_constant_of (K : ℝ) : - ∀ᶠ ε in 𝓝[>] 0, ∀ x t (ht : t ≤ K), - μ (closed_ball x (t * ε)) ≤ (scaling_constant_of μ K) * μ (closed_ball x ε) := -classical.some_spec $ exists_eventually_forall_measure_closed_ball_le_mul μ K +lemma measure_mul_le_scaling_constant_of_mul {K : ℝ} {x : α} {t r : ℝ} + (ht : t ∈ Ioc 0 K) (hr : r ≤ scaling_scale_of μ K) : + μ (closed_ball x (t * r)) ≤ scaling_constant_of μ K * μ (closed_ball x r) := +(eventually_measure_mul_le_scaling_constant_of_mul μ K).some_spec.2 x t r ht hr -variables [proper_space α] [borel_space α] [is_locally_finite_measure μ] +section +variables [second_countable_topology α] [borel_space α] [is_locally_finite_measure μ] -/-- A Vitali family in space with doubling measure with a covering proportion controlled by `K`. -/ -def vitali_family (K : ℝ) (hK : 6 ≤ K) : vitali_family μ := -vitali.vitali_family μ (scaling_constant_of μ K) $ λ x ε hε, +open_locale topological_space + +/-- A Vitali family in a space with a doubling measure, designed so that the sets at `x` contain +all `closed_ball y r` when `dist x y ≤ K * r`. -/ +@[irreducible] def vitali_family (K : ℝ) : vitali_family μ := begin - have h := eventually_scaling_constant_of μ K, - replace h := forall_eventually_of_eventually_forall (forall_eventually_of_eventually_forall h x), - replace h := eventually_imp_distrib_left.mp (h 6) hK, - simpa only [exists_prop] using ((eventually_nhds_within_pos_mem_Ioc hε).and h).exists, + /- the Vitali covering theorem gives a family that works well at small scales, thanks to the + doubling property. We enlarge this family to add large sets, to make sure that all balls and not + only small ones belong to the family, for convenience. -/ + let R := scaling_scale_of μ (max (4 * K + 3) 3), + have Rpos : 0 < R := scaling_scale_of_pos _ _, + have A : ∀ (x : α), ∃ᶠ r in 𝓝[>] (0 : ℝ), + μ (closed_ball x (3 * r)) ≤ scaling_constant_of μ (max (4 * K + 3) 3) * μ (closed_ball x r), + { assume x, + apply frequently_iff.2 (λ U hU, _), + obtain ⟨ε, εpos, hε⟩ := mem_nhds_within_Ioi_iff_exists_Ioc_subset.1 hU, + refine ⟨min ε R, hε ⟨lt_min εpos Rpos, min_le_left _ _⟩, _⟩, + exact measure_mul_le_scaling_constant_of_mul μ + ⟨zero_lt_three, le_max_right _ _⟩ (min_le_right _ _) }, + exact (vitali.vitali_family μ (scaling_constant_of μ (max (4 * K + 3) 3)) A).enlarge + (R / 4) (by linarith), end -/-- A version of *Lebesgue's density theorem* for a sequence of closed balls whose centres are -not required to be fixed. +/-- In the Vitali family `is_doubling_measure.vitali_family K`, the sets based at `x` contain all +balls `closed_ball y r` when `dist x y ≤ K * r`. -/ +lemma closed_ball_mem_vitali_family_of_dist_le_mul + {K : ℝ} {x y : α} {r : ℝ} (h : dist x y ≤ K * r) (rpos : 0 < r) : + closed_ball y r ∈ (vitali_family μ K).sets_at x := +begin + let R := scaling_scale_of μ (max (4 * K + 3) 3), + simp only [vitali_family, vitali_family.enlarge, vitali.vitali_family, mem_union, mem_set_of_eq, + is_closed_ball, true_and, (nonempty_ball.2 rpos).mono ball_subset_interior_closed_ball, + measurable_set_closed_ball], + /- The measure is doubling on scales smaller than `R`. Therefore, we treat differently small + and large balls. For large balls, this follows directly from the enlargement we used in the + definition. -/ + by_cases H : closed_ball y r ⊆ closed_ball x (R / 4), + swap, { exact or.inr H }, + left, + /- For small balls, there is the difficulty that `r` could be large but still the ball could be + small, if the annulus `{y | ε ≤ dist y x ≤ R/4}` is empty. We split between the cases `r ≤ R` + and `r < R`, and use the doubling for the former and rough estimates for the latter. -/ + rcases le_or_lt r R with hr|hr, + { refine ⟨(K + 1) * r, _⟩, + split, + { apply closed_ball_subset_closed_ball', + rw dist_comm, + linarith }, + { have I1 : closed_ball x (3 * ((K + 1) * r)) ⊆ closed_ball y ((4 * K + 3) * r), + { apply closed_ball_subset_closed_ball', + linarith }, + have I2 : closed_ball y ((4 * K + 3) * r) ⊆ closed_ball y ((max (4 * K + 3) 3) * r), + { apply closed_ball_subset_closed_ball, + exact mul_le_mul_of_nonneg_right (le_max_left _ _) rpos.le }, + apply (measure_mono (I1.trans I2)).trans, + exact measure_mul_le_scaling_constant_of_mul _ + ⟨zero_lt_three.trans_le (le_max_right _ _), le_rfl⟩ hr } }, + { refine ⟨R / 4, H, _⟩, + have : closed_ball x (3 * (R / 4)) ⊆ closed_ball y r, + { apply closed_ball_subset_closed_ball', + have A : y ∈ closed_ball y r, from mem_closed_ball_self rpos.le, + have B := mem_closed_ball'.1 (H A), + linarith }, + apply (measure_mono this).trans _, + refine le_mul_of_one_le_left (zero_le _) _, + exact ennreal.one_le_coe_iff.2 (le_max_right _ _) } +end -See also `besicovitch.ae_tendsto_measure_inter_div`. -/ -lemma ae_tendsto_measure_inter_div (S : set α) (K : ℝ) (hK : K ∈ unit_interval) : - ∀ᵐ x ∂μ.restrict S, ∀ {ι : Type*} {l : filter ι} (w : ι → α) (δ : ι → ℝ) - (δlim : tendsto δ l (𝓝[>] 0)) - (xmem : ∀ᶠ j in l, x ∈ closed_ball (w j) (K * δ j)), - tendsto (λ j, μ (S ∩ closed_ball (w j) (δ j)) / μ (closed_ball (w j) (δ j))) l (𝓝 1) := +lemma tendsto_closed_ball_filter_at {K : ℝ} {x : α} {ι : Type*} {l : filter ι} + (w : ι → α) (δ : ι → ℝ) (δlim : tendsto δ l (𝓝[>] 0)) + (xmem : ∀ᶠ j in l, x ∈ closed_ball (w j) (K * δ j)) : + tendsto (λ j, closed_ball (w j) (δ j)) l ((vitali_family μ K).filter_at x) := begin - let v := is_doubling_measure.vitali_family μ 7 (by norm_num), - filter_upwards [v.ae_tendsto_measure_inter_div S] with x hx ι l w δ δlim xmem, - suffices : tendsto (λ j, closed_ball (w j) (δ j)) l (v.filter_at x), { exact hx.comp this, }, - refine v.tendsto_filter_at_iff.mpr ⟨_, λ ε hε, _⟩, - { simp only [v, vitali.vitali_family], - have δpos : ∀ᶠ j in l, 0 < δ j := eventually_mem_of_tendsto_nhds_within δlim, - replace xmem : ∀ᶠ (j : ι) in l, x ∈ closed_ball (w j) (δ j) := (δpos.and xmem).mono - (λ j hj, closed_ball_subset_closed_ball (by nlinarith [hj.1, hK.2]) hj.2), - apply ((δlim.eventually (eventually_scaling_constant_of μ 7)).and (xmem.and δpos)).mono, - rintros j ⟨hjC, hjx, hjδ⟩, - have hdiam : 3 * diam (closed_ball (w j) (δ j)) ≤ 6 * δ j, - { linarith [@diam_closed_ball _ _ (w j) _ hjδ.le], }, - refine ⟨hjx, is_closed_ball, (nonempty_ball.mpr hjδ).mono ball_subset_interior_closed_ball, - (measure_mono (closed_ball_subset_closed_ball hdiam)).trans _⟩, - suffices : closed_ball x (6 * δ j) ⊆ closed_ball (w j) (7 * δ j), - { exact (measure_mono this).trans ((hjC (w j) 7 (by norm_num)).trans $ le_refl _), }, - intros y hy, - simp only [mem_closed_ball, dist_comm x (w j)] at hjx hy ⊢, - linarith [dist_triangle_right y (w j) x], }, - { have δpos := eventually_mem_of_tendsto_nhds_within δlim, + refine (vitali_family μ K).tendsto_filter_at_iff.mpr ⟨_, λ ε hε, _⟩, + { filter_upwards [xmem, δlim self_mem_nhds_within] with j hj h'j, + exact closed_ball_mem_vitali_family_of_dist_le_mul μ hj h'j }, + { by_cases l.ne_bot, + swap, { simp [not_ne_bot.1 h] }, + have hK : 0 ≤ K, + { resetI, + rcases (xmem.and (δlim self_mem_nhds_within)).exists with ⟨j, hj, h'j⟩, + have : 0 ≤ K * δ j := nonempty_closed_ball.1 ⟨x, hj⟩, + exact (mul_nonneg_iff_left_nonneg_of_pos (mem_Ioi.1 h'j)).1 this }, + have δpos := eventually_mem_of_tendsto_nhds_within δlim, replace δlim := tendsto_nhds_of_tendsto_nhds_within δlim, - replace hK : 0 < K + 1 := by linarith [hK.1], + replace hK : 0 < K + 1, by linarith, apply (((metric.tendsto_nhds.mp δlim _ (div_pos hε hK)).and δpos).and xmem).mono, rintros j ⟨⟨hjε, hj₀ : 0 < δ j⟩, hx⟩ y hy, replace hjε : (K + 1) * δ j < ε := by simpa [abs_eq_self.mpr hj₀.le] using (lt_div_iff' hK).mp hjε, simp only [mem_closed_ball] at hx hy ⊢, - linarith [dist_triangle_right y x (w j)], }, + linarith [dist_triangle_right y x (w j)] } end +end + +/-- A version of *Lebesgue's density theorem* for a sequence of closed balls whose centres are +not required to be fixed. + +See also `besicovitch.ae_tendsto_measure_inter_div`. -/ +lemma ae_tendsto_measure_inter_div + [sigma_compact_space α] [borel_space α] [is_locally_finite_measure μ] (S : set α) (K : ℝ) : + ∀ᵐ x ∂μ.restrict S, ∀ {ι : Type*} {l : filter ι} (w : ι → α) (δ : ι → ℝ) + (δlim : tendsto δ l (𝓝[>] 0)) + (xmem : ∀ᶠ j in l, x ∈ closed_ball (w j) (K * δ j)), + tendsto (λ j, μ (S ∩ closed_ball (w j) (δ j)) / μ (closed_ball (w j) (δ j))) l (𝓝 1) := +by filter_upwards [(vitali_family μ K).ae_tendsto_measure_inter_div S] with x hx ι l w δ δlim xmem +using hx.comp (tendsto_closed_ball_filter_at μ _ _ δlim xmem) + end is_doubling_measure diff --git a/src/measure_theory/covering/vitali.lean b/src/measure_theory/covering/vitali.lean index 6a692820e1aee..c8ca760e001c4 100644 --- a/src/measure_theory/covering/vitali.lean +++ b/src/measure_theory/covering/vitali.lean @@ -33,7 +33,7 @@ covering a fixed proportion `1/C` of the ball `closed_ball x (3 * diam a)` forms This version is given in `vitali.vitali_family`. -/ -variables {α : Type*} +variables {α ι : Type*} open set metric measure_theory topological_space filter open_locale nnreal classical ennreal topological_space @@ -48,12 +48,15 @@ When `t` is a family of balls, the `τ`-enlargment of `ball x r` is `ball x ((1+ 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 / τ`. + +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 - (t : set (set α)) (δ : set α → ℝ) (τ : ℝ) (hτ : 1 < τ) (δnonneg : ∀ a ∈ t, 0 ≤ δ a) - (R : ℝ) (δle : ∀ a ∈ t, δ a ≤ R) (hne : ∀ a ∈ t, set.nonempty a) : - ∃ u ⊆ t, u.pairwise_disjoint id ∧ - ∀ a ∈ t, ∃ b ∈ u, set.nonempty (a ∩ b) ∧ δ a ≤ τ * δ b := + (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, u.pairwise_disjoint B ∧ + ∀ a ∈ t, ∃ b ∈ u, (B a ∩ B b).nonempty ∧ δ a ≤ τ * δ b := begin /- The proof could be formulated as a transfinite induction. First pick an element of `t` with `δ` as large as possible (up to a factor of `τ`). Then among the remaining elements not intersecting @@ -68,8 +71,8 @@ begin that `u ∪ {a'}` still has this property, contradicting the maximality. Therefore, `u` intersects all elements of `t`, and by definition it satisfies all the desired properties. -/ - let T : set (set (set α)) := {u | u ⊆ t ∧ u.pairwise_disjoint id - ∧ ∀ a ∈ t, ∀ b ∈ u, set.nonempty (a ∩ b) → ∃ c ∈ u, (a ∩ c).nonempty ∧ δ a ≤ τ * δ c}, + let T : set (set ι) := {u | u ⊆ t ∧ u.pairwise_disjoint B + ∧ ∀ a ∈ t, ∀ b ∈ u, (B a ∩ B b).nonempty → ∃ c ∈ u, (B a ∩ B c).nonempty ∧ δ a ≤ τ * δ c}, -- By Zorn, choose a maximal family in the good set `T` of disjoint families. obtain ⟨u, uT, hu⟩ : ∃ u ∈ T, ∀ v ∈ T, u ⊆ v → v = u, { refine zorn_subset _ (λ U UT hU, _), @@ -78,23 +81,23 @@ begin set.mem_set_of_eq], refine ⟨λ u hu, (UT hu).1, (pairwise_disjoint_sUnion hU.directed_on).2 (λ u hu, (UT hu).2.1), λ a hat b u uU hbu hab, _⟩, - obtain ⟨c, cu, ac, hc⟩ : ∃ (c : set α) (H : c ∈ u), (a ∩ c).nonempty ∧ δ a ≤ τ * δ c := + obtain ⟨c, cu, ac, hc⟩ : ∃ (c : ι) (H : c ∈ u), (B a ∩ B c).nonempty ∧ δ a ≤ τ * δ c := (UT uU).2.2 a hat b hbu hab, exact ⟨c, ⟨u, uU, cu⟩, ac, hc⟩ }, -- the only nontrivial bit is to check that every `a ∈ t` intersects an element `b ∈ u` with -- comparatively large `δ b`. Assume this is not the case, then we will contradict the maximality. refine ⟨u, uT.1, uT.2.1, λ a hat, _⟩, contrapose! hu, - have a_disj : ∀ c ∈ u, disjoint a c, + have a_disj : ∀ c ∈ u, disjoint (B a) (B c), { assume c hc, by_contra, rw not_disjoint_iff_nonempty_inter at h, - obtain ⟨d, du, ad, hd⟩ : ∃ (d : set α) (H : d ∈ u), (a ∩ d).nonempty ∧ δ a ≤ τ * δ d := + obtain ⟨d, du, ad, hd⟩ : ∃ (d : ι) (H : d ∈ u), (B a ∩ B d).nonempty ∧ δ a ≤ τ * δ d := uT.2.2 a hat c hc h, exact lt_irrefl _ ((hu d du ad).trans_le hd) }, -- Let `A` be all the elements of `t` which do not intersect the family `u`. It is nonempty as it -- contains `a`. We will pick an element `a'` of `A` with `δ a'` almost as large as possible. - let A := {a' | a' ∈ t ∧ ∀ c ∈ u, disjoint a' c}, + let A := {a' | a' ∈ t ∧ ∀ c ∈ u, disjoint (B a') (B c)}, have Anonempty : A.nonempty := ⟨a, hat, a_disj⟩, let m := Sup (δ '' A), have bddA : bdd_above (δ '' A), @@ -128,7 +131,7 @@ begin { assume c ct b ba'u hcb, -- if `c` already intersects an element of `u`, then it intersects an element of `u` with -- large `δ` by the assumption on `u`, and there is nothing left to do. - by_cases H : ∃ d ∈ u, set.nonempty (c ∩ d), + by_cases H : ∃ d ∈ u, (B c ∩ B d).nonempty, { rcases H with ⟨d, du, hd⟩, rcases uT.2.2 c ct d du hd with ⟨d', d'u, hd'⟩, exact ⟨d', mem_insert_of_mem _ d'u, hd'⟩ }, @@ -150,101 +153,65 @@ end extract a disjoint subfamily `u ⊆ t` so that all balls in `t` are covered by the 5-times dilations of balls in `u`. -/ theorem exists_disjoint_subfamily_covering_enlargment_closed_ball [metric_space α] - (t : set (set α)) (R : ℝ) (ht : ∀ s ∈ t, ∃ x r, s = closed_ball x r ∧ r ≤ R) : - ∃ u ⊆ t, u.pairwise_disjoint id ∧ - ∀ a ∈ t, ∃ x r, closed_ball x r ∈ u ∧ a ⊆ closed_ball x (5 * r) := + (t : set ι) (x : ι → α) (r : ι → ℝ) (R : ℝ) (hr : ∀ a ∈ t, r a ≤ R) : + ∃ u ⊆ t, u.pairwise_disjoint (λ a, closed_ball (x a) (r a)) ∧ + ∀ a ∈ t, ∃ b ∈ u, closed_ball (x a) (r a) ⊆ closed_ball (x b) (5 * r b) := begin rcases eq_empty_or_nonempty t with rfl|tnonempty, { exact ⟨∅, subset.refl _, pairwise_disjoint_empty, by simp⟩ }, - haveI : inhabited α, - { choose s hst using tnonempty, - choose x r hxr using ht s hst, - exact ⟨x⟩ }, - -- Exclude the trivial case where `t` is reduced to the empty set. - rcases eq_or_ne t {∅} with rfl|t_ne_empty, - { refine ⟨{∅}, subset.refl _, _⟩, - simp only [true_and, closed_ball_eq_empty, mem_singleton_iff, and_true, empty_subset, forall_eq, - pairwise_disjoint_singleton, exists_const], - exact ⟨-1, by simp only [right.neg_neg_iff, zero_lt_one]⟩ }, - -- The real proof starts now. Since the center or the radius of a ball is not uniquely defined - -- in a general metric space, we just choose one for definiteness. - choose! x r hxr using ht, - have r_nonneg : ∀ (a : set α), a ∈ t → a.nonempty → 0 ≤ r a, - { assume a hat a_nonempty, - rw (hxr a hat).1 at a_nonempty, - simpa only [nonempty_closed_ball] using a_nonempty }, - -- The difference with the generic version is that we are not excluding empty sets in our family - -- (which would correspond to `r a < 0`). To compensate for this, we apply the generic version - -- to the subfamily `t'` made of nonempty sets, and we use `δ = r` there. This gives a disjointed - -- subfamily `u'`. + by_cases ht : ∀ a ∈ t, r a < 0, + { exact ⟨t, subset.rfl, λ a ha b hb hab, + by simp only [function.on_fun, closed_ball_eq_empty.2 (ht a ha), empty_disjoint], + λ a ha, ⟨a, ha, by simp only [closed_ball_eq_empty.2 (ht a ha), empty_subset]⟩⟩ }, + push_neg at ht, let t' := {a ∈ t | 0 ≤ r a}, - obtain ⟨u', u't', u'_disj, hu'⟩ : ∃ u' ⊆ t', u'.pairwise_disjoint id ∧ - ∀ a ∈ t', ∃ b ∈ u', set.nonempty (a ∩ b) ∧ r a ≤ 2 * r b, - { refine exists_disjoint_subfamily_covering_enlargment t' r 2 one_lt_two - (λ a ha, ha.2) R (λ a ha, (hxr a ha.1).2) (λ a ha, _), - rw [(hxr a ha.1).1], - simp only [ha.2, nonempty_closed_ball] }, - -- this subfamily is nonempty, as we have excluded the situation `t = {∅}`. - have u'_nonempty : u'.nonempty, - { have : ∃ a ∈ t, a ≠ ∅, - { contrapose! t_ne_empty, - apply subset.antisymm, - { simpa only using t_ne_empty }, - { rcases tnonempty with ⟨a, hat⟩, - have := t_ne_empty a hat, - simpa only [this, singleton_subset_iff] using hat } }, - rcases this with ⟨a, hat, a_nonempty⟩, - have ranonneg : 0 ≤ r a := r_nonneg a hat (ne_empty_iff_nonempty.1 a_nonempty), - rcases hu' a ⟨hat, ranonneg⟩ with ⟨b, bu', hb⟩, - exact ⟨b, bu'⟩ }, - -- check that the family `u'` gives the desired disjoint covering. - refine ⟨u', λ a ha, (u't' ha).1, u'_disj, λ a hat, _⟩, - -- it remains to check that any ball in `t` is contained in the 5-dilation of a ball - -- in `u'`. This depends on whether the ball is empty of not. - rcases eq_empty_or_nonempty a with rfl|a_nonempty, - -- if the ball is empty, use any element of `u'` (since we know that `u'` is nonempty). - { rcases u'_nonempty with ⟨b, hb⟩, - refine ⟨x b, r b, _, empty_subset _⟩, - rwa ← (hxr b (u't' hb).1).1 }, - -- if the ball is not empty, it belongs to `t'`. Then it intersects a ball `a'` in `u'` with - -- controlled radius, by definition of `u'`. It is straightforward to check that this ball - -- satisfies all the desired properties. - { have hat' : a ∈ t' := ⟨hat, r_nonneg a hat a_nonempty⟩, - obtain ⟨a', a'u', aa', raa'⟩ : - (∃ (a' : set α) (H : a' ∈ u'), (a ∩ a').nonempty ∧ r a ≤ 2 * r a') := hu' a hat', - refine ⟨x a', r a', _, _⟩, - { convert a'u', - exact (hxr a' (u't' a'u').1).1.symm }, - { rw (hxr a hat'.1).1 at aa' ⊢, - rw (hxr a' (u't' a'u').1).1 at aa', - have : dist (x a) (x a') ≤ r a + r a' := - dist_le_add_of_nonempty_closed_ball_inter_closed_ball aa', - apply closed_ball_subset_closed_ball', - linarith } } + rcases exists_disjoint_subfamily_covering_enlargment (λ a, closed_ball (x a) (r a)) t' r + 2 one_lt_two (λ a ha, ha.2) R (λ a ha, hr a ha.1) (λ a ha, ⟨x a, mem_closed_ball_self ha.2⟩) + with ⟨u, ut', u_disj, hu⟩, + have A : ∀ a ∈ t', ∃ b ∈ u, closed_ball (x a) (r a) ⊆ closed_ball (x b) (5 * r b), + { assume a ha, + rcases hu a ha with ⟨b, bu, hb, rb⟩, + refine ⟨b, bu, _⟩, + have : dist (x a) (x b) ≤ r a + r b := + dist_le_add_of_nonempty_closed_ball_inter_closed_ball hb, + apply closed_ball_subset_closed_ball', + linarith }, + refine ⟨u, ut'.trans (λ a ha, ha.1), u_disj, λ a ha, _⟩, + rcases le_or_lt 0 (r a) with h'a|h'a, + { exact A a ⟨ha, h'a⟩ }, + { rcases ht with ⟨b, rb⟩, + rcases A b ⟨rb.1, rb.2⟩ with ⟨c, cu, hc⟩, + refine ⟨c, cu, by simp only [closed_ball_eq_empty.2 h'a, empty_subset]⟩ }, end + /-- 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, 6 r)` for a given measure `μ` (think of the situation where `μ` is -a doubling measure and `t` is a family of balls). Consider a (possible non-measurable) set `s` +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`. -/ +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. +-/ theorem exists_disjoint_covering_ae [metric_space α] [measurable_space α] [opens_measurable_space α] [second_countable_topology α] (μ : measure α) [is_locally_finite_measure μ] (s : set α) - (t : set (set α)) (hf : ∀ x ∈ s, ∀ (ε > (0 : ℝ)), ∃ a ∈ t, x ∈ a ∧ a ⊆ closed_ball x ε) - (ht : ∀ a ∈ t, (interior a).nonempty) (h't : ∀ a ∈ t, is_closed a) - (C : ℝ≥0) (h : ∀ a ∈ t, ∃ x ∈ a, μ (closed_ball x (3 * diam a)) ≤ C * μ a) : - ∃ u ⊆ t, u.countable ∧ u.pairwise_disjoint id ∧ μ (s \ ⋃ (a ∈ u), a) = 0 := + (t : set ι) (C : ℝ≥0) (r : ι → ℝ) (c : ι → α) (B : ι → set α) + (hB : ∀ a ∈ t, B a ⊆ closed_ball (c a) (r a)) + (μB : ∀ a ∈ t, μ (closed_ball (c a) (3 * r a)) ≤ C * μ (B a)) + (ht : ∀ a ∈ t, (interior (B a)).nonempty) (h't : ∀ a ∈ t, is_closed (B a)) + (hf : ∀ x ∈ s, ∀ (ε > (0 : ℝ)), ∃ a ∈ t, r a ≤ ε ∧ c a = x) : + ∃ u ⊆ t, u.countable ∧ u.pairwise_disjoint B ∧ μ (s \ ⋃ a ∈ u, B a) = 0 := begin /- The idea of the proof is the following. Assume for simplicity that `μ` is finite. Applying the - abstract Vitali covering theorem with `δ = diam`, one obtains a disjoint subfamily `u`, such - that any element of `t` intersects an element of `u` with comparable diameter. Fix `ε > 0`. + abstract Vitali covering theorem with `δ = r` given by `hf`, one obtains a disjoint subfamily `u`, + such that any element of `t` intersects an element of `u` with comparable radius. Fix `ε > 0`. Since the elements of `u` have summable measure, one can remove finitely elements `w_1, ..., w_n`. so that the measure of the remaining elements is `< ε`. Consider now a point `z` not in the `w_i`. There is a small ball around `z` not intersecting the `w_i` (as they are closed), an element `a ∈ t` contained in this small ball (as the family `t` is fine at `z`) and an element - `b ∈ u` intersecting `a`, with comparable diameter (by definition of `u`). Then `z` belongs to the + `b ∈ u` intersecting `a`, with comparable radius (by definition of `u`). Then `z` belongs to the enlargement of `b`. This shows that `s \ (w_1 ∪ ... ∪ w_n)` is contained in `⋃ (b ∈ u \ {w_1, ... w_n}) (enlargement of b)`. The measure of the latter set is bounded by `∑ (b ∈ u \ {w_1, ... w_n}) C * μ b` (by the doubling property of the measure), which is at most @@ -255,14 +222,8 @@ begin use the whole family `t`, but a subfamily `t'` supported on small balls (which is possible since the family is assumed to be fine at every point of `s`). -/ - rcases eq_empty_or_nonempty s with rfl|nonempty, - { refine ⟨∅, empty_subset _, countable_empty, pairwise_disjoint_empty, - by simp only [measure_empty, Union_false, Union_empty, diff_self]⟩ }, - haveI : inhabited α, - { choose x hx using nonempty, - exact ⟨x⟩ }, -- choose around each `x` a small ball on which the measure is finite - have : ∀ x, ∃ r, 0 < r ∧ r ≤ 1 ∧ μ (closed_ball x (20 * r)) < ∞, + have : ∀ x, ∃ R, 0 < R ∧ R ≤ 1 ∧ μ (closed_ball x (20 * R)) < ∞, { assume x, obtain ⟨R, Rpos, μR⟩ : ∃ (R : ℝ) (hR : 0 < R), μ (closed_ball x R) < ∞ := (μ.finite_at_nhds x).exists_mem_basis nhds_basis_closed_ball, @@ -274,125 +235,123 @@ begin calc 20 * min 1 (R / 20) ≤ 20 * (R/20) : mul_le_mul_of_nonneg_left (min_le_right _ _) (by norm_num) ... = R : by ring } }, - choose r hr0 hr1 hrμ, + choose R hR0 hR1 hRμ, -- we restrict to a subfamily `t'` of `t`, made of elements small enough to ensure that - -- they only see a finite part of the measure. - let t' := {a ∈ t | ∃ x, a ⊆ closed_ball x (r x)}, + -- they only see a finite part of the measure, and with a doubling property + let t' := {a ∈ t | r a ≤ R (c a)}, -- extract a disjoint subfamily `u` of `t'` thanks to the abstract Vitali covering theorem. - obtain ⟨u, ut', u_disj, hu⟩ : ∃ u ⊆ t', u.pairwise_disjoint id ∧ - ∀ a ∈ t', ∃ b ∈ u, set.nonempty (a ∩ b) ∧ diam a ≤ 2 * diam b, - { have A : ∀ (a : set α), a ∈ t' → diam a ≤ 2, - { rintros a ⟨hat, ⟨x, hax⟩⟩, - calc diam a ≤ 2 * 1 : diam_le_of_subset_closed_ball zero_le_one - (hax.trans $ closed_ball_subset_closed_ball $ hr1 x) - ... = 2 : mul_one _ }, - have B : ∀ (a : set α), a ∈ t' → a.nonempty := + obtain ⟨u, ut', u_disj, hu⟩ : ∃ u ⊆ t', u.pairwise_disjoint B ∧ + ∀ a ∈ t', ∃ b ∈ u, (B a ∩ B b).nonempty ∧ r a ≤ 2 * r b, + { have A : ∀ a ∈ t', r a ≤ 1, + { assume a ha, + apply ha.2.trans (hR1 (c a)), }, + have A' : ∀ a ∈ t', (B a).nonempty := λ a hat', set.nonempty.mono interior_subset (ht a hat'.1), - exact exists_disjoint_subfamily_covering_enlargment t' diam 2 one_lt_two - (λ a ha, diam_nonneg) 2 A B }, + refine exists_disjoint_subfamily_covering_enlargment B t' r 2 one_lt_two + (λ a ha, _) 1 A A', + exact nonempty_closed_ball.1 ((A' a ha).mono (hB a ha.1)) }, have ut : u ⊆ t := λ a hau, (ut' hau).1, -- As the space is second countable, the family is countable since all its sets have nonempty -- interior. - have u_count : u.countable := - u_disj.countable_of_nonempty_interior (λ a ha, ht a (ut ha)), + have u_count : u.countable := u_disj.countable_of_nonempty_interior (λ a ha, ht a (ut ha)), -- the family `u` will be the desired family refine ⟨u, λ a hat', (ut' hat').1, u_count, u_disj, _⟩, -- it suffices to show that it covers almost all `s` locally around each point `x`. refine null_of_locally_null _ (λ x hx, _), -- let `v` be the subfamily of `u` made of those sets intersecting the small ball `ball x (r x)` - let v := {a ∈ u | (a ∩ ball x (r x)).nonempty }, + let v := {a ∈ u | (B a ∩ ball x (R x)).nonempty }, have vu : v ⊆ u := λ a ha, ha.1, -- they are all contained in a fixed ball of finite measure, thanks to our choice of `t'` - - obtain ⟨R, μR, hR⟩ : ∃ R, μ (closed_ball x R) < ∞ ∧ - ∀ a ∈ u, (a ∩ ball x (r x)).nonempty → a ⊆ closed_ball x R, - { have : ∀ a ∈ u, ∃ y, a ⊆ closed_ball y (r y) := λ a hau, (ut' hau).2, - choose! y hy using this, - have Idist_v : ∀ a ∈ v, dist (y a) x ≤ r (y a) + r x, + obtain ⟨K, μK, hK⟩ : ∃ K, μ (closed_ball x K) < ∞ ∧ + ∀ a ∈ u, (B a ∩ ball x (R x)).nonempty → B a ⊆ closed_ball x K, + { have Idist_v : ∀ a ∈ v, dist (c a) x ≤ r a + R x, { assume a hav, apply dist_le_add_of_nonempty_closed_ball_inter_closed_ball, - exact hav.2.mono (inter_subset_inter (hy a hav.1) ball_subset_closed_ball) }, - set R0 := Sup ((λ a, r (y a)) '' v) with hR0, - have R0_bdd : bdd_above ((λ a, r (y a)) '' v), + refine hav.2.mono _, + apply inter_subset_inter _ ball_subset_closed_ball, + exact hB a (ut (vu hav)) }, + set R0 := Sup (r '' v) with R0_def, + have R0_bdd : bdd_above (r '' v), { refine ⟨1, λ r' hr', _⟩, rcases (mem_image _ _ _).1 hr' with ⟨b, hb, rfl⟩, - exact hr1 _ }, - rcases le_total R0 (r x) with H|H, - { refine ⟨20 * r x, hrμ x, λ a au hax, _⟩, - refine (hy a au).trans _, + exact le_trans (ut' (vu hb)).2 (hR1 (c b)) }, + rcases le_total R0 (R x) with H|H, + { refine ⟨20 * R x, hRμ x, λ a au hax, _⟩, + refine (hB a (ut au)).trans _, apply closed_ball_subset_closed_ball', - have : r (y a) ≤ R0 := le_cSup R0_bdd (mem_image_of_mem _ ⟨au, hax⟩), - linarith [(hr0 (y a)).le, (hr0 x).le, Idist_v a ⟨au, hax⟩] }, - { have R0pos : 0 < R0 := (hr0 x).trans_le H, + have : r a ≤ R0 := le_cSup R0_bdd (mem_image_of_mem _ ⟨au, hax⟩), + linarith [Idist_v a ⟨au, hax⟩, hR0 x] }, + { have R0pos : 0 < R0 := (hR0 x).trans_le H, have vnonempty : v.nonempty, { by_contra, rw [← ne_empty_iff_nonempty, not_not] at h, - simp only [h, real.Sup_empty, image_empty] at hR0, - exact lt_irrefl _ (R0pos.trans_le (le_of_eq hR0)) }, - obtain ⟨a, hav, R0a⟩ : ∃ a ∈ v, R0/2 < r (y a), - { obtain ⟨r', r'mem, hr'⟩ : ∃ r' ∈ ((λ a, r (y a)) '' v), R0 / 2 < r' := + simp only [h, real.Sup_empty, image_empty] at R0_def, + exact lt_irrefl _ (R0pos.trans_le (le_of_eq R0_def)) }, + obtain ⟨a, hav, R0a⟩ : ∃ a ∈ v, R0/2 < r a, + { obtain ⟨r', r'mem, hr'⟩ : ∃ r' ∈ r '' v, R0 / 2 < r' := exists_lt_of_lt_cSup (nonempty_image_iff.2 vnonempty) (half_lt_self R0pos), rcases (mem_image _ _ _).1 r'mem with ⟨a, hav, rfl⟩, exact ⟨a, hav, hr'⟩ }, refine ⟨8 * R0, _, _⟩, - { apply lt_of_le_of_lt (measure_mono _) (hrμ (y a)), + { apply lt_of_le_of_lt (measure_mono _) (hRμ (c a)), apply closed_ball_subset_closed_ball', rw dist_comm, - linarith [Idist_v a hav] }, + linarith [Idist_v a hav, (ut' (vu hav)).2] }, { assume b bu hbx, - refine (hy b bu).trans _, + refine (hB b (ut bu)).trans _, apply closed_ball_subset_closed_ball', - have : r (y b) ≤ R0 := le_cSup R0_bdd (mem_image_of_mem _ ⟨bu, hbx⟩), + have : r b ≤ R0 := le_cSup R0_bdd (mem_image_of_mem _ ⟨bu, hbx⟩), linarith [Idist_v b ⟨bu, hbx⟩] } } }, - -- we will show that, in `ball x (r x)`, almost all `s` is covered by the family `u`. - refine ⟨_ ∩ ball x (r x), inter_mem_nhds_within _ (ball_mem_nhds _ (hr0 _)), + -- we will show that, in `ball x (R x)`, almost all `s` is covered by the family `u`. + refine ⟨_ ∩ ball x (R x), inter_mem_nhds_within _ (ball_mem_nhds _ (hR0 _)), nonpos_iff_eq_zero.mp (le_of_forall_le_of_dense (λ ε εpos, _))⟩, -- the elements of `v` are disjoint and all contained in a finite volume ball, hence the sum -- of their measures is finite. - have I : ∑' (a : v), μ (↑a) < ∞, - { calc ∑' (a : v), μ (↑a) = μ (⋃ (a ∈ v), a) : begin + have I : ∑' (a : v), μ (B a) < ∞, + { calc ∑' (a : v), μ (B a) = μ (⋃ (a ∈ v), B a) : begin rw measure_bUnion (u_count.mono vu) _ (λ a ha, (h't _ (vu.trans ut ha)).measurable_set), exact u_disj.subset vu end - ... ≤ μ (closed_ball x R) : measure_mono (Union₂_subset (λ a ha, hR a (vu ha) ha.2)) - ... < ∞ : μR }, + ... ≤ μ (closed_ball x K) : measure_mono (Union₂_subset (λ a ha, hK a (vu ha) ha.2)) + ... < ∞ : μK }, -- we can obtain a finite subfamily of `v`, such that the measures of the remaining elements -- add up to an arbitrarily small number, say `ε / C`. - obtain ⟨w, hw⟩ : ∃ (w : finset ↥v), ∑' (a : {a // a ∉ w}), μ (↑a) < ε / C, - { haveI : ne_bot (at_top : filter (finset v)) := at_top_ne_bot, - have : 0 < ε / C, by simp only [ennreal.div_pos_iff, εpos.ne', ennreal.coe_ne_top, ne.def, + obtain ⟨w, hw⟩ : ∃ (w : finset ↥v), ∑' (a : {a // a ∉ w}), μ (B a) < ε / C, + { have : 0 < ε / C, by simp only [ennreal.div_pos_iff, εpos.ne', ennreal.coe_ne_top, ne.def, not_false_iff, and_self], exact ((tendsto_order.1 (ennreal.tendsto_tsum_compl_at_top_zero I.ne)).2 _ this).exists }, - choose! y hy using h, -- main property: the points `z` of `s` which are not covered by `u` are contained in the -- enlargements of the elements not in `w`. - have M : (s \ ⋃ (a : set α) (H : a ∈ u), a) ∩ ball x (r x) - ⊆ ⋃ (a : {a // a ∉ w}), closed_ball (y a) (3 * diam (a : set α)), + have M : (s \ ⋃ a ∈ u, B a) ∩ ball x (R x) + ⊆ ⋃ (a : {a // a ∉ w}), closed_ball (c a) (3 * r a), { assume z hz, - set k := ⋃ (a : v) (ha : a ∈ w), (↑a : set α) with hk, + set k := ⋃ (a : v) (ha : a ∈ w), B a with hk, have k_closed : is_closed k := is_closed_bUnion w.finite_to_set (λ i hi, h't _ (ut (vu i.2))), have z_notmem_k : z ∉ k, { simp only [not_exists, exists_prop, mem_Union, mem_sep_iff, forall_exists_index, set_coe.exists, not_and, exists_and_distrib_right, subtype.coe_mk], assume b hbv h'b h'z, - have : z ∈ (s \ ⋃ (a : set α) (H : a ∈ u), a) ∩ (⋃ (a : set α) (H : a ∈ u), a) := + have : z ∈ (s \ ⋃ a ∈ u, B a) ∩ (⋃ a ∈ u, B a) := mem_inter (mem_of_mem_inter_left hz) (mem_bUnion (vu hbv) h'z), simpa only [diff_inter_self] }, -- since the elements of `w` are closed and finitely many, one can find a small ball around `z` -- not intersecting them - have : ball x (r x) \ k ∈ 𝓝 z, + have : ball x (R x) \ k ∈ 𝓝 z, { apply is_open.mem_nhds (is_open_ball.sdiff k_closed) _, exact (mem_diff _).2 ⟨mem_of_mem_inter_right hz, z_notmem_k⟩ }, - obtain ⟨d, dpos, hd⟩ : ∃ (d : ℝ) (dpos : 0 < d), closed_ball z d ⊆ ball x (r x) \ k := + obtain ⟨d, dpos, hd⟩ : ∃ (d : ℝ) (dpos : 0 < d), closed_ball z d ⊆ ball x (R x) \ k := nhds_basis_closed_ball.mem_iff.1 this, -- choose an element `a` of the family `t` contained in this small ball - obtain ⟨a, hat, za, ad⟩ : ∃ a ∈ t, z ∈ a ∧ a ⊆ closed_ball z d := - hf z ((mem_diff _).1 (mem_of_mem_inter_left hz)).1 d dpos, - have ax : a ⊆ ball x (r x) := ad.trans (hd.trans (diff_subset (ball x (r x)) k)), + obtain ⟨a, hat, ad, rfl⟩ : ∃ a ∈ t, r a ≤ min d (R z) ∧ c a = z, + from hf z ((mem_diff _).1 (mem_of_mem_inter_left hz)).1 (min d (R z)) (lt_min dpos (hR0 z)), + have ax : B a ⊆ ball x (R x), + { refine (hB a hat).trans _, + refine subset.trans _ (hd.trans (diff_subset (ball x (R x)) k)), + exact closed_ball_subset_closed_ball (ad.trans (min_le_left _ _)), }, -- it intersects an element `b` of `u` with comparable diameter, by definition of `u` - obtain ⟨b, bu, ab, bdiam⟩ : ∃ (b : set α) (H : b ∈ u), (a ∩ b).nonempty ∧ diam a ≤ 2 * diam b := - hu a ⟨hat, ⟨x, ax.trans ball_subset_closed_ball⟩⟩, + obtain ⟨b, bu, ab, bdiam⟩ : ∃ b ∈ u, (B a ∩ B b).nonempty ∧ r a ≤ 2 * r b, + from hu a ⟨hat, ad.trans (min_le_right _ _)⟩, have bv : b ∈ v, { refine ⟨bu, ab.mono _⟩, rw inter_comm, @@ -402,110 +361,82 @@ begin -- contrary to `b` have b'_notmem_w : b' ∉ w, { assume b'w, - have b'k : (↑b' : set α) ⊆ k := finset.subset_set_bUnion_of_mem b'w, - have : ((ball x (r x) \ k) ∩ k).nonempty := ab.mono (inter_subset_inter (ad.trans hd) b'k), + have b'k : B b' ⊆ k, from @finset.subset_set_bUnion_of_mem _ _ _ (λ (y : v), B y) _ b'w, + have : ((ball x (R x) \ k) ∩ k).nonempty, + { apply ab.mono (inter_subset_inter _ b'k), + refine ((hB _ hat).trans _).trans hd, + exact (closed_ball_subset_closed_ball (ad.trans (min_le_left _ _))) }, simpa only [diff_inter_self, not_nonempty_empty] }, let b'' : {a // a ∉ w} := ⟨b', b'_notmem_w⟩, -- since `a` and `b` have comparable diameters, it follows that `z` belongs to the -- enlargement of `b` - have zb : z ∈ closed_ball (y b) (3 * diam b), + have zb : c a ∈ closed_ball (c b) (3 * r b), { rcases ab with ⟨e, ⟨ea, eb⟩⟩, - have A : dist z e ≤ diam a := dist_le_diam_of_mem (bounded_closed_ball.mono ad) za ea, - have B : dist e (y b) ≤ diam b, - { rcases (ut' bu).2 with ⟨c, hc⟩, - apply dist_le_diam_of_mem (bounded_closed_ball.mono hc) eb (hy b (ut bu)).1 }, + have A : dist (c a) e ≤ r a, from mem_closed_ball'.1 (hB a hat ea), + have B : dist e (c b) ≤ r b, from mem_closed_ball.1 (hB b (ut bu) eb), simp only [mem_closed_ball], - linarith [dist_triangle z e (y b)] }, - suffices H : closed_ball (y (↑b'')) (3 * diam (↑b'' : set α)) - ⊆ ⋃ (a : {a // a ∉ w}), closed_ball (y (↑a)) (3 * diam (↑a : set α)), from H zb, - exact subset_Union (λ (a : {a // a ∉ w}), closed_ball (y (↑a)) (3 * diam (↑a : set α))) b'' }, + linarith [dist_triangle (c a) e (c b)] }, + suffices H : closed_ball (c b'') (3 * r b'') + ⊆ ⋃ (a : {a // a ∉ w}), closed_ball (c a) (3 * r a), from H zb, + exact subset_Union (λ (a : {a // a ∉ w}), closed_ball (c a) (3 * r a)) b'' }, -- now that we have proved our main inclusion, we can use it to estimate the measure of the points -- in `ball x (r x)` not covered by `u`. haveI : encodable v := (u_count.mono vu).to_encodable, - calc μ ((s \ ⋃ (a : set α) (H : a ∈ u), a) ∩ ball x (r x)) - ≤ μ (⋃ (a : {a // a ∉ w}), closed_ball (y (↑a)) (3 * diam (↑a : set α))) : measure_mono M - ... ≤ ∑' (a : {a // a ∉ w}), μ (closed_ball (y (↑a)) (3 * diam (↑a : set α))) : + calc μ ((s \ ⋃ a ∈ u, B a) ∩ ball x (R x)) + ≤ μ (⋃ (a : {a // a ∉ w}), closed_ball (c a) (3 * r a)) : measure_mono M + ... ≤ ∑' (a : {a // a ∉ w}), μ (closed_ball (c a) (3 * r a)) : measure_Union_le _ - ... ≤ ∑' (a : {a // a ∉ w}), C * μ (↑a) : ennreal.tsum_le_tsum (λ a, (hy a (ut (vu a.1.2))).2) - ... = C * ∑' (a : {a // a ∉ w}), μ (↑a) : ennreal.tsum_mul_left + ... ≤ ∑' (a : {a // a ∉ w}), C * μ (B a) : ennreal.tsum_le_tsum (λ a, μB a (ut (vu a.1.2))) + ... = C * ∑' (a : {a // a ∉ w}), μ (B a) : ennreal.tsum_mul_left ... ≤ C * (ε / C) : ennreal.mul_le_mul le_rfl hw.le ... ≤ ε : ennreal.mul_div_le end /-- Assume that around every point there are arbitrarily small scales at which the measure is -doubling. Then the set of closed sets `a` with nonempty interior covering a fixed proportion `1/C` -of the ball `closed_ball x (3 * diam a)` forms a Vitali family. This is essentially a restatement -of the measurable Vitali theorem. -/ +doubling. Then the set of closed sets `a` with nonempty interior contained in `closed_ball x r` and +covering a fixed proportion `1/C` of the ball `closed_ball x (3 * r)` forms a Vitali family. +This is essentially a restatement of the measurable Vitali theorem. -/ protected def vitali_family [metric_space α] [measurable_space α] [opens_measurable_space α] [second_countable_topology α] (μ : measure α) [is_locally_finite_measure μ] (C : ℝ≥0) - (h : ∀ x (ε > 0), ∃ r ∈ Ioc (0 : ℝ) ε, μ (closed_ball x (6 * r)) ≤ C * μ (closed_ball x r)) : + (h : ∀ x, ∃ᶠ r in 𝓝[>] 0, μ (closed_ball x (3 * r)) ≤ C * μ (closed_ball x r)) : vitali_family μ := -{ sets_at := λ x, {a | x ∈ a ∧ is_closed a ∧ (interior a).nonempty ∧ - μ (closed_ball x (3 * diam a)) ≤ C * μ a}, - measurable_set' := λ x a ha, ha.2.1.measurable_set, - nonempty_interior := λ x a ha, ha.2.2.1, +{ sets_at := λ x, {a | is_closed a ∧ (interior a).nonempty ∧ ∃ r, (a ⊆ closed_ball x r ∧ + μ (closed_ball x (3 * r)) ≤ C * μ a)}, + measurable_set' := λ x a ha, ha.1.measurable_set, + nonempty_interior := λ x a ha, ha.2.1, nontrivial := λ x ε εpos, begin - obtain ⟨r, ⟨rpos, rε⟩, μr⟩ : ∃ r ∈ Ioc (0 : ℝ) ε, - μ (closed_ball x (6 * r)) ≤ C * μ (closed_ball x r) := h x ε εpos, - refine ⟨closed_ball x r, ⟨_, is_closed_ball, _, _⟩, closed_ball_subset_closed_ball rε⟩, - { simp only [rpos.le, mem_closed_ball, dist_self] }, - { exact (nonempty_ball.2 rpos).mono (ball_subset_interior_closed_ball) }, - { apply le_trans (measure_mono (closed_ball_subset_closed_ball _)) μr, - have : diam (closed_ball x r) ≤ 2 * r := diam_closed_ball rpos.le, - linarith } + obtain ⟨r, μr, rpos, rε⟩ : ∃ r, + μ (closed_ball x (3 * r)) ≤ C * μ (closed_ball x r) ∧ r ∈ Ioc (0 : ℝ) ε := + ((h x).and_eventually (Ioc_mem_nhds_within_Ioi ⟨le_rfl, εpos⟩)).exists, + refine ⟨closed_ball x r, ⟨is_closed_ball, _, ⟨r, subset.rfl, μr⟩⟩, + closed_ball_subset_closed_ball rε⟩, + exact (nonempty_ball.2 rpos).mono (ball_subset_interior_closed_ball) end, covering := begin assume s f fsubset ffine, - rcases eq_empty_or_nonempty s with rfl|H, - { exact ⟨∅, λ _, ∅, by simp, by simp⟩ }, - haveI : inhabited α, { choose x hx using H, exact ⟨x⟩ }, - let t := ⋃ (x ∈ s), f x, - have A₁ : ∀ x ∈ s, ∀ (ε : ℝ), 0 < ε → (∃ a ∈ t, x ∈ a ∧ a ⊆ closed_ball x ε), + let t : set (ℝ × α × set α) := + {p | p.2.2 ⊆ closed_ball p.2.1 p.1 ∧ μ (closed_ball p.2.1 (3 * p.1)) ≤ C * μ p.2.2 + ∧ (interior p.2.2).nonempty ∧ is_closed p.2.2 ∧ p.2.2 ∈ f p.2.1 ∧ p.2.1 ∈ s}, + have A : ∀ x ∈ s, ∀ (ε : ℝ), ε > 0 → (∃ (p : ℝ × α × set α) (Hp : p ∈ t), p.1 ≤ ε ∧ p.2.1 = x), { assume x xs ε εpos, - rcases ffine x xs ε εpos with ⟨a, xa, hax⟩, - exact ⟨a, mem_bUnion xs xa, (fsubset x xs xa).1, hax⟩ }, - have A₂ : ∀ a ∈ t, (interior a).nonempty, - { rintros a ha, - rcases mem_Union₂.1 ha with ⟨x, xs, xa⟩, - exact (fsubset x xs xa).2.2.1 }, - have A₃ : ∀ a ∈ t, is_closed a, - { rintros a ha, - rcases mem_Union₂.1 ha with ⟨x, xs, xa⟩, - exact (fsubset x xs xa).2.1 }, - have A₄ : ∀ a ∈ t, ∃ x ∈ a, μ (closed_ball x (3 * diam a)) ≤ C * μ a, - { rintros a ha, - rcases mem_Union₂.1 ha with ⟨x, xs, xa⟩, - exact ⟨x, (fsubset x xs xa).1, (fsubset x xs xa).2.2.2⟩ }, - obtain ⟨u, ut, u_count, u_disj, μu⟩ : - ∃ u ⊆ t, u.countable ∧ u.pairwise disjoint ∧ μ (s \ ⋃ a ∈ u, a) = 0 := - exists_disjoint_covering_ae μ s t A₁ A₂ A₃ C A₄, - have : ∀ a ∈ u, ∃ x ∈ s, a ∈ f x := λ a ha, mem_Union₂.1 (ut ha), - choose! x hx using this, - have inj_on_x : inj_on x u, - { assume a ha b hb hab, - have A : (a ∩ b).nonempty, - { refine ⟨x a, mem_inter ((fsubset _ (hx a ha).1 (hx a ha).2).1) _⟩, - rw hab, - exact (fsubset _ (hx b hb).1 (hx b hb).2).1 }, - contrapose A, - have : disjoint a b := u_disj ha hb A, - simpa only [← not_disjoint_iff_nonempty_inter] }, - refine ⟨x '' u, function.inv_fun_on x u, _, _, _, _⟩, - { assume y hy, - rcases (mem_image _ _ _).1 hy with ⟨a, au, rfl⟩, - exact (hx a au).1 }, - { rw [inj_on_x.pairwise_disjoint_image], - assume a ha b hb hab, - simp only [function.on_fun, inj_on_x.left_inv_on_inv_fun_on ha, - inj_on_x.left_inv_on_inv_fun_on hb, (∘)], - exact u_disj ha hb hab }, - { assume y hy, - rcases (mem_image _ _ _).1 hy with ⟨a, ha, rfl⟩, - rw inj_on_x.left_inv_on_inv_fun_on ha, - exact (hx a ha).2 }, - { rw [bUnion_image], - convert μu using 3, - exact Union₂_congr (λ a ha, inj_on_x.left_inv_on_inv_fun_on ha) } + rcases ffine x xs ε εpos with ⟨a, ha, h'a⟩, + rcases fsubset x xs ha with ⟨a_closed, a_int, ⟨r, ar, μr⟩⟩, + refine ⟨⟨min r ε, x, a⟩, ⟨_, _, a_int, a_closed, ha, xs⟩, min_le_right _ _, rfl⟩, + { rcases min_cases r ε with h'|h'; rwa h'.1 }, + { apply le_trans (measure_mono (closed_ball_subset_closed_ball _)) μr, + exact mul_le_mul_of_nonneg_left (min_le_left _ _) zero_le_three } }, + rcases exists_disjoint_covering_ae μ s t C (λ p, p.1) (λ p, p.2.1) (λ p, p.2.2) (λ p hp, hp.1) + (λ p hp, hp.2.1) (λ p hp, hp.2.2.1) (λ p hp, hp.2.2.2.1) A + with ⟨t', t't, t'_count, t'_disj, μt'⟩, + refine ⟨(λ (p : ℝ × α × set α), p.2) '' t', _, _, _, _⟩, + { rintros - ⟨q, hq, rfl⟩, + exact (t't hq).2.2.2.2.2 }, + { rintros p ⟨q, hq, rfl⟩ p' ⟨q', hq', rfl⟩ hqq', + exact t'_disj hq hq' (ne_of_apply_ne _ hqq') }, + { rintros - ⟨q, hq, rfl⟩, + exact (t't hq).2.2.2.2.1 }, + { convert μt' using 3, + rw bUnion_image } end } end vitali diff --git a/src/measure_theory/covering/vitali_family.lean b/src/measure_theory/covering/vitali_family.lean index 852e5353ae45f..2aad85dc9503d 100644 --- a/src/measure_theory/covering/vitali_family.lean +++ b/src/measure_theory/covering/vitali_family.lean @@ -67,9 +67,9 @@ structure vitali_family {m : measurable_space α} (μ : measure α) := (nonempty_interior : ∀ (x : α), ∀ (y : set α), y ∈ sets_at x → (interior y).nonempty) (nontrivial : ∀ (x : α) (ε > (0 : ℝ)), ∃ y ∈ sets_at x, y ⊆ closed_ball x ε) (covering : ∀ (s : set α) (f : Π (x : α), set (set α)), (∀ x ∈ s, f x ⊆ sets_at x) → - (∀ (x ∈ s) (ε > (0 : ℝ)), ∃ a ∈ f x, a ⊆ closed_ball x ε) → - ∃ (t : set α) (u : α → set α), t ⊆ s ∧ t.pairwise_disjoint u ∧ (∀ x ∈ t, u x ∈ f x) - ∧ μ (s \ ⋃ x ∈ t, u x) = 0) + (∀ (x ∈ s) (ε > (0 : ℝ)), ∃ a ∈ f x, a ⊆ closed_ball x ε) → ∃ (t : set (α × set α)), + (∀ (p : α × set α), p ∈ t → p.1 ∈ s) ∧ t.pairwise_disjoint (λ p, p.2) ∧ + (∀ (p : α × set α), p ∈ t → p.2 ∈ f p.1) ∧ μ (s \ ⋃ (p : α × set α) (hp : p ∈ t), p.2) = 0) namespace vitali_family @@ -85,8 +85,8 @@ def mono (v : vitali_family μ) (ν : measure α) (hν : ν ≪ μ) : nonempty_interior := v.nonempty_interior, nontrivial := v.nontrivial, covering := λ s f h h', begin - rcases v.covering s f h h' with ⟨t, u, ts, u_disj, uf, μu⟩, - exact ⟨t, u, ts, u_disj, uf, hν μu⟩ + rcases v.covering s f h h' with ⟨t, ts, disj, mem_f, hμ⟩, + exact ⟨t, ts, disj, mem_f, hν hμ⟩ end } /-- Given a Vitali family `v` for a measure `μ`, a family `f` is a fine subfamily on a set `s` if @@ -102,53 +102,56 @@ variables {v : vitali_family μ} {f : α → set (set α)} {s : set α} (h : v.f include h theorem exists_disjoint_covering_ae : - ∃ (t : set α) (u : α → set α), t ⊆ s ∧ t.pairwise_disjoint u ∧ - (∀ x ∈ t, u x ∈ v.sets_at x ∩ f x) ∧ μ (s \ ⋃ x ∈ t, u x) = 0 := + ∃ (t : set (α × set α)), (∀ (p : α × set α), p ∈ t → p.1 ∈ s) ∧ + t.pairwise_disjoint (λ p, p.2) ∧ + (∀ (p : α × set α), p ∈ t → p.2 ∈ v.sets_at p.1 ∩ f p.1) + ∧ μ (s \ ⋃ (p : α × set α) (hp : p ∈ t), p.2) = 0 := v.covering s (λ x, v.sets_at x ∩ f x) (λ x hx, inter_subset_left _ _) h -/-- Given `h : v.fine_subfamily_on f s`, then `h.index` is a subset of `s` parametrizing a disjoint +/-- Given `h : v.fine_subfamily_on f s`, then `h.index` is a set parametrizing a disjoint covering of almost every `s`. -/ -protected def index : set α := +protected def index : set (α × set α) := h.exists_disjoint_covering_ae.some -/-- Given `h : v.fine_subfamily_on f s`, then `h.covering x` is a set in the family, -for `x ∈ h.index`, such that these sets form a disjoint covering of almost every `s`. -/ -protected def covering : α → set α := -h.exists_disjoint_covering_ae.some_spec.some +/-- Given `h : v.fine_subfamily_on f s`, then `h.covering p` is a set in the family, +for `p ∈ h.index`, such that these sets form a disjoint covering of almost every `s`. -/ +@[nolint unused_arguments] protected def covering : α × set α → set α := +λ p, p.2 -lemma index_subset : h.index ⊆ s := -h.exists_disjoint_covering_ae.some_spec.some_spec.1 +lemma index_subset : ∀ (p : α × set α), p ∈ h.index → p.1 ∈ s := +h.exists_disjoint_covering_ae.some_spec.1 lemma covering_disjoint : h.index.pairwise_disjoint h.covering := -h.exists_disjoint_covering_ae.some_spec.some_spec.2.1 +h.exists_disjoint_covering_ae.some_spec.2.1 lemma covering_disjoint_subtype : pairwise (disjoint on (λ x : h.index, h.covering x)) := (pairwise_subtype_iff_pairwise_set _ _).2 h.covering_disjoint -lemma covering_mem {x : α} (hx : x ∈ h.index) : h.covering x ∈ f x := -(h.exists_disjoint_covering_ae.some_spec.some_spec.2.2.1 x hx).2 +lemma covering_mem {p : α × set α} (hp : p ∈ h.index) : h.covering p ∈ f p.1 := +(h.exists_disjoint_covering_ae.some_spec.2.2.1 p hp).2 -lemma covering_mem_family {x : α} (hx : x ∈ h.index) : h.covering x ∈ v.sets_at x := -(h.exists_disjoint_covering_ae.some_spec.some_spec.2.2.1 x hx).1 +lemma covering_mem_family {p : α × set α} (hp : p ∈ h.index) : h.covering p ∈ v.sets_at p.1 := +(h.exists_disjoint_covering_ae.some_spec.2.2.1 p hp).1 -lemma measure_diff_bUnion : μ (s \ ⋃ x ∈ h.index, h.covering x) = 0 := -h.exists_disjoint_covering_ae.some_spec.some_spec.2.2.2 +lemma measure_diff_bUnion : μ (s \ ⋃ p ∈ h.index, h.covering p) = 0 := +h.exists_disjoint_covering_ae.some_spec.2.2.2 lemma index_countable [second_countable_topology α] : h.index.countable := h.covering_disjoint.countable_of_nonempty_interior (λ x hx, v.nonempty_interior _ _ (h.covering_mem_family hx)) -protected lemma measurable_set_u {x : α} (hx : x ∈ h.index) : measurable_set (h.covering x) := -v.measurable_set' x _ (h.covering_mem_family hx) +protected lemma measurable_set_u {p : α × set α} (hp : p ∈ h.index) : + measurable_set (h.covering p) := +v.measurable_set' p.1 _ (h.covering_mem_family hp) lemma measure_le_tsum_of_absolutely_continuous [second_countable_topology α] {ρ : measure α} (hρ : ρ ≪ μ) : - ρ s ≤ ∑' (x : h.index), ρ (h.covering x) := -calc ρ s ≤ ρ ((s \ ⋃ (x ∈ h.index), h.covering x) ∪ (⋃ (x ∈ h.index), h.covering x)) : + ρ s ≤ ∑' (p : h.index), ρ (h.covering p) := +calc ρ s ≤ ρ ((s \ ⋃ (p ∈ h.index), h.covering p) ∪ (⋃ (p ∈ h.index), h.covering p)) : measure_mono (by simp only [subset_union_left, diff_union_self]) - ... ≤ ρ (s \ ⋃ (x ∈ h.index), h.covering x) + ρ (⋃ (x ∈ h.index), h.covering x) : + ... ≤ ρ (s \ ⋃ (p ∈ h.index), h.covering p) + ρ (⋃ (p ∈ h.index), h.covering p) : measure_union_le _ _ - ... = ∑' (x : h.index), ρ (h.covering x) : by rw [hρ h.measure_diff_bUnion, + ... = ∑' (p : h.index), ρ (h.covering p) : by rw [hρ h.measure_diff_bUnion, measure_bUnion h.index_countable h.covering_disjoint (λ x hx, h.measurable_set_u hx), zero_add] @@ -158,6 +161,34 @@ h.measure_le_tsum_of_absolutely_continuous measure.absolutely_continuous.rfl end fine_subfamily_on +/-- One can enlarge a Vitali family by adding to the sets `f x` at `x` all sets which are not +contained in a `δ`-neighborhood on `x`. This does not change the local filter at a point, but it +can be convenient to get a nicer global behavior. -/ +def enlarge (v : vitali_family μ) (δ : ℝ) (δpos : 0 < δ) : vitali_family μ := +{ sets_at := λ x, v.sets_at x ∪ + {a | measurable_set a ∧ (interior a).nonempty ∧ ¬(a ⊆ closed_ball x δ)}, + measurable_set' := λ x a ha, by { cases ha, exacts [v.measurable_set' _ _ ha, ha.1] }, + nonempty_interior := λ x a ha, by { cases ha, exacts [v.nonempty_interior _ _ ha, ha.2.1] }, + nontrivial := begin + assume x ε εpos, + rcases v.nontrivial x ε εpos with ⟨a, ha, h'a⟩, + exact ⟨a, mem_union_left _ ha, h'a⟩, + end, + covering := begin + assume s f fset ffine, + let g : α → set (set α) := λ x, f x ∩ v.sets_at x, + have : ∀ x ∈ s, ∀ (ε : ℝ), ε > 0 → (∃ (a : set α) (H : a ∈ g x), a ⊆ closed_ball x ε), + { assume x hx ε εpos, + obtain ⟨a, af, ha⟩ : ∃ a ∈ f x, a ⊆ closed_ball x (min ε δ), + from ffine x hx (min ε δ) (lt_min εpos δpos), + rcases fset x hx af with h'a|h'a, + { exact ⟨a, ⟨af, h'a⟩, ha.trans (closed_ball_subset_closed_ball (min_le_left _ _))⟩ }, + { refine false.elim (h'a.2.2 _), + exact ha.trans (closed_ball_subset_closed_ball (min_le_right _ _)) } }, + rcases v.covering s g (λ x hx, inter_subset_right _ _) this with ⟨t, ts, tdisj, tg, μt⟩, + exact ⟨t, ts, tdisj, λ p hp, (tg p hp).1, μt⟩, + end } + variable (v : vitali_family μ) include v diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index a5d2c5fda17de..53ef1a5ebe410 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -86,6 +86,14 @@ calc (⨅ z ∈ s, edist x z) ≤ ⨅ z ∈ s, edist y z + edist x y : lemma inf_edist_le_edist_add_inf_edist : inf_edist x s ≤ edist x y + inf_edist y s := by { rw add_comm, exact inf_edist_le_inf_edist_add_edist } +lemma edist_le_inf_edist_add_ediam (hy : y ∈ s) : edist x y ≤ inf_edist x s + diam s := +begin + simp_rw [inf_edist, ennreal.infi_add], + refine le_infi (λ i, le_infi (λ hi, _)), + calc edist x y ≤ edist x i + edist i y : edist_triangle _ _ _ + ... ≤ edist x i + diam s : add_le_add le_rfl (edist_le_diam_of_mem hi hy) +end + /-- The edist to a set depends continuously on the point -/ @[continuity] lemma continuous_inf_edist : continuous (λx, inf_edist x s) := @@ -475,6 +483,17 @@ lemma disjoint_closed_ball_of_lt_inf_dist {r : ℝ} (h : r < inf_dist x s) : disjoint (closed_ball x r) s := disjoint_ball_inf_dist.mono_left $ closed_ball_subset_ball h +lemma dist_le_inf_dist_add_diam (hs : bounded s) (hy : y ∈ s) : dist x y ≤ inf_dist x s + diam s := +begin + have A : inf_edist x s ≠ ∞, from inf_edist_ne_top ⟨y, hy⟩, + have B : emetric.diam s ≠ ∞, from hs.ediam_ne_top, + rw [inf_dist, diam, ← ennreal.to_real_add A B, dist_edist], + apply (ennreal.to_real_le_to_real _ _).2, + { exact edist_le_inf_edist_add_ediam hy }, + { rw edist_dist, exact ennreal.of_real_ne_top }, + { exact ennreal.add_ne_top.2 ⟨A, B⟩ } +end + variable (s) /-- The minimal distance to a set is Lipschitz in point with constant 1 -/ From fdb930f36205c4febd263c2a4a93edaaf5f83454 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 8 Oct 2022 05:41:13 +0000 Subject: [PATCH 629/828] feat(data/bool/basic): add `bool.cond_eq_ite` (#16834) --- src/data/bool/basic.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/data/bool/basic.lean b/src/data/bool/basic.lean index 1997c7369338d..42b46ddfbcfcc 100644 --- a/src/data/bool/basic.lean +++ b/src/data/bool/basic.lean @@ -91,9 +91,11 @@ decidable_of_decidable_of_iff or.decidable exists_bool.symm @[simp] theorem cond_tt {α} (t e : α) : cond tt t e = t := rfl +theorem cond_eq_ite {α} (b : bool) (t e : α) : cond b t e = if b then t else e := by cases b; simp + @[simp] theorem cond_to_bool {α} (p : Prop) [decidable p] (t e : α) : cond (to_bool p) t e = if p then t else e := -by by_cases p; simp * +by simp [cond_eq_ite] @[simp] theorem cond_bnot {α} (b : bool) (t e : α) : cond (!b) t e = cond b e t := by cases b; refl From f5a600f8102c8bfdbd22781968a20a539304c1b4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 8 Oct 2022 09:14:22 +0000 Subject: [PATCH 630/828] feat(algebra/order/ring): Non-cancellative ordered semirings (#16172) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ## Historical background This tackles a problem that we have had for six years (#2697 for its move to mathlib, before it was in core): `ordered_semiring` assumes that addition and multiplication are strictly monotone. This led to weirdness within the algebraic order hierarchy: * `ennreal`/`ereal`/`enat` needed custom lemmas (`∞ + 0 = ∞ = ∞ + 1`, `1 * ∞ = ∞ = 2 * ∞`). * A `canonically_ordered_comm_semiring` was not an `ordered_comm_semiring`! ## What this PR does This PR solves the problem minimally by renaming `ordered_(comm_)semiring` to `strict_ordered_(comm_)semiring` and adding a new class `ordered_(comm_)semiring` that doesn't assume strict monotonicity of addition and multiplication but only monotonicity: ``` class ordered_semiring (α : Type*) extends semiring α, ordered_add_comm_monoid α := (zero_le_one : (0 : α) ≤ 1) (mul_le_mul_of_nonneg_left : ∀ a b c : α, a ≤ b → 0 ≤ c → c * a ≤ c * b) (mul_le_mul_of_nonneg_right : ∀ a b c : α, a ≤ b → 0 ≤ c → a * c ≤ b * c) class ordered_comm_semiring (α : Type*) extends ordered_semiring α, comm_semiring α class ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α := (zero_le_one : 0 ≤ (1 : α)) (mul_nonneg : ∀ a b : α, 0 ≤ a → 0 ≤ b → 0 ≤ a * b) class ordered_comm_ring (α : Type u) extends ordered_ring α, comm_ring α class strict_ordered_semiring (α : Type*) extends semiring α, ordered_cancel_add_comm_monoid α := (zero_le_one : (0 : α) ≤ 1) (mul_lt_mul_of_pos_left : ∀ a b c : α, a < b → 0 < c → c * a < c * b) (mul_lt_mul_of_pos_right : ∀ a b c : α, a < b → 0 < c → a * c < b * c) class strict_ordered_comm_semiring (α : Type*) extends strict_ordered_semiring α, comm_semiring α class strict_ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α := (zero_le_one : 0 ≤ (1 : α)) (mul_pos : ∀ a b : α, 0 < a → 0 < b → 0 < a * b) class strict_ordered_comm_ring (α : Type*) extends strict_ordered_ring α, comm_ring α ``` ## Whatever happened to the `decidable` lemmas? Scrolling through the diff, you will see that only one lemma in the `decidable` namespace out of many survives this PR. Those lemmas were originally meant to avoid using choice in the proof of `nat` and `int` lemmas. The need for decidability originated from the proofs of `mul_le_mul_of_nonneg_left`/`mul_le_mul_of_nonneg_right`. ``` protected lemma decidable.mul_le_mul_of_nonneg_left [@decidable_rel α (≤)] (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c * a ≤ c * b := begin by_cases ba : b ≤ a, { simp [ba.antisymm h₁] }, by_cases c0 : c ≤ 0, { simp [c0.antisymm h₂] }, exact (mul_lt_mul_of_pos_left (h₁.lt_of_not_le ba) (h₂.lt_of_not_le c0)).le, end ``` Now that these are fields to `ordered_semiring`, they are already choice-free. Instead, choice is now used in showing that an `ordered_cancel_semiring` is an `ordered_semiring`. ``` @[priority 100] -- see Note [lower instance priority] instance ordered_cancel_semiring.to_ordered_semiring : ordered_semiring α := { mul_le_mul_of_nonneg_left := λ a b c hab hc, begin obtain rfl | hab := hab.eq_or_lt, { refl }, obtain rfl | hc := hc.eq_or_lt, { simp }, { exact (mul_lt_mul_of_pos_left hab hc).le } end, mul_le_mul_of_nonneg_right := λ a b c hab hc, begin obtain rfl | hab := hab.eq_or_lt, { refl }, obtain rfl | hc := hc.eq_or_lt, { simp }, { exact (mul_lt_mul_of_pos_right hab hc).le } end, ..‹ordered_cancel_semiring α› } ``` It seems unreasonable to make that instance depend on `@decidable_rel α (≤)` even though it's only needed for the proofs. ## Other changes To have some lemmas in the generality of the new `ordered_semiring`, I needed a few new lemmas: * `bit0_mono` * `bit0_strict_mono` It was also simpler to golf `analysis.special_functions.stirling` using `positivity` rather than fixing it so I introduce the following (simple) `positivity` extensions: * `positivity_succ` * `positivity_factorial` * `positivity_asc_factorial` --- archive/imo/imo2008_q3.lean | 4 +- ...nically_ordered_comm_semiring_two_mul.lean | 2 +- src/algebra/algebra/subalgebra/basic.lean | 6 + src/algebra/big_operators/order.lean | 24 +- src/algebra/geom_sum.lean | 93 +- src/algebra/group_power/lemmas.lean | 6 +- src/algebra/group_power/order.lean | 67 +- src/algebra/ne_zero.lean | 2 +- src/algebra/order/algebra.lean | 10 - src/algebra/order/archimedean.lean | 12 +- src/algebra/order/field.lean | 5 +- src/algebra/order/invertible.lean | 2 +- src/algebra/order/monoid_lemmas.lean | 11 + src/algebra/order/nonneg.lean | 12 +- src/algebra/order/pi.lean | 22 +- src/algebra/order/positive/ring.lean | 4 +- src/algebra/order/ring.lean | 1404 ++++++++--------- src/algebra/order/ring_lemmas.lean | 4 + src/algebra/ring/prod.lean | 23 +- src/analysis/convex/basic.lean | 26 +- src/analysis/convex/between.lean | 7 +- src/analysis/convex/segment.lean | 2 +- src/analysis/convex/specific_functions.lean | 4 +- src/analysis/special_functions/stirling.lean | 55 +- src/combinatorics/set_family/kleitman.lean | 4 +- src/data/complex/basic.lean | 8 +- src/data/int/basic.lean | 22 +- src/data/list/big_operators.lean | 2 +- src/data/nat/basic.lean | 38 +- src/data/nat/cast.lean | 21 +- src/data/nat/choose/central.lean | 2 +- src/data/nat/prime.lean | 4 +- src/data/real/basic.lean | 19 +- src/data/real/ennreal.lean | 18 +- src/data/real/hyperreal.lean | 34 +- src/data/real/nnreal.lean | 2 +- src/data/set/intervals/instances.lean | 33 +- src/linear_algebra/orientation.lean | 4 +- src/linear_algebra/ray.lean | 12 +- src/number_theory/cyclotomic/rat.lean | 4 +- .../liouville/liouville_with.lean | 2 - src/order/filter/archimedean.lean | 16 +- src/order/filter/at_top_bot.lean | 13 +- src/order/filter/filter_product.lean | 136 +- src/order/filter/germ.lean | 124 +- src/ring_theory/class_group.lean | 1 - src/ring_theory/hahn_series.lean | 9 +- .../polynomial/cyclotomic/eval.lean | 6 +- src/ring_theory/polynomial/pochhammer.lean | 6 +- src/ring_theory/subsemiring/basic.lean | 32 +- src/set_theory/game/nim.lean | 15 +- src/tactic/linarith/lemmas.lean | 2 +- src/tactic/norm_num.lean | 4 +- src/topology/algebra/order/liminf_limsup.lean | 2 +- .../metric_space/hausdorff_dimension.lean | 2 +- test/instance_diamonds.lean | 6 +- test/positivity.lean | 4 +- 57 files changed, 1237 insertions(+), 1177 deletions(-) diff --git a/archive/imo/imo2008_q3.lean b/archive/imo/imo2008_q3.lean index e4c807c149fa4..07addcdb172bb 100644 --- a/archive/imo/imo2008_q3.lean +++ b/archive/imo/imo2008_q3.lean @@ -66,11 +66,11 @@ begin have hreal₃ : (k:ℝ) ^ 2 + 4 ≥ p, { assumption_mod_cast }, have hreal₅ : (k:ℝ) > 4, - { apply lt_of_pow_lt_pow 2 k.cast_nonneg, + { refine lt_of_pow_lt_pow 2 k.cast_nonneg _, linarith only [hreal₂, hreal₃] }, have hreal₆ : (k:ℝ) > sqrt (2 * n), - { apply lt_of_pow_lt_pow 2 k.cast_nonneg, + { refine lt_of_pow_lt_pow 2 k.cast_nonneg _, rw sq_sqrt (mul_nonneg zero_le_two n.cast_nonneg), linarith only [hreal₁, hreal₃, hreal₅] }, diff --git a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean index 92eb998214db3..119c01715536e 100644 --- a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean +++ b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean @@ -132,7 +132,7 @@ lemma mul_lt_mul_of_pos_left : ∀ (a b c : ℕ × zmod 2), a < b → 0 < c → lemma mul_lt_mul_of_pos_right : ∀ (a b c : ℕ × zmod 2), a < b → 0 < c → a * c < b * c := λ a b c ab c0, lt_def.mpr ((mul_lt_mul_right (lt_def.mp c0)).mpr (lt_def.mp ab)) -instance ocsN2 : ordered_comm_semiring (ℕ × zmod 2) := +instance socsN2 : strict_ordered_comm_semiring (ℕ × zmod 2) := { add_le_add_left := add_le_add_left, le_of_add_le_add_left := le_of_add_le_add_left, zero_le_one := zero_le_one, diff --git a/src/algebra/algebra/subalgebra/basic.lean b/src/algebra/algebra/subalgebra/basic.lean index 620ae80cf9b58..4179a9e59f0a1 100644 --- a/src/algebra/algebra/subalgebra/basic.lean +++ b/src/algebra/algebra/subalgebra/basic.lean @@ -188,9 +188,15 @@ instance to_comm_ring {R A} instance to_ordered_semiring {R A} [comm_semiring R] [ordered_semiring A] [algebra R A] (S : subalgebra R A) : ordered_semiring S := S.to_subsemiring.to_ordered_semiring +instance to_strict_ordered_semiring {R A} + [comm_semiring R] [strict_ordered_semiring A] [algebra R A] (S : subalgebra R A) : + strict_ordered_semiring S := S.to_subsemiring.to_strict_ordered_semiring instance to_ordered_comm_semiring {R A} [comm_semiring R] [ordered_comm_semiring A] [algebra R A] (S : subalgebra R A) : ordered_comm_semiring S := S.to_subsemiring.to_ordered_comm_semiring +instance to_strict_ordered_comm_semiring {R A} + [comm_semiring R] [strict_ordered_comm_semiring A] [algebra R A] (S : subalgebra R A) : + strict_ordered_comm_semiring S := S.to_subsemiring.to_strict_ordered_comm_semiring instance to_ordered_ring {R A} [comm_ring R] [ordered_ring A] [algebra R A] (S : subalgebra R A) : ordered_ring S := S.to_subring.to_ordered_ring diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 68bb20bb7cc68..f52200dda4fd0 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -411,6 +411,14 @@ lt_of_le_of_lt (by rw prod_const_one) $ prod_lt_prod_of_nonempty' hs h (∏ i in s, f i) < 1 := (prod_lt_prod_of_nonempty' hs h).trans_le (by rw prod_const_one) +@[to_additive sum_pos'] lemma one_lt_prod' (h : ∀ i ∈ s, 1 ≤ f i) (hs : ∃ i ∈ s, 1 < f i) : + 1 < (∏ i in s, f i) := +prod_const_one.symm.trans_lt $ prod_lt_prod' h hs + +@[to_additive] lemma prod_lt_one' (h : ∀ i ∈ s, f i ≤ 1) (hs : ∃ i ∈ s, f i < 1) : + ∏ i in s, f i < 1 := +prod_const_one.le.trans_lt' $ prod_lt_prod' h hs + @[to_additive] lemma prod_eq_prod_iff_of_le {f g : ι → M} (h : ∀ i ∈ s, f i ≤ g i) : ∏ i in s, f i = ∏ i in s, g i ↔ ∀ i ∈ s, f i = g i := begin @@ -464,15 +472,10 @@ section ordered_comm_semiring variables [ordered_comm_semiring R] {f g : ι → R} {s t : finset ι} open_locale classical -/- this is also true for a ordered commutative multiplicative monoid -/ +/- this is also true for a ordered commutative multiplicative monoid with zero -/ lemma prod_nonneg (h0 : ∀ i ∈ s, 0 ≤ f i) : 0 ≤ ∏ i in s, f i := prod_induction f (λ i, 0 ≤ i) (λ _ _ ha hb, mul_nonneg ha hb) zero_le_one h0 -/- this is also true for a ordered commutative multiplicative monoid -/ -lemma prod_pos [nontrivial R] (h0 : ∀ i ∈ s, 0 < f i) : - 0 < ∏ i in s, f i := -prod_induction f (λ x, 0 < x) (λ _ _ ha hb, mul_pos ha hb) zero_lt_one h0 - /-- If all `f i`, `i ∈ s`, are nonnegative and each `f i` is less than or equal to `g i`, then the product of `f i` is less than or equal to the product of `g i`. See also `finset.prod_le_prod''` for the case of an ordered commutative multiplicative monoid. -/ @@ -515,6 +518,15 @@ end end ordered_comm_semiring +section strict_ordered_comm_semiring +variables [strict_ordered_comm_semiring R] [nontrivial R] {f : ι → R} {s : finset ι} + +/- This is also true for a ordered commutative multiplicative monoid with zero -/ +lemma prod_pos (h0 : ∀ i ∈ s, 0 < f i) : 0 < ∏ i in s, f i := +prod_induction f (λ x, 0 < x) (λ _ _ ha hb, mul_pos ha hb) zero_lt_one h0 + +end strict_ordered_comm_semiring + section canonically_ordered_comm_semiring variables [canonically_ordered_comm_semiring R] {f g h : ι → R} {s : finset ι} {i : ι} diff --git a/src/algebra/geom_sum.lean b/src/algebra/geom_sum.lean index 666185670230d..f2da01a0e2d53 100644 --- a/src/algebra/geom_sum.lean +++ b/src/algebra/geom_sum.lean @@ -409,10 +409,11 @@ section order variables {n : ℕ} {x : α} -lemma geom_sum_pos [ordered_semiring α] (hx : 0 < x) (hn : n ≠ 0) : 0 < ∑ i in range n, x ^ i := -sum_pos (λ k hk, pow_pos hx _) $ nonempty_range_iff.2 hn +lemma geom_sum_pos [strict_ordered_semiring α] [nontrivial α] (hx : 0 ≤ x) (hn : n ≠ 0) : + 0 < ∑ i in range n, x ^ i := +sum_pos' (λ k hk, pow_nonneg hx _) ⟨0, mem_range.2 hn.bot_lt, by simp⟩ -lemma geom_sum_pos_and_lt_one [ordered_ring α] (hx : x < 0) (hx' : 0 < x + 1) (hn : 1 < n) : +lemma geom_sum_pos_and_lt_one [strict_ordered_ring α] (hx : x < 0) (hx' : 0 < x + 1) (hn : 1 < n) : 0 < ∑ i in range n, x ^ i ∧ ∑ i in range n, x ^ i < 1 := begin refine nat.le_induction _ _ n (show 2 ≤ n, from hn), @@ -425,7 +426,22 @@ begin (neg_lt_iff_pos_add'.2 hx') ihn.2.le, mul_neg_of_neg_of_pos hx ihn.1⟩ end -lemma geom_sum_alternating_of_lt_neg_one [ordered_ring α] (hx : x + 1 < 0) (hn : 1 < n) : +lemma geom_sum_alternating_of_le_neg_one [strict_ordered_ring α] (hx : x + 1 ≤ 0) (n : ℕ) : + if even n then ∑ i in range n, x ^ i ≤ 0 else 1 ≤ ∑ i in range n, x ^ i := +begin + have hx0 : x ≤ 0 := (le_add_of_nonneg_right zero_le_one).trans hx, + induction n with n ih, + { simp only [even_zero, geom_sum_zero, le_refl] }, + simp only [nat.even_add_one, geom_sum_succ], + split_ifs at ih, + { rw [if_neg (not_not_intro h), le_add_iff_nonneg_left], + exact mul_nonneg_of_nonpos_of_nonpos hx0 ih }, + { rw if_pos h, + refine (add_le_add_right _ _).trans hx, + simpa only [mul_one] using mul_le_mul_of_nonpos_left ih hx0 } +end + +lemma geom_sum_alternating_of_lt_neg_one [strict_ordered_ring α] (hx : x + 1 < 0) (hn : 1 < n) : if even n then ∑ i in range n, x ^ i < 0 else 1 < ∑ i in range n, x ^ i := begin have hx0 : x < 0, from ((le_add_iff_nonneg_right _).2 zero_le_one).trans_lt hx, @@ -443,6 +459,17 @@ begin exact this.trans hx } end +lemma geom_sum_pos' [linear_ordered_ring α] (hx : 0 < x + 1) (hn : n ≠ 0) : + 0 < ∑ i in range n, x ^ i := +begin + obtain _ | _ | n := n, + { cases hn rfl }, + { simp }, + obtain hx' | hx' := lt_or_le x 0, + { exact (geom_sum_pos_and_lt_one hx' hx n.one_lt_succ_succ).1 }, + { exact geom_sum_pos hx' (by simp only [nat.succ_ne_zero, ne.def, not_false_iff]) } +end + lemma odd.geom_sum_pos [linear_ordered_ring α] (h : odd n) : 0 < ∑ i in range n, x ^ i := begin @@ -455,55 +482,47 @@ begin simp only [h, if_false] at this, exact zero_lt_one.trans this }, { simp only [eq_neg_of_add_eq_zero_left hx, h, neg_one_geom_sum, if_false, zero_lt_one] }, - rcases lt_trichotomy x 0 with hx' | rfl | hx', - { exact (geom_sum_pos_and_lt_one hx' hx k.one_lt_succ_succ).1 }, - { simp only [zero_geom_sum, nat.succ_ne_zero, if_false, zero_lt_one] }, - { exact geom_sum_pos hx' (by simp only [nat.succ_ne_zero, ne.def, not_false_iff]) } + { exact geom_sum_pos' hx k.succ.succ_ne_zero } end -lemma geom_sum_pos_iff [linear_ordered_ring α] (hn : 1 < n) : +lemma geom_sum_pos_iff [linear_ordered_ring α] (hn : n ≠ 0) : 0 < ∑ i in range n, x ^ i ↔ odd n ∨ 0 < x + 1 := begin refine ⟨λ h, _, _⟩, - { suffices : ¬ 0 < x + 1 → odd n, by tauto, - intro hx, - rw not_lt at hx, - contrapose! h, - rw [←nat.even_iff_not_odd] at h, - rcases hx.eq_or_lt with hx | hx, - { rw [←neg_neg (1 : α), add_neg_eq_iff_eq_add, zero_add] at hx, - simp only [hx, neg_one_geom_sum, h, if_true] }, - apply le_of_lt, - simpa [h] using geom_sum_alternating_of_lt_neg_one hx hn }, + { rw [or_iff_not_imp_left, ←not_le, ←nat.even_iff_not_odd], + refine λ hn hx, h.not_le _, + simpa [if_pos hn] using geom_sum_alternating_of_le_neg_one hx n }, { rintro (hn | hx'), { exact hn.geom_sum_pos }, - rcases lt_trichotomy x 0 with hx | rfl | hx, - { exact (geom_sum_pos_and_lt_one hx hx' hn).1 }, - { simp only [(zero_lt_one.trans hn).ne', zero_geom_sum, if_false, zero_lt_one] }, - { exact geom_sum_pos hx (zero_lt_one.trans hn).ne' } } + { exact geom_sum_pos' hx' hn } } end -lemma geom_sum_eq_zero_iff_neg_one [linear_ordered_ring α] (hn : 1 < n) : +lemma geom_sum_ne_zero [linear_ordered_ring α] (hx : x ≠ -1) (hn : n ≠ 0) : + ∑ i in range n, x ^ i ≠ 0 := +begin + obtain _ | _ | n := n, + { cases hn rfl }, + { simp }, + rw [ne.def, eq_neg_iff_add_eq_zero, ←ne.def] at hx, + obtain h | h := hx.lt_or_lt, + { have := geom_sum_alternating_of_lt_neg_one h n.one_lt_succ_succ, + split_ifs at this, + { exact this.ne }, + { exact (zero_lt_one.trans this).ne' } }, + { exact (geom_sum_pos' h n.succ.succ_ne_zero).ne' } +end + +lemma geom_sum_eq_zero_iff_neg_one [linear_ordered_ring α] (hn : n ≠ 0) : ∑ i in range n, x ^ i = 0 ↔ x = -1 ∧ even n := begin refine ⟨λ h, _, λ ⟨h, hn⟩, by simp only [h, hn, neg_one_geom_sum, if_true]⟩, contrapose! h, - rcases eq_or_ne x (-1) with rfl | h, + obtain rfl | hx := eq_or_ne x (-1), { simp only [h rfl, neg_one_geom_sum, if_false, ne.def, not_false_iff, one_ne_zero] }, - rw [ne.def, eq_neg_iff_add_eq_zero, ←ne.def] at h, - rcases h.lt_or_lt with h | h, - { have := geom_sum_alternating_of_lt_neg_one h hn, - split_ifs at this, - { exact this.ne }, - { exact (zero_lt_one.trans this).ne' } }, - apply ne_of_gt, - rcases lt_trichotomy x 0 with h' | rfl | h', - { exact (geom_sum_pos_and_lt_one h' h hn).1 }, - { simp only [(pos_of_gt hn).ne', zero_geom_sum, if_false, zero_lt_one] }, - { exact geom_sum_pos h' (pos_of_gt hn).ne' } + { exact geom_sum_ne_zero hx hn } end -lemma geom_sum_neg_iff [linear_ordered_ring α] (hn : 1 < n) : +lemma geom_sum_neg_iff [linear_ordered_ring α] (hn : n ≠ 0) : ∑ i in range n, x ^ i < 0 ↔ even n ∧ x + 1 < 0 := by rw [← not_iff_not, not_lt, le_iff_lt_or_eq, eq_comm, or_congr (geom_sum_pos_iff hn) (geom_sum_eq_zero_iff_neg_one hn), nat.odd_iff_not_even, diff --git a/src/algebra/group_power/lemmas.lean b/src/algebra/group_power/lemmas.lean index 142103b04c506..956c140d90a50 100644 --- a/src/algebra/group_power/lemmas.lean +++ b/src/algebra/group_power/lemmas.lean @@ -423,8 +423,8 @@ end lemma neg_one_pow_eq_pow_mod_two [ring R] {n : ℕ} : (-1 : R) ^ n = (-1) ^ (n % 2) := by rw [← nat.mod_add_div n 2, pow_add, pow_mul]; simp [sq] -section ordered_semiring -variables [ordered_semiring R] {a : R} +section strict_ordered_semiring +variables [strict_ordered_semiring R] {a : R} /-- Bernoulli's inequality. This version works for semirings but requires additional hypotheses `0 ≤ a * a` and `0 ≤ (1 + a) * (1 + a)`. -/ @@ -461,7 +461,7 @@ lemma pow_le_of_le_one (h₀ : 0 ≤ a) (h₁ : a ≤ 1) {n : ℕ} (hn : n ≠ 0 lemma sq_le (h₀ : 0 ≤ a) (h₁ : a ≤ 1) : a ^ 2 ≤ a := pow_le_of_le_one h₀ h₁ two_ne_zero -end ordered_semiring +end strict_ordered_semiring section linear_ordered_semiring diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index 58d0cbda3e30a..8ec4e2e08b05a 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -222,6 +222,15 @@ begin by { rw [pow_succ _ n], exact mul_le_mul_of_nonneg_left (ih (nat.succ_ne_zero k)) h2 } end +lemma pow_le_one : ∀ (n : ℕ) (h₀ : 0 ≤ a) (h₁ : a ≤ 1), a ^ n ≤ 1 +| 0 h₀ h₁ := (pow_zero a).le +| (n + 1) h₀ h₁ := (pow_succ' a n).le.trans (mul_le_one (pow_le_one n h₀ h₁) h₀ h₁) + +lemma pow_lt_one (h₀ : 0 ≤ a) (h₁ : a < 1) : ∀ {n : ℕ} (hn : n ≠ 0), a ^ n < 1 +| 0 h := (h rfl).elim +| (n + 1) h := + by { rw pow_succ, exact mul_lt_one_of_nonneg_of_lt_one_left h₀ h₁ (pow_le_one _ h₀ h₁.le) } + theorem one_le_pow_of_one_le (H : 1 ≤ a) : ∀ (n : ℕ), 1 ≤ a ^ n | 0 := by rw [pow_zero] | (n+1) := by { rw pow_succ, simpa only [mul_one] using mul_le_mul H (one_le_pow_of_one_le n) @@ -237,6 +246,29 @@ pow_mono ha h theorem le_self_pow (ha : 1 ≤ a) (h : m ≠ 0) : a ≤ a ^ m := (pow_one a).symm.trans_le (pow_le_pow ha $ pos_iff_ne_zero.mpr h) +@[mono] lemma pow_le_pow_of_le_left {a b : R} (ha : 0 ≤ a) (hab : a ≤ b) : ∀ i : ℕ, a^i ≤ b^i +| 0 := by simp +| (k+1) := by { rw [pow_succ, pow_succ], + exact mul_le_mul hab (pow_le_pow_of_le_left _) (pow_nonneg ha _) (le_trans ha hab) } + +lemma one_lt_pow (ha : 1 < a) : ∀ {n : ℕ} (hn : n ≠ 0), 1 < a ^ n +| 0 h := (h rfl).elim +| (n + 1) h := + by { rw pow_succ, exact one_lt_mul_of_lt_of_le ha (one_le_pow_of_one_le ha.le _) } + +end ordered_semiring + +section strict_ordered_semiring +variables [strict_ordered_semiring R] {a x y : R} {n m : ℕ} + +lemma pow_lt_pow_of_lt_left (h : x < y) (hx : 0 ≤ x) : ∀ {n : ℕ}, 0 < n → x ^ n < y ^ n +| 0 hn := hn.false.elim +| (n + 1) _ := by simpa only [pow_succ'] using + mul_lt_mul_of_le_of_le' (pow_le_pow_of_le_left hx h.le _) h (pow_pos (hx.trans_lt h) _) hx + +lemma strict_mono_on_pow (hn : 0 < n) : strict_mono_on (λ x : R, x ^ n) (set.Ici 0) := +λ x hx y hy h, pow_lt_pow_of_lt_left h hx hn + lemma strict_mono_pow (h : 1 < a) : strict_mono (λ n : ℕ, a ^ n) := have 0 < a := zero_le_one.trans_lt h, strict_mono_nat_of_lt_succ $ λ n, by simpa only [one_mul, pow_succ] @@ -261,37 +293,12 @@ lemma pow_lt_pow_iff_of_lt_one (h₀ : 0 < a) (h₁ : a < 1) : a ^ m < a ^ n ↔ lemma pow_lt_pow_of_lt_one (h : 0 < a) (ha : a < 1) {i j : ℕ} (hij : i < j) : a ^ j < a ^ i := (pow_lt_pow_iff_of_lt_one h ha).2 hij -@[mono] lemma pow_le_pow_of_le_left {a b : R} (ha : 0 ≤ a) (hab : a ≤ b) : ∀ i : ℕ, a^i ≤ b^i -| 0 := by simp -| (k+1) := by { rw [pow_succ, pow_succ], - exact mul_le_mul hab (pow_le_pow_of_le_left _) (pow_nonneg ha _) (le_trans ha hab) } - -lemma pow_lt_pow_of_lt_left (h : x < y) (hx : 0 ≤ x) : ∀ {n : ℕ}, 0 < n → x ^ n < y ^ n -| 0 hn := hn.false.elim -| (n + 1) _ := by simpa only [pow_succ'] using - mul_lt_mul' (pow_le_pow_of_le_left hx h.le _) h hx (pow_pos (hx.trans_lt h) _) - -lemma pow_lt_one (h₀ : 0 ≤ a) (h₁ : a < 1) {n : ℕ} (hn : n ≠ 0) : a ^ n < 1 := -(one_pow n).subst (pow_lt_pow_of_lt_left h₁ h₀ (nat.pos_of_ne_zero hn)) - -lemma one_lt_pow (ha : 1 < a) {n : ℕ} (hn : n ≠ 0) : 1 < a ^ n := -pow_zero a ▸ pow_lt_pow ha (pos_iff_ne_zero.2 hn) - -lemma pow_le_one : ∀ (n : ℕ) (h₀ : 0 ≤ a) (h₁ : a ≤ 1), a ^ n ≤ 1 -| 0 h₀ h₁ := (pow_zero a).le -| (n + 1) h₀ h₁ := (pow_succ' a n).le.trans (mul_le_one (pow_le_one n h₀ h₁) h₀ h₁) - -lemma strict_mono_on_pow (hn : 0 < n) : strict_mono_on (λ x : R, x ^ n) (set.Ici 0) := -λ x hx y hy h, pow_lt_pow_of_lt_left h hx hn - lemma sq_pos_of_pos (ha : 0 < a) : 0 < a ^ 2 := by { rw sq, exact mul_pos ha ha } -end ordered_semiring +end strict_ordered_semiring -section ordered_ring -variables [ordered_ring R] {a : R} - -lemma sq_pos_of_neg (ha : a < 0) : 0 < a ^ 2 := by { rw sq, exact mul_pos_of_neg_of_neg ha ha } +section strict_ordered_ring +variables [strict_ordered_ring R] {a : R} lemma pow_bit0_pos_of_neg (ha : a < 0) (n : ℕ) : 0 < a ^ bit0 n := begin @@ -305,7 +312,9 @@ begin exact mul_neg_of_neg_of_pos ha (pow_bit0_pos_of_neg ha n), end -end ordered_ring +lemma sq_pos_of_neg (ha : a < 0) : 0 < a ^ 2 := pow_bit0_pos_of_neg ha _ + +end strict_ordered_ring section linear_ordered_semiring variables [linear_ordered_semiring R] {a b : R} diff --git a/src/algebra/ne_zero.lean b/src/algebra/ne_zero.lean index cb460646b7071..0380e84d878d1 100644 --- a/src/algebra/ne_zero.lean +++ b/src/algebra/ne_zero.lean @@ -62,7 +62,7 @@ instance bit0 [canonically_ordered_add_monoid M] [ne_zero x] : ne_zero (bit0 x) of_pos $ bit0_pos $ ne_zero.pos x instance bit1 [canonically_ordered_comm_semiring M] [nontrivial M] : ne_zero (bit1 x) := -⟨mt (λ h, le_iff_exists_add'.2 ⟨_, h.symm⟩) canonically_ordered_comm_semiring.zero_lt_one.not_le⟩ +⟨mt (λ h, le_iff_exists_add'.2 ⟨_, h.symm⟩) zero_lt_one.not_le⟩ instance pow [monoid_with_zero M] [no_zero_divisors M] [ne_zero x] : ne_zero (x ^ n) := ⟨pow_ne_zero n out⟩ diff --git a/src/algebra/order/algebra.lean b/src/algebra/order/algebra.lean index cc23ac67fb579..c78a8c6f615bd 100644 --- a/src/algebra/order/algebra.lean +++ b/src/algebra/order/algebra.lean @@ -45,13 +45,3 @@ begin end end ordered_algebra - -section instances - -variables {R : Type*} [linear_ordered_comm_ring R] - -instance linear_ordered_comm_ring.to_ordered_smul : ordered_smul R R := -{ smul_lt_smul_of_pos := ordered_semiring.mul_lt_mul_of_pos_left, - lt_of_smul_lt_smul_of_pos := λ a b c w₁ w₂, (mul_lt_mul_left w₂).mp w₁ } - -end instances diff --git a/src/algebra/order/archimedean.lean b/src/algebra/order/archimedean.lean index c618b3a236034..9f1e1a0ac3863 100644 --- a/src/algebra/order/archimedean.lean +++ b/src/algebra/order/archimedean.lean @@ -83,20 +83,20 @@ lemma exists_unique_add_zsmul_mem_Ioc {a : α} (ha : 0 < a) (b c : α) : end linear_ordered_add_comm_group -theorem exists_nat_gt [ordered_semiring α] [nontrivial α] [archimedean α] +theorem exists_nat_gt [strict_ordered_semiring α] [nontrivial α] [archimedean α] (x : α) : ∃ n : ℕ, x < n := let ⟨n, h⟩ := archimedean.arch x zero_lt_one in ⟨n+1, lt_of_le_of_lt (by rwa ← nsmul_one) (nat.cast_lt.2 (nat.lt_succ_self _))⟩ -theorem exists_nat_ge [ordered_semiring α] [archimedean α] (x : α) : +theorem exists_nat_ge [strict_ordered_semiring α] [archimedean α] (x : α) : ∃ n : ℕ, x ≤ n := begin nontriviality α, exact (exists_nat_gt x).imp (λ n, le_of_lt) end -lemma add_one_pow_unbounded_of_pos [ordered_semiring α] [nontrivial α] [archimedean α] +lemma add_one_pow_unbounded_of_pos [strict_ordered_semiring α] [nontrivial α] [archimedean α] (x : α) {y : α} (hy : 0 < y) : ∃ n : ℕ, x < (y + 1) ^ n := have 0 ≤ 1 + y, from add_nonneg zero_le_one hy.le, @@ -108,8 +108,8 @@ let ⟨n, h⟩ := archimedean.arch x hy in (add_nonneg zero_le_two hy.le) _ ... = (y + 1) ^ n : by rw [add_comm]⟩ -section ordered_ring -variables [ordered_ring α] [nontrivial α] [archimedean α] +section strict_ordered_ring +variables [strict_ordered_ring α] [nontrivial α] [archimedean α] lemma pow_unbounded_of_one_lt (x : α) {y : α} (hy1 : 1 < y) : ∃ n : ℕ, x < y ^ n := @@ -135,7 +135,7 @@ begin exact ⟨λ h, le_trans (int.cast_le.2 h) h₁, h₂ z⟩, end -end ordered_ring +end strict_ordered_ring section linear_ordered_ring variables [linear_ordered_ring α] [archimedean α] diff --git a/src/algebra/order/field.lean b/src/algebra/order/field.lean index b04e1c3153d83..84b6bc911076e 100644 --- a/src/algebra/order/field.lean +++ b/src/algebra/order/field.lean @@ -560,8 +560,7 @@ lemma monotone.div_const {β : Type*} [preorder β] {f : β → α} (hf : monoto {c : α} (hc : 0 ≤ c) : monotone (λ x, (f x) / c) := begin haveI := @linear_order.decidable_le α _, - simpa only [div_eq_mul_inv] - using (decidable.monotone_mul_right_of_nonneg (inv_nonneg.2 hc)).comp hf + simpa only [div_eq_mul_inv] using (monotone_mul_right_of_nonneg (inv_nonneg.2 hc)).comp hf end lemma strict_mono.div_const {β : Type*} [preorder β] {f : β → α} (hf : strict_mono f) @@ -867,7 +866,7 @@ end lemma sub_one_div_inv_le_two (a2 : 2 ≤ a) : (1 - 1 / a)⁻¹ ≤ 2 := begin -- Take inverses on both sides to obtain `2⁻¹ ≤ 1 - 1 / a` - refine (inv_le_inv_of_le (inv_pos.2 zero_lt_two) _).trans_eq (inv_inv (2 : α)), + refine (inv_le_inv_of_le (inv_pos.2 $ zero_lt_two' α) _).trans_eq (inv_inv (2 : α)), -- move `1 / a` to the left and `1 - 1 / 2 = 1 / 2` to the right to obtain `1 / a ≤ ⅟ 2` refine (le_sub_iff_add_le.2 (_ : _ + 2⁻¹ = _ ).le).trans ((sub_le_sub_iff_left 1).2 _), { -- show 2⁻¹ + 2⁻¹ = 1 diff --git a/src/algebra/order/invertible.lean b/src/algebra/order/invertible.lean index 77460f8c5b6af..bc3687733354b 100644 --- a/src/algebra/order/invertible.lean +++ b/src/algebra/order/invertible.lean @@ -32,4 +32,4 @@ by simp only [← not_le, inv_of_nonneg] @[simp] lemma inv_of_le_one [invertible a] (h : 1 ≤ a) : ⅟a ≤ 1 := by haveI := @linear_order.decidable_le α _; exact -mul_inv_of_self a ▸ decidable.le_mul_of_one_le_left (inv_of_nonneg.2 $ zero_le_one.trans h) h +mul_inv_of_self a ▸ le_mul_of_one_le_left (inv_of_nonneg.2 $ zero_le_one.trans h) h diff --git a/src/algebra/order/monoid_lemmas.lean b/src/algebra/order/monoid_lemmas.lean index dede9ac014afc..c89f86faf4a00 100644 --- a/src/algebra/order/monoid_lemmas.lean +++ b/src/algebra/order/monoid_lemmas.lean @@ -1161,3 +1161,14 @@ protected lemma mul_le_iff_le_one_left [comm_monoid α] [covariant_class α α ( by rw [mul_comm, ha.mul_le_iff_le_one_right] end mul_le_cancellable + +section bit +variables [has_add α] [preorder α] + +lemma bit0_mono [covariant_class α α (+) (≤)] [covariant_class α α (swap (+)) (≤)] : + monotone (bit0 : α → α) := λ a b h, add_le_add h h + +lemma bit0_strict_mono [covariant_class α α (+) (<)] [covariant_class α α (swap (+)) (<)] : + strict_mono (bit0 : α → α) := λ a b h, add_lt_add h h + +end bit diff --git a/src/algebra/order/nonneg.lean b/src/algebra/order/nonneg.lean index e0d7c627300e7..0cbe1e34384d1 100644 --- a/src/algebra/order/nonneg.lean +++ b/src/algebra/order/nonneg.lean @@ -179,7 +179,7 @@ instance add_monoid_with_one [ordered_semiring α] : add_monoid_with_one {x : α { nat_cast := λ n, ⟨n, nat.cast_nonneg n⟩, nat_cast_zero := by simp [nat.cast], nat_cast_succ := λ _, by simp [nat.cast]; refl, - .. nonneg.has_one, .. nonneg.ordered_cancel_add_comm_monoid } + .. nonneg.has_one, .. nonneg.ordered_add_comm_monoid } instance has_pow [ordered_semiring α] : has_pow {x : α // 0 ≤ x} ℕ := { pow := λ x n, ⟨x ^ n, pow_nonneg x.2 n⟩ } @@ -196,10 +196,20 @@ instance ordered_semiring [ordered_semiring α] : ordered_semiring {x : α // 0 subtype.coe_injective.ordered_semiring _ rfl rfl (λ x y, rfl) (λ x y, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) +instance strict_ordered_semiring [strict_ordered_semiring α] : + strict_ordered_semiring {x : α // 0 ≤ x} := +subtype.coe_injective.strict_ordered_semiring _ + rfl rfl (λ x y, rfl) (λ x y, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + instance ordered_comm_semiring [ordered_comm_semiring α] : ordered_comm_semiring {x : α // 0 ≤ x} := subtype.coe_injective.ordered_comm_semiring _ rfl rfl (λ x y, rfl) (λ x y, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) +instance strict_ordered_comm_semiring [strict_ordered_comm_semiring α] : + strict_ordered_comm_semiring {x : α // 0 ≤ x} := +subtype.coe_injective.strict_ordered_comm_semiring _ + rfl rfl (λ x y, rfl) (λ x y, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + -- These prevent noncomputable instances being found, as it does not require `linear_order` which -- is frequently non-computable. instance monoid_with_zero [ordered_semiring α] : monoid_with_zero {x : α // 0 ≤ x} := diff --git a/src/algebra/order/pi.lean b/src/algebra/order/pi.lean index 1b4e1f3e1ac5c..ef6b67d2fb7cd 100644 --- a/src/algebra/order/pi.lean +++ b/src/algebra/order/pi.lean @@ -3,9 +3,8 @@ Copyright (c) 2018 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot -/ -import algebra.group.pi -import algebra.order.group -import tactic.pi_instances +import algebra.ring.pi + /-! # Pi instances for ordered groups and monoids @@ -56,4 +55,21 @@ instance ordered_comm_group [∀ i, ordered_comm_group $ f i] : ..pi.comm_group, ..pi.ordered_comm_monoid, } +instance [Π i, ordered_semiring (f i)] : ordered_semiring (Π i, f i) := +{ add_le_add_left := λ a b hab c i, add_le_add_left (hab _) _, + zero_le_one := λ _, zero_le_one, + mul_le_mul_of_nonneg_left := λ a b c hab hc i, mul_le_mul_of_nonneg_left (hab _) $ hc _, + mul_le_mul_of_nonneg_right := λ a b c hab hc i, mul_le_mul_of_nonneg_right (hab _) $ hc _, + ..pi.semiring, ..pi.partial_order } + +instance [Π i, ordered_comm_semiring (f i)] : ordered_comm_semiring (Π i, f i) := +{ ..pi.comm_semiring, ..pi.ordered_semiring } + +instance [Π i, ordered_ring (f i)] : ordered_ring (Π i, f i) := +{ mul_nonneg := λ a b ha hb i, mul_nonneg (ha _) (hb _), + ..pi.ring, ..pi.ordered_semiring } + +instance [Π i, ordered_comm_ring (f i)] : ordered_comm_ring (Π i, f i) := +{ ..pi.comm_ring, ..pi.ordered_ring } + end pi diff --git a/src/algebra/order/positive/ring.lean b/src/algebra/order/positive/ring.lean index 3a8990a88684f..a4fbe05a928d8 100644 --- a/src/algebra/order/positive/ring.lean +++ b/src/algebra/order/positive/ring.lean @@ -71,7 +71,7 @@ instance covariant_class_add_le [add_monoid M] [partial_order M] [covariant_clas section mul -variables [ordered_semiring R] +variables [strict_ordered_semiring R] instance : has_mul {x : R // 0 < x} := ⟨λ x y, ⟨x * y, mul_pos x.2 y.2⟩⟩ @@ -96,7 +96,7 @@ end mul section mul_comm -instance [ordered_comm_semiring R] [nontrivial R] : ordered_comm_monoid {x : R // 0 < x} := +instance [strict_ordered_comm_semiring R] [nontrivial R] : ordered_comm_monoid {x : R // 0 < x} := { mul_le_mul_left := λ x y hxy c, subtype.coe_le_coe.1 $ mul_le_mul_of_nonneg_left hxy c.2.le, .. subtype.partial_order _, .. subtype.coe_injective.comm_monoid (coe : {x : R // 0 < x} → R) coe_one coe_mul coe_pow } diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 207719e1f75cb..052d9d104eb99 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -21,13 +21,18 @@ Each typeclass here comprises For short, * "`+` respects `≤`" means "monotonicity of addition" +* "`+` respects `<`" means "strict monotonicity of addition" +* "`*` respects `≤`" means "monotonicity of multiplication by a nonnegative number". * "`*` respects `<`" means "strict monotonicity of multiplication by a positive number". ## Typeclasses -* `ordered_semiring`: Semiring with a partial order such that `+` respects `≤` and `*` respects `<`. -* `ordered_comm_semiring`: Commutative semiring with a partial order such that `+` respects `≤` and - `*` respects `<`. +* `ordered_semiring`: Semiring with a partial order such that `+` and `*` respect `≤`. +* `strict_ordered_semiring`: Semiring with a partial order such that `+` and `*` respects `<`. +* `ordered_comm_semiring`: Commutative semiring with a partial order such that `+` and `*` respect + `≤`. +* `strict_ordered_comm_semiring`: Commutative semiring with a partial order such that `+` and `*` + respect `<`. * `ordered_ring`: Ring with a partial order such that `+` respects `≤` and `*` respects `<`. * `ordered_comm_ring`: Commutative ring with a partial order such that `+` respects `≤` and `*` respects `<`. @@ -41,7 +46,7 @@ For short, * `canonically_ordered_comm_semiring`: Commutative semiring with a partial order such that `+` respects `≤`, `*` respects `<`, and `a ≤ b ↔ ∃ c, b = a + c`. -and some typeclasses to define ordered rings by specifying their nonegative elements: +and some typeclasses to define ordered rings by specifying their nonnegative elements: * `nonneg_ring`: To define `ordered_ring`s. * `linear_nonneg_ring`: To define `linear_ordered_ring`s. @@ -52,13 +57,19 @@ corresponding typeclass. Here's an attempt at demystifying it. For each typeclas immediate predecessors and what conditions are added to each of them. * `ordered_semiring` + - `ordered_add_comm_monoid` & multiplication & `*` respects `≤` + - `semiring` & partial order structure & `+` respects `≤` & `*` respects `≤` +* `strict_ordered_semiring` - `ordered_cancel_add_comm_monoid` & multiplication & `*` respects `<` - - `semiring` & partial order structure & `+` respects `≤` & `*` respects `<` + - `ordered_semiring` & `+` respects `<` & `*` respects `<` * `ordered_comm_semiring` - `ordered_semiring` & commutativity of multiplication - `comm_semiring` & partial order structure & `+` respects `≤` & `*` respects `<` +* `strict_ordered_comm_semiring` + - `strict_ordered_semiring` & commutativity of multiplication + - `ordered_comm_semiring` & `+` respects `<` & `*` respects `<` * `ordered_ring` - - `ordered_semiring` & additive inverses + - `strict_ordered_semiring` & additive inverses - `ordered_add_comm_group` & multiplication & `*` respects `<` - `ring` & partial order structure & `+` respects `≤` & `*` respects `<` * `ordered_comm_ring` @@ -66,11 +77,11 @@ immediate predecessors and what conditions are added to each of them. - `ordered_comm_semiring` & additive inverses - `comm_ring` & partial order structure & `+` respects `≤` & `*` respects `<` * `linear_ordered_semiring` - - `ordered_semiring` & totality of the order & nontriviality + - `strict_ordered_semiring` & totality of the order & nontriviality + - `linear_ordered_add_comm_monoid` & multiplication & nontriviality & `*` respects `<` * `linear_ordered_comm_semiring` - - `ordered_comm_semiring` & totality of the order & nontriviality + - `strict_ordered_comm_semiring` & totality of the order & nontriviality - `linear_ordered_semiring` & commutativity of multiplication - - `linear_ordered_add_comm_monoid` & multiplication & nontriviality & `*` respects `<` * `linear_ordered_ring` - `ordered_ring` & totality of the order & nontriviality - `linear_ordered_semiring` & additive inverses @@ -82,7 +93,7 @@ immediate predecessors and what conditions are added to each of them. - `linear_ordered_comm_semiring` & additive inverses - `is_domain` & linear order structure * `canonically_ordered_comm_semiring` - - `canonically_ordered_add_monoid` & multiplication & `*` respects `<` & no zero divisors + - `canonically_ordered_add_monoid` & multiplication & `*` respects `≤` & no zero divisors - `comm_semiring` & `a ≤ b ↔ ∃ c, b = a + c` & no zero divisors ## TODO @@ -96,8 +107,10 @@ open function set_option old_structure_cmd true +open function + universe u -variable {α : Type u} +variables {α : Type u} {β : Type*} /-! Note that `order_dual` does not satisfy any of the ordered ring typeclasses due to the `zero_le_one` field. -/ @@ -108,283 +121,381 @@ lemma add_one_le_two_mul [has_le α] [semiring α] [covariant_class α α (+) ( calc a + 1 ≤ a + a : add_le_add_left a1 a ... = 2 * a : (two_mul _).symm -/-- An `ordered_semiring α` is a semiring `α` with a partial order such that -addition is monotone and multiplication by a positive number is strictly monotone. -/ +/-- An `ordered_semiring` is a semiring with a partial order such that addition is monotone and +multiplication by a nonnegative number is monotone. -/ @[protect_proj] -class ordered_semiring (α : Type u) extends semiring α, ordered_cancel_add_comm_monoid α := +class ordered_semiring (α : Type u) extends semiring α, ordered_add_comm_monoid α := +(zero_le_one : (0 : α) ≤ 1) +(mul_le_mul_of_nonneg_left : ∀ a b c : α, a ≤ b → 0 ≤ c → c * a ≤ c * b) +(mul_le_mul_of_nonneg_right : ∀ a b c : α, a ≤ b → 0 ≤ c → a * c ≤ b * c) + +/-- An `ordered_comm_semiring` is a commutative semiring with a partial order such that addition is +monotone and multiplication by a nonnegative number is monotone. -/ +@[protect_proj] +class ordered_comm_semiring (α : Type u) extends ordered_semiring α, comm_semiring α + +/-- An `ordered_ring` is a ring with a partial order such that addition is monotone and +multiplication by a nonnegative number is monotone. -/ +@[protect_proj] +class ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α := +(zero_le_one : 0 ≤ (1 : α)) +(mul_nonneg : ∀ a b : α, 0 ≤ a → 0 ≤ b → 0 ≤ a * b) + +/-- An `ordered_comm_ring` is a commutative ring with a partial order such that addition is monotone +and multiplication by a nonnegative number is monotone. -/ +@[protect_proj] +class ordered_comm_ring (α : Type u) extends ordered_ring α, comm_ring α + +/-- A `strict_ordered_semiring` is a semiring with a partial order such that addition is strictly +monotone and multiplication by a positive number is strictly monotone. -/ +@[protect_proj] +class strict_ordered_semiring (α : Type u) extends semiring α, ordered_cancel_add_comm_monoid α := (zero_le_one : (0 : α) ≤ 1) (mul_lt_mul_of_pos_left : ∀ a b c : α, a < b → 0 < c → c * a < c * b) (mul_lt_mul_of_pos_right : ∀ a b c : α, a < b → 0 < c → a * c < b * c) +/-- A `strict_ordered_comm_semiring` is a commutative semiring with a partial order such that +addition is strictly monotone and multiplication by a positive number is strictly monotone. -/ +@[protect_proj] +class strict_ordered_comm_semiring (α : Type u) extends strict_ordered_semiring α, comm_semiring α + +/-- A `strict_ordered_ring` is a ring with a partial order such that addition is strictly monotone +and multiplication by a positive number is strictly monotone. -/ +@[protect_proj] +class strict_ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α := +(zero_le_one : 0 ≤ (1 : α)) +(mul_pos : ∀ a b : α, 0 < a → 0 < b → 0 < a * b) + +/-- A `strict_ordered_comm_ring` is a commutative ring with a partial order such that addition is +strictly monotone and multiplication by a positive number is strictly monotone. -/ +@[protect_proj] +class strict_ordered_comm_ring (α : Type*) extends strict_ordered_ring α, comm_ring α + +/-- A `linear_ordered_semiring` is a nontrivial semiring with a linear order such that +addition is monotone and multiplication by a positive number is strictly monotone. -/ +/- It's not entirely clear we should assume `nontrivial` at this point; it would be reasonable to +explore changing this, but be warned that the instances involving `domain` may cause typeclass +search loops. -/ +@[protect_proj] +class linear_ordered_semiring (α : Type u) + extends strict_ordered_semiring α, linear_ordered_add_comm_monoid α, nontrivial α + +/-- A `linear_ordered_comm_semiring` is a nontrivial commutative semiring with a linear order such +that addition is monotone and multiplication by a positive number is strictly monotone. -/ +@[protect_proj, ancestor ordered_comm_semiring linear_ordered_semiring] +class linear_ordered_comm_semiring (α : Type*) + extends strict_ordered_comm_semiring α, linear_ordered_semiring α + +/-- A `linear_ordered_ring` is a ring with a linear order such that addition is monotone and +multiplication by a positive number is strictly monotone. -/ +@[protect_proj] +class linear_ordered_ring (α : Type u) extends strict_ordered_ring α, linear_order α, nontrivial α + +/-- A `linear_ordered_comm_ring` is a commutative ring with a linear order such that addition is +monotone and multiplication by a positive number is strictly monotone. -/ +@[protect_proj] +class linear_ordered_comm_ring (α : Type u) extends linear_ordered_ring α, comm_monoid α + +/-- A canonically ordered commutative semiring is an ordered, commutative semiring in which `a ≤ b` +iff there exists `c` with `b = a + c`. This is satisfied by the natural numbers, for example, but +not the integers or other ordered groups. -/ +@[protect_proj] +class canonically_ordered_comm_semiring (α : Type*) extends + canonically_ordered_add_monoid α, comm_semiring α := +(eq_zero_or_eq_zero_of_mul_eq_zero : ∀ a b : α, a * b = 0 → a = 0 ∨ b = 0) + section ordered_semiring variables [ordered_semiring α] {a b c d : α} +@[priority 100] -- see Note [lower instance priority] +instance ordered_semiring.zero_le_one_class : zero_le_one_class α := +{ ..‹ordered_semiring α› } + @[priority 200] -- see Note [lower instance priority] -instance ordered_semiring.pos_mul_strict_mono : pos_mul_strict_mono α := -⟨λ x a b h, ordered_semiring.mul_lt_mul_of_pos_left _ _ _ h x.prop⟩ +instance ordered_semiring.to_pos_mul_mono : pos_mul_mono α := +⟨λ x a b h, ordered_semiring.mul_le_mul_of_nonneg_left _ _ _ h x.2⟩ @[priority 200] -- see Note [lower instance priority] -instance ordered_semiring.mul_pos_strict_mono : mul_pos_strict_mono α := -⟨λ x a b h, ordered_semiring.mul_lt_mul_of_pos_right _ _ _ h x.prop⟩ +instance ordered_semiring.to_mul_pos_mono : mul_pos_mono α := +⟨λ x a b h, ordered_semiring.mul_le_mul_of_nonneg_right _ _ _ h x.2⟩ -@[priority 100] -- see Note [lower instance priority] -instance ordered_semiring.zero_le_one_class : zero_le_one_class α := -{ ..‹ordered_semiring α› } +lemma bit1_mono : monotone (bit1 : α → α) := λ a b h, add_le_add_right (bit0_mono h) _ -section nontrivial +@[simp] lemma pow_nonneg (H : 0 ≤ a) : ∀ (n : ℕ), 0 ≤ a ^ n +| 0 := by { rw pow_zero, exact zero_le_one} +| (n+1) := by { rw pow_succ, exact mul_nonneg H (pow_nonneg _) } -variables [nontrivial α] +lemma add_le_mul_two_add (a2 : 2 ≤ a) (b0 : 0 ≤ b) : a + (2 + b) ≤ a * (2 + b) := +calc a + (2 + b) ≤ a + (a + a * b) : + add_le_add_left (add_le_add a2 $ le_mul_of_one_le_left b0 $ one_le_two.trans a2) a + ... ≤ a * (2 + b) : by rw [mul_add, mul_two, add_assoc] -@[simp] lemma zero_lt_one : 0 < (1 : α) := -lt_of_le_of_ne zero_le_one zero_ne_one +lemma one_le_mul_of_one_le_of_one_le (ha : 1 ≤ a) (hb : 1 ≤ b) : (1 : α) ≤ a * b := +left.one_le_mul_of_le_of_le ha hb $ zero_le_one.trans ha -lemma zero_lt_two : 0 < (2:α) := add_pos zero_lt_one zero_lt_one +section monotone +variables [preorder β] {f g : β → α} -@[field_simps] lemma two_ne_zero : (2:α) ≠ 0 := -zero_lt_two.ne' +lemma monotone_mul_left_of_nonneg (ha : 0 ≤ a) : monotone (λ x, a * x) := +λ b c h, mul_le_mul_of_nonneg_left h ha -lemma one_lt_two : 1 < (2:α) := -calc (2:α) = 1+1 : one_add_one_eq_two - ... > 1+0 : add_lt_add_left zero_lt_one _ - ... = 1 : add_zero 1 +lemma monotone_mul_right_of_nonneg (ha : 0 ≤ a) : monotone (λ x, x * a) := +λ b c h, mul_le_mul_of_nonneg_right h ha -lemma zero_lt_three : 0 < (3:α) := add_pos zero_lt_two zero_lt_one +lemma monotone.mul_const (hf : monotone f) (ha : 0 ≤ a) : monotone (λ x, f x * a) := +(monotone_mul_right_of_nonneg ha).comp hf + +lemma monotone.const_mul (hf : monotone f) (ha : 0 ≤ a) : monotone (λ x, a * f x) := +(monotone_mul_left_of_nonneg ha).comp hf -@[field_simps] lemma three_ne_zero : (3:α) ≠ 0 := -zero_lt_three.ne' +lemma antitone.mul_const (hf : antitone f) (ha : 0 ≤ a) : antitone (λ x, f x * a) := +(monotone_mul_right_of_nonneg ha).comp_antitone hf -lemma zero_lt_four : 0 < (4:α) := add_pos zero_lt_two zero_lt_two +lemma antitone.const_mul (hf : antitone f) (ha : 0 ≤ a) : antitone (λ x, a * f x) := +(monotone_mul_left_of_nonneg ha).comp_antitone hf -@[field_simps] lemma four_ne_zero : (4:α) ≠ 0 := -zero_lt_four.ne' +lemma monotone.mul (hf : monotone f) (hg : monotone g) (hf₀ : ∀ x, 0 ≤ f x) (hg₀ : ∀ x, 0 ≤ g x) : + monotone (f * g) := +λ b c h, mul_le_mul (hf h) (hg h) (hg₀ _) (hf₀ _) + +end monotone + +section nontrivial +variables [nontrivial α] + +/-- See `zero_lt_one'` for a version with the type explicit. -/ +@[simp] lemma zero_lt_one : (0 : α) < 1 := zero_le_one.lt_of_ne zero_ne_one +/-- See `zero_lt_two'` for a version with the type explicit. -/ +@[simp] lemma zero_lt_two : (0 : α) < 2 := zero_lt_one.trans_le one_le_two +/-- See `zero_lt_three'` for a version with the type explicit. -/ +@[simp] lemma zero_lt_three : (0 : α) < 3 := +zero_lt_one.trans_le $ bit1_zero.symm.trans_le $ bit1_mono zero_le_one +/-- See `zero_lt_four'` for a version with the type explicit. -/ +@[simp] lemma zero_lt_four : (0 : α) < 4 := zero_lt_two.trans_le $ bit0_mono one_le_two + +@[field_simps] lemma two_ne_zero : (2 : α) ≠ 0 := zero_lt_two.ne' +@[field_simps] lemma three_ne_zero : (3 : α) ≠ 0 := zero_lt_three.ne' +@[field_simps] lemma four_ne_zero : (4 : α) ≠ 0 := zero_lt_four.ne' alias zero_lt_one ← one_pos alias zero_lt_two ← two_pos alias zero_lt_three ← three_pos alias zero_lt_four ← four_pos -@[priority 100] -- see Note [lower instance priority] -instance ordered_semiring.to_no_max_order : no_max_order α := -⟨assume a, ⟨a + 1, lt_add_of_pos_right _ one_pos⟩⟩ +lemma bit1_pos (h : 0 ≤ a) : 0 < bit1 a := +zero_lt_one.trans_le $ bit1_zero.symm.trans_le $ bit1_mono h + +variables (α) + +/-- See `zero_lt_one` for a version with the type implicit. -/ +lemma zero_lt_one' : (0 : α) < 1 := zero_lt_one +/-- See `zero_lt_two` for a version with the type implicit. -/ +lemma zero_lt_two' : (0 : α) < 2 := zero_lt_two +/-- See `zero_lt_three` for a version with the type implicit. -/ +lemma zero_lt_three' : (0 : α) < 3 := zero_lt_three +/-- See `zero_lt_four` for a version with the type implicit. -/ +lemma zero_lt_four' : (0 : α) < 4 := zero_lt_four end nontrivial --- See Note [decidable namespace] -protected lemma decidable.mul_le_mul_of_nonneg_left [@decidable_rel α (≤)] - (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c * a ≤ c * b := -begin - by_cases ba : b ≤ a, { simp [ba.antisymm h₁] }, - by_cases c0 : c ≤ 0, { simp [c0.antisymm h₂] }, - exact (mul_lt_mul_of_pos_left (h₁.lt_of_not_le ba) (h₂.lt_of_not_le c0)).le, -end +lemma bit1_pos' (h : 0 < a) : 0 < bit1 a := by { nontriviality, exact bit1_pos h.le } --- See Note [decidable namespace] -protected lemma decidable.mul_le_mul_of_nonneg_right [@decidable_rel α (≤)] - (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a * c ≤ b * c := -begin - by_cases ba : b ≤ a, { simp [ba.antisymm h₁] }, - by_cases c0 : c ≤ 0, { simp [c0.antisymm h₂] }, - exact (mul_lt_mul_of_pos_right (h₁.lt_of_not_le ba) (h₂.lt_of_not_le c0)).le, -end +lemma mul_le_one (ha : a ≤ 1) (hb' : 0 ≤ b) (hb : b ≤ 1) : a * b ≤ 1 := +one_mul (1 : α) ▸ mul_le_mul ha hb hb' zero_le_one --- TODO: there are four variations, depending on which variables we assume to be nonneg --- See Note [decidable namespace] -protected lemma decidable.mul_le_mul [@decidable_rel α (≤)] - (hac : a ≤ c) (hbd : b ≤ d) (nn_b : 0 ≤ b) (nn_c : 0 ≤ c) : a * b ≤ c * d := -calc - a * b ≤ c * b : decidable.mul_le_mul_of_nonneg_right hac nn_b - ... ≤ c * d : decidable.mul_le_mul_of_nonneg_left hbd nn_c +lemma one_lt_mul_of_le_of_lt (ha : 1 ≤ a) (hb : 1 < b) : 1 < a * b := +hb.trans_le $ le_mul_of_one_le_left (zero_le_one.trans hb.le) ha --- See Note [decidable namespace] -protected lemma decidable.mul_nonneg_le_one_le {α : Type*} [ordered_semiring α] - [@decidable_rel α (≤)] {a b c : α} - (h₁ : 0 ≤ c) (h₂ : a ≤ c) (h₃ : 0 ≤ b) (h₄ : b ≤ 1) : a * b ≤ c := -by simpa only [mul_one] using decidable.mul_le_mul h₂ h₄ h₃ h₁ +lemma one_lt_mul_of_lt_of_le (ha : 1 < a) (hb : 1 ≤ b) : 1 < a * b := +ha.trans_le $ le_mul_of_one_le_right (zero_le_one.trans ha.le) hb --- See Note [decidable namespace] -protected lemma decidable.mul_nonneg [@decidable_rel α (≤)] - (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := -have h : 0 * b ≤ a * b, from decidable.mul_le_mul_of_nonneg_right ha hb, -by rwa [zero_mul] at h +alias one_lt_mul_of_le_of_lt ← one_lt_mul -@[simp] theorem pow_nonneg (H : 0 ≤ a) : ∀ (n : ℕ), 0 ≤ a ^ n -| 0 := by { rw pow_zero, exact zero_le_one} -| (n+1) := by { rw pow_succ, exact mul_nonneg H (pow_nonneg _) } +lemma mul_lt_one_of_nonneg_of_lt_one_left (ha₀ : 0 ≤ a) (ha : a < 1) (hb : b ≤ 1) : a * b < 1 := +(mul_le_of_le_one_right ha₀ hb).trans_lt ha --- See Note [decidable namespace] -protected lemma decidable.mul_nonpos_of_nonneg_of_nonpos [@decidable_rel α (≤)] - (ha : 0 ≤ a) (hb : b ≤ 0) : a * b ≤ 0 := -have h : a * b ≤ a * 0, from decidable.mul_le_mul_of_nonneg_left hb ha, -by rwa mul_zero at h +lemma mul_lt_one_of_nonneg_of_lt_one_right (ha : a ≤ 1) (hb₀ : 0 ≤ b) (hb : b < 1) : a * b < 1 := +(mul_le_of_le_one_left hb₀ ha).trans_lt hb --- See Note [decidable namespace] -protected lemma decidable.mul_nonpos_of_nonpos_of_nonneg [@decidable_rel α (≤)] - (ha : a ≤ 0) (hb : 0 ≤ b) : a * b ≤ 0 := -have h : a * b ≤ 0 * b, from decidable.mul_le_mul_of_nonneg_right ha hb, -by rwa zero_mul at h +end ordered_semiring --- See Note [decidable namespace] -protected lemma decidable.mul_lt_mul [@decidable_rel α (≤)] - (hac : a < c) (hbd : b ≤ d) (pos_b : 0 < b) (nn_c : 0 ≤ c) : a * b < c * d := -calc - a * b < c * b : mul_lt_mul_of_pos_right hac pos_b - ... ≤ c * d : decidable.mul_le_mul_of_nonneg_left hbd nn_c +section ordered_ring +variables [ordered_ring α] {a b c d : α} -lemma mul_lt_mul : a < c → b ≤ d → 0 < b → 0 ≤ c → a * b < c * d := -by classical; exact decidable.mul_lt_mul +@[priority 100] -- see Note [lower instance priority] +instance ordered_ring.to_ordered_semiring : ordered_semiring α := +{ mul_le_mul_of_nonneg_left := λ a b c h hc, + by simpa only [mul_sub, sub_nonneg] using ordered_ring.mul_nonneg _ _ hc (sub_nonneg.2 h), + mul_le_mul_of_nonneg_right := λ a b c h hc, + by simpa only [sub_mul, sub_nonneg] using ordered_ring.mul_nonneg _ _ (sub_nonneg.2 h) hc, + ..‹ordered_ring α›, ..ring.to_semiring } --- See Note [decidable namespace] -protected lemma decidable.mul_lt_mul' [@decidable_rel α (≤)] - (h1 : a ≤ c) (h2 : b < d) (h3 : 0 ≤ b) (h4 : 0 < c) : a * b < c * d := -calc - a * b ≤ c * b : decidable.mul_le_mul_of_nonneg_right h1 h3 - ... < c * d : mul_lt_mul_of_pos_left h2 h4 +lemma mul_le_mul_of_nonpos_left (h : b ≤ a) (hc : c ≤ 0) : c * a ≤ c * b := +by simpa only [neg_mul, neg_le_neg_iff] using mul_le_mul_of_nonneg_left h (neg_nonneg.2 hc) -lemma mul_lt_mul' : a ≤ c → b < d → 0 ≤ b → 0 < c → a * b < c * d := -by classical; exact decidable.mul_lt_mul' +lemma mul_le_mul_of_nonpos_right (h : b ≤ a) (hc : c ≤ 0) : a * c ≤ b * c := +by simpa only [mul_neg, neg_le_neg_iff] using mul_le_mul_of_nonneg_right h (neg_nonneg.2 hc) -@[simp] theorem pow_pos (H : 0 < a) : ∀ (n : ℕ), 0 < a ^ n -| 0 := by { nontriviality, rw pow_zero, exact zero_lt_one } -| (n+1) := by { rw pow_succ, exact mul_pos H (pow_pos _) } +lemma mul_nonneg_of_nonpos_of_nonpos (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b := +by simpa only [zero_mul] using mul_le_mul_of_nonpos_right ha hb --- See Note [decidable namespace] -protected lemma decidable.mul_self_lt_mul_self [@decidable_rel α (≤)] - (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b := -decidable.mul_lt_mul' h2.le h2 h1 $ h1.trans_lt h2 +lemma mul_le_mul_of_nonneg_of_nonpos (hca : c ≤ a) (hbd : b ≤ d) (hc : 0 ≤ c) (hb : b ≤ 0) : + a * b ≤ c * d := +(mul_le_mul_of_nonpos_right hca hb).trans $ mul_le_mul_of_nonneg_left hbd hc -lemma mul_self_lt_mul_self (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b := -mul_lt_mul' h2.le h2 h1 $ h1.trans_lt h2 +lemma mul_le_mul_of_nonneg_of_nonpos' (hca : c ≤ a) (hbd : b ≤ d) (ha : 0 ≤ a) (hd : d ≤ 0) : + a * b ≤ c * d := +(mul_le_mul_of_nonneg_left hbd ha).trans $ mul_le_mul_of_nonpos_right hca hd --- See Note [decidable namespace] -protected lemma decidable.strict_mono_on_mul_self [@decidable_rel α (≤)] : - strict_mono_on (λ x : α, x * x) (set.Ici 0) := -λ x hx y hy hxy, decidable.mul_self_lt_mul_self hx hxy +lemma mul_le_mul_of_nonpos_of_nonneg (hac : a ≤ c) (hdb : d ≤ b) (hc : c ≤ 0) (hb : 0 ≤ b) : + a * b ≤ c * d := +(mul_le_mul_of_nonneg_right hac hb).trans $ mul_le_mul_of_nonpos_left hdb hc -lemma strict_mono_on_mul_self : strict_mono_on (λ x : α, x * x) (set.Ici 0) := -λ x hx y hy hxy, mul_self_lt_mul_self hx hxy +lemma mul_le_mul_of_nonpos_of_nonneg' (hca : c ≤ a) (hbd : b ≤ d) (ha : 0 ≤ a) (hd : d ≤ 0) : + a * b ≤ c * d := +(mul_le_mul_of_nonneg_left hbd ha).trans $ mul_le_mul_of_nonpos_right hca hd --- See Note [decidable namespace] -protected lemma decidable.mul_self_le_mul_self [@decidable_rel α (≤)] - (h1 : 0 ≤ a) (h2 : a ≤ b) : a * a ≤ b * b := -decidable.mul_le_mul h2 h2 h1 $ h1.trans h2 +lemma mul_le_mul_of_nonpos_of_nonpos (hca : c ≤ a) (hdb : d ≤ b) (hc : c ≤ 0) (hb : b ≤ 0) : + a * b ≤ c * d := +(mul_le_mul_of_nonpos_right hca hb).trans $ mul_le_mul_of_nonpos_left hdb hc -lemma mul_self_le_mul_self (h1 : 0 ≤ a) (h2 : a ≤ b) : a * a ≤ b * b := -mul_le_mul h2 h2 h1 $ h1.trans h2 +lemma mul_le_mul_of_nonpos_of_nonpos' (hca : c ≤ a) (hdb : d ≤ b) (ha : a ≤ 0) (hd : d ≤ 0) : + a * b ≤ c * d := +(mul_le_mul_of_nonpos_left hdb ha).trans $ mul_le_mul_of_nonpos_right hca hd --- See Note [decidable namespace] -protected lemma decidable.mul_lt_mul'' [@decidable_rel α (≤)] - (h1 : a < c) (h2 : b < d) (h3 : 0 ≤ a) (h4 : 0 ≤ b) : a * b < c * d := -h4.lt_or_eq_dec.elim - (λ b0, decidable.mul_lt_mul h1 h2.le b0 $ h3.trans h1.le) - (λ b0, by rw [← b0, mul_zero]; exact - mul_pos (h3.trans_lt h1) (h4.trans_lt h2)) +section monotone +variables [preorder β] {f g : β → α} -lemma mul_lt_mul'' : a < c → b < d → 0 ≤ a → 0 ≤ b → a * b < c * d := -by classical; exact decidable.mul_lt_mul'' +lemma antitone_mul_left {a : α} (ha : a ≤ 0) : antitone ((*) a) := +λ b c b_le_c, mul_le_mul_of_nonpos_left b_le_c ha --- See Note [decidable namespace] -protected lemma decidable.le_mul_of_one_le_right [@decidable_rel α (≤)] - (hb : 0 ≤ b) (h : 1 ≤ a) : b ≤ b * a := -suffices b * 1 ≤ b * a, by rwa mul_one at this, -decidable.mul_le_mul_of_nonneg_left h hb +lemma antitone_mul_right {a : α} (ha : a ≤ 0) : antitone (λ x, x * a) := +λ b c b_le_c, mul_le_mul_of_nonpos_right b_le_c ha --- See Note [decidable namespace] -protected lemma decidable.le_mul_of_one_le_left [@decidable_rel α (≤)] - (hb : 0 ≤ b) (h : 1 ≤ a) : b ≤ a * b := -suffices 1 * b ≤ a * b, by rwa one_mul at this, -decidable.mul_le_mul_of_nonneg_right h hb +lemma monotone.const_mul_of_nonpos (hf : monotone f) (ha : a ≤ 0) : antitone (λ x, a * f x) := +(antitone_mul_left ha).comp_monotone hf --- See Note [decidable namespace] -protected lemma decidable.lt_mul_of_one_lt_right [@decidable_rel α (≤)] - (hb : 0 < b) (h : 1 < a) : b < b * a := -suffices b * 1 < b * a, by rwa mul_one at this, -decidable.mul_lt_mul' le_rfl h zero_le_one hb +lemma monotone.mul_const_of_nonpos (hf : monotone f) (ha : a ≤ 0) : antitone (λ x, f x * a) := +(antitone_mul_right ha).comp_monotone hf --- See Note [decidable namespace] -protected lemma decidable.lt_mul_of_one_lt_left [@decidable_rel α (≤)] - (hb : 0 < b) (h : 1 < a) : b < a * b := -suffices 1 * b < a * b, by rwa one_mul at this, -decidable.mul_lt_mul h le_rfl hb (zero_le_one.trans h.le) +lemma antitone.const_mul_of_nonpos (hf : antitone f) (ha : a ≤ 0) : monotone (λ x, a * f x) := +(antitone_mul_left ha).comp hf -lemma lt_two_mul_self (ha : 0 < a) : a < 2 * a := -lt_mul_of_one_lt_left ha (@one_lt_two α _ (nontrivial_of_ne 0 a ha.ne)) +lemma antitone.mul_const_of_nonpos (hf : antitone f) (ha : a ≤ 0) : monotone (λ x, f x * a) := +(antitone_mul_right ha).comp hf -lemma lt_mul_self (hn : 1 < a) : a < a * a := -lt_mul_of_one_lt_left (hn.trans_le' zero_le_one) hn +lemma antitone.mul_monotone (hf : antitone f) (hg : monotone g) (hf₀ : ∀ x, f x ≤ 0) + (hg₀ : ∀ x, 0 ≤ g x) : + antitone (f * g) := +λ b c h, mul_le_mul_of_nonpos_of_nonneg (hf h) (hg h) (hf₀ _) (hg₀ _) --- See Note [decidable namespace] -protected lemma decidable.add_le_mul_two_add [@decidable_rel α (≤)] {a b : α} - (a2 : 2 ≤ a) (b0 : 0 ≤ b) : a + (2 + b) ≤ a * (2 + b) := -calc a + (2 + b) ≤ a + (a + a * b) : - add_le_add_left (add_le_add a2 (decidable.le_mul_of_one_le_left b0 (one_le_two.trans a2))) a - ... ≤ a * (2 + b) : by rw [mul_add, mul_two, add_assoc] +lemma monotone.mul_antitone (hf : monotone f) (hg : antitone g) (hf₀ : ∀ x, 0 ≤ f x) + (hg₀ : ∀ x, g x ≤ 0) : + antitone (f * g) := +λ b c h, mul_le_mul_of_nonneg_of_nonpos (hf h) (hg h) (hf₀ _) (hg₀ _) -lemma add_le_mul_two_add {a b : α} : 2 ≤ a → 0 ≤ b → a + (2 + b) ≤ a * (2 + b) := -by classical; exact decidable.add_le_mul_two_add +lemma antitone.mul (hf : antitone f) (hg : antitone g) (hf₀ : ∀ x, f x ≤ 0) (hg₀ : ∀ x, g x ≤ 0) : + monotone (f * g) := +λ b c h, mul_le_mul_of_nonpos_of_nonpos (hf h) (hg h) (hf₀ _) (hg₀ _) --- See Note [decidable namespace] -protected lemma decidable.one_le_mul_of_one_le_of_one_le [@decidable_rel α (≤)] - {a b : α} (a1 : 1 ≤ a) (b1 : 1 ≤ b) : (1 : α) ≤ a * b := -(mul_one (1 : α)).symm.le.trans (decidable.mul_le_mul a1 b1 zero_le_one (zero_le_one.trans a1)) +end monotone -lemma one_le_mul_of_one_le_of_one_le {a b : α} : 1 ≤ a → 1 ≤ b → (1 : α) ≤ a * b := -by classical; exact decidable.one_le_mul_of_one_le_of_one_le +lemma le_iff_exists_nonneg_add (a b : α) : a ≤ b ↔ ∃ c ≥ 0, b = a + c := +⟨λ h, ⟨b - a, sub_nonneg.mpr h, by simp⟩, + λ ⟨c, hc, h⟩, by { rw [h, le_add_iff_nonneg_right], exact hc }⟩ -namespace decidable +end ordered_ring -variables {β : Type*} [@decidable_rel α (≤)] [preorder β] {f g : β → α} +section ordered_comm_ring +variables [ordered_comm_ring α] -lemma monotone_mul_left_of_nonneg (ha : 0 ≤ a) : monotone (λ x, a*x) := -assume b c b_le_c, decidable.mul_le_mul_of_nonneg_left b_le_c ha +@[priority 100] -- See note [lower instance priority] +instance ordered_comm_ring.to_ordered_comm_semiring : ordered_comm_semiring α := +{ ..ordered_ring.to_ordered_semiring, ..‹ordered_comm_ring α› } -lemma monotone_mul_right_of_nonneg (ha : 0 ≤ a) : monotone (λ x, x*a) := -assume b c b_le_c, decidable.mul_le_mul_of_nonneg_right b_le_c ha +end ordered_comm_ring -lemma monotone_mul {β : Type*} [preorder β] {f g : β → α} - (hf : monotone f) (hg : monotone g) (hf0 : ∀ x, 0 ≤ f x) (hg0 : ∀ x, 0 ≤ g x) : - monotone (λ x, f x * g x) := -λ x y h, decidable.mul_le_mul (hf h) (hg h) (hg0 x) (hf0 y) +section strict_ordered_semiring +variables [strict_ordered_semiring α] {a b c d : α} -lemma strict_mono_mul_monotone (hf : strict_mono f) (hg : monotone g) (hf0 : ∀ x, 0 ≤ f x) - (hg0 : ∀ x, 0 < g x) : - strict_mono (λ x, f x * g x) := -λ x y h, decidable.mul_lt_mul (hf h) (hg h.le) (hg0 x) (hf0 y) +@[priority 200] -- see Note [lower instance priority] +instance strict_ordered_semiring.to_pos_mul_strict_mono : pos_mul_strict_mono α := +⟨λ x a b h, strict_ordered_semiring.mul_lt_mul_of_pos_left _ _ _ h x.prop⟩ -lemma monotone_mul_strict_mono (hf : monotone f) (hg : strict_mono g) (hf0 : ∀ x, 0 < f x) - (hg0 : ∀ x, 0 ≤ g x) : - strict_mono (λ x, f x * g x) := -λ x y h, decidable.mul_lt_mul' (hf h.le) (hg h) (hg0 x) (hf0 y) +@[priority 200] -- see Note [lower instance priority] +instance strict_ordered_semiring.to_mul_pos_strict_mono : mul_pos_strict_mono α := +⟨λ x a b h, strict_ordered_semiring.mul_lt_mul_of_pos_right _ _ _ h x.prop⟩ + +/-- A choice-free version of `strict_ordered_semiring.to_ordered_semiring` to avoid using choice in +basic `nat` lemmas. -/ +@[reducible] -- See note [reducible non-instances] +def strict_ordered_semiring.to_ordered_semiring' [@decidable_rel α (≤)] : ordered_semiring α := +{ mul_le_mul_of_nonneg_left := λ a b c hab hc, begin + obtain rfl | hab := decidable.eq_or_lt_of_le hab, + { refl }, + obtain rfl | hc := decidable.eq_or_lt_of_le hc, + { simp }, + { exact (mul_lt_mul_of_pos_left hab hc).le } + end, + mul_le_mul_of_nonneg_right := λ a b c hab hc, begin + obtain rfl | hab := decidable.eq_or_lt_of_le hab, + { refl }, + obtain rfl | hc := decidable.eq_or_lt_of_le hc, + { simp }, + { exact (mul_lt_mul_of_pos_right hab hc).le } + end, + ..‹strict_ordered_semiring α› } -lemma strict_mono_mul (hf : strict_mono f) (hg : strict_mono g) (hf0 : ∀ x, 0 ≤ f x) - (hg0 : ∀ x, 0 ≤ g x) : - strict_mono (λ x, f x * g x) := -λ x y h, decidable.mul_lt_mul'' (hf h) (hg h) (hf0 x) (hg0 x) +@[priority 100] -- see Note [lower instance priority] +instance strict_ordered_semiring.to_ordered_semiring : ordered_semiring α := +{ mul_le_mul_of_nonneg_left := λ _ _ _, begin + letI := @strict_ordered_semiring.to_ordered_semiring' α _ (classical.dec_rel _), + exact mul_le_mul_of_nonneg_left, + end, + mul_le_mul_of_nonneg_right := λ _ _ _, begin + letI := @strict_ordered_semiring.to_ordered_semiring' α _ (classical.dec_rel _), + exact mul_le_mul_of_nonneg_right, + end, + ..‹strict_ordered_semiring α› } -end decidable +lemma mul_lt_mul (hac : a < c) (hbd : b ≤ d) (hb : 0 < b) (hc : 0 ≤ c) : a * b < c * d := +(mul_lt_mul_of_pos_right hac hb).trans_le $ mul_le_mul_of_nonneg_left hbd hc -section mono +lemma mul_lt_mul' (hac : a ≤ c) (hbd : b < d) (hb : 0 ≤ b) (hc : 0 < c) : a * b < c * d := +(mul_le_mul_of_nonneg_right hac hb).trans_lt $ mul_lt_mul_of_pos_left hbd hc -open_locale classical +@[simp] theorem pow_pos (H : 0 < a) : ∀ (n : ℕ), 0 < a ^ n +| 0 := by { nontriviality, rw pow_zero, exact zero_lt_one } +| (n+1) := by { rw pow_succ, exact mul_pos H (pow_pos _) } -variables {β : Type*} [preorder β] {f g : β → α} +lemma mul_self_lt_mul_self (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b := +mul_lt_mul' h2.le h2 h1 $ h1.trans_lt h2 -lemma monotone_mul_left_of_nonneg (ha : 0 ≤ a) : monotone (λ x, a*x) := -decidable.monotone_mul_left_of_nonneg ha +lemma strict_mono_on_mul_self : strict_mono_on (λ x : α, x * x) (set.Ici 0) := +λ x hx y hy hxy, mul_self_lt_mul_self hx hxy -lemma monotone_mul_right_of_nonneg (ha : 0 ≤ a) : monotone (λ x, x*a) := -decidable.monotone_mul_right_of_nonneg ha +-- See Note [decidable namespace] +protected lemma decidable.mul_lt_mul'' [@decidable_rel α (≤)] + (h1 : a < c) (h2 : b < d) (h3 : 0 ≤ a) (h4 : 0 ≤ b) : a * b < c * d := +h4.lt_or_eq_dec.elim + (λ b0, mul_lt_mul h1 h2.le b0 $ h3.trans h1.le) + (λ b0, by rw [← b0, mul_zero]; exact + mul_pos (h3.trans_lt h1) (h4.trans_lt h2)) -lemma monotone.mul_const (hf : monotone f) (ha : 0 ≤ a) : - monotone (λ x, (f x) * a) := -(monotone_mul_right_of_nonneg ha).comp hf +lemma mul_lt_mul'' : a < c → b < d → 0 ≤ a → 0 ≤ b → a * b < c * d := +by classical; exact decidable.mul_lt_mul'' -lemma monotone.const_mul (hf : monotone f) (ha : 0 ≤ a) : - monotone (λ x, a * (f x)) := -(monotone_mul_left_of_nonneg ha).comp hf +lemma lt_mul_left (hn : 0 < a) (hm : 1 < b) : a < b * a := +by { convert mul_lt_mul_of_pos_right hm hn, rw one_mul } -lemma monotone.mul (hf : monotone f) (hg : monotone g) (hf0 : ∀ x, 0 ≤ f x) (hg0 : ∀ x, 0 ≤ g x) : - monotone (λ x, f x * g x) := -decidable.monotone_mul hf hg hf0 hg0 +lemma lt_mul_right (hn : 0 < a) (hm : 1 < b) : a < a * b := +by { convert mul_lt_mul_of_pos_left hm hn, rw mul_one } + +lemma lt_mul_self (hn : 1 < a) : a < a * a := +lt_mul_left (hn.trans_le' zero_le_one) hn + +section monotone +variables [preorder β] {f g : β → α} lemma strict_mono_mul_left_of_pos (ha : 0 < a) : strict_mono (λ x, a * x) := assume b c b_lt_c, mul_lt_mul_of_pos_left b_lt_c ha @@ -400,150 +511,51 @@ lemma strict_mono.const_mul (hf : strict_mono f) (ha : 0 < a) : strict_mono (λ x, a * (f x)) := (strict_mono_mul_left_of_pos ha).comp hf -lemma strict_mono.mul_monotone (hf : strict_mono f) (hg : monotone g) (hf0 : ∀ x, 0 ≤ f x) - (hg0 : ∀ x, 0 < g x) : - strict_mono (λ x, f x * g x) := -decidable.strict_mono_mul_monotone hf hg hf0 hg0 - -lemma monotone.mul_strict_mono (hf : monotone f) (hg : strict_mono g) (hf0 : ∀ x, 0 < f x) - (hg0 : ∀ x, 0 ≤ g x) : - strict_mono (λ x, f x * g x) := -decidable.monotone_mul_strict_mono hf hg hf0 hg0 - -lemma strict_mono.mul (hf : strict_mono f) (hg : strict_mono g) (hf0 : ∀ x, 0 ≤ f x) - (hg0 : ∀ x, 0 ≤ g x) : - strict_mono (λ x, f x * g x) := -decidable.strict_mono_mul hf hg hf0 hg0 - -end mono - -/-- Pullback an `ordered_semiring` under an injective map. -See note [reducible non-instances]. -/ -@[reducible] -def function.injective.ordered_semiring {β : Type*} - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] - [has_smul ℕ β] [has_nat_cast β] - (f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) : - ordered_semiring β := -{ zero_le_one := show f 0 ≤ f 1, by simp only [zero, one, zero_le_one], - mul_lt_mul_of_pos_left := λ a b c ab c0, show f (c * a) < f (c * b), - begin - rw [mul, mul], - refine mul_lt_mul_of_pos_left ab _, - rwa ← zero, - end, - mul_lt_mul_of_pos_right := λ a b c ab c0, show f (a * c) < f (b * c), - begin - rw [mul, mul], - refine mul_lt_mul_of_pos_right ab _, - rwa ← zero, - end, - ..hf.ordered_cancel_add_comm_monoid f zero add nsmul, - ..hf.semiring f zero one add mul nsmul npow nat_cast } - -section -variable [nontrivial α] - -lemma bit1_pos (h : 0 ≤ a) : 0 < bit1 a := -lt_add_of_le_of_pos (add_nonneg h h) zero_lt_one - -lemma lt_add_one (a : α) : a < a + 1 := -lt_add_of_le_of_pos le_rfl zero_lt_one - -lemma lt_one_add (a : α) : a < 1 + a := -by { rw [add_comm], apply lt_add_one } - -end - -lemma bit1_pos' (h : 0 < a) : 0 < bit1 a := -begin - nontriviality, - exact bit1_pos h.le, -end - --- See Note [decidable namespace] -protected lemma decidable.one_lt_mul [@decidable_rel α (≤)] - (ha : 1 ≤ a) (hb : 1 < b) : 1 < a * b := -begin - nontriviality, - exact (one_mul (1 : α)) ▸ decidable.mul_lt_mul' ha hb zero_le_one (zero_lt_one.trans_le ha) -end - -lemma one_lt_mul : 1 ≤ a → 1 < b → 1 < a * b := -by classical; exact decidable.one_lt_mul +lemma strict_anti.mul_const (hf : strict_anti f) (ha : 0 < a) : strict_anti (λ x, f x * a) := +(strict_mono_mul_right_of_pos ha).comp_strict_anti hf --- See Note [decidable namespace] -protected lemma decidable.mul_le_one [@decidable_rel α (≤)] - (ha : a ≤ 1) (hb' : 0 ≤ b) (hb : b ≤ 1) : a * b ≤ 1 := -one_mul (1 : α) ▸ decidable.mul_le_mul ha hb hb' zero_le_one +lemma strict_anti.const_mul (hf : strict_anti f) (ha : 0 < a) : strict_anti (λ x, a * f x) := +(strict_mono_mul_left_of_pos ha).comp_strict_anti hf -lemma mul_le_one : a ≤ 1 → 0 ≤ b → b ≤ 1 → a * b ≤ 1 := -by classical; exact decidable.mul_le_one +lemma strict_mono.mul_monotone (hf : strict_mono f) (hg : monotone g) (hf₀ : ∀ x, 0 ≤ f x) + (hg₀ : ∀ x, 0 < g x) : + strict_mono (f * g) := +λ b c h, mul_lt_mul (hf h) (hg h.le) (hg₀ _) (hf₀ _) --- See Note [decidable namespace] -protected lemma decidable.one_lt_mul_of_le_of_lt [@decidable_rel α (≤)] - (ha : 1 ≤ a) (hb : 1 < b) : 1 < a * b := -begin - nontriviality, - calc 1 = 1 * 1 : by rw one_mul - ... < a * b : decidable.mul_lt_mul' ha hb zero_le_one (zero_lt_one.trans_le ha) -end +lemma monotone.mul_strict_mono (hf : monotone f) (hg : strict_mono g) (hf₀ : ∀ x, 0 < f x) + (hg₀ : ∀ x, 0 ≤ g x) : + strict_mono (f * g) := +λ b c h, mul_lt_mul' (hf h.le) (hg h) (hg₀ _) (hf₀ _) -lemma one_lt_mul_of_le_of_lt : 1 ≤ a → 1 < b → 1 < a * b := -by classical; exact decidable.one_lt_mul_of_le_of_lt +lemma strict_mono.mul (hf : strict_mono f) (hg : strict_mono g) (hf₀ : ∀ x, 0 ≤ f x) + (hg₀ : ∀ x, 0 ≤ g x) : + strict_mono (f * g) := +λ b c h, mul_lt_mul'' (hf h) (hg h) (hf₀ _) (hg₀ _) --- See Note [decidable namespace] -protected lemma decidable.one_lt_mul_of_lt_of_le [@decidable_rel α (≤)] - (ha : 1 < a) (hb : 1 ≤ b) : 1 < a * b := -begin - nontriviality, - calc 1 = 1 * 1 : by rw one_mul - ... < a * b : decidable.mul_lt_mul ha hb zero_lt_one $ zero_le_one.trans ha.le -end +end monotone -lemma one_lt_mul_of_lt_of_le : 1 < a → 1 ≤ b → 1 < a * b := -by classical; exact decidable.one_lt_mul_of_lt_of_le +section nontrivial +variables [nontrivial α] --- See Note [decidable namespace] -protected lemma decidable.mul_le_of_le_one_right [@decidable_rel α (≤)] - (ha : 0 ≤ a) (hb1 : b ≤ 1) : a * b ≤ a := -calc a * b ≤ a * 1 : decidable.mul_le_mul_of_nonneg_left hb1 ha -... = a : mul_one a +lemma lt_one_add (a : α) : a < 1 + a := lt_add_of_pos_left _ zero_lt_one +lemma lt_add_one (a : α) : a < a + 1 := lt_add_of_pos_right _ zero_lt_one --- See Note [decidable namespace] -protected lemma decidable.mul_le_of_le_one_left [@decidable_rel α (≤)] - (hb : 0 ≤ b) (ha1 : a ≤ 1) : a * b ≤ b := -calc a * b ≤ 1 * b : decidable.mul_le_mul ha1 le_rfl hb zero_le_one -... = b : one_mul b +lemma one_lt_two : (1 : α) < 2 := lt_add_one _ --- See Note [decidable namespace] -protected lemma decidable.mul_lt_one_of_nonneg_of_lt_one_left [@decidable_rel α (≤)] - (ha0 : 0 ≤ a) (ha : a < 1) (hb : b ≤ 1) : a * b < 1 := -calc a * b ≤ a : decidable.mul_le_of_le_one_right ha0 hb -... < 1 : ha +lemma lt_two_mul_self (ha : 0 < a) : a < 2 * a := lt_mul_of_one_lt_left ha one_lt_two -lemma mul_lt_one_of_nonneg_of_lt_one_left : 0 ≤ a → a < 1 → b ≤ 1 → a * b < 1 := -by classical; exact decidable.mul_lt_one_of_nonneg_of_lt_one_left +lemma nat.strict_mono_cast : strict_mono (coe : ℕ → α) := +strict_mono_nat_of_lt_succ $ λ n, by rw [nat.cast_succ]; apply lt_add_one --- See Note [decidable namespace] -protected lemma decidable.mul_lt_one_of_nonneg_of_lt_one_right [@decidable_rel α (≤)] - (ha : a ≤ 1) (hb0 : 0 ≤ b) (hb : b < 1) : a * b < 1 := -calc a * b ≤ b : decidable.mul_le_of_le_one_left hb0 ha -... < 1 : hb +@[priority 100] -- see Note [lower instance priority] +instance strict_ordered_semiring.to_no_max_order : no_max_order α := +⟨λ a, ⟨a + 1, lt_add_of_pos_right _ one_pos⟩⟩ -lemma mul_lt_one_of_nonneg_of_lt_one_right : a ≤ 1 → 0 ≤ b → b < 1 → a * b < 1 := -by classical; exact decidable.mul_lt_one_of_nonneg_of_lt_one_right +/-- Note this is not an instance as `char_zero` implies `nontrivial`, and this would risk forming a +loop. -/ +lemma strict_ordered_semiring.to_char_zero : char_zero α := ⟨nat.strict_mono_cast.injective⟩ -theorem nat.strict_mono_cast [nontrivial α] : strict_mono (coe : ℕ → α) := -strict_mono_nat_of_lt_succ $ λ n, by rw [nat.cast_succ]; apply lt_add_one - -/-- Note this is not an instance as `char_zero` implies `nontrivial`, -and this would risk forming a loop. -/ -lemma ordered_semiring.to_char_zero [nontrivial α] : char_zero α := -⟨nat.strict_mono_cast.injective⟩ +end nontrivial section has_exists_add_of_le variables [has_exists_add_of_le α] @@ -575,53 +587,123 @@ lemma mul_add_mul_lt_mul_add_mul' (hba : b < a) (hdc : d < c) : a • d + b • by { rw [add_comm (a • d), add_comm (a • c)], exact mul_add_mul_lt_mul_add_mul hba hdc } end has_exists_add_of_le -end ordered_semiring +end strict_ordered_semiring -section ordered_comm_semiring +section strict_ordered_comm_semiring +variables [strict_ordered_comm_semiring α] -/-- An `ordered_comm_semiring α` is a commutative semiring `α` with a partial order such that -addition is monotone and multiplication by a positive number is strictly monotone. -/ -@[protect_proj] -class ordered_comm_semiring (α : Type u) extends ordered_semiring α, comm_semiring α +/-- A choice-free version of `strict_ordered_comm_semiring.to_ordered_comm_semiring` to avoid using +choice in basic `nat` lemmas. -/ +@[reducible] -- See note [reducible non-instances] +def strict_ordered_comm_semiring.to_ordered_comm_semiring' [@decidable_rel α (≤)] : + ordered_comm_semiring α := +{ ..‹strict_ordered_comm_semiring α›, ..strict_ordered_semiring.to_ordered_semiring' } -/-- Pullback an `ordered_comm_semiring` under an injective map. -See note [reducible non-instances]. -/ -@[reducible] -def function.injective.ordered_comm_semiring [ordered_comm_semiring α] {β : Type*} - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] - (f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) : - ordered_comm_semiring β := -{ ..hf.comm_semiring f zero one add mul nsmul npow nat_cast, - ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast } +@[priority 100] -- see Note [lower instance priority] +instance strict_ordered_comm_semiring.to_ordered_comm_semiring : ordered_comm_semiring α := +{ ..‹strict_ordered_comm_semiring α›, ..strict_ordered_semiring.to_ordered_semiring } -end ordered_comm_semiring +end strict_ordered_comm_semiring -/-- -A `linear_ordered_semiring α` is a nontrivial semiring `α` with a linear order -such that addition is monotone and multiplication by a positive number is strictly monotone. --/ --- It's not entirely clear we should assume `nontrivial` at this point; --- it would be reasonable to explore changing this, --- but be warned that the instances involving `domain` may cause --- typeclass search loops. -@[protect_proj] -class linear_ordered_semiring (α : Type u) - extends ordered_semiring α, linear_ordered_add_comm_monoid α, nontrivial α +section strict_ordered_ring +variables [strict_ordered_ring α] {a b c : α} + +@[priority 100] -- see Note [lower instance priority] +instance strict_ordered_ring.to_strict_ordered_semiring : strict_ordered_semiring α := +{ le_of_add_le_add_left := @le_of_add_le_add_left α _ _ _, + mul_lt_mul_of_pos_left := λ a b c h hc, + by simpa only [mul_sub, sub_pos] using strict_ordered_ring.mul_pos _ _ hc (sub_pos.2 h), + mul_lt_mul_of_pos_right := λ a b c h hc, + by simpa only [sub_mul, sub_pos] using strict_ordered_ring.mul_pos _ _ (sub_pos.2 h) hc, + ..‹strict_ordered_ring α›, ..ring.to_semiring } + +/-- A choice-free version of `strict_ordered_ring.to_ordered_ring` to avoid using choice in basic +`int` lemmas. -/ +@[reducible] -- See note [reducible non-instances] +def strict_ordered_ring.to_ordered_ring' [@decidable_rel α (≤)] : ordered_ring α := +{ mul_nonneg := λ a b ha hb, begin + cases decidable.eq_or_lt_of_le ha with ha ha, + { rw [←ha, zero_mul] }, + cases decidable.eq_or_lt_of_le hb with hb hb, + { rw [←hb, mul_zero] }, + { exact (mul_pos ha hb).le } + end, + ..‹strict_ordered_ring α›, ..ring.to_semiring } + + +@[priority 100] -- see Note [lower instance priority] +instance strict_ordered_ring.to_ordered_ring : ordered_ring α := +{ mul_nonneg := λ a b, begin + letI := @strict_ordered_ring.to_ordered_ring' α _ (classical.dec_rel _), + exact mul_nonneg, + end, + ..‹strict_ordered_ring α› } + +lemma mul_lt_mul_of_neg_left (h : b < a) (hc : c < 0) : c * a < c * b := +by simpa only [neg_mul, neg_lt_neg_iff] using mul_lt_mul_of_pos_left h (neg_pos_of_neg hc) + +lemma mul_lt_mul_of_neg_right (h : b < a) (hc : c < 0) : a * c < b * c := +by simpa only [mul_neg, neg_lt_neg_iff] using mul_lt_mul_of_pos_right h (neg_pos_of_neg hc) + +lemma mul_pos_of_neg_of_neg {a b : α} (ha : a < 0) (hb : b < 0) : 0 < a * b := +by simpa only [zero_mul] using mul_lt_mul_of_neg_right ha hb + +section monotone +variables [preorder β] {f g : β → α} + +lemma strict_anti_mul_left {a : α} (ha : a < 0) : strict_anti ((*) a) := +λ b c b_lt_c, mul_lt_mul_of_neg_left b_lt_c ha + +lemma strict_anti_mul_right {a : α} (ha : a < 0) : strict_anti (λ x, x * a) := +λ b c b_lt_c, mul_lt_mul_of_neg_right b_lt_c ha + +lemma strict_mono.const_mul_of_neg (hf : strict_mono f) (ha : a < 0) : strict_anti (λ x, a * f x) := +(strict_anti_mul_left ha).comp_strict_mono hf + +lemma strict_mono.mul_const_of_neg (hf : strict_mono f) (ha : a < 0) : strict_anti (λ x, f x * a) := +(strict_anti_mul_right ha).comp_strict_mono hf + +lemma strict_anti.const_mul_of_neg (hf : strict_anti f) (ha : a < 0) : strict_mono (λ x, a * f x) := +(strict_anti_mul_left ha).comp hf + +lemma strict_anti.mul_const_of_neg (hf : strict_anti f) (ha : a < 0) : strict_mono (λ x, f x * a) := +(strict_anti_mul_right ha).comp hf + +end monotone +end strict_ordered_ring + +section strict_ordered_comm_ring +variables [strict_ordered_comm_ring α] + +/-- A choice-free version of `strict_ordered_comm_ring.to_ordered_comm_semiring'` to avoid using +choice in basic `int` lemmas. -/ +@[reducible] -- See note [reducible non-instances] +def strict_ordered_comm_ring.to_ordered_comm_ring' [@decidable_rel α (≤)] : ordered_comm_ring α := +{ ..‹strict_ordered_comm_ring α›, ..strict_ordered_ring.to_ordered_ring' } + +@[priority 100] -- See note [lower instance priority] +instance strict_ordered_comm_ring.to_strict_ordered_comm_semiring : + strict_ordered_comm_semiring α := +{ ..‹strict_ordered_comm_ring α›, ..strict_ordered_ring.to_strict_ordered_semiring } + +@[priority 100] -- See note [lower instance priority] +instance strict_ordered_comm_ring.to_ordered_comm_ring : ordered_comm_ring α := +{ ..‹strict_ordered_comm_ring α›, ..strict_ordered_ring.to_ordered_ring } + +end strict_ordered_comm_ring section linear_ordered_semiring variables [linear_ordered_semiring α] {a b c d : α} -local attribute [instance] linear_ordered_semiring.decidable_le +@[priority 200] -- see Note [lower instance priority] +instance linear_ordered_semiring.to_pos_mul_reflect_lt : pos_mul_reflect_lt α := +⟨λ a b c, (monotone_mul_left_of_nonneg a.2).reflect_lt⟩ + +@[priority 200] -- see Note [lower instance priority] +instance linear_ordered_semiring.to_mul_pos_reflect_lt : mul_pos_reflect_lt α := +⟨λ a b c, (monotone_mul_right_of_nonneg a.2).reflect_lt⟩ --- `norm_num` expects the lemma stating `0 < 1` to have a single typeclass argument --- (see `norm_num.prove_pos_nat`). --- Rather than working out how to relax that assumption, --- we provide a synonym for `zero_lt_one` (which needs both `ordered_semiring α` and `nontrivial α`) --- with only a `linear_ordered_semiring` typeclass argument. -lemma zero_lt_one' : 0 < (1 : α) := zero_lt_one +local attribute [instance] linear_ordered_semiring.decidable_le linear_ordered_semiring.decidable_lt lemma nonneg_and_nonneg_or_nonpos_and_nonpos_of_mul_nnonneg (hab : 0 ≤ a * b) : (0 ≤ a ∧ 0 ≤ b) ∨ (a ≤ 0 ∧ b ≤ 0) := @@ -640,10 +722,10 @@ lemma nonneg_of_mul_nonneg_right (h : 0 ≤ a * b) (ha : 0 < a) : 0 ≤ b := le_of_not_gt $ λ hb, (mul_neg_of_pos_of_neg ha hb).not_le h lemma neg_of_mul_neg_left (h : a * b < 0) (hb : 0 ≤ b) : a < 0 := -lt_of_not_ge $ λ ha, (decidable.mul_nonneg ha hb).not_lt h +lt_of_not_ge $ λ ha, (mul_nonneg ha hb).not_lt h lemma neg_of_mul_neg_right (h : a * b < 0) (ha : 0 ≤ a) : b < 0 := -lt_of_not_ge $ λ hb, (decidable.mul_nonneg ha hb).not_lt h +lt_of_not_ge $ λ hb, (mul_nonneg ha hb).not_lt h lemma nonpos_of_mul_nonpos_left (h : a * b ≤ 0) (hb : 0 < b) : a ≤ 0 := le_of_not_gt (assume ha : a > 0, (mul_pos ha hb).not_le h) @@ -711,12 +793,10 @@ by rw [bit0, ← two_mul, zero_lt_mul_left (zero_lt_two : 0 < (2:α))] end theorem mul_nonneg_iff_right_nonneg_of_pos (ha : 0 < a) : 0 ≤ a * b ↔ 0 ≤ b := -by haveI := @linear_order.decidable_le α _; exact -⟨λ h, nonneg_of_mul_nonneg_right h ha, λ h, decidable.mul_nonneg ha.le h⟩ +⟨λ h, nonneg_of_mul_nonneg_right h ha, mul_nonneg ha.le⟩ theorem mul_nonneg_iff_left_nonneg_of_pos (hb : 0 < b) : 0 ≤ a * b ↔ 0 ≤ a := -by haveI := @linear_order.decidable_le α _; exact -⟨λ h, nonneg_of_mul_nonneg_left h hb, λ h, decidable.mul_nonneg h hb.le⟩ +⟨λ h, nonneg_of_mul_nonneg_left h hb, λ h, mul_nonneg h hb.le⟩ lemma nonpos_of_mul_nonneg_left (h : 0 ≤ a * b) (hb : b < 0) : a ≤ 0 := le_of_not_gt (λ ha, absurd h (mul_neg_of_pos_of_neg ha hb).not_le) @@ -724,22 +804,6 @@ le_of_not_gt (λ ha, absurd h (mul_neg_of_pos_of_neg ha hb).not_le) lemma nonpos_of_mul_nonneg_right (h : 0 ≤ a * b) (ha : a < 0) : b ≤ 0 := le_of_not_gt (λ hb, absurd h (mul_neg_of_neg_of_pos ha hb).not_le) -/-- Pullback a `linear_ordered_semiring` under an injective map. -See note [reducible non-instances]. -/ -@[reducible] -def function.injective.linear_ordered_semiring {β : Type*} - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] - [has_sup β] [has_inf β] - (f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) - (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_semiring β := -{ .. linear_order.lift f hf hsup hinf, - .. pullback_nonzero f zero one, - .. hf.ordered_semiring f zero one add mul nsmul npow nat_cast } - @[simp] lemma units.inv_pos {u : αˣ} : (0 : α) < ↑u⁻¹ ↔ (0 : α) < u := have ∀ {u : αˣ}, (0 : α) < u → (0 : α) < ↑u⁻¹ := λ u h, (zero_lt_mul_left h).mp $ u.mul_inv.symm ▸ zero_lt_one, @@ -752,28 +816,14 @@ have ∀ {u : αˣ}, ↑u < (0 : α) → ↑u⁻¹ < (0 : α) := λ u h, @[priority 100] -- see Note [lower instance priority] instance linear_ordered_semiring.to_char_zero : char_zero α := -ordered_semiring.to_char_zero +strict_ordered_semiring.to_char_zero -end linear_ordered_semiring - -section mono -variables [linear_ordered_semiring α] {a : α} - -local attribute [instance] linear_ordered_semiring.decidable_lt - -lemma cmp_mul_pos_left (ha : 0 < a) (b c : α) : - cmp (a * b) (a * c) = cmp b c := +lemma cmp_mul_pos_left (ha : 0 < a) (b c : α) : cmp (a * b) (a * c) = cmp b c := (strict_mono_mul_left_of_pos ha).cmp_map_eq b c -lemma cmp_mul_pos_right (ha : 0 < a) (b c : α) : - cmp (b * a) (c * a) = cmp b c := +lemma cmp_mul_pos_right (ha : 0 < a) (b c : α) : cmp (b * a) (c * a) = cmp b c := (strict_mono_mul_right_of_pos ha).cmp_map_eq b c -end mono - -section linear_ordered_semiring -variables [linear_ordered_semiring α] {a b c : α} - lemma mul_max_of_nonneg (b c : α) (ha : 0 ≤ a) : a * max b c = max (a * b) (a * c) := (monotone_mul_left_of_nonneg ha).map_max @@ -786,253 +836,28 @@ lemma max_mul_of_nonneg (a b : α) (hc : 0 ≤ c) : max a b * c = max (a * c) (b lemma min_mul_of_nonneg (a b : α) (hc : 0 ≤ c) : min a b * c = min (a * c) (b * c) := (monotone_mul_right_of_nonneg hc).map_min -end linear_ordered_semiring - -/-- A `linear_ordered_comm_semiring` is a nontrivial commutative semiring with a linear order such -that addition is monotone and multiplication by a positive number is strictly monotone. -/ -@[protect_proj, ancestor ordered_comm_semiring linear_ordered_semiring] -class linear_ordered_comm_semiring (α : Type*) - extends ordered_comm_semiring α, linear_ordered_semiring α - -@[priority 100] -- See note [lower instance priority] -instance linear_ordered_comm_semiring.to_linear_ordered_cancel_add_comm_monoid - [linear_ordered_comm_semiring α] : linear_ordered_cancel_add_comm_monoid α := -{ ..‹linear_ordered_comm_semiring α› } - -/-- Pullback a `linear_ordered_semiring` under an injective map. -See note [reducible non-instances]. -/ -@[reducible] -def function.injective.linear_ordered_comm_semiring [linear_ordered_comm_semiring α] {β : Type*} - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] - [has_sup β] [has_inf β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) - (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_comm_semiring β := -{ ..hf.linear_ordered_semiring f zero one add mul nsmul npow nat_cast hsup hinf, - ..hf.ordered_comm_semiring f zero one add mul nsmul npow nat_cast } - -/-- An `ordered_ring α` is a ring `α` with a partial order such that -addition is monotone and multiplication by a positive number is strictly monotone. -/ -@[protect_proj] -class ordered_ring (α : Type u) extends ring α, ordered_add_comm_group α := -(zero_le_one : 0 ≤ (1 : α)) -(mul_pos : ∀ a b : α, 0 < a → 0 < b → 0 < a * b) - -section ordered_ring -variables [ordered_ring α] {a b c : α} - --- See Note [decidable namespace] -protected lemma decidable.ordered_ring.mul_nonneg [@decidable_rel α (≤)] - {a b : α} (h₁ : 0 ≤ a) (h₂ : 0 ≤ b) : 0 ≤ a * b := -begin - by_cases ha : a ≤ 0, { simp [le_antisymm ha h₁] }, - by_cases hb : b ≤ 0, { simp [le_antisymm hb h₂] }, - exact (le_not_le_of_lt (ordered_ring.mul_pos a b (h₁.lt_of_not_le ha) (h₂.lt_of_not_le hb))).1, -end - -lemma ordered_ring.mul_nonneg : 0 ≤ a → 0 ≤ b → 0 ≤ a * b := -by classical; exact decidable.ordered_ring.mul_nonneg - --- See Note [decidable namespace] -protected lemma decidable.ordered_ring.mul_le_mul_of_nonneg_left - [@decidable_rel α (≤)] (h₁ : a ≤ b) (h₂ : 0 ≤ c) : c * a ≤ c * b := -begin - rw [← sub_nonneg, ← mul_sub], - exact decidable.ordered_ring.mul_nonneg h₂ (sub_nonneg.2 h₁), -end - -lemma ordered_ring.mul_le_mul_of_nonneg_left : a ≤ b → 0 ≤ c → c * a ≤ c * b := -by classical; exact decidable.ordered_ring.mul_le_mul_of_nonneg_left - --- See Note [decidable namespace] -protected lemma decidable.ordered_ring.mul_le_mul_of_nonneg_right - [@decidable_rel α (≤)] (h₁ : a ≤ b) (h₂ : 0 ≤ c) : a * c ≤ b * c := -begin - rw [← sub_nonneg, ← sub_mul], - exact decidable.ordered_ring.mul_nonneg (sub_nonneg.2 h₁) h₂, -end - -lemma ordered_ring.mul_le_mul_of_nonneg_right : a ≤ b → 0 ≤ c → a * c ≤ b * c := -by classical; exact decidable.ordered_ring.mul_le_mul_of_nonneg_right - -lemma ordered_ring.mul_lt_mul_of_pos_left (h₁ : a < b) (h₂ : 0 < c) : c * a < c * b := -begin - rw [← sub_pos, ← mul_sub], - exact ordered_ring.mul_pos _ _ h₂ (sub_pos.2 h₁), -end - -lemma ordered_ring.mul_lt_mul_of_pos_right (h₁ : a < b) (h₂ : 0 < c) : a * c < b * c := -begin - rw [← sub_pos, ← sub_mul], - exact ordered_ring.mul_pos _ _ (sub_pos.2 h₁) h₂, -end - -@[priority 100] -- see Note [lower instance priority] -instance ordered_ring.to_ordered_semiring : ordered_semiring α := -{ le_of_add_le_add_left := @le_of_add_le_add_left α _ _ _, - mul_lt_mul_of_pos_left := @ordered_ring.mul_lt_mul_of_pos_left α _, - mul_lt_mul_of_pos_right := @ordered_ring.mul_lt_mul_of_pos_right α _, - ..‹ordered_ring α›, ..ring.to_semiring } - --- See Note [decidable namespace] -protected lemma decidable.mul_le_mul_of_nonpos_left [@decidable_rel α (≤)] - {a b c : α} (h : b ≤ a) (hc : c ≤ 0) : c * a ≤ c * b := -have -c ≥ 0, from neg_nonneg_of_nonpos hc, -have -c * b ≤ -c * a, from decidable.mul_le_mul_of_nonneg_left h this, -have -(c * b) ≤ -(c * a), by rwa [neg_mul, neg_mul] at this, -le_of_neg_le_neg this - -lemma mul_le_mul_of_nonpos_left {a b c : α} : b ≤ a → c ≤ 0 → c * a ≤ c * b := -by classical; exact decidable.mul_le_mul_of_nonpos_left - --- See Note [decidable namespace] -protected lemma decidable.mul_le_mul_of_nonpos_right [@decidable_rel α (≤)] - {a b c : α} (h : b ≤ a) (hc : c ≤ 0) : a * c ≤ b * c := -have -c ≥ 0, from neg_nonneg_of_nonpos hc, -have b * -c ≤ a * -c, from decidable.mul_le_mul_of_nonneg_right h this, -have -(b * c) ≤ -(a * c), by rwa [mul_neg, mul_neg] at this, -le_of_neg_le_neg this - -lemma mul_le_mul_of_nonpos_right {a b c : α} : b ≤ a → c ≤ 0 → a * c ≤ b * c := -by classical; exact decidable.mul_le_mul_of_nonpos_right - --- See Note [decidable namespace] -protected lemma decidable.mul_nonneg_of_nonpos_of_nonpos [@decidable_rel α (≤)] - {a b : α} (ha : a ≤ 0) (hb : b ≤ 0) : 0 ≤ a * b := -have 0 * b ≤ a * b, from decidable.mul_le_mul_of_nonpos_right ha hb, -by rwa zero_mul at this - -lemma mul_nonneg_of_nonpos_of_nonpos {a b : α} : a ≤ 0 → b ≤ 0 → 0 ≤ a * b := -by classical; exact decidable.mul_nonneg_of_nonpos_of_nonpos - -lemma mul_lt_mul_of_neg_left {a b c : α} (h : b < a) (hc : c < 0) : c * a < c * b := -have -c > 0, from neg_pos_of_neg hc, -have -c * b < -c * a, from mul_lt_mul_of_pos_left h this, -have -(c * b) < -(c * a), by rwa [neg_mul, neg_mul] at this, -lt_of_neg_lt_neg this - -lemma mul_lt_mul_of_neg_right {a b c : α} (h : b < a) (hc : c < 0) : a * c < b * c := -have -c > 0, from neg_pos_of_neg hc, -have b * -c < a * -c, from mul_lt_mul_of_pos_right h this, -have -(b * c) < -(a * c), by rwa [mul_neg, mul_neg] at this, -lt_of_neg_lt_neg this - -lemma mul_pos_of_neg_of_neg {a b : α} (ha : a < 0) (hb : b < 0) : 0 < a * b := -have 0 * b < a * b, from mul_lt_mul_of_neg_right ha hb, -by rwa zero_mul at this - -lemma decidable.antitone_mul_left [@decidable_rel α (≤)] {a : α} (ha : a ≤ 0) : - antitone ((*) a) := -λ b c b_le_c, decidable.mul_le_mul_of_nonpos_left b_le_c ha - -lemma antitone_mul_left {a : α} (ha : a ≤ 0) : antitone ((*) a) := -λ b c b_le_c, mul_le_mul_of_nonpos_left b_le_c ha - -lemma decidable.antitone_mul_right [@decidable_rel α (≤)] {a : α} (ha : a ≤ 0) : - antitone (λ x, x * a) := -λ b c b_le_c, decidable.mul_le_mul_of_nonpos_right b_le_c ha - -lemma antitone_mul_right {a : α} (ha : a ≤ 0) : antitone (λ x, x * a) := -λ b c b_le_c, mul_le_mul_of_nonpos_right b_le_c ha - -lemma strict_anti_mul_left {a : α} (ha : a < 0) : strict_anti ((*) a) := -λ b c b_lt_c, mul_lt_mul_of_neg_left b_lt_c ha - -lemma strict_anti_mul_right {a : α} (ha : a < 0) : strict_anti (λ x, x * a) := -λ b c b_lt_c, mul_lt_mul_of_neg_right b_lt_c ha - -/-- Pullback an `ordered_ring` under an injective map. -See note [reducible non-instances]. -/ -@[reducible] -def function.injective.ordered_ring {β : Type*} - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] - (f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - ordered_ring β := -{ mul_pos := λ a b a0 b0, show f 0 < f (a * b), by { rw [zero, mul], apply mul_pos; rwa ← zero }, - ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast, - ..hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } - -lemma le_iff_exists_nonneg_add (a b : α) : a ≤ b ↔ ∃ c ≥ 0, b = a + c := -⟨λ h, ⟨b - a, sub_nonneg.mpr h, by simp⟩, - λ ⟨c, hc, h⟩, by { rw [h, le_add_iff_nonneg_right], exact hc }⟩ - -end ordered_ring - -section ordered_comm_ring - -/-- An `ordered_comm_ring α` is a commutative ring `α` with a partial order such that -addition is monotone and multiplication by a positive number is strictly monotone. -/ -@[protect_proj] -class ordered_comm_ring (α : Type u) extends ordered_ring α, comm_ring α - -@[priority 100] -- See note [lower instance priority] -instance ordered_comm_ring.to_ordered_comm_semiring {α : Type u} [ordered_comm_ring α] : - ordered_comm_semiring α := -{ .. (by apply_instance : ordered_semiring α), - .. ‹ordered_comm_ring α› } - -/-- Pullback an `ordered_comm_ring` under an injective map. -See note [reducible non-instances]. -/ -@[reducible] -def function.injective.ordered_comm_ring [ordered_comm_ring α] {β : Type*} - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_pow β ℕ] [has_smul ℕ β] [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] - (f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : - ordered_comm_ring β := -{ ..hf.ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, - ..hf.comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } - -end ordered_comm_ring - -/-- A `linear_ordered_ring α` is a ring `α` with a linear order such that -addition is monotone and multiplication by a positive number is strictly monotone. -/ -@[protect_proj] class linear_ordered_ring (α : Type u) - extends ordered_ring α, linear_order α, nontrivial α - -@[priority 100] -- see Note [lower instance priority] -instance linear_ordered_ring.to_linear_ordered_add_comm_group [s : linear_ordered_ring α] : - linear_ordered_add_comm_group α := -{ .. s } - -section linear_ordered_semiring -variables [linear_ordered_semiring α] {a b c : α} - -local attribute [instance] linear_ordered_semiring.decidable_le - lemma le_of_mul_le_of_one_le {a b c : α} (h : a * c ≤ b) (hb : 0 ≤ b) (hc : 1 ≤ c) : a ≤ b := -have h' : a * c ≤ b * c, from calc - a * c ≤ b : h - ... = b * 1 : by rewrite mul_one - ... ≤ b * c : decidable.mul_le_mul_of_nonneg_left hc hb, -le_of_mul_le_mul_right h' (zero_lt_one.trans_le hc) +le_of_mul_le_mul_right (h.trans $ le_mul_of_one_le_right hb hc) $ zero_lt_one.trans_le hc lemma nonneg_le_nonneg_of_sq_le_sq {a b : α} (hb : 0 ≤ b) (h : a * a ≤ b * b) : a ≤ b := -le_of_not_gt (λhab, (decidable.mul_self_lt_mul_self hb hab).not_le h) +le_of_not_gt $ λ hab, (mul_self_lt_mul_self hb hab).not_le h lemma mul_self_le_mul_self_iff {a b : α} (h1 : 0 ≤ a) (h2 : 0 ≤ b) : a ≤ b ↔ a * a ≤ b * b := -⟨decidable.mul_self_le_mul_self h1, nonneg_le_nonneg_of_sq_le_sq h2⟩ +⟨mul_self_le_mul_self h1, nonneg_le_nonneg_of_sq_le_sq h2⟩ lemma mul_self_lt_mul_self_iff {a b : α} (h1 : 0 ≤ a) (h2 : 0 ≤ b) : a < b ↔ a * a < b * b := -((@decidable.strict_mono_on_mul_self α _ _).lt_iff_lt h1 h2).symm +((@strict_mono_on_mul_self α _).lt_iff_lt h1 h2).symm lemma mul_self_inj {a b : α} (h1 : 0 ≤ a) (h2 : 0 ≤ b) : a * a = b * b ↔ a = b := -(@decidable.strict_mono_on_mul_self α _ _).inj_on.eq_iff h1 h2 +(@strict_mono_on_mul_self α _).inj_on.eq_iff h1 h2 end linear_ordered_semiring +@[priority 100] -- See note [lower instance priority] +instance linear_ordered_comm_semiring.to_linear_ordered_cancel_add_comm_monoid + [linear_ordered_comm_semiring α] : linear_ordered_cancel_add_comm_monoid α := +{ ..‹linear_ordered_comm_semiring α› } + section linear_ordered_ring variables [linear_ordered_ring α] {a b c : α} @@ -1040,7 +865,11 @@ local attribute [instance] linear_ordered_ring.decidable_le linear_ordered_ring. @[priority 100] -- see Note [lower instance priority] instance linear_ordered_ring.to_linear_ordered_semiring : linear_ordered_semiring α := -{ ..‹linear_ordered_ring α›, ..ordered_ring.to_ordered_semiring } +{ ..‹linear_ordered_ring α›, ..strict_ordered_ring.to_strict_ordered_semiring } + +@[priority 100] -- see Note [lower instance priority] +instance linear_ordered_ring.to_linear_ordered_add_comm_group : linear_ordered_add_comm_group α := +{ ..‹linear_ordered_ring α› } @[priority 100] -- see Note [lower instance priority] instance linear_ordered_ring.is_domain : is_domain α := @@ -1059,8 +888,7 @@ instance linear_ordered_ring.is_domain : is_domain α := lemma abs_mul (a b : α) : |a * b| = |a| * |b| := begin - haveI := @linear_order.decidable_le α _, - rw [abs_eq (decidable.mul_nonneg (abs_nonneg a) (abs_nonneg b))], + rw [abs_eq (mul_nonneg (abs_nonneg a) (abs_nonneg b))], cases le_total a 0 with ha ha; cases le_total b 0 with hb hb; simp only [abs_of_nonpos, abs_of_nonneg, true_or, or_true, eq_self_iff_true, neg_mul, mul_neg, neg_neg, *] @@ -1083,9 +911,8 @@ lemma mul_neg_iff : a * b < 0 ↔ 0 < a ∧ b < 0 ∨ a < 0 ∧ 0 < b := by rw [← neg_pos, neg_mul_eq_mul_neg, mul_pos_iff, neg_pos, neg_lt_zero] lemma mul_nonneg_iff : 0 ≤ a * b ↔ 0 ≤ a ∧ 0 ≤ b ∨ a ≤ 0 ∧ b ≤ 0 := -by haveI := @linear_order.decidable_le α _; exact ⟨nonneg_and_nonneg_or_nonpos_and_nonpos_of_mul_nnonneg, - λ h, h.elim (and_imp.2 decidable.mul_nonneg) (and_imp.2 decidable.mul_nonneg_of_nonpos_of_nonpos)⟩ + λ h, h.elim (and_imp.2 mul_nonneg) (and_imp.2 mul_nonneg_of_nonpos_of_nonpos)⟩ /-- Out of three elements of a `linear_ordered_ring`, two must have the same sign. -/ lemma mul_nonneg_of_three (a b c : α) : @@ -1140,15 +967,6 @@ begin simp [ha], end -lemma gt_of_mul_lt_mul_neg_left (h : c * a < c * b) (hc : c ≤ 0) : b < a := -have nhc : 0 ≤ -c, from neg_nonneg_of_nonpos hc, -have h2 : -(c * b) < -(c * a), from neg_lt_neg h, -have h3 : (-c) * b < (-c) * a, from calc - (-c) * b = - (c * b) : by rewrite neg_mul_eq_neg_mul - ... < -(c * a) : h2 - ... = (-c) * a : by rewrite neg_mul_eq_neg_mul, -lt_of_mul_lt_mul_left h3 nhc - lemma neg_one_lt_zero : -1 < (0:α) := neg_lt_zero.2 zero_lt_one @[simp] lemma mul_le_mul_left_of_neg {a b c : α} (h : c < 0) : c * a ≤ c * b ↔ b ≤ a := @@ -1189,9 +1007,8 @@ end lemma mul_self_le_mul_self_of_le_of_neg_le {x y : α} (h₁ : x ≤ y) (h₂ : -x ≤ y) : x * x ≤ y * y := begin - haveI := @linear_order.decidable_le α _, rw [← abs_mul_abs_self x], - exact decidable.mul_self_le_mul_self (abs_nonneg x) (abs_le.2 ⟨neg_le.2 h₂, h₁⟩) + exact mul_self_le_mul_self (abs_nonneg x) (abs_le.2 ⟨neg_le.2 h₂, h₁⟩) end lemma nonneg_of_mul_nonpos_left {a b : α} (h : a * b ≤ 0) (hb : b < 0) : 0 ≤ a := @@ -1201,12 +1018,10 @@ lemma nonneg_of_mul_nonpos_right {a b : α} (h : a * b ≤ 0) (ha : a < 0) : 0 le_of_not_gt (λ hb, absurd h (mul_pos_of_neg_of_neg ha hb).not_le) lemma pos_of_mul_neg_left {a b : α} (h : a * b < 0) (hb : b ≤ 0) : 0 < a := -by haveI := @linear_order.decidable_le α _; exact -lt_of_not_ge (λ ha, absurd h (decidable.mul_nonneg_of_nonpos_of_nonpos ha hb).not_lt) +lt_of_not_ge (λ ha, absurd h (mul_nonneg_of_nonpos_of_nonpos ha hb).not_lt) lemma pos_of_mul_neg_right {a b : α} (h : a * b < 0) (ha : a ≤ 0) : 0 < b := -by haveI := @linear_order.decidable_le α _; exact -lt_of_not_ge (λ hb, absurd h (decidable.mul_nonneg_of_nonpos_of_nonpos ha hb).not_lt) +lt_of_not_ge (λ hb, absurd h (mul_nonneg_of_nonpos_of_nonpos ha hb).not_lt) lemma neg_iff_pos_of_mul_neg (hab : a * b < 0) : a < 0 ↔ 0 < b := ⟨pos_of_mul_neg_right hab ∘ le_of_lt, neg_of_mul_neg_left hab ∘ le_of_lt⟩ @@ -1242,35 +1057,11 @@ end lemma abs_le_one_iff_mul_self_le_one : |a| ≤ 1 ↔ a * a ≤ 1 := by simpa only [abs_one, one_mul] using @abs_le_iff_mul_self_le α _ a 1 -/-- Pullback a `linear_ordered_ring` under an injective map. -See note [reducible non-instances]. -/ -@[reducible] -def function.injective.linear_ordered_ring {β : Type*} - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] [has_int_cast β] - [has_sup β] [has_inf β] - (f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1) - (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) - (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) - (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) - (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) - (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_ring β := -{ .. linear_order.lift f hf hsup hinf, - .. pullback_nonzero f zero one, - .. hf.ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } - end linear_ordered_ring -/-- A `linear_ordered_comm_ring α` is a commutative ring `α` with a linear order -such that addition is monotone and multiplication by a positive number is strictly monotone. -/ -@[protect_proj] -class linear_ordered_comm_ring (α : Type u) extends linear_ordered_ring α, comm_monoid α - @[priority 100] -- see Note [lower instance priority] -instance linear_ordered_comm_ring.to_ordered_comm_ring [d : linear_ordered_comm_ring α] : - ordered_comm_ring α := +instance linear_ordered_comm_ring.to_strict_ordered_comm_ring [d : linear_ordered_comm_ring α] : + strict_ordered_comm_ring α := { ..d } @[priority 100] -- see Note [lower instance priority] @@ -1284,11 +1075,10 @@ variables [linear_ordered_comm_ring α] {a b c d : α} lemma max_mul_mul_le_max_mul_max (b c : α) (ha : 0 ≤ a) (hd: 0 ≤ d) : max (a * b) (d * c) ≤ max a c * max d b := -by haveI := @linear_order.decidable_le α _; exact have ba : b * a ≤ max d b * max c a, from - decidable.mul_le_mul (le_max_right d b) (le_max_right c a) ha (le_trans hd (le_max_left d b)), + mul_le_mul (le_max_right d b) (le_max_right c a) ha (le_trans hd (le_max_left d b)), have cd : c * d ≤ max a c * max b d, from - decidable.mul_le_mul (le_max_right a c) (le_max_right b d) hd (le_trans ha (le_max_left a c)), + mul_le_mul (le_max_right a c) (le_max_right b d) hd (le_trans ha (le_max_left a c)), max_le (by simpa [mul_comm, max_comm] using ba) (by simpa [mul_comm, max_comm] using cd) @@ -1321,30 +1111,182 @@ lemma abs_dvd_abs (a b : α) : |a| ∣ |b| ↔ a ∣ b := end -section linear_ordered_comm_ring +namespace function.injective + +/-- Pullback an `ordered_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def ordered_semiring [ordered_semiring α] [has_zero β] [has_one β] [has_add β] [has_mul β] + [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) + (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) : + ordered_semiring β := +{ zero_le_one := show f 0 ≤ f 1, by simp only [zero, one, zero_le_one], + mul_le_mul_of_nonneg_left := λ a b c h hc, show f (c * a) ≤ f (c * b), + by { rw [mul, mul], refine mul_le_mul_of_nonneg_left h _, rwa ←zero }, + mul_le_mul_of_nonneg_right := λ a b c h hc, show f (a * c) ≤ f (b * c), + by { rw [mul, mul], refine mul_le_mul_of_nonneg_right h _, rwa ←zero }, + ..hf.ordered_add_comm_monoid f zero add nsmul, + ..hf.semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback an `ordered_comm_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def ordered_comm_semiring [ordered_comm_semiring α] [has_zero β] [has_one β] [has_add β] + [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) (hf : injective f) + (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) : + ordered_comm_semiring β := +{ ..hf.comm_semiring f zero one add mul nsmul npow nat_cast, + ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback an `ordered_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def ordered_ring [ordered_ring α] [has_zero β] [has_one β] [has_add β] [has_mul β] + [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] + [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + ordered_ring β := +{ mul_nonneg := λ a b ha hb, show f 0 ≤ f (a * b), + by { rw [zero, mul], apply mul_nonneg; rwa ← zero }, + ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast, + ..hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } -variables [linear_ordered_comm_ring α] +/-- Pullback an `ordered_comm_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def ordered_comm_ring [ordered_comm_ring α] [has_zero β] [has_one β] [has_add β] + [has_mul β] [has_neg β] [has_sub β] [has_pow β ℕ] [has_smul ℕ β] [has_smul ℤ β] [has_nat_cast β] + [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + ordered_comm_ring β := +{ ..hf.ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, + ..hf.comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } -/-- Pullback a `linear_ordered_comm_ring` under an injective map. -See note [reducible non-instances]. -/ -@[reducible] -def function.injective.linear_ordered_comm_ring {β : Type*} - [has_zero β] [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] - [has_pow β ℕ] [has_smul ℕ β] [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] - [has_sup β] [has_inf β] - (f : β → α) (hf : function.injective f) (zero : f 0 = 0) (one : f 1 = 1) +/-- Pullback a `strict_ordered_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def strict_ordered_semiring [strict_ordered_semiring α] [has_zero β] [has_one β] + [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) + (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) : + strict_ordered_semiring β := +{ mul_lt_mul_of_pos_left := λ a b c h hc, show f (c * a) < f (c * b), + by simpa only [mul, zero] using mul_lt_mul_of_pos_left ‹f a < f b› (by rwa ←zero), + mul_lt_mul_of_pos_right := λ a b c h hc, show f (a * c) < f (b * c), + by simpa only [mul, zero] using mul_lt_mul_of_pos_right ‹f a < f b› (by rwa ←zero), + ..hf.ordered_cancel_add_comm_monoid f zero add nsmul, + ..hf.ordered_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback a `strict_ordered_comm_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def strict_ordered_comm_semiring [strict_ordered_comm_semiring α] [has_zero β] [has_one β] + [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] (f : β → α) + (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) : + strict_ordered_comm_semiring β := +{ ..hf.comm_semiring f zero one add mul nsmul npow nat_cast, + ..hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback a `strict_ordered_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def strict_ordered_ring [strict_ordered_ring α] [has_zero β] [has_one β] [has_add β] + [has_mul β] [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] + [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + strict_ordered_ring β := +{ mul_pos := λ a b a0 b0, show f 0 < f (a * b), by { rw [zero, mul], apply mul_pos; rwa ← zero }, + ..hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast, + ..hf.ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + +/-- Pullback a `strict_ordered_comm_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def strict_ordered_comm_ring [strict_ordered_comm_ring α] [has_zero β] + [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_pow β ℕ] [has_smul ℕ β] + [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] (f : β → α) (hf : injective f) (zero : f 0 = 0) + (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + strict_ordered_comm_ring β := +{ ..hf.strict_ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast, + ..hf.comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + +/-- Pullback a `linear_ordered_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def linear_ordered_semiring [linear_ordered_semiring α] [has_zero β] [has_one β] + [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] [has_sup β] [has_inf β] + (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) + (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_semiring β := +{ .. linear_order.lift f hf hsup hinf, + .. pullback_nonzero f zero one, + .. hf.strict_ordered_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback a `linear_ordered_semiring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def linear_ordered_comm_semiring [linear_ordered_comm_semiring α] + [has_zero β] [has_one β] [has_add β] [has_mul β] [has_pow β ℕ] [has_smul ℕ β] [has_nat_cast β] + [has_sup β] [has_inf β] (f : β → α) (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) + (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) + (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_comm_semiring β := +{ ..hf.linear_ordered_semiring f zero one add mul nsmul npow nat_cast hsup hinf, + ..hf.strict_ordered_comm_semiring f zero one add mul nsmul npow nat_cast } + +/-- Pullback a `linear_ordered_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +def linear_ordered_ring [linear_ordered_ring α] [has_zero β] [has_one β] [has_add β] [has_mul β] + [has_neg β] [has_sub β] [has_smul ℕ β] [has_smul ℤ β] [has_pow β ℕ] [has_nat_cast β] + [has_int_cast β] [has_sup β] [has_inf β] (f : β → α) (hf : injective f) (zero : f 0 = 0) + (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (mul : ∀ x y, f (x * y) = f x * f y) (neg : ∀ x, f (-x) = -f x) (sub : ∀ x y, f (x - y) = f x - f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_ring β := +{ .. linear_order.lift f hf hsup hinf, + .. pullback_nonzero f zero one, + .. hf.strict_ordered_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + +/-- Pullback a `linear_ordered_comm_ring` under an injective map. -/ +@[reducible] -- See note [reducible non-instances] +protected def linear_ordered_comm_ring [linear_ordered_comm_ring α] [has_zero β] + [has_one β] [has_add β] [has_mul β] [has_neg β] [has_sub β] [has_pow β ℕ] [has_smul ℕ β] + [has_smul ℤ β] [has_nat_cast β] [has_int_cast β] [has_sup β] [has_inf β] (f : β → α) + (hf : injective f) (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (mul : ∀ x y, f (x * y) = f x * f y) (neg : ∀ x, f (-x) = -f x) + (sub : ∀ x y, f (x - y) = f x - f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) + (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : linear_ordered_comm_ring β := { .. linear_order.lift f hf hsup hinf, .. pullback_nonzero f zero one, - .. hf.ordered_comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } + .. hf.strict_ordered_comm_ring f zero one add mul neg sub nsmul zsmul npow nat_cast int_cast } -end linear_ordered_comm_ring +end function.injective + +/-! ### Positive cones -/ namespace ring @@ -1372,14 +1314,12 @@ add_decl_doc total_positive_cone.to_total_positive_cone end ring -namespace ordered_ring +namespace strict_ordered_ring open ring -/-- Construct an `ordered_ring` by -designating a positive cone in an existing `ring`. -/ -def mk_of_positive_cone {α : Type*} [ring α] (C : positive_cone α) : - ordered_ring α := +/-- Construct a `strict_ordered_ring` by designating a positive cone in an existing `ring`. -/ +def mk_of_positive_cone {α : Type*} [ring α] (C : positive_cone α) : strict_ordered_ring α := { zero_le_one := by { change C.nonneg (1 - 0), convert C.one_nonneg, simp, }, mul_pos := λ x y xp yp, begin change C.pos (x*y - 0), @@ -1389,7 +1329,7 @@ def mk_of_positive_cone {α : Type*} [ring α] (C : positive_cone α) : ..‹ring α›, ..ordered_add_comm_group.mk_of_positive_cone C.to_positive_cone } -end ordered_ring +end strict_ordered_ring namespace linear_ordered_ring @@ -1405,19 +1345,11 @@ def mk_of_positive_cone {α : Type*} [ring α] (C : total_positive_cone α) : rw [←h, C.pos_iff] at one_pos, simpa using one_pos, end⟩, - ..ordered_ring.mk_of_positive_cone C.to_positive_cone, + ..strict_ordered_ring.mk_of_positive_cone C.to_positive_cone, ..linear_ordered_add_comm_group.mk_of_positive_cone C.to_total_positive_cone, } end linear_ordered_ring -/-- A canonically ordered commutative semiring is an ordered, commutative semiring -in which `a ≤ b` iff there exists `c` with `b = a + c`. This is satisfied by the -natural numbers, for example, but not the integers or other ordered groups. -/ -@[protect_proj] -class canonically_ordered_comm_semiring (α : Type*) extends - canonically_ordered_add_monoid α, comm_semiring α := -(eq_zero_or_eq_zero_of_mul_eq_zero : ∀ a b : α, a * b = 0 → a = 0 ∨ b = 0) - namespace canonically_ordered_comm_semiring variables [canonically_ordered_comm_semiring α] {a b : α} @@ -1434,16 +1366,12 @@ begin apply self_le_add_right end -@[priority 200] -- see Note [lower instance priority] -instance canonically_ordered_comm_semiring.to_pos_mul_mono : pos_mul_mono α := -⟨λ x a b h, by { obtain ⟨d, rfl⟩ := exists_add_of_le h, simp_rw [left_distrib, le_self_add], }⟩ - -@[priority 200] -- see Note [lower instance priority] -instance canonically_ordered_comm_semiring.to_mul_pos_mono : mul_pos_mono α := -⟨λ x a b h, by { obtain ⟨d, rfl⟩ := exists_add_of_le h, simp_rw [right_distrib, le_self_add], }⟩ - -/-- A version of `zero_lt_one : 0 < 1` for a `canonically_ordered_comm_semiring`. -/ -lemma zero_lt_one [nontrivial α] : (0:α) < 1 := (zero_le 1).lt_of_ne zero_ne_one +@[priority 100] -- see Note [lower instance priority] +instance to_ordered_comm_semiring : ordered_comm_semiring α := +{ zero_le_one := zero_le _, + mul_le_mul_of_nonneg_left := λ a b c h _, mul_le_mul_left' h _, + mul_le_mul_of_nonneg_right := λ a b c h _, mul_le_mul_right' h _, + ..‹canonically_ordered_comm_semiring α› } @[simp] lemma mul_pos : 0 < a * b ↔ (0 < a) ∧ (0 < b) := by simp only [pos_iff_ne_zero, ne.def, mul_eq_zero, not_or_distrib] diff --git a/src/algebra/order/ring_lemmas.lean b/src/algebra/order/ring_lemmas.lean index 0df10df40ffd6..9415bda48665b 100644 --- a/src/algebra/order/ring_lemmas.lean +++ b/src/algebra/order/ring_lemmas.lean @@ -316,6 +316,10 @@ lemma mul_le_mul [pos_mul_mono α] [mul_pos_mono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (c0 : 0 ≤ c) (b0 : 0 ≤ b) : a * c ≤ b * d := (mul_le_mul_of_nonneg_right h₁ c0).trans $ mul_le_mul_of_nonneg_left h₂ b0 +lemma mul_self_le_mul_self [pos_mul_mono α] [mul_pos_mono α] (ha : 0 ≤ a) (hab : a ≤ b) : + a * a ≤ b * b := +mul_le_mul hab hab ha $ ha.trans hab + lemma mul_le_of_mul_le_of_nonneg_left [pos_mul_mono α] (h : a * b ≤ c) (hle : d ≤ b) (a0 : 0 ≤ a) : a * d ≤ c := (mul_le_mul_of_nonneg_left hle a0).trans h diff --git a/src/algebra/ring/prod.lean b/src/algebra/ring/prod.lean index 894e8bc5fa37e..41946099d8546 100644 --- a/src/algebra/ring/prod.lean +++ b/src/algebra/ring/prod.lean @@ -21,7 +21,7 @@ trivial `simp` lemmas, and define the following operations on `ring_hom`s and si sends `(x, y)` to `(f x, g y)`. -/ -variables {R : Type*} {R' : Type*} {S : Type*} {S' : Type*} {T : Type*} {T' : Type*} +variables {α β R R' S S' T T' : Type*} namespace prod @@ -258,3 +258,24 @@ begin { exact zero_ne_one h.symm }, { exact zero_ne_one h.symm } end + +/-! ### Order -/ + +instance [ordered_semiring α] [ordered_semiring β] : ordered_semiring (α × β) := +{ add_le_add_left := λ _ _, add_le_add_left, + zero_le_one := ⟨zero_le_one, zero_le_one⟩, + mul_le_mul_of_nonneg_left := λ a b c hab hc, + ⟨mul_le_mul_of_nonneg_left hab.1 hc.1, mul_le_mul_of_nonneg_left hab.2 hc.2⟩, + mul_le_mul_of_nonneg_right := λ a b c hab hc, + ⟨mul_le_mul_of_nonneg_right hab.1 hc.1, mul_le_mul_of_nonneg_right hab.2 hc.2⟩, + ..prod.semiring, ..prod.partial_order _ _ } + +instance [ordered_comm_semiring α] [ordered_comm_semiring β] : ordered_comm_semiring (α × β) := +{ ..prod.comm_semiring, ..prod.ordered_semiring } + +instance [ordered_ring α] [ordered_ring β] : ordered_ring (α × β) := +{ mul_nonneg := λ a b ha hb, ⟨mul_nonneg ha.1 hb.1, mul_nonneg ha.2 hb.2⟩, + ..prod.ring, ..prod.ordered_semiring } + +instance [ordered_comm_ring α] [ordered_comm_ring β] : ordered_comm_ring (α × β) := +{ ..prod.comm_ring, ..prod.ordered_ring } diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index 0ebac40c09dc4..a3d0acf4558cf 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -160,18 +160,6 @@ begin exact add_add_add_comm _ _ _ _ } end -lemma convex_open_segment (a b : E) : convex 𝕜 (open_segment 𝕜 a b) := -begin - rw convex_iff_open_segment_subset, - rintro p ⟨ap, bp, hap, hbp, habp, rfl⟩ q ⟨aq, bq, haq, hbq, habq, rfl⟩ z ⟨a, b, ha, hb, hab, rfl⟩, - refine ⟨a * ap + b * aq, a * bp + b * bq, - add_pos (mul_pos ha hap) (mul_pos hb haq), - add_pos (mul_pos ha hbp) (mul_pos hb hbq), _, _⟩, - { rw [add_add_add_comm, ←mul_add, ←mul_add, habp, habq, mul_one, mul_one, hab] }, - { simp_rw [add_smul, mul_smul, smul_add], - exact add_add_add_comm _ _ _ _ } -end - lemma convex.linear_image (hs : convex 𝕜 s) (f : E →ₗ[𝕜] F) : convex 𝕜 (f '' s) := begin intros x hx y hy a b ha hb hab, @@ -390,6 +378,20 @@ by simpa only [←image_smul, ←image_vadd, image_image] using (hs.smul c).vadd end add_comm_monoid end ordered_comm_semiring +section strict_ordered_comm_semiring +variables [strict_ordered_comm_semiring 𝕜] [add_comm_group E] [module 𝕜 E] + +lemma convex_open_segment (a b : E) : convex 𝕜 (open_segment 𝕜 a b) := +begin + rw convex_iff_open_segment_subset, + rintro p ⟨ap, bp, hap, hbp, habp, rfl⟩ q ⟨aq, bq, haq, hbq, habq, rfl⟩ z ⟨a, b, ha, hb, hab, rfl⟩, + refine ⟨a * ap + b * aq, a * bp + b * bq, by positivity, by positivity, _, _⟩, + { rw [add_add_add_comm, ←mul_add, ←mul_add, habp, habq, mul_one, mul_one, hab] }, + { simp_rw [add_smul, mul_smul, smul_add, add_add_add_comm] } +end + +end strict_ordered_comm_semiring + section ordered_ring variables [ordered_ring 𝕜] diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index 638f58f508156..cb21a9e52c901 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -379,9 +379,8 @@ h₁.wbtw.trans_sbtw_right h₂ end ordered_ring -section ordered_comm_ring - -variables [ordered_comm_ring R] [add_comm_group V] [module R V] [add_torsor V P] +section strict_ordered_comm_ring +variables [strict_ordered_comm_ring R] [add_comm_group V] [module R V] [add_torsor V P] include V @@ -411,7 +410,7 @@ begin same_ray_nonneg_smul_right (z -ᵥ x) (sub_nonneg.2 ht1) end -end ordered_comm_ring +end strict_ordered_comm_ring section linear_ordered_field diff --git a/src/analysis/convex/segment.lean b/src/analysis/convex/segment.lean index 15a02d5fbdf9d..049ce613bbf9c 100644 --- a/src/analysis/convex/segment.lean +++ b/src/analysis/convex/segment.lean @@ -221,7 +221,7 @@ open_segment_translate_preimage 𝕜 a b c ▸ image_preimage_eq _ $ add_left_su end ordered_ring -lemma same_ray_of_mem_segment [ordered_comm_ring 𝕜] [add_comm_group E] [module 𝕜 E] +lemma same_ray_of_mem_segment [strict_ordered_comm_ring 𝕜] [add_comm_group E] [module 𝕜 E] {x y z : E} (h : x ∈ [y -[𝕜] z]) : same_ray 𝕜 (x - y) (z - x) := begin rw segment_eq_image' at h, diff --git a/src/analysis/convex/specific_functions.lean b/src/analysis/convex/specific_functions.lean index f8ad2ed96ed8b..0f2cd196c667d 100644 --- a/src/analysis/convex/specific_functions.lean +++ b/src/analysis/convex/specific_functions.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov, Sébastien Gouëzel import analysis.calculus.mean_value import analysis.special_functions.pow_deriv import analysis.special_functions.sqrt +import data.int.basic /-! # Collection of convex functions @@ -128,7 +129,8 @@ begin refine mul_nonneg ihn _, generalize : (1 + 1) * n = k, cases le_or_lt m k with hmk hmk, { have : m ≤ k + 1, from hmk.trans (lt_add_one ↑k).le, - exact mul_nonneg_of_nonpos_of_nonpos (sub_nonpos_of_le hmk) (sub_nonpos_of_le this) }, + convert mul_nonneg_of_nonpos_of_nonpos (sub_nonpos_of_le hmk) _, + convert sub_nonpos_of_le this }, { exact mul_nonneg (sub_nonneg_of_le hmk.le) (sub_nonneg_of_le hmk) } end diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index 272bd7a5d729c..21351f9a4e86a 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -60,14 +60,8 @@ We have the expression -/ lemma log_stirling_seq_formula (n : ℕ) : log (stirling_seq n.succ) = log n.succ!- 1 / 2 * log (2 * n.succ) - n.succ * log (n.succ / exp 1) := -begin - have h1 : (0 : ℝ) < n.succ! := cast_pos.mpr n.succ.factorial_pos, - have h2 : (0 : ℝ) < (2 * n.succ) := mul_pos two_pos (cast_pos.mpr (succ_pos n)), - have h3 := real.sqrt_pos.mpr h2, - have h4 := pow_pos (div_pos (cast_pos.mpr n.succ_pos ) (exp_pos 1)) n.succ, - have h5 := mul_pos h3 h4, - rw [stirling_seq, log_div, log_mul, sqrt_eq_rpow, log_rpow, log_pow]; linarith, -end +by rw [stirling_seq, log_div, log_mul, sqrt_eq_rpow, log_rpow, log_pow, tsub_tsub]; + try { apply ne_of_gt }; positivity -- TODO: Make `positivity` handle `≠ 0` goals /-- The sequence `log (stirling_seq (m + 1)) - log (stirling_seq (m + 2))` has the series expansion @@ -97,12 +91,8 @@ end /-- The sequence `log ∘ stirling_seq ∘ succ` is monotone decreasing -/ lemma log_stirling_seq'_antitone : antitone (log ∘ stirling_seq ∘ succ) := -begin - have : ∀ {k : ℕ}, 0 < (1 : ℝ) / (2 * k.succ + 1) := - λ k, one_div_pos.mpr (add_pos (mul_pos two_pos (cast_pos.mpr k.succ_pos)) one_pos), - exact antitone_nat_of_succ_le (λ n, sub_nonneg.mp ((log_stirling_seq_diff_has_sum n).nonneg - (λ m, (mul_pos this (pow_pos (pow_pos this 2) m.succ)).le))), -end +antitone_nat_of_succ_le $ λ n, sub_nonneg.mp $ (log_stirling_seq_diff_has_sum n).nonneg $ λ m, + by positivity /-- We have a bound for successive elements in the sequence `log (stirling_seq k)`. @@ -116,13 +106,12 @@ begin ((1 / (2 * n.succ + 1)) ^ 2 / (1 - (1 / (2 * n.succ + 1)) ^ 2)), { refine (has_sum_geometric_of_lt_1 h_nonneg _).mul_left ((1 / (2 * (n.succ : ℝ) + 1)) ^ 2), rw [one_div, inv_pow], - refine inv_lt_one (one_lt_pow ((lt_add_iff_pos_left 1).mpr - (mul_pos two_pos (cast_pos.mpr n.succ_pos))) two_ne_zero) }, + exact inv_lt_one (one_lt_pow ((lt_add_iff_pos_left 1).mpr $ by positivity) two_ne_zero) }, have hab : ∀ (k : ℕ), (1 / (2 * (k.succ : ℝ) + 1)) * ((1 / (2 * n.succ + 1)) ^ 2) ^ k.succ ≤ ((1 / (2 * n.succ + 1)) ^ 2) ^ k.succ, { refine λ k, mul_le_of_le_one_left (pow_nonneg h_nonneg k.succ) _, rw one_div, - exact inv_le_one (le_add_of_nonneg_left (mul_pos two_pos (cast_pos.mpr k.succ_pos)).le) }, + exact inv_le_one (le_add_of_nonneg_left $ by positivity) }, exact has_sum_le hab (log_stirling_seq_diff_has_sum n) g, end @@ -159,18 +148,15 @@ begin have h₁ : ∀ k, log_stirling_seq' k - log_stirling_seq' (k + 1) ≤ 1 / 4 * (1 / k.succ ^ 2) := by { intro k, convert log_stirling_seq_sub_log_stirling_seq_succ k using 1, field_simp, }, have h₂ : ∑ (k : ℕ) in range n, (1 : ℝ) / (k.succ) ^ 2 ≤ d := by - { refine sum_le_tsum (range n) (λ k _, _) - ((summable_nat_add_iff 1).mpr (real.summable_one_div_nat_pow.mpr one_lt_two)), - apply le_of_lt, - rw [one_div_pos, sq_pos_iff], - exact nonzero_of_invertible (succ k), }, + { exact sum_le_tsum (range n) (λ k _, by positivity) + ((summable_nat_add_iff 1).mpr $ real.summable_one_div_nat_pow.mpr one_lt_two) }, calc log (stirling_seq 1) - log (stirling_seq n.succ) = log_stirling_seq' 0 - log_stirling_seq' n : rfl ... = ∑ k in range n, (log_stirling_seq' k - log_stirling_seq' (k + 1)) : by rw ← sum_range_sub' log_stirling_seq' n ... ≤ ∑ k in range n, (1/4) * (1 / k.succ^2) : sum_le_sum (λ k _, h₁ k) ... = 1 / 4 * ∑ k in range n, 1 / k.succ ^ 2 : by rw mul_sum - ... ≤ 1 / 4 * d : (mul_le_mul_left (one_div_pos.mpr four_pos)).mpr h₂, + ... ≤ 1 / 4 * d : mul_le_mul_of_nonneg_left h₂ $ by positivity, end /-- The sequence `log_stirling_seq` is bounded below for `n ≥ 1`. -/ @@ -181,9 +167,7 @@ begin end /-- The sequence `stirling_seq` is positive for `n > 0` -/ -lemma stirling_seq'_pos (n : ℕ) : 0 < stirling_seq n.succ := -div_pos (cast_pos.mpr n.succ.factorial_pos) (mul_pos (real.sqrt_pos.mpr (mul_pos two_pos - (cast_pos.mpr n.succ_pos))) (pow_pos (div_pos (cast_pos.mpr n.succ_pos) (exp_pos 1)) n.succ)) +lemma stirling_seq'_pos (n : ℕ) : 0 < stirling_seq n.succ := by { unfold stirling_seq, positivity } /-- The sequence `stirling_seq` has a positive lower bound. @@ -229,16 +213,11 @@ begin { rw [w, prod_range_zero, cast_zero, mul_zero, pow_zero, one_mul, mul_zero, factorial_zero, cast_one, one_pow, one_pow, one_mul, mul_zero, zero_add, div_one] }, rw [w, prod_range_succ, ←ih, w, _root_.div_mul_div_comm, _root_.div_mul_div_comm], - refine (div_eq_div_iff (ne_of_gt _) (ne_of_gt _)).mpr _, - { exact mul_pos (pow_pos (cast_pos.mpr (2 * n.succ).factorial_pos) 2) - (add_pos (mul_pos two_pos (cast_pos.mpr n.succ_pos)) one_pos) }, - { have h : 0 < 2 * (n : ℝ) + 1 := - add_pos_of_nonneg_of_pos (mul_nonneg zero_le_two n.cast_nonneg) one_pos, - exact mul_pos (mul_pos (pow_pos (cast_pos.mpr (2 * n).factorial_pos) 2) h) - (mul_pos h (add_pos_of_nonneg_of_pos (mul_nonneg zero_le_two n.cast_nonneg) three_pos)) }, - { simp_rw [nat.mul_succ, factorial_succ, succ_eq_add_one, pow_succ], - push_cast, - ring_nf }, + refine (div_eq_div_iff _ _).mpr _, + any_goals { exact ne_of_gt (by positivity) }, + simp_rw [nat.mul_succ, factorial_succ, pow_succ], + push_cast, + ring_nf, end /-- The sequence `n / (2 * n + 1)` tends to `1/2` -/ @@ -259,8 +238,8 @@ lemma stirling_seq_pow_four_div_stirling_seq_pow_two_eq (n : ℕ) (hn : n ≠ 0) begin rw [bit0_eq_two_mul, stirling_seq, pow_mul, stirling_seq, w], simp_rw [div_pow, mul_pow], - rw [sq_sqrt (mul_nonneg two_pos.le n.cast_nonneg), - sq_sqrt (mul_nonneg two_pos.le (2 * n).cast_nonneg)], + rw [sq_sqrt, sq_sqrt], + any_goals { positivity }, have : (n : ℝ) ≠ 0, from cast_ne_zero.mpr hn, have : (exp 1) ≠ 0, from exp_ne_zero 1, have : ((2 * n)!: ℝ) ≠ 0, from cast_ne_zero.mpr (factorial_ne_zero (2 * n)), diff --git a/src/combinatorics/set_family/kleitman.lean b/src/combinatorics/set_family/kleitman.lean index 6949daf8fde3e..9c8e39d497dcf 100644 --- a/src/combinatorics/set_family/kleitman.lean +++ b/src/combinatorics/set_family/kleitman.lean @@ -59,8 +59,8 @@ begin rw [union_sdiff_left, sdiff_eq_inter_compl], refine le_of_mul_le_mul_left _ (pow_pos zero_lt_two $ card α + 1), rw [pow_succ', mul_add, mul_assoc, mul_comm _ 2, mul_assoc], - refine (add_le_add ((mul_le_mul_left $ pow_pos two_pos _).2 - (hf₁ _ $ mem_cons_self _ _).2.2.card_le) $ (mul_le_mul_left two_pos).2 $ + refine (add_le_add ((mul_le_mul_left $ pow_pos (zero_lt_two' ℕ) _).2 + (hf₁ _ $ mem_cons_self _ _).2.2.card_le) $ (mul_le_mul_left $ zero_lt_two' ℕ).2 $ is_upper_set.card_inter_le_finset _ _).trans _, { rw coe_bUnion, exact is_upper_set_Union₂ (λ i hi, hf₂ _ $ subset_cons _ hi) }, diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index 58860cb1368cc..13e65af8fb228 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -645,9 +645,9 @@ lemma not_le_zero_iff {z : ℂ} : ¬z ≤ 0 ↔ 0 < z.re ∨ z.im ≠ 0 := not_l lemma not_lt_zero_iff {z : ℂ} : ¬z < 0 ↔ 0 ≤ z.re ∨ z.im ≠ 0 := not_lt_iff /-- -With `z ≤ w` iff `w - z` is real and nonnegative, `ℂ` is an ordered ring. +With `z ≤ w` iff `w - z` is real and nonnegative, `ℂ` is a strictly ordered ring. -/ -protected def ordered_comm_ring : ordered_comm_ring ℂ := +protected def strict_ordered_comm_ring : strict_ordered_comm_ring ℂ := { zero_le_one := ⟨zero_le_one, rfl⟩, add_le_add_left := λ w z h y, ⟨add_le_add_left h.1 _, congr_arg2 (+) rfl h.2⟩, mul_pos := λ z w hz hw, @@ -655,7 +655,7 @@ protected def ordered_comm_ring : ordered_comm_ring ℂ := .. complex.partial_order, .. complex.comm_ring } -localized "attribute [instance] complex.ordered_comm_ring" in complex_order +localized "attribute [instance] complex.strict_ordered_comm_ring" in complex_order /-- With `z ≤ w` iff `w - z` is real and nonnegative, `ℂ` is a star ordered ring. @@ -673,7 +673,7 @@ protected def star_ordered_ring : star_ordered_ring ℂ := mul_im, mul_zero, neg_zero] } }, { obtain ⟨s, rfl⟩ := h, simp only [←norm_sq_eq_conj_mul_self, norm_sq_nonneg, zero_le_real, star_def] } }, - ..complex.ordered_comm_ring } + ..complex.strict_ordered_comm_ring } localized "attribute [instance] complex.star_ordered_ring" in complex_order diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index 875ee85a9249c..ad162953bd48d 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -65,6 +65,12 @@ instance : comm_ring int := int.one_mul], zsmul_neg' := λ n x, int.neg_mul_eq_neg_mul_symm (n.succ : ℤ) x } +instance : linear_ordered_comm_ring int := +{ add_le_add_left := @int.add_le_add_left, + mul_pos := @int.mul_pos, + zero_le_one := le_of_lt int.zero_lt_one, + .. int.comm_ring, .. int.linear_order, .. int.nontrivial } + /-! ### Extra instances to short-circuit type class resolution These also prevent non-computable instances like `int.normed_comm_ring` being used to construct @@ -85,15 +91,9 @@ instance : comm_semiring int := by apply_instance instance : semiring int := by apply_instance instance : ring int := by apply_instance instance : distrib int := by apply_instance - -instance : linear_ordered_comm_ring int := -{ add_le_add_left := @int.add_le_add_left, - mul_pos := @int.mul_pos, - zero_le_one := le_of_lt int.zero_lt_one, - .. int.comm_ring, .. int.linear_order, .. int.nontrivial } - -instance : linear_ordered_add_comm_group int := -by apply_instance +instance : ordered_comm_ring ℤ := strict_ordered_comm_ring.to_ordered_comm_ring' +instance : ordered_ring ℤ := strict_ordered_ring.to_ordered_ring' +instance : linear_ordered_add_comm_group ℤ := by apply_instance @[simp] lemma add_neg_one (i : ℤ) : i + -1 = i - 1 := rfl @@ -1079,7 +1079,7 @@ protected theorem mul_lt_of_lt_div {a b c : ℤ} (H : 0 < c) (H3 : a < b / c) : lt_of_not_ge $ mt (int.div_le_of_le_mul H) (not_le_of_gt H3) protected theorem mul_le_of_le_div {a b c : ℤ} (H1 : 0 < c) (H2 : a ≤ b / c) : a * c ≤ b := -le_trans (decidable.mul_le_mul_of_nonneg_right H2 (le_of_lt H1)) (int.div_mul_le _ (ne_of_gt H1)) +le_trans (mul_le_mul_of_nonneg_right H2 (le_of_lt H1)) (int.div_mul_le _ (ne_of_gt H1)) protected theorem le_div_of_mul_le {a b c : ℤ} (H1 : 0 < c) (H2 : a * c ≤ b) : a ≤ b / c := le_of_lt_add_one $ lt_of_mul_lt_mul_right @@ -1102,7 +1102,7 @@ protected theorem div_lt_iff_lt_mul {a b c : ℤ} (H : 0 < c) : a / c < b ↔ a protected theorem le_mul_of_div_le {a b c : ℤ} (H1 : 0 ≤ b) (H2 : b ∣ a) (H3 : a / b ≤ c) : a ≤ c * b := -by rw [← int.div_mul_cancel H2]; exact decidable.mul_le_mul_of_nonneg_right H3 H1 +by rw [← int.div_mul_cancel H2]; exact mul_le_mul_of_nonneg_right H3 H1 protected theorem lt_div_of_mul_lt {a b c : ℤ} (H1 : 0 ≤ b) (H2 : b ∣ c) (H3 : a * b < c) : a < c / b := diff --git a/src/data/list/big_operators.lean b/src/data/list/big_operators.lean index 345c581d5a53a..5ddb9c831b581 100644 --- a/src/data/list/big_operators.lean +++ b/src/data/list/big_operators.lean @@ -509,7 +509,7 @@ end /-- The product of a list of positive natural numbers is positive, and likewise for any nontrivial ordered semiring. -/ -lemma prod_pos [ordered_semiring R] [nontrivial R] (l : list R) (h : ∀ a ∈ l, (0 : R) < a) : +lemma prod_pos [strict_ordered_semiring R] [nontrivial R] (l : list R) (h : ∀ a ∈ l, (0 : R) < a) : 0 < l.prod := begin induction l with a l ih, diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index e1eaeba817a19..5a97adb64eb0a 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -69,19 +69,23 @@ instance : linear_ordered_comm_monoid_with_zero ℕ := ..nat.linear_ordered_comm_semiring, ..(infer_instance : comm_monoid_with_zero ℕ)} -/-! Extra instances to short-circuit type class resolution -/ -instance : add_comm_monoid ℕ := infer_instance -instance : add_monoid ℕ := infer_instance -instance : monoid ℕ := infer_instance -instance : comm_monoid ℕ := infer_instance -instance : comm_semigroup ℕ := infer_instance -instance : semigroup ℕ := infer_instance -instance : add_comm_semigroup ℕ := infer_instance -instance : add_semigroup ℕ := infer_instance -instance : distrib ℕ := infer_instance -instance : semiring ℕ := infer_instance -instance : ordered_semiring ℕ := infer_instance -instance : ordered_comm_semiring ℕ := infer_instance +/-! Extra instances to short-circuit type class resolution and ensure computability -/ +-- Not using `infer_instance` avoids `classical.choice` in the following two +instance : linear_ordered_semiring ℕ := infer_instance +instance : strict_ordered_semiring ℕ := infer_instance +instance : strict_ordered_comm_semiring ℕ := infer_instance +instance : ordered_semiring ℕ := strict_ordered_semiring.to_ordered_semiring' +instance : ordered_comm_semiring ℕ := strict_ordered_comm_semiring.to_ordered_comm_semiring' +instance : add_comm_monoid ℕ := infer_instance +instance : add_monoid ℕ := infer_instance +instance : monoid ℕ := infer_instance +instance : comm_monoid ℕ := infer_instance +instance : comm_semigroup ℕ := infer_instance +instance : semigroup ℕ := infer_instance +instance : add_comm_semigroup ℕ := infer_instance +instance : add_semigroup ℕ := infer_instance +instance : distrib ℕ := infer_instance +instance : semiring ℕ := infer_instance instance : linear_ordered_cancel_add_comm_monoid ℕ := infer_instance instance nat.order_bot : order_bot ℕ := @@ -521,11 +525,11 @@ lemma succ_mul_pos (m : ℕ) (hn : 0 < n) : 0 < (succ m) * n := mul_pos (succ_pos m) hn theorem mul_self_le_mul_self {n m : ℕ} (h : n ≤ m) : n * n ≤ m * m := -decidable.mul_le_mul h h (zero_le _) (zero_le _) +mul_le_mul h h (zero_le _) (zero_le _) theorem mul_self_lt_mul_self : Π {n m : ℕ}, n < m → n * n < m * m | 0 m h := mul_pos h h -| (succ n) m h := decidable.mul_lt_mul h (le_of_lt h) (succ_pos _) (zero_le _) +| (succ n) m h := mul_lt_mul h (le_of_lt h) (succ_pos _) (zero_le _) theorem mul_self_le_mul_self_iff {n m : ℕ} : n ≤ m ↔ n * n ≤ m * m := ⟨mul_self_le_mul_self, le_imp_le_of_lt_imp_lt mul_self_lt_mul_self⟩ @@ -540,13 +544,13 @@ theorem le_mul_self : Π (n : ℕ), n ≤ n * n lemma le_mul_of_pos_left {m n : ℕ} (h : 0 < n) : m ≤ n * m := begin conv {to_lhs, rw [← one_mul(m)]}, - exact decidable.mul_le_mul_of_nonneg_right h.nat_succ_le dec_trivial, + exact mul_le_mul_of_nonneg_right h.nat_succ_le dec_trivial, end lemma le_mul_of_pos_right {m n : ℕ} (h : 0 < n) : m ≤ m * n := begin conv {to_lhs, rw [← mul_one(m)]}, - exact decidable.mul_le_mul_of_nonneg_left h.nat_succ_le dec_trivial, + exact mul_le_mul_of_nonneg_left h.nat_succ_le dec_trivial, end theorem two_mul_ne_two_mul_add_one {n m} : 2 * n ≠ 2 * m + 1 := diff --git a/src/data/nat/cast.lean b/src/data/nat/cast.lean index 8283d0dab77b1..f4081bf2945c6 100644 --- a/src/data/nat/cast.lean +++ b/src/data/nat/cast.lean @@ -57,8 +57,7 @@ lemma cast_comm [non_assoc_semiring α] (n : ℕ) (x : α) : (n : α) * x = x * lemma commute_cast [non_assoc_semiring α] (x : α) (n : ℕ) : commute x n := (n.cast_commute x).symm -section - +section ordered_semiring variables [ordered_semiring α] @[mono] theorem mono_cast : monotone (coe : ℕ → α) := @@ -69,6 +68,16 @@ monotone_nat_of_le_succ $ λ n, by rw [nat.cast_succ]; exact le_add_of_nonneg_ri variable [nontrivial α] +lemma cast_add_one_pos (n : ℕ) : 0 < (n : α) + 1 := +zero_lt_one.trans_le $ le_add_of_nonneg_left n.cast_nonneg + +@[simp] lemma cast_pos {n : ℕ} : (0 : α) < n ↔ 0 < n := by cases n; simp [cast_add_one_pos] + +end ordered_semiring + +section strict_ordered_semiring +variables [strict_ordered_semiring α] [nontrivial α] + @[simp, norm_cast] theorem cast_le {m n : ℕ} : (m : α) ≤ n ↔ m ≤ n := strict_mono_cast.le_iff_le @@ -76,12 +85,6 @@ strict_mono_cast.le_iff_le @[simp, norm_cast, mono] theorem cast_lt {m n : ℕ} : (m : α) < n ↔ m < n := strict_mono_cast.lt_iff_lt -@[simp] theorem cast_pos {n : ℕ} : (0 : α) < n ↔ 0 < n := -by rw [← cast_zero, cast_lt] - -lemma cast_add_one_pos (n : ℕ) : 0 < (n : α) + 1 := - add_pos_of_nonneg_of_pos n.cast_nonneg zero_lt_one - @[simp, norm_cast] theorem one_lt_cast {n : ℕ} : 1 < (n : α) ↔ 1 < n := by rw [← cast_one, cast_lt] @@ -94,7 +97,7 @@ by rw [← cast_one, cast_lt, lt_succ_iff, le_zero_iff] @[simp, norm_cast] theorem cast_le_one {n : ℕ} : (n : α) ≤ 1 ↔ n ≤ 1 := by rw [← cast_one, cast_le] -end +end strict_ordered_semiring @[simp, norm_cast] theorem cast_min [linear_ordered_semiring α] {a b : ℕ} : (↑(min a b) : α) = min a b := diff --git a/src/data/nat/choose/central.lean b/src/data/nat/choose/central.lean index d5e9bebe5f13c..7548eb02c4091 100644 --- a/src/data/nat/choose/central.lean +++ b/src/data/nat/choose/central.lean @@ -79,7 +79,7 @@ begin { norm_num [central_binom, choose] }, obtain ⟨n, rfl⟩ : ∃ m, n = m + 1 := nat.exists_eq_succ_of_ne_zero (zero_lt_four.trans hn).ne', calc 4 ^ (n + 1) < 4 * (n * central_binom n) : - (mul_lt_mul_left zero_lt_four).mpr (IH n n.lt_succ_self (nat.le_of_lt_succ hn)) + (mul_lt_mul_left $ zero_lt_four' ℕ).mpr (IH n n.lt_succ_self (nat.le_of_lt_succ hn)) ... ≤ 2 * (2 * n + 1) * central_binom n : by { rw ← mul_assoc, linarith } ... = (n + 1) * central_binom (n + 1) : (succ_mul_central_binom_succ n).symm, end diff --git a/src/data/nat/prime.lean b/src/data/nat/prime.lean index 6baff00b813e8..69d63b49f863c 100644 --- a/src/data/nat/prime.lean +++ b/src/data/nat/prime.lean @@ -858,8 +858,6 @@ end nat /-! ### Primality prover -/ -open norm_num - namespace tactic namespace norm_num @@ -941,6 +939,8 @@ begin exact not_le_of_lt hd this end +open _root_.norm_num + /-- Given `e` a natural numeral and `d : nat` a factor of it, return `⊢ ¬ prime e`. -/ meta def prove_non_prime (e : expr) (n d₁ : ℕ) : tactic expr := do let e₁ := reflect d₁, diff --git a/src/data/real/basic.lean b/src/data/real/basic.lean index bca0c89a201bf..0b1ff25721612 100644 --- a/src/data/real/basic.lean +++ b/src/data/real/basic.lean @@ -238,7 +238,7 @@ begin simpa only [mk_lt, mk_pos, ← mk_mul] using cau_seq.mul_pos end -instance : ordered_comm_ring ℝ := +instance : strict_ordered_comm_ring ℝ := { add_le_add_left := begin simp only [le_iff_eq_or_lt], @@ -250,12 +250,15 @@ instance : ordered_comm_ring ℝ := mul_pos := @real.mul_pos, .. real.comm_ring, .. real.partial_order, .. real.semiring } -instance : ordered_ring ℝ := by apply_instance -instance : ordered_semiring ℝ := by apply_instance -instance : ordered_add_comm_group ℝ := by apply_instance -instance : ordered_cancel_add_comm_monoid ℝ := by apply_instance -instance : ordered_add_comm_monoid ℝ := by apply_instance -instance : nontrivial ℝ := ⟨⟨0, 1, ne_of_lt real.zero_lt_one⟩⟩ +instance : strict_ordered_ring ℝ := infer_instance +instance : strict_ordered_comm_semiring ℝ := infer_instance +instance : strict_ordered_semiring ℝ := infer_instance +instance : ordered_ring ℝ := infer_instance +instance : ordered_semiring ℝ := infer_instance +instance : ordered_add_comm_group ℝ := infer_instance +instance : ordered_cancel_add_comm_monoid ℝ := infer_instance +instance : ordered_add_comm_monoid ℝ := infer_instance +instance : nontrivial ℝ := ⟨⟨0, 1, ne_of_lt real.zero_lt_one⟩⟩ @[irreducible] private def sup : ℝ → ℝ → ℝ | ⟨x⟩ ⟨y⟩ := @@ -323,7 +326,7 @@ noncomputable instance : linear_order ℝ := lattice.to_linear_order _ noncomputable instance : linear_ordered_comm_ring ℝ := -{ .. real.nontrivial, .. real.ordered_ring, .. real.comm_ring, .. real.linear_order } +{ .. real.nontrivial, .. real.strict_ordered_ring, .. real.comm_ring, .. real.linear_order } /- Extra instances to short-circuit type class resolution -/ noncomputable instance : linear_ordered_ring ℝ := by apply_instance diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 7e1f7dc313b3d..3aa0589e4c1a6 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -233,7 +233,7 @@ lemma coe_ne_zero : (r : ℝ≥0∞) ≠ 0 ↔ r ≠ 0 := not_congr coe_eq_coe lemma coe_two : ((2:ℝ≥0) : ℝ≥0∞) = 2 := by norm_cast lemma to_nnreal_eq_to_nnreal_iff (x y : ℝ≥0∞) : - x.to_nnreal = y.to_nnreal ↔ x = y ∨ (x = 0 ∧ y = ⊤) ∨ (x = ⊤ ∧ y = 0) := + x.to_nnreal = y.to_nnreal ↔ x = y ∨ x = 0 ∧ y = ⊤ ∨ x = ⊤ ∧ y = 0 := begin cases x, { simp only [@eq_comm ℝ≥0 _ y.to_nnreal, @eq_comm ℝ≥0∞ _ y, to_nnreal_eq_zero_iff, @@ -254,13 +254,11 @@ lemma to_real_eq_to_real_iff' {x y : ℝ≥0∞} (hx : x ≠ ⊤) (hy : y ≠ x.to_real = y.to_real ↔ x = y := by simp only [ennreal.to_real, nnreal.coe_eq, to_nnreal_eq_to_nnreal_iff' hx hy] - -protected lemma zero_lt_one : 0 < (1 : ℝ≥0∞) := - canonically_ordered_comm_semiring.zero_lt_one +protected lemma zero_lt_one : 0 < (1 : ℝ≥0∞) := zero_lt_one @[simp] lemma one_lt_two : (1 : ℝ≥0∞) < 2 := coe_one ▸ coe_two ▸ by exact_mod_cast (@one_lt_two ℕ _ _) -@[simp] lemma zero_lt_two : (0:ℝ≥0∞) < 2 := lt_trans ennreal.zero_lt_one one_lt_two +@[simp] lemma zero_lt_two : (0:ℝ≥0∞) < 2 := lt_trans zero_lt_one one_lt_two lemma two_ne_zero : (2:ℝ≥0∞) ≠ 0 := (ne_of_lt zero_lt_two).symm lemma two_ne_top : (2:ℝ≥0∞) ≠ ∞ := coe_two ▸ coe_ne_top @@ -755,7 +753,7 @@ section cancel lemma add_le_cancellable_iff_ne {a : ℝ≥0∞} : add_le_cancellable a ↔ a ≠ ∞ := begin split, - { rintro h rfl, refine ennreal.zero_lt_one.not_le (h _), simp, }, + { rintro h rfl, refine zero_lt_one.not_le (h _), simp, }, { rintro h b c hbc, apply ennreal.le_of_add_le_add_left h hbc } end @@ -1432,7 +1430,7 @@ lemma exists_mem_Ico_zpow begin lift x to ℝ≥0 using h'x, lift y to ℝ≥0 using h'y, - have A : y ≠ 0, { simpa only [ne.def, coe_eq_zero] using (ennreal.zero_lt_one.trans hy).ne' }, + have A : y ≠ 0, { simpa only [ne.def, coe_eq_zero] using (zero_lt_one.trans hy).ne' }, obtain ⟨n, hn, h'n⟩ : ∃ n : ℤ, y ^ n ≤ x ∧ x < y ^ (n + 1), { refine nnreal.exists_mem_Ico_zpow _ (one_lt_coe_iff.1 hy), simpa only [ne.def, coe_eq_zero] using hx }, @@ -1447,7 +1445,7 @@ lemma exists_mem_Ioc_zpow begin lift x to ℝ≥0 using h'x, lift y to ℝ≥0 using h'y, - have A : y ≠ 0, { simpa only [ne.def, coe_eq_zero] using (ennreal.zero_lt_one.trans hy).ne' }, + have A : y ≠ 0, { simpa only [ne.def, coe_eq_zero] using (zero_lt_one.trans hy).ne' }, obtain ⟨n, hn, h'n⟩ : ∃ n : ℤ, y ^ n < x ∧ x ≤ y ^ (n + 1), { refine nnreal.exists_mem_Ioc_zpow _ (one_lt_coe_iff.1 hy), simpa only [ne.def, coe_eq_zero] using hx }, @@ -1467,9 +1465,9 @@ begin { rintros ⟨n, hn, h'n⟩, split, { apply lt_of_lt_of_le _ hn, - exact ennreal.zpow_pos (ennreal.zero_lt_one.trans hy).ne' h'y _ }, + exact ennreal.zpow_pos (zero_lt_one.trans hy).ne' h'y _ }, { apply lt_trans h'n _, - exact ennreal.zpow_lt_top (ennreal.zero_lt_one.trans hy).ne' h'y _ } } + exact ennreal.zpow_lt_top (zero_lt_one.trans hy).ne' h'y _ } } end lemma zpow_le_of_le {x : ℝ≥0∞} (hx : 1 ≤ x) {a b : ℤ} (h : a ≤ b) : x ^ a ≤ x ^ b := diff --git a/src/data/real/hyperreal.lean b/src/data/real/hyperreal.lean index 9c8b7679f2ebb..5f8f7ae1e1ace 100644 --- a/src/data/real/hyperreal.lean +++ b/src/data/real/hyperreal.lean @@ -42,15 +42,11 @@ lemma coe_ne_coe {x y : ℝ} : (x : ℝ*) ≠ y ↔ x ≠ y := coe_eq_coe.not @[simp, norm_cast] lemma coe_div (x y : ℝ) : ↑(x / y) = (x / y : ℝ*) := rfl @[simp, norm_cast] lemma coe_sub (x y : ℝ) : ↑(x - y) = (x - y : ℝ*) := rfl -@[simp, norm_cast] lemma coe_lt_coe {x y : ℝ} : (x : ℝ*) < y ↔ x < y := germ.const_lt -@[simp, norm_cast] lemma coe_pos {x : ℝ} : 0 < (x : ℝ*) ↔ 0 < x := coe_lt_coe @[simp, norm_cast] lemma coe_le_coe {x y : ℝ} : (x : ℝ*) ≤ y ↔ x ≤ y := germ.const_le_iff +@[simp, norm_cast] lemma coe_lt_coe {x y : ℝ} : (x : ℝ*) < y ↔ x < y := germ.const_lt_iff @[simp, norm_cast] lemma coe_nonneg {x : ℝ} : 0 ≤ (x : ℝ*) ↔ 0 ≤ x := coe_le_coe -@[simp, norm_cast] lemma coe_abs (x : ℝ) : ((|x| : ℝ) : ℝ*) = |x| := -begin - convert const_abs x, - apply linear_order.to_lattice_eq_filter_germ_lattice, -end +@[simp, norm_cast] lemma coe_pos {x : ℝ} : 0 < (x : ℝ*) ↔ 0 < x := coe_lt_coe +@[simp, norm_cast] lemma coe_abs (x : ℝ) : ((|x| : ℝ) : ℝ*) = |x| := const_abs x @[simp, norm_cast] lemma coe_max (x y : ℝ) : ((max x y : ℝ) : ℝ*) = max x y := germ.const_max _ _ @[simp, norm_cast] lemma coe_min (x y : ℝ) : ((min x y : ℝ) : ℝ*) = min x y := germ.const_min _ _ @@ -66,24 +62,18 @@ noncomputable def omega : ℝ* := of_seq coe localized "notation (name := hyperreal.epsilon) `ε` := hyperreal.epsilon" in hyperreal localized "notation (name := hyperreal.omega) `ω` := hyperreal.omega" in hyperreal -lemma epsilon_eq_inv_omega : ε = ω⁻¹ := rfl +@[simp] lemma inv_omega : ω⁻¹ = ε := rfl +@[simp] lemma inv_epsilon : ε⁻¹ = ω := @inv_inv _ _ ω -lemma inv_epsilon_eq_omega : ε⁻¹ = ω := @inv_inv _ _ ω - -lemma epsilon_pos : 0 < ε := -suffices ∀ᶠ i in hyperfilter ℕ, (0 : ℝ) < (i : ℕ)⁻¹, by rwa lt_def, -have h0' : {n : ℕ | ¬ 0 < n} = {0} := -by simp only [not_lt, (set.set_of_eq_eq_singleton).symm]; ext; exact le_bot_iff, -begin - simp only [inv_pos, nat.cast_pos], - exact mem_hyperfilter_of_finite_compl (by convert set.finite_singleton _), +lemma omega_pos : 0 < ω := germ.coe_pos.2 $ mem_hyperfilter_of_finite_compl $ begin + convert set.finite_singleton 0, + simp [set.eq_singleton_iff_unique_mem], end -lemma epsilon_ne_zero : ε ≠ 0 := ne_of_gt epsilon_pos - -lemma omega_pos : 0 < ω := by rw ←inv_epsilon_eq_omega; exact inv_pos.2 epsilon_pos +lemma epsilon_pos : 0 < ε := inv_pos_of_pos omega_pos -lemma omega_ne_zero : ω ≠ 0 := ne_of_gt omega_pos +lemma epsilon_ne_zero : ε ≠ 0 := epsilon_pos.ne' +lemma omega_ne_zero : ω ≠ 0 := omega_pos.ne' theorem epsilon_mul_omega : ε * ω = 1 := @inv_mul_cancel _ _ ω omega_ne_zero @@ -786,7 +776,7 @@ lemma infinite_mul_of_not_infinitesimal_infinite {x y : ℝ*} : ¬ infinitesimal x → infinite y → infinite (x * y) := λ hx hy, by rw [mul_comm]; exact infinite_mul_of_infinite_not_infinitesimal hy hx -lemma infinite_mul_infinite {x y : ℝ*} : infinite x → infinite y → infinite (x * y) := +lemma infinite.mul {x y : ℝ*} : infinite x → infinite y → infinite (x * y) := λ hx hy, infinite_mul_of_infinite_not_infinitesimal hx (not_infinitesimal_of_infinite hy) end hyperreal diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index a8e2aa4935eef..5d1e550c56236 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -50,7 +50,7 @@ open_locale classical big_operators /-- Nonnegative real numbers. -/ @[derive [ - ordered_semiring, comm_monoid_with_zero, -- to ensure these instance are computable + strict_ordered_semiring, comm_monoid_with_zero, -- to ensure these instances are computable floor_semiring, comm_semiring, semiring, semilattice_inf, semilattice_sup, distrib_lattice, densely_ordered, order_bot, diff --git a/src/data/set/intervals/instances.lean b/src/data/set/intervals/instances.lean index 5cb2ca4d67010..d15ea9da11266 100644 --- a/src/data/set/intervals/instances.lean +++ b/src/data/set/intervals/instances.lean @@ -32,10 +32,13 @@ The strongest typeclass provided on each interval is: -/ -variables {α : Type*} [ordered_semiring α] - open set +variables {α : Type*} + +section ordered_semiring +variables [ordered_semiring α] + /-! ### Instances for `↥(set.Icc 0 1)` -/ namespace set.Icc @@ -154,6 +157,10 @@ subtype.coe_injective.comm_semigroup _ coe_mul end set.Ico +end ordered_semiring + +variables [strict_ordered_semiring α] + /-! ### Instances for `↥(set.Ioc 0 1)` -/ namespace set.Ioc @@ -192,7 +199,15 @@ subtype.coe_injective.semigroup _ coe_mul instance monoid [nontrivial α] : monoid (Ioc (0:α) 1) := subtype.coe_injective.monoid _ coe_one coe_mul coe_pow -instance cancel_monoid {α : Type*} [ordered_comm_ring α] [is_domain α] : +instance comm_semigroup {α : Type*} [strict_ordered_comm_semiring α] : + comm_semigroup (Ioc (0:α) 1) := +subtype.coe_injective.comm_semigroup _ coe_mul + +instance comm_monoid {α : Type*} [strict_ordered_comm_semiring α] [nontrivial α] : + comm_monoid (Ioc (0:α) 1) := +subtype.coe_injective.comm_monoid _ coe_one coe_mul coe_pow + +instance cancel_monoid {α : Type*} [strict_ordered_ring α] [is_domain α] : cancel_monoid (Ioc (0:α) 1) := { mul_left_cancel := λ a b c h, subtype.ext $ mul_left_cancel₀ a.prop.1.ne' $ (congr_arg subtype.val h : _), @@ -200,14 +215,7 @@ instance cancel_monoid {α : Type*} [ordered_comm_ring α] [is_domain α] : subtype.ext $ mul_right_cancel₀ b.prop.1.ne' $ (congr_arg subtype.val h : _), ..set.Ioc.monoid} -instance comm_semigroup {α : Type*} [ordered_comm_semiring α] : comm_semigroup (Ioc (0:α) 1) := -subtype.coe_injective.comm_semigroup _ coe_mul - -instance comm_monoid {α : Type*} [ordered_comm_semiring α] [nontrivial α] : - comm_monoid (Ioc (0:α) 1) := -subtype.coe_injective.comm_monoid _ coe_one coe_mul coe_pow - -instance cancel_comm_monoid {α : Type*} [ordered_comm_ring α] [is_domain α] : +instance cancel_comm_monoid {α : Type*} [strict_ordered_comm_ring α] [is_domain α] : cancel_comm_monoid (Ioc (0:α) 1) := { ..set.Ioc.cancel_monoid, ..set.Ioc.comm_monoid } @@ -228,7 +236,8 @@ instance has_mul : has_mul (Ioo (0:α) 1) := { mul := λ p q, ⟨p.1 * q.1, ⟨m instance semigroup : semigroup (Ioo (0:α) 1) := subtype.coe_injective.semigroup _ coe_mul -instance comm_semigroup {α : Type*} [ordered_comm_semiring α] : comm_semigroup (Ioo (0:α) 1) := +instance comm_semigroup {α : Type*} [strict_ordered_comm_semiring α] : + comm_semigroup (Ioo (0:α) 1) := subtype.coe_injective.comm_semigroup _ coe_mul variables {β : Type*} [ordered_ring β] diff --git a/src/linear_algebra/orientation.lean b/src/linear_algebra/orientation.lean index 88990c4ec8947..52be381be8078 100644 --- a/src/linear_algebra/orientation.lean +++ b/src/linear_algebra/orientation.lean @@ -37,7 +37,7 @@ open_locale big_operators section ordered_comm_semiring -variables (R : Type*) [ordered_comm_semiring R] +variables (R : Type*) [strict_ordered_comm_semiring R] variables (M : Type*) [add_comm_monoid M] [module R M] variables {N : Type*} [add_comm_monoid N] [module R N] variables (ι : Type*) [decidable_eq ι] @@ -72,7 +72,7 @@ end ordered_comm_semiring section ordered_comm_ring -variables {R : Type*} [ordered_comm_ring R] +variables {R : Type*} [strict_ordered_comm_ring R] variables {M N : Type*} [add_comm_group M] [add_comm_group N] [module R M] [module R N] namespace basis diff --git a/src/linear_algebra/ray.lean b/src/linear_algebra/ray.lean index ea44a62866284..ac0ce2d200480 100644 --- a/src/linear_algebra/ray.lean +++ b/src/linear_algebra/ray.lean @@ -23,9 +23,9 @@ noncomputable theory open_locale big_operators -section ordered_comm_semiring +section strict_ordered_comm_semiring -variables (R : Type*) [ordered_comm_semiring R] +variables (R : Type*) [strict_ordered_comm_semiring R] variables {M : Type*} [add_comm_monoid M] [module R M] variables {N : Type*} [add_comm_monoid N] [module R N] variables (ι : Type*) [decidable_eq ι] @@ -297,11 +297,11 @@ x.some_ray_vector.property end module.ray -end ordered_comm_semiring +end strict_ordered_comm_semiring -section ordered_comm_ring +section strict_ordered_comm_ring -variables {R : Type*} [ordered_comm_ring R] +variables {R : Type*} [strict_ordered_comm_ring R] variables {M N : Type*} [add_comm_group M] [add_comm_group N] [module R M] [module R N] {x y : M} /-- `same_ray.neg` as an `iff`. -/ @@ -399,7 +399,7 @@ end end module.ray -end ordered_comm_ring +end strict_ordered_comm_ring section linear_ordered_comm_ring diff --git a/src/number_theory/cyclotomic/rat.lean b/src/number_theory/cyclotomic/rat.lean index e7dfbe00438df..f61362fd4b4e9 100644 --- a/src/number_theory/cyclotomic/rat.lean +++ b/src/number_theory/cyclotomic/rat.lean @@ -111,7 +111,7 @@ begin rw [is_primitive_root.sub_one_power_basis_gen] at h₁, rw [h₁, ← map_cyclotomic_int, show int.cast_ring_hom ℚ = algebra_map ℤ ℚ, by refl, show ((X + 1)) = map (algebra_map ℤ ℚ) (X + 1), by simp, ← map_comp] at h₂, - haveI : char_zero ℚ := ordered_semiring.to_char_zero, + haveI : char_zero ℚ := strict_ordered_semiring.to_char_zero, rw [is_primitive_root.sub_one_power_basis_gen, map_injective (algebra_map ℤ ℚ) ((algebra_map ℤ ℚ).injective_int) h₂], exact cyclotomic_prime_pow_comp_X_add_one_is_eisenstein_at _ _ }, @@ -136,7 +136,7 @@ local attribute [-instance] cyclotomic_field.algebra lemma cyclotomic_ring_is_integral_closure_of_prime_pow : is_integral_closure (cyclotomic_ring (p ^ k) ℤ ℚ) ℤ (cyclotomic_field (p ^ k) ℚ) := begin - haveI : char_zero ℚ := ordered_semiring.to_char_zero, + haveI : char_zero ℚ := strict_ordered_semiring.to_char_zero, haveI : is_cyclotomic_extension {p ^ k} ℚ (cyclotomic_field (p ^ k) ℚ), { convert cyclotomic_field.is_cyclotomic_extension (p ^ k) _, { exact subsingleton.elim _ _ }, diff --git a/src/number_theory/liouville/liouville_with.lean b/src/number_theory/liouville/liouville_with.lean index 9ebc0dcefe9ee..087dce9d8126e 100644 --- a/src/number_theory/liouville/liouville_with.lean +++ b/src/number_theory/liouville/liouville_with.lean @@ -286,9 +286,7 @@ begin rcases H with ⟨N, hN⟩, have : ∀ b > (1 : ℕ), ∀ᶠ m : ℕ in at_top, ∀ a : ℤ, (1 / b ^ m : ℝ) ≤ |x - a / b|, { intros b hb, - have hb0' : (b : ℚ) ≠ 0 := (zero_lt_one.trans (nat.one_lt_cast.2 hb)).ne', replace hb : (1 : ℝ) < b := nat.one_lt_cast.2 hb, - have hb0 : (0 : ℝ) < b := zero_lt_one.trans hb, have H : tendsto (λ m, 1 / b ^ m : ℕ → ℝ) at_top (𝓝 0), { simp only [one_div], exact tendsto_inv_at_top_zero.comp (tendsto_pow_at_top_at_top_of_one_lt hb) }, diff --git a/src/order/filter/archimedean.lean b/src/order/filter/archimedean.lean index 11cb64425b1f3..5fadc0546fa6d 100644 --- a/src/order/filter/archimedean.lean +++ b/src/order/filter/archimedean.lean @@ -19,40 +19,40 @@ variables {α R : Type*} open filter set -@[simp] lemma nat.comap_coe_at_top [ordered_semiring R] [nontrivial R] [archimedean R] : +@[simp] lemma nat.comap_coe_at_top [strict_ordered_semiring R] [nontrivial R] [archimedean R] : comap (coe : ℕ → R) at_top = at_top := comap_embedding_at_top (λ _ _, nat.cast_le) exists_nat_ge -lemma tendsto_coe_nat_at_top_iff [ordered_semiring R] [nontrivial R] [archimedean R] +lemma tendsto_coe_nat_at_top_iff [strict_ordered_semiring R] [nontrivial R] [archimedean R] {f : α → ℕ} {l : filter α} : tendsto (λ n, (f n : R)) l at_top ↔ tendsto f l at_top := tendsto_at_top_embedding (assume a₁ a₂, nat.cast_le) exists_nat_ge -lemma tendsto_coe_nat_at_top_at_top [ordered_semiring R] [archimedean R] : +lemma tendsto_coe_nat_at_top_at_top [strict_ordered_semiring R] [archimedean R] : tendsto (coe : ℕ → R) at_top at_top := nat.mono_cast.tendsto_at_top_at_top exists_nat_ge -@[simp] lemma int.comap_coe_at_top [ordered_ring R] [nontrivial R] [archimedean R] : +@[simp] lemma int.comap_coe_at_top [strict_ordered_ring R] [nontrivial R] [archimedean R] : comap (coe : ℤ → R) at_top = at_top := comap_embedding_at_top (λ _ _, int.cast_le) $ λ r, let ⟨n, hn⟩ := exists_nat_ge r in ⟨n, by exact_mod_cast hn⟩ -@[simp] lemma int.comap_coe_at_bot [ordered_ring R] [nontrivial R] [archimedean R] : +@[simp] lemma int.comap_coe_at_bot [strict_ordered_ring R] [nontrivial R] [archimedean R] : comap (coe : ℤ → R) at_bot = at_bot := comap_embedding_at_bot (λ _ _, int.cast_le) $ λ r, let ⟨n, hn⟩ := exists_nat_ge (-r) in ⟨-n, by simpa [neg_le] using hn⟩ -lemma tendsto_coe_int_at_top_iff [ordered_ring R] [nontrivial R] [archimedean R] +lemma tendsto_coe_int_at_top_iff [strict_ordered_ring R] [nontrivial R] [archimedean R] {f : α → ℤ} {l : filter α} : tendsto (λ n, (f n : R)) l at_top ↔ tendsto f l at_top := by rw [← tendsto_comap_iff, int.comap_coe_at_top] -lemma tendsto_coe_int_at_bot_iff [ordered_ring R] [nontrivial R] [archimedean R] +lemma tendsto_coe_int_at_bot_iff [strict_ordered_ring R] [nontrivial R] [archimedean R] {f : α → ℤ} {l : filter α} : tendsto (λ n, (f n : R)) l at_bot ↔ tendsto f l at_bot := by rw [← tendsto_comap_iff, int.comap_coe_at_bot] -lemma tendsto_coe_int_at_top_at_top [ordered_ring R] [archimedean R] : +lemma tendsto_coe_int_at_top_at_top [strict_ordered_ring R] [archimedean R] : tendsto (coe : ℤ → R) at_top at_top := int.cast_mono.tendsto_at_top_at_top $ λ b, let ⟨n, hn⟩ := exists_nat_ge b in ⟨n, by exact_mod_cast hn⟩ diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index e76e002fb8699..d473d68e67169 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -696,9 +696,9 @@ variable {l} end ordered_group -section ordered_semiring +section strict_ordered_semiring -variables [ordered_semiring α] {l : filter β} {f g : β → α} +variables [strict_ordered_semiring α] {l : filter β} {f g : β → α} lemma tendsto_bit1_at_top : tendsto bit1 (at_top : filter α) at_top := tendsto_at_top_add_nonneg_right tendsto_bit0_at_top (λ _, zero_le_one) @@ -719,15 +719,14 @@ A version for positive real powers exists as `tendsto_rpow_at_top`. -/ lemma tendsto_pow_at_top {n : ℕ} (hn : n ≠ 0) : tendsto (λ x : α, x ^ n) at_top at_top := tendsto_at_top_mono' _ ((eventually_ge_at_top 1).mono $ λ x hx, le_self_pow hx hn) tendsto_id -end ordered_semiring +end strict_ordered_semiring lemma zero_pow_eventually_eq [monoid_with_zero α] : (λ n : ℕ, (0 : α) ^ n) =ᶠ[at_top] (λ n, 0) := eventually_at_top.2 ⟨1, λ n hn, zero_pow (zero_lt_one.trans_le hn)⟩ -section ordered_ring - -variables [ordered_ring α] {l : filter β} {f g : β → α} +section strict_ordered_ring +variables [strict_ordered_ring α] {l : filter β} {f g : β → α} lemma tendsto.at_top_mul_at_bot (hf : tendsto f l at_top) (hg : tendsto g l at_bot) : tendsto (λ x, f x * g x) l at_bot := @@ -746,7 +745,7 @@ have tendsto (λ x, (-f x) * (-g x)) l at_top := (tendsto_neg_at_bot_at_top.comp hf).at_top_mul_at_top (tendsto_neg_at_bot_at_top.comp hg), by simpa only [neg_mul_neg] using this -end ordered_ring +end strict_ordered_ring section linear_ordered_add_comm_group diff --git a/src/order/filter/filter_product.lean b/src/order/filter/filter_product.lean index 523775e202746..0855cf8eb3dd3 100644 --- a/src/order/filter/filter_product.lean +++ b/src/order/filter/filter_product.lean @@ -32,63 +32,135 @@ open ultrafilter local notation `β*` := germ (φ : filter α) β -/-- If `φ` is an ultrafilter then the ultraproduct is a division ring. -/ -instance [division_ring β] : division_ring β* := +instance [division_semiring β] : division_semiring β* := { mul_inv_cancel := λ f, induction_on f $ λ f hf, coe_eq.2 $ (φ.em (λ y, f y = 0)).elim (λ H, (hf $ coe_eq.2 H).elim) (λ H, H.mono $ λ x, mul_inv_cancel), inv_zero := coe_eq.2 $ by simp only [(∘), inv_zero], - .. germ.ring, .. germ.div_inv_monoid, .. germ.nontrivial } + ..germ.semiring, ..germ.div_inv_monoid, ..germ.nontrivial } -/-- If `φ` is an ultrafilter then the ultraproduct is a field. -/ -instance [field β] : field β* := -{ .. germ.comm_ring, .. germ.division_ring } - -/-- If `φ` is an ultrafilter then the ultraproduct is a linear order. -/ -noncomputable instance [linear_order β] : linear_order β* := -{ le_total := λ f g, induction_on₂ f g $ λ f g, eventually_or.1 $ eventually_of_forall $ - λ x, le_total _ _, - decidable_le := by apply_instance, - .. germ.partial_order } - -@[simp, norm_cast] lemma const_div [division_ring β] (x y : β) : (↑(x / y) : β*) = ↑x / ↑y := rfl +instance [division_ring β] : division_ring β* := { ..germ.ring, ..germ.division_semiring } +instance [semifield β] : semifield β* := { ..germ.comm_semiring, ..germ.division_semiring } +instance [field β] : field β* := { ..germ.comm_ring, ..germ.division_ring } lemma coe_lt [preorder β] {f g : α → β} : (f : β*) < g ↔ ∀* x, f x < g x := by simp only [lt_iff_le_not_le, eventually_and, coe_le, eventually_not, eventually_le] -lemma coe_pos [preorder β] [has_zero β] {f : α → β} : 0 < (f : β*) ↔ ∀* x, 0 < f x := -coe_lt +lemma coe_pos [preorder β] [has_zero β] {f : α → β} : 0 < (f : β*) ↔ ∀* x, 0 < f x := coe_lt -lemma const_lt [preorder β] {x y : β} : (↑x : β*) < ↑y ↔ x < y := +lemma const_lt [preorder β] {x y : β} : x ≤ y → (↑x : β*) ≤ ↑y := lift_rel_const + +lemma const_lt_iff [preorder β] {x y : β} : (↑x : β*) < ↑y ↔ x < y := coe_lt.trans lift_rel_const_iff lemma lt_def [preorder β] : ((<) : β* → β* → Prop) = lift_rel (<) := by { ext ⟨f⟩ ⟨g⟩, exact coe_lt } -/-- If `φ` is an ultrafilter then the ultraproduct is an ordered ring. -/ +instance [has_sup β] : has_sup β* := ⟨map₂ (⊔)⟩ +instance [has_inf β] : has_inf β* := ⟨map₂ (⊓)⟩ + +@[simp, norm_cast] lemma const_sup [has_sup β] (a b : β) : ↑(a ⊔ b) = (↑a ⊔ ↑b : β*) := rfl +@[simp, norm_cast] lemma const_inf [has_inf β] (a b : β) : ↑(a ⊓ b) = (↑a ⊓ ↑b : β*) := rfl + +instance [semilattice_sup β] : semilattice_sup β* := +{ sup := (⊔), + le_sup_left := λ f g, induction_on₂ f g $ λ f g, + eventually_of_forall $ λ x, le_sup_left, + le_sup_right := λ f g, induction_on₂ f g $ λ f g, + eventually_of_forall $ λ x, le_sup_right, + sup_le := λ f₁ f₂ g, induction_on₃ f₁ f₂ g $ λ f₁ f₂ g h₁ h₂, + h₂.mp $ h₁.mono $ λ x, sup_le, + .. germ.partial_order } + +instance [semilattice_inf β] : semilattice_inf β* := +{ inf := (⊓), + inf_le_left := λ f g, induction_on₂ f g $ λ f g, + eventually_of_forall $ λ x, inf_le_left, + inf_le_right := λ f g, induction_on₂ f g $ λ f g, + eventually_of_forall $ λ x, inf_le_right, + le_inf := λ f₁ f₂ g, induction_on₃ f₁ f₂ g $ λ f₁ f₂ g h₁ h₂, + h₂.mp $ h₁.mono $ λ x, le_inf, + .. germ.partial_order } + +instance [lattice β] : lattice β* := +{ .. germ.semilattice_sup, .. germ.semilattice_inf } + +instance [distrib_lattice β] : distrib_lattice β* := +{ le_sup_inf := λ f g h, induction_on₃ f g h $ λ f g h, eventually_of_forall $ λ _, le_sup_inf, + .. germ.semilattice_sup, .. germ.semilattice_inf } + +instance [has_le β] [is_total β (≤)] : is_total β* (≤) := +⟨λ f g, induction_on₂ f g $ λ f g, eventually_or.1 $ eventually_of_forall $ λ x, total_of _ _ _⟩ + +/-- If `φ` is an ultrafilter then the ultraproduct is a linear order. -/ +noncomputable instance [linear_order β] : linear_order β* := lattice.to_linear_order _ + +@[to_additive] +instance [ordered_comm_monoid β] : ordered_comm_monoid β* := +{ mul_le_mul_left := λ f g, induction_on₂ f g $ λ f g H h, induction_on h $ λ h, + H.mono $ λ x H, mul_le_mul_left' H _, + .. germ.partial_order, .. germ.comm_monoid } + +@[to_additive] +instance [ordered_cancel_comm_monoid β] : ordered_cancel_comm_monoid β* := +{ le_of_mul_le_mul_left := λ f g h, induction_on₃ f g h $ λ f g h H, + H.mono $ λ x, le_of_mul_le_mul_left', + .. germ.partial_order, .. germ.ordered_comm_monoid } + +@[to_additive] +instance [ordered_comm_group β] : ordered_comm_group β* := +{ .. germ.ordered_cancel_comm_monoid, .. germ.comm_group } + +@[to_additive] +noncomputable instance [linear_ordered_comm_group β] : linear_ordered_comm_group β* := +{ .. germ.ordered_comm_group, .. germ.linear_order } + +instance [ordered_semiring β] : ordered_semiring β* := +{ zero_le_one := const_le zero_le_one, + mul_le_mul_of_nonneg_left := λ x y z, induction_on₃ x y z $ λ f g h hfg hh, hh.mp $ + hfg.mono $ λ a, mul_le_mul_of_nonneg_left, + mul_le_mul_of_nonneg_right := λ x y z, induction_on₃ x y z $ λ f g h hfg hh, hh.mp $ + hfg.mono $ λ a, mul_le_mul_of_nonneg_right, + ..germ.semiring, ..germ.ordered_add_comm_monoid } + +instance [ordered_comm_semiring β] : ordered_comm_semiring β* := +{ ..germ.ordered_semiring, ..germ.comm_semiring } + instance [ordered_ring β] : ordered_ring β* := +{ zero_le_one := const_le zero_le_one, + mul_nonneg := λ x y, induction_on₂ x y $ λ f g hf hg, hg.mp $ hf.mono $ λ a, mul_nonneg, + ..germ.ring, ..germ.ordered_add_comm_group } + +instance [ordered_comm_ring β] : ordered_comm_ring β* := +{ ..germ.ordered_ring, ..germ.ordered_comm_semiring } + +instance [strict_ordered_semiring β] : strict_ordered_semiring β* := +{ mul_lt_mul_of_pos_left := λ x y z, induction_on₃ x y z $ λ f g h hfg hh, coe_lt.2 $ + (coe_lt.1 hh).mp $ (coe_lt.1 hfg).mono $ λ a, mul_lt_mul_of_pos_left, + mul_lt_mul_of_pos_right := λ x y z, induction_on₃ x y z $ λ f g h hfg hh, coe_lt.2 $ + (coe_lt.1 hh).mp $ (coe_lt.1 hfg).mono $ λ a, mul_lt_mul_of_pos_right, + .. germ.ordered_semiring, ..germ.ordered_cancel_add_comm_monoid } + +instance [strict_ordered_comm_semiring β] : strict_ordered_comm_semiring β* := +{ .. germ.strict_ordered_semiring, ..germ.ordered_comm_semiring } + +instance [strict_ordered_ring β] : strict_ordered_ring β* := { zero_le_one := const_le zero_le_one, mul_pos := λ x y, induction_on₂ x y $ λ f g hf hg, coe_pos.2 $ (coe_pos.1 hg).mp $ (coe_pos.1 hf).mono $ λ x, mul_pos, .. germ.ring, .. germ.ordered_add_comm_group, .. germ.nontrivial } -/-- If `φ` is an ultrafilter then the ultraproduct is a linear ordered ring. -/ +instance [strict_ordered_comm_ring β] : strict_ordered_comm_ring β* := +{ .. germ.strict_ordered_ring, ..germ.ordered_comm_ring } + noncomputable instance [linear_ordered_ring β] : linear_ordered_ring β* := -{ .. germ.ordered_ring, .. germ.linear_order, .. germ.nontrivial } +{ ..germ.strict_ordered_ring, ..germ.linear_order, ..germ.nontrivial } -/-- If `φ` is an ultrafilter then the ultraproduct is a linear ordered field. -/ noncomputable instance [linear_ordered_field β] : linear_ordered_field β* := { .. germ.linear_ordered_ring, .. germ.field } -/-- If `φ` is an ultrafilter then the ultraproduct is a linear ordered commutative ring. -/ -noncomputable instance [linear_ordered_comm_ring β] : - linear_ordered_comm_ring β* := +noncomputable instance [linear_ordered_comm_ring β] : linear_ordered_comm_ring β* := { .. germ.linear_ordered_ring, .. germ.comm_monoid } -/-- If `φ` is an ultrafilter then the ultraproduct is a decidable linear ordered commutative -group. -/ -noncomputable instance [linear_ordered_add_comm_group β] : linear_ordered_add_comm_group β* := -{ .. germ.ordered_add_comm_group, .. germ.linear_order } - lemma max_def [linear_order β] (x y : β*) : max x y = map₂ max x y := induction_on₂ x y $ λ a b, begin @@ -118,10 +190,6 @@ by rw [min_def, map₂_const] (↑(|x|) : β*) = |↑x| := by rw [abs_def, map_const] -lemma linear_order.to_lattice_eq_filter_germ_lattice [linear_order β] : - (@linear_order.to_lattice (filter.germ ↑φ β) filter.germ.linear_order) = filter.germ.lattice := -lattice.ext (λ x y, iff.rfl) - end germ end filter diff --git a/src/order/filter/germ.lean b/src/order/filter/germ.lean index c2baca43deeaa..2a6c20199d729 100644 --- a/src/order/filter/germ.lean +++ b/src/order/filter/germ.lean @@ -294,36 +294,30 @@ instance [right_cancel_semigroup M] : right_cancel_semigroup (germ l M) := coe_eq.2 $ (coe_eq.1 H).mono $ λ x, mul_right_cancel, .. germ.semigroup } -instance has_nat_pow [monoid G] : has_pow (germ l G) ℕ := ⟨λ f n, map (^ n) f⟩ +instance [has_vadd M G] : has_vadd M (germ l G) := ⟨λ n, map ((+ᵥ) n)⟩ +@[to_additive] instance [has_smul M G] : has_smul M (germ l G) := ⟨λ n, map ((•) n)⟩ +@[to_additive has_smul] instance [has_pow G M] : has_pow (germ l G) M := ⟨λ f n, map (^ n) f⟩ -@[simp] lemma coe_pow [monoid G] (f : α → G) (n : ℕ) : ↑(f ^ n) = (f ^ n : germ l G) := rfl - -instance has_int_pow [div_inv_monoid G] : has_pow (germ l G) ℤ := ⟨λ f z, map (^ z) f⟩ - -@[simp] lemma coe_zpow [div_inv_monoid G] (f : α → G) (z : ℤ) : ↑(f ^ z) = (f ^ z : germ l G) := -rfl +@[simp, norm_cast, to_additive] +lemma coe_smul [has_smul M G] (n : M) (f : α → G) : ↑(n • f) = (n • f : germ l G) := rfl -instance [has_smul M β] : has_smul M (germ l β) := -⟨λ c, map ((•) c)⟩ +@[simp, norm_cast, to_additive] +lemma const_smul [has_smul M G] (n : M) (a : G) : (↑(n • a) : germ l G) = n • ↑a := rfl -@[simp, norm_cast] lemma coe_smul [has_smul M β] (c : M) (f : α → β) : - ↑(c • f) = (c • f : germ l β) := -rfl +@[simp, norm_cast, to_additive coe_smul] +lemma coe_pow [has_pow G M] (f : α → G) (n : M) : ↑(f ^ n) = (f ^ n : germ l G) := rfl -instance [add_monoid M] : add_monoid (germ l M) := -function.surjective.add_monoid coe (surjective_quot_mk _) rfl (λ a b, coe_add a b) (λ _ _, rfl) +@[simp, norm_cast, to_additive const_smul] +lemma const_pow [has_pow G M] (a : G) (n : M) : (↑(a ^ n) : germ l G) = ↑a ^ n := rfl @[to_additive] instance [monoid M] : monoid (germ l M) := -function.surjective.monoid coe (surjective_quot_mk _) rfl (λ a b, coe_mul a b) coe_pow +function.surjective.monoid coe (surjective_quot_mk _) rfl (λ _ _, rfl) (λ _ _, rfl) -/-- coercion from functions to germs as a monoid homomorphism. -/ -@[to_additive] +/-- Coercion from functions to germs as a monoid homomorphism. -/ +@[to_additive "Coercion from functions to germs as an additive monoid homomorphism."] def coe_mul_hom [monoid M] (l : filter α) : (α → M) →* germ l M := ⟨coe, rfl, λ f g, rfl⟩ -/-- coercion from functions to germs as an additive monoid homomorphism. -/ -add_decl_doc coe_add_hom - @[simp, to_additive] lemma coe_coe_mul_hom [monoid M] : (coe_mul_hom l : (α → M) → germ l M) = coe := rfl @@ -345,15 +339,17 @@ instance [has_inv G] : has_inv (germ l G) := ⟨map has_inv.inv⟩ @[simp, norm_cast, to_additive] lemma coe_inv [has_inv G] (f : α → G) : ↑f⁻¹ = (f⁻¹ : germ l G) := rfl +@[simp, norm_cast, to_additive] +lemma const_inv [has_inv G] (a : G) : (↑a⁻¹ : germ l G) = (↑a)⁻¹ := rfl + @[to_additive] instance [has_div M] : has_div (germ l M) := ⟨map₂ (/)⟩ @[simp, norm_cast, to_additive] lemma coe_div [has_div M] (f g : α → M) : ↑(f / g) = (f / g : germ l M) := rfl -instance [sub_neg_monoid G] : sub_neg_monoid (germ l G) := -function.surjective.sub_neg_monoid coe (surjective_quot_mk _) rfl (λ _ _, rfl) - (λ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) +@[simp, norm_cast, to_additive] +lemma const_div [has_div M] (a b : M) : (↑(a / b) : germ l M) = ↑a / ↑b := rfl @[to_additive sub_neg_monoid] instance [div_inv_monoid G] : div_inv_monoid (germ l G) := @@ -420,18 +416,20 @@ section module variables {M N R : Type*} -instance has_smul' [has_smul M β] : has_smul (germ l M) (germ l β) := -⟨map₂ (•)⟩ +@[to_additive] +instance has_smul' [has_smul M β] : has_smul (germ l M) (germ l β) := ⟨map₂ (•)⟩ -@[simp, norm_cast] lemma coe_smul' [has_smul M β] (c : α → M) (f : α → β) : +@[simp, norm_cast, to_additive] lemma coe_smul' [has_smul M β] (c : α → M) (f : α → β) : ↑(c • f) = (c : germ l M) • (f : germ l β) := rfl -instance [monoid M] [mul_action M β] : mul_action M (germ l β) := +@[to_additive] +instance [monoid M] [mul_action M β] : mul_action M (germ l β) := { one_smul := λ f, induction_on f $ λ f, by { norm_cast, simp only [one_smul] }, mul_smul := λ c₁ c₂ f, induction_on f $ λ f, by { norm_cast, simp only [mul_smul] } } -instance mul_action' [monoid M] [mul_action M β] : mul_action (germ l M) (germ l β) := +@[to_additive] +instance mul_action' [monoid M] [mul_action M β] : mul_action (germ l M) (germ l β) := { one_smul := λ f, induction_on f $ λ f, by simp only [← coe_one, ← coe_smul', one_smul], mul_smul := λ c₁ c₂ f, induction_on₃ c₁ c₂ f $ λ c₁ c₂ f, by { norm_cast, simp only [mul_smul] } } @@ -457,15 +455,16 @@ instance module' [semiring R] [add_comm_monoid M] [module R M] : end module -instance [has_le β] : has_le (germ l β) := -⟨lift_rel (≤)⟩ - -@[simp] lemma coe_le [has_le β] : (f : germ l β) ≤ g ↔ (f ≤ᶠ[l] g) := iff.rfl +instance [has_le β] : has_le (germ l β) := ⟨lift_rel (≤)⟩ lemma le_def [has_le β] : ((≤) : germ l β → germ l β → Prop) = lift_rel (≤) := rfl -lemma const_le [has_le β] {x y : β} (h : x ≤ y) : (↑x : germ l β) ≤ ↑y := -lift_rel_const h +@[simp] lemma coe_le [has_le β] : (f : germ l β) ≤ g ↔ f ≤ᶠ[l] g := iff.rfl + +lemma coe_nonneg [has_le β] [has_zero β] {f : α → β} : 0 ≤ (f : germ l β) ↔ ∀ᶠ x in l, 0 ≤ f x := +iff.rfl + +lemma const_le [has_le β] {x y : β} : x ≤ y → (↑x : germ l β) ≤ ↑y := lift_rel_const @[simp, norm_cast] lemma const_le_iff [has_le β] [ne_bot l] {x y : β} : (↑x : germ l β) ≤ ↑y ↔ x ≤ y := @@ -481,69 +480,22 @@ instance [partial_order β] : partial_order (germ l β) := le_antisymm := λ f g, induction_on₂ f g $ λ f g h₁ h₂, (eventually_le.antisymm h₁ h₂).germ_eq, .. germ.preorder } -instance [has_bot β] : has_bot (germ l β) := ⟨↑(⊥:β)⟩ +instance [has_bot β] : has_bot (germ l β) := ⟨↑(⊥ : β)⟩ +instance [has_top β] : has_top (germ l β) := ⟨↑(⊤ : β)⟩ -@[simp, norm_cast] lemma const_bot [has_bot β] : (↑(⊥:β) : germ l β) = ⊥ := rfl +@[simp, norm_cast] lemma const_bot [has_bot β] : (↑(⊥ : β) : germ l β) = ⊥ := rfl +@[simp, norm_cast] lemma const_top [has_top β] : (↑(⊤ : β) : germ l β) = ⊤ := rfl instance [has_le β] [order_bot β] : order_bot (germ l β) := { bot := ⊥, bot_le := λ f, induction_on f $ λ f, eventually_of_forall $ λ x, bot_le } -instance [has_top β] : has_top (germ l β) := ⟨↑(⊤:β)⟩ - -@[simp, norm_cast] lemma const_top [has_top β] : (↑(⊤:β) : germ l β) = ⊤ := rfl - instance [has_le β] [order_top β] : order_top (germ l β) := { top := ⊤, le_top := λ f, induction_on f $ λ f, eventually_of_forall $ λ x, le_top } -instance [has_sup β] : has_sup (germ l β) := ⟨map₂ (⊔)⟩ - -@[simp, norm_cast] lemma const_sup [has_sup β] (a b : β) : ↑(a ⊔ b) = (↑a ⊔ ↑b : germ l β) := rfl - -instance [has_inf β] : has_inf (germ l β) := ⟨map₂ (⊓)⟩ - -@[simp, norm_cast] lemma const_inf [has_inf β] (a b : β) : ↑(a ⊓ b) = (↑a ⊓ ↑b : germ l β) := rfl - -instance [semilattice_sup β] : semilattice_sup (germ l β) := -{ sup := (⊔), - le_sup_left := λ f g, induction_on₂ f g $ λ f g, - eventually_of_forall $ λ x, le_sup_left, - le_sup_right := λ f g, induction_on₂ f g $ λ f g, - eventually_of_forall $ λ x, le_sup_right, - sup_le := λ f₁ f₂ g, induction_on₃ f₁ f₂ g $ λ f₁ f₂ g h₁ h₂, - h₂.mp $ h₁.mono $ λ x, sup_le, - .. germ.partial_order } - -instance [semilattice_inf β] : semilattice_inf (germ l β) := -{ inf := (⊓), - inf_le_left := λ f g, induction_on₂ f g $ λ f g, - eventually_of_forall $ λ x, inf_le_left, - inf_le_right := λ f g, induction_on₂ f g $ λ f g, - eventually_of_forall $ λ x, inf_le_right, - le_inf := λ f₁ f₂ g, induction_on₃ f₁ f₂ g $ λ f₁ f₂ g h₁ h₂, - h₂.mp $ h₁.mono $ λ x, le_inf, - .. germ.partial_order } - -instance [lattice β] : lattice (germ l β) := -{ .. germ.semilattice_sup, .. germ.semilattice_inf } - instance [has_le β] [bounded_order β] : bounded_order (germ l β) := -{ .. germ.order_bot, .. germ.order_top } - -@[to_additive] -instance [ordered_cancel_comm_monoid β] : ordered_cancel_comm_monoid (germ l β) := -{ mul_le_mul_left := λ f g, induction_on₂ f g $ λ f g H h, induction_on h $ λ h, - H.mono $ λ x H, mul_le_mul_left' H _, - le_of_mul_le_mul_left := λ f g h, induction_on₃ f g h $ λ f g h H, - H.mono $ λ x, le_of_mul_le_mul_left', - .. germ.partial_order, .. germ.comm_monoid } - -@[to_additive] -instance ordered_comm_group [ordered_comm_group β] : ordered_comm_group (germ l β) := -{ mul_le_mul_left := λ f g, induction_on₂ f g $ λ f g H h, induction_on h $ λ h, - H.mono $ λ x H, mul_le_mul_left' H _, - .. germ.partial_order, .. germ.comm_group } +{ ..germ.order_bot, ..germ.order_top } end germ diff --git a/src/ring_theory/class_group.lean b/src/ring_theory/class_group.lean index fb8c62f2999b4..47490ebbb5fbb 100644 --- a/src/ring_theory/class_group.lean +++ b/src/ring_theory/class_group.lean @@ -89,7 +89,6 @@ noncomputable def fractional_ideal.mk0 [is_dedekind_domain R] : map_mul' := λ x y, by simp } /-- Send a nonzero ideal to the corresponding class in the class group. -/ -@[simps] noncomputable def class_group.mk0 [is_dedekind_domain R] : (ideal R)⁰ →* class_group R K := (quotient_group.mk' _).comp (fractional_ideal.mk0 K) diff --git a/src/ring_theory/hahn_series.lean b/src/ring_theory/hahn_series.lean index b12508fac023a..220fff76edbd5 100644 --- a/src/ring_theory/hahn_series.lean +++ b/src/ring_theory/hahn_series.lean @@ -1012,8 +1012,9 @@ power_series.coeff_mk _ _ lemma coeff_to_power_series_symm {f : power_series R} {n : ℕ} : (hahn_series.to_power_series.symm f).coeff n = power_series.coeff R n f := rfl -variables (Γ) (R) [ordered_semiring Γ] [nontrivial Γ] -/-- Casts a power series as a Hahn series with coefficients from an `ordered_semiring`. -/ +variables (Γ) (R) [strict_ordered_semiring Γ] [nontrivial Γ] + +/-- Casts a power series as a Hahn series with coefficients from an `strict_ordered_semiring`. -/ def of_power_series : (power_series R) →+* hahn_series Γ R := (hahn_series.emb_domain_ring_hom (nat.cast_add_monoid_hom Γ) nat.strict_mono_cast.injective (λ _ _, nat.cast_le)).comp @@ -1130,8 +1131,8 @@ variables (R) [comm_semiring R] {A : Type*} [semiring A] [algebra R A] end, .. to_power_series } -variables (Γ) (R) [ordered_semiring Γ] [nontrivial Γ] -/-- Casting a power series as a Hahn series with coefficients from an `ordered_semiring` +variables (Γ) (R) [strict_ordered_semiring Γ] [nontrivial Γ] +/-- Casting a power series as a Hahn series with coefficients from an `strict_ordered_semiring` is an algebra homomorphism. -/ @[simps] def of_power_series_alg : (power_series A) →ₐ[R] hahn_series Γ A := (hahn_series.emb_domain_alg_hom (nat.cast_add_monoid_hom Γ) nat.strict_mono_cast.injective diff --git a/src/ring_theory/polynomial/cyclotomic/eval.lean b/src/ring_theory/polynomial/cyclotomic/eval.lean index 8cd844817de6a..1cd3e317d7103 100644 --- a/src/ring_theory/polynomial/cyclotomic/eval.lean +++ b/src/ring_theory/polynomial/cyclotomic/eval.lean @@ -83,7 +83,7 @@ begin rw eval_prod, refine finset.prod_nonneg (λ i hi, _), simp only [finset.mem_sdiff, mem_proper_divisors, finset.mem_singleton] at hi, - rw geom_sum_pos_iff hn'' at h, + rw geom_sum_pos_iff hn'.ne' at h, cases h with hk hx, { refine (ih _ hi.1.2 (nat.two_lt_of_ne _ hi.2 _)).le; rintro rfl, { exact hn'.ne' (zero_dvd_iff.mp hi.1.1) }, @@ -93,11 +93,11 @@ begin refine (ih _ hi.1.2 (nat.two_lt_of_ne _ hi.2 hk)).le, rintro rfl, exact (hn'.ne' $ zero_dvd_iff.mp hi.1.1) } }, - { rw [eq_comm, geom_sum_eq_zero_iff_neg_one hn''] at h, + { rw [eq_comm, geom_sum_eq_zero_iff_neg_one hn'.ne'] at h, exact h.1.symm ▸ cyclotomic_neg_one_pos hn }, { apply pos_of_mul_neg_left, { rwa this }, - rw [geom_sum_neg_iff hn''] at h, + rw geom_sum_neg_iff hn'.ne' at h, have h2 : {2} ⊆ n.proper_divisors \ {1}, { rw [finset.singleton_subset_iff, finset.mem_sdiff, mem_proper_divisors, finset.not_mem_singleton], diff --git a/src/ring_theory/polynomial/pochhammer.lean b/src/ring_theory/polynomial/pochhammer.lean index 7f36bd5948bbc..1a331fa73ca11 100644 --- a/src/ring_theory/polynomial/pochhammer.lean +++ b/src/ring_theory/polynomial/pochhammer.lean @@ -145,8 +145,8 @@ end end semiring -section ordered_semiring -variables {S : Type*} [ordered_semiring S] [nontrivial S] +section strict_ordered_semiring +variables {S : Type*} [strict_ordered_semiring S] [nontrivial S] lemma pochhammer_pos (n : ℕ) (s : S) (h : 0 < s) : 0 < (pochhammer S n).eval s := begin @@ -158,7 +158,7 @@ begin (lt_of_lt_of_le h ((le_add_iff_nonneg_right _).mpr (nat.cast_nonneg n))), } end -end ordered_semiring +end strict_ordered_semiring section factorial diff --git a/src/ring_theory/subsemiring/basic.lean b/src/ring_theory/subsemiring/basic.lean index 41640480f0cf6..7e3306c721e55 100644 --- a/src/ring_theory/subsemiring/basic.lean +++ b/src/ring_theory/subsemiring/basic.lean @@ -121,12 +121,26 @@ instance to_ordered_semiring {R} [ordered_semiring R] [set_like S R] [subsemirin subtype.coe_injective.ordered_semiring coe rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) +/-- A subsemiring of an `strict_ordered_semiring` is an `strict_ordered_semiring`. -/ +instance to_strict_ordered_semiring {R} [strict_ordered_semiring R] [set_like S R] + [subsemiring_class S R] : + strict_ordered_semiring s := +subtype.coe_injective.strict_ordered_semiring coe + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + /-- A subsemiring of an `ordered_comm_semiring` is an `ordered_comm_semiring`. -/ instance to_ordered_comm_semiring {R} [ordered_comm_semiring R] [set_like S R] [subsemiring_class S R] : ordered_comm_semiring s := subtype.coe_injective.ordered_comm_semiring coe rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) +/-- A subsemiring of an `strict_ordered_comm_semiring` is an `strict_ordered_comm_semiring`. -/ +instance to_strict_ordered_comm_semiring {R} [strict_ordered_comm_semiring R] [set_like S R] + [subsemiring_class S R] : + strict_ordered_comm_semiring s := +subtype.coe_injective.strict_ordered_comm_semiring coe + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + /-- A subsemiring of a `linear_ordered_semiring` is a `linear_ordered_semiring`. -/ instance to_linear_ordered_semiring {R} [linear_ordered_semiring R] [set_like S R] [subsemiring_class S R] : linear_ordered_semiring s := @@ -343,12 +357,24 @@ instance to_ordered_semiring {R} [ordered_semiring R] (s : subsemiring R) : orde subtype.coe_injective.ordered_semiring coe rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) +/-- A subsemiring of a `strict_ordered_semiring` is a `strict_ordered_semiring`. -/ +instance to_strict_ordered_semiring {R} [strict_ordered_semiring R] (s : subsemiring R) : + strict_ordered_semiring s := +subtype.coe_injective.strict_ordered_semiring coe + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + /-- A subsemiring of an `ordered_comm_semiring` is an `ordered_comm_semiring`. -/ instance to_ordered_comm_semiring {R} [ordered_comm_semiring R] (s : subsemiring R) : ordered_comm_semiring s := subtype.coe_injective.ordered_comm_semiring coe rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) +/-- A subsemiring of a `strict_ordered_comm_semiring` is a `strict_ordered_comm_semiring`. -/ +instance to_strict_ordered_comm_semiring {R} [strict_ordered_comm_semiring R] (s : subsemiring R) : + strict_ordered_comm_semiring s := +subtype.coe_injective.strict_ordered_comm_semiring coe + rfl rfl (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _ _, rfl) (λ _, rfl) + /-- A subsemiring of a `linear_ordered_semiring` is a `linear_ordered_semiring`. -/ instance to_linear_ordered_semiring {R} [linear_ordered_semiring R] (s : subsemiring R) : linear_ordered_semiring s := @@ -1077,14 +1103,14 @@ end subsemiring end actions -- While this definition is not about `subsemiring`s, this is the earliest we have --- both `ordered_semiring` and `submonoid` available. +-- both `strict_ordered_semiring` and `submonoid` available. /-- Submonoid of positive elements of an ordered semiring. -/ -def pos_submonoid (R : Type*) [ordered_semiring R] [nontrivial R] : submonoid R := +def pos_submonoid (R : Type*) [strict_ordered_semiring R] [nontrivial R] : submonoid R := { carrier := {x | 0 < x}, one_mem' := show (0 : R) < 1, from zero_lt_one, mul_mem' := λ x y (hx : 0 < x) (hy : 0 < y), mul_pos hx hy } -@[simp] lemma mem_pos_monoid {R : Type*} [ordered_semiring R] [nontrivial R] (u : Rˣ) : +@[simp] lemma mem_pos_monoid {R : Type*} [strict_ordered_semiring R] [nontrivial R] (u : Rˣ) : ↑u ∈ pos_submonoid R ↔ (0 : R) < u := iff.rfl diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index 07e25164cf376..6c70b45933468 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -198,17 +198,16 @@ end @[simp] lemma nim_add_equiv_zero_iff (o₁ o₂ : ordinal) : nim o₁ + nim o₂ ≈ 0 ↔ o₁ = o₂ := begin split, - { contrapose, - intro h, - rw [impartial.not_equiv_zero_iff], - wlog h' : o₁ ≤ o₂ using [o₁ o₂, o₂ o₁], - { exact le_total o₁ o₂ }, - { have h : o₁ < o₂ := lt_of_le_of_ne h' h, - rw [impartial.fuzzy_zero_iff_gf, zero_lf_le, nim_def o₂], + { refine not_imp_not.1 (λ (h : _ ≠ _), (impartial.not_equiv_zero_iff _).2 _), + obtain h | h := h.lt_or_lt, + { rw [impartial.fuzzy_zero_iff_gf, zero_lf_le, nim_def o₂], refine ⟨to_left_moves_add (sum.inr _), _⟩, { exact (ordinal.principal_seg_out h).top }, { simpa using (impartial.add_self (nim o₁)).2 } }, - { exact (fuzzy_congr_left add_comm_equiv).1 (this (ne.symm h)) } }, + { rw [impartial.fuzzy_zero_iff_gf, zero_lf_le, nim_def o₁], + refine ⟨to_left_moves_add (sum.inl _), _⟩, + { exact (ordinal.principal_seg_out h).top }, + { simpa using (impartial.add_self (nim o₂)).2 } } }, { rintro rfl, exact impartial.add_self (nim o₁) } end diff --git a/src/tactic/linarith/lemmas.lean b/src/tactic/linarith/lemmas.lean index d14c9a2497a22..c5285ae9c6d32 100644 --- a/src/tactic/linarith/lemmas.lean +++ b/src/tactic/linarith/lemmas.lean @@ -59,7 +59,7 @@ by simp * lemma lt_of_lt_of_eq {α} [ordered_semiring α] {a b : α} (ha : a < 0) (hb : b = 0) : a + b < 0 := by simp * -lemma mul_neg {α} [ordered_ring α] {a b : α} (ha : a < 0) (hb : 0 < b) : b * a < 0 := +lemma mul_neg {α} [strict_ordered_ring α] {a b : α} (ha : a < 0) (hb : 0 < b) : b * a < 0 := have (-b)*a > 0, from mul_pos_of_neg_of_neg (neg_neg_of_pos hb) ha, neg_of_neg_pos (by simpa) diff --git a/src/tactic/norm_num.lean b/src/tactic/norm_num.lean index 38083d87ff708..0567a82827779 100644 --- a/src/tactic/norm_num.lean +++ b/src/tactic/norm_num.lean @@ -249,6 +249,8 @@ meta def prove_mul_nat : instance_cache → expr → expr → tactic (instance_c end +lemma zero_lt_one [linear_ordered_semiring α] : (0 : α) < 1 := zero_lt_one + section open match_numeral_result @@ -256,7 +258,7 @@ open match_numeral_result meta def prove_pos_nat (c : instance_cache) : expr → tactic (instance_cache × expr) | e := match match_numeral e with - | one := c.mk_app ``zero_lt_one' [] + | one := c.mk_app ``zero_lt_one [] | bit0 e := do (c, p) ← prove_pos_nat e, c.mk_app ``bit0_pos [e, p] | bit1 e := do (c, p) ← prove_pos_nat e, c.mk_app ``bit1_pos' [e, p] | _ := failed diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index 80a2df2c1a6b1..a50d7c1808ed9 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -382,7 +382,7 @@ begin end lemma limsup_eq_tendsto_sum_indicator_at_top - (R : Type*) [ordered_semiring R] [nontrivial R] [archimedean R] (s : ℕ → set α) : + (R : Type*) [strict_ordered_semiring R] [nontrivial R] [archimedean R] (s : ℕ → set α) : limsup at_top s = {ω | tendsto (λ n, ∑ k in finset.range n, (s (k + 1)).indicator (1 : α → R) ω) at_top at_top} := begin diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index 9c5c3117657dd..fcfcd22bc4213 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -436,7 +436,7 @@ begin have : μH[fintype.card ι] (metric.ball x r) = ennreal.of_real ((2 * r) ^ fintype.card ι), by rw [hausdorff_measure_pi_real, real.volume_pi_ball _ hr], refine dimH_of_hausdorff_measure_ne_zero_ne_top _ _; rw [nnreal.coe_nat_cast, this], - { simp [pow_pos (mul_pos zero_lt_two hr)] }, + { simp [pow_pos (mul_pos (zero_lt_two' ℝ) hr)] }, { exact ennreal.of_real_ne_top } } end diff --git a/test/instance_diamonds.lean b/test/instance_diamonds.lean index f6f3e1f2ba90a..9e0db87bd6486 100644 --- a/test/instance_diamonds.lean +++ b/test/instance_diamonds.lean @@ -108,17 +108,17 @@ end has_smul /-! ## `with_top` (Type with point at infinity) instances -/ section with_top -example (R : Type*) [h : ordered_semiring R] : +example (R : Type*) [h : strict_ordered_semiring R] : (@with_top.add_comm_monoid R (@non_unital_non_assoc_semiring.to_add_comm_monoid R (@non_assoc_semiring.to_non_unital_non_assoc_semiring R (@semiring.to_non_assoc_semiring R - (@ordered_semiring.to_semiring R h))))) + (@strict_ordered_semiring.to_semiring R h))))) = (@ordered_add_comm_monoid.to_add_comm_monoid (with_top R) (@with_top.ordered_add_comm_monoid R (@ordered_cancel_add_comm_monoid.to_ordered_add_comm_monoid R - (@ordered_semiring.to_ordered_cancel_add_comm_monoid R h)))) := + (@strict_ordered_semiring.to_ordered_cancel_add_comm_monoid R h)))) := rfl end with_top diff --git a/test/positivity.lean b/test/positivity.lean index 05bfb619c2efe..3a5f4d112fa24 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -136,7 +136,7 @@ example {a : ℤ} (ha : 0 < a) : 0 < a / a := by positivity example [ordered_semiring α] [nontrivial α] (a : α) : 0 < a ^ 0 := by positivity example [linear_ordered_ring α] (a : α) (n : ℕ) : 0 ≤ a ^ bit0 n := by positivity example [ordered_semiring α] {a : α} {n : ℕ} (ha : 0 ≤ a) : 0 ≤ a ^ n := by positivity -example [ordered_semiring α] {a : α} {n : ℕ} (ha : 0 < a) : 0 < a ^ n := by positivity +example [strict_ordered_semiring α] {a : α} {n : ℕ} (ha : 0 < a) : 0 < a ^ n := by positivity example [linear_ordered_semifield α] (a : α) : 0 < a ^ (0 : ℤ) := by positivity example [linear_ordered_field α] (a : α) (n : ℤ) : 0 ≤ a ^ bit0 n := by positivity @@ -235,7 +235,9 @@ example {a : ℝ≥0∞} : 0 ≤ a := by positivity /- ### Coercions -/ example {a : ℕ} : (0 : ℤ) ≤ a := by positivity +example {a : ℕ} : (0 : ℚ) ≤ a := by positivity example {a : ℕ} (ha : 0 < a) : (0 : ℤ) < a := by positivity +example {a : ℕ} (ha : 0 < a) : (0 : ℚ) < a := by positivity example {a : ℤ} (ha : a ≠ 0) : (a : ℚ) ≠ 0 := by positivity example {a : ℤ} (ha : 0 ≤ a) : (0 : ℚ) ≤ a := by positivity example {a : ℤ} (ha : 0 < a) : (0 : ℚ) < a := by positivity From 05b820ec79b3c98a7dbf1cb32e181584166da2ca Mon Sep 17 00:00:00 2001 From: Jujian Zhang Date: Sat, 8 Oct 2022 12:41:18 +0000 Subject: [PATCH 631/828] feat(topology/sheaves/stalks): stalk functor preserves monomorphism (#16797) --- .../concrete_category/basic.lean | 19 +++++++++++++++ src/topology/sheaves/stalks.lean | 24 +++++++++++++++++++ 2 files changed, 43 insertions(+) diff --git a/src/category_theory/concrete_category/basic.lean b/src/category_theory/concrete_category/basic.lean index f39794e77d776..4a74dcdba59c6 100644 --- a/src/category_theory/concrete_category/basic.lean +++ b/src/category_theory/concrete_category/basic.lean @@ -5,6 +5,7 @@ Authors: Scott Morrison, Johannes Hölzl, Reid Barton, Sean Leather, Yury Kudrya -/ import category_theory.types import category_theory.functor.epi_mono +import category_theory.limits.constructions.epi_mono /-! # Concrete categories @@ -38,6 +39,8 @@ universes w v v' u namespace category_theory +open category_theory.limits + /-- A concrete category is a category `C` with a fixed faithful functor `forget : C ⥤ Type`. @@ -129,11 +132,27 @@ lemma concrete_category.mono_of_injective {X Y : C} (f : X ⟶ Y) (i : function. mono f := (forget C).mono_of_mono_map ((mono_iff_injective f).2 i) +lemma concrete_category.injective_of_mono_of_preserves_pullback {X Y : C} (f : X ⟶ Y) [mono f] + [preserves_limits_of_shape walking_cospan (forget C)] : function.injective f := +(mono_iff_injective ((forget C).map f)).mp infer_instance + +lemma concrete_category.mono_iff_injective_of_preserves_pullback {X Y : C} (f : X ⟶ Y) + [preserves_limits_of_shape walking_cospan (forget C)] : mono f ↔ function.injective f := +((forget C).mono_map_iff_mono _).symm.trans (mono_iff_injective _) + /-- In any concrete category, surjective morphisms are epimorphisms. -/ lemma concrete_category.epi_of_surjective {X Y : C} (f : X ⟶ Y) (s : function.surjective f) : epi f := (forget C).epi_of_epi_map ((epi_iff_surjective f).2 s) +lemma concrete_category.surjective_of_epi_of_preserves_pushout {X Y : C} (f : X ⟶ Y) [epi f] + [preserves_colimits_of_shape walking_span (forget C)] : function.surjective f := +(epi_iff_surjective ((forget C).map f)).mp infer_instance + +lemma concrete_category.epi_iff_surjective_of_preserves_pushout {X Y : C} (f : X ⟶ Y) + [preserves_colimits_of_shape walking_span (forget C)] : epi f ↔ function.surjective f := +((forget C).epi_map_iff_epi _).symm.trans (epi_iff_surjective _) + lemma concrete_category.bijective_of_is_iso {X Y : C} (f : X ⟶ Y) [is_iso f] : function.bijective ((forget C).map f) := by { rw ← is_iso_iff_bijective, apply_instance, } diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 90f4f1ce8cb0c..91abec167dd3f 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -6,6 +6,7 @@ Authors: Scott Morrison, Justus Springer import topology.category.Top.open_nhds import topology.sheaves.presheaf import topology.sheaves.sheaf_condition.unique_gluing +import category_theory.adjunction.evaluation import category_theory.limits.types import category_theory.limits.preserves.filtered import category_theory.limits.final @@ -421,6 +422,29 @@ lemma app_injective_iff_stalk_functor_map_injective {F : sheaf C X} ⟨λ h U, app_injective_of_stalk_functor_map_injective f U (λ x, h x.1), stalk_functor_map_injective_of_app_injective f⟩ +instance stalk_functor_preserves_mono (x : X) : + functor.preserves_monomorphisms (sheaf.forget C X ⋙ stalk_functor C x) := +⟨λ 𝓐 𝓑 f m, concrete_category.mono_of_injective _ $ + (app_injective_iff_stalk_functor_map_injective f.1).mpr + (λ c, (@@concrete_category.mono_iff_injective_of_preserves_pullback _ _ (f.1.app (op c)) _).mp + ((nat_trans.mono_iff_mono_app _ f.1).mp + (@@category_theory.presheaf_mono_of_mono _ _ _ _ _ _ _ _ _ _ _ m) $ op c)) x⟩ + +lemma stalk_mono_of_mono {F G : sheaf C X} (f : F ⟶ G) [mono f] : + Π x, mono $ (stalk_functor C x).map f.1 := +λ x, by convert functor.map_mono (sheaf.forget.{v} C X ⋙ stalk_functor C x) f + +lemma mono_of_stalk_mono {F G : sheaf C X} (f : F ⟶ G) + [Π x, mono $ (stalk_functor C x).map f.1] : mono f := +(Sheaf.hom.mono_iff_presheaf_mono _ _ _).mpr $ (nat_trans.mono_iff_mono_app _ _).mpr $ λ U, + (concrete_category.mono_iff_injective_of_preserves_pullback _).mpr $ + app_injective_of_stalk_functor_map_injective f.1 U.unop $ λ ⟨x, hx⟩, + (concrete_category.mono_iff_injective_of_preserves_pullback _).mp $ infer_instance + +lemma mono_iff_stalk_mono {F G : sheaf C X} (f : F ⟶ G) : + mono f ↔ ∀ x, mono ((stalk_functor C x).map f.1) := +⟨by { introI m, exact stalk_mono_of_mono _ }, by { introI m, exact mono_of_stalk_mono _ }⟩ + /-- For surjectivity, we are given an arbitrary section `t` and need to find a preimage for it. We claim that it suffices to find preimages *locally*. That is, for each `x : U` we construct a neighborhood `V ≤ U` and a section `s : F.obj (op V))` such that `f.app (op V) s` and `t` From 0ff989e5d31dd73bcea3c52be7b601c95d4e3930 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 8 Oct 2022 15:49:24 +0000 Subject: [PATCH 632/828] chore(algebra/group/basic): split file (#16847) Motivated by porting material on ordered_ring for linarith. Just stripping out tangential imports from files at the very bottom of the import thicket. Co-authored-by: Scott Morrison --- src/algebra/group/basic.lean | 118 ------------------------- src/algebra/group/commutator.lean | 19 ++++ src/algebra/group/order_synonym.lean | 116 ++++++++++++++++++++++++ src/algebra/group_with_zero/basic.lean | 1 + src/tactic/group.lean | 1 + 5 files changed, 137 insertions(+), 118 deletions(-) create mode 100644 src/algebra/group/commutator.lean create mode 100644 src/algebra/group/order_synonym.lean diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index 097f8c2fd4ba0..c33fe507ca94d 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -5,9 +5,6 @@ Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro -/ import algebra.group.defs -import data.bracket -import logic.function.basic -import order.synonym /-! # Basic lemmas about semigroups, monoids, and groups @@ -650,118 +647,3 @@ lemma bit1_sub [has_one M] (a b : M) : bit1 (a - b) = bit1 a - bit0 b := (congr_arg (+ (1 : M)) $ bit0_sub a b : _).trans $ sub_add_eq_add_sub _ _ _ end subtraction_comm_monoid - -section commutator - -/-- The commutator of two elements `g₁` and `g₂`. -/ -instance commutator_element {G : Type*} [group G] : has_bracket G G := -⟨λ g₁ g₂, g₁ * g₂ * g₁⁻¹ * g₂⁻¹⟩ - -lemma commutator_element_def {G : Type*} [group G] (g₁ g₂ : G) : - ⁅g₁, g₂⁆ = g₁ * g₂ * g₁⁻¹ * g₂⁻¹ := rfl - -end commutator - -/-! ### Order dual -/ - -open order_dual - -@[to_additive] instance [h : has_one α] : has_one αᵒᵈ := h -@[to_additive] instance [h : has_mul α] : has_mul αᵒᵈ := h -@[to_additive] instance [h : has_inv α] : has_inv αᵒᵈ := h -@[to_additive] instance [h : has_div α] : has_div αᵒᵈ := h -@[to_additive] instance [h : has_smul α β] : has_smul α βᵒᵈ := h -@[to_additive] instance order_dual.has_pow [h : has_pow α β] : has_pow αᵒᵈ β := h -@[to_additive] instance [h : semigroup α] : semigroup αᵒᵈ := h -@[to_additive] instance [h : comm_semigroup α] : comm_semigroup αᵒᵈ := h -@[to_additive] instance [h : left_cancel_semigroup α] : left_cancel_semigroup αᵒᵈ := h -@[to_additive] instance [h : right_cancel_semigroup α] : right_cancel_semigroup αᵒᵈ := h -@[to_additive] instance [h : mul_one_class α] : mul_one_class αᵒᵈ := h -@[to_additive] instance [h : monoid α] : monoid αᵒᵈ := h -@[to_additive] instance [h : comm_monoid α] : comm_monoid αᵒᵈ := h -@[to_additive] instance [h : left_cancel_monoid α] : left_cancel_monoid αᵒᵈ := h -@[to_additive] instance [h : right_cancel_monoid α] : right_cancel_monoid αᵒᵈ := h -@[to_additive] instance [h : cancel_monoid α] : cancel_monoid αᵒᵈ := h -@[to_additive] instance [h : cancel_comm_monoid α] : cancel_comm_monoid αᵒᵈ := h -@[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv αᵒᵈ := h -@[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid αᵒᵈ := h -@[to_additive order_dual.subtraction_monoid] -instance [h : division_monoid α] : division_monoid αᵒᵈ := h -@[to_additive order_dual.subtraction_comm_monoid] -instance [h : division_comm_monoid α] : division_comm_monoid αᵒᵈ := h -@[to_additive] instance [h : group α] : group αᵒᵈ := h -@[to_additive] instance [h : comm_group α] : comm_group αᵒᵈ := h - -@[simp, to_additive] lemma to_dual_one [has_one α] : to_dual (1 : α) = 1 := rfl -@[simp, to_additive] lemma of_dual_one [has_one α] : (of_dual 1 : α) = 1 := rfl -@[simp, to_additive] -lemma to_dual_mul [has_mul α] (a b : α) : to_dual (a * b) = to_dual a * to_dual b := rfl -@[simp, to_additive] -lemma of_dual_mul [has_mul α] (a b : αᵒᵈ) : of_dual (a * b) = of_dual a * of_dual b := rfl -@[simp, to_additive] lemma to_dual_inv [has_inv α] (a : α) : to_dual a⁻¹ = (to_dual a)⁻¹ := rfl -@[simp, to_additive] lemma of_dual_inv [has_inv α] (a : αᵒᵈ) : of_dual a⁻¹ = (of_dual a)⁻¹ := rfl -@[simp, to_additive] -lemma to_dual_div [has_div α] (a b : α) : to_dual (a / b) = to_dual a / to_dual b := rfl -@[simp, to_additive] -lemma of_dual_div [has_div α] (a b : αᵒᵈ) : of_dual (a / b) = of_dual a / of_dual b := rfl -lemma to_dual_vadd [has_vadd α β] (a : α) (b : β) : to_dual (a +ᵥ b) = a +ᵥ to_dual b := rfl -lemma of_dual_vadd [has_vadd α β] (a : α) (b : βᵒᵈ) : of_dual (a +ᵥ b) = a +ᵥ of_dual b := rfl -@[simp, to_additive] -lemma to_dual_smul [has_smul α β] (a : α) (b : β) : to_dual (a • b) = a • to_dual b := rfl -@[simp, to_additive] -lemma of_dual_smul [has_smul α β] (a : α) (b : βᵒᵈ) : of_dual (a • b) = a • of_dual b := rfl -@[simp, to_additive to_dual_smul] -lemma to_dual_pow [has_pow α β] (a : α) (b : β) : to_dual (a ^ b) = to_dual a ^ b := rfl -@[simp, to_additive of_dual_smul] -lemma of_dual_pow [has_pow α β] (a : αᵒᵈ) (b : β) : of_dual (a ^ b) = of_dual a ^ b := rfl - -/-! ### Lexicographical order -/ - -@[to_additive] instance [h : has_one α] : has_one (lex α) := h -@[to_additive] instance [h : has_mul α] : has_mul (lex α) := h -@[to_additive] instance [h : has_inv α] : has_inv (lex α) := h -@[to_additive] instance [h : has_div α] : has_div (lex α) := h -@[to_additive] instance [h : has_smul α β] : has_smul α (lex β) := h -@[to_additive] instance lex.has_pow [h : has_pow α β] : has_pow (lex α) β := h -@[to_additive] instance [h : semigroup α] : semigroup (lex α) := h -@[to_additive] instance [h : comm_semigroup α] : comm_semigroup (lex α) := h -@[to_additive] instance [h : left_cancel_semigroup α] : left_cancel_semigroup (lex α) := h -@[to_additive] instance [h : right_cancel_semigroup α] : right_cancel_semigroup (lex α) := h -@[to_additive] instance [h : mul_one_class α] : mul_one_class (lex α) := h -@[to_additive] instance [h : monoid α] : monoid (lex α) := h -@[to_additive] instance [h : comm_monoid α] : comm_monoid (lex α) := h -@[to_additive] instance [h : left_cancel_monoid α] : left_cancel_monoid (lex α) := h -@[to_additive] instance [h : right_cancel_monoid α] : right_cancel_monoid (lex α) := h -@[to_additive] instance [h : cancel_monoid α] : cancel_monoid (lex α) := h -@[to_additive] instance [h : cancel_comm_monoid α] : cancel_comm_monoid (lex α) := h -@[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv (lex α) := h -@[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid (lex α) := h -@[to_additive order_dual.subtraction_monoid] -instance [h : division_monoid α] : division_monoid (lex α) := h -@[to_additive order_dual.subtraction_comm_monoid] -instance [h : division_comm_monoid α] : division_comm_monoid (lex α) := h -@[to_additive] instance [h : group α] : group (lex α) := h -@[to_additive] instance [h : comm_group α] : comm_group (lex α) := h - -@[simp, to_additive] lemma to_lex_one [has_one α] : to_lex (1 : α) = 1 := rfl -@[simp, to_additive] lemma of_lex_one [has_one α] : (of_lex 1 : α) = 1 := rfl -@[simp, to_additive] -lemma to_lex_mul [has_mul α] (a b : α) : to_lex (a * b) = to_lex a * to_lex b := rfl -@[simp, to_additive] -lemma of_lex_mul [has_mul α] (a b : αᵒᵈ) : of_lex (a * b) = of_lex a * of_lex b := rfl -@[simp, to_additive] lemma to_lex_inv [has_inv α] (a : α) : to_lex a⁻¹ = (to_lex a)⁻¹ := rfl -@[simp, to_additive] lemma of_lex_inv [has_inv α] (a : αᵒᵈ) : of_lex a⁻¹ = (of_lex a)⁻¹ := rfl -@[simp, to_additive] -lemma to_lex_div [has_div α] (a b : α) : to_lex (a / b) = to_lex a / to_lex b := rfl -@[simp, to_additive] -lemma of_lex_div [has_div α] (a b : αᵒᵈ) : of_lex (a / b) = of_lex a / of_lex b := rfl -lemma to_lex_vadd [has_vadd α β] (a : α) (b : β) : to_lex (a +ᵥ b) = a +ᵥ to_lex b := rfl -lemma of_lex_vadd [has_vadd α β] (a : α) (b : βᵒᵈ) : of_lex (a +ᵥ b) = a +ᵥ of_lex b := rfl -@[simp, to_additive] -lemma to_lex_smul [has_smul α β] (a : α) (b : β) : to_lex (a • b) = a • to_lex b := rfl -@[simp, to_additive] -lemma of_lex_smul [has_smul α β] (a : α) (b : βᵒᵈ) : of_lex (a • b) = a • of_lex b := rfl -@[simp, to_additive to_lex_smul, to_additive_reorder 1 4] -lemma to_lex_pow [has_pow α β] (a : α) (b : β) : to_lex (a ^ b) = to_lex a ^ b := rfl -@[simp, to_additive of_lex_smul, to_additive_reorder 1 4] -lemma of_lex_pow [has_pow α β] (a : αᵒᵈ) (b : β) : of_lex (a ^ b) = of_lex a ^ b := rfl diff --git a/src/algebra/group/commutator.lean b/src/algebra/group/commutator.lean new file mode 100644 index 0000000000000..2b25672d50d3e --- /dev/null +++ b/src/algebra/group/commutator.lean @@ -0,0 +1,19 @@ +/- +Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro +-/ + +import algebra.group.defs +import data.bracket + +/-! +# The bracket on a group given by commutator. +-/ + +/-- The commutator of two elements `g₁` and `g₂`. -/ +instance commutator_element {G : Type*} [group G] : has_bracket G G := +⟨λ g₁ g₂, g₁ * g₂ * g₁⁻¹ * g₂⁻¹⟩ + +lemma commutator_element_def {G : Type*} [group G] (g₁ g₂ : G) : + ⁅g₁, g₂⁆ = g₁ * g₂ * g₁⁻¹ * g₂⁻¹ := rfl diff --git a/src/algebra/group/order_synonym.lean b/src/algebra/group/order_synonym.lean new file mode 100644 index 0000000000000..5020ec4fc9ed1 --- /dev/null +++ b/src/algebra/group/order_synonym.lean @@ -0,0 +1,116 @@ +/- +Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro +-/ + +import algebra.group.defs +import order.synonym + +/-! # Group structure on the type synonyms -/ + +variables {α β : Type*} + +/-! ### order_dual -/ + +open order_dual + +@[to_additive] instance [h : has_one α] : has_one αᵒᵈ := h +@[to_additive] instance [h : has_mul α] : has_mul αᵒᵈ := h +@[to_additive] instance [h : has_inv α] : has_inv αᵒᵈ := h +@[to_additive] instance [h : has_div α] : has_div αᵒᵈ := h +@[to_additive] instance [h : has_smul α β] : has_smul α βᵒᵈ := h +@[to_additive] instance order_dual.has_pow [h : has_pow α β] : has_pow αᵒᵈ β := h +@[to_additive] instance [h : semigroup α] : semigroup αᵒᵈ := h +@[to_additive] instance [h : comm_semigroup α] : comm_semigroup αᵒᵈ := h +@[to_additive] instance [h : left_cancel_semigroup α] : left_cancel_semigroup αᵒᵈ := h +@[to_additive] instance [h : right_cancel_semigroup α] : right_cancel_semigroup αᵒᵈ := h +@[to_additive] instance [h : mul_one_class α] : mul_one_class αᵒᵈ := h +@[to_additive] instance [h : monoid α] : monoid αᵒᵈ := h +@[to_additive] instance [h : comm_monoid α] : comm_monoid αᵒᵈ := h +@[to_additive] instance [h : left_cancel_monoid α] : left_cancel_monoid αᵒᵈ := h +@[to_additive] instance [h : right_cancel_monoid α] : right_cancel_monoid αᵒᵈ := h +@[to_additive] instance [h : cancel_monoid α] : cancel_monoid αᵒᵈ := h +@[to_additive] instance [h : cancel_comm_monoid α] : cancel_comm_monoid αᵒᵈ := h +@[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv αᵒᵈ := h +@[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid αᵒᵈ := h +@[to_additive order_dual.subtraction_monoid] +instance [h : division_monoid α] : division_monoid αᵒᵈ := h +@[to_additive order_dual.subtraction_comm_monoid] +instance [h : division_comm_monoid α] : division_comm_monoid αᵒᵈ := h +@[to_additive] instance [h : group α] : group αᵒᵈ := h +@[to_additive] instance [h : comm_group α] : comm_group αᵒᵈ := h + +@[simp, to_additive] lemma to_dual_one [has_one α] : to_dual (1 : α) = 1 := rfl +@[simp, to_additive] lemma of_dual_one [has_one α] : (of_dual 1 : α) = 1 := rfl +@[simp, to_additive] +lemma to_dual_mul [has_mul α] (a b : α) : to_dual (a * b) = to_dual a * to_dual b := rfl +@[simp, to_additive] +lemma of_dual_mul [has_mul α] (a b : αᵒᵈ) : of_dual (a * b) = of_dual a * of_dual b := rfl +@[simp, to_additive] lemma to_dual_inv [has_inv α] (a : α) : to_dual a⁻¹ = (to_dual a)⁻¹ := rfl +@[simp, to_additive] lemma of_dual_inv [has_inv α] (a : αᵒᵈ) : of_dual a⁻¹ = (of_dual a)⁻¹ := rfl +@[simp, to_additive] +lemma to_dual_div [has_div α] (a b : α) : to_dual (a / b) = to_dual a / to_dual b := rfl +@[simp, to_additive] +lemma of_dual_div [has_div α] (a b : αᵒᵈ) : of_dual (a / b) = of_dual a / of_dual b := rfl +lemma to_dual_vadd [has_vadd α β] (a : α) (b : β) : to_dual (a +ᵥ b) = a +ᵥ to_dual b := rfl +lemma of_dual_vadd [has_vadd α β] (a : α) (b : βᵒᵈ) : of_dual (a +ᵥ b) = a +ᵥ of_dual b := rfl +@[simp, to_additive] +lemma to_dual_smul [has_smul α β] (a : α) (b : β) : to_dual (a • b) = a • to_dual b := rfl +@[simp, to_additive] +lemma of_dual_smul [has_smul α β] (a : α) (b : βᵒᵈ) : of_dual (a • b) = a • of_dual b := rfl +@[simp, to_additive to_dual_smul] +lemma to_dual_pow [has_pow α β] (a : α) (b : β) : to_dual (a ^ b) = to_dual a ^ b := rfl +@[simp, to_additive of_dual_smul] +lemma of_dual_pow [has_pow α β] (a : αᵒᵈ) (b : β) : of_dual (a ^ b) = of_dual a ^ b := rfl + +/-! ### Lexicographical order -/ + +@[to_additive] instance [h : has_one α] : has_one (lex α) := h +@[to_additive] instance [h : has_mul α] : has_mul (lex α) := h +@[to_additive] instance [h : has_inv α] : has_inv (lex α) := h +@[to_additive] instance [h : has_div α] : has_div (lex α) := h +@[to_additive] instance [h : has_smul α β] : has_smul α (lex β) := h +@[to_additive] instance lex.has_pow [h : has_pow α β] : has_pow (lex α) β := h +@[to_additive] instance [h : semigroup α] : semigroup (lex α) := h +@[to_additive] instance [h : comm_semigroup α] : comm_semigroup (lex α) := h +@[to_additive] instance [h : left_cancel_semigroup α] : left_cancel_semigroup (lex α) := h +@[to_additive] instance [h : right_cancel_semigroup α] : right_cancel_semigroup (lex α) := h +@[to_additive] instance [h : mul_one_class α] : mul_one_class (lex α) := h +@[to_additive] instance [h : monoid α] : monoid (lex α) := h +@[to_additive] instance [h : comm_monoid α] : comm_monoid (lex α) := h +@[to_additive] instance [h : left_cancel_monoid α] : left_cancel_monoid (lex α) := h +@[to_additive] instance [h : right_cancel_monoid α] : right_cancel_monoid (lex α) := h +@[to_additive] instance [h : cancel_monoid α] : cancel_monoid (lex α) := h +@[to_additive] instance [h : cancel_comm_monoid α] : cancel_comm_monoid (lex α) := h +@[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv (lex α) := h +@[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid (lex α) := h +@[to_additive order_dual.subtraction_monoid] +instance [h : division_monoid α] : division_monoid (lex α) := h +@[to_additive order_dual.subtraction_comm_monoid] +instance [h : division_comm_monoid α] : division_comm_monoid (lex α) := h +@[to_additive] instance [h : group α] : group (lex α) := h +@[to_additive] instance [h : comm_group α] : comm_group (lex α) := h + +@[simp, to_additive] lemma to_lex_one [has_one α] : to_lex (1 : α) = 1 := rfl +@[simp, to_additive] lemma of_lex_one [has_one α] : (of_lex 1 : α) = 1 := rfl +@[simp, to_additive] +lemma to_lex_mul [has_mul α] (a b : α) : to_lex (a * b) = to_lex a * to_lex b := rfl +@[simp, to_additive] +lemma of_lex_mul [has_mul α] (a b : αᵒᵈ) : of_lex (a * b) = of_lex a * of_lex b := rfl +@[simp, to_additive] lemma to_lex_inv [has_inv α] (a : α) : to_lex a⁻¹ = (to_lex a)⁻¹ := rfl +@[simp, to_additive] lemma of_lex_inv [has_inv α] (a : αᵒᵈ) : of_lex a⁻¹ = (of_lex a)⁻¹ := rfl +@[simp, to_additive] +lemma to_lex_div [has_div α] (a b : α) : to_lex (a / b) = to_lex a / to_lex b := rfl +@[simp, to_additive] +lemma of_lex_div [has_div α] (a b : αᵒᵈ) : of_lex (a / b) = of_lex a / of_lex b := rfl +lemma to_lex_vadd [has_vadd α β] (a : α) (b : β) : to_lex (a +ᵥ b) = a +ᵥ to_lex b := rfl +lemma of_lex_vadd [has_vadd α β] (a : α) (b : βᵒᵈ) : of_lex (a +ᵥ b) = a +ᵥ of_lex b := rfl +@[simp, to_additive] +lemma to_lex_smul [has_smul α β] (a : α) (b : β) : to_lex (a • b) = a • to_lex b := rfl +@[simp, to_additive] +lemma of_lex_smul [has_smul α β] (a : α) (b : βᵒᵈ) : of_lex (a • b) = a • of_lex b := rfl +@[simp, to_additive to_lex_smul, to_additive_reorder 1 4] +lemma to_lex_pow [has_pow α β] (a : α) (b : β) : to_lex (a ^ b) = to_lex a ^ b := rfl +@[simp, to_additive of_lex_smul, to_additive_reorder 1 4] +lemma of_lex_pow [has_pow α β] (a : αᵒᵈ) (b : β) : of_lex (a ^ b) = of_lex a ^ b := rfl diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index e642256f50ef4..b2adb88d9eab1 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -5,6 +5,7 @@ Authors: Johan Commelin -/ import algebra.group.inj_surj import algebra.group_with_zero.defs +import algebra.group.order_synonym import algebra.hom.units import logic.nontrivial import group_theory.group_action.units diff --git a/src/tactic/group.lean b/src/tactic/group.lean index 99a8ebb16824b..1cfe207169a8e 100644 --- a/src/tactic/group.lean +++ b/src/tactic/group.lean @@ -5,6 +5,7 @@ Authors: Patrick Massot -/ import tactic.ring import tactic.doc_commands +import algebra.group.commutator /-! # `group` From 675192fccd9c2d7f7f4b655797d70f1715720eca Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Sat, 8 Oct 2022 15:49:24 +0000 Subject: [PATCH 633/828] feat(topology): accumulation in infinite compact sets. (#16862) --- src/order/filter/basic.lean | 15 +++++++++++++++ src/topology/basic.lean | 7 ++++--- src/topology/constructions.lean | 17 +++++++++++++++++ src/topology/order.lean | 25 +++++++++++++++++++++++++ src/topology/separation.lean | 21 --------------------- src/topology/subset_properties.lean | 20 ++++++++++++++++++++ 6 files changed, 81 insertions(+), 24 deletions(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index fcf4c0a0c7d52..6d3e18b988a20 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1791,6 +1791,14 @@ lemma le_comap_map : f ≤ comap m (map m f) := (gc_map_comap m).le_u_l _ @[simp] lemma comap_bot : comap m ⊥ = ⊥ := bot_unique $ λ s _, ⟨∅, mem_bot, by simp only [empty_subset, preimage_empty]⟩ +lemma ne_bot_of_comap (h : (comap m g).ne_bot) : g.ne_bot := +begin + rw ne_bot_iff at *, + contrapose! h, + rw h, + exact comap_bot +end + lemma comap_inf_principal_range : comap m (g ⊓ 𝓟 (range m)) = comap m g := by simp lemma disjoint_comap (h : disjoint g₁ g₂) : disjoint (comap m g₁) (comap m g₂) := @@ -2111,6 +2119,13 @@ protected lemma push_pull' (f : α → β) (F : filter α) (G : filter β) : map f (comap f G ⊓ F) = G ⊓ map f F := by simp only [filter.push_pull, inf_comm] +lemma principal_eq_map_coe_top (s : set α) : 𝓟 s = map (coe : s → α) ⊤ := +by simp + +lemma inf_principal_eq_bot_iff_comap {F : filter α} {s : set α} : + F ⊓ 𝓟 s = ⊥ ↔ comap (coe : s → α) F = ⊥ := +by rw [principal_eq_map_coe_top s, ← filter.push_pull',inf_top_eq, map_eq_bot_iff] + section applicative lemma singleton_mem_pure {a : α} : {a} ∈ (pure a : filter α) := diff --git a/src/topology/basic.lean b/src/topology/basic.lean index cd10a32dfe068..c63e3dc256205 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -857,8 +857,9 @@ In this section we define [cluster points](https://en.wikipedia.org/wiki/Limit_p (also known as limit points and accumulation points) of a filter and of a sequence. -/ -/-- A point `x` is a cluster point of a filter `F` if 𝓝 x ⊓ F ≠ ⊥. Also known as -an accumulation point or a limit point. -/ +/-- A point `x` is a cluster point of a filter `F` if `𝓝 x ⊓ F ≠ ⊥`. Also known as +an accumulation point or a limit point, but beware that terminology varies. This +is *not* the same as asking `𝓝[≠] x ⊓ F ≠ ⊥`. See `mem_closure_iff_cluster_pt` in particular. -/ def cluster_pt (x : α) (F : filter α) : Prop := ne_bot (𝓝 x ⊓ F) lemma cluster_pt.ne_bot {x : α} {F : filter α} (h : cluster_pt x F) : ne_bot (𝓝 x ⊓ F) := h @@ -874,7 +875,7 @@ lemma cluster_pt_iff {x : α} {F : filter α} : inf_ne_bot_iff /-- `x` is a cluster point of a set `s` if every neighbourhood of `x` meets `s` on a nonempty -set. -/ +set. See also `mem_closure_iff_cluster_pt`. -/ lemma cluster_pt_principal_iff {x : α} {s : set α} : cluster_pt x (𝓟 s) ↔ ∀ U ∈ 𝓝 x, (U ∩ s).nonempty := inf_principal_ne_bot_iff diff --git a/src/topology/constructions.lean b/src/topology/constructions.lean index cb915182bc25a..de8ade7757520 100644 --- a/src/topology/constructions.lean +++ b/src/topology/constructions.lean @@ -180,6 +180,23 @@ theorem nhds_subtype (s : set α) (a : {x // x ∈ s}) : 𝓝 a = comap coe (𝓝 (a : α)) := nhds_induced coe a +lemma nhds_within_subtype_eq_bot_iff {s t : set α} {x : s} : + 𝓝[(coe : s → α) ⁻¹' t] x = ⊥ ↔ 𝓝[t] (x : α) ⊓ 𝓟 s = ⊥ := +by rw [inf_principal_eq_bot_iff_comap, nhds_within, nhds_within, comap_inf, comap_principal, + nhds_induced] + +lemma nhds_ne_subtype_eq_bot_iff {S : set α} {x : S} : 𝓝[{x}ᶜ] x = ⊥ ↔ 𝓝[{x}ᶜ] (x : α) ⊓ 𝓟 S = ⊥ := +by rw [← nhds_within_subtype_eq_bot_iff, preimage_compl, ← image_singleton, + subtype.coe_injective.preimage_image ] + +lemma nhds_ne_subtype_ne_bot_iff {S : set α} {x : S} : + (𝓝[{x}ᶜ] x).ne_bot ↔ (𝓝[{x}ᶜ] (x : α) ⊓ 𝓟 S).ne_bot := +by rw [ne_bot_iff, ne_bot_iff, not_iff_not, nhds_ne_subtype_eq_bot_iff] + +lemma discrete_topology_subtype_iff {S : set α} : + discrete_topology S ↔ ∀ x ∈ S, 𝓝[≠] x ⊓ 𝓟 S = ⊥ := +by simp_rw [discrete_topology_iff_nhds_ne, set_coe.forall', nhds_ne_subtype_eq_bot_iff] + end topα /-- A type synonym equiped with the topology whose open sets are the empty set and the sets with diff --git a/src/topology/order.lean b/src/topology/order.lean index 5ab285f5c4d75..facb9f1143ecc 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -327,6 +327,31 @@ lemma singletons_open_iff_discrete {X : Type*} [topological_space X] : (∀ a : X, is_open ({a} : set X)) ↔ discrete_topology X := ⟨λ h, ⟨eq_bot_of_singletons_open h⟩, λ a _, @is_open_discrete _ _ a _⟩ +/-- This lemma characterizes discrete topological spaces as those whose singletons are +neighbourhoods. -/ +lemma discrete_topology_iff_nhds [topological_space α] : + discrete_topology α ↔ ∀ x : α, 𝓝 x = pure x := +begin + split ; introI h, + { intro x, + rw nhds_discrete }, + { constructor, + apply eq_of_nhds_eq_nhds, + simp [h, nhds_discrete] }, +end + +lemma discrete_topology_iff_nhds_ne [topological_space α] : + discrete_topology α ↔ ∀ x : α, 𝓝[≠] x = ⊥ := +begin + rw discrete_topology_iff_nhds, + apply forall_congr (λ x, _), + rw [nhds_within, inf_principal_eq_bot, compl_compl], + split ; intro h, + { rw h, + exact singleton_mem_pure }, + { exact le_antisymm (le_pure_iff.mpr h) (pure_le_nhds x) } +end + end lattice section galois_connection diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 7588d5d98802d..166be11df5040 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -79,13 +79,6 @@ If the space is also compact: * `disjoint_nested_nhds`: Given two points `x ≠ y`, we can find neighbourhoods `x ∈ V₁ ⊆ U₁` and `y ∈ V₂ ⊆ U₂`, with the `Vₖ` closed and the `Uₖ` open, such that the `Uₖ` are disjoint. -### Discrete spaces - -* `discrete_topology_iff_nhds`: Discrete topological spaces are those whose neighbourhood - filters are the `pure` filter (which is the principal filter at a singleton). -* `induced_bot`/`discrete_topology_induced`: The pullback of the discrete topology - under an inclusion is the discrete topology. - ## References https://en.wikipedia.org/wiki/Separation_axiom @@ -741,20 +734,6 @@ begin rw ← induced_compose, end -/-- This lemma characterizes discrete topological spaces as those whose singletons are -neighbourhoods. -/ -lemma discrete_topology_iff_nhds {X : Type*} [topological_space X] : - discrete_topology X ↔ (nhds : X → filter X) = pure := -begin - split, - { introI hX, - exact nhds_discrete X }, - { intro h, - constructor, - apply eq_of_nhds_eq_nhds, - simp [h, nhds_bot] } -end - /-- The topology pulled-back under an inclusion `f : X → Y` from the discrete topology (`⊥`) is the discrete topology. This version does not assume the choice of a topology on either the source `X` diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index b824ae53ca921..b8706ea8dc33c 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -776,6 +776,15 @@ noncompact_space_of_ne_bot $ by simp only [filter.cocompact_eq_cofinite, filter. lemma finite_of_compact_of_discrete [compact_space α] [discrete_topology α] : finite α := finite.of_finite_univ $ compact_univ.finite_of_discrete +lemma exists_nhds_ne_ne_bot (α : Type*) [topological_space α] [compact_space α] [infinite α] : + ∃ z : α, (𝓝[≠] z).ne_bot := +begin + by_contra' H, + simp_rw not_ne_bot at H, + haveI := discrete_topology_iff_nhds_ne.mpr H, + exact infinite.not_finite (finite_of_compact_of_discrete : finite α), +end + lemma finite_cover_nhds_interior [compact_space α] {U : α → set α} (hU : ∀ x, U x ∈ 𝓝 x) : ∃ t : finset α, (⋃ x ∈ t, interior (U x)) = univ := let ⟨t, ht⟩ := compact_univ.elim_finite_subcover (λ x, interior (U x)) (λ x, is_open_interior) @@ -911,6 +920,17 @@ by rw [compact_iff_compact_in_subtype, image_univ, subtype.range_coe]; refl lemma is_compact_iff_compact_space {s : set α} : is_compact s ↔ compact_space s := is_compact_iff_is_compact_univ.trans ⟨λ h, ⟨h⟩, @compact_space.compact_univ _ _⟩ +lemma is_compact.finite {s : set α} (hs : is_compact s) (hs' : discrete_topology s) : s.finite := +finite_coe_iff.mp (@finite_of_compact_of_discrete _ _ (is_compact_iff_compact_space.mp hs) hs') + +lemma exists_nhds_ne_inf_principal_ne_bot {s : set α} (hs : is_compact s) (hs' : s.infinite) : + ∃ z ∈ s, (𝓝[≠] z ⊓ 𝓟 s).ne_bot := +begin + by_contra' H, + simp_rw not_ne_bot at H, + exact hs' (hs.finite $ discrete_topology_subtype_iff.mpr H), +end + protected lemma closed_embedding.noncompact_space [noncompact_space α] {f : α → β} (hf : closed_embedding f) : noncompact_space β := noncompact_space_of_ne_bot hf.tendsto_cocompact.ne_bot From 446eb51ce0a90f8385f260d2b52e760e2004246b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sat, 8 Oct 2022 20:02:40 +0000 Subject: [PATCH 634/828] feat(linear_algebra/clifford_algebra): the clifford algebra is isomorphic as a module to the exterior algebra (#11468) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The key result here is ```lean /-- The module isomorphism to the exterior algebra -/ def equiv_exterior [invertible (2 : R)] : clifford_algebra Q ≃ₗ[R] exterior_algebra R M := ``` There are a handful of intermediate definitions that are needed to get here that are missing lots of useful (but difficult) API lemmas, but I don't expect to have time to address those for a while. Probably the main missing result is that `equiv_exterior` preserves reversion. --- docs/references.bib | 21 ++ .../clifford_algebra/contraction.lean | 351 ++++++++++++++++++ src/linear_algebra/clifford_algebra/fold.lean | 61 +++ 3 files changed, 433 insertions(+) create mode 100644 src/linear_algebra/clifford_algebra/contraction.lean diff --git a/docs/references.bib b/docs/references.bib index 135bb7eddf5fe..f52e984a5f582 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -343,6 +343,18 @@ @Book{ bourbaki1987 url = {https://doi.org/10.1007/978-3-642-61715-7} } +@Book{ bourbaki2007, + author = {Bourbaki, Nicolas}, + edition = {Réimpression inchangée de l'édition originale de 1959}, + series = {Eléments de mathématique}, + title = {Algèbre. {C}hapitre IX}, + isbn = {978-3-540-35338-6}, + language = {fr}, + number = {2}, + publisher = {Springer}, + year = {2007} +} + @Book{ boydVandenberghe2004, author = {Stephen P. Boyd and Lieven Vandenberghe}, title = {Convex Optimization}, @@ -931,6 +943,15 @@ @Book{ Gratzer2011 mrnumber = {2768581} } +@Unpublished{ grinberg_clifford_2016, + title = {The {Clifford} algebra and the {Chevalley} map- a + computational approach (summary version 1)}, + url = {http://mit.edu/~darij/www/algebra/chevalleys.pdf}, + author = {Grinberg, D.}, + month = jun, + year = {2016} +} + @Book{ gunter1992, title = {Semantics of Programming Languages: Structures and Techniques}, diff --git a/src/linear_algebra/clifford_algebra/contraction.lean b/src/linear_algebra/clifford_algebra/contraction.lean new file mode 100644 index 0000000000000..893c6246ccefa --- /dev/null +++ b/src/linear_algebra/clifford_algebra/contraction.lean @@ -0,0 +1,351 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ + +import linear_algebra.exterior_algebra.basic +import linear_algebra.clifford_algebra.fold +import linear_algebra.clifford_algebra.grading +import linear_algebra.clifford_algebra.conjugation + +/-! +# Contraction in Clifford Algebras + +This file contains some of the results from [grinberg_clifford_2016][]. +The key result is `clifford_algebra.equiv_exterior`. + +## Main definitions + +* `clifford_algebra.contract_left`: contract a multivector by a `module.dual R M` on the left. +* `clifford_algebra.contract_right`: contract a multivector by a `module.dual R M` on the right. +* `clifford_algebra.change_form`: convert between two algebras of different quadratic form, sending + vectors to vectors. The difference of the quadratic forms must be a bilinear form. +* `clifford_algebra.equiv_exterior`: in characteristic not-two, the `clifford_algebra Q` is + isomorphic as a module to the exterior algebra. + +## Implementation notes + +This file somewhat follows [grinberg_clifford_2016][], although we are missing some of the induction +principles needed to prove many of the results. Here, we avoid the quotient-based approach described +in [grinberg_clifford_2016][], instead directly constructing our objects using the universal +property. + +Note that [grinberg_clifford_2016][] concludes that its contents are not novel, and are in fact just +a rehash of parts of [bourbaki2007][]; we should at some point consider swapping our references to +refer to the latter. + +Within this file, we use the local notation +* `x ⌊ d` for `contract_right x d` +* `d ⌋ x` for `contract_left d x` + +-/ + +universes u1 u2 u3 + +variables {R : Type u1} [comm_ring R] +variables {M : Type u2} [add_comm_group M] [module R M] +variables (Q : quadratic_form R M) + +namespace clifford_algebra + +section contract_left + +variables (d d' : module.dual R M) + +/-- Auxiliary construction for `clifford_algebra.contract_left` -/ +@[simps] +def contract_left_aux (d : module.dual R M) : + M →ₗ[R] clifford_algebra Q × clifford_algebra Q →ₗ[R] clifford_algebra Q := +begin + have v_mul := (algebra.lmul R (clifford_algebra Q)).to_linear_map ∘ₗ (ι Q), + exact d.smul_right (linear_map.fst _ (clifford_algebra Q) (clifford_algebra Q)) - + v_mul.compl₂ (linear_map.snd _ (clifford_algebra Q) _), +end + +lemma contract_left_aux_contract_left_aux (v : M) (x : clifford_algebra Q) + (fx : clifford_algebra Q) : + contract_left_aux Q d v (ι Q v * x, contract_left_aux Q d v (x, fx)) = Q v • fx := +begin + simp only [contract_left_aux_apply_apply], + rw [mul_sub, ←mul_assoc, ι_sq_scalar, ←algebra.smul_def, ←sub_add, mul_smul_comm, sub_self, + zero_add], +end + +variables {Q} + +/-- Contract an element of the clifford algebra with an element `d : module.dual R M` from the left. + +Note that $v ⌋ x$ is spelt `contract_left (Q.associated v) x`. + +This includes [grinberg_clifford_2016][] Theorem 10.75 -/ +def contract_left : module.dual R M →ₗ[R] clifford_algebra Q →ₗ[R] clifford_algebra Q := +{ to_fun := λ d, foldr' Q (contract_left_aux Q d) (contract_left_aux_contract_left_aux Q d) 0, + map_add' := λ d₁ d₂, linear_map.ext $ λ x, begin + rw linear_map.add_apply, + induction x using clifford_algebra.left_induction with r x y hx hy m x hx, + { simp_rw [foldr'_algebra_map, smul_zero, zero_add] }, + { rw [map_add, map_add, map_add, add_add_add_comm, hx, hy] }, + { rw [foldr'_ι_mul, foldr'_ι_mul, foldr'_ι_mul, hx], + dsimp only [contract_left_aux_apply_apply], + rw [sub_add_sub_comm, mul_add, linear_map.add_apply, add_smul] } + end, + map_smul' := λ c d, linear_map.ext $ λ x, begin + rw [linear_map.smul_apply, ring_hom.id_apply], + induction x using clifford_algebra.left_induction with r x y hx hy m x hx, + { simp_rw [foldr'_algebra_map, smul_zero] }, + { rw [map_add, map_add, smul_add, hx, hy] }, + { rw [foldr'_ι_mul, foldr'_ι_mul, hx], + dsimp only [contract_left_aux_apply_apply], + rw [linear_map.smul_apply, smul_assoc, mul_smul_comm, smul_sub], } + end } + +/-- Contract an element of the clifford algebra with an element `d : module.dual R M` from the +right. + +Note that $x ⌊ v$ is spelt `contract_right x (Q.associated v)`. + +This includes [grinberg_clifford_2016][] Theorem 16.75 -/ +def contract_right : clifford_algebra Q →ₗ[R] module.dual R M →ₗ[R] clifford_algebra Q := +linear_map.flip (linear_map.compl₂ (linear_map.compr₂ contract_left reverse) reverse) + +lemma contract_right_eq (x : clifford_algebra Q) : + contract_right x d = reverse (contract_left d $ reverse x) := rfl + +local infix `⌋`:70 := contract_left +local infix `⌊`:70 := contract_right + +/-- This is [grinberg_clifford_2016][] Theorem 6 -/ +lemma contract_left_ι_mul (a : M) (b : clifford_algebra Q) : + d ⌋ (ι Q a * b) = d a • b - ι Q a * (d ⌋ b) := +foldr'_ι_mul _ _ _ _ _ _ + +/-- This is [grinberg_clifford_2016][] Theorem 12 -/ +lemma contract_right_mul_ι (a : M) (b : clifford_algebra Q) : + (b * ι Q a) ⌊ d = d a • b - (b ⌊ d) * ι Q a := +by rw [contract_right_eq, reverse.map_mul, reverse_ι, contract_left_ι_mul, map_sub, map_smul, + reverse_reverse, reverse.map_mul, reverse_ι, contract_right_eq] + +lemma contract_left_algebra_map_mul (r : R) (b : clifford_algebra Q) : + d ⌋ (algebra_map _ _ r * b) = algebra_map _ _ r * (d ⌋ b) := +by rw [←algebra.smul_def, map_smul, algebra.smul_def] + +lemma contract_left_mul_algebra_map (a : clifford_algebra Q) (r : R) : + d ⌋ (a * algebra_map _ _ r) = (d ⌋ a) * algebra_map _ _ r := +by rw [←algebra.commutes, contract_left_algebra_map_mul, algebra.commutes] + +lemma contract_right_algebra_map_mul (r : R) (b : clifford_algebra Q) : + (algebra_map _ _ r * b) ⌊ d = algebra_map _ _ r * (b ⌊ d) := +by rw [←algebra.smul_def, linear_map.map_smul₂, algebra.smul_def] + +lemma contract_right_mul_algebra_map (a : clifford_algebra Q) (r : R) : + (a * algebra_map _ _ r) ⌊ d = (a ⌊ d) * algebra_map _ _ r := +by rw [←algebra.commutes, contract_right_algebra_map_mul, algebra.commutes] + +variables (Q) + +@[simp] lemma contract_left_ι (x : M) : d ⌋ ι Q x = algebra_map R _ (d x) := +(foldr'_ι _ _ _ _ _).trans $ + by simp_rw [contract_left_aux_apply_apply, mul_zero, sub_zero, algebra.algebra_map_eq_smul_one] + +@[simp] lemma contract_right_ι (x : M) : ι Q x ⌊ d = algebra_map R _ (d x) := +by rw [contract_right_eq, reverse_ι, contract_left_ι, reverse.commutes] + +@[simp] lemma contract_left_algebra_map (r : R) : + d ⌋ (algebra_map R (clifford_algebra Q) r) = 0 := +(foldr'_algebra_map _ _ _ _ _).trans $ smul_zero _ + +@[simp] lemma contract_right_algebra_map (r : R) : + (algebra_map R (clifford_algebra Q) r) ⌊ d = 0 := +by rw [contract_right_eq, reverse.commutes, contract_left_algebra_map, map_zero] + +@[simp] lemma contract_left_one : d ⌋ (1 : clifford_algebra Q) = 0 := +by simpa only [map_one] using contract_left_algebra_map Q d 1 + +@[simp] lemma contract_right_one : (1 : clifford_algebra Q) ⌊ d = 0 := +by simpa only [map_one] using contract_right_algebra_map Q d 1 + +variables {Q} + +/-- This is [grinberg_clifford_2016][] Theorem 7 -/ +lemma contract_left_contract_left (x : clifford_algebra Q) : + d ⌋ (d ⌋ x) = 0 := +begin + induction x using clifford_algebra.left_induction with r x y hx hy m x hx, + { simp_rw [contract_left_algebra_map, map_zero] }, + { rw [map_add, map_add, hx, hy, add_zero] }, + { rw [contract_left_ι_mul, map_sub, contract_left_ι_mul, hx, linear_map.map_smul, mul_zero, + sub_zero, sub_self], } +end + +/-- This is [grinberg_clifford_2016][] Theorem 13 -/ +lemma contract_right_contract_right (x : clifford_algebra Q) : + (x ⌊ d) ⌊ d = 0 := +by rw [contract_right_eq, contract_right_eq, reverse_reverse, contract_left_contract_left, + map_zero] + +/-- This is [grinberg_clifford_2016][] Theorem 8 -/ +lemma contract_left_comm (x : clifford_algebra Q) : d ⌋ (d' ⌋ x) = -(d' ⌋ (d ⌋ x)) := +begin + induction x using clifford_algebra.left_induction with r x y hx hy m x hx, + { simp_rw [contract_left_algebra_map, map_zero, neg_zero] }, + { rw [map_add, map_add, map_add, map_add, hx, hy, neg_add] }, + { simp only [contract_left_ι_mul, map_sub, linear_map.map_smul], + rw [neg_sub, sub_sub_eq_add_sub, hx, mul_neg, ←sub_eq_add_neg] } +end + +/-- This is [grinberg_clifford_2016][] Theorem 14 -/ +lemma contract_right_comm (x : clifford_algebra Q) : (x ⌊ d) ⌊ d' = -((x ⌊ d') ⌊ d) := +by rw [contract_right_eq, contract_right_eq, contract_right_eq, contract_right_eq, + reverse_reverse, reverse_reverse, contract_left_comm, map_neg] + +/- TODO: +lemma contract_right_contract_left (x : clifford_algebra Q) : (d ⌋ x) ⌊ d' = d ⌋ (x ⌊ d') := +-/ + +end contract_left + +local infix `⌋`:70 := contract_left +local infix `⌊`:70 := contract_right + +/-- Auxiliary construction for `clifford_algebra.change_form` -/ +@[simps] +def change_form_aux (B : bilin_form R M) : M →ₗ[R] clifford_algebra Q →ₗ[R] clifford_algebra Q := +begin + have v_mul := (algebra.lmul R (clifford_algebra Q)).to_linear_map ∘ₗ ι Q, + exact v_mul - (contract_left ∘ₗ B.to_lin) , +end + +lemma change_form_aux_change_form_aux (B : bilin_form R M) (v : M) (x : clifford_algebra Q) : + change_form_aux Q B v (change_form_aux Q B v x) = (Q v - B v v) • x := +begin + simp only [change_form_aux_apply_apply], + rw [mul_sub, ←mul_assoc, ι_sq_scalar, map_sub, contract_left_ι_mul, ←sub_add, sub_sub_sub_comm, + ←algebra.smul_def, bilin_form.to_lin_apply, sub_self, sub_zero, contract_left_contract_left, + add_zero, sub_smul], +end + +variables {Q} + +variables {Q' Q'' : quadratic_form R M} {B B' : bilin_form R M} +variables (h : B.to_quadratic_form = Q' - Q) (h' : B'.to_quadratic_form = Q'' - Q') + +/-- Convert between two algebras of different quadratic form, sending vector to vectors, scalars to +scalars, and adjusting products by a contraction term. + +This is $\lambda_B$ from [bourbaki2007][] $9 Lemma 2. -/ +def change_form (h : B.to_quadratic_form = Q' - Q) : + clifford_algebra Q →ₗ[R] clifford_algebra Q' := +foldr Q (change_form_aux Q' B) (λ m x, (change_form_aux_change_form_aux Q' B m x).trans $ + begin + dsimp [←bilin_form.to_quadratic_form_apply], + rw [h, quadratic_form.sub_apply, sub_sub_cancel], + end) 1 + +/-- Auxiliary lemma used as an argument to `clifford_algebra.change_form` -/ +lemma change_form.zero_proof : (0 : bilin_form R M).to_quadratic_form = Q - Q := +(sub_self _).symm + +/-- Auxiliary lemma used as an argument to `clifford_algebra.change_form` -/ +lemma change_form.add_proof : (B + B').to_quadratic_form = Q'' - Q := +(congr_arg2 (+) h h').trans $ sub_add_sub_cancel' _ _ _ + +/-- Auxiliary lemma used as an argument to `clifford_algebra.change_form` -/ +lemma change_form.neg_proof : (-B).to_quadratic_form = Q - Q' := +(congr_arg has_neg.neg h).trans $ neg_sub _ _ + +lemma change_form.associated_neg_proof [invertible (2 : R)] : + (-Q).associated.to_quadratic_form = 0 - Q := +by simp [quadratic_form.to_quadratic_form_associated] + +@[simp] +lemma change_form_algebra_map (r : R) : change_form h (algebra_map R _ r) = algebra_map R _ r := +(foldr_algebra_map _ _ _ _ _).trans $ eq.symm $ algebra.algebra_map_eq_smul_one r + +@[simp] lemma change_form_one : change_form h (1 : clifford_algebra Q) = 1 := +by simpa using change_form_algebra_map h (1 : R) + +@[simp] +lemma change_form_ι (m : M) : change_form h (ι _ m) = ι _ m := +(foldr_ι _ _ _ _ _).trans $ eq.symm $ + by rw [change_form_aux_apply_apply, mul_one, contract_left_one, sub_zero] + +lemma change_form_ι_mul (m : M) (x : clifford_algebra Q) : + change_form h (ι _ m * x) = ι _ m * change_form h x - bilin_form.to_lin B m ⌋ change_form h x := +(foldr_mul _ _ _ _ _ _).trans $ begin rw foldr_ι, refl, end + +lemma change_form_ι_mul_ι (m₁ m₂ : M) : + change_form h (ι _ m₁ * ι _ m₂) = ι _ m₁ * ι _ m₂ - algebra_map _ _ (B m₁ m₂) := +by rw [change_form_ι_mul, change_form_ι, contract_left_ι, bilin_form.to_lin_apply] + +/-- Theorem 23 of [grinberg_clifford_2016][] -/ +lemma change_form_contract_left (d : module.dual R M) (x : clifford_algebra Q) : + change_form h (d ⌋ x) = d ⌋ change_form h x := +begin + induction x using clifford_algebra.left_induction with r x y hx hy m x hx, + { simp only [contract_left_algebra_map, change_form_algebra_map, map_zero] }, + { rw [map_add, map_add, map_add, map_add, hx, hy] }, + { simp only [contract_left_ι_mul, change_form_ι_mul, map_sub, linear_map.map_smul], + rw [←hx, contract_left_comm, ←sub_add, sub_neg_eq_add, ←hx] } +end + +lemma change_form_self_apply (x : clifford_algebra Q) : + change_form (change_form.zero_proof) x = x := +begin + induction x using clifford_algebra.left_induction with r x y hx hy m x hx, + { simp_rw [change_form_algebra_map] }, + { rw [map_add, hx, hy] }, + { rw [change_form_ι_mul, hx, map_zero, linear_map.zero_apply, map_zero, linear_map.zero_apply, + sub_zero] } +end + +@[simp] +lemma change_form_self : + change_form change_form.zero_proof = (linear_map.id : clifford_algebra Q →ₗ[R] _) := +linear_map.ext $ change_form_self_apply + +/-- This is [bourbaki2007][] $9 Lemma 3. -/ +lemma change_form_change_form (x : clifford_algebra Q) : + change_form h' (change_form h x) = change_form (change_form.add_proof h h') x := +begin + induction x using clifford_algebra.left_induction with r x y hx hy m x hx, + { simp_rw [change_form_algebra_map] }, + { rw [map_add, map_add, map_add, hx, hy] }, + { rw [change_form_ι_mul, map_sub, change_form_ι_mul, change_form_ι_mul, hx, sub_sub, map_add, + linear_map.add_apply, map_add, linear_map.add_apply, change_form_contract_left, hx, + add_comm (_ : clifford_algebra Q'')] } +end + +lemma change_form_comp_change_form : + (change_form h').comp (change_form h) = change_form (change_form.add_proof h h') := +linear_map.ext $ change_form_change_form _ _ + +/-- Any two algebras whose quadratic forms differ by a bilinear form are isomorphic as modules. + +This is $\bar \lambda_B$ from [bourbaki2007][] $9 Proposition 3. -/ +@[simps apply] +def change_form_equiv : clifford_algebra Q ≃ₗ[R] clifford_algebra Q' := +{ to_fun := change_form h, + inv_fun := change_form (change_form.neg_proof h), + left_inv := λ x, (change_form_change_form _ _ x).trans $ + by simp_rw [add_right_neg, change_form_self_apply], + right_inv := λ x, (change_form_change_form _ _ x).trans $ + by simp_rw [add_left_neg, change_form_self_apply], + ..change_form h } + +@[simp] +lemma change_form_equiv_symm : + (change_form_equiv h).symm = change_form_equiv (change_form.neg_proof h) := +linear_equiv.ext $ λ x, (rfl : change_form _ x = change_form _ x) + +variables (Q) + +/-- The module isomorphism to the exterior algebra. + +Note that this holds more generally when `Q` is divisible by two, rather than only when `1` is +divisible by two; but that would be more awkward to use. -/ +@[simp] +def equiv_exterior [invertible (2 : R)] : clifford_algebra Q ≃ₗ[R] exterior_algebra R M := +change_form_equiv change_form.associated_neg_proof + +end clifford_algebra diff --git a/src/linear_algebra/clifford_algebra/fold.lean b/src/linear_algebra/clifford_algebra/fold.lean index 86d1432eb9908..625219a22945a 100644 --- a/src/linear_algebra/clifford_algebra/fold.lean +++ b/src/linear_algebra/clifford_algebra/fold.lean @@ -145,4 +145,65 @@ begin { simpa only [reverse.map_mul, reverse_ι] using h_mul_ι _ _ hx }, end +/-! ### Versions with extra state -/ +/-- Auxiliary definition for `clifford_algebra.foldr'` -/ +def foldr'_aux (f : M →ₗ[R] clifford_algebra Q × N →ₗ[R] N) : + M →ₗ[R] module.End R (clifford_algebra Q × N) := +begin + have v_mul := (algebra.lmul R (clifford_algebra Q)).to_linear_map ∘ₗ (ι Q), + have l := v_mul.compl₂ (linear_map.fst _ _ N), + exact { to_fun := λ m, (l m).prod (f m), + map_add' := λ v₂ v₂, linear_map.ext $ λ x, prod.ext + (linear_map.congr_fun (l.map_add _ _) x) (linear_map.congr_fun (f.map_add _ _) x), + map_smul' := λ c v, linear_map.ext $ λ x, prod.ext + (linear_map.congr_fun (l.map_smul _ _) x) (linear_map.congr_fun (f.map_smul _ _) x), }, +end + +lemma foldr'_aux_apply_apply (f : M →ₗ[R] clifford_algebra Q × N →ₗ[R] N) (m : M) (x_fx) : + foldr'_aux Q f m x_fx = (ι Q m * x_fx.1, f m x_fx) := rfl + +lemma foldr'_aux_foldr'_aux (f : M →ₗ[R] clifford_algebra Q × N →ₗ[R] N) + (hf : ∀ m x fx, f m (ι Q m * x, f m (x, fx)) = Q m • fx) + (v : M) (x_fx) : + foldr'_aux Q f v (foldr'_aux Q f v x_fx) = Q v • x_fx := +begin + cases x_fx with x fx, + simp only [foldr'_aux_apply_apply], + rw [←mul_assoc, ι_sq_scalar, ← algebra.smul_def, hf, prod.smul_mk], +end + +/-- Fold a bilinear map along the generators of a term of the clifford algebra, with the rule +given by `foldr' Q f hf n (ι Q m * x) = f m (x, foldr' Q f hf n x)`. +Note this is like `clifford_algebra.foldr`, but with an extra `x` argument. +Implement the recursion scheme `F[n0](m * x) = f(m, (x, F[n0](x)))`. -/ +def foldr' (f : M →ₗ[R] clifford_algebra Q × N →ₗ[R] N) + (hf : ∀ m x fx, f m (ι Q m * x, f m (x, fx)) = Q m • fx) + (n : N) : + clifford_algebra Q →ₗ[R] N := +linear_map.snd _ _ _ ∘ₗ foldr Q (foldr'_aux Q f) (foldr'_aux_foldr'_aux Q _ hf) (1, n) + +lemma foldr'_algebra_map (f : M →ₗ[R] clifford_algebra Q × N →ₗ[R] N) + (hf : ∀ m x fx, f m (ι Q m * x, f m (x, fx)) = Q m • fx) (n r) : + foldr' Q f hf n (algebra_map R _ r) = r • n := +congr_arg prod.snd (foldr_algebra_map _ _ _ _ _) + +lemma foldr'_ι (f : M →ₗ[R] clifford_algebra Q × N →ₗ[R] N) + (hf : ∀ m x fx, f m (ι Q m * x, f m (x, fx)) = Q m • fx) (n m) : + foldr' Q f hf n (ι Q m) = f m (1, n) := +congr_arg prod.snd (foldr_ι _ _ _ _ _) + +lemma foldr'_ι_mul (f : M →ₗ[R] clifford_algebra Q × N →ₗ[R] N) + (hf : ∀ m x fx, f m (ι Q m * x, f m (x, fx)) = Q m • fx) (n m) (x) : + foldr' Q f hf n (ι Q m * x) = f m (x, foldr' Q f hf n x) := +begin + dsimp [foldr'], + rw [foldr_mul, foldr_ι, foldr'_aux_apply_apply], + refine congr_arg (f m) (prod.mk.eta.symm.trans _), + congr' 1, + induction x using clifford_algebra.left_induction with r x y hx hy m x hx, + { simp_rw [foldr_algebra_map, prod.smul_mk, algebra.algebra_map_eq_smul_one] }, + { rw [map_add, prod.fst_add, hx, hy] }, + { rw [foldr_mul, foldr_ι, foldr'_aux_apply_apply, hx], }, +end + end clifford_algebra From d0259b01c82eed3f50390a60404c63faf9e60b1f Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 8 Oct 2022 20:02:41 +0000 Subject: [PATCH 635/828] refactor(ring_theory/class_group): redefine class_group without fraction field (#16727) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Although the definition of the class group of a ring `R` involves the field of fractions, the definition does not depend (up to isomorphism) on the choice of field of fractions. This PR proposes to always choose `fraction_ring R` as that field, so that we don't need to carry around the mathematically unnecessary parameters `(K) [field K] [algebra R K] [is_fraction_ring R K]`. Instead, we insert the isomorphism between the definitions of class group at the API boundaries. This was inspired by our work on quadratic rings: it gets even more annoying when you need to start carrying around a proof that the field of fractions of `ℤ[1/2√-7]` is `ℚ(√-7)`. Co-authored-by: Anne Baanen Co-authored-by: Anne Baanen --- src/number_theory/class_number/finite.lean | 27 +- .../class_number/function_field.lean | 4 +- .../number_field/class_number.lean | 6 +- src/ring_theory/class_group.lean | 248 +++++++++++++----- 4 files changed, 196 insertions(+), 89 deletions(-) diff --git a/src/number_theory/class_number/finite.lean b/src/number_theory/class_number/finite.lean index 332cfe561384c..8406e7cdc4ba9 100644 --- a/src/number_theory/class_number/finite.lean +++ b/src/number_theory/class_number/finite.lean @@ -292,9 +292,8 @@ include ist iic /-- Each class in the class group contains an ideal `J` such that `M := Π m ∈ finset_approx` is in `J`. -/ -theorem exists_mk0_eq_mk0 [is_dedekind_domain S] [is_fraction_ring S L] - (h : algebra.is_algebraic R L) (I : (ideal S)⁰) : - ∃ (J : (ideal S)⁰), class_group.mk0 L I = class_group.mk0 L J ∧ +theorem exists_mk0_eq_mk0 [is_dedekind_domain S] (h : algebra.is_algebraic R L) (I : (ideal S)⁰) : + ∃ (J : (ideal S)⁰), class_group.mk0 I = class_group.mk0 J ∧ algebra_map _ _ (∏ m in finset_approx bS adm, m) ∈ (J : ideal S) := begin set M := ∏ m in finset_approx bS adm, m with M_eq, @@ -332,17 +331,16 @@ omit iic ist /-- `class_group.mk_M_mem` is a specialization of `class_group.mk0` to (the finite set of) ideals that contain `M := ∏ m in finset_approx L f abs, m`. By showing this function is surjective, we prove that the class group is finite. -/ -noncomputable def mk_M_mem [is_fraction_ring S L] [is_dedekind_domain S] +noncomputable def mk_M_mem [is_dedekind_domain S] (J : {J : ideal S // algebra_map _ _ (∏ m in finset_approx bS adm, m) ∈ J}) : - class_group S L := -class_group.mk0 _ ⟨J.1, mem_non_zero_divisors_iff_ne_zero.mpr + class_group S := +class_group.mk0 ⟨J.1, mem_non_zero_divisors_iff_ne_zero.mpr (ne_bot_of_prod_finset_approx_mem bS adm J.1 J.2)⟩ include iic ist -lemma mk_M_mem_surjective [is_fraction_ring S L] [is_dedekind_domain S] - (h : algebra.is_algebraic R L) : - function.surjective (class_group.mk_M_mem L bS adm) := +lemma mk_M_mem_surjective [is_dedekind_domain S] (h : algebra.is_algebraic R L) : + function.surjective (class_group.mk_M_mem bS adm) := begin intro I', obtain ⟨⟨I, hI⟩, rfl⟩ := class_group.mk0_surjective I', @@ -358,8 +356,8 @@ algebraic extension `L` is finite if there is an admissible absolute value. See also `class_group.fintype_of_admissible_of_finite` where `L` is a finite extension of `K = Frac(R)`, supplying most of the required assumptions automatically. -/ -noncomputable def fintype_of_admissible_of_algebraic [is_fraction_ring S L] [is_dedekind_domain S] - (h : algebra.is_algebraic R L) : fintype (class_group S L) := +noncomputable def fintype_of_admissible_of_algebraic [is_dedekind_domain S] + (h : algebra.is_algebraic R L) : fintype (class_group S) := @fintype.of_surjective _ _ _ (@fintype.of_equiv _ {J // J ∣ ideal.span ({algebra_map R S (∏ (m : R) in finset_approx bS adm, m)} : set S)} @@ -368,9 +366,11 @@ noncomputable def fintype_of_admissible_of_algebraic [is_fraction_ring S L] [is_ exact prod_finset_approx_ne_zero bS adm })) ((equiv.refl _).subtype_equiv (λ I, ideal.dvd_iff_le.trans (by rw [equiv.refl_apply, ideal.span_le, set.singleton_subset_iff])))) - (class_group.mk_M_mem L bS adm) + (class_group.mk_M_mem bS adm) (class_group.mk_M_mem_surjective L bS adm h) +include K + /-- The main theorem: the class group of an integral closure `S` of `R` in a finite extension `L` of `K = Frac(R)` is finite if there is an admissible absolute value. @@ -379,8 +379,7 @@ See also `class_group.fintype_of_admissible_of_algebraic` where `L` is an algebraic extension of `R`, that includes some extra assumptions. -/ noncomputable def fintype_of_admissible_of_finite [is_dedekind_domain R] : - fintype (@class_group S L _ _ _ - (is_integral_closure.is_fraction_ring_of_finite_extension R K L S)) := + fintype (class_group S) := begin letI := classical.dec_eq L, letI := is_integral_closure.is_fraction_ring_of_finite_extension R K L S, diff --git a/src/number_theory/class_number/function_field.lean b/src/number_theory/class_number/function_field.lean index 1ee867b982a75..b56a75a38d58a 100644 --- a/src/number_theory/class_number/function_field.lean +++ b/src/number_theory/class_number/function_field.lean @@ -33,7 +33,7 @@ namespace ring_of_integers open function_field -noncomputable instance : fintype (class_group (ring_of_integers Fq F) F) := +noncomputable instance : fintype (class_group (ring_of_integers Fq F)) := class_group.fintype_of_admissible_of_finite (ratfunc Fq) F (polynomial.card_pow_degree_is_admissible : absolute_value.is_admissible (polynomial.card_pow_degree : absolute_value Fq[X] ℤ)) @@ -41,7 +41,7 @@ class_group.fintype_of_admissible_of_finite (ratfunc Fq) F end ring_of_integers /-- The class number in a function field is the (finite) cardinality of the class group. -/ -noncomputable def class_number : ℕ := fintype.card (class_group (ring_of_integers Fq F) F) +noncomputable def class_number : ℕ := fintype.card (class_group (ring_of_integers Fq F)) /-- The class number of a function field is `1` iff the ring of integers is a PID. -/ theorem class_number_eq_one_iff : diff --git a/src/number_theory/number_field/class_number.lean b/src/number_theory/number_field/class_number.lean index 4b8bc9a7b94df..9b1bb64056c76 100644 --- a/src/number_theory/number_field/class_number.lean +++ b/src/number_theory/number_field/class_number.lean @@ -25,13 +25,13 @@ variables (K : Type*) [field K] [number_field K] namespace ring_of_integers -noncomputable instance : fintype (class_group (ring_of_integers K) K) := -class_group.fintype_of_admissible_of_finite ℚ _ absolute_value.abs_is_admissible +noncomputable instance : fintype (class_group (ring_of_integers K)) := +class_group.fintype_of_admissible_of_finite ℚ K absolute_value.abs_is_admissible end ring_of_integers /-- The class number of a number field is the (finite) cardinality of the class group. -/ -noncomputable def class_number : ℕ := fintype.card (class_group (ring_of_integers K) K) +noncomputable def class_number : ℕ := fintype.card (class_group (ring_of_integers K)) variables {K} diff --git a/src/ring_theory/class_group.lean b/src/ring_theory/class_group.lean index 47490ebbb5fbb..59466077dfe9a 100644 --- a/src/ring_theory/class_group.lean +++ b/src/ring_theory/class_group.lean @@ -10,8 +10,8 @@ import ring_theory.dedekind_domain.ideal /-! # The ideal class group -This file defines the ideal class group `class_group R K` of fractional ideals of `R` -inside `A`'s field of fractions `K`. +This file defines the ideal class group `class_group R` of fractional ideals of `R` +inside its field of fractions. ## Main definitions - `to_principal_ideal` sends an invertible `x : K` to an invertible fractional ideal @@ -21,6 +21,11 @@ inside `A`'s field of fractions `K`. ## Main results - `class_group.mk0_eq_mk0_iff` shows the equivalence with the "classical" definition, where `I ~ J` iff `x I = y J` for `x y ≠ (0 : R)` + +## Implementation details + +The definition of `class_group R` involves `fraction_ring R`. However, the API should be completely +identical no matter the choice of field of fractions for `R`. -/ variables {R K L : Type*} [comm_ring R] @@ -61,26 +66,97 @@ rfl to_principal_ideal R K x = I ↔ span_singleton R⁰ (x : K) = I := units.ext_iff +lemma mem_principal_ideals_iff {I : (fractional_ideal R⁰ K)ˣ} : + I ∈ (to_principal_ideal R K).range ↔ ∃ x : K, span_singleton R⁰ x = I := +begin + simp only [monoid_hom.mem_range, to_principal_ideal_eq_iff], + split; rintros ⟨x, hx⟩, + { exact ⟨x, hx⟩ }, + { refine ⟨units.mk0 x _, hx⟩, + rintro rfl, + simpa [I.ne_zero.symm] using hx }, + end instance principal_ideals.normal : (to_principal_ideal R K).range.normal := subgroup.normal_of_comm _ +end -section - -variables (R K) +variables (R) [is_domain R] -/-- The ideal class group of `R` in a field of fractions `K` -is the group of invertible fractional ideals modulo the principal ideals. -/ +/-- The ideal class group of `R` is the group of invertible fractional ideals +modulo the principal ideals. -/ @[derive(comm_group)] -def class_group := (fractional_ideal R⁰ K)ˣ ⧸ (to_principal_ideal R K).range +def class_group := +(fractional_ideal R⁰ (fraction_ring R))ˣ ⧸ (to_principal_ideal R (fraction_ring R)).range + +noncomputable instance : inhabited (class_group R) := ⟨1⟩ -instance : inhabited (class_group R K) := ⟨1⟩ +variables {R K} + +/-- Send a nonzero fractional ideal to the corresponding class in the class group. -/ +noncomputable def class_group.mk : (fractional_ideal R⁰ K)ˣ →* class_group R := +(quotient_group.mk' (to_principal_ideal R (fraction_ring R)).range).comp + (units.map (fractional_ideal.canonical_equiv R⁰ K (fraction_ring R))) + +variables (K) + +/-- Induction principle for the class group: to show something holds for all `x : class_group R`, +we can choose a fraction field `K` and show it holds for the equivalence class of each +`I : fractional_ideal R⁰ K`. -/ +@[elab_as_eliminator] lemma class_group.induction {P : class_group R → Prop} + (h : ∀ (I : (fractional_ideal R⁰ K)ˣ), P (class_group.mk I)) (x : class_group R) : P x := +quotient_group.induction_on x (λ I, begin + convert h (units.map_equiv ↑(canonical_equiv R⁰ (fraction_ring R) K) I), + ext : 1, + rw [units.coe_map, units.coe_map_equiv], + exact (canonical_equiv_flip R⁰ K (fraction_ring R) I).symm +end) + +/-- The definition of the class group does not depend on the choice of field of fractions. -/ +noncomputable def class_group.equiv : + class_group R ≃* (fractional_ideal R⁰ K)ˣ ⧸ (to_principal_ideal R K).range := +quotient_group.congr _ _ + (units.map_equiv (fractional_ideal.canonical_equiv R⁰ (fraction_ring R) K : + fractional_ideal R⁰ (fraction_ring R) ≃* fractional_ideal R⁰ K)) $ +begin + ext I, + simp only [subgroup.mem_map, mem_principal_ideals_iff, monoid_hom.coe_coe], + split, + { rintro ⟨I, ⟨x, hx⟩, rfl⟩, + refine ⟨fraction_ring.alg_equiv R K x, _⟩, + rw [units.coe_map_equiv, ← hx, ring_equiv.coe_to_mul_equiv, canonical_equiv_span_singleton], + refl }, + { rintro ⟨x, hx⟩, + refine ⟨units.map_equiv ↑(canonical_equiv R⁰ K (fraction_ring R)) I, + ⟨(fraction_ring.alg_equiv R K).symm x, _⟩, + units.ext _⟩, + { rw [units.coe_map_equiv, ← hx, ring_equiv.coe_to_mul_equiv, canonical_equiv_span_singleton], + refl }, + simp only [ring_equiv.coe_to_mul_equiv, canonical_equiv_flip, units.coe_map_equiv] }, +end + +@[simp] lemma class_group.equiv_mk (K' : Type*) [field K'] [algebra R K'] [is_fraction_ring R K'] + (I : (fractional_ideal R⁰ K)ˣ) : + class_group.equiv K' (class_group.mk I) = + quotient_group.mk' _ (units.map_equiv ↑(fractional_ideal.canonical_equiv R⁰ K K') I) := +begin + rw [class_group.equiv, class_group.mk, monoid_hom.comp_apply, quotient_group.congr_mk'], + congr, + ext : 1, + rw [units.coe_map_equiv, units.coe_map_equiv, units.coe_map], + exact fractional_ideal.canonical_equiv_canonical_equiv _ _ _ _ _ +end -variables {R} [is_domain R] +@[simp] lemma class_group.mk_canonical_equiv (K' : Type*) [field K'] [algebra R K'] + [is_fraction_ring R K'] (I : (fractional_ideal R⁰ K)ˣ) : + class_group.mk (units.map ↑(canonical_equiv R⁰ K K') I : (fractional_ideal R⁰ K')ˣ) = + class_group.mk I := +by rw [class_group.mk, monoid_hom.comp_apply, ← monoid_hom.comp_apply (units.map _), + ← units.map_comp, ← ring_equiv.coe_monoid_hom_trans, + fractional_ideal.canonical_equiv_trans_canonical_equiv]; refl /-- Send a nonzero integral ideal to an invertible fractional ideal. -/ -@[simps] noncomputable def fractional_ideal.mk0 [is_dedekind_domain R] : (ideal R)⁰ →* (fractional_ideal R⁰ K)ˣ := { to_fun := λ I, units.mk0 I ((fractional_ideal.coe_to_fractional_ideal_ne_zero (le_refl R⁰)).mpr @@ -88,107 +164,138 @@ noncomputable def fractional_ideal.mk0 [is_dedekind_domain R] : map_one' := by simp, map_mul' := λ x y, by simp } +@[simp] lemma fractional_ideal.coe_mk0 [is_dedekind_domain R] (I : (ideal R)⁰) : + (fractional_ideal.mk0 K I : fractional_ideal R⁰ K) = I := +rfl + +lemma fractional_ideal.canonical_equiv_mk0 [is_dedekind_domain R] + (K' : Type*) [field K'] [algebra R K'] [is_fraction_ring R K'] (I : (ideal R)⁰) : + fractional_ideal.canonical_equiv R⁰ K K' (fractional_ideal.mk0 K I) = + fractional_ideal.mk0 K' I := +by simp only [fractional_ideal.coe_mk0, coe_coe, fractional_ideal.canonical_equiv_coe_ideal] + +@[simp] lemma fractional_ideal.map_canonical_equiv_mk0 [is_dedekind_domain R] + (K' : Type*) [field K'] [algebra R K'] [is_fraction_ring R K'] (I : (ideal R)⁰) : + units.map ↑(fractional_ideal.canonical_equiv R⁰ K K') (fractional_ideal.mk0 K I) = + fractional_ideal.mk0 K' I := +units.ext (fractional_ideal.canonical_equiv_mk0 K K' I) + /-- Send a nonzero ideal to the corresponding class in the class group. -/ noncomputable def class_group.mk0 [is_dedekind_domain R] : - (ideal R)⁰ →* class_group R K := -(quotient_group.mk' _).comp (fractional_ideal.mk0 K) - -variables {K} + (ideal R)⁰ →* class_group R := +class_group.mk.comp (fractional_ideal.mk0 (fraction_ring R)) + +@[simp] lemma class_group.mk_mk0 [is_dedekind_domain R] (I : (ideal R)⁰): + class_group.mk (fractional_ideal.mk0 K I) = class_group.mk0 I := +by rw [class_group.mk0, monoid_hom.comp_apply, + ← class_group.mk_canonical_equiv K (fraction_ring R), + fractional_ideal.map_canonical_equiv_mk0] + +@[simp] lemma class_group.equiv_mk0 [is_dedekind_domain R] (I : (ideal R)⁰): + class_group.equiv K (class_group.mk0 I) = + quotient_group.mk' (to_principal_ideal R K).range (fractional_ideal.mk0 K I) := +begin + rw [class_group.mk0, monoid_hom.comp_apply, class_group.equiv_mk], + congr, + ext, + simp +end lemma class_group.mk0_eq_mk0_iff_exists_fraction_ring [is_dedekind_domain R] {I J : (ideal R)⁰} : - class_group.mk0 K I = class_group.mk0 K J ↔ + class_group.mk0 I = class_group.mk0 J ↔ ∃ (x ≠ (0 : K)), span_singleton R⁰ x * I = J := begin - simp only [class_group.mk0, monoid_hom.comp_apply, quotient_group.mk'_eq_mk'], + refine (class_group.equiv K).injective.eq_iff.symm.trans _, + simp only [class_group.equiv_mk0, quotient_group.mk'_eq_mk', mem_principal_ideals_iff, + coe_coe, units.ext_iff, units.coe_mul, fractional_ideal.coe_mk0, exists_prop], split, - { rintros ⟨_, ⟨x, rfl⟩, hx⟩, - refine ⟨x, x.ne_zero, _⟩, - simpa only [mul_comm, coe_mk0, monoid_hom.to_fun_eq_coe, coe_to_principal_ideal, units.coe_mul] - using congr_arg (coe : _ → fractional_ideal R⁰ K) hx }, + { rintros ⟨X, ⟨x, hX⟩, hx⟩, + refine ⟨x, _, _⟩, + { rintro rfl, simpa [X.ne_zero.symm] using hX }, + simpa only [hX, mul_comm] using hx }, { rintros ⟨x, hx, eq_J⟩, - refine ⟨_, ⟨units.mk0 x hx, rfl⟩, units.ext _⟩, - simpa only [fractional_ideal.mk0_apply, units.coe_mk0, mul_comm, coe_to_principal_ideal, - coe_coe, units.coe_mul] using eq_J } + refine ⟨units.mk0 _ (span_singleton_ne_zero_iff.mpr hx), ⟨x, rfl⟩, _⟩, + simpa only [mul_comm] using eq_J } end +variables {K} + lemma class_group.mk0_eq_mk0_iff [is_dedekind_domain R] {I J : (ideal R)⁰} : - class_group.mk0 K I = class_group.mk0 K J ↔ + class_group.mk0 I = class_group.mk0 J ↔ ∃ (x y : R) (hx : x ≠ 0) (hy : y ≠ 0), ideal.span {x} * (I : ideal R) = ideal.span {y} * J := begin - refine class_group.mk0_eq_mk0_iff_exists_fraction_ring.trans ⟨_, _⟩, + refine (class_group.mk0_eq_mk0_iff_exists_fraction_ring (fraction_ring R)).trans ⟨_, _⟩, { rintros ⟨z, hz, h⟩, obtain ⟨x, ⟨y, hy⟩, rfl⟩ := is_localization.mk'_surjective R⁰ z, refine ⟨x, y, _, mem_non_zero_divisors_iff_ne_zero.mp hy, _⟩, { rintro hx, apply hz, - rw [hx, is_fraction_ring.mk'_eq_div, (algebra_map R K).map_zero, zero_div] }, - { exact (fractional_ideal.mk'_mul_coe_ideal_eq_coe_ideal K hy).mp h } }, + rw [hx, is_fraction_ring.mk'_eq_div, _root_.map_zero, zero_div] }, + { exact (fractional_ideal.mk'_mul_coe_ideal_eq_coe_ideal _ hy).mp h } }, { rintros ⟨x, y, hx, hy, h⟩, have hy' : y ∈ R⁰ := mem_non_zero_divisors_iff_ne_zero.mpr hy, - refine ⟨is_localization.mk' K x ⟨y, hy'⟩, _, _⟩, + refine ⟨is_localization.mk' _ x ⟨y, hy'⟩, _, _⟩, { contrapose! hx, - rwa [is_localization.mk'_eq_iff_eq_mul, zero_mul, ← (algebra_map R K).map_zero, - (is_fraction_ring.injective R K).eq_iff] at hx }, - { exact (fractional_ideal.mk'_mul_coe_ideal_eq_coe_ideal K hy').mpr h } }, + rwa [mk'_eq_iff_eq_mul, zero_mul, ← (algebra_map R (fraction_ring R)).map_zero, + (is_fraction_ring.injective R (fraction_ring R)).eq_iff] + at hx }, + { exact (fractional_ideal.mk'_mul_coe_ideal_eq_coe_ideal _ hy').mpr h } }, end lemma class_group.mk0_surjective [is_dedekind_domain R] : - function.surjective (class_group.mk0 K : (ideal R)⁰ → class_group R K) := + function.surjective (class_group.mk0 : (ideal R)⁰ → class_group R) := begin rintros ⟨I⟩, obtain ⟨a, a_ne_zero', ha⟩ := I.1.2, have a_ne_zero := mem_non_zero_divisors_iff_ne_zero.mp a_ne_zero', - have fa_ne_zero : (algebra_map R K) a ≠ 0 := + have fa_ne_zero : (algebra_map R (fraction_ring R)) a ≠ 0 := is_fraction_ring.to_map_ne_zero_of_mem_non_zero_divisors a_ne_zero', - refine ⟨⟨{ carrier := { x | (algebra_map R K a)⁻¹ * algebra_map R K x ∈ I.1 }, .. }, _⟩, _⟩, + refine ⟨⟨{ carrier := { x | (algebra_map R _ a)⁻¹ * algebra_map R _ x ∈ I.1 }, .. }, _⟩, _⟩, { simp only [ring_hom.map_add, set.mem_set_of_eq, mul_zero, ring_hom.map_mul, mul_add], exact λ _ _ ha hb, submodule.add_mem I ha hb }, { simp only [ring_hom.map_zero, set.mem_set_of_eq, mul_zero, ring_hom.map_mul], exact submodule.zero_mem I }, { intros c _ hb, simp only [smul_eq_mul, set.mem_set_of_eq, mul_zero, ring_hom.map_mul, mul_add, - mul_left_comm ((algebra_map R K) a)⁻¹], + mul_left_comm ((algebra_map R (fraction_ring R)) a)⁻¹], rw ← algebra.smul_def c, exact submodule.smul_mem I c hb }, { rw [mem_non_zero_divisors_iff_ne_zero, submodule.zero_eq_bot, submodule.ne_bot_iff], obtain ⟨x, x_ne, x_mem⟩ := exists_ne_zero_mem_is_integer I.ne_zero, refine ⟨a * x, _, mul_ne_zero a_ne_zero x_ne⟩, - change ((algebra_map R K) a)⁻¹ * (algebra_map R K) (a * x) ∈ I.1, + change ((algebra_map R _) a)⁻¹ * (algebra_map R _) (a * x) ∈ I.1, rwa [ring_hom.map_mul, ← mul_assoc, inv_mul_cancel fa_ne_zero, one_mul] }, { symmetry, apply quotient.sound, change setoid.r _ _, rw quotient_group.left_rel_apply, - refine ⟨units.mk0 (algebra_map R K a) fa_ne_zero, _⟩, + refine ⟨units.mk0 (algebra_map R _ a) fa_ne_zero, _⟩, apply @mul_left_cancel _ _ I, rw [← mul_assoc, mul_right_inv, one_mul, eq_comm, mul_comm I], apply units.ext, - simp only [monoid_hom.coe_mk, subtype.coe_mk, ring_hom.map_mul, coe_coe, - units.coe_mul, coe_to_principal_ideal, coe_mk0, - fractional_ideal.eq_span_singleton_mul], + simp only [fractional_ideal.coe_mk0, fractional_ideal.map_canonical_equiv_mk0, set_like.coe_mk, + units.coe_mk0, coe_to_principal_ideal, coe_coe, units.coe_mul, + fractional_ideal.eq_span_singleton_mul], split, { intros zJ' hzJ', - obtain ⟨zJ, hzJ : (algebra_map R K a)⁻¹ * algebra_map R K zJ ∈ ↑I, rfl⟩ := + obtain ⟨zJ, hzJ : (algebra_map R _ a)⁻¹ * algebra_map R _ zJ ∈ ↑I, rfl⟩ := (mem_coe_ideal R⁰).mp hzJ', refine ⟨_, hzJ, _⟩, rw [← mul_assoc, mul_inv_cancel fa_ne_zero, one_mul] }, { intros zI' hzI', obtain ⟨y, hy⟩ := ha zI' hzI', - rw [← algebra.smul_def, fractional_ideal.mk0_apply, coe_mk0, coe_coe, mem_coe_ideal], + rw [← algebra.smul_def, mem_coe_ideal], refine ⟨y, _, hy⟩, - show (algebra_map R K a)⁻¹ * algebra_map R K y ∈ (I : fractional_ideal R⁰ K), + show (algebra_map R _ a)⁻¹ * algebra_map R _ y ∈ (I : fractional_ideal R⁰ (fraction_ring R)), rwa [hy, algebra.smul_def, ← mul_assoc, inv_mul_cancel fa_ne_zero, one_mul] } } end -end - -lemma class_group.mk_eq_one_iff - {I : (fractional_ideal R⁰ K)ˣ} : - quotient_group.mk' (to_principal_ideal R K).range I = 1 ↔ - (I : submodule R K).is_principal := +lemma class_group.mk_eq_one_iff {I : (fractional_ideal R⁰ K)ˣ} : + class_group.mk I = 1 ↔ (I : submodule R K).is_principal := begin - rw [← (quotient_group.mk' _).map_one, eq_comm, quotient_group.mk'_eq_mk'], - simp only [exists_prop, one_mul, exists_eq_right, to_principal_ideal_eq_iff, - monoid_hom.mem_range, coe_coe], + simp only [← (class_group.equiv K).injective.eq_iff, _root_.map_one, class_group.equiv_mk, + quotient_group.mk'_apply, quotient_group.eq_one_iff, monoid_hom.mem_range, units.ext_iff, + coe_to_principal_ideal, units.coe_map_equiv, fractional_ideal.canonical_equiv_self, coe_coe, + ring_equiv.coe_mul_equiv_refl, mul_equiv.refl_apply], refine ⟨λ ⟨x, hx⟩, ⟨⟨x, by rw [← hx, coe_span_singleton]⟩⟩, _⟩, unfreezingI { intros hI }, obtain ⟨x, hx⟩ := @submodule.is_principal.principal _ _ _ _ _ _ hI, @@ -199,43 +306,44 @@ begin simp [hx'] end -variables [is_domain R] - lemma class_group.mk0_eq_one_iff [is_dedekind_domain R] {I : ideal R} (hI : I ∈ (ideal R)⁰) : - class_group.mk0 K ⟨I, hI⟩ = 1 ↔ I.is_principal := -class_group.mk_eq_one_iff.trans (coe_submodule_is_principal R K) + class_group.mk0 ⟨I, hI⟩ = 1 ↔ I.is_principal := +class_group.mk_eq_one_iff.trans (coe_submodule_is_principal R _) /-- The class group of principal ideal domain is finite (in fact a singleton). -TODO: generalize to Dedekind domains -/ -instance [is_principal_ideal_ring R] : - fintype (class_group R K) := + +See `class_group.fintype_of_admissible` for a finiteness proof that works for rings of integers +of global fields. +-/ +noncomputable instance [is_principal_ideal_ring R] : + fintype (class_group R) := { elems := {1}, complete := begin - rintros ⟨I⟩, - rw [finset.mem_singleton], - exact class_group.mk_eq_one_iff.mpr (I : fractional_ideal R⁰ K).is_principal + refine class_group.induction (fraction_ring R) (λ I, _), + rw finset.mem_singleton, + exact class_group.mk_eq_one_iff.mpr (I : fractional_ideal R⁰ (fraction_ring R)).is_principal end } /-- The class number of a principal ideal domain is `1`. -/ lemma card_class_group_eq_one [is_principal_ideal_ring R] : - fintype.card (class_group R K) = 1 := + fintype.card (class_group R) = 1 := begin rw fintype.card_eq_one_iff, use 1, - rintros ⟨I⟩, - exact class_group.mk_eq_one_iff.mpr (I : fractional_ideal R⁰ K).is_principal + refine class_group.induction (fraction_ring R) (λ I, _), + exact class_group.mk_eq_one_iff.mpr (I : fractional_ideal R⁰ (fraction_ring R)).is_principal end /-- The class number is `1` iff the ring of integers is a principal ideal domain. -/ -lemma card_class_group_eq_one_iff [is_dedekind_domain R] [fintype (class_group R K)] : - fintype.card (class_group R K) = 1 ↔ is_principal_ideal_ring R := +lemma card_class_group_eq_one_iff [is_dedekind_domain R] [fintype (class_group R)] : + fintype.card (class_group R) = 1 ↔ is_principal_ideal_ring R := begin - split, swap, { introsI, convert card_class_group_eq_one, assumption, assumption, }, + split, swap, { introsI, convert card_class_group_eq_one, assumption, }, rw fintype.card_eq_one_iff, rintros ⟨I, hI⟩, - have eq_one : ∀ J : class_group R K, J = 1 := λ J, trans (hI J) (hI 1).symm, + have eq_one : ∀ J : class_group R, J = 1 := λ J, trans (hI J) (hI 1).symm, refine ⟨λ I, _⟩, by_cases hI : I = ⊥, { rw hI, exact bot_is_principal }, From 324a7502510e835cdbd3de1519b6c66b51fb2467 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Sat, 8 Oct 2022 20:02:42 +0000 Subject: [PATCH 636/828] feat(algebraic_topology): extra degeneracies and homotopy equivalences (#16855) In this PR, it is shown that if an augmented simplicial object `X` in a preadditive category has an extra degeneracy, then the augmentation is a homotopy equivalence on the alternating face map complex. --- src/algebra/homology/single.lean | 32 +++++++++ .../alternating_face_map_complex.lean | 22 ++++++ src/algebraic_topology/extra_degeneracy.lean | 70 +++++++++++++++++-- .../preadditive/projective_resolution.lean | 6 +- 4 files changed, 119 insertions(+), 11 deletions(-) diff --git a/src/algebra/homology/single.lean b/src/algebra/homology/single.lean index 82ff2975d8030..dee21c6c21614 100644 --- a/src/algebra/homology/single.lean +++ b/src/algebra/homology/single.lean @@ -179,6 +179,7 @@ Morphisms from a `ℕ`-indexed chain complex `C` to a single object chain complex with `X` concentrated in degree 0 are the same as morphisms `f : C.X 0 ⟶ X` such that `C.d 1 0 ≫ f = 0`. -/ +@[simps] def to_single₀_equiv (C : chain_complex V ℕ) (X : V) : (C ⟶ (single₀ V).obj X) ≃ { f : C.X 0 ⟶ X // C.d 1 0 ≫ f = 0 } := { to_fun := λ f, ⟨f.f 0, by { rw ←f.comm 1 0, simp, }⟩, @@ -201,6 +202,37 @@ def to_single₀_equiv (C : chain_complex V ℕ) (X : V) : end, right_inv := by tidy, } +@[ext] +lemma to_single₀_ext {C : chain_complex V ℕ} {X : V} + (f g : (C ⟶ (single₀ V).obj X)) (h : f.f 0 = g.f 0) : f = g := +(to_single₀_equiv C X).injective (by { ext, exact h, }) + +/-- +Morphisms from a single object chain complex with `X` concentrated in degree 0 +to a `ℕ`-indexed chain complex `C` are the same as morphisms `f : X → C.X`. +-/ +@[simps] +def from_single₀_equiv (C : chain_complex V ℕ) (X : V) : + ((single₀ V).obj X ⟶ C) ≃ (X ⟶ C.X 0) := +{ to_fun := λ f, f.f 0, + inv_fun := λ f, + { f := λ i, match i with + | 0 := f + | (n+1) := 0 + end, + comm' := λ i j h, begin + cases i; cases j; unfold_aux; + simp only [shape, complex_shape.down_rel, nat.one_ne_zero, not_false_iff, + comp_zero, zero_comp, nat.succ_ne_zero, single₀_obj_X_d], + end }, + left_inv := λ f, begin + ext i, + cases i, + { refl, }, + { ext, }, + end, + right_inv := λ g, rfl, } + variables (V) /-- `single₀` is the same as `single V _ 0`. -/ diff --git a/src/algebraic_topology/alternating_face_map_complex.lean b/src/algebraic_topology/alternating_face_map_complex.lean index 2a571c2195c97..9c05ce8e2a5c8 100644 --- a/src/algebraic_topology/alternating_face_map_complex.lean +++ b/src/algebraic_topology/alternating_face_map_complex.lean @@ -8,6 +8,7 @@ import algebra.homology.additive import algebraic_topology.Moore_complex import algebra.big_operators.fin import category_theory.preadditive.opposite +import tactic.equiv_rw /-! @@ -203,6 +204,27 @@ begin refl, }, }, end +namespace alternating_face_map_complex + +/-- The natural transformation which gives the augmentation of the alternating face map +complex attached to an augmented simplicial object. -/ +@[simps] +def ε [limits.has_zero_object C] : + simplicial_object.augmented.drop ⋙ algebraic_topology.alternating_face_map_complex C ⟶ + simplicial_object.augmented.point ⋙ chain_complex.single₀ C := +{ app := λ X, begin + equiv_rw chain_complex.to_single₀_equiv _ _, + refine ⟨X.hom.app (op [0]), _⟩, + dsimp, + simp only [alternating_face_map_complex_obj_d, obj_d, fin.sum_univ_two, + fin.coe_zero, pow_zero, one_zsmul, fin.coe_one, pow_one, neg_smul, add_comp, + simplicial_object.δ_naturality, neg_comp], + apply add_right_neg, + end, + naturality' := λ X Y f, by { ext, exact congr_app f.w _, }, } + +end alternating_face_map_complex + /-! ## Construction of the natural inclusion of the normalized Moore complex -/ diff --git a/src/algebraic_topology/extra_degeneracy.lean b/src/algebraic_topology/extra_degeneracy.lean index 414c52dbb93e4..0499cbf96e582 100644 --- a/src/algebraic_topology/extra_degeneracy.lean +++ b/src/algebraic_topology/extra_degeneracy.lean @@ -4,8 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ +import algebraic_topology.alternating_face_map_complex import algebraic_topology.simplicial_set import algebraic_topology.cech_nerve +import algebra.homology.homotopy import tactic.fin_cases /-! @@ -31,12 +33,10 @@ functor `C ⥤ D` an extra degeneracy - `arrow.augmented_cech_nerve.extra_degeneracy`: the Čech nerve of a split epimorphism has an extra degeneracy - -TODO @joelriou: -1) when the category `C` is preadditive and has a zero object, and -`X : simplicial_object.augmented C` has an extra degeneracy, then the augmentation -on the alternating face map complex of `X` is a homotopy equivalence of chain -complexes. +- `extra_degeneracy.homotopy_equiv`: in the case the category `C` is preadditive, +if we have an extra degeneracy on `X : simplicial_object.augmented C`, then +the augmentation on the alternating face map complex of `X` is a homotopy +equivalence. ## References * [Paul G. Goerss, John F. Jardine, *Simplical Homotopy Theory*][goerss-jardine-2009] @@ -337,3 +337,61 @@ end augmented_cech_nerve end arrow end category_theory + +namespace simplicial_object + +namespace augmented + +namespace extra_degeneracy + +open algebraic_topology category_theory category_theory.limits + +/-- If `C` is a preadditive category and `X` is an augmented simplicial object +in `C` that has an extra degeneracy, then the augmentation on the alternating +face map complex of `X` is an homotopy equivalence. -/ +noncomputable +def homotopy_equiv {C : Type*} [category C] + [preadditive C] [has_zero_object C] {X : simplicial_object.augmented C} + (ed : extra_degeneracy X) : + homotopy_equiv (algebraic_topology.alternating_face_map_complex.obj (drop.obj X)) + ((chain_complex.single₀ C).obj (point.obj X)) := +{ hom := alternating_face_map_complex.ε.app X, + inv := (chain_complex.from_single₀_equiv _ _).inv_fun ed.s', + homotopy_inv_hom_id := homotopy.of_eq (by { ext, exact ed.s'_comp_ε, }), + homotopy_hom_inv_id := + { hom := λ i j, begin + by_cases i+1 = j, + { exact (-ed.s i) ≫ eq_to_hom (by congr'), }, + { exact 0, }, + end, + zero' := λ i j hij, begin + split_ifs, + { exfalso, exact hij h, }, + { simp only [eq_self_iff_true], }, + end, + comm := λ i, begin + cases i, + { rw [homotopy.prev_d_chain_complex, homotopy.d_next_zero_chain_complex, zero_add], + dsimp [chain_complex.from_single₀_equiv, chain_complex.to_single₀_equiv], + simp only [zero_add, eq_self_iff_true, preadditive.neg_comp, comp_id, if_true, + alternating_face_map_complex.obj_d_eq, fin.sum_univ_two, fin.coe_zero, pow_zero, + one_zsmul, fin.coe_one, pow_one, neg_smul, preadditive.comp_add, ← s₀_comp_δ₁, + s_comp_δ₀, preadditive.comp_neg, neg_add_rev, neg_neg, neg_add_cancel_right, + neg_add_cancel_comm], }, + { rw [homotopy.prev_d_chain_complex, homotopy.d_next_succ_chain_complex], + dsimp [chain_complex.to_single₀_equiv, chain_complex.from_single₀_equiv], + simp only [zero_comp, alternating_face_map_complex.obj_d_eq, eq_self_iff_true, + preadditive.neg_comp, comp_id, if_true, preadditive.comp_neg, + @fin.sum_univ_succ _ _ (i+2), preadditive.comp_add, fin.coe_zero, pow_zero, one_zsmul, + s_comp_δ₀, fin.coe_succ, pow_add, pow_one, mul_neg, neg_zsmul, + preadditive.comp_sum, preadditive.sum_comp, neg_neg, mul_one, + preadditive.comp_zsmul, preadditive.zsmul_comp, s_comp_δ, zsmul_neg], + rw [add_comm (-𝟙 _), add_assoc, add_assoc, add_left_neg, add_zero, + finset.sum_neg_distrib, add_left_neg], }, + end, }, } + +end extra_degeneracy + +end augmented + +end simplicial_object diff --git a/src/category_theory/preadditive/projective_resolution.lean b/src/category_theory/preadditive/projective_resolution.lean index d9f1d8c7636c5..d04bcf9f2bbb5 100644 --- a/src/category_theory/preadditive/projective_resolution.lean +++ b/src/category_theory/preadditive/projective_resolution.lean @@ -172,11 +172,7 @@ chain_complex.mk_hom _ _ (lift_f_zero f _ _) (lift_f_one f _ _) (lift_f_one_zero lemma lift_commutes {Y Z : C} (f : Y ⟶ Z) (P : ProjectiveResolution Y) (Q : ProjectiveResolution Z) : lift f P Q ≫ Q.π = P.π ≫ (chain_complex.single₀ C).map f := -begin - ext n, - rcases n with (_|_|n); - { dsimp [lift, lift_f_zero, lift_f_one], simp, } -end +by { ext, dsimp [lift, lift_f_zero], apply factor_thru_comp, } -- Now that we've checked this property of the lift, -- we can seal away the actual definition. From 619fd4d5030c549880e71b13ba3a2db75616a6f7 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sat, 8 Oct 2022 20:02:43 +0000 Subject: [PATCH 637/828] feat(data/finite/card): Add `finite.card_pos` (#16858) We already have `finite.card_pos_iff` analogous to `fintype.card_pos_iff`. This PR adds `finite.card_pos` analogous to `fintype.card_pos`. --- src/data/finite/card.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/src/data/finite/card.lean b/src/data/finite/card.lean index 20d2f09883f4c..a785977b0e09d 100644 --- a/src/data/finite/card.lean +++ b/src/data/finite/card.lean @@ -51,14 +51,15 @@ begin { simp [*, not_finite_iff_infinite.mpr h] }, end -lemma finite.card_pos_iff [finite α] : - 0 < nat.card α ↔ nonempty α := +lemma finite.card_pos_iff [finite α] : 0 < nat.card α ↔ nonempty α := begin haveI := fintype.of_finite α, - simp only [nat.card_eq_fintype_card], - exact fintype.card_pos_iff, + rw [nat.card_eq_fintype_card, fintype.card_pos_iff], end +lemma finite.card_pos [finite α] [h : nonempty α] : 0 < nat.card α := +finite.card_pos_iff.mpr h + namespace finite lemma card_eq [finite α] [finite β] : nat.card α = nat.card β ↔ nonempty (α ≃ β) := From c0e00a871b9f6d3aca7c10fb3abdc8720a2c5313 Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Sat, 8 Oct 2022 22:24:11 +0000 Subject: [PATCH 638/828] feat(category_theory/closed/functor_category): the functor category from a groupoid to a monoidal closed category is monoidal closed (#15643) Co-authored-by: antoinelab01 <66086247+antoinelab01@users.noreply.github.com> --- .../closed/functor_category.lean | 73 +++++++++++++++++++ src/category_theory/closed/monoidal.lean | 8 ++ src/category_theory/groupoid.lean | 7 ++ 3 files changed, 88 insertions(+) create mode 100644 src/category_theory/closed/functor_category.lean diff --git a/src/category_theory/closed/functor_category.lean b/src/category_theory/closed/functor_category.lean new file mode 100644 index 0000000000000..079a52c0f2982 --- /dev/null +++ b/src/category_theory/closed/functor_category.lean @@ -0,0 +1,73 @@ +/- +Copyright (c) 2022 Antoine Labelle. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Labelle +-/ +import category_theory.closed.monoidal +import category_theory.monoidal.functor_category + +/-! +# Functors from a groupoid into a monoidal closed category form a monoidal closed category. + +(Using the pointwise monoidal structure on the functor category.) +-/ + +noncomputable theory + +open category_theory +open category_theory.monoidal_category +open category_theory.monoidal_closed + +namespace category_theory.functor + +variables {C D : Type*} [groupoid D] [category C] [monoidal_category C] [monoidal_closed C] + +/-- Auxiliary definition for `category_theory.monoidal_closed.functor_closed`. +The internal hom functor `F ⟶[C] -` -/ +@[simps] def closed_ihom (F : D ⥤ C) : (D ⥤ C) ⥤ (D ⥤ C) := +((whiskering_right₂ D Cᵒᵖ C C).obj internal_hom).obj (groupoid.inv_functor D ⋙ F.op) + +/-- Auxiliary definition for `category_theory.monoidal_closed.functor_closed`. +The unit for the adjunction `(tensor_left F) ⊣ (ihom F)`. -/ +@[simps] +def closed_unit (F : D ⥤ C) : 𝟭 (D ⥤ C) ⟶ (tensor_left F) ⋙ (closed_ihom F) := +{ app := λ G, + { app := λ X, (ihom.coev (F.obj X)).app (G.obj X), + naturality' := begin + intros X Y f, + dsimp, + simp only [ihom.coev_naturality, closed_ihom_obj_map, monoidal.tensor_obj_map], + dsimp, + rw [coev_app_comp_pre_app_assoc, ←functor.map_comp], + simp, + end } } + +/-- Auxiliary definition for `category_theory.monoidal_closed.functor_closed`. +The counit for the adjunction `(tensor_left F) ⊣ (ihom F)`. -/ +@[simps] +def closed_counit (F : D ⥤ C) : (closed_ihom F) ⋙ (tensor_left F) ⟶ 𝟭 (D ⥤ C) := +{ app := λ G, + { app := λ X, (ihom.ev (F.obj X)).app (G.obj X), + naturality' := begin + intros X Y f, + dsimp, + simp only [closed_ihom_obj_map, pre_comm_ihom_map], + rw [←tensor_id_comp_id_tensor, id_tensor_comp], + simp, + end } } + +/-- If `C` is a monoidal closed category and `D` is groupoid, then every functor `F : D ⥤ C` is +closed in the functor category `F : D ⥤ C` with the pointwise monoidal structure. -/ +@[simps] instance closed (F : D ⥤ C) : closed F := +{ is_adj := + { right := closed_ihom F, + adj := adjunction.mk_of_unit_counit + { unit := closed_unit F, + counit := closed_counit F } } } + +/-- If `C` is a monoidal closed category and `D` is groupoid, then the functor category `D ⥤ C`, +with the pointwise monoidal structure, is monoidal closed. -/ +@[simps] instance monoidal_closed : monoidal_closed (D ⥤ C) := +{ closed' := by apply_instance } + +end category_theory.functor diff --git a/src/category_theory/closed/monoidal.lean b/src/category_theory/closed/monoidal.lean index 433aec0d14dce..f583bab91441f 100644 --- a/src/category_theory/closed/monoidal.lean +++ b/src/category_theory/closed/monoidal.lean @@ -206,15 +206,18 @@ variables {A B} [closed B] def pre (f : B ⟶ A) : ihom A ⟶ ihom B := transfer_nat_trans_self (ihom.adjunction _) (ihom.adjunction _) ((tensoring_left C).map f) +@[simp, reassoc] lemma id_tensor_pre_app_comp_ev (f : B ⟶ A) (X : C) : (𝟙 B ⊗ ((pre f).app X)) ≫ (ihom.ev B).app X = (f ⊗ (𝟙 (A ⟶[C] X))) ≫ (ihom.ev A).app X := transfer_nat_trans_self_counit _ _ ((tensoring_left C).map f) X +@[simp] lemma uncurry_pre (f : B ⟶ A) (X : C) : monoidal_closed.uncurry ((pre f).app X) = (f ⊗ 𝟙 _) ≫ (ihom.ev A).app X := by rw [uncurry_eq, id_tensor_pre_app_comp_ev] +@[simp, reassoc] lemma coev_app_comp_pre_app (f : B ⟶ A) : (ihom.coev A).app X ≫ (pre f).app (A ⊗ X) = (ihom.coev B).app X ≫ (ihom B).map (f ⊗ (𝟙 _)) := @@ -230,9 +233,14 @@ lemma pre_map {A₁ A₂ A₃ : C} [closed A₁] [closed A₂] [closed A₃] pre (f ≫ g) = pre g ≫ pre f := by rw [pre, pre, pre, transfer_nat_trans_self_comp, (tensoring_left C).map_comp] +lemma pre_comm_ihom_map {W X Y Z : C} [closed W] [closed X] + (f : W ⟶ X) (g : Y ⟶ Z) : + (pre f).app Y ≫ (ihom W).map g = (ihom X).map g ≫ (pre f).app Z := by simp + end pre /-- The internal hom functor given by the monoidal closed structure. -/ +@[simps] def internal_hom [monoidal_closed C] : Cᵒᵖ ⥤ C ⥤ C := { obj := λ X, ihom X.unop, map := λ X Y f, pre f.unop } diff --git a/src/category_theory/groupoid.lean b/src/category_theory/groupoid.lean index bd01812fbbaba..65b4775c15776 100644 --- a/src/category_theory/groupoid.lean +++ b/src/category_theory/groupoid.lean @@ -83,6 +83,13 @@ def groupoid.iso_equiv_hom : (X ≅ Y) ≃ (X ⟶ Y) := left_inv := λ i, iso.ext rfl, right_inv := λ f, rfl } +variables (C) + +/-- The functor from a groupoid `C` to its opposite sending every morphism to its inverse. -/ +@[simps] noncomputable def groupoid.inv_functor : C ⥤ Cᵒᵖ := +{ obj := opposite.op, + map := λ {X Y} f, (inv f).op } + end section From a64e3646da34466de13153495381fc412bd30b02 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 9 Oct 2022 00:39:12 +0000 Subject: [PATCH 639/828] feat(order/cover): congr lemmas for `covby` (#15663) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove that `antisymm_rel (≤) a b` and `a ⋖ c` implies `b ⋖ c`, and variations thereof. Co-authored-by: Yaël Dillies Co-authored-by: Junyan Xu Co-authored-by: Scott Morrison --- src/order/cover.lean | 34 +++++++++++++++++-- .../sheaves/presheaf_of_functions.lean | 9 +++-- 2 files changed, 39 insertions(+), 4 deletions(-) diff --git a/src/order/cover.lean b/src/order/cover.lean index 1c64915559291..250089f164cdb 100644 --- a/src/order/cover.lean +++ b/src/order/cover.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Violeta Hernández Palacios, Grayson Burton, Floris van Doorn -/ import data.set.intervals.ord_connected +import order.antisymmetrization /-! # The covering relation @@ -26,7 +27,7 @@ variables {α β : Type*} section weakly_covers section preorder -variables [preorder α] [preorder β] {a b c: α} +variables [preorder α] [preorder β] {a b c : α} /-- `wcovby a b` means that `a = b` or `b` covers `a`. This means that `a ≤ b` and there is no element in between. @@ -47,12 +48,26 @@ lemma wcovby_of_le_of_le (h1 : a ≤ b) (h2 : b ≤ a) : a ⩿ b := alias wcovby_of_le_of_le ← has_le.le.wcovby_of_le +lemma antisymm_rel.wcovby (h : antisymm_rel (≤) a b) : a ⩿ b := wcovby_of_le_of_le h.1 h.2 + lemma wcovby.wcovby_iff_le (hab : a ⩿ b) : b ⩿ a ↔ b ≤ a := ⟨λ h, h.le, λ h, h.wcovby_of_le hab.le⟩ lemma wcovby_of_eq_or_eq (hab : a ≤ b) (h : ∀ c, a ≤ c → c ≤ b → c = a ∨ c = b) : a ⩿ b := ⟨hab, λ c ha hb, (h c ha.le hb.le).elim ha.ne' hb.ne⟩ +lemma antisymm_rel.trans_wcovby (hab : antisymm_rel (≤) a b) (hbc : b ⩿ c) : a ⩿ c := +⟨hab.1.trans hbc.le, λ d had hdc, hbc.2 (hab.2.trans_lt had) hdc⟩ + +lemma wcovby_congr_left (hab : antisymm_rel (≤) a b) : a ⩿ c ↔ b ⩿ c := +⟨hab.symm.trans_wcovby, hab.trans_wcovby⟩ + +lemma wcovby.trans_antisymm_rel (hab : a ⩿ b) (hbc : antisymm_rel (≤) b c) : a ⩿ c := +⟨hab.le.trans hbc.1, λ d had hdc, hab.2 had $ hdc.trans_le hbc.2⟩ + +lemma wcovby_congr_right (hab : antisymm_rel (≤) a b) : c ⩿ a ↔ c ⩿ b := +⟨λ h, h.trans_antisymm_rel hab, λ h, h.trans_antisymm_rel hab.symm⟩ + /-- If `a ≤ b`, then `b` does not cover `a` iff there's an element in between. -/ lemma not_wcovby_iff (h : a ≤ b) : ¬ a ⩿ b ↔ ∃ c, a < c ∧ c < b := by simp_rw [wcovby, h, true_and, not_forall, exists_prop, not_not] @@ -157,7 +172,7 @@ alias of_dual_covby_of_dual_iff ↔ _ covby.of_dual end has_lt section preorder -variables [preorder α] [preorder β] {a b : α} +variables [preorder α] [preorder β] {a b c : α} lemma covby.le (h : a ⋖ b) : a ≤ b := h.1.le protected lemma covby.ne (h : a ⋖ b) : a ≠ b := h.lt.ne @@ -167,6 +182,9 @@ protected lemma covby.wcovby (h : a ⋖ b) : a ⩿ b := ⟨h.le, h.2⟩ lemma wcovby.covby_of_not_le (h : a ⩿ b) (h2 : ¬ b ≤ a) : a ⋖ b := ⟨h.le.lt_of_not_le h2, h.2⟩ lemma wcovby.covby_of_lt (h : a ⩿ b) (h2 : a < b) : a ⋖ b := ⟨h2, h.2⟩ +lemma not_covby_of_lt_of_lt (h₁ : a < b) (h₂ : b < c) : ¬ a ⋖ c := +(not_covby_iff (h₁.trans h₂)).2 ⟨b, h₁, h₂⟩ + lemma covby_iff_wcovby_and_lt : a ⋖ b ↔ a ⩿ b ∧ a < b := ⟨λ h, ⟨h.wcovby, h.lt⟩, λ h, h.1.covby_of_lt h.2⟩ @@ -177,6 +195,18 @@ lemma wcovby_iff_covby_or_le_and_le : a ⩿ b ↔ a ⋖ b ∨ (a ≤ b ∧ b ≤ ⟨λ h, or_iff_not_imp_right.mpr $ λ h', h.covby_of_not_le $ λ hba, h' ⟨h.le, hba⟩, λ h', h'.elim (λ h, h.wcovby) (λ h, h.1.wcovby_of_le h.2)⟩ +lemma antisymm_rel.trans_covby (hab : antisymm_rel (≤) a b) (hbc : b ⋖ c) : a ⋖ c := +⟨hab.1.trans_lt hbc.lt, λ d had hdc, hbc.2 (hab.2.trans_lt had) hdc⟩ + +lemma covby_congr_left (hab : antisymm_rel (≤) a b) : a ⋖ c ↔ b ⋖ c := +⟨hab.symm.trans_covby, hab.trans_covby⟩ + +lemma covby.trans_antisymm_rel (hab : a ⋖ b) (hbc : antisymm_rel (≤) b c) : a ⋖ c := +⟨hab.lt.trans_le hbc.1, λ d had hdb, hab.2 had $ hdb.trans_le hbc.2⟩ + +lemma covby_congr_right (hab : antisymm_rel (≤) a b) : c ⋖ a ↔ c ⋖ b := +⟨λ h, h.trans_antisymm_rel hab, λ h, h.trans_antisymm_rel hab.symm⟩ + instance : is_nonstrict_strict_order α (⩿) (⋖) := ⟨λ a b, covby_iff_wcovby_and_not_le.trans $ and_congr_right $ λ h, h.wcovby_iff_le.not.symm⟩ diff --git a/src/topology/sheaves/presheaf_of_functions.lean b/src/topology/sheaves/presheaf_of_functions.lean index dc45e84c8b10d..db5f92930671d 100644 --- a/src/topology/sheaves/presheaf_of_functions.lean +++ b/src/topology/sheaves/presheaf_of_functions.lean @@ -119,9 +119,14 @@ from `X : Top` to `R : TopCommRing` form a commutative ring, functorial in both def CommRing_yoneda : TopCommRing.{u} ⥤ (Top.{u}ᵒᵖ ⥤ CommRing.{u}) := { obj := λ R, { obj := λ X, continuous_functions X R, - map := λ X Y f, continuous_functions.pullback f R }, + map := λ X Y f, continuous_functions.pullback f R, + map_id' := λ X, by { ext, refl }, + map_comp' := λ X Y Z f g, rfl }, map := λ R S φ, - { app := λ X, continuous_functions.map X φ } } + { app := λ X, continuous_functions.map X φ, + naturality' := λ X Y f, rfl }, + map_id' := λ X, by { ext, refl }, + map_comp' := λ X Y Z f g, rfl } /-- The presheaf (of commutative rings), consisting of functions on an open set `U ⊆ X` with From e71a12cda08a56f6de36d2e988312408f4de1cf1 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Sun, 9 Oct 2022 00:39:13 +0000 Subject: [PATCH 640/828] feat(analysis/seminorm): closed balls for a seminorm (#16676) This introduces the closed ball version of `seminorm.ball` and duplicates a bunch of API. My motivation is the general Banach-Steinhaus theorem in barreled space (one characterization of barrels is that they are exactly closed unit balls of lower semicontinuous seminorms), and I didn't see any reason not do go full duplication here. Co-authored-by: Patrick Massot --- src/analysis/seminorm.lean | 151 +++++++++++++++++++++++++-- src/topology/metric_space/basic.lean | 3 + 2 files changed, 148 insertions(+), 6 deletions(-) diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index e686006c2c223..c10ddad892510 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -399,18 +399,32 @@ section has_smul variables [has_smul 𝕜 E] (p : seminorm 𝕜 E) /-- The ball of radius `r` at `x` with respect to seminorm `p` is the set of elements `y` with -`p (y - x) < `r`. -/ +`p (y - x) < r`. -/ def ball (x : E) (r : ℝ) := { y : E | p (y - x) < r } +/-- The closed ball of radius `r` at `x` with respect to seminorm `p` is the set of elements `y` +with `p (y - x) ≤ r`. -/ +def closed_ball (x : E) (r : ℝ) := { y : E | p (y - x) ≤ r } + variables {x y : E} {r : ℝ} @[simp] lemma mem_ball : y ∈ ball p x r ↔ p (y - x) < r := iff.rfl +@[simp] lemma mem_closed_ball : y ∈ closed_ball p x r ↔ p (y - x) ≤ r := iff.rfl lemma mem_ball_self (hr : 0 < r) : x ∈ ball p x r := by simp [hr] +lemma mem_closed_ball_self (hr : 0 ≤ r) : x ∈ closed_ball p x r := by simp [hr] lemma mem_ball_zero : y ∈ ball p 0 r ↔ p y < r := by rw [mem_ball, sub_zero] +lemma mem_closed_ball_zero : y ∈ closed_ball p 0 r ↔ p y ≤ r := by rw [mem_closed_ball, sub_zero] lemma ball_zero_eq : ball p 0 r = { y : E | p y < r } := set.ext $ λ x, p.mem_ball_zero +lemma closed_ball_zero_eq : closed_ball p 0 r = { y : E | p y ≤ r } := +set.ext $ λ x, p.mem_closed_ball_zero + +lemma ball_subset_closed_ball (x r) : ball p x r ⊆ closed_ball p x r := λ y (hy : _ < _), hy.le + +lemma closed_ball_eq_bInter_ball (x r) : closed_ball p x r = ⋂ ρ > r, ball p x ρ := +by ext y; simp_rw [mem_closed_ball, mem_Inter₂, mem_ball, ← forall_lt_iff_le'] @[simp] lemma ball_zero' (x : E) (hr : 0 < r) : ball (0 : seminorm 𝕜 E) x r = set.univ := begin @@ -418,15 +432,28 @@ begin simp [hr], end +@[simp] lemma closed_ball_zero' (x : E) (hr : 0 < r) : + closed_ball (0 : seminorm 𝕜 E) x r = set.univ := +eq_univ_of_subset (ball_subset_closed_ball _ _ _) (ball_zero' x hr) + lemma ball_smul (p : seminorm 𝕜 E) {c : nnreal} (hc : 0 < c) (r : ℝ) (x : E) : (c • p).ball x r = p.ball x (r / c) := by { ext, rw [mem_ball, mem_ball, smul_apply, nnreal.smul_def, smul_eq_mul, mul_comm, lt_div_iff (nnreal.coe_pos.mpr hc)] } +lemma closed_ball_smul (p : seminorm 𝕜 E) {c : nnreal} (hc : 0 < c) (r : ℝ) (x : E) : + (c • p).closed_ball x r = p.closed_ball x (r / c) := +by { ext, rw [mem_closed_ball, mem_closed_ball, smul_apply, nnreal.smul_def, smul_eq_mul, mul_comm, + le_div_iff (nnreal.coe_pos.mpr hc)] } + lemma ball_sup (p : seminorm 𝕜 E) (q : seminorm 𝕜 E) (e : E) (r : ℝ) : ball (p ⊔ q) e r = ball p e r ∩ ball q e r := by simp_rw [ball, ←set.set_of_and, coe_sup, pi.sup_apply, sup_lt_iff] +lemma closed_ball_sup (p : seminorm 𝕜 E) (q : seminorm 𝕜 E) (e : E) (r : ℝ) : + closed_ball (p ⊔ q) e r = closed_ball p e r ∩ closed_ball q e r := +by simp_rw [closed_ball, ←set.set_of_and, coe_sup, pi.sup_apply, sup_le_iff] + lemma ball_finset_sup' (p : ι → seminorm 𝕜 E) (s : finset ι) (H : s.nonempty) (e : E) (r : ℝ) : ball (s.sup' H p) e r = s.inf' H (λ i, ball (p i) e r) := begin @@ -435,12 +462,28 @@ begin { rw [finset.sup'_cons hs, finset.inf'_cons hs, ball_sup, inf_eq_inter, ih] }, end +lemma closed_ball_finset_sup' (p : ι → seminorm 𝕜 E) (s : finset ι) (H : s.nonempty) (e : E) + (r : ℝ) : closed_ball (s.sup' H p) e r = s.inf' H (λ i, closed_ball (p i) e r) := +begin + induction H using finset.nonempty.cons_induction with a a s ha hs ih, + { classical, simp }, + { rw [finset.sup'_cons hs, finset.inf'_cons hs, closed_ball_sup, inf_eq_inter, ih] }, +end + lemma ball_mono {p : seminorm 𝕜 E} {r₁ r₂ : ℝ} (h : r₁ ≤ r₂) : p.ball x r₁ ⊆ p.ball x r₂ := λ _ (hx : _ < _), hx.trans_le h +lemma closed_ball_mono {p : seminorm 𝕜 E} {r₁ r₂ : ℝ} (h : r₁ ≤ r₂) : + p.closed_ball x r₁ ⊆ p.closed_ball x r₂ := +λ _ (hx : _ ≤ _), hx.trans h + lemma ball_antitone {p q : seminorm 𝕜 E} (h : q ≤ p) : p.ball x r ⊆ q.ball x r := λ _, (h _).trans_lt +lemma closed_ball_antitone {p q : seminorm 𝕜 E} (h : q ≤ p) : + p.closed_ball x r ⊆ q.closed_ball x r := +λ _, (h _).trans + lemma ball_add_ball_subset (p : seminorm 𝕜 E) (r₁ r₂ : ℝ) (x₁ x₂ : E): p.ball (x₁ : E) r₁ + p.ball (x₂ : E) r₂ ⊆ p.ball (x₁ + x₂) (r₁ + r₂) := begin @@ -449,6 +492,14 @@ begin exact (map_add_le_add p _ _).trans_lt (add_lt_add hy₁ hy₂), end +lemma closed_ball_add_closed_ball_subset (p : seminorm 𝕜 E) (r₁ r₂ : ℝ) (x₁ x₂ : E) : + p.closed_ball (x₁ : E) r₁ + p.closed_ball (x₂ : E) r₂ ⊆ p.closed_ball (x₁ + x₂) (r₁ + r₂) := +begin + rintros x ⟨y₁, y₂, hy₁, hy₂, rfl⟩, + rw [mem_closed_ball, add_sub_add_comm], + exact (map_add_le_add p _ _).trans (add_le_add hy₁ hy₂) +end + end has_smul section module @@ -463,21 +514,46 @@ begin simp_rw [ball, mem_preimage, comp_apply, set.mem_set_of_eq, map_sub], end +lemma closed_ball_comp (p : seminorm 𝕜 F) (f : E →ₗ[𝕜] F) (x : E) (r : ℝ) : + (p.comp f).closed_ball x r = f ⁻¹' (p.closed_ball (f x) r) := +begin + ext, + simp_rw [closed_ball, mem_preimage, comp_apply, set.mem_set_of_eq, map_sub], +end + variables (p : seminorm 𝕜 E) -lemma ball_zero_eq_preimage_ball {r : ℝ} : - p.ball 0 r = p ⁻¹' (metric.ball 0 r) := +lemma preimage_metric_ball {r : ℝ} : + p ⁻¹' (metric.ball 0 r) = {x | p x < r} := +begin + ext x, + simp only [mem_set_of, mem_preimage, mem_ball_zero_iff, real.norm_of_nonneg (map_nonneg p _)] +end + +lemma preimage_metric_closed_ball {r : ℝ} : + p ⁻¹' (metric.closed_ball 0 r) = {x | p x ≤ r} := begin ext x, - simp only [mem_ball, sub_zero, mem_preimage, mem_ball_zero_iff], - rw real.norm_of_nonneg, - exact map_nonneg p _, + simp only [mem_set_of, mem_preimage, mem_closed_ball_zero_iff, + real.norm_of_nonneg (map_nonneg p _)] end +lemma ball_zero_eq_preimage_ball {r : ℝ} : + p.ball 0 r = p ⁻¹' (metric.ball 0 r) := +by rw [ball_zero_eq, preimage_metric_ball] + +lemma closed_ball_zero_eq_preimage_closed_ball {r : ℝ} : + p.closed_ball 0 r = p ⁻¹' (metric.closed_ball 0 r) := +by rw [closed_ball_zero_eq, preimage_metric_closed_ball] + @[simp] lemma ball_bot {r : ℝ} (x : E) (hr : 0 < r) : ball (⊥ : seminorm 𝕜 E) x r = set.univ := ball_zero' x hr +@[simp] lemma closed_ball_bot {r : ℝ} (x : E) (hr : 0 < r) : + closed_ball (⊥ : seminorm 𝕜 E) x r = set.univ := +closed_ball_zero' x hr + /-- Seminorm-balls at the origin are balanced. -/ lemma balanced_ball_zero (r : ℝ) : balanced 𝕜 (ball p 0 r) := begin @@ -487,6 +563,15 @@ begin ... < r : by rwa mem_ball_zero at hy, end +/-- Closed seminorm-balls at the origin are balanced. -/ +lemma balanced_closed_ball_zero (r : ℝ) : balanced 𝕜 (closed_ball p 0 r) := +begin + rintro a ha x ⟨y, hy, hx⟩, + rw [mem_closed_ball_zero, ←hx, map_smul_eq_mul], + calc _ ≤ p y : mul_le_of_le_one_left (map_nonneg p _) ha + ... ≤ r : by rwa mem_closed_ball_zero at hy +end + lemma ball_finset_sup_eq_Inter (p : ι → seminorm 𝕜 E) (s : finset ι) (x : E) {r : ℝ} (hr : 0 < r) : ball (s.sup p) x r = ⋂ (i ∈ s), ball (p i) x r := begin @@ -495,6 +580,14 @@ begin finset.sup_lt_iff (show ⊥ < r, from hr), ←nnreal.coe_lt_coe, subtype.coe_mk], end +lemma closed_ball_finset_sup_eq_Inter (p : ι → seminorm 𝕜 E) (s : finset ι) (x : E) {r : ℝ} + (hr : 0 ≤ r) : closed_ball (s.sup p) x r = ⋂ (i ∈ s), closed_ball (p i) x r := +begin + lift r to nnreal using hr, + simp_rw [closed_ball, Inter_set_of, finset_sup_apply, nnreal.coe_le_coe, + finset.sup_le_iff, ←nnreal.coe_le_coe, subtype.coe_mk] +end + lemma ball_finset_sup (p : ι → seminorm 𝕜 E) (s : finset ι) (x : E) {r : ℝ} (hr : 0 < r) : ball (s.sup p) x r = s.inf (λ i, ball (p i) x r) := begin @@ -502,6 +595,13 @@ begin exact ball_finset_sup_eq_Inter _ _ _ hr, end +lemma closed_ball_finset_sup (p : ι → seminorm 𝕜 E) (s : finset ι) (x : E) {r : ℝ} (hr : 0 ≤ r) : + closed_ball (s.sup p) x r = s.inf (λ i, closed_ball (p i) x r) := +begin + rw finset.inf_eq_infi, + exact closed_ball_finset_sup_eq_Inter _ _ _ hr, +end + lemma ball_smul_ball (p : seminorm 𝕜 E) (r₁ r₂ : ℝ) : metric.ball (0 : 𝕜) r₁ • p.ball 0 r₂ ⊆ p.ball 0 (r₁ * r₂) := begin @@ -514,6 +614,18 @@ begin (map_nonneg p y), end +lemma closed_ball_smul_closed_ball (p : seminorm 𝕜 E) (r₁ r₂ : ℝ) : + metric.closed_ball (0 : 𝕜) r₁ • p.closed_ball 0 r₂ ⊆ p.closed_ball 0 (r₁ * r₂) := +begin + rw set.subset_def, + intros x hx, + rw set.mem_smul at hx, + rcases hx with ⟨a, y, ha, hy, hx⟩, + rw [←hx, mem_closed_ball_zero, map_smul_eq_mul], + rw mem_closed_ball_zero_iff at ha, + exact mul_le_mul ha (p.mem_closed_ball_zero.mp hy) (map_nonneg _ y) ((norm_nonneg a).trans ha) +end + @[simp] lemma ball_eq_emptyset (p : seminorm 𝕜 E) {x : E} {r : ℝ} (hr : r ≤ 0) : p.ball x r = ∅ := begin ext, @@ -521,6 +633,14 @@ begin exact hr.trans (map_nonneg p _), end +@[simp] lemma closed_ball_eq_emptyset (p : seminorm 𝕜 E) {x : E} {r : ℝ} (hr : r < 0) : + p.closed_ball x r = ∅ := +begin + ext, + rw [seminorm.mem_closed_ball, set.mem_empty_iff_false, iff_false, not_le], + exact hr.trans_le (map_nonneg _ _), +end + end module end add_comm_group end semi_normed_ring @@ -575,6 +695,10 @@ begin rwa [mem_ball_zero, map_smul_eq_mul, norm_inv, inv_mul_lt_iff ha₀, ←div_lt_iff hr], end +/-- Closed seminorm-balls at the origin are absorbent. -/ +protected lemma absorbent_closed_ball_zero (hr : 0 < r) : absorbent 𝕜 (closed_ball p (0 : E) r) := +(p.absorbent_ball_zero hr).subset (p.ball_subset_closed_ball _ _) + /-- Seminorm-balls containing the origin are absorbent. -/ protected lemma absorbent_ball (hpr : p x < r) : absorbent 𝕜 (ball p x r) := begin @@ -583,6 +707,14 @@ begin exact p.mem_ball.2 ((map_sub_le_add p _ _).trans_lt $ add_lt_of_lt_sub_right hy), end +/-- Seminorm-balls containing the origin are absorbent. -/ +protected lemma absorbent_closed_ball (hpr : p x < r) : absorbent 𝕜 (closed_ball p x r) := +begin + refine (p.absorbent_closed_ball_zero $ sub_pos.2 hpr).subset (λ y hy, _), + rw p.mem_closed_ball_zero at hy, + exact p.mem_closed_ball.2 ((map_sub_le_add p _ _).trans $ add_le_of_le_sub_right hy), +end + lemma symmetric_ball_zero (r : ℝ) (hx : x ∈ ball p 0 r) : -x ∈ ball p 0 r := balanced_ball_zero p r (-1) (by rw [norm_neg, norm_one]) ⟨x, hx, by rw [neg_smul, one_smul]⟩ @@ -631,6 +763,13 @@ begin refl, end +/-- Closed seminorm-balls are convex. -/ +lemma convex_closed_ball : convex ℝ (closed_ball p x r) := +begin + rw closed_ball_eq_bInter_ball, + exact convex_Inter₂ (λ _ _, convex_ball _ _ _) +end + end module end convex diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 8c22175a37fc6..4843a13956b29 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -537,6 +537,9 @@ by rw [mem_sphere', mem_sphere] theorem ball_subset_ball (h : ε₁ ≤ ε₂) : ball x ε₁ ⊆ ball x ε₂ := λ y (yx : _ < ε₁), lt_of_lt_of_le yx h +lemma closed_ball_eq_bInter_ball : closed_ball x ε = ⋂ δ > ε, ball x δ := +by ext y; rw [mem_closed_ball, ← forall_lt_iff_le', mem_Inter₂]; refl + lemma ball_subset_ball' (h : ε₁ + dist x y ≤ ε₂) : ball x ε₁ ⊆ ball y ε₂ := λ z hz, calc dist z y ≤ dist z x + dist x y : dist_triangle _ _ _ From b8d96d7778110fb801dbdf4216aa3d8a1ae0e11a Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 9 Oct 2022 00:39:14 +0000 Subject: [PATCH 641/828] chore(algebra/ring/basic): postpone unneeded imports (#16863) No rearranging of content, just deferring some imports that aren't actually needed until later files. Co-authored-by: Scott Morrison --- src/algebra/hom/ring.lean | 1 + src/algebra/ring/basic.lean | 2 -- 2 files changed, 1 insertion(+), 2 deletions(-) diff --git a/src/algebra/hom/ring.lean b/src/algebra/hom/ring.lean index 1181c9e69f957..3ba8dfd690db1 100644 --- a/src/algebra/hom/ring.lean +++ b/src/algebra/hom/ring.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Jireh Loreaux -/ import algebra.ring.basic +import data.pi.algebra /-! # Homomorphisms of semirings and rings diff --git a/src/algebra/ring/basic.lean b/src/algebra/ring/basic.lean index 956571775f59d..5fa9588c53a54 100644 --- a/src/algebra/ring/basic.lean +++ b/src/algebra/ring/basic.lean @@ -5,8 +5,6 @@ Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Ne -/ import algebra.divisibility import algebra.regular.basic -import data.int.cast.defs -import data.pi.algebra /-! # Semirings and rings From 3bd3840f6a9a3896aba0356118a0bddc08006fab Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 9 Oct 2022 03:38:15 +0000 Subject: [PATCH 642/828] chore(algebra/group/defs): golf a proof (#16845) Co-authored-by: Scott Morrison --- src/algebra/group/defs.lean | 16 ++-------------- 1 file changed, 2 insertions(+), 14 deletions(-) diff --git a/src/algebra/group/defs.lean b/src/algebra/group/defs.lean index edcd8148d4736..8348d576871af 100644 --- a/src/algebra/group/defs.lean +++ b/src/algebra/group/defs.lean @@ -814,13 +814,7 @@ end group @[to_additive] lemma group.to_div_inv_monoid_injective {G : Type*} : function.injective (@group.to_div_inv_monoid G) := -begin - rintros ⟨⟩ ⟨⟩ h, - replace h := div_inv_monoid.mk.inj h, - dsimp at h, - rcases h with ⟨rfl, rfl, rfl, rfl, rfl, rfl⟩, - refl -end +by { rintros ⟨⟩ ⟨⟩ ⟨⟩, refl } /-- A commutative group is a group with commutative `(*)`. -/ @[protect_proj, ancestor group comm_monoid] @@ -834,13 +828,7 @@ attribute [instance, priority 300] add_comm_group.to_add_comm_monoid @[to_additive] lemma comm_group.to_group_injective {G : Type u} : function.injective (@comm_group.to_group G) := -begin - rintros ⟨⟩ ⟨⟩ h, - replace h := group.mk.inj h, - dsimp at h, - rcases h with ⟨rfl, rfl, rfl, rfl, rfl, rfl⟩, - refl -end +by { rintros ⟨⟩ ⟨⟩ ⟨⟩, refl } section comm_group From 8fb5c46f41ea2ace212026adf02766e45eac425c Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 9 Oct 2022 08:02:37 +0000 Subject: [PATCH 643/828] chore(algebra/ring/basic): move results about regular elements (#16865) Co-authored-by: Scott Morrison --- src/algebra/ring/basic.lean | 64 ---------------------- src/algebra/ring/regular.lean | 77 +++++++++++++++++++++++++++ src/data/int/basic.lean | 1 + src/data/set/intervals/instances.lean | 1 + 4 files changed, 79 insertions(+), 64 deletions(-) create mode 100644 src/algebra/ring/regular.lean diff --git a/src/algebra/ring/basic.lean b/src/algebra/ring/basic.lean index 5fa9588c53a54..56ef280a61db4 100644 --- a/src/algebra/ring/basic.lean +++ b/src/algebra/ring/basic.lean @@ -1047,56 +1047,6 @@ lemma succ_ne_self [non_assoc_ring α] [nontrivial α] (a : α) : a + 1 ≠ a := lemma pred_ne_self [non_assoc_ring α] [nontrivial α] (a : α) : a - 1 ≠ a := λ h, one_ne_zero (neg_injective ((add_right_inj a).mp (by simpa [sub_eq_add_neg] using h))) -/-- Left `mul` by a `k : α` over `[ring α]` is injective, if `k` is not a zero divisor. -The typeclass that restricts all terms of `α` to have this property is `no_zero_divisors`. -/ -lemma is_left_regular_of_non_zero_divisor [non_unital_non_assoc_ring α] (k : α) - (h : ∀ (x : α), k * x = 0 → x = 0) : is_left_regular k := -begin - refine λ x y (h' : k * x = k * y), sub_eq_zero.mp (h _ _), - rw [mul_sub, sub_eq_zero, h'] -end - -/-- Right `mul` by a `k : α` over `[ring α]` is injective, if `k` is not a zero divisor. -The typeclass that restricts all terms of `α` to have this property is `no_zero_divisors`. -/ -lemma is_right_regular_of_non_zero_divisor [non_unital_non_assoc_ring α] (k : α) - (h : ∀ (x : α), x * k = 0 → x = 0) : is_right_regular k := -begin - refine λ x y (h' : x * k = y * k), sub_eq_zero.mp (h _ _), - rw [sub_mul, sub_eq_zero, h'] -end - -lemma is_regular_of_ne_zero' [non_unital_non_assoc_ring α] [no_zero_divisors α] {k : α} - (hk : k ≠ 0) : is_regular k := -⟨is_left_regular_of_non_zero_divisor k - (λ x h, (no_zero_divisors.eq_zero_or_eq_zero_of_mul_eq_zero h).resolve_left hk), - is_right_regular_of_non_zero_divisor k - (λ x h, (no_zero_divisors.eq_zero_or_eq_zero_of_mul_eq_zero h).resolve_right hk)⟩ - -lemma is_regular_iff_ne_zero' [nontrivial α] [non_unital_non_assoc_ring α] [no_zero_divisors α] - {k : α} : is_regular k ↔ k ≠ 0 := -⟨λ h, by { rintro rfl, exact not_not.mpr h.left not_is_left_regular_zero }, is_regular_of_ne_zero'⟩ - -/-- A ring with no zero divisors is a `cancel_monoid_with_zero`. - -Note this is not an instance as it forms a typeclass loop. -/ -@[reducible] -def no_zero_divisors.to_cancel_monoid_with_zero [ring α] [no_zero_divisors α] : - cancel_monoid_with_zero α := -{ mul_left_cancel_of_ne_zero := λ a b c ha, - @is_regular.left _ _ _ (is_regular_of_ne_zero' ha) _ _, - mul_right_cancel_of_ne_zero := λ a b c hb, - @is_regular.right _ _ _ (is_regular_of_ne_zero' hb) _ _, - .. (by apply_instance : monoid_with_zero α) } - -/-- A commutative ring with no zero divisors is a `cancel_comm_monoid_with_zero`. - -Note this is not an instance as it forms a typeclass loop. -/ -@[reducible] -def no_zero_divisors.to_cancel_comm_monoid_with_zero [comm_ring α] [no_zero_divisors α] : - cancel_comm_monoid_with_zero α := -{ .. no_zero_divisors.to_cancel_monoid_with_zero, - .. (by apply_instance : comm_monoid_with_zero α) } - /-- A domain is a nontrivial ring with no zero divisors, i.e. satisfying the condition `a * b = 0 ↔ a = 0 ∨ b = 0`. @@ -1105,20 +1055,6 @@ def no_zero_divisors.to_cancel_comm_monoid_with_zero [comm_ring α] [no_zero_div @[protect_proj] class is_domain (α : Type u) [ring α] extends no_zero_divisors α, nontrivial α : Prop -section is_domain - -@[priority 100] -- see Note [lower instance priority] -instance is_domain.to_cancel_monoid_with_zero [ring α] [is_domain α] : cancel_monoid_with_zero α := -no_zero_divisors.to_cancel_monoid_with_zero - -variables [comm_ring α] [is_domain α] - -@[priority 100] -- see Note [lower instance priority] -instance is_domain.to_cancel_comm_monoid_with_zero : cancel_comm_monoid_with_zero α := -no_zero_divisors.to_cancel_comm_monoid_with_zero - -end is_domain - namespace semiconj_by @[simp] lemma add_right [distrib R] {a x y x' y' : R} diff --git a/src/algebra/ring/regular.lean b/src/algebra/ring/regular.lean new file mode 100644 index 0000000000000..0558de1481990 --- /dev/null +++ b/src/algebra/ring/regular.lean @@ -0,0 +1,77 @@ +/- +Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland +-/ +import algebra.ring.basic +import algebra.regular.basic + +/-! +# Lemmas about regular elements in rings. +-/ + +variables {α : Type*} + +/-- Left `mul` by a `k : α` over `[ring α]` is injective, if `k` is not a zero divisor. +The typeclass that restricts all terms of `α` to have this property is `no_zero_divisors`. -/ +lemma is_left_regular_of_non_zero_divisor [non_unital_non_assoc_ring α] (k : α) + (h : ∀ (x : α), k * x = 0 → x = 0) : is_left_regular k := +begin + refine λ x y (h' : k * x = k * y), sub_eq_zero.mp (h _ _), + rw [mul_sub, sub_eq_zero, h'] +end + +/-- Right `mul` by a `k : α` over `[ring α]` is injective, if `k` is not a zero divisor. +The typeclass that restricts all terms of `α` to have this property is `no_zero_divisors`. -/ +lemma is_right_regular_of_non_zero_divisor [non_unital_non_assoc_ring α] (k : α) + (h : ∀ (x : α), x * k = 0 → x = 0) : is_right_regular k := +begin + refine λ x y (h' : x * k = y * k), sub_eq_zero.mp (h _ _), + rw [sub_mul, sub_eq_zero, h'] +end + +lemma is_regular_of_ne_zero' [non_unital_non_assoc_ring α] [no_zero_divisors α] {k : α} + (hk : k ≠ 0) : is_regular k := +⟨is_left_regular_of_non_zero_divisor k + (λ x h, (no_zero_divisors.eq_zero_or_eq_zero_of_mul_eq_zero h).resolve_left hk), + is_right_regular_of_non_zero_divisor k + (λ x h, (no_zero_divisors.eq_zero_or_eq_zero_of_mul_eq_zero h).resolve_right hk)⟩ + +lemma is_regular_iff_ne_zero' [nontrivial α] [non_unital_non_assoc_ring α] [no_zero_divisors α] + {k : α} : is_regular k ↔ k ≠ 0 := +⟨λ h, by { rintro rfl, exact not_not.mpr h.left not_is_left_regular_zero }, is_regular_of_ne_zero'⟩ + +/-- A ring with no zero divisors is a `cancel_monoid_with_zero`. + +Note this is not an instance as it forms a typeclass loop. -/ +@[reducible] +def no_zero_divisors.to_cancel_monoid_with_zero [ring α] [no_zero_divisors α] : + cancel_monoid_with_zero α := +{ mul_left_cancel_of_ne_zero := λ a b c ha, + @is_regular.left _ _ _ (is_regular_of_ne_zero' ha) _ _, + mul_right_cancel_of_ne_zero := λ a b c hb, + @is_regular.right _ _ _ (is_regular_of_ne_zero' hb) _ _, + .. (by apply_instance : monoid_with_zero α) } + +/-- A commutative ring with no zero divisors is a `cancel_comm_monoid_with_zero`. + +Note this is not an instance as it forms a typeclass loop. -/ +@[reducible] +def no_zero_divisors.to_cancel_comm_monoid_with_zero [comm_ring α] [no_zero_divisors α] : + cancel_comm_monoid_with_zero α := +{ .. no_zero_divisors.to_cancel_monoid_with_zero, + .. (by apply_instance : comm_monoid_with_zero α) } + +section is_domain + +@[priority 100] -- see Note [lower instance priority] +instance is_domain.to_cancel_monoid_with_zero [ring α] [is_domain α] : cancel_monoid_with_zero α := +no_zero_divisors.to_cancel_monoid_with_zero + +variables [comm_ring α] [is_domain α] + +@[priority 100] -- see Note [lower instance priority] +instance is_domain.to_cancel_comm_monoid_with_zero : cancel_comm_monoid_with_zero α := +no_zero_divisors.to_cancel_comm_monoid_with_zero + +end is_domain diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index ad162953bd48d..403a0ab879577 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -6,6 +6,7 @@ Authors: Jeremy Avigad import data.nat.pow import order.min_max import data.nat.cast +import algebra.ring.regular /-! # Basic operations on the integers diff --git a/src/data/set/intervals/instances.lean b/src/data/set/intervals/instances.lean index d15ea9da11266..ae33770106b68 100644 --- a/src/data/set/intervals/instances.lean +++ b/src/data/set/intervals/instances.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Stuart Presnell, Eric Wieser, Yaël Dillies, Patrick Massot, Scott Morrison -/ import algebra.group_power.order +import algebra.ring.regular import data.set.intervals.proj_Icc /-! From d7c8521e5ff42ce5ce89489ff4f9442ac8e39d7c Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 9 Oct 2022 11:15:16 +0000 Subject: [PATCH 644/828] chore(algebra/group_with_zero/basic): split into two files (#16866) Co-authored-by: Scott Morrison --- src/algebra/divisibility.lean | 1 + src/algebra/group_with_zero/basic.lean | 489 ----------------------- src/algebra/group_with_zero/units.lean | 514 +++++++++++++++++++++++++ src/algebra/hom/equiv.lean | 2 +- src/algebra/hom/ring.lean | 1 + src/algebra/regular/basic.lean | 1 + src/algebra/ring/basic.lean | 1 + 7 files changed, 519 insertions(+), 490 deletions(-) create mode 100644 src/algebra/group_with_zero/units.lean diff --git a/src/algebra/divisibility.lean b/src/algebra/divisibility.lean index aaf3eddfce219..a969a9f77de59 100644 --- a/src/algebra/divisibility.lean +++ b/src/algebra/divisibility.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Amelia Livingston, Neil Strickland, Aaron Anderson -/ +import algebra.hom.group import algebra.group_with_zero.basic /-! diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index b2adb88d9eab1..559c2d1dacfee 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -6,9 +6,6 @@ Authors: Johan Commelin import algebra.group.inj_surj import algebra.group_with_zero.defs import algebra.group.order_synonym -import algebra.hom.units -import logic.nontrivial -import group_theory.group_action.units /-! # Groups with an adjoined zero element @@ -306,132 +303,6 @@ protected def function.surjective.comm_monoid_with_zero [has_zero M₀'] [has_mu comm_monoid_with_zero M₀' := { .. hf.comm_monoid f one mul npow, .. hf.mul_zero_class f zero mul } -variables [monoid_with_zero M₀] - -namespace units - -/-- An element of the unit group of a nonzero monoid with zero represented as an element - of the monoid is nonzero. -/ -@[simp] lemma ne_zero [nontrivial M₀] (u : M₀ˣ) : - (u : M₀) ≠ 0 := -left_ne_zero_of_mul_eq_one u.mul_inv - --- We can't use `mul_eq_zero` + `units.ne_zero` in the next two lemmas because we don't assume --- `nonzero M₀`. - -@[simp] lemma mul_left_eq_zero (u : M₀ˣ) {a : M₀} : a * u = 0 ↔ a = 0 := -⟨λ h, by simpa using mul_eq_zero_of_left h ↑u⁻¹, λ h, mul_eq_zero_of_left h u⟩ - -@[simp] lemma mul_right_eq_zero (u : M₀ˣ) {a : M₀} : ↑u * a = 0 ↔ a = 0 := -⟨λ h, by simpa using mul_eq_zero_of_right ↑u⁻¹ h, mul_eq_zero_of_right u⟩ - -end units - -namespace is_unit - -lemma ne_zero [nontrivial M₀] {a : M₀} (ha : is_unit a) : a ≠ 0 := let ⟨u, hu⟩ := -ha in hu ▸ u.ne_zero - -lemma mul_right_eq_zero {a b : M₀} (ha : is_unit a) : a * b = 0 ↔ b = 0 := -let ⟨u, hu⟩ := ha in hu ▸ u.mul_right_eq_zero - -lemma mul_left_eq_zero {a b : M₀} (hb : is_unit b) : a * b = 0 ↔ a = 0 := -let ⟨u, hu⟩ := hb in hu ▸ u.mul_left_eq_zero - -end is_unit - -@[simp] theorem is_unit_zero_iff : is_unit (0 : M₀) ↔ (0:M₀) = 1 := -⟨λ ⟨⟨_, a, (a0 : 0 * a = 1), _⟩, rfl⟩, by rwa zero_mul at a0, - λ h, @is_unit_of_subsingleton _ _ (subsingleton_of_zero_eq_one h) 0⟩ - -@[simp] theorem not_is_unit_zero [nontrivial M₀] : ¬ is_unit (0 : M₀) := -mt is_unit_zero_iff.1 zero_ne_one - -namespace ring -open_locale classical - -/-- Introduce a function `inverse` on a monoid with zero `M₀`, which sends `x` to `x⁻¹` if `x` is -invertible and to `0` otherwise. This definition is somewhat ad hoc, but one needs a fully (rather -than partially) defined inverse function for some purposes, including for calculus. - -Note that while this is in the `ring` namespace for brevity, it requires the weaker assumption -`monoid_with_zero M₀` instead of `ring M₀`. -/ -noncomputable def inverse : M₀ → M₀ := -λ x, if h : is_unit x then ((h.unit⁻¹ : M₀ˣ) : M₀) else 0 - -/-- By definition, if `x` is invertible then `inverse x = x⁻¹`. -/ -@[simp] lemma inverse_unit (u : M₀ˣ) : inverse (u : M₀) = (u⁻¹ : M₀ˣ) := -begin - simp only [units.is_unit, inverse, dif_pos], - exact units.inv_unique rfl -end - -/-- By definition, if `x` is not invertible then `inverse x = 0`. -/ -@[simp] lemma inverse_non_unit (x : M₀) (h : ¬(is_unit x)) : inverse x = 0 := dif_neg h - -lemma mul_inverse_cancel (x : M₀) (h : is_unit x) : x * inverse x = 1 := -by { rcases h with ⟨u, rfl⟩, rw [inverse_unit, units.mul_inv], } - -lemma inverse_mul_cancel (x : M₀) (h : is_unit x) : inverse x * x = 1 := -by { rcases h with ⟨u, rfl⟩, rw [inverse_unit, units.inv_mul], } - -lemma mul_inverse_cancel_right (x y : M₀) (h : is_unit x) : y * x * inverse x = y := -by rw [mul_assoc, mul_inverse_cancel x h, mul_one] - -lemma inverse_mul_cancel_right (x y : M₀) (h : is_unit x) : y * inverse x * x = y := -by rw [mul_assoc, inverse_mul_cancel x h, mul_one] - -lemma mul_inverse_cancel_left (x y : M₀) (h : is_unit x) : x * (inverse x * y) = y := -by rw [← mul_assoc, mul_inverse_cancel x h, one_mul] - -lemma inverse_mul_cancel_left (x y : M₀) (h : is_unit x) : inverse x * (x * y) = y := -by rw [← mul_assoc, inverse_mul_cancel x h, one_mul] - -variables (M₀) - -@[simp] lemma inverse_one : inverse (1 : M₀) = 1 := -inverse_unit 1 - -@[simp] lemma inverse_zero : inverse (0 : M₀) = 0 := -by { nontriviality, exact inverse_non_unit _ not_is_unit_zero } - -variables {M₀} - -lemma mul_inverse_rev' {a b : M₀} (h : commute a b) : inverse (a * b) = inverse b * inverse a := -begin - by_cases hab : is_unit (a * b), - { obtain ⟨⟨a, rfl⟩, b, rfl⟩ := h.is_unit_mul_iff.mp hab, - rw [←units.coe_mul, inverse_unit, inverse_unit, inverse_unit, ←units.coe_mul, - mul_inv_rev], }, - obtain ha | hb := not_and_distrib.mp (mt h.is_unit_mul_iff.mpr hab), - { rw [inverse_non_unit _ hab, inverse_non_unit _ ha, mul_zero]}, - { rw [inverse_non_unit _ hab, inverse_non_unit _ hb, zero_mul]}, -end - -lemma mul_inverse_rev {M₀} [comm_monoid_with_zero M₀] (a b : M₀) : - ring.inverse (a * b) = inverse b * inverse a := -mul_inverse_rev' (commute.all _ _) - -end ring - -lemma is_unit.ring_inverse {a : M₀} : is_unit a → is_unit (ring.inverse a) -| ⟨u, hu⟩ := hu ▸ ⟨u⁻¹, (ring.inverse_unit u).symm⟩ - -@[simp] lemma is_unit_ring_inverse {a : M₀} : is_unit (ring.inverse a) ↔ is_unit a := -⟨λ h, begin - casesI subsingleton_or_nontrivial M₀, - { convert h }, - { contrapose h, - rw ring.inverse_non_unit _ h, - exact not_is_unit_zero, }, -end, is_unit.ring_inverse⟩ - -lemma commute.ring_inverse_ring_inverse {a b : M₀} (h : commute a b) : - commute (ring.inverse a) (ring.inverse b) := -(ring.mul_inverse_rev' h.symm).symm.trans $ (congr_arg _ h.symm.eq).trans $ ring.mul_inverse_rev' h - -variable (M₀) - end monoid_with_zero section cancel_monoid_with_zero @@ -597,150 +468,9 @@ instance group_with_zero.to_division_monoid : division_monoid G₀ := end group_with_zero -namespace units -variables [group_with_zero G₀] -variables {a b : G₀} - -/-- Embed a non-zero element of a `group_with_zero` into the unit group. - By combining this function with the operations on units, - or the `/ₚ` operation, it is possible to write a division - as a partial function with three arguments. -/ -def mk0 (a : G₀) (ha : a ≠ 0) : G₀ˣ := -⟨a, a⁻¹, mul_inv_cancel ha, inv_mul_cancel ha⟩ - -@[simp] lemma mk0_one (h := one_ne_zero) : - mk0 (1 : G₀) h = 1 := -by { ext, refl } - -@[simp] lemma coe_mk0 {a : G₀} (h : a ≠ 0) : (mk0 a h : G₀) = a := rfl - -@[simp] lemma mk0_coe (u : G₀ˣ) (h : (u : G₀) ≠ 0) : mk0 (u : G₀) h = u := -units.ext rfl - -@[simp] lemma mul_inv' (u : G₀ˣ) : (u : G₀) * u⁻¹ = 1 := mul_inv_cancel u.ne_zero - -@[simp] lemma inv_mul' (u : G₀ˣ) : (u⁻¹ : G₀) * u = 1 := inv_mul_cancel u.ne_zero - -@[simp] lemma mk0_inj {a b : G₀} (ha : a ≠ 0) (hb : b ≠ 0) : - units.mk0 a ha = units.mk0 b hb ↔ a = b := -⟨λ h, by injection h, λ h, units.ext h⟩ - -/-- In a group with zero, an existential over a unit can be rewritten in terms of `units.mk0`. -/ -lemma exists0 {p : G₀ˣ → Prop} : (∃ g : G₀ˣ, p g) ↔ ∃ (g : G₀) (hg : g ≠ 0), p (units.mk0 g hg) := -⟨λ ⟨g, pg⟩, ⟨g, g.ne_zero, (g.mk0_coe g.ne_zero).symm ▸ pg⟩, λ ⟨g, hg, pg⟩, ⟨units.mk0 g hg, pg⟩⟩ - -/-- An alternative version of `units.exists0`. This one is useful if Lean cannot -figure out `p` when using `units.exists0` from right to left. -/ -lemma exists0' {p : Π g : G₀, g ≠ 0 → Prop} : - (∃ (g : G₀) (hg : g ≠ 0), p g hg) ↔ ∃ g : G₀ˣ, p g g.ne_zero := -iff.trans (by simp_rw [coe_mk0]) exists0.symm - -@[simp] lemma exists_iff_ne_zero {x : G₀} : (∃ u : G₀ˣ, ↑u = x) ↔ x ≠ 0 := -by simp [exists0] - -lemma _root_.group_with_zero.eq_zero_or_unit (a : G₀) : - a = 0 ∨ ∃ u : G₀ˣ, a = u := -begin - by_cases h : a = 0, - { left, - exact h }, - { right, - simpa only [eq_comm] using units.exists_iff_ne_zero.mpr h } -end - -@[simp] lemma smul_mk0 {α : Type*} [has_smul G₀ α] {g : G₀} (hg : g ≠ 0) (a : α) : - (mk0 g hg) • a = g • a := -rfl - -end units - section group_with_zero variables [group_with_zero G₀] {a b c : G₀} -lemma is_unit.mk0 (x : G₀) (hx : x ≠ 0) : is_unit x := (units.mk0 x hx).is_unit - -lemma is_unit_iff_ne_zero : is_unit a ↔ a ≠ 0 := units.exists_iff_ne_zero - -alias is_unit_iff_ne_zero ↔ _ ne.is_unit - -attribute [protected] ne.is_unit - -@[priority 10] -- see Note [lower instance priority] -instance group_with_zero.no_zero_divisors : no_zero_divisors G₀ := -{ eq_zero_or_eq_zero_of_mul_eq_zero := λ a b h, - begin - contrapose! h, - exact ((units.mk0 a h.1) * (units.mk0 b h.2)).ne_zero - end, - .. (‹_› : group_with_zero G₀) } - -@[priority 10] -- see Note [lower instance priority] -instance group_with_zero.cancel_monoid_with_zero : cancel_monoid_with_zero G₀ := -{ mul_left_cancel_of_ne_zero := λ x y z hx h, - by rw [← inv_mul_cancel_left₀ hx y, h, inv_mul_cancel_left₀ hx z], - mul_right_cancel_of_ne_zero := λ x y z hy h, - by rw [← mul_inv_cancel_right₀ hy x, h, mul_inv_cancel_right₀ hy z], - .. (‹_› : group_with_zero G₀) } - --- Can't be put next to the other `mk0` lemmas because it depends on the --- `no_zero_divisors` instance, which depends on `mk0`. -@[simp] lemma units.mk0_mul (x y : G₀) (hxy) : - units.mk0 (x * y) hxy = - units.mk0 x (mul_ne_zero_iff.mp hxy).1 * units.mk0 y (mul_ne_zero_iff.mp hxy).2 := -by { ext, refl } - -@[simp] lemma div_self (h : a ≠ 0) : a / a = 1 := h.is_unit.div_self - -lemma eq_mul_inv_iff_mul_eq₀ (hc : c ≠ 0) : a = b * c⁻¹ ↔ a * c = b := -hc.is_unit.eq_mul_inv_iff_mul_eq - -lemma eq_inv_mul_iff_mul_eq₀ (hb : b ≠ 0) : a = b⁻¹ * c ↔ b * a = c := -hb.is_unit.eq_inv_mul_iff_mul_eq - -lemma inv_mul_eq_iff_eq_mul₀ (ha : a ≠ 0) : a⁻¹ * b = c ↔ b = a * c := -ha.is_unit.inv_mul_eq_iff_eq_mul - -lemma mul_inv_eq_iff_eq_mul₀ (hb : b ≠ 0) : a * b⁻¹ = c ↔ a = c * b := -hb.is_unit.mul_inv_eq_iff_eq_mul - -lemma mul_inv_eq_one₀ (hb : b ≠ 0) : a * b⁻¹ = 1 ↔ a = b := hb.is_unit.mul_inv_eq_one -lemma inv_mul_eq_one₀ (ha : a ≠ 0) : a⁻¹ * b = 1 ↔ a = b := ha.is_unit.inv_mul_eq_one - -lemma mul_eq_one_iff_eq_inv₀ (hb : b ≠ 0) : a * b = 1 ↔ a = b⁻¹ := hb.is_unit.mul_eq_one_iff_eq_inv -lemma mul_eq_one_iff_inv_eq₀ (ha : a ≠ 0) : a * b = 1 ↔ a⁻¹ = b := ha.is_unit.mul_eq_one_iff_inv_eq - -@[simp] lemma div_mul_cancel (a : G₀) (h : b ≠ 0) : a / b * b = a := h.is_unit.div_mul_cancel _ -@[simp] lemma mul_div_cancel (a : G₀) (h : b ≠ 0) : a * b / b = a := h.is_unit.mul_div_cancel _ - -lemma mul_one_div_cancel (h : a ≠ 0) : a * (1 / a) = 1 := h.is_unit.mul_one_div_cancel -lemma one_div_mul_cancel (h : a ≠ 0) : (1 / a) * a = 1 := h.is_unit.one_div_mul_cancel - -lemma div_left_inj' (hc : c ≠ 0) : a / c = b / c ↔ a = b := hc.is_unit.div_left_inj - -@[field_simps] lemma div_eq_iff (hb : b ≠ 0) : a / b = c ↔ a = c * b := hb.is_unit.div_eq_iff -@[field_simps] lemma eq_div_iff (hb : b ≠ 0) : c = a / b ↔ c * b = a := hb.is_unit.eq_div_iff - -lemma div_eq_iff_mul_eq (hb : b ≠ 0) : a / b = c ↔ c * b = a := hb.is_unit.div_eq_iff.trans eq_comm -lemma eq_div_iff_mul_eq (hc : c ≠ 0) : a = b / c ↔ a * c = b := hc.is_unit.eq_div_iff - -lemma div_eq_of_eq_mul (hb : b ≠ 0) : a = c * b → a / b = c := hb.is_unit.div_eq_of_eq_mul -lemma eq_div_of_mul_eq (hc : c ≠ 0) : a * c = b → a = b / c := hc.is_unit.eq_div_of_mul_eq - -lemma div_eq_one_iff_eq (hb : b ≠ 0) : a / b = 1 ↔ a = b := hb.is_unit.div_eq_one_iff_eq - -lemma div_mul_left (hb : b ≠ 0) : b / (a * b) = 1 / a := hb.is_unit.div_mul_left - -lemma mul_div_mul_right (a b : G₀) (hc : c ≠ 0) : (a * c) / (b * c) = a / b := -hc.is_unit.mul_div_mul_right _ _ - -lemma mul_mul_div (a : G₀) (hb : b ≠ 0) : a = a * b * (1 / b) := (hb.is_unit.mul_mul_div _).symm - -lemma div_div_div_cancel_right (a : G₀) (hc : c ≠ 0) : (a / c) / (b / c) = a / b := -by rw [div_div_eq_mul_div, div_mul_cancel _ hc] - -lemma div_mul_div_cancel (a : G₀) (hc : c ≠ 0) : (a / c) * (c / b) = a / b := -by rw [← mul_div_assoc, div_mul_cancel _ hc] - @[simp] lemma zero_div (a : G₀) : 0 / a = 0 := by rw [div_eq_mul_inv, zero_mul] @@ -784,12 +514,6 @@ zero. -/ @[simp] lemma div_self_mul_self (a : G₀) : a / a * a = a := by rw [div_eq_mul_inv, mul_inv_mul_self a] -lemma div_mul_cancel_of_imp {a b : G₀} (h : b = 0 → a = 0) : a / b * b = a := -classical.by_cases (λ hb : b = 0, by simp [*]) (div_mul_cancel a) - -lemma mul_div_cancel_of_imp {a b : G₀} (h : b = 0 → a = 0) : a * b / b = a := -classical.by_cases (λ hb : b = 0, by simp [*]) (mul_div_cancel a) - local attribute [simp] div_eq_mul_inv mul_comm mul_assoc mul_left_comm @[simp] lemma div_self_mul_self' (a : G₀) : a / (a * a) = a⁻¹ := @@ -805,29 +529,6 @@ by rw [inv_eq_iff_inv_eq, inv_zero, eq_comm] @[simp] lemma zero_eq_inv {a : G₀} : 0 = a⁻¹ ↔ 0 = a := eq_comm.trans $ inv_eq_zero.trans eq_comm -@[simp] theorem divp_mk0 (a : G₀) {b : G₀} (hb : b ≠ 0) : - a /ₚ units.mk0 b hb = a / b := -divp_eq_div _ _ - -lemma div_ne_zero (ha : a ≠ 0) (hb : b ≠ 0) : a / b ≠ 0 := -by { rw div_eq_mul_inv, exact mul_ne_zero ha (inv_ne_zero hb) } - -@[simp] lemma div_eq_zero_iff : a / b = 0 ↔ a = 0 ∨ b = 0:= -by simp [div_eq_mul_inv] - -lemma div_ne_zero_iff : a / b ≠ 0 ↔ a ≠ 0 ∧ b ≠ 0 := -div_eq_zero_iff.not.trans not_or_distrib - -lemma ring.inverse_eq_inv (a : G₀) : ring.inverse a = a⁻¹ := -begin - obtain rfl | ha := eq_or_ne a 0, - { simp }, - { exact ring.inverse_unit (units.mk0 a ha) } -end - -@[simp] lemma ring.inverse_eq_inv' : (ring.inverse : G₀ → G₀) = has_inv.inv := -funext ring.inverse_eq_inv - /-- Dividing `a` by the result of dividing `a` by itself results in `a` (whether or not `a` is zero). -/ @[simp] lemma div_div_self (a : G₀) : a / (a / a) = a := @@ -855,14 +556,6 @@ end group_with_zero section comm_group_with_zero -- comm variables [comm_group_with_zero G₀] {a b c d : G₀} -@[priority 10] -- see Note [lower instance priority] -instance comm_group_with_zero.cancel_comm_monoid_with_zero : cancel_comm_monoid_with_zero G₀ := -{ ..group_with_zero.cancel_monoid_with_zero, ..comm_group_with_zero.to_comm_monoid_with_zero G₀ } - -@[priority 100] -- See note [lower instance priority] -instance comm_group_with_zero.to_division_comm_monoid : division_comm_monoid G₀ := -{ ..‹comm_group_with_zero G₀›, ..group_with_zero.to_division_monoid } - /-- Pullback a `comm_group_with_zero` class along an injective function. See note [reducible non-instances]. -/ @[reducible] @@ -885,193 +578,11 @@ protected def function.surjective.comm_group_with_zero [has_zero G₀'] [has_mul comm_group_with_zero G₀' := { .. hf.group_with_zero h01 f zero one mul inv div npow zpow, .. hf.comm_semigroup f mul } -lemma div_mul_right (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := ha.is_unit.div_mul_right _ - -lemma mul_div_cancel_left_of_imp {a b : G₀} (h : a = 0 → b = 0) : a * b / a = b := -by rw [mul_comm, mul_div_cancel_of_imp h] - -lemma mul_div_cancel_left (b : G₀) (ha : a ≠ 0) : a * b / a = b := ha.is_unit.mul_div_cancel_left _ - -lemma mul_div_cancel_of_imp' {a b : G₀} (h : b = 0 → a = 0) : b * (a / b) = a := -by rw [mul_comm, div_mul_cancel_of_imp h] - -lemma mul_div_cancel' (a : G₀) (hb : b ≠ 0) : b * (a / b) = a := hb.is_unit.mul_div_cancel' _ - -lemma mul_div_mul_left (a b : G₀) (hc : c ≠ 0) : (c * a) / (c * b) = a / b := -hc.is_unit.mul_div_mul_left _ _ - -lemma mul_eq_mul_of_div_eq_div (a : G₀) {b : G₀} (c : G₀) {d : G₀} (hb : b ≠ 0) (hd : d ≠ 0) - (h : a / b = c / d) : a * d = c * b := -by rw [←mul_one a, ←div_self hb, ←mul_comm_div, h, div_mul_eq_mul_div, div_mul_cancel _ hd] - -@[field_simps] lemma div_eq_div_iff (hb : b ≠ 0) (hd : d ≠ 0) : a / b = c / d ↔ a * d = c * b := -hb.is_unit.div_eq_div_iff hd.is_unit - -lemma div_div_cancel' (ha : a ≠ 0) : a / (a / b) = b := ha.is_unit.div_div_cancel - -lemma div_helper (b : G₀) (h : a ≠ 0) : 1 / (a * b) * a = 1 / b := -by rw [div_mul_eq_mul_div, one_mul, div_mul_right _ h] - lemma div_mul_eq_mul_div₀ (a b c : G₀) : (a / c) * b = a * b / c := by simp_rw [div_eq_mul_inv, mul_assoc, mul_comm c⁻¹] end comm_group_with_zero -namespace semiconj_by - -@[simp] lemma zero_right [mul_zero_class G₀] (a : G₀) : semiconj_by a 0 0 := -by simp only [semiconj_by, mul_zero, zero_mul] - -@[simp] lemma zero_left [mul_zero_class G₀] (x y : G₀) : semiconj_by 0 x y := -by simp only [semiconj_by, mul_zero, zero_mul] - -variables [group_with_zero G₀] {a x y x' y' : G₀} - -@[simp] lemma inv_symm_left_iff₀ : semiconj_by a⁻¹ x y ↔ semiconj_by a y x := -classical.by_cases - (λ ha : a = 0, by simp only [ha, inv_zero, semiconj_by.zero_left]) - (λ ha, @units_inv_symm_left_iff _ _ (units.mk0 a ha) _ _) - -lemma inv_symm_left₀ (h : semiconj_by a x y) : semiconj_by a⁻¹ y x := -semiconj_by.inv_symm_left_iff₀.2 h - -lemma inv_right₀ (h : semiconj_by a x y) : semiconj_by a x⁻¹ y⁻¹ := -begin - by_cases ha : a = 0, - { simp only [ha, zero_left] }, - by_cases hx : x = 0, - { subst x, - simp only [semiconj_by, mul_zero, @eq_comm _ _ (y * a), mul_eq_zero] at h, - simp [h.resolve_right ha] }, - { have := mul_ne_zero ha hx, - rw [h.eq, mul_ne_zero_iff] at this, - exact @units_inv_right _ _ _ (units.mk0 x hx) (units.mk0 y this.1) h }, -end - -@[simp] lemma inv_right_iff₀ : semiconj_by a x⁻¹ y⁻¹ ↔ semiconj_by a x y := -⟨λ h, inv_inv x ▸ inv_inv y ▸ h.inv_right₀, inv_right₀⟩ - -lemma div_right (h : semiconj_by a x y) (h' : semiconj_by a x' y') : - semiconj_by a (x / x') (y / y') := -by { rw [div_eq_mul_inv, div_eq_mul_inv], exact h.mul_right h'.inv_right₀ } - -end semiconj_by - -namespace commute - -@[simp] theorem zero_right [mul_zero_class G₀] (a : G₀) :commute a 0 := semiconj_by.zero_right a -@[simp] theorem zero_left [mul_zero_class G₀] (a : G₀) : commute 0 a := semiconj_by.zero_left a a - -variables [group_with_zero G₀] {a b c : G₀} - -@[simp] theorem inv_left_iff₀ : commute a⁻¹ b ↔ commute a b := -semiconj_by.inv_symm_left_iff₀ - -theorem inv_left₀ (h : commute a b) : commute a⁻¹ b := inv_left_iff₀.2 h - -@[simp] theorem inv_right_iff₀ : commute a b⁻¹ ↔ commute a b := -semiconj_by.inv_right_iff₀ - -theorem inv_right₀ (h : commute a b) : commute a b⁻¹ := inv_right_iff₀.2 h - -@[simp] theorem div_right (hab : commute a b) (hac : commute a c) : - commute a (b / c) := -hab.div_right hac - -@[simp] theorem div_left (hac : commute a c) (hbc : commute b c) : - commute (a / b) c := -by { rw div_eq_mul_inv, exact hac.mul_left hbc.inv_left₀ } - -end commute - -section monoid_with_zero -variables [group_with_zero G₀] [monoid_with_zero M₀] [nontrivial M₀] - [monoid_with_zero M₀'] [monoid_with_zero_hom_class F G₀ M₀] - [monoid_with_zero_hom_class F' G₀ M₀'] (f : F) {a : G₀} -include M₀ - -lemma map_ne_zero : f a ≠ 0 ↔ a ≠ 0 := -⟨λ hfa ha, hfa $ ha.symm ▸ map_zero f, λ ha, ((is_unit.mk0 a ha).map f).ne_zero⟩ - -@[simp] lemma map_eq_zero : f a = 0 ↔ a = 0 := not_iff_not.1 (map_ne_zero f) - -omit M₀ -include M₀' - -lemma eq_on_inv₀ (f g : F') (h : f a = g a) : f a⁻¹ = g a⁻¹ := -begin - rcases eq_or_ne a 0 with rfl|ha, - { rw [inv_zero, map_zero, map_zero] }, - { exact (is_unit.mk0 a ha).eq_on_inv f g h } -end - -end monoid_with_zero - -section group_with_zero - -variables [group_with_zero G₀] [group_with_zero G₀'] [monoid_with_zero_hom_class F G₀ G₀'] - (f : F) (a b : G₀) -include G₀' - -/-- A monoid homomorphism between groups with zeros sending `0` to `0` sends `a⁻¹` to `(f a)⁻¹`. -/ -@[simp] lemma map_inv₀ : f a⁻¹ = (f a)⁻¹ := -begin - by_cases h : a = 0, by simp [h], - apply eq_inv_of_mul_eq_one_left, - rw [← map_mul, inv_mul_cancel h, map_one] -end - -@[simp] lemma map_div₀ : f (a / b) = f a / f b := map_div' f (map_inv₀ f) a b - -end group_with_zero - -/-- We define the inverse as a `monoid_with_zero_hom` by extending the inverse map by zero -on non-units. -/ -noncomputable -def monoid_with_zero.inverse {M : Type*} [comm_monoid_with_zero M] : - M →*₀ M := -{ to_fun := ring.inverse, - map_zero' := ring.inverse_zero _, - map_one' := ring.inverse_one _, - map_mul' := λ x y, (ring.mul_inverse_rev x y).trans (mul_comm _ _) } - -@[simp] -lemma monoid_with_zero.coe_inverse {M : Type*} [comm_monoid_with_zero M] : - (monoid_with_zero.inverse : M → M) = ring.inverse := rfl - -@[simp] -lemma monoid_with_zero.inverse_apply {M : Type*} [comm_monoid_with_zero M] (a : M) : - monoid_with_zero.inverse a = ring.inverse a := rfl - -/-- Inversion on a commutative group with zero, considered as a monoid with zero homomorphism. -/ -def inv_monoid_with_zero_hom {G₀ : Type*} [comm_group_with_zero G₀] : G₀ →*₀ G₀ := -{ map_zero' := inv_zero, - ..inv_monoid_hom } - -section noncomputable_defs - -variables {M : Type*} [nontrivial M] - -/-- Constructs a `group_with_zero` structure on a `monoid_with_zero` - consisting only of units and 0. -/ -noncomputable def group_with_zero_of_is_unit_or_eq_zero [hM : monoid_with_zero M] - (h : ∀ (a : M), is_unit a ∨ a = 0) : group_with_zero M := -{ inv := λ a, if h0 : a = 0 then 0 else ↑((h a).resolve_right h0).unit⁻¹, - inv_zero := dif_pos rfl, - mul_inv_cancel := λ a h0, by - { change a * (if h0 : a = 0 then 0 else ↑((h a).resolve_right h0).unit⁻¹) = 1, - rw [dif_neg h0, units.mul_inv_eq_iff_eq_mul, one_mul, is_unit.unit_spec] }, - exists_pair_ne := nontrivial.exists_pair_ne, -.. hM } - -/-- Constructs a `comm_group_with_zero` structure on a `comm_monoid_with_zero` - consisting only of units and 0. -/ -noncomputable def comm_group_with_zero_of_is_unit_or_eq_zero [hM : comm_monoid_with_zero M] - (h : ∀ (a : M), is_unit a ∨ a = 0) : comm_group_with_zero M := -{ .. (group_with_zero_of_is_unit_or_eq_zero h), .. hM } - -end noncomputable_defs - /-! ### Order dual -/ open order_dual diff --git a/src/algebra/group_with_zero/units.lean b/src/algebra/group_with_zero/units.lean new file mode 100644 index 0000000000000..00589fb05a5f1 --- /dev/null +++ b/src/algebra/group_with_zero/units.lean @@ -0,0 +1,514 @@ +/- +Copyright (c) 2020 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ +import algebra.group_with_zero.basic +import algebra.hom.units +import group_theory.group_action.units + +/-! +# Lemmas about units in a `monoid_with_zero` or a `group_with_zero`. + +We also define `ring.inverse`, a globally defined function on any ring +(in fact any `monoid_with_zero`), which inverts units and sends non-units to zero. +-/ + +variables {α M₀ G₀ M₀' G₀' F F' : Type*} + +variables [monoid_with_zero M₀] + +namespace units + +/-- An element of the unit group of a nonzero monoid with zero represented as an element + of the monoid is nonzero. -/ +@[simp] lemma ne_zero [nontrivial M₀] (u : M₀ˣ) : + (u : M₀) ≠ 0 := +left_ne_zero_of_mul_eq_one u.mul_inv + +-- We can't use `mul_eq_zero` + `units.ne_zero` in the next two lemmas because we don't assume +-- `nonzero M₀`. + +@[simp] lemma mul_left_eq_zero (u : M₀ˣ) {a : M₀} : a * u = 0 ↔ a = 0 := +⟨λ h, by simpa using mul_eq_zero_of_left h ↑u⁻¹, λ h, mul_eq_zero_of_left h u⟩ + +@[simp] lemma mul_right_eq_zero (u : M₀ˣ) {a : M₀} : ↑u * a = 0 ↔ a = 0 := +⟨λ h, by simpa using mul_eq_zero_of_right ↑u⁻¹ h, mul_eq_zero_of_right u⟩ + +end units + +namespace is_unit + +lemma ne_zero [nontrivial M₀] {a : M₀} (ha : is_unit a) : a ≠ 0 := let ⟨u, hu⟩ := +ha in hu ▸ u.ne_zero + +lemma mul_right_eq_zero {a b : M₀} (ha : is_unit a) : a * b = 0 ↔ b = 0 := +let ⟨u, hu⟩ := ha in hu ▸ u.mul_right_eq_zero + +lemma mul_left_eq_zero {a b : M₀} (hb : is_unit b) : a * b = 0 ↔ a = 0 := +let ⟨u, hu⟩ := hb in hu ▸ u.mul_left_eq_zero + +end is_unit + +@[simp] theorem is_unit_zero_iff : is_unit (0 : M₀) ↔ (0:M₀) = 1 := +⟨λ ⟨⟨_, a, (a0 : 0 * a = 1), _⟩, rfl⟩, by rwa zero_mul at a0, + λ h, @is_unit_of_subsingleton _ _ (subsingleton_of_zero_eq_one h) 0⟩ + +@[simp] theorem not_is_unit_zero [nontrivial M₀] : ¬ is_unit (0 : M₀) := +mt is_unit_zero_iff.1 zero_ne_one + +namespace ring +open_locale classical + +/-- Introduce a function `inverse` on a monoid with zero `M₀`, which sends `x` to `x⁻¹` if `x` is +invertible and to `0` otherwise. This definition is somewhat ad hoc, but one needs a fully (rather +than partially) defined inverse function for some purposes, including for calculus. + +Note that while this is in the `ring` namespace for brevity, it requires the weaker assumption +`monoid_with_zero M₀` instead of `ring M₀`. -/ +noncomputable def inverse : M₀ → M₀ := +λ x, if h : is_unit x then ((h.unit⁻¹ : M₀ˣ) : M₀) else 0 + +/-- By definition, if `x` is invertible then `inverse x = x⁻¹`. -/ +@[simp] lemma inverse_unit (u : M₀ˣ) : inverse (u : M₀) = (u⁻¹ : M₀ˣ) := +begin + simp only [units.is_unit, inverse, dif_pos], + exact units.inv_unique rfl +end + +/-- By definition, if `x` is not invertible then `inverse x = 0`. -/ +@[simp] lemma inverse_non_unit (x : M₀) (h : ¬(is_unit x)) : inverse x = 0 := dif_neg h + +lemma mul_inverse_cancel (x : M₀) (h : is_unit x) : x * inverse x = 1 := +by { rcases h with ⟨u, rfl⟩, rw [inverse_unit, units.mul_inv], } + +lemma inverse_mul_cancel (x : M₀) (h : is_unit x) : inverse x * x = 1 := +by { rcases h with ⟨u, rfl⟩, rw [inverse_unit, units.inv_mul], } + +lemma mul_inverse_cancel_right (x y : M₀) (h : is_unit x) : y * x * inverse x = y := +by rw [mul_assoc, mul_inverse_cancel x h, mul_one] + +lemma inverse_mul_cancel_right (x y : M₀) (h : is_unit x) : y * inverse x * x = y := +by rw [mul_assoc, inverse_mul_cancel x h, mul_one] + +lemma mul_inverse_cancel_left (x y : M₀) (h : is_unit x) : x * (inverse x * y) = y := +by rw [← mul_assoc, mul_inverse_cancel x h, one_mul] + +lemma inverse_mul_cancel_left (x y : M₀) (h : is_unit x) : inverse x * (x * y) = y := +by rw [← mul_assoc, inverse_mul_cancel x h, one_mul] + +variables (M₀) + +@[simp] lemma inverse_one : inverse (1 : M₀) = 1 := +inverse_unit 1 + +@[simp] lemma inverse_zero : inverse (0 : M₀) = 0 := +by { nontriviality, exact inverse_non_unit _ not_is_unit_zero } + +variables {M₀} + +lemma mul_inverse_rev' {a b : M₀} (h : commute a b) : inverse (a * b) = inverse b * inverse a := +begin + by_cases hab : is_unit (a * b), + { obtain ⟨⟨a, rfl⟩, b, rfl⟩ := h.is_unit_mul_iff.mp hab, + rw [←units.coe_mul, inverse_unit, inverse_unit, inverse_unit, ←units.coe_mul, + mul_inv_rev], }, + obtain ha | hb := not_and_distrib.mp (mt h.is_unit_mul_iff.mpr hab), + { rw [inverse_non_unit _ hab, inverse_non_unit _ ha, mul_zero]}, + { rw [inverse_non_unit _ hab, inverse_non_unit _ hb, zero_mul]}, +end + +lemma mul_inverse_rev {M₀} [comm_monoid_with_zero M₀] (a b : M₀) : + ring.inverse (a * b) = inverse b * inverse a := +mul_inverse_rev' (commute.all _ _) + +end ring + +lemma is_unit.ring_inverse {a : M₀} : is_unit a → is_unit (ring.inverse a) +| ⟨u, hu⟩ := hu ▸ ⟨u⁻¹, (ring.inverse_unit u).symm⟩ + +@[simp] lemma is_unit_ring_inverse {a : M₀} : is_unit (ring.inverse a) ↔ is_unit a := +⟨λ h, begin + casesI subsingleton_or_nontrivial M₀, + { convert h }, + { contrapose h, + rw ring.inverse_non_unit _ h, + exact not_is_unit_zero, }, +end, is_unit.ring_inverse⟩ + +lemma commute.ring_inverse_ring_inverse {a b : M₀} (h : commute a b) : + commute (ring.inverse a) (ring.inverse b) := +(ring.mul_inverse_rev' h.symm).symm.trans $ (congr_arg _ h.symm.eq).trans $ ring.mul_inverse_rev' h + + +namespace units +variables [group_with_zero G₀] +variables {a b : G₀} + +/-- Embed a non-zero element of a `group_with_zero` into the unit group. + By combining this function with the operations on units, + or the `/ₚ` operation, it is possible to write a division + as a partial function with three arguments. -/ +def mk0 (a : G₀) (ha : a ≠ 0) : G₀ˣ := +⟨a, a⁻¹, mul_inv_cancel ha, inv_mul_cancel ha⟩ + +@[simp] lemma mk0_one (h := one_ne_zero) : + mk0 (1 : G₀) h = 1 := +by { ext, refl } + +@[simp] lemma coe_mk0 {a : G₀} (h : a ≠ 0) : (mk0 a h : G₀) = a := rfl + +@[simp] lemma mk0_coe (u : G₀ˣ) (h : (u : G₀) ≠ 0) : mk0 (u : G₀) h = u := +units.ext rfl + +@[simp] lemma mul_inv' (u : G₀ˣ) : (u : G₀) * u⁻¹ = 1 := mul_inv_cancel u.ne_zero + +@[simp] lemma inv_mul' (u : G₀ˣ) : (u⁻¹ : G₀) * u = 1 := inv_mul_cancel u.ne_zero + +@[simp] lemma mk0_inj {a b : G₀} (ha : a ≠ 0) (hb : b ≠ 0) : + units.mk0 a ha = units.mk0 b hb ↔ a = b := +⟨λ h, by injection h, λ h, units.ext h⟩ + +/-- In a group with zero, an existential over a unit can be rewritten in terms of `units.mk0`. -/ +lemma exists0 {p : G₀ˣ → Prop} : (∃ g : G₀ˣ, p g) ↔ ∃ (g : G₀) (hg : g ≠ 0), p (units.mk0 g hg) := +⟨λ ⟨g, pg⟩, ⟨g, g.ne_zero, (g.mk0_coe g.ne_zero).symm ▸ pg⟩, λ ⟨g, hg, pg⟩, ⟨units.mk0 g hg, pg⟩⟩ + +/-- An alternative version of `units.exists0`. This one is useful if Lean cannot +figure out `p` when using `units.exists0` from right to left. -/ +lemma exists0' {p : Π g : G₀, g ≠ 0 → Prop} : + (∃ (g : G₀) (hg : g ≠ 0), p g hg) ↔ ∃ g : G₀ˣ, p g g.ne_zero := +iff.trans (by simp_rw [coe_mk0]) exists0.symm + +@[simp] lemma exists_iff_ne_zero {x : G₀} : (∃ u : G₀ˣ, ↑u = x) ↔ x ≠ 0 := +by simp [exists0] + +lemma _root_.group_with_zero.eq_zero_or_unit (a : G₀) : + a = 0 ∨ ∃ u : G₀ˣ, a = u := +begin + by_cases h : a = 0, + { left, + exact h }, + { right, + simpa only [eq_comm] using units.exists_iff_ne_zero.mpr h } +end + +@[simp] lemma smul_mk0 {α : Type*} [has_smul G₀ α] {g : G₀} (hg : g ≠ 0) (a : α) : + (mk0 g hg) • a = g • a := +rfl + +end units + +section group_with_zero +variables [group_with_zero G₀] {a b c : G₀} + +lemma is_unit.mk0 (x : G₀) (hx : x ≠ 0) : is_unit x := (units.mk0 x hx).is_unit + +lemma is_unit_iff_ne_zero : is_unit a ↔ a ≠ 0 := units.exists_iff_ne_zero + +alias is_unit_iff_ne_zero ↔ _ ne.is_unit + +attribute [protected] ne.is_unit + +@[priority 10] -- see Note [lower instance priority] +instance group_with_zero.no_zero_divisors : no_zero_divisors G₀ := +{ eq_zero_or_eq_zero_of_mul_eq_zero := λ a b h, + begin + contrapose! h, + exact ((units.mk0 a h.1) * (units.mk0 b h.2)).ne_zero + end, + .. (‹_› : group_with_zero G₀) } + +@[priority 10] -- see Note [lower instance priority] +instance group_with_zero.cancel_monoid_with_zero : cancel_monoid_with_zero G₀ := +{ mul_left_cancel_of_ne_zero := λ x y z hx h, + by rw [← inv_mul_cancel_left₀ hx y, h, inv_mul_cancel_left₀ hx z], + mul_right_cancel_of_ne_zero := λ x y z hy h, + by rw [← mul_inv_cancel_right₀ hy x, h, mul_inv_cancel_right₀ hy z], + .. (‹_› : group_with_zero G₀) } + +-- Can't be put next to the other `mk0` lemmas because it depends on the +-- `no_zero_divisors` instance, which depends on `mk0`. +@[simp] lemma units.mk0_mul (x y : G₀) (hxy) : + units.mk0 (x * y) hxy = + units.mk0 x (mul_ne_zero_iff.mp hxy).1 * units.mk0 y (mul_ne_zero_iff.mp hxy).2 := +by { ext, refl } + +@[simp] lemma div_self (h : a ≠ 0) : a / a = 1 := h.is_unit.div_self + +lemma eq_mul_inv_iff_mul_eq₀ (hc : c ≠ 0) : a = b * c⁻¹ ↔ a * c = b := +hc.is_unit.eq_mul_inv_iff_mul_eq + +lemma eq_inv_mul_iff_mul_eq₀ (hb : b ≠ 0) : a = b⁻¹ * c ↔ b * a = c := +hb.is_unit.eq_inv_mul_iff_mul_eq + +lemma inv_mul_eq_iff_eq_mul₀ (ha : a ≠ 0) : a⁻¹ * b = c ↔ b = a * c := +ha.is_unit.inv_mul_eq_iff_eq_mul + +lemma mul_inv_eq_iff_eq_mul₀ (hb : b ≠ 0) : a * b⁻¹ = c ↔ a = c * b := +hb.is_unit.mul_inv_eq_iff_eq_mul + +lemma mul_inv_eq_one₀ (hb : b ≠ 0) : a * b⁻¹ = 1 ↔ a = b := hb.is_unit.mul_inv_eq_one +lemma inv_mul_eq_one₀ (ha : a ≠ 0) : a⁻¹ * b = 1 ↔ a = b := ha.is_unit.inv_mul_eq_one + +lemma mul_eq_one_iff_eq_inv₀ (hb : b ≠ 0) : a * b = 1 ↔ a = b⁻¹ := hb.is_unit.mul_eq_one_iff_eq_inv +lemma mul_eq_one_iff_inv_eq₀ (ha : a ≠ 0) : a * b = 1 ↔ a⁻¹ = b := ha.is_unit.mul_eq_one_iff_inv_eq + +@[simp] lemma div_mul_cancel (a : G₀) (h : b ≠ 0) : a / b * b = a := h.is_unit.div_mul_cancel _ +@[simp] lemma mul_div_cancel (a : G₀) (h : b ≠ 0) : a * b / b = a := h.is_unit.mul_div_cancel _ + +lemma mul_one_div_cancel (h : a ≠ 0) : a * (1 / a) = 1 := h.is_unit.mul_one_div_cancel +lemma one_div_mul_cancel (h : a ≠ 0) : (1 / a) * a = 1 := h.is_unit.one_div_mul_cancel + +lemma div_left_inj' (hc : c ≠ 0) : a / c = b / c ↔ a = b := hc.is_unit.div_left_inj + +@[field_simps] lemma div_eq_iff (hb : b ≠ 0) : a / b = c ↔ a = c * b := hb.is_unit.div_eq_iff +@[field_simps] lemma eq_div_iff (hb : b ≠ 0) : c = a / b ↔ c * b = a := hb.is_unit.eq_div_iff + +lemma div_eq_iff_mul_eq (hb : b ≠ 0) : a / b = c ↔ c * b = a := hb.is_unit.div_eq_iff.trans eq_comm +lemma eq_div_iff_mul_eq (hc : c ≠ 0) : a = b / c ↔ a * c = b := hc.is_unit.eq_div_iff + +lemma div_eq_of_eq_mul (hb : b ≠ 0) : a = c * b → a / b = c := hb.is_unit.div_eq_of_eq_mul +lemma eq_div_of_mul_eq (hc : c ≠ 0) : a * c = b → a = b / c := hc.is_unit.eq_div_of_mul_eq + +lemma div_eq_one_iff_eq (hb : b ≠ 0) : a / b = 1 ↔ a = b := hb.is_unit.div_eq_one_iff_eq + +lemma div_mul_left (hb : b ≠ 0) : b / (a * b) = 1 / a := hb.is_unit.div_mul_left + +lemma mul_div_mul_right (a b : G₀) (hc : c ≠ 0) : (a * c) / (b * c) = a / b := +hc.is_unit.mul_div_mul_right _ _ + +lemma mul_mul_div (a : G₀) (hb : b ≠ 0) : a = a * b * (1 / b) := (hb.is_unit.mul_mul_div _).symm + +lemma div_div_div_cancel_right (a : G₀) (hc : c ≠ 0) : (a / c) / (b / c) = a / b := +by rw [div_div_eq_mul_div, div_mul_cancel _ hc] + +lemma div_mul_div_cancel (a : G₀) (hc : c ≠ 0) : (a / c) * (c / b) = a / b := +by rw [← mul_div_assoc, div_mul_cancel _ hc] + +lemma div_mul_cancel_of_imp {a b : G₀} (h : b = 0 → a = 0) : a / b * b = a := +classical.by_cases (λ hb : b = 0, by simp [*]) (div_mul_cancel a) + +lemma mul_div_cancel_of_imp {a b : G₀} (h : b = 0 → a = 0) : a * b / b = a := +classical.by_cases (λ hb : b = 0, by simp [*]) (mul_div_cancel a) + +@[simp] theorem divp_mk0 (a : G₀) {b : G₀} (hb : b ≠ 0) : + a /ₚ units.mk0 b hb = a / b := +divp_eq_div _ _ + +lemma div_ne_zero (ha : a ≠ 0) (hb : b ≠ 0) : a / b ≠ 0 := +by { rw div_eq_mul_inv, exact mul_ne_zero ha (inv_ne_zero hb) } + +@[simp] lemma div_eq_zero_iff : a / b = 0 ↔ a = 0 ∨ b = 0:= +by simp [div_eq_mul_inv] + +lemma div_ne_zero_iff : a / b ≠ 0 ↔ a ≠ 0 ∧ b ≠ 0 := +div_eq_zero_iff.not.trans not_or_distrib + +lemma ring.inverse_eq_inv (a : G₀) : ring.inverse a = a⁻¹ := +begin + obtain rfl | ha := eq_or_ne a 0, + { simp }, + { exact ring.inverse_unit (units.mk0 a ha) } +end + +@[simp] lemma ring.inverse_eq_inv' : (ring.inverse : G₀ → G₀) = has_inv.inv := +funext ring.inverse_eq_inv + +end group_with_zero + +section comm_group_with_zero -- comm +variables [comm_group_with_zero G₀] {a b c d : G₀} + +@[priority 10] -- see Note [lower instance priority] +instance comm_group_with_zero.cancel_comm_monoid_with_zero : cancel_comm_monoid_with_zero G₀ := +{ ..group_with_zero.cancel_monoid_with_zero, ..comm_group_with_zero.to_comm_monoid_with_zero G₀ } + +@[priority 100] -- See note [lower instance priority] +instance comm_group_with_zero.to_division_comm_monoid : division_comm_monoid G₀ := +{ ..‹comm_group_with_zero G₀›, ..group_with_zero.to_division_monoid } + +lemma div_mul_right (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := ha.is_unit.div_mul_right _ + +lemma mul_div_cancel_left_of_imp {a b : G₀} (h : a = 0 → b = 0) : a * b / a = b := +by rw [mul_comm, mul_div_cancel_of_imp h] + +lemma mul_div_cancel_left (b : G₀) (ha : a ≠ 0) : a * b / a = b := ha.is_unit.mul_div_cancel_left _ + +lemma mul_div_cancel_of_imp' {a b : G₀} (h : b = 0 → a = 0) : b * (a / b) = a := +by rw [mul_comm, div_mul_cancel_of_imp h] + +lemma mul_div_cancel' (a : G₀) (hb : b ≠ 0) : b * (a / b) = a := hb.is_unit.mul_div_cancel' _ + +lemma mul_div_mul_left (a b : G₀) (hc : c ≠ 0) : (c * a) / (c * b) = a / b := +hc.is_unit.mul_div_mul_left _ _ + +lemma mul_eq_mul_of_div_eq_div (a : G₀) {b : G₀} (c : G₀) {d : G₀} (hb : b ≠ 0) (hd : d ≠ 0) + (h : a / b = c / d) : a * d = c * b := +by rw [←mul_one a, ←div_self hb, ←mul_comm_div, h, div_mul_eq_mul_div, div_mul_cancel _ hd] + +@[field_simps] lemma div_eq_div_iff (hb : b ≠ 0) (hd : d ≠ 0) : a / b = c / d ↔ a * d = c * b := +hb.is_unit.div_eq_div_iff hd.is_unit + +lemma div_div_cancel' (ha : a ≠ 0) : a / (a / b) = b := ha.is_unit.div_div_cancel + +lemma div_helper (b : G₀) (h : a ≠ 0) : 1 / (a * b) * a = 1 / b := +by rw [div_mul_eq_mul_div, one_mul, div_mul_right _ h] + +end comm_group_with_zero + + +namespace semiconj_by + +@[simp] lemma zero_right [mul_zero_class G₀] (a : G₀) : semiconj_by a 0 0 := +by simp only [semiconj_by, mul_zero, zero_mul] + +@[simp] lemma zero_left [mul_zero_class G₀] (x y : G₀) : semiconj_by 0 x y := +by simp only [semiconj_by, mul_zero, zero_mul] + +variables [group_with_zero G₀] {a x y x' y' : G₀} + +@[simp] lemma inv_symm_left_iff₀ : semiconj_by a⁻¹ x y ↔ semiconj_by a y x := +classical.by_cases + (λ ha : a = 0, by simp only [ha, inv_zero, semiconj_by.zero_left]) + (λ ha, @units_inv_symm_left_iff _ _ (units.mk0 a ha) _ _) + +lemma inv_symm_left₀ (h : semiconj_by a x y) : semiconj_by a⁻¹ y x := +semiconj_by.inv_symm_left_iff₀.2 h + +lemma inv_right₀ (h : semiconj_by a x y) : semiconj_by a x⁻¹ y⁻¹ := +begin + by_cases ha : a = 0, + { simp only [ha, zero_left] }, + by_cases hx : x = 0, + { subst x, + simp only [semiconj_by, mul_zero, @eq_comm _ _ (y * a), mul_eq_zero] at h, + simp [h.resolve_right ha] }, + { have := mul_ne_zero ha hx, + rw [h.eq, mul_ne_zero_iff] at this, + exact @units_inv_right _ _ _ (units.mk0 x hx) (units.mk0 y this.1) h }, +end + +@[simp] lemma inv_right_iff₀ : semiconj_by a x⁻¹ y⁻¹ ↔ semiconj_by a x y := +⟨λ h, inv_inv x ▸ inv_inv y ▸ h.inv_right₀, inv_right₀⟩ + +lemma div_right (h : semiconj_by a x y) (h' : semiconj_by a x' y') : + semiconj_by a (x / x') (y / y') := +by { rw [div_eq_mul_inv, div_eq_mul_inv], exact h.mul_right h'.inv_right₀ } + +end semiconj_by + +namespace commute + +@[simp] theorem zero_right [mul_zero_class G₀] (a : G₀) :commute a 0 := semiconj_by.zero_right a +@[simp] theorem zero_left [mul_zero_class G₀] (a : G₀) : commute 0 a := semiconj_by.zero_left a a + +variables [group_with_zero G₀] {a b c : G₀} + +@[simp] theorem inv_left_iff₀ : commute a⁻¹ b ↔ commute a b := +semiconj_by.inv_symm_left_iff₀ + +theorem inv_left₀ (h : commute a b) : commute a⁻¹ b := inv_left_iff₀.2 h + +@[simp] theorem inv_right_iff₀ : commute a b⁻¹ ↔ commute a b := +semiconj_by.inv_right_iff₀ + +theorem inv_right₀ (h : commute a b) : commute a b⁻¹ := inv_right_iff₀.2 h + +@[simp] theorem div_right (hab : commute a b) (hac : commute a c) : + commute a (b / c) := +hab.div_right hac + +@[simp] theorem div_left (hac : commute a c) (hbc : commute b c) : + commute (a / b) c := +by { rw div_eq_mul_inv, exact hac.mul_left hbc.inv_left₀ } + +end commute + +section monoid_with_zero +variables [group_with_zero G₀] [nontrivial M₀] + [monoid_with_zero M₀'] [monoid_with_zero_hom_class F G₀ M₀] + [monoid_with_zero_hom_class F' G₀ M₀'] (f : F) {a : G₀} +include M₀ + +lemma map_ne_zero : f a ≠ 0 ↔ a ≠ 0 := +⟨λ hfa ha, hfa $ ha.symm ▸ map_zero f, λ ha, ((is_unit.mk0 a ha).map f).ne_zero⟩ + +@[simp] lemma map_eq_zero : f a = 0 ↔ a = 0 := not_iff_not.1 (map_ne_zero f) + +omit M₀ +include M₀' + +lemma eq_on_inv₀ (f g : F') (h : f a = g a) : f a⁻¹ = g a⁻¹ := +begin + rcases eq_or_ne a 0 with rfl|ha, + { rw [inv_zero, map_zero, map_zero] }, + { exact (is_unit.mk0 a ha).eq_on_inv f g h } +end + +end monoid_with_zero + +section group_with_zero + +variables [group_with_zero G₀] [group_with_zero G₀'] [monoid_with_zero_hom_class F G₀ G₀'] + (f : F) (a b : G₀) +include G₀' + +/-- A monoid homomorphism between groups with zeros sending `0` to `0` sends `a⁻¹` to `(f a)⁻¹`. -/ +@[simp] lemma map_inv₀ : f a⁻¹ = (f a)⁻¹ := +begin + by_cases h : a = 0, by simp [h], + apply eq_inv_of_mul_eq_one_left, + rw [← map_mul, inv_mul_cancel h, map_one] +end + +@[simp] lemma map_div₀ : f (a / b) = f a / f b := map_div' f (map_inv₀ f) a b + +end group_with_zero + +/-- We define the inverse as a `monoid_with_zero_hom` by extending the inverse map by zero +on non-units. -/ +noncomputable +def monoid_with_zero.inverse {M : Type*} [comm_monoid_with_zero M] : + M →*₀ M := +{ to_fun := ring.inverse, + map_zero' := ring.inverse_zero _, + map_one' := ring.inverse_one _, + map_mul' := λ x y, (ring.mul_inverse_rev x y).trans (mul_comm _ _) } + +@[simp] +lemma monoid_with_zero.coe_inverse {M : Type*} [comm_monoid_with_zero M] : + (monoid_with_zero.inverse : M → M) = ring.inverse := rfl + +@[simp] +lemma monoid_with_zero.inverse_apply {M : Type*} [comm_monoid_with_zero M] (a : M) : + monoid_with_zero.inverse a = ring.inverse a := rfl + +/-- Inversion on a commutative group with zero, considered as a monoid with zero homomorphism. -/ +def inv_monoid_with_zero_hom {G₀ : Type*} [comm_group_with_zero G₀] : G₀ →*₀ G₀ := +{ map_zero' := inv_zero, + ..inv_monoid_hom } + +section noncomputable_defs + +open_locale classical +variables {M : Type*} [nontrivial M] + +/-- Constructs a `group_with_zero` structure on a `monoid_with_zero` + consisting only of units and 0. -/ +noncomputable def group_with_zero_of_is_unit_or_eq_zero [hM : monoid_with_zero M] + (h : ∀ (a : M), is_unit a ∨ a = 0) : group_with_zero M := +{ inv := λ a, if h0 : a = 0 then 0 else ↑((h a).resolve_right h0).unit⁻¹, + inv_zero := dif_pos rfl, + mul_inv_cancel := λ a h0, by + { change a * (if h0 : a = 0 then 0 else ↑((h a).resolve_right h0).unit⁻¹) = 1, + rw [dif_neg h0, units.mul_inv_eq_iff_eq_mul, one_mul, is_unit.unit_spec] }, + exists_pair_ne := nontrivial.exists_pair_ne, +.. hM } + +/-- Constructs a `comm_group_with_zero` structure on a `comm_monoid_with_zero` + consisting only of units and 0. -/ +noncomputable def comm_group_with_zero_of_is_unit_or_eq_zero [hM : comm_monoid_with_zero M] + (h : ∀ (a : M), is_unit a ∨ a = 0) : comm_group_with_zero M := +{ .. (group_with_zero_of_is_unit_or_eq_zero h), .. hM } + +end noncomputable_defs diff --git a/src/algebra/hom/equiv.lean b/src/algebra/hom/equiv.lean index d97a8cacf05eb..b3b89a0c53361 100644 --- a/src/algebra/hom/equiv.lean +++ b/src/algebra/hom/equiv.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ import algebra.group.type_tags -import algebra.group_with_zero.basic +import algebra.group_with_zero.units import data.pi.algebra /-! diff --git a/src/algebra/hom/ring.lean b/src/algebra/hom/ring.lean index 3ba8dfd690db1..05d874afee47e 100644 --- a/src/algebra/hom/ring.lean +++ b/src/algebra/hom/ring.lean @@ -5,6 +5,7 @@ Authors: Amelia Livingston, Jireh Loreaux -/ import algebra.ring.basic import data.pi.algebra +import algebra.hom.units /-! # Homomorphisms of semirings and rings diff --git a/src/algebra/regular/basic.lean b/src/algebra/regular/basic.lean index 231870d1342e7..2deee3838b6f4 100644 --- a/src/algebra/regular/basic.lean +++ b/src/algebra/regular/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import algebra.group.commute import algebra.order.monoid_lemmas import algebra.group_with_zero.basic import logic.embedding diff --git a/src/algebra/ring/basic.lean b/src/algebra/ring/basic.lean index 56ef280a61db4..346b69fa94b13 100644 --- a/src/algebra/ring/basic.lean +++ b/src/algebra/ring/basic.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Ne -/ import algebra.divisibility import algebra.regular.basic +import algebra.opposites /-! # Semirings and rings From 4fc47d5f294b2d2b9cc11b505e36eeb826ec2362 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 9 Oct 2022 14:41:38 +0000 Subject: [PATCH 645/828] feat(set_theory/*): Miscellaneous lemmas (#16854) A few simple lemmas about cardinals and ordinals. --- src/data/set/basic.lean | 2 ++ src/set_theory/cardinal/basic.lean | 11 +++++++++++ src/set_theory/ordinal/basic.lean | 5 +++++ 3 files changed, 18 insertions(+) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index a42e241734b0b..8503fe1dfff55 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1767,6 +1767,8 @@ begin { exact λ h, subsingleton.intro (λ a b, set_coe.ext (h a.property b.property)) } end +lemma subsingleton.coe_sort {s : set α} : s.subsingleton → subsingleton s := s.subsingleton_coe.2 + /-- The `coe_sort` of a set `s` in a subsingleton type is a subsingleton. For the corresponding result for `subtype`, see `subtype.subsingleton`. -/ instance subsingleton_coe_of_subsingleton [subsingleton α] {s : set α} : subsingleton s := diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index 58e2f04368d08..03faec351a871 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -298,6 +298,11 @@ theorem le_one_iff_subsingleton {α : Type u} : #α ≤ 1 ↔ subsingleton α := ⟨λ ⟨f⟩, ⟨λ a b, f.injective (subsingleton.elim _ _)⟩, λ ⟨h⟩, ⟨⟨λ a, punit.star, λ a b _, h _ _⟩⟩⟩ +@[simp] lemma mk_le_one_iff_set_subsingleton {s : set α} : #s ≤ 1 ↔ s.subsingleton := +le_one_iff_subsingleton.trans s.subsingleton_coe + +alias mk_le_one_iff_set_subsingleton ↔ _ _root_.set.subsingleton.cardinal_mk_le_one + instance : has_add cardinal.{u} := ⟨map₂ sum $ λ α β γ δ, equiv.sum_congr⟩ theorem add_def (α β : Type u) : #α + #β = #(α ⊕ β) := rfl @@ -1420,6 +1425,12 @@ lemma mk_le_mk_of_subset {α} {s t : set α} (h : s ⊆ t) : #s ≤ #t := lemma mk_subtype_mono {p q : α → Prop} (h : ∀ x, p x → q x) : #{x // p x} ≤ #{x // q x} := ⟨embedding_of_subset _ _ h⟩ +lemma le_mk_diff_add_mk (S T : set α) : #S ≤ #(S \ T : set α) + #T := +(mk_le_mk_of_subset $ subset_diff_union _ _).trans $ mk_union_le _ _ + +lemma mk_diff_add_mk {S T : set α} (h : T ⊆ S) : #(S \ T : set α) + #T = #S := +(mk_union_of_disjoint $ by exact disjoint_sdiff_self_left).symm.trans $ by rw diff_union_of_subset h + lemma mk_union_le_aleph_0 {α} {P Q : set α} : #((P ∪ Q : set α)) ≤ ℵ₀ ↔ #P ≤ ℵ₀ ∧ #Q ≤ ℵ₀ := by simp diff --git a/src/set_theory/ordinal/basic.lean b/src/set_theory/ordinal/basic.lean index 6180ef4ff8fcb..fc5a299311c1a 100644 --- a/src/set_theory/ordinal/basic.lean +++ b/src/set_theory/ordinal/basic.lean @@ -103,6 +103,8 @@ namespace Well_order instance : inhabited Well_order := ⟨⟨pempty, _, empty_relation.is_well_order⟩⟩ +@[simp] lemma eta (o : Well_order) : mk o.α o.r o.wo = o := by { cases o, refl } + end Well_order /-- Equivalence relation on well orders on arbitrary types in universe `u`, given by order @@ -141,6 +143,9 @@ by { cases w, refl } @[simp] theorem type_def (r) [wo : is_well_order α r] : (⟦⟨α, r, wo⟩⟧ : ordinal) = type r := rfl +@[simp] lemma type_out (o : ordinal) : ordinal.type o.out.r = o := +by rw [ordinal.type, Well_order.eta, quotient.out_eq] + theorem type_eq {α β} {r : α → α → Prop} {s : β → β → Prop} [is_well_order α r] [is_well_order β s] : type r = type s ↔ nonempty (r ≃r s) := quotient.eq From 8ef4f3fa22fbf945bb5097f30288724115e63cdf Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 9 Oct 2022 22:10:26 +0000 Subject: [PATCH 646/828] chore(algebra/order/monoid): split into smaller files (#16861) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This, along later PRs for `algebra.order.group` and `algebra.order.ring`, will replace #16792. - [x] depends on: #16172 [![Open in Gitpod](https://gitpod.io/button/open-in-gitpod.svg)](https://gitpod.io/from-referrer/) Co-authored-by: Yaël Dillies --- ...nically_ordered_comm_semiring_two_mul.lean | 2 +- ...linear_order_with_pos_mul_pos_eq_zero.lean | 2 +- src/algebra/order/group.lean | 4 + src/algebra/order/monoid.lean | 1463 ----------------- src/algebra/order/monoid/basic.lean | 171 ++ src/algebra/order/monoid/cancel.lean | 120 ++ src/algebra/order/monoid/canonical.lean | 221 +++ .../lemmas.lean} | 0 src/algebra/order/monoid/min_max.lean | 98 ++ src/algebra/order/monoid/order_dual.lean | 89 + src/algebra/order/monoid/prod.lean | 37 + src/algebra/order/monoid/to_mul_bot.lean | 42 + src/algebra/order/monoid/type_tags.lean | 108 ++ src/algebra/order/monoid/units.lean | 51 + src/algebra/order/monoid/with_top.lean | 475 ++++++ src/algebra/order/monoid/with_zero.lean | 155 ++ src/algebra/order/sub.lean | 2 +- src/algebra/order/with_zero.lean | 1 + src/algebra/order/zero_le_one.lean | 49 + src/algebra/regular/basic.lean | 2 +- 20 files changed, 1625 insertions(+), 1467 deletions(-) delete mode 100644 src/algebra/order/monoid.lean create mode 100644 src/algebra/order/monoid/basic.lean create mode 100644 src/algebra/order/monoid/cancel.lean create mode 100644 src/algebra/order/monoid/canonical.lean rename src/algebra/order/{monoid_lemmas.lean => monoid/lemmas.lean} (100%) create mode 100644 src/algebra/order/monoid/min_max.lean create mode 100644 src/algebra/order/monoid/order_dual.lean create mode 100644 src/algebra/order/monoid/prod.lean create mode 100644 src/algebra/order/monoid/to_mul_bot.lean create mode 100644 src/algebra/order/monoid/type_tags.lean create mode 100644 src/algebra/order/monoid/units.lean create mode 100644 src/algebra/order/monoid/with_top.lean create mode 100644 src/algebra/order/monoid/with_zero.lean create mode 100644 src/algebra/order/zero_le_one.lean diff --git a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean index 119c01715536e..efc8bddf63aaf 100644 --- a/counterexamples/canonically_ordered_comm_semiring_two_mul.lean +++ b/counterexamples/canonically_ordered_comm_semiring_two_mul.lean @@ -5,7 +5,7 @@ Authors: Damiano Testa -/ import data.zmod.basic import ring_theory.subsemiring.basic -import algebra.order.monoid +import algebra.order.monoid.basic /-! A `canonically_ordered_comm_semiring` with two different elements `a` and `b` such that diff --git a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean index 8484b0d9cf917..25b019fcc0266 100644 --- a/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean +++ b/counterexamples/linear_order_with_pos_mul_pos_eq_zero.lean @@ -4,7 +4,7 @@ All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Damiano Testa, Kevin Buzzard -/ -import algebra.order.monoid +import algebra.order.monoid.basic /-! An example of a `linear_ordered_comm_monoid_with_zero` in which the product of two positive diff --git a/src/algebra/order/group.lean b/src/algebra/order/group.lean index 081f861d78bae..003f998f8f090 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group.lean @@ -5,6 +5,10 @@ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import algebra.abs import algebra.order.sub +import algebra.order.monoid.min_max +import algebra.order.monoid.prod +import algebra.order.monoid.type_tags +import algebra.order.monoid.units /-! # Ordered groups diff --git a/src/algebra/order/monoid.lean b/src/algebra/order/monoid.lean deleted file mode 100644 index 5d592a6fcfc04..0000000000000 --- a/src/algebra/order/monoid.lean +++ /dev/null @@ -1,1463 +0,0 @@ -/- -Copyright (c) 2016 Jeremy Avigad. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl --/ -import algebra.group.with_one -import algebra.group.prod -import algebra.hom.equiv -import algebra.order.monoid_lemmas -import order.min_max -import order.hom.basic - -/-! -# Ordered monoids - -This file develops the basics of ordered monoids. - -## Implementation details - -Unfortunately, the number of `'` appended to lemmas in this file -may differ between the multiplicative and the additive version of a lemma. -The reason is that we did not want to change existing names in the library. --/ - -set_option old_structure_cmd true -open function - -universe u -variables {α : Type u} {β : Type*} - -/-- An ordered commutative monoid is a commutative monoid -with a partial order such that `a ≤ b → c * a ≤ c * b` (multiplication is monotone) --/ -@[protect_proj, ancestor comm_monoid partial_order] -class ordered_comm_monoid (α : Type*) extends comm_monoid α, partial_order α := -(mul_le_mul_left : ∀ a b : α, a ≤ b → ∀ c : α, c * a ≤ c * b) - -/-- An ordered (additive) commutative monoid is a commutative monoid - with a partial order such that `a ≤ b → c + a ≤ c + b` (addition is monotone) --/ -@[protect_proj, ancestor add_comm_monoid partial_order] -class ordered_add_comm_monoid (α : Type*) extends add_comm_monoid α, partial_order α := -(add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) - -attribute [to_additive] ordered_comm_monoid - -section ordered_instances - -@[to_additive] -instance ordered_comm_monoid.to_covariant_class_left (M : Type*) [ordered_comm_monoid M] : - covariant_class M M (*) (≤) := -{ elim := λ a b c bc, ordered_comm_monoid.mul_le_mul_left _ _ bc a } - -/- This instance can be proven with `by apply_instance`. However, `with_bot ℕ` does not -pick up a `covariant_class M M (function.swap (*)) (≤)` instance without it (see PR #7940). -/ -@[to_additive] -instance ordered_comm_monoid.to_covariant_class_right (M : Type*) [ordered_comm_monoid M] : - covariant_class M M (swap (*)) (≤) := -covariant_swap_mul_le_of_covariant_mul_le M - -/- This is not an instance, to avoid creating a loop in the type-class system: in a -`left_cancel_semigroup` with a `partial_order`, assuming `covariant_class M M (*) (≤)` implies -`covariant_class M M (*) (<)`, see `left_cancel_semigroup.covariant_mul_lt_of_covariant_mul_le`. -/ -@[to_additive] lemma has_mul.to_covariant_class_left - (M : Type*) [has_mul M] [partial_order M] [covariant_class M M (*) (<)] : - covariant_class M M (*) (≤) := -⟨covariant_le_of_covariant_lt _ _ _ covariant_class.elim⟩ - -/- This is not an instance, to avoid creating a loop in the type-class system: in a -`right_cancel_semigroup` with a `partial_order`, assuming `covariant_class M M (swap (*)) (<)` -implies `covariant_class M M (swap (*)) (≤)`, see -`right_cancel_semigroup.covariant_swap_mul_lt_of_covariant_swap_mul_le`. -/ -@[to_additive] lemma has_mul.to_covariant_class_right - (M : Type*) [has_mul M] [partial_order M] [covariant_class M M (swap (*)) (<)] : - covariant_class M M (swap (*)) (≤) := -⟨covariant_le_of_covariant_lt _ _ _ covariant_class.elim⟩ - -end ordered_instances - -/-- An `ordered_comm_monoid` with one-sided 'division' in the sense that -if `a ≤ b`, there is some `c` for which `a * c = b`. This is a weaker version -of the condition on canonical orderings defined by `canonically_ordered_monoid`. -/ -class has_exists_mul_of_le (α : Type u) [has_mul α] [has_le α] : Prop := -(exists_mul_of_le : ∀ {a b : α}, a ≤ b → ∃ (c : α), b = a * c) - -/-- An `ordered_add_comm_monoid` with one-sided 'subtraction' in the sense that -if `a ≤ b`, then there is some `c` for which `a + c = b`. This is a weaker version -of the condition on canonical orderings defined by `canonically_ordered_add_monoid`. -/ -class has_exists_add_of_le (α : Type u) [has_add α] [has_le α] : Prop := -(exists_add_of_le : ∀ {a b : α}, a ≤ b → ∃ (c : α), b = a + c) - -attribute [to_additive] has_exists_mul_of_le - -export has_exists_mul_of_le (exists_mul_of_le) - -export has_exists_add_of_le (exists_add_of_le) - -section has_exists_mul_of_le -variables [linear_order α] [densely_ordered α] [monoid α] [has_exists_mul_of_le α] - [covariant_class α α (*) (<)] [contravariant_class α α (*) (<)] {a b : α} - -@[to_additive] -lemma le_of_forall_one_lt_le_mul (h : ∀ ε : α, 1 < ε → a ≤ b * ε) : a ≤ b := -le_of_forall_le_of_dense $ λ x hxb, by { obtain ⟨ε, rfl⟩ := exists_mul_of_le hxb.le, - exact h _ ((lt_mul_iff_one_lt_right' b).1 hxb) } - -@[to_additive] -lemma le_of_forall_one_lt_lt_mul' (h : ∀ ε : α, 1 < ε → a < b * ε) : a ≤ b := -le_of_forall_one_lt_le_mul $ λ ε hε, (h _ hε).le - -@[to_additive] -lemma le_iff_forall_one_lt_lt_mul' : a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε := -⟨λ h ε, lt_mul_of_le_of_one_lt h, le_of_forall_one_lt_lt_mul'⟩ - -end has_exists_mul_of_le - -/-- A linearly ordered additive commutative monoid. -/ -@[protect_proj, ancestor linear_order ordered_add_comm_monoid] -class linear_ordered_add_comm_monoid (α : Type*) - extends linear_order α, ordered_add_comm_monoid α. - -/-- A linearly ordered commutative monoid. -/ -@[protect_proj, ancestor linear_order ordered_comm_monoid, to_additive] -class linear_ordered_comm_monoid (α : Type*) - extends linear_order α, ordered_comm_monoid α. - -/-- Typeclass for expressing that the `0` of a type is less or equal to its `1`. -/ -class zero_le_one_class (α : Type*) [has_zero α] [has_one α] [has_le α] := -(zero_le_one : (0 : α) ≤ 1) - -@[simp] lemma zero_le_one [has_zero α] [has_one α] [has_le α] [zero_le_one_class α] : (0 : α) ≤ 1 := -zero_le_one_class.zero_le_one - -/- `zero_le_one` with an explicit type argument. -/ -lemma zero_le_one' (α) [has_zero α] [has_one α] [has_le α] [zero_le_one_class α] : (0 : α) ≤ 1 := -zero_le_one - -lemma zero_le_two [preorder α] [has_one α] [add_zero_class α] [zero_le_one_class α] - [covariant_class α α (+) (≤)] : (0 : α) ≤ 2 := -add_nonneg zero_le_one zero_le_one - -lemma zero_le_three [preorder α] [has_one α] [add_zero_class α] [zero_le_one_class α] - [covariant_class α α (+) (≤)] : (0 : α) ≤ 3 := -add_nonneg zero_le_two zero_le_one - -lemma zero_le_four [preorder α] [has_one α] [add_zero_class α] [zero_le_one_class α] - [covariant_class α α (+) (≤)] : (0 : α) ≤ 4 := -add_nonneg zero_le_two zero_le_two - -lemma one_le_two [has_le α] [has_one α] [add_zero_class α] [zero_le_one_class α] - [covariant_class α α (+) (≤)] : (1 : α) ≤ 2 := -calc 1 = 1 + 0 : (add_zero 1).symm - ... ≤ 1 + 1 : add_le_add_left zero_le_one _ - -lemma one_le_two' [has_le α] [has_one α] [add_zero_class α] [zero_le_one_class α] - [covariant_class α α (swap (+)) (≤)] : (1 : α) ≤ 2 := -calc 1 = 0 + 1 : (zero_add 1).symm - ... ≤ 1 + 1 : add_le_add_right zero_le_one _ - -/-- A linearly ordered commutative monoid with a zero element. -/ -class linear_ordered_comm_monoid_with_zero (α : Type*) - extends linear_ordered_comm_monoid α, comm_monoid_with_zero α := -(zero_le_one : (0 : α) ≤ 1) - -@[priority 100] -instance linear_ordered_comm_monoid_with_zero.zero_le_one_class - [h : linear_ordered_comm_monoid_with_zero α] : zero_le_one_class α := -{ ..h } - -/-- A linearly ordered commutative monoid with an additively absorbing `⊤` element. - Instances should include number systems with an infinite element adjoined.` -/ -@[protect_proj, ancestor linear_ordered_add_comm_monoid has_top] -class linear_ordered_add_comm_monoid_with_top (α : Type*) - extends linear_ordered_add_comm_monoid α, has_top α := -(le_top : ∀ x : α, x ≤ ⊤) -(top_add' : ∀ x : α, ⊤ + x = ⊤) - -@[priority 100] -- see Note [lower instance priority] -instance linear_ordered_add_comm_monoid_with_top.to_order_top (α : Type u) - [h : linear_ordered_add_comm_monoid_with_top α] : order_top α := -{ ..h } - -section linear_ordered_add_comm_monoid_with_top -variables [linear_ordered_add_comm_monoid_with_top α] {a b : α} - -@[simp] -lemma top_add (a : α) : ⊤ + a = ⊤ := linear_ordered_add_comm_monoid_with_top.top_add' a - -@[simp] -lemma add_top (a : α) : a + ⊤ = ⊤ := -trans (add_comm _ _) (top_add _) - -end linear_ordered_add_comm_monoid_with_top - -/-- Pullback an `ordered_comm_monoid` under an injective map. -See note [reducible non-instances]. -/ -@[reducible, to_additive function.injective.ordered_add_comm_monoid -"Pullback an `ordered_add_comm_monoid` under an injective map."] -def function.injective.ordered_comm_monoid [ordered_comm_monoid α] {β : Type*} - [has_one β] [has_mul β] [has_pow β ℕ] - (f : β → α) (hf : function.injective f) (one : f 1 = 1) - (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) : - ordered_comm_monoid β := -{ mul_le_mul_left := λ a b ab c, show f (c * a) ≤ f (c * b), by - { rw [mul, mul], apply mul_le_mul_left', exact ab }, - ..partial_order.lift f hf, - ..hf.comm_monoid f one mul npow } - -/-- Pullback a `linear_ordered_comm_monoid` under an injective map. -See note [reducible non-instances]. -/ -@[reducible, to_additive function.injective.linear_ordered_add_comm_monoid -"Pullback an `ordered_add_comm_monoid` under an injective map."] -def function.injective.linear_ordered_comm_monoid [linear_ordered_comm_monoid α] {β : Type*} - [has_one β] [has_mul β] [has_pow β ℕ] [has_sup β] [has_inf β] - (f : β → α) (hf : function.injective f) (one : f 1 = 1) - (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_comm_monoid β := -{ .. hf.ordered_comm_monoid f one mul npow, - .. linear_order.lift f hf hsup hinf } - -lemma bit0_pos [ordered_add_comm_monoid α] {a : α} (h : 0 < a) : 0 < bit0 a := -add_pos' h h - -namespace units - -@[to_additive] -instance [monoid α] [preorder α] : preorder αˣ := -preorder.lift (coe : αˣ → α) - -@[simp, norm_cast, to_additive] -theorem coe_le_coe [monoid α] [preorder α] {a b : αˣ} : - (a : α) ≤ b ↔ a ≤ b := iff.rfl - -@[simp, norm_cast, to_additive] -theorem coe_lt_coe [monoid α] [preorder α] {a b : αˣ} : - (a : α) < b ↔ a < b := iff.rfl - -@[to_additive] -instance [monoid α] [partial_order α] : partial_order αˣ := -partial_order.lift coe units.ext - -@[to_additive] -instance [monoid α] [linear_order α] : linear_order αˣ := -linear_order.lift' coe units.ext - -/-- `coe : αˣ → α` as an order embedding. -/ -@[to_additive "`coe : add_units α → α` as an order embedding.", simps { fully_applied := ff }] -def order_embedding_coe [monoid α] [linear_order α] : αˣ ↪o α := ⟨⟨coe, ext⟩, λ _ _, iff.rfl⟩ - -@[simp, norm_cast, to_additive] -theorem max_coe [monoid α] [linear_order α] {a b : αˣ} : - (↑(max a b) : α) = max a b := -monotone.map_max order_embedding_coe.monotone - -@[simp, norm_cast, to_additive] -theorem min_coe [monoid α] [linear_order α] {a b : αˣ} : - (↑(min a b) : α) = min a b := -monotone.map_min order_embedding_coe.monotone - -end units - -namespace with_zero - -local attribute [semireducible] with_zero - -instance [preorder α] : preorder (with_zero α) := with_bot.preorder - -instance [partial_order α] : partial_order (with_zero α) := with_bot.partial_order - -instance [preorder α] : order_bot (with_zero α) := with_bot.order_bot - -lemma zero_le [preorder α] (a : with_zero α) : 0 ≤ a := bot_le - -lemma zero_lt_coe [preorder α] (a : α) : (0 : with_zero α) < a := with_bot.bot_lt_coe a - -lemma zero_eq_bot [preorder α] : (0 : with_zero α) = ⊥ := rfl - -@[simp, norm_cast] lemma coe_lt_coe [preorder α] {a b : α} : (a : with_zero α) < b ↔ a < b := -with_bot.coe_lt_coe - -@[simp, norm_cast] lemma coe_le_coe [preorder α] {a b : α} : (a : with_zero α) ≤ b ↔ a ≤ b := -with_bot.coe_le_coe - -instance [lattice α] : lattice (with_zero α) := with_bot.lattice - -instance [linear_order α] : linear_order (with_zero α) := with_bot.linear_order - -instance covariant_class_mul_le {α : Type u} [has_mul α] [preorder α] - [covariant_class α α (*) (≤)] : - covariant_class (with_zero α) (with_zero α) (*) (≤) := -begin - refine ⟨λ a b c hbc, _⟩, - induction a using with_zero.rec_zero_coe, { exact zero_le _ }, - induction b using with_zero.rec_zero_coe, { exact zero_le _ }, - rcases with_bot.coe_le_iff.1 hbc with ⟨c, rfl, hbc'⟩, - rw [← coe_mul, ← coe_mul, coe_le_coe], - exact mul_le_mul_left' hbc' a -end - -instance contravariant_class_mul_lt {α : Type u} [has_mul α] [partial_order α] - [contravariant_class α α (*) (<)] : - contravariant_class (with_zero α) (with_zero α) (*) (<) := -begin - refine ⟨λ a b c h, _⟩, - have := ((zero_le _).trans_lt h).ne', - lift a to α using left_ne_zero_of_mul this, - lift c to α using right_ne_zero_of_mul this, - induction b using with_zero.rec_zero_coe, - exacts [zero_lt_coe _, coe_lt_coe.mpr (lt_of_mul_lt_mul_left' $ coe_lt_coe.mp h)] -end - -@[simp] lemma le_max_iff [linear_order α] {a b c : α} : - (a : with_zero α) ≤ max b c ↔ a ≤ max b c := -by simp only [with_zero.coe_le_coe, le_max_iff] - -@[simp] lemma min_le_iff [linear_order α] {a b c : α} : - min (a : with_zero α) b ≤ c ↔ min a b ≤ c := -by simp only [with_zero.coe_le_coe, min_le_iff] - -instance [ordered_comm_monoid α] : ordered_comm_monoid (with_zero α) := -{ mul_le_mul_left := λ _ _, mul_le_mul_left', - ..with_zero.comm_monoid_with_zero, - ..with_zero.partial_order } - -protected lemma covariant_class_add_le [add_zero_class α] [preorder α] - [covariant_class α α (+) (≤)] (h : ∀ a : α, 0 ≤ a) : - covariant_class (with_zero α) (with_zero α) (+) (≤) := -begin - refine ⟨λ a b c hbc, _⟩, - induction a using with_zero.rec_zero_coe, - { rwa [zero_add, zero_add] }, - induction b using with_zero.rec_zero_coe, - { rw [add_zero], - induction c using with_zero.rec_zero_coe, - { rw [add_zero], exact le_rfl }, - { rw [← coe_add, coe_le_coe], - exact le_add_of_nonneg_right (h _) } }, - { rcases with_bot.coe_le_iff.1 hbc with ⟨c, rfl, hbc'⟩, - rw [← coe_add, ← coe_add, coe_le_coe], - exact add_le_add_left hbc' a } -end - -/- -Note 1 : the below is not an instance because it requires `zero_le`. It seems -like a rather pathological definition because α already has a zero. -Note 2 : there is no multiplicative analogue because it does not seem necessary. -Mathematicians might be more likely to use the order-dual version, where all -elements are ≤ 1 and then 1 is the top element. --/ - -/-- -If `0` is the least element in `α`, then `with_zero α` is an `ordered_add_comm_monoid`. -See note [reducible non-instances]. --/ -@[reducible] protected def ordered_add_comm_monoid [ordered_add_comm_monoid α] - (zero_le : ∀ a : α, 0 ≤ a) : ordered_add_comm_monoid (with_zero α) := -{ add_le_add_left := @add_le_add_left _ _ _ (with_zero.covariant_class_add_le zero_le), - ..with_zero.partial_order, - ..with_zero.add_comm_monoid, .. } - -end with_zero - -/-- A canonically ordered additive monoid is an ordered commutative additive monoid - in which the ordering coincides with the subtractibility relation, - which is to say, `a ≤ b` iff there exists `c` with `b = a + c`. - This is satisfied by the natural numbers, for example, but not - the integers or other nontrivial `ordered_add_comm_group`s. -/ -@[protect_proj, ancestor ordered_add_comm_monoid has_bot] -class canonically_ordered_add_monoid (α : Type*) extends ordered_add_comm_monoid α, has_bot α := -(bot_le : ∀ x : α, ⊥ ≤ x) -(exists_add_of_le : ∀ {a b : α}, a ≤ b → ∃ c, b = a + c) -(le_self_add : ∀ a b : α, a ≤ a + b) - -@[priority 100] -- see Note [lower instance priority] -instance canonically_ordered_add_monoid.to_order_bot (α : Type u) - [h : canonically_ordered_add_monoid α] : order_bot α := -{ ..h } - -/-- A canonically ordered monoid is an ordered commutative monoid - in which the ordering coincides with the divisibility relation, - which is to say, `a ≤ b` iff there exists `c` with `b = a * c`. - Examples seem rare; it seems more likely that the `order_dual` - of a naturally-occurring lattice satisfies this than the lattice - itself (for example, dual of the lattice of ideals of a PID or - Dedekind domain satisfy this; collections of all things ≤ 1 seem to - be more natural that collections of all things ≥ 1). --/ -@[protect_proj, ancestor ordered_comm_monoid has_bot, to_additive] -class canonically_ordered_monoid (α : Type*) extends ordered_comm_monoid α, has_bot α := -(bot_le : ∀ x : α, ⊥ ≤ x) -(exists_mul_of_le : ∀ {a b : α}, a ≤ b → ∃ c, b = a * c) -(le_self_mul : ∀ a b : α, a ≤ a * b) - -@[priority 100, to_additive] -- see Note [lower instance priority] -instance canonically_ordered_monoid.to_order_bot (α : Type u) - [h : canonically_ordered_monoid α] : order_bot α := -{ ..h } - -@[priority 100, to_additive] -- see Note [lower instance priority] -instance canonically_ordered_monoid.has_exists_mul_of_le (α : Type u) - [h : canonically_ordered_monoid α] : has_exists_mul_of_le α := -{ ..h } - -section canonically_ordered_monoid - -variables [canonically_ordered_monoid α] {a b c d : α} - -@[to_additive] lemma le_self_mul : a ≤ a * c := canonically_ordered_monoid.le_self_mul _ _ -@[to_additive] lemma le_mul_self : a ≤ b * a := by { rw mul_comm, exact le_self_mul } - -@[to_additive] lemma self_le_mul_right (a b : α) : a ≤ a * b := le_self_mul -@[to_additive] lemma self_le_mul_left (a b : α) : a ≤ b * a := le_mul_self - -@[to_additive] lemma le_of_mul_le_left : a * b ≤ c → a ≤ c := le_self_mul.trans -@[to_additive] lemma le_of_mul_le_right : a * b ≤ c → b ≤ c := le_mul_self.trans - -@[to_additive] -lemma le_iff_exists_mul : a ≤ b ↔ ∃ c, b = a * c := -⟨exists_mul_of_le, by { rintro ⟨c, rfl⟩, exact le_self_mul }⟩ - -@[to_additive] -lemma le_iff_exists_mul' : a ≤ b ↔ ∃ c, b = c * a := -by simpa only [mul_comm _ a] using le_iff_exists_mul - -@[simp, to_additive zero_le] lemma one_le (a : α) : 1 ≤ a := -le_iff_exists_mul.mpr ⟨a, (one_mul _).symm⟩ - -@[to_additive] lemma bot_eq_one : (⊥ : α) = 1 := -le_antisymm bot_le (one_le ⊥) - -@[simp, to_additive] lemma mul_eq_one_iff : a * b = 1 ↔ a = 1 ∧ b = 1 := -mul_eq_one_iff' (one_le _) (one_le _) - -@[simp, to_additive] lemma le_one_iff_eq_one : a ≤ 1 ↔ a = 1 := -(one_le a).le_iff_eq - -@[to_additive] lemma one_lt_iff_ne_one : 1 < a ↔ a ≠ 1 := -(one_le a).lt_iff_ne.trans ne_comm - -@[to_additive] lemma eq_one_or_one_lt : a = 1 ∨ 1 < a := -(one_le a).eq_or_lt.imp_left eq.symm - -@[simp, to_additive add_pos_iff] lemma one_lt_mul_iff : 1 < a * b ↔ 1 < a ∨ 1 < b := -by simp only [one_lt_iff_ne_one, ne.def, mul_eq_one_iff, not_and_distrib] - -@[to_additive] lemma exists_one_lt_mul_of_lt (h : a < b) : ∃ c (hc : 1 < c), a * c = b := -begin - obtain ⟨c, hc⟩ := le_iff_exists_mul.1 h.le, - refine ⟨c, one_lt_iff_ne_one.2 _, hc.symm⟩, - rintro rfl, - simpa [hc, lt_irrefl] using h -end - -@[to_additive] lemma le_mul_left (h : a ≤ c) : a ≤ b * c := -calc a = 1 * a : by simp - ... ≤ b * c : mul_le_mul' (one_le _) h - -@[to_additive] lemma le_mul_right (h : a ≤ b) : a ≤ b * c := -calc a = a * 1 : by simp - ... ≤ b * c : mul_le_mul' h (one_le _) - -@[to_additive] -lemma lt_iff_exists_mul [covariant_class α α (*) (<)] : a < b ↔ ∃ c > 1, b = a * c := -begin - simp_rw [lt_iff_le_and_ne, and_comm, le_iff_exists_mul, ← exists_and_distrib_left, exists_prop], - apply exists_congr, intro c, - rw [and.congr_left_iff, gt_iff_lt], rintro rfl, - split, - { rw [one_lt_iff_ne_one], apply mt, rintro rfl, rw [mul_one] }, - { rw [← (self_le_mul_right a c).lt_iff_ne], apply lt_mul_of_one_lt_right' } -end - -instance with_zero.has_exists_add_of_le {α} [has_add α] [preorder α] [has_exists_add_of_le α] : - has_exists_add_of_le (with_zero α) := -⟨λ a b, begin - apply with_zero.cases_on a, - { exact λ _, ⟨b, (zero_add b).symm⟩ }, - apply with_zero.cases_on b, - { exact λ b' h, (with_bot.not_coe_le_bot _ h).elim }, - rintro a' b' h, - obtain ⟨c, rfl⟩ := exists_add_of_le (with_zero.coe_le_coe.1 h), - exact ⟨c, rfl⟩, -end⟩ - --- This instance looks absurd: a monoid already has a zero -/-- Adding a new zero to a canonically ordered additive monoid produces another one. -/ -instance with_zero.canonically_ordered_add_monoid {α : Type u} [canonically_ordered_add_monoid α] : - canonically_ordered_add_monoid (with_zero α) := -{ le_self_add := λ a b, begin - apply with_zero.cases_on a, - { exact bot_le }, - apply with_zero.cases_on b, - { exact λ b', le_rfl }, - { exact λ a' b', with_zero.coe_le_coe.2 le_self_add } - end, - .. with_zero.order_bot, - .. with_zero.ordered_add_comm_monoid zero_le, ..with_zero.has_exists_add_of_le } - -end canonically_ordered_monoid - -lemma pos_of_gt {M : Type*} [canonically_ordered_add_monoid M] {n m : M} (h : n < m) : 0 < m := -lt_of_le_of_lt (zero_le _) h - -@[priority 100] instance canonically_ordered_add_monoid.zero_le_one_class {M : Type*} - [canonically_ordered_add_monoid M] [has_one M] : zero_le_one_class M := -⟨zero_le 1⟩ - -/-- A canonically linear-ordered additive monoid is a canonically ordered additive monoid - whose ordering is a linear order. -/ -@[protect_proj, ancestor canonically_ordered_add_monoid linear_order] -class canonically_linear_ordered_add_monoid (α : Type*) - extends canonically_ordered_add_monoid α, linear_order α - -/-- A canonically linear-ordered monoid is a canonically ordered monoid - whose ordering is a linear order. -/ -@[protect_proj, ancestor canonically_ordered_monoid linear_order, to_additive] -class canonically_linear_ordered_monoid (α : Type*) - extends canonically_ordered_monoid α, linear_order α - -section canonically_linear_ordered_monoid -variables [canonically_linear_ordered_monoid α] - -@[priority 100, to_additive] -- see Note [lower instance priority] -instance canonically_linear_ordered_monoid.semilattice_sup : semilattice_sup α := -{ ..linear_order.to_lattice } - -instance with_zero.canonically_linear_ordered_add_monoid - (α : Type*) [canonically_linear_ordered_add_monoid α] : - canonically_linear_ordered_add_monoid (with_zero α) := -{ .. with_zero.canonically_ordered_add_monoid, - .. with_zero.linear_order } - -@[to_additive] -lemma min_mul_distrib (a b c : α) : min a (b * c) = min a (min a b * min a c) := -begin - cases le_total a b with hb hb, - { simp [hb, le_mul_right] }, - { cases le_total a c with hc hc, - { simp [hc, le_mul_left] }, - { simp [hb, hc] } } -end - -@[to_additive] -lemma min_mul_distrib' (a b c : α) : min (a * b) c = min (min a c * min b c) c := -by simpa [min_comm _ c] using min_mul_distrib c a b - -@[simp, to_additive] -lemma one_min (a : α) : min 1 a = 1 := -min_eq_left (one_le a) - -@[simp, to_additive] -lemma min_one (a : α) : min a 1 = 1 := -min_eq_right (one_le a) - -/-- In a linearly ordered monoid, we are happy for `bot_eq_one` to be a `@[simp]` lemma. -/ -@[simp, to_additive - "In a linearly ordered monoid, we are happy for `bot_eq_zero` to be a `@[simp]` lemma"] -lemma bot_eq_one' : (⊥ : α) = 1 := -bot_eq_one - -end canonically_linear_ordered_monoid - -/-- An ordered cancellative additive commutative monoid -is an additive commutative monoid with a partial order, -in which addition is cancellative and monotone. -/ -@[protect_proj, ancestor add_comm_monoid partial_order] -class ordered_cancel_add_comm_monoid (α : Type u) extends add_comm_monoid α, partial_order α := -(add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) -(le_of_add_le_add_left : ∀ a b c : α, a + b ≤ a + c → b ≤ c) - -/-- An ordered cancellative commutative monoid -is a commutative monoid with a partial order, -in which multiplication is cancellative and monotone. -/ -@[protect_proj, ancestor comm_monoid partial_order, to_additive] -class ordered_cancel_comm_monoid (α : Type u) extends comm_monoid α, partial_order α := -(mul_le_mul_left : ∀ a b : α, a ≤ b → ∀ c : α, c * a ≤ c * b) -(le_of_mul_le_mul_left : ∀ a b c : α, a * b ≤ a * c → b ≤ c) - -section ordered_cancel_comm_monoid -variables [ordered_cancel_comm_monoid α] {a b c d : α} - -@[priority 200, to_additive] -- see Note [lower instance priority] -instance ordered_cancel_comm_monoid.to_contravariant_class_le_left : - contravariant_class α α (*) (≤) := -⟨ordered_cancel_comm_monoid.le_of_mul_le_mul_left⟩ - -@[to_additive] -lemma ordered_cancel_comm_monoid.lt_of_mul_lt_mul_left : ∀ a b c : α, a * b < a * c → b < c := -λ a b c h, lt_of_le_not_le - (ordered_cancel_comm_monoid.le_of_mul_le_mul_left a b c h.le) $ - mt (λ h, ordered_cancel_comm_monoid.mul_le_mul_left _ _ h _) (not_le_of_gt h) - -@[to_additive] -instance ordered_cancel_comm_monoid.to_contravariant_class_left - (M : Type*) [ordered_cancel_comm_monoid M] : - contravariant_class M M (*) (<) := -{ elim := λ a b c, ordered_cancel_comm_monoid.lt_of_mul_lt_mul_left _ _ _ } - -/- This instance can be proven with `by apply_instance`. However, by analogy with the -instance `ordered_cancel_comm_monoid.to_covariant_class_right` above, I imagine that without -this instance, some Type would not have a `contravariant_class M M (function.swap (*)) (<)` -instance. -/ -@[to_additive] -instance ordered_cancel_comm_monoid.to_contravariant_class_right - (M : Type*) [ordered_cancel_comm_monoid M] : - contravariant_class M M (swap (*)) (<) := -contravariant_swap_mul_lt_of_contravariant_mul_lt M - -@[priority 100, to_additive] -- see Note [lower instance priority] -instance ordered_cancel_comm_monoid.to_ordered_comm_monoid : ordered_comm_monoid α := -{ ..‹ordered_cancel_comm_monoid α› } - -@[priority 100, to_additive] -- see Note [lower instance priority] -instance ordered_cancel_comm_monoid.to_cancel_comm_monoid : cancel_comm_monoid α := -{ mul_left_cancel := λ a b c h, - (le_of_mul_le_mul_left' h.le).antisymm $ le_of_mul_le_mul_left' h.ge, - ..‹ordered_cancel_comm_monoid α› } - -/-- Pullback an `ordered_cancel_comm_monoid` under an injective map. -See note [reducible non-instances]. -/ -@[reducible, to_additive function.injective.ordered_cancel_add_comm_monoid -"Pullback an `ordered_cancel_add_comm_monoid` under an injective map."] -def function.injective.ordered_cancel_comm_monoid {β : Type*} - [has_one β] [has_mul β] [has_pow β ℕ] - (f : β → α) (hf : function.injective f) (one : f 1 = 1) - (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) : - ordered_cancel_comm_monoid β := -{ le_of_mul_le_mul_left := λ a b c (bc : f (a * b) ≤ f (a * c)), - (mul_le_mul_iff_left (f a)).mp (by rwa [← mul, ← mul]), - ..hf.ordered_comm_monoid f one mul npow } - -end ordered_cancel_comm_monoid - -/-! Some lemmas about types that have an ordering and a binary operation, with no - rules relating them. -/ -@[to_additive] -lemma fn_min_mul_fn_max {β} [linear_order α] [comm_semigroup β] (f : α → β) (n m : α) : - f (min n m) * f (max n m) = f n * f m := -by { cases le_total n m with h h; simp [h, mul_comm] } - -@[to_additive] -lemma min_mul_max [linear_order α] [comm_semigroup α] (n m : α) : - min n m * max n m = n * m := -fn_min_mul_fn_max id n m - -/-- A linearly ordered cancellative additive commutative monoid -is an additive commutative monoid with a decidable linear order -in which addition is cancellative and monotone. -/ -@[protect_proj, ancestor ordered_cancel_add_comm_monoid linear_ordered_add_comm_monoid] -class linear_ordered_cancel_add_comm_monoid (α : Type u) - extends ordered_cancel_add_comm_monoid α, linear_ordered_add_comm_monoid α - -/-- A linearly ordered cancellative commutative monoid -is a commutative monoid with a linear order -in which multiplication is cancellative and monotone. -/ -@[protect_proj, ancestor ordered_cancel_comm_monoid linear_ordered_comm_monoid, to_additive] -class linear_ordered_cancel_comm_monoid (α : Type u) - extends ordered_cancel_comm_monoid α, linear_ordered_comm_monoid α - -section covariant_class_mul_le -variables [linear_order α] - -section has_mul -variable [has_mul α] - -section left -variable [covariant_class α α (*) (≤)] - -@[to_additive] lemma min_mul_mul_left (a b c : α) : min (a * b) (a * c) = a * min b c := -(monotone_id.const_mul' a).map_min.symm - -@[to_additive] -lemma max_mul_mul_left (a b c : α) : max (a * b) (a * c) = a * max b c := -(monotone_id.const_mul' a).map_max.symm - -@[to_additive] -lemma lt_or_lt_of_mul_lt_mul [covariant_class α α (function.swap (*)) (≤)] - {a b m n : α} (h : m * n < a * b) : - m < a ∨ n < b := -by { contrapose! h, exact mul_le_mul' h.1 h.2 } - -@[to_additive] -lemma mul_lt_mul_iff_of_le_of_le - [covariant_class α α (function.swap (*)) (<)] - [covariant_class α α (*) (<)] - [covariant_class α α (function.swap (*)) (≤)] - {a b c d : α} (ac : a ≤ c) (bd : b ≤ d) : - a * b < c * d ↔ (a < c) ∨ (b < d) := -begin - refine ⟨lt_or_lt_of_mul_lt_mul, λ h, _⟩, - cases h with ha hb, - { exact mul_lt_mul_of_lt_of_le ha bd }, - { exact mul_lt_mul_of_le_of_lt ac hb } -end - -end left - -section right -variable [covariant_class α α (function.swap (*)) (≤)] - -@[to_additive] -lemma min_mul_mul_right (a b c : α) : min (a * c) (b * c) = min a b * c := -(monotone_id.mul_const' c).map_min.symm - -@[to_additive] -lemma max_mul_mul_right (a b c : α) : max (a * c) (b * c) = max a b * c := -(monotone_id.mul_const' c).map_max.symm - -end right - -end has_mul - -variable [mul_one_class α] - -@[to_additive] -lemma min_le_mul_of_one_le_right [covariant_class α α (*) (≤)] {a b : α} (hb : 1 ≤ b) : - min a b ≤ a * b := -min_le_iff.2 $ or.inl $ le_mul_of_one_le_right' hb - -@[to_additive] -lemma min_le_mul_of_one_le_left [covariant_class α α (function.swap (*)) (≤)] {a b : α} - (ha : 1 ≤ a) : min a b ≤ a * b := -min_le_iff.2 $ or.inr $ le_mul_of_one_le_left' ha - -@[to_additive] -lemma max_le_mul_of_one_le [covariant_class α α (*) (≤)] - [covariant_class α α (function.swap (*)) (≤)] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : - max a b ≤ a * b := -max_le_iff.2 ⟨le_mul_of_one_le_right' hb, le_mul_of_one_le_left' ha⟩ - -end covariant_class_mul_le - -section linear_ordered_cancel_comm_monoid -variables [linear_ordered_cancel_comm_monoid α] - -/-- Pullback a `linear_ordered_cancel_comm_monoid` under an injective map. -See note [reducible non-instances]. -/ -@[reducible, to_additive function.injective.linear_ordered_cancel_add_comm_monoid -"Pullback a `linear_ordered_cancel_add_comm_monoid` under an injective map."] -def function.injective.linear_ordered_cancel_comm_monoid {β : Type*} - [has_one β] [has_mul β] [has_pow β ℕ] [has_sup β] [has_inf β] - (f : β → α) (hf : function.injective f) (one : f 1 = 1) - (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) - (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : - linear_ordered_cancel_comm_monoid β := -{ ..hf.linear_ordered_comm_monoid f one mul npow hsup hinf, - ..hf.ordered_cancel_comm_monoid f one mul npow } - -end linear_ordered_cancel_comm_monoid - -/-! ### Order dual -/ - -namespace order_dual - -@[to_additive] -instance contravariant_class_mul_le [has_le α] [has_mul α] [c : contravariant_class α α (*) (≤)] : - contravariant_class αᵒᵈ αᵒᵈ (*) (≤) := -⟨c.1.flip⟩ - -@[to_additive] -instance covariant_class_mul_le [has_le α] [has_mul α] [c : covariant_class α α (*) (≤)] : - covariant_class αᵒᵈ αᵒᵈ (*) (≤) := -⟨c.1.flip⟩ - -@[to_additive] instance contravariant_class_swap_mul_le [has_le α] [has_mul α] - [c : contravariant_class α α (swap (*)) (≤)] : - contravariant_class αᵒᵈ αᵒᵈ (swap (*)) (≤) := -⟨c.1.flip⟩ - -@[to_additive] -instance covariant_class_swap_mul_le [has_le α] [has_mul α] - [c : covariant_class α α (swap (*)) (≤)] : - covariant_class αᵒᵈ αᵒᵈ (swap (*)) (≤) := -⟨c.1.flip⟩ - -@[to_additive] -instance contravariant_class_mul_lt [has_lt α] [has_mul α] [c : contravariant_class α α (*) (<)] : - contravariant_class αᵒᵈ αᵒᵈ (*) (<) := -⟨c.1.flip⟩ - -@[to_additive] -instance covariant_class_mul_lt [has_lt α] [has_mul α] [c : covariant_class α α (*) (<)] : - covariant_class αᵒᵈ αᵒᵈ (*) (<) := -⟨c.1.flip⟩ - -@[to_additive] instance contravariant_class_swap_mul_lt [has_lt α] [has_mul α] - [c : contravariant_class α α (swap (*)) (<)] : - contravariant_class αᵒᵈ αᵒᵈ (swap (*)) (<) := -⟨c.1.flip⟩ - -@[to_additive] -instance covariant_class_swap_mul_lt [has_lt α] [has_mul α] - [c : covariant_class α α (swap (*)) (<)] : - covariant_class αᵒᵈ αᵒᵈ (swap (*)) (<) := -⟨c.1.flip⟩ - -@[to_additive] -instance [ordered_comm_monoid α] : ordered_comm_monoid αᵒᵈ := -{ mul_le_mul_left := λ a b h c, mul_le_mul_left' h c, - .. order_dual.partial_order α, - .. order_dual.comm_monoid } - -@[to_additive ordered_cancel_add_comm_monoid.to_contravariant_class] -instance ordered_cancel_comm_monoid.to_contravariant_class [ordered_cancel_comm_monoid α] : - contravariant_class αᵒᵈ αᵒᵈ has_mul.mul has_le.le := -{ elim := λ a b c, ordered_cancel_comm_monoid.le_of_mul_le_mul_left a c b } - -@[to_additive] -instance [ordered_cancel_comm_monoid α] : ordered_cancel_comm_monoid αᵒᵈ := -{ le_of_mul_le_mul_left := λ a b c : α, le_of_mul_le_mul_left', - .. order_dual.ordered_comm_monoid, .. order_dual.cancel_comm_monoid } - -@[to_additive] -instance [linear_ordered_cancel_comm_monoid α] : - linear_ordered_cancel_comm_monoid αᵒᵈ := -{ .. order_dual.linear_order α, - .. order_dual.ordered_cancel_comm_monoid } - -@[to_additive] -instance [linear_ordered_comm_monoid α] : - linear_ordered_comm_monoid αᵒᵈ := -{ .. order_dual.linear_order α, - .. order_dual.ordered_comm_monoid } - -end order_dual - -/-! ### Product -/ - -namespace prod - -variables {M N : Type*} - -@[to_additive] -instance [ordered_comm_monoid α] [ordered_comm_monoid β] : ordered_comm_monoid (α × β) := -{ mul_le_mul_left := λ a b h c, ⟨mul_le_mul_left' h.1 _, mul_le_mul_left' h.2 _⟩, - .. prod.comm_monoid, .. prod.partial_order _ _ } - -@[to_additive] -instance [ordered_cancel_comm_monoid M] [ordered_cancel_comm_monoid N] : - ordered_cancel_comm_monoid (M × N) := -{ le_of_mul_le_mul_left := λ a b c h, ⟨le_of_mul_le_mul_left' h.1, le_of_mul_le_mul_left' h.2⟩, - .. prod.ordered_comm_monoid } - -@[to_additive] instance [has_le α] [has_le β] [has_mul α] [has_mul β] [has_exists_mul_of_le α] - [has_exists_mul_of_le β] : has_exists_mul_of_le (α × β) := -⟨λ a b h, let ⟨c, hc⟩ := exists_mul_of_le h.1, ⟨d, hd⟩ := exists_mul_of_le h.2 in - ⟨(c, d), ext hc hd⟩⟩ - -@[to_additive] instance [canonically_ordered_monoid α] [canonically_ordered_monoid β] : - canonically_ordered_monoid (α × β) := -{ le_self_mul := λ a b, ⟨le_self_mul, le_self_mul⟩, - ..prod.ordered_comm_monoid, ..prod.order_bot _ _, ..prod.has_exists_mul_of_le } - -end prod - -/-! ### `with_bot`/`with_top`-/ - -namespace with_top - -section has_one - -variables [has_one α] - -@[to_additive] instance : has_one (with_top α) := ⟨(1 : α)⟩ - -@[simp, norm_cast, to_additive] lemma coe_one : ((1 : α) : with_top α) = 1 := rfl - -@[simp, norm_cast, to_additive] lemma coe_eq_one {a : α} : (a : with_top α) = 1 ↔ a = 1 := -coe_eq_coe - -@[simp, to_additive] protected lemma map_one {β} (f : α → β) : - (1 : with_top α).map f = (f 1 : with_top β) := rfl - -@[simp, norm_cast, to_additive] theorem one_eq_coe {a : α} : 1 = (a : with_top α) ↔ a = 1 := -trans eq_comm coe_eq_one - -@[simp, to_additive] theorem top_ne_one : ⊤ ≠ (1 : with_top α) . -@[simp, to_additive] theorem one_ne_top : (1 : with_top α) ≠ ⊤ . - -instance [has_zero α] [has_le α] [zero_le_one_class α] : zero_le_one_class (with_top α) := -⟨some_le_some.2 zero_le_one⟩ - -end has_one - -section has_add -variables [has_add α] {a b c d : with_top α} {x y : α} - -instance : has_add (with_top α) := ⟨λ o₁ o₂, o₁.bind $ λ a, o₂.map $ (+) a⟩ - -@[norm_cast] lemma coe_add : ((x + y : α) : with_top α) = x + y := rfl -@[norm_cast] lemma coe_bit0 : ((bit0 x : α) : with_top α) = bit0 x := rfl -@[norm_cast] lemma coe_bit1 [has_one α] {a : α} : ((bit1 a : α) : with_top α) = bit1 a := rfl - -@[simp] lemma top_add (a : with_top α) : ⊤ + a = ⊤ := rfl -@[simp] lemma add_top (a : with_top α) : a + ⊤ = ⊤ := by cases a; refl - -@[simp] lemma add_eq_top : a + b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := -by cases a; cases b; simp [none_eq_top, some_eq_coe, ←with_top.coe_add, ←with_zero.coe_add] - -lemma add_ne_top : a + b ≠ ⊤ ↔ a ≠ ⊤ ∧ b ≠ ⊤ := add_eq_top.not.trans not_or_distrib - -lemma add_lt_top [partial_order α] {a b : with_top α} : a + b < ⊤ ↔ a < ⊤ ∧ b < ⊤ := -by simp_rw [lt_top_iff_ne_top, add_ne_top] - -lemma add_eq_coe : ∀ {a b : with_top α} {c : α}, - a + b = c ↔ ∃ (a' b' : α), ↑a' = a ∧ ↑b' = b ∧ a' + b' = c -| none b c := by simp [none_eq_top] -| (some a) none c := by simp [none_eq_top] -| (some a) (some b) c := - by simp only [some_eq_coe, ← coe_add, coe_eq_coe, exists_and_distrib_left, exists_eq_left] - -@[simp] lemma add_coe_eq_top_iff {x : with_top α} {y : α} : x + y = ⊤ ↔ x = ⊤ := -by { induction x using with_top.rec_top_coe; simp [← coe_add, -with_zero.coe_add] } - -@[simp] lemma coe_add_eq_top_iff {y : with_top α} : ↑x + y = ⊤ ↔ y = ⊤ := -by { induction y using with_top.rec_top_coe; simp [← coe_add, -with_zero.coe_add] } - -instance covariant_class_add_le [has_le α] [covariant_class α α (+) (≤)] : - covariant_class (with_top α) (with_top α) (+) (≤) := -⟨λ a b c h, begin - cases a; cases c; try { exact le_top }, - rcases le_coe_iff.1 h with ⟨b, rfl, h'⟩, - exact coe_le_coe.2 (add_le_add_left (coe_le_coe.1 h) _) -end⟩ - -instance covariant_class_swap_add_le [has_le α] [covariant_class α α (swap (+)) (≤)] : - covariant_class (with_top α) (with_top α) (swap (+)) (≤) := -⟨λ a b c h, begin - cases a; cases c; try { exact le_top }, - rcases le_coe_iff.1 h with ⟨b, rfl, h'⟩, - exact coe_le_coe.2 (add_le_add_right (coe_le_coe.1 h) _) -end⟩ - -instance contravariant_class_add_lt [has_lt α] [contravariant_class α α (+) (<)] : - contravariant_class (with_top α) (with_top α) (+) (<) := -⟨λ a b c h, begin - induction a using with_top.rec_top_coe, { exact (not_none_lt _ h).elim }, - induction b using with_top.rec_top_coe, { exact (not_none_lt _ h).elim }, - induction c using with_top.rec_top_coe, - { exact coe_lt_top _ }, - { exact coe_lt_coe.2 (lt_of_add_lt_add_left $ coe_lt_coe.1 h) } -end⟩ - -instance contravariant_class_swap_add_lt [has_lt α] [contravariant_class α α (swap (+)) (<)] : - contravariant_class (with_top α) (with_top α) (swap (+)) (<) := -⟨λ a b c h, begin - cases a; cases b; try { exact (not_none_lt _ h).elim }, - cases c, - { exact coe_lt_top _ }, - { exact coe_lt_coe.2 (lt_of_add_lt_add_right $ coe_lt_coe.1 h) } -end⟩ - -protected lemma le_of_add_le_add_left [has_le α] [contravariant_class α α (+) (≤)] (ha : a ≠ ⊤) - (h : a + b ≤ a + c) : b ≤ c := -begin - lift a to α using ha, - induction c using with_top.rec_top_coe, { exact le_top }, - induction b using with_top.rec_top_coe, { exact (not_top_le_coe _ h).elim }, - simp only [← coe_add, coe_le_coe] at h ⊢, - exact le_of_add_le_add_left h -end - -protected lemma le_of_add_le_add_right [has_le α] [contravariant_class α α (swap (+)) (≤)] - (ha : a ≠ ⊤) (h : b + a ≤ c + a) : b ≤ c := -begin - lift a to α using ha, - cases c, - { exact le_top }, - cases b, - { exact (not_top_le_coe _ h).elim }, - { exact coe_le_coe.2 (le_of_add_le_add_right $ coe_le_coe.1 h) } -end - -protected lemma add_lt_add_left [has_lt α] [covariant_class α α (+) (<)] (ha : a ≠ ⊤) (h : b < c) : - a + b < a + c := -begin - lift a to α using ha, - rcases lt_iff_exists_coe.1 h with ⟨b, rfl, h'⟩, - cases c, - { exact coe_lt_top _ }, - { exact coe_lt_coe.2 (add_lt_add_left (coe_lt_coe.1 h) _) } -end - -protected lemma add_lt_add_right [has_lt α] [covariant_class α α (swap (+)) (<)] - (ha : a ≠ ⊤) (h : b < c) : - b + a < c + a := -begin - lift a to α using ha, - rcases lt_iff_exists_coe.1 h with ⟨b, rfl, h'⟩, - cases c, - { exact coe_lt_top _ }, - { exact coe_lt_coe.2 (add_lt_add_right (coe_lt_coe.1 h) _) } -end - -protected lemma add_le_add_iff_left [has_le α] [covariant_class α α (+) (≤)] - [contravariant_class α α (+) (≤)] - (ha : a ≠ ⊤) : a + b ≤ a + c ↔ b ≤ c := -⟨with_top.le_of_add_le_add_left ha, λ h, add_le_add_left h a⟩ - -protected lemma add_le_add_iff_right [has_le α] [covariant_class α α (swap (+)) (≤)] - [contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊤) : b + a ≤ c + a ↔ b ≤ c := -⟨with_top.le_of_add_le_add_right ha, λ h, add_le_add_right h a⟩ - -protected lemma add_lt_add_iff_left [has_lt α] [covariant_class α α (+) (<)] - [contravariant_class α α (+) (<)] (ha : a ≠ ⊤) : a + b < a + c ↔ b < c := -⟨lt_of_add_lt_add_left, with_top.add_lt_add_left ha⟩ - -protected lemma add_lt_add_iff_right [has_lt α] [covariant_class α α (swap (+)) (<)] - [contravariant_class α α (swap (+)) (<)] (ha : a ≠ ⊤) : b + a < c + a ↔ b < c := -⟨lt_of_add_lt_add_right, with_top.add_lt_add_right ha⟩ - -protected lemma add_lt_add_of_le_of_lt [preorder α] [covariant_class α α (+) (<)] - [covariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊤) (hab : a ≤ b) (hcd : c < d) : a + c < b + d := -(with_top.add_lt_add_left ha hcd).trans_le $ add_le_add_right hab _ - -protected lemma add_lt_add_of_lt_of_le [preorder α] [covariant_class α α (+) (≤)] - [covariant_class α α (swap (+)) (<)] (hc : c ≠ ⊤) (hab : a < b) (hcd : c ≤ d) : a + c < b + d := -(with_top.add_lt_add_right hc hab).trans_le $ add_le_add_left hcd _ - -/- There is no `with_top.map_mul_of_mul_hom`, since `with_top` does not have a multiplication. -/ -@[simp] protected lemma map_add {F} [has_add β] [add_hom_class F α β] (f : F) (a b : with_top α) : - (a + b).map f = a.map f + b.map f := -begin - induction a using with_top.rec_top_coe, - { exact (top_add _).symm }, - { induction b using with_top.rec_top_coe, - { exact (add_top _).symm }, - { rw [map_coe, map_coe, ← coe_add, ← coe_add, ← map_add], - refl } }, -end - -end has_add - -instance [add_semigroup α] : add_semigroup (with_top α) := -{ add_assoc := begin - repeat { refine with_top.rec_top_coe _ _; try { intro }}; - simp [←with_top.coe_add, add_assoc] - end, - ..with_top.has_add } - -instance [add_comm_semigroup α] : add_comm_semigroup (with_top α) := -{ add_comm := - begin - repeat { refine with_top.rec_top_coe _ _; try { intro }}; - simp [←with_top.coe_add, add_comm] - end, - ..with_top.add_semigroup } - -instance [add_zero_class α] : add_zero_class (with_top α) := -{ zero_add := - begin - refine with_top.rec_top_coe _ _, - { simp }, - { intro, - rw [←with_top.coe_zero, ←with_top.coe_add, zero_add] } - end, - add_zero := - begin - refine with_top.rec_top_coe _ _, - { simp }, - { intro, - rw [←with_top.coe_zero, ←with_top.coe_add, add_zero] } - end, - ..with_top.has_zero, - ..with_top.has_add } - -instance [add_monoid α] : add_monoid (with_top α) := -{ ..with_top.add_zero_class, - ..with_top.has_zero, - ..with_top.add_semigroup } - -instance [add_comm_monoid α] : add_comm_monoid (with_top α) := -{ ..with_top.add_monoid, ..with_top.add_comm_semigroup } - -instance [add_monoid_with_one α] : add_monoid_with_one (with_top α) := -{ nat_cast := λ n, ↑(n : α), - nat_cast_zero := by rw [nat.cast_zero, with_top.coe_zero], - nat_cast_succ := λ n, by rw [nat.cast_add_one, with_top.coe_add, with_top.coe_one], - .. with_top.has_one, .. with_top.add_monoid } - -instance [add_comm_monoid_with_one α] : add_comm_monoid_with_one (with_top α) := -{ .. with_top.add_monoid_with_one, .. with_top.add_comm_monoid } - -instance [ordered_add_comm_monoid α] : ordered_add_comm_monoid (with_top α) := -{ add_le_add_left := - begin - rintros a b h (_|c), { simp [none_eq_top] }, - rcases b with (_|b), { simp [none_eq_top] }, - rcases le_coe_iff.1 h with ⟨a, rfl, h⟩, - simp only [some_eq_coe, ← coe_add, coe_le_coe] at h ⊢, - exact add_le_add_left h c - end, - ..with_top.partial_order, ..with_top.add_comm_monoid } - -instance [linear_ordered_add_comm_monoid α] : - linear_ordered_add_comm_monoid_with_top (with_top α) := -{ top_add' := with_top.top_add, - ..with_top.order_top, - ..with_top.linear_order, - ..with_top.ordered_add_comm_monoid, - ..option.nontrivial } - -instance [has_le α] [has_add α] [has_exists_add_of_le α] : has_exists_add_of_le (with_top α) := -⟨λ a b, match a, b with - | ⊤, ⊤ := by simp - | (a : α), ⊤ := λ _, ⟨⊤, rfl⟩ - | (a : α), (b : α) := λ h, begin - obtain ⟨c, rfl⟩ := exists_add_of_le (with_top.coe_le_coe.1 h), - exact ⟨c, rfl⟩ - end - | ⊤, (b : α) := λ h, (not_top_le_coe _ h).elim -end⟩ - -instance [canonically_ordered_add_monoid α] : canonically_ordered_add_monoid (with_top α) := -{ le_self_add := λ a b, match a, b with - | ⊤, ⊤ := le_rfl - | (a : α), ⊤ := le_top - | (a : α), (b : α) := with_top.coe_le_coe.2 le_self_add - | ⊤, (b : α) := le_rfl - end, - ..with_top.order_bot, ..with_top.ordered_add_comm_monoid, ..with_top.has_exists_add_of_le } - -instance [canonically_linear_ordered_add_monoid α] : - canonically_linear_ordered_add_monoid (with_top α) := -{ ..with_top.canonically_ordered_add_monoid, ..with_top.linear_order } - -@[simp, norm_cast] lemma coe_nat [add_monoid_with_one α] (n : ℕ) : ((n : α) : with_top α) = n := rfl -@[simp] lemma nat_ne_top [add_monoid_with_one α] (n : ℕ) : (n : with_top α) ≠ ⊤ := coe_ne_top -@[simp] lemma top_ne_nat [add_monoid_with_one α] (n : ℕ) : (⊤ : with_top α) ≠ n := top_ne_coe - -/-- Coercion from `α` to `with_top α` as an `add_monoid_hom`. -/ -def coe_add_hom [add_monoid α] : α →+ with_top α := -⟨coe, rfl, λ _ _, rfl⟩ - -@[simp] lemma coe_coe_add_hom [add_monoid α] : ⇑(coe_add_hom : α →+ with_top α) = coe := rfl - -@[simp] lemma zero_lt_top [ordered_add_comm_monoid α] : (0 : with_top α) < ⊤ := -coe_lt_top 0 - -@[simp, norm_cast] lemma zero_lt_coe [ordered_add_comm_monoid α] (a : α) : - (0 : with_top α) < a ↔ 0 < a := -coe_lt_coe - -/-- A version of `with_top.map` for `one_hom`s. -/ -@[to_additive "A version of `with_top.map` for `zero_hom`s", simps { fully_applied := ff }] -protected def _root_.one_hom.with_top_map {M N : Type*} [has_one M] [has_one N] (f : one_hom M N) : - one_hom (with_top M) (with_top N) := -{ to_fun := with_top.map f, - map_one' := by rw [with_top.map_one, map_one, coe_one] } - -/-- A version of `with_top.map` for `add_hom`s. -/ -@[simps { fully_applied := ff }] protected def _root_.add_hom.with_top_map - {M N : Type*} [has_add M] [has_add N] (f : add_hom M N) : - add_hom (with_top M) (with_top N) := -{ to_fun := with_top.map f, - map_add' := with_top.map_add f } - -/-- A version of `with_top.map` for `add_monoid_hom`s. -/ -@[simps { fully_applied := ff }] protected def _root_.add_monoid_hom.with_top_map - {M N : Type*} [add_zero_class M] [add_zero_class N] (f : M →+ N) : - with_top M →+ with_top N := -{ to_fun := with_top.map f, - .. f.to_zero_hom.with_top_map, .. f.to_add_hom.with_top_map } - -end with_top - -namespace with_bot - -@[to_additive] instance [has_one α] : has_one (with_bot α) := with_top.has_one -instance [has_add α] : has_add (with_bot α) := with_top.has_add -instance [add_semigroup α] : add_semigroup (with_bot α) := with_top.add_semigroup -instance [add_comm_semigroup α] : add_comm_semigroup (with_bot α) := with_top.add_comm_semigroup -instance [add_zero_class α] : add_zero_class (with_bot α) := with_top.add_zero_class -instance [add_monoid α] : add_monoid (with_bot α) := with_top.add_monoid -instance [add_comm_monoid α] : add_comm_monoid (with_bot α) := with_top.add_comm_monoid -instance [add_monoid_with_one α] : add_monoid_with_one (with_bot α) := with_top.add_monoid_with_one - -instance [add_comm_monoid_with_one α] : add_comm_monoid_with_one (with_bot α) := -with_top.add_comm_monoid_with_one - -instance [has_zero α] [has_one α] [has_le α] [zero_le_one_class α] : - zero_le_one_class (with_bot α) := -⟨some_le_some.2 zero_le_one⟩ - --- `by norm_cast` proves this lemma, so I did not tag it with `norm_cast` -@[to_additive] -lemma coe_one [has_one α] : ((1 : α) : with_bot α) = 1 := rfl - --- `by norm_cast` proves this lemma, so I did not tag it with `norm_cast` -@[to_additive] -lemma coe_eq_one [has_one α] {a : α} : (a : with_bot α) = 1 ↔ a = 1 := -with_top.coe_eq_one - -@[to_additive] protected lemma map_one {β} [has_one α] (f : α → β) : - (1 : with_bot α).map f = (f 1 : with_bot β) := rfl - -@[norm_cast] lemma coe_nat [add_monoid_with_one α] (n : ℕ) : ((n : α) : with_bot α) = n := rfl -@[simp] lemma nat_ne_bot [add_monoid_with_one α] (n : ℕ) : (n : with_bot α) ≠ ⊥ := coe_ne_bot -@[simp] lemma bot_ne_nat [add_monoid_with_one α] (n : ℕ) : (⊥ : with_bot α) ≠ n := bot_ne_coe - -section has_add -variables [has_add α] {a b c d : with_bot α} {x y : α} - --- `norm_cast` proves those lemmas, because `with_top`/`with_bot` are reducible -lemma coe_add (a b : α) : ((a + b : α) : with_bot α) = a + b := rfl -lemma coe_bit0 : ((bit0 x : α) : with_bot α) = bit0 x := rfl -lemma coe_bit1 [has_one α] {a : α} : ((bit1 a : α) : with_bot α) = bit1 a := rfl - -@[simp] lemma bot_add (a : with_bot α) : ⊥ + a = ⊥ := rfl -@[simp] lemma add_bot (a : with_bot α) : a + ⊥ = ⊥ := by cases a; refl - -@[simp] lemma add_eq_bot : a + b = ⊥ ↔ a = ⊥ ∨ b = ⊥ := with_top.add_eq_top -lemma add_ne_bot : a + b ≠ ⊥ ↔ a ≠ ⊥ ∧ b ≠ ⊥ := with_top.add_ne_top - -lemma bot_lt_add [partial_order α] {a b : with_bot α} : ⊥ < a + b ↔ ⊥ < a ∧ ⊥ < b := -@with_top.add_lt_top αᵒᵈ _ _ _ _ - -lemma add_eq_coe : a + b = x ↔ ∃ (a' b' : α), ↑a' = a ∧ ↑b' = b ∧ a' + b' = x := with_top.add_eq_coe - -@[simp] lemma add_coe_eq_bot_iff : a + y = ⊥ ↔ a = ⊥ := with_top.add_coe_eq_top_iff -@[simp] lemma coe_add_eq_bot_iff : ↑x + b = ⊥ ↔ b = ⊥ := with_top.coe_add_eq_top_iff - -/- There is no `with_bot.map_mul_of_mul_hom`, since `with_bot` does not have a multiplication. -/ -@[simp] protected lemma map_add {F} [has_add β] [add_hom_class F α β] (f : F) (a b : with_bot α) : - (a + b).map f = a.map f + b.map f := -with_top.map_add f a b - -/-- A version of `with_bot.map` for `one_hom`s. -/ -@[to_additive "A version of `with_bot.map` for `zero_hom`s", simps { fully_applied := ff }] -protected def _root_.one_hom.with_bot_map {M N : Type*} [has_one M] [has_one N] (f : one_hom M N) : - one_hom (with_bot M) (with_bot N) := -{ to_fun := with_bot.map f, - map_one' := by rw [with_bot.map_one, map_one, coe_one] } - -/-- A version of `with_bot.map` for `add_hom`s. -/ -@[simps { fully_applied := ff }] protected def _root_.add_hom.with_bot_map - {M N : Type*} [has_add M] [has_add N] (f : add_hom M N) : - add_hom (with_bot M) (with_bot N) := -{ to_fun := with_bot.map f, - map_add' := with_bot.map_add f } - -/-- A version of `with_bot.map` for `add_monoid_hom`s. -/ -@[simps { fully_applied := ff }] protected def _root_.add_monoid_hom.with_bot_map - {M N : Type*} [add_zero_class M] [add_zero_class N] (f : M →+ N) : - with_bot M →+ with_bot N := -{ to_fun := with_bot.map f, - .. f.to_zero_hom.with_bot_map, .. f.to_add_hom.with_bot_map } - -variables [preorder α] - -instance covariant_class_add_le [covariant_class α α (+) (≤)] : - covariant_class (with_bot α) (with_bot α) (+) (≤) := -@order_dual.covariant_class_add_le (with_top αᵒᵈ) _ _ _ - -instance covariant_class_swap_add_le [covariant_class α α (swap (+)) (≤)] : - covariant_class (with_bot α) (with_bot α) (swap (+)) (≤) := -@order_dual.covariant_class_swap_add_le (with_top αᵒᵈ) _ _ _ - -instance contravariant_class_add_lt [contravariant_class α α (+) (<)] : - contravariant_class (with_bot α) (with_bot α) (+) (<) := -@order_dual.contravariant_class_add_lt (with_top αᵒᵈ) _ _ _ - -instance contravariant_class_swap_add_lt [contravariant_class α α (swap (+)) (<)] : - contravariant_class (with_bot α) (with_bot α) (swap (+)) (<) := -@order_dual.contravariant_class_swap_add_lt (with_top αᵒᵈ) _ _ _ - -protected lemma le_of_add_le_add_left [contravariant_class α α (+) (≤)] (ha : a ≠ ⊥) - (h : a + b ≤ a + c) : b ≤ c := -@with_top.le_of_add_le_add_left αᵒᵈ _ _ _ _ _ _ ha h - -protected lemma le_of_add_le_add_right [contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊥) - (h : b + a ≤ c + a) : b ≤ c := -@with_top.le_of_add_le_add_right αᵒᵈ _ _ _ _ _ _ ha h - -protected lemma add_lt_add_left [covariant_class α α (+) (<)] (ha : a ≠ ⊥) (h : b < c) : - a + b < a + c := -@with_top.add_lt_add_left αᵒᵈ _ _ _ _ _ _ ha h - -protected lemma add_lt_add_right [covariant_class α α (swap (+)) (<)] (ha : a ≠ ⊥) (h : b < c) : - b + a < c + a := -@with_top.add_lt_add_right αᵒᵈ _ _ _ _ _ _ ha h - -protected lemma add_le_add_iff_left [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤)] - (ha : a ≠ ⊥) : a + b ≤ a + c ↔ b ≤ c := -⟨with_bot.le_of_add_le_add_left ha, λ h, add_le_add_left h a⟩ - -protected lemma add_le_add_iff_right [covariant_class α α (swap (+)) (≤)] - [contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊥) : b + a ≤ c + a ↔ b ≤ c := -⟨with_bot.le_of_add_le_add_right ha, λ h, add_le_add_right h a⟩ - -protected lemma add_lt_add_iff_left [covariant_class α α (+) (<)] [contravariant_class α α (+) (<)] - (ha : a ≠ ⊥) : a + b < a + c ↔ b < c := -⟨lt_of_add_lt_add_left, with_bot.add_lt_add_left ha⟩ - -protected lemma add_lt_add_iff_right [covariant_class α α (swap (+)) (<)] - [contravariant_class α α (swap (+)) (<)] (ha : a ≠ ⊥) : b + a < c + a ↔ b < c := -⟨lt_of_add_lt_add_right, with_bot.add_lt_add_right ha⟩ - -protected lemma add_lt_add_of_le_of_lt [covariant_class α α (+) (<)] - [covariant_class α α (swap (+)) (≤)] (hb : b ≠ ⊥) (hab : a ≤ b) (hcd : c < d) : a + c < b + d := -@with_top.add_lt_add_of_le_of_lt αᵒᵈ _ _ _ _ _ _ _ _ hb hab hcd - -protected lemma add_lt_add_of_lt_of_le [covariant_class α α (+) (≤)] - [covariant_class α α (swap (+)) (<)] (hd : d ≠ ⊥) (hab : a < b) (hcd : c ≤ d) : a + c < b + d := -@with_top.add_lt_add_of_lt_of_le αᵒᵈ _ _ _ _ _ _ _ _ hd hab hcd - -end has_add - -instance [ordered_add_comm_monoid α] : ordered_add_comm_monoid (with_bot α) := -{ add_le_add_left := λ a b h c, add_le_add_left h c, - ..with_bot.partial_order, - ..with_bot.add_comm_monoid } - -instance [linear_ordered_add_comm_monoid α] : linear_ordered_add_comm_monoid (with_bot α) := -{ ..with_bot.linear_order, ..with_bot.ordered_add_comm_monoid } - -end with_bot - -/-! ### `additive`/`multiplicative` -/ - -section type_tags - -instance : Π [has_le α], has_le (multiplicative α) := id -instance : Π [has_le α], has_le (additive α) := id -instance : Π [has_lt α], has_lt (multiplicative α) := id -instance : Π [has_lt α], has_lt (additive α) := id -instance : Π [preorder α], preorder (multiplicative α) := id -instance : Π [preorder α], preorder (additive α) := id -instance : Π [partial_order α], partial_order (multiplicative α) := id -instance : Π [partial_order α], partial_order (additive α) := id -instance : Π [linear_order α], linear_order (multiplicative α) := id -instance : Π [linear_order α], linear_order (additive α) := id -instance [has_le α] : Π [order_bot α], order_bot (multiplicative α) := id -instance [has_le α] : Π [order_bot α], order_bot (additive α) := id -instance [has_le α] : Π [order_top α], order_top (multiplicative α) := id -instance [has_le α] : Π [order_top α], order_top (additive α) := id -instance [has_le α] : Π [bounded_order α], bounded_order (multiplicative α) := id -instance [has_le α] : Π [bounded_order α], bounded_order (additive α) := id - -instance [ordered_add_comm_monoid α] : ordered_comm_monoid (multiplicative α) := -{ mul_le_mul_left := @ordered_add_comm_monoid.add_le_add_left α _, - ..multiplicative.partial_order, - ..multiplicative.comm_monoid } - -instance [ordered_comm_monoid α] : ordered_add_comm_monoid (additive α) := -{ add_le_add_left := @ordered_comm_monoid.mul_le_mul_left α _, - ..additive.partial_order, - ..additive.add_comm_monoid } - -instance [ordered_cancel_add_comm_monoid α] : ordered_cancel_comm_monoid (multiplicative α) := -{ le_of_mul_le_mul_left := @ordered_cancel_add_comm_monoid.le_of_add_le_add_left α _, - ..multiplicative.ordered_comm_monoid } - -instance [ordered_cancel_comm_monoid α] : ordered_cancel_add_comm_monoid (additive α) := -{ le_of_add_le_add_left := @ordered_cancel_comm_monoid.le_of_mul_le_mul_left α _, - ..additive.ordered_add_comm_monoid } - -instance [linear_ordered_add_comm_monoid α] : linear_ordered_comm_monoid (multiplicative α) := -{ ..multiplicative.linear_order, - ..multiplicative.ordered_comm_monoid } - -instance [linear_ordered_comm_monoid α] : linear_ordered_add_comm_monoid (additive α) := -{ ..additive.linear_order, - ..additive.ordered_add_comm_monoid } - -instance [has_add α] [has_le α] [has_exists_add_of_le α] : - has_exists_mul_of_le (multiplicative α) := -⟨@exists_add_of_le α _ _ _⟩ - -instance [has_mul α] [has_le α] [has_exists_mul_of_le α] : has_exists_add_of_le (additive α) := -⟨@exists_mul_of_le α _ _ _⟩ - -instance [canonically_ordered_add_monoid α] : canonically_ordered_monoid (multiplicative α) := -{ le_self_mul := @le_self_add α _, - ..multiplicative.ordered_comm_monoid, ..multiplicative.order_bot, - ..multiplicative.has_exists_mul_of_le } - -instance [canonically_ordered_monoid α] : canonically_ordered_add_monoid (additive α) := -{ le_self_add := @le_self_mul α _, - ..additive.ordered_add_comm_monoid, ..additive.order_bot, ..additive.has_exists_add_of_le } - -instance [canonically_linear_ordered_add_monoid α] : - canonically_linear_ordered_monoid (multiplicative α) := -{ ..multiplicative.canonically_ordered_monoid, ..multiplicative.linear_order } - -instance [canonically_linear_ordered_monoid α] : - canonically_linear_ordered_add_monoid (additive α) := -{ ..additive.canonically_ordered_add_monoid, ..additive.linear_order } - -namespace additive - -variables [preorder α] - -@[simp] lemma of_mul_le {a b : α} : of_mul a ≤ of_mul b ↔ a ≤ b := iff.rfl - -@[simp] lemma of_mul_lt {a b : α} : of_mul a < of_mul b ↔ a < b := iff.rfl - -@[simp] lemma to_mul_le {a b : additive α} : to_mul a ≤ to_mul b ↔ a ≤ b := iff.rfl - -@[simp] lemma to_mul_lt {a b : additive α} : to_mul a < to_mul b ↔ a < b := iff.rfl - -end additive - -namespace multiplicative - -variables [preorder α] - -@[simp] lemma of_add_le {a b : α} : of_add a ≤ of_add b ↔ a ≤ b := iff.rfl - -@[simp] lemma of_add_lt {a b : α} : of_add a < of_add b ↔ a < b := iff.rfl - -@[simp] lemma to_add_le {a b : multiplicative α} : to_add a ≤ to_add b ↔ a ≤ b := iff.rfl - -@[simp] lemma to_add_lt {a b : multiplicative α} : to_add a < to_add b ↔ a < b := iff.rfl - -end multiplicative - -end type_tags - -namespace with_zero - -local attribute [semireducible] with_zero -variables [has_add α] - -/-- Making an additive monoid multiplicative then adding a zero is the same as adding a bottom -element then making it multiplicative. -/ -def to_mul_bot : with_zero (multiplicative α) ≃* multiplicative (with_bot α) := -by exact mul_equiv.refl _ - -@[simp] lemma to_mul_bot_zero : - to_mul_bot (0 : with_zero (multiplicative α)) = multiplicative.of_add ⊥ := rfl -@[simp] lemma to_mul_bot_coe (x : multiplicative α) : - to_mul_bot ↑x = multiplicative.of_add (x.to_add : with_bot α) := rfl -@[simp] lemma to_mul_bot_symm_bot : - to_mul_bot.symm (multiplicative.of_add (⊥ : with_bot α)) = 0 := rfl -@[simp] lemma to_mul_bot_coe_of_add (x : α) : - to_mul_bot.symm (multiplicative.of_add (x : with_bot α)) = multiplicative.of_add x := rfl - -variables [preorder α] (a b : with_zero (multiplicative α)) - -lemma to_mul_bot_strict_mono : strict_mono (@to_mul_bot α _) := λ x y, id -@[simp] lemma to_mul_bot_le : to_mul_bot a ≤ to_mul_bot b ↔ a ≤ b := iff.rfl -@[simp] lemma to_mul_bot_lt : to_mul_bot a < to_mul_bot b ↔ a < b := iff.rfl - -end with_zero - -/-- The order embedding sending `b` to `a * b`, for some fixed `a`. -See also `order_iso.mul_left` when working in an ordered group. -/ -@[to_additive "The order embedding sending `b` to `a + b`, for some fixed `a`. - See also `order_iso.add_left` when working in an additive ordered group.", simps] -def order_embedding.mul_left - {α : Type*} [has_mul α] [linear_order α] [covariant_class α α (*) (<)] (m : α) : α ↪o α := -order_embedding.of_strict_mono (λ n, m * n) (λ a b w, mul_lt_mul_left' w m) - -/-- The order embedding sending `b` to `b * a`, for some fixed `a`. -See also `order_iso.mul_right` when working in an ordered group. -/ -@[to_additive "The order embedding sending `b` to `b + a`, for some fixed `a`. - See also `order_iso.add_right` when working in an additive ordered group.", simps] -def order_embedding.mul_right - {α : Type*} [has_mul α] [linear_order α] [covariant_class α α (swap (*)) (<)] (m : α) : - α ↪o α := -order_embedding.of_strict_mono (λ n, n * m) (λ a b w, mul_lt_mul_right' w m) diff --git a/src/algebra/order/monoid/basic.lean b/src/algebra/order/monoid/basic.lean new file mode 100644 index 0000000000000..7bd4c83de9d34 --- /dev/null +++ b/src/algebra/order/monoid/basic.lean @@ -0,0 +1,171 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.group.units +import algebra.group_with_zero.defs +import algebra.group.inj_surj +import algebra.order.zero_le_one +import order.hom.basic + +/-! +# Ordered monoids + +This file develops the basics of ordered monoids. + +## Implementation details + +Unfortunately, the number of `'` appended to lemmas in this file +may differ between the multiplicative and the additive version of a lemma. +The reason is that we did not want to change existing names in the library. +-/ + +set_option old_structure_cmd true +open function + +universe u +variables {α : Type u} {β : Type*} + +/-- An ordered commutative monoid is a commutative monoid +with a partial order such that `a ≤ b → c * a ≤ c * b` (multiplication is monotone) +-/ +@[protect_proj, ancestor comm_monoid partial_order] +class ordered_comm_monoid (α : Type*) extends comm_monoid α, partial_order α := +(mul_le_mul_left : ∀ a b : α, a ≤ b → ∀ c : α, c * a ≤ c * b) + +/-- An ordered (additive) commutative monoid is a commutative monoid + with a partial order such that `a ≤ b → c + a ≤ c + b` (addition is monotone) +-/ +@[protect_proj, ancestor add_comm_monoid partial_order] +class ordered_add_comm_monoid (α : Type*) extends add_comm_monoid α, partial_order α := +(add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) + +attribute [to_additive] ordered_comm_monoid + +section ordered_instances + +@[to_additive] +instance ordered_comm_monoid.to_covariant_class_left (M : Type*) [ordered_comm_monoid M] : + covariant_class M M (*) (≤) := +{ elim := λ a b c bc, ordered_comm_monoid.mul_le_mul_left _ _ bc a } + +/- This instance can be proven with `by apply_instance`. However, `with_bot ℕ` does not +pick up a `covariant_class M M (function.swap (*)) (≤)` instance without it (see PR #7940). -/ +@[to_additive] +instance ordered_comm_monoid.to_covariant_class_right (M : Type*) [ordered_comm_monoid M] : + covariant_class M M (swap (*)) (≤) := +covariant_swap_mul_le_of_covariant_mul_le M + +/- This is not an instance, to avoid creating a loop in the type-class system: in a +`left_cancel_semigroup` with a `partial_order`, assuming `covariant_class M M (*) (≤)` implies +`covariant_class M M (*) (<)`, see `left_cancel_semigroup.covariant_mul_lt_of_covariant_mul_le`. -/ +@[to_additive] lemma has_mul.to_covariant_class_left + (M : Type*) [has_mul M] [partial_order M] [covariant_class M M (*) (<)] : + covariant_class M M (*) (≤) := +⟨covariant_le_of_covariant_lt _ _ _ covariant_class.elim⟩ + +/- This is not an instance, to avoid creating a loop in the type-class system: in a +`right_cancel_semigroup` with a `partial_order`, assuming `covariant_class M M (swap (*)) (<)` +implies `covariant_class M M (swap (*)) (≤)`, see +`right_cancel_semigroup.covariant_swap_mul_lt_of_covariant_swap_mul_le`. -/ +@[to_additive] lemma has_mul.to_covariant_class_right + (M : Type*) [has_mul M] [partial_order M] [covariant_class M M (swap (*)) (<)] : + covariant_class M M (swap (*)) (≤) := +⟨covariant_le_of_covariant_lt _ _ _ covariant_class.elim⟩ + +end ordered_instances + +lemma bit0_pos [ordered_add_comm_monoid α] {a : α} (h : 0 < a) : 0 < bit0 a := +add_pos' h h + +/-- A linearly ordered additive commutative monoid. -/ +@[protect_proj, ancestor linear_order ordered_add_comm_monoid] +class linear_ordered_add_comm_monoid (α : Type*) + extends linear_order α, ordered_add_comm_monoid α. + +/-- A linearly ordered commutative monoid. -/ +@[protect_proj, ancestor linear_order ordered_comm_monoid, to_additive] +class linear_ordered_comm_monoid (α : Type*) + extends linear_order α, ordered_comm_monoid α. + +/-- A linearly ordered commutative monoid with a zero element. -/ +class linear_ordered_comm_monoid_with_zero (α : Type*) + extends linear_ordered_comm_monoid α, comm_monoid_with_zero α := +(zero_le_one : (0 : α) ≤ 1) + +@[priority 100] +instance linear_ordered_comm_monoid_with_zero.zero_le_one_class + [h : linear_ordered_comm_monoid_with_zero α] : zero_le_one_class α := +{ ..h } + +/-- A linearly ordered commutative monoid with an additively absorbing `⊤` element. + Instances should include number systems with an infinite element adjoined.` -/ +@[protect_proj, ancestor linear_ordered_add_comm_monoid has_top] +class linear_ordered_add_comm_monoid_with_top (α : Type*) + extends linear_ordered_add_comm_monoid α, has_top α := +(le_top : ∀ x : α, x ≤ ⊤) +(top_add' : ∀ x : α, ⊤ + x = ⊤) + +@[priority 100] -- see Note [lower instance priority] +instance linear_ordered_add_comm_monoid_with_top.to_order_top (α : Type u) + [h : linear_ordered_add_comm_monoid_with_top α] : order_top α := +{ ..h } + +section linear_ordered_add_comm_monoid_with_top +variables [linear_ordered_add_comm_monoid_with_top α] {a b : α} + +@[simp] +lemma top_add (a : α) : ⊤ + a = ⊤ := linear_ordered_add_comm_monoid_with_top.top_add' a + +@[simp] +lemma add_top (a : α) : a + ⊤ = ⊤ := +trans (add_comm _ _) (top_add _) + +end linear_ordered_add_comm_monoid_with_top + +/-- Pullback an `ordered_comm_monoid` under an injective map. +See note [reducible non-instances]. -/ +@[reducible, to_additive function.injective.ordered_add_comm_monoid +"Pullback an `ordered_add_comm_monoid` under an injective map."] +def function.injective.ordered_comm_monoid [ordered_comm_monoid α] {β : Type*} + [has_one β] [has_mul β] [has_pow β ℕ] + (f : β → α) (hf : function.injective f) (one : f 1 = 1) + (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) : + ordered_comm_monoid β := +{ mul_le_mul_left := λ a b ab c, show f (c * a) ≤ f (c * b), by + { rw [mul, mul], apply mul_le_mul_left', exact ab }, + ..partial_order.lift f hf, + ..hf.comm_monoid f one mul npow } + +/-- Pullback a `linear_ordered_comm_monoid` under an injective map. +See note [reducible non-instances]. -/ +@[reducible, to_additive function.injective.linear_ordered_add_comm_monoid +"Pullback an `ordered_add_comm_monoid` under an injective map."] +def function.injective.linear_ordered_comm_monoid [linear_ordered_comm_monoid α] {β : Type*} + [has_one β] [has_mul β] [has_pow β ℕ] [has_sup β] [has_inf β] + (f : β → α) (hf : function.injective f) (one : f 1 = 1) + (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_comm_monoid β := +{ .. hf.ordered_comm_monoid f one mul npow, + .. linear_order.lift f hf hsup hinf } + +-- TODO find a better home for the next two constructions. + +/-- The order embedding sending `b` to `a * b`, for some fixed `a`. +See also `order_iso.mul_left` when working in an ordered group. -/ +@[to_additive "The order embedding sending `b` to `a + b`, for some fixed `a`. + See also `order_iso.add_left` when working in an additive ordered group.", simps] +def order_embedding.mul_left + {α : Type*} [has_mul α] [linear_order α] [covariant_class α α (*) (<)] (m : α) : α ↪o α := +order_embedding.of_strict_mono (λ n, m * n) (λ a b w, mul_lt_mul_left' w m) + +/-- The order embedding sending `b` to `b * a`, for some fixed `a`. +See also `order_iso.mul_right` when working in an ordered group. -/ +@[to_additive "The order embedding sending `b` to `b + a`, for some fixed `a`. + See also `order_iso.add_right` when working in an additive ordered group.", simps] +def order_embedding.mul_right + {α : Type*} [has_mul α] [linear_order α] [covariant_class α α (swap (*)) (<)] (m : α) : + α ↪o α := +order_embedding.of_strict_mono (λ n, n * m) (λ a b w, mul_lt_mul_right' w m) diff --git a/src/algebra/order/monoid/cancel.lean b/src/algebra/order/monoid/cancel.lean new file mode 100644 index 0000000000000..02bc3b4bca0dd --- /dev/null +++ b/src/algebra/order/monoid/cancel.lean @@ -0,0 +1,120 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.order.monoid.basic + +/-! +# Ordered cancellative monoids +-/ + +universe u +variables {α : Type u} + +open function + +set_option old_structure_cmd true + +/-- An ordered cancellative additive commutative monoid +is an additive commutative monoid with a partial order, +in which addition is cancellative and monotone. -/ +@[protect_proj, ancestor add_comm_monoid partial_order] +class ordered_cancel_add_comm_monoid (α : Type u) extends add_comm_monoid α, partial_order α := +(add_le_add_left : ∀ a b : α, a ≤ b → ∀ c : α, c + a ≤ c + b) +(le_of_add_le_add_left : ∀ a b c : α, a + b ≤ a + c → b ≤ c) + +/-- An ordered cancellative commutative monoid +is a commutative monoid with a partial order, +in which multiplication is cancellative and monotone. -/ +@[protect_proj, ancestor comm_monoid partial_order, to_additive] +class ordered_cancel_comm_monoid (α : Type u) extends comm_monoid α, partial_order α := +(mul_le_mul_left : ∀ a b : α, a ≤ b → ∀ c : α, c * a ≤ c * b) +(le_of_mul_le_mul_left : ∀ a b c : α, a * b ≤ a * c → b ≤ c) + +section ordered_cancel_comm_monoid +variables [ordered_cancel_comm_monoid α] {a b c d : α} + +@[priority 200, to_additive] -- see Note [lower instance priority] +instance ordered_cancel_comm_monoid.to_contravariant_class_le_left : + contravariant_class α α (*) (≤) := +⟨ordered_cancel_comm_monoid.le_of_mul_le_mul_left⟩ + +@[to_additive] +lemma ordered_cancel_comm_monoid.lt_of_mul_lt_mul_left : ∀ a b c : α, a * b < a * c → b < c := +λ a b c h, lt_of_le_not_le + (ordered_cancel_comm_monoid.le_of_mul_le_mul_left a b c h.le) $ + mt (λ h, ordered_cancel_comm_monoid.mul_le_mul_left _ _ h _) (not_le_of_gt h) + +@[to_additive] +instance ordered_cancel_comm_monoid.to_contravariant_class_left + (M : Type*) [ordered_cancel_comm_monoid M] : + contravariant_class M M (*) (<) := +{ elim := λ a b c, ordered_cancel_comm_monoid.lt_of_mul_lt_mul_left _ _ _ } + +/- This instance can be proven with `by apply_instance`. However, by analogy with the +instance `ordered_cancel_comm_monoid.to_covariant_class_right` above, I imagine that without +this instance, some Type would not have a `contravariant_class M M (function.swap (*)) (<)` +instance. -/ +@[to_additive] +instance ordered_cancel_comm_monoid.to_contravariant_class_right + (M : Type*) [ordered_cancel_comm_monoid M] : + contravariant_class M M (swap (*)) (<) := +contravariant_swap_mul_lt_of_contravariant_mul_lt M + +@[priority 100, to_additive] -- see Note [lower instance priority] +instance ordered_cancel_comm_monoid.to_ordered_comm_monoid : ordered_comm_monoid α := +{ ..‹ordered_cancel_comm_monoid α› } + +@[priority 100, to_additive] -- see Note [lower instance priority] +instance ordered_cancel_comm_monoid.to_cancel_comm_monoid : cancel_comm_monoid α := +{ mul_left_cancel := λ a b c h, + (le_of_mul_le_mul_left' h.le).antisymm $ le_of_mul_le_mul_left' h.ge, + ..‹ordered_cancel_comm_monoid α› } + +/-- Pullback an `ordered_cancel_comm_monoid` under an injective map. +See note [reducible non-instances]. -/ +@[reducible, to_additive function.injective.ordered_cancel_add_comm_monoid +"Pullback an `ordered_cancel_add_comm_monoid` under an injective map."] +def function.injective.ordered_cancel_comm_monoid {β : Type*} + [has_one β] [has_mul β] [has_pow β ℕ] + (f : β → α) (hf : function.injective f) (one : f 1 = 1) + (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) : + ordered_cancel_comm_monoid β := +{ le_of_mul_le_mul_left := λ a b c (bc : f (a * b) ≤ f (a * c)), + (mul_le_mul_iff_left (f a)).mp (by rwa [← mul, ← mul]), + ..hf.ordered_comm_monoid f one mul npow } + +end ordered_cancel_comm_monoid + +/-- A linearly ordered cancellative additive commutative monoid +is an additive commutative monoid with a decidable linear order +in which addition is cancellative and monotone. -/ +@[protect_proj, ancestor ordered_cancel_add_comm_monoid linear_ordered_add_comm_monoid] +class linear_ordered_cancel_add_comm_monoid (α : Type u) + extends ordered_cancel_add_comm_monoid α, linear_ordered_add_comm_monoid α + +/-- A linearly ordered cancellative commutative monoid +is a commutative monoid with a linear order +in which multiplication is cancellative and monotone. -/ +@[protect_proj, ancestor ordered_cancel_comm_monoid linear_ordered_comm_monoid, to_additive] +class linear_ordered_cancel_comm_monoid (α : Type u) + extends ordered_cancel_comm_monoid α, linear_ordered_comm_monoid α + +section linear_ordered_cancel_comm_monoid +variables [linear_ordered_cancel_comm_monoid α] + +/-- Pullback a `linear_ordered_cancel_comm_monoid` under an injective map. +See note [reducible non-instances]. -/ +@[reducible, to_additive function.injective.linear_ordered_cancel_add_comm_monoid +"Pullback a `linear_ordered_cancel_add_comm_monoid` under an injective map."] +def function.injective.linear_ordered_cancel_comm_monoid {β : Type*} + [has_one β] [has_mul β] [has_pow β ℕ] [has_sup β] [has_inf β] + (f : β → α) (hf : function.injective f) (one : f 1 = 1) + (mul : ∀ x y, f (x * y) = f x * f y) (npow : ∀ x (n : ℕ), f (x ^ n) = f x ^ n) + (hsup : ∀ x y, f (x ⊔ y) = max (f x) (f y)) (hinf : ∀ x y, f (x ⊓ y) = min (f x) (f y)) : + linear_ordered_cancel_comm_monoid β := +{ ..hf.linear_ordered_comm_monoid f one mul npow hsup hinf, + ..hf.ordered_cancel_comm_monoid f one mul npow } + +end linear_ordered_cancel_comm_monoid diff --git a/src/algebra/order/monoid/canonical.lean b/src/algebra/order/monoid/canonical.lean new file mode 100644 index 0000000000000..88e4ec9edfb57 --- /dev/null +++ b/src/algebra/order/monoid/canonical.lean @@ -0,0 +1,221 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.order.monoid.basic +import order.min_max + +/-! +# Canonical ordered monoids +-/ + +universe u +variables {α : Type u} + +set_option old_structure_cmd true + +/-- An `ordered_comm_monoid` with one-sided 'division' in the sense that +if `a ≤ b`, there is some `c` for which `a * c = b`. This is a weaker version +of the condition on canonical orderings defined by `canonically_ordered_monoid`. -/ +class has_exists_mul_of_le (α : Type u) [has_mul α] [has_le α] : Prop := +(exists_mul_of_le : ∀ {a b : α}, a ≤ b → ∃ (c : α), b = a * c) + +/-- An `ordered_add_comm_monoid` with one-sided 'subtraction' in the sense that +if `a ≤ b`, then there is some `c` for which `a + c = b`. This is a weaker version +of the condition on canonical orderings defined by `canonically_ordered_add_monoid`. -/ +class has_exists_add_of_le (α : Type u) [has_add α] [has_le α] : Prop := +(exists_add_of_le : ∀ {a b : α}, a ≤ b → ∃ (c : α), b = a + c) + +attribute [to_additive] has_exists_mul_of_le + +export has_exists_mul_of_le (exists_mul_of_le) + +export has_exists_add_of_le (exists_add_of_le) + +section has_exists_mul_of_le +variables [linear_order α] [densely_ordered α] [monoid α] [has_exists_mul_of_le α] + [covariant_class α α (*) (<)] [contravariant_class α α (*) (<)] {a b : α} + +@[to_additive] +lemma le_of_forall_one_lt_le_mul (h : ∀ ε : α, 1 < ε → a ≤ b * ε) : a ≤ b := +le_of_forall_le_of_dense $ λ x hxb, by { obtain ⟨ε, rfl⟩ := exists_mul_of_le hxb.le, + exact h _ ((lt_mul_iff_one_lt_right' b).1 hxb) } + +@[to_additive] +lemma le_of_forall_one_lt_lt_mul' (h : ∀ ε : α, 1 < ε → a < b * ε) : a ≤ b := +le_of_forall_one_lt_le_mul $ λ ε hε, (h _ hε).le + +@[to_additive] +lemma le_iff_forall_one_lt_lt_mul' : a ≤ b ↔ ∀ ε, 1 < ε → a < b * ε := +⟨λ h ε, lt_mul_of_le_of_one_lt h, le_of_forall_one_lt_lt_mul'⟩ + +end has_exists_mul_of_le + +/-- A canonically ordered additive monoid is an ordered commutative additive monoid + in which the ordering coincides with the subtractibility relation, + which is to say, `a ≤ b` iff there exists `c` with `b = a + c`. + This is satisfied by the natural numbers, for example, but not + the integers or other nontrivial `ordered_add_comm_group`s. -/ +@[protect_proj, ancestor ordered_add_comm_monoid has_bot] +class canonically_ordered_add_monoid (α : Type*) extends ordered_add_comm_monoid α, has_bot α := +(bot_le : ∀ x : α, ⊥ ≤ x) +(exists_add_of_le : ∀ {a b : α}, a ≤ b → ∃ c, b = a + c) +(le_self_add : ∀ a b : α, a ≤ a + b) + +@[priority 100] -- see Note [lower instance priority] +instance canonically_ordered_add_monoid.to_order_bot (α : Type u) + [h : canonically_ordered_add_monoid α] : order_bot α := +{ ..h } + +/-- A canonically ordered monoid is an ordered commutative monoid + in which the ordering coincides with the divisibility relation, + which is to say, `a ≤ b` iff there exists `c` with `b = a * c`. + Examples seem rare; it seems more likely that the `order_dual` + of a naturally-occurring lattice satisfies this than the lattice + itself (for example, dual of the lattice of ideals of a PID or + Dedekind domain satisfy this; collections of all things ≤ 1 seem to + be more natural that collections of all things ≥ 1). +-/ +@[protect_proj, ancestor ordered_comm_monoid has_bot, to_additive] +class canonically_ordered_monoid (α : Type*) extends ordered_comm_monoid α, has_bot α := +(bot_le : ∀ x : α, ⊥ ≤ x) +(exists_mul_of_le : ∀ {a b : α}, a ≤ b → ∃ c, b = a * c) +(le_self_mul : ∀ a b : α, a ≤ a * b) + +@[priority 100, to_additive] -- see Note [lower instance priority] +instance canonically_ordered_monoid.to_order_bot (α : Type u) + [h : canonically_ordered_monoid α] : order_bot α := +{ ..h } + +@[priority 100, to_additive] -- see Note [lower instance priority] +instance canonically_ordered_monoid.has_exists_mul_of_le (α : Type u) + [h : canonically_ordered_monoid α] : has_exists_mul_of_le α := +{ ..h } + +section canonically_ordered_monoid + +variables [canonically_ordered_monoid α] {a b c d : α} + +@[to_additive] lemma le_self_mul : a ≤ a * c := canonically_ordered_monoid.le_self_mul _ _ +@[to_additive] lemma le_mul_self : a ≤ b * a := by { rw mul_comm, exact le_self_mul } + +@[to_additive] lemma self_le_mul_right (a b : α) : a ≤ a * b := le_self_mul +@[to_additive] lemma self_le_mul_left (a b : α) : a ≤ b * a := le_mul_self + +@[to_additive] lemma le_of_mul_le_left : a * b ≤ c → a ≤ c := le_self_mul.trans +@[to_additive] lemma le_of_mul_le_right : a * b ≤ c → b ≤ c := le_mul_self.trans + +@[to_additive] +lemma le_iff_exists_mul : a ≤ b ↔ ∃ c, b = a * c := +⟨exists_mul_of_le, by { rintro ⟨c, rfl⟩, exact le_self_mul }⟩ + +@[to_additive] +lemma le_iff_exists_mul' : a ≤ b ↔ ∃ c, b = c * a := +by simpa only [mul_comm _ a] using le_iff_exists_mul + +@[simp, to_additive zero_le] lemma one_le (a : α) : 1 ≤ a := +le_iff_exists_mul.mpr ⟨a, (one_mul _).symm⟩ + +@[to_additive] lemma bot_eq_one : (⊥ : α) = 1 := +le_antisymm bot_le (one_le ⊥) + +@[simp, to_additive] lemma mul_eq_one_iff : a * b = 1 ↔ a = 1 ∧ b = 1 := +mul_eq_one_iff' (one_le _) (one_le _) + +@[simp, to_additive] lemma le_one_iff_eq_one : a ≤ 1 ↔ a = 1 := +(one_le a).le_iff_eq + +@[to_additive] lemma one_lt_iff_ne_one : 1 < a ↔ a ≠ 1 := +(one_le a).lt_iff_ne.trans ne_comm + +@[to_additive] lemma eq_one_or_one_lt : a = 1 ∨ 1 < a := +(one_le a).eq_or_lt.imp_left eq.symm + +@[simp, to_additive add_pos_iff] lemma one_lt_mul_iff : 1 < a * b ↔ 1 < a ∨ 1 < b := +by simp only [one_lt_iff_ne_one, ne.def, mul_eq_one_iff, not_and_distrib] + +@[to_additive] lemma exists_one_lt_mul_of_lt (h : a < b) : ∃ c (hc : 1 < c), a * c = b := +begin + obtain ⟨c, hc⟩ := le_iff_exists_mul.1 h.le, + refine ⟨c, one_lt_iff_ne_one.2 _, hc.symm⟩, + rintro rfl, + simpa [hc, lt_irrefl] using h +end + +@[to_additive] lemma le_mul_left (h : a ≤ c) : a ≤ b * c := +calc a = 1 * a : by simp + ... ≤ b * c : mul_le_mul' (one_le _) h + +@[to_additive] lemma le_mul_right (h : a ≤ b) : a ≤ b * c := +calc a = a * 1 : by simp + ... ≤ b * c : mul_le_mul' h (one_le _) + +@[to_additive] +lemma lt_iff_exists_mul [covariant_class α α (*) (<)] : a < b ↔ ∃ c > 1, b = a * c := +begin + simp_rw [lt_iff_le_and_ne, and_comm, le_iff_exists_mul, ← exists_and_distrib_left, exists_prop], + apply exists_congr, intro c, + rw [and.congr_left_iff, gt_iff_lt], rintro rfl, + split, + { rw [one_lt_iff_ne_one], apply mt, rintro rfl, rw [mul_one] }, + { rw [← (self_le_mul_right a c).lt_iff_ne], apply lt_mul_of_one_lt_right' } +end + +end canonically_ordered_monoid + +lemma pos_of_gt {M : Type*} [canonically_ordered_add_monoid M] {n m : M} (h : n < m) : 0 < m := +lt_of_le_of_lt (zero_le _) h + +@[priority 100] instance canonically_ordered_add_monoid.zero_le_one_class {M : Type*} + [canonically_ordered_add_monoid M] [has_one M] : zero_le_one_class M := +⟨zero_le 1⟩ + +/-- A canonically linear-ordered additive monoid is a canonically ordered additive monoid + whose ordering is a linear order. -/ +@[protect_proj, ancestor canonically_ordered_add_monoid linear_order] +class canonically_linear_ordered_add_monoid (α : Type*) + extends canonically_ordered_add_monoid α, linear_order α + +/-- A canonically linear-ordered monoid is a canonically ordered monoid + whose ordering is a linear order. -/ +@[protect_proj, ancestor canonically_ordered_monoid linear_order, to_additive] +class canonically_linear_ordered_monoid (α : Type*) + extends canonically_ordered_monoid α, linear_order α + +section canonically_linear_ordered_monoid +variables [canonically_linear_ordered_monoid α] + +@[priority 100, to_additive] -- see Note [lower instance priority] +instance canonically_linear_ordered_monoid.semilattice_sup : semilattice_sup α := +{ ..linear_order.to_lattice } + +@[to_additive] +lemma min_mul_distrib (a b c : α) : min a (b * c) = min a (min a b * min a c) := +begin + cases le_total a b with hb hb, + { simp [hb, le_mul_right] }, + { cases le_total a c with hc hc, + { simp [hc, le_mul_left] }, + { simp [hb, hc] } } +end + +@[to_additive] +lemma min_mul_distrib' (a b c : α) : min (a * b) c = min (min a c * min b c) c := +by simpa [min_comm _ c] using min_mul_distrib c a b + +@[simp, to_additive] +lemma one_min (a : α) : min 1 a = 1 := +min_eq_left (one_le a) + +@[simp, to_additive] +lemma min_one (a : α) : min a 1 = 1 := +min_eq_right (one_le a) + +/-- In a linearly ordered monoid, we are happy for `bot_eq_one` to be a `@[simp]` lemma. -/ +@[simp, to_additive + "In a linearly ordered monoid, we are happy for `bot_eq_zero` to be a `@[simp]` lemma"] +lemma bot_eq_one' : (⊥ : α) = 1 := +bot_eq_one + +end canonically_linear_ordered_monoid diff --git a/src/algebra/order/monoid_lemmas.lean b/src/algebra/order/monoid/lemmas.lean similarity index 100% rename from src/algebra/order/monoid_lemmas.lean rename to src/algebra/order/monoid/lemmas.lean diff --git a/src/algebra/order/monoid/min_max.lean b/src/algebra/order/monoid/min_max.lean new file mode 100644 index 0000000000000..a94dda7230c1d --- /dev/null +++ b/src/algebra/order/monoid/min_max.lean @@ -0,0 +1,98 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.order.monoid.basic +import order.min_max + +/-! +# Lemmas about `min` and `max` in an ordered monoid. +-/ + +variables {α β : Type*} + +/-! Some lemmas about types that have an ordering and a binary operation, with no + rules relating them. -/ +@[to_additive] +lemma fn_min_mul_fn_max [linear_order α] [comm_semigroup β] (f : α → β) (n m : α) : + f (min n m) * f (max n m) = f n * f m := +by { cases le_total n m with h h; simp [h, mul_comm] } + +@[to_additive] +lemma min_mul_max [linear_order α] [comm_semigroup α] (n m : α) : + min n m * max n m = n * m := +fn_min_mul_fn_max id n m + +section covariant_class_mul_le +variables [linear_order α] + +section has_mul +variable [has_mul α] + +section left +variable [covariant_class α α (*) (≤)] + +@[to_additive] lemma min_mul_mul_left (a b c : α) : min (a * b) (a * c) = a * min b c := +(monotone_id.const_mul' a).map_min.symm + +@[to_additive] +lemma max_mul_mul_left (a b c : α) : max (a * b) (a * c) = a * max b c := +(monotone_id.const_mul' a).map_max.symm + +@[to_additive] +lemma lt_or_lt_of_mul_lt_mul [covariant_class α α (function.swap (*)) (≤)] + {a b m n : α} (h : m * n < a * b) : + m < a ∨ n < b := +by { contrapose! h, exact mul_le_mul' h.1 h.2 } + +@[to_additive] +lemma mul_lt_mul_iff_of_le_of_le + [covariant_class α α (function.swap (*)) (<)] + [covariant_class α α (*) (<)] + [covariant_class α α (function.swap (*)) (≤)] + {a b c d : α} (ac : a ≤ c) (bd : b ≤ d) : + a * b < c * d ↔ (a < c) ∨ (b < d) := +begin + refine ⟨lt_or_lt_of_mul_lt_mul, λ h, _⟩, + cases h with ha hb, + { exact mul_lt_mul_of_lt_of_le ha bd }, + { exact mul_lt_mul_of_le_of_lt ac hb } +end + +end left + +section right +variable [covariant_class α α (function.swap (*)) (≤)] + +@[to_additive] +lemma min_mul_mul_right (a b c : α) : min (a * c) (b * c) = min a b * c := +(monotone_id.mul_const' c).map_min.symm + +@[to_additive] +lemma max_mul_mul_right (a b c : α) : max (a * c) (b * c) = max a b * c := +(monotone_id.mul_const' c).map_max.symm + +end right + +end has_mul + +variable [mul_one_class α] + +@[to_additive] +lemma min_le_mul_of_one_le_right [covariant_class α α (*) (≤)] {a b : α} (hb : 1 ≤ b) : + min a b ≤ a * b := +min_le_iff.2 $ or.inl $ le_mul_of_one_le_right' hb + +@[to_additive] +lemma min_le_mul_of_one_le_left [covariant_class α α (function.swap (*)) (≤)] {a b : α} + (ha : 1 ≤ a) : min a b ≤ a * b := +min_le_iff.2 $ or.inr $ le_mul_of_one_le_left' ha + +@[to_additive] +lemma max_le_mul_of_one_le [covariant_class α α (*) (≤)] + [covariant_class α α (function.swap (*)) (≤)] {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : + max a b ≤ a * b := +max_le_iff.2 ⟨le_mul_of_one_le_right' hb, le_mul_of_one_le_left' ha⟩ + +end covariant_class_mul_le diff --git a/src/algebra/order/monoid/order_dual.lean b/src/algebra/order/monoid/order_dual.lean new file mode 100644 index 0000000000000..6802456dfa4ff --- /dev/null +++ b/src/algebra/order/monoid/order_dual.lean @@ -0,0 +1,89 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ + +import algebra.group.order_synonym +import algebra.order.monoid.cancel + +/-! # Ordered monoid structures on the order dual. -/ + +universes u +variables {α : Type u} + +open function + +namespace order_dual + +@[to_additive] +instance contravariant_class_mul_le [has_le α] [has_mul α] [c : contravariant_class α α (*) (≤)] : + contravariant_class αᵒᵈ αᵒᵈ (*) (≤) := +⟨c.1.flip⟩ + +@[to_additive] +instance covariant_class_mul_le [has_le α] [has_mul α] [c : covariant_class α α (*) (≤)] : + covariant_class αᵒᵈ αᵒᵈ (*) (≤) := +⟨c.1.flip⟩ + +@[to_additive] instance contravariant_class_swap_mul_le [has_le α] [has_mul α] + [c : contravariant_class α α (swap (*)) (≤)] : + contravariant_class αᵒᵈ αᵒᵈ (swap (*)) (≤) := +⟨c.1.flip⟩ + +@[to_additive] +instance covariant_class_swap_mul_le [has_le α] [has_mul α] + [c : covariant_class α α (swap (*)) (≤)] : + covariant_class αᵒᵈ αᵒᵈ (swap (*)) (≤) := +⟨c.1.flip⟩ + +@[to_additive] +instance contravariant_class_mul_lt [has_lt α] [has_mul α] [c : contravariant_class α α (*) (<)] : + contravariant_class αᵒᵈ αᵒᵈ (*) (<) := +⟨c.1.flip⟩ + +@[to_additive] +instance covariant_class_mul_lt [has_lt α] [has_mul α] [c : covariant_class α α (*) (<)] : + covariant_class αᵒᵈ αᵒᵈ (*) (<) := +⟨c.1.flip⟩ + +@[to_additive] instance contravariant_class_swap_mul_lt [has_lt α] [has_mul α] + [c : contravariant_class α α (swap (*)) (<)] : + contravariant_class αᵒᵈ αᵒᵈ (swap (*)) (<) := +⟨c.1.flip⟩ + +@[to_additive] +instance covariant_class_swap_mul_lt [has_lt α] [has_mul α] + [c : covariant_class α α (swap (*)) (<)] : + covariant_class αᵒᵈ αᵒᵈ (swap (*)) (<) := +⟨c.1.flip⟩ + +@[to_additive] +instance [ordered_comm_monoid α] : ordered_comm_monoid αᵒᵈ := +{ mul_le_mul_left := λ a b h c, mul_le_mul_left' h c, + .. order_dual.partial_order α, + .. order_dual.comm_monoid } + +@[to_additive ordered_cancel_add_comm_monoid.to_contravariant_class] +instance ordered_cancel_comm_monoid.to_contravariant_class [ordered_cancel_comm_monoid α] : + contravariant_class αᵒᵈ αᵒᵈ has_mul.mul has_le.le := +{ elim := λ a b c, ordered_cancel_comm_monoid.le_of_mul_le_mul_left a c b } + +@[to_additive] +instance [ordered_cancel_comm_monoid α] : ordered_cancel_comm_monoid αᵒᵈ := +{ le_of_mul_le_mul_left := λ a b c : α, le_of_mul_le_mul_left', + .. order_dual.ordered_comm_monoid, .. order_dual.cancel_comm_monoid } + +@[to_additive] +instance [linear_ordered_cancel_comm_monoid α] : + linear_ordered_cancel_comm_monoid αᵒᵈ := +{ .. order_dual.linear_order α, + .. order_dual.ordered_cancel_comm_monoid } + +@[to_additive] +instance [linear_ordered_comm_monoid α] : + linear_ordered_comm_monoid αᵒᵈ := +{ .. order_dual.linear_order α, + .. order_dual.ordered_comm_monoid } + +end order_dual diff --git a/src/algebra/order/monoid/prod.lean b/src/algebra/order/monoid/prod.lean new file mode 100644 index 0000000000000..af0535d8324d8 --- /dev/null +++ b/src/algebra/order/monoid/prod.lean @@ -0,0 +1,37 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.group.prod +import algebra.order.monoid.cancel +import algebra.order.monoid.canonical + +/-! # Products of ordered monoids -/ + +namespace prod + +variables {α β M N : Type*} + +@[to_additive] +instance [ordered_comm_monoid α] [ordered_comm_monoid β] : ordered_comm_monoid (α × β) := +{ mul_le_mul_left := λ a b h c, ⟨mul_le_mul_left' h.1 _, mul_le_mul_left' h.2 _⟩, + .. prod.comm_monoid, .. prod.partial_order _ _ } + +@[to_additive] +instance [ordered_cancel_comm_monoid M] [ordered_cancel_comm_monoid N] : + ordered_cancel_comm_monoid (M × N) := +{ le_of_mul_le_mul_left := λ a b c h, ⟨le_of_mul_le_mul_left' h.1, le_of_mul_le_mul_left' h.2⟩, + .. prod.ordered_comm_monoid } + +@[to_additive] instance [has_le α] [has_le β] [has_mul α] [has_mul β] [has_exists_mul_of_le α] + [has_exists_mul_of_le β] : has_exists_mul_of_le (α × β) := +⟨λ a b h, let ⟨c, hc⟩ := exists_mul_of_le h.1, ⟨d, hd⟩ := exists_mul_of_le h.2 in + ⟨(c, d), ext hc hd⟩⟩ + +@[to_additive] instance [canonically_ordered_monoid α] [canonically_ordered_monoid β] : + canonically_ordered_monoid (α × β) := +{ le_self_mul := λ a b, ⟨le_self_mul, le_self_mul⟩, + ..prod.ordered_comm_monoid, ..prod.order_bot _ _, ..prod.has_exists_mul_of_le } + +end prod diff --git a/src/algebra/order/monoid/to_mul_bot.lean b/src/algebra/order/monoid/to_mul_bot.lean new file mode 100644 index 0000000000000..7c7777d7561d8 --- /dev/null +++ b/src/algebra/order/monoid/to_mul_bot.lean @@ -0,0 +1,42 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.order.monoid.with_top +import algebra.order.monoid.type_tags + +/-! +Making an additive monoid multiplicative then adding a zero is the same as adding a bottom +element then making it multiplicative. +-/ + +universe u +variables {α : Type u} + +namespace with_zero + +local attribute [semireducible] with_zero +variables [has_add α] + +/-- Making an additive monoid multiplicative then adding a zero is the same as adding a bottom +element then making it multiplicative. -/ +def to_mul_bot : with_zero (multiplicative α) ≃* multiplicative (with_bot α) := +by exact mul_equiv.refl _ + +@[simp] lemma to_mul_bot_zero : + to_mul_bot (0 : with_zero (multiplicative α)) = multiplicative.of_add ⊥ := rfl +@[simp] lemma to_mul_bot_coe (x : multiplicative α) : + to_mul_bot ↑x = multiplicative.of_add (x.to_add : with_bot α) := rfl +@[simp] lemma to_mul_bot_symm_bot : + to_mul_bot.symm (multiplicative.of_add (⊥ : with_bot α)) = 0 := rfl +@[simp] lemma to_mul_bot_coe_of_add (x : α) : + to_mul_bot.symm (multiplicative.of_add (x : with_bot α)) = multiplicative.of_add x := rfl + +variables [preorder α] (a b : with_zero (multiplicative α)) + +lemma to_mul_bot_strict_mono : strict_mono (@to_mul_bot α _) := λ x y, id +@[simp] lemma to_mul_bot_le : to_mul_bot a ≤ to_mul_bot b ↔ a ≤ b := iff.rfl +@[simp] lemma to_mul_bot_lt : to_mul_bot a < to_mul_bot b ↔ a < b := iff.rfl + +end with_zero diff --git a/src/algebra/order/monoid/type_tags.lean b/src/algebra/order/monoid/type_tags.lean new file mode 100644 index 0000000000000..da4904966ce75 --- /dev/null +++ b/src/algebra/order/monoid/type_tags.lean @@ -0,0 +1,108 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.group.type_tags +import algebra.order.monoid.cancel +import algebra.order.monoid.canonical + +/-! # Ordered monoid structures on `multiplicative α` and `additive α`. -/ + +universes u +variables {α : Type u} + +instance : Π [has_le α], has_le (multiplicative α) := id +instance : Π [has_le α], has_le (additive α) := id +instance : Π [has_lt α], has_lt (multiplicative α) := id +instance : Π [has_lt α], has_lt (additive α) := id +instance : Π [preorder α], preorder (multiplicative α) := id +instance : Π [preorder α], preorder (additive α) := id +instance : Π [partial_order α], partial_order (multiplicative α) := id +instance : Π [partial_order α], partial_order (additive α) := id +instance : Π [linear_order α], linear_order (multiplicative α) := id +instance : Π [linear_order α], linear_order (additive α) := id +instance [has_le α] : Π [order_bot α], order_bot (multiplicative α) := id +instance [has_le α] : Π [order_bot α], order_bot (additive α) := id +instance [has_le α] : Π [order_top α], order_top (multiplicative α) := id +instance [has_le α] : Π [order_top α], order_top (additive α) := id +instance [has_le α] : Π [bounded_order α], bounded_order (multiplicative α) := id +instance [has_le α] : Π [bounded_order α], bounded_order (additive α) := id + +instance [ordered_add_comm_monoid α] : ordered_comm_monoid (multiplicative α) := +{ mul_le_mul_left := @ordered_add_comm_monoid.add_le_add_left α _, + ..multiplicative.partial_order, + ..multiplicative.comm_monoid } + +instance [ordered_comm_monoid α] : ordered_add_comm_monoid (additive α) := +{ add_le_add_left := @ordered_comm_monoid.mul_le_mul_left α _, + ..additive.partial_order, + ..additive.add_comm_monoid } + +instance [ordered_cancel_add_comm_monoid α] : ordered_cancel_comm_monoid (multiplicative α) := +{ le_of_mul_le_mul_left := @ordered_cancel_add_comm_monoid.le_of_add_le_add_left α _, + ..multiplicative.ordered_comm_monoid } + +instance [ordered_cancel_comm_monoid α] : ordered_cancel_add_comm_monoid (additive α) := +{ le_of_add_le_add_left := @ordered_cancel_comm_monoid.le_of_mul_le_mul_left α _, + ..additive.ordered_add_comm_monoid } + +instance [linear_ordered_add_comm_monoid α] : linear_ordered_comm_monoid (multiplicative α) := +{ ..multiplicative.linear_order, + ..multiplicative.ordered_comm_monoid } + +instance [linear_ordered_comm_monoid α] : linear_ordered_add_comm_monoid (additive α) := +{ ..additive.linear_order, + ..additive.ordered_add_comm_monoid } + +instance [has_add α] [has_le α] [has_exists_add_of_le α] : + has_exists_mul_of_le (multiplicative α) := +⟨@exists_add_of_le α _ _ _⟩ + +instance [has_mul α] [has_le α] [has_exists_mul_of_le α] : has_exists_add_of_le (additive α) := +⟨@exists_mul_of_le α _ _ _⟩ + +instance [canonically_ordered_add_monoid α] : canonically_ordered_monoid (multiplicative α) := +{ le_self_mul := @le_self_add α _, + ..multiplicative.ordered_comm_monoid, ..multiplicative.order_bot, + ..multiplicative.has_exists_mul_of_le } + +instance [canonically_ordered_monoid α] : canonically_ordered_add_monoid (additive α) := +{ le_self_add := @le_self_mul α _, + ..additive.ordered_add_comm_monoid, ..additive.order_bot, ..additive.has_exists_add_of_le } + +instance [canonically_linear_ordered_add_monoid α] : + canonically_linear_ordered_monoid (multiplicative α) := +{ ..multiplicative.canonically_ordered_monoid, ..multiplicative.linear_order } + +instance [canonically_linear_ordered_monoid α] : + canonically_linear_ordered_add_monoid (additive α) := +{ ..additive.canonically_ordered_add_monoid, ..additive.linear_order } + +namespace additive + +variables [preorder α] + +@[simp] lemma of_mul_le {a b : α} : of_mul a ≤ of_mul b ↔ a ≤ b := iff.rfl + +@[simp] lemma of_mul_lt {a b : α} : of_mul a < of_mul b ↔ a < b := iff.rfl + +@[simp] lemma to_mul_le {a b : additive α} : to_mul a ≤ to_mul b ↔ a ≤ b := iff.rfl + +@[simp] lemma to_mul_lt {a b : additive α} : to_mul a < to_mul b ↔ a < b := iff.rfl + +end additive + +namespace multiplicative + +variables [preorder α] + +@[simp] lemma of_add_le {a b : α} : of_add a ≤ of_add b ↔ a ≤ b := iff.rfl + +@[simp] lemma of_add_lt {a b : α} : of_add a < of_add b ↔ a < b := iff.rfl + +@[simp] lemma to_add_le {a b : multiplicative α} : to_add a ≤ to_add b ↔ a ≤ b := iff.rfl + +@[simp] lemma to_add_lt {a b : multiplicative α} : to_add a < to_add b ↔ a < b := iff.rfl + +end multiplicative diff --git a/src/algebra/order/monoid/units.lean b/src/algebra/order/monoid/units.lean new file mode 100644 index 0000000000000..d632ac1820813 --- /dev/null +++ b/src/algebra/order/monoid/units.lean @@ -0,0 +1,51 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.order.monoid.basic +import order.min_max + +/-! +# Units in ordered monoids +-/ + +variables {α : Type*} + +namespace units + +@[to_additive] +instance [monoid α] [preorder α] : preorder αˣ := +preorder.lift (coe : αˣ → α) + +@[simp, norm_cast, to_additive] +theorem coe_le_coe [monoid α] [preorder α] {a b : αˣ} : + (a : α) ≤ b ↔ a ≤ b := iff.rfl + +@[simp, norm_cast, to_additive] +theorem coe_lt_coe [monoid α] [preorder α] {a b : αˣ} : + (a : α) < b ↔ a < b := iff.rfl + +@[to_additive] +instance [monoid α] [partial_order α] : partial_order αˣ := +partial_order.lift coe units.ext + +@[to_additive] +instance [monoid α] [linear_order α] : linear_order αˣ := +linear_order.lift' coe units.ext + +/-- `coe : αˣ → α` as an order embedding. -/ +@[to_additive "`coe : add_units α → α` as an order embedding.", simps { fully_applied := ff }] +def order_embedding_coe [monoid α] [linear_order α] : αˣ ↪o α := ⟨⟨coe, ext⟩, λ _ _, iff.rfl⟩ + +@[simp, norm_cast, to_additive] +theorem max_coe [monoid α] [linear_order α] {a b : αˣ} : + (↑(max a b) : α) = max a b := +monotone.map_max order_embedding_coe.monotone + +@[simp, norm_cast, to_additive] +theorem min_coe [monoid α] [linear_order α] {a b : αˣ} : + (↑(min a b) : α) = min a b := +monotone.map_min order_embedding_coe.monotone + +end units diff --git a/src/algebra/order/monoid/with_top.lean b/src/algebra/order/monoid/with_top.lean new file mode 100644 index 0000000000000..9b9875b1bcc07 --- /dev/null +++ b/src/algebra/order/monoid/with_top.lean @@ -0,0 +1,475 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.order.monoid.with_zero +import algebra.order.monoid.order_dual + +/-! # Adjoining top/bottom elements to ordered monoids. -/ + +universes u v +variables {α : Type u} {β : Type v} + +open function + +namespace with_top + +section has_one + +variables [has_one α] + +@[to_additive] instance : has_one (with_top α) := ⟨(1 : α)⟩ + +@[simp, norm_cast, to_additive] lemma coe_one : ((1 : α) : with_top α) = 1 := rfl + +@[simp, norm_cast, to_additive] lemma coe_eq_one {a : α} : (a : with_top α) = 1 ↔ a = 1 := +coe_eq_coe + +@[simp, to_additive] protected lemma map_one {β} (f : α → β) : + (1 : with_top α).map f = (f 1 : with_top β) := rfl + +@[simp, norm_cast, to_additive] theorem one_eq_coe {a : α} : 1 = (a : with_top α) ↔ a = 1 := +trans eq_comm coe_eq_one + +@[simp, to_additive] theorem top_ne_one : ⊤ ≠ (1 : with_top α) . +@[simp, to_additive] theorem one_ne_top : (1 : with_top α) ≠ ⊤ . + +instance [has_zero α] [has_le α] [zero_le_one_class α] : zero_le_one_class (with_top α) := +⟨some_le_some.2 zero_le_one⟩ + +end has_one + +section has_add +variables [has_add α] {a b c d : with_top α} {x y : α} + +instance : has_add (with_top α) := ⟨λ o₁ o₂, o₁.bind $ λ a, o₂.map $ (+) a⟩ + +@[norm_cast] lemma coe_add : ((x + y : α) : with_top α) = x + y := rfl +@[norm_cast] lemma coe_bit0 : ((bit0 x : α) : with_top α) = bit0 x := rfl +@[norm_cast] lemma coe_bit1 [has_one α] {a : α} : ((bit1 a : α) : with_top α) = bit1 a := rfl + +@[simp] lemma top_add (a : with_top α) : ⊤ + a = ⊤ := rfl +@[simp] lemma add_top (a : with_top α) : a + ⊤ = ⊤ := by cases a; refl + +@[simp] lemma add_eq_top : a + b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := +by cases a; cases b; simp [none_eq_top, some_eq_coe, ←with_top.coe_add, ←with_zero.coe_add] + +lemma add_ne_top : a + b ≠ ⊤ ↔ a ≠ ⊤ ∧ b ≠ ⊤ := add_eq_top.not.trans not_or_distrib + +lemma add_lt_top [partial_order α] {a b : with_top α} : a + b < ⊤ ↔ a < ⊤ ∧ b < ⊤ := +by simp_rw [lt_top_iff_ne_top, add_ne_top] + +lemma add_eq_coe : ∀ {a b : with_top α} {c : α}, + a + b = c ↔ ∃ (a' b' : α), ↑a' = a ∧ ↑b' = b ∧ a' + b' = c +| none b c := by simp [none_eq_top] +| (some a) none c := by simp [none_eq_top] +| (some a) (some b) c := + by simp only [some_eq_coe, ← coe_add, coe_eq_coe, exists_and_distrib_left, exists_eq_left] + +@[simp] lemma add_coe_eq_top_iff {x : with_top α} {y : α} : x + y = ⊤ ↔ x = ⊤ := +by { induction x using with_top.rec_top_coe; simp [← coe_add, -with_zero.coe_add] } + +@[simp] lemma coe_add_eq_top_iff {y : with_top α} : ↑x + y = ⊤ ↔ y = ⊤ := +by { induction y using with_top.rec_top_coe; simp [← coe_add, -with_zero.coe_add] } + +instance covariant_class_add_le [has_le α] [covariant_class α α (+) (≤)] : + covariant_class (with_top α) (with_top α) (+) (≤) := +⟨λ a b c h, begin + cases a; cases c; try { exact le_top }, + rcases le_coe_iff.1 h with ⟨b, rfl, h'⟩, + exact coe_le_coe.2 (add_le_add_left (coe_le_coe.1 h) _) +end⟩ + +instance covariant_class_swap_add_le [has_le α] [covariant_class α α (swap (+)) (≤)] : + covariant_class (with_top α) (with_top α) (swap (+)) (≤) := +⟨λ a b c h, begin + cases a; cases c; try { exact le_top }, + rcases le_coe_iff.1 h with ⟨b, rfl, h'⟩, + exact coe_le_coe.2 (add_le_add_right (coe_le_coe.1 h) _) +end⟩ + +instance contravariant_class_add_lt [has_lt α] [contravariant_class α α (+) (<)] : + contravariant_class (with_top α) (with_top α) (+) (<) := +⟨λ a b c h, begin + induction a using with_top.rec_top_coe, { exact (not_none_lt _ h).elim }, + induction b using with_top.rec_top_coe, { exact (not_none_lt _ h).elim }, + induction c using with_top.rec_top_coe, + { exact coe_lt_top _ }, + { exact coe_lt_coe.2 (lt_of_add_lt_add_left $ coe_lt_coe.1 h) } +end⟩ + +instance contravariant_class_swap_add_lt [has_lt α] [contravariant_class α α (swap (+)) (<)] : + contravariant_class (with_top α) (with_top α) (swap (+)) (<) := +⟨λ a b c h, begin + cases a; cases b; try { exact (not_none_lt _ h).elim }, + cases c, + { exact coe_lt_top _ }, + { exact coe_lt_coe.2 (lt_of_add_lt_add_right $ coe_lt_coe.1 h) } +end⟩ + +protected lemma le_of_add_le_add_left [has_le α] [contravariant_class α α (+) (≤)] (ha : a ≠ ⊤) + (h : a + b ≤ a + c) : b ≤ c := +begin + lift a to α using ha, + induction c using with_top.rec_top_coe, { exact le_top }, + induction b using with_top.rec_top_coe, { exact (not_top_le_coe _ h).elim }, + simp only [← coe_add, coe_le_coe] at h ⊢, + exact le_of_add_le_add_left h +end + +protected lemma le_of_add_le_add_right [has_le α] [contravariant_class α α (swap (+)) (≤)] + (ha : a ≠ ⊤) (h : b + a ≤ c + a) : b ≤ c := +begin + lift a to α using ha, + cases c, + { exact le_top }, + cases b, + { exact (not_top_le_coe _ h).elim }, + { exact coe_le_coe.2 (le_of_add_le_add_right $ coe_le_coe.1 h) } +end + +protected lemma add_lt_add_left [has_lt α] [covariant_class α α (+) (<)] (ha : a ≠ ⊤) (h : b < c) : + a + b < a + c := +begin + lift a to α using ha, + rcases lt_iff_exists_coe.1 h with ⟨b, rfl, h'⟩, + cases c, + { exact coe_lt_top _ }, + { exact coe_lt_coe.2 (add_lt_add_left (coe_lt_coe.1 h) _) } +end + +protected lemma add_lt_add_right [has_lt α] [covariant_class α α (swap (+)) (<)] + (ha : a ≠ ⊤) (h : b < c) : + b + a < c + a := +begin + lift a to α using ha, + rcases lt_iff_exists_coe.1 h with ⟨b, rfl, h'⟩, + cases c, + { exact coe_lt_top _ }, + { exact coe_lt_coe.2 (add_lt_add_right (coe_lt_coe.1 h) _) } +end + +protected lemma add_le_add_iff_left [has_le α] [covariant_class α α (+) (≤)] + [contravariant_class α α (+) (≤)] + (ha : a ≠ ⊤) : a + b ≤ a + c ↔ b ≤ c := +⟨with_top.le_of_add_le_add_left ha, λ h, add_le_add_left h a⟩ + +protected lemma add_le_add_iff_right [has_le α] [covariant_class α α (swap (+)) (≤)] + [contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊤) : b + a ≤ c + a ↔ b ≤ c := +⟨with_top.le_of_add_le_add_right ha, λ h, add_le_add_right h a⟩ + +protected lemma add_lt_add_iff_left [has_lt α] [covariant_class α α (+) (<)] + [contravariant_class α α (+) (<)] (ha : a ≠ ⊤) : a + b < a + c ↔ b < c := +⟨lt_of_add_lt_add_left, with_top.add_lt_add_left ha⟩ + +protected lemma add_lt_add_iff_right [has_lt α] [covariant_class α α (swap (+)) (<)] + [contravariant_class α α (swap (+)) (<)] (ha : a ≠ ⊤) : b + a < c + a ↔ b < c := +⟨lt_of_add_lt_add_right, with_top.add_lt_add_right ha⟩ + +protected lemma add_lt_add_of_le_of_lt [preorder α] [covariant_class α α (+) (<)] + [covariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊤) (hab : a ≤ b) (hcd : c < d) : a + c < b + d := +(with_top.add_lt_add_left ha hcd).trans_le $ add_le_add_right hab _ + +protected lemma add_lt_add_of_lt_of_le [preorder α] [covariant_class α α (+) (≤)] + [covariant_class α α (swap (+)) (<)] (hc : c ≠ ⊤) (hab : a < b) (hcd : c ≤ d) : a + c < b + d := +(with_top.add_lt_add_right hc hab).trans_le $ add_le_add_left hcd _ + +/- There is no `with_top.map_mul_of_mul_hom`, since `with_top` does not have a multiplication. -/ +@[simp] protected lemma map_add {F} [has_add β] [add_hom_class F α β] (f : F) (a b : with_top α) : + (a + b).map f = a.map f + b.map f := +begin + induction a using with_top.rec_top_coe, + { exact (top_add _).symm }, + { induction b using with_top.rec_top_coe, + { exact (add_top _).symm }, + { rw [map_coe, map_coe, ← coe_add, ← coe_add, ← map_add], + refl } }, +end + +end has_add + +instance [add_semigroup α] : add_semigroup (with_top α) := +{ add_assoc := begin + repeat { refine with_top.rec_top_coe _ _; try { intro }}; + simp [←with_top.coe_add, add_assoc] + end, + ..with_top.has_add } + +instance [add_comm_semigroup α] : add_comm_semigroup (with_top α) := +{ add_comm := + begin + repeat { refine with_top.rec_top_coe _ _; try { intro }}; + simp [←with_top.coe_add, add_comm] + end, + ..with_top.add_semigroup } + +instance [add_zero_class α] : add_zero_class (with_top α) := +{ zero_add := + begin + refine with_top.rec_top_coe _ _, + { simp }, + { intro, + rw [←with_top.coe_zero, ←with_top.coe_add, zero_add] } + end, + add_zero := + begin + refine with_top.rec_top_coe _ _, + { simp }, + { intro, + rw [←with_top.coe_zero, ←with_top.coe_add, add_zero] } + end, + ..with_top.has_zero, + ..with_top.has_add } + +instance [add_monoid α] : add_monoid (with_top α) := +{ ..with_top.add_zero_class, + ..with_top.has_zero, + ..with_top.add_semigroup } + +instance [add_comm_monoid α] : add_comm_monoid (with_top α) := +{ ..with_top.add_monoid, ..with_top.add_comm_semigroup } + +instance [add_monoid_with_one α] : add_monoid_with_one (with_top α) := +{ nat_cast := λ n, ↑(n : α), + nat_cast_zero := by rw [nat.cast_zero, with_top.coe_zero], + nat_cast_succ := λ n, by rw [nat.cast_add_one, with_top.coe_add, with_top.coe_one], + .. with_top.has_one, .. with_top.add_monoid } + +instance [add_comm_monoid_with_one α] : add_comm_monoid_with_one (with_top α) := +{ .. with_top.add_monoid_with_one, .. with_top.add_comm_monoid } + +instance [ordered_add_comm_monoid α] : ordered_add_comm_monoid (with_top α) := +{ add_le_add_left := + begin + rintros a b h (_|c), { simp [none_eq_top] }, + rcases b with (_|b), { simp [none_eq_top] }, + rcases le_coe_iff.1 h with ⟨a, rfl, h⟩, + simp only [some_eq_coe, ← coe_add, coe_le_coe] at h ⊢, + exact add_le_add_left h c + end, + ..with_top.partial_order, ..with_top.add_comm_monoid } + +instance [linear_ordered_add_comm_monoid α] : + linear_ordered_add_comm_monoid_with_top (with_top α) := +{ top_add' := with_top.top_add, + ..with_top.order_top, + ..with_top.linear_order, + ..with_top.ordered_add_comm_monoid, + ..option.nontrivial } + +instance [has_le α] [has_add α] [has_exists_add_of_le α] : has_exists_add_of_le (with_top α) := +⟨λ a b, match a, b with + | ⊤, ⊤ := by simp + | (a : α), ⊤ := λ _, ⟨⊤, rfl⟩ + | (a : α), (b : α) := λ h, begin + obtain ⟨c, rfl⟩ := exists_add_of_le (with_top.coe_le_coe.1 h), + exact ⟨c, rfl⟩ + end + | ⊤, (b : α) := λ h, (not_top_le_coe _ h).elim +end⟩ + +instance [canonically_ordered_add_monoid α] : canonically_ordered_add_monoid (with_top α) := +{ le_self_add := λ a b, match a, b with + | ⊤, ⊤ := le_rfl + | (a : α), ⊤ := le_top + | (a : α), (b : α) := with_top.coe_le_coe.2 le_self_add + | ⊤, (b : α) := le_rfl + end, + ..with_top.order_bot, ..with_top.ordered_add_comm_monoid, ..with_top.has_exists_add_of_le } + +instance [canonically_linear_ordered_add_monoid α] : + canonically_linear_ordered_add_monoid (with_top α) := +{ ..with_top.canonically_ordered_add_monoid, ..with_top.linear_order } + +@[simp, norm_cast] lemma coe_nat [add_monoid_with_one α] (n : ℕ) : ((n : α) : with_top α) = n := rfl +@[simp] lemma nat_ne_top [add_monoid_with_one α] (n : ℕ) : (n : with_top α) ≠ ⊤ := coe_ne_top +@[simp] lemma top_ne_nat [add_monoid_with_one α] (n : ℕ) : (⊤ : with_top α) ≠ n := top_ne_coe + +/-- Coercion from `α` to `with_top α` as an `add_monoid_hom`. -/ +def coe_add_hom [add_monoid α] : α →+ with_top α := +⟨coe, rfl, λ _ _, rfl⟩ + +@[simp] lemma coe_coe_add_hom [add_monoid α] : ⇑(coe_add_hom : α →+ with_top α) = coe := rfl + +@[simp] lemma zero_lt_top [ordered_add_comm_monoid α] : (0 : with_top α) < ⊤ := +coe_lt_top 0 + +@[simp, norm_cast] lemma zero_lt_coe [ordered_add_comm_monoid α] (a : α) : + (0 : with_top α) < a ↔ 0 < a := +coe_lt_coe + +/-- A version of `with_top.map` for `one_hom`s. -/ +@[to_additive "A version of `with_top.map` for `zero_hom`s", simps { fully_applied := ff }] +protected def _root_.one_hom.with_top_map {M N : Type*} [has_one M] [has_one N] (f : one_hom M N) : + one_hom (with_top M) (with_top N) := +{ to_fun := with_top.map f, + map_one' := by rw [with_top.map_one, map_one, coe_one] } + +/-- A version of `with_top.map` for `add_hom`s. -/ +@[simps { fully_applied := ff }] protected def _root_.add_hom.with_top_map + {M N : Type*} [has_add M] [has_add N] (f : add_hom M N) : + add_hom (with_top M) (with_top N) := +{ to_fun := with_top.map f, + map_add' := with_top.map_add f } + +/-- A version of `with_top.map` for `add_monoid_hom`s. -/ +@[simps { fully_applied := ff }] protected def _root_.add_monoid_hom.with_top_map + {M N : Type*} [add_zero_class M] [add_zero_class N] (f : M →+ N) : + with_top M →+ with_top N := +{ to_fun := with_top.map f, + .. f.to_zero_hom.with_top_map, .. f.to_add_hom.with_top_map } + +end with_top + +namespace with_bot + +@[to_additive] instance [has_one α] : has_one (with_bot α) := with_top.has_one +instance [has_add α] : has_add (with_bot α) := with_top.has_add +instance [add_semigroup α] : add_semigroup (with_bot α) := with_top.add_semigroup +instance [add_comm_semigroup α] : add_comm_semigroup (with_bot α) := with_top.add_comm_semigroup +instance [add_zero_class α] : add_zero_class (with_bot α) := with_top.add_zero_class +instance [add_monoid α] : add_monoid (with_bot α) := with_top.add_monoid +instance [add_comm_monoid α] : add_comm_monoid (with_bot α) := with_top.add_comm_monoid +instance [add_monoid_with_one α] : add_monoid_with_one (with_bot α) := with_top.add_monoid_with_one + +instance [add_comm_monoid_with_one α] : add_comm_monoid_with_one (with_bot α) := +with_top.add_comm_monoid_with_one + +instance [has_zero α] [has_one α] [has_le α] [zero_le_one_class α] : + zero_le_one_class (with_bot α) := +⟨some_le_some.2 zero_le_one⟩ + +-- `by norm_cast` proves this lemma, so I did not tag it with `norm_cast` +@[to_additive] +lemma coe_one [has_one α] : ((1 : α) : with_bot α) = 1 := rfl + +-- `by norm_cast` proves this lemma, so I did not tag it with `norm_cast` +@[to_additive] +lemma coe_eq_one [has_one α] {a : α} : (a : with_bot α) = 1 ↔ a = 1 := +with_top.coe_eq_one + +@[to_additive] protected lemma map_one {β} [has_one α] (f : α → β) : + (1 : with_bot α).map f = (f 1 : with_bot β) := rfl + +@[norm_cast] lemma coe_nat [add_monoid_with_one α] (n : ℕ) : ((n : α) : with_bot α) = n := rfl +@[simp] lemma nat_ne_bot [add_monoid_with_one α] (n : ℕ) : (n : with_bot α) ≠ ⊥ := coe_ne_bot +@[simp] lemma bot_ne_nat [add_monoid_with_one α] (n : ℕ) : (⊥ : with_bot α) ≠ n := bot_ne_coe + +section has_add +variables [has_add α] {a b c d : with_bot α} {x y : α} + +-- `norm_cast` proves those lemmas, because `with_top`/`with_bot` are reducible +lemma coe_add (a b : α) : ((a + b : α) : with_bot α) = a + b := rfl +lemma coe_bit0 : ((bit0 x : α) : with_bot α) = bit0 x := rfl +lemma coe_bit1 [has_one α] {a : α} : ((bit1 a : α) : with_bot α) = bit1 a := rfl + +@[simp] lemma bot_add (a : with_bot α) : ⊥ + a = ⊥ := rfl +@[simp] lemma add_bot (a : with_bot α) : a + ⊥ = ⊥ := by cases a; refl + +@[simp] lemma add_eq_bot : a + b = ⊥ ↔ a = ⊥ ∨ b = ⊥ := with_top.add_eq_top +lemma add_ne_bot : a + b ≠ ⊥ ↔ a ≠ ⊥ ∧ b ≠ ⊥ := with_top.add_ne_top + +lemma bot_lt_add [partial_order α] {a b : with_bot α} : ⊥ < a + b ↔ ⊥ < a ∧ ⊥ < b := +@with_top.add_lt_top αᵒᵈ _ _ _ _ + +lemma add_eq_coe : a + b = x ↔ ∃ (a' b' : α), ↑a' = a ∧ ↑b' = b ∧ a' + b' = x := with_top.add_eq_coe + +@[simp] lemma add_coe_eq_bot_iff : a + y = ⊥ ↔ a = ⊥ := with_top.add_coe_eq_top_iff +@[simp] lemma coe_add_eq_bot_iff : ↑x + b = ⊥ ↔ b = ⊥ := with_top.coe_add_eq_top_iff + +/- There is no `with_bot.map_mul_of_mul_hom`, since `with_bot` does not have a multiplication. -/ +@[simp] protected lemma map_add {F} [has_add β] [add_hom_class F α β] (f : F) (a b : with_bot α) : + (a + b).map f = a.map f + b.map f := +with_top.map_add f a b + +/-- A version of `with_bot.map` for `one_hom`s. -/ +@[to_additive "A version of `with_bot.map` for `zero_hom`s", simps { fully_applied := ff }] +protected def _root_.one_hom.with_bot_map {M N : Type*} [has_one M] [has_one N] (f : one_hom M N) : + one_hom (with_bot M) (with_bot N) := +{ to_fun := with_bot.map f, + map_one' := by rw [with_bot.map_one, map_one, coe_one] } + +/-- A version of `with_bot.map` for `add_hom`s. -/ +@[simps { fully_applied := ff }] protected def _root_.add_hom.with_bot_map + {M N : Type*} [has_add M] [has_add N] (f : add_hom M N) : + add_hom (with_bot M) (with_bot N) := +{ to_fun := with_bot.map f, + map_add' := with_bot.map_add f } + +/-- A version of `with_bot.map` for `add_monoid_hom`s. -/ +@[simps { fully_applied := ff }] protected def _root_.add_monoid_hom.with_bot_map + {M N : Type*} [add_zero_class M] [add_zero_class N] (f : M →+ N) : + with_bot M →+ with_bot N := +{ to_fun := with_bot.map f, + .. f.to_zero_hom.with_bot_map, .. f.to_add_hom.with_bot_map } + +variables [preorder α] + +instance covariant_class_add_le [covariant_class α α (+) (≤)] : + covariant_class (with_bot α) (with_bot α) (+) (≤) := +@order_dual.covariant_class_add_le (with_top αᵒᵈ) _ _ _ + +instance covariant_class_swap_add_le [covariant_class α α (swap (+)) (≤)] : + covariant_class (with_bot α) (with_bot α) (swap (+)) (≤) := +@order_dual.covariant_class_swap_add_le (with_top αᵒᵈ) _ _ _ + +instance contravariant_class_add_lt [contravariant_class α α (+) (<)] : + contravariant_class (with_bot α) (with_bot α) (+) (<) := +@order_dual.contravariant_class_add_lt (with_top αᵒᵈ) _ _ _ + +instance contravariant_class_swap_add_lt [contravariant_class α α (swap (+)) (<)] : + contravariant_class (with_bot α) (with_bot α) (swap (+)) (<) := +@order_dual.contravariant_class_swap_add_lt (with_top αᵒᵈ) _ _ _ + +protected lemma le_of_add_le_add_left [contravariant_class α α (+) (≤)] (ha : a ≠ ⊥) + (h : a + b ≤ a + c) : b ≤ c := +@with_top.le_of_add_le_add_left αᵒᵈ _ _ _ _ _ _ ha h + +protected lemma le_of_add_le_add_right [contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊥) + (h : b + a ≤ c + a) : b ≤ c := +@with_top.le_of_add_le_add_right αᵒᵈ _ _ _ _ _ _ ha h + +protected lemma add_lt_add_left [covariant_class α α (+) (<)] (ha : a ≠ ⊥) (h : b < c) : + a + b < a + c := +@with_top.add_lt_add_left αᵒᵈ _ _ _ _ _ _ ha h + +protected lemma add_lt_add_right [covariant_class α α (swap (+)) (<)] (ha : a ≠ ⊥) (h : b < c) : + b + a < c + a := +@with_top.add_lt_add_right αᵒᵈ _ _ _ _ _ _ ha h + +protected lemma add_le_add_iff_left [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤)] + (ha : a ≠ ⊥) : a + b ≤ a + c ↔ b ≤ c := +⟨with_bot.le_of_add_le_add_left ha, λ h, add_le_add_left h a⟩ + +protected lemma add_le_add_iff_right [covariant_class α α (swap (+)) (≤)] + [contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊥) : b + a ≤ c + a ↔ b ≤ c := +⟨with_bot.le_of_add_le_add_right ha, λ h, add_le_add_right h a⟩ + +protected lemma add_lt_add_iff_left [covariant_class α α (+) (<)] [contravariant_class α α (+) (<)] + (ha : a ≠ ⊥) : a + b < a + c ↔ b < c := +⟨lt_of_add_lt_add_left, with_bot.add_lt_add_left ha⟩ + +protected lemma add_lt_add_iff_right [covariant_class α α (swap (+)) (<)] + [contravariant_class α α (swap (+)) (<)] (ha : a ≠ ⊥) : b + a < c + a ↔ b < c := +⟨lt_of_add_lt_add_right, with_bot.add_lt_add_right ha⟩ + +protected lemma add_lt_add_of_le_of_lt [covariant_class α α (+) (<)] + [covariant_class α α (swap (+)) (≤)] (hb : b ≠ ⊥) (hab : a ≤ b) (hcd : c < d) : a + c < b + d := +@with_top.add_lt_add_of_le_of_lt αᵒᵈ _ _ _ _ _ _ _ _ hb hab hcd + +protected lemma add_lt_add_of_lt_of_le [covariant_class α α (+) (≤)] + [covariant_class α α (swap (+)) (<)] (hd : d ≠ ⊥) (hab : a < b) (hcd : c ≤ d) : a + c < b + d := +@with_top.add_lt_add_of_lt_of_le αᵒᵈ _ _ _ _ _ _ _ _ hd hab hcd + +end has_add + +instance [ordered_add_comm_monoid α] : ordered_add_comm_monoid (with_bot α) := +{ add_le_add_left := λ a b h c, add_le_add_left h c, + ..with_bot.partial_order, + ..with_bot.add_comm_monoid } + +instance [linear_ordered_add_comm_monoid α] : linear_ordered_add_comm_monoid (with_bot α) := +{ ..with_bot.linear_order, ..with_bot.ordered_add_comm_monoid } + +end with_bot diff --git a/src/algebra/order/monoid/with_zero.lean b/src/algebra/order/monoid/with_zero.lean new file mode 100644 index 0000000000000..38922f65d2e3b --- /dev/null +++ b/src/algebra/order/monoid/with_zero.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.order.monoid.canonical +import algebra.group.with_one + +/-! +# Adjoining a zero element to an ordered monoid. +-/ + +universe u +variables {α : Type u} + +namespace with_zero + +local attribute [semireducible] with_zero + +instance [preorder α] : preorder (with_zero α) := with_bot.preorder + +instance [partial_order α] : partial_order (with_zero α) := with_bot.partial_order + +instance [preorder α] : order_bot (with_zero α) := with_bot.order_bot + +lemma zero_le [preorder α] (a : with_zero α) : 0 ≤ a := bot_le + +lemma zero_lt_coe [preorder α] (a : α) : (0 : with_zero α) < a := with_bot.bot_lt_coe a + +lemma zero_eq_bot [preorder α] : (0 : with_zero α) = ⊥ := rfl + +@[simp, norm_cast] lemma coe_lt_coe [preorder α] {a b : α} : (a : with_zero α) < b ↔ a < b := +with_bot.coe_lt_coe + +@[simp, norm_cast] lemma coe_le_coe [preorder α] {a b : α} : (a : with_zero α) ≤ b ↔ a ≤ b := +with_bot.coe_le_coe + +instance [lattice α] : lattice (with_zero α) := with_bot.lattice + +instance [linear_order α] : linear_order (with_zero α) := with_bot.linear_order + +instance covariant_class_mul_le {α : Type u} [has_mul α] [preorder α] + [covariant_class α α (*) (≤)] : + covariant_class (with_zero α) (with_zero α) (*) (≤) := +begin + refine ⟨λ a b c hbc, _⟩, + induction a using with_zero.rec_zero_coe, { exact zero_le _ }, + induction b using with_zero.rec_zero_coe, { exact zero_le _ }, + rcases with_bot.coe_le_iff.1 hbc with ⟨c, rfl, hbc'⟩, + rw [← coe_mul, ← coe_mul, coe_le_coe], + exact mul_le_mul_left' hbc' a +end + +instance contravariant_class_mul_lt {α : Type u} [has_mul α] [partial_order α] + [contravariant_class α α (*) (<)] : + contravariant_class (with_zero α) (with_zero α) (*) (<) := +begin + refine ⟨λ a b c h, _⟩, + have := ((zero_le _).trans_lt h).ne', + lift a to α using left_ne_zero_of_mul this, + lift c to α using right_ne_zero_of_mul this, + induction b using with_zero.rec_zero_coe, + exacts [zero_lt_coe _, coe_lt_coe.mpr (lt_of_mul_lt_mul_left' $ coe_lt_coe.mp h)] +end + +@[simp] lemma le_max_iff [linear_order α] {a b c : α} : + (a : with_zero α) ≤ max b c ↔ a ≤ max b c := +by simp only [with_zero.coe_le_coe, le_max_iff] + +@[simp] lemma min_le_iff [linear_order α] {a b c : α} : + min (a : with_zero α) b ≤ c ↔ min a b ≤ c := +by simp only [with_zero.coe_le_coe, min_le_iff] + +instance [ordered_comm_monoid α] : ordered_comm_monoid (with_zero α) := +{ mul_le_mul_left := λ _ _, mul_le_mul_left', + ..with_zero.comm_monoid_with_zero, + ..with_zero.partial_order } + +protected lemma covariant_class_add_le [add_zero_class α] [preorder α] + [covariant_class α α (+) (≤)] (h : ∀ a : α, 0 ≤ a) : + covariant_class (with_zero α) (with_zero α) (+) (≤) := +begin + refine ⟨λ a b c hbc, _⟩, + induction a using with_zero.rec_zero_coe, + { rwa [zero_add, zero_add] }, + induction b using with_zero.rec_zero_coe, + { rw [add_zero], + induction c using with_zero.rec_zero_coe, + { rw [add_zero], exact le_rfl }, + { rw [← coe_add, coe_le_coe], + exact le_add_of_nonneg_right (h _) } }, + { rcases with_bot.coe_le_iff.1 hbc with ⟨c, rfl, hbc'⟩, + rw [← coe_add, ← coe_add, coe_le_coe], + exact add_le_add_left hbc' a } +end + +/- +Note 1 : the below is not an instance because it requires `zero_le`. It seems +like a rather pathological definition because α already has a zero. +Note 2 : there is no multiplicative analogue because it does not seem necessary. +Mathematicians might be more likely to use the order-dual version, where all +elements are ≤ 1 and then 1 is the top element. +-/ + +/-- +If `0` is the least element in `α`, then `with_zero α` is an `ordered_add_comm_monoid`. +See note [reducible non-instances]. +-/ +@[reducible] protected def ordered_add_comm_monoid [ordered_add_comm_monoid α] + (zero_le : ∀ a : α, 0 ≤ a) : ordered_add_comm_monoid (with_zero α) := +{ add_le_add_left := @add_le_add_left _ _ _ (with_zero.covariant_class_add_le zero_le), + ..with_zero.partial_order, + ..with_zero.add_comm_monoid, .. } + +end with_zero + +section canonically_ordered_monoid + +instance with_zero.has_exists_add_of_le {α} [has_add α] [preorder α] [has_exists_add_of_le α] : + has_exists_add_of_le (with_zero α) := +⟨λ a b, begin + apply with_zero.cases_on a, + { exact λ _, ⟨b, (zero_add b).symm⟩ }, + apply with_zero.cases_on b, + { exact λ b' h, (with_bot.not_coe_le_bot _ h).elim }, + rintro a' b' h, + obtain ⟨c, rfl⟩ := exists_add_of_le (with_zero.coe_le_coe.1 h), + exact ⟨c, rfl⟩, +end⟩ + +-- This instance looks absurd: a monoid already has a zero +/-- Adding a new zero to a canonically ordered additive monoid produces another one. -/ +instance with_zero.canonically_ordered_add_monoid {α : Type u} [canonically_ordered_add_monoid α] : + canonically_ordered_add_monoid (with_zero α) := +{ le_self_add := λ a b, begin + apply with_zero.cases_on a, + { exact bot_le }, + apply with_zero.cases_on b, + { exact λ b', le_rfl }, + { exact λ a' b', with_zero.coe_le_coe.2 le_self_add } + end, + .. with_zero.order_bot, + .. with_zero.ordered_add_comm_monoid zero_le, ..with_zero.has_exists_add_of_le } + +end canonically_ordered_monoid + +section canonically_linear_ordered_monoid + +instance with_zero.canonically_linear_ordered_add_monoid + (α : Type*) [canonically_linear_ordered_add_monoid α] : + canonically_linear_ordered_add_monoid (with_zero α) := +{ .. with_zero.canonically_ordered_add_monoid, + .. with_zero.linear_order } + +end canonically_linear_ordered_monoid diff --git a/src/algebra/order/sub.lean b/src/algebra/order/sub.lean index e1b90310e8cd6..4657f9907cd89 100644 --- a/src/algebra/order/sub.lean +++ b/src/algebra/order/sub.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import algebra.order.monoid +import algebra.order.monoid.with_top /-! # Ordered Subtraction diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index 9d019991854ea..0173838fee2d8 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Johan Commelin, Patrick Massot -/ import algebra.order.group +import algebra.order.monoid.with_zero /-! # Linearly ordered commutative groups and monoids with a zero element adjoined diff --git a/src/algebra/order/zero_le_one.lean b/src/algebra/order/zero_le_one.lean new file mode 100644 index 0000000000000..bf909b2eda78e --- /dev/null +++ b/src/algebra/order/zero_le_one.lean @@ -0,0 +1,49 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.group.defs +import algebra.covariant_and_contravariant +import algebra.order.monoid.lemmas + +/-! +# Typeclass expressing `0 ≤ 1`. +-/ + +variables {α : Type*} + +open function + +/-- Typeclass for expressing that the `0` of a type is less or equal to its `1`. -/ +class zero_le_one_class (α : Type*) [has_zero α] [has_one α] [has_le α] := +(zero_le_one : (0 : α) ≤ 1) + +@[simp] lemma zero_le_one [has_zero α] [has_one α] [has_le α] [zero_le_one_class α] : (0 : α) ≤ 1 := +zero_le_one_class.zero_le_one + +/- `zero_le_one` with an explicit type argument. -/ +lemma zero_le_one' (α) [has_zero α] [has_one α] [has_le α] [zero_le_one_class α] : (0 : α) ≤ 1 := +zero_le_one + +lemma zero_le_two [preorder α] [has_one α] [add_zero_class α] [zero_le_one_class α] + [covariant_class α α (+) (≤)] : (0 : α) ≤ 2 := +add_nonneg zero_le_one zero_le_one + +lemma zero_le_three [preorder α] [has_one α] [add_zero_class α] [zero_le_one_class α] + [covariant_class α α (+) (≤)] : (0 : α) ≤ 3 := +add_nonneg zero_le_two zero_le_one + +lemma zero_le_four [preorder α] [has_one α] [add_zero_class α] [zero_le_one_class α] + [covariant_class α α (+) (≤)] : (0 : α) ≤ 4 := +add_nonneg zero_le_two zero_le_two + +lemma one_le_two [has_le α] [has_one α] [add_zero_class α] [zero_le_one_class α] + [covariant_class α α (+) (≤)] : (1 : α) ≤ 2 := +calc 1 = 1 + 0 : (add_zero 1).symm + ... ≤ 1 + 1 : add_le_add_left zero_le_one _ + +lemma one_le_two' [has_le α] [has_one α] [add_zero_class α] [zero_le_one_class α] + [covariant_class α α (swap (+)) (≤)] : (1 : α) ≤ 2 := +calc 1 = 0 + 1 : (zero_add 1).symm + ... ≤ 1 + 1 : add_le_add_right zero_le_one _ diff --git a/src/algebra/regular/basic.lean b/src/algebra/regular/basic.lean index 2deee3838b6f4..a5bd132fc2a7b 100644 --- a/src/algebra/regular/basic.lean +++ b/src/algebra/regular/basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ import algebra.group.commute -import algebra.order.monoid_lemmas +import algebra.order.monoid.lemmas import algebra.group_with_zero.basic import logic.embedding From ceecc887fe7d4ec3d28df56613f8759f398da0ad Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 9 Oct 2022 22:10:27 +0000 Subject: [PATCH 647/828] chore(algebra/ring/order): avoid import of data.set.intervals (#16869) Co-authored-by: Scott Morrison --- src/algebra/group_power/order.lean | 1 + src/algebra/order/ring.lean | 7 +++++-- 2 files changed, 6 insertions(+), 2 deletions(-) diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index 8ec4e2e08b05a..9c9fbd50f7a90 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -5,6 +5,7 @@ Authors: Jeremy Avigad, Robert Y. Lewis -/ import algebra.order.ring import algebra.group_power.ring +import data.set.intervals.basic /-! # Lemmas about the interaction of power operations with order diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 052d9d104eb99..5fb15cbba7f82 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -7,7 +7,6 @@ import algebra.char_zero.defs import algebra.hom.ring import algebra.order.group import algebra.order.ring_lemmas -import data.set.intervals.basic /-! # Ordered rings and semirings @@ -471,7 +470,11 @@ lemma mul_lt_mul' (hac : a ≤ c) (hbd : b < d) (hb : 0 ≤ b) (hc : 0 < c) : a lemma mul_self_lt_mul_self (h1 : 0 ≤ a) (h2 : a < b) : a * a < b * b := mul_lt_mul' h2.le h2 h1 $ h1.trans_lt h2 -lemma strict_mono_on_mul_self : strict_mono_on (λ x : α, x * x) (set.Ici 0) := +-- In the next lemma, we used to write `set.Ici 0` instead of `{x | 0 ≤ x}`. +-- As this lemma is not used outside this file, +-- and the import for `set.Ici` is not otherwise needed until later, +-- we choose not to use it here. +lemma strict_mono_on_mul_self : strict_mono_on (λ x : α, x * x) {x | 0 ≤ x} := λ x hx y hy hxy, mul_self_lt_mul_self hx hxy -- See Note [decidable namespace] From 9fc0f22750a86ace4302a7d941b446aec896924b Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 10 Oct 2022 03:41:49 +0000 Subject: [PATCH 648/828] feat(ring_theory/noetherian): The nilradical of a noetherian ring is nilpotent (#16881) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/ring_theory/ideal/operations.lean | 2 ++ src/ring_theory/noetherian.lean | 28 ++++++++++++++++++++++++++- 2 files changed, 29 insertions(+), 1 deletion(-) diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 5a111557b1ab6..0f70cfb04772f 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -310,6 +310,8 @@ variables {R : Type u} [semiring R] @[simp] lemma add_eq_sup {I J : ideal R} : I + J = I ⊔ J := rfl @[simp] lemma zero_eq_bot : (0 : ideal R) = ⊥ := rfl +@[simp] lemma sum_eq_sup {ι : Type*} (s : finset ι) (f : ι → ideal R) : s.sum f = s.sup f := rfl + end add section mul_and_radical diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index d5f7472f8141b..907a45df05994 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -7,7 +7,7 @@ import group_theory.finiteness import data.multiset.finset_ops import algebra.algebra.tower import order.order_iso_nat -import ring_theory.ideal.operations +import ring_theory.nilpotent import order.compactly_generated import linear_algebra.linear_independent import algebra.ring.idempotents @@ -427,6 +427,25 @@ begin { rintro (rfl|rfl); simp [is_idempotent_elem] } end +lemma exists_radical_pow_le_of_fg {R : Type*} [comm_semiring R] (I : ideal R) (h : I.radical.fg) : + ∃ n : ℕ, I.radical ^ n ≤ I := +begin + have := le_refl I.radical, revert this, + refine submodule.fg_induction _ _ (λ J, J ≤ I.radical → ∃ n : ℕ, J ^ n ≤ I) _ _ _ h, + { intros x hx, obtain ⟨n, hn⟩ := hx (subset_span (set.mem_singleton x)), + exact ⟨n, by rwa [← ideal.span, span_singleton_pow, span_le, set.singleton_subset_iff]⟩ }, + { intros J K hJ hK hJK, + obtain ⟨n, hn⟩ := hJ (λ x hx, hJK $ ideal.mem_sup_left hx), + obtain ⟨m, hm⟩ := hK (λ x hx, hJK $ ideal.mem_sup_right hx), + use n + m, + rw [← ideal.add_eq_sup, add_pow, ideal.sum_eq_sup, finset.sup_le_iff], + refine λ i hi, ideal.mul_le_right.trans _, + obtain h | h := le_or_lt n i, + { exact ideal.mul_le_right.trans ((ideal.pow_le_pow h).trans hn) }, + { refine ideal.mul_le_left.trans ((ideal.pow_le_pow _).trans hm), + rw [add_comm, nat.add_sub_assoc h.le], apply nat.le_add_right } }, +end + end ideal /-- @@ -888,6 +907,13 @@ theorem is_noetherian_ring_of_ring_equiv (R) [ring R] {S} [ring S] (f : R ≃+* S) [is_noetherian_ring R] : is_noetherian_ring S := is_noetherian_ring_of_surjective R S f.to_ring_hom f.to_equiv.surjective +lemma is_noetherian_ring.is_nilpotent_nilradical (R : Type*) [comm_ring R] [is_noetherian_ring R] : + is_nilpotent (nilradical R) := +begin + obtain ⟨n, hn⟩ := ideal.exists_radical_pow_le_of_fg (⊥ : ideal R) (is_noetherian.noetherian _), + exact ⟨n, eq_bot_iff.mpr hn⟩ +end + namespace submodule section map₂ From 91db83321de2352cb26a4ec60f6f5ad568b54463 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Mon, 10 Oct 2022 05:50:05 +0000 Subject: [PATCH 649/828] feat(data/set/lattice): Product of `sInter` (#16692) Lemma to show that the product of two non-empty set intersections is equal to the intersection of the set of products. Co-authored-by: Christopher Hoskin Co-authored-by: Christopher Hoskin --- src/data/set/lattice.lean | 27 +++++++++++++++++++++++++++ 1 file changed, 27 insertions(+) diff --git a/src/data/set/lattice.lean b/src/data/set/lattice.lean index 60165659b97db..69f34fce9cbdf 100644 --- a/src/data/set/lattice.lean +++ b/src/data/set/lattice.lean @@ -1377,6 +1377,33 @@ begin { intros x hz x' hw, exact ⟨x ⊔ x', hs le_sup_left hz, ht le_sup_right hw⟩ } end +lemma sInter_prod_sInter_subset (S : set (set α)) (T : set (set β)) : + ⋂₀ S ×ˢ ⋂₀ T ⊆ ⋂ r ∈ S ×ˢ T, r.1 ×ˢ r.2 := +subset_Inter₂ (λ x hx y hy, ⟨hy.1 x.1 hx.1, hy.2 x.2 hx.2⟩) + +lemma sInter_prod_sInter {S : set (set α)} {T : set (set β)} (hS : S.nonempty) (hT : T.nonempty) : + ⋂₀ S ×ˢ ⋂₀ T = ⋂ r ∈ S ×ˢ T, r.1 ×ˢ r.2 := +begin + obtain ⟨s₁, h₁⟩ := hS, + obtain ⟨s₂, h₂⟩ := hT, + refine set.subset.antisymm (sInter_prod_sInter_subset S T) (λ x hx, _), + rw mem_Inter₂ at hx, + exact ⟨λ s₀ h₀, (hx (s₀, s₂) ⟨h₀, h₂⟩).1, λ s₀ h₀, (hx (s₁, s₀) ⟨h₁, h₀⟩).2⟩, +end + +lemma sInter_prod {S : set (set α)} (hS : S.nonempty) (t : set β) : + ⋂₀ S ×ˢ t = ⋂ s ∈ S, s ×ˢ t := +begin + rw [←sInter_singleton t, sInter_prod_sInter hS (singleton_nonempty t), sInter_singleton], + simp_rw [prod_singleton, mem_image, Inter_exists, bInter_and', Inter_Inter_eq_right], +end + +lemma prod_sInter {T : set (set β)} (hT : T.nonempty) (s : set α) : + s ×ˢ ⋂₀ T = ⋂ t ∈ T, s ×ˢ t := +begin + rw [←sInter_singleton s, sInter_prod_sInter (singleton_nonempty s) hT, sInter_singleton], + simp_rw [singleton_prod, mem_image, Inter_exists, bInter_and', Inter_Inter_eq_right], +end end prod section image2 From 8434f767b0a7b54ce4077e13da190ae51ff353bf Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 10 Oct 2022 05:50:07 +0000 Subject: [PATCH 650/828] chore(analysis/normed_space/star/complex): delete file (#16871) We now have two versions of this. I am keeping the other one instead of this because: (a) it lives in the right place, (b) it has associated notation, (c) it has slightly more API, and (d) it has already been used somewhere. --- src/analysis/normed_space/star/complex.lean | 59 --------------------- 1 file changed, 59 deletions(-) delete mode 100644 src/analysis/normed_space/star/complex.lean diff --git a/src/analysis/normed_space/star/complex.lean b/src/analysis/normed_space/star/complex.lean deleted file mode 100644 index 512dd0d99440f..0000000000000 --- a/src/analysis/normed_space/star/complex.lean +++ /dev/null @@ -1,59 +0,0 @@ -/- -Copyright (c) 2022 Frédéric Dupuis. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Frédéric Dupuis --/ - -import analysis.normed_space.star.basic -import algebra.star.module -import analysis.complex.basic - -/-! -# Complex normed star modules and algebras - -Facts about star modules and star algebras over the complex numbers. - -## Main definitions - -* `star_module.mul_neg_I_lin`: multiplication by -I as a real-linear equivalence between the - skew-adjoint and self-adjoint elements of a star module. -* `star_module.im`: the imaginary part of an element of a star module, defined via - `skew_adjoint_part`. - --/ - -variables {E : Type*} - -namespace star_module -open_locale complex_conjugate -open complex - -variables [add_comm_group E] [star_add_monoid E] [module ℂ E] [star_module ℂ E] - -/-- Multiplication by -I as a real-linear equivalence between the skew-adjoint and self-adjoint -elements of a star module. -/ -@[simps] def mul_neg_I_lin : skew_adjoint E ≃ₗ[ℝ] self_adjoint E := -{ to_fun := λ x, ⟨-I • x, by simp [self_adjoint.mem_iff]⟩, - inv_fun := λ x, ⟨I • x, by simp [skew_adjoint.mem_iff]⟩, - map_add' := λ x y, by { ext, simp only [add_subgroup.coe_add, smul_add, add_subgroup.coe_mk] }, - map_smul' := λ r x, by { ext, simp only [neg_smul, neg_inj, skew_adjoint.coe_smul, - add_subgroup.coe_mk, ring_hom.id_apply, self_adjoint.coe_smul, smul_neg, smul_comm I], }, - left_inv := λ x, by simp only [neg_smul, add_subgroup.coe_mk, smul_neg, ←mul_smul, I_mul_I, - neg_neg, one_smul, set_like.eta], - right_inv := λ x, by simp only [←mul_smul, I_mul_I, add_subgroup.coe_mk, neg_mul, neg_neg, - one_smul, set_like.eta] } - -/-- The imaginary part of an element of a star module, as a real-linear map. -/ -@[simps] noncomputable def im : E →ₗ[ℝ] self_adjoint E := - mul_neg_I_lin.to_linear_map.comp (skew_adjoint_part ℝ) - -/-- The real part of an element of a star module, as a real-linear map. This is simply an -abbreviation for `self_adjoint_part ℝ`. -/ -@[simps] noncomputable abbreviation re : E →ₗ[ℝ] self_adjoint E := self_adjoint_part ℝ - -/-- An element of a complex star module can be decomposed into self-adjoint "real" and -"imaginary" parts -/ -lemma re_add_im (x : E) : (re x : E) + I • im x = x := -by simp [←mul_smul, I_mul_I, ←smul_add, ←two_smul ℝ] - -end star_module From 6c31dd6563a3745bf8e0b80bdd077167583ebb8f Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 10 Oct 2022 05:50:08 +0000 Subject: [PATCH 651/828] perf(topology/sheaves/presheaf_of_functions): speedups (#16873) Speed up`CommRing_yoneda` (which caused [a bors failure](https://github.com/leanprover-community/mathlib/actions/runs/3211416019/jobs/5249570763)) from >20s to <2.5s by filling in automatic fields. (merged in another PR) Speed up `presheaf_to_Types` from 13.3s to 0.3s and `presheaf_to_Type` from 4s to 0.1s by filling in automatic fields. --- src/topology/sheaves/presheaf_of_functions.lean | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/src/topology/sheaves/presheaf_of_functions.lean b/src/topology/sheaves/presheaf_of_functions.lean index db5f92930671d..216bb9f08f765 100644 --- a/src/topology/sheaves/presheaf_of_functions.lean +++ b/src/topology/sheaves/presheaf_of_functions.lean @@ -41,7 +41,9 @@ There is no requirement that the functions are continuous, here. -/ def presheaf_to_Types (T : X → Type v) : X.presheaf (Type v) := { obj := λ U, Π x : (unop U), T x, - map := λ U V i g, λ (x : unop V), g (i.unop x) } + map := λ U V i g, λ (x : unop V), g (i.unop x), + map_id' := λ U, by { ext g ⟨x, hx⟩, refl }, + map_comp' := λ U V W i j, rfl } @[simp] lemma presheaf_to_Types_obj {T : X → Type v} {U : (opens X)ᵒᵖ} : @@ -65,7 +67,9 @@ There is no requirement that the functions are continuous, here. -- written as an equality of functions (rather than being applied to some argument). def presheaf_to_Type (T : Type v) : X.presheaf (Type v) := { obj := λ U, (unop U) → T, - map := λ U V i g, g ∘ i.unop } + map := λ U V i g, g ∘ i.unop, + map_id' := λ U, by { ext g ⟨x, hx⟩, refl }, + map_comp' := λ U V W i j, rfl } @[simp] lemma presheaf_to_Type_obj {T : Type v} {U : (opens X)ᵒᵖ} : From ad6bbaf4bb517f8d038e571a1d00024bd8460c0c Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 10 Oct 2022 05:50:09 +0000 Subject: [PATCH 652/828] chore(topology/continuous_function/algebra): fix an unexplained issue with `continuous_map.coe_fn_alg_hom` (#16876) --- src/topology/continuous_function/algebra.lean | 8 ++------ 1 file changed, 2 insertions(+), 6 deletions(-) diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index 1fdde95a636b9..f4ac6d39dbd4d 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -406,7 +406,7 @@ coe_injective.comm_ring _ coe_zero coe_one coe_add coe_mul coe_neg coe_sub coe_n /-- Coercion to a function as a `ring_hom`. -/ @[simps] def coe_fn_ring_hom {α : Type*} {β : Type*} [topological_space α] [topological_space β] - [ring β] [topological_ring β] : C(α, β) →+* (α → β) := + [semiring β] [topological_semiring β] : C(α, β) →+* (α → β) := { to_fun := coe_fn, ..(coe_fn_monoid_hom : C(α, β) →* _), ..(coe_fn_add_monoid_hom : C(α, β) →+ _) } @@ -597,11 +597,7 @@ variables (R) def continuous_map.coe_fn_alg_hom : C(α, A) →ₐ[R] (α → A) := { to_fun := coe_fn, commutes' := λ r, rfl, - -- `..(continuous_map.coe_fn_ring_hom : C(α, A) →+* _)` times out for some reason - map_zero' := continuous_map.coe_zero, - map_one' := continuous_map.coe_one, - map_add' := continuous_map.coe_add, - map_mul' := continuous_map.coe_mul } + ..(continuous_map.coe_fn_ring_hom : C(α, A) →+* _) } variables {R} From 2af511a73a0fb758dde3c62f53ea05b7d3c4746e Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 10 Oct 2022 05:50:10 +0000 Subject: [PATCH 653/828] chore(order/lattice): remove redundant imports (#16877) Co-authored-by: Scott Morrison --- src/order/bounded_order.lean | 4 ---- src/order/lattice.lean | 1 - 2 files changed, 5 deletions(-) diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index 6bb58fcc28940..a6e39ab08d3da 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -3,11 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import data.option.basic -import logic.nontrivial import order.lattice -import order.max -import tactic.pi_instances /-! # ⊤ and ⊥, bounded lattices and variants diff --git a/src/order/lattice.lean b/src/order/lattice.lean index 7a6ea3f91cb96..63ae03cc0a8a4 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import order.monotone -import order.rel_classes import tactic.simps import tactic.pi_instances From b381419c18566398f9a7f92c75a18256519ce928 Mon Sep 17 00:00:00 2001 From: Andrew Yang <36414270+erdOne@users.noreply.github.com> Date: Mon, 10 Oct 2022 05:50:11 +0000 Subject: [PATCH 654/828] feat(ring_theory/integral_closure): Generalize `is_integral_algebra_map_iff` to noncommutative setting. (#16880) --- src/ring_theory/integral_closure.lean | 23 +++++++++++++++-------- 1 file changed, 15 insertions(+), 8 deletions(-) diff --git a/src/ring_theory/integral_closure.lean b/src/ring_theory/integral_closure.lean index 6b1dfa0366afd..4a83b5a97404b 100644 --- a/src/ring_theory/integral_closure.lean +++ b/src/ring_theory/integral_closure.lean @@ -114,12 +114,24 @@ variables {R A B S : Type*} variables [comm_ring R] [comm_ring A] [comm_ring B] [comm_ring S] variables [algebra R A] [algebra R B] (f : R →+* S) -theorem is_integral_alg_hom (f : A →ₐ[R] B) {x : A} (hx : is_integral R x) : is_integral R (f x) := +theorem is_integral_alg_hom {A B : Type*} [ring A] [ring B] [algebra R A] [algebra R B] + (f : A →ₐ[R] B) {x : A} (hx : is_integral R x) : is_integral R (f x) := let ⟨p, hp, hpx⟩ := hx in ⟨p, hp, by rw [← aeval_def, aeval_alg_hom_apply, aeval_def, hpx, f.map_zero]⟩ +theorem is_integral_alg_hom_iff {A B : Type*} [ring A] [ring B] [algebra R A] [algebra R B] + (f : A →ₐ[R] B) (hf : function.injective f) {x : A} : is_integral R (f x) ↔ is_integral R x := +begin + refine ⟨_, is_integral_alg_hom f⟩, + rintros ⟨p, hp, hx⟩, + use [p, hp], + rwa [← f.comp_algebra_map, ← alg_hom.coe_to_ring_hom, ← polynomial.hom_eval₂, + alg_hom.coe_to_ring_hom, map_eq_zero_iff f hf] at hx +end + @[simp] -theorem is_integral_alg_equiv (f : A ≃ₐ[R] B) {x : A} : is_integral R (f x) ↔ is_integral R x := +theorem is_integral_alg_equiv {A B : Type*} [ring A] [ring B] [algebra R A] [algebra R B] + (f : A ≃ₐ[R] B) {x : A} : is_integral R (f x) ↔ is_integral R x := ⟨λ h, by simpa using is_integral_alg_hom f.symm.to_alg_hom h, is_integral_alg_hom f.to_alg_hom⟩ theorem is_integral_of_is_scalar_tower [algebra A B] [is_scalar_tower R A B] @@ -144,12 +156,7 @@ end lemma is_integral_algebra_map_iff [algebra A B] [is_scalar_tower R A B] {x : A} (hAB : function.injective (algebra_map A B)) : is_integral R (algebra_map A B x) ↔ is_integral R x := -begin - refine ⟨_, λ h, h.algebra_map⟩, - rintros ⟨f, hf, hx⟩, - use [f, hf], - exact (aeval_algebra_map_eq_zero_iff_of_injective hAB).mp hx, -end +is_integral_alg_hom_iff (is_scalar_tower.to_alg_hom R A B) hAB theorem is_integral_iff_is_integral_closure_finite {r : A} : is_integral R r ↔ ∃ s : set R, s.finite ∧ is_integral (subring.closure s) r := From 103252aeb19cd2de4ab2ad5fa0ad3539faf14560 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 10 Oct 2022 08:47:32 +0000 Subject: [PATCH 655/828] chore(data/int/gcd): streamline imports (#16811) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The file on gcds of integers is fundamentally very elementary, but it contained two lemmas about prime numbers, and `data.nat.prime` seems to import everything (modules! half the order library!). Co-authored-by: Andrew Yang Co-authored-by: Yaël Dillies Co-authored-by: Joël Riou Co-authored-by: Junyan Xu --- src/combinatorics/pigeonhole.lean | 1 + src/data/int/gcd.lean | 23 ++-------------------- src/data/int/nat_prime.lean | 19 ++++++++++++++++++ src/data/nat/modeq.lean | 14 ++++++------- src/data/nat/parity.lean | 1 + src/number_theory/pythagorean_triples.lean | 1 + src/ring_theory/int/basic.lean | 1 + 7 files changed, 32 insertions(+), 28 deletions(-) diff --git a/src/combinatorics/pigeonhole.lean b/src/combinatorics/pigeonhole.lean index 08307cf5dddab..9c89ce592f3ff 100644 --- a/src/combinatorics/pigeonhole.lean +++ b/src/combinatorics/pigeonhole.lean @@ -6,6 +6,7 @@ Authors: Kyle Miller, Yury Kudryashov import data.set.finite import data.nat.modeq import algebra.big_operators.order +import algebra.module.basic /-! # Pigeonhole principles diff --git a/src/data/int/gcd.lean b/src/data/int/gcd.lean index fd2612cb1d56a..ef153ac08f22d 100644 --- a/src/data/int/gcd.lean +++ b/src/data/int/gcd.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Guy Leroy. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sangwoo Jo (aka Jason), Guy Leroy, Johannes Hölzl, Mario Carneiro -/ -import data.nat.prime -import data.int.order +import data.nat.gcd.basic +import tactic.norm_num /-! # Extended GCD and divisibility over ℤ @@ -160,31 +160,12 @@ begin ... = nat_abs a / nat_abs b : by rw int.div_mul_cancel H, end -lemma succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul {p : ℕ} (p_prime : nat.prime p) {m n : ℤ} {k l : ℕ} - (hpm : ↑(p ^ k) ∣ m) - (hpn : ↑(p ^ l) ∣ n) (hpmn : ↑(p ^ (k+l+1)) ∣ m*n) : ↑(p ^ (k+1)) ∣ m ∨ ↑(p ^ (l+1)) ∣ n := -have hpm' : p ^ k ∣ m.nat_abs, from int.coe_nat_dvd.1 $ int.dvd_nat_abs.2 hpm, -have hpn' : p ^ l ∣ n.nat_abs, from int.coe_nat_dvd.1 $ int.dvd_nat_abs.2 hpn, -have hpmn' : (p ^ (k+l+1)) ∣ m.nat_abs*n.nat_abs, - by rw ←int.nat_abs_mul; apply (int.coe_nat_dvd.1 $ int.dvd_nat_abs.2 hpmn), -let hsd := nat.succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul p_prime hpm' hpn' hpmn' in -hsd.elim - (λ hsd1, or.inl begin apply int.dvd_nat_abs.1, apply int.coe_nat_dvd.2 hsd1 end) - (λ hsd2, or.inr begin apply int.dvd_nat_abs.1, apply int.coe_nat_dvd.2 hsd2 end) - theorem dvd_of_mul_dvd_mul_left {i j k : ℤ} (k_non_zero : k ≠ 0) (H : k * i ∣ k * j) : i ∣ j := dvd.elim H (λl H1, by rw mul_assoc at H1; exact ⟨_, mul_left_cancel₀ k_non_zero H1⟩) theorem dvd_of_mul_dvd_mul_right {i j k : ℤ} (k_non_zero : k ≠ 0) (H : i * k ∣ j * k) : i ∣ j := by rw [mul_comm i k, mul_comm j k] at H; exact dvd_of_mul_dvd_mul_left k_non_zero H -lemma prime.dvd_nat_abs_of_coe_dvd_sq {p : ℕ} (hp : p.prime) (k : ℤ) (h : ↑p ∣ k ^ 2) : - p ∣ k.nat_abs := -begin - apply @nat.prime.dvd_of_dvd_pow _ _ 2 hp, - rwa [sq, ← nat_abs_mul, ← coe_nat_dvd_left, ← sq] -end - /-- ℤ specific version of least common multiple. -/ def lcm (i j : ℤ) : ℕ := nat.lcm (nat_abs i) (nat_abs j) diff --git a/src/data/int/nat_prime.lean b/src/data/int/nat_prime.lean index ad6fae2fe0bd1..7bbbd5ef5a03a 100644 --- a/src/data/int/nat_prime.lean +++ b/src/data/int/nat_prime.lean @@ -17,4 +17,23 @@ lemma not_prime_of_int_mul {a b : ℤ} {c : ℕ} (ha : 1 < a.nat_abs) (hb : 1 < b.nat_abs) (hc : a*b = (c : ℤ)) : ¬ nat.prime c := not_prime_mul' (nat_abs_mul_nat_abs_eq hc) ha hb +lemma succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul {p : ℕ} (p_prime : nat.prime p) {m n : ℤ} {k l : ℕ} + (hpm : ↑(p ^ k) ∣ m) + (hpn : ↑(p ^ l) ∣ n) (hpmn : ↑(p ^ (k+l+1)) ∣ m*n) : ↑(p ^ (k+1)) ∣ m ∨ ↑(p ^ (l+1)) ∣ n := +have hpm' : p ^ k ∣ m.nat_abs, from int.coe_nat_dvd.1 $ int.dvd_nat_abs.2 hpm, +have hpn' : p ^ l ∣ n.nat_abs, from int.coe_nat_dvd.1 $ int.dvd_nat_abs.2 hpn, +have hpmn' : (p ^ (k+l+1)) ∣ m.nat_abs*n.nat_abs, + by rw ←int.nat_abs_mul; apply (int.coe_nat_dvd.1 $ int.dvd_nat_abs.2 hpmn), +let hsd := nat.succ_dvd_or_succ_dvd_of_succ_sum_dvd_mul p_prime hpm' hpn' hpmn' in +hsd.elim + (λ hsd1, or.inl begin apply int.dvd_nat_abs.1, apply int.coe_nat_dvd.2 hsd1 end) + (λ hsd2, or.inr begin apply int.dvd_nat_abs.1, apply int.coe_nat_dvd.2 hsd2 end) + +lemma prime.dvd_nat_abs_of_coe_dvd_sq {p : ℕ} (hp : p.prime) (k : ℤ) (h : ↑p ∣ k ^ 2) : + p ∣ k.nat_abs := +begin + apply @nat.prime.dvd_of_dvd_pow _ _ 2 hp, + rwa [sq, ← nat_abs_mul, ← coe_nat_dvd_left, ← sq] +end + end int diff --git a/src/data/nat/modeq.lean b/src/data/nat/modeq.lean index 53ee3cded32c9..c4b9d36e129f6 100644 --- a/src/data/nat/modeq.lean +++ b/src/data/nat/modeq.lean @@ -309,13 +309,13 @@ lemma modeq_and_modeq_iff_modeq_mul {a b m n : ℕ} (hmn : coprime m n) : λ h, ⟨h.of_modeq_mul_right _, h.of_modeq_mul_left _⟩⟩ lemma coprime_of_mul_modeq_one (b : ℕ) {a n : ℕ} (h : a * b ≡ 1 [MOD n]) : coprime a n := -nat.coprime_of_dvd' (λ k kp ⟨ka, hka⟩ ⟨kb, hkb⟩, int.coe_nat_dvd.1 begin - rw [hka, hkb, modeq_iff_dvd] at h, - cases h with z hz, - rw [sub_eq_iff_eq_add] at hz, - rw [hz, int.coe_nat_mul, mul_assoc, mul_assoc, int.coe_nat_mul, ← mul_add], - exact dvd_mul_right _ _, -end) +begin + obtain ⟨g, hh⟩ := nat.gcd_dvd_right a n, + rw [nat.coprime_iff_gcd_eq_one, ← nat.dvd_one, ← nat.modeq_zero_iff_dvd], + calc 1 ≡ a * b [MOD a.gcd n] : nat.modeq.of_modeq_mul_right g (hh.subst h).symm + ... ≡ 0 * b [MOD a.gcd n] : (nat.modeq_zero_iff_dvd.mpr (nat.gcd_dvd_left _ _)).mul_right b + ... = 0 : by rw zero_mul, +end @[simp] lemma mod_mul_right_mod (a b c : ℕ) : a % (b * c) % b = a % b := (mod_modeq _ _).of_modeq_mul_right _ diff --git a/src/data/nat/parity.lean b/src/data/nat/parity.lean index ee1e8e91211e2..02800e885ca39 100644 --- a/src/data/nat/parity.lean +++ b/src/data/nat/parity.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Benjamin Davidson -/ import data.nat.modeq +import data.nat.prime import algebra.parity /-! diff --git a/src/number_theory/pythagorean_triples.lean b/src/number_theory/pythagorean_triples.lean index 98342e8198626..77b3ce978967d 100644 --- a/src/number_theory/pythagorean_triples.lean +++ b/src/number_theory/pythagorean_triples.lean @@ -9,6 +9,7 @@ import algebra.group_with_zero.power import tactic.ring import tactic.ring_exp import tactic.field_simp +import data.int.nat_prime import data.zmod.basic /-! diff --git a/src/ring_theory/int/basic.lean b/src/ring_theory/int/basic.lean index ac14779ac532b..a01b11a68459b 100644 --- a/src/ring_theory/int/basic.lean +++ b/src/ring_theory/int/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jens Wagemaker, Aaron Anderson -/ +import data.nat.prime import ring_theory.coprime.basic import ring_theory.principal_ideal_domain From 11156f5a400abf860a1f763e722e8c1f3c8f59b4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 10 Oct 2022 08:47:33 +0000 Subject: [PATCH 656/828] feat(data/list/basic): Alias for `length_le_of_sublist` (#16841) Make an alias `list.sublist.length_le` of `list.length_le_of_sublist` and use it. Rename `list.eq_of_sublist_of_length_eq` and `list.eq_of_sublist_of_length_le` to use dot notation. --- src/data/list/basic.lean | 32 ++++++++++------------ src/data/list/count.lean | 5 ++-- src/data/list/infix.lean | 14 +++++----- src/data/list/perm.lean | 8 ++---- src/data/list/range.lean | 2 +- src/data/list/sublists.lean | 2 +- src/data/multiset/basic.lean | 6 ++-- src/group_theory/free_group.lean | 3 +- src/number_theory/arithmetic_function.lean | 2 +- 9 files changed, 34 insertions(+), 40 deletions(-) diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 45d696485b69b..9e25340644993 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -13,7 +13,7 @@ open function nat (hiding one_pos) namespace list universes u v w x -variables {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {δ : Type x} +variables {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {δ : Type x} {l₁ l₂ : list α} attribute [inline] list.head @@ -277,6 +277,8 @@ lemma length_eq_two {l : list α} : l.length = 2 ↔ ∃ a b, l = [a, b] := lemma length_eq_three {l : list α} : l.length = 3 ↔ ∃ a b c, l = [a, b, c] := ⟨match l with [a, b, c], _ := ⟨a, b, c, rfl⟩ end, λ ⟨a, b, c, e⟩, e.symm ▸ rfl⟩ +alias length_le_of_sublist ← sublist.length_le + /-! ### set-theoretic notation of lists -/ lemma empty_eq : (∅ : list α) = [] := by refl @@ -1090,22 +1092,19 @@ eq_nil_of_subset_nil $ s.subset ⟨eq_nil_of_sublist_nil, λ H, H ▸ sublist.refl _⟩ @[simp] theorem repeat_sublist_repeat (a : α) {m n} : repeat a m <+ repeat a n ↔ m ≤ n := -⟨λ h, by simpa only [length_repeat] using length_le_of_sublist h, +⟨λ h, by simpa only [length_repeat] using h.length_le, λ h, by induction h; [refl, simp only [*, repeat_succ, sublist.cons]] ⟩ -theorem eq_of_sublist_of_length_eq : ∀ {l₁ l₂ : list α}, l₁ <+ l₂ → length l₁ = length l₂ → l₁ = l₂ +theorem sublist.eq_of_length : ∀ {l₁ l₂ : list α}, l₁ <+ l₂ → length l₁ = length l₂ → l₁ = l₂ | ._ ._ sublist.slnil h := rfl -| ._ ._ (sublist.cons l₁ l₂ a s) h := - absurd (length_le_of_sublist s) $ not_le_of_gt $ by rw h; apply lt_succ_self +| ._ ._ (sublist.cons l₁ l₂ a s) h := by cases s.length_le.not_lt (by rw h; apply lt_succ_self) | ._ ._ (sublist.cons2 l₁ l₂ a s) h := - by rw [length, length] at h; injection h with h; rw eq_of_sublist_of_length_eq s h + by rw [length, length] at h; injection h with h; rw s.eq_of_length h -theorem eq_of_sublist_of_length_le {l₁ l₂ : list α} (s : l₁ <+ l₂) (h : length l₂ ≤ length l₁) : - l₁ = l₂ := -eq_of_sublist_of_length_eq s (le_antisymm (length_le_of_sublist s) h) +theorem sublist.eq_of_length_le (s : l₁ <+ l₂) (h : length l₂ ≤ length l₁) : l₁ = l₂ := +s.eq_of_length $ s.length_le.antisymm h -theorem sublist.antisymm {l₁ l₂ : list α} (s₁ : l₁ <+ l₂) (s₂ : l₂ <+ l₁) : l₁ = l₂ := -eq_of_sublist_of_length_le s₁ (length_le_of_sublist s₂) +lemma sublist.antisymm (s₁ : l₁ <+ l₂) (s₂ : l₂ <+ l₁) : l₁ = l₂ := s₁.eq_of_length_le s₂.length_le instance decidable_sublist [decidable_eq α] : ∀ (l₁ l₂ : list α), decidable (l₁ <+ l₂) | [] l₂ := is_true $ nil_sublist _ @@ -3003,7 +3002,7 @@ by simp only [map_filter_map, H, filter_map_some] theorem length_filter_le (p : α → Prop) [decidable_pred p] (l : list α) : (l.filter p).length ≤ l.length := -list.length_le_of_sublist (list.filter_sublist _) +(list.filter_sublist _).length_le theorem length_filter_map_le (f : α → option β) (l : list α) : (list.filter_map f l).length ≤ l.length := @@ -3169,14 +3168,13 @@ begin { exact iff_of_true rfl (forall_mem_nil _) }, rw forall_mem_cons, by_cases p a, { rw [filter_cons_of_pos _ h, cons_inj, ih, and_iff_right h] }, - { rw [filter_cons_of_neg _ h], - refine iff_of_false _ (mt and.left h), intro e, - have := filter_sublist l, rw e at this, - exact not_lt_of_ge (length_le_of_sublist this) (lt_succ_self _) } + { refine iff_of_false (λ hl, h $ of_mem_filter (_ : a ∈ filter p (a :: l))) (mt and.left h), + rw hl, + exact mem_cons_self _ _ } end theorem filter_length_eq_length {l} : (filter p l).length = l.length ↔ ∀ a ∈ l, p a := -iff.trans ⟨eq_of_sublist_of_length_eq l.filter_sublist, congr_arg list.length⟩ filter_eq_self +iff.trans ⟨l.filter_sublist.eq_of_length, congr_arg list.length⟩ filter_eq_self theorem filter_eq_nil {l} : filter p l = [] ↔ ∀ a ∈ l, ¬p a := by simp only [eq_nil_iff_forall_not_mem, mem_filter, not_and] diff --git a/src/data/list/count.lean b/src/data/list/count.lean index acff9e9fbff55..930e01323d41c 100644 --- a/src/data/list/count.lean +++ b/src/data/list/count.lean @@ -45,7 +45,7 @@ by induction l with x l ih; [refl, by_cases (p x)]; simp only [countp_cons_of_neg _ _ h, ih, filter_cons_of_neg _ h]]; refl lemma countp_le_length : countp p l ≤ l.length := -by simpa only [countp_eq_length_filter] using length_le_of_sublist (filter_sublist _) +by simpa only [countp_eq_length_filter] using length_filter_le _ _ @[simp] lemma countp_append (l₁ l₂) : countp p (l₁ ++ l₂) = countp p l₁ + countp p l₂ := by simp only [countp_eq_length_filter, filter_append, length_append] @@ -175,8 +175,7 @@ lemma le_count_iff_repeat_sublist {a : α} {l : list α} {n : ℕ} : lemma repeat_count_eq_of_count_eq_length {a : α} {l : list α} (h : count a l = length l) : repeat a (count a l) = l := -eq_of_sublist_of_length_eq (le_count_iff_repeat_sublist.mp (le_refl (count a l))) - (eq.trans (length_repeat a (count a l)) h) +(le_count_iff_repeat_sublist.mp le_rfl).eq_of_length $ (length_repeat a (count a l)).trans h @[simp] lemma count_filter {p} [decidable_pred p] {a} {l : list α} (h : p a) : count a (filter p l) = count a l := diff --git a/src/data/list/infix.lean b/src/data/list/infix.lean index 135f4112838e0..1d8d58e50fd64 100644 --- a/src/data/list/infix.lean +++ b/src/data/list/infix.lean @@ -113,9 +113,9 @@ alias reverse_prefix ↔ _ is_suffix.reverse alias reverse_suffix ↔ _ is_prefix.reverse alias reverse_infix ↔ _ is_infix.reverse -lemma is_infix.length_le (h : l₁ <:+: l₂) : l₁.length ≤ l₂.length := length_le_of_sublist h.sublist -lemma is_prefix.length_le (h : l₁ <+: l₂) : l₁.length ≤ l₂.length := length_le_of_sublist h.sublist -lemma is_suffix.length_le (h : l₁ <:+ l₂) : l₁.length ≤ l₂.length := length_le_of_sublist h.sublist +lemma is_infix.length_le (h : l₁ <:+: l₂) : l₁.length ≤ l₂.length := h.sublist.length_le +lemma is_prefix.length_le (h : l₁ <+: l₂) : l₁.length ≤ l₂.length := h.sublist.length_le +lemma is_suffix.length_le (h : l₁ <:+ l₂) : l₁.length ≤ l₂.length := h.sublist.length_le lemma eq_nil_of_infix_nil (h : l <:+: []) : l = [] := eq_nil_of_sublist_nil h.sublist @@ -138,13 +138,13 @@ lemma infix_iff_prefix_suffix (l₁ l₂ : list α) : l₁ <:+: l₂ ↔ ∃ t, λ ⟨._, ⟨t, rfl⟩, s, e⟩, ⟨s, t, by rw append_assoc; exact e⟩⟩ lemma eq_of_infix_of_length_eq (h : l₁ <:+: l₂) : l₁.length = l₂.length → l₁ = l₂ := -eq_of_sublist_of_length_eq h.sublist +h.sublist.eq_of_length lemma eq_of_prefix_of_length_eq (h : l₁ <+: l₂) : l₁.length = l₂.length → l₁ = l₂ := -eq_of_sublist_of_length_eq h.sublist +h.sublist.eq_of_length lemma eq_of_suffix_of_length_eq (h : l₁ <:+ l₂) : l₁.length = l₂.length → l₁ = l₂ := -eq_of_sublist_of_length_eq h.sublist +h.sublist.eq_of_length lemma prefix_of_prefix_length_le : ∀ {l₁ l₂ l₃ : list α}, l₁ <+: l₃ → l₂ <+: l₃ → length l₁ ≤ length l₂ → l₁ <+: l₂ @@ -259,7 +259,7 @@ instance decidable_prefix [decidable_eq α] : ∀ (l₁ l₂ : list α), decidab -- Alternatively, use mem_tails instance decidable_suffix [decidable_eq α] : ∀ (l₁ l₂ : list α), decidable (l₁ <:+ l₂) | [] l₂ := is_true ⟨l₂, append_nil _⟩ -| (a :: l₁) [] := is_false $ mt (length_le_of_sublist ∘ is_suffix.sublist) dec_trivial +| (a :: l₁) [] := is_false $ mt (sublist.length_le ∘ is_suffix.sublist) dec_trivial | l₁ (b :: l₂) := decidable_of_decidable_of_iff (@or.decidable _ _ _ (l₁.decidable_suffix l₂)) suffix_cons_iff.symm diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index 73c8662d851c8..b0053623b7fa7 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -352,12 +352,10 @@ theorem perm.subperm {l₁ l₂ : list α} (p : l₁ ~ l₂) : l₁ <+~ l₂ := let ⟨l₁', p₁, s₁⟩ := p₂.subperm_left.2 s in ⟨l₁', p₁, s₁.trans s₂⟩ theorem subperm.length_le {l₁ l₂ : list α} : l₁ <+~ l₂ → length l₁ ≤ length l₂ -| ⟨l, p, s⟩ := p.length_eq ▸ length_le_of_sublist s +| ⟨l, p, s⟩ := p.length_eq ▸ s.length_le theorem subperm.perm_of_length_le {l₁ l₂ : list α} : l₁ <+~ l₂ → length l₂ ≤ length l₁ → l₁ ~ l₂ -| ⟨l, p, s⟩ h := - suffices l = l₂, from this ▸ p.symm, - eq_of_sublist_of_length_le s $ p.symm.length_eq ▸ h +| ⟨l, p, s⟩ h := (s.eq_of_length_le $ p.symm.length_eq ▸ h) ▸ p.symm theorem subperm.antisymm {l₁ l₂ : list α} (h₁ : l₁ <+~ l₂) (h₂ : l₂ <+~ l₁) : l₁ ~ l₂ := h₁.perm_of_length_le h₂.length_le @@ -591,7 +589,7 @@ theorem subperm.exists_of_length_lt {l₁ l₂ : list α} : { cases h }, { cases lt_or_eq_of_le (nat.le_of_lt_succ h : length l₁ ≤ length l₂) with h h, { exact (IH h).imp (λ a s, s.trans (sublist_cons _ _).subperm) }, - { exact ⟨a, eq_of_sublist_of_length_eq s h ▸ subperm.refl _⟩ } }, + { exact ⟨a, s.eq_of_length h ▸ subperm.refl _⟩ } }, { exact (IH $ nat.lt_of_succ_lt_succ h).imp (λ a s, (swap _ _ _).subperm_right.1 $ (subperm_cons _).2 s) } end diff --git a/src/data/list/range.lean b/src/data/list/range.lean index aa0b3fc912bdb..4b44d73b49d2f 100644 --- a/src/data/list/range.lean +++ b/src/data/list/range.lean @@ -77,7 +77,7 @@ theorem nodup_range' (s n : ℕ) : nodup (range' s n) := by rw [add_right_comm, range'_append] theorem range'_sublist_right {s m n : ℕ} : range' s m <+ range' s n ↔ m ≤ n := -⟨λ h, by simpa only [length_range'] using length_le_of_sublist h, +⟨λ h, by simpa only [length_range'] using h.length_le, λ h, by rw [← tsub_add_cancel_of_le h, ← range'_append]; apply sublist_append_left⟩ theorem range'_subset_right {s m n : ℕ} : range' s m ⊆ range' s n ↔ m ≤ n := diff --git a/src/data/list/sublists.lean b/src/data/list/sublists.lean index a23405897b56a..a7dd13f3da481 100644 --- a/src/data/list/sublists.lean +++ b/src/data/list/sublists.lean @@ -284,7 +284,7 @@ end lemma sublists_len_of_length_lt {n} {l : list α} (h : l.length < n) : sublists_len n l = [] := eq_nil_iff_forall_not_mem.mpr $ λ x, mem_sublists_len.not.mpr $ λ ⟨hs, hl⟩, - (h.trans_eq hl.symm).not_le (length_le_of_sublist hs) + (h.trans_eq hl.symm).not_le (sublist.length_le hs) @[simp] lemma sublists_len_length : ∀ (l : list α), sublists_len l.length l = [l] | [] := rfl diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index b304b1e89c5d0..7ed374f032db8 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -517,12 +517,12 @@ theorem card_eq_one {s : multiset α} : card s = 1 ↔ ∃ a, s = {a} := λ ⟨a, e⟩, e.symm ▸ rfl⟩ theorem card_le_of_le {s t : multiset α} (h : s ≤ t) : card s ≤ card t := -le_induction_on h $ λ l₁ l₂, length_le_of_sublist +le_induction_on h $ λ l₁ l₂, sublist.length_le @[mono] theorem card_mono : monotone (@card α) := λ a b, card_le_of_le theorem eq_of_le_of_card_le {s t : multiset α} (h : s ≤ t) : card t ≤ card s → s = t := -le_induction_on h $ λ l₁ l₂ s h₂, congr_arg coe $ eq_of_sublist_of_length_le s h₂ +le_induction_on h $ λ l₁ l₂ s h₂, congr_arg coe $ s.eq_of_length_le h₂ theorem card_lt_of_lt {s t : multiset α} (h : s < t) : card s < card t := lt_of_not_ge $ λ h₂, ne_of_lt h $ eq_of_le_of_card_le (le_of_lt h) h₂ @@ -1413,7 +1413,7 @@ mem_filter.2 ⟨m, h⟩ theorem filter_eq_self {s} : filter p s = s ↔ ∀ a ∈ s, p a := quot.induction_on s $ λ l, iff.trans ⟨λ h, - eq_of_sublist_of_length_eq (filter_sublist _) (@congr_arg _ _ _ _ card h), + (filter_sublist _).eq_of_length (@congr_arg _ _ _ _ card h), congr_arg coe⟩ filter_eq_self theorem filter_eq_nil {s} : filter p s = 0 ↔ ∀ a ∈ s, ¬p a := diff --git a/src/group_theory/free_group.lean b/src/group_theory/free_group.lean index fcca4d3aee03e..9415ffd2406fa 100644 --- a/src/group_theory/free_group.lean +++ b/src/group_theory/free_group.lean @@ -275,8 +275,7 @@ protected theorem sublist : red L₁ L₂ → L₂ <+ L₁ := refl_trans_gen_of_transitive_reflexive (λl, list.sublist.refl l) (λa b c hab hbc, list.sublist.trans hbc hab) (λa b, red.step.sublist) -theorem length_le (h : red L₁ L₂) : L₂.length ≤ L₁.length := -list.length_le_of_sublist h.sublist +theorem length_le (h : red L₁ L₂) : L₂.length ≤ L₁.length := h.sublist.length_le theorem sizeof_of_step : ∀ {L₁ L₂ : list (α × bool)}, step L₁ L₂ → L₂.sizeof < L₁.sizeof | _ _ (@step.bnot _ L1 L2 x b) := diff --git a/src/number_theory/arithmetic_function.lean b/src/number_theory/arithmetic_function.lean index 8cf8cf809b847..fe0b8160f2e7b 100644 --- a/src/number_theory/arithmetic_function.lean +++ b/src/number_theory/arithmetic_function.lean @@ -797,7 +797,7 @@ lemma card_distinct_factors_eq_card_factors_iff_squarefree {n : ℕ} (h0 : n ≠ begin rw [squarefree_iff_nodup_factors h0, card_distinct_factors_apply], split; intro h, - { rw ← list.eq_of_sublist_of_length_eq n.factors.dedup_sublist h, + { rw ←n.factors.dedup_sublist.eq_of_length h, apply list.nodup_dedup }, { rw h.dedup, refl } From 3441411d7e97040f5c459d851d4e49181edc1d0e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 10 Oct 2022 08:47:34 +0000 Subject: [PATCH 657/828] feat(topology/metric_space/basic): `positivity` extension for `diam` (#16842) Add `positivity_diam`, a `positivity` extension for `metric.diam`. --- src/topology/metric_space/basic.lean | 11 +++++++++++ test/positivity.lean | 3 ++- 2 files changed, 13 insertions(+), 1 deletion(-) diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 4843a13956b29..30fd0843e0972 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -2495,6 +2495,17 @@ end diam end metric +namespace tactic +open positivity + +/-- Extension for the `positivity` tactic: the diameter of a set is always nonnegative. -/ +@[positivity] +meta def positivity_diam : expr → tactic strictness +| `(metric.diam %%s) := nonnegative <$> mk_app ``metric.diam_nonneg [s] +| e := pp e >>= fail ∘ format.bracket "The expression " " is not of the form `metric.diam s`" + +end tactic + lemma comap_dist_right_at_top_le_cocompact (x : α) : comap (λ y, dist y x) at_top ≤ cocompact α := begin refine filter.has_basis_cocompact.ge_iff.2 (λ s hs, mem_comap.2 _), diff --git a/test/positivity.lean b/test/positivity.lean index 3a5f4d112fa24..aac26ac9d9ab2 100644 --- a/test/positivity.lean +++ b/test/positivity.lean @@ -215,7 +215,8 @@ example {r : ℝ} : 0 < real.exp r := by positivity example {V : Type*} [normed_add_comm_group V] (x : V) : 0 ≤ ∥x∥ := by positivity -example {X : Type*} [metric_space X] (x y : X) : 0 ≤ dist x y := by positivity +example [metric_space α] (x y : α) : 0 ≤ dist x y := by positivity +example [metric_space α] {s : set α} : 0 ≤ metric.diam s := by positivity example {E : Type*} [add_group E] {p : add_group_seminorm E} {x : E} : 0 ≤ p x := by positivity example {E : Type*} [group E] {p : group_seminorm E} {x : E} : 0 ≤ p x := by positivity From 4f569375c00a55f25469e5d4c663bab3dcb1bdf9 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 10 Oct 2022 08:47:36 +0000 Subject: [PATCH 658/828] feat(topology/basic): Singleton sets are not open (#16843) An immediate consequence of `dense_compl_singleton_iff_not_open` and `dense_compl_singleton`. --- src/topology/basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/topology/basic.lean b/src/topology/basic.lean index c63e3dc256205..f9ea347e286b9 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -1041,6 +1041,9 @@ space. -/ interior {x} = (∅ : set α) := interior_eq_empty_iff_dense_compl.2 (dense_compl_singleton x) +lemma not_is_open_singleton (x : α) [ne_bot (𝓝[≠] x)] : ¬ is_open ({x} : set α) := +dense_compl_singleton_iff_not_open.1 (dense_compl_singleton x) + lemma closure_eq_cluster_pts {s : set α} : closure s = {a | cluster_pt a (𝓟 s)} := set.ext $ λ x, mem_closure_iff_cluster_pt From 874b4ab774afb578fb2102701aba238eda6526a1 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 10 Oct 2022 08:47:37 +0000 Subject: [PATCH 659/828] chore(.github): add CODEOWNERS stub (#16846) This is a first very minimal stub of a CODEOWNERS file. I suggest that we try this out for a few days, and then organically refine it. --- .github/CODEOWNERS | 12 ++++++++++++ 1 file changed, 12 insertions(+) create mode 100644 .github/CODEOWNERS diff --git a/.github/CODEOWNERS b/.github/CODEOWNERS new file mode 100644 index 0000000000000..2bee90cc9e6fc --- /dev/null +++ b/.github/CODEOWNERS @@ -0,0 +1,12 @@ +# By default, the mathlib maintainers own everything in the repo. +# Later matches will take precedence over this match. +* @leanprover-community/mathlib-maintainers + +src/tactic/ @leanprover-community/mathlib-meta + +src/category_theory/ @leanprover-community/mathlib-CT +src/algebraic_topology/ @leanprover-community/mathlib-CT # for now the category theory team can take care of this + +src/algebraic_geometry/ @leanprover-community/mathlib-AG + +src/combinatorics/ @leanprover-community/mathlib-CO From 07a2c0d8ac660586743c8603d01a132e39c99a81 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Mon, 10 Oct 2022 08:47:38 +0000 Subject: [PATCH 660/828] doc(order/filter/zero_and_bounded_at_filter): tinker with docstrings (#16874) I found them a bit hard to understand. --- src/order/filter/zero_and_bounded_at_filter.lean | 16 ++++++++++------ 1 file changed, 10 insertions(+), 6 deletions(-) diff --git a/src/order/filter/zero_and_bounded_at_filter.lean b/src/order/filter/zero_and_bounded_at_filter.lean index 3fec25c7b891b..d8029af39f217 100644 --- a/src/order/filter/zero_and_bounded_at_filter.lean +++ b/src/order/filter/zero_and_bounded_at_filter.lean @@ -25,14 +25,16 @@ variables {α β : Type*} open_locale topological_space -/-- A function `f : α → β` is `zero_at_filter` if in the limit it is zero. -/ +/-- If `l` is a filter on `α`, then a function `f : α → β` is `zero_at_filter l` + if it tends to zero along `l`. -/ def zero_at_filter [has_zero β] [topological_space β] (l : filter α) (f : α → β) : Prop := filter.tendsto f l (𝓝 0) lemma zero_is_zero_at_filter [has_zero β] [topological_space β] (l : filter α) : zero_at_filter l (0 : α → β) := tendsto_const_nhds -/-- The submodule of funtions that are `zero_at_filter`. -/ +/-- `zero_at_filter_submodule l` is the submodule of `f : α → β` which +tend to zero along `l`. -/ def zero_at_filter_submodule [topological_space β] [semiring β] [has_continuous_add β] [has_continuous_mul β] (l : filter α) : submodule β (α → β) := { carrier := zero_at_filter l, @@ -40,14 +42,16 @@ def zero_at_filter_submodule [topological_space β] [semiring β] add_mem' := by { intros a b ha hb, simpa using ha.add hb, }, smul_mem' := by { intros c f hf, simpa using hf.const_mul c }, } -/-- The additive submonoid of funtions that are `zero_at_filter`. -/ +/-- `zero_at_filter_add_submonoid l` is the additive submonoid of `f : α → β` +which tend to zero along `l`. -/ def zero_at_filter_add_submonoid [topological_space β] [add_zero_class β] [has_continuous_add β] (l : filter α) : add_submonoid (α → β) := { carrier := zero_at_filter l, add_mem' := by { intros a b ha hb, simpa using ha.add hb }, zero_mem' := zero_is_zero_at_filter l, } -/-- A function `f: α → β` is `bounded_at_filter` if `f =O[l] 1`. -/ +/-- If `l` is a filter on `α`, then a function `f: α → β` is `bounded_at_filter l` +if `f =O[l] 1`. -/ def bounded_at_filter [has_norm β] [has_one (α → β)] (l : filter α) (f : α → β) : Prop := asymptotics.is_O l f (1 : α → β) @@ -59,14 +63,14 @@ lemma zero_is_bounded_at_filter [normed_field β] (l : filter α) : bounded_at_filter l (0 : α → β) := (zero_at_filter_is_bounded_at_filter l _) (zero_is_zero_at_filter l) -/-- The submodule of funtions that are `bounded_at_filter`. -/ +/-- The submodule of functions that are bounded along a filter `l`. -/ def bounded_filter_submodule [normed_field β] (l : filter α) : submodule β (α → β) := { carrier := bounded_at_filter l, zero_mem' := zero_is_bounded_at_filter l, add_mem' := by { intros f g hf hg, simpa using hf.add hg, }, smul_mem' := by { intros c f hf, simpa using hf.const_mul_left c }, } -/-- The subalgebra of funtions that are `bounded_at_filter`. -/ +/-- The subalgebra of functions that are bounded along a filter `l`. -/ def bounded_filter_subalgebra [normed_field β] (l : filter α) : subalgebra β (α → β) := begin From 09b58946c2bbda1b73165474c0c565d40a13cbec Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 10 Oct 2022 13:36:05 +0000 Subject: [PATCH 661/828] feat(measure_theory/measure/haar_lebesgue): the Lebesgue measure is doubling (#16885) Also split the file `measure_theory/covering/density_theorem` in two, to avoid importing all the differentiation theory into `haar_lebesgue.lean`. --- src/analysis/special_functions/stirling.lean | 4 +- src/data/real/nnreal.lean | 4 + .../covering/density_theorem.lean | 98 +------------- src/measure_theory/measure/doubling.lean | 121 ++++++++++++++++++ src/measure_theory/measure/haar_lebesgue.lean | 14 +- 5 files changed, 141 insertions(+), 100 deletions(-) create mode 100644 src/measure_theory/measure/doubling.lean diff --git a/src/analysis/special_functions/stirling.lean b/src/analysis/special_functions/stirling.lean index 21351f9a4e86a..720d295c46bac 100644 --- a/src/analysis/special_functions/stirling.lean +++ b/src/analysis/special_functions/stirling.lean @@ -60,7 +60,7 @@ We have the expression -/ lemma log_stirling_seq_formula (n : ℕ) : log (stirling_seq n.succ) = log n.succ!- 1 / 2 * log (2 * n.succ) - n.succ * log (n.succ / exp 1) := -by rw [stirling_seq, log_div, log_mul, sqrt_eq_rpow, log_rpow, log_pow, tsub_tsub]; +by rw [stirling_seq, log_div, log_mul, sqrt_eq_rpow, log_rpow, real.log_pow, tsub_tsub]; try { apply ne_of_gt }; positivity -- TODO: Make `positivity` handle `≠ 0` goals /-- @@ -90,7 +90,7 @@ begin end /-- The sequence `log ∘ stirling_seq ∘ succ` is monotone decreasing -/ -lemma log_stirling_seq'_antitone : antitone (log ∘ stirling_seq ∘ succ) := +lemma log_stirling_seq'_antitone : antitone (real.log ∘ stirling_seq ∘ succ) := antitone_nat_of_succ_le $ λ n, sub_nonneg.mp $ (log_stirling_seq_diff_has_sum n).nonneg $ λ m, by positivity diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 5d1e550c56236..420fcd8776851 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -485,6 +485,10 @@ end real.to_nnreal (bit1 r) = bit1 (real.to_nnreal r) := (real.to_nnreal_add (by simp [hr]) zero_le_one).trans (by simp [bit1]) +lemma to_nnreal_pow {x : ℝ} (hx : 0 ≤ x) (n : ℕ) : (x ^ n).to_nnreal = (x.to_nnreal) ^ n := +by rw [← nnreal.coe_eq, nnreal.coe_pow, real.coe_to_nnreal _ (pow_nonneg hx _), + real.coe_to_nnreal x hx] + end to_nnreal end real diff --git a/src/measure_theory/covering/density_theorem.lean b/src/measure_theory/covering/density_theorem.lean index fe98c8be7ad7a..cd4078adfa1b8 100644 --- a/src/measure_theory/covering/density_theorem.lean +++ b/src/measure_theory/covering/density_theorem.lean @@ -3,17 +3,13 @@ Copyright (c) 2022 Oliver Nash. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash -/ +import measure_theory.measure.doubling import measure_theory.covering.vitali import measure_theory.covering.differentiation -import analysis.special_functions.log.base /-! # Doubling measures and Lebesgue's density theorem -A doubling measure `μ` on a metric space is a measure for which there exists a constant `C` such -that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius -`2 * ε` is bounded by `C` times the measure of the concentric ball of radius `ε`. - Lebesgue's density theorem states that given a set `S` in a sigma compact metric space with locally-finite doubling measure `μ` then for almost all points `x` in `S`, for any sequence of closed balls `B₀, B₁, B₂, ...` containing `x`, the limit `μ (S ∩ Bⱼ) / μ (Bⱼ) → 1` as `j → ∞`. @@ -23,12 +19,9 @@ with results about differentiation along a Vitali family to obtain an explicit f density theorem. ## Main results - - * `is_doubling_measure`: the definition of a doubling measure (as a typeclass). - * `is_doubling_measure.doubling_constant`: a function yielding the doubling constant `C` appearing - in the definition of a doubling measure. * `is_doubling_measure.ae_tendsto_measure_inter_div`: a version of Lebesgue's density theorem for sequences of balls converging on a point but whose centres are not required to be fixed. + -/ noncomputable theory @@ -38,97 +31,10 @@ open_locale nnreal topological_space local attribute [instance] emetric.second_countable_of_sigma_compact -/-- A measure `μ` is said to be a doubling measure if there exists a constant `C` such that for -all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius `2 * ε` is -bounded by `C` times the measure of the concentric ball of radius `ε`. - -Note: it is important that this definition makes a demand only for sufficiently small `ε`. For -example we want hyperbolic space to carry the instance `is_doubling_measure volume` but volumes grow -exponentially in hyperbolic space. To be really explicit, consider the hyperbolic plane of -curvature -1, the area of a disc of radius `ε` is `A(ε) = 2π(cosh(ε) - 1)` so `A(2ε)/A(ε) ~ exp(ε)`. --/ -class is_doubling_measure {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) := -(exists_measure_closed_ball_le_mul [] : - ∃ (C : ℝ≥0), ∀ᶠ ε in 𝓝[>] 0, ∀ x, μ (closed_ball x (2 * ε)) ≤ C * μ (closed_ball x ε)) - namespace is_doubling_measure variables {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) [is_doubling_measure μ] -/-- A doubling constant for a doubling measure. - -See also `is_doubling_measure.scaling_constant_of`. -/ -def doubling_constant : ℝ≥0 := classical.some $ exists_measure_closed_ball_le_mul μ - -lemma exists_measure_closed_ball_le_mul' : - ∀ᶠ ε in 𝓝[>] 0, ∀ x, μ (closed_ball x (2 * ε)) ≤ doubling_constant μ * μ (closed_ball x ε) := -classical.some_spec $ exists_measure_closed_ball_le_mul μ - -lemma exists_eventually_forall_measure_closed_ball_le_mul (K : ℝ) : - ∃ (C : ℝ≥0), ∀ᶠ ε in 𝓝[>] 0, ∀ x t (ht : t ≤ K), - μ (closed_ball x (t * ε)) ≤ C * μ (closed_ball x ε) := -begin - let C := doubling_constant μ, - have hμ : ∀ (n : ℕ), ∀ᶠ ε in 𝓝[>] 0, ∀ x, - μ (closed_ball x (2^n * ε)) ≤ ↑(C^n) * μ (closed_ball x ε), - { intros n, - induction n with n ih, { simp, }, - replace ih := eventually_nhds_within_pos_mul_left (two_pos : 0 < (2 : ℝ)) ih, - refine (ih.and (exists_measure_closed_ball_le_mul' μ)).mono (λ ε hε x, _), - calc μ (closed_ball x (2^(n + 1) * ε)) - = μ (closed_ball x (2^n * (2 * ε))) : by rw [pow_succ', mul_assoc] - ... ≤ ↑(C^n) * μ (closed_ball x (2 * ε)) : hε.1 x - ... ≤ ↑(C^n) * (C * μ (closed_ball x ε)) : ennreal.mul_left_mono (hε.2 x) - ... = ↑(C^(n + 1)) * μ (closed_ball x ε) : by rw [← mul_assoc, pow_succ', ennreal.coe_mul], }, - rcases lt_or_le K 1 with hK | hK, - { refine ⟨1, _⟩, - simp only [ennreal.coe_one, one_mul], - exact eventually_mem_nhds_within.mono (λ ε hε x t ht, - measure_mono $ closed_ball_subset_closed_ball (by nlinarith [mem_Ioi.mp hε])), }, - { refine ⟨C^⌈real.logb 2 K⌉₊, ((hμ ⌈real.logb 2 K⌉₊).and eventually_mem_nhds_within).mono - (λ ε hε x t ht, le_trans (measure_mono $ closed_ball_subset_closed_ball _) (hε.1 x))⟩, - refine mul_le_mul_of_nonneg_right (ht.trans _) (mem_Ioi.mp hε.2).le, - conv_lhs { rw ← real.rpow_logb two_pos (by norm_num) (by linarith : 0 < K), }, - rw ← real.rpow_nat_cast, - exact real.rpow_le_rpow_of_exponent_le one_le_two (nat.le_ceil (real.logb 2 K)), }, -end - -/-- A variant of `is_doubling_measure.doubling_constant` which allows for scaling the radius by -values other than `2`. -/ -def scaling_constant_of (K : ℝ) : ℝ≥0 := -max (classical.some $ exists_eventually_forall_measure_closed_ball_le_mul μ K) 1 - -lemma eventually_measure_mul_le_scaling_constant_of_mul (K : ℝ) : - ∃ (R : ℝ), 0 < R ∧ ∀ x t r (ht : t ∈ Ioc 0 K) (hr : r ≤ R), - μ (closed_ball x (t * r)) ≤ scaling_constant_of μ K * μ (closed_ball x r) := -begin - have h := classical.some_spec (exists_eventually_forall_measure_closed_ball_le_mul μ K), - rcases mem_nhds_within_Ioi_iff_exists_Ioc_subset.1 h with ⟨R, Rpos, hR⟩, - refine ⟨R, Rpos, λ x t r ht hr, _⟩, - rcases lt_trichotomy r 0 with rneg|rfl|rpos, - { have : t * r < 0, from mul_neg_of_pos_of_neg ht.1 rneg, - simp only [closed_ball_eq_empty.2 this, measure_empty, zero_le'] }, - { simp only [mul_zero, closed_ball_zero], - refine le_mul_of_one_le_of_le _ le_rfl, - apply ennreal.one_le_coe_iff.2 (le_max_right _ _) }, - { apply (hR ⟨rpos, hr⟩ x t ht.2).trans _, - exact ennreal.mul_le_mul (ennreal.coe_le_coe.2 (le_max_left _ _)) le_rfl } -end - -/-- A scale below which the doubling measure `μ` satisfies good rescaling properties when one -multiplies the radius of balls by at most `K`, as stated -in `measure_mul_le_scaling_constant_of_mul`. -/ -def scaling_scale_of (K : ℝ) : ℝ := -(eventually_measure_mul_le_scaling_constant_of_mul μ K).some - -lemma scaling_scale_of_pos (K : ℝ) : 0 < scaling_scale_of μ K := -(eventually_measure_mul_le_scaling_constant_of_mul μ K).some_spec.1 - -lemma measure_mul_le_scaling_constant_of_mul {K : ℝ} {x : α} {t r : ℝ} - (ht : t ∈ Ioc 0 K) (hr : r ≤ scaling_scale_of μ K) : - μ (closed_ball x (t * r)) ≤ scaling_constant_of μ K * μ (closed_ball x r) := -(eventually_measure_mul_le_scaling_constant_of_mul μ K).some_spec.2 x t r ht hr - section variables [second_countable_topology α] [borel_space α] [is_locally_finite_measure μ] diff --git a/src/measure_theory/measure/doubling.lean b/src/measure_theory/measure/doubling.lean new file mode 100644 index 0000000000000..ca8ba6b8d9483 --- /dev/null +++ b/src/measure_theory/measure/doubling.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2022 Oliver Nash. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Oliver Nash +-/ +import measure_theory.measure.measure_space +import analysis.special_functions.log.base + +/-! +# Doubling measures + +A doubling measure `μ` on a metric space is a measure for which there exists a constant `C` such +that for all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius +`2 * ε` is bounded by `C` times the measure of the concentric ball of radius `ε`. + +This file records basic files on doubling measures. + +## Main definitions + + * `is_doubling_measure`: the definition of a doubling measure (as a typeclass). + * `is_doubling_measure.doubling_constant`: a function yielding the doubling constant `C` appearing + in the definition of a doubling measure. +-/ + +noncomputable theory + +open set filter metric measure_theory topological_space +open_locale nnreal topological_space + +/-- A measure `μ` is said to be a doubling measure if there exists a constant `C` such that for +all sufficiently small radii `ε`, and for any centre, the measure of a ball of radius `2 * ε` is +bounded by `C` times the measure of the concentric ball of radius `ε`. + +Note: it is important that this definition makes a demand only for sufficiently small `ε`. For +example we want hyperbolic space to carry the instance `is_doubling_measure volume` but volumes grow +exponentially in hyperbolic space. To be really explicit, consider the hyperbolic plane of +curvature -1, the area of a disc of radius `ε` is `A(ε) = 2π(cosh(ε) - 1)` so `A(2ε)/A(ε) ~ exp(ε)`. +-/ +class is_doubling_measure {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) := +(exists_measure_closed_ball_le_mul [] : + ∃ (C : ℝ≥0), ∀ᶠ ε in 𝓝[>] 0, ∀ x, μ (closed_ball x (2 * ε)) ≤ C * μ (closed_ball x ε)) + +namespace is_doubling_measure + +variables {α : Type*} [metric_space α] [measurable_space α] (μ : measure α) [is_doubling_measure μ] + +/-- A doubling constant for a doubling measure. + +See also `is_doubling_measure.scaling_constant_of`. -/ +def doubling_constant : ℝ≥0 := classical.some $ exists_measure_closed_ball_le_mul μ + +lemma exists_measure_closed_ball_le_mul' : + ∀ᶠ ε in 𝓝[>] 0, ∀ x, μ (closed_ball x (2 * ε)) ≤ doubling_constant μ * μ (closed_ball x ε) := +classical.some_spec $ exists_measure_closed_ball_le_mul μ + +lemma exists_eventually_forall_measure_closed_ball_le_mul (K : ℝ) : + ∃ (C : ℝ≥0), ∀ᶠ ε in 𝓝[>] 0, ∀ x t (ht : t ≤ K), + μ (closed_ball x (t * ε)) ≤ C * μ (closed_ball x ε) := +begin + let C := doubling_constant μ, + have hμ : ∀ (n : ℕ), ∀ᶠ ε in 𝓝[>] 0, ∀ x, + μ (closed_ball x (2^n * ε)) ≤ ↑(C^n) * μ (closed_ball x ε), + { intros n, + induction n with n ih, { simp, }, + replace ih := eventually_nhds_within_pos_mul_left (two_pos : 0 < (2 : ℝ)) ih, + refine (ih.and (exists_measure_closed_ball_le_mul' μ)).mono (λ ε hε x, _), + calc μ (closed_ball x (2^(n + 1) * ε)) + = μ (closed_ball x (2^n * (2 * ε))) : by rw [pow_succ', mul_assoc] + ... ≤ ↑(C^n) * μ (closed_ball x (2 * ε)) : hε.1 x + ... ≤ ↑(C^n) * (C * μ (closed_ball x ε)) : ennreal.mul_left_mono (hε.2 x) + ... = ↑(C^(n + 1)) * μ (closed_ball x ε) : by rw [← mul_assoc, pow_succ', ennreal.coe_mul], }, + rcases lt_or_le K 1 with hK | hK, + { refine ⟨1, _⟩, + simp only [ennreal.coe_one, one_mul], + exact eventually_mem_nhds_within.mono (λ ε hε x t ht, + measure_mono $ closed_ball_subset_closed_ball (by nlinarith [mem_Ioi.mp hε])), }, + { refine ⟨C^⌈real.logb 2 K⌉₊, ((hμ ⌈real.logb 2 K⌉₊).and eventually_mem_nhds_within).mono + (λ ε hε x t ht, le_trans (measure_mono $ closed_ball_subset_closed_ball _) (hε.1 x))⟩, + refine mul_le_mul_of_nonneg_right (ht.trans _) (mem_Ioi.mp hε.2).le, + conv_lhs { rw ← real.rpow_logb two_pos (by norm_num) (by linarith : 0 < K), }, + rw ← real.rpow_nat_cast, + exact real.rpow_le_rpow_of_exponent_le one_le_two (nat.le_ceil (real.logb 2 K)), }, +end + +/-- A variant of `is_doubling_measure.doubling_constant` which allows for scaling the radius by +values other than `2`. -/ +def scaling_constant_of (K : ℝ) : ℝ≥0 := +max (classical.some $ exists_eventually_forall_measure_closed_ball_le_mul μ K) 1 + +lemma eventually_measure_mul_le_scaling_constant_of_mul (K : ℝ) : + ∃ (R : ℝ), 0 < R ∧ ∀ x t r (ht : t ∈ Ioc 0 K) (hr : r ≤ R), + μ (closed_ball x (t * r)) ≤ scaling_constant_of μ K * μ (closed_ball x r) := +begin + have h := classical.some_spec (exists_eventually_forall_measure_closed_ball_le_mul μ K), + rcases mem_nhds_within_Ioi_iff_exists_Ioc_subset.1 h with ⟨R, Rpos, hR⟩, + refine ⟨R, Rpos, λ x t r ht hr, _⟩, + rcases lt_trichotomy r 0 with rneg|rfl|rpos, + { have : t * r < 0, from mul_neg_of_pos_of_neg ht.1 rneg, + simp only [closed_ball_eq_empty.2 this, measure_empty, zero_le'] }, + { simp only [mul_zero, closed_ball_zero], + refine le_mul_of_one_le_of_le _ le_rfl, + apply ennreal.one_le_coe_iff.2 (le_max_right _ _) }, + { apply (hR ⟨rpos, hr⟩ x t ht.2).trans _, + exact ennreal.mul_le_mul (ennreal.coe_le_coe.2 (le_max_left _ _)) le_rfl } +end + +/-- A scale below which the doubling measure `μ` satisfies good rescaling properties when one +multiplies the radius of balls by at most `K`, as stated +in `measure_mul_le_scaling_constant_of_mul`. -/ +def scaling_scale_of (K : ℝ) : ℝ := +(eventually_measure_mul_le_scaling_constant_of_mul μ K).some + +lemma scaling_scale_of_pos (K : ℝ) : 0 < scaling_scale_of μ K := +(eventually_measure_mul_le_scaling_constant_of_mul μ K).some_spec.1 + +lemma measure_mul_le_scaling_constant_of_mul {K : ℝ} {x : α} {t r : ℝ} + (ht : t ∈ Ioc 0 K) (hr : r ≤ scaling_scale_of μ K) : + μ (closed_ball x (t * r)) ≤ scaling_constant_of μ K * μ (closed_ball x r) := +(eventually_measure_mul_le_scaling_constant_of_mul μ K).some_spec.2 x t r ht hr + +end is_doubling_measure diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index b42d1140e8211..c84c51cb624f7 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -8,6 +8,7 @@ import measure_theory.measure.haar import linear_algebra.finite_dimensional import analysis.normed_space.pointwise import measure_theory.group.pointwise +import measure_theory.measure.doubling /-! # Relationship between the Haar and Lebesgue measures @@ -21,7 +22,7 @@ We deduce basic properties of any Haar measure on a finite dimensional real vect * `add_haar_preimage_linear_map` : when `f` is a linear map with nonzero determinant, the measure of `f ⁻¹' s` is the measure of `s` multiplied by the absolute value of the inverse of the determinant of `f`. -* `add_haar_image_linear_map` : when `f` is a linear map, the measure of `f '' s` is the +* `add_haar_image_linear_map` : when `f` is a linear map, the measure of `f '' s` is the measure of `s` multiplied by the absolute value of the determinant of `f`. * `add_haar_submodule` : a strict submodule has measure `0`. * `add_haar_smul` : the measure of `r • s` is `|r| ^ dim * μ s`. @@ -36,7 +37,7 @@ small `r`, see `eventually_nonempty_inter_smul_of_density_one`. -/ open topological_space set filter metric -open_locale ennreal pointwise topological_space +open_locale ennreal pointwise topological_space nnreal /-- The interval `[0,1]` as a compact set with non-empty interior. -/ def topological_space.positive_compacts.Icc01 : positive_compacts ℝ := @@ -501,6 +502,15 @@ calc { simp only [ennreal.of_real_ne_top, ne.def, not_false_iff] } end +@[priority 100] instance is_doubling_measure_of_is_add_haar_measure : is_doubling_measure μ := +begin + refine ⟨⟨(2 : ℝ≥0) ^ (finrank ℝ E), _⟩⟩, + filter_upwards [self_mem_nhds_within] with r hr x, + rw [add_haar_closed_ball_mul μ x zero_le_two (le_of_lt hr), add_haar_closed_ball_center μ x, + ennreal.of_real, real.to_nnreal_pow zero_le_two], + simp only [real.to_nnreal_bit0, real.to_nnreal_one, le_refl], +end + /-! ### Density points From fbfef5a154a37b7a02288393919adfd667b104c9 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 10 Oct 2022 15:22:05 +0000 Subject: [PATCH 662/828] fix(.github/CODEOWNERS): disable catch-all pattern for now (#16892) --- .github/CODEOWNERS | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.github/CODEOWNERS b/.github/CODEOWNERS index 2bee90cc9e6fc..7a40559f48dc7 100644 --- a/.github/CODEOWNERS +++ b/.github/CODEOWNERS @@ -1,6 +1,7 @@ # By default, the mathlib maintainers own everything in the repo. # Later matches will take precedence over this match. -* @leanprover-community/mathlib-maintainers +# Disabled until we have better coverage by other patterns. Reenable in the future. +# * @leanprover-community/mathlib-maintainers src/tactic/ @leanprover-community/mathlib-meta From c476f200c379dbbc8ff5aaac5390ed69ee582099 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Mon, 10 Oct 2022 16:26:00 +0000 Subject: [PATCH 663/828] feat(topology/algebra/group): discrete subgroup acts properly discontinuously (#16593) A discrete subgroup of a topological group acts properly discontinuously on the left and on the right. Here discrete means finite intersection with any compact subset. Co-authored-by: Alex Kontorovich <58564076+AlexKontorovich@users.noreply.github.com> --- src/data/set/pointwise.lean | 32 +++++++++++++++++++++++ src/topology/algebra/group.lean | 46 +++++++++++++++++++++++++++++++++ 2 files changed, 78 insertions(+) diff --git a/src/data/set/pointwise.lean b/src/data/set/pointwise.lean index d8ea09db90344..cab46b8a64e00 100644 --- a/src/data/set/pointwise.lean +++ b/src/data/set/pointwise.lean @@ -1302,6 +1302,38 @@ lemma subset_set_smul_iff : A ⊆ a • B ↔ a⁻¹ • A ⊆ B := iff.symm $ (image_subset_iff).trans $ iff.symm $ iff_of_eq $ congr_arg _ $ image_equiv_eq_preimage_symm _ $ mul_action.to_perm _ +@[to_additive] +lemma smul_inter_ne_empty_iff {s t : set α} {x : α} : + x • s ∩ t ≠ ∅ ↔ ∃ a b, (a ∈ t ∧ b ∈ s) ∧ a * b⁻¹ = x := +begin + rw ne_empty_iff_nonempty, + split, + { rintros ⟨a, h, ha⟩, + obtain ⟨b, hb, rfl⟩ := mem_smul_set.mp h, + exact ⟨x • b, b, ⟨ha, hb⟩, by simp⟩, }, + { rintros ⟨a, b, ⟨ha, hb⟩, rfl⟩, + exact ⟨a, mem_inter (mem_smul_set.mpr ⟨b, hb, by simp⟩) ha⟩, }, +end + +@[to_additive] +lemma smul_inter_ne_empty_iff' {s t : set α} {x : α} : + x • s ∩ t ≠ ∅ ↔ ∃ a b, (a ∈ t ∧ b ∈ s) ∧ a / b = x := +by simp_rw [smul_inter_ne_empty_iff, div_eq_mul_inv] + +@[to_additive] +lemma op_smul_inter_ne_empty_iff {s t : set α} {x : αᵐᵒᵖ} : + x • s ∩ t ≠ ∅ ↔ ∃ a b, (a ∈ s ∧ b ∈ t) ∧ a⁻¹ * b = mul_opposite.unop x := +begin + rw ne_empty_iff_nonempty, + split, + { rintros ⟨a, h, ha⟩, + obtain ⟨b, hb, rfl⟩ := mem_smul_set.mp h, + exact ⟨b, x • b, ⟨hb, ha⟩, by simp⟩, }, + { rintros ⟨a, b, ⟨ha, hb⟩, H⟩, + have : mul_opposite.op (a⁻¹ * b) = x := congr_arg mul_opposite.op H, + exact ⟨b, mem_inter (mem_smul_set.mpr ⟨a, ha, by simp [← this]⟩) hb⟩, }, +end + end group section group_with_zero diff --git a/src/topology/algebra/group.lean b/src/topology/algebra/group.lean index 2d89b6131302d..449af7ee8a905 100644 --- a/src/topology/algebra/group.lean +++ b/src/topology/algebra/group.lean @@ -1036,6 +1036,52 @@ begin exact topological_group.t1_space (G ⧸ S) ((quotient_map_quotient_mk.is_closed_preimage).mp hS), end +/-- A subgroup `S` of a topological group `G` acts on `G` properly discontinuously on the left, if +it is discrete in the sense that `S ∩ K` is finite for all compact `K`. (See also +`discrete_topology`.) -/ +@[to_additive "A subgroup `S` of an additive topological group `G` acts on `G` properly +discontinuously on the left, if it is discrete in the sense that `S ∩ K` is finite for all compact +`K`. (See also `discrete_topology`."] +lemma subgroup.properly_discontinuous_smul_of_tendsto_cofinite + (S : subgroup G) (hS : tendsto S.subtype cofinite (cocompact G)) : + properly_discontinuous_smul S G := +{ finite_disjoint_inter_image := begin + intros K L hK hL, + have H : set.finite _ := hS ((hL.prod hK).image continuous_div').compl_mem_cocompact, + rw [preimage_compl, compl_compl] at H, + convert H, + ext x, + simpa only [image_smul, mem_image, prod.exists] using set.smul_inter_ne_empty_iff', + end } + +local attribute [semireducible] mul_opposite + +/-- A subgroup `S` of a topological group `G` acts on `G` properly discontinuously on the right, if +it is discrete in the sense that `S ∩ K` is finite for all compact `K`. (See also +`discrete_topology`.) + +If `G` is Hausdorff, this can be combined with `t2_space_of_properly_discontinuous_smul_of_t2_space` +to show that the quotient group `G ⧸ S` is Hausdorff. -/ +@[to_additive "A subgroup `S` of an additive topological group `G` acts on `G` properly +discontinuously on the right, if it is discrete in the sense that `S ∩ K` is finite for all compact +`K`. (See also `discrete_topology`.) + +If `G` is Hausdorff, this can be combined with `t2_space_of_properly_discontinuous_vadd_of_t2_space` +to show that the quotient group `G ⧸ S` is Hausdorff."] +lemma subgroup.properly_discontinuous_smul_opposite_of_tendsto_cofinite + (S : subgroup G) (hS : tendsto S.subtype cofinite (cocompact G)) : + properly_discontinuous_smul S.opposite G := +{ finite_disjoint_inter_image := begin + intros K L hK hL, + have : continuous (λ p : G × G, (p.1⁻¹, p.2)) := continuous_inv.prod_map continuous_id, + have H : set.finite _ := + hS ((hK.prod hL).image (continuous_mul.comp this)).compl_mem_cocompact, + rw [preimage_compl, compl_compl] at H, + convert H, + ext x, + simpa only [image_smul, mem_image, prod.exists] using set.op_smul_inter_ne_empty_iff, + end } + end section From d20f598ea7bbf11ae0f59ef4b95adb413ef7698d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Mon, 10 Oct 2022 16:26:02 +0000 Subject: [PATCH 664/828] chore(ring_theory.dedekind_domain.adic_valuation): delete outdated comment (#16886) --- src/ring_theory/dedekind_domain/adic_valuation.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/src/ring_theory/dedekind_domain/adic_valuation.lean b/src/ring_theory/dedekind_domain/adic_valuation.lean index 68fb94cfe8eb2..5098e9e5faf79 100644 --- a/src/ring_theory/dedekind_domain/adic_valuation.lean +++ b/src/ring_theory/dedekind_domain/adic_valuation.lean @@ -35,8 +35,6 @@ We define the completion of `K` with respect to the `v`-adic valuation, denoted ideal `(r)`. - `is_dedekind_domain.height_one_spectrum.int_valuation_exists_uniformizer` : There exists `π ∈ R` with `v`-adic valuation `multiplicative.of_add (-1)`. -- `is_dedekind_domain.height_one_spectrum.int_valuation_div_eq_div` : The valuation of `k ∈ K` is - independent on how we express `k` as a fraction. - `is_dedekind_domain.height_one_spectrum.valuation_of_mk'` : The `v`-adic valuation of `r/s ∈ K` is the valuation of `r` divided by the valuation of `s`. - `is_dedekind_domain.height_one_spectrum.valuation_of_algebra_map` : The `v`-adic valuation on `K` From 71a08c308e7b7b0e810768e25c34644614871123 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Mon, 10 Oct 2022 20:56:51 +0000 Subject: [PATCH 665/828] refactor(analysis/normed_space/spectrum): remove `norm_one_class` hypotheses (#16891) When this file and its main results were created, mathlib had a different definition of `normed_algebra` which automatically entailed `norm_one_class`. When `normed_algebra` was refactored, this hypothesis was added everywhere it was needed, including this file. However, most of the main results here are independent of that assumption, but the proof requires slightly more work. Here, we do that work so that we can remove these hypotheses here and a few places downstream. For a few results, we provide both the `norm_one_class` version, and the version without it for convenience (especially because any nontrivial cstar_ring automatically satisfies `norm_one_class`). --- src/analysis/normed_space/algebra.lean | 20 ++-- src/analysis/normed_space/spectrum.lean | 104 ++++++++++++------ .../normed_space/star/gelfand_duality.lean | 13 +-- src/analysis/normed_space/star/spectrum.lean | 23 ++-- src/analysis/special_functions/pow.lean | 22 ++++ .../algebra/module/character_space.lean | 3 + 6 files changed, 121 insertions(+), 64 deletions(-) diff --git a/src/analysis/normed_space/algebra.lean b/src/analysis/normed_space/algebra.lean index c76180d6d6f24..c935af4a5bcf7 100644 --- a/src/analysis/normed_space/algebra.lean +++ b/src/analysis/normed_space/algebra.lean @@ -33,25 +33,21 @@ namespace weak_dual namespace character_space variables [nontrivially_normed_field 𝕜] [normed_ring A] - [normed_algebra 𝕜 A] [complete_space A] [norm_one_class A] + [normed_algebra 𝕜 A] [complete_space A] -lemma norm_one (φ : character_space 𝕜 A) : ∥to_normed_dual (φ : weak_dual 𝕜 A)∥ = 1 := -begin - refine continuous_linear_map.op_norm_eq_of_bounds zero_le_one (λ a, _) (λ x hx h, _), - { rw [one_mul], - exact spectrum.norm_le_norm_of_mem (apply_mem_spectrum φ a) }, - { have : ∥φ 1∥ ≤ x * ∥(1 : A)∥ := h 1, - simpa only [norm_one, mul_one, map_one] using this }, -end +lemma norm_le_norm_one (φ : character_space 𝕜 A) : + ∥to_normed_dual (φ : weak_dual 𝕜 A)∥ ≤ ∥(1 : A)∥ := +continuous_linear_map.op_norm_le_bound _ (norm_nonneg (1 : A)) $ + λ a, mul_comm (∥a∥) (∥(1 : A)∥) ▸ spectrum.norm_le_norm_mul_of_mem (apply_mem_spectrum φ a) instance [proper_space 𝕜] : compact_space (character_space 𝕜 A) := begin rw [←is_compact_iff_compact_space], - have h : character_space 𝕜 A ⊆ to_normed_dual ⁻¹' metric.closed_ball 0 1, + have h : character_space 𝕜 A ⊆ to_normed_dual ⁻¹' metric.closed_ball 0 (∥(1 : A)∥), { intros φ hφ, rw [set.mem_preimage, mem_closed_ball_zero_iff], - exact (le_of_eq $ norm_one ⟨φ, ⟨hφ.1, hφ.2⟩⟩ : _), }, - exact compact_of_is_closed_subset (is_compact_closed_ball 𝕜 0 1) character_space.is_closed h, + exact (norm_le_norm_one ⟨φ, ⟨hφ.1, hφ.2⟩⟩ : _), }, + exact compact_of_is_closed_subset (is_compact_closed_ball 𝕜 0 _) character_space.is_closed h, end end character_space diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 995ef43a4dc3b..367d2fcbb12c6 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -58,12 +58,21 @@ namespace spectrum section spectrum_compact +open filter + variables [normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] local notation `σ` := spectrum 𝕜 local notation `ρ` := resolvent_set 𝕜 local notation `↑ₐ` := algebra_map 𝕜 A +@[simp] lemma spectral_radius.of_subsingleton [subsingleton A] (a : A) : + spectral_radius 𝕜 a = 0 := +by simp [spectral_radius] + +@[simp] lemma spectral_radius_zero : spectral_radius 𝕜 (0 : A) = 0 := +by { nontriviality A, simp [spectral_radius] } + lemma mem_resolvent_set_of_spectral_radius_lt {a : A} {k : 𝕜} (h : spectral_radius 𝕜 a < ∥k∥₊) : k ∈ ρ a := not_not.mp $ λ hn, h.not_le $ le_supr₂ k hn @@ -73,32 +82,48 @@ variable [complete_space A] lemma is_open_resolvent_set (a : A) : is_open (ρ a) := units.is_open.preimage ((continuous_algebra_map 𝕜 A).sub continuous_const) -lemma is_closed (a : A) : is_closed (σ a) := +protected lemma is_closed (a : A) : is_closed (σ a) := (is_open_resolvent_set a).is_closed_compl -lemma mem_resolvent_of_norm_lt [norm_one_class A] {a : A} {k : 𝕜} (h : ∥a∥ < ∥k∥) : +lemma mem_resolvent_set_of_norm_lt_mul {a : A} {k : 𝕜} (h : ∥a∥ * ∥(1 : A)∥ < ∥k∥) : k ∈ ρ a := begin rw [resolvent_set, set.mem_set_of_eq, algebra.algebra_map_eq_smul_one], - have hk : k ≠ 0 := ne_zero_of_norm_ne_zero (by linarith [norm_nonneg a]), + nontriviality A, + have hk : k ≠ 0, + from ne_zero_of_norm_ne_zero ((mul_nonneg (norm_nonneg _) (norm_nonneg _)).trans_lt h).ne', let ku := units.map (↑ₐ).to_monoid_hom (units.mk0 k hk), - have hku : ∥-a∥ < ∥(↑ku⁻¹:A)∥⁻¹ := by simpa [ku, algebra_map_isometry] using h, + rw [←inv_inv (∥(1 : A)∥), mul_inv_lt_iff (inv_pos.2 $ norm_pos_iff.2 (one_ne_zero : (1 : A) ≠ 0))] + at h, + have hku : ∥-a∥ < ∥(↑ku⁻¹:A)∥⁻¹ := by simpa [ku, norm_algebra_map] using h, simpa [ku, sub_eq_add_neg, algebra.algebra_map_eq_smul_one] using (ku.add (-a) hku).is_unit, end +lemma mem_resolvent_set_of_norm_lt [norm_one_class A] {a : A} {k : 𝕜} (h : ∥a∥ < ∥k∥) : + k ∈ ρ a := +mem_resolvent_set_of_norm_lt_mul (by rwa [norm_one, mul_one]) + +lemma norm_le_norm_mul_of_mem {a : A} {k : 𝕜} (hk : k ∈ σ a) : + ∥k∥ ≤ ∥a∥ * ∥(1 : A)∥ := +le_of_not_lt $ mt mem_resolvent_set_of_norm_lt_mul hk + lemma norm_le_norm_of_mem [norm_one_class A] {a : A} {k : 𝕜} (hk : k ∈ σ a) : ∥k∥ ≤ ∥a∥ := -le_of_not_lt $ mt mem_resolvent_of_norm_lt hk +le_of_not_lt $ mt mem_resolvent_set_of_norm_lt hk + +lemma subset_closed_ball_norm_mul (a : A) : + σ a ⊆ metric.closed_ball (0 : 𝕜) (∥a∥ * ∥(1 : A)∥) := +λ k hk, by simp [norm_le_norm_mul_of_mem hk] lemma subset_closed_ball_norm [norm_one_class A] (a : A) : σ a ⊆ metric.closed_ball (0 : 𝕜) (∥a∥) := λ k hk, by simp [norm_le_norm_of_mem hk] -lemma is_bounded [norm_one_class A] (a : A) : metric.bounded (σ a) := -(metric.bounded_iff_subset_ball 0).mpr ⟨∥a∥, subset_closed_ball_norm a⟩ +lemma is_bounded (a : A) : metric.bounded (σ a) := +(metric.bounded_iff_subset_ball 0).mpr ⟨∥a∥ * ∥(1 : A)∥, subset_closed_ball_norm_mul a⟩ -theorem is_compact [norm_one_class A] [proper_space 𝕜] (a : A) : is_compact (σ a) := -metric.is_compact_of_is_closed_bounded (is_closed a) (is_bounded a) +protected theorem is_compact [proper_space 𝕜] (a : A) : is_compact (σ a) := +metric.is_compact_of_is_closed_bounded (spectrum.is_closed a) (is_bounded a) theorem spectral_radius_le_nnnorm [norm_one_class A] (a : A) : spectral_radius 𝕜 a ≤ ∥a∥₊ := @@ -107,8 +132,8 @@ by { refine supr₂_le (λ k hk, _), exact_mod_cast norm_le_norm_of_mem hk } open ennreal polynomial variable (𝕜) -theorem spectral_radius_le_pow_nnnorm_pow_one_div [norm_one_class A] (a : A) (n : ℕ) : - spectral_radius 𝕜 a ≤ ∥a ^ (n + 1)∥₊ ^ (1 / (n + 1) : ℝ) := +theorem spectral_radius_le_pow_nnnorm_pow_one_div (a : A) (n : ℕ) : + spectral_radius 𝕜 a ≤ (∥a ^ (n + 1)∥₊) ^ (1 / (n + 1) : ℝ) * (∥(1 : A)∥₊) ^ (1 / (n + 1) : ℝ) := begin refine supr₂_le (λ k hk, _), /- apply easy direction of the spectral mapping theorem for polynomials -/ @@ -116,13 +141,35 @@ begin by simpa only [one_mul, algebra.algebra_map_eq_smul_one, one_smul, aeval_monomial, one_mul, eval_monomial] using subset_polynomial_aeval a (monomial (n + 1) (1 : 𝕜)) ⟨k, hk, rfl⟩, /- power of the norm is bounded by norm of the power -/ - have nnnorm_pow_le : (↑(∥k∥₊ ^ (n + 1)) : ℝ≥0∞) ≤ ↑∥a ^ (n + 1)∥₊, - by simpa only [norm_to_nnreal, nnnorm_pow k (n+1)] - using coe_mono (real.to_nnreal_mono (norm_le_norm_of_mem pow_mem)), + have nnnorm_pow_le : (↑(∥k∥₊ ^ (n + 1)) : ℝ≥0∞) ≤ ∥a ^ (n + 1)∥₊ * ∥(1 : A)∥₊, + { simpa only [real.to_nnreal_mul (norm_nonneg _), norm_to_nnreal, nnnorm_pow k (n + 1), + ennreal.coe_mul] using coe_mono (real.to_nnreal_mono (norm_le_norm_mul_of_mem pow_mem)) }, /- take (n + 1)ᵗʰ roots and clean up the left-hand side -/ have hn : 0 < ((n + 1 : ℕ) : ℝ), by exact_mod_cast nat.succ_pos', convert monotone_rpow_of_nonneg (one_div_pos.mpr hn).le nnnorm_pow_le, - erw [coe_pow, ←rpow_nat_cast, ←rpow_mul, mul_one_div_cancel hn.ne', rpow_one], rw nat.cast_succ, + erw [coe_pow, ←rpow_nat_cast, ←rpow_mul, mul_one_div_cancel hn.ne', rpow_one], + rw [nat.cast_succ, ennreal.coe_mul_rpow], +end + +theorem spectral_radius_le_liminf_pow_nnnorm_pow_one_div (a : A) : + spectral_radius 𝕜 a ≤ at_top.liminf (λ n : ℕ, (∥a ^ n∥₊ : ℝ≥0∞) ^ (1 / n : ℝ)) := +begin + refine ennreal.le_of_forall_lt_one_mul_le (λ ε hε, _), + by_cases ε = 0, + { simp only [h, zero_mul, zero_le'] }, + have hε' : ε⁻¹ ≠ ∞, + from λ h', h (by simpa only [inv_inv, inv_top] using congr_arg (λ (x : ℝ≥0∞), x⁻¹) h'), + simp only [ennreal.mul_le_iff_le_inv h (hε.trans_le le_top).ne, mul_comm ε⁻¹, + liminf_eq_supr_infi_of_nat', ennreal.supr_mul, ennreal.infi_mul hε'], + rw [←ennreal.inv_lt_inv, inv_one] at hε, + obtain ⟨N, hN⟩ := eventually_at_top.mp + (ennreal.eventually_pow_one_div_le (ennreal.coe_ne_top : ↑∥(1 : A)∥₊ ≠ ∞) hε), + refine (le_trans _ (le_supr _ (N + 1))), + refine le_infi (λ n, _), + simp only [←add_assoc], + refine (spectral_radius_le_pow_nnnorm_pow_one_div 𝕜 a (n + N)).trans _, + norm_cast, + exact mul_le_mul_left' (hN (n + N + 1) (by linarith)) _, end end spectrum_compact @@ -283,22 +330,16 @@ end /-- **Gelfand's formula**: Given an element `a : A` of a complex Banach algebra, the `spectral_radius` of `a` is the limit of the sequence `∥a ^ n∥₊ ^ (1 / n)` -/ -theorem pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius [norm_one_class A] (a : A) : +theorem pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius (a : A) : tendsto (λ n : ℕ, ((∥a ^ n∥₊ ^ (1 / n : ℝ)) : ℝ≥0∞)) at_top (𝓝 (spectral_radius ℂ a)) := -begin - refine tendsto_of_le_liminf_of_limsup_le _ _ (by apply_auto_param) (by apply_auto_param), - { rw [←liminf_nat_add _ 1, liminf_eq_supr_infi_of_nat], - refine le_trans _ (le_supr _ 0), - simp only [nat.cast_succ], - exact le_infi₂ (λ i hi, spectral_radius_le_pow_nnnorm_pow_one_div ℂ a i) }, - { exact limsup_pow_nnnorm_pow_one_div_le_spectral_radius a }, -end +tendsto_of_le_liminf_of_limsup_le (spectral_radius_le_liminf_pow_nnnorm_pow_one_div ℂ a) + (limsup_pow_nnnorm_pow_one_div_le_spectral_radius a) /- This is the same as `pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius` but for `norm` instead of `nnnorm`. -/ /-- **Gelfand's formula**: Given an element `a : A` of a complex Banach algebra, the `spectral_radius` of `a` is the limit of the sequence `∥a ^ n∥₊ ^ (1 / n)` -/ -theorem pow_norm_pow_one_div_tendsto_nhds_spectral_radius [norm_one_class A] (a : A) : +theorem pow_norm_pow_one_div_tendsto_nhds_spectral_radius (a : A) : tendsto (λ n : ℕ, ennreal.of_real (∥a ^ n∥ ^ (1 / n : ℝ))) at_top (𝓝 (spectral_radius ℂ a)) := begin convert pow_nnnorm_pow_one_div_tendsto_nhds_spectral_radius a, @@ -423,20 +464,19 @@ section normed_field variables [normed_field 𝕜] [normed_ring A] [normed_algebra 𝕜 A] [complete_space A] local notation `↑ₐ` := algebra_map 𝕜 A - /-- An algebra homomorphism into the base field, as a continuous linear map (since it is automatically bounded). -/ -instance [norm_one_class A] : continuous_linear_map_class (A →ₐ[𝕜] 𝕜) 𝕜 A 𝕜 := -{ map_continuous := λ φ, add_monoid_hom_class.continuous_of_bound φ 1 $ - λ a, (one_mul ∥a∥).symm ▸ spectrum.norm_le_norm_of_mem (apply_mem_spectrum φ _), +instance : continuous_linear_map_class (A →ₐ[𝕜] 𝕜) 𝕜 A 𝕜 := +{ map_continuous := λ φ, add_monoid_hom_class.continuous_of_bound φ ∥(1 : A)∥ $ + λ a, (mul_comm ∥a∥ ∥(1 : A)∥) ▸ spectrum.norm_le_norm_mul_of_mem (apply_mem_spectrum φ _), .. alg_hom_class.linear_map_class } /-- An algebra homomorphism into the base field, as a continuous linear map (since it is automatically bounded). -/ -def to_continuous_linear_map [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : A →L[𝕜] 𝕜 := +def to_continuous_linear_map (φ : A →ₐ[𝕜] 𝕜) : A →L[𝕜] 𝕜 := { cont := map_continuous φ, .. φ.to_linear_map } -@[simp] lemma coe_to_continuous_linear_map [norm_one_class A] (φ : A →ₐ[𝕜] 𝕜) : +@[simp] lemma coe_to_continuous_linear_map (φ : A →ₐ[𝕜] 𝕜) : ⇑φ.to_continuous_linear_map = φ := rfl end normed_field @@ -459,7 +499,7 @@ namespace weak_dual namespace character_space -variables [normed_field 𝕜] [normed_ring A] [complete_space A] [norm_one_class A] +variables [nontrivially_normed_field 𝕜] [normed_ring A] [complete_space A] variables [normed_algebra 𝕜 A] /-- The equivalence between characters and algebra homomorphisms into the base field. -/ diff --git a/src/analysis/normed_space/star/gelfand_duality.lean b/src/analysis/normed_space/star/gelfand_duality.lean index 9c79754e4b49d..501f3d88a7b3a 100644 --- a/src/analysis/normed_space/star/gelfand_duality.lean +++ b/src/analysis/normed_space/star/gelfand_duality.lean @@ -59,7 +59,7 @@ section complex_banach_algebra open ideal variables {A : Type*} [normed_comm_ring A] [normed_algebra ℂ A] [complete_space A] - [norm_one_class A] (I : ideal A) [ideal.is_maximal I] + (I : ideal A) [ideal.is_maximal I] /-- Every maximal ideal in a commutative complex Banach algebra gives rise to a character on that algebra. In particular, the character, which may be identified as an algebra homomorphism due to @@ -101,23 +101,20 @@ begin exact (continuous_map.spectrum_eq_range (gelfand_transform ℂ A a)).symm ▸ ⟨f, hf.symm⟩, end -instance : nonempty (character_space ℂ A) := -begin - haveI := norm_one_class.nontrivial A, - exact ⟨classical.some $ - weak_dual.character_space.exists_apply_eq_zero (zero_mem_nonunits.mpr zero_ne_one)⟩, -end +instance [nontrivial A] : nonempty (character_space ℂ A) := +⟨classical.some $ weak_dual.character_space.exists_apply_eq_zero $ zero_mem_nonunits.2 zero_ne_one⟩ end complex_banach_algebra section complex_cstar_algebra variables (A : Type*) [normed_comm_ring A] [normed_algebra ℂ A] [complete_space A] -variables [star_ring A] [cstar_ring A] [star_module ℂ A] [nontrivial A] +variables [star_ring A] [cstar_ring A] [star_module ℂ A] /-- The Gelfand transform is an isometry when the algebra is a C⋆-algebra over `ℂ`. -/ lemma gelfand_transform_isometry : isometry (gelfand_transform ℂ A) := begin + nontriviality A, refine add_monoid_hom_class.isometry_of_norm (gelfand_transform ℂ A) (λ a, _), have gt_map_star : gelfand_transform ℂ A (star a) = star (gelfand_transform ℂ A a), from continuous_map.ext (λ φ, map_star φ a), diff --git a/src/analysis/normed_space/star/spectrum.lean b/src/analysis/normed_space/star/spectrum.lean index 863448cb51d90..0073d0d87ff0c 100644 --- a/src/analysis/normed_space/star/spectrum.lean +++ b/src/analysis/normed_space/star/spectrum.lean @@ -25,11 +25,12 @@ section unitary_spectrum variables {𝕜 : Type*} [normed_field 𝕜] {E : Type*} [normed_ring E] [star_ring E] [cstar_ring E] -[normed_algebra 𝕜 E] [complete_space E] [nontrivial E] +[normed_algebra 𝕜 E] [complete_space E] lemma unitary.spectrum_subset_circle (u : unitary E) : spectrum 𝕜 (u : E) ⊆ metric.sphere 0 1 := begin + nontriviality E, refine λ k hk, mem_sphere_zero_iff_norm.mpr (le_antisymm _ _), { simpa only [cstar_ring.norm_coe_unitary u] using norm_le_norm_of_mem hk }, { rw ←unitary.coe_to_units_apply u at hk, @@ -54,7 +55,7 @@ variables {A : Type*} local notation `↑ₐ` := algebra_map ℂ A -lemma is_self_adjoint.spectral_radius_eq_nnnorm [norm_one_class A] {a : A} +lemma is_self_adjoint.spectral_radius_eq_nnnorm {a : A} (ha : is_self_adjoint a) : spectral_radius ℂ a = ∥a∥₊ := begin @@ -68,7 +69,7 @@ begin simp, end -lemma is_star_normal.spectral_radius_eq_nnnorm [norm_one_class A] (a : A) [is_star_normal a] : +lemma is_star_normal.spectral_radius_eq_nnnorm (a : A) [is_star_normal a] : spectral_radius ℂ a = ∥a∥₊ := begin refine (ennreal.pow_strict_mono two_ne_zero).injective _, @@ -86,7 +87,7 @@ begin end /-- Any element of the spectrum of a selfadjoint is real. -/ -theorem is_self_adjoint.mem_spectrum_eq_re [star_module ℂ A] [nontrivial A] {a : A} +theorem is_self_adjoint.mem_spectrum_eq_re [star_module ℂ A] {a : A} (ha : is_self_adjoint a) {z : ℂ} (hz : z ∈ spectrum ℂ a) : z = z.re := begin let Iu := units.mk0 I I_ne_zero, @@ -100,19 +101,19 @@ begin end /-- Any element of the spectrum of a selfadjoint is real. -/ -theorem self_adjoint.mem_spectrum_eq_re [star_module ℂ A] [nontrivial A] +theorem self_adjoint.mem_spectrum_eq_re [star_module ℂ A] (a : self_adjoint A) {z : ℂ} (hz : z ∈ spectrum ℂ (a : A)) : z = z.re := a.prop.mem_spectrum_eq_re hz /-- The spectrum of a selfadjoint is real -/ -theorem is_self_adjoint.coe_re_map_spectrum [star_module ℂ A] [nontrivial A] {a : A} +theorem is_self_adjoint.coe_re_map_spectrum [star_module ℂ A] {a : A} (ha : is_self_adjoint a) : spectrum ℂ a = (coe ∘ re '' (spectrum ℂ a) : set ℂ) := le_antisymm (λ z hz, ⟨z, hz, (ha.mem_spectrum_eq_re hz).symm⟩) (λ z, by { rintros ⟨z, hz, rfl⟩, simpa only [(ha.mem_spectrum_eq_re hz).symm, function.comp_app] using hz }) /-- The spectrum of a selfadjoint is real -/ -theorem self_adjoint.coe_re_map_spectrum [star_module ℂ A] [nontrivial A] (a : self_adjoint A) : +theorem self_adjoint.coe_re_map_spectrum [star_module ℂ A] (a : self_adjoint A) : spectrum ℂ (a : A) = (coe ∘ re '' (spectrum ℂ (a : A)) : set ℂ) := a.property.coe_re_map_spectrum @@ -121,10 +122,8 @@ end complex_scalars namespace star_alg_hom variables {F A B : Type*} -[normed_ring A] [normed_algebra ℂ A] [norm_one_class A] -[complete_space A] [star_ring A] [cstar_ring A] -[normed_ring B] [normed_algebra ℂ B] [norm_one_class B] -[complete_space B] [star_ring B] [cstar_ring B] +[normed_ring A] [normed_algebra ℂ A] [complete_space A] [star_ring A] [cstar_ring A] +[normed_ring B] [normed_algebra ℂ B] [complete_space B] [star_ring B] [cstar_ring B] [hF : star_alg_hom_class F ℂ A B] (φ : F) include hF @@ -161,7 +160,7 @@ namespace weak_dual open continuous_map complex open_locale complex_star_module -variables {F A : Type*} [normed_ring A] [normed_algebra ℂ A] [nontrivial A] [complete_space A] +variables {F A : Type*} [normed_ring A] [normed_algebra ℂ A] [complete_space A] [star_ring A] [cstar_ring A] [star_module ℂ A] [hF : alg_hom_class F ℂ A ℂ] include hF diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 300519f9ebd3b..1b248e80db550 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -1442,6 +1442,16 @@ begin rw [←nnreal.coe_rpow, real.to_nnreal_coe], end +lemma eventually_pow_one_div_le (x : ℝ≥0) {y : ℝ≥0} (hy : 1 < y) : + ∀ᶠ (n : ℕ) in at_top, x ^ (1 / n : ℝ) ≤ y := +begin + obtain ⟨m, hm⟩ := add_one_pow_unbounded_of_pos x (tsub_pos_of_lt hy), + rw [tsub_add_cancel_of_le hy.le] at hm, + refine eventually_at_top.2 ⟨m + 1, λ n hn, _⟩, + simpa only [nnreal.rpow_one_div_le_iff (nat.cast_pos.2 $ m.succ_pos.trans_le hn), + nnreal.rpow_nat_cast] using hm.le.trans (pow_le_pow hy.le (m.le_succ.trans hn)), +end + end nnreal namespace real @@ -2003,6 +2013,18 @@ begin exact_mod_cast hc a (by exact_mod_cast ha), end +lemma eventually_pow_one_div_le {x : ℝ≥0∞} (hx : x ≠ ∞) {y : ℝ≥0∞} (hy : 1 < y) : + ∀ᶠ (n : ℕ) in at_top, x ^ (1 / n : ℝ) ≤ y := +begin + lift x to ℝ≥0 using hx, + by_cases y = ∞, + { exact eventually_of_forall (λ n, h.symm ▸ le_top) }, + { lift y to ℝ≥0 using h, + have := nnreal.eventually_pow_one_div_le x (by exact_mod_cast hy : 1 < y), + refine this.congr (eventually_of_forall $ λ n, _), + rw [coe_rpow_of_nonneg x (by positivity : 0 ≤ (1 / n : ℝ)), coe_le_coe] }, +end + private lemma continuous_at_rpow_const_of_pos {x : ℝ≥0∞} {y : ℝ} (h : 0 < y) : continuous_at (λ a : ℝ≥0∞, a ^ y) x := begin diff --git a/src/topology/algebra/module/character_space.lean b/src/topology/algebra/module/character_space.lean index 46744ed368b82..1257a49d62540 100644 --- a/src/topology/algebra/module/character_space.lean +++ b/src/topology/algebra/module/character_space.lean @@ -87,6 +87,9 @@ def to_non_unital_alg_hom (φ : character_space 𝕜 A) : A →ₙₐ[𝕜] 𝕜 @[simp] lemma coe_to_non_unital_alg_hom (φ : character_space 𝕜 A) : ⇑(to_non_unital_alg_hom φ) = φ := rfl +instance [subsingleton A] : is_empty (character_space 𝕜 A) := +⟨λ φ, φ.prop.1 $ continuous_linear_map.ext (λ x, by simp only [subsingleton.elim x 0, map_zero])⟩ + variables (𝕜 A) lemma union_zero : From 10a3d03bc8b3e3bf65c8bb14a06a718cc580f19c Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 11 Oct 2022 00:11:04 +0000 Subject: [PATCH 666/828] feat(ring_theory/valuation): Equivalent conditions of DVRs. (#15081) --- src/ring_theory/discrete_valuation_ring.lean | 24 +- src/ring_theory/ideal/local_ring.lean | 10 + src/ring_theory/ideal/operations.lean | 11 + src/ring_theory/integral_closure.lean | 41 +++- src/ring_theory/valuation/tfae.lean | 245 +++++++++++++++++++ 5 files changed, 319 insertions(+), 12 deletions(-) create mode 100644 src/ring_theory/valuation/tfae.lean diff --git a/src/ring_theory/discrete_valuation_ring.lean b/src/ring_theory/discrete_valuation_ring.lean index 1d6cc27d1c22d..412da64f8312d 100644 --- a/src/ring_theory/discrete_valuation_ring.lean +++ b/src/ring_theory/discrete_valuation_ring.lean @@ -64,13 +64,9 @@ variable {R} open principal_ideal_ring -/-- An element of a DVR is irreducible iff it is a uniformizer, that is, generates the - maximal ideal of R -/ -theorem irreducible_iff_uniformizer (ϖ : R) : - irreducible ϖ ↔ maximal_ideal R = ideal.span {ϖ} := -⟨λ hϖ, (eq_maximal_ideal (is_maximal_of_irreducible hϖ)).symm, +theorem irreducible_of_span_eq_maximal_ideal {R : Type*} [comm_ring R] [local_ring R] [is_domain R] + (ϖ : R) (hϖ : ϖ ≠ 0) (h : maximal_ideal R = ideal.span {ϖ}) : irreducible ϖ := begin - intro h, have h2 : ¬(is_unit ϖ) := show ϖ ∈ maximal_ideal R, from h.symm ▸ submodule.mem_span_singleton_self ϖ, refine ⟨h2, _⟩, @@ -81,11 +77,17 @@ begin rcases ha with ⟨a, rfl⟩, rcases hb with ⟨b, rfl⟩, rw (show a * ϖ * (b * ϖ) = ϖ * (ϖ * (a * b)), by ring) at hab, - have h3 := eq_zero_of_mul_eq_self_right _ hab.symm, - { apply not_a_field R, - simp [h, h3] }, - { exact λ hh, h2 (is_unit_of_dvd_one ϖ ⟨_, hh.symm⟩) } -end⟩ + apply hϖ, + apply eq_zero_of_mul_eq_self_right _ hab.symm, + exact λ hh, h2 (is_unit_of_dvd_one ϖ ⟨_, hh.symm⟩) +end + +/-- An element of a DVR is irreducible iff it is a uniformizer, that is, generates the + maximal ideal of R -/ +theorem irreducible_iff_uniformizer (ϖ : R) : + irreducible ϖ ↔ maximal_ideal R = ideal.span {ϖ} := +⟨λ hϖ, (eq_maximal_ideal (is_maximal_of_irreducible hϖ)).symm, λ h, + irreducible_of_span_eq_maximal_ideal ϖ (λ e, not_a_field R $ by rwa [h, span_singleton_eq_bot]) h⟩ lemma _root_.irreducible.maximal_ideal_eq {ϖ : R} (h : irreducible ϖ) : maximal_ideal R = ideal.span {ϖ} := diff --git a/src/ring_theory/ideal/local_ring.lean b/src/ring_theory/ideal/local_ring.lean index c1dccf7c0f3c7..6f2836a06e23b 100644 --- a/src/ring_theory/ideal/local_ring.lean +++ b/src/ring_theory/ideal/local_ring.lean @@ -7,6 +7,7 @@ Authors: Kenny Lau, Chris Hughes, Mario Carneiro import algebra.algebra.basic import algebra.category.Ring.basic import ring_theory.ideal.operations +import ring_theory.jacobson_ideal /-! @@ -167,6 +168,15 @@ begin apply f.is_unit_map, end +lemma jacobson_eq_maximal_ideal (I : ideal R) (h : I ≠ ⊤) : + I.jacobson = local_ring.maximal_ideal R := +begin + apply le_antisymm, + { exact Inf_le ⟨local_ring.le_maximal_ideal h, local_ring.maximal_ideal.is_maximal R⟩ }, + { exact le_Inf (λ J (hJ : I ≤ J ∧ J.is_maximal), + le_of_eq (local_ring.eq_maximal_ideal hJ.2).symm) } +end + end local_ring end comm_ring diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 0f70cfb04772f..81495a3e4a4ee 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -259,6 +259,17 @@ theorem mem_ideal_smul_span_iff_exists_sum' {ι : Type*} (s : set ι) (f : ι ∃ (a : s →₀ R) (ha : ∀ i, a i ∈ I), a.sum (λ i c, c • f i) = x := by rw [← submodule.mem_ideal_smul_span_iff_exists_sum, ← set.image_eq_range] +lemma mem_smul_top_iff (N : submodule R M) (x : N) : + x ∈ I • (⊤ : submodule R N) ↔ (x : M) ∈ I • N := +begin + change _ ↔ N.subtype x ∈ I • N, + have : submodule.map N.subtype (I • ⊤) = I • N, + { rw [submodule.map_smul'', submodule.map_top, submodule.range_subtype] }, + rw ← this, + convert (function.injective.mem_set_image N.injective_subtype).symm using 1, + refl, +end + @[simp] lemma smul_comap_le_comap_smul (f : M →ₗ[R] M') (S : submodule R M') (I : ideal R) : I • S.comap f ≤ (I • S).comap f := begin diff --git a/src/ring_theory/integral_closure.lean b/src/ring_theory/integral_closure.lean index 4a83b5a97404b..83a6380292f74 100644 --- a/src/ring_theory/integral_closure.lean +++ b/src/ring_theory/integral_closure.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau -/ import data.polynomial.expand import linear_algebra.finite_dimensional -import linear_algebra.matrix.determinant +import linear_algebra.matrix.charpoly.linear_map import ring_theory.adjoin.fg import ring_theory.polynomial.scale_roots import ring_theory.polynomial.tower @@ -295,6 +295,45 @@ begin exact subalgebra.smul_mem _ (algebra.subset_adjoin $ hlx1 hr) _ end +lemma module.End.is_integral {M : Type*} [add_comm_group M] [module R M] [module.finite R M] : + algebra.is_integral R (module.End R M) := +linear_map.exists_monic_and_aeval_eq_zero R + +/-- Suppose `A` is an `R`-algebra, `M` is an `A`-module such that `a • m ≠ 0` for all non-zero `a` +and `m`. If `x : A` fixes a nontrivial f.g. `R`-submodule `N` of `M`, then `x` is `R`-integral. -/ +lemma is_integral_of_smul_mem_submodule {M : Type*} [add_comm_group M] [module R M] + [module A M] [is_scalar_tower R A M] [no_zero_smul_divisors A M] + (N : submodule R M) (hN : N ≠ ⊥) (hN' : N.fg) (x : A) + (hx : ∀ n ∈ N, x • n ∈ N) : is_integral R x := +begin + let A' : subalgebra R A := + { carrier := { x | ∀ n ∈ N, x • n ∈ N }, + mul_mem' := λ a b ha hb n hn, smul_smul a b n ▸ ha _ (hb _ hn), + one_mem' := λ n hn, (one_smul A n).symm ▸ hn, + add_mem' := λ a b ha hb n hn, (add_smul a b n).symm ▸ N.add_mem (ha _ hn) (hb _ hn), + zero_mem' := λ n hn, (zero_smul A n).symm ▸ N.zero_mem, + algebra_map_mem' := λ r n hn, (algebra_map_smul A r n).symm ▸ N.smul_mem r hn }, + let f : A' →ₐ[R] module.End R N := alg_hom.of_linear_map + { to_fun := λ x, (distrib_mul_action.to_linear_map R M x).restrict x.prop, + map_add' := λ x y, linear_map.ext $ λ n, subtype.ext $ add_smul x y n, + map_smul' := λ r s, linear_map.ext $ λ n, subtype.ext $ smul_assoc r s n } + (linear_map.ext $ λ n, subtype.ext $ one_smul _ _) + (λ x y, linear_map.ext $ λ n, subtype.ext $ mul_smul x y n), + obtain ⟨a, ha₁, ha₂⟩ : ∃ a ∈ N, a ≠ (0 : M), + { by_contra h', push_neg at h', apply hN, rwa eq_bot_iff }, + have : function.injective f, + { show function.injective f.to_linear_map, + rw [← linear_map.ker_eq_bot, eq_bot_iff], + intros s hs, + have : s.1 • a = 0 := congr_arg subtype.val (linear_map.congr_fun hs ⟨a, ha₁⟩), + exact subtype.ext ((eq_zero_or_eq_zero_of_smul_eq_zero this).resolve_right ha₂) }, + show is_integral R (A'.val ⟨x, hx⟩), + rw [is_integral_alg_hom_iff A'.val subtype.val_injective, + ← is_integral_alg_hom_iff f this], + haveI : module.finite R N := by rwa [module.finite_def, submodule.fg_top], + apply module.End.is_integral, +end + variables {f} lemma ring_hom.finite.to_is_integral (h : f.finite) : f.is_integral := diff --git a/src/ring_theory/valuation/tfae.lean b/src/ring_theory/valuation/tfae.lean new file mode 100644 index 0000000000000..b3ba9f5a435eb --- /dev/null +++ b/src/ring_theory/valuation/tfae.lean @@ -0,0 +1,245 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import ring_theory.valuation.valuation_subring +import ring_theory.ideal.cotangent +import ring_theory.dedekind_domain.basic + +/-! + +# Equivalent conditions for DVR + +In `discrete_valuation_ring.tfae`, we show that the following are equivalent for a +noetherian local domain `(R, m, k)`: +- `R` is a discrete valuation ring +- `R` is a valuation ring +- `R` is a dedekind domain +- `R` is integrally closed with a unique prime ideal +- `m` is principal +- `dimₖ m/m² = 1` +- Every nonzero ideal is a power of `m`. + +-/ + + +variables (R : Type*) [comm_ring R] (K : Type*) [field K] [algebra R K] [is_fraction_ring R K] + +open_locale discrete_valuation +open local_ring + +open_locale big_operators + +lemma exists_maximal_ideal_pow_eq_of_principal [is_noetherian_ring R] [local_ring R] [is_domain R] + (h : ¬ is_field R) (h' : (maximal_ideal R).is_principal) (I : ideal R) (hI : I ≠ ⊥) : + ∃ n : ℕ, I = (maximal_ideal R) ^ n := +begin + classical, + unfreezingI { obtain ⟨x, hx : _ = ideal.span _⟩ := h' }, + by_cases hI' : I = ⊤, { use 0, rw [pow_zero, hI', ideal.one_eq_top] }, + have H : ∀ r : R, ¬ (is_unit r) ↔ x ∣ r := + λ r, (set_like.ext_iff.mp hx r).trans ideal.mem_span_singleton, + have : x ≠ 0, + { rintro rfl, + apply ring.ne_bot_of_is_maximal_of_not_is_field (maximal_ideal.is_maximal R) h, + simp [hx] }, + have hx' := discrete_valuation_ring.irreducible_of_span_eq_maximal_ideal x this hx, + have H' : ∀ r : R, r ≠ 0 → r ∈ nonunits R → ∃ (n : ℕ), associated (x ^ n) r, + { intros r hr₁ hr₂, + obtain ⟨f, hf₁, rfl, hf₂⟩ := (wf_dvd_monoid.not_unit_iff_exists_factors_eq r hr₁).mp hr₂, + have : ∀ b ∈ f, associated x b, + { intros b hb, + exact irreducible.associated_of_dvd hx' (hf₁ b hb) ((H b).mp (hf₁ b hb).1) }, + clear hr₁ hr₂ hf₁, + induction f using multiset.induction with fa fs fh, + { exact (hf₂ rfl).elim }, + rcases eq_or_ne fs ∅ with rfl|hf', + { use 1, + rw [pow_one, multiset.prod_cons, multiset.empty_eq_zero, multiset.prod_zero, mul_one], + exact this _ (multiset.mem_cons_self _ _) }, + { obtain ⟨n, hn⟩ := fh hf' (λ b hb, this _ (multiset.mem_cons_of_mem hb)), + use n + 1, + rw [pow_add, multiset.prod_cons, mul_comm, pow_one], + exact associated.mul_mul (this _ (multiset.mem_cons_self _ _)) hn } }, + have : ∃ n : ℕ, x ^ n ∈ I, + { obtain ⟨r, hr₁, hr₂⟩ : ∃ r : R, r ∈ I ∧ r ≠ 0, + { by_contra h, push_neg at h, apply hI, rw eq_bot_iff, exact h }, + obtain ⟨n, u, rfl⟩ := H' r hr₂ (le_maximal_ideal hI' hr₁), + use n, + rwa [← I.unit_mul_mem_iff_mem u.is_unit, mul_comm] }, + use nat.find this, + apply le_antisymm, + { change ∀ s ∈ I, s ∈ _, + by_contra hI'', + push_neg at hI'', + obtain ⟨s, hs₁, hs₂⟩ := hI'', + apply hs₂, + by_cases hs₃ : s = 0, { rw hs₃, exact zero_mem _ }, + obtain ⟨n, u, rfl⟩ := H' s hs₃ (le_maximal_ideal hI' hs₁), + rw [mul_comm, ideal.unit_mul_mem_iff_mem _ u.is_unit] at ⊢ hs₁, + apply ideal.pow_le_pow (nat.find_min' this hs₁), + apply ideal.pow_mem_pow, + exact (H _).mpr (dvd_refl _) }, + { rw [hx, ideal.span_singleton_pow, ideal.span_le, set.singleton_subset_iff], + exact nat.find_spec this } +end + +lemma maximal_ideal_is_principal_of_is_dedekind_domain + [local_ring R] [is_domain R] [is_dedekind_domain R] : (maximal_ideal R).is_principal := +begin + classical, + by_cases ne_bot : maximal_ideal R = ⊥, + { rw ne_bot, apply_instance }, + obtain ⟨a, ha₁, ha₂⟩ : ∃ a ∈ maximal_ideal R, a ≠ (0 : R), + { by_contra h', push_neg at h', apply ne_bot, rwa eq_bot_iff }, + have hle : ideal.span {a} ≤ maximal_ideal R, + { rwa [ideal.span_le, set.singleton_subset_iff] }, + have : (ideal.span {a}).radical = maximal_ideal R, + { rw ideal.radical_eq_Inf, + apply le_antisymm, + { exact Inf_le ⟨hle, infer_instance⟩ }, + { refine le_Inf (λ I hI, (eq_maximal_ideal $ + is_dedekind_domain.dimension_le_one _ (λ e, ha₂ _) hI.2).ge), + rw [← ideal.span_singleton_eq_bot, eq_bot_iff, ← e], exact hI.1 } }, + have : ∃ n, maximal_ideal R ^ n ≤ ideal.span {a}, + { rw ← this, apply ideal.exists_radical_pow_le_of_fg, exact is_noetherian.noetherian _ }, + cases hn : nat.find this, + { have := nat.find_spec this, + rw [hn, pow_zero, ideal.one_eq_top] at this, + exact (ideal.is_maximal.ne_top infer_instance (eq_top_iff.mpr $ this.trans hle)).elim }, + obtain ⟨b, hb₁, hb₂⟩ : ∃ b ∈ maximal_ideal R ^ n, ¬ b ∈ ideal.span {a}, + { by_contra h', push_neg at h', rw nat.find_eq_iff at hn, + exact hn.2 n n.lt_succ_self (λ x hx, not_not.mp (h' x hx)) }, + have hb₃ : ∀ m ∈ maximal_ideal R, ∃ k : R, k * a = b * m, + { intros m hm, rw ← ideal.mem_span_singleton', apply nat.find_spec this, + rw [hn, pow_succ'], exact ideal.mul_mem_mul hb₁ hm }, + have hb₄ : b ≠ 0, + { rintro rfl, apply hb₂, exact zero_mem _ }, + let K := fraction_ring R, + let x : K := algebra_map R K b / algebra_map R K a, + let M := submodule.map (algebra.of_id R K).to_linear_map (maximal_ideal R), + have ha₃ : algebra_map R K a ≠ 0 := is_fraction_ring.to_map_eq_zero_iff.not.mpr ha₂, + by_cases hx : ∀ y ∈ M, x * y ∈ M, + { have := is_integral_of_smul_mem_submodule M _ _ x hx, + { obtain ⟨y, e⟩ := is_integrally_closed.algebra_map_eq_of_integral this, + refine (hb₂ (ideal.mem_span_singleton'.mpr ⟨y, _⟩)).elim, + apply is_fraction_ring.injective R K, + rw [map_mul, e, div_mul_cancel _ ha₃] }, + { rw submodule.ne_bot_iff, refine ⟨_, ⟨a, ha₁, rfl⟩, _⟩, + exact is_fraction_ring.to_map_eq_zero_iff.not.mpr ha₂ }, + { apply submodule.fg.map, exact is_noetherian.noetherian _ } }, + { have : (M.map (distrib_mul_action.to_linear_map R K x)).comap + (algebra.of_id R K).to_linear_map = ⊤, + { by_contra h, apply hx, + rintros m' ⟨m, hm, (rfl : algebra_map R K m = m')⟩, + obtain ⟨k, hk⟩ := hb₃ m hm, + have hk' : x * algebra_map R K m = algebra_map R K k, + { rw [← mul_div_right_comm, ← map_mul, ← hk, map_mul, mul_div_cancel _ ha₃] }, + exact ⟨k, le_maximal_ideal h ⟨_, ⟨_, hm, rfl⟩, hk'⟩, hk'.symm⟩ }, + obtain ⟨y, hy₁, hy₂⟩ : ∃ y ∈ maximal_ideal R, b * y = a, + { rw [ideal.eq_top_iff_one, submodule.mem_comap] at this, + obtain ⟨_, ⟨y, hy, rfl⟩, hy' : x * algebra_map R K y = algebra_map R K 1⟩ := this, + rw [map_one, ← mul_div_right_comm, div_eq_one_iff_eq ha₃, ← map_mul] at hy', + exact ⟨y, hy, is_fraction_ring.injective R K hy'⟩ }, + refine ⟨⟨y, _⟩⟩, + apply le_antisymm, + { intros m hm, obtain ⟨k, hk⟩ := hb₃ m hm, rw [← hy₂, mul_comm, mul_assoc] at hk, + rw [← mul_left_cancel₀ hb₄ hk, mul_comm], exact ideal.mem_span_singleton'.mpr ⟨_, rfl⟩ }, + { rwa [submodule.span_le, set.singleton_subset_iff] } } +end + +lemma discrete_valuation_ring.tfae [is_noetherian_ring R] [local_ring R] [is_domain R] + (h : ¬ is_field R) : + tfae [discrete_valuation_ring R, + valuation_ring R, + is_dedekind_domain R, + is_integrally_closed R ∧ ∃! P : ideal R, P ≠ ⊥ ∧ P.is_prime, + (maximal_ideal R).is_principal, + finite_dimensional.finrank (residue_field R) (cotangent_space R) = 1, + ∀ I ≠ ⊥, ∃ n : ℕ, I = (maximal_ideal R) ^ n] := +begin + have ne_bot := ring.ne_bot_of_is_maximal_of_not_is_field (maximal_ideal.is_maximal R) h, + classical, + rw finrank_eq_one_iff', + tfae_have : 1 → 2, + { introI _, apply_instance }, + tfae_have : 2 → 1, + { introI _, + haveI := is_bezout.to_gcd_domain R, + haveI : unique_factorization_monoid R := ufm_of_gcd_of_wf_dvd_monoid, + apply discrete_valuation_ring.of_ufd_of_unique_irreducible, + { obtain ⟨x, hx₁, hx₂⟩ := ring.exists_not_is_unit_of_not_is_field h, + obtain ⟨p, hp₁, hp₂⟩ := wf_dvd_monoid.exists_irreducible_factor hx₂ hx₁, + exact ⟨p, hp₁⟩ }, + { exact valuation_ring.unique_irreducible } }, + tfae_have : 1 → 4, + { introI H, + exact ⟨infer_instance, ((discrete_valuation_ring.iff_pid_with_one_nonzero_prime R).mp H).2⟩ }, + tfae_have : 4 → 3, + { rintros ⟨h₁, h₂⟩, exact ⟨infer_instance, λ I hI hI', unique_of_exists_unique h₂ + ⟨ne_bot, infer_instance⟩ ⟨hI, hI'⟩ ▸ maximal_ideal.is_maximal R, h₁⟩ }, + tfae_have : 3 → 5, + { introI h, exact maximal_ideal_is_principal_of_is_dedekind_domain R }, + tfae_have : 5 → 6, + { rintro ⟨x, hx⟩, + have : x ∈ maximal_ideal R := by { rw hx, exact submodule.subset_span (set.mem_singleton x) }, + let x' : maximal_ideal R := ⟨x, this⟩, + use submodule.quotient.mk x', + split, + { intro e, + rw submodule.quotient.mk_eq_zero at e, + apply ring.ne_bot_of_is_maximal_of_not_is_field (maximal_ideal.is_maximal R) h, + apply submodule.eq_bot_of_le_smul_of_le_jacobson_bot (maximal_ideal R), + { exact ⟨{x}, (finset.coe_singleton x).symm ▸ hx.symm⟩ }, + { conv_lhs { rw hx }, + rw submodule.mem_smul_top_iff at e, + rwa [submodule.span_le, set.singleton_subset_iff] }, + { rw local_ring.jacobson_eq_maximal_ideal (⊥ : ideal R) bot_ne_top, exact le_refl _ } }, + { refine λ w, quotient.induction_on' w $ λ y, _, + obtain ⟨y, hy⟩ := y, + rw [hx, submodule.mem_span_singleton] at hy, + obtain ⟨a, rfl⟩ := hy, + exact ⟨ideal.quotient.mk _ a, rfl⟩ } }, + tfae_have : 6 → 5, + { rintro ⟨x, hx, hx'⟩, + induction x using quotient.induction_on', + use x, + apply le_antisymm, + swap, { rw [submodule.span_le, set.singleton_subset_iff], exact x.prop }, + have h₁ : (ideal.span {x} : ideal R) ⊔ maximal_ideal R ≤ + ideal.span {x} ⊔ (maximal_ideal R) • (maximal_ideal R), + { refine sup_le le_sup_left _, + rintros m hm, + obtain ⟨c, hc⟩ := hx' (submodule.quotient.mk ⟨m, hm⟩), + induction c using quotient.induction_on', + rw ← sub_sub_cancel (c * x) m, + apply sub_mem _ _, + { apply_instance }, + { refine ideal.mem_sup_left (ideal.mem_span_singleton'.mpr ⟨c, rfl⟩) }, + { have := (submodule.quotient.eq _).mp hc, + rw [submodule.mem_smul_top_iff] at this, + exact ideal.mem_sup_right this } }, + have h₂ : maximal_ideal R ≤ (⊥ : ideal R).jacobson, + { rw local_ring.jacobson_eq_maximal_ideal, exacts [le_refl _, bot_ne_top] }, + have := submodule.smul_sup_eq_smul_sup_of_le_smul_of_le_jacobson + (is_noetherian.noetherian _) h₂ h₁, + rw [submodule.bot_smul, sup_bot_eq] at this, + rw [← sup_eq_left, eq_comm], + exact le_sup_left.antisymm (h₁.trans $ le_of_eq this) }, + tfae_have : 5 → 7, + { exact exists_maximal_ideal_pow_eq_of_principal R h }, + tfae_have : 7 → 2, + { rw valuation_ring.iff_ideal_total, + intro H, + constructor, + intros I J, + by_cases hI : I = ⊥, { subst hI, left, exact bot_le }, + by_cases hJ : J = ⊥, { subst hJ, right, exact bot_le }, + obtain ⟨n, rfl⟩ := H I hI, + obtain ⟨m, rfl⟩ := H J hJ, + cases le_total m n with h' h', + { left, exact ideal.pow_le_pow h' }, + { right, exact ideal.pow_le_pow h' } }, + tfae_finish, +end From d840349db150e13809a02a4f3021cbf0850c9b1c Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Tue, 11 Oct 2022 03:45:35 +0000 Subject: [PATCH 667/828] feat(algebra/algebra/basic): make algebra_map a low prio has_lift_t (#16567) Make `algebra_map` a coercion (or more precisely a `has_lift_t`). I have an example where making algebra_map a coercion makes a lot of code far less messy and I'm wondering whether it's a good idea in general. Co-authored-by: Vierkantor --- src/algebra/algebra/basic.lean | 68 +++++++++++++++++++++++++ src/data/complex/is_R_or_C.lean | 10 ++-- src/topology/algebra/uniform_field.lean | 3 +- 3 files changed, 74 insertions(+), 7 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index c81a009eb39b1..29a545f6e3660 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -124,6 +124,74 @@ end prio def algebra_map (R : Type u) (A : Type v) [comm_semiring R] [semiring A] [algebra R A] : R →+* A := algebra.to_ring_hom +namespace algebra_map + +def has_lift_t (R A : Type*) [comm_semiring R] [semiring A] [algebra R A] : + has_lift_t R A := ⟨λ r, algebra_map R A r⟩ + +attribute [instance, priority 900] algebra_map.has_lift_t + +section comm_semiring_semiring + +variables {R A : Type*} [comm_semiring R] [semiring A] [algebra R A] + +@[simp, norm_cast] lemma coe_zero : (↑(0 : R) : A) = 0 := map_zero (algebra_map R A) +@[simp, norm_cast] lemma coe_one : (↑(1 : R) : A) = 1 := map_one (algebra_map R A) +@[norm_cast] lemma coe_add (a b : R) : (↑(a + b : R) : A) = ↑a + ↑b := +map_add (algebra_map R A) a b +@[norm_cast] lemma coe_mul (a b : R) : (↑(a * b : R) : A) = ↑a * ↑b := +map_mul (algebra_map R A) a b +@[norm_cast] lemma coe_pow (a : R) (n : ℕ) : (↑(a ^ n : R) : A) = ↑a ^ n := +map_pow (algebra_map R A) _ _ + +end comm_semiring_semiring + +section comm_ring_ring + +variables {R A : Type*} [comm_ring R] [ring A] [algebra R A] + +@[norm_cast] lemma coe_neg (x : R) : (↑(-x : R) : A) = -↑x := +map_neg (algebra_map R A) x + +end comm_ring_ring + +section comm_semiring_comm_semiring + +variables {R A : Type*} [comm_semiring R] [comm_semiring A] [algebra R A] + +open_locale big_operators + +-- direct to_additive fails because of some mix-up with polynomials +@[norm_cast] lemma coe_prod {ι : Type*} {s : finset ι} (a : ι → R) : + (↑( ∏ (i : ι) in s, a i : R) : A) = ∏ (i : ι) in s, (↑(a i) : A) := +map_prod (algebra_map R A) a s + +-- to_additive fails for some reason +@[norm_cast] lemma coe_sum {ι : Type*} {s : finset ι} (a : ι → R) : + ↑(( ∑ (i : ι) in s, a i)) = ∑ (i : ι) in s, (↑(a i) : A) := +map_sum (algebra_map R A) a s + +attribute [to_additive] coe_prod + +end comm_semiring_comm_semiring + +section field_nontrivial + +variables {R A : Type*} [field R] [comm_semiring A] [nontrivial A] [algebra R A] + +@[norm_cast, simp] lemma coe_inj {a b : R} : (↑a : A) = ↑b ↔ a = b := +⟨λ h, (algebra_map R A).injective h, by rintro rfl; refl⟩ + +@[norm_cast, simp] lemma lift_map_eq_zero_iff (a : R) : (↑a : A) = 0 ↔ a = 0 := +begin + rw (show (0 : A) = ↑(0 : R), from (map_zero (algebra_map R A)).symm), + norm_cast, +end + +end field_nontrivial + +end algebra_map + /-- Creating an algebra from a morphism to the center of a semiring. -/ def ring_hom.to_algebra' {R S} [comm_semiring R] [semiring S] (i : R →+* S) (h : ∀ c x, i c * x = x * i c) : diff --git a/src/data/complex/is_R_or_C.lean b/src/data/complex/is_R_or_C.lean index f8b57dac29b84..c517bf360b99e 100644 --- a/src/data/complex/is_R_or_C.lean +++ b/src/data/complex/is_R_or_C.lean @@ -114,17 +114,17 @@ theorem ext : ∀ {z w : K}, re z = re w → im z = im w → z = w := by { simp_rw ext_iff, cc } -@[simp, norm_cast, is_R_or_C_simps, priority 900] lemma of_real_zero : ((0 : ℝ) : K) = 0 := +@[norm_cast] lemma of_real_zero : ((0 : ℝ) : K) = 0 := by rw [of_real_alg, zero_smul] @[simp, is_R_or_C_simps] lemma zero_re' : re (0 : K) = (0 : ℝ) := re.map_zero -@[simp, norm_cast, is_R_or_C_simps, priority 900] lemma of_real_one : ((1 : ℝ) : K) = 1 := +@[norm_cast] lemma of_real_one : ((1 : ℝ) : K) = 1 := by rw [of_real_alg, one_smul] @[simp, is_R_or_C_simps] lemma one_re : re (1 : K) = 1 := by rw [←of_real_one, of_real_re] @[simp, is_R_or_C_simps] lemma one_im : im (1 : K) = 0 := by rw [←of_real_one, of_real_im] -@[simp, norm_cast, priority 900] theorem of_real_inj {z w : ℝ} : (z : K) = (w : K) ↔ z = w := +@[norm_cast] theorem of_real_inj {z w : ℝ} : (z : K) = (w : K) ↔ z = w := { mp := λ h, by { convert congr_arg re h; simp only [of_real_re] }, mpr := λ h, by rw h } @@ -137,7 +137,6 @@ by simp only [bit0, map_add] @[simp, is_R_or_C_simps] lemma bit1_im (z : K) : im (bit1 z) = bit0 (im z) := by simp only [bit1, add_right_eq_self, add_monoid_hom.map_add, bit0_im, one_im] -@[simp, is_R_or_C_simps, priority 900] theorem of_real_eq_zero {z : ℝ} : (z : K) = 0 ↔ z = 0 := by rw [←of_real_zero]; exact of_real_inj @@ -558,7 +557,8 @@ begin end lemma re_eq_self_of_le {a : K} (h : abs a ≤ re a) : (re a : K) = a := -by { rw ← re_add_im a, simp only [im_eq_zero_of_le h, add_zero, zero_mul] with is_R_or_C_simps } +by { rw ← re_add_im a, simp only [im_eq_zero_of_le h, add_zero, zero_mul, algebra_map.coe_zero] + with is_R_or_C_simps, } lemma abs_add (z w : K) : abs (z + w) ≤ abs z + abs w := (mul_self_le_mul_self_iff (abs_nonneg _) diff --git a/src/topology/algebra/uniform_field.lean b/src/topology/algebra/uniform_field.lean index 7539f91be14b0..f8a53e5876f01 100644 --- a/src/topology/algebra/uniform_field.lean +++ b/src/topology/algebra/uniform_field.lean @@ -144,8 +144,7 @@ begin dsimp [c, f], rw hat_inv_extends z_ne, norm_cast, - rw mul_inv_cancel z_ne, - norm_cast }, + rw mul_inv_cancel z_ne, }, replace fxclo := closure_mono this fxclo, rwa [closure_singleton, mem_singleton_iff] at fxclo end From aadde30334fb810adb293b0d895ce78ad7e394a5 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 11 Oct 2022 03:45:36 +0000 Subject: [PATCH 668/828] feat(analysis/calculus): bounded variation functions (#16683) We define bounded variation functions, and prove that such a function is the difference of two monotone functions. We also show that Lipschitz functions have bounded variation. --- src/analysis/bounded_variation.lean | 653 ++++++++++++++++++++++++++++ 1 file changed, 653 insertions(+) create mode 100644 src/analysis/bounded_variation.lean diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean new file mode 100644 index 0000000000000..a9a24529d5bad --- /dev/null +++ b/src/analysis/bounded_variation.lean @@ -0,0 +1,653 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import data.set.intervals.monotone +import measure_theory.measure.lebesgue + +/-! +# Functions of bounded variation + +We study functions of bounded variation. In particular, we show that a bounded variation function +is a difference of monotone functions, and that Lipschitz functions have bounded variation + +## Main definitions and results + +* `evariation_on f s` is the total variation of the function `f` on the set `s`, in `ℝ≥0∞`. +* `has_bounded_variation_on f s` registers that the variation of `f` on `s` is finite. +* `has_locally_bounded_variation f s` registers that `f` has finite variation on any compact + subinterval of `s`. + +* `evariation_on.Icc_add_Icc` states that the variation of `f` on `[a, c]` is the sum of its + variations on `[a, b]` and `[b, c]`. +* `has_locally_bounded_variation_on.exists_monotone_on_sub_monotone_on` proves that a function + with locally bounded variation is the difference of two monotone functions. +* `lipschitz_with.has_locally_bounded_variation_on` shows that a Lipschitz function has locally + bounded variation. + +We also give several variations around these results. + +## Implementation + +We define the variation as an extended nonnegative real, to allow for infinite variation. This makes +it possible to use the complete linear order structure of `ℝ≥0∞`. The proofs would be much +more tedious with an `ℝ`-valued or `ℝ≥0`-valued variation, since one would always need to check +that the sets one uses are nonempty and bounded above as these are only conditionally complete. +-/ + +open_locale big_operators nnreal ennreal +open set measure_theory + +variables {α : Type*} [linear_order α] +{E F : Type*} [pseudo_emetric_space E] [pseudo_emetric_space F] +{V : Type*} [normed_add_comm_group V] [normed_space ℝ V] [finite_dimensional ℝ V] + +/-- The (extended real valued) variation of a function `f` on a set `s` inside a linear order is +the supremum of the sum of `edist (f (u (i+1))) (f (u i))` over all finite increasing +sequences `u` in `s`. -/ +noncomputable def evariation_on (f : α → E) (s : set α) : ℝ≥0∞ := +⨆ (p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s}), + ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i)) + +/-- A function has bounded variation on a set `s` if its total variation there is finite. -/ +def has_bounded_variation_on (f : α → E) (s : set α) := +evariation_on f s ≠ ∞ + +/-- A function has locally bounded variation on a set `s` if, given any interval `[a, b]` with +endpoints in `s`, then the function has finite variation on `s ∩ [a, b]`. -/ +def has_locally_bounded_variation_on (f : α → E) (s : set α) := +∀ a b, a ∈ s → b ∈ s → has_bounded_variation_on f (s ∩ Icc a b) + +/-! ## Basic computations of variation -/ + +namespace evariation_on + +lemma nonempty_monotone_mem {s : set α} (hs : s.nonempty) : + nonempty {u // monotone u ∧ ∀ (i : ℕ), u i ∈ s} := +begin + obtain ⟨x, hx⟩ := hs, + exact ⟨⟨λ i, x, λ i j hij, le_rfl, λ i, hx⟩⟩, +end + +lemma sum_le + (f : α → E) {s : set α} (n : ℕ) {u : ℕ → α} (hu : monotone u) (us : ∀ i, u i ∈ s) : + ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) ≤ evariation_on f s := +begin + let p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s} := (n, ⟨u, hu, us⟩), + change ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i)) + ≤ evariation_on f s, + exact le_supr (λ (p : ℕ × {u : ℕ → α // monotone u ∧ ∀ i, u i ∈ s}), + ∑ i in finset.range p.1, edist (f ((p.2 : ℕ → α) (i+1))) (f ((p.2 : ℕ → α) i))) _, +end + +lemma sum_le_of_monotone_on_Iic + (f : α → E) {s : set α} {n : ℕ} {u : ℕ → α} (hu : monotone_on u (Iic n)) + (us : ∀ i ≤ n, u i ∈ s) : + ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) ≤ evariation_on f s := +begin + let v := λ i, if i ≤ n then u i else u n, + have vs : ∀ i, v i ∈ s, + { assume i, + simp only [v], + split_ifs, + { exact us i h }, + { exact us n le_rfl } }, + have hv : monotone v, + { apply monotone_nat_of_le_succ (λ i, _), + simp only [v], + rcases lt_trichotomy i n with hi|rfl|hi, + { have : i + 1 ≤ n, by linarith, + simp only [hi.le, this, if_true], + exact hu hi.le this (nat.le_succ i) }, + { simp only [le_refl, if_true, add_le_iff_nonpos_right, le_zero_iff, nat.one_ne_zero, + if_false] }, + { have A : ¬(i ≤ n), by linarith, + have B : ¬(i + 1 ≤ n), by linarith, + simp [A, B] } }, + convert sum_le f n hv vs using 1, + apply finset.sum_congr rfl (λ i hi, _), + simp only [finset.mem_range] at hi, + have : i + 1 ≤ n, by linarith, + simp only [v], + simp [this, hi.le], +end + +lemma sum_le_of_monotone_on_Icc + (f : α → E) {s : set α} {m n : ℕ} {u : ℕ → α} (hu : monotone_on u (Icc m n)) + (us : ∀ i ∈ Icc m n, u i ∈ s) : + ∑ i in finset.Ico m n, edist (f (u (i+1))) (f (u i)) ≤ evariation_on f s := +begin + rcases le_or_lt n m with hnm|hmn, + { simp only [finset.Ico_eq_empty_of_le hnm, finset.sum_empty, zero_le'] }, + let v := λ i, u (m + i), + have hv : monotone_on v (Iic (n - m)), + { assume a ha b hb hab, + simp only [le_tsub_iff_left hmn.le, mem_Iic] at ha hb, + exact hu ⟨le_add_right le_rfl, ha⟩ ⟨le_add_right le_rfl, hb⟩ (add_le_add le_rfl hab) }, + have vs : ∀ i ∈ Iic (n - m), v i ∈ s, + { assume i hi, + simp only [le_tsub_iff_left hmn.le, mem_Iic] at hi, + exact us _ ⟨le_add_right le_rfl, hi⟩ }, + calc ∑ i in finset.Ico m n, edist (f (u (i + 1))) (f (u i)) + = ∑ i in finset.range (n - m), edist (f (u (m + i + 1))) (f (u (m + i))) : + begin + rw [finset.range_eq_Ico], + convert (finset.sum_Ico_add (λ i, edist (f (u (i + 1))) (f (u i))) 0 (n - m) m).symm, + { rw [zero_add] }, + { rw tsub_add_cancel_of_le hmn.le } + end + ... = ∑ i in finset.range (n - m), edist (f (v (i + 1))) (f (v i)) : + begin + apply finset.sum_congr rfl (λ i hi, _), + simp only [v, add_assoc], + end + ... ≤ evariation_on f s : sum_le_of_monotone_on_Iic f hv vs, +end + +lemma mono (f : α → E) {s t : set α} (hst : t ⊆ s) : + evariation_on f t ≤ evariation_on f s := +begin + apply supr_le _, + rintros ⟨n, ⟨u, hu, ut⟩⟩, + exact sum_le f n hu (λ i, hst (ut i)), +end + +lemma _root_.has_bounded_variation_on.mono {f : α → E} {s : set α} + (h : has_bounded_variation_on f s) {t : set α} (ht : t ⊆ s) : + has_bounded_variation_on f t := +(lt_of_le_of_lt (evariation_on.mono f ht) (lt_top_iff_ne_top.2 h)).ne + +lemma _root_.has_bounded_variation_on.has_locally_bounded_variation_on {f : α → E} {s : set α} + (h : has_bounded_variation_on f s) : has_locally_bounded_variation_on f s := +λ x y hx hy, h.mono (inter_subset_left _ _) + +@[simp] protected lemma subsingleton (f : α → E) {s : set α} (hs : s.subsingleton) : + evariation_on f s = 0 := +begin + apply le_antisymm _ (zero_le _), + apply supr_le _, + rintros ⟨n, ⟨u, hu, ut⟩⟩, + have : ∀ i, u i = u 0, from λ i, hs (ut _) (ut _), + simp [subtype.coe_mk, le_zero_iff, finset.sum_eq_zero_iff, finset.mem_range, this], +end + +lemma edist_le (f : α → E) {s : set α} {x y : α} (hx : x ∈ s) (hy : y ∈ s) : + edist (f x) (f y) ≤ evariation_on f s := +begin + wlog hxy : x ≤ y := le_total x y using [x y, y x] tactic.skip, swap, + { assume hx hy, + rw edist_comm, + exact this hy hx }, + let u : ℕ → α := λ n, if n = 0 then x else y, + have hu : monotone u, + { assume m n hmn, + dsimp only [u], + split_ifs, + exacts [le_rfl, hxy, by linarith [pos_iff_ne_zero.2 h], le_rfl] }, + have us : ∀ i, u i ∈ s, + { assume i, + dsimp only [u], + split_ifs, + exacts [hx, hy] }, + convert sum_le f 1 hu us, + simp [u, edist_comm], +end + +lemma _root_.has_bounded_variation_on.dist_le {E : Type*} [pseudo_metric_space E] + {f : α → E} {s : set α} (h : has_bounded_variation_on f s) {x y : α} (hx : x ∈ s) (hy : y ∈ s) : + dist (f x) (f y) ≤ (evariation_on f s).to_real := +begin + rw [← ennreal.of_real_le_of_real_iff ennreal.to_real_nonneg, ennreal.of_real_to_real h, + ← edist_dist], + exact edist_le f hx hy +end + +lemma _root_.has_bounded_variation_on.sub_le + {f : α → ℝ} {s : set α} (h : has_bounded_variation_on f s) {x y : α} (hx : x ∈ s) (hy : y ∈ s) : + f x - f y ≤ (evariation_on f s).to_real := +begin + apply (le_abs_self _).trans, + rw ← real.dist_eq, + exact h.dist_le hx hy +end + +/-- Consider a monotone function `u` parameterizing some points of a set `s`. Given `x ∈ s`, then +one can find another monotone function `v` parameterizing the same points as `u`, with `x` added. +In particular, the variation of a function along `u` is bounded by its variation along `v`. -/ +lemma add_point (f : α → E) {s : set α} {x : α} (hx : x ∈ s) + (u : ℕ → α) (hu : monotone u) (us : ∀ i, u i ∈ s) (n : ℕ) : + ∃ (v : ℕ → α) (m : ℕ), monotone v ∧ (∀ i, v i ∈ s) ∧ x ∈ v '' (Iio m) ∧ + ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) ≤ + ∑ j in finset.range m, edist (f (v (j+1))) (f (v j)) := +begin + rcases le_or_lt (u n) x with h|h, + { let v := λ i, if i ≤ n then u i else x, + have vs : ∀ i, v i ∈ s, + { assume i, + simp only [v], + split_ifs, + { exact us i }, + { exact hx } }, + have hv : monotone v, + { apply monotone_nat_of_le_succ (λ i, _), + simp only [v], + rcases lt_trichotomy i n with hi|rfl|hi, + { have : i + 1 ≤ n, by linarith, + simp only [hi.le, this, if_true], + exact hu (nat.le_succ i) }, + { simp only [le_refl, if_true, add_le_iff_nonpos_right, le_zero_iff, nat.one_ne_zero, + if_false, h], }, + { have A : ¬(i ≤ n), by linarith, + have B : ¬(i + 1 ≤ n), by linarith, + simp [A, B] } }, + refine ⟨v, n+2, hv, vs, (mem_image _ _ _).2 ⟨n+1, _, _⟩, _⟩, + { rw mem_Iio, exact nat.lt_succ_self (n+1) }, + { have : ¬(n + 1 ≤ n), by linarith, + simp only [this, ite_eq_right_iff, is_empty.forall_iff] }, + { calc + ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) + = ∑ i in finset.range n, edist (f (v (i+1))) (f (v i)) : + begin + apply finset.sum_congr rfl (λ i hi, _), + simp only [finset.mem_range] at hi, + have : i + 1 ≤ n, by linarith, + dsimp only [v], + simp only [hi.le, this, if_true], + end + ... ≤ ∑ j in finset.range (n + 2), edist (f (v (j+1))) (f (v j)) : + finset.sum_le_sum_of_subset (finset.range_mono (by linarith)) } }, + have exists_N : ∃ N, N ≤ n ∧ x < u N, from ⟨n, le_rfl, h⟩, + let N := nat.find exists_N, + have hN : N ≤ n ∧ x < u N := nat.find_spec exists_N, + let w : ℕ → α := λ i, if i < N then u i else if i = N then x else u (i - 1), + have ws : ∀ i, w i ∈ s, + { dsimp only [w], + assume i, + split_ifs, + exacts [us _, hx, us _] }, + have hw : monotone w, + { apply monotone_nat_of_le_succ (λ i, _), + dsimp [w], + rcases lt_trichotomy (i + 1) N with hi|hi|hi, + { have : i < N, by linarith, + simp [hi, this], + exact hu (nat.le_succ _) }, + { have A : i < N, by linarith, + have B : ¬(i + 1 < N), by linarith, + rw [if_pos A, if_neg B, if_pos hi], + have T := nat.find_min exists_N A, + push_neg at T, + exact T (A.le.trans hN.1) }, + { have A : ¬(i < N), by linarith, + have B : ¬(i + 1 < N), by linarith, + have C : ¬(i + 1 = N), by linarith, + have D : i + 1 - 1 = i := nat.pred_succ i, + rw [if_neg A, if_neg B, if_neg C, D], + split_ifs, + { exact hN.2.le.trans (hu (by linarith)) }, + { exact hu (nat.pred_le _) } } }, + refine ⟨w, n+1, hw, ws, (mem_image _ _ _).2 ⟨N, hN.1.trans_lt (nat.lt_succ_self n), _⟩, _⟩, + { dsimp only [w], rw [if_neg (lt_irrefl N), if_pos rfl] }, + rcases eq_or_lt_of_le (zero_le N) with Npos|Npos, + { calc ∑ i in finset.range n, edist (f (u (i + 1))) (f (u i)) + = ∑ i in finset.range n, edist (f (w (1 + i + 1))) (f (w (1 + i))) : + begin + apply finset.sum_congr rfl (λ i hi, _), + dsimp only [w], + simp only [← Npos, nat.not_lt_zero, nat.add_succ_sub_one, add_zero, if_false, + add_eq_zero_iff, nat.one_ne_zero, false_and, nat.succ_add_sub_one, zero_add], + rw add_comm 1 i, + end + ... = ∑ i in finset.Ico 1 (n + 1), edist (f (w (i + 1))) (f (w i)) : + begin + rw finset.range_eq_Ico, + exact finset.sum_Ico_add (λ i, edist (f (w (i + 1))) (f (w i))) 0 n 1, + end + ... ≤ ∑ j in finset.range (n + 1), edist (f (w (j + 1))) (f (w j)) : + begin + apply finset.sum_le_sum_of_subset _, + rw finset.range_eq_Ico, + exact finset.Ico_subset_Ico zero_le_one le_rfl, + end }, + { calc ∑ i in finset.range n, edist (f (u (i + 1))) (f (u i)) + = ∑ i in finset.Ico 0 (N-1), edist (f (u (i + 1))) (f (u i)) + + ∑ i in finset.Ico (N-1) N, edist (f (u (i + 1))) (f (u i)) + + ∑ i in finset.Ico N n, edist (f (u (i + 1))) (f (u i)) : + begin + rw [finset.sum_Ico_consecutive, finset.sum_Ico_consecutive, finset.range_eq_Ico], + { exact zero_le _ }, + { exact hN.1 }, + { exact zero_le _}, + { exact nat.pred_le _ } + end + ... = ∑ i in finset.Ico 0 (N-1), edist (f (w (i + 1))) (f (w i)) + + edist (f (u N)) (f (u (N - 1))) + + ∑ i in finset.Ico N n, edist (f (w (1 + i + 1))) (f (w (1 + i))) : + begin + congr' 1, congr' 1, + { apply finset.sum_congr rfl (λ i hi, _), + simp only [finset.mem_Ico, zero_le', true_and] at hi, + dsimp only [w], + have A : i + 1 < N, from nat.lt_pred_iff.1 hi, + have B : i < N, by linarith, + rw [if_pos A, if_pos B] }, + { have A : N - 1 + 1 = N, from nat.succ_pred_eq_of_pos Npos, + have : finset.Ico (N - 1) N = {N - 1}, by rw [← nat.Ico_succ_singleton, A], + simp only [this, A, finset.sum_singleton] }, + { apply finset.sum_congr rfl (λ i hi, _), + simp only [finset.mem_Ico] at hi, + dsimp only [w], + have A : ¬(1 + i + 1 < N), by linarith, + have B : ¬(1 + i + 1 = N), by linarith, + have C : ¬(1 + i < N), by linarith, + have D : ¬(1 + i = N), by linarith, + rw [if_neg A, if_neg B, if_neg C, if_neg D], + congr' 3; + { rw eq_tsub_iff_add_eq_of_le, { abel }, { linarith } } } + end + ... = ∑ i in finset.Ico 0 (N-1), edist (f (w (i + 1))) (f (w i)) + + edist (f (w (N + 1))) (f (w (N - 1))) + + ∑ i in finset.Ico (N + 1) (n + 1), edist (f (w (i + 1))) (f (w (i))) : + begin + congr' 1, congr' 1, + { dsimp only [w], + have A : ¬(N + 1 < N), by linarith, + have B : N - 1 < N := nat.pred_lt Npos.ne', + simp only [A, not_and, not_lt, nat.succ_ne_self, nat.add_succ_sub_one, add_zero, if_false, + B, if_true] }, + { exact finset.sum_Ico_add (λ i, edist (f (w (i + 1))) (f (w i))) N n 1 } + end + ... ≤ ∑ i in finset.Ico 0 (N - 1), edist (f (w (i + 1))) (f (w i)) + + ∑ i in finset.Ico (N - 1) (N + 1), edist (f (w (i + 1))) (f (w i)) + + ∑ i in finset.Ico (N + 1) (n + 1), edist (f (w (i + 1))) (f (w i)) : + begin + refine add_le_add (add_le_add le_rfl _) le_rfl, + have A : N - 1 + 1 = N, from nat.succ_pred_eq_of_pos Npos, + have B : N - 1 + 1 < N + 1, by linarith, + have C : N - 1 < N + 1, by linarith, + rw [finset.sum_eq_sum_Ico_succ_bot C, finset.sum_eq_sum_Ico_succ_bot B, A, finset.Ico_self, + finset.sum_empty, add_zero, add_comm (edist _ _)], + exact edist_triangle _ _ _, + end + ... = ∑ j in finset.range (n + 1), edist (f (w (j + 1))) (f (w j)) : + begin + rw [finset.sum_Ico_consecutive, finset.sum_Ico_consecutive, finset.range_eq_Ico], + { exact zero_le _ }, + { linarith }, + { exact zero_le _ }, + { linarith } + end } +end + +/-- The variation of a function on the union of two sets `s` and `t`, with `s` to the left of `t`, +bounds the sum of the variations along `s` and `t`. -/ +lemma add_le_union (f : α → E) {s t : set α} (h : ∀ x ∈ s, ∀ y ∈ t, x ≤ y) : + evariation_on f s + evariation_on f t ≤ evariation_on f (s ∪ t) := +begin + by_cases hs : s = ∅, + { simp [hs] }, + haveI : nonempty {u // monotone u ∧ ∀ (i : ℕ), u i ∈ s}, + from nonempty_monotone_mem (ne_empty_iff_nonempty.1 hs), + by_cases ht : t = ∅, + { simp [ht] }, + haveI : nonempty {u // monotone u ∧ ∀ (i : ℕ), u i ∈ t}, + from nonempty_monotone_mem (ne_empty_iff_nonempty.1 ht), + refine ennreal.supr_add_supr_le _, + /- We start from two sequences `u` and `v` along `s` and `t` respectively, and we build a new + sequence `w` along `s ∪ t` by juxtaposing them. Its variation is larger than the sum of the + variations. -/ + rintros ⟨n, ⟨u, hu, us⟩⟩ ⟨m, ⟨v, hv, vt⟩⟩, + let w := λ i, if i ≤ n then u i else v (i - (n+1)), + have wst : ∀ i, w i ∈ s ∪ t, + { assume i, + by_cases hi : i ≤ n, + { simp [w, hi, us] }, + { simp [w, hi, vt] } }, + have hw : monotone w, + { assume i j hij, + dsimp only [w], + split_ifs, + { exact hu hij }, + { apply h _ (us _) _ (vt _) }, + { linarith }, + { apply hv (tsub_le_tsub hij le_rfl) } }, + calc ∑ i in finset.range n, edist (f (u (i + 1))) (f (u i)) + + ∑ (i : ℕ) in finset.range m, edist (f (v (i + 1))) (f (v i)) + = ∑ i in finset.range n, edist (f (w (i + 1))) (f (w i)) + + ∑ (i : ℕ) in finset.range m, edist (f (w ((n+1) + i + 1))) (f (w ((n+1) + i))) : + begin + dsimp only [w], + congr' 1, + { apply finset.sum_congr rfl (λ i hi, _), + simp only [finset.mem_range] at hi, + have : i + 1 ≤ n, by linarith, + simp [hi.le, this] }, + { apply finset.sum_congr rfl (λ i hi, _), + simp only [finset.mem_range] at hi, + have A : ¬(n + 1 + i + 1 ≤ n), by linarith, + have B : ¬(n + 1 + i ≤ n), by linarith, + have C : n + 1 + i - n = i + 1, + { rw tsub_eq_iff_eq_add_of_le, + { abel }, + { linarith } }, + simp only [A, B, C, nat.succ_sub_succ_eq_sub, if_false, add_tsub_cancel_left] } + end + ... = ∑ i in finset.range n, edist (f (w (i + 1))) (f (w i)) + + ∑ (i : ℕ) in finset.Ico (n+1) ((n+1)+m), edist (f (w (i + 1))) (f (w i)) : + begin + congr' 1, + rw finset.range_eq_Ico, + convert finset.sum_Ico_add (λ (i : ℕ), edist (f (w (i + 1))) (f (w i))) 0 m (n+1) using 3; + abel, + end + ... ≤ ∑ i in finset.range ((n+1) + m), edist (f (w (i + 1))) (f (w i)) : + begin + rw ← finset.sum_union, + { apply finset.sum_le_sum_of_subset _, + rintros i hi, + simp only [finset.mem_union, finset.mem_range, finset.mem_Ico] at hi ⊢, + cases hi, + { linarith }, + { exact hi.2 } }, + { apply finset.disjoint_left.2 (λ i hi h'i, _), + simp only [finset.mem_Ico, finset.mem_range] at hi h'i, + linarith [h'i.1] } + end + ... ≤ evariation_on f (s ∪ t) : sum_le f _ hw wst +end + +/-- If a set `s` is to the left of a set `t`, and both contain the boundary point `x`, then +the variation of `f` along `s ∪ t` is the sum of the variations. -/ +lemma union (f : α → E) {s t : set α} {x : α} (hs : is_greatest s x) (ht : is_least t x) : + evariation_on f (s ∪ t) = evariation_on f s + evariation_on f t := +begin + classical, + apply le_antisymm _ (evariation_on.add_le_union f (λ a ha b hb, le_trans (hs.2 ha) (ht.2 hb))), + apply supr_le _, + rintros ⟨n, ⟨u, hu, ust⟩⟩, + obtain ⟨v, m, hv, vst, xv, huv⟩ : ∃ (v : ℕ → α) (m : ℕ), monotone v ∧ (∀ i, v i ∈ s ∪ t) ∧ + x ∈ v '' (Iio m) ∧ ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) ≤ + ∑ j in finset.range m, edist (f (v (j+1))) (f (v j)), + from evariation_on.add_point f (mem_union_left t hs.1) u hu ust n, + obtain ⟨N, hN, Nx⟩ : ∃ N, N < m ∧ v N = x, from xv, + calc ∑ j in finset.range n, edist (f (u (j + 1))) (f (u j)) + ≤ ∑ j in finset.range m, edist (f (v (j + 1))) (f (v j)) : huv + ... = ∑ j in finset.Ico 0 N , edist (f (v (j + 1))) (f (v j)) + + ∑ j in finset.Ico N m , edist (f (v (j + 1))) (f (v j)) : + by rw [finset.range_eq_Ico, finset.sum_Ico_consecutive _ (zero_le _) hN.le] + ... ≤ evariation_on f s + evariation_on f t : + begin + refine add_le_add _ _, + { apply sum_le_of_monotone_on_Icc _ (hv.monotone_on _) (λ i hi, _), + rcases vst i with h|h, { exact h }, + have : v i = x, + { apply le_antisymm, + { rw ← Nx, exact hv hi.2 }, + { exact ht.2 h } }, + rw this, + exact hs.1 }, + { apply sum_le_of_monotone_on_Icc _ (hv.monotone_on _) (λ i hi, _), + rcases vst i with h|h, swap, { exact h }, + have : v i = x, + { apply le_antisymm, + { exact hs.2 h }, + { rw ← Nx, exact hv hi.1 } }, + rw this, + exact ht.1 } + end +end + +lemma Icc_add_Icc (f : α → E) {s : set α} {a b c : α} + (hab : a ≤ b) (hbc : b ≤ c) (hb : b ∈ s) : + evariation_on f (s ∩ Icc a b) + evariation_on f (s ∩ Icc b c) = evariation_on f (s ∩ Icc a c) := +begin + have A : is_greatest (s ∩ Icc a b) b := + ⟨⟨hb, hab, le_rfl⟩, (inter_subset_right _ _).trans (Icc_subset_Iic_self)⟩, + have B : is_least (s ∩ Icc b c) b := + ⟨⟨hb, le_rfl, hbc⟩, (inter_subset_right _ _).trans (Icc_subset_Ici_self)⟩, + rw [← evariation_on.union f A B, ← inter_union_distrib_left, Icc_union_Icc_eq_Icc hab hbc], +end + +end evariation_on + +/-! ## Monotone functions and bounded variation -/ + +lemma monotone_on.evariation_on_le {f : α → ℝ} {s : set α} (hf : monotone_on f s) {a b : α} + (as : a ∈ s) (bs : b ∈ s) : + evariation_on f (s ∩ Icc a b) ≤ ennreal.of_real (f b - f a) := +begin + apply supr_le _, + rintros ⟨n, ⟨u, hu, us⟩⟩, + calc + ∑ i in finset.range n, edist (f (u (i+1))) (f (u i)) + = ∑ i in finset.range n, ennreal.of_real (f (u (i + 1)) - f (u i)) : + begin + apply finset.sum_congr rfl (λ i hi, _), + simp only [finset.mem_range] at hi, + rw [edist_dist, real.dist_eq, abs_of_nonneg], + exact sub_nonneg_of_le (hf (us i).1 (us (i+1)).1 (hu (nat.le_succ _))), + end + ... = ennreal.of_real (∑ i in finset.range n, (f (u (i + 1)) - f (u i))) : + begin + rw [ennreal.of_real_sum_of_nonneg], + assume i hi, + exact sub_nonneg_of_le (hf (us i).1 (us (i+1)).1 (hu (nat.le_succ _))) + end + ... = ennreal.of_real (f (u n) - f (u 0)) : by rw finset.sum_range_sub (λ i, f (u i)) + ... ≤ ennreal.of_real (f b - f a) : + begin + apply ennreal.of_real_le_of_real, + exact sub_le_sub (hf (us n).1 bs (us n).2.2) (hf as (us 0).1 (us 0).2.1), + end +end + +lemma monotone_on.has_locally_bounded_variation_on {f : α → ℝ} {s : set α} (hf : monotone_on f s) : + has_locally_bounded_variation_on f s := +λ a b as bs, ((hf.evariation_on_le as bs).trans_lt ennreal.of_real_lt_top).ne + +/-- If a real valued function has bounded variation on a set, then it is a difference of monotone +functions there. -/ +lemma has_locally_bounded_variation_on.exists_monotone_on_sub_monotone_on {f : α → ℝ} {s : set α} + (h : has_locally_bounded_variation_on f s) : + ∃ (p q : α → ℝ), monotone_on p s ∧ monotone_on q s ∧ f = p - q := +begin + rcases eq_empty_or_nonempty s with rfl|hs, + { exact ⟨f, 0, subsingleton_empty.monotone_on _, subsingleton_empty.monotone_on _, + by simp only [tsub_zero]⟩ }, + rcases hs with ⟨c, cs⟩, + let p := λ x, if c ≤ x then (evariation_on f (s ∩ Icc c x)).to_real + else -(evariation_on f (s ∩ Icc x c)).to_real, + have hp : monotone_on p s, + { assume x xs y ys hxy, + dsimp only [p], + split_ifs with hcx hcy hcy, + { have : evariation_on f (s ∩ Icc c x) + evariation_on f (s ∩ Icc x y) + = evariation_on f (s ∩ Icc c y), from evariation_on.Icc_add_Icc f hcx hxy xs, + rw [← this, ennreal.to_real_add (h c x cs xs) (h x y xs ys)], + exact le_add_of_le_of_nonneg le_rfl ennreal.to_real_nonneg }, + { exact (lt_irrefl _ ((not_le.1 hcy).trans_le (hcx.trans hxy))).elim }, + { exact (neg_nonpos.2 ennreal.to_real_nonneg).trans ennreal.to_real_nonneg }, + { simp only [neg_le_neg_iff], + have : evariation_on f (s ∩ Icc x y) + evariation_on f (s ∩ Icc y c) + = evariation_on f (s ∩ Icc x c), from evariation_on.Icc_add_Icc f hxy (not_le.1 hcy).le ys, + rw [← this, ennreal.to_real_add (h x y xs ys) (h y c ys cs), add_comm], + exact le_add_of_le_of_nonneg le_rfl ennreal.to_real_nonneg } }, + have hq : monotone_on (λ x, p x - f x) s, + { assume x xs y ys hxy, + dsimp only [p], + split_ifs with hcx hcy hcy, + { have : evariation_on f (s ∩ Icc c x) + evariation_on f (s ∩ Icc x y) + = evariation_on f (s ∩ Icc c y), from evariation_on.Icc_add_Icc f hcx hxy xs, + rw [← this, ennreal.to_real_add (h c x cs xs) (h x y xs ys)], + suffices : f y - f x ≤ (evariation_on f (s ∩ Icc x y)).to_real, by linarith, + exact (h x y xs ys).sub_le ⟨ys, hxy, le_rfl⟩ ⟨xs, le_rfl, hxy⟩ }, + { exact (lt_irrefl _ ((not_le.1 hcy).trans_le (hcx.trans hxy))).elim }, + { suffices : f y - f x ≤ (evariation_on f (s ∩ Icc x c)).to_real + + (evariation_on f (s ∩ Icc c y)).to_real, by linarith, + rw [← ennreal.to_real_add (h x c xs cs) (h c y cs ys), + evariation_on.Icc_add_Icc f (not_le.1 hcx).le hcy cs], + exact (h x y xs ys).sub_le ⟨ys, hxy, le_rfl⟩ ⟨xs, le_rfl, hxy⟩ }, + { have : evariation_on f (s ∩ Icc x y) + evariation_on f (s ∩ Icc y c) + = evariation_on f (s ∩ Icc x c), from evariation_on.Icc_add_Icc f hxy (not_le.1 hcy).le ys, + rw [← this, ennreal.to_real_add (h x y xs ys) (h y c ys cs)], + suffices : f y - f x ≤ (evariation_on f (s ∩ Icc x y)).to_real, by linarith, + exact (h x y xs ys).sub_le ⟨ys, hxy, le_rfl⟩ ⟨xs, le_rfl, hxy⟩ } }, + refine ⟨p, λ x, p x - f x, hp, hq, _⟩, + ext x, + dsimp, + abel, +end + + +/-! ## Lipschitz functions and bounded variation -/ + +lemma lipschitz_on_with.comp_evariation_on_le {f : E → F} {C : ℝ≥0} {t : set E} + (h : lipschitz_on_with C f t) {g : α → E} {s : set α} (hg : maps_to g s t) : + evariation_on (f ∘ g) s ≤ C * evariation_on g s := +begin + apply supr_le _, + rintros ⟨n, ⟨u, hu, us⟩⟩, + calc + ∑ i in finset.range n, edist (f (g (u (i+1)))) (f (g (u i))) + ≤ ∑ i in finset.range n, C * edist (g (u (i+1))) (g (u i)) : + finset.sum_le_sum (λ i hi, h (hg (us _)) (hg (us _))) + ... = C * ∑ i in finset.range n, edist (g (u (i+1))) (g (u i)) : by rw finset.mul_sum + ... ≤ C * evariation_on g s : mul_le_mul_left' (evariation_on.sum_le _ _ hu us) _ +end + +lemma lipschitz_on_with.comp_has_bounded_variation_on {f : E → F} {C : ℝ≥0} {t : set E} + (hf : lipschitz_on_with C f t) {g : α → E} {s : set α} (hg : maps_to g s t) + (h : has_bounded_variation_on g s) : + has_bounded_variation_on (f ∘ g) s := +begin + dsimp [has_bounded_variation_on] at h, + apply ne_of_lt, + apply (hf.comp_evariation_on_le hg).trans_lt, + simp [lt_top_iff_ne_top, h], +end + +lemma lipschitz_on_with.comp_has_locally_bounded_variation_on {f : E → F} {C : ℝ≥0} {t : set E} + (hf : lipschitz_on_with C f t) {g : α → E} {s : set α} (hg : maps_to g s t) + (h : has_locally_bounded_variation_on g s) : + has_locally_bounded_variation_on (f ∘ g) s := +λ x y xs ys, hf.comp_has_bounded_variation_on (hg.mono_left (inter_subset_left _ _)) (h x y xs ys) + +lemma lipschitz_with.comp_has_bounded_variation_on {f : E → F} {C : ℝ≥0} + (hf : lipschitz_with C f) {g : α → E} {s : set α} (h : has_bounded_variation_on g s) : + has_bounded_variation_on (f ∘ g) s := +(hf.lipschitz_on_with univ).comp_has_bounded_variation_on (maps_to_univ _ _) h + +lemma lipschitz_with.comp_has_locally_bounded_variation_on {f : E → F} {C : ℝ≥0} + (hf : lipschitz_with C f) {g : α → E} {s : set α} (h : has_locally_bounded_variation_on g s) : + has_locally_bounded_variation_on (f ∘ g) s := +(hf.lipschitz_on_with univ).comp_has_locally_bounded_variation_on (maps_to_univ _ _) h + +lemma lipschitz_on_with.has_locally_bounded_variation_on {f : ℝ → E} {C : ℝ≥0} {s : set ℝ} + (hf : lipschitz_on_with C f s) : has_locally_bounded_variation_on f s := +hf.comp_has_locally_bounded_variation_on (maps_to_id _) + (@monotone_on_id ℝ _ s).has_locally_bounded_variation_on + +lemma lipschitz_with.has_locally_bounded_variation_on {f : ℝ → E} {C : ℝ≥0} + (hf : lipschitz_with C f) (s : set ℝ) : has_locally_bounded_variation_on f s := +(hf.lipschitz_on_with s).has_locally_bounded_variation_on From 476f28f8481aea7227a6acc2888b76b4d5188dc3 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Tue, 11 Oct 2022 03:45:37 +0000 Subject: [PATCH 669/828] feat(analysis/analytic/uniqueness): two analytic functions that coincide locally coincide globally (#16723) Higher dimensional version of #16489. --- src/analysis/analytic/isolated_zeros.lean | 62 ++++++------- src/analysis/analytic/uniqueness.lean | 106 ++++++++++++++++++++++ src/topology/connected.lean | 16 ++++ 3 files changed, 152 insertions(+), 32 deletions(-) create mode 100644 src/analysis/analytic/uniqueness.lean diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean index d60bf934f46f4..0a4d40c6a4363 100644 --- a/src/analysis/analytic/isolated_zeros.lean +++ b/src/analysis/analytic/isolated_zeros.lean @@ -9,6 +9,7 @@ import analysis.calculus.fderiv_analytic import analysis.calculus.formal_multilinear_series import analysis.complex.basic import topology.algebra.infinite_sum +import analysis.analytic.uniqueness /-! # Principle of isolated zeros @@ -149,46 +150,43 @@ end analytic_at namespace analytic_on -variables {U : set 𝕜} {w : 𝕜} +variables {U : set 𝕜} -theorem eq_on_of_preconnected_of_frequently_eq (hf : analytic_on 𝕜 f U) (hU : is_preconnected U) - (hw : w ∈ U) (hfw : ∃ᶠ z in 𝓝[≠] w, f z = 0) : +/-- The *principle of isolated zeros* for an analytic function, global version: if a function is +analytic on a connected set `U` and vanishes in arbitrary neighborhoods of a point `z₀ ∈ U`, then +it is identically zero in `U`. +For higher-dimensional versions requiring that the function vanishes in a neighborhood of `z₀`, +see `eq_on_zero_of_preconnected_of_eventually_eq_zero`. -/ +theorem eq_on_zero_of_preconnected_of_frequently_eq_zero + (hf : analytic_on 𝕜 f U) (hU : is_preconnected U) + (h₀ : z₀ ∈ U) (hfw : ∃ᶠ z in 𝓝[≠] z₀, f z = 0) : eq_on f 0 U := -begin - by_contra, - simp only [eq_on, not_forall] at h, - obtain ⟨x, hx1, hx2⟩ := h, - let u := {z | f =ᶠ[𝓝 z] 0}, - have hu : is_open u := is_open_set_of_eventually_nhds, - have hu' : (U ∩ u).nonempty := ⟨w, hw, (hf w hw).frequently_zero_iff_eventually_zero.mp hfw⟩, - let v := {z | ∀ᶠ w in 𝓝[≠] z, f w ≠ 0}, - have hv : is_open v := by apply is_open_set_of_eventually_nhds_within, - have hv' : (U ∩ v).nonempty, - from ⟨x, hx1, ((hf x hx1).continuous_at.eventually_ne hx2).filter_mono nhds_within_le_nhds⟩, - have huv : U ⊆ u ∪ v := λ z hz, (hf z hz).eventually_eq_zero_or_eventually_ne_zero, - have huv' : u ∩ v = ∅, - by { ext z, - simp only [mem_inter_iff, mem_empty_iff_false, iff_false, not_and], - exact λ h, (h.filter_mono nhds_within_le_nhds).frequently }, - simpa [huv'] using hU u v hu hv huv hu' hv' -end +hf.eq_on_zero_of_preconnected_of_eventually_eq_zero hU h₀ + ((hf z₀ h₀).frequently_zero_iff_eventually_zero.1 hfw) -theorem eq_on_of_preconnected_of_mem_closure (hf : analytic_on 𝕜 f U) (hU : is_preconnected U) - (hw : w ∈ U) (hfw : w ∈ closure ({z | f z = 0} \ {w})) : +theorem eq_on_zero_of_preconnected_of_mem_closure (hf : analytic_on 𝕜 f U) (hU : is_preconnected U) + (h₀ : z₀ ∈ U) (hfz₀ : z₀ ∈ closure ({z | f z = 0} \ {z₀})) : eq_on f 0 U := -hf.eq_on_of_preconnected_of_frequently_eq hU hw (mem_closure_ne_iff_frequently_within.mp hfw) - -theorem eq_on_of_preconnected_of_frequently_eq' (hf : analytic_on 𝕜 f U) (hg : analytic_on 𝕜 g U) - (hU : is_preconnected U) (hw : w ∈ U) (hfg : ∃ᶠ z in 𝓝[≠] w, f z = g z) : +hf.eq_on_zero_of_preconnected_of_frequently_eq_zero hU h₀ + (mem_closure_ne_iff_frequently_within.mp hfz₀) + +/-- The *identity principle* for analytic functions, global version: if two functions are +analytic on a connected set `U` and coincide at points which accumulate to a point `z₀ ∈ U`, then +they coincide globally in `U`. +For higher-dimensional versions requiring that the functions coincide in a neighborhood of `z₀`, +see `eq_on_of_preconnected_of_eventually_eq`. -/ +theorem eq_on_of_preconnected_of_frequently_eq (hf : analytic_on 𝕜 f U) (hg : analytic_on 𝕜 g U) + (hU : is_preconnected U) (h₀ : z₀ ∈ U) (hfg : ∃ᶠ z in 𝓝[≠] z₀, f z = g z) : eq_on f g U := begin - have hfg' : ∃ᶠ z in 𝓝[≠] w, (f - g) z = 0 := hfg.mono (λ z h, by rw [pi.sub_apply, h, sub_self]), - simpa [sub_eq_zero] using λ z hz, (hf.sub hg).eq_on_of_preconnected_of_frequently_eq hU hw hfg' hz + have hfg' : ∃ᶠ z in 𝓝[≠] z₀, (f - g) z = 0 := hfg.mono (λ z h, by rw [pi.sub_apply, h, sub_self]), + simpa [sub_eq_zero] using + λ z hz, (hf.sub hg).eq_on_zero_of_preconnected_of_frequently_eq_zero hU h₀ hfg' hz end -theorem eq_on_of_preconnected_of_mem_closure' (hf : analytic_on 𝕜 f U) (hg : analytic_on 𝕜 g U) - (hU : is_preconnected U) (hw : w ∈ U) (hfw : w ∈ closure ({z | f z = g z} \ {w})) : +theorem eq_on_of_preconnected_of_mem_closure (hf : analytic_on 𝕜 f U) (hg : analytic_on 𝕜 g U) + (hU : is_preconnected U) (h₀ : z₀ ∈ U) (hfg : z₀ ∈ closure ({z | f z = g z} \ {z₀})) : eq_on f g U := -hf.eq_on_of_preconnected_of_frequently_eq' hg hU hw (mem_closure_ne_iff_frequently_within.mp hfw) +hf.eq_on_of_preconnected_of_frequently_eq hg hU h₀ (mem_closure_ne_iff_frequently_within.mp hfg) end analytic_on diff --git a/src/analysis/analytic/uniqueness.lean b/src/analysis/analytic/uniqueness.lean new file mode 100644 index 0000000000000..dce9486e217ba --- /dev/null +++ b/src/analysis/analytic/uniqueness.lean @@ -0,0 +1,106 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import analysis.analytic.linear +import analysis.analytic.composition +import analysis.normed_space.completion + +/-! +# Uniqueness principle for analytic functions + +We show that two analytic functions which coincide around a point coincide on whole connected sets, +in `analytic_on.eq_on_of_preconnected_of_eventually_eq`. +-/ + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] + +open set +open_locale topological_space ennreal + +namespace analytic_on + +/-- If an analytic function vanishes around a point, then it is uniformly zero along +a connected set. Superseded by `eq_on_zero_of_preconnected_of_locally_zero` which does not assume +completeness of the target space. -/ +theorem eq_on_zero_of_preconnected_of_eventually_eq_zero_aux [complete_space F] + {f : E → F} {U : set E} (hf : analytic_on 𝕜 f U) (hU : is_preconnected U) {z₀ : E} + (h₀ : z₀ ∈ U) (hfz₀ : f =ᶠ[𝓝 z₀] 0) : eq_on f 0 U := +begin + /- Let `u` be the set of points around which `f` vanishes. It is clearly open. We have to show + that its limit points in `U` still belong to it, from which the inclusion `U ⊆ u` will follow + by connectedness. -/ + let u := {x | f =ᶠ[𝓝 x] 0}, + suffices main : closure u ∩ U ⊆ u, + { have Uu : U ⊆ u, from + hU.subset_of_closure_inter_subset is_open_set_of_eventually_nhds ⟨z₀, h₀, hfz₀⟩ main, + assume z hz, + simpa using mem_of_mem_nhds (Uu hz) }, + /- Take a limit point `x`, then a ball `B (x, r)` on which it has a power series expansion, and + then `y ∈ B (x, r/2) ∩ u`. Then `f` has a power series expansion on `B (y, r/2)` as it is + contained in `B (x, r)`. All the coefficients in this series expansion vanish, as `f` is zero on a + neighborhood of `y`. Therefore, `f` is zero on `B (y, r/2)`. As this ball contains `x`, it follows + that `f` vanishes on a neighborhood of `x`, proving the claim. -/ + rintros x ⟨xu, xU⟩, + rcases hf x xU with ⟨p, r, hp⟩, + obtain ⟨y, yu, hxy⟩ : ∃ y ∈ u, edist x y < r / 2, + from emetric.mem_closure_iff.1 xu (r / 2) (ennreal.half_pos hp.r_pos.ne'), + let q := p.change_origin (y - x), + have has_series : has_fpower_series_on_ball f q y (r / 2), + { have A : (∥y - x∥₊ : ℝ≥0∞) < r / 2, by rwa [edist_comm, edist_eq_coe_nnnorm_sub] at hxy, + have := hp.change_origin (A.trans_le ennreal.half_le_self), + simp only [add_sub_cancel'_right] at this, + apply this.mono (ennreal.half_pos hp.r_pos.ne'), + apply ennreal.le_sub_of_add_le_left ennreal.coe_ne_top, + apply (add_le_add (A.le) (le_refl (r / 2))).trans (le_of_eq _), + exact ennreal.add_halves _ }, + have M : emetric.ball y (r / 2) ∈ 𝓝 x, from emetric.is_open_ball.mem_nhds hxy, + filter_upwards [M] with z hz, + have A : has_sum (λ (n : ℕ), q n (λ (i : fin n), z - y)) (f z) := has_series.has_sum_sub hz, + have B : has_sum (λ (n : ℕ), q n (λ (i : fin n), z - y)) (0), + { have : has_fpower_series_at 0 q y, from has_series.has_fpower_series_at.congr yu, + convert has_sum_zero, + ext n, + exact this.apply_eq_zero n _ }, + exact has_sum.unique A B +end + +/-- The *identity principle* for analytic functions: If an analytic function vanishes in a whole +neighborhood of a point `z₀`, then it is uniformly zero along a connected set. For a one-dimensional +version assuming only that the function vanishes at some points arbitrarily close to `z₀`, see +`eq_on_zero_of_preconnected_of_frequently_eq_zero`. -/ +theorem eq_on_zero_of_preconnected_of_eventually_eq_zero + {f : E → F} {U : set E} (hf : analytic_on 𝕜 f U) (hU : is_preconnected U) {z₀ : E} + (h₀ : z₀ ∈ U) (hfz₀ : f =ᶠ[𝓝 z₀] 0) : + eq_on f 0 U := +begin + let F' := uniform_space.completion F, + set e : F →L[𝕜] F' := uniform_space.completion.to_complL, + have : analytic_on 𝕜 (e ∘ f) U := λ x hx, (e.analytic_at _).comp (hf x hx), + have A : eq_on (e ∘ f) 0 U, + { apply eq_on_zero_of_preconnected_of_eventually_eq_zero_aux this hU h₀, + filter_upwards [hfz₀] with x hx, + simp only [hx, function.comp_app, pi.zero_apply, map_zero] }, + assume z hz, + have : e (f z) = e 0, by simpa only using A hz, + exact uniform_space.completion.coe_injective F this, +end + +/-- The *identity principle* for analytic functions: If two analytic function coincide in a whole +neighborhood of a point `z₀`, then they coincide globally along a connected set. +For a one-dimensional version assuming only that the functions coincide at some points +arbitrarily close to `z₀`, see `eq_on_of_preconnected_of_frequently_eq`. -/ +theorem eq_on_of_preconnected_of_eventually_eq + {f g : E → F} {U : set E} (hf : analytic_on 𝕜 f U) (hg : analytic_on 𝕜 g U) + (hU : is_preconnected U) {z₀ : E} (h₀ : z₀ ∈ U) (hfg : f =ᶠ[𝓝 z₀] g) : + eq_on f g U := +begin + have hfg' : (f - g) =ᶠ[𝓝 z₀] 0 := hfg.mono (λ z h, by simp [h]), + simpa [sub_eq_zero] using + λ z hz, (hf.sub hg).eq_on_zero_of_preconnected_of_eventually_eq_zero hU h₀ hfg' hz, +end + +end analytic_on diff --git a/src/topology/connected.lean b/src/topology/connected.lean index 509bd405dfcde..a7ec33f07d78c 100644 --- a/src/topology/connected.lean +++ b/src/topology/connected.lean @@ -441,6 +441,22 @@ lemma is_preconnected.subset_right_of_subset_union (hu : is_open u) (hv : is_ope s ⊆ v := hs.subset_left_of_subset_union hv hu huv.symm (union_comm u v ▸ hsuv) hsv +/-- If a preconnected set `s` intersects an open set `u`, and limit points of `u` inside `s` are +contained in `u`, then the whole set `s` is contained in `u`. -/ +lemma is_preconnected.subset_of_closure_inter_subset (hs : is_preconnected s) + (hu : is_open u) (h'u : (s ∩ u).nonempty) (h : closure u ∩ s ⊆ u) : s ⊆ u := +begin + have A : s ⊆ u ∪ (closure u)ᶜ, + { assume x hx, + by_cases xu : x ∈ u, + { exact or.inl xu }, + { right, + assume h'x, + exact xu (h (mem_inter h'x hx)) } }, + apply hs.subset_left_of_subset_union hu is_closed_closure.is_open_compl _ A h'u, + exact disjoint_compl_right.mono_right (compl_subset_compl.2 subset_closure), +end + theorem is_preconnected.prod [topological_space β] {s : set α} {t : set β} (hs : is_preconnected s) (ht : is_preconnected t) : is_preconnected (s ×ˢ t) := From 3348bed93f19048ce2a6055eee86ee570cd5e86e Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 11 Oct 2022 03:45:38 +0000 Subject: [PATCH 670/828] chore(tactic/linarith): remove unneeded code (#16791) Removes some unneeded lemmas (and consequently unneeded imports) in `tactic/linarith/lemmas.lean`, in preparation for porting. Co-authored-by: Scott Morrison --- src/algebra/order/ring.lean | 2 +- src/data/int/basic.lean | 1 - src/tactic/linarith/lemmas.lean | 30 ------------------------------ 3 files changed, 1 insertion(+), 32 deletions(-) diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 5fb15cbba7f82..8148b5a398eff 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -1490,7 +1490,7 @@ end end mul_zero_class -/-- `nontrivial α` is needed here as otherwise we have `1 * ⊤ = ⊤` but also `= 0 * ⊤ = 0`. -/ +/-- `nontrivial α` is needed here as otherwise we have `1 * ⊤ = ⊤` but also `0 * ⊤ = 0`. -/ instance [mul_zero_one_class α] [nontrivial α] : mul_zero_one_class (with_top α) := { mul := (*), one := 1, diff --git a/src/data/int/basic.lean b/src/data/int/basic.lean index 403a0ab879577..811960830e6bc 100644 --- a/src/data/int/basic.lean +++ b/src/data/int/basic.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ import data.nat.pow -import order.min_max import data.nat.cast import algebra.ring.regular diff --git a/src/tactic/linarith/lemmas.lean b/src/tactic/linarith/lemmas.lean index c5285ae9c6d32..83d7faf69ead9 100644 --- a/src/tactic/linarith/lemmas.lean +++ b/src/tactic/linarith/lemmas.lean @@ -5,8 +5,6 @@ Authors: Robert Y. Lewis -/ import algebra.order.ring -import data.int.basic -import tactic.norm_num /-! # Lemmas for `linarith` @@ -17,33 +15,6 @@ If you find yourself looking for a theorem here, you might be in the wrong place namespace linarith -lemma int.coe_nat_bit0 (n : ℕ) : (↑(bit0 n : ℕ) : ℤ) = bit0 (↑n : ℤ) := by simp [bit0] -lemma int.coe_nat_bit1 (n : ℕ) : (↑(bit1 n : ℕ) : ℤ) = bit1 (↑n : ℤ) := by simp [bit1, bit0] -lemma int.coe_nat_bit0_mul (n : ℕ) (x : ℕ) : (↑(bit0 n * x) : ℤ) = (↑(bit0 n) : ℤ) * (↑x : ℤ) := -by simp -lemma int.coe_nat_bit1_mul (n : ℕ) (x : ℕ) : (↑(bit1 n * x) : ℤ) = (↑(bit1 n) : ℤ) * (↑x : ℤ) := -by simp -lemma int.coe_nat_one_mul (x : ℕ) : (↑(1 * x) : ℤ) = 1 * (↑x : ℤ) := by simp -lemma int.coe_nat_zero_mul (x : ℕ) : (↑(0 * x) : ℤ) = 0 * (↑x : ℤ) := by simp -lemma int.coe_nat_mul_bit0 (n : ℕ) (x : ℕ) : (↑(x * bit0 n) : ℤ) = (↑x : ℤ) * (↑(bit0 n) : ℤ) := -by simp -lemma int.coe_nat_mul_bit1 (n : ℕ) (x : ℕ) : (↑(x * bit1 n) : ℤ) = (↑x : ℤ) * (↑(bit1 n) : ℤ) := -by simp -lemma int.coe_nat_mul_one (x : ℕ) : (↑(x * 1) : ℤ) = (↑x : ℤ) * 1 := by simp -lemma int.coe_nat_mul_zero (x : ℕ) : (↑(x * 0) : ℤ) = (↑x : ℤ) * 0 := by simp - -lemma nat_eq_subst {n1 n2 : ℕ} {z1 z2 : ℤ} (hn : n1 = n2) (h1 : ↑n1 = z1) (h2 : ↑n2 = z2) : - z1 = z2 := -by simpa [eq.symm h1, eq.symm h2, int.coe_nat_eq_coe_nat_iff] - -lemma nat_le_subst {n1 n2 : ℕ} {z1 z2 : ℤ} (hn : n1 ≤ n2) (h1 : ↑n1 = z1) (h2 : ↑n2 = z2) : - z1 ≤ z2 := -by simpa [eq.symm h1, eq.symm h2, int.coe_nat_le] - -lemma nat_lt_subst {n1 n2 : ℕ} {z1 z2 : ℤ} (hn : n1 < n2) (h1 : ↑n1 = z1) (h2 : ↑n2 = z2) : - z1 < z2 := -by simpa [eq.symm h1, eq.symm h2, int.coe_nat_lt] - lemma eq_of_eq_of_eq {α} [ordered_semiring α] {a b : α} (ha : a = 0) (hb : b = 0) : a + b = 0 := by simp * @@ -75,7 +46,6 @@ by simp * lemma eq_of_not_lt_of_not_gt {α} [linear_order α] (a b : α) (h1 : ¬ a < b) (h2 : ¬ b < a) : a = b := le_antisymm (le_of_not_gt h2) (le_of_not_gt h1) - -- used in the `nlinarith` normalization steps. The `_` argument is for uniformity. @[nolint unused_arguments] lemma mul_zero_eq {α} {R : α → α → Prop} [semiring α] {a b : α} (_ : R a 0) (h : b = 0) : From 419c83f6c89789ac582cf29c742c5ef856bdfd93 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 11 Oct 2022 03:45:39 +0000 Subject: [PATCH 671/828] feat(linear_algebra/affine_space/finite_dimensional): more `collinear` lemmas (#16806) Add a further batch of lemmas about collinear sets of points, in particular relating collinearity to any pair of distinct points in the set. --- .../affine_space/finite_dimensional.lean | 57 +++++++++++++++++++ 1 file changed, 57 insertions(+) diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index 35208c1d243d4..35c7a8d8a803d 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -442,6 +442,63 @@ lemma collinear_iff_not_affine_independent {p : fin 3 → P} : by rw [collinear_iff_finrank_le_one, finrank_vector_span_le_iff_not_affine_independent k p (fintype.card_fin 3)] +/-- If three points are not collinear, the first and second are different. -/ +lemma ne₁₂_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear k ({p₁, p₂, p₃} : set P)) : p₁ ≠ p₂ := +by { rintro rfl, simpa [collinear_pair] using h } + +/-- If three points are not collinear, the first and third are different. -/ +lemma ne₁₃_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear k ({p₁, p₂, p₃} : set P)) : p₁ ≠ p₃ := +by { rintro rfl, simpa [collinear_pair] using h } + +/-- If three points are not collinear, the second and third are different. -/ +lemma ne₂₃_of_not_collinear {p₁ p₂ p₃ : P} (h : ¬collinear k ({p₁, p₂, p₃} : set P)) : p₂ ≠ p₃ := +by { rintro rfl, simpa [collinear_pair] using h } + +/-- A point in a collinear set of points lies in the affine span of any two distinct points of +that set. -/ +lemma collinear.mem_affine_span_of_mem_of_ne {s : set P} (h : collinear k s) {p₁ p₂ p₃ : P} + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∈ s) (hp₁p₂ : p₁ ≠ p₂) : + p₃ ∈ affine_span k ({p₁, p₂} : set P) := +begin + rw collinear_iff_of_mem hp₁ at h, + rcases h with ⟨v, h⟩, + rcases h p₂ hp₂ with ⟨r₂, rfl⟩, + rcases h p₃ hp₃ with ⟨r₃, rfl⟩, + rw vadd_left_mem_affine_span_pair, + refine ⟨r₃ / r₂, _⟩, + have h₂ : r₂ ≠ 0, + { rintro rfl, + simpa using hp₁p₂ }, + simp [smul_smul, h₂] +end + +/-- The affine span of any two distinct points of a collinear set of points equals the affine +span of the whole set. -/ +lemma collinear.affine_span_eq_of_ne {s : set P} (h : collinear k s) {p₁ p₂ : P} + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) (hp₁p₂ : p₁ ≠ p₂) : + affine_span k ({p₁, p₂} : set P) = affine_span k s := +le_antisymm (affine_span_mono _ + (set.insert_subset.2 ⟨hp₁, set.singleton_subset_iff.2 hp₂⟩)) + (affine_span_le.2 (λ p hp, h.mem_affine_span_of_mem_of_ne hp₁ hp₂ hp hp₁p₂)) + +/-- Given a collinear set of points, and two distinct points `p₂` and `p₃` in it, a point `p₁` is +collinear with the set if and only if it is collinear with `p₂` and `p₃`. -/ +lemma collinear.collinear_insert_iff_of_ne {s : set P} (h : collinear k s) {p₁ p₂ p₃ : P} + (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∈ s) (hp₂p₃ : p₂ ≠ p₃) : + collinear k (insert p₁ s) ↔ collinear k ({p₁, p₂, p₃} : set P) := +begin + have hv : vector_span k (insert p₁ s) = vector_span k ({p₁, p₂, p₃} : set P), + { conv_lhs { rw [←direction_affine_span, ←affine_span_insert_affine_span] }, + conv_rhs { rw [←direction_affine_span, ←affine_span_insert_affine_span] }, + rw h.affine_span_eq_of_ne hp₂ hp₃ hp₂p₃ }, + rw [collinear, collinear, hv] +end + +/-- Adding a point in the affine span of a set does not change whether that set is collinear. -/ +lemma collinear_insert_iff_of_mem_affine_span {s : set P} {p : P} (h : p ∈ affine_span k s) : + collinear k (insert p s) ↔ collinear k s := +by rw [collinear, collinear, vector_span_insert_eq_vector_span h] + end affine_space' section field From c8243a6a5c3273e983606a70bf8913911c5cea3b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 11 Oct 2022 03:45:40 +0000 Subject: [PATCH 672/828] feat(group_theory/submonoid/basic): add `monoid_hom.of_mclosure_eq_top_left` (#16809) * rename `monoid_hom.of_mdense` to `monoid_hom.of_mclosure_eq_top_right`; * add `monoid_hom.of_mclosure_eq_top_left`. --- src/group_theory/submonoid/basic.lean | 50 ++++++++++++++++++--------- 1 file changed, 33 insertions(+), 17 deletions(-) diff --git a/src/group_theory/submonoid/basic.lean b/src/group_theory/submonoid/basic.lean index 442244b297116..970bd72c486f6 100644 --- a/src/group_theory/submonoid/basic.lean +++ b/src/group_theory/submonoid/basic.lean @@ -13,7 +13,7 @@ This file defines bundled multiplicative and additive submonoids. We also define a `complete_lattice` structure on `submonoid`s, define the closure of a set as the minimal submonoid that includes this set, and prove a few results about extending properties from a dense set (i.e. a set with `closure s = ⊤`) to the whole monoid, see `submonoid.dense_induction` and -`monoid_hom.of_mdense`. +`monoid_hom.of_mclosure_eq_top_left`/`monoid_hom.of_mclosure_eq_top_right`. ## Main definitions @@ -30,10 +30,10 @@ definition in the `add_submonoid` namespace. * `submonoid.gi` : `closure : set M → submonoid M` and coercion `coe : submonoid M → set M` form a `galois_insertion`; * `monoid_hom.eq_mlocus`: the submonoid of elements `x : M` such that `f x = g x`; -* `monoid_hom.of_mdense`: if a map `f : M → N` between two monoids satisfies `f 1 = 1` and - `f (x * y) = f x * f y` for `y` from some dense set `s`, then `f` is a monoid homomorphism. - E.g., if `f : ℕ → M` satisfies `f 0 = 0` and `f (x + 1) = f x + f 1`, then `f` is an additive - monoid homomorphism. +* `monoid_hom.of_mclosure_eq_top_right`: if a map `f : M → N` between two monoids satisfies + `f 1 = 1` and `f (x * y) = f x * f y` for `y` from some dense set `s`, then `f` is a monoid + homomorphism. E.g., if `f : ℕ → M` satisfies `f 0 = 0` and `f (x + 1) = f x + f 1`, then `f` is + an additive monoid homomorphism. ## Implementation notes @@ -498,24 +498,40 @@ namespace monoid_hom open submonoid /-- Let `s` be a subset of a monoid `M` such that the closure of `s` is the whole monoid. -Then `monoid_hom.of_mdense` defines a monoid homomorphism from `M` asking for a proof -of `f (x * y) = f x * f y` only for `y ∈ s`. -/ -@[to_additive] -def of_mdense {M N} [monoid M] [monoid N] {s : set M} (f : M → N) (hs : closure s = ⊤) - (h1 : f 1 = 1) (hmul : ∀ x (y ∈ s), f (x * y) = f x * f y) : +Then `monoid_hom.of_mclosure_eq_top_left` defines a monoid homomorphism from `M` asking for +a proof of `f (x * y) = f x * f y` only for `x ∈ s`. -/ +@[to_additive "/-- Let `s` be a subset of an additive monoid `M` such that the closure of `s` is +the whole monoid. Then `add_monoid_hom.of_mclosure_eq_top_left` defines an additive monoid +homomorphism from `M` asking for a proof of `f (x + y) = f x + f y` only for `x ∈ s`. -/"] +def of_mclosure_eq_top_left {M N} [monoid M] [monoid N] {s : set M} (f : M → N) + (hs : closure s = ⊤) (h1 : f 1 = 1) (hmul : ∀ (x ∈ s) y, f (x * y) = f x * f y) : + M →* N := +{ to_fun := f, + map_one' := h1, + map_mul' := λ x, dense_induction x hs hmul (λ y, by rw [one_mul, h1, one_mul]) $ λ a b ha hb y, + by rw [mul_assoc, ha, ha, hb, mul_assoc] } + +@[simp, norm_cast, to_additive] lemma coe_of_mclosure_eq_top_left (f : M → N) (hs : closure s = ⊤) + (h1 hmul) : ⇑(of_mclosure_eq_top_left f hs h1 hmul) = f := +rfl + +/-- Let `s` be a subset of a monoid `M` such that the closure of `s` is the whole monoid. +Then `monoid_hom.of_mclosure_eq_top_right` defines a monoid homomorphism from `M` asking for +a proof of `f (x * y) = f x * f y` only for `y ∈ s`. -/ +@[to_additive "/-- Let `s` be a subset of an additive monoid `M` such that the closure of `s` is +the whole monoid. Then `add_monoid_hom.of_mclosure_eq_top_right` defines an additive monoid +homomorphism from `M` asking for a proof of `f (x + y) = f x + f y` only for `y ∈ s`. -/"] +def of_mclosure_eq_top_right {M N} [monoid M] [monoid N] {s : set M} (f : M → N) + (hs : closure s = ⊤) (h1 : f 1 = 1) (hmul : ∀ x (y ∈ s), f (x * y) = f x * f y) : M →* N := { to_fun := f, map_one' := h1, map_mul' := λ x y, dense_induction y hs (λ y hy x, hmul x y hy) (by simp [h1]) (λ y₁ y₂ h₁ h₂ x, by simp only [← mul_assoc, h₁, h₂]) x } -/-- Let `s` be a subset of an additive monoid `M` such that the closure of `s` is the whole monoid. -Then `add_monoid_hom.of_mdense` defines an additive monoid homomorphism from `M` asking for a proof -of `f (x + y) = f x + f y` only for `y ∈ s`. -/ -add_decl_doc add_monoid_hom.of_mdense - -@[simp, norm_cast, to_additive] lemma coe_of_mdense (f : M → N) (hs : closure s = ⊤) (h1 hmul) : - ⇑(of_mdense f hs h1 hmul) = f := rfl +@[simp, norm_cast, to_additive] lemma coe_of_mclosure_eq_top_right (f : M → N) (hs : closure s = ⊤) + (h1 hmul) : ⇑(of_mclosure_eq_top_right f hs h1 hmul) = f := +rfl end monoid_hom From 069c3d1ea82d52abc801b5ae6b506cdabb78b923 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 11 Oct 2022 03:45:41 +0000 Subject: [PATCH 673/828] feat(linear_algebra/affine_space/independent): affine independence of all but one point (#16815) Add the lemma that, in an affine space for a module over a division ring, if all but one point of a family are affinely independent, and that point does not lie in the affine span of the other points, the family is affinely independent. Thus, deduce lemmas that three points are affinely independent if two are distinct and the third does not lie in an affine subspace containing those two. --- .../affine_space/independent.lean | 95 +++++++++++++++++++ 1 file changed, 95 insertions(+) diff --git a/src/linear_algebra/affine_space/independent.lean b/src/linear_algebra/affine_space/independent.lean index 94e3251552e58..fa328b91b4a9f 100644 --- a/src/linear_algebra/affine_space/independent.lean +++ b/src/linear_algebra/affine_space/independent.lean @@ -559,6 +559,101 @@ begin exact linear_independent_unique _ hz end +variables {k V P} + +/-- If all but one point of a family are affinely independent, and that point does not lie in +the affine span of that family, the family is affinely independent. -/ +lemma affine_independent.affine_independent_of_not_mem_span {p : ι → P} {i : ι} + (ha : affine_independent k (λ x : {y // y ≠ i}, p x)) + (hi : p i ∉ affine_span k (p '' {x | x ≠ i})) : affine_independent k p := +begin + classical, + intros s w hw hs, + let s' : finset {y // y ≠ i} := s.subtype (≠ i), + let p' : {y // y ≠ i} → P := λ x, p x, + by_cases his : i ∈ s ∧ w i ≠ 0, + { refine false.elim (hi _), + let wm : ι → k := -(w i)⁻¹ • w, + have hms : s.weighted_vsub p wm = (0 : V), { simp [wm, hs] }, + have hwm : ∑ i in s, wm i = 0, { simp [wm, ←finset.mul_sum, hw] }, + have hwmi : wm i = -1, { simp [wm, his.2] }, + let w' : {y // y ≠ i} → k := λ x, wm x, + have hw' : ∑ x in s', w' x = 1, + { simp_rw [w', finset.sum_subtype_eq_sum_filter], + rw ←s.sum_filter_add_sum_filter_not (≠ i) at hwm, + simp_rw [not_not, finset.filter_eq', if_pos his.1, finset.sum_singleton, ←wm, hwmi, + ←sub_eq_add_neg, sub_eq_zero] at hwm, + exact hwm }, + rw [←s.affine_combination_eq_of_weighted_vsub_eq_zero_of_eq_neg_one hms his.1 hwmi, + ←(subtype.range_coe : _ = {x | x ≠ i}), ←set.range_comp, + ←s.affine_combination_subtype_eq_filter], + exact affine_combination_mem_affine_span hw' p' }, + { rw [not_and_distrib, not_not] at his, + let w' : {y // y ≠ i} → k := λ x, w x, + have hw' : ∑ x in s', w' x = 0, + { simp_rw [finset.sum_subtype_eq_sum_filter], + rw [finset.sum_filter_of_ne, hw], + rintro x hxs hwx rfl, + exact hwx (his.neg_resolve_left hxs) }, + have hs' : s'.weighted_vsub p' w' = (0 : V), + { simp_rw finset.weighted_vsub_subtype_eq_filter, + rw [finset.weighted_vsub_filter_of_ne, hs], + rintro x hxs hwx rfl, + exact hwx (his.neg_resolve_left hxs) }, + intros j hj, + by_cases hji : j = i, + { rw hji at hj, + exact hji.symm ▸ (his.neg_resolve_left hj) }, + { exact ha s' w' hw' hs' ⟨j, hji⟩ (finset.mem_subtype.2 hj) } } +end + +/-- If distinct points `p₁` and `p₂` lie in `s` but `p₃` does not, the three points are affinely +independent. -/ +lemma affine_independent_of_ne_of_mem_of_mem_of_not_mem {s : affine_subspace k P} {p₁ p₂ p₃ : P} + (hp₁p₂ : p₁ ≠ p₂) (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∉ s) : + affine_independent k ![p₁, p₂, p₃] := +begin + have ha : affine_independent k (λ x : {x : fin 3 // x ≠ 2}, ![p₁, p₂, p₃] x), + { rw ←affine_independent_equiv ((fin_succ_above_equiv (2 : fin 3)).to_equiv), + convert affine_independent_of_ne k hp₁p₂, + ext x, + fin_cases x; + refl }, + refine ha.affine_independent_of_not_mem_span _, + intro h, + refine hp₃ ((affine_subspace.le_def' _ s).1 _ p₃ h), + simp_rw [affine_span_le, set.image_subset_iff, set.subset_def, set.mem_preimage], + intro x, + fin_cases x; + simp [hp₁, hp₂] +end + +/-- If distinct points `p₁` and `p₃` lie in `s` but `p₂` does not, the three points are affinely +independent. -/ +lemma affine_independent_of_ne_of_mem_of_not_mem_of_mem {s : affine_subspace k P} {p₁ p₂ p₃ : P} + (hp₁p₃ : p₁ ≠ p₃) (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∉ s) (hp₃ : p₃ ∈ s) : + affine_independent k ![p₁, p₂, p₃] := +begin + rw ←affine_independent_equiv (equiv.swap (1 : fin 3) 2), + convert affine_independent_of_ne_of_mem_of_mem_of_not_mem hp₁p₃ hp₁ hp₃ hp₂ using 1, + ext x, + fin_cases x; + refl +end + +/-- If distinct points `p₂` and `p₃` lie in `s` but `p₁` does not, the three points are affinely +independent. -/ +lemma affine_independent_of_ne_of_not_mem_of_mem_of_mem {s : affine_subspace k P} {p₁ p₂ p₃ : P} + (hp₂p₃ : p₂ ≠ p₃) (hp₁ : p₁ ∉ s) (hp₂ : p₂ ∈ s) (hp₃ : p₃ ∈ s) : + affine_independent k ![p₁, p₂, p₃] := +begin + rw ←affine_independent_equiv (equiv.swap (0 : fin 3) 2), + convert affine_independent_of_ne_of_mem_of_mem_of_not_mem hp₂p₃.symm hp₃ hp₂ hp₁ using 1, + ext x, + fin_cases x; + refl +end + end division_ring namespace affine From abed0373a51594b6126f86306a26f08a5d402b73 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 11 Oct 2022 03:45:42 +0000 Subject: [PATCH 674/828] feat(geometry/euclidean/oriented_angle): collinearity and affine independence (#16822) Add two lemmas relating the oriented angle between three points to those points being collinear or affinely independent. --- src/geometry/euclidean/oriented_angle.lean | 23 ++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index f98d366ec9e62..37273cb70b561 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -1845,6 +1845,29 @@ lemma oangle_eq_zero_iff_oangle_rev_eq_zero {p₁ p₂ p₃ : P} : ∡ p₁ p₂ lemma oangle_eq_pi_iff_oangle_rev_eq_pi {p₁ p₂ p₃ : P} : ∡ p₁ p₂ p₃ = π ↔ ∡ p₃ p₂ p₁ = π := (o).oangle_eq_pi_iff_oangle_rev_eq_pi +/-- An oriented angle is not zero or `π` if and only if the three points are affinely +independent. -/ +lemma oangle_ne_zero_and_ne_pi_iff_affine_independent {p₁ p₂ p₃ : P} : + (∡ p₁ p₂ p₃ ≠ 0 ∧ ∡ p₁ p₂ p₃ ≠ π) ↔ affine_independent ℝ ![p₁, p₂, p₃] := +begin + rw [oangle, (o).oangle_ne_zero_and_ne_pi_iff_linear_independent, + affine_independent_iff_linear_independent_vsub ℝ _ (1 : fin 3), + ←linear_independent_equiv (fin_succ_above_equiv (1 : fin 3)).to_equiv], + convert iff.rfl, + ext i, + fin_cases i; + refl +end + +/-- An oriented angle is zero or `π` if and only if the three points are collinear. -/ +lemma oangle_eq_zero_or_eq_pi_iff_collinear {p₁ p₂ p₃ : P} : + (∡ p₁ p₂ p₃ = 0 ∨ ∡ p₁ p₂ p₃ = π) ↔ collinear ℝ ({p₁, p₂, p₃} : set P) := +begin + rw [←not_iff_not, not_or_distrib, oangle_ne_zero_and_ne_pi_iff_affine_independent, + affine_independent_iff_not_collinear], + simp [-set.union_singleton] +end + /-- Given three points not equal to `p`, the angle between the first and the second at `p` plus the angle between the second and the third equals the angle between the first and the third. -/ @[simp] lemma oangle_add {p p₁ p₂ p₃ : P} (hp₁ : p₁ ≠ p) (hp₂ : p₂ ≠ p) (hp₃ : p₃ ≠ p) : From 34c70db665acbced1de5edacd4e444d800af00e6 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Tue, 11 Oct 2022 03:45:43 +0000 Subject: [PATCH 675/828] feat(geometry/euclidean/oriented_angle): oriented angles and betweenness (#16824) Add lemmas relating oriented angles to strict and weak betweenness of points. --- src/geometry/euclidean/oriented_angle.lean | 137 +++++++++++++++++++++ 1 file changed, 137 insertions(+) diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 37273cb70b561..26d495b857624 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -1994,4 +1994,141 @@ by rw [oangle_swap₁₃_sign, ←oangle_swap₁₂_sign, oangle_swap₁₃_sign lemma oangle_rotate_sign (p₁ p₂ p₃ : P) : (∡ p₂ p₃ p₁).sign = (∡ p₁ p₂ p₃).sign := by rw [←oangle_swap₁₂_sign, oangle_swap₁₃_sign] +/-- The oriented angle between three points is π if and only if the second point is strictly +between the other two. -/ +lemma oangle_eq_pi_iff_sbtw {p₁ p₂ p₃ : P} : ∡ p₁ p₂ p₃ = π ↔ sbtw ℝ p₁ p₂ p₃ := +by rw [oangle_eq_pi_iff_angle_eq_pi, angle_eq_pi_iff_sbtw] + +/-- If the second of three points is strictly between the other two, the oriented angle at that +point is π. -/ +lemma _root_.sbtw.oangle₁₂₃_eq_pi {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∡ p₁ p₂ p₃ = π := +oangle_eq_pi_iff_sbtw.2 h + +/-- If the second of three points is strictly between the other two, the oriented angle at that +point (reversed) is π. -/ +lemma _root_.sbtw.oangle₃₂₁_eq_pi {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∡ p₃ p₂ p₁ = π := +by rw [oangle_eq_pi_iff_oangle_rev_eq_pi, ←h.oangle₁₂₃_eq_pi] + +/-- If the second of three points is weakly between the other two, the oriented angle at the +first point is zero. -/ +lemma _root_.wbtw.oangle₂₁₃_eq_zero {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) : ∡ p₂ p₁ p₃ = 0 := +begin + by_cases hp₂p₁ : p₂ = p₁, { simp [hp₂p₁] }, + by_cases hp₃p₁ : p₃ = p₁, { simp [hp₃p₁] }, + rw oangle_eq_zero_iff_angle_eq_zero hp₂p₁ hp₃p₁, + exact h.angle₂₁₃_eq_zero_of_ne hp₂p₁ +end + +/-- If the second of three points is strictly between the other two, the oriented angle at the +first point is zero. -/ +lemma _root_.sbtw.oangle₂₁₃_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∡ p₂ p₁ p₃ = 0 := +h.wbtw.oangle₂₁₃_eq_zero + +/-- If the second of three points is weakly between the other two, the oriented angle at the +first point (reversed) is zero. -/ +lemma _root_.wbtw.oangle₃₁₂_eq_zero {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) : ∡ p₃ p₁ p₂ = 0 := +by rw [oangle_eq_zero_iff_oangle_rev_eq_zero, h.oangle₂₁₃_eq_zero] + +/-- If the second of three points is strictly between the other two, the oriented angle at the +first point (reversed) is zero. -/ +lemma _root_.sbtw.oangle₃₁₂_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∡ p₃ p₁ p₂ = 0 := +h.wbtw.oangle₃₁₂_eq_zero + +/-- If the second of three points is weakly between the other two, the oriented angle at the +third point is zero. -/ +lemma _root_.wbtw.oangle₂₃₁_eq_zero {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) : ∡ p₂ p₃ p₁ = 0 := +h.symm.oangle₂₁₃_eq_zero + +/-- If the second of three points is strictly between the other two, the oriented angle at the +third point is zero. -/ +lemma _root_.sbtw.oangle₂₃₁_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∡ p₂ p₃ p₁ = 0 := +h.wbtw.oangle₂₃₁_eq_zero + +/-- If the second of three points is weakly between the other two, the oriented angle at the +third point (reversed) is zero. -/ +lemma _root_.wbtw.oangle₁₃₂_eq_zero {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) : ∡ p₁ p₃ p₂ = 0 := +h.symm.oangle₃₁₂_eq_zero + +/-- If the second of three points is strictly between the other two, the oriented angle at the +third point (reversed) is zero. -/ +lemma _root_.sbtw.oangle₁₃₂_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∡ p₁ p₃ p₂ = 0 := +h.wbtw.oangle₁₃₂_eq_zero + +/-- The oriented angle between three points is zero if and only if one of the first and third +points is weakly between the other two. -/ +lemma oangle_eq_zero_iff_wbtw {p₁ p₂ p₃ : P} : + ∡ p₁ p₂ p₃ = 0 ↔ wbtw ℝ p₂ p₁ p₃ ∨ wbtw ℝ p₂ p₃ p₁ := +begin + by_cases hp₁p₂ : p₁ = p₂, { simp [hp₁p₂] }, + by_cases hp₃p₂ : p₃ = p₂, { simp [hp₃p₂] }, + rw [oangle_eq_zero_iff_angle_eq_zero hp₁p₂ hp₃p₂, angle_eq_zero_iff_ne_and_wbtw], + simp [hp₁p₂, hp₃p₂] +end + +/-- An oriented angle is unchanged by replacing the first point by one weakly further away on the +same ray. -/ +lemma _root_.wbtw.oangle_eq_left {p₁ p₁' p₂ p₃ : P} (h : wbtw ℝ p₂ p₁ p₁') (hp₁p₂ : p₁ ≠ p₂) : + ∡ p₁ p₂ p₃ = ∡ p₁' p₂ p₃ := +begin + by_cases hp₃p₂ : p₃ = p₂, { simp [hp₃p₂] }, + by_cases hp₁'p₂ : p₁' = p₂, { rw [hp₁'p₂, wbtw_self_iff] at h, exact false.elim (hp₁p₂ h) }, + rw [←oangle_add hp₁'p₂ hp₁p₂ hp₃p₂, h.oangle₃₁₂_eq_zero, zero_add] +end + +/-- An oriented angle is unchanged by replacing the first point by one strictly further away on +the same ray. -/ +lemma _root_.sbtw.oangle_eq_left {p₁ p₁' p₂ p₃ : P} (h : sbtw ℝ p₂ p₁ p₁') : + ∡ p₁ p₂ p₃ = ∡ p₁' p₂ p₃ := +h.wbtw.oangle_eq_left h.ne_left + +/-- An oriented angle is unchanged by replacing the third point by one weakly further away on the +same ray. -/ +lemma _root_.wbtw.oangle_eq_right {p₁ p₂ p₃ p₃' : P} (h : wbtw ℝ p₂ p₃ p₃') (hp₃p₂ : p₃ ≠ p₂) : + ∡ p₁ p₂ p₃ = ∡ p₁ p₂ p₃' := +by rw [oangle_rev, h.oangle_eq_left hp₃p₂, ←oangle_rev] + +/-- An oriented angle is unchanged by replacing the third point by one strictly further away on +the same ray. -/ +lemma _root_.sbtw.oangle_eq_right {p₁ p₂ p₃ p₃' : P} (h : sbtw ℝ p₂ p₃ p₃') : + ∡ p₁ p₂ p₃ = ∡ p₁ p₂ p₃' := +h.wbtw.oangle_eq_right h.ne_left + +/-- Replacing the first point by one on the same line but the opposite ray adds π to the oriented +angle. -/ +lemma _root_.sbtw.oangle_eq_add_pi_left {p₁ p₁' p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₁') + (hp₃p₂ : p₃ ≠ p₂) : ∡ p₁ p₂ p₃ = ∡ p₁' p₂ p₃ + π := +by rw [←h.oangle₁₂₃_eq_pi, oangle_add_swap h.left_ne h.right_ne hp₃p₂] + +/-- Replacing the third point by one on the same line but the opposite ray adds π to the oriented +angle. -/ +lemma _root_.sbtw.oangle_eq_add_pi_right {p₁ p₂ p₃ p₃' : P} (h : sbtw ℝ p₃ p₂ p₃') + (hp₁p₂ : p₁ ≠ p₂) : ∡ p₁ p₂ p₃ = ∡ p₁ p₂ p₃' + π := +by rw [←h.oangle₃₂₁_eq_pi, oangle_add hp₁p₂ h.right_ne h.left_ne] + +/-- Replacing both the first and third points by ones on the same lines but the opposite rays +does not change the oriented angle (vertically opposite angles). -/ +lemma _root_.sbtw.oangle_eq_left_right {p₁ p₁' p₂ p₃ p₃' : P} (h₁ : sbtw ℝ p₁ p₂ p₁') + (h₃ : sbtw ℝ p₃ p₂ p₃') : ∡ p₁ p₂ p₃ = ∡ p₁' p₂ p₃' := +by rw [h₁.oangle_eq_add_pi_left h₃.left_ne, h₃.oangle_eq_add_pi_right h₁.right_ne, add_assoc, + real.angle.coe_pi_add_coe_pi, add_zero] + +/-- Replacing the first point by one on the same line does not change twice the oriented angle. -/ +lemma _root_.collinear.two_zsmul_oangle_eq_left {p₁ p₁' p₂ p₃ : P} + (h : collinear ℝ ({p₁, p₂, p₁'} : set P)) (hp₁p₂ : p₁ ≠ p₂) (hp₁'p₂ : p₁' ≠ p₂) : + (2 : ℤ) • ∡ p₁ p₂ p₃ = (2 : ℤ) • ∡ p₁' p₂ p₃ := +begin + by_cases hp₃p₂ : p₃ = p₂, { simp [hp₃p₂] }, + rcases h.wbtw_or_wbtw_or_wbtw with hw | hw | hw, + { have hw' : sbtw ℝ p₁ p₂ p₁' := ⟨hw, hp₁p₂.symm, hp₁'p₂.symm⟩, + rw [hw'.oangle_eq_add_pi_left hp₃p₂, smul_add, real.angle.two_zsmul_coe_pi, add_zero] }, + { rw hw.oangle_eq_left hp₁'p₂ }, + { rw hw.symm.oangle_eq_left hp₁p₂ } +end + +/-- Replacing the third point by one on the same line does not change twice the oriented angle. -/ +lemma _root_.collinear.two_zsmul_oangle_eq_right {p₁ p₂ p₃ p₃' : P} + (h : collinear ℝ ({p₃, p₂, p₃'} : set P)) (hp₃p₂ : p₃ ≠ p₂) (hp₃'p₂ : p₃' ≠ p₂) : + (2 : ℤ) • ∡ p₁ p₂ p₃ = (2 : ℤ) • ∡ p₁ p₂ p₃' := +by rw [oangle_rev, smul_neg, h.two_zsmul_oangle_eq_left hp₃p₂ hp₃'p₂, ←smul_neg, ←oangle_rev] + end euclidean_geometry From 7bbc52b758daaa5dbe02d472dd22e1f8e3af42bd Mon Sep 17 00:00:00 2001 From: ReimannJ Date: Tue, 11 Oct 2022 09:55:39 +0000 Subject: [PATCH 676/828] feat(measure_theory/integral/riesz_markov_kakutani): begin construction of content for Riesz-Markov-Kakutani (#16367) Begin construction of content for Riesz-Markov-Kakutani Construct a map on the compact sets and prove monotonicity and subadditivity. Co-authored-by: ReimannJ <105789986+ReimannJ@users.noreply.github.com> --- .../integral/riesz_markov_kakutani.lean | 111 ++++++++++++++++++ 1 file changed, 111 insertions(+) create mode 100644 src/measure_theory/integral/riesz_markov_kakutani.lean diff --git a/src/measure_theory/integral/riesz_markov_kakutani.lean b/src/measure_theory/integral/riesz_markov_kakutani.lean new file mode 100644 index 0000000000000..cd3cbbbe813a8 --- /dev/null +++ b/src/measure_theory/integral/riesz_markov_kakutani.lean @@ -0,0 +1,111 @@ +/- +Copyright (c) 2022 Jesse Reimann. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jesse Reimann, Kalle Kytölä +-/ +import topology.continuous_function.bounded + +/-! +# Riesz–Markov–Kakutani representation theorem + +This file will prove different versions of the Riesz-Markov-Kakutani representation theorem. +The theorem is first proven for compact spaces, from which the statements about linear functionals +on bounded continuous functions or compactly supported functions on locally compact spaces follow. + +To make use of the existing API, the measure is constructed from a content `λ` on the +compact subsets of the space X, rather than the usual construction of open sets in the literature. + +## References + +* [Walter Rudin, Real and Complex Analysis.][Rud87] + +-/ + +noncomputable theory + +open_locale bounded_continuous_function nnreal ennreal +open set function topological_space + +variables {X : Type*} [topological_space X] +variables (Λ : (X →ᵇ ℝ≥0) →ₗ[ℝ≥0] ℝ≥0) + +/-! ### Construction of the content: -/ + +/-- Given a positive linear functional Λ on X, for `K ⊆ X` compact define +`λ(K) = inf {Λf | 1≤f on K}`. When X is a compact Hausdorff space, this will be shown to be a +content, and will be shown to agree with the Riesz measure on the compact subsets `K ⊆ X`. -/ +def riesz_content_aux : (compacts X) → ℝ≥0 := +λ K, Inf (Λ '' {f : X →ᵇ ℝ≥0 | ∀ x ∈ K, (1 : ℝ≥0) ≤ f x}) + +section riesz_monotone + +/-- For any compact subset `K ⊆ X`, there exist some bounded continuous nonnegative +functions f on X such that `f ≥ 1` on K. -/ +lemma riesz_content_aux_image_nonempty (K : compacts X) : + (Λ '' {f : X →ᵇ ℝ≥0 | ∀ x ∈ K, (1 : ℝ≥0) ≤ f x}).nonempty := +begin + rw nonempty_image_iff, + use (1 : X →ᵇ ℝ≥0), + intros x x_in_K, + simp only [bounded_continuous_function.coe_one, pi.one_apply], +end + +/-- Riesz content λ (associated with a positive linear functional Λ) is +monotone: if `K₁ ⊆ K₂` are compact subsets in X, then `λ(K₁) ≤ λ(K₂)`. -/ +lemma riesz_content_aux_mono {K₁ K₂ : compacts X} (h : K₁ ≤ K₂) : + riesz_content_aux Λ K₁ ≤ riesz_content_aux Λ K₂ := +cInf_le_cInf (order_bot.bdd_below _) (riesz_content_aux_image_nonempty Λ K₂) + (image_subset Λ (set_of_subset_set_of.mpr (λ f f_hyp x x_in_K₁, f_hyp x (h x_in_K₁)))) + +end riesz_monotone + +section riesz_subadditive + +/-- Any bounded continuous nonnegative f such that `f ≥ 1` on K gives an upper bound on the +content of K; namely `λ(K) ≤ Λ f`. -/ +lemma riesz_content_aux_le {K : compacts X} + {f : X →ᵇ ℝ≥0} (h : ∀ x ∈ K, (1 : ℝ≥0) ≤ f x) : + riesz_content_aux Λ K ≤ Λ f := cInf_le (order_bot.bdd_below _) ⟨f, ⟨h, rfl⟩⟩ + +/-- The Riesz content can be approximated arbitrarily well by evaluating the positive linear +functional on test functions: for any `ε > 0`, there exists a bounded continuous nonnegative +function f on X such that `f ≥ 1` on K and such that `λ(K) ≤ Λ f < λ(K) + ε`. -/ +lemma exists_lt_riesz_content_aux_add_pos (K : compacts X) + {ε : ℝ≥0} (εpos : 0 < ε) : + ∃ (f : X →ᵇ ℝ≥0), (∀ x ∈ K, (1 : ℝ≥0) ≤ f x) ∧ Λ f < riesz_content_aux Λ K + ε := +begin + --choose a test function `f` s.t. `Λf = α < λ(K) + ε` + obtain ⟨α, ⟨⟨f, f_hyp⟩, α_hyp⟩⟩ := + exists_lt_of_cInf_lt (riesz_content_aux_image_nonempty Λ K) + (lt_add_of_pos_right (riesz_content_aux Λ K) εpos), + refine ⟨f, f_hyp.left, _ ⟩, + rw f_hyp.right, + exact α_hyp, +end + +/-- The Riesz content λ associated to a given positive linear functional Λ is +finitely subadditive: `λ(K₁ ∪ K₂) ≤ λ(K₁) + λ(K₂)` for any compact subsets `K₁, K₂ ⊆ X`. -/ +lemma riesz_content_aux_sup_le (K1 K2 : compacts X) : + riesz_content_aux Λ (K1 ⊔ K2) ≤ riesz_content_aux Λ (K1) + riesz_content_aux Λ (K2) := +begin + apply nnreal.le_of_forall_pos_le_add, + intros ε εpos, + --get test functions s.t. `λ(Ki) ≤ Λfi ≤ λ(Ki) + ε/2, i=1,2` + obtain ⟨f1, f_test_function_K1⟩ := exists_lt_riesz_content_aux_add_pos Λ K1 + (nnreal.half_pos εpos), + obtain ⟨f2, f_test_function_K2⟩ := exists_lt_riesz_content_aux_add_pos Λ K2 + (nnreal.half_pos εpos), + --let `f := f1 + f2` test function for the content of `K` + have f_test_function_union : (∀ x ∈ (K1 ⊔ K2), (1 : ℝ≥0) ≤ (f1 + f2) x), + { rintros x (x_in_K1 | x_in_K2), + { exact le_add_right (f_test_function_K1.left x x_in_K1) }, + { exact le_add_left (f_test_function_K2.left x x_in_K2) }}, + --use that `Λf` is an upper bound for `λ(K1⊔K2)` + apply (riesz_content_aux_le Λ f_test_function_union).trans (le_of_lt _), + rw map_add, + --use that `Λfi` are lower bounds for `λ(Ki) + ε/2` + apply lt_of_lt_of_le (add_lt_add f_test_function_K1.right f_test_function_K2.right) (le_of_eq _), + rw [add_assoc, add_comm (ε/2), add_assoc, nnreal.add_halves ε, add_assoc], +end + +end riesz_subadditive From 43e95eefc5ef91202effd28c04de51283c2bcf1b Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 11 Oct 2022 09:55:41 +0000 Subject: [PATCH 677/828] feat(topology/separation): add instances about `separation_quotient` (#16579) --- src/order/filter/basic.lean | 11 +++++++++ src/topology/inseparable.lean | 29 ++++++++++++++++++++++-- src/topology/maps.lean | 2 +- src/topology/separation.lean | 42 +++++++++++++++++++++++++++++++++++ 4 files changed, 81 insertions(+), 3 deletions(-) diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 6d3e18b988a20..67cf2f8971247 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -1921,6 +1921,17 @@ lemma comap_ne_bot_iff_compl_range {f : filter β} {m : α → β} : ne_bot (comap m f) ↔ (range m)ᶜ ∉ f := comap_ne_bot_iff_frequently +lemma comap_eq_bot_iff_compl_range {f : filter β} {m : α → β} : + comap m f = ⊥ ↔ (range m)ᶜ ∈ f := +not_iff_not.mp $ ne_bot_iff.symm.trans comap_ne_bot_iff_compl_range + +lemma comap_surjective_eq_bot {f : filter β} {m : α → β} (hm : surjective m) : + comap m f = ⊥ ↔ f = ⊥ := +by rw [comap_eq_bot_iff_compl_range, hm.range_eq, compl_univ, empty_mem_iff_bot] + +lemma disjoint_comap_iff (h : surjective m) : disjoint (comap m g₁) (comap m g₂) ↔ disjoint g₁ g₂ := +by rw [disjoint_iff, disjoint_iff, ← comap_inf, comap_surjective_eq_bot h] + lemma ne_bot.comap_of_range_mem {f : filter β} {m : α → β} (hf : ne_bot f) (hm : range m ∈ f) : ne_bot (comap m f) := comap_ne_bot_iff_frequently.2 $ eventually.frequently hm diff --git a/src/topology/inseparable.lean b/src/topology/inseparable.lean index 77338de425f79..6e95d75c61c72 100644 --- a/src/topology/inseparable.lean +++ b/src/topology/inseparable.lean @@ -301,7 +301,7 @@ be a T₀ space. -/ @[derive topological_space] def separation_quotient := quotient (inseparable_setoid X) -variable {X} +variables {X} {t : set (separation_quotient X)} namespace separation_quotient @@ -350,7 +350,32 @@ lemma inducing_mk : inducing (mk : X → separation_quotient X) := lemma is_closed_map_mk : is_closed_map (mk : X → separation_quotient X) := inducing_mk.is_closed_map $ by { rw [range_mk], exact is_closed_univ } +@[simp] lemma comap_mk_nhds_mk : comap mk (𝓝 (mk x)) = 𝓝 x := +(inducing_mk.nhds_eq_comap _).symm + +@[simp] lemma comap_mk_nhds_set_image : comap mk (𝓝ˢ (mk '' s)) = 𝓝ˢ s := +(inducing_mk.nhds_set_eq_comap _).symm + lemma map_mk_nhds : map mk (𝓝 x) = 𝓝 (mk x) := -by rw [inducing_mk.nhds_eq_comap, map_comap_of_surjective surjective_mk] +by rw [← comap_mk_nhds_mk, map_comap_of_surjective surjective_mk] + +lemma map_mk_nhds_set : map mk (𝓝ˢ s) = 𝓝ˢ (mk '' s) := +by rw [← comap_mk_nhds_set_image, map_comap_of_surjective surjective_mk] + +lemma comap_mk_nhds_set : comap mk (𝓝ˢ t) = 𝓝ˢ (mk ⁻¹' t) := +by conv_lhs { rw [← image_preimage_eq t surjective_mk, comap_mk_nhds_set_image] } + +lemma preimage_mk_closure : mk ⁻¹' (closure t) = closure (mk ⁻¹' t) := +is_open_map_mk.preimage_closure_eq_closure_preimage continuous_mk t + +lemma preimage_mk_interior : mk ⁻¹' (interior t) = interior (mk ⁻¹' t) := +is_open_map_mk.preimage_interior_eq_interior_preimage continuous_mk t + +lemma preimage_mk_frontier : mk ⁻¹' (frontier t) = frontier (mk ⁻¹' t) := +is_open_map_mk.preimage_frontier_eq_frontier_preimage continuous_mk t + +lemma image_mk_closure : mk '' closure s = closure (mk '' s) := +(image_closure_subset_closure_image continuous_mk).antisymm $ + is_closed_map_mk.closure_image_subset _ end separation_quotient diff --git a/src/topology/maps.lean b/src/topology/maps.lean index a7252ab91a5ef..458748676cac0 100644 --- a/src/topology/maps.lean +++ b/src/topology/maps.lean @@ -580,7 +580,7 @@ lemma closed_embedding.comp {g : β → γ} {f : α → β} lemma closed_embedding.closure_image_eq {f : α → β} (hf : closed_embedding f) (s : set α) : closure (f '' s) = f '' closure s := -le_antisymm (is_closed_map_iff_closure_image.mp hf.is_closed_map _) +(hf.is_closed_map.closure_image_subset _).antisymm (image_closure_subset_closure_image hf.continuous) end closed_embedding diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 166be11df5040..5d40e4cbbec0c 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -1487,6 +1487,14 @@ begin V₁_closed, V₂_closed, U₁_op, U₂_op, h₁, h₂, H⟩ end +open separation_quotient + +/-- The `separation_quotient` of a regular space is a T₃ space. -/ +instance [regular_space α] : t3_space (separation_quotient α) := +{ regular := λ s, surjective_mk.forall.2 $ λ a hs ha, + by { rw [← disjoint_comap_iff surjective_mk, comap_mk_nhds_mk, comap_mk_nhds_set], + exact regular_space.regular (hs.preimage continuous_mk) ha } } + end t3 section normality @@ -1536,6 +1544,24 @@ protected lemma closed_embedding.normal_space [topological_space β] [normal_spa exact (H.preimage hf.continuous).mono (subset_preimage_image _ _) (subset_preimage_image _ _) end } +namespace separation_quotient + +/-- The `separation_quotient` of a normal space is a T₄ space. We don't have separate typeclasses +for normal spaces (without T₁ assumption) and T₄ spaces, so we use the same class for assumption +and for conclusion. + +One can prove this using a homeomorphism between `α` and `separation_quotient α`. We give an +alternative proof that works without assuming that `α` is a T₁ space. -/ +instance [normal_space α] : normal_space (separation_quotient α) := +{ normal := λ s t hs ht hd, separated_nhds_iff_disjoint.2 $ + begin + rw [← disjoint_comap_iff surjective_mk, comap_mk_nhds_set, comap_mk_nhds_set], + exact separated_nhds_iff_disjoint.1 (normal_separation (hs.preimage continuous_mk) + (ht.preimage continuous_mk) (hd.preimage mk)) + end } + +end separation_quotient + variable (α) /-- A T₃ topological space with second countable topology is a normal space. @@ -1622,6 +1648,22 @@ instance t5_space.to_normal_space [t5_space α] : normal_space α := ⟨λ s t hs ht hd, separated_nhds_iff_disjoint.2 $ completely_normal (by rwa [hs.closure_eq]) (by rwa [ht.closure_eq])⟩ +open separation_quotient + +/-- The `separation_quotient` of a completely normal space is a T₅ space. We don't have separate +typeclasses for completely normal spaces (without T₁ assumption) and T₅ spaces, so we use the same +class for assumption and for conclusion. + +One can prove this using a homeomorphism between `α` and `separation_quotient α`. We give an +alternative proof that works without assuming that `α` is a T₁ space. -/ +instance [t5_space α] : t5_space (separation_quotient α) := +{ completely_normal := λ s t hd₁ hd₂, + begin + rw [← disjoint_comap_iff surjective_mk, comap_mk_nhds_set, comap_mk_nhds_set], + apply t5_space.completely_normal; rw [← preimage_mk_closure], + exacts [hd₁.preimage mk, hd₂.preimage mk] + end } + end completely_normal /-- In a compact t2 space, the connected component of a point equals the intersection of all From 0fbd703b3da57ec0476dc1bc06931994055ae2f5 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 11 Oct 2022 09:55:42 +0000 Subject: [PATCH 678/828] =?UTF-8?q?feat(topology/instances/add=5Fcircle):?= =?UTF-8?q?=20the=20additive=20circle=20`=E2=84=9D=20/=20(a=20=E2=88=99=20?= =?UTF-8?q?=E2=84=A4)`=20is=20normal=20(T4)=20and=20second-countable=20(#1?= =?UTF-8?q?6594)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The T4 fact is in mathlib already, via the `normed_add_comm_group` instance on `ℝ / (a ∙ ℤ)`, but this gives it earlier in the hierarchy. Co-authored-by: Alex Kontorovich <58564076+AlexKontorovich@users.noreply.github.com> --- src/topology/instances/add_circle.lean | 56 +++++++++++++++++--------- 1 file changed, 36 insertions(+), 20 deletions(-) diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index cc7d4a0d9c703..dbc36f21a8f70 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -38,7 +38,7 @@ the rational circle `add_circle (1 : ℚ)`, and so we set things up more general noncomputable theory -open set int (hiding mem_zmultiples_iff) add_subgroup +open set int (hiding mem_zmultiples_iff) add_subgroup topological_space variables {𝕜 : Type*} @@ -80,55 +80,71 @@ rfl (equiv_add_circle p q hp hq).symm (x : 𝕜) = (x * (q⁻¹ * p) : 𝕜) := rfl -variables [floor_ring 𝕜] +variables [floor_ring 𝕜] [hp : fact (0 < p)] +include hp /-- The natural equivalence between `add_circle p` and the half-open interval `[0, p)`. -/ -def equiv_Ico (hp : 0 < p) : add_circle p ≃ Ico 0 p := +def equiv_Ico : add_circle p ≃ Ico 0 p := { inv_fun := quotient_add_group.mk' _ ∘ coe, - to_fun := λ x, ⟨(to_Ico_mod_periodic 0 hp).lift x, quot.induction_on x $ to_Ico_mod_mem_Ico' hp⟩, + to_fun := λ x, ⟨(to_Ico_mod_periodic 0 hp.out).lift x, + quot.induction_on x $ to_Ico_mod_mem_Ico' hp.out⟩, right_inv := by { rintros ⟨x, hx⟩, ext, simp [to_Ico_mod_eq_self, hx.1, hx.2], }, left_inv := begin rintros ⟨x⟩, - change quotient_add_group.mk (to_Ico_mod 0 hp x) = quotient_add_group.mk x, + change quotient_add_group.mk (to_Ico_mod 0 hp.out x) = quotient_add_group.mk x, rw [quotient_add_group.eq', neg_add_eq_sub, self_sub_to_Ico_mod, zsmul_eq_mul], apply int_cast_mul_mem_zmultiples, end } -@[simp] lemma coe_equiv_Ico_mk_apply (hp : 0 < p) (x : 𝕜) : - (equiv_Ico p hp $ quotient_add_group.mk x : 𝕜) = fract (x / p) * p := -to_Ico_mod_eq_fract_mul hp x +@[simp] lemma coe_equiv_Ico_mk_apply (x : 𝕜) : + (equiv_Ico p $ quotient_add_group.mk x : 𝕜) = fract (x / p) * p := +to_Ico_mod_eq_fract_mul _ x -@[continuity] lemma continuous_equiv_Ico_symm (hp : 0 < p) : continuous (equiv_Ico p hp).symm := +@[continuity] lemma continuous_equiv_Ico_symm : continuous (equiv_Ico p).symm := continuous_coinduced_rng.comp continuous_induced_dom /-- The image of the closed interval `[0, p]` under the quotient map `𝕜 → add_circle p` is the entire space. -/ -@[simp] lemma coe_image_Icc_eq (hp : 0 < p) : +@[simp] lemma coe_image_Icc_eq : (coe : 𝕜 → add_circle p) '' (Icc 0 p) = univ := begin refine eq_univ_iff_forall.mpr (λ x, _), - let y := equiv_Ico p hp x, - exact ⟨y, ⟨y.2.1, y.2.2.le⟩, (equiv_Ico p hp).symm_apply_apply x⟩, + let y := equiv_Ico p x, + exact ⟨y, ⟨y.2.1, y.2.2.le⟩, (equiv_Ico p).symm_apply_apply x⟩, end end linear_ordered_field variables (p : ℝ) -lemma compact_space (hp : 0 < p) : compact_space $ add_circle p := +/-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is compact. -/ +instance compact_space [fact (0 < p)] : compact_space $ add_circle p := begin - rw [← is_compact_univ_iff, ← coe_image_Icc_eq p hp], + rw [← is_compact_univ_iff, ← coe_image_Icc_eq p], exact is_compact_Icc.image (add_circle.continuous_mk' p), end -end add_circle +/-- The action on `ℝ` by right multiplication of its the subgroup `zmultiples p` (the multiples of +`p:ℝ`) is properly discontinuous. -/ +instance : properly_discontinuous_vadd (zmultiples p).opposite ℝ := +(zmultiples p).properly_discontinuous_vadd_opposite_of_tendsto_cofinite + (add_subgroup.tendsto_zmultiples_subtype_cofinite p) -/-- The unit circle `ℝ ⧸ ℤ`. -/ -abbreviation unit_add_circle := add_circle (1 : ℝ) +/-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is Hausdorff. -/ +instance : t2_space (add_circle p) := t2_space_of_properly_discontinuous_vadd_of_t2_space + +/-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is normal. -/ +instance [fact (0 < p)] : normal_space (add_circle p) := normal_of_compact_t2 -namespace unit_add_circle +/-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is second-countable. -/ +instance : second_countable_topology (add_circle p) := quotient_add_group.second_countable_topology -instance : compact_space unit_add_circle := add_circle.compact_space _ zero_lt_one +end add_circle + +private lemma fact_zero_lt_one : fact ((0:ℝ) < 1) := ⟨zero_lt_one⟩ +local attribute [instance] fact_zero_lt_one -end unit_add_circle +/-- The unit circle `ℝ ⧸ ℤ`. -/ +@[derive [compact_space, normal_space, second_countable_topology]] +abbreviation unit_add_circle := add_circle (1 : ℝ) From 45ef183d061eb4f6d900a8d15f985be3899ccf01 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Tue, 11 Oct 2022 09:55:44 +0000 Subject: [PATCH 679/828] feat(group_theory/subgroup/basic): The center can be written as the intersection of centralizers (#16837) This PRs adds a lemma that writes the center as the intersection of centralizers. I kept `S` explicit so that you can `rw center_eq_infi S` (specifying `S` without having to immediately supply a proof of `hS`). --- src/group_theory/subgroup/basic.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index dcbc5dc1bb42a..ad42a5666e90a 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2772,6 +2772,14 @@ le_antisymm (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ zpowers_le.2 $ su $ le_centralizer_iff.1 $ (closure_le _).2 $ λ g, set_like.mem_coe.2 ∘ zpowers_le.1 ∘ le_centralizer_iff.1 ∘ infi_le_of_le g ∘ infi_le _ +@[to_additive] lemma center_eq_infi (S : set G) (hS : closure S = ⊤) : + center G = ⨅ g ∈ S, centralizer (zpowers g) := +by rw [←centralizer_top, ←hS, centralizer_closure] + +@[to_additive] lemma center_eq_infi' (S : set G) (hS : closure S = ⊤) : + center G = ⨅ g : S, centralizer (zpowers g) := +by rw [center_eq_infi S hS, ←infi_subtype''] + end subgroup namespace monoid_hom From 03eaefc4f874b8dd5c6998bb799f8e23c01631b1 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 11 Oct 2022 09:55:46 +0000 Subject: [PATCH 680/828] =?UTF-8?q?feat(ring=5Ftheory/etale):=20Formally?= =?UTF-8?q?=20unramified=20iff=20`=CE=A9[S=E2=81=84R]=20=3D=200`=20(#16849?= =?UTF-8?q?)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/ring_theory/etale.lean | 61 ++++++++++++++++++++++++++++++++++++-- 1 file changed, 58 insertions(+), 3 deletions(-) diff --git a/src/ring_theory/etale.lean b/src/ring_theory/etale.lean index a3f16ba25edf1..c73ad75f99cc9 100644 --- a/src/ring_theory/etale.lean +++ b/src/ring_theory/etale.lean @@ -3,11 +3,9 @@ Copyright (c) 2022 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import ring_theory.ideal.operations import ring_theory.nilpotent -import ring_theory.tensor_product import linear_algebra.isomorphisms -import ring_theory.ideal.cotangent +import ring_theory.derivation /-! @@ -106,6 +104,27 @@ lemma formally_unramified.ext [formally_unramified R A] (hI : is_nilpotent I) g₁ = g₂ := formally_unramified.lift_unique I hI g₁ g₂ (alg_hom.ext H) +lemma formally_unramified.lift_unique_of_ring_hom [formally_unramified R A] + {C : Type u} [comm_ring C] (f : B →+* C) (hf : is_nilpotent f.ker) + (g₁ g₂ : A →ₐ[R] B) (h : f.comp ↑g₁ = f.comp (g₂ : A →+* B)) : g₁ = g₂ := +formally_unramified.lift_unique _ hf _ _ +begin + ext x, + have := ring_hom.congr_fun h x, + simpa only [ideal.quotient.eq, function.comp_app, alg_hom.coe_comp, ideal.quotient.mkₐ_eq_mk, + ring_hom.mem_ker, map_sub, sub_eq_zero], +end + +lemma formally_unramified.ext' [formally_unramified R A] + {C : Type u} [comm_ring C] (f : B →+* C) (hf : is_nilpotent f.ker) + (g₁ g₂ : A →ₐ[R] B) (h : ∀ x, f (g₁ x) = f (g₂ x)) : g₁ = g₂ := +formally_unramified.lift_unique_of_ring_hom f hf g₁ g₂ (ring_hom.ext h) + +lemma formally_unramified.lift_unique' [formally_unramified R A] + {C : Type u} [comm_ring C] [algebra R C] (f : B →ₐ[R] C) (hf : is_nilpotent (f : B →+* C).ker) + (g₁ g₂ : A →ₐ[R] B) (h : f.comp g₁ = f.comp g₂) : g₁ = g₂ := +formally_unramified.ext' _ hf g₁ g₂ (alg_hom.congr_fun h) + lemma formally_smooth.exists_lift {B : Type u} [comm_ring B] [_RB : algebra R B] [formally_smooth R A] (I : ideal B) (hI : is_nilpotent I) (g : A →ₐ[R] B ⧸ I) : @@ -355,6 +374,42 @@ end end of_surjective +section unramified_derivation + +open_locale tensor_product + +variables {R S : Type u} [comm_ring R] [comm_ring S] [algebra R S] + +instance formally_unramified.subsingleton_kaehler_differential [formally_unramified R S] : + subsingleton Ω[S⁄R] := +begin + rw ← not_nontrivial_iff_subsingleton, + introsI h, + obtain ⟨f₁, f₂, e⟩ := (kaehler_differential.End_equiv R S).injective.nontrivial, + apply e, + ext1, + apply formally_unramified.lift_unique' _ _ _ _ (f₁.2.trans f₂.2.symm), + rw [← alg_hom.to_ring_hom_eq_coe, alg_hom.ker_ker_sqare_lift], + exact ⟨_, ideal.cotangent_ideal_square _⟩, +end + +lemma formally_unramified.iff_subsingleton_kaehler_differential : + formally_unramified R S ↔ subsingleton Ω[S⁄R] := +begin + split, + { introsI, apply_instance }, + { introI H, + constructor, + introsI B _ _ I hI f₁ f₂ e, + letI := f₁.to_ring_hom.to_algebra, + haveI := is_scalar_tower.of_algebra_map_eq' (f₁.comp_algebra_map).symm, + have := ((kaehler_differential.linear_map_equiv_derivation R S).to_equiv.trans + (derivation_to_square_zero_equiv_lift I hI)).surjective.subsingleton, + exact subtype.ext_iff.mp (@@subsingleton.elim this ⟨f₁, rfl⟩ ⟨f₂, e.symm⟩) } +end + +end unramified_derivation + section base_change open_locale tensor_product From 21926ffe1255d452bf4ebb873ac16eedfc98f66b Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Tue, 11 Oct 2022 09:55:47 +0000 Subject: [PATCH 681/828] feat(group_theory/subgroup/pointwise): Basic lemmas (#16856) This PR adds some basic lemmas `smul_bot`, `smul_inf`, and `smul_normal` (a restatement of `normal.conj_act` in terms of `mul_aut.conj`). --- src/group_theory/subgroup/pointwise.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/group_theory/subgroup/pointwise.lean b/src/group_theory/subgroup/pointwise.lean index ee899b4054625..7d5b3aa2b3e24 100644 --- a/src/group_theory/subgroup/pointwise.lean +++ b/src/group_theory/subgroup/pointwise.lean @@ -60,6 +60,9 @@ lemma mem_smul_pointwise_iff_exists (m : G) (a : α) (S : subgroup G) : m ∈ a • S ↔ ∃ (s : G), s ∈ S ∧ a • s = m := (set.mem_smul_set : m ∈ a • (S : set G) ↔ _) +@[simp] lemma smul_bot (a : α) : a • (⊥ : subgroup G) = ⊥ := +by simp [set_like.ext_iff, mem_smul_pointwise_iff_exists, eq_comm] + instance pointwise_central_scalar [mul_distrib_mul_action αᵐᵒᵖ G] [is_central_scalar α G] : is_central_scalar α (subgroup G) := ⟨λ a S, congr_arg (λ f, S.map f) $ monoid_hom.ext $ by exact op_smul_eq_smul _⟩ @@ -109,6 +112,9 @@ set_smul_subset_iff lemma subset_pointwise_smul_iff {a : α} {S T : subgroup G} : S ≤ a • T ↔ a⁻¹ • S ≤ T := subset_set_smul_iff +@[simp] lemma smul_inf (a : α) (S T : subgroup G) : a • (S ⊓ T) = a • S ⊓ a • T := +by simp [set_like.ext_iff, mem_pointwise_smul_iff_inv_smul_mem] + /-- Applying a `mul_distrib_mul_action` results in an isomorphic subgroup -/ @[simps] def equiv_smul (a : α) (H : subgroup G) : H ≃* (a • H : subgroup G) := (mul_distrib_mul_action.to_mul_equiv G a).subgroup_map H @@ -152,6 +158,9 @@ begin exact h} end +@[simp] lemma smul_normal (g : G) (H : subgroup G) [h : normal H] : mul_aut.conj g • H = H := +h.conj_act g + end group section group_with_zero From b1aef4335385d6e551e83837615c534b916ac32a Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 11 Oct 2022 09:55:49 +0000 Subject: [PATCH 682/828] chore(algebra/ring/basic): move results on dvd (#16864) Move some results in `algebra.ring.basic` about `dvd` to `algebra.ring.divisibility`. Co-authored-by: Scott Morrison --- src/algebra/hom/ring.lean | 1 + src/algebra/order/group.lean | 3 + src/algebra/order/ring.lean | 1 + src/algebra/ring/basic.lean | 102 +----------------------- src/algebra/ring/divisibility.lean | 124 +++++++++++++++++++++++++++++ 5 files changed, 130 insertions(+), 101 deletions(-) create mode 100644 src/algebra/ring/divisibility.lean diff --git a/src/algebra/hom/ring.lean b/src/algebra/hom/ring.lean index 05d874afee47e..2d551a48d6c9f 100644 --- a/src/algebra/hom/ring.lean +++ b/src/algebra/hom/ring.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston, Jireh Loreaux -/ import algebra.ring.basic +import algebra.divisibility import data.pi.algebra import algebra.hom.units diff --git a/src/algebra/order/group.lean b/src/algebra/order/group.lean index 003f998f8f090..9aa93a76c2d33 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group.lean @@ -46,6 +46,9 @@ instance ordered_comm_group.to_covariant_class_left_le (α : Type u) [ordered_co covariant_class α α (*) (≤) := { elim := λ a b c bc, ordered_comm_group.mul_le_mul_left b c bc a } +example (α : Type u) [ordered_add_comm_group α] : covariant_class α α (swap (+)) (<) := +add_right_cancel_semigroup.covariant_swap_add_lt_of_covariant_swap_add_le α + /--The units of an ordered commutative monoid form an ordered commutative group. -/ @[to_additive "The units of an ordered commutative additive monoid form an ordered commutative additive group."] diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 8148b5a398eff..6bec664ae1ad2 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro -/ import algebra.char_zero.defs +import algebra.ring.divisibility import algebra.hom.ring import algebra.order.group import algebra.order.ring_lemmas diff --git a/src/algebra/ring/basic.lean b/src/algebra/ring/basic.lean index 346b69fa94b13..afef8dabffe13 100644 --- a/src/algebra/ring/basic.lean +++ b/src/algebra/ring/basic.lean @@ -3,8 +3,8 @@ Copyright (c) 2014 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland -/ -import algebra.divisibility import algebra.regular.basic +import algebra.hom.group import algebra.opposites /-! @@ -249,14 +249,6 @@ lemma one_add_one_eq_two : 1 + 1 = (2 : α) := rfl end has_one_has_add -section distrib_semigroup -variables [has_add α] [semigroup α] - -theorem dvd_add [left_distrib_class α] {a b c : α} (h₁ : a ∣ b) (h₂ : a ∣ c) : a ∣ b + c := -dvd.elim h₁ (λ d hd, dvd.elim h₂ (λ e he, dvd.intro (d + e) (by simp [left_distrib, hd, he]))) - -end distrib_semigroup - section distrib_mul_one_class variables [has_add α] [mul_one_class α] @@ -371,8 +363,6 @@ lemma mul_right_apply {R : Type*} [non_unital_non_assoc_semiring R] (a r : R) : end add_monoid_hom -@[simp] theorem two_dvd_bit0 [semiring α] {a : α} : 2 ∣ bit0 a := ⟨a, bit0_eq_two_mul _⟩ - /-- A non-unital commutative semiring is a `non_unital_semiring` with commutative multiplication. In other words, it is a type with the following structures: additive commutative monoid (`add_comm_monoid`), commutative semigroup (`comm_semigroup`), distributive laws (`distrib`), and @@ -403,10 +393,6 @@ protected def function.surjective.non_unital_comm_semiring [has_zero γ] [has_ad non_unital_comm_semiring γ := { .. hf.non_unital_semiring f zero add mul nsmul, .. hf.comm_semigroup f mul } -lemma has_dvd.dvd.linear_comb {d x y : α} (hdx : d ∣ x) (hdy : d ∣ y) (a b : α) : - d ∣ (a * x + b * y) := -dvd_add (hdx.mul_left a) (hdy.mul_left b) - end non_unital_comm_semiring /-- A commutative semiring is a `semiring` with commutative multiplication. In other words, it is a @@ -549,34 +535,6 @@ instance mul_zero_class.neg_zero_class : neg_zero_class α := end mul_zero_class -section semigroup - -variables [semigroup α] [has_distrib_neg α] {a b c : α} - -theorem dvd_neg_of_dvd (h : a ∣ b) : (a ∣ -b) := -let ⟨c, hc⟩ := h in ⟨-c, by simp [hc]⟩ - -theorem dvd_of_dvd_neg (h : a ∣ -b) : (a ∣ b) := -let t := dvd_neg_of_dvd h in by rwa neg_neg at t - -/-- An element a of a semigroup with a distributive negation divides the negation of an element b -iff a divides b. -/ -@[simp] lemma dvd_neg (a b : α) : (a ∣ -b) ↔ (a ∣ b) := -⟨dvd_of_dvd_neg, dvd_neg_of_dvd⟩ - -theorem neg_dvd_of_dvd (h : a ∣ b) : -a ∣ b := -let ⟨c, hc⟩ := h in ⟨-c, by simp [hc]⟩ - -theorem dvd_of_neg_dvd (h : -a ∣ b) : a ∣ b := -let t := neg_dvd_of_dvd h in by rwa neg_neg at t - -/-- The negation of an element a of a semigroup with a distributive negation divides -another element b iff a divides b. -/ -@[simp] lemma neg_dvd (a b : α) : (-a ∣ b) ↔ (a ∣ b) := -⟨dvd_of_neg_dvd, neg_dvd_of_dvd⟩ - -end semigroup - section group variables [group α] [has_distrib_neg α] @@ -902,56 +860,6 @@ instance comm_ring.to_comm_semiring [s : comm_ring α] : comm_semiring α := instance comm_ring.to_non_unital_comm_ring [s : comm_ring α] : non_unital_comm_ring α := { mul_zero := mul_zero, zero_mul := zero_mul, ..s } -section non_unital_ring -variables [non_unital_ring α] {a b c : α} - -theorem dvd_sub (h₁ : a ∣ b) (h₂ : a ∣ c) : a ∣ b - c := -by { rw sub_eq_add_neg, exact dvd_add h₁ (dvd_neg_of_dvd h₂) } - -theorem dvd_add_iff_left (h : a ∣ c) : a ∣ b ↔ a ∣ b + c := -⟨λh₂, dvd_add h₂ h, λH, by have t := dvd_sub H h; rwa add_sub_cancel at t⟩ - -theorem dvd_add_iff_right (h : a ∣ b) : a ∣ c ↔ a ∣ b + c := -by rw add_comm; exact dvd_add_iff_left h - -/-- If an element a divides another element c in a commutative ring, a divides the sum of another - element b with c iff a divides b. -/ -theorem dvd_add_left (h : a ∣ c) : a ∣ b + c ↔ a ∣ b := -(dvd_add_iff_left h).symm - -/-- If an element a divides another element b in a commutative ring, a divides the sum of b and - another element c iff a divides c. -/ -theorem dvd_add_right (h : a ∣ b) : a ∣ b + c ↔ a ∣ c := -(dvd_add_iff_right h).symm - -lemma dvd_iff_dvd_of_dvd_sub {a b c : α} (h : a ∣ (b - c)) : (a ∣ b ↔ a ∣ c) := -begin - split, - { intro h', - convert dvd_sub h' h, - exact eq.symm (sub_sub_self b c) }, - { intro h', - convert dvd_add h h', - exact eq_add_of_sub_eq rfl } -end - -end non_unital_ring - -section ring -variables [ring α] {a b c : α} - -theorem two_dvd_bit1 : 2 ∣ bit1 a ↔ (2 : α) ∣ 1 := (dvd_add_iff_right (@two_dvd_bit0 _ _ a)).symm - -/-- An element a divides the sum a + b if and only if a divides b.-/ -@[simp] lemma dvd_add_self_left {a b : α} : a ∣ a + b ↔ a ∣ b := -dvd_add_right (dvd_refl a) - -/-- An element a divides the sum b + a if and only if a divides b.-/ -@[simp] lemma dvd_add_self_right {a b : α} : a ∣ b + a ↔ a ∣ b := -dvd_add_left (dvd_refl a) - -end ring - section non_unital_comm_ring variables [non_unital_comm_ring α] {a b c : α} @@ -995,14 +903,6 @@ begin rw [this, sub_add, ← sub_mul, sub_self] end -lemma dvd_mul_sub_mul {k a b x y : α} (hab : k ∣ a - b) (hxy : k ∣ x - y) : - k ∣ a * x - b * y := -begin - convert dvd_add (hxy.mul_left a) (hab.mul_right y), - rw [mul_sub_left_distrib, mul_sub_right_distrib], - simp only [sub_eq_add_neg, add_assoc, neg_add_cancel_left], -end - end non_unital_comm_ring section comm_ring diff --git a/src/algebra/ring/divisibility.lean b/src/algebra/ring/divisibility.lean new file mode 100644 index 0000000000000..329a97b6611e6 --- /dev/null +++ b/src/algebra/ring/divisibility.lean @@ -0,0 +1,124 @@ +/- +Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Floris van Doorn, Yury Kudryashov, Neil Strickland +-/ +import algebra.ring.basic +import algebra.divisibility + +/-! +# Lemmas about divisibility in rings +-/ + +variables {α β : Type*} + +section distrib_semigroup +variables [has_add α] [semigroup α] + +theorem dvd_add [left_distrib_class α] {a b c : α} (h₁ : a ∣ b) (h₂ : a ∣ c) : a ∣ b + c := +dvd.elim h₁ (λ d hd, dvd.elim h₂ (λ e he, dvd.intro (d + e) (by simp [left_distrib, hd, he]))) + +end distrib_semigroup + +@[simp] theorem two_dvd_bit0 [semiring α] {a : α} : 2 ∣ bit0 a := ⟨a, bit0_eq_two_mul _⟩ + +section non_unital_comm_semiring +variables [non_unital_comm_semiring α] [non_unital_comm_semiring β] {a b c : α} + +lemma has_dvd.dvd.linear_comb {d x y : α} (hdx : d ∣ x) (hdy : d ∣ y) (a b : α) : + d ∣ (a * x + b * y) := +dvd_add (hdx.mul_left a) (hdy.mul_left b) + +end non_unital_comm_semiring + + +section semigroup + +variables [semigroup α] [has_distrib_neg α] {a b c : α} + +theorem dvd_neg_of_dvd (h : a ∣ b) : (a ∣ -b) := +let ⟨c, hc⟩ := h in ⟨-c, by simp [hc]⟩ + +theorem dvd_of_dvd_neg (h : a ∣ -b) : (a ∣ b) := +let t := dvd_neg_of_dvd h in by rwa neg_neg at t + +/-- An element a of a semigroup with a distributive negation divides the negation of an element b +iff a divides b. -/ +@[simp] lemma dvd_neg (a b : α) : (a ∣ -b) ↔ (a ∣ b) := +⟨dvd_of_dvd_neg, dvd_neg_of_dvd⟩ + +theorem neg_dvd_of_dvd (h : a ∣ b) : -a ∣ b := +let ⟨c, hc⟩ := h in ⟨-c, by simp [hc]⟩ + +theorem dvd_of_neg_dvd (h : -a ∣ b) : a ∣ b := +let t := neg_dvd_of_dvd h in by rwa neg_neg at t + +/-- The negation of an element a of a semigroup with a distributive negation divides +another element b iff a divides b. -/ +@[simp] lemma neg_dvd (a b : α) : (-a ∣ b) ↔ (a ∣ b) := +⟨dvd_of_neg_dvd, neg_dvd_of_dvd⟩ + +end semigroup + +section non_unital_ring +variables [non_unital_ring α] {a b c : α} + +theorem dvd_sub (h₁ : a ∣ b) (h₂ : a ∣ c) : a ∣ b - c := +by { rw sub_eq_add_neg, exact dvd_add h₁ (dvd_neg_of_dvd h₂) } + +theorem dvd_add_iff_left (h : a ∣ c) : a ∣ b ↔ a ∣ b + c := +⟨λh₂, dvd_add h₂ h, λH, by have t := dvd_sub H h; rwa add_sub_cancel at t⟩ + +theorem dvd_add_iff_right (h : a ∣ b) : a ∣ c ↔ a ∣ b + c := +by rw add_comm; exact dvd_add_iff_left h + +/-- If an element a divides another element c in a commutative ring, a divides the sum of another + element b with c iff a divides b. -/ +theorem dvd_add_left (h : a ∣ c) : a ∣ b + c ↔ a ∣ b := +(dvd_add_iff_left h).symm + +/-- If an element a divides another element b in a commutative ring, a divides the sum of b and + another element c iff a divides c. -/ +theorem dvd_add_right (h : a ∣ b) : a ∣ b + c ↔ a ∣ c := +(dvd_add_iff_right h).symm + +lemma dvd_iff_dvd_of_dvd_sub {a b c : α} (h : a ∣ (b - c)) : (a ∣ b ↔ a ∣ c) := +begin + split, + { intro h', + convert dvd_sub h' h, + exact eq.symm (sub_sub_self b c) }, + { intro h', + convert dvd_add h h', + exact eq_add_of_sub_eq rfl } +end + +end non_unital_ring + +section ring +variables [ring α] {a b c : α} + +theorem two_dvd_bit1 : 2 ∣ bit1 a ↔ (2 : α) ∣ 1 := (dvd_add_iff_right (@two_dvd_bit0 _ _ a)).symm + +/-- An element a divides the sum a + b if and only if a divides b.-/ +@[simp] lemma dvd_add_self_left {a b : α} : a ∣ a + b ↔ a ∣ b := +dvd_add_right (dvd_refl a) + +/-- An element a divides the sum b + a if and only if a divides b.-/ +@[simp] lemma dvd_add_self_right {a b : α} : a ∣ b + a ↔ a ∣ b := +dvd_add_left (dvd_refl a) + +end ring + +section non_unital_comm_ring +variables [non_unital_comm_ring α] {a b c : α} + +lemma dvd_mul_sub_mul {k a b x y : α} (hab : k ∣ a - b) (hxy : k ∣ x - y) : + k ∣ a * x - b * y := +begin + convert dvd_add (hxy.mul_left a) (hab.mul_right y), + rw [mul_sub_left_distrib, mul_sub_right_distrib], + simp only [sub_eq_add_neg, add_assoc, neg_add_cancel_left], +end + +end non_unital_comm_ring From 23d1b33cd318c0f0e3be6446e45ab66c3885b3dc Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 11 Oct 2022 09:55:50 +0000 Subject: [PATCH 683/828] feat(category_theory): categories where inclusions into coproducts are mono (#16867) --- src/category_theory/limits/mono_coprod.lean | 115 ++++++++++++++++++++ 1 file changed, 115 insertions(+) create mode 100644 src/category_theory/limits/mono_coprod.lean diff --git a/src/category_theory/limits/mono_coprod.lean b/src/category_theory/limits/mono_coprod.lean new file mode 100644 index 0000000000000..5ba59c1ad670b --- /dev/null +++ b/src/category_theory/limits/mono_coprod.lean @@ -0,0 +1,115 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import category_theory.limits.types +import category_theory.morphism_property + +/-! + +# Categories where inclusions into coproducts are monomorphisms + +If `C` is a category, the class `mono_coprod C` expresses that left +inclusions `A ⟶ A ⨿ B` are monomorphisms when `has_coproduct A B` +is satisfied. If it is so, it is shown that right inclusions are +also monomorphisms. + +TODO @joelriou: show that if `X : I → C` and `ι : J → I` is an injective map, +then the canonical morphism `∐ (X ∘ ι) ⟶ ∐ X` is a monomorphism. + +TODO: define distributive categories, and show that they satisfy `mono_coprod`, see + + +-/ + +noncomputable theory + +open category_theory category_theory.category category_theory.limits + +universe u + +namespace category_theory + +namespace limits + +variables (C : Type*) [category C] + +/-- This condition expresses that inclusion morphisms into coproducts are monomorphisms. -/ +class mono_coprod : Prop := +(binary_cofan_inl : ∀ ⦃A B : C⦄ (c : binary_cofan A B) (hc : is_colimit c), mono c.inl) + +variable {C} + +@[priority 100] +instance mono_coprod_of_has_zero_morphisms + [has_zero_morphisms C] : mono_coprod C := +⟨λ A B c hc, begin + haveI : is_split_mono c.inl := is_split_mono.mk' + (split_mono.mk (hc.desc (binary_cofan.mk (𝟙 A) 0)) (is_colimit.fac _ _ _)), + apply_instance, +end⟩ + +namespace mono_coprod + +lemma binary_cofan_inr {A B : C}[mono_coprod C] (c : binary_cofan A B) (hc : is_colimit c) : + mono c.inr := +begin + have hc' : is_colimit (binary_cofan.mk c.inr c.inl) := + binary_cofan.is_colimit_mk (λ s, hc.desc (binary_cofan.mk s.inr s.inl)) (by tidy) (by tidy) + (λ s m h₁ h₂, binary_cofan.is_colimit.hom_ext hc + (by simp only [h₂, is_colimit.fac, binary_cofan.ι_app_left, binary_cofan.mk_inl]) + (by simp only [h₁, is_colimit.fac, binary_cofan.ι_app_right, binary_cofan.mk_inr])), + exact binary_cofan_inl _ hc', +end + +instance {A B : C} [mono_coprod C] [has_binary_coproduct A B] : + mono (coprod.inl : A ⟶ A ⨿ B) := +binary_cofan_inl _ (colimit.is_colimit _) + +instance {A B : C} [mono_coprod C] [has_binary_coproduct A B] : + mono (coprod.inr : B ⟶ A ⨿ B) := +binary_cofan_inr _ (colimit.is_colimit _) + +lemma mono_inl_iff {A B : C} {c₁ c₂ : binary_cofan A B} (hc₁ : is_colimit c₁) + (hc₂ : is_colimit c₂) : mono c₁.inl ↔ mono c₂.inl := +begin + suffices : ∀ (c₁ c₂ : binary_cofan A B) (hc₁ : is_colimit c₁) (hc₂ : is_colimit c₂) + (h : mono c₁.inl), mono c₂.inl, + { exact ⟨λ h₁, this _ _ hc₁ hc₂ h₁, λ h₂, this _ _ hc₂ hc₁ h₂⟩, }, + intros c₁ c₂ hc₁ hc₂, + introI, + simpa only [is_colimit.comp_cocone_point_unique_up_to_iso_hom] + using mono_comp c₁.inl (hc₁.cocone_point_unique_up_to_iso hc₂).hom, +end + +lemma mk' (h : ∀ (A B : C), ∃ (c : binary_cofan A B) (hc : is_colimit c), mono c.inl) : + mono_coprod C := +⟨λ A B c' hc', begin + obtain ⟨c, hc₁, hc₂⟩ := h A B, + simpa only [mono_inl_iff hc' hc₁] using hc₂, +end⟩ + +instance mono_coprod_type : mono_coprod (Type u) := +mono_coprod.mk' (λ A B, begin + refine ⟨binary_cofan.mk (sum.inl : A ⟶ A ⊕ B) sum.inr, _, _⟩, + { refine binary_cofan.is_colimit.mk _ (λ Y f₁ f₂ x, by { cases x, exacts [f₁ x, f₂ x], }) + (λ Y f₁ f₂, rfl) (λ Y f₁ f₂, rfl) _, + intros Y f₁ f₂ m h₁ h₂, + ext x, + cases x, + { dsimp, exact congr_fun h₁ x, }, + { dsimp, exact congr_fun h₂ x, }, }, + { rw mono_iff_injective, + intros a₁ a₂ h, + simp only [binary_cofan.mk_inl] at h, + dsimp at h, + simpa only using h, }, +end) + +end mono_coprod + +end limits + +end category_theory From 556f35346ebc651dbc65ae6deb3a031f3d399e9d Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 11 Oct 2022 09:55:52 +0000 Subject: [PATCH 684/828] perf(algebra/category/Module/monoidal): fix timeout (#16903) elaboration of category_theory.monoidal_preadditive took 23.2s -> 2.93s elaboration of category_theory.monoidal_linear took 12.5s -> 2.89s [timeout in last bors batch](https://github.com/leanprover-community/mathlib/actions/runs/3223730534/jobs/5274107482) --- src/algebra/category/Module/monoidal.lean | 15 +++++++++------ 1 file changed, 9 insertions(+), 6 deletions(-) diff --git a/src/algebra/category/Module/monoidal.lean b/src/algebra/category/Module/monoidal.lean index dea21e486a38a..0f80c5596f23f 100644 --- a/src/algebra/category/Module/monoidal.lean +++ b/src/algebra/category/Module/monoidal.lean @@ -269,14 +269,17 @@ end monoidal_category open opposite instance : monoidal_preadditive (Module.{u} R) := -{ tensor_zero' := by { intros, ext, simp, }, - zero_tensor' := by { intros, ext, simp, }, - tensor_add' := by { intros, ext, simp [tensor_product.tmul_add], }, - add_tensor' := by { intros, ext, simp [tensor_product.add_tmul], }, } +by refine ⟨_, _, _, _⟩; dsimp only [auto_param]; intros; + refine tensor_product.ext (linear_map.ext $ λ x, linear_map.ext $ λ y, _); + simp only [linear_map.compr₂_apply, tensor_product.mk_apply, monoidal_category.hom_apply, + linear_map.zero_apply, tensor_product.tmul_zero, tensor_product.zero_tmul, + linear_map.add_apply, tensor_product.tmul_add, tensor_product.add_tmul] instance : monoidal_linear R (Module.{u} R) := -{ tensor_smul' := by { intros, ext, simp, }, - smul_tensor' := by { intros, ext, simp [tensor_product.smul_tmul], }, } +by refine ⟨_, _⟩; dsimp only [auto_param]; intros; + refine tensor_product.ext (linear_map.ext $ λ x, linear_map.ext $ λ y, _); + simp only [linear_map.compr₂_apply, tensor_product.mk_apply, monoidal_category.hom_apply, + linear_map.smul_apply, tensor_product.tmul_smul, tensor_product.smul_tmul] /-- Auxiliary definition for the `monoidal_closed` instance on `Module R`. From 348289cc3e8ea3f1dcddad5fd52c10156435b17f Mon Sep 17 00:00:00 2001 From: Amelia Livingston <101damnations@github.com> Date: Tue, 11 Oct 2022 13:09:26 +0000 Subject: [PATCH 685/828] feat(representation_theory/group_cohomology_resolution): add chain complex underlying the standard resolution (#16258) --- src/algebra/algebra/basic.lean | 4 + src/representation_theory/Action.lean | 56 +++++- src/representation_theory/Rep.lean | 45 +++++ .../group_cohomology_resolution.lean | 170 ++++++++++++++++-- 4 files changed, 264 insertions(+), 11 deletions(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index 29a545f6e3660..c4eda7542b773 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -1726,6 +1726,10 @@ by rw [←(one_smul A m), ←smul_assoc, algebra.smul_def, mul_one, one_smul] @[simp] lemma algebra_map_smul (r : R) (m : M) : ((algebra_map R A) r) • m = r • m := (algebra_compatible_smul A r m).symm +lemma int_cast_smul {k V : Type*} [comm_ring k] [add_comm_group V] [module k V] (r : ℤ) (x : V) : + (r : k) • x = r • x := +algebra_map_smul k r x + lemma no_zero_smul_divisors.trans (R A M : Type*) [comm_ring R] [ring A] [is_domain A] [algebra R A] [add_comm_group M] [module R M] [module A M] [is_scalar_tower R A M] [no_zero_smul_divisors R A] [no_zero_smul_divisors A M] : no_zero_smul_divisors R M := diff --git a/src/representation_theory/Action.lean b/src/representation_theory/Action.lean index 85475cb9c4d61..feacd272fbe35 100644 --- a/src/representation_theory/Action.lean +++ b/src/representation_theory/Action.lean @@ -33,7 +33,7 @@ and construct the restriction functors `res {G H : Mon} (f : G ⟶ H) : Action V * When `V` is preadditive, linear, or abelian so is `Action V G`. -/ -universes u +universes u v open category_theory open category_theory.limits @@ -74,6 +74,8 @@ variable (G : Mon.{u}) section +instance inhabited' : inhabited (Action (Type u) G) := ⟨⟨punit, 1⟩⟩ + /-- The trivial representation of a group. -/ def trivial : Action AddCommGroup G := { V := AddCommGroup.of punit, @@ -288,10 +290,13 @@ instance : preadditive (Action V G) := comp_add' := by { intros, ext, exact preadditive.comp_add _ _ _ _ _ _, }, } instance : functor.additive (functor_category_equivalence V G).functor := {} +instance forget_additive : functor.additive (forget V G) := {} @[simp] lemma zero_hom {X Y : Action V G} : (0 : X ⟶ Y).hom = 0 := rfl @[simp] lemma neg_hom {X Y : Action V G} (f : X ⟶ Y) : (-f).hom = -f.hom := rfl @[simp] lemma add_hom {X Y : Action V G} (f g : X ⟶ Y) : (f + g).hom = f.hom + g.hom := rfl +@[simp] lemma sum_hom {X Y : Action V G} {ι : Type*} (f : ι → (X ⟶ Y)) (s : finset ι) : + (s.sum f).hom = s.sum (λ i, (f i).hom) := (forget V G).map_sum f s end preadditive @@ -521,6 +526,55 @@ variables {R : Type*} [semiring R] instance res_linear [preadditive V] [linear R V] : (res V f).linear R := {} +/-- Bundles a type `H` with a multiplicative action of `G` as an `Action`. -/ +def of_mul_action (G H : Type u) [monoid G] [mul_action G H] : Action (Type u) (Mon.of G) := +{ V := H, + ρ := @mul_action.to_End_hom _ _ _ (by assumption) } + +@[simp] lemma of_mul_action_apply {G H : Type u} [monoid G] [mul_action G H] (g : G) (x : H) : + (of_mul_action G H).ρ g x = (g • x : H) := +rfl + +/-- Given a family `F` of types with `G`-actions, this is the limit cone demonstrating that the +product of `F` as types is a product in the category of `G`-sets. -/ +def of_mul_action_limit_cone {ι : Type v} (G : Type (max v u)) [monoid G] + (F : ι → Type (max v u)) [Π i : ι, mul_action G (F i)] : + limit_cone (discrete.functor (λ i : ι, Action.of_mul_action G (F i))) := +{ cone := + { X := Action.of_mul_action G (Π i : ι, F i), + π := + { app := λ i, ⟨λ x, x i.as, λ g, by ext; refl⟩, + naturality' := λ i j x, + begin + ext, + discrete_cases, + cases x, + congr + end } }, + is_limit := + { lift := λ s, + { hom := λ x i, (s.π.app ⟨i⟩).hom x, + comm' := λ g, + begin + ext x j, + dsimp, + exact congr_fun ((s.π.app ⟨j⟩).comm g) x, + end }, + fac' := λ s j, + begin + ext, + dsimp, + congr, + rw discrete.mk_as, + end, + uniq' := λ s f h, + begin + ext x j, + dsimp at *, + rw ←h ⟨j⟩, + congr, + end } } + end Action namespace category_theory.functor diff --git a/src/representation_theory/Rep.lean b/src/representation_theory/Rep.lean index 21b1c5278e5d3..376a1fc0447e1 100644 --- a/src/representation_theory/Rep.lean +++ b/src/representation_theory/Rep.lean @@ -8,6 +8,7 @@ import representation_theory.Action import algebra.category.Module.abelian import algebra.category.Module.colimits import algebra.category.Module.monoidal +import algebra.category.Module.adjunctions /-! # `Rep k G` is the category of `k`-linear representations of `G`. @@ -71,6 +72,50 @@ by apply_instance noncomputable example : preserves_colimits (forget₂ (Rep k G) (Module.{u} k)) := by apply_instance +section linearization + +variables (k G) + +/-- The monoidal functor sending a type `H` with a `G`-action to the induced `k`-linear +`G`-representation on `k[H].` -/ +noncomputable def linearization : + monoidal_functor (Action (Type u) (Mon.of G)) (Rep k G) := +(Module.monoidal_free k).map_Action (Mon.of G) + +variables {k G} + +@[simp] lemma linearization_obj_ρ (X : Action (Type u) (Mon.of G)) (g : G) (x : X.V →₀ k) : + ((linearization k G).1.1.obj X).ρ g x = finsupp.lmap_domain k k (X.ρ g) x := rfl + +@[simp] lemma linearization_of (X : Action (Type u) (Mon.of G)) (g : G) (x : X.V) : + ((linearization k G).1.1.obj X).ρ g (finsupp.single x (1 : k)) + = finsupp.single (X.ρ g x) (1 : k) := +by rw [linearization_obj_ρ, finsupp.lmap_domain_apply, finsupp.map_domain_single] + +variables (X Y : Action (Type u) (Mon.of G)) (f : X ⟶ Y) + +@[simp] lemma linearization_map_hom : + ((linearization k G).1.1.map f).hom = finsupp.lmap_domain k k f.hom := rfl + +lemma linearization_map_hom_of (x : X.V) : + ((linearization k G).1.1.map f).hom (finsupp.single x (1 : k)) + = finsupp.single (f.hom x) (1 : k) := +by rw [linearization_map_hom, finsupp.lmap_domain_apply, finsupp.map_domain_single] + +variables (k G) + +/-- Given a `G`-action on `H`, this is `k[H]` bundled with the natural representation +`G →* End(k[H])` as a term of type `Rep k G`. -/ +noncomputable abbreviation of_mul_action (H : Type u) [mul_action G H] : Rep k G := +of $ representation.of_mul_action k G H + +/-- The linearization of a type `H` with a `G`-action is definitionally isomorphic to the +`k`-linear `G`-representation on `k[H]` induced by the `G`-action on `H`. -/ +noncomputable def linearization_of_mul_action_iso (n : ℕ) : + (linearization k G).1.1.obj (Action.of_mul_action G (fin n → G)) + ≅ of_mul_action k G (fin n → G) := iso.refl _ + +end linearization end Rep /-! diff --git a/src/representation_theory/group_cohomology_resolution.lean b/src/representation_theory/group_cohomology_resolution.lean index 0c953185cae1c..d2eb3d6ca7ad6 100644 --- a/src/representation_theory/group_cohomology_resolution.lean +++ b/src/representation_theory/group_cohomology_resolution.lean @@ -3,8 +3,11 @@ Copyright (c) 2022 Amelia Livingston. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Amelia Livingston -/ -import representation_theory.Rep + +import algebraic_topology.alternating_face_map_complex +import algebraic_topology.cech_nerve import representation_theory.basic +import representation_theory.Rep /-! # The structure of the `k[G]`-module `k[Gⁿ]` @@ -19,6 +22,12 @@ In particular, we define an isomorphism of `k`-linear `G`-representations betwee This allows us to define a `k[G]`-basis on `k[Gⁿ⁺¹]`, by mapping the natural `k[G]`-basis of `k[G] ⊗ₖ k[Gⁿ]` along the isomorphism. +We then define the standard resolution of `k` as a trivial representation, by +taking the alternating face map complex associated to an appropriate simplicial `k`-linear +`G`-representation. This simplicial object is the `linearization` of the simplicial `G`-set given +by the universal cover of the classifying space of `G`. We prove this `G`-set is isomorphic to the +Čech nerve of the natural arrow of `G`-sets `G ⟶ {pt}`. + ## Main definitions * `group_cohomology.resolution.to_tensor` @@ -26,6 +35,9 @@ This allows us to define a `k[G]`-basis on `k[Gⁿ⁺¹]`, by mapping the natura * `Rep.of_mul_action` * `group_cohomology.resolution.equiv_tensor` * `group_cohomology.resolution.of_mul_action_basis` + * `classifying_space_universal_cover` + * `classifying_space_universal_cover.cech_nerve_iso` + * `group_cohomology.resolution` ## TODO @@ -47,11 +59,12 @@ over `k`. noncomputable theory -universes u +universes u v w variables {k G : Type u} [comm_ring k] {n : ℕ} open_locale tensor_product +open category_theory local notation `Gⁿ` := fin n → G local notation `Gⁿ⁺¹` := fin (n + 1) → G @@ -60,6 +73,7 @@ namespace group_cohomology.resolution open finsupp (hiding lift) fin (partial_prod) representation +section basis variables (k G n) [group G] /-- The `k`-linear map from `k[Gⁿ⁺¹]` to `k[G] ⊗ₖ k[Gⁿ]` sending `(g₀, ..., gₙ)` @@ -132,12 +146,6 @@ end variables (k G n) -/-- Given a `G`-action on `H`, this is `k[H]` bundled with the natural representation -`G →* End(k[H])` as a term of type `Rep k G`. -/ -abbreviation _root_.Rep.of_mul_action (G : Type u) [monoid G] (H : Type u) [mul_action G H] : - Rep k G := -Rep.of $ representation.of_mul_action k G H - /-- A hom of `k`-linear representations of `G` from `k[Gⁿ⁺¹]` to `k[G] ⊗ₖ k[Gⁿ]` (on which `G` acts by `ρ(g₁)(g₂ ⊗ x) = (g₁ * g₂) ⊗ x`) sending `(g₀, ..., gₙ)` to `g₀ ⊗ (g₀⁻¹g₁, g₁⁻¹g₂, ..., gₙ₋₁⁻¹gₙ)`. -/ @@ -184,7 +192,6 @@ Action.mk_iso (linear_equiv.to_Module_iso right_inv := λ x, by convert to_tensor_aux_right_inv x, ..to_tensor_aux k G n }) (to_tensor k G n).comm --- not quite sure which simp lemmas to make here @[simp] lemma equiv_tensor_def : (equiv_tensor k G n).hom = to_tensor k G n := rfl @@ -216,10 +223,153 @@ def of_mul_action_basis : basis (fin n → G) (monoid_algebra k G) (of_mul_action k G (fin (n + 1) → G)).as_module := @basis.map _ (monoid_algebra k G) (monoid_algebra k G ⊗[k] ((fin n → G) →₀ k)) _ _ _ _ _ _ (@algebra.tensor_product.basis k _ (monoid_algebra k G) _ _ ((fin n → G) →₀ k) _ _ - (fin n → G) (⟨linear_equiv.refl k _⟩)) (of_mul_action_basis_aux k G n) + (fin n → G) ⟨linear_equiv.refl k _⟩) (of_mul_action_basis_aux k G n) lemma of_mul_action_free : module.free (monoid_algebra k G) (of_mul_action k G (fin (n + 1) → G)).as_module := module.free.of_basis (of_mul_action_basis k G n) +end basis +end group_cohomology.resolution +variables (G) + +/-- The simplicial `G`-set sending `[n]` to `Gⁿ⁺¹` equipped with the diagonal action of `G`. -/ +def classifying_space_universal_cover [monoid G] : + simplicial_object (Action (Type u) $ Mon.of G) := +{ obj := λ n, Action.of_mul_action G (fin (n.unop.len + 1) → G), + map := λ m n f, + { hom := λ x, x ∘ f.unop.to_order_hom, + comm' := λ g, rfl }, + map_id' := λ n, rfl, + map_comp' := λ i j k f g, rfl } + +namespace classifying_space_universal_cover + +open category_theory.limits category_theory.arrow + +variables (ι : Type w) {C : Type u} [category.{v} C] [has_terminal C] + +/-- The diagram `option ι ⥤ C` sending `none` to the terminal object and `some j` to `X`. -/ +def wide_cospan (X : C) : wide_pullback_shape ι ⥤ C := +wide_pullback_shape.wide_cospan (terminal C) (λ i : ι, X) (λ i, terminal.from X) + +instance unique_to_wide_cospan_none (X Y : C) : unique (Y ⟶ (wide_cospan ι X).obj none) := +by unfold wide_cospan; dsimp; apply_instance + +variables [has_finite_products C] + +/-- The product `Xᶥ` is the vertex of a limit cone on `wide_cospan ι X`. -/ +def wide_cospan.limit_cone [fintype ι] (X : C) : limit_cone (wide_cospan ι X) := +{ cone := + { X := ∏ (λ i : ι, X), + π := + { app := λ X, option.cases_on X (terminal.from _) (λ i, limit.π _ ⟨i⟩), + naturality' := λ i j f, + begin + cases f, + { cases i, + all_goals { dsimp, simp }}, + { dsimp, + simp only [terminal.comp_from], + exact subsingleton.elim _ _ } + end } }, + is_limit := + { lift := λ s, limits.pi.lift (λ j, s.π.app (some j)), + fac' := λ s j, option.cases_on j (subsingleton.elim _ _) (λ j, limit.lift_π _ _), + uniq' := λ s f h, + begin + ext j, + dunfold limits.pi.lift, + rw limit.lift_π, + dsimp, + rw ←h (some j.as), + congr, + ext, + refl, + end } } + +instance has_wide_pullback [finite ι] (X : C) : + has_wide_pullback (arrow.mk (terminal.from X)).right + (λ i : ι, (arrow.mk (terminal.from X)).left) (λ i, (arrow.mk (terminal.from X)).hom) := +begin + casesI nonempty_fintype ι, + exact ⟨⟨wide_cospan.limit_cone ι X⟩⟩, +end + +variables [monoid G] (m : simplex_categoryᵒᵖ) + +/-- The `m`th object of the Čech nerve of the `G`-set hom `G ⟶ {pt}` is isomorphic to `Gᵐ⁺¹` with +the diagonal action. -/ +def cech_nerve_obj_iso : + (cech_nerve (arrow.mk (terminal.from (Action.of_mul_action G G)))).obj m + ≅ Action.of_mul_action G (fin (m.unop.len + 1) → G) := +(is_limit.cone_point_unique_up_to_iso (limit.is_limit _) (wide_cospan.limit_cone + (fin (m.unop.len + 1)) (Action.of_mul_action G G)).2).trans $ + limit.iso_limit_cone (Action.of_mul_action_limit_cone _ _) + +/-- The Čech nerve of the `G`-set hom `G ⟶ {pt}` is naturally isomorphic to `EG`, the universal +cover of the classifying space of `G` as a simplicial `G`-set. -/ +def cech_nerve_iso : cech_nerve (arrow.mk (terminal.from (Action.of_mul_action G G))) + ≅ classifying_space_universal_cover G := +(@nat_iso.of_components _ _ _ _ (classifying_space_universal_cover G) + (cech_nerve (arrow.mk (terminal.from (Action.of_mul_action G G)))) +(λ n, (cech_nerve_obj_iso G n).symm) $ λ m n f, wide_pullback.hom_ext + (λ j, (arrow.mk (terminal.from (Action.of_mul_action G G))).hom) _ _ +(λ j, begin + simp only [category.assoc], + dsimp [cech_nerve_obj_iso], + rw [wide_pullback.lift_π, category.assoc], + erw [limit.cone_point_unique_up_to_iso_hom_comp (wide_cospan.limit_cone _ _).2, + limit.cone_point_unique_up_to_iso_hom_comp (Action.of_mul_action_limit_cone _ _).2, + category.assoc], + dunfold wide_pullback.π, + erw [limit.cone_point_unique_up_to_iso_inv_comp, limit.iso_limit_cone_inv_π], + refl, +end) +(@subsingleton.elim _ (@unique.subsingleton _ (limits.unique_to_terminal _)) _ _)).symm + +end classifying_space_universal_cover +namespace group_cohomology.resolution + +section differential +variables (k G) +/-- The `k`-linear map underlying the differential in the standard resolution of `k` as a trivial +`k`-linear `G`-representation. It sends `(g₀, ..., gₙ) ↦ ∑ (-1)ⁱ • (g₀, ..., ĝᵢ, ..., gₙ)`. -/ +def d (n : ℕ) : ((fin (n + 1) → G) →₀ k) →ₗ[k] ((fin n → G) →₀ k) := +finsupp.lift ((fin n → G) →₀ k) k (fin (n + 1) → G) (λ g, (@finset.univ (fin (n + 1)) _).sum + (λ p, finsupp.single (g ∘ p.succ_above) ((-1 : k) ^ (p : ℕ)))) + +variables {k G} + +@[simp] lemma d_of {n : ℕ} (c : fin (n + 1) → G) : + d k G n (finsupp.single c 1) = finset.univ.sum (λ p : fin (n + 1), finsupp.single + (c ∘ p.succ_above) ((-1 : k) ^ (p : ℕ))) := +by simp [d] + +end differential +end group_cohomology.resolution +variables (k G) [monoid G] + +/-- The standard resolution of `k` as a trivial representation, defined as the alternating +face map complex of a simplicial `k`-linear `G`-representation. -/ +def group_cohomology.resolution := (algebraic_topology.alternating_face_map_complex (Rep k G)).obj + (classifying_space_universal_cover G ⋙ (Rep.linearization k G).1.1) + +namespace group_cohomology.resolution + +/-- The `n`th object of the standard resolution of `k` is definitionally isomorphic to `k[Gⁿ⁺¹]` +equipped with the representation induced by the diagonal action of `G`. -/ +def X_iso (n : ℕ) : + (group_cohomology.resolution k G).X n ≅ Rep.of_mul_action k G (fin (n + 1) → G) := iso.refl _ + +/-- Simpler expression for the differential in the standard resolution of `k` as a +`G`-representation. It sends `(g₀, ..., gₙ₊₁) ↦ ∑ (-1)ⁱ • (g₀, ..., ĝᵢ, ..., gₙ₊₁)`. -/ +theorem d_eq (n : ℕ) : + ((group_cohomology.resolution k G).d (n + 1) n).hom = d k G (n + 1) := +begin + ext x y, + dsimp [group_cohomology.resolution], + simpa [←@int_cast_smul k, simplicial_object.δ], +end + end group_cohomology.resolution From c80bc30e35e88dfdd51589b4126e65e2ee101297 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 11 Oct 2022 13:09:27 +0000 Subject: [PATCH 686/828] feat(analysis/locally_convex/with_seminorms): the topology induced by a family of seminorms is the infimum of the ones induced by each seminorm (#16669) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Yaël Dillies --- .../locally_convex/with_seminorms.lean | 36 +++++++++++++++++-- 1 file changed, 34 insertions(+), 2 deletions(-) diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index d295bea227674..9d64fcd3f0bff 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -288,14 +288,16 @@ end topology section topological_add_group variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] -variables [topological_space E] [topological_add_group E] +variables [t : topological_space E] [topological_add_group E] variables [nonempty ι] +include t + lemma seminorm_family.with_seminorms_of_nhds (p : seminorm_family 𝕜 E ι) (h : 𝓝 (0 : E) = p.module_filter_basis.to_filter_basis.filter) : with_seminorms p := begin - refine ⟨topological_add_group.ext (by apply_instance) + refine ⟨topological_add_group.ext infer_instance (p.add_group_filter_basis.is_topological_add_group) _⟩, rw add_group_filter_basis.nhds_zero_eq, exact h, @@ -325,6 +327,36 @@ begin exact filter.mem_infi_of_mem i (filter.preimage_mem_comap $ metric.ball_mem_nhds _ one_pos) end +/-- The topology induced by a family of seminorms is exactly the infimum of the ones induced by +each seminorm individually. We express this as a characterization of `with_seminorms p`. -/ +lemma seminorm_family.with_seminorms_iff_topological_space_eq_infi (p : seminorm_family 𝕜 E ι) : + with_seminorms p ↔ t = ⨅ i, (p i).to_add_group_seminorm.to_seminormed_add_comm_group + .to_uniform_space.to_topological_space := +begin + rw [p.with_seminorms_iff_nhds_eq_infi, topological_add_group.ext_iff infer_instance + (topological_add_group_infi $ λ i, infer_instance), nhds_infi], + congrm (_ = ⨅ i, _), + exact @comap_norm_nhds_zero _ (p i).to_add_group_seminorm.to_seminormed_add_group, + all_goals {apply_instance} +end + +omit t + +/-- The uniform structure induced by a family of seminorms is exactly the infimum of the ones +induced by each seminorm individually. We express this as a characterization of +`with_seminorms p`. -/ +lemma seminorm_family.with_seminorms_iff_uniform_space_eq_infi [u : uniform_space E] + [uniform_add_group E] (p : seminorm_family 𝕜 E ι) : + with_seminorms p ↔ u = ⨅ i, (p i).to_add_group_seminorm.to_seminormed_add_comm_group + .to_uniform_space := +begin + rw [p.with_seminorms_iff_nhds_eq_infi, uniform_add_group.ext_iff infer_instance + (uniform_add_group_infi $ λ i, infer_instance), to_topological_space_infi, nhds_infi], + congrm (_ = ⨅ i, _), + exact @comap_norm_nhds_zero _ (p i).to_add_group_seminorm.to_seminormed_add_group, + all_goals {apply_instance} +end + end topological_add_group section normed_space From 695b3191d1d49fb6624694dbfef9bc1303bfcfb6 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Tue, 11 Oct 2022 13:09:28 +0000 Subject: [PATCH 687/828] feat(measure_theory/measure): add a measure_space instance for subtype (#16708) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Using the comap definition added in #15343. This was predominantly written by Rémy Degenne, and has applications in #2819. Co-authored-by: Rémy Degenne Co-authored-by: RemyDegenne --- src/measure_theory/measure/measure_space.lean | 118 ++++++++++++++++++ 1 file changed, 118 insertions(+) diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 8ebb80a41982f..eaf74338f9bc9 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -1066,6 +1066,11 @@ begin rw [to_measure_apply₀ _ _ hs, outer_measure.comap_apply, coe_to_outer_measure] end +lemma le_comap_apply {β} [measurable_space α] {mβ : measurable_space β} (f : α → β) (μ : measure β) + (hfi : injective f) (hf : ∀ s, measurable_set s → null_measurable_set (f '' s) μ) (s : set α) : + μ (f '' s) ≤ comap f μ s := +by { rw [comap, dif_pos (and.intro hfi hf)], exact le_to_measure_apply _ _ _, } + lemma comap_apply {β} [measurable_space α] {mβ : measurable_space β} (f : α → β) (hfi : injective f) (hf : ∀ s, measurable_set s → measurable_set (f '' s)) (μ : measure β) (hs : measurable_set s) : comap f μ s = μ (f '' s) := @@ -1077,6 +1082,119 @@ lemma comapₗ_eq_comap {β} [measurable_space α] {mβ : measurable_space β} ( comapₗ f μ s = comap f μ s := (comapₗ_apply f hfi hf μ hs).trans (comap_apply f hfi hf μ hs).symm +lemma measure_image_eq_zero_of_comap_eq_zero {β} [measurable_space α] {mβ : measurable_space β} + (f : α → β) (μ : measure β) (hfi : injective f) + (hf : ∀ s, measurable_set s → null_measurable_set (f '' s) μ) {s : set α} (hs : comap f μ s = 0) : + μ (f '' s) = 0 := +le_antisymm ((le_comap_apply f μ hfi hf s).trans hs.le) (zero_le _) + +lemma ae_eq_image_of_ae_eq_comap {β} [measurable_space α] {mβ : measurable_space β} + (f : α → β) (μ : measure β) (hfi : injective f) + (hf : ∀ s, measurable_set s → null_measurable_set (f '' s) μ) {s t : set α} + (hst : s =ᵐ[comap f μ] t) : + f '' s =ᵐ[μ] f '' t := +begin + rw [eventually_eq, ae_iff] at hst ⊢, + have h_eq_α : {a : α | ¬s a = t a} = s \ t ∪ t \ s, + { ext1 x, simp only [eq_iff_iff, mem_set_of_eq, mem_union, mem_diff], tauto, }, + have h_eq_β : {a : β | ¬(f '' s) a = (f '' t) a} = f '' s \ f '' t ∪ f '' t \ f '' s, + { ext1 x, simp only [eq_iff_iff, mem_set_of_eq, mem_union, mem_diff], tauto, }, + rw [← set.image_diff hfi, ← set.image_diff hfi, ← set.image_union] at h_eq_β, + rw h_eq_β, + rw h_eq_α at hst, + exact measure_image_eq_zero_of_comap_eq_zero f μ hfi hf hst, +end + +lemma null_measurable_set.image {β} [measurable_space α] {mβ : measurable_space β} + (f : α → β) (μ : measure β) (hfi : injective f) + (hf : ∀ s, measurable_set s → null_measurable_set (f '' s) μ) {s : set α} + (hs : null_measurable_set s (μ.comap f)) : + null_measurable_set (f '' s) μ := +begin + refine ⟨to_measurable μ (f '' (to_measurable (μ.comap f) s)), + measurable_set_to_measurable _ _, _⟩, + refine eventually_eq.trans _ (null_measurable_set.to_measurable_ae_eq _).symm, + swap, { exact hf _ (measurable_set_to_measurable _ _), }, + have h : to_measurable (comap f μ) s =ᵐ[comap f μ] s, + from @null_measurable_set.to_measurable_ae_eq _ _ (μ.comap f : measure α) s hs, + exact ae_eq_image_of_ae_eq_comap f μ hfi hf h.symm, +end + +section subtype + +/-! ### Subtype of a measure space -/ + +section comap_any_measure + +lemma measurable_set.null_measurable_set_subtype_coe + {t : set s} (hs : null_measurable_set s μ) (ht : measurable_set t) : + null_measurable_set ((coe : s → α) '' t) μ := +begin + rw [subtype.measurable_space, comap_eq_generate_from] at ht, + refine generate_from_induction + (λ t : set s, null_measurable_set (coe '' t) μ) + {t : set s | ∃ (s' : set α), measurable_set s' ∧ coe ⁻¹' s' = t} _ _ _ _ ht, + { rintros t' ⟨s', hs', rfl⟩, + rw [subtype.image_preimage_coe], + exact hs'.null_measurable_set.inter hs, }, + { simp only [image_empty, null_measurable_set_empty], }, + { intro t', + simp only [←range_diff_image subtype.coe_injective, subtype.range_coe_subtype, set_of_mem_eq], + exact hs.diff, }, + { intro f, + rw image_Union, + exact null_measurable_set.Union, }, +end + +lemma null_measurable_set.subtype_coe {t : set s} (hs : null_measurable_set s μ) + (ht : null_measurable_set t (μ.comap subtype.val)) : + null_measurable_set ((coe : s → α) '' t) μ := +null_measurable_set.image coe μ subtype.coe_injective + (λ t, measurable_set.null_measurable_set_subtype_coe hs) ht + +lemma measure_subtype_coe_le_comap (hs : null_measurable_set s μ) (t : set s) : + μ ((coe : s → α) '' t) ≤ μ.comap subtype.val t := +le_comap_apply _ _ subtype.coe_injective (λ t, measurable_set.null_measurable_set_subtype_coe hs) _ + +lemma measure_subtype_coe_eq_zero_of_comap_eq_zero (hs : null_measurable_set s μ) + {t : set s} (ht : μ.comap subtype.val t = 0) : + μ ((coe : s → α) '' t) = 0 := +eq_bot_iff.mpr $ (measure_subtype_coe_le_comap hs t).trans ht.le + +end comap_any_measure + +section measure_space +variables [measure_space α] {p : α → Prop} + +instance subtype.measure_space : measure_space (subtype p) := +{ volume := measure.comap subtype.val volume, + ..subtype.measurable_space } + +lemma subtype.volume_def : (volume : measure s) = volume.comap subtype.val := rfl + +lemma subtype.volume_univ (hs : null_measurable_set s) : + volume (univ : set s) = volume s := +begin + rw [subtype.volume_def, comap_apply₀ _ _ _ _ measurable_set.univ.null_measurable_set], + { congr, simp only [subtype.val_eq_coe, image_univ, subtype.range_coe_subtype, set_of_mem_eq], }, + { exact subtype.coe_injective, }, + { exact λ t, measurable_set.null_measurable_set_subtype_coe hs, }, +end + +lemma volume_subtype_coe_le_volume (hs : null_measurable_set s) (t : set s) : + volume ((coe : s → α) '' t) ≤ volume t := +measure_subtype_coe_le_comap hs t + +lemma volume_subtype_coe_eq_zero_of_volume_eq_zero (hs : null_measurable_set s) + {t : set s} (ht : volume t = 0) : + volume ((coe : s → α) '' t) = 0 := +measure_subtype_coe_eq_zero_of_comap_eq_zero hs ht + +end measure_space + +end subtype + + /-! ### Restricting a measure -/ /-- Restrict a measure `μ` to a set `s` as an `ℝ≥0∞`-linear map. -/ From 94fc87ee359ec535ccb659f36c59e4a9a3664843 Mon Sep 17 00:00:00 2001 From: bottine Date: Tue, 11 Oct 2022 13:09:30 +0000 Subject: [PATCH 688/828] feat(category_theory/groupoid/free): functoriality of the free groupoid plus some cleanup (#16907) Add lemmas witnessing functoriality of the free groupoid construction, and clean up some style/doc issues. Co-authored-by: Junyan Xu --- .../groupoid/free_groupoid.lean | 62 +++++++++++++------ src/category_theory/path_category.lean | 20 +++--- src/category_theory/quotient.lean | 2 +- src/combinatorics/quiver/basic.lean | 2 +- .../quiver/connected_component.lean | 10 +-- 5 files changed, 60 insertions(+), 36 deletions(-) diff --git a/src/category_theory/groupoid/free_groupoid.lean b/src/category_theory/groupoid/free_groupoid.lean index 87c0f3d05bbdd..5fe970564e6db 100644 --- a/src/category_theory/groupoid/free_groupoid.lean +++ b/src/category_theory/groupoid/free_groupoid.lean @@ -37,7 +37,6 @@ and finally quotienting by the reducibility relation. -/ - open set classical function relation local attribute [instance] prop_decidable @@ -45,7 +44,7 @@ namespace category_theory namespace groupoid namespace free -universes u v u' v' +universes u v u' v' u'' v'' variables {V : Type u} [quiver.{v+1} V] @@ -71,16 +70,16 @@ inductive red_step : hom_rel (paths (quiver.symmetrify V)) red_step (𝟙 X) (f.to_path ≫ (quiver.reverse f).to_path) /-- The underlying vertices of the free groupoid -/ -def free_groupoid (V) [Q : quiver.{v+1} V] := quotient (@red_step V Q) +def _root_.category_theory.free_groupoid (V) [Q : quiver.{v+1} V] := quotient (@red_step V Q) instance {V} [Q : quiver.{v+1} V] [h : nonempty V] : nonempty (free_groupoid V) := ⟨⟨h.some⟩⟩ lemma congr_reverse {X Y : paths $ quiver.symmetrify V} (p q : X ⟶ Y) : quotient.comp_closure red_step p q → - quotient.comp_closure red_step (p.reverse) (q.reverse) := + quotient.comp_closure red_step (p.reverse) (q.reverse) := begin - rintros ⟨U, W, XW, pp, qq, WY, ⟨_, Z, f⟩⟩, - have : quotient.comp_closure red_step (WY.reverse ≫ 𝟙 _ ≫ XW.reverse) + rintros ⟨U, W, XW, pp, qq, WY, _, Z, f⟩, + have : quotient.comp_closure red_step (WY.reverse ≫ 𝟙 _ ≫ XW.reverse) (WY.reverse ≫ (f.to_path ≫ (quiver.reverse f).to_path) ≫ XW.reverse), { apply quotient.comp_closure.intro, apply red_step.step, }, @@ -128,16 +127,16 @@ quot.lift_on f (λ pp qq con, quot.sound $ congr_reverse pp qq con) instance : groupoid (free_groupoid V) := -{ inv := λ X Y f, quot_inv f -, inv_comp' := λ X Y p, quot.induction_on p $ λ pp, congr_reverse_comp pp -, comp_inv' := λ X Y p, quot.induction_on p $ λ pp, congr_comp_reverse pp } +{ inv := λ X Y f, quot_inv f, + inv_comp' := λ X Y p, quot.induction_on p $ λ pp, congr_reverse_comp pp, + comp_inv' := λ X Y p, quot.induction_on p $ λ pp, congr_comp_reverse pp } /-- The inclusion of the quiver on `V` to the underlying quiver on `free_groupoid V`-/ -def of : prefunctor V (free_groupoid V) := -{ obj := λ X, ⟨X⟩ -, map := λ X Y f, quot.mk _ f.to_pos_path} +def of (V) [quiver.{v+1} V] : prefunctor V (free_groupoid V) := +{ obj := λ X, ⟨X⟩, + map := λ X Y f, quot.mk _ f.to_pos_path } -lemma of_eq : of = +lemma of_eq : of V = ((quiver.symmetrify.of).comp paths.of).comp (quotient.functor $ @red_step V _).to_prefunctor := begin @@ -161,19 +160,19 @@ quotient.lift _ symmetry, apply groupoid.comp_inv, }) -lemma lift_spec (φ : prefunctor V V') : of.comp (lift φ).to_prefunctor = φ := +lemma lift_spec (φ : prefunctor V V') : (of V).comp (lift φ).to_prefunctor = φ := begin rw [of_eq, prefunctor.comp_assoc, prefunctor.comp_assoc, functor.to_prefunctor_comp], dsimp [lift], rw [quotient.lift_spec, paths.lift_spec, quiver.symmetrify.lift_spec], end -lemma lift_unique_spec (φ : prefunctor V V') (Φ : free_groupoid V ⥤ V') - (hΦ : of.comp Φ.to_prefunctor = φ) : Φ = (lift φ) := +lemma lift_unique (φ : prefunctor V V') (Φ : free_groupoid V ⥤ V') + (hΦ : (of V).comp Φ.to_prefunctor = φ) : Φ = (lift φ) := begin - apply quotient.lift_spec_unique, - apply paths.lift_spec_unique, - apply quiver.symmetrify.lift_spec_unique, + apply quotient.lift_unique, + apply paths.lift_unique, + apply quiver.symmetrify.lift_unique, { rw ←functor.to_prefunctor_comp, exact hΦ, }, { rintros X Y f, simp [←functor.to_prefunctor_comp,prefunctor.comp_map, paths.of_map, inv_eq_inv], @@ -185,6 +184,31 @@ end end universal_property +section functoriality + +variables {V' : Type u'} [quiver.{v'+1} V'] {V'' : Type u''} [quiver.{v''+1} V''] + +/-- The functor of free groupoid induced by a prefunctor of quivers -/ +def _root_.category_theory.free_groupoid_functor (φ : prefunctor V V') : + free_groupoid V ⥤ free_groupoid V' := lift (φ.comp (of V')) + +lemma free_groupoid_functor_id : + free_groupoid_functor (prefunctor.id V) = functor.id (free_groupoid V) := +begin + dsimp only [free_groupoid_functor], symmetry, + apply lift_unique, refl, +end + +lemma free_groupoid_functor_comp + (φ : prefunctor V V') (φ' : prefunctor V' V'') : + free_groupoid_functor (φ.comp φ') = (free_groupoid_functor φ) ⋙ (free_groupoid_functor φ') := +begin + dsimp only [free_groupoid_functor], symmetry, + apply lift_unique, refl, +end + +end functoriality + end free end groupoid end category_theory diff --git a/src/category_theory/path_category.lean b/src/category_theory/path_category.lean index 60339f3cd66a5..28b11bba0740e 100644 --- a/src/category_theory/path_category.lean +++ b/src/category_theory/path_category.lean @@ -54,20 +54,20 @@ local attribute [ext] functor.ext /-- Any prefunctor from `V` lifts to a functor from `paths V` -/ def lift {C} [category C] (φ : prefunctor V C) : (paths V) ⥤ C := -{ obj := φ.obj -, map := λ X Y f, @quiver.path.rec V _ X (λ Y f, φ.obj X ⟶ φ.obj Y) (𝟙 $ φ.obj X) - (λ Y Z p f ihp, ihp ≫ (φ.map f)) Y f -, map_id' := λ X, by { refl, } -, map_comp' := λ X Y Z f g, by +{ obj := φ.obj, + map := λ X Y f, @quiver.path.rec V _ X (λ Y f, φ.obj X ⟶ φ.obj Y) (𝟙 $ φ.obj X) + (λ Y Z p f ihp, ihp ≫ (φ.map f)) Y f, + map_id' := λ X, by { refl, }, + map_comp' := λ X Y Z f g, by { induction g with _ _ g' p ih _ _ _, { rw category.comp_id, refl, }, { have : f ≫ g'.cons p = (f ≫ g').cons p, by apply quiver.path.comp_cons, - rw this, simp only, rw [ih, category.assoc], }} } + rw this, simp only, rw [ih, category.assoc], } } } -@[simp] lemma lift_nil {C} [category C] (φ : prefunctor V C) (X : V) : +@[simp] lemma lift_nil {C} [category C] (φ : prefunctor V C) (X : V) : (lift φ).map (quiver.path.nil) = 𝟙 (φ.obj X) := rfl -@[simp] lemma lift_cons {C} [category C] (φ : prefunctor V C) {X Y Z: V} +@[simp] lemma lift_cons {C} [category C] (φ : prefunctor V C) {X Y Z : V} (p : quiver.path X Y) (f : Y ⟶ Z) : (lift φ).map (p.cons f) = (lift φ).map p ≫ (φ.map f) := rfl @@ -84,7 +84,7 @@ begin simp only [category.id_comp], }, end -lemma lift_spec_unique {C} [category C] (φ : prefunctor V C) (Φ : paths V ⥤ C) +lemma lift_unique {C} [category C] (φ : prefunctor V C) (Φ : paths V ⥤ C) (hΦ : of.comp Φ.to_prefunctor = φ) : Φ = lift φ := begin subst_vars, @@ -97,7 +97,7 @@ begin { simp only [category.comp_id, category.id_comp] at ih ⊢, have : Φ.map (p.cons f') = Φ.map p ≫ (Φ.map (f'.to_path)), by { convert functor.map_comp Φ p (f'.to_path), }, - rw [this,ih], }, }, + rw [this, ih], }, }, end /-- Two functors out of a path category are equal when they agree on singleton paths. -/ diff --git a/src/category_theory/quotient.lean b/src/category_theory/quotient.lean index 1505a2c65360b..0d68d8d19c9ef 100644 --- a/src/category_theory/quotient.lean +++ b/src/category_theory/quotient.lean @@ -137,7 +137,7 @@ begin { rintro X Y f, simp, }, end -lemma lift_spec_unique (Φ : quotient r ⥤ D) (hΦ : (functor r) ⋙ Φ = F) : Φ = lift r F H := +lemma lift_unique (Φ : quotient r ⥤ D) (hΦ : (functor r) ⋙ Φ = F) : Φ = lift r F H := begin subst_vars, apply functor.hext, diff --git a/src/combinatorics/quiver/basic.lean b/src/combinatorics/quiver/basic.lean index 24b02fbed757e..b8d528af34ff4 100644 --- a/src/combinatorics/quiver/basic.lean +++ b/src/combinatorics/quiver/basic.lean @@ -88,7 +88,7 @@ def comp {U : Type*} [quiver U] {V : Type*} [quiver V] {W : Type*} [quiver W] @[simp] lemma comp_assoc - {U : Type*} [quiver U] {V : Type*} [quiver V] {W : Type*} [quiver W] {Z : Type*} [quiver Z] + {U V W Z : Type*} [quiver U] [quiver V] [quiver W] [quiver Z] (F : prefunctor U V) (G : prefunctor V W) (H : prefunctor W Z) : (F.comp G).comp H = F.comp (G.comp H) := begin diff --git a/src/combinatorics/quiver/connected_component.lean b/src/combinatorics/quiver/connected_component.lean index 9d3404adf7159..e1ed5dd20a23b 100644 --- a/src/combinatorics/quiver/connected_component.lean +++ b/src/combinatorics/quiver/connected_component.lean @@ -76,15 +76,15 @@ instance : has_involutive_reverse (symmetrify V) := /-- The inclusion of a quiver in its symmetrification -/ def symmetrify.of : prefunctor V (symmetrify V) := -{ obj := id -, map := λ X Y f, sum.inl f } +{ obj := id, + map := λ X Y f, sum.inl f } /-- Given a quiver `V'` with reversible arrows, a prefunctor to `V'` can be lifted to one from `symmetrify V` to `V'` -/ def symmetrify.lift {V' : Type*} [quiver V'] [has_reverse V'] (φ : prefunctor V V') : prefunctor (symmetrify V) V' := -{ obj := φ.obj -, map := λ X Y f, sum.rec (λ fwd, φ.map fwd) (λ bwd, reverse (φ.map bwd)) f } +{ obj := φ.obj, + map := λ X Y f, sum.rec (λ fwd, φ.map fwd) (λ bwd, reverse (φ.map bwd)) f } lemma symmetrify.lift_spec (V' : Type*) [quiver V'] [has_reverse V'] (φ : prefunctor V V') : symmetrify.of.comp (symmetrify.lift φ) = φ := @@ -105,7 +105,7 @@ begin end /-- `lift φ` is the only prefunctor extending `φ` and preserving reverses. -/ -lemma symmetrify.lift_spec_unique (V' : Type*) [quiver V'] [has_reverse V'] +lemma symmetrify.lift_unique (V' : Type*) [quiver V'] [has_reverse V'] (φ : prefunctor V V') (Φ : prefunctor (symmetrify V) V') (hΦ : symmetrify.of.comp Φ = φ) From f737ca11965922b3e9269f26258c409e9f6c2552 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Tue, 11 Oct 2022 18:36:20 +0000 Subject: [PATCH 689/828] feat(data/option/basic): add `bind_some'` (#16641) We already have `some_bind` `bind_some` `some_bind'`, but there is no `bind_some'`. --- src/data/option/basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/data/option/basic.lean b/src/data/option/basic.lean index 16ff17eba180f..e02a2de4d7611 100644 --- a/src/data/option/basic.lean +++ b/src/data/option/basic.lean @@ -110,6 +110,9 @@ theorem eq_none_iff_forall_not_mem {o : option α} : @[simp] theorem bind_some : ∀ x : option α, x >>= some = x := @bind_pure α option _ _ +@[simp] theorem bind_some' : ∀ x : option α, x.bind some = x := +bind_some + @[simp] theorem bind_eq_some {α β} {x : option α} {f : α → option β} {b : β} : x >>= f = some b ↔ ∃ a, x = some a ∧ f a = some b := by cases x; simp From 300cf4dc1dc92a040ffb3d65ac94cbe0a84f6021 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Tue, 11 Oct 2022 18:36:21 +0000 Subject: [PATCH 690/828] feat(topology/uniform_space/basic): separation of compacts and closed sets in a uniform space (#16747) --- src/topology/uniform_space/basic.lean | 26 ++++++++++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/src/topology/uniform_space/basic.lean b/src/topology/uniform_space/basic.lean index 0ec0af2ac8d0b..8314bd8d5ee3e 100644 --- a/src/topology/uniform_space/basic.lean +++ b/src/topology/uniform_space/basic.lean @@ -709,6 +709,32 @@ begin ... ⊆ U : hI z (htK z hzt), end +lemma disjoint.exists_uniform_thickening {A B : set α} + (hA : is_compact A) (hB : is_closed B) (h : disjoint A B) : + ∃ V ∈ 𝓤 α, disjoint (⋃ x ∈ A, ball x V) (⋃ x ∈ B, ball x V) := +begin + have : Bᶜ ∈ 𝓝ˢ A := hB.is_open_compl.mem_nhds_set.mpr h.le_compl_right, + rw (hA.nhds_set_basis_uniformity (filter.basis_sets _)).mem_iff at this, + rcases this with ⟨U, hU, hUAB⟩, + rcases comp_symm_mem_uniformity_sets hU with ⟨V, hV, hVsymm, hVU⟩, + refine ⟨V, hV, λ x, _⟩, + simp only [inf_eq_inter, mem_inter_iff, mem_Union₂], + rintro ⟨⟨a, ha, hxa⟩, ⟨b, hb, hxb⟩⟩, + rw mem_ball_symmetry hVsymm at hxa hxb, + exact hUAB (mem_Union₂_of_mem ha $ hVU $ mem_comp_of_mem_ball hVsymm hxa hxb) hb +end + +lemma disjoint.exists_uniform_thickening_of_basis {p : ι → Prop} {s : ι → set (α × α)} + (hU : (𝓤 α).has_basis p s) {A B : set α} + (hA : is_compact A) (hB : is_closed B) (h : disjoint A B) : + ∃ i, p i ∧ disjoint (⋃ x ∈ A, ball x (s i)) (⋃ x ∈ B, ball x (s i)) := +begin + rcases h.exists_uniform_thickening hA hB with ⟨V, hV, hVAB⟩, + rcases hU.mem_iff.1 hV with ⟨i, hi, hiV⟩, + exact ⟨i, hi, hVAB.mono + (Union₂_mono $ λ a _, ball_mono hiV a) (Union₂_mono $ λ b _, ball_mono hiV b)⟩, +end + lemma tendsto_right_nhds_uniformity {a : α} : tendsto (λa', (a', a)) (𝓝 a) (𝓤 α) := assume s, mem_nhds_right a From 93def16e7b9e0b95945d73ae1e3b58244c922708 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 11 Oct 2022 18:36:23 +0000 Subject: [PATCH 691/828] feat(analysis/normed/ring/seminorm): Multiplicative ring norms (#16766) Define `mul_ring_seminorm`/`mul_ring_norm`, the type of multiplicative `ring_seminorm`/`ring_norm`. --- src/analysis/normed/ring/seminorm.lean | 127 +++++++++++++++++++++++-- 1 file changed, 120 insertions(+), 7 deletions(-) diff --git a/src/analysis/normed/ring/seminorm.lean b/src/analysis/normed/ring/seminorm.lean index dfc994e58364d..61b22a3650e81 100644 --- a/src/analysis/normed/ring/seminorm.lean +++ b/src/analysis/normed/ring/seminorm.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2022 María Inés de Frutos-Fernández. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: María Inés de Frutos-Fernández +Authors: María Inés de Frutos-Fernández, Yaël Dillies -/ import analysis.seminorm @@ -18,6 +18,9 @@ For a ring `R`: nonnegative values, is subadditive and submultiplicative and such that `f (-x) = f x` for all `x ∈ R`. * `ring_norm`: A seminorm `f` is a norm if `f x = 0` if and only if `x = 0`. +* `mul_ring_seminorm`: A multiplicative seminorm on a ring `R` is a ring seminorm that preserves + multiplication. +* `mul_ring_norm`: A multiplicative norm on a ring `R` is a ring norm that preserves multiplication. ## References @@ -31,33 +34,70 @@ set_option old_structure_cmd true open_locale nnreal -variables {R S : Type*} (x y : R) (r : ℝ) +variables {F R S : Type*} (x y : R) (r : ℝ) /-- A seminorm on a ring `R` is a function `f : R → ℝ` that preserves zero, takes nonnegative values, is subadditive and submultiplicative and such that `f (-x) = f x` for all `x ∈ R`. -/ -structure ring_seminorm (R : Type*) [non_unital_ring R] +structure ring_seminorm (R : Type*) [non_unital_non_assoc_ring R] extends add_group_seminorm R := (mul_le' : ∀ x y : R, to_fun (x * y) ≤ to_fun x * to_fun y) /-- A function `f : R → ℝ` is a norm on a (nonunital) ring if it is a seminorm and `f x = 0` implies `x = 0`. -/ -structure ring_norm (R : Type*) [non_unital_ring R] extends add_group_norm R, ring_seminorm R +structure ring_norm (R : Type*) [non_unital_non_assoc_ring R] + extends ring_seminorm R, add_group_norm R + +/-- A multiplicative seminorm on a ring `R` is a function `f : R → ℝ` that preserves zero and +multiplication, takes nonnegative values, is subadditive and such that `f (-x) = f x` for all `x`. +-/ +structure mul_ring_seminorm (R : Type*) [non_assoc_ring R] + extends add_group_seminorm R, monoid_with_zero_hom R ℝ + +/-- A multiplicative norm on a ring `R` is a multiplicative ring seminorm such that `f x = 0` +implies `x = 0`. -/ +structure mul_ring_norm (R : Type*) [non_assoc_ring R] extends mul_ring_seminorm R, add_group_norm R attribute [nolint doc_blame] ring_seminorm.to_add_group_seminorm ring_norm.to_add_group_norm - ring_norm.to_ring_seminorm + ring_norm.to_ring_seminorm mul_ring_seminorm.to_add_group_seminorm + mul_ring_seminorm.to_monoid_with_zero_hom mul_ring_norm.to_add_group_norm + mul_ring_norm.to_mul_ring_seminorm /-- `ring_seminorm_class F α` states that `F` is a type of seminorms on the ring `α`. You should extend this class when you extend `ring_seminorm`. -/ -class ring_seminorm_class (F : Type*) (α : out_param $ Type*) [non_unital_ring α] +class ring_seminorm_class (F : Type*) (α : out_param $ Type*) [non_unital_non_assoc_ring α] extends add_group_seminorm_class F α, submultiplicative_hom_class F α ℝ /-- `ring_norm_class F α` states that `F` is a type of norms on the ring `α`. You should extend this class when you extend `ring_norm`. -/ -class ring_norm_class (F : Type*) (α : out_param $ Type*) [non_unital_ring α] +class ring_norm_class (F : Type*) (α : out_param $ Type*) [non_unital_non_assoc_ring α] extends ring_seminorm_class F α, add_group_norm_class F α +/-- `mul_ring_seminorm_class F α` states that `F` is a type of multiplicative seminorms on the ring +`α`. + +You should extend this class when you extend `mul_ring_seminorm`. -/ +class mul_ring_seminorm_class (F : Type*) (α : out_param $ Type*) [non_assoc_ring α] + extends add_group_seminorm_class F α, monoid_with_zero_hom_class F α ℝ + +/-- `mul_ring_norm_class F α` states that `F` is a type of multiplicative norms on the ring `α`. + +You should extend this class when you extend `mul_ring_norm`. -/ +class mul_ring_norm_class (F : Type*) (α : out_param $ Type*) [non_assoc_ring α] + extends mul_ring_seminorm_class F α, add_group_norm_class F α + +@[priority 100] -- See note [lower instance priority] +instance mul_ring_seminorm_class.to_ring_seminorm_class [non_assoc_ring R] + [mul_ring_seminorm_class F R] : ring_seminorm_class F R := +{ map_mul_le_mul := λ f a b, (map_mul _ _ _).le, + ..‹mul_ring_seminorm_class F R› } + +@[priority 100] -- See note [lower instance priority] +instance mul_ring_norm_class.to_ring_norm_class [non_assoc_ring R] [mul_ring_norm_class F R] : + ring_norm_class F R := +{ ..‹mul_ring_norm_class F R›, ..mul_ring_seminorm_class.to_ring_seminorm_class } + namespace ring_seminorm section non_unital_ring @@ -167,6 +207,79 @@ instance [decidable_eq R] : inhabited (ring_norm R) := ⟨1⟩ end ring_norm +namespace mul_ring_seminorm +variables [non_assoc_ring R] + +instance mul_ring_seminorm_class : mul_ring_seminorm_class (mul_ring_seminorm R) R := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_zero := λ f, f.map_zero', + map_one := λ f, f.map_one', + map_add_le_add := λ f, f.add_le', + map_mul := λ f, f.map_mul', + map_neg_eq_map := λ f, f.neg' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ +instance : has_coe_to_fun (mul_ring_seminorm R) (λ _, R → ℝ) := fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe (p : mul_ring_seminorm R) : p.to_fun = p := rfl + +@[ext] lemma ext {p q : mul_ring_seminorm R} : (∀ x, p x = q x) → p = q := fun_like.ext p q + +variables [decidable_eq R] [no_zero_divisors R] [nontrivial R] + +/-- The trivial seminorm on a ring `R` is the `mul_ring_seminorm` taking value `0` at `0` and `1` at +every other element. -/ +instance : has_one (mul_ring_seminorm R) := +⟨{ map_one' := if_neg one_ne_zero, + map_mul' := λ x y, begin + obtain rfl | hx := eq_or_ne x 0, + { simp }, + obtain rfl | hy := eq_or_ne y 0, + { simp }, + { simp [hx, hy] } + end, + ..(1 : add_group_seminorm R) }⟩ + +@[simp] lemma apply_one (x : R) : (1 : mul_ring_seminorm R) x = if x = 0 then 0 else 1 := rfl + +instance : inhabited (mul_ring_seminorm R) := ⟨1⟩ + +end mul_ring_seminorm + +namespace mul_ring_norm +variable [non_assoc_ring R] + +instance mul_ring_norm_class : mul_ring_norm_class (mul_ring_norm R) R := +{ coe := λ f, f.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_zero := λ f, f.map_zero', + map_one := λ f, f.map_one', + map_add_le_add := λ f, f.add_le', + map_mul := λ f, f.map_mul', + map_neg_eq_map := λ f, f.neg', + eq_zero_of_map_eq_zero := λ f, f.eq_zero_of_map_eq_zero' } + +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun`. -/ +instance : has_coe_to_fun (mul_ring_norm R) (λ _, R → ℝ) := ⟨λ p, p.to_fun⟩ + +@[simp] lemma to_fun_eq_coe (p : mul_ring_norm R) : p.to_fun = p := rfl + +@[ext] lemma ext {p q : mul_ring_norm R} : (∀ x, p x = q x) → p = q := fun_like.ext p q + +variables (R) [decidable_eq R] [no_zero_divisors R] [nontrivial R] + +/-- The trivial norm on a ring `R` is the `mul_ring_norm` taking value `0` at `0` and `1` at every +other element. -/ +instance : has_one (mul_ring_norm R) := +⟨{ ..(1 : mul_ring_seminorm R), ..(1 : add_group_norm R) }⟩ + +@[simp] lemma apply_one (x : R) : (1 : mul_ring_norm R) x = if x = 0 then 0 else 1 := rfl + +instance : inhabited (mul_ring_norm R) := ⟨1⟩ + +end mul_ring_norm + /-- A nonzero ring seminorm on a field `K` is a ring norm. -/ def ring_seminorm.to_ring_norm {K : Type*} [field K] (f : ring_seminorm K) (hnt : f ≠ 0) : ring_norm K := From 0a0be05e42dc5ef1217ca92fadf40de857e81602 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 11 Oct 2022 21:54:55 +0000 Subject: [PATCH 692/828] feat(data/set/intervals/unordered_interval): Intervals are injective in both endpoints (#16168) `set.interval a` and `set.interval_oc a` are injective. --- .../set/intervals/unordered_interval.lean | 87 ++++++++++++++----- 1 file changed, 66 insertions(+), 21 deletions(-) diff --git a/src/data/set/intervals/unordered_interval.lean b/src/data/set/intervals/unordered_interval.lean index 33d7df4ff9b6f..e9e5239712fcb 100644 --- a/src/data/set/intervals/unordered_interval.lean +++ b/src/data/set/intervals/unordered_interval.lean @@ -26,6 +26,8 @@ make the notation available. -/ universe u + +open function open_locale pointwise open order_dual (to_dual of_dual) @@ -63,17 +65,18 @@ interval_of_gt (lt_of_not_ge h) lemma interval_of_not_ge (h : ¬ b ≤ a) : [a, b] = Icc a b := interval_of_lt (lt_of_not_ge h) +lemma interval_eq_union : [a, b] = Icc a b ∪ Icc b a := by rw [Icc_union_Icc', max_comm]; refl + +lemma mem_interval : a ∈ [b, c] ↔ b ≤ a ∧ a ≤ c ∨ c ≤ a ∧ a ≤ b := by simp [interval_eq_union] + @[simp] lemma interval_self : [a, a] = {a} := set.ext $ by simp [le_antisymm_iff, and_comm] @[simp] lemma nonempty_interval : set.nonempty [a, b] := by { simp only [interval, min_le_iff, le_max_iff, nonempty_Icc], left, left, refl } -@[simp] lemma left_mem_interval : a ∈ [a, b] := -by { rw [interval, mem_Icc], exact ⟨min_le_left _ _, le_max_left _ _⟩ } - -@[simp] lemma right_mem_interval : b ∈ [a, b] := -by { rw interval_swap, exact left_mem_interval } +@[simp] lemma left_mem_interval : a ∈ [a, b] := by simp [mem_interval, le_total] +@[simp] lemma right_mem_interval : b ∈ [a, b] := by simp [mem_interval, le_total] lemma Icc_subset_interval : Icc a b ⊆ [a, b] := Icc_subset_Icc (min_le_left _ _) (le_max_right _ _) @@ -87,10 +90,10 @@ Icc_subset_interval ⟨ha, hb⟩ lemma mem_interval_of_ge (hb : b ≤ x) (ha : x ≤ a) : x ∈ [a, b] := Icc_subset_interval' ⟨hb, ha⟩ -lemma not_mem_interval_of_lt (ha : c < a) (hb : c < b) : c ∉ interval a b := +lemma not_mem_interval_of_lt (ha : c < a) (hb : c < b) : c ∉ [a, b] := not_mem_Icc_of_lt $ lt_min_iff.mpr ⟨ha, hb⟩ -lemma not_mem_interval_of_gt (ha : a < c) (hb : b < c) : c ∉ interval a b := +lemma not_mem_interval_of_gt (ha : a < c) (hb : b < c) : c ∉ [a, b] := not_mem_Icc_of_gt $ max_lt_iff.mpr ⟨ha, hb⟩ lemma interval_subset_interval (h₁ : a₁ ∈ [a₂, b₂]) (h₂ : b₁ ∈ [a₂, b₂]) : [a₁, b₁] ⊆ [a₂, b₂] := @@ -114,18 +117,21 @@ interval_subset_interval left_mem_interval h /-- A sort of triangle inequality. -/ lemma interval_subset_interval_union_interval : [a, c] ⊆ [a, b] ∪ [b, c] := -begin - rintro x hx, - obtain hac | hac := le_total a c, - { rw interval_of_le hac at hx, - obtain hb | hb := le_total x b, - { exact or.inl (mem_interval_of_le hx.1 hb) }, - { exact or.inr (mem_interval_of_le hb hx.2) } }, - { rw interval_of_ge hac at hx, - obtain hb | hb := le_total x b, - { exact or.inr (mem_interval_of_ge hx.1 hb) }, - { exact or.inl (mem_interval_of_ge hb hx.2) } } -end +λ x, by simp only [mem_interval, mem_union]; cases le_total a c; cases le_total x b; tauto + +lemma eq_of_mem_interval_of_mem_interval : a ∈ [b, c] → b ∈ [a, c] → a = b := +by simp_rw mem_interval; rintro (⟨_, _⟩ | ⟨_, _⟩) (⟨_, _⟩ | ⟨_, _⟩); apply le_antisymm; + assumption <|> { exact le_trans ‹_› ‹_› } + +lemma eq_of_mem_interval_of_mem_interval' : b ∈ [a, c] → c ∈ [a, b] → b = c := +by simpa only [interval_swap a] using eq_of_mem_interval_of_mem_interval + +lemma interval_injective_right (a : α) : injective (λ b, interval b a) := +λ b c h, by { rw ext_iff at h, + exact eq_of_mem_interval_of_mem_interval ((h _).1 left_mem_interval) ((h _).2 left_mem_interval) } + +lemma interval_injective_left (a : α) : injective (interval a) := +by simpa only [interval_swap] using interval_injective_right a lemma bdd_below_bdd_above_iff_subset_interval (s : set α) : bdd_below s ∧ bdd_above s ↔ ∃ a b, s ⊆ [a, b] := @@ -142,15 +148,24 @@ def interval_oc : α → α → set α := λ a b, Ioc (min a b) (max a b) -- Below is a capital iota localized "notation `Ι` := set.interval_oc" in interval -lemma interval_oc_of_le (h : a ≤ b) : Ι a b = Ioc a b := +@[simp] lemma interval_oc_of_le (h : a ≤ b) : Ι a b = Ioc a b := by simp [interval_oc, h] -lemma interval_oc_of_lt (h : b < a) : Ι a b = Ioc b a := +@[simp] lemma interval_oc_of_lt (h : b < a) : Ι a b = Ioc b a := by simp [interval_oc, le_of_lt h] lemma interval_oc_eq_union : Ι a b = Ioc a b ∪ Ioc b a := by cases le_total a b; simp [interval_oc, *] +lemma mem_interval_oc : a ∈ Ι b c ↔ b < a ∧ a ≤ c ∨ c < a ∧ a ≤ b := +by simp only [interval_oc_eq_union, mem_union, mem_Ioc] + +lemma not_mem_interval_oc : a ∉ Ι b c ↔ a ≤ b ∧ a ≤ c ∨ c < a ∧ b < a := +by { simp only [interval_oc_eq_union, mem_union, mem_Ioc, not_lt, ←not_le], tauto } + +@[simp] lemma left_mem_interval_oc : a ∈ Ι a b ↔ b < a := by simp [mem_interval_oc] +@[simp] lemma right_mem_interval_oc : b ∈ Ι a b ↔ a < b := by simp [mem_interval_oc] + lemma forall_interval_oc_iff {P : α → Prop} : (∀ x ∈ Ι a b, P x) ↔ (∀ x ∈ Ioc a b, P x) ∧ (∀ x ∈ Ioc b a, P x) := by simp only [interval_oc_eq_union, mem_union, or_imp_distrib, forall_and_distrib] @@ -168,6 +183,36 @@ Ioc_subset_Ioc (min_le_left _ _) (le_max_right _ _) lemma Ioc_subset_interval_oc' : Ioc a b ⊆ Ι b a := Ioc_subset_Ioc (min_le_right _ _) (le_max_left _ _) +lemma eq_of_mem_interval_oc_of_mem_interval_oc : a ∈ Ι b c → b ∈ Ι a c → a = b := +by simp_rw mem_interval_oc; rintro (⟨_, _⟩ | ⟨_, _⟩) (⟨_, _⟩ | ⟨_, _⟩); apply le_antisymm; + assumption <|> exact le_of_lt ‹_› <|> exact le_trans ‹_› (le_of_lt ‹_›) + +lemma eq_of_mem_interval_oc_of_mem_interval_oc' : b ∈ Ι a c → c ∈ Ι a b → b = c := +by simpa only [interval_oc_swap a] using eq_of_mem_interval_oc_of_mem_interval_oc + +lemma eq_of_not_mem_interval_oc_of_not_mem_interval_oc (ha : a ≤ c) (hb : b ≤ c) : + a ∉ Ι b c → b ∉ Ι a c → a = b := +by simp_rw not_mem_interval_oc; rintro (⟨_, _⟩ | ⟨_, _⟩) (⟨_, _⟩ | ⟨_, _⟩); apply le_antisymm; + assumption <|> exact le_of_lt ‹_› <|> cases not_le_of_lt ‹_› ‹_› + +lemma interval_oc_injective_right (a : α) : injective (λ b, Ι b a) := +begin + rintro b c h, + rw ext_iff at h, + obtain ha | ha := le_or_lt b a, + { have hb := (h b).not, + simp only [ha, left_mem_interval_oc, not_lt, true_iff, not_mem_interval_oc, ←not_le, and_true, + not_true, false_and, not_false_iff, true_iff, or_false] at hb, + refine hb.eq_of_not_lt (λ hc, _), + simpa [ha, and_iff_right hc, ←@not_le _ _ _ a, -not_le] using h c }, + { refine eq_of_mem_interval_oc_of_mem_interval_oc ((h _).1 $ left_mem_interval_oc.2 ha) + ((h _).2 $ left_mem_interval_oc.2 $ ha.trans_le _), + simpa [ha, ha.not_le, mem_interval_oc] using h b } +end + +lemma interval_oc_injective_left (a : α) : injective (Ι a) := +by simpa only [interval_oc_swap] using interval_oc_injective_right a + end linear_order open_locale interval From ccf8aa04f4748e4153bf7df029e5e43af306ef8b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 11 Oct 2022 21:54:56 +0000 Subject: [PATCH 693/828] feat(data/finset/basic): disj_Union (#16421) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `finset.disj_Union` is to `finset.bUnion` as: * `finset.disj_union` is to `finset.union` * `finset.cons` is to `insert` * `finset.map` is to `finset.image` Co-authored-by: Yaël Dillies --- src/algebra/big_operators/basic.lean | 14 +++++ src/data/finset/basic.lean | 80 ++++++++++++++++++++++++++++ src/data/finset/fold.lean | 4 ++ src/data/list/basic.lean | 12 ++++- src/data/multiset/basic.lean | 17 +++--- src/data/multiset/bind.lean | 5 ++ src/data/multiset/fold.lean | 8 +++ src/data/multiset/nodup.lean | 10 ++-- 8 files changed, 137 insertions(+), 13 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index 13446fadb00c1..a99b0f85fcf5f 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -261,6 +261,16 @@ attribute [congr] finset.sum_congr lemma prod_disj_union (h) : ∏ x in s₁.disj_union s₂ h, f x = (∏ x in s₁, f x) * ∏ x in s₂, f x := by { refine eq.trans _ (fold_disj_union h), rw one_mul, refl } +@[to_additive] +lemma prod_disj_Union (s : finset ι) (t : ι → finset α) (h) : + ∏ x in s.disj_Union t h, f x = ∏ i in s, ∏ x in t i, f x := +begin + refine eq.trans _ (fold_disj_Union h), + dsimp [finset.prod, multiset.prod, multiset.fold, finset.disj_Union, finset.fold], + congr', + exact prod_const_one.symm, +end + @[to_additive] lemma prod_union_inter [decidable_eq α] : (∏ x in (s₁ ∪ s₂), f x) * (∏ x in (s₁ ∩ s₂), f x) = (∏ x in s₁, f x) * (∏ x in s₂, f x) := @@ -1469,6 +1479,10 @@ end comm_group card (s.sigma t) = ∑ a in s, card (t a) := multiset.card_sigma _ _ +@[simp] lemma card_disj_Union (s : finset α) (t : α → finset β) (h) : + (s.disj_Union t h).card = s.sum (λ i, (t i).card) := +multiset.card_bind _ _ + lemma card_bUnion [decidable_eq β] {s : finset α} {t : α → finset β} (h : ∀ x ∈ s, ∀ y ∈ s, x ≠ y → disjoint (t x) (t y)) : (s.bUnion t).card = ∑ u in s, card (t u) := diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 3f1252d41f0ba..ff7ef1ae6aec9 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -2623,6 +2623,78 @@ cons_eq_insert _ _ h ▸ to_list_cons _ end to_list +/-! +### disj_Union + +This section is about the bounded union of a disjoint indexed family `t : α → finset β` of finite +sets over a finite set `s : finset α`. In most cases `finset.bUnion` should be preferred. +-/ +section disj_Union + +variables {s s₁ s₂ : finset α} {t t₁ t₂ : α → finset β} + +/-- `disj_Union s f h` is the set such that `a ∈ disj_Union s f` iff `a ∈ f i` for some `i ∈ s`. +It is the same as `s.bUnion f`, but it does not require decidable equality on the type. The +hypothesis ensures that the sets are disjoint. -/ +def disj_Union (s : finset α) (t : α → finset β) + (hf : (s : set α).pairwise $ λ a b, ∀ x, x ∈ t a → x ∉ t b) : finset β := +⟨(s.val.bind (finset.val ∘ t)), multiset.nodup_bind.mpr + ⟨λ a ha, (t a).nodup, s.nodup.pairwise hf⟩⟩ + +@[simp] theorem disj_Union_val (s : finset α) (t : α → finset β) (h) : + (s.disj_Union t h).1 = (s.1.bind (λ a, (t a).1)) := rfl + +@[simp] theorem disj_Union_empty (t : α → finset β) : disj_Union ∅ t (by simp) = ∅ := rfl + +@[simp] lemma mem_disj_Union {b : β} {h} : + b ∈ s.disj_Union t h ↔ ∃ a ∈ s, b ∈ t a := +by simp only [mem_def, disj_Union_val, mem_bind, exists_prop] + +@[simp, norm_cast] lemma coe_disj_Union {h} : (s.disj_Union t h : set β) = ⋃ x ∈ (s : set α), t x := +by simp only [set.ext_iff, mem_disj_Union, set.mem_Union, iff_self, mem_coe, implies_true_iff] + +@[simp] theorem disj_Union_cons (a : α) (s : finset α) (ha : a ∉ s) (f : α → finset β) (H) : + disj_Union (cons a s ha) f H = (f a).disj_union + (s.disj_Union f $ + λ b hb c hc, H (mem_cons_of_mem hb) (mem_cons_of_mem hc)) + (λ b hb h, let ⟨c, hc, h⟩ := mem_disj_Union.mp h in + H (mem_cons_self a s) (mem_cons_of_mem hc) (ne_of_mem_of_not_mem hc ha).symm b hb h) := +eq_of_veq $ multiset.cons_bind _ _ _ + +@[simp] lemma singleton_disj_Union (a : α) {h} : finset.disj_Union {a} t h = t a := +eq_of_veq $ multiset.singleton_bind _ _ + +theorem map_disj_Union {f : α ↪ β} {s : finset α} {t : β → finset γ} {h} : + (s.map f).disj_Union t h = s.disj_Union (λa, t (f a)) + (λ a ha b hb hab, h (mem_map_of_mem _ ha) (mem_map_of_mem _ hb) (f.injective.ne hab)) := +eq_of_veq $ multiset.bind_map _ _ _ + +theorem disj_Union_map {s : finset α} {t : α → finset β} {f : β ↪ γ} {h} : + (s.disj_Union t h).map f = s.disj_Union (λa, (t a).map f) + (λ a ha b hb hab x hxa hxb, begin + obtain ⟨xa, hfa, rfl⟩ := mem_map.mp hxa, + obtain ⟨xb, hfb, hfab⟩ := mem_map.mp hxb, + obtain rfl := f.injective hfab, + exact h ha hb hab _ hfa hfb, + end) := +eq_of_veq $ multiset.map_bind _ _ _ + +lemma disj_Union_disj_Union (s : finset α) (f : α → finset β) (g : β → finset γ) (h1 h2) : + (s.disj_Union f h1).disj_Union g h2 = + s.attach.disj_Union (λ a, (f a).disj_Union g $ + λ b hb c hc, h2 (mem_disj_Union.mpr ⟨_, a.prop, hb⟩) (mem_disj_Union.mpr ⟨_, a.prop, hc⟩)) + (λ a ha b hb hab x hxa hxb, begin + obtain ⟨xa, hfa, hga⟩ := mem_disj_Union.mp hxa, + obtain ⟨xb, hfb, hgb⟩ := mem_disj_Union.mp hxb, + refine h2 + (mem_disj_Union.mpr ⟨_, a.prop, hfa⟩) (mem_disj_Union.mpr ⟨_, b.prop, hfb⟩) _ _ hga hgb, + rintro rfl, + exact h1 a.prop b.prop (subtype.coe_injective.ne hab) _ hfa hfb, + end) := +eq_of_veq $ multiset.bind_assoc.trans (multiset.attach_bind_coe _ _).symm + +end disj_Union + section bUnion /-! ### bUnion @@ -2657,6 +2729,14 @@ ext $ λ x, by simp only [mem_bUnion, exists_prop, mem_union, mem_insert, lemma bUnion_congr (hs : s₁ = s₂) (ht : ∀ a ∈ s₁, t₁ a = t₂ a) : s₁.bUnion t₁ = s₂.bUnion t₂ := ext $ λ x, by simp [hs, ht] { contextual := tt } +@[simp] lemma disj_Union_eq_bUnion (s : finset α) (f : α → finset β) (hf) : + s.disj_Union f hf = s.bUnion f := +begin + dsimp [disj_Union, finset.bUnion, function.comp], + generalize_proofs h, + exact eq_of_veq h.dedup.symm, +end + theorem bUnion_subset {s' : finset β} : s.bUnion t ⊆ s' ↔ ∀ x ∈ s, t x ⊆ s' := by simp only [subset_iff, mem_bUnion]; exact ⟨λ H a ha b hb, H ⟨a, ha, hb⟩, λ H b ⟨a, ha, hb⟩, H a ha hb⟩ diff --git a/src/data/finset/fold.lean b/src/data/finset/fold.lean index 31120d6e4bc95..f18afefa04224 100644 --- a/src/data/finset/fold.lean +++ b/src/data/finset/fold.lean @@ -74,6 +74,10 @@ theorem fold_disj_union {s₁ s₂ : finset α} {b₁ b₂ : β} (h) : (s₁.disj_union s₂ h).fold op (b₁ * b₂) f = s₁.fold op b₁ f * s₂.fold op b₂ f := (congr_arg _ $ multiset.map_add _ _ _).trans (multiset.fold_add _ _ _ _ _) +theorem fold_disj_Union {ι : Type*} {s : finset ι} {t : ι → finset α} {b : ι → β} {b₀ : β} (h) : + (s.disj_Union t h).fold op (s.fold op b₀ b) f = s.fold op b₀ (λ i, (t i).fold op (b i) f) := +(congr_arg _ $ multiset.map_bind _ _ _).trans (multiset.fold_bind _ _ _ _ _) + theorem fold_union_inter [decidable_eq α] {s₁ s₂ : finset α} {b₁ b₂ : β} : (s₁ ∪ s₂).fold op b₁ f * (s₁ ∩ s₂).fold op b₂ f = s₁.fold op b₂ f * s₂.fold op b₁ f := by unfold fold; rw [← fold_add op, ← multiset.map_add, union_val, diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 9e25340644993..d7b95c043908d 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -2710,8 +2710,16 @@ theorem pmap_eq_map_attach {p : α → Prop} (f : Π a, p a → β) (l H) : pmap f l H = l.attach.map (λ x, f x.1 (H _ x.2)) := by rw [attach, map_pmap]; exact pmap_congr l (λ _ _ _ _, rfl) -theorem attach_map_val (l : list α) : l.attach.map subtype.val = l := -by rw [attach, map_pmap]; exact (pmap_eq_map _ _ _ _).trans (map_id l) +@[simp] lemma attach_map_coe' (l : list α) (f : α → β) : l.attach.map (λ i, f i) = l.map f := +by rw [attach, map_pmap]; exact (pmap_eq_map _ _ _ _) + +lemma attach_map_val' (l : list α) (f : α → β) : l.attach.map (λ i, f i.val) = l.map f := +attach_map_coe' _ _ + +@[simp] lemma attach_map_coe (l : list α) : l.attach.map (coe : _ → α) = l := +(attach_map_coe' _ _).trans l.map_id + +lemma attach_map_val (l : list α) : l.attach.map subtype.val = l := attach_map_coe _ @[simp] theorem mem_attach (l : list α) : ∀ x, x ∈ l.attach | ⟨a, h⟩ := by have := mem_map.1 (by rw [attach_map_val]; exact h); diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 7ed374f032db8..040a00767e83a 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -1077,11 +1077,19 @@ theorem map_pmap {p : α → Prop} (g : β → γ) (f : Π a, p a → β) quot.induction_on s $ λ l H, congr_arg coe $ map_pmap g f l H theorem pmap_eq_map_attach {p : α → Prop} (f : Π a, p a → β) - (s) : ∀ H, pmap f s H = s.attach.map (λ x, f x.1 (H _ x.2)) := + (s) : ∀ H, pmap f s H = s.attach.map (λ x, f x (H _ x.prop)) := quot.induction_on s $ λ l H, congr_arg coe $ pmap_eq_map_attach f l H -theorem attach_map_val (s : multiset α) : s.attach.map subtype.val = s := -quot.induction_on s $ λ l, congr_arg coe $ attach_map_val l +@[simp] lemma attach_map_coe' (s : multiset α) (f : α → β) : s.attach.map (λ i, f i) = s.map f := +quot.induction_on s $ λ l, congr_arg coe $ attach_map_coe' l f + +lemma attach_map_val' (s : multiset α) (f : α → β) : s.attach.map (λ i, f i.val) = s.map f := +attach_map_coe' _ _ + +@[simp] lemma attach_map_coe (s : multiset α) : s.attach.map (coe : _ → α) = s := +(attach_map_coe' _ _).trans s.map_id + +lemma attach_map_val (s : multiset α) : s.attach.map subtype.val = s := attach_map_coe _ @[simp] theorem mem_attach (s : multiset α) : ∀ x, x ∈ s.attach := quot.induction_on s $ λ l, mem_attach _ @@ -1103,9 +1111,6 @@ lemma attach_cons (a : α) (m : multiset α) : quotient.induction_on m $ assume l, congr_arg coe $ congr_arg (list.cons _) $ by rw [list.map_pmap]; exact list.pmap_congr _ (λ _ _ _ _, subtype.eq rfl) -@[simp] -lemma attach_map_coe (m : multiset α) : multiset.map (coe : _ → α) m.attach = m := m.attach_map_val - section decidable_pi_exists variables {m : multiset α} diff --git a/src/data/multiset/bind.lean b/src/data/multiset/bind.lean index 1942de94367d9..120f923eaa064 100644 --- a/src/data/multiset/bind.lean +++ b/src/data/multiset/bind.lean @@ -139,6 +139,11 @@ begin rw count_bind, apply le_sum_of_mem, rw mem_map, exact ⟨x, hx, rfl⟩ end + +@[simp] theorem attach_bind_coe (s : multiset α) (f : α → multiset β) : + s.attach.bind (λ i, f i) = s.bind f := +congr_arg join $ attach_map_coe' _ _ + end bind /-! ### Product of two multisets -/ diff --git a/src/data/multiset/fold.lean b/src/data/multiset/fold.lean index 94e79d0543047..9adcacde43569 100644 --- a/src/data/multiset/fold.lean +++ b/src/data/multiset/fold.lean @@ -55,6 +55,14 @@ multiset.induction_on s₂ (by rw [add_zero, fold_zero, ← fold_cons'_right, ← fold_cons_right op]) (by simp {contextual := tt}; cc) +theorem fold_bind {ι : Type*} (s : multiset ι) (t : ι → multiset α) (b : ι → α) (b₀ : α) : + (s.bind t).fold op ((s.map b).fold op b₀) = (s.map (λ i, (t i).fold op (b i))).fold op b₀ := +begin + induction s using multiset.induction_on with a ha ih, + { rw [zero_bind, map_zero, map_zero, fold_zero] }, + { rw [cons_bind, map_cons, map_cons, fold_cons_left, fold_cons_left, fold_add, ih] }, +end + theorem fold_singleton (b a : α) : ({a} : multiset α).fold op b = a * b := foldr_singleton _ _ _ _ theorem fold_distrib {f g : β → α} (u₁ u₂ : α) (s : multiset β) : diff --git a/src/data/multiset/nodup.lean b/src/data/multiset/nodup.lean index 1b2d7d382837f..73f744131d197 100644 --- a/src/data/multiset/nodup.lean +++ b/src/data/multiset/nodup.lean @@ -205,14 +205,14 @@ lemma map_eq_map_of_bij_of_nodup (f : α → γ) (g : β → γ) {s : multiset (i_inj : ∀a₁ a₂ ha₁ ha₂, i a₁ ha₁ = i a₂ ha₂ → a₁ = a₂) (i_surj : ∀b∈t, ∃a ha, b = i a ha) : s.map f = t.map g := -have t = s.attach.map (λ x, i x.1 x.2), +have t = s.attach.map (λ x, i x x.2), from (ht.ext $ (nodup_attach.2 hs).map $ - show injective (λ x : {x // x ∈ s}, i x.1 x.2), from λ x y hxy, - subtype.eq $ i_inj x.1 y.1 x.2 y.2 hxy).2 + show injective (λ x : {x // x ∈ s}, i x x.2), from λ x y hxy, + subtype.ext $ i_inj x y x.2 y.2 hxy).2 (λ x, by simp only [mem_map, true_and, subtype.exists, eq_comm, mem_attach]; exact ⟨i_surj _, λ ⟨y, hy⟩, hy.snd.symm ▸ hi _ _⟩), -calc s.map f = s.pmap (λ x _, f x) (λ _, id) : by rw [pmap_eq_map] -... = s.attach.map (λ x, f x.1) : by rw [pmap_eq_map_attach] +calc s.map f = s.pmap (λ x _, f x) (λ _, id) : by rw [pmap_eq_map] +... = s.attach.map (λ x, f x) : by rw [pmap_eq_map_attach] ... = t.map g : by rw [this, multiset.map_map]; exact map_congr rfl (λ x _, h _ _) end multiset From 7507876b7eb105ccb92729ec6bbb2a6a30392e1d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 11 Oct 2022 21:54:57 +0000 Subject: [PATCH 694/828] =?UTF-8?q?feat(data/real/ereal):=20`coe=20:=20?= =?UTF-8?q?=E2=84=9D=20=E2=86=92=20ereal`=20is=20strictly=20monotone=20(#1?= =?UTF-8?q?6607)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and golf existing lemmas with it. Also delete `ereal.bot_lt_top` and friends in terms of the more general `bot_lt_top`. --- src/data/real/ereal.lean | 107 +++++++++++++++++------------- src/topology/instances/ereal.lean | 2 +- 2 files changed, 63 insertions(+), 46 deletions(-) diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index a29a8597cdebe..0585a16df30db 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -53,7 +53,7 @@ open function open_locale ennreal nnreal /-- ereal : The type `[-∞, ∞]` -/ -@[derive [has_top, comm_monoid_with_zero, nontrivial, +@[derive [has_top, comm_monoid_with_zero, nontrivial, add_monoid, has_Sup, has_Inf, complete_linear_order, linear_ordered_add_comm_monoid_with_top]] def ereal := with_top (with_bot ℝ) @@ -66,9 +66,6 @@ namespace ereal -- TODO: Provide explicitly, otherwise it is inferred noncomputably from `complete_linear_order` instance : has_bot ereal := ⟨some ⊥⟩ -@[simp] lemma bot_lt_top : (⊥ : ereal) < ⊤ := with_top.coe_lt_top _ -@[simp] lemma bot_ne_top : (⊥ : ereal) ≠ ⊤ := bot_lt_top.ne - instance : has_coe ℝ ereal := ⟨real.to_ereal⟩ lemma coe_strict_mono : strict_mono (coe : ℝ → ereal) := @@ -157,10 +154,18 @@ by { apply with_top.coe_lt_coe.2, exact with_bot.bot_lt_coe _ } @[simp] lemma top_ne_zero : (⊤ : ereal) ≠ 0 := (coe_ne_top 0).symm -@[simp, norm_cast] lemma coe_add (x y : ℝ) : ((x + y : ℝ) : ereal) = (x : ereal) + (y : ereal) := -rfl - @[simp, norm_cast] lemma coe_zero : ((0 : ℝ) : ereal) = 0 := rfl +@[simp, norm_cast] lemma coe_one : ((1 : ℝ) : ereal) = 1 := rfl +@[simp, norm_cast] lemma coe_add (x y : ℝ) : (↑(x + y) : ereal) = x + y := rfl +@[simp, norm_cast] lemma coe_mul (x y : ℝ) : (↑(x * y) : ereal) = x * y := +(with_top.coe_eq_coe.2 with_bot.coe_mul).trans with_top.coe_mul +@[norm_cast] lemma coe_nsmul (n : ℕ) (x : ℝ) : (↑(n • x) : ereal) = n • x := +map_nsmul (⟨coe, coe_zero, coe_add⟩ : ℝ →+ ereal) _ _ +@[simp, norm_cast] lemma coe_pow (x : ℝ) (n : ℕ) : (↑(x ^ n) : ereal) = x ^ n := +map_pow (⟨coe, coe_one, coe_mul⟩ : ℝ →* ereal) _ _ + +@[simp, norm_cast] lemma coe_bit0 (x : ℝ) : (↑(bit0 x) : ereal) = bit0 x := rfl +@[simp, norm_cast] lemma coe_bit1 (x : ℝ) : (↑(bit1 x) : ereal) = bit1 x := rfl @[simp, norm_cast] lemma coe_eq_zero {x : ℝ} : (x : ereal) = 0 ↔ x = 0 := ereal.coe_eq_coe_iff @[simp, norm_cast] lemma coe_eq_one {x : ℝ} : (x : ereal) = 1 ↔ x = 1 := ereal.coe_eq_coe_iff @@ -249,25 +254,28 @@ lemma coe_nnreal_ne_top (x : ℝ≥0) : ((x : ℝ≥0∞) : ereal) ≠ ⊤ := de @[simp] lemma coe_nnreal_lt_top (x : ℝ≥0) : ((x : ℝ≥0∞) : ereal) < ⊤ := dec_trivial -@[simp, norm_cast] lemma coe_ennreal_le_coe_ennreal_iff : ∀ {x y : ℝ≥0∞}, - (x : ereal) ≤ (y : ereal) ↔ x ≤ y -| x ⊤ := by simp -| ⊤ (some y) := by simp -| (some x) (some y) := by simp [coe_nnreal_eq_coe_real] - -@[simp, norm_cast] lemma coe_ennreal_lt_coe_ennreal_iff : ∀ {x y : ℝ≥0∞}, - (x : ereal) < (y : ereal) ↔ x < y +lemma coe_ennreal_strict_mono : strict_mono (coe : ℝ≥0∞ → ereal) | ⊤ ⊤ := by simp | (some x) ⊤ := by simp | ⊤ (some y) := by simp | (some x) (some y) := by simp [coe_nnreal_eq_coe_real] -@[simp, norm_cast] lemma coe_ennreal_eq_coe_ennreal_iff : ∀ {x y : ℝ≥0∞}, - (x : ereal) = (y : ereal) ↔ x = y -| ⊤ ⊤ := by simp -| (some x) ⊤ := by simp -| ⊤ (some y) := by simp [(coe_nnreal_lt_top y).ne'] -| (some x) (some y) := by simp [coe_nnreal_eq_coe_real] +lemma coe_ennreal_injective : injective (coe : ℝ≥0∞ → ereal) := coe_ennreal_strict_mono.injective + +@[simp, norm_cast] lemma coe_ennreal_le_coe_ennreal_iff {x y : ℝ≥0∞} : + (x : ereal) ≤ (y : ereal) ↔ x ≤ y := +coe_ennreal_strict_mono.le_iff_le + +@[simp, norm_cast] lemma coe_ennreal_lt_coe_ennreal_iff {x y : ℝ≥0∞} : + (x : ereal) < (y : ereal) ↔ x < y := +coe_ennreal_strict_mono.lt_iff_lt + +@[simp, norm_cast] lemma coe_ennreal_eq_coe_ennreal_iff {x y : ℝ≥0∞} : + (x : ereal) = (y : ereal) ↔ x = y := +coe_ennreal_injective.eq_iff + +lemma coe_ennreal_ne_coe_ennreal_iff {x y : ℝ≥0∞} : (x : ereal) ≠ (y : ereal) ↔ x ≠ y := +coe_ennreal_injective.ne_iff @[simp, norm_cast] lemma coe_ennreal_eq_zero {x : ℝ≥0∞} : (x : ereal) = 0 ↔ x = 0 := by rw [←coe_ennreal_eq_coe_ennreal_iff, coe_ennreal_zero] @@ -296,6 +304,21 @@ by rw [←coe_ennreal_zero, coe_ennreal_lt_coe_ennreal_iff] | x ⊤ := by simp | (some x) (some y) := rfl +@[simp, norm_cast] lemma coe_ennreal_mul : ∀ (x y : ℝ≥0∞), ((x * y : ℝ≥0∞) : ereal) = x * y +| ⊤ y := by { rw ennreal.top_mul, split_ifs; simp [h] } +| x ⊤ := by { rw ennreal.mul_top, split_ifs; simp [h] } +| (some x) (some y) := by simp [←ennreal.coe_mul, coe_nnreal_eq_coe_real] + +@[norm_cast] lemma coe_ennreal_nsmul (n : ℕ) (x : ℝ≥0∞) : (↑(n • x) : ereal) = n • x := +map_nsmul (⟨coe, coe_ennreal_zero, coe_ennreal_add⟩ : ℝ≥0∞ →+ ereal) _ _ + +@[simp, norm_cast] lemma coe_ennreal_pow (x : ℝ≥0∞) (n : ℕ) : (↑(x ^ n) : ereal) = x ^ n := +map_pow (⟨coe, coe_ennreal_one, coe_ennreal_mul⟩ : ℝ≥0∞ →* ereal) _ _ + +@[simp, norm_cast] lemma coe_ennreal_bit0 (x : ℝ≥0∞) : (↑(bit0 x) : ereal) = bit0 x := +coe_ennreal_add _ _ +@[simp, norm_cast] lemma coe_ennreal_bit1 (x : ℝ≥0∞) : (↑(bit1 x) : ereal) = bit1 x := +by simp_rw [bit1, coe_ennreal_add, coe_ennreal_bit0, coe_ennreal_one] /-! ### Order -/ @@ -401,11 +424,20 @@ protected def neg : ereal → ereal instance : has_neg ereal := ⟨ereal.neg⟩ +instance : sub_neg_zero_monoid ereal := +{ neg_zero := by { change ((-0 : ℝ) : ereal) = 0, simp }, + ..ereal.add_monoid, ..ereal.has_neg } + @[norm_cast] protected lemma neg_def (x : ℝ) : ((-x : ℝ) : ereal) = -x := rfl @[simp] lemma neg_top : - (⊤ : ereal) = ⊥ := rfl @[simp] lemma neg_bot : - (⊥ : ereal) = ⊤ := rfl +@[simp, norm_cast] lemma coe_neg (x : ℝ) : (↑(-x) : ereal) = -x := rfl +@[simp, norm_cast] lemma coe_sub (x y : ℝ) : (↑(x - y) : ereal) = x - y := rfl +@[norm_cast] lemma coe_zsmul (n : ℤ) (x : ℝ) : (↑(n • x) : ereal) = n • x := +map_zsmul' (⟨coe, coe_zero, coe_add⟩ : ℝ →+ ereal) coe_neg _ _ + instance : has_involutive_neg ereal := { neg := has_neg.neg, neg_neg := λ a, match a with @@ -419,14 +451,14 @@ instance : has_involutive_neg ereal := | ⊥ := by simp | (x : ℝ) := rfl -@[simp] lemma neg_eg_top_iff {x : ereal} : - x = ⊤ ↔ x = ⊥ := +@[simp] lemma neg_eq_top_iff {x : ereal} : - x = ⊤ ↔ x = ⊥ := by { rw neg_eq_iff_neg_eq, simp [eq_comm] } -@[simp] lemma neg_eg_bot_iff {x : ereal} : - x = ⊥ ↔ x = ⊤ := +@[simp] lemma neg_eq_bot_iff {x : ereal} : - x = ⊥ ↔ x = ⊤ := by { rw neg_eq_iff_neg_eq, simp [eq_comm] } -@[simp] lemma neg_eg_zero_iff {x : ereal} : - x = 0 ↔ x = 0 := -by { rw neg_eq_iff_neg_eq, change ((-0 : ℝ) : ereal) = _ ↔ _, simp [eq_comm] } +@[simp] lemma neg_eq_zero_iff {x : ereal} : - x = 0 ↔ x = 0 := +by { rw neg_eq_iff_neg_eq, simp [eq_comm] } /-- if `-a ≤ b` then `-b ≤ a` on `ereal`. -/ protected theorem neg_le_of_neg_le : ∀ {a b : ereal} (h : -a ≤ b), -b ≤ a @@ -448,8 +480,6 @@ by rwa [←neg_neg b, ereal.neg_le, neg_neg] @[simp] lemma neg_le_neg_iff {a b : ereal} : - a ≤ - b ↔ b ≤ a := by conv_lhs { rw [ereal.neg_le, neg_neg] } -@[simp, norm_cast] lemma coe_neg (x : ℝ) : ((- x : ℝ) : ereal) = - (x : ereal) := rfl - /-- Negation as an order reversing isomorphism on `ereal`. -/ def neg_order_iso : ereal ≃o erealᵒᵈ := { to_fun := λ x, order_dual.to_dual (-x), @@ -468,19 +498,13 @@ end lemma neg_lt_iff_neg_lt {a b : ereal} : -a < b ↔ -b < a := ⟨λ h, ereal.neg_lt_of_neg_lt h, λ h, ereal.neg_lt_of_neg_lt h⟩ -/-! ### Subtraction -/ +/-! +### Subtraction -/-- Subtraction on `ereal`, defined by `x - y = x + (-y)`. Since addition is badly behaved at some +Subtraction on `ereal` is defined by `x - y = x + (-y)`. Since addition is badly behaved at some points, so is subtraction. There is no standard algebraic typeclass involving subtraction that is -registered on `ereal`, beyond `sub_neg_zero_monoid`, because of this bad behavior. -/ -protected noncomputable def sub (x y : ereal) : ereal := x + (-y) - -noncomputable instance : has_sub ereal := ⟨ereal.sub⟩ - -noncomputable instance : sub_neg_zero_monoid ereal := -{ neg_zero := by { change ((-0 : ℝ) : ereal) = 0, simp }, - ..(infer_instance : add_monoid ereal), - ..ereal.has_neg } +registered on `ereal`, beyond `sub_neg_zero_monoid`, because of this bad behavior. +-/ @[simp] lemma top_sub (x : ereal) : ⊤ - x = ⊤ := top_add x @[simp] lemma sub_bot (x : ereal) : x - ⊥ = ⊤ := add_top x @@ -489,8 +513,6 @@ noncomputable instance : sub_neg_zero_monoid ereal := @[simp] lemma bot_sub_coe (x : ℝ) : (⊥ : ereal) - x = ⊥ := rfl @[simp] lemma coe_sub_bot (x : ℝ) : (x : ereal) - ⊤ = ⊥ := rfl -@[simp] lemma zero_sub (x : ereal) : 0 - x = - x := by { change 0 + (-x) = - x, simp } - lemma sub_le_sub {x y z t : ereal} (h : x ≤ y) (h' : t ≤ z) : x - z ≤ y - t := add_le_add h (neg_le_neg_iff.2 h') @@ -525,11 +547,6 @@ end /-! ### Multiplication -/ -@[simp] lemma coe_one : ((1 : ℝ) : ereal) = 1 := rfl - -@[simp, norm_cast] lemma coe_mul (x y : ℝ) : ((x * y : ℝ) : ereal) = (x : ereal) * (y : ereal) := -eq.trans (with_bot.coe_eq_coe.mpr with_bot.coe_mul) with_top.coe_mul - @[simp] lemma mul_top (x : ereal) (h : x ≠ 0) : x * ⊤ = ⊤ := with_top.mul_top h @[simp] lemma top_mul (x : ereal) (h : x ≠ 0) : ⊤ * x = ⊤ := with_top.top_mul h diff --git a/src/topology/instances/ereal.lean b/src/topology/instances/ereal.lean index 504077e238b47..942acbe39b44b 100644 --- a/src/topology/instances/ereal.lean +++ b/src/topology/instances/ereal.lean @@ -259,7 +259,7 @@ begin assume r, rw eventually_prod_iff, refine ⟨λ z, ((r - (a - 1): ℝ) : ereal) < z, Ioi_mem_nhds (coe_lt_top _), - λ z, ((a - 1 : ℝ) : ereal) < z, Ioi_mem_nhds (by simp [zero_lt_one]), + λ z, ((a - 1 : ℝ) : ereal) < z, Ioi_mem_nhds (by simp [-ereal.coe_sub]), λ x hx y hy, _⟩, dsimp, convert add_lt_add hx hy, From 5fd24816689e85003d855aea49bcff358db53049 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 11 Oct 2022 21:54:58 +0000 Subject: [PATCH 695/828] feat(logic/basic): add four lemmas about (d)ite (#16879) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ```lean lemma dite_eq_iff' : dite P A B = c ↔ (∀ h, A h = c) ∧ (∀ h, B h = c) := lemma ite_eq_iff' : ite P a b = c ↔ (P → a = c) ∧ (¬ P → b = c) := dite_eq_iff' lemma dite_dite_comm {B : Q → α} {C : ¬P → ¬Q → α} (h : P → ¬Q) : (if p : P then A p else if q : Q then B q else C p q) = (if q : Q then B q else if p : P then A p else C p q) := lemma ite_ite_comm (h : P → ¬Q) : (if P then a else if Q then b else c) = (if Q then b else if P then a else c) := ``` for #15681 --- src/logic/basic.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/logic/basic.lean b/src/logic/basic.lean index 7cafe15c710c1..ad3276f1ff2bc 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -1508,6 +1508,12 @@ by by_cases P; simp [*, exists_prop_of_false not_false] lemma ite_eq_iff : ite P a b = c ↔ P ∧ a = c ∨ ¬ P ∧ b = c := dite_eq_iff.trans $ by rw [exists_prop, exists_prop] +lemma dite_eq_iff' : dite P A B = c ↔ (∀ h, A h = c) ∧ (∀ h, B h = c) := +⟨λ he, ⟨λ h, (dif_pos h).symm.trans he, λ h, (dif_neg h).symm.trans he⟩, + λ he, (em P).elim (λ h, (dif_pos h).trans $ he.1 h) (λ h, (dif_neg h).trans $ he.2 h)⟩ + +lemma ite_eq_iff' : ite P a b = c ↔ (P → a = c) ∧ (¬ P → b = c) := dite_eq_iff' + @[simp] lemma dite_eq_left_iff : dite P (λ _, a) B = a ↔ ∀ h, B h = a := by by_cases P; simp [*, forall_prop_of_false not_false] @@ -1599,4 +1605,14 @@ by by_cases h : P; simp [h] lemma ite_and : ite (P ∧ Q) a b = ite P (ite Q a b) b := by by_cases hp : P; by_cases hq : Q; simp [hp, hq] +lemma dite_dite_comm {B : Q → α} {C : ¬P → ¬Q → α} (h : P → ¬Q) : + (if p : P then A p else if q : Q then B q else C p q) = + (if q : Q then B q else if p : P then A p else C p q) := +dite_eq_iff'.2 ⟨λ p, by rw [dif_neg (h p), dif_pos p], λ np, by { congr, funext, rw dif_neg np }⟩ + +lemma ite_ite_comm (h : P → ¬Q) : + (if P then a else if Q then b else c) = + (if Q then b else if P then a else c) := +dite_dite_comm P Q h + end ite From a4efbb0edc5c2c46fd7c86cc862e53dc1d83cc22 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Tue, 11 Oct 2022 21:54:59 +0000 Subject: [PATCH 696/828] feat(topology/uniform_space): local uniform convergence on union (#16883) I'm pretty confident these need some assumption like open-ness, but maybe something weaker is enough? Co-authored-by: Bhavik Mehta --- .../uniform_space/uniform_convergence.lean | 26 +++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/src/topology/uniform_space/uniform_convergence.lean b/src/topology/uniform_space/uniform_convergence.lean index b351b64278cb5..6a4d072ca3f6b 100644 --- a/src/topology/uniform_space/uniform_convergence.lean +++ b/src/topology/uniform_space/uniform_convergence.lean @@ -682,6 +682,32 @@ begin exact ⟨t, nhds_within_mono x h' ht, H.mono (λ n, id)⟩ end +lemma tendsto_locally_uniformly_on_Union {S : γ → set α} (hS : ∀ i, is_open (S i)) + (h : ∀ i, tendsto_locally_uniformly_on F f p (S i)) : + tendsto_locally_uniformly_on F f p (⋃ i, S i) := +begin + rintro v hv x ⟨_, ⟨i, rfl⟩, hi : x ∈ S i⟩, + obtain ⟨t, ht, ht'⟩ := h i v hv x hi, + refine ⟨t, _, ht'⟩, + rw (hS _).nhds_within_eq hi at ht, + exact mem_nhds_within_of_mem_nhds ht, +end + +lemma tendsto_locally_uniformly_on_bUnion {s : set γ} {S : γ → set α} + (hS : ∀ i ∈ s, is_open (S i)) (h : ∀ i ∈ s, tendsto_locally_uniformly_on F f p (S i)) : + tendsto_locally_uniformly_on F f p (⋃ i ∈ s, S i) := +by { rw bUnion_eq_Union, exact tendsto_locally_uniformly_on_Union (λ i, hS _ i.2) (λ i, h _ i.2) } + +lemma tendsto_locally_uniformly_on_sUnion (S : set (set α)) (hS : ∀ s ∈ S, is_open s) + (h : ∀ s ∈ S, tendsto_locally_uniformly_on F f p s) : + tendsto_locally_uniformly_on F f p (⋃₀ S) := +by { rw sUnion_eq_bUnion, exact tendsto_locally_uniformly_on_bUnion hS h } + +lemma tendsto_locally_uniformly_on.union {s₁ s₂ : set α} (hs₁ : is_open s₁) (hs₂ : is_open s₂) + (h₁ : tendsto_locally_uniformly_on F f p s₁) (h₂ : tendsto_locally_uniformly_on F f p s₂) : + tendsto_locally_uniformly_on F f p (s₁ ∪ s₂) := +by { rw ←sUnion_pair, refine tendsto_locally_uniformly_on_sUnion _ _ _; simp [*] } + lemma tendsto_locally_uniformly_on_univ : tendsto_locally_uniformly_on F f p univ ↔ tendsto_locally_uniformly F f p := by simp [tendsto_locally_uniformly_on, tendsto_locally_uniformly, nhds_within_univ] From 2f93d81cd3b14ef981a94a4dd5082d5259ac8b80 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 11 Oct 2022 21:55:01 +0000 Subject: [PATCH 697/828] feat(logic/equiv/transfer_instances, group_theory/eckmann_hilton): make definitions reducible (#16884) Make `equiv.has_one` etc. reducible, which is necessary to use these defs to construct instances; the check_reducibility linter will complain about "semireducible non-instances" if these are not marked reducible. [function.injective.comm_group](https://leanprover-community.github.io/mathlib_docs/algebra/group/inj_surj.html#function.injective.comm_group) etc. are already made reducible and used to construct instances. --- src/group_theory/eckmann_hilton.lean | 8 +-- src/logic/equiv/transfer_instance.lean | 96 ++++++++++++++------------ 2 files changed, 55 insertions(+), 49 deletions(-) diff --git a/src/group_theory/eckmann_hilton.lean b/src/group_theory/eckmann_hilton.lean index 3cbc0aadc8db5..f39129abc072a 100644 --- a/src/group_theory/eckmann_hilton.lean +++ b/src/group_theory/eckmann_hilton.lean @@ -81,8 +81,8 @@ omit h₁ h₂ distrib /-- If a type carries a unital magma structure that distributes over a unital binary operations, then the magma structure is a commutative monoid. -/ -@[to_additive "If a type carries a unital additive magma structure that distributes over a -unital binary operations, then the additive magma structure is a commutative additive monoid."] +@[reducible, to_additive "If a type carries a unital additive magma structure that distributes over +a unital binary operations, then the additive magma structure is a commutative additive monoid."] def comm_monoid [h : mul_one_class X] (distrib : ∀ a b c d, ((a * b) (c * d)) = ((a c) * (b d))) : comm_monoid X := { mul := (*), @@ -93,8 +93,8 @@ def comm_monoid [h : mul_one_class X] /-- If a type carries a group structure that distributes over a unital binary operation, then the group is commutative. -/ -@[to_additive "If a type carries an additive group structure that distributes -over a unital binary operation, then the additive group is commutative."] +@[reducible, to_additive "If a type carries an additive group structure that +distributes over a unital binary operation, then the additive group is commutative."] def comm_group [G : group X] (distrib : ∀ a b c d, ((a * b) (c * d)) = ((a c) * (b d))) : comm_group X := { ..(eckmann_hilton.comm_monoid h₁ distrib), diff --git a/src/logic/equiv/transfer_instance.lean b/src/logic/equiv/transfer_instance.lean index 82d66799c069d..c5cdb3b84e8bc 100644 --- a/src/logic/equiv/transfer_instance.lean +++ b/src/logic/equiv/transfer_instance.lean @@ -19,6 +19,11 @@ so on. Note that most of these constructions can also be obtained using the `transport` tactic. +### Implementation details + +When adding new definitions that transfer type-classes across an equivalence, please mark them +`@[reducible]`. See note [reducible non-instances]. + ## Tags equiv, group, ring, field, module, algebra @@ -35,39 +40,39 @@ section instances variables (e : α ≃ β) /-- Transfer `has_one` across an `equiv` -/ -@[to_additive "Transfer `has_zero` across an `equiv`"] +@[reducible, to_additive "Transfer `has_zero` across an `equiv`"] protected def has_one [has_one β] : has_one α := ⟨e.symm 1⟩ @[to_additive] lemma one_def [has_one β] : @has_one.one _ (equiv.has_one e) = e.symm 1 := rfl /-- Transfer `has_mul` across an `equiv` -/ -@[to_additive "Transfer `has_add` across an `equiv`"] +@[reducible, to_additive "Transfer `has_add` across an `equiv`"] protected def has_mul [has_mul β] : has_mul α := ⟨λ x y, e.symm (e x * e y)⟩ @[to_additive] lemma mul_def [has_mul β] (x y : α) : @has_mul.mul _ (equiv.has_mul e) x y = e.symm (e x * e y) := rfl /-- Transfer `has_div` across an `equiv` -/ -@[to_additive "Transfer `has_sub` across an `equiv`"] +@[reducible, to_additive "Transfer `has_sub` across an `equiv`"] protected def has_div [has_div β] : has_div α := ⟨λ x y, e.symm (e x / e y)⟩ @[to_additive] lemma div_def [has_div β] (x y : α) : @has_div.div _ (equiv.has_div e) x y = e.symm (e x / e y) := rfl /-- Transfer `has_inv` across an `equiv` -/ -@[to_additive "Transfer `has_neg` across an `equiv`"] +@[reducible, to_additive "Transfer `has_neg` across an `equiv`"] protected def has_inv [has_inv β] : has_inv α := ⟨λ x, e.symm (e x)⁻¹⟩ @[to_additive] lemma inv_def [has_inv β] (x : α) : @has_inv.inv _ (equiv.has_inv e) x = e.symm (e x)⁻¹ := rfl /-- Transfer `has_smul` across an `equiv` -/ -protected def has_smul (R : Type*) [has_smul R β] : has_smul R α := +@[reducible] protected def has_smul (R : Type*) [has_smul R β] : has_smul R α := ⟨λ r x, e.symm (r • (e x))⟩ lemma smul_def {R : Type*} [has_smul R β] (r : R) (x : α) : @has_smul.smul _ _ (e.has_smul R) r x = e.symm (r • (e x)) := rfl /-- Transfer `has_pow` across an `equiv` -/ -@[to_additive has_smul] +@[reducible, to_additive has_smul] protected def has_pow (N : Type*) [has_pow β N] : has_pow α N := ⟨λ x n, e.symm (e x ^ n)⟩ lemma pow_def {N : Type*} [has_pow β N] (n : N) (x : α) : @@ -87,7 +92,7 @@ def mul_equiv (e : α ≃ β) [has_mul β] : begin introsI, exact - { map_mul' := λ x y, by { apply e.symm.injective, simp, refl, }, + { map_mul' := λ x y, by { apply e.symm.injective, simp, }, ..e } end @@ -110,8 +115,8 @@ def ring_equiv (e : α ≃ β) [has_add β] [has_mul β] : begin introsI, exact - { map_add' := λ x y, by { apply e.symm.injective, simp, refl, }, - map_mul' := λ x y, by { apply e.symm.injective, simp, refl, }, + { map_add' := λ x y, by { apply e.symm.injective, simp, }, + map_mul' := λ x y, by { apply e.symm.injective, simp, }, ..e } end @@ -125,84 +130,84 @@ begin end /-- Transfer `semigroup` across an `equiv` -/ -@[to_additive "Transfer `add_semigroup` across an `equiv`"] +@[reducible, to_additive "Transfer `add_semigroup` across an `equiv`"] protected def semigroup [semigroup β] : semigroup α := let mul := e.has_mul in by resetI; apply e.injective.semigroup _; intros; exact e.apply_symm_apply _ /-- Transfer `semigroup_with_zero` across an `equiv` -/ -protected def semigroup_with_zero [semigroup_with_zero β] : semigroup_with_zero α := +@[reducible] protected def semigroup_with_zero [semigroup_with_zero β] : semigroup_with_zero α := let mul := e.has_mul, zero := e.has_zero in by resetI; apply e.injective.semigroup_with_zero _; intros; exact e.apply_symm_apply _ /-- Transfer `comm_semigroup` across an `equiv` -/ -@[to_additive "Transfer `add_comm_semigroup` across an `equiv`"] +@[reducible, to_additive "Transfer `add_comm_semigroup` across an `equiv`"] protected def comm_semigroup [comm_semigroup β] : comm_semigroup α := let mul := e.has_mul in by resetI; apply e.injective.comm_semigroup _; intros; exact e.apply_symm_apply _ /-- Transfer `mul_zero_class` across an `equiv` -/ -protected def mul_zero_class [mul_zero_class β] : mul_zero_class α := +@[reducible] protected def mul_zero_class [mul_zero_class β] : mul_zero_class α := let zero := e.has_zero, mul := e.has_mul in by resetI; apply e.injective.mul_zero_class _; intros; exact e.apply_symm_apply _ /-- Transfer `mul_one_class` across an `equiv` -/ -@[to_additive "Transfer `add_zero_class` across an `equiv`"] +@[reducible, to_additive "Transfer `add_zero_class` across an `equiv`"] protected def mul_one_class [mul_one_class β] : mul_one_class α := let one := e.has_one, mul := e.has_mul in by resetI; apply e.injective.mul_one_class _; intros; exact e.apply_symm_apply _ /-- Transfer `mul_zero_one_class` across an `equiv` -/ -protected def mul_zero_one_class [mul_zero_one_class β] : mul_zero_one_class α := +@[reducible] protected def mul_zero_one_class [mul_zero_one_class β] : mul_zero_one_class α := let zero := e.has_zero, one := e.has_one,mul := e.has_mul in by resetI; apply e.injective.mul_zero_one_class _; intros; exact e.apply_symm_apply _ /-- Transfer `monoid` across an `equiv` -/ -@[to_additive "Transfer `add_monoid` across an `equiv`"] +@[reducible, to_additive "Transfer `add_monoid` across an `equiv`"] protected def monoid [monoid β] : monoid α := let one := e.has_one, mul := e.has_mul, pow := e.has_pow ℕ in by resetI; apply e.injective.monoid _; intros; exact e.apply_symm_apply _ /-- Transfer `comm_monoid` across an `equiv` -/ -@[to_additive "Transfer `add_comm_monoid` across an `equiv`"] +@[reducible, to_additive "Transfer `add_comm_monoid` across an `equiv`"] protected def comm_monoid [comm_monoid β] : comm_monoid α := let one := e.has_one, mul := e.has_mul, pow := e.has_pow ℕ in by resetI; apply e.injective.comm_monoid _; intros; exact e.apply_symm_apply _ /-- Transfer `group` across an `equiv` -/ -@[to_additive "Transfer `add_group` across an `equiv`"] +@[reducible, to_additive "Transfer `add_group` across an `equiv`"] protected def group [group β] : group α := let one := e.has_one, mul := e.has_mul, inv := e.has_inv, div := e.has_div, npow := e.has_pow ℕ, zpow := e.has_pow ℤ in by resetI; apply e.injective.group _; intros; exact e.apply_symm_apply _ /-- Transfer `comm_group` across an `equiv` -/ -@[to_additive "Transfer `add_comm_group` across an `equiv`"] +@[reducible, to_additive "Transfer `add_comm_group` across an `equiv`"] protected def comm_group [comm_group β] : comm_group α := let one := e.has_one, mul := e.has_mul, inv := e.has_inv, div := e.has_div, npow := e.has_pow ℕ, zpow := e.has_pow ℤ in by resetI; apply e.injective.comm_group _; intros; exact e.apply_symm_apply _ /-- Transfer `non_unital_non_assoc_semiring` across an `equiv` -/ -protected def non_unital_non_assoc_semiring [non_unital_non_assoc_semiring β] : +@[reducible] protected def non_unital_non_assoc_semiring [non_unital_non_assoc_semiring β] : non_unital_non_assoc_semiring α := let zero := e.has_zero, add := e.has_add, mul := e.has_mul, nsmul := e.has_smul ℕ in by resetI; apply e.injective.non_unital_non_assoc_semiring _; intros; exact e.apply_symm_apply _ /-- Transfer `non_unital_semiring` across an `equiv` -/ -protected def non_unital_semiring [non_unital_semiring β] : non_unital_semiring α := +@[reducible] protected def non_unital_semiring [non_unital_semiring β] : non_unital_semiring α := let zero := e.has_zero, add := e.has_add, mul := e.has_mul, nsmul := e.has_smul ℕ in by resetI; apply e.injective.non_unital_semiring _; intros; exact e.apply_symm_apply _ /-- Transfer `add_monoid_with_one` across an `equiv` -/ -protected def add_monoid_with_one [add_monoid_with_one β] : add_monoid_with_one α := +@[reducible] protected def add_monoid_with_one [add_monoid_with_one β] : add_monoid_with_one α := { nat_cast := λ n, e.symm n, nat_cast_zero := show e.symm _ = _, by simp [zero_def], nat_cast_succ := λ n, show e.symm _ = e.symm (e (e.symm _) + _), by simp [add_def, one_def], .. e.add_monoid, .. e.has_one } /-- Transfer `add_group_with_one` across an `equiv` -/ -protected def add_group_with_one [add_group_with_one β] : add_group_with_one α := +@[reducible] protected def add_group_with_one [add_group_with_one β] : add_group_with_one α := { int_cast := λ n, e.symm n, int_cast_of_nat := λ n, by rw [int.cast_coe_nat]; refl, int_cast_neg_succ_of_nat := λ n, congr_arg e.symm $ @@ -210,82 +215,83 @@ protected def add_group_with_one [add_group_with_one β] : add_group_with_one α .. e.add_monoid_with_one, .. e.add_group } /-- Transfer `non_assoc_semiring` across an `equiv` -/ -protected def non_assoc_semiring [non_assoc_semiring β] : non_assoc_semiring α := +@[reducible] protected def non_assoc_semiring [non_assoc_semiring β] : non_assoc_semiring α := let mul := e.has_mul, add_monoid_with_one := e.add_monoid_with_one in by resetI; apply e.injective.non_assoc_semiring _; intros; exact e.apply_symm_apply _ /-- Transfer `semiring` across an `equiv` -/ -protected def semiring [semiring β] : semiring α := +@[reducible] protected def semiring [semiring β] : semiring α := let mul := e.has_mul, add_monoid_with_one := e.add_monoid_with_one, npow := e.has_pow ℕ in by resetI; apply e.injective.semiring _; intros; exact e.apply_symm_apply _ /-- Transfer `non_unital_comm_semiring` across an `equiv` -/ -protected def non_unital_comm_semiring [non_unital_comm_semiring β] : non_unital_comm_semiring α := +@[reducible] protected def non_unital_comm_semiring [non_unital_comm_semiring β] : + non_unital_comm_semiring α := let zero := e.has_zero, add := e.has_add, mul := e.has_mul, nsmul := e.has_smul ℕ in by resetI; apply e.injective.non_unital_comm_semiring _; intros; exact e.apply_symm_apply _ /-- Transfer `comm_semiring` across an `equiv` -/ -protected def comm_semiring [comm_semiring β] : comm_semiring α := +@[reducible] protected def comm_semiring [comm_semiring β] : comm_semiring α := let mul := e.has_mul, add_monoid_with_one := e.add_monoid_with_one, npow := e.has_pow ℕ in by resetI; apply e.injective.comm_semiring _; intros; exact e.apply_symm_apply _ /-- Transfer `non_unital_non_assoc_ring` across an `equiv` -/ -protected def non_unital_non_assoc_ring [non_unital_non_assoc_ring β] : +@[reducible] protected def non_unital_non_assoc_ring [non_unital_non_assoc_ring β] : non_unital_non_assoc_ring α := let zero := e.has_zero, add := e.has_add, mul := e.has_mul, neg := e.has_neg, sub := e.has_sub, nsmul := e.has_smul ℕ, zsmul := e.has_smul ℤ in by resetI; apply e.injective.non_unital_non_assoc_ring _; intros; exact e.apply_symm_apply _ /-- Transfer `non_unital_ring` across an `equiv` -/ -protected def non_unital_ring [non_unital_ring β] : +@[reducible] protected def non_unital_ring [non_unital_ring β] : non_unital_ring α := let zero := e.has_zero, add := e.has_add, mul := e.has_mul, neg := e.has_neg, sub := e.has_sub, nsmul := e.has_smul ℕ, zsmul := e.has_smul ℤ in by resetI; apply e.injective.non_unital_ring _; intros; exact e.apply_symm_apply _ /-- Transfer `non_assoc_ring` across an `equiv` -/ -protected def non_assoc_ring [non_assoc_ring β] : +@[reducible] protected def non_assoc_ring [non_assoc_ring β] : non_assoc_ring α := let add_group_with_one := e.add_group_with_one, mul := e.has_mul in by resetI; apply e.injective.non_assoc_ring _; intros; exact e.apply_symm_apply _ /-- Transfer `ring` across an `equiv` -/ -protected def ring [ring β] : ring α := +@[reducible] protected def ring [ring β] : ring α := let mul := e.has_mul, add_group_with_one := e.add_group_with_one, npow := e.has_pow ℕ in by resetI; apply e.injective.ring _; intros; exact e.apply_symm_apply _ /-- Transfer `non_unital_comm_ring` across an `equiv` -/ -protected def non_unital_comm_ring [non_unital_comm_ring β] : non_unital_comm_ring α := +@[reducible] protected def non_unital_comm_ring [non_unital_comm_ring β] : non_unital_comm_ring α := let zero := e.has_zero, add := e.has_add, mul := e.has_mul, neg := e.has_neg, sub := e.has_sub, nsmul := e.has_smul ℕ, zsmul := e.has_smul ℤ in by resetI; apply e.injective.non_unital_comm_ring _; intros; exact e.apply_symm_apply _ /-- Transfer `comm_ring` across an `equiv` -/ -protected def comm_ring [comm_ring β] : comm_ring α := +@[reducible] protected def comm_ring [comm_ring β] : comm_ring α := let mul := e.has_mul, add_group_with_one := e.add_group_with_one, npow := e.has_pow ℕ in by resetI; apply e.injective.comm_ring _; intros; exact e.apply_symm_apply _ /-- Transfer `nontrivial` across an `equiv` -/ -protected theorem nontrivial [nontrivial β] : nontrivial α := +@[reducible] protected theorem nontrivial [nontrivial β] : nontrivial α := e.surjective.nontrivial /-- Transfer `is_domain` across an `equiv` -/ -protected theorem is_domain [ring α] [ring β] [is_domain β] (e : α ≃+* β) : is_domain α := -function.injective.is_domain e.to_ring_hom e.injective +@[reducible] protected theorem is_domain [ring α] [ring β] [is_domain β] (e : α ≃+* β) : + is_domain α := function.injective.is_domain e.to_ring_hom e.injective /-- Transfer `has_rat_cast` across an `equiv` -/ -protected def has_rat_cast [has_rat_cast β] : has_rat_cast α := +@[reducible] protected def has_rat_cast [has_rat_cast β] : has_rat_cast α := { rat_cast := λ n, e.symm n } /-- Transfer `division_ring` across an `equiv` -/ -protected def division_ring [division_ring β] : division_ring α := +@[reducible] protected def division_ring [division_ring β] : division_ring α := let add_group_with_one := e.add_group_with_one, mul := e.has_mul, inv := e.has_inv, div := e.has_div, mul := e.has_mul, npow := e.has_pow ℕ, zpow := e.has_pow ℤ, rat_cast := e.has_rat_cast, qsmul := e.has_smul ℚ in by resetI; apply e.injective.division_ring _; intros; exact e.apply_symm_apply _ /-- Transfer `field` across an `equiv` -/ -protected def field [field β] : field α := +@[reducible] protected def field [field β] : field α := let add_group_with_one := e.add_group_with_one, mul := e.has_mul, neg := e.has_neg, inv := e.has_inv, div := e.has_div, mul := e.has_mul, npow := e.has_pow ℕ, zpow := e.has_pow ℤ, rat_cast := e.has_rat_cast, qsmul := e.has_smul ℚ in @@ -299,13 +305,13 @@ section variables [monoid R] /-- Transfer `mul_action` across an `equiv` -/ -protected def mul_action (e : α ≃ β) [mul_action R β] : mul_action R α := +@[reducible] protected def mul_action (e : α ≃ β) [mul_action R β] : mul_action R α := { one_smul := by simp [smul_def], mul_smul := by simp [smul_def, mul_smul], ..e.has_smul R } /-- Transfer `distrib_mul_action` across an `equiv` -/ -protected def distrib_mul_action (e : α ≃ β) [add_comm_monoid β] : +@[reducible] protected def distrib_mul_action (e : α ≃ β) [add_comm_monoid β] : begin letI := equiv.add_comm_monoid e, exact Π [distrib_mul_action R β], distrib_mul_action R α @@ -325,7 +331,7 @@ section variables [semiring R] /-- Transfer `module` across an `equiv` -/ -protected def module (e : α ≃ β) [add_comm_monoid β] : +@[reducible] protected def module (e : α ≃ β) [add_comm_monoid β] : begin letI := equiv.add_comm_monoid e, exact Π [module R β], module R α @@ -362,7 +368,7 @@ section variables [comm_semiring R] /-- Transfer `algebra` across an `equiv` -/ -protected def algebra (e : α ≃ β) [semiring β] : +@[reducible] protected def algebra (e : α ≃ β) [semiring β] : begin letI := equiv.semiring e, exact Π [algebra R β], algebra R α @@ -407,8 +413,8 @@ end equiv namespace ring_equiv -protected lemma local_ring {A B : Type*} [comm_semiring A] [local_ring A] [comm_semiring B] - (e : A ≃+* B) : local_ring B := +@[reducible] protected lemma local_ring {A B : Type*} [comm_semiring A] [local_ring A] + [comm_semiring B] (e : A ≃+* B) : local_ring B := begin haveI := e.symm.to_equiv.nontrivial, exact local_ring.of_surjective (e : A →+* B) e.surjective From bbcc67f5c5af1d13d36046e0ea9be9d10de37087 Mon Sep 17 00:00:00 2001 From: mkaratarakis Date: Tue, 11 Oct 2022 21:55:02 +0000 Subject: [PATCH 698/828] =?UTF-8?q?feat(ring=5Ftheory/ideal/local=5Fring):?= =?UTF-8?q?=20add=20=E2=80=8Blocal=5Fring.residue=5Ffield.map=5Fid=20&=20l?= =?UTF-8?q?ocal=5Fring.residue=5Ffield.map=5Fcomp=20(#16916)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Applying `map` to the identity ring homomorphism gives the identity ring homomorphism. The composite of two `map`s is the `map` of the composite. I need these for my study of the inertia group Co-authored-by: mkaratarakis <40603357+mkaratarakis@users.noreply.github.com> --- src/ring_theory/ideal/local_ring.lean | 14 +++++++++++++- 1 file changed, 13 insertions(+), 1 deletion(-) diff --git a/src/ring_theory/ideal/local_ring.lean b/src/ring_theory/ideal/local_ring.lean index 6f2836a06e23b..6b6ef0552529c 100644 --- a/src/ring_theory/ideal/local_ring.lean +++ b/src/ring_theory/ideal/local_ring.lean @@ -328,7 +328,7 @@ begin end section -variables (R) [comm_ring R] [local_ring R] [comm_ring S] [local_ring S] +variables (R) [comm_ring R] [local_ring R] [comm_ring S] [local_ring S] [comm_ring T] [local_ring T] /-- The residue field of a local ring is the quotient of the ring by its maximal ideal. -/ def residue_field := R ⧸ maximal_ideal R @@ -359,6 +359,18 @@ begin exact map_nonunit f a ha end +/-- Applying `residue_field.map` to the identity ring homomorphism gives the identity +ring homomorphism. -/ +lemma map_id : + local_ring.residue_field.map (ring_hom.id R) = ring_hom.id (local_ring.residue_field R) := +ideal.quotient.ring_hom_ext $ ring_hom.ext $ λx, rfl + +/-- The composite of two `residue_field.map`s is the `residue_field.map` of the composite. -/ +lemma map_comp (f : T →+* R) (g : R →+* S) [is_local_ring_hom f] [is_local_ring_hom g] : + local_ring.residue_field.map (g.comp f) = + (local_ring.residue_field.map g).comp (local_ring.residue_field.map f) := +ideal.quotient.ring_hom_ext $ ring_hom.ext $ λx, rfl + end residue_field lemma ker_eq_maximal_ideal [field K] (φ : R →+* K) (hφ : function.surjective φ) : From 24f5150e2584aab5c421725fa2db9170ad0b4cd7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 12 Oct 2022 01:16:27 +0000 Subject: [PATCH 699/828] feat(algebraic_topology/dold_kan): very basic defs for split simplicial objects in preadditive categories (#16888) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR introduces very basic definitions for split simplicial objects in preadditive category, mostly the projections `s.π_summand A` on a summand (with index `A`) in the coproduct decomposition given by `s : splitting X` where `X` is a simplicial object. --- .../dold_kan/split_simplicial_object.lean | 99 +++++++++++++++++++ .../split_simplicial_object.lean | 54 ++++++++++ 2 files changed, 153 insertions(+) create mode 100644 src/algebraic_topology/dold_kan/split_simplicial_object.lean diff --git a/src/algebraic_topology/dold_kan/split_simplicial_object.lean b/src/algebraic_topology/dold_kan/split_simplicial_object.lean new file mode 100644 index 0000000000000..a071d2799a7c4 --- /dev/null +++ b/src/algebraic_topology/dold_kan/split_simplicial_object.lean @@ -0,0 +1,99 @@ +/- +Copyright (c) 2022 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import algebraic_topology.split_simplicial_object +import category_theory.preadditive + +/-! + +# Split simplicial objects in preadditive categories + +TODO @joelriou: Define a functor `N' : simplicial_object.split C ⥤ chain_complex C ℕ` +when `C` is a preadditive category, and get an isomorphism +`N' ⋙ to_karoubi (chain_complex C ℕ) ≅ forget C ⋙ dold_kan.N₁` + +-/ + +noncomputable theory + +open category_theory category_theory.limits category_theory.category + category_theory.preadditive opposite + +open_locale big_operators simplicial + +namespace simplicial_object + +namespace splitting + +variables {C : Type*} [category C] [has_finite_coproducts C] + {X : simplicial_object C} (s : splitting X) + +/-- The projection on a summand of the coproduct decomposition given +by a splitting of a simplicial object. -/ +def π_summand [has_zero_morphisms C] {Δ : simplex_categoryᵒᵖ} (A : index_set Δ) : + X.obj Δ ⟶ s.N A.1.unop.len := +begin + refine (s.iso Δ).inv ≫ sigma.desc (λ B, _), + by_cases B = A, + { exact eq_to_hom (by { subst h, refl, }), }, + { exact 0, }, +end + +@[simp, reassoc] +lemma ι_π_summand_eq_id [has_zero_morphisms C] {Δ : simplex_categoryᵒᵖ} (A : index_set Δ) : + s.ι_summand A ≫ s.π_summand A = 𝟙 _ := +begin + dsimp [ι_summand, π_summand], + simp only [summand, assoc, is_iso.hom_inv_id_assoc], + erw [colimit.ι_desc, cofan.mk_ι_app], + dsimp, + simp only [eq_self_iff_true, if_true], +end + +@[simp, reassoc] +lemma ι_π_summand_eq_zero [has_zero_morphisms C] {Δ : simplex_categoryᵒᵖ} (A B : index_set Δ) + (h : B ≠ A) : s.ι_summand A ≫ s.π_summand B = 0 := +begin + dsimp [ι_summand, π_summand], + simp only [summand, assoc, is_iso.hom_inv_id_assoc], + erw [colimit.ι_desc, cofan.mk_ι_app], + apply dif_neg, + exact h.symm, +end + +variable [preadditive C] + +lemma decomposition_id (Δ : simplex_categoryᵒᵖ) : + 𝟙 (X.obj Δ) = ∑ (A : index_set Δ), s.π_summand A ≫ s.ι_summand A := +begin + apply s.hom_ext', + intro A, + rw [comp_id, comp_sum, finset.sum_eq_single A, ι_π_summand_eq_id_assoc], + { intros B h₁ h₂, + rw [s.ι_π_summand_eq_zero_assoc _ _ h₂, zero_comp], }, + { simp only [finset.mem_univ, not_true, is_empty.forall_iff], }, +end + +@[simp, reassoc] +lemma σ_comp_π_summand_id_eq_zero {n : ℕ} (i : fin (n+1)) : + X.σ i ≫ s.π_summand (index_set.id (op [n+1])) = 0 := +begin + apply s.hom_ext', + intro A, + dsimp only [simplicial_object.σ], + rw [comp_zero, s.ι_summand_epi_naturality_assoc A (simplex_category.σ i).op, + ι_π_summand_eq_zero], + symmetry, + change ¬ (A.epi_comp (simplex_category.σ i).op).eq_id, + rw index_set.eq_id_iff_len_eq, + have h := simplex_category.len_le_of_epi (infer_instance : epi A.e), + dsimp at ⊢ h, + linarith, +end + +end splitting + +end simplicial_object diff --git a/src/algebraic_topology/split_simplicial_object.lean b/src/algebraic_topology/split_simplicial_object.lean index 30b18af9c349e..9bb7d7de43d88 100644 --- a/src/algebraic_topology/split_simplicial_object.lean +++ b/src/algebraic_topology/split_simplicial_object.lean @@ -104,6 +104,49 @@ def id : index_set Δ := ⟨Δ, ⟨𝟙 _, by apply_instance,⟩⟩ instance : inhabited (index_set Δ) := ⟨id Δ⟩ +variable {Δ} + +/-- The condition that an element `splitting.index_set Δ` is the distinguished +element `splitting.index_set.id Δ`. -/ +@[simp] +def eq_id : Prop := A = id _ + +lemma eq_id_iff_eq : A.eq_id ↔ A.1 = Δ := +begin + split, + { intro h, + dsimp at h, + rw h, + refl, }, + { intro h, + rcases A with ⟨Δ', ⟨f, hf⟩⟩, + simp only at h, + subst h, + refine ext _ _ rfl _, + { haveI := hf, + simp only [eq_to_hom_refl, comp_id], + exact simplex_category.eq_id_of_epi f, }, }, +end + +lemma eq_id_iff_len_eq : A.eq_id ↔ A.1.unop.len = Δ.unop.len := +begin + rw eq_id_iff_eq, + split, + { intro h, + rw h, }, + { intro h, + rw ← unop_inj_iff, + ext, + exact h, }, +end + +/-- Given `A : index_set Δ₁`, if `p.unop : unop Δ₂ ⟶ unop Δ₁` is an epi, this +is the obvious element in `A : index_set Δ₂` associated to the composition +of epimorphisms `p.unop ≫ A.e`. -/ +@[simps] +def epi_comp {Δ₁ Δ₂ : simplex_categoryᵒᵖ} (A : index_set Δ₁) (p : Δ₁ ⟶ Δ₂) [epi p.unop] : + index_set Δ₂ := ⟨A.1, ⟨p.unop ≫ A.e, epi_comp _ _⟩⟩ + end index_set variables (N : ℕ → C) (Δ : simplex_categoryᵒᵖ) @@ -237,6 +280,17 @@ def of_iso (e : X ≅ Y) : splitting Y := tidy, end, } +@[reassoc] +lemma ι_summand_epi_naturality {Δ₁ Δ₂ : simplex_categoryᵒᵖ} (A : index_set Δ₁) + (p : Δ₁ ⟶ Δ₂) [epi p.unop] : + s.ι_summand A ≫ X.map p = s.ι_summand (A.epi_comp p) := +begin + dsimp [ι_summand], + erw [colimit.ι_desc, colimit.ι_desc, cofan.mk_ι_app, cofan.mk_ι_app], + dsimp only [index_set.epi_comp, index_set.e], + rw [op_comp, X.map_comp, assoc, quiver.hom.op_unop], +end + end splitting variable (C) From 8853975ed779665849d3936f1663d917c25f4a68 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Wed, 12 Oct 2022 01:16:49 +0000 Subject: [PATCH 700/828] feat(category_theory): yoneda functor on abelian category preserves finite colimits iff object is injective (#16893) --- src/category_theory/abelian/injective.lean | 47 +++++++++++++++++++++ src/category_theory/abelian/projective.lean | 27 ++++++++---- 2 files changed, 66 insertions(+), 8 deletions(-) create mode 100644 src/category_theory/abelian/injective.lean diff --git a/src/category_theory/abelian/injective.lean b/src/category_theory/abelian/injective.lean new file mode 100644 index 0000000000000..c4dabce7e96aa --- /dev/null +++ b/src/category_theory/abelian/injective.lean @@ -0,0 +1,47 @@ +/- +Copyright (c) 2022 Jakob von Raumer. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jakob von Raumer +-/ + +import category_theory.abelian.exact +import category_theory.preadditive.injective +import category_theory.preadditive.yoneda + +/-! +# Injective objects in abelian categories + +* Objects in an abelian categories are injective if and only if the preadditive Yoneda functor + on them preserves finite colimits. +-/ + +noncomputable theory + +open category_theory +open category_theory.limits +open category_theory.injective +open opposite + +universes v u + +namespace category_theory + +variables {C : Type u} [category.{v} C] [abelian C] + +/-- The preadditive Yoneda functor on `J` preserves colimits if `J` is injective. -/ +def preserves_finite_colimits_preadditive_yoneda_obj_of_injective (J : C) + [hP : injective J] : preserves_finite_colimits (preadditive_yoneda_obj J) := +begin + letI := (injective_iff_preserves_epimorphisms_preadditive_yoneda_obj' J).mp hP, + apply functor.preserves_finite_colimits_of_preserves_epis_and_kernels, +end + +/-- An object is injective if its preadditive Yoneda functor preserves finite colimits. -/ +lemma injective_of_preserves_finite_colimits_preadditive_yoneda_obj (J : C) + [hP : preserves_finite_colimits (preadditive_yoneda_obj J)] : injective J := +begin + rw injective_iff_preserves_epimorphisms_preadditive_yoneda_obj', + apply_instance +end + +end category_theory diff --git a/src/category_theory/abelian/projective.lean b/src/category_theory/abelian/projective.lean index 2404efdbfe317..78217b91189f4 100644 --- a/src/category_theory/abelian/projective.lean +++ b/src/category_theory/abelian/projective.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Markus Himmel, Scott Morrison +Authors: Markus Himmel, Scott Morrison, Jakob von Raumer -/ import category_theory.abelian.exact import category_theory.preadditive.projective_resolution @@ -18,6 +18,7 @@ noncomputable theory open category_theory open category_theory.limits +open opposite universes v u @@ -25,18 +26,29 @@ namespace category_theory open category_theory.projective -variables {C : Type u} [category.{v} C] - -section -variables [enough_projectives C] [abelian C] +variables {C : Type u} [category.{v} C] [abelian C] /-- When `C` is abelian, `projective.d f` and `f` are exact. -/ -lemma exact_d_f {X Y : C} (f : X ⟶ Y) : exact (d f) f := +lemma exact_d_f [enough_projectives C] {X Y : C} (f : X ⟶ Y) : exact (d f) f := (abelian.exact_iff _ _).2 $ ⟨by simp, zero_of_epi_comp (π _) $ by rw [←category.assoc, cokernel.condition]⟩ +/-- The preadditive Co-Yoneda functor on `P` preserves colimits if `P` is projective. -/ +def preserves_finite_colimits_preadditive_coyoneda_obj_of_projective (P : C) + [hP : projective P] : preserves_finite_colimits (preadditive_coyoneda_obj (op P)) := +begin + letI := (projective_iff_preserves_epimorphisms_preadditive_coyoneda_obj' P).mp hP, + apply functor.preserves_finite_colimits_of_preserves_epis_and_kernels, +end + +/-- An object is projective if its preadditive Co-Yoneda functor preserves finite colimits. -/ +lemma projective_of_preserves_finite_colimits_preadditive_coyoneda_obj (P : C) + [hP : preserves_finite_colimits (preadditive_coyoneda_obj (op P))] : projective P := +begin + rw projective_iff_preserves_epimorphisms_preadditive_coyoneda_obj', + apply_instance end namespace ProjectiveResolution @@ -49,8 +61,7 @@ After that, we build the `n+1`-st object as `projective.syzygies` applied to the previously constructed morphism, and the map to the `n`-th object as `projective.d`. -/ - -variables [abelian C] [enough_projectives C] +variables [enough_projectives C] /-- Auxiliary definition for `ProjectiveResolution.of`. -/ @[simps] From 4d4167104581a21259f7f448e1972a63a4546be7 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Wed, 12 Oct 2022 01:16:52 +0000 Subject: [PATCH 701/828] chore(*): remove uses of with_cases (#16894) For the port, resolves #16568. I did a pretty crude job, so the proofs are likely more brittle and ugly than before, but at least for the `rbtree` file it doesn't seem anoyone really needs to read these proofs often. One fun side-effect of this is that the linter started complaining about unused arguments, so we fix those too. --- src/data/rbtree/insert.lean | 124 ++++++++++++---------------- src/data/rbtree/main.lean | 4 +- src/linear_algebra/alternating.lean | 21 +++-- 3 files changed, 66 insertions(+), 83 deletions(-) diff --git a/src/data/rbtree/insert.lean b/src/data/rbtree/insert.lean index 955270781fd37..565f6f98c7eb8 100644 --- a/src/data/rbtree/insert.lean +++ b/src/data/rbtree/insert.lean @@ -149,20 +149,20 @@ lemma is_searchable_ins [decidable_rel lt] {t x} [is_strict_weak_order α lt] : ∀ {lo hi} (h : is_searchable lt t lo hi), lift lt lo (some x) → lift lt (some x) hi → is_searchable lt (ins lt t x) lo hi := begin - with_cases { apply ins.induction lt t x; intros; simp! [*] at * {eta := ff}; - is_searchable_tactic }, - case is_red_lt hs₁ { apply ih h_hs₁, assumption, simp [*] }, - case is_red_eq hs₁ { apply is_searchable_of_is_searchable_of_incomp hc, assumption }, - case is_red_eq hs₂ { apply is_searchable_of_incomp_of_is_searchable hc, assumption }, - case is_red_gt hs₂ { apply ih h_hs₂, cases hi; simp [*], assumption }, - case is_black_lt_red { apply is_searchable_balance1_node, apply ih h_hs₁, assumption, simp [*], + apply ins.induction lt t x; intros; simp! [*] at * {eta := ff}; + is_searchable_tactic, + { apply ih h_hs₁, assumption, simp [*] }, + { apply is_searchable_of_is_searchable_of_incomp hc, assumption }, + { apply is_searchable_of_incomp_of_is_searchable hc, assumption }, + { apply ih h_hs₂, cases hi; simp [*], assumption }, + { apply is_searchable_balance1_node, apply ih h_hs₁, assumption, simp [*], assumption }, - case is_black_lt_not_red hs₁ { apply ih h_hs₁, assumption, simp [*] }, - case is_black_eq hs₁ { apply is_searchable_of_is_searchable_of_incomp hc, assumption }, - case is_black_eq hs₂ { apply is_searchable_of_incomp_of_is_searchable hc, assumption }, - case is_black_gt_red { apply is_searchable_balance2_node, assumption, apply ih h_hs₂, simp [*], + { apply ih h_hs₁, assumption, simp [*] }, + { apply is_searchable_of_is_searchable_of_incomp hc, assumption }, + { apply is_searchable_of_incomp_of_is_searchable hc, assumption }, + { apply is_searchable_balance2_node, assumption, apply ih h_hs₂, simp [*], assumption }, - case is_black_gt_not_red hs₂ { apply ih h_hs₂, assumption, simp [*] } + { apply ih h_hs₂, assumption, simp [*] } end lemma is_searchable_mk_insert_result {c t} : is_searchable lt t none none → @@ -251,36 +251,27 @@ end lemma mem_ins_of_incomp [decidable_rel lt] (t : rbnode α) {x y : α} : ∀ h : ¬ lt x y ∧ ¬ lt y x, x ∈ t.ins lt y := begin - with_cases { apply ins.induction lt t y; intros; simp [ins, *] }, - case is_black_lt_red { have := ih h, apply mem_balance1_node_of_mem_left, assumption }, - case is_black_gt_red { have := ih h, apply mem_balance2_node_of_mem_left, assumption } + apply ins.induction lt t y; intros; simp [ins, *], + { have := ih h, apply mem_balance1_node_of_mem_left, assumption }, + { have := ih h, apply mem_balance2_node_of_mem_left, assumption } end lemma mem_ins_of_mem [decidable_rel lt] [is_strict_weak_order α lt] {t : rbnode α} (z : α) : ∀ {x} (h : x ∈ t), x ∈ t.ins lt z := begin - with_cases { apply ins.induction lt t z; intros; simp [ins, *] at *; try { contradiction }; - blast_disjs }, - case is_red_eq or.inr or.inl + apply ins.induction lt t z; intros; simp [ins, *] at *; try { contradiction }; + blast_disjs, + any_goals { intros, simp [h], done }, + any_goals { intros, simp [ih h], done }, { have := incomp_trans_of lt h ⟨hc.2, hc.1⟩, simp [this] }, - case is_black_lt_red or.inl { apply mem_balance1_node_of_mem_left, apply ih h }, - case is_black_lt_red or.inr or.inl { apply mem_balance1_node_of_incomp, cases h, all_goals { simp [*, ins_ne_leaf lt a z] } }, - case is_black_lt_red or.inr or.inr { apply mem_balance1_node_of_mem_right, assumption }, - case is_black_eq or.inr or.inl { have := incomp_trans_of lt hc ⟨h.2, h.1⟩, simp [this] }, - case is_black_gt_red or.inl { apply mem_balance2_node_of_mem_right, assumption }, - case is_black_gt_red or.inr or.inl { have := ins_ne_leaf lt a z, apply mem_balance2_node_of_incomp, cases h, simp [*], apply ins_ne_leaf }, - case is_black_gt_red or.inr or.inr { apply mem_balance2_node_of_mem_left, apply ih h }, - -- remaining cases are easy - any_goals { intros, simp [h], done }, - all_goals { intros, simp [ih h], done }, end lemma mem_mk_insert_result {a t} (c) : mem lt a t → mem lt a (mk_insert_result c t) := @@ -315,26 +306,22 @@ begin simp [*] } end -lemma equiv_or_mem_of_mem_ins [decidable_rel lt] [is_strict_weak_order α lt] {t : rbnode α} {x z} : +lemma equiv_or_mem_of_mem_ins [decidable_rel lt] {t : rbnode α} {x z} : ∀ (h : x ∈ t.ins lt z), x ≈[lt] z ∨ x ∈ t := begin - with_cases { apply ins.induction lt t z; intros; simp [ins, strict_weak_order.equiv, *] at *; - blast_disjs }, - case is_black_lt_red - { have h' := of_mem_balance1_node lt h, blast_disjs, - have := ih h', blast_disjs, - all_goals { simp [h, *] } }, - case is_black_gt_red - { have h' := of_mem_balance2_node lt h, blast_disjs, - have := ih h', blast_disjs, - all_goals { simp [h, *] }}, - -- All other goals can be solved by the following tactics + apply ins.induction lt t z; intros; simp [ins, strict_weak_order.equiv, *] at *; + blast_disjs, any_goals { intros, simp [h] }, - all_goals { intros, have ih := ih h, cases ih; simp [*], done }, + any_goals { intros, have ih := ih h, cases ih; simp [*], done }, + { have h' := of_mem_balance1_node lt h, blast_disjs, + have := ih h', blast_disjs, + all_goals { simp [h, *] } }, + { have h' := of_mem_balance2_node lt h, blast_disjs, + have := ih h', blast_disjs, + all_goals { simp [h, *] }}, end -lemma equiv_or_mem_of_mem_insert [decidable_rel lt] [is_strict_weak_order α lt] {t : rbnode α} - {x z} : +lemma equiv_or_mem_of_mem_insert [decidable_rel lt] {t : rbnode α} {x z} : ∀ (h : x ∈ t.insert lt z), x ≈[lt] z ∨ x ∈ t := begin simp [insert], intros, apply equiv_or_mem_of_mem_ins, exact mem_of_mem_mk_insert_result lt h @@ -489,13 +476,12 @@ lemma find_balance1_lt {l r t v x y lo hi} (ht : is_searchable lt t (some y) hi) : find lt (balance1 l v r y t) x = find lt (red_node l v r) x := begin - with_cases { revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic }, - case red_left : _ _ _ z r { apply weak_trichotomous lt z x; intros; simp [*] }, - case red_right : l_left l_val l_right z r - { with_cases { apply weak_trichotomous lt z x; intro h' }, - case is_lt { have := trans_of lt (lo_lt_hi hr_hs₁) h', simp [*] }, - case is_eqv { have : lt l_val x := lt_of_lt_of_incomp (lo_lt_hi hr_hs₁) h', simp [*] }, - case is_gt { apply weak_trichotomous lt l_val x; intros; simp [*] } } + revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic, + { apply weak_trichotomous lt y_1 x; intros; simp [*] }, + { apply weak_trichotomous lt x_1 x; intro h', + { have := trans_of lt (lo_lt_hi hr_hs₁) h', simp [*] }, + { have : lt y_1 x := lt_of_lt_of_incomp (lo_lt_hi hr_hs₁) h', simp [*] }, + { apply weak_trichotomous lt y_1 x; intros; simp [*] } } end meta def ins_ne_leaf_tac := `[apply ins_ne_leaf] @@ -518,10 +504,8 @@ lemma find_balance1_gt {l r t v x y lo hi} (ht : is_searchable lt t (some y) hi) : find lt (balance1 l v r y t) x = find lt t x := begin - with_cases { revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic }, - case red_left : _ _ _ z + revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic, { have := trans_of lt (lo_lt_hi hr) h, simp [*] }, - case red_right : _ _ _ z { have := trans_of lt (lo_lt_hi hr_hs₂) h, simp [*] } end @@ -542,12 +526,10 @@ lemma find_balance1_eqv {l r t v x y lo hi} (ht : is_searchable lt t (some y) hi) : find lt (balance1 l v r y t) x = some y := begin - with_cases { revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic }, - case red_left : _ _ _ z - { have : lt z x := lt_of_lt_of_incomp (lo_lt_hi hr) h.swap, + revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic, + { have : lt y_1 x := lt_of_lt_of_incomp (lo_lt_hi hr) h.swap, simp [*] }, - case red_right : _ _ _ z - { have : lt z x := lt_of_lt_of_incomp (lo_lt_hi hr_hs₂) h.swap, + { have : lt x_1 x := lt_of_lt_of_incomp (lo_lt_hi hr_hs₂) h.swap, simp [*] } end @@ -570,9 +552,9 @@ lemma find_balance2_lt {l v r t x y lo hi} (ht : is_searchable lt t lo (some y)) : find lt (balance2 l v r y t) x = find lt t x := begin - with_cases { revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic }, - case red_left { have := trans h (lo_lt_hi hl_hs₁), simp [*] }, - case red_right { have := trans h (lo_lt_hi hl), simp [*] } + revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic, + { have := trans h (lo_lt_hi hl_hs₁), simp [*] }, + { have := trans h (lo_lt_hi hl), simp [*] } end lemma find_balance2_node_lt {s t x y lo hi} @@ -593,14 +575,12 @@ lemma find_balance2_gt {l v r t x y lo hi} (ht : is_searchable lt t lo (some y)) : find lt (balance2 l v r y t) x = find lt (red_node l v r) x := begin - with_cases { revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic }, - case red_left : _ val _ z - { with_cases { apply weak_trichotomous lt val x; intro h'; simp [*] }, - case is_lt { apply weak_trichotomous lt z x; intros; simp [*] }, - case is_eqv { have : lt x z := lt_of_incomp_of_lt h'.swap (lo_lt_hi hl_hs₂), simp [*] }, - case is_gt { have := trans h' (lo_lt_hi hl_hs₂), simp [*] } }, - case red_right : _ val - { apply weak_trichotomous lt val x; intros; simp [*] } + revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic, + { apply weak_trichotomous lt x_1 x; intro h'; simp [*], + { apply weak_trichotomous lt y_1 x; intros; simp [*] }, + { have : lt x _ := lt_of_incomp_of_lt h'.swap (lo_lt_hi hl_hs₂), simp [*] }, + { have := trans h' (lo_lt_hi hl_hs₂), simp [*] } }, + { apply weak_trichotomous lt y_1 x; intros; simp [*] } end lemma find_balance2_node_gt {s t x y lo hi} @@ -622,9 +602,9 @@ lemma find_balance2_eqv {l v r t x y lo hi} (ht : is_searchable lt t lo (some y)) : find lt (balance2 l v r y t) x = some y := begin - with_cases { revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic }, - case red_left { have := lt_of_incomp_of_lt h (lo_lt_hi hl_hs₁), simp [*] }, - case red_right { have := lt_of_incomp_of_lt h (lo_lt_hi hl), simp [*] } + revert hl hr ht, apply balance.cases l v r; intros; simp [*]; is_searchable_tactic, + { have := lt_of_incomp_of_lt h (lo_lt_hi hl_hs₁), simp [*] }, + { have := lt_of_incomp_of_lt h (lo_lt_hi hl), simp [*] } end lemma find_balance2_node_eqv {t s x y lo hi} diff --git a/src/data/rbtree/main.lean b/src/data/rbtree/main.lean index f8a148a4a12f1..8c49c3943b5a8 100644 --- a/src/data/rbtree/main.lean +++ b/src/data/rbtree/main.lean @@ -178,11 +178,11 @@ lemma mem_insert_of_mem [is_strict_weak_order α lt] {a : α} {t : rbtree α lt} a ∈ t → a ∈ t.insert b := begin cases t, apply rbnode.mem_insert_of_mem end -lemma equiv_or_mem_of_mem_insert [is_strict_weak_order α lt] {a b : α} {t : rbtree α lt} : +lemma equiv_or_mem_of_mem_insert {a b : α} {t : rbtree α lt} : a ∈ t.insert b → a ≈[lt] b ∨ a ∈ t := begin cases t, apply rbnode.equiv_or_mem_of_mem_insert end -lemma incomp_or_mem_of_mem_ins [is_strict_weak_order α lt] {a b : α} {t : rbtree α lt} : +lemma incomp_or_mem_of_mem_ins {a b : α} {t : rbtree α lt} : a ∈ t.insert b → (¬ lt a b ∧ ¬ lt b a) ∨ a ∈ t := equiv_or_mem_of_mem_insert diff --git a/src/linear_algebra/alternating.lean b/src/linear_algebra/alternating.lean index 9a9988761210b..ae56a066a3d5b 100644 --- a/src/linear_algebra/alternating.lean +++ b/src/linear_algebra/alternating.lean @@ -769,22 +769,25 @@ begin dsimp only [quotient.lift_on'_mk', quotient.map'_mk', multilinear_map.smul_apply, multilinear_map.dom_dom_congr_apply, multilinear_map.dom_coprod_apply, dom_coprod.summand], intro hσ, - with_cases - { cases hi : σ⁻¹ i; - cases hj : σ⁻¹ j; - rw perm.inv_eq_iff_eq at hi hj; - substs hi hj, }, - case [sum.inl sum.inr : i' j', sum.inr sum.inl : i' j'] + cases hi : σ⁻¹ i; + cases hj : σ⁻¹ j; + rw perm.inv_eq_iff_eq at hi hj; + substs hi hj; revert val val_1, + case [sum.inl sum.inr, sum.inr sum.inl] { -- the term pairs with and cancels another term - all_goals { obtain ⟨⟨sl, sr⟩, hσ⟩ := quotient_group.left_rel_apply.mp (quotient.exact' hσ), }, + all_goals { + intros i' j' hv hij hσ, + obtain ⟨⟨sl, sr⟩, hσ⟩ := quotient_group.left_rel_apply.mp (quotient.exact' hσ), }, work_on_goal 1 { replace hσ := equiv.congr_fun hσ (sum.inl i'), }, work_on_goal 2 { replace hσ := equiv.congr_fun hσ (sum.inr i'), }, all_goals { rw [smul_eq_mul, ←mul_swap_eq_swap_mul, mul_inv_rev, swap_inv, inv_mul_cancel_right] at hσ, simpa using hσ, }, }, - case [sum.inr sum.inr : i' j', sum.inl sum.inl : i' j'] + case [sum.inr sum.inr, sum.inl sum.inl] { -- the term does not pair but is zero - all_goals { convert smul_zero _, }, + all_goals { + intros i' j' hv hij hσ, + convert smul_zero _, }, work_on_goal 1 { convert tensor_product.tmul_zero _ _, }, work_on_goal 2 { convert tensor_product.zero_tmul _ _, }, all_goals { exact alternating_map.map_eq_zero_of_eq _ _ hv (λ hij', hij (hij' ▸ rfl)), } }, From c5549d56d24132ba4db3f53ffa3083d2f40cae8c Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 12 Oct 2022 01:16:53 +0000 Subject: [PATCH 702/828] feat(data/option/basic): add `option.{forall,exists}_mem_map` (#16897) --- src/data/option/basic.lean | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/src/data/option/basic.lean b/src/data/option/basic.lean index e02a2de4d7611..914c176d53e2e 100644 --- a/src/data/option/basic.lean +++ b/src/data/option/basic.lean @@ -214,9 +214,19 @@ lemma comp_map (h : β → γ) (g : α → β) (x : option α) : option.map g ∘ option.map f = option.map (g ∘ f) := by { ext x, rw comp_map } -lemma mem_map_of_mem {α β : Type*} {a : α} {x : option α} (g : α → β) (h : a ∈ x) : g a ∈ x.map g := +lemma mem_map_of_mem {a : α} {x : option α} (g : α → β) (h : a ∈ x) : g a ∈ x.map g := mem_def.mpr ((mem_def.mp h).symm ▸ map_some') +lemma mem_map {f : α → β} {y : β} {o : option α} : y ∈ o.map f ↔ ∃ x ∈ o, f x = y := by simp + +lemma forall_mem_map {f : α → β} {o : option α} {p : β → Prop} : + (∀ y ∈ o.map f, p y) ↔ ∀ x ∈ o, p (f x) := +by simp + +lemma exists_mem_map {f : α → β} {o : option α} {p : β → Prop} : + (∃ y ∈ o.map f, p y) ↔ ∃ x ∈ o, p (f x) := +by simp + lemma bind_map_comm {α β} {x : option (option α) } {f : α → β} : x >>= option.map f = x.map (option.map f) >>= id := by { cases x; simp } From 113568c9586be44bf6ba1181bfc11aeb21380463 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 12 Oct 2022 01:16:56 +0000 Subject: [PATCH 703/828] refactor(archive/100-theorems-list/02_cubing_a_cube): review API (#16898) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * reuse lemmas about `set.pi`; * fix some `def`/`lemma`; * use `nontrivial` and `set.nontrivial` instead of `2 ≤ cardinal.mk ι`; * redefine `correct` as a `structure`. --- .../100-theorems-list/82_cubing_a_cube.lean | 177 ++++++++---------- 1 file changed, 83 insertions(+), 94 deletions(-) diff --git a/archive/100-theorems-list/82_cubing_a_cube.lean b/archive/100-theorems-list/82_cubing_a_cube.lean index f1e94be0ad83c..b9a83c4c8b248 100644 --- a/archive/100-theorems-list/82_cubing_a_cube.lean +++ b/archive/100-theorems-list/82_cubing_a_cube.lean @@ -7,7 +7,6 @@ import data.fin.tuple import data.real.basic import data.set.intervals import data.set.pairwise -import set_theory.cardinal.basic /-! Proof that a cube (in dimension n ≥ 3) cannot be cubed: @@ -20,7 +19,6 @@ http://www.alaricstephen.com/main-featured/2017/9/28/cubing-a-cube-proof open real set function fin -open_locale cardinal noncomputable theory @@ -63,26 +61,17 @@ by simp [side, cube.hw, le_refl] def to_set (c : cube n) : set (fin n → ℝ) := { x | ∀j, x j ∈ side c j } -def to_set_subset {c c' : cube n} : c.to_set ⊆ c'.to_set ↔ ∀j, c.side j ⊆ c'.side j := -begin - split, intros h j x hx, - let f : fin n → ℝ := λ j', if j' = j then x else c.b j', - have : f ∈ c.to_set, - { intro j', by_cases hj' : j' = j; simp [f, hj', if_pos, if_neg, hx] }, - convert h this j, { simp [f, if_pos] }, - intros h f hf j, exact h j (hf j) -end +lemma side_nonempty (c : cube n) (i : fin n) : (side c i).nonempty := by simp [side, c.hw] -def to_set_disjoint {c c' : cube n} : disjoint c.to_set c'.to_set ↔ - ∃j, disjoint (c.side j) (c'.side j) := -begin - split, intros h, classical, by_contra h', - simp only [not_disjoint_iff, classical.skolem, not_exists] at h', - cases h' with f hf, - apply not_disjoint_iff.mpr ⟨f, _, _⟩ h; intro j, exact (hf j).1, exact (hf j).2, - rintro ⟨j, hj⟩, rw [set.disjoint_iff], rintros f ⟨h1f, h2f⟩, - apply not_disjoint_iff.mpr ⟨f j, h1f j, h2f j⟩ hj -end +lemma univ_pi_side (c : cube n) : pi univ (side c) = c.to_set := ext $ λ x, mem_univ_pi + +lemma to_set_subset {c c' : cube n} : c.to_set ⊆ c'.to_set ↔ ∀j, c.side j ⊆ c'.side j := +by simp only [← univ_pi_side, univ_pi_subset_univ_pi_iff, (c.side_nonempty _).ne_empty, + exists_false, or_false] + +lemma to_set_disjoint {c c' : cube n} : disjoint c.to_set c'.to_set ↔ + ∃ j, disjoint (c.side j) (c'.side j) := +by simp only [← univ_pi_side, disjoint_univ_pi] lemma b_mem_to_set (c : cube n) : c.b ∈ c.to_set := by simp [to_set] @@ -121,25 +110,25 @@ by norm_num [unit_cube, side] end cube open cube -variables {ι : Type} [fintype ι] {cs : ι → cube (n+1)} {i i' : ι} +variables {ι : Type} {cs : ι → cube (n+1)} {i i' : ι} /-- A finite family of (at least 2) cubes partitioning the unit cube with different sizes -/ -def correct (cs : ι → cube n) : Prop := -pairwise (disjoint on (cube.to_set ∘ cs)) ∧ -(⋃(i : ι), (cs i).to_set) = unit_cube.to_set ∧ -injective (cube.w ∘ cs) ∧ -2 ≤ #ι ∧ -3 ≤ n +@[protect_proj] structure correct (cs : ι → cube n) : Prop := +(pairwise_disjoint : pairwise (disjoint on (cube.to_set ∘ cs))) +(Union_eq : (⋃(i : ι), (cs i).to_set) = unit_cube.to_set) +(injective : injective (cube.w ∘ cs)) +(three_le : 3 ≤ n) -variable (h : correct cs) +namespace correct +variable (h : correct cs) include h + lemma to_set_subset_unit_cube {i} : (cs i).to_set ⊆ unit_cube.to_set := -by { rw [←h.2.1], exact subset_Union _ i } +h.Union_eq ▸ subset_Union _ i lemma side_subset {i j} : (cs i).side j ⊆ Ico 0 1 := -by { have := to_set_subset_unit_cube h, rw [to_set_subset] at this, - convert this j, norm_num [unit_cube] } +by simpa only [side_unit_cube] using to_set_subset.1 h.to_set_subset_unit_cube j lemma zero_le_of_mem_side {i j x} (hx : x ∈ (cs i).side j) : 0 ≤ x := (side_subset h hx).1 @@ -153,11 +142,14 @@ zero_le_of_mem h (cs i).b_mem_to_set j lemma b_add_w_le_one {j} : (cs i).b j + (cs i).w ≤ 1 := by { have := side_subset h, rw [side, Ico_subset_Ico_iff] at this, convert this.2, simp [hw] } +lemma nontrivial_fin : nontrivial (fin n) := +fin.nontrivial_iff_two_le.2 (nat.le_of_succ_le_succ h.three_le) + /-- The width of any cube in the partition cannot be 1. -/ -lemma w_ne_one (i : ι) : (cs i).w ≠ 1 := +lemma w_ne_one [nontrivial ι] (i : ι) : (cs i).w ≠ 1 := begin intro hi, - have := h.2.2.2.1, rw [cardinal.two_le_iff' i] at this, cases this with i' hi', + cases exists_ne i with i' hi', let p := (cs i').b, have hp : p ∈ (cs i').to_set := (cs i').b_mem_to_set, have h2p : p ∈ (cs i).to_set, @@ -166,8 +158,7 @@ begin { rw [←add_le_add_iff_right (1 : ℝ)], convert b_add_w_le_one h, rw hi, rw zero_add }, apply zero_le_b h, apply lt_of_lt_of_le (side_subset h $ (cs i').b_mem_side j).2, simp [hi, zero_le_b h] }, - apply not_disjoint_iff.mpr ⟨p, hp, h2p⟩, - apply h.1, exact hi'.symm + exact h.pairwise_disjoint i' i hi' ⟨hp, h2p⟩ end /-- The top of a cube (which is the bottom of the cube shifted up by its width) must be covered by @@ -182,7 +173,7 @@ begin { rw [←zero_add (0 : ℝ)], apply add_le_add, apply zero_le_b h, apply (cs i).hw' }, { exact lt_of_le_of_ne (b_add_w_le_one h) hc }, intro j, exact side_subset h (hps j) }, - rw [←h.2.1] at this, rcases this with ⟨_, ⟨i', rfl⟩, hi'⟩, + rw [← h.2, mem_Union] at this, rcases this with ⟨i', hi'⟩, rw [mem_Union], use i', refine ⟨_, λ j, hi' j.succ⟩, have : i ≠ i', { rintro rfl, apply not_le_of_lt (hi' 0).2, rw [hp0], refl }, have := h.1 i i' this, rw [on_fun, to_set_disjoint, exists_fin_succ] at this, @@ -191,7 +182,8 @@ begin convert hi' 0, rw [hp0], refl, exfalso, apply not_disjoint_iff.mpr ⟨tail p j, hps j, hi' j.succ⟩ hj end -omit h + +end correct /-- A valley is a square on which cubes in the family of cubes are placed, so that the cubes completely cover the valley and none of those cubes is partially outside the square. @@ -208,10 +200,10 @@ c.bottom ⊆ (⋃(i : ι), (cs i).bottom) ∧ (cs i).tail.to_set ⊆ c.tail.to_set) ∧ ∀(i : ι), (cs i).b 0 = c.b 0 → (cs i).w ≠ c.w -variables {c : cube (n+1)} (v : valley cs c) +variables {c : cube (n+1)} (h : correct cs) (v : valley cs c) /-- The bottom of the unit cube is a valley -/ -lemma valley_unit_cube (h : correct cs) : valley cs unit_cube := +lemma valley_unit_cube [nontrivial ι] (h : correct cs) : valley cs unit_cube := begin refine ⟨_, _, _⟩, { intro v, @@ -220,12 +212,12 @@ begin have : v ∈ (unit_cube : cube (n+1)).to_set, { dsimp only [to_set, unit_cube, mem_set_of_eq], rw [forall_fin_succ, h0], split, norm_num [side, unit_cube], exact hv }, - rw [←h.2.1] at this, rcases this with ⟨_, ⟨i, rfl⟩, hi⟩, + rw [← h.2, mem_Union] at this, rcases this with ⟨i, hi⟩, use i, - split, { apply le_antisymm, rw h0, exact zero_le_b h, exact (hi 0).1 }, + split, { apply le_antisymm, rw h0, exact h.zero_le_b, exact (hi 0).1 }, intro j, exact hi _ }, - { intros i hi h', rw to_set_subset, intro j, convert side_subset h using 1, simp [side_tail] }, - { intros i hi, exact w_ne_one h i } + { intros i hi h', rw to_set_subset, intro j, convert h.side_subset using 1, simp [side_tail] }, + { intros i hi, exact h.w_ne_one i } end /-- the cubes which lie in the valley `c` -/ @@ -257,20 +249,18 @@ include h v lemma w_lt_w (hi : i ∈ bcubes cs c) : (cs i).w < c.w := begin apply lt_of_le_of_ne _ (v.2.2 i hi.1), - have j : fin n := ⟨1, nat.le_of_succ_le_succ h.2.2.2.2⟩, + have j : fin n := ⟨1, nat.le_of_succ_le_succ h.three_le⟩, rw [←add_le_add_iff_left ((cs i).b j.succ)], apply le_trans (t_le_t hi j), rw [add_le_add_iff_right], apply b_le_b hi, end -open cardinal /-- There are at least two cubes in a valley -/ -lemma two_le_mk_bcubes : 2 ≤ #(bcubes cs c) := +lemma nontrivial_bcubes : (bcubes cs c).nontrivial := begin - rw [two_le_iff], rcases v.1 c.b_mem_bottom with ⟨_, ⟨i, rfl⟩, hi⟩, have h2i : i ∈ bcubes cs c := ⟨hi.1.symm, v.2.1 i hi.1.symm ⟨tail c.b, hi.2, λ j, c.b_mem_side j.succ⟩⟩, - let j : fin (n+1) := ⟨2, h.2.2.2.2⟩, + let j : fin (n+1) := ⟨2, h.three_le⟩, have hj : 0 ≠ j := by { simp only [fin.ext_iff, ne.def], contradiction }, let p : fin (n+1) → ℝ := λ j', if j' = j then c.b j + (cs i).w else c.b j', have hp : p ∈ c.bottom, @@ -281,10 +271,10 @@ begin { simp [p, -add_comm, if_neg hj'] }}, rcases v.1 hp with ⟨_, ⟨i', rfl⟩, hi'⟩, have h2i' : i' ∈ bcubes cs c := ⟨hi'.1.symm, v.2.1 i' hi'.1.symm ⟨tail p, hi'.2, hp.2⟩⟩, - refine ⟨⟨i, h2i⟩, ⟨i', h2i'⟩, _⟩, - intro hii', cases congr_arg subtype.val hii', - apply not_le_of_lt (hi'.2 ⟨1, nat.le_of_succ_le_succ h.2.2.2.2⟩).2, - simp only [-add_comm, tail, cube.tail, p], + refine ⟨i, h2i, i', h2i', _⟩, + rintro rfl, + apply not_le_of_lt (hi'.2 ⟨1, nat.le_of_succ_le_succ h.three_le⟩).2, + simp only [tail, cube.tail, p], rw [if_pos, add_le_add_iff_right], { exact (hi.2 _).1 }, refl @@ -292,39 +282,36 @@ end /-- There is a cube in the valley -/ lemma nonempty_bcubes : (bcubes cs c).nonempty := -begin - rw [←set.ne_empty_iff_nonempty], intro h', have := two_le_mk_bcubes h v, rw h' at this, - apply not_lt_of_le this, rw mk_emptyc, norm_cast, norm_num -end +(nontrivial_bcubes h v).nonempty + +variables [finite ι] /-- There is a smallest cube in the valley -/ -lemma exists_mi : ∃(i : ι), i ∈ bcubes cs c ∧ ∀(i' ∈ bcubes cs c), +lemma exists_mi : ∃ i ∈ bcubes cs c, ∀ i' ∈ bcubes cs c, (cs i).w ≤ (cs i').w := -by simpa - using (bcubes cs c).exists_min_image (λ i, (cs i).w) (set.to_finite _) (nonempty_bcubes h v) +(bcubes cs c).exists_min_image (λ i, (cs i).w) (set.to_finite _) (nonempty_bcubes h v) /-- We let `mi` be the (index for the) smallest cube in the valley `c` -/ def mi : ι := classical.some $ exists_mi h v variables {h v} lemma mi_mem_bcubes : mi h v ∈ bcubes cs c := -(classical.some_spec $ exists_mi h v).1 +(classical.some_spec $ exists_mi h v).fst lemma mi_minimal (hi : i ∈ bcubes cs c) : (cs $ mi h v).w ≤ (cs i).w := -(classical.some_spec $ exists_mi h v).2 i hi +(classical.some_spec $ exists_mi h v).snd i hi lemma mi_strict_minimal (hii' : mi h v ≠ i) (hi : i ∈ bcubes cs c) : (cs $ mi h v).w < (cs i).w := -by { apply lt_of_le_of_ne (mi_minimal hi), apply h.2.2.1.ne, apply hii' } +(mi_minimal hi).lt_of_ne $ h.injective.ne hii' /-- The top of `mi` cannot be 1, since there is a larger cube in the valley -/ lemma mi_xm_ne_one : (cs $ mi h v).xm ≠ 1 := begin - apply ne_of_lt, rcases (two_le_iff' _).mp (two_le_mk_bcubes h v) with ⟨⟨i, hi⟩, h2i⟩, - swap, exact ⟨mi h v, mi_mem_bcubes⟩, - apply lt_of_lt_of_le _ (b_add_w_le_one h), exact i, exact 0, + apply ne_of_lt, rcases (nontrivial_bcubes h v).exists_ne (mi h v) with ⟨i, hi, h2i⟩, + apply lt_of_lt_of_le _ h.b_add_w_le_one, exact i, exact 0, rw [xm, mi_mem_bcubes.1, hi.1, _root_.add_lt_add_iff_left], - apply mi_strict_minimal _ hi, intro h', apply h2i, rw subtype.ext_iff_val, exact h' + exact mi_strict_minimal h2i.symm hi end /-- If `mi` lies on the boundary of the valley in dimension j, then this lemma expresses that all @@ -334,7 +321,7 @@ end coordinate `x` -/ lemma smallest_on_boundary {j} (bi : on_boundary (mi_mem_bcubes : mi h v ∈ _) j) : ∃(x : ℝ), x ∈ c.side j.succ \ (cs $ mi h v).side j.succ ∧ - ∀{{i'}} (hi' : i' ∈ bcubes cs c), i' ≠ mi h v → + ∀ ⦃i'⦄ (hi' : i' ∈ bcubes cs c), i' ≠ mi h v → (cs $ mi h v).b j.succ ∈ (cs i').side j.succ → x ∈ (cs i').side j.succ := begin let i := mi h v, have hi : i ∈ bcubes cs c := mi_mem_bcubes, @@ -348,8 +335,8 @@ begin simp [bi.symm, b_le_b hi'] }, let s := bcubes cs c \ { i }, have hs : s.nonempty, - { rcases (two_le_iff' (⟨i, hi⟩ : bcubes cs c)).mp (two_le_mk_bcubes h v) with ⟨⟨i', hi'⟩, h2i'⟩, - refine ⟨i', hi', _⟩, simp only [mem_singleton_iff], intro h, apply h2i', simp [h] }, + { rcases (nontrivial_bcubes h v).exists_ne i with ⟨i', hi', h2i'⟩, + exact ⟨i', hi', h2i'⟩ }, rcases set.exists_min_image s (w ∘ cs) (set.to_finite _) hs with ⟨i', ⟨hi', h2i'⟩, h3i'⟩, rw [mem_singleton_iff] at h2i', let x := c.b j.succ + c.w - (cs i').w, @@ -374,8 +361,8 @@ variables (h v) lemma mi_not_on_boundary (j : fin n) : ¬on_boundary (mi_mem_bcubes : mi h v ∈ _) j := begin let i := mi h v, have hi : i ∈ bcubes cs c := mi_mem_bcubes, - rcases (two_le_iff' j).mp _ with ⟨j', hj'⟩, swap, - { rw [mk_fin, ←nat.cast_two, nat_cast_le], apply nat.le_of_succ_le_succ h.2.2.2.2 }, + haveI := h.nontrivial_fin, + rcases exists_ne j with ⟨j', hj'⟩, swap, intro hj, rcases smallest_on_boundary hj with ⟨x, ⟨hx, h2x⟩, h3x⟩, let p : fin (n+1) → ℝ := cons (c.b 0) (λ j₂, if j₂ = j then x else (cs i).b j₂.succ), @@ -406,7 +393,7 @@ begin have : (cs i).b ∈ (cs i').to_set, { simp only [to_set, forall_fin_succ, hi.1, bottom_mem_side h2i', true_and, mem_set_of_eq], intro j₂, by_cases hj₂ : j₂ = j, - { simpa [side_tail, p', hj', hj₂] using hi''.2 j }, + { simpa [side_tail, p', hj'.symm, hj₂] using hi''.2 j }, { simpa [hj₂] using hi'.2 j₂ } }, apply not_disjoint_iff.mpr ⟨(cs i).b, (cs i).b_mem_to_set, this⟩ (h.1 i i' i_i') }, have i_i'' : i ≠ i'', { intro h, induction h, simpa [hx'.2] using hi''.2 j' }, @@ -417,7 +404,7 @@ begin by_cases hj₂ : j₂ = j, { cases hj₂, refine ⟨x, _, _⟩, { convert hi'.2 j, simp [p] }, - apply h3x h2i'' i_i''.symm, convert hi''.2 j, simp [p', hj'] }, + apply h3x h2i'' i_i''.symm, convert hi''.2 j, simp [p', hj'.symm] }, by_cases h2j₂ : j₂ = j', { cases h2j₂, refine ⟨x', hx'.1, _⟩, convert hi''.2 j', simp }, refine ⟨(cs i).b j₂.succ, _, _⟩, @@ -441,11 +428,11 @@ end /-- The top of `mi` gives rise to a new valley, since the neighbouring cubes extend further upward than `mi`. -/ -def valley_mi : valley cs ((cs (mi h v)).shift_up) := +lemma valley_mi : valley cs ((cs (mi h v)).shift_up) := begin let i := mi h v, have hi : i ∈ bcubes cs c := mi_mem_bcubes, refine ⟨_, _, _⟩, - { intro p, apply shift_up_bottom_subset_bottoms h mi_xm_ne_one }, + { intro p, apply h.shift_up_bottom_subset_bottoms mi_xm_ne_one }, { rintros i' hi' ⟨p2, hp2, h2p2⟩, simp only [head_shift_up] at hi', classical, by_contra h2i', rw [tail_shift_up] at h2p2, simp only [not_subset, tail_shift_up] at h2i', rcases h2i' with ⟨p1, hp1, h2p1⟩, @@ -483,12 +470,12 @@ begin apply h.1, rintro rfl, apply (cs i).b_ne_xm, rw [←hi', ←hi''.1, hi.1], refl }, { intros i' hi' h2i', dsimp only [shift_up] at h2i', - replace h2i' := h.2.2.1 h2i'.symm, + replace h2i' := h.injective h2i'.symm, induction h2i', exact b_ne_xm (cs i) hi' } end -variables (h) +variables (h) [nontrivial ι] omit v /-- We get a sequence of cubes whose size is decreasing -/ @@ -505,27 +492,29 @@ begin apply w_lt_w h v (mi_mem_bcubes : mi h v ∈ _), end +lemma injective_sequence_of_cubes : injective (sequence_of_cubes h) := +@injective.of_comp _ _ _ (λ x : {i : ι // _}, (cs x.1).w) _ + (strict_anti_sequence_of_cubes h).injective + omit h + /-- The infinite sequence of cubes contradicts the finiteness of the family. -/ theorem not_correct : ¬correct cs := -begin - intro h, apply (lt_aleph_0_of_finite ι).not_le, - rw [aleph_0, lift_id], fapply mk_le_of_injective, exact λ n, (sequence_of_cubes h n).1, - intros n m hnm, apply (strict_anti_sequence_of_cubes h).injective, - dsimp only [decreasing_sequence], rw hnm -end +λ h, (finite.of_injective _ $ injective_sequence_of_cubes h).false /-- **Dissection of Cubes**: A cube cannot be cubed. -/ theorem cannot_cube_a_cube : - ∀{n : ℕ}, n ≥ 3 → -- In ℝ^n for n ≥ 3 - ∀{ι : Type} [fintype ι] {cs : ι → cube n}, -- given a finite collection of (hyper)cubes - 2 ≤ #ι → -- containing at least two elements - pairwise (disjoint on (cube.to_set ∘ cs)) → -- which is pairwise disjoint - (⋃(i : ι), (cs i).to_set) = unit_cube.to_set → -- whose union is the unit cube - injective (cube.w ∘ cs) → -- such that the widths of all cubes are different - false := -- then we can derive a contradiction + ∀ {n : ℕ}, n ≥ 3 → -- In ℝ^n for n ≥ 3 + ∀ {s : set (cube n)}, s.finite → -- given a finite collection of (hyper)cubes + s.nontrivial → -- containing at least two elements + s.pairwise_disjoint cube.to_set → -- which is pairwise disjoint + (⋃ c ∈ s, cube.to_set c) = unit_cube.to_set → -- whose union is the unit cube + inj_on cube.w s → -- such that the widths of all cubes are different + false := -- then we can derive a contradiction begin - intros n hn ι hι cs h1 h2 h3 h4, resetI, - rcases n, cases hn, - exact not_correct ⟨h2, h3, h4, h1, hn⟩ + intros n hn s hfin h2 hd hU hinj, + cases n, + { cases hn }, + exact @not_correct n s coe hfin.to_subtype ((nontrivial_coe _).2 h2) + ⟨hd.subtype _ _, (Union_subtype _ _).trans hU, hinj.injective, hn⟩ end From ea3dac316422054dbfc6bba967746c6216ac1666 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 12 Oct 2022 01:16:58 +0000 Subject: [PATCH 704/828] fix(number_theory/cyclotomic/basic): change variable names, fix doc (#16901) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The docstring of `is_cyclotomic` was using `a ∈ S` and `n` (`∈ S`) interchangeably. I've switched to `n` uniformly, because I think `a` risks being confused as a term of the ring `A`. --- src/number_theory/cyclotomic/basic.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index 4e562176e7713..654922e343927 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -72,11 +72,11 @@ variables [field K] [field L] [algebra K L] noncomputable theory /-- Given an `A`-algebra `B` and `S : set ℕ+`, we define `is_cyclotomic_extension S A B` requiring -that there is a `a`-th primitive root of unity in `B` for all `a ∈ S` and that `B` is generated +that there is a `n`-th primitive root of unity in `B` for all `n ∈ S` and that `B` is generated over `A` by the roots of `X ^ n - 1`. -/ @[mk_iff] class is_cyclotomic_extension : Prop := -(exists_prim_root {a : ℕ+} (ha : a ∈ S) : ∃ r : B, is_primitive_root r a) -(adjoin_roots : ∀ (x : B), x ∈ adjoin A { b : B | ∃ a : ℕ+, a ∈ S ∧ b ^ (a : ℕ) = 1 }) +(exists_prim_root {n : ℕ+} (ha : n ∈ S) : ∃ r : B, is_primitive_root r n) +(adjoin_roots : ∀ (x : B), x ∈ adjoin A { b : B | ∃ n : ℕ+, n ∈ S ∧ b ^ (n : ℕ) = 1 }) namespace is_cyclotomic_extension @@ -84,8 +84,8 @@ section basic /-- A reformulation of `is_cyclotomic_extension` that uses `⊤`. -/ lemma iff_adjoin_eq_top : is_cyclotomic_extension S A B ↔ - (∀ (a : ℕ+), a ∈ S → ∃ r : B, is_primitive_root r a) ∧ - (adjoin A { b : B | ∃ a : ℕ+, a ∈ S ∧ b ^ (a : ℕ) = 1 } = ⊤) := + (∀ (n : ℕ+), n ∈ S → ∃ r : B, is_primitive_root r n) ∧ + (adjoin A { b : B | ∃ n : ℕ+, n ∈ S ∧ b ^ (n : ℕ) = 1 } = ⊤) := ⟨λ h, ⟨λ _, h.exists_prim_root, algebra.eq_top_iff.2 h.adjoin_roots⟩, λ h, ⟨h.1, algebra.eq_top_iff.1 h.2⟩⟩ From 43a6243f33e0891a9488d34065aa07002ce3ed8e Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Wed, 12 Oct 2022 01:17:00 +0000 Subject: [PATCH 705/828] =?UTF-8?q?feat(data/finset/basic):=20add=20instan?= =?UTF-8?q?ce=20`is=5Fempty=20=E2=86=A5=E2=88=85`=20(#16902)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/data/finset/basic.lean | 2 ++ src/ring_theory/noetherian.lean | 3 +-- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index ff7ef1ae6aec9..36c919d376986 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -407,6 +407,8 @@ by rw [← coe_empty, coe_inj] @[simp] lemma is_empty_coe_sort {s : finset α} : is_empty ↥s ↔ s = ∅ := by simpa using @set.is_empty_coe_sort α s +instance : is_empty (∅ : finset α) := is_empty_coe_sort.2 rfl + /-- A `finset` for an empty type is empty. -/ lemma eq_empty_of_is_empty [is_empty α] (s : finset α) : s = ∅ := finset.eq_empty_of_forall_not_mem is_empty_elim diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index 907a45df05994..4bd06aac77755 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -558,8 +558,7 @@ begin exact is_noetherian_of_linear_equiv (linear_equiv.Pi_congr_left R M coe_e), }, intro s, induction s using finset.induction with a s has ih, - { split, intro s, convert submodule.fg_bot, apply eq_bot_iff.2, - intros x hx, refine (submodule.mem_bot R).2 _, ext i, cases i.2 }, + { exact ⟨λ s, by convert submodule.fg_bot⟩ }, refine @is_noetherian_of_linear_equiv _ _ _ _ _ _ _ _ _ (@is_noetherian_prod _ (M a) _ _ _ _ _ _ _ ih), fconstructor, From 8e9df64538594fd76d825c1add2f37b177ca2add Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 12 Oct 2022 01:17:04 +0000 Subject: [PATCH 706/828] chore(analysis/calculus/diff_ont_int_cont): rename to diff_cont_on_cl to match content (#16908) --- .../calculus/{diff_on_int_cont.lean => diff_cont_on_cl.lean} | 0 src/analysis/complex/cauchy_integral.lean | 2 +- 2 files changed, 1 insertion(+), 1 deletion(-) rename src/analysis/calculus/{diff_on_int_cont.lean => diff_cont_on_cl.lean} (100%) diff --git a/src/analysis/calculus/diff_on_int_cont.lean b/src/analysis/calculus/diff_cont_on_cl.lean similarity index 100% rename from src/analysis/calculus/diff_on_int_cont.lean rename to src/analysis/calculus/diff_cont_on_cl.lean diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index 02821474cb302..777b13bd4428d 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -9,7 +9,7 @@ import measure_theory.integral.circle_integral import analysis.calculus.dslope import analysis.analytic.basic import analysis.complex.re_im_topology -import analysis.calculus.diff_on_int_cont +import analysis.calculus.diff_cont_on_cl import data.real.cardinality /-! From 2e348d408dc0d73e99b4cced0c3beddf83a855ef Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 12 Oct 2022 04:37:47 +0000 Subject: [PATCH 707/828] feat(topology/sets/closeds): atoms in a t1_space are singletons (#16721) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is helpful to show that maximal ideals of `C(X, 𝕜)` with `X` compact Hausdorff and `𝕜` either `ℝ` or `ℂ` correspond to complements of singletons with respect to the Galois connection between `continuous_map.ideal_to_set` and `continuous_map.set_to_ideal`. --- src/topology/sets/closeds.lean | 48 ++++++++++++++++++++++++++++++++++ 1 file changed, 48 insertions(+) diff --git a/src/topology/sets/closeds.lean b/src/topology/sets/closeds.lean index 4ca4f1b2c2ff5..9cc9db98e3f94 100644 --- a/src/topology/sets/closeds.lean +++ b/src/topology/sets/closeds.lean @@ -108,6 +108,10 @@ instance : coframe (closeds α) := (set_like.coe_injective $ by simp only [coe_sup, coe_infi, coe_Inf, set.union_Inter₂]).le, ..closeds.complete_lattice } +/-- The term of `closeds α` corresponding to a singleton. -/ +@[simps] def singleton [t1_space α] (x : α) : closeds α := +⟨{x}, is_closed_singleton⟩ + end closeds /-- The complement of a closed set as an open set. -/ @@ -124,6 +128,50 @@ function.bijective_iff_has_inverse.mpr ⟨opens.compl, closeds.compl_compl, open lemma opens.compl_bijective : function.bijective (@opens.compl α _) := function.bijective_iff_has_inverse.mpr ⟨closeds.compl, opens.compl_compl, closeds.compl_compl⟩ +variables (α) + +/-- `closeds.compl` as an `order_iso` to the order dual of `opens α`. -/ +@[simps] def closeds.compl_order_iso : closeds α ≃o (opens α)ᵒᵈ := +{ to_fun := order_dual.to_dual ∘ closeds.compl, + inv_fun := opens.compl ∘ order_dual.of_dual, + left_inv := λ s, by simp [closeds.compl_compl], + right_inv := λ s, by simp [opens.compl_compl], + map_rel_iff' := λ s t, by simpa only [equiv.coe_fn_mk, function.comp_app, + order_dual.to_dual_le_to_dual] using compl_subset_compl } + +/-- `opens.compl` as an `order_iso` to the order dual of `closeds α`. -/ +@[simps] def opens.compl_order_iso : opens α ≃o (closeds α)ᵒᵈ := +{ to_fun := order_dual.to_dual ∘ opens.compl, + inv_fun := closeds.compl ∘ order_dual.of_dual, + left_inv := λ s, by simp [opens.compl_compl], + right_inv := λ s, by simp [closeds.compl_compl], + map_rel_iff' := λ s t, by simpa only [equiv.coe_fn_mk, function.comp_app, + order_dual.to_dual_le_to_dual] using compl_subset_compl } + +variables {α} + +/-- in a `t1_space`, atoms of `closeds α` are precisely the `closeds.singleton`s. -/ +lemma closeds.is_atom_iff [t1_space α] {s : closeds α} : is_atom s ↔ ∃ x, s = closeds.singleton x := +begin + have : is_atom (s : set α) ↔ is_atom s, + { refine closeds.gi.is_atom_iff' rfl (λ t ht, _) s, + obtain ⟨x, rfl⟩ := t.is_atom_iff.mp ht, + exact closure_singleton }, + simpa only [← this, (s : set α).is_atom_iff, set_like.ext_iff, set.ext_iff] +end + +/-- in a `t1_space`, coatoms of `opens α` are precisely complements of singletons: +`(closeds.singleton x).compl`. -/ +lemma opens.is_coatom_iff [t1_space α] {s : opens α} : + is_coatom s ↔ ∃ x, s = (closeds.singleton x).compl := +begin + rw [←s.compl_compl, ←is_atom_dual_iff_is_coatom], + change is_atom (closeds.compl_order_iso α s.compl) ↔ _, + rw [(closeds.compl_order_iso α).is_atom_iff, closeds.is_atom_iff], + congrm ∃ x, _, + exact closeds.compl_bijective.injective.eq_iff.symm, +end + /-! ### Clopen sets -/ /-- The type of clopen sets of a topological space. -/ From f1138dc36c9712ebe491a4002704e89cdc597b86 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 12 Oct 2022 04:37:48 +0000 Subject: [PATCH 708/828] =?UTF-8?q?feat(measure=5Ftheory/integral/periodic?= =?UTF-8?q?):=20relate=20integrals=20over=20`add=5Fcircle`=20with=20integr?= =?UTF-8?q?als=20upstairs=20in=20`=E2=84=9D`=20(#16836)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove that the covering map from `ℝ` to `add_circle` is measure-preserving, with respect to Lebesgue measure restricted to an interval (upstairs) and Haar measure (downstairs). As corollaries, lemmas relating the various integrals upstairs and downstairs. Co-authored-by: Alex Kontorovich <58564076+AlexKontorovich@users.noreply.github.com> --- src/group_theory/subgroup/basic.lean | 6 + .../constructions/borel_space.lean | 10 ++ src/measure_theory/integral/periodic.lean | 144 ++++++++++++++++-- src/measure_theory/measure/haar_quotient.lean | 19 ++- src/topology/instances/add_circle.lean | 1 - 5 files changed, 169 insertions(+), 11 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index ad42a5666e90a..70f36e771b93f 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2683,6 +2683,9 @@ lemma exists_mem_zpowers {x : G} {p : G → Prop} : (∃ g ∈ zpowers x, p g) ↔ ∃ m : ℤ, p (x ^ m) := set.exists_range_iff +instance (a : G) : countable (zpowers a) := +((zpowers_hom G a).range_restrict_surjective.comp multiplicative.of_add.surjective).countable + end subgroup namespace add_subgroup @@ -2705,6 +2708,9 @@ attribute [to_additive add_subgroup.forall_mem_zmultiples] subgroup.forall_mem_z attribute [to_additive add_subgroup.exists_zmultiples] subgroup.exists_zpowers attribute [to_additive add_subgroup.exists_mem_zmultiples] subgroup.exists_mem_zpowers +instance (a : A) : countable (zmultiples a) := +(zmultiples_hom A a).range_restrict_surjective.countable + section ring variables {R : Type*} [ring R] (r : R) (k : ℤ) diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 10052329d7aa0..58fb2f7ca3b3d 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -11,6 +11,7 @@ import measure_theory.lattice import measure_theory.measure.open_pos import topology.algebra.order.liminf_limsup import topology.continuous_function.basic +import topology.instances.add_circle import topology.instances.ereal import topology.G_delta import topology.order.lattice @@ -1343,6 +1344,15 @@ instance ereal.borel_space : borel_space ereal := ⟨rfl⟩ instance complex.measurable_space : measurable_space ℂ := borel ℂ instance complex.borel_space : borel_space ℂ := ⟨rfl⟩ +instance add_circle.measurable_space {a : ℝ} : measurable_space (add_circle a) := +borel (add_circle a) + +instance add_circle.borel_space {a : ℝ} : borel_space (add_circle a) := ⟨rfl⟩ + +@[measurability] protected lemma add_circle.measurable_mk' {a : ℝ} : + measurable (coe : ℝ → add_circle a) := +continuous.measurable $ add_circle.continuous_mk' a + /-- One can cut out `ℝ≥0∞` into the sets `{0}`, `Ico (t^n) (t^(n+1))` for `n : ℤ` and `{∞}`. This gives a way to compute the measure of a set in terms of sets on which a given function `f` does not fluctuate by more than `t`. -/ diff --git a/src/measure_theory/integral/periodic.lean b/src/measure_theory/integral/periodic.lean index 9e9acb3aaca47..bc019f2c1f0b1 100644 --- a/src/measure_theory/integral/periodic.lean +++ b/src/measure_theory/integral/periodic.lean @@ -1,21 +1,33 @@ /- Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Yury Kudryashov +Authors: Yury Kudryashov, Alex Kontorovich, Heather Macbeth -/ -import measure_theory.group.fundamental_domain +import measure_theory.measure.haar_quotient import measure_theory.integral.interval_integral import topology.algebra.order.floor /-! # Integrals of periodic functions -In this file we prove that `∫ x in t..t + T, f x = ∫ x in s..s + T, f x` for any (not necessarily -measurable) function periodic function with period `T`. +In this file we prove that the half-open interval `Ioc t (t + T)` in `ℝ` is a fundamental domain of +the action of the subgroup `ℤ ∙ T` on `ℝ`. + +A consequence is `add_circle.measure_preserving_mk`: the covering map from `ℝ` to the "additive +circle" `ℝ ⧸ (ℤ ∙ T)` is measure-preserving, with respect to the restriction of Lebesgue measure to +`Ioc t (t + T)` (upstairs) and with respect to Haar measure (downstairs). + +Another consequence (`function.periodic.interval_integral_add_eq` and related declarations) is that +`∫ x in t..t + T, f x = ∫ x in s..s + T, f x` for any (not necessarily measurable) function with +period `T`. -/ -open set function measure_theory measure_theory.measure topological_space -open_locale measure_theory +open set function measure_theory measure_theory.measure topological_space add_subgroup + interval_integral + +open_locale measure_theory nnreal ennreal + +local attribute [-instance] quotient_add_group.measurable_space quotient.measurable_space lemma is_add_fundamental_domain_Ioc {T : ℝ} (hT : 0 < T) (t : ℝ) (μ : measure ℝ . volume_tac) : is_add_fundamental_domain (add_subgroup.zmultiples T) (Ioc t (t + T)) μ := @@ -27,21 +39,135 @@ begin simpa only [add_comm x] using exists_unique_add_zsmul_mem_Ioc hT x t end +lemma is_add_fundamental_domain_Ioc' {T : ℝ} (hT : 0 < T) (t : ℝ) (μ : measure ℝ . volume_tac) : + is_add_fundamental_domain (add_subgroup.zmultiples T).opposite (Ioc t (t + T)) μ := +begin + refine is_add_fundamental_domain.mk' measurable_set_Ioc.null_measurable_set (λ x, _), + have : bijective (cod_restrict (λ n : ℤ, n • T) (add_subgroup.zmultiples T) _), + from (equiv.of_injective (λ n : ℤ, n • T) (zsmul_strict_mono_left hT).injective).bijective, + refine this.exists_unique_iff.2 _, + simpa using exists_unique_add_zsmul_mem_Ioc hT x t, +end + +namespace add_circle +variables (T : ℝ) [hT : fact (0 < T)] +include hT + +/-- Equip the "additive circle" `ℝ ⧸ (ℤ ∙ T)` with, as a standard measure, the Haar measure of total +mass `T` -/ +noncomputable instance measure_space : measure_space (add_circle T) := +{ volume := (ennreal.of_real T) • add_haar_measure ⊤, + .. add_circle.measurable_space } + +@[simp] protected lemma measure_univ : volume (set.univ : set (add_circle T)) = ennreal.of_real T := +begin + dsimp [(volume)], + rw ← positive_compacts.coe_top, + simp [add_haar_measure_self, -positive_compacts.coe_top], +end + +instance is_finite_measure : is_finite_measure (volume : measure (add_circle T)) := +{ measure_univ_lt_top := by simp } + +/-- The covering map from `ℝ` to the "additive circle" `ℝ ⧸ (ℤ ∙ T)` is measure-preserving, +considered with respect to the standard measure (defined to be the Haar measure of total mass `T`) +on the additive circle, and with respect to the restriction of Lebsegue measure on `ℝ` to an +interval (t, t + T]. -/ +protected lemma measure_preserving_mk (t : ℝ) : + measure_preserving (coe : ℝ → add_circle T) (volume.restrict (Ioc t (t + T))) := +measure_preserving_quotient_add_group.mk' + (is_add_fundamental_domain_Ioc' hT.out t) + (⊤ : positive_compacts (add_circle T)) + (by simp) + T.to_nnreal + (by simp [← ennreal.of_real_coe_nnreal, real.coe_to_nnreal T hT.out.le]) + +/-- The integral of a measurable function over `add_circle T` is equal to the integral over an +interval (t, t + T] in `ℝ` of its lift to `ℝ`. -/ +protected lemma lintegral_preimage (t : ℝ) {f : add_circle T → ℝ≥0∞} (hf : measurable f) : + ∫⁻ a in Ioc t (t + T), f a = ∫⁻ b : add_circle T, f b := +by rw [← lintegral_map hf add_circle.measurable_mk', + (add_circle.measure_preserving_mk T t).map_eq] + +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + +/-- The integral of an almost-everywhere strongly measurable function over `add_circle T` is equal +to the integral over an interval (t, t + T] in `ℝ` of its lift to `ℝ`. -/ +protected lemma integral_preimage (t : ℝ) {f : add_circle T → E} + (hf : ae_strongly_measurable f volume) : + ∫ a in Ioc t (t + T), f a = ∫ b : add_circle T, f b := +begin + rw ← (add_circle.measure_preserving_mk T t).map_eq at ⊢ hf, + rw integral_map add_circle.measurable_mk'.ae_measurable hf, +end + +/-- The integral of an almost-everywhere strongly measurable function over `add_circle T` is equal +to the integral over an interval (t, t + T] in `ℝ` of its lift to `ℝ`. -/ +protected lemma interval_integral_preimage (t : ℝ) {f : add_circle T → E} + (hf : ae_strongly_measurable f volume) : + ∫ a in t..(t + T), f a = ∫ b : add_circle T, f b := +begin + rw [integral_of_le, add_circle.integral_preimage T t hf], + linarith [hT.out], +end + +end add_circle + +namespace unit_add_circle +private lemma fact_zero_lt_one : fact ((0:ℝ) < 1) := ⟨zero_lt_one⟩ +local attribute [instance] fact_zero_lt_one + +noncomputable instance measure_space : measure_space unit_add_circle := add_circle.measure_space 1 + +@[simp] protected lemma measure_univ : volume (set.univ : set unit_add_circle) = 1 := by simp + +instance is_finite_measure : is_finite_measure (volume : measure unit_add_circle) := +add_circle.is_finite_measure 1 + +/-- The covering map from `ℝ` to the "unit additive circle" `ℝ ⧸ ℤ` is measure-preserving, +considered with respect to the standard measure (defined to be the Haar measure of total mass 1) +on the additive circle, and with respect to the restriction of Lebsegue measure on `ℝ` to an +interval (t, t + 1]. -/ +protected lemma measure_preserving_mk (t : ℝ) : + measure_preserving (coe : ℝ → unit_add_circle) (volume.restrict (Ioc t (t + 1))) := +add_circle.measure_preserving_mk 1 t + +/-- The integral of a measurable function over `unit_add_circle` is equal to the integral over an +interval (t, t + 1] in `ℝ` of its lift to `ℝ`. -/ +protected lemma lintegral_preimage (t : ℝ) {f : unit_add_circle → ℝ≥0∞} (hf : measurable f) : + ∫⁻ a in Ioc t (t + 1), f a = ∫⁻ b : unit_add_circle, f b := +add_circle.lintegral_preimage 1 t hf + +variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] + +/-- The integral of an almost-everywhere strongly measurable function over `unit_add_circle` is +equal to the integral over an interval (t, t + 1] in `ℝ` of its lift to `ℝ`. -/ +protected lemma integral_preimage (t : ℝ) {f : unit_add_circle → E} + (hf : ae_strongly_measurable f volume) : + ∫ a in Ioc t (t + 1), f a = ∫ b : unit_add_circle, f b := +add_circle.integral_preimage 1 t hf + +/-- The integral of an almost-everywhere strongly measurable function over `unit_add_circle` is +equal to the integral over an interval (t, t + 1] in `ℝ` of its lift to `ℝ`. -/ +protected lemma interval_integral_preimage (t : ℝ) {f : unit_add_circle → E} + (hf : ae_strongly_measurable f volume) : + ∫ a in t..(t + 1), f a = ∫ b : unit_add_circle, f b := +add_circle.interval_integral_preimage 1 t hf + +end unit_add_circle + variables {E : Type*} [normed_add_comm_group E] [normed_space ℝ E] [complete_space E] namespace function namespace periodic -open interval_integral - variables {f : ℝ → E} {T : ℝ} /-- An auxiliary lemma for a more general `function.periodic.interval_integral_add_eq`. -/ lemma interval_integral_add_eq_of_pos (hf : periodic f T) (hT : 0 < T) (t s : ℝ) : ∫ x in t..t + T, f x = ∫ x in s..s + T, f x := begin - haveI : encodable (add_subgroup.zmultiples T) := (countable_range _).to_encodable, simp only [integral_of_le, hT.le, le_add_iff_nonneg_right], haveI : vadd_invariant_measure (add_subgroup.zmultiples T) ℝ volume := ⟨λ c s hs, measure_preimage_add _ _ _⟩, diff --git a/src/measure_theory/measure/haar_quotient.lean b/src/measure_theory/measure/haar_quotient.lean index 84cbc237b70ba..688188c0760e0 100644 --- a/src/measure_theory/measure/haar_quotient.lean +++ b/src/measure_theory/measure/haar_quotient.lean @@ -32,7 +32,7 @@ Note that a group `G` with Haar measure that is both left and right invariant is -/ open set measure_theory topological_space measure_theory.measure -open_locale pointwise +open_locale pointwise nnreal variables {G : Type*} [group G] [measurable_space G] [topological_space G] [topological_group G] [borel_space G] @@ -158,3 +158,20 @@ begin measure.map_apply meas_π, measure.restrict_apply₀' 𝓕meas, inter_comm], exact K.compact.measurable_set, end + +/-- Given a normal subgroup `Γ` of a topological group `G` with Haar measure `μ`, which is also + right-invariant, and a finite volume fundamental domain `𝓕`, the quotient map to `G ⧸ Γ` is + measure-preserving between appropriate multiples of Haar measure on `G` and `G ⧸ Γ`. -/ +@[to_additive measure_preserving_quotient_add_group.mk' "Given a normal subgroup `Γ` of an additive + topological group `G` with Haar measure `μ`, which is also right-invariant, and a finite volume + fundamental domain `𝓕`, the quotient map to `G ⧸ Γ` is measure-preserving between appropriate + multiples of Haar measure on `G` and `G ⧸ Γ`."] +lemma measure_preserving_quotient_group.mk' [subgroup.normal Γ] + [measure_theory.measure.is_haar_measure μ] [μ.is_mul_right_invariant] + (h𝓕_finite : μ 𝓕 < ⊤) (c : ℝ≥0) (h : μ (𝓕 ∩ (quotient_group.mk' Γ) ⁻¹' K) = c) : + measure_preserving + (quotient_group.mk' Γ) + (μ.restrict 𝓕) + (c • (measure_theory.measure.haar_measure K)) := +{ measurable := continuous_quotient_mk.measurable, + map_eq := by rw [h𝓕.map_restrict_quotient K h𝓕_finite, h]; refl } diff --git a/src/topology/instances/add_circle.lean b/src/topology/instances/add_circle.lean index dbc36f21a8f70..92dfd4a82f499 100644 --- a/src/topology/instances/add_circle.lean +++ b/src/topology/instances/add_circle.lean @@ -30,7 +30,6 @@ the rational circle `add_circle (1 : ℚ)`, and so we set things up more general ## TODO * Link with periodicity - * Measure space structure * Lie group structure * Exponential equivalence to `circle` From 2c8826210f746dff7a6ec5fed96ea5dfd4ff508f Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 12 Oct 2022 07:21:05 +0000 Subject: [PATCH 709/828] feat(category_theory/sites): Surjective presheaf morphisms. (#15283) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/category_theory/sites/subsheaf.lean | 28 +++- src/category_theory/sites/surjective.lean | 181 ++++++++++++++++++++++ 2 files changed, 207 insertions(+), 2 deletions(-) create mode 100644 src/category_theory/sites/surjective.lean diff --git a/src/category_theory/sites/subsheaf.lean b/src/category_theory/sites/subsheaf.lean index c3081cd5705f4..b83afcef8c5dd 100644 --- a/src/category_theory/sites/subsheaf.lean +++ b/src/category_theory/sites/subsheaf.lean @@ -89,6 +89,21 @@ lemma subpresheaf.hom_of_le_ι {G G' : subpresheaf F} (h : G ≤ G') : subpresheaf.hom_of_le h ≫ G'.ι = G.ι := by { ext, refl } +instance : is_iso (subpresheaf.ι (⊤ : subpresheaf F)) := +begin + apply_with nat_iso.is_iso_of_is_iso_app { instances := ff }, + { intro X, rw is_iso_iff_bijective, + exact ⟨subtype.coe_injective, λ x, ⟨⟨x, _root_.trivial⟩, rfl⟩⟩ } +end + +lemma subpresheaf.eq_top_iff_is_iso : G = ⊤ ↔ is_iso G.ι := +begin + split, + { rintro rfl, apply_instance }, + { introI H, ext U x, apply (iff_true _).mpr, rw ← is_iso.inv_hom_id_apply (G.ι.app U) x, + exact ((inv (G.ι.app U)) x).2 } +end + /-- If the image of a morphism falls in a subpresheaf, then the morphism factors through it. -/ @[simps] def subpresheaf.lift (f : F' ⟶ F) (hf : ∀ U x, f.app U x ∈ G.obj U) : F' ⟶ G.to_presheaf := @@ -308,6 +323,15 @@ by { ext, simp } def to_image_presheaf (f : F' ⟶ F) : F' ⟶ (image_presheaf f).to_presheaf := (image_presheaf f).lift f (λ U x, set.mem_range_self _) +variables (J) + +/-- A morphism factors through the sheafification of the image presheaf. -/ +@[simps] +def to_image_presheaf_sheafify (f : F' ⟶ F) : F' ⟶ ((image_presheaf f).sheafify J).to_presheaf := + to_image_presheaf f ≫ subpresheaf.hom_of_le ((image_presheaf f).le_sheafify J) + +variables {J} + @[simp, reassoc] lemma to_image_presheaf_ι (f : F' ⟶ F) : to_image_presheaf f ≫ (image_presheaf f).ι = f := (image_presheaf f).lift_ι _ _ @@ -341,7 +365,7 @@ def image_sheaf {F F' : Sheaf J (Type w)} (f : F ⟶ F') : Sheaf J (Type w) := /-- A morphism factors through the image sheaf. -/ @[simps] def to_image_sheaf {F F' : Sheaf J (Type w)} (f : F ⟶ F') : F ⟶ image_sheaf f := -⟨to_image_presheaf f.1 ≫ subpresheaf.hom_of_le ((image_presheaf f.1).le_sheafify J)⟩ +⟨to_image_presheaf_sheafify J f.1⟩ /-- The inclusion of the image sheaf to the target. -/ @[simps] @@ -351,7 +375,7 @@ def image_sheaf_ι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : image_sheaf f ⟶ @[simp, reassoc] lemma to_image_sheaf_ι {F F' : Sheaf J (Type w)} (f : F ⟶ F') : to_image_sheaf f ≫ image_sheaf_ι f = f := -by { ext1, simp } +by { ext1, simp [to_image_presheaf_sheafify] } instance {F F' : Sheaf J (Type w)} (f : F ⟶ F') : mono (image_sheaf_ι f) := (Sheaf_to_presheaf J _).mono_of_mono_map (by { dsimp, apply_instance }) diff --git a/src/category_theory/sites/surjective.lean b/src/category_theory/sites/surjective.lean new file mode 100644 index 0000000000000..7aa508c04c973 --- /dev/null +++ b/src/category_theory/sites/surjective.lean @@ -0,0 +1,181 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import category_theory.sites.subsheaf +import category_theory.adjunction.evaluation + +/-! + +# Locally surjective morphisms + +## Main definitions + +- `is_locally_surjective` : A morphism of presheaves valued in a concrete category is locally + surjective with respect to a grothendieck topology if every section in the target is locally + in the set-theoretic image, i.e. the image sheaf coincides with the target. + +## Main results + +- `to_sheafify_is_locally_surjective` : `to_sheafify` is locally surjective. + +-/ + +universes v u w v' u' w' + +open opposite category_theory category_theory.grothendieck_topology + +namespace category_theory + +variables {C : Type u} [category.{v} C] (J : grothendieck_topology C) + +local attribute [instance] concrete_category.has_coe_to_sort concrete_category.has_coe_to_fun + +variables {A : Type u'} [category.{v'} A] [concrete_category.{w'} A] + +/-- Given `f : F ⟶ G`, a morphism between presieves, and `s : G.obj (op U)`, this is the sieve +of `U` consisting of the `i : V ⟶ U` such that `s` restricted along `i` is in the image of `f`. -/ +@[simps (lemmas_only)] +def image_sieve {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : G.obj (op U)) : sieve U := +{ arrows := λ V i, ∃ t : F.obj (op V), f.app _ t = G.map i.op s, + downward_closed' := begin + rintros V W i ⟨t, ht⟩ j, + refine ⟨F.map j.op t, _⟩, + rw [op_comp, G.map_comp, comp_apply, ← ht, elementwise_of f.naturality], + end } + +lemma image_sieve_eq_sieve_of_section {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : G.obj (op U)) : + image_sieve f s = (image_presheaf (whisker_right f (forget A))).sieve_of_section s := rfl + +lemma image_sieve_whisker_forget {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : G.obj (op U)) : + image_sieve (whisker_right f (forget A)) s = image_sieve f s := rfl + +lemma image_sieve_app {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) {U : C} (s : F.obj (op U)) : + image_sieve f (f.app _ s) = ⊤ := +begin + ext V i, + simp only [sieve.top_apply, iff_true, image_sieve_apply], + have := elementwise_of (f.naturality i.op), + exact ⟨F.map i.op s, this s⟩, +end + +/-- A morphism of presheaves `f : F ⟶ G` is locally surjective with respect to a grothendieck +topology if every section of `G` is locally in the image of `f`. -/ +def is_locally_surjective {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) : Prop := +∀ (U : C) (s : G.obj (op U)), image_sieve f s ∈ J U + +lemma is_locally_surjective_iff_image_presheaf_sheafify_eq_top {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) : + is_locally_surjective J f ↔ (image_presheaf (whisker_right f (forget A))).sheafify J = ⊤ := +begin + simp only [subpresheaf.ext_iff, function.funext_iff, set.ext_iff, top_subpresheaf_obj, + set.top_eq_univ, set.mem_univ, iff_true], + exact ⟨λ H U, H (unop U), λ H U, H (op U)⟩ +end + +lemma is_locally_surjective_iff_image_presheaf_sheafify_eq_top' + {F G : Cᵒᵖ ⥤ (Type w)} (f : F ⟶ G) : + is_locally_surjective J f ↔ (image_presheaf f).sheafify J = ⊤ := +begin + simp only [subpresheaf.ext_iff, function.funext_iff, set.ext_iff, top_subpresheaf_obj, + set.top_eq_univ, set.mem_univ, iff_true], + exact ⟨λ H U, H (unop U), λ H U, H (op U)⟩ +end + +lemma is_locally_surjective_iff_is_iso + {F G : Sheaf J (Type w)} (f : F ⟶ G) : + is_locally_surjective J f.1 ↔ is_iso (image_sheaf_ι f) := +begin + rw [image_sheaf_ι, is_locally_surjective_iff_image_presheaf_sheafify_eq_top', + subpresheaf.eq_top_iff_is_iso], + exact ⟨λ h, @@is_iso_of_reflects_iso _ _ (image_sheaf_ι f) (Sheaf_to_presheaf J _) h _, + λ h, @@functor.map_is_iso _ _ (Sheaf_to_presheaf J _) _ h⟩, +end + +lemma is_locally_surjective_iff_whisker_forget {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) : + is_locally_surjective J f ↔ is_locally_surjective J (whisker_right f (forget A)) := +begin + simpa only [is_locally_surjective_iff_image_presheaf_sheafify_eq_top] +end + +lemma is_locally_surjective_of_surjective {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) + (H : ∀ U, function.surjective (f.app U)) : is_locally_surjective J f := +begin + intros U s, + obtain ⟨t, rfl⟩ := H _ s, + rw image_sieve_app, + exact J.top_mem _ +end + +lemma is_locally_surjective_of_iso {F G : Cᵒᵖ ⥤ A} (f : F ⟶ G) [is_iso f] : + is_locally_surjective J f := +begin + apply is_locally_surjective_of_surjective, + intro U, + apply function.bijective.surjective, + rw ← is_iso_iff_bijective, + apply_instance +end + +lemma is_locally_surjective.comp {F₁ F₂ F₃ : Cᵒᵖ ⥤ A} {f₁ : F₁ ⟶ F₂} {f₂ : F₂ ⟶ F₃} + (h₁ : is_locally_surjective J f₁) (h₂ : is_locally_surjective J f₂) : + is_locally_surjective J (f₁ ≫ f₂) := +begin + intros U s, + have : sieve.bind (image_sieve f₂ s) (λ _ _ h, image_sieve f₁ h.some) ≤ image_sieve (f₁ ≫ f₂) s, + { rintros V i ⟨W, i, j, H, ⟨t', ht'⟩, rfl⟩, + refine ⟨t', _⟩, + rw [op_comp, F₃.map_comp, nat_trans.comp_app, comp_apply, comp_apply, ht', + elementwise_of f₂.naturality, H.some_spec] }, + apply J.superset_covering this, + apply J.bind_covering, + { apply h₂ }, + { intros, apply h₁ } +end + +section + +variables (F : Cᵒᵖ ⥤ Type (max u v)) + +/-- The image of `F` in `J.sheafify F` is isomorphic to the sheafification. -/ +noncomputable +def sheafification_iso_image_presheaf : + J.sheafify F ≅ ((image_presheaf (J.to_sheafify F)).sheafify J).to_presheaf := +{ hom := J.sheafify_lift (to_image_presheaf_sheafify J _) + ((is_sheaf_iff_is_sheaf_of_type J _).mpr $ subpresheaf.sheafify_is_sheaf _ $ + (is_sheaf_iff_is_sheaf_of_type J _).mp $ sheafify_is_sheaf J _), + inv := subpresheaf.ι _, + hom_inv_id' := J.sheafify_hom_ext _ _ (J.sheafify_is_sheaf _) + (by simp [to_image_presheaf_sheafify]), + inv_hom_id' := begin + rw [← cancel_mono (subpresheaf.ι _), category.id_comp, category.assoc], + refine eq.trans _ (category.comp_id _), + congr' 1, + exact J.sheafify_hom_ext _ _ (J.sheafify_is_sheaf _) (by simp [to_image_presheaf_sheafify]), + apply_instance + end } + +-- We need to sheafify +variables {B : Type w} [category.{max u v} B] + [concrete_category.{max u v} B] + [∀ (X : C), limits.has_colimits_of_shape (J.cover X)ᵒᵖ B] + [∀ (P : Cᵒᵖ ⥤ B) (X : C) (S : J.cover X), limits.has_multiequalizer (S.index P)] + [Π (X : C) (W : J.cover X) (P : Cᵒᵖ ⥤ B), + limits.preserves_limit (W.index P).multicospan (forget B)] + [Π (X : C), limits.preserves_colimits_of_shape (J.cover X)ᵒᵖ (forget B)] + [∀ (α β : Type (max u v)) (fst snd : β → α), + limits.has_limits_of_shape (limits.walking_multicospan fst snd) B] + +lemma to_sheafify_is_locally_surjective (F : Cᵒᵖ ⥤ B) : + is_locally_surjective J (J.to_sheafify F) := +begin + rw [is_locally_surjective_iff_whisker_forget, ← to_sheafify_comp_sheafify_comp_iso_inv], + apply is_locally_surjective.comp, + { rw [is_locally_surjective_iff_image_presheaf_sheafify_eq_top, subpresheaf.eq_top_iff_is_iso], + exact is_iso.of_iso_inv (sheafification_iso_image_presheaf J (F ⋙ forget B)) }, + { exact is_locally_surjective_of_iso _ _ } +end + +end + +end category_theory From cbaea5ddfa6db4b4a7de7995db9f04bf6c7e6731 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 12 Oct 2022 07:21:06 +0000 Subject: [PATCH 710/828] feat(polynomial/degree/trailing_degree): redefine `trailing_degree` as `p.support.min` (#16337) --- .../polynomial/degree/trailing_degree.lean | 22 +++++++++---------- 1 file changed, 10 insertions(+), 12 deletions(-) diff --git a/src/data/polynomial/degree/trailing_degree.lean b/src/data/polynomial/degree/trailing_degree.lean index bef7f53b802b2..439d3c9a8220a 100644 --- a/src/data/polynomial/degree/trailing_degree.lean +++ b/src/data/polynomial/degree/trailing_degree.lean @@ -36,7 +36,7 @@ variables [semiring R] {p q r : R[X]} `trailing_degree p = some n` when `p ≠ 0` and `n` is the smallest power of `X` that appears in `p`, otherwise `trailing_degree 0 = ⊤`. -/ -def trailing_degree (p : R[X]) : ℕ∞ := p.support.inf some +def trailing_degree (p : R[X]) : ℕ∞ := p.support.min lemma trailing_degree_lt_wf : well_founded (λp q : R[X], trailing_degree p < trailing_degree q) := @@ -67,14 +67,13 @@ by unfold trailing_monic; apply_instance @[simp] lemma nat_trailing_degree_zero : nat_trailing_degree (0 : R[X]) = 0 := rfl lemma trailing_degree_eq_top : trailing_degree p = ⊤ ↔ p = 0 := -⟨λ h, support_eq_empty.1 (finset.min_eq_top.1 h), -λ h, by simp [h]⟩ +⟨λ h, support_eq_empty.1 (finset.min_eq_top.1 h), λ h, by simp [h]⟩ lemma trailing_degree_eq_nat_trailing_degree (hp : p ≠ 0) : trailing_degree p = (nat_trailing_degree p : ℕ∞) := let ⟨n, hn⟩ := not_forall.1 (mt option.eq_none_iff_forall_not_mem.2 (mt trailing_degree_eq_top.1 hp)) in -have hn : trailing_degree p = some n := not_not.1 hn, +have hn : trailing_degree p = n := not_not.1 hn, by rw [nat_trailing_degree, hn]; refl lemma trailing_degree_eq_iff_nat_trailing_degree_eq {p : R[X]} {n : ℕ} (hp : p ≠ 0) : @@ -111,8 +110,8 @@ nat_trailing_degree p = nat_trailing_degree q := by unfold nat_trailing_degree; rw h lemma le_trailing_degree_of_ne_zero (h : coeff p n ≠ 0) : trailing_degree p ≤ n := -show @has_le.le (ℕ∞) _ (p.support.inf some : ℕ∞) (some n : ℕ∞), -from finset.inf_le (mem_support_iff.2 h) +show @has_le.le ℕ∞ _ p.support.min n, +from min_le (mem_support_iff.2 h) lemma nat_trailing_degree_le_of_ne_zero (h : coeff p n ≠ 0) : nat_trailing_degree p ≤ n := begin @@ -150,7 +149,7 @@ begin end @[simp] lemma trailing_degree_monomial (ha : a ≠ 0) : trailing_degree (monomial n a) = n := -by rw [trailing_degree, support_monomial n ha, inf_singleton, with_top.some_eq_coe] +by rw [trailing_degree, support_monomial n ha, min_singleton] lemma nat_trailing_degree_monomial (ha : a ≠ 0) : nat_trailing_degree (monomial n a) = n := by rw [nat_trailing_degree, trailing_degree_monomial ha]; refl @@ -274,13 +273,12 @@ end lemma le_trailing_degree_mul : p.trailing_degree + q.trailing_degree ≤ (p * q).trailing_degree := begin - refine le_inf (λ n hn, _), + refine le_min (λ n hn, _), rw [mem_support_iff, coeff_mul] at hn, obtain ⟨⟨i, j⟩, hij, hpq⟩ := exists_ne_zero_of_sum_ne_zero hn, - refine (add_le_add (inf_le (mem_support_iff.mpr (left_ne_zero_of_mul hpq))) - (inf_le (mem_support_iff.mpr (right_ne_zero_of_mul hpq)))).trans (le_of_eq _), - rwa [with_top.some_eq_coe, with_top.some_eq_coe, with_top.some_eq_coe, - ← with_top.coe_add, with_top.coe_eq_coe, ←nat.mem_antidiagonal], + refine (add_le_add (min_le (mem_support_iff.mpr (left_ne_zero_of_mul hpq))) + (min_le (mem_support_iff.mpr (right_ne_zero_of_mul hpq)))).trans (le_of_eq _), + rwa [← with_top.coe_add, with_top.coe_eq_coe, ←nat.mem_antidiagonal], end lemma le_nat_trailing_degree_mul (h : p * q ≠ 0) : From 59de3c16f446563afe32307dc2c32df9f1ba0775 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 12 Oct 2022 07:21:07 +0000 Subject: [PATCH 711/828] refactor(measure_theory/group/prod): rename the variables (#16707) * We don't usually use `E` and `F` as set variables. Especially if we want to talk about the Bochner integral in this file, it is convenient to call the codomain `E`. --- src/measure_theory/group/prod.lean | 134 ++++++++++++++--------------- 1 file changed, 67 insertions(+), 67 deletions(-) diff --git a/src/measure_theory/group/prod.lean b/src/measure_theory/group/prod.lean index 960ee25a60a95..351eb20e19929 100644 --- a/src/measure_theory/group/prod.lean +++ b/src/measure_theory/group/prod.lean @@ -14,17 +14,17 @@ and properties of iterated integrals in measurable groups. These lemmas show the uniqueness of left invariant measures on measurable groups, up to scaling. In this file we follow the proof and refer to the book *Measure Theory* by Paul Halmos. -The idea of the proof is to use the translation invariance of measures to prove `μ(F) = c * μ(E)` -for two sets `E` and `F`, where `c` is a constant that does not depend on `μ`. Let `e` and `f` be -the characteristic functions of `E` and `F`. +The idea of the proof is to use the translation invariance of measures to prove `μ(t) = c * μ(s)` +for two sets `s` and `t`, where `c` is a constant that does not depend on `μ`. Let `e` and `f` be +the characteristic functions of `s` and `t`. Assume that `μ` and `ν` are left-invariant measures. Then the map `(x, y) ↦ (y * x, x⁻¹)` preserves the measure `μ.prod ν`, which means that ``` ∫ x, ∫ y, h x y ∂ν ∂μ = ∫ x, ∫ y, h (y * x) x⁻¹ ∂ν ∂μ ``` -If we apply this to `h x y := e x * f y⁻¹ / ν ((λ h, h * y⁻¹) ⁻¹' E)`, we can rewrite the RHS to -`μ(F)`, and the LHS to `c * μ(E)`, where `c = c(ν)` does not depend on `μ`. -Applying this to `μ` and to `ν` gives `μ (F) / μ (E) = ν (F) / ν (E)`, which is the uniqueness up to +If we apply this to `h x y := e x * f y⁻¹ / ν ((λ h, h * y⁻¹) ⁻¹' s)`, we can rewrite the RHS to +`μ(t)`, and the LHS to `c * μ(s)`, where `c = c(ν)` does not depend on `μ`. +Applying this to `μ` and to `ν` gives `μ (t) / μ (s) = ν (t) / ν (s)`, which is the uniqueness up to scalar multiplication. The proof in [Halmos] seems to contain an omission in §60 Th. A, see @@ -38,7 +38,7 @@ open_locale classical ennreal pointwise measure_theory variables (G : Type*) [measurable_space G] variables [group G] [has_measurable_mul₂ G] -variables (μ ν : measure G) [sigma_finite ν] [sigma_finite μ] {E : set G} +variables (μ ν : measure G) [sigma_finite ν] [sigma_finite μ] {s : set G} /-- The map `(x, y) ↦ (x, xy)` as a `measurable_equiv`. This is a shear mapping. -/ @[to_additive "The map `(x, y) ↦ (x, x + y)` as a `measurable_equiv`. @@ -74,14 +74,14 @@ lemma measure_preserving_prod_mul_swap [is_mul_left_invariant μ] : (measure_preserving_prod_mul ν μ).comp measure_preserving_swap @[to_additive] -lemma measurable_measure_mul_right (hE : measurable_set E) : - measurable (λ x, μ ((λ y, y * x) ⁻¹' E)) := +lemma measurable_measure_mul_right (hs : measurable_set s) : + measurable (λ x, μ ((λ y, y * x) ⁻¹' s)) := begin suffices : measurable (λ y, - μ ((λ x, (x, y)) ⁻¹' ((λ z : G × G, ((1 : G), z.1 * z.2)) ⁻¹' (univ ×ˢ E)))), + μ ((λ x, (x, y)) ⁻¹' ((λ z : G × G, ((1 : G), z.1 * z.2)) ⁻¹' (univ ×ˢ s)))), { convert this, ext1 x, congr' 1 with y : 1, simp }, apply measurable_measure_prod_mk_right, - exact measurable_const.prod_mk measurable_mul (measurable_set.univ.prod hE) + exact measurable_const.prod_mk measurable_mul (measurable_set.univ.prod hs) end variables [has_measurable_inv G] @@ -144,10 +144,10 @@ begin end @[to_additive] -lemma measure_inv_null : μ E⁻¹ = 0 ↔ μ E = 0 := +lemma measure_inv_null : μ s⁻¹ = 0 ↔ μ s = 0 := begin - refine ⟨λ hE, _, (quasi_measure_preserving_inv μ).preimage_null⟩, - convert (quasi_measure_preserving_inv μ).preimage_null hE, + refine ⟨λ hs, _, (quasi_measure_preserving_inv μ).preimage_null⟩, + convert (quasi_measure_preserving_inv μ).preimage_null hs, exact (inv_inv _).symm end @@ -176,15 +176,15 @@ end @[to_additive] lemma measure_mul_right_null (y : G) : - μ ((λ x, x * y) ⁻¹' E) = 0 ↔ μ E = 0 := -calc μ ((λ x, x * y) ⁻¹' E) = 0 ↔ μ ((λ x, y⁻¹ * x) ⁻¹' E⁻¹)⁻¹ = 0 : + μ ((λ x, x * y) ⁻¹' s) = 0 ↔ μ s = 0 := +calc μ ((λ x, x * y) ⁻¹' s) = 0 ↔ μ ((λ x, y⁻¹ * x) ⁻¹' s⁻¹)⁻¹ = 0 : by simp_rw [← inv_preimage, preimage_preimage, mul_inv_rev, inv_inv] -... ↔ μ E = 0 : by simp only [measure_inv_null μ, measure_preimage_mul] +... ↔ μ s = 0 : by simp only [measure_inv_null μ, measure_preimage_mul] @[to_additive] lemma measure_mul_right_ne_zero - (h2E : μ E ≠ 0) (y : G) : μ ((λ x, x * y) ⁻¹' E) ≠ 0 := -(not_iff_not_of_iff (measure_mul_right_null μ y)).mpr h2E + (h2s : μ s ≠ 0) (y : G) : μ ((λ x, x * y) ⁻¹' s) ≠ 0 := +(not_iff_not_of_iff (measure_mul_right_null μ y)).mpr h2s /-- A *left*-invariant measure is quasi-preserved by *right*-multiplication. This should not be confused with `(measure_preserving_mul_right μ g).quasi_measure_preserving`. -/ @@ -224,20 +224,20 @@ end /-- This is the computation performed in the proof of [Halmos, §60 Th. A]. -/ @[to_additive "This is the computation performed in the proof of [Halmos, §60 Th. A]."] lemma measure_mul_lintegral_eq - [is_mul_left_invariant ν] (Em : measurable_set E) (f : G → ℝ≥0∞) (hf : measurable f) : - μ E * ∫⁻ y, f y ∂ν = ∫⁻ x, ν ((λ z, z * x) ⁻¹' E) * f (x⁻¹) ∂μ := + [is_mul_left_invariant ν] (sm : measurable_set s) (f : G → ℝ≥0∞) (hf : measurable f) : + μ s * ∫⁻ y, f y ∂ν = ∫⁻ x, ν ((λ z, z * x) ⁻¹' s) * f (x⁻¹) ∂μ := begin - rw [← set_lintegral_one, ← lintegral_indicator _ Em, - ← lintegral_lintegral_mul (measurable_const.indicator Em).ae_measurable hf.ae_measurable, + rw [← set_lintegral_one, ← lintegral_indicator _ sm, + ← lintegral_lintegral_mul (measurable_const.indicator sm).ae_measurable hf.ae_measurable, ← lintegral_lintegral_mul_inv μ ν], - swap, { exact (((measurable_const.indicator Em).comp measurable_fst).mul + swap, { exact (((measurable_const.indicator sm).comp measurable_fst).mul (hf.comp measurable_snd)).ae_measurable }, - have mE : ∀ x : G, measurable (λ y, ((λ z, z * x) ⁻¹' E).indicator (λ z, (1 : ℝ≥0∞)) y) := - λ x, measurable_const.indicator (measurable_mul_const _ Em), - have : ∀ x y, E.indicator (λ (z : G), (1 : ℝ≥0∞)) (y * x) = - ((λ z, z * x) ⁻¹' E).indicator (λ (b : G), 1) y, + have ms : ∀ x : G, measurable (λ y, ((λ z, z * x) ⁻¹' s).indicator (λ z, (1 : ℝ≥0∞)) y) := + λ x, measurable_const.indicator (measurable_mul_const _ sm), + have : ∀ x y, s.indicator (λ (z : G), (1 : ℝ≥0∞)) (y * x) = + ((λ z, z * x) ⁻¹' s).indicator (λ (b : G), 1) y, { intros x y, symmetry, convert indicator_comp_right (λ y, y * x), ext1 z, refl }, - simp_rw [this, lintegral_mul_const _ (mE _), lintegral_indicator _ (measurable_mul_const _ Em), + simp_rw [this, lintegral_mul_const _ (ms _), lintegral_indicator _ (measurable_mul_const _ sm), set_lintegral_one], end @@ -247,84 +247,84 @@ other. "-/] lemma absolutely_continuous_of_is_mul_left_invariant [is_mul_left_invariant ν] (hν : ν ≠ 0) : μ ≪ ν := begin - refine absolutely_continuous.mk (λ E Em hνE, _), - have h1 := measure_mul_lintegral_eq μ ν Em 1 measurable_one, - simp_rw [pi.one_apply, lintegral_one, mul_one, (measure_mul_right_null ν _).mpr hνE, + refine absolutely_continuous.mk (λ s sm hνs, _), + have h1 := measure_mul_lintegral_eq μ ν sm 1 measurable_one, + simp_rw [pi.one_apply, lintegral_one, mul_one, (measure_mul_right_null ν _).mpr hνs, lintegral_zero, mul_eq_zero, measure_univ_eq_zero.not.mpr hν, or_false] at h1, exact h1 end @[to_additive] lemma ae_measure_preimage_mul_right_lt_top [is_mul_left_invariant ν] - (Em : measurable_set E) (hμE : μ E ≠ ∞) : - ∀ᵐ x ∂μ, ν ((λ y, y * x) ⁻¹' E) < ∞ := + (sm : measurable_set s) (hμs : μ s ≠ ∞) : + ∀ᵐ x ∂μ, ν ((λ y, y * x) ⁻¹' s) < ∞ := begin refine ae_of_forall_measure_lt_top_ae_restrict' ν.inv _ _, intros A hA h2A h3A, simp only [ν.inv_apply] at h3A, - apply ae_lt_top (measurable_measure_mul_right ν Em), - have h1 := measure_mul_lintegral_eq μ ν Em (A⁻¹.indicator 1) (measurable_one.indicator hA.inv), + apply ae_lt_top (measurable_measure_mul_right ν sm), + have h1 := measure_mul_lintegral_eq μ ν sm (A⁻¹.indicator 1) (measurable_one.indicator hA.inv), rw [lintegral_indicator _ hA.inv] at h1, simp_rw [pi.one_apply, set_lintegral_one, ← image_inv, indicator_image inv_injective, image_inv, - ← indicator_mul_right _ (λ x, ν ((λ y, y * x) ⁻¹' E)), function.comp, pi.one_apply, + ← indicator_mul_right _ (λ x, ν ((λ y, y * x) ⁻¹' s)), function.comp, pi.one_apply, mul_one] at h1, rw [← lintegral_indicator _ hA, ← h1], - exact ennreal.mul_ne_top hμE h3A.ne, + exact ennreal.mul_ne_top hμs h3A.ne, end @[to_additive] lemma ae_measure_preimage_mul_right_lt_top_of_ne_zero [is_mul_left_invariant ν] - (Em : measurable_set E) (h2E : ν E ≠ 0) (h3E : ν E ≠ ∞) : - ∀ᵐ x ∂μ, ν ((λ y, y * x) ⁻¹' E) < ∞ := + (sm : measurable_set s) (h2s : ν s ≠ 0) (h3s : ν s ≠ ∞) : + ∀ᵐ x ∂μ, ν ((λ y, y * x) ⁻¹' s) < ∞ := begin - refine (ae_measure_preimage_mul_right_lt_top ν ν Em h3E).filter_mono _, + refine (ae_measure_preimage_mul_right_lt_top ν ν sm h3s).filter_mono _, refine (absolutely_continuous_of_is_mul_left_invariant μ ν _).ae_le, - refine mt _ h2E, + refine mt _ h2s, intro hν, rw [hν, measure.coe_zero, pi.zero_apply] end /-- A technical lemma relating two different measures. This is basically [Halmos, §60 Th. A]. - Note that if `f` is the characteristic function of a measurable set `F` this states that - `μ F = c * μ E` for a constant `c` that does not depend on `μ`. + Note that if `f` is the characteristic function of a measurable set `t` this states that + `μ t = c * μ s` for a constant `c` that does not depend on `μ`. Note: There is a gap in the last step of the proof in [Halmos]. - In the last line, the equality `g(x⁻¹)ν(Ex⁻¹) = f(x)` holds if we can prove that - `0 < ν(Ex⁻¹) < ∞`. The first inequality follows from §59, Th. D, but the second inequality is + In the last line, the equality `g(x⁻¹)ν(sx⁻¹) = f(x)` holds if we can prove that + `0 < ν(sx⁻¹) < ∞`. The first inequality follows from §59, Th. D, but the second inequality is not justified. We prove this inequality for almost all `x` in `measure_theory.ae_measure_preimage_mul_right_lt_top_of_ne_zero`. -/ @[to_additive "A technical lemma relating two different measures. This is basically -[Halmos, §60 Th. A]. Note that if `f` is the characteristic function of a measurable set `F` this -states that `μ F = c * μ E` for a constant `c` that does not depend on `μ`. +[Halmos, §60 Th. A]. Note that if `f` is the characteristic function of a measurable set `t` this +states that `μ t = c * μ s` for a constant `c` that does not depend on `μ`. Note: There is a gap in the last step of the proof in [Halmos]. In the last line, the equality -`g(-x) + ν(E - x) = f(x)` holds if we can prove that `0 < ν(E - x) < ∞`. The first inequality +`g(-x) + ν(s - x) = f(x)` holds if we can prove that `0 < ν(s - x) < ∞`. The first inequality follows from §59, Th. D, but the second inequality is not justified. We prove this inequality for almost all `x` in `measure_theory.ae_measure_preimage_add_right_lt_top_of_ne_zero`."] lemma measure_lintegral_div_measure [is_mul_left_invariant ν] - (Em : measurable_set E) (h2E : ν E ≠ 0) (h3E : ν E ≠ ∞) + (sm : measurable_set s) (h2s : ν s ≠ 0) (h3s : ν s ≠ ∞) (f : G → ℝ≥0∞) (hf : measurable f) : - μ E * ∫⁻ y, f y⁻¹ / ν ((λ x, x * y⁻¹) ⁻¹' E) ∂ν = ∫⁻ x, f x ∂μ := + μ s * ∫⁻ y, f y⁻¹ / ν ((λ x, x * y⁻¹) ⁻¹' s) ∂ν = ∫⁻ x, f x ∂μ := begin - set g := λ y, f y⁻¹ / ν ((λ x, x * y⁻¹) ⁻¹' E), + set g := λ y, f y⁻¹ / ν ((λ x, x * y⁻¹) ⁻¹' s), have hg : measurable g := (hf.comp measurable_inv).div - ((measurable_measure_mul_right ν Em).comp measurable_inv), - simp_rw [measure_mul_lintegral_eq μ ν Em g hg, g, inv_inv], + ((measurable_measure_mul_right ν sm).comp measurable_inv), + simp_rw [measure_mul_lintegral_eq μ ν sm g hg, g, inv_inv], refine lintegral_congr_ae _, - refine (ae_measure_preimage_mul_right_lt_top_of_ne_zero μ ν Em h2E h3E).mono (λ x hx , _), - simp_rw [ennreal.mul_div_cancel' (measure_mul_right_ne_zero ν h2E _) hx.ne] + refine (ae_measure_preimage_mul_right_lt_top_of_ne_zero μ ν sm h2s h3s).mono (λ x hx , _), + simp_rw [ennreal.mul_div_cancel' (measure_mul_right_ne_zero ν h2s _) hx.ne] end @[to_additive] -lemma measure_mul_measure_eq [is_mul_left_invariant ν] {E F : set G} - (hE : measurable_set E) (hF : measurable_set F) (h2E : ν E ≠ 0) (h3E : ν E ≠ ∞) : - μ E * ν F = ν E * μ F := +lemma measure_mul_measure_eq [is_mul_left_invariant ν] {s t : set G} + (hs : measurable_set s) (ht : measurable_set t) (h2s : ν s ≠ 0) (h3s : ν s ≠ ∞) : + μ s * ν t = ν s * μ t := begin - have h1 := measure_lintegral_div_measure ν ν hE h2E h3E (F.indicator (λ x, 1)) - (measurable_const.indicator hF), - have h2 := measure_lintegral_div_measure μ ν hE h2E h3E (F.indicator (λ x, 1)) - (measurable_const.indicator hF), - rw [lintegral_indicator _ hF, set_lintegral_one] at h1 h2, + have h1 := measure_lintegral_div_measure ν ν hs h2s h3s (t.indicator (λ x, 1)) + (measurable_const.indicator ht), + have h2 := measure_lintegral_div_measure μ ν hs h2s h3s (t.indicator (λ x, 1)) + (measurable_const.indicator ht), + rw [lintegral_indicator _ ht, set_lintegral_one] at h1 h2, rw [← h1, mul_left_comm, h2], end @@ -332,11 +332,11 @@ end @[to_additive /-" Left invariant Borel measures on an additive measurable group are unique (up to a scalar). "-/] lemma measure_eq_div_smul [is_mul_left_invariant ν] - (hE : measurable_set E) (h2E : ν E ≠ 0) (h3E : ν E ≠ ∞) : μ = (μ E / ν E) • ν := + (hs : measurable_set s) (h2s : ν s ≠ 0) (h3s : ν s ≠ ∞) : μ = (μ s / ν s) • ν := begin - ext1 F hF, + ext1 t ht, rw [smul_apply, smul_eq_mul, mul_comm, ← mul_div_assoc, mul_comm, - measure_mul_measure_eq μ ν hE hF h2E h3E, mul_div_assoc, ennreal.mul_div_cancel' h2E h3E] + measure_mul_measure_eq μ ν hs ht h2s h3s, mul_div_assoc, ennreal.mul_div_cancel' h2s h3s] end end measure_theory From 5b4d38a292535f15a6732c22246887a6d753a967 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Wed, 12 Oct 2022 07:21:08 +0000 Subject: [PATCH 712/828] feat(algebra/order/sub): add tsub_tsub_tsub_cancel_left (#16887) --- src/algebra/order/sub.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/algebra/order/sub.lean b/src/algebra/order/sub.lean index 4657f9907cd89..c9cc427b90096 100644 --- a/src/algebra/order/sub.lean +++ b/src/algebra/order/sub.lean @@ -532,6 +532,10 @@ protected lemma tsub_tsub_cancel_of_le (hba : add_le_cancellable (b - a)) (h : a b - (b - a) = a := hba.tsub_eq_of_eq_add (add_tsub_cancel_of_le h).symm +protected lemma tsub_tsub_tsub_cancel_left (hab : add_le_cancellable (a - b)) (h : b ≤ a) : + a - c - (a - b) = b - c := +by rw [tsub_right_comm, hab.tsub_tsub_cancel_of_le h] + end add_le_cancellable section contra @@ -610,6 +614,9 @@ contravariant.add_le_cancellable.add_tsub_tsub_cancel h lemma tsub_tsub_cancel_of_le (h : a ≤ b) : b - (b - a) = a := contravariant.add_le_cancellable.tsub_tsub_cancel_of_le h +lemma tsub_tsub_tsub_cancel_left (h : b ≤ a) : a - c - (a - b) = b - c := +contravariant.add_le_cancellable.tsub_tsub_tsub_cancel_left h + end contra end has_exists_add_of_le From 01788b5e4cd1bef928f706c815c30fa9d376a089 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 12 Oct 2022 07:21:10 +0000 Subject: [PATCH 713/828] feat(group_theory/subgroup/basic): add some trivial lemmas (#16915) * add `subgroup.subtype_injective`, `subgroup.le_closure_to_submonoid`, `subgroup.closure_eq_top_of_mclosure_eq_top`; * golf `subgroup.closure_inv`, `subgroup.map_le_map_iff_of_injective`, and `subgroup.map_injective`. --- src/group_theory/subgroup/basic.lean | 25 +++++++++++++++---------- 1 file changed, 15 insertions(+), 10 deletions(-) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 70f36e771b93f..b504acc242a35 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -586,6 +586,8 @@ def subtype : H →* G := ⟨coe, rfl, λ _ _, rfl⟩ @[simp, to_additive] theorem coe_subtype : ⇑H.subtype = coe := rfl +@[to_additive] lemma subtype_injective : function.injective (subtype H) := subtype.coe_injective + @[simp, norm_cast, to_additive] theorem coe_list_prod (l : list H) : (l.prod : G) = (l.map coe).prod := submonoid_class.coe_list_prod l @@ -999,13 +1001,6 @@ begin exact subset_closure (mem_inv.mp hs), end -@[simp, to_additive] lemma closure_inv (S : set G) : closure S⁻¹ = closure S := -begin - refine le_antisymm ((subgroup.closure_le _).2 _) ((subgroup.closure_le _).2 _), - { exact inv_subset_closure S }, - { simpa only [inv_inv] using inv_subset_closure S⁻¹ }, -end - @[to_additive] lemma closure_to_submonoid (S : set G) : (closure S).to_submonoid = submonoid.closure (S ∪ S⁻¹) := @@ -1019,6 +1014,17 @@ begin { simp only [true_and, coe_to_submonoid, union_subset_iff, subset_closure, inv_subset_closure] } end +@[to_additive] +lemma le_closure_to_submonoid (S : set G) : submonoid.closure S ≤ (closure S).to_submonoid := +submonoid.closure_le.2 subset_closure + +@[simp, to_additive] lemma closure_inv (S : set G) : closure S⁻¹ = closure S := +by simp only [← to_submonoid_eq, closure_to_submonoid, inv_inv, union_comm] + +@[to_additive] lemma closure_eq_top_of_mclosure_eq_top {S : set G} (h : submonoid.closure S = ⊤) : + closure S = ⊤ := +(eq_top_iff' _).2 $ λ x, le_closure_to_submonoid _ $ h.symm ▸ trivial + @[to_additive] lemma closure_induction_left {p : G → Prop} {x : G} (h : x ∈ closure k) (H1 : p 1) (Hmul : ∀ (x ∈ k) y, p y → p (x * y)) (Hinv : ∀ (x ∈ k) y, p y → p (x⁻¹ * y)) : p x := @@ -2382,8 +2388,7 @@ comap_map_eq_self (((ker_eq_bot_iff _).mpr h).symm ▸ bot_le) @[to_additive] lemma map_le_map_iff_of_injective {f : G →* N} (hf : function.injective f) {H K : subgroup G} : H.map f ≤ K.map f ↔ H ≤ K := -⟨(congr_arg2 (≤) (H.comap_map_eq_self_of_injective hf) - (K.comap_map_eq_self_of_injective hf)).mp ∘ comap_mono, map_mono⟩ +by rw [map_le_iff_le_comap, comap_map_eq_self_of_injective hf] @[simp, to_additive] lemma map_subtype_le_map_subtype {G' : subgroup G} {H K : subgroup G'} : @@ -2392,7 +2397,7 @@ map_le_map_iff_of_injective subtype.coe_injective @[to_additive] lemma map_injective {f : G →* N} (h : function.injective f) : function.injective (map f) := -λ K L hKL, by { apply_fun comap f at hKL, simpa [comap_map_eq_self_of_injective h] using hKL } +function.left_inverse.injective $ comap_map_eq_self_of_injective h @[to_additive] lemma map_eq_comap_of_inverse {f : G →* N} {g : N →* G} (hl : function.left_inverse g f) From 265fe3d814d6eb043ec6c269196df77f30c605fc Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Wed, 12 Oct 2022 10:35:36 +0000 Subject: [PATCH 714/828] feat(analysis/inner_product_space/orientation): volume form (#16487) Construct the volume form, a uniquely-determined top-dimensional alternating form on an oriented real inner product space. Formalized as part of the Sphere Eversion project. --- .../inner_product_space/orientation.lean | 218 ++++++++++++++++-- src/linear_algebra/determinant.lean | 6 + src/linear_algebra/orientation.lean | 40 ++++ 3 files changed, 249 insertions(+), 15 deletions(-) diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index f83291043780d..b0daf2a16f765 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -1,9 +1,9 @@ /- Copyright (c) 2022 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Joseph Myers +Authors: Joseph Myers, Heather Macbeth -/ -import analysis.inner_product_space.pi_L2 +import analysis.inner_product_space.gram_schmidt_ortho import linear_algebra.orientation /-! @@ -13,8 +13,23 @@ This file provides definitions and proves lemmas about orientations of real inne ## Main definitions +* `orthonormal_basis.adjust_to_orientation` takes an orthonormal basis and an orientation, and + returns an orthonormal basis with that orientation: either the original orthonormal basis, or one + constructed by negating a single (arbitrary) basis vector. * `orientation.fin_orthonormal_basis` is an orthonormal basis, indexed by `fin n`, with the given -orientation. + orientation. +* `orientation.volume_form` is a nonvanishing top-dimensional alternating form on an oriented real + inner product space, uniquely defined by compatibility with the orientation and inner product + structure. + +## Main theorems + +* `orientation.volume_form_apply_le` states that the result of applying the volume form to a set of + `n` vectors, where `n` is the dimension the inner product space, is bounded by the product of the + lengths of the vectors. +* `orientation.abs_volume_form_apply_of_pairwise_orthogonal` states that the result of applying the + volume form to a set of `n` orthogonal vectors, where `n` is the dimension the inner product + space, is equal up to sign to the product of the lengths of the vectors. -/ @@ -23,15 +38,49 @@ noncomputable theory variables {E : Type*} [inner_product_space ℝ E] open finite_dimensional +open_locale big_operators real_inner_product_space -section adjust_to_orientation -variables {ι : Type*} [fintype ι] [decidable_eq ι] [nonempty ι] (e : orthonormal_basis ι ℝ E) +namespace orthonormal_basis +variables {ι : Type*} [fintype ι] [decidable_eq ι] [ne : nonempty ι] (e f : orthonormal_basis ι ℝ E) (x : orientation ℝ E ι) -/-- `orthonormal_basis.adjust_to_orientation`, applied to an orthonormal basis, produces an -orthonormal basis. -/ -lemma orthonormal_basis.orthonormal_adjust_to_orientation : - orthonormal ℝ (e.to_basis.adjust_to_orientation x) := +/-- The change-of-basis matrix between two orthonormal bases with the same orientation has +determinant 1. -/ +lemma det_to_matrix_orthonormal_basis_of_same_orientation + (h : e.to_basis.orientation = f.to_basis.orientation) : + e.to_basis.det f = 1 := +begin + apply (e.det_to_matrix_orthonormal_basis_real f).resolve_right, + have : 0 < e.to_basis.det f, + { rw e.to_basis.orientation_eq_iff_det_pos at h, + simpa using h }, + linarith, +end + +variables {e f} + +/-- Two orthonormal bases with the same orientation determine the same "determinant" top-dimensional +form on `E`, and conversely. -/ +lemma same_orientation_iff_det_eq_det : + e.to_basis.det = f.to_basis.det ↔ e.to_basis.orientation = f.to_basis.orientation := +begin + split, + { intros h, + dsimp [basis.orientation], + congr' }, + { intros h, + rw e.to_basis.det.eq_smul_basis_det f.to_basis, + simp [e.det_to_matrix_orthonormal_basis_of_same_orientation f h], }, +end + +variables (e) + +section adjust_to_orientation +include ne + +/-- `orthonormal_basis.adjust_to_orientation`, applied to an orthonormal basis, preserves the +property of orthonormality. -/ +lemma orthonormal_adjust_to_orientation : orthonormal ℝ (e.to_basis.adjust_to_orientation x) := begin apply e.orthonormal.orthonormal_of_forall_eq_or_eq_neg, simpa using e.to_basis.adjust_to_orientation_apply_eq_or_eq_neg x @@ -40,15 +89,15 @@ end /-- Given an orthonormal basis and an orientation, return an orthonormal basis giving that orientation: either the original basis, or one constructed by negating a single (arbitrary) basis vector. -/ -def orthonormal_basis.adjust_to_orientation : orthonormal_basis ι ℝ E := +def adjust_to_orientation : orthonormal_basis ι ℝ E := (e.to_basis.adjust_to_orientation x).to_orthonormal_basis (e.orthonormal_adjust_to_orientation x) -lemma orthonormal_basis.to_basis_adjust_to_orientation : +lemma to_basis_adjust_to_orientation : (e.adjust_to_orientation x).to_basis = e.to_basis.adjust_to_orientation x := (e.to_basis.adjust_to_orientation x).to_basis_to_orthonormal_basis _ /-- `adjust_to_orientation` gives an orthonormal basis with the required orientation. -/ -@[simp] lemma orthonormal_basis.orientation_adjust_to_orientation : +@[simp] lemma orientation_adjust_to_orientation : (e.adjust_to_orientation x).to_basis.orientation = x := begin rw e.to_basis_adjust_to_orientation, @@ -57,15 +106,31 @@ end /-- Every basis vector from `adjust_to_orientation` is either that from the original basis or its negation. -/ -lemma orthonormal_basis.adjust_to_orientation_apply_eq_or_eq_neg (i : ι) : +lemma adjust_to_orientation_apply_eq_or_eq_neg (i : ι) : e.adjust_to_orientation x i = e i ∨ e.adjust_to_orientation x i = -(e i) := by simpa [← e.to_basis_adjust_to_orientation] using e.to_basis.adjust_to_orientation_apply_eq_or_eq_neg x i +lemma det_adjust_to_orientation : + (e.adjust_to_orientation x).to_basis.det = e.to_basis.det + ∨ (e.adjust_to_orientation x).to_basis.det = -e.to_basis.det := +by simpa using e.to_basis.det_adjust_to_orientation x + +lemma abs_det_adjust_to_orientation (v : ι → E) : + |(e.adjust_to_orientation x).to_basis.det v| = |e.to_basis.det v| := +by simp [to_basis_adjust_to_orientation] + end adjust_to_orientation +end orthonormal_basis + +namespace orientation +variables {n : ℕ} + +open orthonormal_basis + /-- An orthonormal basis, indexed by `fin n`, with the given orientation. -/ -protected def orientation.fin_orthonormal_basis {n : ℕ} (hn : 0 < n) (h : finrank ℝ E = n) +protected def fin_orthonormal_basis (hn : 0 < n) (h : finrank ℝ E = n) (x : orientation ℝ E (fin n)) : orthonormal_basis (fin n) ℝ E := begin haveI := fin.pos_iff_nonempty.1 hn, @@ -74,7 +139,7 @@ begin end /-- `orientation.fin_orthonormal_basis` gives a basis with the required orientation. -/ -@[simp] lemma orientation.fin_orthonormal_basis_orientation {n : ℕ} (hn : 0 < n) +@[simp] lemma fin_orthonormal_basis_orientation (hn : 0 < n) (h : finrank ℝ E = n) (x : orientation ℝ E (fin n)) : (x.fin_orthonormal_basis hn h).to_basis.orientation = x := begin @@ -82,3 +147,126 @@ begin haveI := finite_dimensional_of_finrank (h.symm ▸ hn : 0 < finrank ℝ E), exact ((std_orthonormal_basis _ _).reindex $ fin_congr h).orientation_adjust_to_orientation x end + +section volume_form +variables [_i : fact (finrank ℝ E = n)] (o : orientation ℝ E (fin n)) + +include _i o + +/-- The volume form on an oriented real inner product space, a nonvanishing top-dimensional +alternating form uniquely defined by compatibility with the orientation and inner product structure. +-/ +@[irreducible] def volume_form : alternating_map ℝ E ℝ (fin n) := +begin + classical, + unfreezingI { cases n }, + { let opos : alternating_map ℝ E ℝ (fin 0) := alternating_map.const_of_is_empty ℝ E (1:ℝ), + let oneg : alternating_map ℝ E ℝ (fin 0) := alternating_map.const_of_is_empty ℝ E (-1:ℝ), + exact o.eq_or_eq_neg_of_is_empty.by_cases (λ _, opos) (λ _, oneg) }, + { exact (o.fin_orthonormal_basis n.succ_pos _i.out).to_basis.det } +end + +omit _i o + +@[simp] lemma volume_form_zero_pos [_i : fact (finrank ℝ E = 0)] : + orientation.volume_form (positive_orientation : orientation ℝ E (fin 0)) + = alternating_map.const_linear_equiv_of_is_empty 1 := +by simp [volume_form, or.by_cases, dif_pos] + +@[simp] lemma volume_form_zero_neg [_i : fact (finrank ℝ E = 0)] : + orientation.volume_form (-positive_orientation : orientation ℝ E (fin 0)) + = alternating_map.const_linear_equiv_of_is_empty (-1) := +begin + dsimp [volume_form, or.by_cases, positive_orientation], + apply if_neg, + rw [ray_eq_iff, same_ray_comm], + intros h, + simpa using + congr_arg alternating_map.const_linear_equiv_of_is_empty.symm (eq_zero_of_same_ray_self_neg h), +end + +include _i o + +/-- The volume form on an oriented real inner product space can be evaluated as the determinant with +respect to any orthonormal basis of the space compatible with the orientation. -/ +lemma volume_form_robust (b : orthonormal_basis (fin n) ℝ E) (hb : b.to_basis.orientation = o) : + o.volume_form = b.to_basis.det := +begin + unfreezingI { cases n }, + { have : o = positive_orientation := hb.symm.trans b.to_basis.orientation_is_empty, + simp [volume_form, or.by_cases, dif_pos this] }, + { dsimp [volume_form], + rw [same_orientation_iff_det_eq_det, hb], + exact o.fin_orthonormal_basis_orientation _ _ }, +end + +lemma volume_form_robust' (b : orthonormal_basis (fin n) ℝ E) (v : fin n → E) : + |o.volume_form v| = |b.to_basis.det v| := +begin + unfreezingI { cases n }, + { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, + { rw [o.volume_form_robust (b.adjust_to_orientation o) (b.orientation_adjust_to_orientation o), + b.abs_det_adjust_to_orientation] }, +end + +/-- Let `v` be an indexed family of `n` vectors in an oriented `n`-dimensional real inner +product space `E`. The output of the volume form of `E` when evaluated on `v` is bounded in absolute +value by the product of the norms of the vectors `v i`. -/ +lemma abs_volume_form_apply_le (v : fin n → E) : |o.volume_form v| ≤ ∏ i : fin n, ∥v i∥ := +begin + unfreezingI { cases n }, + { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, + haveI : finite_dimensional ℝ E := fact_finite_dimensional_of_finrank_eq_succ n, + have : finrank ℝ E = fintype.card (fin n.succ) := by simpa using _i.out, + let b : orthonormal_basis (fin n.succ) ℝ E := gram_schmidt_orthonormal_basis this v, + have hb : b.to_basis.det v = ∏ i, ⟪b i, v i⟫ := gram_schmidt_orthonormal_basis_det this v, + rw [o.volume_form_robust' b, hb, finset.abs_prod], + apply finset.prod_le_prod, + { intros i hi, + positivity }, + intros i hi, + convert abs_real_inner_le_norm (b i) (v i), + simp [b.orthonormal.1 i], +end + +lemma volume_form_apply_le (v : fin n → E) : o.volume_form v ≤ ∏ i : fin n, ∥v i∥ := +(le_abs_self _).trans (o.abs_volume_form_apply_le v) + +/-- Let `v` be an indexed family of `n` orthogonal vectors in an oriented `n`-dimensional +real inner product space `E`. The output of the volume form of `E` when evaluated on `v` is, up to +sign, the product of the norms of the vectors `v i`. -/ +lemma abs_volume_form_apply_of_pairwise_orthogonal + {v : fin n → E} (hv : pairwise (λ i j, ⟪v i, v j⟫ = 0)) : + |o.volume_form v| = ∏ i : fin n, ∥v i∥ := +begin + unfreezingI { cases n }, + { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, + haveI : finite_dimensional ℝ E := fact_finite_dimensional_of_finrank_eq_succ n, + have hdim : finrank ℝ E = fintype.card (fin n.succ) := by simpa using _i.out, + let b : orthonormal_basis (fin n.succ) ℝ E := gram_schmidt_orthonormal_basis hdim v, + have hb : b.to_basis.det v = ∏ i, ⟪b i, v i⟫ := gram_schmidt_orthonormal_basis_det hdim v, + rw [o.volume_form_robust' b, hb, finset.abs_prod], + by_cases h : ∃ i, v i = 0, + obtain ⟨i, hi⟩ := h, + { rw [finset.prod_eq_zero (finset.mem_univ i), finset.prod_eq_zero (finset.mem_univ i)]; + simp [hi] }, + push_neg at h, + congr, + ext i, + have hb : b i = ∥v i∥⁻¹ • v i := gram_schmidt_orthonormal_basis_apply_of_orthogonal hdim hv (h i), + simp only [hb, inner_smul_left, real_inner_self_eq_norm_mul_norm, is_R_or_C.conj_to_real], + rw abs_of_nonneg, + { have : ∥v i∥ ≠ 0 := by simpa using h i, + field_simp }, + { positivity }, +end + +/-- The output of the volume form of an oriented real inner product space `E` when evaluated on an +orthonormal basis is ±1. -/ +lemma abs_volume_form_apply_of_orthonormal (v : orthonormal_basis (fin n) ℝ E) : + |o.volume_form v| = 1 := +by simpa [o.volume_form_robust' v v] using congr_arg abs v.to_basis.det_self + +end volume_form + +end orientation diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index 3894ad659cc10..0347ab5baa19f 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -454,6 +454,12 @@ lemma basis.det_apply (v : ι → M) : e.det v = det (e.to_matrix v) := rfl lemma basis.det_self : e.det e = 1 := by simp [e.det_apply] +@[simp] lemma basis.det_is_empty [is_empty ι] : e.det = alternating_map.const_of_is_empty R M 1 := +begin + ext v, + exact matrix.det_is_empty, +end + /-- `basis.det` is not the zero map. -/ lemma basis.det_ne_zero [nontrivial R] : e.det ≠ 0 := λ h, by simpa [h] using e.det_self diff --git a/src/linear_algebra/orientation.lean b/src/linear_algebra/orientation.lean index 52be381be8078..bc5c073ef61af 100644 --- a/src/linear_algebra/orientation.lean +++ b/src/linear_algebra/orientation.lean @@ -50,6 +50,8 @@ abbreviation orientation := module.ray R (alternating_map R M R ι) class module.oriented := (positive_orientation : orientation R M ι) +export module.oriented (positive_orientation) + variables {R M} /-- An equivalence between modules implies an equivalence between orientations. -/ @@ -68,6 +70,12 @@ by rw [orientation.map, alternating_map.dom_lcongr_refl, module.ray.map_refl] @[simp] lemma orientation.map_symm (e : M ≃ₗ[R] N) : (orientation.map ι e).symm = orientation.map ι e.symm := rfl +/-- A module is canonically oriented with respect to an empty index type. -/ +@[priority 100] instance is_empty.oriented [nontrivial R] [is_empty ι] : + module.oriented R M ι := +{ positive_orientation := ray_of_ne_zero R (alternating_map.const_linear_equiv_of_is_empty 1) $ + alternating_map.const_linear_equiv_of_is_empty.injective.ne (by simp) } + end ordered_comm_semiring section ordered_comm_ring @@ -113,6 +121,13 @@ begin simp end +@[simp] lemma orientation_is_empty [nontrivial R] [is_empty ι] (b : basis ι R M) : + b.orientation = positive_orientation := +begin + congrm ray_of_ne_zero _ _ _, + convert b.det_is_empty, +end + end basis end ordered_comm_ring @@ -123,6 +138,31 @@ variables {R : Type*} [linear_ordered_comm_ring R] variables {M : Type*} [add_comm_group M] [module R M] variables {ι : Type*} [decidable_eq ι] +namespace orientation + +/-- A module `M` over a linearly ordered commutative ring has precisely two "orientations" with +respect to an empty index type. (Note that these are only orientations of `M` of in the conventional +mathematical sense if `M` is zero-dimensional.) -/ +lemma eq_or_eq_neg_of_is_empty [nontrivial R] [is_empty ι] (o : orientation R M ι) : + o = positive_orientation ∨ o = - positive_orientation := +begin + induction o using module.ray.ind with x hx, + dsimp [positive_orientation], + simp only [ray_eq_iff, same_ray_neg_swap], + rw same_ray_or_same_ray_neg_iff_not_linear_independent, + intros h, + let a : R := alternating_map.const_linear_equiv_of_is_empty.symm x, + have H : linear_independent R ![a, 1], + { convert h.map' ↑alternating_map.const_linear_equiv_of_is_empty.symm (linear_equiv.ker _), + ext i, + fin_cases i; + simp [a] }, + rw linear_independent_iff' at H, + simpa using H finset.univ ![1, -a] (by simp [fin.sum_univ_succ]) 0 (by simp), +end + +end orientation + namespace basis variables [fintype ι] From 032f11f8f575876899a9b5d16b8f4b65c2bd6f41 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 12 Oct 2022 10:35:37 +0000 Subject: [PATCH 715/828] feat(category_theory/localization): developing the predicate is_localization (#16890) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When `L : C ⥤ D` and `W : morphism_property C`, a constructor for the predicate `L.is_localization W` is introduced: it takes as inputs the universal property of the localized category. In this PR, it is also shown that is `L.is_localization W`, the functor `L` is essentially surjective. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- .../localization/predicate.lean | 121 +++++++++++++++++- src/category_theory/whiskering.lean | 3 + 2 files changed, 122 insertions(+), 2 deletions(-) diff --git a/src/category_theory/localization/predicate.lean b/src/category_theory/localization/predicate.lean index dec36bd151ce2..f14aa0be1d8bc 100644 --- a/src/category_theory/localization/predicate.lean +++ b/src/category_theory/localization/predicate.lean @@ -11,15 +11,23 @@ import category_theory.localization.construction # Predicate for localized categories In this file, a predicate `L.is_localization W` is introduced for a functor `L : C ⥤ D` -and `W : morphism_property C`. It expresses that `L` identifies `D` with the localized +and `W : morphism_property C`: it expresses that `L` identifies `D` with the localized category of `C` with respect to `W` (up to equivalence). +We introduce a universal property `strict_universal_property_fixed_target L W E` which +states that `L` inverts the morphisms in `W` and that all functors `C ⥤ E` inverting +`W` uniquely factors as a composition of `L ⋙ G` with `G : D ⥤ E`. Such universal +properties are inputs for the constructor `is_localization.mk'` for `L.is_localization W`. + -/ +noncomputable theory + namespace category_theory variables {C D : Type*} [category C] [category D] -variables (L : C ⥤ D) (W : morphism_property C) + (L : C ⥤ D) (W : morphism_property C) + (E : Type*) [category E] namespace functor @@ -41,4 +49,113 @@ instance Q_is_localization : W.Q.is_localization W := end functor +namespace localization + +/-- This universal property states that a functor `L : C ⥤ D` inverts morphisms +in `W` and the all functors `D ⥤ E` (for a fixed category `E`) uniquely factors +through `L`. -/ +structure strict_universal_property_fixed_target := +(inverts : W.is_inverted_by L) +(lift : Π (F : C ⥤ E) (hF : W.is_inverted_by F), D ⥤ E) +(fac : Π (F : C ⥤ E) (hF : W.is_inverted_by F), L ⋙ lift F hF = F) +(uniq : Π (F₁ F₂ : D ⥤ E) (h : L ⋙ F₁ = L ⋙ F₂), F₁ = F₂) + +/-- The localized category `W.localization` that was constructed satisfies +the universal property of the localization. -/ +@[simps] +def strict_universal_property_fixed_target_Q : + strict_universal_property_fixed_target W.Q W E := +{ inverts := W.Q_inverts, + lift := construction.lift, + fac := construction.fac, + uniq := construction.uniq, } + +instance : inhabited (strict_universal_property_fixed_target W.Q W E) := +⟨strict_universal_property_fixed_target_Q _ _⟩ + +/-- When `W` consists of isomorphisms, the identity satisfies the universal property +of the localization. -/ +@[simps] +def strict_universal_property_fixed_target_id (hW : W ⊆ morphism_property.isomorphisms C): + strict_universal_property_fixed_target (𝟭 C) W E := +{ inverts := λ X Y f hf, hW f hf, + lift := λ F hF, F, + fac := λ F hF, by { cases F, refl, }, + uniq := λ F₁ F₂ eq, by { cases F₁, cases F₂, exact eq, }, } + +end localization + +namespace functor + +lemma is_localization.mk' + (h₁ : localization.strict_universal_property_fixed_target L W D) + (h₂ : localization.strict_universal_property_fixed_target L W W.localization) : + is_localization L W := +{ inverts := h₁.inverts, + nonempty_is_equivalence := nonempty.intro + { inverse := h₂.lift W.Q W.Q_inverts, + unit_iso := eq_to_iso (localization.construction.uniq _ _ + (by simp only [← functor.assoc, localization.construction.fac, h₂.fac, functor.comp_id])), + counit_iso := eq_to_iso (h₁.uniq _ _ (by simp only [← functor.assoc, h₂.fac, + localization.construction.fac, functor.comp_id])), + functor_unit_iso_comp' := λ X, by simpa only [eq_to_iso.hom, eq_to_hom_app, + eq_to_hom_map, eq_to_hom_trans, eq_to_hom_refl], }, } + +lemma is_localization.for_id (hW : W ⊆ morphism_property.isomorphisms C): + (𝟭 C).is_localization W := +is_localization.mk' _ _ + (localization.strict_universal_property_fixed_target_id W _ hW) + (localization.strict_universal_property_fixed_target_id W _ hW) + +end functor + +namespace localization + +variable [L.is_localization W] +include L W + +lemma inverts : W.is_inverted_by L := (infer_instance : L.is_localization W).inverts + +variable {W} + +/-- The isomorphism `L.obj X ≅ L.obj Y` that is deduced from a morphism `f : X ⟶ Y` which +belongs to `W`, when `L.is_localization W`. -/ +@[simps] +def iso_of_hom {X Y : C} (f : X ⟶ Y) (hf : W f) : L.obj X ≅ L.obj Y := +by { haveI : is_iso (L.map f) := inverts L W f hf, exact as_iso (L.map f), } + +variable (W) + +instance : is_equivalence (localization.construction.lift L (inverts L W)) := +(infer_instance : L.is_localization W).nonempty_is_equivalence.some + +/-- A chosen equivalence of categories `W.localization ≅ D` for a functor +`L : C ⥤ D` which satisfies `L.is_localization W`. This shall be used in +order to deduce properties of `L` from properties of `W.Q`. -/ +def equivalence_from_model : W.localization ≌ D := +(localization.construction.lift L (inverts L W)).as_equivalence + +/-- Via the equivalence of categories `equivalence_from_model L W : W.localization ≌ D`, +one may identify the functors `W.Q` and `L`. -/ +def Q_comp_equivalence_from_model_functor_iso : + W.Q ⋙ (equivalence_from_model L W).functor ≅ L := eq_to_iso (construction.fac _ _) + +/-- Via the equivalence of categories `equivalence_from_model L W : W.localization ≌ D`, +one may identify the functors `L` and `W.Q`. -/ +def comp_equivalence_from_model_inverse_iso : + L ⋙ (equivalence_from_model L W).inverse ≅ W.Q := +calc L ⋙ (equivalence_from_model L W).inverse ≅ _ : + iso_whisker_right (Q_comp_equivalence_from_model_functor_iso L W).symm _ +... ≅ W.Q ⋙ ((equivalence_from_model L W).functor ⋙ (equivalence_from_model L W).inverse) : + functor.associator _ _ _ +... ≅ W.Q ⋙ 𝟭 _ : iso_whisker_left _ ((equivalence_from_model L W).unit_iso.symm) +... ≅ W.Q : functor.right_unitor _ + +lemma ess_surj : ess_surj L := +⟨λ X, ⟨(construction.obj_equiv W).inv_fun ((equivalence_from_model L W).inverse.obj X), + nonempty.intro ((Q_comp_equivalence_from_model_functor_iso L W).symm.app _ ≪≫ + (equivalence_from_model L W).counit_iso.app X)⟩⟩ + +end localization + end category_theory diff --git a/src/category_theory/whiskering.lean b/src/category_theory/whiskering.lean index 3983f85813602..b50448b1a49db 100644 --- a/src/category_theory/whiskering.lean +++ b/src/category_theory/whiskering.lean @@ -201,6 +201,9 @@ and it's usually best to insert explicit associators.) { hom := { app := λ _, 𝟙 _ }, inv := { app := λ _, 𝟙 _ } } +@[protected] +lemma assoc (F : A ⥤ B) (G : B ⥤ C) (H : C ⥤ D) : ((F ⋙ G) ⋙ H) = (F ⋙ (G ⋙ H)) := rfl + lemma triangle (F : A ⥤ B) (G : B ⥤ C) : (associator F (𝟭 B) G).hom ≫ (whisker_left F (left_unitor G).hom) = (whisker_right (right_unitor F).hom G) := From 12edc98821c3d52388864279d27d01449bfca411 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Wed, 12 Oct 2022 10:35:39 +0000 Subject: [PATCH 716/828] feat(algebraic_geometry/prime_spectrum/maximal): maximal spectrum (#16905) --- .../prime_spectrum/maximal.lean | 76 +++++++++++++++++++ 1 file changed, 76 insertions(+) create mode 100644 src/algebraic_geometry/prime_spectrum/maximal.lean diff --git a/src/algebraic_geometry/prime_spectrum/maximal.lean b/src/algebraic_geometry/prime_spectrum/maximal.lean new file mode 100644 index 0000000000000..4b2ad313271b9 --- /dev/null +++ b/src/algebraic_geometry/prime_spectrum/maximal.lean @@ -0,0 +1,76 @@ +/- +Copyright (c) 2022 David Kurniadi Angdinata. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: David Kurniadi Angdinata +-/ + +import algebraic_geometry.prime_spectrum.basic + +/-! +# Maximal spectrum of a commutative ring + +The maximal spectrum of a commutative ring is the type of all maximal ideals. +It is naturally a subset of the prime spectrum endowed with the subspace topology. + +## Main definitions + +* `maximal_spectrum R`: The maximal spectrum of a commutative ring `R`, + i.e., the set of all maximal ideals of `R`. + +## Implementation notes + +The Zariski topology on the maximal spectrum is defined as the subspace topology induced by the +natural inclusion into the prime spectrum to avoid API duplication for zero loci. +-/ + +noncomputable theory +open_locale classical + +universes u v + +variables (R : Type u) [comm_ring R] + +/-- The maximal spectrum of a commutative ring `R` is the type of all maximal ideals of `R`. -/ +@[ext] structure maximal_spectrum := +(as_ideal : ideal R) +(is_maximal : as_ideal.is_maximal) + +attribute [instance] maximal_spectrum.is_maximal + +variable {R} + +namespace maximal_spectrum + +instance [nontrivial R] : nonempty $ maximal_spectrum R := +⟨⟨(ideal.exists_maximal R).some, (ideal.exists_maximal R).some_spec⟩⟩ + +/-- The natural inclusion from the maximal spectrum to the prime spectrum. -/ +def to_prime_spectrum (x : maximal_spectrum R) : prime_spectrum R := +⟨x.as_ideal, x.is_maximal.is_prime⟩ + +lemma to_prime_spectrum_injective : (@to_prime_spectrum R _).injective := +λ ⟨_, _⟩ ⟨_, _⟩ h, by simpa only [maximal_spectrum.mk.inj_eq] using subtype.mk.inj h + +open prime_spectrum set + +lemma to_prime_spectrum_range : + set.range (@to_prime_spectrum R _) = {x | is_closed ({x} : set $ prime_spectrum R)} := +begin + simp only [is_closed_singleton_iff_is_maximal], + ext ⟨x, _⟩, + exact ⟨λ ⟨y, hy⟩, hy ▸ y.is_maximal, λ hx, ⟨⟨x, hx⟩, rfl⟩⟩ +end + +/-- The Zariski topology on the maximal spectrum of a commutative ring is defined as the subspace +topology induced by the natural inclusion into the prime spectrum. -/ +instance zariski_topology : topological_space $ maximal_spectrum R := +prime_spectrum.zariski_topology.induced to_prime_spectrum + +instance : t1_space $ maximal_spectrum R := +⟨λ x, is_closed_induced_iff.mpr + ⟨{to_prime_spectrum x}, (is_closed_singleton_iff_is_maximal _).mpr x.is_maximal, + by simpa only [← image_singleton] using preimage_image_eq {x} to_prime_spectrum_injective⟩⟩ + +lemma to_prime_spectrum_continuous : continuous $ @to_prime_spectrum R _ := continuous_induced_dom + +end maximal_spectrum From b30449d60b503696fd87a8e92a778b10a4653153 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 12 Oct 2022 10:35:44 +0000 Subject: [PATCH 717/828] feat(geometry/euclidean/oriented_angle): signs of angles and `same_ray` (#16909) Add more lemmas about signs of angles in relation to the order of points on a line. --- src/geometry/euclidean/oriented_angle.lean | 130 +++++++++++++++++++++ 1 file changed, 130 insertions(+) diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 26d495b857624..2616f17c64193 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -2131,4 +2131,134 @@ lemma _root_.collinear.two_zsmul_oangle_eq_right {p₁ p₂ p₃ p₃' : P} (2 : ℤ) • ∡ p₁ p₂ p₃ = (2 : ℤ) • ∡ p₁ p₂ p₃' := by rw [oangle_rev, smul_neg, h.two_zsmul_oangle_eq_left hp₃p₂ hp₃'p₂, ←smul_neg, ←oangle_rev] +open affine_subspace + +/-- Given two pairs of distinct points on the same line, such that the vectors between those +pairs of points are on the same ray (oriented in the same direction on that line), and a fifth +point, the angles at the fifth point between each of those two pairs of points have the same +sign. -/ +lemma _root_.collinear.oangle_sign_of_same_ray_vsub {p₁ p₂ p₃ p₄ : P} (p₅ : P) (hp₁p₂ : p₁ ≠ p₂) + (hp₃p₄ : p₃ ≠ p₄) (hc : collinear ℝ ({p₁, p₂, p₃, p₄} : set P)) + (hr : same_ray ℝ (p₂ -ᵥ p₁) (p₄ -ᵥ p₃)) : (∡ p₁ p₅ p₂).sign = (∡ p₃ p₅ p₄).sign := +begin + by_cases hc₅₁₂ : collinear ℝ ({p₅, p₁, p₂} : set P), + { have hc₅₁₂₃₄ : collinear ℝ ({p₅, p₁, p₂, p₃, p₄} : set P) := + (hc.collinear_insert_iff_of_ne (set.mem_insert _ _) + (set.mem_insert_of_mem _ (set.mem_insert _ _)) hp₁p₂).2 hc₅₁₂, + have hc₅₃₄ : collinear ℝ ({p₅, p₃, p₄} : set P) := + (hc.collinear_insert_iff_of_ne + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_insert _ _))) + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ + (set.mem_singleton _)))) hp₃p₄).1 hc₅₁₂₃₄, + rw set.insert_comm at hc₅₁₂ hc₅₃₄, + have hs₁₅₂ := oangle_eq_zero_or_eq_pi_iff_collinear.2 hc₅₁₂, + have hs₃₅₄ := oangle_eq_zero_or_eq_pi_iff_collinear.2 hc₅₃₄, + rw ←real.angle.sign_eq_zero_iff at hs₁₅₂ hs₃₅₄, + rw [hs₁₅₂, hs₃₅₄] }, + { let s : set (P × P × P) := + (λ x : affine_span ℝ ({p₁, p₂} : set P) × V, (x.1, p₅, x.2 +ᵥ x.1)) '' + set.univ ×ˢ {v | same_ray ℝ (p₂ -ᵥ p₁) v ∧ v ≠ 0}, + have hco : is_connected s, + { haveI : connected_space (affine_span ℝ ({p₁, p₂} : set P)) := add_torsor.connected_space _ _, + exact (is_connected_univ.prod (is_connected_set_of_same_ray_and_ne_zero + (vsub_ne_zero.2 hp₁p₂.symm))).image _ + ((continuous_fst.subtype_coe.prod_mk + (continuous_const.prod_mk + (continuous_snd.vadd continuous_fst.subtype_coe))).continuous_on) }, + have hf : continuous_on (λ p : P × P × P, ∡ p.1 p.2.1 p.2.2) s, + { refine continuous_at.continuous_on (λ p hp, continuous_at_oangle _ _), + all_goals { simp_rw [s, set.mem_image, set.mem_prod, set.mem_univ, true_and, + prod.ext_iff] at hp, + obtain ⟨q₁, q₅, q₂⟩ := p, + dsimp only at ⊢ hp, + obtain ⟨⟨⟨q, hq⟩, v⟩, hv, rfl, rfl, rfl⟩ := hp, + dsimp only [subtype.coe_mk, set.mem_set_of] at ⊢ hv, + obtain ⟨hvr, -⟩ := hv, + rintro rfl, + refine hc₅₁₂ ((collinear_insert_iff_of_mem_affine_span _).2 + (collinear_pair _ _ _)) }, + { exact hq }, + { refine vadd_mem_of_mem_direction _ hq, + rw ←exists_nonneg_left_iff_same_ray (vsub_ne_zero.2 hp₁p₂.symm) at hvr, + obtain ⟨r, -, rfl⟩ := hvr, + rw direction_affine_span, + exact smul_vsub_rev_mem_vector_span_pair _ _ _ } }, + have hsp : ∀ p : P × P × P, p ∈ s → ∡ p.1 p.2.1 p.2.2 ≠ 0 ∧ ∡ p.1 p.2.1 p.2.2 ≠ π, + { intros p hp, + simp_rw [s, set.mem_image, set.mem_prod, set.mem_set_of, set.mem_univ, true_and, + prod.ext_iff] at hp, + obtain ⟨q₁, q₅, q₂⟩ := p, + dsimp only at ⊢ hp, + obtain ⟨⟨⟨q, hq⟩, v⟩, hv, rfl, rfl, rfl⟩ := hp, + dsimp only [subtype.coe_mk, set.mem_set_of] at ⊢ hv, + obtain ⟨hvr, hv0⟩ := hv, + rw ←exists_nonneg_left_iff_same_ray (vsub_ne_zero.2 hp₁p₂.symm) at hvr, + obtain ⟨r, -, rfl⟩ := hvr, + change q ∈ affine_span ℝ ({p₁, p₂} : set P) at hq, + rw [oangle_ne_zero_and_ne_pi_iff_affine_independent], + refine affine_independent_of_ne_of_mem_of_not_mem_of_mem _ hq + (λ h, hc₅₁₂ ((collinear_insert_iff_of_mem_affine_span h).2 (collinear_pair _ _ _))) _, + { rwa [←@vsub_ne_zero V, vsub_vadd_eq_vsub_sub, vsub_self, zero_sub, neg_ne_zero] }, + { refine vadd_mem_of_mem_direction _ hq, + rw direction_affine_span, + exact smul_vsub_rev_mem_vector_span_pair _ _ _ } }, + have hp₁p₂s : (p₁, p₅, p₂) ∈ s, + { simp_rw [s, set.mem_image, set.mem_prod, set.mem_set_of, set.mem_univ, true_and, + prod.ext_iff], + refine ⟨⟨⟨p₁, left_mem_affine_span_pair _ _ _⟩, p₂ -ᵥ p₁⟩, + ⟨same_ray.rfl, vsub_ne_zero.2 hp₁p₂.symm⟩, _⟩, + simp }, + have hp₃p₄s : (p₃, p₅, p₄) ∈ s, + { simp_rw [s, set.mem_image, set.mem_prod, set.mem_set_of, set.mem_univ, true_and, + prod.ext_iff], + refine ⟨⟨⟨p₃, + hc.mem_affine_span_of_mem_of_ne + (set.mem_insert _ _) + (set.mem_insert_of_mem _ (set.mem_insert _ _)) + (set.mem_insert_of_mem _ (set.mem_insert_of_mem _ (set.mem_insert _ _))) + hp₁p₂⟩, p₄ -ᵥ p₃⟩, ⟨hr, vsub_ne_zero.2 hp₃p₄.symm⟩, _⟩, + simp }, + convert real.angle.sign_eq_of_continuous_on hco hf hsp hp₃p₄s hp₁p₂s } +end + +/-- Given three points in strict order on the same line, and a fourth point, the angles at the +fourth point between the first and second or second and third points have the same sign. -/ +lemma _root_.sbtw.oangle_sign_eq {p₁ p₂ p₃ : P} (p₄ : P) (h : sbtw ℝ p₁ p₂ p₃) : + (∡ p₁ p₄ p₂).sign = (∡ p₂ p₄ p₃).sign := +begin + have hc : collinear ℝ ({p₁, p₂, p₂, p₃} : set P), { simpa using h.wbtw.collinear }, + exact hc.oangle_sign_of_same_ray_vsub _ h.left_ne h.ne_right h.wbtw.same_ray_vsub +end + +/-- Given three points in weak order on the same line, with the first not equal to the second, +and a fourth point, the angles at the fourth point between the first and second or first and +third points have the same sign. -/ +lemma _root_.wbtw.oangle_sign_eq_of_ne_left {p₁ p₂ p₃ : P} (p₄ : P) (h : wbtw ℝ p₁ p₂ p₃) + (hne : p₁ ≠ p₂) : (∡ p₁ p₄ p₂).sign = (∡ p₁ p₄ p₃).sign := +begin + have hc : collinear ℝ ({p₁, p₂, p₁, p₃} : set P), + { simpa [set.insert_comm p₂] using h.collinear }, + exact hc.oangle_sign_of_same_ray_vsub _ hne (h.left_ne_right_of_ne_left hne.symm) + h.same_ray_vsub_left +end + +/-- Given three points in strict order on the same line, and a fourth point, the angles at the +fourth point between the first and second or first and third points have the same sign. -/ +lemma _root_.sbtw.oangle_sign_eq_left {p₁ p₂ p₃ : P} (p₄ : P) (h : sbtw ℝ p₁ p₂ p₃) : + (∡ p₁ p₄ p₂).sign = (∡ p₁ p₄ p₃).sign := +h.wbtw.oangle_sign_eq_of_ne_left _ h.left_ne + +/-- Given three points in weak order on the same line, with the second not equal to the third, +and a fourth point, the angles at the fourth point between the second and third or first and +third points have the same sign. -/ +lemma _root_.wbtw.oangle_sign_eq_of_ne_right {p₁ p₂ p₃ : P} (p₄ : P) (h : wbtw ℝ p₁ p₂ p₃) + (hne : p₂ ≠ p₃) : (∡ p₂ p₄ p₃).sign = (∡ p₁ p₄ p₃).sign := +by simp_rw [oangle_rev p₃, real.angle.sign_neg, h.symm.oangle_sign_eq_of_ne_left _ hne.symm] + +/-- Given three points in strict order on the same line, and a fourth point, the angles at the +fourth point between the second and third or first and third points have the same sign. -/ +lemma _root_.sbtw.oangle_sign_eq_right {p₁ p₂ p₃ : P} (p₄ : P) (h : sbtw ℝ p₁ p₂ p₃) : + (∡ p₂ p₄ p₃).sign = (∡ p₁ p₄ p₃).sign := +h.wbtw.oangle_sign_eq_of_ne_right _ h.ne_right + end euclidean_geometry From 4d5491dc685e8a14f3ea4d747b0ee3a2ff769b6e Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 12 Oct 2022 10:35:53 +0000 Subject: [PATCH 718/828] feat(linear_algebra/affine_space/finite_dimensional): `coplanar` (#16910) Add a definition of coplanar sets of points, and some basic API similar to parts of that for `collinear`. --- .../affine_space/finite_dimensional.lean | 80 +++++++++++++++++++ 1 file changed, 80 insertions(+) diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index 35c7a8d8a803d..0b7fc15c9e2f3 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -499,6 +499,69 @@ lemma collinear_insert_iff_of_mem_affine_span {s : set P} {p : P} (h : p ∈ aff collinear k (insert p s) ↔ collinear k s := by rw [collinear, collinear, vector_span_insert_eq_vector_span h] +variables (k) + +/-- A set of points is coplanar if their `vector_span` has dimension at most `2`. -/ +def coplanar (s : set P) : Prop := module.rank k (vector_span k s) ≤ 2 + +variables {k} + +/-- The `vector_span` of coplanar points is finite-dimensional. -/ +lemma coplanar.finite_dimensional_vector_span {s : set P} (h : coplanar k s) : + finite_dimensional k (vector_span k s) := +begin + refine is_noetherian.iff_fg.1 (is_noetherian.iff_dim_lt_aleph_0.2 (lt_of_le_of_lt h _)), + simp, +end + +/-- The direction of the affine span of coplanar points is finite-dimensional. -/ +lemma coplanar.finite_dimensional_direction_affine_span {s : set P} (h : coplanar k s) : + finite_dimensional k (affine_span k s).direction := +(direction_affine_span k s).symm ▸ h.finite_dimensional_vector_span + +/-- A set of points, whose `vector_span` is finite-dimensional, is coplanar if and only if their +`vector_span` has dimension at most `2`. -/ +lemma coplanar_iff_finrank_le_two {s : set P} [finite_dimensional k (vector_span k s)] : + coplanar k s ↔ finrank k (vector_span k s) ≤ 2 := +begin + have h : coplanar k s ↔ module.rank k (vector_span k s) ≤ 2 := iff.rfl, + rw ←finrank_eq_dim at h, + exact_mod_cast h +end + +alias coplanar_iff_finrank_le_two ↔ coplanar.finrank_le_two _ + +/-- A subset of a coplanar set is coplanar. -/ +lemma coplanar.subset {s₁ s₂ : set P} (hs : s₁ ⊆ s₂) (h : coplanar k s₂) : coplanar k s₁ := +(dim_le_of_submodule (vector_span k s₁) (vector_span k s₂) (vector_span_mono k hs)).trans h + +/-- Collinear points are coplanar. -/ +lemma collinear.coplanar {s : set P} (h : collinear k s) : coplanar k s := +le_trans h one_le_two + +variables (k) (P) + +/-- The empty set is coplanar. -/ +lemma coplanar_empty : coplanar k (∅ : set P) := +(collinear_empty k P).coplanar + +variables {P} + +/-- A single point is coplanar. -/ +lemma coplanar_singleton (p : P) : coplanar k ({p} : set P) := +(collinear_singleton k p).coplanar + +/-- Two points are coplanar. -/ +lemma coplanar_pair (p₁ p₂ : P) : coplanar k ({p₁, p₂} : set P) := +(collinear_pair k p₁ p₂).coplanar + +variables {k} + +/-- Adding a point in the affine span of a set does not change whether that set is coplanar. -/ +lemma coplanar_insert_iff_of_mem_affine_span {s : set P} {p : P} (h : p ∈ affine_span k s) : + coplanar k (insert p s) ↔ coplanar k s := +by rw [coplanar, coplanar, vector_span_insert_eq_vector_span h] + end affine_space' section field @@ -556,4 +619,21 @@ begin rw direction_affine_span end +variables {k} + +/-- Adding a point to a collinear set produces a coplanar set. -/ +lemma collinear.coplanar_insert {s : set P} (h : collinear k s) (p : P) : + coplanar k (insert p s) := +begin + haveI := h.finite_dimensional_vector_span, + rw [coplanar_iff_finrank_le_two], + exact (finrank_vector_span_insert_le_set k s p).trans (add_le_add_right h.finrank_le_one _) +end + +variables (k) + +/-- Three points are coplanar. -/ +lemma coplanar_triple (p₁ p₂ p₃ : P) : coplanar k ({p₁, p₂, p₃} : set P) := +(collinear_pair k p₂ p₃).coplanar_insert p₁ + end field From 68ead6e664f4a86b2405055689683ebdcf44c177 Mon Sep 17 00:00:00 2001 From: Alexander Bentkamp Date: Wed, 12 Oct 2022 13:05:02 +0000 Subject: [PATCH 719/828] feat(linear_algebra/matrix): Schur complement (#15270) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit If a matrix `A` is positive definite, then `[A B; Bᴴ D]` is postive semidefinite if and only if `D - Bᴴ A⁻¹ B` is postive semidefinite. --- src/algebra/star/pi.lean | 4 + src/linear_algebra/matrix/hermitian.lean | 46 ++++++- .../matrix/nonsingular_inverse.lean | 6 + src/linear_algebra/matrix/pos_def.lean | 54 ++++++-- .../matrix/schur_complement.lean | 125 ++++++++++++++++++ 5 files changed, 223 insertions(+), 12 deletions(-) create mode 100644 src/linear_algebra/matrix/schur_complement.lean diff --git a/src/algebra/star/pi.lean b/src/algebra/star/pi.lean index 3f4a1e23184bd..727fe48d4c1f3 100644 --- a/src/algebra/star/pi.lean +++ b/src/algebra/star/pi.lean @@ -58,4 +58,8 @@ lemma update_star [Π i, has_star (f i)] [decidable_eq I] function.update (star h) i (star a) = star (function.update h i a) := funext $ λ j, (apply_update (λ i, star) h i a j).symm +lemma star_sum_elim {I J α : Type*} (x : I → α) (y : J → α) [has_star α] : + star (sum.elim x y) = sum.elim (star x) (star y) := +by { ext x, cases x; simp } + end function diff --git a/src/linear_algebra/matrix/hermitian.lean b/src/linear_algebra/matrix/hermitian.lean index fbf01cba1c774..51fa8ee4ceae6 100644 --- a/src/linear_algebra/matrix/hermitian.lean +++ b/src/linear_algebra/matrix/hermitian.lean @@ -55,6 +55,16 @@ lemma is_hermitian_transpose_mul_self [fintype n] (A : matrix n n α) : (Aᴴ ⬝ A).is_hermitian := by rw [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose] +lemma is_hermitian_conj_transpose_mul_mul [fintype m] {A : matrix m m α} (B : matrix m n α) + (hA : A.is_hermitian) : (Bᴴ ⬝ A ⬝ B).is_hermitian := +by simp only [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose, hA.eq, + matrix.mul_assoc] + +lemma is_hermitian_mul_mul_conj_transpose [fintype m] {A : matrix m m α} (B : matrix n m α) + (hA : A.is_hermitian) : (B ⬝ A ⬝ Bᴴ).is_hermitian := +by simp only [is_hermitian, conj_transpose_mul, conj_transpose_conj_transpose, hA.eq, + matrix.mul_assoc] + lemma is_hermitian_add_transpose_self (A : matrix n n α) : (A + Aᴴ).is_hermitian := by simp [is_hermitian, add_comm] @@ -72,14 +82,24 @@ conj_transpose_zero (A.map f).is_hermitian := (conj_transpose_map f hf).symm.trans $ h.eq.symm ▸ rfl -@[simp] lemma is_hermitian.transpose {A : matrix n n α} (h : A.is_hermitian) : +lemma is_hermitian.transpose {A : matrix n n α} (h : A.is_hermitian) : Aᵀ.is_hermitian := by { rw [is_hermitian, conj_transpose, transpose_map], congr, exact h } -@[simp] lemma is_hermitian.conj_transpose {A : matrix n n α} (h : A.is_hermitian) : +@[simp] lemma is_hermitian_transpose_iff (A : matrix n n α) : + Aᵀ.is_hermitian ↔ A.is_hermitian := +⟨by { intro h, rw [← transpose_transpose A], exact is_hermitian.transpose h }, + is_hermitian.transpose⟩ + +lemma is_hermitian.conj_transpose {A : matrix n n α} (h : A.is_hermitian) : Aᴴ.is_hermitian := h.transpose.map _ $ λ _, rfl +@[simp] lemma is_hermitian_conj_transpose_iff (A : matrix n n α) : + Aᴴ.is_hermitian ↔ A.is_hermitian := +⟨by { intro h, rw [← conj_transpose_conj_transpose A], exact is_hermitian.conj_transpose h }, + is_hermitian.conj_transpose⟩ + @[simp] lemma is_hermitian.add {A B : matrix n n α} (hA : A.is_hermitian) (hB : B.is_hermitian) : (A + B).is_hermitian := (conj_transpose_add _ _).trans (hA.symm ▸ hB.symm ▸ rfl) @@ -88,6 +108,10 @@ h.transpose.map _ $ λ _, rfl (A.submatrix f f).is_hermitian := (conj_transpose_submatrix _ _ _).trans (h.symm ▸ rfl) +@[simp] lemma is_hermitian_submatrix_equiv {A : matrix n n α} (e : m ≃ n) : + (A.submatrix e e).is_hermitian ↔ A.is_hermitian := +⟨λ h, by simpa using h.submatrix e.symm, λ h, h.submatrix _⟩ + /-- The real diagonal matrix `diagonal v` is hermitian. -/ @[simp] lemma is_hermitian_diagonal [decidable_eq n] (v : n → ℝ) : (diagonal v).is_hermitian := @@ -141,6 +165,24 @@ variables [ring α] [star_ring α] [ring β] [star_ring β] end ring +section comm_ring + +variables [comm_ring α] [star_ring α] + +lemma is_hermitian.inv [fintype m] [decidable_eq m] {A : matrix m m α} + (hA : A.is_hermitian) : A⁻¹.is_hermitian := +by simp [is_hermitian, conj_transpose_nonsing_inv, hA.eq] + +@[simp] lemma is_hermitian_inv [fintype m] [decidable_eq m] (A : matrix m m α) [invertible A]: + (A⁻¹).is_hermitian ↔ A.is_hermitian := +⟨λ h, by {rw [← inv_inv_of_invertible A], exact is_hermitian.inv h }, is_hermitian.inv⟩ + +lemma is_hermitian.adjugate [fintype m] [decidable_eq m] {A : matrix m m α} + (hA : A.is_hermitian) : A.adjugate.is_hermitian := +by simp [is_hermitian, adjugate_conj_transpose, hA.eq] + +end comm_ring + section is_R_or_C open is_R_or_C diff --git a/src/linear_algebra/matrix/nonsingular_inverse.lean b/src/linear_algebra/matrix/nonsingular_inverse.lean index 5375985da2046..a8c4d8aeb5952 100644 --- a/src/linear_algebra/matrix/nonsingular_inverse.lean +++ b/src/linear_algebra/matrix/nonsingular_inverse.lean @@ -275,6 +275,12 @@ begin rw [←inv_of_eq_nonsing_inv, matrix.inv_of_mul_self], end +instance [invertible A] : invertible A⁻¹ := +by { rw ← inv_of_eq_nonsing_inv, apply_instance } + +@[simp] lemma inv_inv_of_invertible [invertible A] : A⁻¹⁻¹ = A := +by simp only [← inv_of_eq_nonsing_inv, inv_of_inv_of] + @[simp] lemma mul_nonsing_inv_cancel_right (B : matrix m n α) (h : is_unit A.det) : B ⬝ A ⬝ A⁻¹ = B := by simp [matrix.mul_assoc, mul_nonsing_inv A h] diff --git a/src/linear_algebra/matrix/pos_def.lean b/src/linear_algebra/matrix/pos_def.lean index 80149bb427397..aa243db83e8a8 100644 --- a/src/linear_algebra/matrix/pos_def.lean +++ b/src/linear_algebra/matrix/pos_def.lean @@ -7,30 +7,64 @@ import linear_algebra.matrix.spectrum import linear_algebra.quadratic_form.basic /-! # Positive Definite Matrices - -This file defines positive definite matrices and connects this notion to positive definiteness of -quadratic forms. - +This file defines positive (semi)definite matrices and connects the notion to positive definiteness +of quadratic forms. ## Main definition - - * `matrix.pos_def` : a matrix `M : matrix n n R` is positive definite if it is hermitian - and `xᴴMx` is greater than zero for all nonzero `x`. - + * `matrix.pos_def` : a matrix `M : matrix n n 𝕜` is positive definite if it is hermitian and `xᴴMx` + is greater than zero for all nonzero `x`. + * `matrix.pos_semidef` : a matrix `M : matrix n n 𝕜` is positive semidefinite if it is hermitian + and `xᴴMx` is nonnegative for all `x`. -/ namespace matrix -variables {𝕜 : Type*} [is_R_or_C 𝕜] {n : Type*} [fintype n] +variables {𝕜 : Type*} [is_R_or_C 𝕜] {m n : Type*} [fintype m] [fintype n] open_locale matrix -/-- A matrix `M : matrix n n R` is positive definite if it is hermitian +/-- A matrix `M : matrix n n 𝕜` is positive definite if it is hermitian and `xᴴMx` is greater than zero for all nonzero `x`. -/ def pos_def (M : matrix n n 𝕜) := M.is_hermitian ∧ ∀ x : n → 𝕜, x ≠ 0 → 0 < is_R_or_C.re (dot_product (star x) (M.mul_vec x)) lemma pos_def.is_hermitian {M : matrix n n 𝕜} (hM : M.pos_def) : M.is_hermitian := hM.1 +/-- A matrix `M : matrix n n 𝕜` is positive semidefinite if it is hermitian + and `xᴴMx` is nonnegative for all `x`. -/ +def pos_semidef (M : matrix n n 𝕜) := +M.is_hermitian ∧ ∀ x : n → 𝕜, 0 ≤ is_R_or_C.re (dot_product (star x) (M.mul_vec x)) + +lemma pos_def.pos_semidef {M : matrix n n 𝕜} (hM : M.pos_def) : M.pos_semidef := +begin + refine ⟨hM.1, _⟩, + intros x, + by_cases hx : x = 0, + { simp only [hx, zero_dot_product, star_zero, is_R_or_C.zero_re'] }, + { exact le_of_lt (hM.2 x hx) } +end + +lemma pos_semidef.submatrix {M : matrix n n 𝕜} (hM : M.pos_semidef) (e : m ≃ n): + (M.submatrix e e).pos_semidef := +begin + refine ⟨hM.1.submatrix e, λ x, _⟩, + have : (M.submatrix ⇑e ⇑e).mul_vec x = M.mul_vec (λ (i : n), x (e.symm i)) ∘ e, + { ext i, + dsimp only [(∘), mul_vec, dot_product], + rw finset.sum_bij' (λ i _, e i) _ _ (λ i _, e.symm i); + simp only [eq_self_iff_true, implies_true_iff, equiv.symm_apply_apply, finset.mem_univ, + submatrix_apply, equiv.apply_symm_apply] }, + rw this, + convert hM.2 (λ i, x (e.symm i)) using 3, + unfold dot_product, + rw [finset.sum_bij' (λ i _, e i) _ _ (λ i _, e.symm i)]; + simp only [eq_self_iff_true, implies_true_iff, equiv.symm_apply_apply, finset.mem_univ, + submatrix_apply, equiv.apply_symm_apply, pi.star_apply], +end + +@[simp] lemma pos_semidef_submatrix_equiv {M : matrix n n 𝕜} (e : m ≃ n) : + (M.submatrix e e).pos_semidef ↔ M.pos_semidef := +⟨λ h, by simpa using h.submatrix e.symm, λ h, h.submatrix _⟩ + lemma pos_def.transpose {M : matrix n n 𝕜} (hM : M.pos_def) : Mᵀ.pos_def := begin refine ⟨is_hermitian.transpose hM.1, λ x hx, _⟩, diff --git a/src/linear_algebra/matrix/schur_complement.lean b/src/linear_algebra/matrix/schur_complement.lean new file mode 100644 index 0000000000000..99de6ff74316e --- /dev/null +++ b/src/linear_algebra/matrix/schur_complement.lean @@ -0,0 +1,125 @@ +/- +Copyright (c) 2022 Alexander Bentkamp. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alexander Bentkamp, Jeremy Avigad, Johan Commelin +-/ +import linear_algebra.matrix.symmetric +import linear_algebra.matrix.nonsingular_inverse +import linear_algebra.matrix.pos_def + +/-! # Schur complement + +This file proves properties of the Schur complement `D - C A⁻¹ B` of a block matrix `[A B; C D]`. + +The determinant of a block matrix in terms of the Schur complement is expressed in the lemmas +`matrix.det_from_blocks₁₁` and `matrix.det_from_blocks₂₂` in the file +`linear_algebra.matrix.nonsingular_inverse`. + +## Main result + + * `matrix.schur_complement_pos_semidef_iff` : If a matrix `A` is positive definite, then `[A B; Bᴴ + D]` is postive semidefinite if and only if `D - Bᴴ A⁻¹ B` is postive semidefinite. + +-/ + +namespace matrix + +open_locale matrix +variables {n : Type*} {m : Type*} {𝕜 : Type*} [is_R_or_C 𝕜] + +localized "infix ` ⊕ᵥ `:65 := sum.elim" in matrix + +lemma schur_complement_eq₁₁ [fintype m] [decidable_eq m] [fintype n] + {A : matrix m m 𝕜} (B : matrix m n 𝕜) (D : matrix n n 𝕜) (x : m → 𝕜) (y : n → 𝕜) + [invertible A] (hA : A.is_hermitian) : +vec_mul (star (x ⊕ᵥ y)) (from_blocks A B Bᴴ D) ⬝ᵥ (x ⊕ᵥ y) = + vec_mul (star (x + (A⁻¹ ⬝ B).mul_vec y)) A ⬝ᵥ (x + (A⁻¹ ⬝ B).mul_vec y) + + vec_mul (star y) (D - Bᴴ ⬝ A⁻¹ ⬝ B) ⬝ᵥ y := +begin + simp [function.star_sum_elim, from_blocks_mul_vec, vec_mul_from_blocks, add_vec_mul, + dot_product_mul_vec, vec_mul_sub, matrix.mul_assoc, vec_mul_mul_vec, hA.eq, + conj_transpose_nonsing_inv, star_mul_vec], + abel +end + +lemma schur_complement_eq₂₂ [fintype m] [fintype n] [decidable_eq n] + (A : matrix m m 𝕜) (B : matrix m n 𝕜) {D : matrix n n 𝕜} (x : m → 𝕜) (y : n → 𝕜) + [invertible D] (hD : D.is_hermitian) : +vec_mul (star (x ⊕ᵥ y)) (from_blocks A B Bᴴ D) ⬝ᵥ (x ⊕ᵥ y) = + vec_mul (star ((D⁻¹ ⬝ Bᴴ).mul_vec x + y)) D ⬝ᵥ ((D⁻¹ ⬝ Bᴴ).mul_vec x + y) + + vec_mul (star x) (A - B ⬝ D⁻¹ ⬝ Bᴴ) ⬝ᵥ x := +begin + simp [function.star_sum_elim, from_blocks_mul_vec, vec_mul_from_blocks, add_vec_mul, + dot_product_mul_vec, vec_mul_sub, matrix.mul_assoc, vec_mul_mul_vec, hD.eq, + conj_transpose_nonsing_inv, star_mul_vec], + abel +end + +end matrix + +namespace matrix + +open_locale matrix +variables {n : Type*} {m : Type*} + {𝕜 : Type*} [is_R_or_C 𝕜] + +lemma is_hermitian.from_blocks₁₁ [fintype m] [decidable_eq m] + {A : matrix m m 𝕜} (B : matrix m n 𝕜) (D : matrix n n 𝕜) + (hA : A.is_hermitian) : + (from_blocks A B Bᴴ D).is_hermitian ↔ (D - Bᴴ ⬝ A⁻¹ ⬝ B).is_hermitian := +begin + have hBAB : (Bᴴ ⬝ A⁻¹ ⬝ B).is_hermitian, + { apply is_hermitian_conj_transpose_mul_mul, + apply hA.inv }, + rw [is_hermitian_from_blocks_iff], + split, + { intro h, + apply is_hermitian.sub h.2.2.2 hBAB }, + { intro h, + refine ⟨hA, rfl, conj_transpose_conj_transpose B, _⟩, + rw ← sub_add_cancel D, + apply is_hermitian.add h hBAB } +end + +lemma is_hermitian.from_blocks₂₂ [fintype n] [decidable_eq n] + (A : matrix m m 𝕜) (B : matrix m n 𝕜) {D : matrix n n 𝕜} + (hD : D.is_hermitian) : + (from_blocks A B Bᴴ D).is_hermitian ↔ (A - B ⬝ D⁻¹ ⬝ Bᴴ).is_hermitian := +begin + rw [←is_hermitian_submatrix_equiv (equiv.sum_comm n m), equiv.sum_comm_apply, + from_blocks_submatrix_sum_swap_sum_swap], + convert is_hermitian.from_blocks₁₁ _ _ hD; simp +end + +lemma pos_semidef.from_blocks₁₁ [fintype m] [decidable_eq m] [fintype n] + {A : matrix m m 𝕜} (B : matrix m n 𝕜) (D : matrix n n 𝕜) + (hA : A.pos_def) [invertible A] : + (from_blocks A B Bᴴ D).pos_semidef ↔ (D - Bᴴ ⬝ A⁻¹ ⬝ B).pos_semidef := +begin + rw [pos_semidef, is_hermitian.from_blocks₁₁ _ _ hA.1], + split, + { refine λ h, ⟨h.1, λ x, _⟩, + have := h.2 (- ((A⁻¹ ⬝ B).mul_vec x) ⊕ᵥ x), + rw [dot_product_mul_vec, schur_complement_eq₁₁ B D _ _ hA.1, neg_add_self, + dot_product_zero, zero_add] at this, + rw [dot_product_mul_vec], exact this }, + { refine λ h, ⟨h.1, λ x, _⟩, + rw [dot_product_mul_vec, ← sum.elim_comp_inl_inr x, schur_complement_eq₁₁ B D _ _ hA.1, + map_add], + apply le_add_of_nonneg_of_le, + { rw ← dot_product_mul_vec, + apply hA.pos_semidef.2, }, + { rw ← dot_product_mul_vec, apply h.2 } } +end + +lemma pos_semidef.from_blocks₂₂ [fintype m] [fintype n] [decidable_eq n] + (A : matrix m m 𝕜) (B : matrix m n 𝕜) {D : matrix n n 𝕜} + (hD : D.pos_def) [invertible D] : + (from_blocks A B Bᴴ D).pos_semidef ↔ (A - B ⬝ D⁻¹ ⬝ Bᴴ).pos_semidef := +begin + rw [←pos_semidef_submatrix_equiv (equiv.sum_comm n m), equiv.sum_comm_apply, + from_blocks_submatrix_sum_swap_sum_swap], + convert pos_semidef.from_blocks₁₁ _ _ hD; apply_instance <|> simp +end + +end matrix From 6e4c1b3115f77493ce87737927f6c0fc9fa5c611 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Wed, 12 Oct 2022 13:05:03 +0000 Subject: [PATCH 720/828] feat(order/well_founded, data/fin/tuple/*, ...): add `well_founded.induction_bot`, `tuple.bubble_sort_induction` (#16536) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main purpose of this PR is to add the following induction principle fir tuples. ```lean lemma tuple.bubble_sort_induction {n : ℕ} {α : Type*} [linear_order α] {f : fin n → α} {P : (fin n → α) → Prop} (hf : P f) (h : ∀ (g : fin n → α) (i j : fin n), i < j → g j < g i → P g → P (g ∘ equiv.swap i j)) : P (f ∘ sort f) ``` This is in a new file `data/fin/tuple/bubble_sort_induction`. The additional API lemmas needed for it are mostly added to `data/fin/tuple/sort` (and some in `data/list/perm` and `data/list/sort`). One of these lemmas is the following induction principle for well-founded relations. ```lean lemma well_founded.induction_bot {α} {r : α → α → Prop} (hwf : well_founded r) {a bot : α} {C : α → Prop} (ha : C a) (ih : ∀ b, b ≠ bot → C b → ∃ c, r c b ∧ C c) : C bot ``` See the discussion on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/.22sort.22.20a.20function/near/299289703). --- src/data/fin/tuple/bubble_sort_induction.lean | 53 ++++++++++++++ src/data/fin/tuple/sort.lean | 73 ++++++++++++++++++- src/data/list/perm.lean | 18 +++++ src/data/list/sort.lean | 17 +++++ src/data/pi/lex.lean | 8 ++ src/order/well_founded.lean | 48 +++++++++++- 6 files changed, 215 insertions(+), 2 deletions(-) create mode 100644 src/data/fin/tuple/bubble_sort_induction.lean diff --git a/src/data/fin/tuple/bubble_sort_induction.lean b/src/data/fin/tuple/bubble_sort_induction.lean new file mode 100644 index 0000000000000..a6d199e700af3 --- /dev/null +++ b/src/data/fin/tuple/bubble_sort_induction.lean @@ -0,0 +1,53 @@ +/- +Copyright (c) 2022 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import data.fin.tuple.sort +import order.well_founded + +/-! +# "Bubble sort" induction + +We implement the following induction principle `tuple.bubble_sort_induction` +on tuples with values in a linear order `α`. + +Let `f : fin n → α` and let `P` be a predicate on `fin n → α`. Then we can show that +`f ∘ sort f` satisfies `P` if `f` satisfies `P`, and whenever some `g : fin n → α` +satisfies `P` and `g i > g j` for some `i < j`, then `g ∘ swap i j` also satisfies `P`. + +We deduce it from a stronger variant `tuple.bubble_sort_induction'`, which +requires the assumption only for `g` that are permutations of `f`. + +The latter is proved by well-founded induction via `well_founded.induction_bot'` +with respect to the lexicographic ordering on the finite set of all permutations of `f`. +-/ + +namespace tuple + +/-- *Bubble sort induction*: Prove that the sorted version of `f` has some property `P` +if `f` satsifies `P` and `P` is preserved on permutations of `f` when swapping two +antitone values. -/ +lemma bubble_sort_induction' {n : ℕ} {α : Type*} [linear_order α] {f : fin n → α} + {P : (fin n → α) → Prop} (hf : P f) + (h : ∀ (σ : equiv.perm (fin n)) (i j : fin n), + i < j → (f ∘ σ) j < (f ∘ σ) i → P (f ∘ σ) → P (f ∘ σ ∘ equiv.swap i j)) : + P (f ∘ sort f) := +begin + letI := @preorder.lift _ (lex (fin n → α)) _ (λ σ : equiv.perm (fin n), to_lex (f ∘ σ)), + refine @well_founded.induction_bot' _ _ _ + (@finite.preorder.well_founded_lt (equiv.perm (fin n)) _ _) + (equiv.refl _) (sort f) P (λ σ, f ∘ σ) (λ σ hσ hfσ, _) hf, + obtain ⟨i, j, hij₁, hij₂⟩ := antitone_pair_of_not_sorted' hσ, + exact ⟨σ * equiv.swap i j, pi.lex_desc hij₁ hij₂, h σ i j hij₁ hij₂ hfσ⟩, +end + +/-- *Bubble sort induction*: Prove that the sorted version of `f` has some property `P` +if `f` satsifies `P` and `P` is preserved when swapping two antitone values. -/ +lemma bubble_sort_induction {n : ℕ} {α : Type*} [linear_order α] {f : fin n → α} + {P : (fin n → α) → Prop} (hf : P f) + (h : ∀ (g : fin n → α) (i j : fin n), i < j → g j < g i → P g → P (g ∘ equiv.swap i j)) : + P (f ∘ sort f) := +bubble_sort_induction' hf (λ σ, h _) + +end tuple diff --git a/src/data/fin/tuple/sort.lean b/src/data/fin/tuple/sort.lean index 9b3dced457ac5..7fa36495cbd35 100644 --- a/src/data/fin/tuple/sort.lean +++ b/src/data/fin/tuple/sort.lean @@ -72,11 +72,14 @@ finset.order_iso_of_fin _ (by simp) def sort (f : fin n → α) : equiv.perm (fin n) := (graph_equiv₂ f).to_equiv.trans (graph_equiv₁ f).symm +lemma graph_equiv₂_apply (f : fin n → α) (i : fin n) : + graph_equiv₂ f i = graph_equiv₁ f (sort f i) := +((graph_equiv₁ f).apply_symm_apply _).symm + lemma self_comp_sort (f : fin n → α) : f ∘ sort f = graph.proj ∘ graph_equiv₂ f := show graph.proj ∘ ((graph_equiv₁ f) ∘ (graph_equiv₁ f).symm) ∘ (graph_equiv₂ f).to_equiv = _, by simp - lemma monotone_proj (f : fin n → α) : monotone (graph.proj : graph f → α) := begin rintro ⟨⟨x, i⟩, hx⟩ ⟨⟨y, j⟩, hy⟩ (h|h), @@ -91,3 +94,71 @@ begin end end tuple + +namespace tuple + +open list + +variables {n : ℕ} {α : Type*} + +/-- If two permutations of a tuple `f` are both monotone, then they are equal. -/ +lemma unique_monotone [partial_order α] {f : fin n → α} {σ τ : equiv.perm (fin n)} + (hfσ : monotone (f ∘ σ)) (hfτ : monotone (f ∘ τ)) : f ∘ σ = f ∘ τ := +of_fn_injective $ eq_of_perm_of_sorted + ((σ.of_fn_comp_perm f).trans (τ.of_fn_comp_perm f).symm) hfσ.of_fn_sorted hfτ.of_fn_sorted + +variables [linear_order α] {f : fin n → α} {σ : equiv.perm (fin n)} + +/-- A permutation `σ` equals `sort f` if and only if the map `i ↦ (f (σ i), σ i)` is +strictly monotone (w.r.t. the lexicographic ordering on the target). -/ +lemma eq_sort_iff' : σ = sort f ↔ strict_mono (σ.trans $ graph_equiv₁ f) := +begin + split; intro h, + { rw [h, sort, equiv.trans_assoc, equiv.symm_trans_self], exact (graph_equiv₂ f).strict_mono }, + { have := subsingleton.elim (graph_equiv₂ f) (h.order_iso_of_surjective _ $ equiv.surjective _), + ext1, exact (graph_equiv₁ f).apply_eq_iff_eq_symm_apply.1 (fun_like.congr_fun this x).symm }, +end + +/-- A permutation `σ` equals `sort f` if and only if `f ∘ σ` is monotone and whenever `i < j` +and `f (σ i) = f (σ j)`, then `σ i < σ j`. This means that `sort f` is the lexicographically +smallest permutation `σ` such that `f ∘ σ` is monotone. -/ +lemma eq_sort_iff : σ = sort f ↔ monotone (f ∘ σ) ∧ ∀ i j, i < j → f (σ i) = f (σ j) → σ i < σ j := +begin + rw eq_sort_iff', + refine ⟨λ h, ⟨(monotone_proj f).comp h.monotone, λ i j hij hfij, _⟩, λ h i j hij, _⟩, + { exact (((prod.lex.lt_iff _ _).1 $ h hij).resolve_left hfij.not_lt).2 }, + { obtain he|hl := (h.1 hij.le).eq_or_lt; apply (prod.lex.lt_iff _ _).2, + exacts [or.inr ⟨he, h.2 i j hij he⟩, or.inl hl] }, +end + +/-- The permutation that sorts `f` is the identity if and only if `f` is monotone. -/ +lemma sort_eq_refl_iff_monotone : sort f = equiv.refl _ ↔ monotone f := +begin + rw [eq_comm, eq_sort_iff, equiv.coe_refl, function.comp.right_id], + simp only [id.def, and_iff_left_iff_imp], + exact λ _ _ _ hij _, hij, +end + +/-- A permutation of a tuple `f` is `f` sorted if and only if it is monotone. -/ +lemma comp_sort_eq_comp_iff_monotone : f ∘ σ = f ∘ sort f ↔ monotone (f ∘ σ) := +⟨λ h, h.symm ▸ monotone_sort f, λ h, unique_monotone h (monotone_sort f)⟩ + +/-- The sorted versions of a tuple `f` and of any permutation of `f` agree. -/ +lemma comp_perm_comp_sort_eq_comp_sort : (f ∘ σ) ∘ (sort (f ∘ σ)) = f ∘ sort f := +begin + rw [function.comp.assoc, ← equiv.perm.coe_mul], + exact unique_monotone (monotone_sort (f ∘ σ)) (monotone_sort f), +end + +/-- If a permutation `f ∘ σ` of the tuple `f` is not the same as `f ∘ sort f`, then `f ∘ σ` +has a pair of strictly decreasing entries. -/ +lemma antitone_pair_of_not_sorted' (h : f ∘ σ ≠ f ∘ sort f) : + ∃ i j, i < j ∧ (f ∘ σ) j < (f ∘ σ) i := +by { contrapose! h, exact comp_sort_eq_comp_iff_monotone.mpr (monotone_iff_forall_lt.mpr h) } + +/-- If the tuple `f` is not the same as `f ∘ sort f`, then `f` has a pair of strictly decreasing +entries. -/ +lemma antitone_pair_of_not_sorted (h : f ≠ f ∘ sort f) : ∃ i j, i < j ∧ f j < f i := +antitone_pair_of_not_sorted' (id h : f ∘ equiv.refl _ ≠ _) + +end tuple diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index b0053623b7fa7..1aaddb4d71628 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -1393,3 +1393,21 @@ end end permutations end list + +open list + +lemma equiv.perm.map_fin_range_perm {n : ℕ} (σ : equiv.perm (fin n)) : + map σ (fin_range n) ~ fin_range n := +begin + rw [perm_ext ((nodup_fin_range n).map σ.injective) $ nodup_fin_range n], + simpa only [mem_map, mem_fin_range, true_and, iff_true] using σ.surjective +end + +/-- The list obtained from a permutation of a tuple `f` is permutation equivalent to +the list obtained from `f`. -/ +lemma equiv.perm.of_fn_comp_perm {n : ℕ} {α : Type uu} (σ : equiv.perm (fin n)) (f : fin n → α) : + of_fn (f ∘ σ) ~ of_fn f := +begin + rw [of_fn_eq_map, of_fn_eq_map, ←map_map], + exact σ.map_fin_range_perm.map f, +end diff --git a/src/data/list/sort.lean b/src/data/list/sort.lean index 4a5e7e1ab8070..92ba2fed19d55 100644 --- a/src/data/list/sort.lean +++ b/src/data/list/sort.lean @@ -104,6 +104,23 @@ end end sorted +section monotone + +variables {n : ℕ} {α : Type uu} [preorder α] {f : fin n → α} + +/-- A tuple is monotone if and only if the list obtained from it is sorted. -/ +lemma monotone_iff_of_fn_sorted : monotone f ↔ (of_fn f).sorted (≤) := +begin + simp_rw [sorted, pairwise_iff_nth_le, length_of_fn, nth_le_of_fn', monotone_iff_forall_lt], + exact ⟨λ h i j hj hij, h $ fin.mk_lt_mk.mpr hij, λ h ⟨i, _⟩ ⟨j, hj⟩ hij, h i j hj hij⟩, +end + +/-- The list obtained from a monotone tuple is sorted. -/ +lemma monotone.of_fn_sorted (h : monotone f) : (of_fn f).sorted (≤) := +monotone_iff_of_fn_sorted.1 h + +end monotone + section sort variables {α : Type uu} (r : α → α → Prop) [decidable_rel r] local infix ` ≼ ` : 50 := r diff --git a/src/data/pi/lex.lean b/src/data/pi/lex.lean index 46b225975a289..c19d46ac522f7 100644 --- a/src/data/pi/lex.lean +++ b/src/data/pi/lex.lean @@ -127,4 +127,12 @@ instance lex.ordered_comm_group [linear_order ι] [∀ a, ordered_comm_group (β ..pi.lex.partial_order, ..pi.comm_group } +/-- If we swap two strictly decreasing values in a function, then the result is lexicographically +smaller than the original function. -/ +lemma lex_desc {α} [preorder ι] [decidable_eq ι] [preorder α] {f : ι → α} {i j : ι} + (h₁ : i < j) (h₂ : f j < f i) : + to_lex (f ∘ equiv.swap i j) < to_lex f := +⟨i, λ k hik, congr_arg f (equiv.swap_apply_of_ne_of_ne hik.ne (hik.trans h₁).ne), + by simpa only [pi.to_lex_apply, function.comp_app, equiv.swap_apply_left] using h₂⟩ + end pi diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index 2ad3a0b402a0d..4ee42fa17d934 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -14,7 +14,8 @@ implies `P x`. Well-founded relations can be used for induction and recursion, i construction of fixed points in the space of dependent functions `Π x : α , β x`. The predicate `well_founded` is defined in the core library. In this file we prove some extra lemmas -and provide a few new definitions: `well_founded.min`, `well_founded.sup`, and `well_founded.succ`. +and provide a few new definitions: `well_founded.min`, `well_founded.sup`, and `well_founded.succ`, +and an induction principle `well_founded.induction_bot`. -/ variables {α : Type*} @@ -211,3 +212,48 @@ not_lt.mp $ not_lt_argmin_on f h s ha hs end linear_order end function + +section induction + +/-- Let `r` be a relation on `α`, let `f : α → β` be a function, let `C : β → Prop`, and +let `bot : α`. This induction principle shows that `C (f bot)` holds, given that +* some `a` that is accessible by `r` satisfies `C (f a)`, and +* for each `b` such that `f b ≠ f bot` and `C (f b)` holds, there is `c` + satisfying `r c b` and `C (f c)`. -/ +lemma acc.induction_bot' {α β} {r : α → α → Prop} {a bot : α} (ha : acc r a) {C : β → Prop} + {f : α → β} (ih : ∀ b, f b ≠ f bot → C (f b) → ∃ c, r c b ∧ C (f c)) : C (f a) → C (f bot) := +@acc.rec_on _ _ (λ x, C (f x) → C (f bot)) _ ha $ λ x ac ih' hC, + (eq_or_ne (f x) (f bot)).elim (λ h, h ▸ hC) + (λ h, let ⟨y, hy₁, hy₂⟩ := ih x h hC in ih' y hy₁ hy₂) + +/-- Let `r` be a relation on `α`, let `C : α → Prop` and let `bot : α`. +This induction principle shows that `C bot` holds, given that +* some `a` that is accessible by `r` satisfies `C a`, and +* for each `b ≠ bot` such that `C b` holds, there is `c` satisfying `r c b` and `C c`. -/ +lemma acc.induction_bot {α} {r : α → α → Prop} {a bot : α} (ha : acc r a) + {C : α → Prop} (ih : ∀ b, b ≠ bot → C b → ∃ c, r c b ∧ C c) : C a → C bot := +ha.induction_bot' ih + +/-- Let `r` be a well-founded relation on `α`, let `f : α → β` be a function, +let `C : β → Prop`, and let `bot : α`. +This induction principle shows that `C (f bot)` holds, given that +* some `a` satisfies `C (f a)`, and +* for each `b` such that `f b ≠ f bot` and `C (f b)` holds, there is `c` + satisfying `r c b` and `C (f c)`. -/ +lemma well_founded.induction_bot' {α β} {r : α → α → Prop} (hwf : well_founded r) {a bot : α} + {C : β → Prop} {f : α → β} (ih : ∀ b, f b ≠ f bot → C (f b) → ∃ c, r c b ∧ C (f c)) : + C (f a) → C (f bot) := +(hwf.apply a).induction_bot' ih + +/-- Let `r` be a well-founded relation on `α`, let `C : α → Prop`, and let `bot : α`. +This induction principle shows that `C bot` holds, given that +* some `a` satisfies `C a`, and +* for each `b` that satisfies `C b`, there is `c` satisfying `r c b` and `C c`. + +The naming is inspired by the fact that when `r` is transitive, it follows that `bot` is +the smallest element w.r.t. `r` that satisfies `C`. -/ +lemma well_founded.induction_bot {α} {r : α → α → Prop} (hwf : well_founded r) {a bot : α} + {C : α → Prop} (ih : ∀ b, b ≠ bot → C b → ∃ c, r c b ∧ C c) : C a → C bot := +hwf.induction_bot' ih + +end induction From 5ef85df4ee84b31b65a7fd3f7115b2c3710c653b Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 12 Oct 2022 13:05:05 +0000 Subject: [PATCH 721/828] feat(topology/basic): add `is_closed.interior_union` and `is_closed.interior_union'` (#16599) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `closure_inter_open` will be renamed to `is_open.closure_inter` in #16598 Already merged. Would it be better to add some lemmas like `interior (s ∪ t) = interior (s ∪ interior t)`? --- src/analysis/calculus/extend_deriv.lean | 2 +- src/topology/basic.lean | 29 +++++++++++++++---------- src/topology/locally_finite.lean | 2 +- 3 files changed, 19 insertions(+), 14 deletions(-) diff --git a/src/analysis/calculus/extend_deriv.lean b/src/analysis/calculus/extend_deriv.lean index 469273820c714..8915ac113fd2d 100644 --- a/src/analysis/calculus/extend_deriv.lean +++ b/src/analysis/calculus/extend_deriv.lean @@ -58,7 +58,7 @@ begin { rw closure_prod_eq at this, intros y y_in, apply this ⟨x, y⟩, - have : B ∩ closure s ⊆ closure (B ∩ s), from is_open_ball.closure_inter, + have : B ∩ closure s ⊆ closure (B ∩ s), from is_open_ball.inter_closure, exact ⟨this ⟨mem_ball_self δ_pos, hx⟩, this y_in⟩ }, have key : ∀ p : E × E, p ∈ (B ∩ s) ×ˢ (B ∩ s) → ∥f p.2 - f p.1 - (f' p.2 - f' p.1)∥ ≤ ε * ∥p.2 - p.1∥, diff --git a/src/topology/basic.lean b/src/topology/basic.lean index f9ea347e286b9..abff57bf24ca8 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -1086,21 +1086,26 @@ calc is_closed s ↔ closure s ⊆ s : closure_subset_iff_is_closed.symm lemma is_closed_iff_nhds {s : set α} : is_closed s ↔ ∀ x, (∀ U ∈ 𝓝 x, (U ∩ s).nonempty) → x ∈ s := by simp_rw [is_closed_iff_cluster_pt, cluster_pt, inf_principal_ne_bot_iff] -lemma is_open.closure_inter {s t : set α} (h : is_open s) : s ∩ closure t ⊆ closure (s ∩ t) := -begin - rintro a ⟨hs, ht⟩, - have : s ∈ 𝓝 a := is_open.mem_nhds h hs, - rw mem_closure_iff_nhds_ne_bot at ht ⊢, - rwa [← inf_principal, ← inf_assoc, inf_eq_left.2 (le_principal_iff.2 this)], -end +lemma is_closed.interior_union_left {s t : set α} (h : is_closed s) : + interior (s ∪ t) ⊆ s ∪ interior t := +λ a ⟨u, ⟨⟨hu₁, hu₂⟩, ha⟩⟩, (classical.em (a ∈ s)).imp_right $ λ h, mem_interior.mpr + ⟨u ∩ sᶜ, λ x hx, (hu₂ hx.1).resolve_left hx.2, is_open.inter hu₁ is_closed.is_open_compl, ⟨ha, h⟩⟩ + +lemma is_closed.interior_union_right {s t : set α} (h : is_closed t) : + interior (s ∪ t) ⊆ interior s ∪ t := +by simpa only [union_comm] using h.interior_union_left + +lemma is_open.inter_closure {s t : set α} (h : is_open s) : s ∩ closure t ⊆ closure (s ∩ t) := +compl_subset_compl.mp $ by simpa only [← interior_compl, compl_inter] + using is_closed.interior_union_left h.is_closed_compl -lemma is_open.closure_inter' {s t : set α} (h : is_open t) : closure s ∩ t ⊆ closure (s ∩ t) := -by simpa only [inter_comm] using h.closure_inter +lemma is_open.closure_inter {s t : set α} (h : is_open t) : closure s ∩ t ⊆ closure (s ∩ t) := +by simpa only [inter_comm] using h.inter_closure lemma dense.open_subset_closure_inter {s t : set α} (hs : dense s) (ht : is_open t) : t ⊆ closure (t ∩ s) := calc t = t ∩ closure s : by rw [hs.closure_eq, inter_univ] - ... ⊆ closure (t ∩ s) : ht.closure_inter + ... ⊆ closure (t ∩ s) : ht.inter_closure lemma mem_closure_of_mem_closure_union {s₁ s₂ : set α} {x : α} (h : x ∈ closure (s₁ ∪ s₂)) (h₁ : s₁ᶜ ∈ 𝓝 x) : x ∈ closure s₂ := @@ -1116,7 +1121,7 @@ end /-- The intersection of an open dense set with a dense set is a dense set. -/ lemma dense.inter_of_open_left {s t : set α} (hs : dense s) (ht : dense t) (hso : is_open s) : dense (s ∩ t) := -λ x, (closure_minimal hso.closure_inter is_closed_closure) $ +λ x, (closure_minimal hso.inter_closure is_closed_closure) $ by simp [hs.closure_eq, ht.closure_eq] /-- The intersection of a dense set with an open dense set is a dense set. -/ @@ -1131,7 +1136,7 @@ let ⟨U, hsub, ho, hx⟩ := mem_nhds_iff.1 ht in lemma closure_diff {s t : set α} : closure s \ closure t ⊆ closure (s \ t) := calc closure s \ closure t = (closure t)ᶜ ∩ closure s : by simp only [diff_eq, inter_comm] - ... ⊆ closure ((closure t)ᶜ ∩ s) : (is_open_compl_iff.mpr $ is_closed_closure).closure_inter + ... ⊆ closure ((closure t)ᶜ ∩ s) : (is_open_compl_iff.mpr $ is_closed_closure).inter_closure ... = closure (s \ closure t) : by simp only [diff_eq, inter_comm] ... ⊆ closure (s \ t) : closure_mono $ diff_subset_diff (subset.refl s) subset_closure diff --git a/src/topology/locally_finite.lean b/src/topology/locally_finite.lean index ef1b37e5b8318..b152ac06d10bf 100644 --- a/src/topology/locally_finite.lean +++ b/src/topology/locally_finite.lean @@ -70,7 +70,7 @@ begin intro x, rcases hf x with ⟨s, hsx, hsf⟩, refine ⟨interior s, interior_mem_nhds.2 hsx, hsf.subset $ λ i hi, _⟩, - exact (hi.mono is_open_interior.closure_inter').of_closure.mono + exact (hi.mono is_open_interior.closure_inter).of_closure.mono (inter_subset_inter_right _ interior_subset) end From 6ef05bb32396902743f6f3f1c8baf6a7ce2e404a Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 12 Oct 2022 13:05:06 +0000 Subject: [PATCH 722/828] feat(topology/algebra/infinite_sum): infinite sums on subsets (#16671) * Add more lemmas about infinite sums on subsets of the domain * Group some existing lemmas together in the same "section" * Also add some lemmas about `set.maps_to.restrict` that I used until I refactored one of the proofs --- src/data/set/function.lean | 11 +++++++ src/topology/algebra/infinite_sum.lean | 41 ++++++++++++++++++-------- 2 files changed, 40 insertions(+), 12 deletions(-) diff --git a/src/data/set/function.lean b/src/data/set/function.lean index f1cb1341f2ac1..6e7b7284f964e 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -261,6 +261,14 @@ subtype.map f h @[simp] lemma maps_to.coe_restrict_apply (h : maps_to f s t) (x : s) : (h.restrict f s t x : β) = f x := rfl +/-- Restricting the domain and then the codomain is the same as `maps_to.restrict`. -/ +@[simp] lemma cod_restrict_restrict (h : ∀ x : s, f x ∈ t) : + cod_restrict (s.restrict f) t h = maps_to.restrict f s t (λ x hx, h ⟨x, hx⟩) := rfl + +/-- Reverse of `set.cod_restrict_restrict`. -/ +lemma maps_to.restrict_eq_cod_restrict (h : maps_to f s t) : + h.restrict f s t = cod_restrict (s.restrict f) t (λ x, h x.2) := rfl + lemma maps_to.coe_restrict (h : set.maps_to f s t) : coe ∘ h.restrict f s t = s.restrict f := rfl @@ -463,6 +471,9 @@ lemma inj_on_iff_injective : inj_on f s ↔ injective (s.restrict f) := alias inj_on_iff_injective ↔ inj_on.injective _ +lemma maps_to.restrict_inj (h : maps_to f s t) : injective (h.restrict f s t) ↔ inj_on f s := +by rw [h.restrict_eq_cod_restrict, injective_cod_restrict, inj_on_iff_injective] + lemma exists_inj_on_iff_injective [nonempty β] : (∃ f : α → β, inj_on f s) ↔ ∃ f : s → β, injective f := ⟨λ ⟨f, hf⟩, ⟨_, hf.injective⟩, diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 118bca4c834b9..1d59f5836bdd5 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -456,14 +456,6 @@ lemma tsum_fintype [fintype β] (f : β → α) : ∑'b, f b = ∑ b, f b := lemma tsum_bool (f : bool → α) : ∑' i : bool, f i = f false + f true := by { rw [tsum_fintype, finset.sum_eq_add]; simp } -@[simp] lemma finset.tsum_subtype (s : finset β) (f : β → α) : - ∑' x : {x // x ∈ s}, f x = ∑ x in s, f x := -(s.has_sum f).tsum_eq - -@[simp] lemma finset.tsum_subtype' (s : finset β) (f : β → α) : - ∑' x : (s : set β), f x = ∑ x in s, f x := -s.tsum_subtype f - lemma tsum_eq_single {f : β → α} (b : β) (hf : ∀b' ≠ b, f b' = 0) : ∑'b, f b = f b := (has_sum_single b hf).tsum_eq @@ -516,10 +508,6 @@ lemma tsum_eq_tsum_of_ne_zero_bij {g : γ → α} (i : support g → β) ∑' x, f x = ∑' y, g y := tsum_eq_tsum_of_has_sum_iff_has_sum $ λ _, has_sum_iff_has_sum_of_ne_zero_bij i hi hf hfg -lemma tsum_subtype (s : set β) (f : β → α) : - ∑' x:s, f x = ∑' x, s.indicator f x := -tsum_eq_tsum_of_has_sum_iff_has_sum $ λ _, has_sum_subtype_iff_indicator - lemma tsum_op : ∑' x, mul_opposite.op (f x) = mul_opposite.op (∑' x, f x) := begin by_cases h : summable f, @@ -531,6 +519,35 @@ end lemma tsum_unop {f : β → αᵐᵒᵖ} : ∑' x, mul_opposite.unop (f x) = mul_opposite.unop (∑' x, f x) := mul_opposite.op_injective tsum_op.symm +/-! ### `tsum` on subsets -/ + +@[simp] lemma finset.tsum_subtype (s : finset β) (f : β → α) : + ∑' x : {x // x ∈ s}, f x = ∑ x in s, f x := +(s.has_sum f).tsum_eq + +@[simp] lemma finset.tsum_subtype' (s : finset β) (f : β → α) : + ∑' x : (s : set β), f x = ∑ x in s, f x := +s.tsum_subtype f + +lemma tsum_subtype (s : set β) (f : β → α) : + ∑' x : s, f x = ∑' x, s.indicator f x := +tsum_eq_tsum_of_has_sum_iff_has_sum $ λ _, has_sum_subtype_iff_indicator + +lemma tsum_subtype_eq_of_support_subset {f : β → α} {s : set β} (hs : support f ⊆ s) : + ∑' x : s, f x = ∑' x, f x := +tsum_eq_tsum_of_has_sum_iff_has_sum $ λ x, has_sum_subtype_iff_of_support_subset hs + +@[simp] lemma tsum_univ (f : β → α) : ∑' x : (set.univ : set β), f x = ∑' x, f x := +tsum_subtype_eq_of_support_subset $ set.subset_univ _ + +lemma tsum_image {g : γ → β} (f : β → α) {s : set γ} (hg : set.inj_on g s) : + ∑' x : g '' s, f x = ∑' x : s, f (g x) := +((equiv.set.image_of_inj_on _ _ hg).tsum_eq (λ x, f x)).symm + +lemma tsum_range {g : γ → β} (f : β → α) (hg : injective g) : + ∑' x : set.range g, f x = ∑' x, f (g x) := +by rw [← set.image_univ, tsum_image f (hg.inj_on _), tsum_univ (f ∘ g)] + section has_continuous_add variable [has_continuous_add α] From 98c944e9331fca265f5dcbcb13b5b30d77df7a84 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 12 Oct 2022 13:05:07 +0000 Subject: [PATCH 723/828] feat(combinatorics/quiver/path): Composition is injective (#16852) `quiver.path.comp` is injective on the left and on the right. --- src/combinatorics/quiver/path.lean | 41 ++++++++++++++++++++++++++++-- src/logic/lemmas.lean | 5 ++++ 2 files changed, 44 insertions(+), 2 deletions(-) diff --git a/src/combinatorics/quiver/path.lean b/src/combinatorics/quiver/path.lean index 3663aaf7f361a..bc230d8317209 100644 --- a/src/combinatorics/quiver/path.lean +++ b/src/combinatorics/quiver/path.lean @@ -5,6 +5,7 @@ Authors: David Wärn, Scott Morrison -/ import combinatorics.quiver.basic import data.list.basic +import logic.lemmas /-! # Paths in quivers @@ -30,7 +31,7 @@ path.nil.cons e namespace path -variables {V : Type u} [quiver V] {a b : V} +variables {V : Type u} [quiver V] {a b c : V} /-- The length of a path is the number of arrows it uses. -/ def length {a : V} : Π {b : V}, path a b → ℕ @@ -65,6 +66,43 @@ def comp {a b : V} : Π {c}, path a b → path b c → path a c | c p q path.nil := rfl | d p q (path.cons r e) := by rw [comp_cons, comp_cons, comp_cons, comp_assoc] +@[simp] lemma length_comp (p : path a b) : + ∀ {c} (q : path b c), (p.comp q).length = p.length + q.length +| c nil := rfl +| c (cons q h) := congr_arg nat.succ q.length_comp + +lemma comp_inj {p₁ p₂ : path a b} {q₁ q₂ : path b c} (hq : q₁.length = q₂.length) : + p₁.comp q₁ = p₂.comp q₂ ↔ p₁ = p₂ ∧ q₁ = q₂ := +begin + refine ⟨λ h, _, by { rintro ⟨rfl, rfl⟩, refl }⟩, + induction q₁ with d₁ e₁ q₁ f₁ ih generalizing q₂; obtain _ | ⟨d₂, e₂, q₂, f₂⟩ := q₂, + { exact ⟨h, rfl⟩ }, + { cases hq }, + { cases hq }, + simp only [comp_cons] at h, + obtain rfl := h.1, + obtain ⟨rfl, rfl⟩ := ih (nat.succ_injective hq) h.2.1.eq, + rw h.2.2.eq, + exact ⟨rfl, rfl⟩, +end + +lemma comp_inj' {p₁ p₂ : path a b} {q₁ q₂ : path b c} (h : p₁.length = p₂.length) : + p₁.comp q₁ = p₂.comp q₂ ↔ p₁ = p₂ ∧ q₁ = q₂ := +⟨λ h_eq, (comp_inj $ add_left_injective p₁.length $ + by simpa [h] using congr_arg length h_eq).1 h_eq, by { rintro ⟨rfl, rfl⟩, refl }⟩ + +lemma comp_injective_left (q : path b c) : injective (λ p : path a b, p.comp q) := +λ p₁ p₂ h, ((comp_inj rfl).1 h).1 + +lemma comp_injective_right (p : path a b) : injective (p.comp : path b c → path a c) := +λ q₁ q₂ h, ((comp_inj' rfl).1 h).2 + +@[simp] lemma comp_inj_left {p₁ p₂ : path a b} {q : path b c} : p₁.comp q = p₂.comp q ↔ p₁ = p₂ := +q.comp_injective_left.eq_iff + +@[simp] lemma comp_inj_right {p : path a b} {q₁ q₂ : path b c} : p.comp q₁ = p.comp q₂ ↔ q₁ = q₂ := +p.comp_injective_right.eq_iff + /-- Turn a path into a list. The list contains `a` at its head, but not `b` a priori. -/ @[simp] def to_list : Π {b : V}, path a b → list V | b nil := [] @@ -126,4 +164,3 @@ def map_path {a : V} : lemma map_path_to_path {a b : V} (f : a ⟶ b) : F.map_path f.to_path = (F.map f).to_path := rfl end prefunctor - diff --git a/src/logic/lemmas.lean b/src/logic/lemmas.lean index 88a0046c82069..c91dd5cc78d0b 100644 --- a/src/logic/lemmas.lean +++ b/src/logic/lemmas.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ +import tactic.protected import tactic.split_ifs /-! @@ -17,6 +18,10 @@ We spell those lemmas out with `dite` and `ite` rather than the `if then else` n would result in less delta-reduced statements. -/ +alias heq_iff_eq ↔ heq.eq eq.heq + +attribute [protected] heq.eq eq.heq + variables {α : Sort*} {p q r : Prop} [decidable p] [decidable q] {a b c : α} lemma dite_dite_distrib_left {a : p → α} {b : ¬ p → q → α} {c : ¬ p → ¬ q → α} : From 73ccd02e060da99a98f9388e96c4e8f142a3f7c8 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 12 Oct 2022 13:05:08 +0000 Subject: [PATCH 724/828] =?UTF-8?q?refactor(set=5Ftheory/cardinal):=20revi?= =?UTF-8?q?ew=20API=20about=20`#=CE=B1=20=3D=202`/`nat.card=20=CE=B1=20=3D?= =?UTF-8?q?=202`=20(#16899)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Rename `cardinal.mk_eq_nat_iff_finset` to `cardinal.mk_set_eq_nat_iff_finset`, add a version for types and `cardinal.mk_eq_nat_iff_fintype`. * Add `cardinal.to_nat_eq_iff`, a more general version of `cardinal.to_nat_eq_one`. * Rename `cardinal.exists_not_mem_of_length_le` to `cardinal.exists_not_mem_of_length_lt`. * Add `cardinal.mk_eq_two_iff`, `cardinal.mk_eq_two_iff'`, `nat.card_eq_two_iff`, and `nat.card_eq_two_iff'`. --- src/linear_algebra/dimension.lean | 4 +- src/set_theory/cardinal/basic.lean | 57 +++++++++++++++++++++-------- src/set_theory/cardinal/finite.lean | 6 +++ 3 files changed, 49 insertions(+), 18 deletions(-) diff --git a/src/linear_algebra/dimension.lean b/src/linear_algebra/dimension.lean index e262a9055ccba..36ae8f23c62f4 100644 --- a/src/linear_algebra/dimension.lean +++ b/src/linear_algebra/dimension.lean @@ -1203,7 +1203,7 @@ lemma le_dim_iff_exists_linear_independent_finset {n : ℕ} : ↑n ≤ module.rank K V ↔ ∃ s : finset V, s.card = n ∧ linear_independent K (coe : (s : set V) → V) := begin - simp only [le_dim_iff_exists_linear_independent, cardinal.mk_eq_nat_iff_finset], + simp only [le_dim_iff_exists_linear_independent, cardinal.mk_set_eq_nat_iff_finset], split, { rintro ⟨s, ⟨t, rfl, rfl⟩, si⟩, exact ⟨t, rfl, si⟩ }, @@ -1332,7 +1332,7 @@ lemma le_rank_iff_exists_linear_independent_finset {n : ℕ} {f : V →ₗ[K] V' ↑n ≤ rank f ↔ ∃ s : finset V, s.card = n ∧ linear_independent K (λ x : (s : set V), f x) := begin simp only [le_rank_iff_exists_linear_independent, cardinal.lift_nat_cast, - cardinal.lift_eq_nat_iff, cardinal.mk_eq_nat_iff_finset], + cardinal.lift_eq_nat_iff, cardinal.mk_set_eq_nat_iff_finset], split, { rintro ⟨s, ⟨t, rfl, rfl⟩, si⟩, exact ⟨t, rfl, si⟩ }, diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index 03faec351a871..28e96f448a936 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -1169,10 +1169,13 @@ by rw [←to_nat_cast 0, nat.cast_zero] @[simp] lemma one_to_nat : to_nat 1 = 1 := by rw [←to_nat_cast 1, nat.cast_one] +lemma to_nat_eq_iff {c : cardinal} {n : ℕ} (hn : n ≠ 0) : to_nat c = n ↔ c = n := +⟨λ h, (cast_to_nat_of_lt_aleph_0 (lt_of_not_ge (hn ∘ h.symm.trans ∘ + to_nat_apply_of_aleph_0_le))).symm.trans (congr_arg coe h), + λ h, (congr_arg to_nat h).trans (to_nat_cast n)⟩ + @[simp] lemma to_nat_eq_one {c : cardinal} : to_nat c = 1 ↔ c = 1 := -⟨λ h, (cast_to_nat_of_lt_aleph_0 (lt_of_not_ge (one_ne_zero ∘ h.symm.trans ∘ - to_nat_apply_of_aleph_0_le))).symm.trans ((congr_arg coe h).trans nat.cast_one), - λ h, (congr_arg to_nat h).trans one_to_nat⟩ +by rw [to_nat_eq_iff one_ne_zero, nat.cast_one] lemma to_nat_eq_one_iff_unique {α : Type*} : (#α).to_nat = 1 ↔ subsingleton α ∧ nonempty α := to_nat_eq_one.trans eq_one_iff_unique @@ -1388,7 +1391,7 @@ by { rw bUnion_eq_Union, apply mk_Union_le } lemma finset_card_lt_aleph_0 (s : finset α) : #(↑s : set α) < ℵ₀ := lt_aleph_0_of_finite _ -theorem mk_eq_nat_iff_finset {α} {s : set α} {n : ℕ} : +theorem mk_set_eq_nat_iff_finset {α} {s : set α} {n : ℕ} : #s = n ↔ ∃ t : finset α, (t : set α) = s ∧ t.card = n := begin split, @@ -1399,6 +1402,19 @@ begin exact mk_coe_finset } end +theorem mk_eq_nat_iff_finset {n : ℕ} : #α = n ↔ ∃ t : finset α, (t : set α) = univ ∧ t.card = n := +by rw [← mk_univ, mk_set_eq_nat_iff_finset] + +theorem mk_eq_nat_iff_fintype {n : ℕ} : #α = n ↔ ∃ (h : fintype α), @fintype.card α h = n := +begin + rw [mk_eq_nat_iff_finset], + split, + { rintro ⟨t, ht, hn⟩, + exact ⟨⟨t, eq_univ_iff_forall.1 ht⟩, hn⟩ }, + { rintro ⟨⟨t, ht⟩, hn⟩, + exact ⟨t, eq_univ_iff_forall.2 ht, hn⟩ } +end + theorem mk_union_add_mk_inter {α : Type u} {S T : set α} : #(S ∪ T : set α) + #(S ∩ T : set α) = #S + #T := quot.sound ⟨equiv.set.union_sum_inter S T⟩ @@ -1507,24 +1523,33 @@ begin end lemma 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] + +lemma two_le_iff' (x : α) : (2 : cardinal) ≤ #α ↔ ∃y : α, y ≠ x := +by rw [two_le_iff, ← nontrivial_iff, nontrivial_iff_exists_ne x] + +lemma mk_eq_two_iff : #α = 2 ↔ ∃ x y : α, x ≠ y ∧ ({x, y} : set α) = univ := begin + simp only [← @nat.cast_two cardinal, mk_eq_nat_iff_finset, finset.card_eq_two], split, - { rintro ⟨f⟩, refine ⟨f $ sum.inl ⟨⟩, f $ sum.inr ⟨⟩, _⟩, intro h, cases f.2 h }, - { rintro ⟨x, y, h⟩, by_contra h', - rw [not_le, ←nat.cast_two, nat_succ, lt_succ_iff, nat.cast_one, le_one_iff_subsingleton] at h', - apply h, exactI subsingleton.elim _ _ } + { rintro ⟨t, ht, x, y, hne, rfl⟩, + exact ⟨x, y, hne, by simpa using ht⟩ }, + { rintro ⟨x, y, hne, h⟩, + exact ⟨{x, y}, by simpa using h, x, y, hne, rfl⟩ } end -lemma two_le_iff' (x : α) : (2 : cardinal) ≤ #α ↔ ∃y : α, x ≠ y := +lemma mk_eq_two_iff' (x : α) : #α = 2 ↔ ∃! y, y ≠ x := begin - rw [two_le_iff], - split, - { rintro ⟨y, z, h⟩, refine classical.by_cases (λ(h' : x = y), _) (λ h', ⟨y, h'⟩), - rw [←h'] at h, exact ⟨z, h⟩ }, - { rintro ⟨y, h⟩, exact ⟨x, y, h⟩ } + rw [mk_eq_two_iff], split, + { rintro ⟨a, b, hne, h⟩, + simp only [eq_univ_iff_forall, mem_insert_iff, mem_singleton_iff] at h, + rcases h x with rfl|rfl, + exacts [⟨b, hne.symm, λ z, (h z).resolve_left⟩, ⟨a, hne, λ z, (h z).resolve_right⟩] }, + { rintro ⟨y, hne, hy⟩, + exact ⟨x, y, hne.symm, eq_univ_of_forall $ λ z, or_iff_not_imp_left.2 (hy z)⟩ } end -lemma exists_not_mem_of_length_le {α : Type*} (l : list α) (h : ↑l.length < # α) : +lemma exists_not_mem_of_length_lt {α : Type*} (l : list α) (h : ↑l.length < # α) : ∃ (z : α), z ∉ l := begin contrapose! h, @@ -1539,7 +1564,7 @@ lemma three_le {α : Type*} (h : 3 ≤ # α) (x : α) (y : α) : begin have : ↑(3 : ℕ) ≤ # α, simpa using h, have : ↑(2 : ℕ) < # α, rwa [← succ_le_iff, ← cardinal.nat_succ], - have := exists_not_mem_of_length_le [x, y] this, + have := exists_not_mem_of_length_lt [x, y] this, simpa [not_or_distrib] using this, end diff --git a/src/set_theory/cardinal/finite.lean b/src/set_theory/cardinal/finite.lean index 0aa0b280c13cd..537d6afb6d7e9 100644 --- a/src/set_theory/cardinal/finite.lean +++ b/src/set_theory/cardinal/finite.lean @@ -70,6 +70,12 @@ card_of_subsingleton default lemma card_eq_one_iff_unique : nat.card α = 1 ↔ subsingleton α ∧ nonempty α := cardinal.to_nat_eq_one_iff_unique +lemma card_eq_two_iff : nat.card α = 2 ↔ ∃ x y : α, x ≠ y ∧ {x, y} = @set.univ α := +(to_nat_eq_iff two_ne_zero).trans $ iff.trans (by rw [nat.cast_two]) mk_eq_two_iff + +lemma card_eq_two_iff' (x : α) : nat.card α = 2 ↔ ∃! y, y ≠ x := +(to_nat_eq_iff two_ne_zero).trans $ iff.trans (by rw [nat.cast_two]) (mk_eq_two_iff' x) + theorem card_of_is_empty [is_empty α] : nat.card α = 0 := by simp @[simp] lemma card_prod (α β : Type*) : nat.card (α × β) = nat.card α * nat.card β := From 6979111c19114a439cc7b7d4b8b450396a81067a Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Wed, 12 Oct 2022 13:05:09 +0000 Subject: [PATCH 725/828] =?UTF-8?q?feat(group=5Ftheory/transfer):=20The=20?= =?UTF-8?q?kernel=20of=20the=20transfer=20homomorphism=20`G=20=E2=86=92*?= =?UTF-8?q?=20P`=20is=20disjoint=20from=20every=20`p`-group=20(#16914)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR proves if `P` is a Sylow `p`-subgroup of `G`, then the kernel of the transfer homomorphism `G →* P` is disjoint from every `p`-group. This is getting very close to Burnside's transfer theorem. I also added variables lines for some typeclass assumptions. --- src/group_theory/transfer.lean | 38 ++++++++++++++++++++++++++++------ 1 file changed, 32 insertions(+), 6 deletions(-) diff --git a/src/group_theory/transfer.lean b/src/group_theory/transfer.lean index 3d8f0070003ea..914a94823d507 100644 --- a/src/group_theory/transfer.lean +++ b/src/group_theory/transfer.lean @@ -172,9 +172,11 @@ noncomputable def transfer_sylow [fintype (G ⧸ (P : subgroup G))] : G →* (P @transfer G _ P P (@subgroup.is_commutative.comm_group G _ P ⟨⟨λ a b, subtype.ext (hP (le_normalizer b.2) a a.2)⟩⟩) (monoid_hom.id P) _ +variables [fact p.prime] [finite (sylow p G)] + /-- Auxillary lemma in order to state `transfer_sylow_eq_pow`. -/ -lemma transfer_sylow_eq_pow_aux [fact p.prime] [finite (sylow p G)] (g : G) (hg : g ∈ P) (k : ℕ) - (g₀ : G) (h : g₀⁻¹ * g ^ k * g₀ ∈ P) : g₀⁻¹ * g ^ k * g₀ = g ^ k := +lemma transfer_sylow_eq_pow_aux (g : G) (hg : g ∈ P) (k : ℕ) (g₀ : G) (h : g₀⁻¹ * g ^ k * g₀ ∈ P) : + g₀⁻¹ * g ^ k * g₀ = g ^ k := begin haveI : (P : subgroup G).is_commutative := ⟨⟨λ a b, subtype.ext (hP (le_normalizer b.2) a a.2)⟩⟩, replace hg := (P : subgroup G).pow_mem hg k, @@ -182,12 +184,36 @@ begin exact h.trans (commute.inv_mul_cancel (hP hn (g ^ k) hg).symm), end -lemma transfer_sylow_eq_pow [fact p.prime] [finite (sylow p G)] [fintype (G ⧸ (P : subgroup G))] - (hP : P.1.normalizer ≤ P.1.centralizer) (g : G) (hg : g ∈ P) : - transfer_sylow P hP g = ⟨g ^ (P : subgroup G).index, - transfer_eq_pow_aux g (transfer_sylow_eq_pow_aux P hP g hg)⟩ := +variables [fintype (G ⧸ (P : subgroup G))] + +lemma transfer_sylow_eq_pow (g : G) (hg : g ∈ P) : transfer_sylow P hP g = + ⟨g ^ (P : subgroup G).index, transfer_eq_pow_aux g (transfer_sylow_eq_pow_aux P hP g hg)⟩ := by apply transfer_eq_pow +lemma ker_transfer_sylow_disjoint : disjoint (transfer_sylow P hP).ker ↑P := +begin + intros g hg, + obtain ⟨j, hj⟩ := P.2 ⟨g, hg.2⟩, + have := pow_gcd_eq_one (⟨g, hg.2⟩ : (P : subgroup G)) + ((transfer_sylow_eq_pow P hP g hg.2).symm.trans hg.1) hj, + rwa [((fact.out p.prime).coprime_pow_of_not_dvd (not_dvd_index_sylow P _)).gcd_eq_one, + pow_one, subtype.ext_iff] at this, + exact index_ne_zero_of_finite ∘ (relindex_top_right (P : subgroup G)).symm.trans ∘ + relindex_eq_zero_of_le_right le_top, +end + +lemma ker_transfer_sylow_disjoint' (Q : sylow p G) : disjoint (transfer_sylow P hP).ker ↑Q := +begin + obtain ⟨g, hg⟩ := exists_smul_eq G Q P, + rw [disjoint_iff, ←smul_left_cancel_iff (mul_aut.conj g), smul_bot, smul_inf, smul_normal, + ←sylow.coe_subgroup_smul, hg, ←disjoint_iff], + exact ker_transfer_sylow_disjoint P hP, +end + +lemma ker_transfer_sylow_disjoint'' (Q : subgroup G) (hQ : is_p_group p Q) : + disjoint (transfer_sylow P hP).ker Q := +let ⟨R, hR⟩ := hQ.exists_le_sylow in (ker_transfer_sylow_disjoint' P hP R).mono_right hR + end burnside_transfer end monoid_hom From b49d5686ebc6a118cacd06139ead6d055b8e3759 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 12 Oct 2022 13:05:11 +0000 Subject: [PATCH 726/828] chore(.github/CODEOWNERS): two new rules (#16921) There is an attempt to order the rules from generic to specific. --- .github/CODEOWNERS | 14 ++++++++++---- 1 file changed, 10 insertions(+), 4 deletions(-) diff --git a/.github/CODEOWNERS b/.github/CODEOWNERS index 7a40559f48dc7..07aee9c516613 100644 --- a/.github/CODEOWNERS +++ b/.github/CODEOWNERS @@ -1,13 +1,19 @@ # By default, the mathlib maintainers own everything in the repo. # Later matches will take precedence over this match. + # Disabled until we have better coverage by other patterns. Reenable in the future. -# * @leanprover-community/mathlib-maintainers +# * @leanprover-community/mathlib-maintainers + +src/category_theory/ @leanprover-community/mathlib-CT -src/tactic/ @leanprover-community/mathlib-meta +src/measure_theory/ @leanprover-community/mathlib-meas -src/category_theory/ @leanprover-community/mathlib-CT src/algebraic_topology/ @leanprover-community/mathlib-CT # for now the category theory team can take care of this src/algebraic_geometry/ @leanprover-community/mathlib-AG -src/combinatorics/ @leanprover-community/mathlib-CO +src/combinatorics/ @leanprover-community/mathlib-CO + +src/probability/ @leanprover-community/mathlib-PR + +src/tactic/ @leanprover-community/mathlib-meta From 357b2966fcd12f837c3141ef400c6a48e620292c Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Wed, 12 Oct 2022 16:16:46 +0000 Subject: [PATCH 727/828] feat(analysis/convex/side): betweenness and affine subspaces (#16733) Define notions of points being (weakly or strictly) on the same side or opposite sides of an affine subspace (e.g. a line), for use in describing geometrical configurations, and provide some associated API. --- src/analysis/convex/side.lean | 803 ++++++++++++++++++++++++++++++++++ 1 file changed, 803 insertions(+) create mode 100644 src/analysis/convex/side.lean diff --git a/src/analysis/convex/side.lean b/src/analysis/convex/side.lean new file mode 100644 index 0000000000000..5b90dcc960f69 --- /dev/null +++ b/src/analysis/convex/side.lean @@ -0,0 +1,803 @@ +/- +Copyright (c) 2022 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers +-/ +import analysis.convex.between + +/-! +# Sides of affine subspaces + +This file defines notions of two points being on the same or opposite sides of an affine subspace. + +## Main definitions + +* `s.w_same_side x y`: The points `x` and `y` are weakly on the same side of the affine + subspace `s`. +* `s.s_same_side x y`: The points `x` and `y` are strictly on the same side of the affine + subspace `s`. +* `s.w_opp_side x y`: The points `x` and `y` are weakly on opposite sides of the affine + subspace `s`. +* `s.s_opp_side x y`: The points `x` and `y` are strictly on opposite sides of the affine + subspace `s`. + +-/ + +variables {R V V' P P' : Type*} + +open affine_map + +namespace affine_subspace + +section strict_ordered_comm_ring + +variables [strict_ordered_comm_ring R] [add_comm_group V] [module R V] [add_torsor V P] +variables [add_comm_group V'] [module R V'] [add_torsor V' P'] + +include V + +/-- The points `x` and `y` are weakly on the same side of `s`. -/ +def w_same_side (s : affine_subspace R P) (x y : P) : Prop := +∃ p₁ p₂ ∈ s, same_ray R (x -ᵥ p₁) (y -ᵥ p₂) + +/-- The points `x` and `y` are strictly on the same side of `s`. -/ +def s_same_side (s : affine_subspace R P) (x y : P) : Prop := +s.w_same_side x y ∧ x ∉ s ∧ y ∉ s + +/-- The points `x` and `y` are weakly on opposite sides of `s`. -/ +def w_opp_side (s : affine_subspace R P) (x y : P) : Prop := +∃ p₁ p₂ ∈ s, same_ray R (x -ᵥ p₁) (p₂ -ᵥ y) + +/-- The points `x` and `y` are strictly on opposite sides of `s`. -/ +def s_opp_side (s : affine_subspace R P) (x y : P) : Prop := +s.w_opp_side x y ∧ x ∉ s ∧ y ∉ s + +include V' + +lemma w_same_side.map {s : affine_subspace R P} {x y : P} (h : s.w_same_side x y) + (f : P →ᵃ[R] P') : (s.map f).w_same_side (f x) (f y) := +begin + rcases h with ⟨p₁, hp₁, p₂, hp₂, h⟩, + refine ⟨f p₁, mem_map_of_mem f hp₁, f p₂, mem_map_of_mem f hp₂, _⟩, + simp_rw [←linear_map_vsub], + exact h.map f.linear +end + +lemma _root_.function.injective.w_same_side_map_iff {s : affine_subspace R P} {x y : P} + {f : P →ᵃ[R] P'} (hf : function.injective f) : + (s.map f).w_same_side (f x) (f y) ↔ s.w_same_side x y := +begin + refine ⟨λ h, _, λ h, h.map _⟩, + rcases h with ⟨fp₁, hfp₁, fp₂, hfp₂, h⟩, + rw mem_map at hfp₁ hfp₂, + rcases hfp₁ with ⟨p₁, hp₁, rfl⟩, + rcases hfp₂ with ⟨p₂, hp₂, rfl⟩, + refine ⟨p₁, hp₁, p₂, hp₂, _⟩, + simp_rw [←linear_map_vsub, (f.injective_iff_linear_injective.2 hf).same_ray_map_iff] at h, + exact h +end + +lemma _root_.function.injective.s_same_side_map_iff {s : affine_subspace R P} {x y : P} + {f : P →ᵃ[R] P'} (hf : function.injective f) : + (s.map f).s_same_side (f x) (f y) ↔ s.s_same_side x y := +by simp_rw [s_same_side, hf.w_same_side_map_iff, mem_map_iff_mem_of_injective hf] + +@[simp] lemma _root_.affine_equiv.w_same_side_map_iff {s : affine_subspace R P} {x y : P} + (f : P ≃ᵃ[R] P') : (s.map ↑f).w_same_side (f x) (f y) ↔ s.w_same_side x y := +(show function.injective f.to_affine_map, from f.injective).w_same_side_map_iff + +@[simp] lemma _root_.affine_equiv.s_same_side_map_iff {s : affine_subspace R P} {x y : P} + (f : P ≃ᵃ[R] P') : (s.map ↑f).s_same_side (f x) (f y) ↔ s.s_same_side x y := +(show function.injective f.to_affine_map, from f.injective).s_same_side_map_iff + +lemma w_opp_side.map {s : affine_subspace R P} {x y : P} (h : s.w_opp_side x y) + (f : P →ᵃ[R] P') : (s.map f).w_opp_side (f x) (f y) := +begin + rcases h with ⟨p₁, hp₁, p₂, hp₂, h⟩, + refine ⟨f p₁, mem_map_of_mem f hp₁, f p₂, mem_map_of_mem f hp₂, _⟩, + simp_rw [←linear_map_vsub], + exact h.map f.linear +end + +lemma _root_.function.injective.w_opp_side_map_iff {s : affine_subspace R P} {x y : P} + {f : P →ᵃ[R] P'} (hf : function.injective f) : + (s.map f).w_opp_side (f x) (f y) ↔ s.w_opp_side x y := +begin + refine ⟨λ h, _, λ h, h.map _⟩, + rcases h with ⟨fp₁, hfp₁, fp₂, hfp₂, h⟩, + rw mem_map at hfp₁ hfp₂, + rcases hfp₁ with ⟨p₁, hp₁, rfl⟩, + rcases hfp₂ with ⟨p₂, hp₂, rfl⟩, + refine ⟨p₁, hp₁, p₂, hp₂, _⟩, + simp_rw [←linear_map_vsub, (f.injective_iff_linear_injective.2 hf).same_ray_map_iff] at h, + exact h +end + +lemma _root_.function.injective.s_opp_side_map_iff {s : affine_subspace R P} {x y : P} + {f : P →ᵃ[R] P'} (hf : function.injective f) : + (s.map f).s_opp_side (f x) (f y) ↔ s.s_opp_side x y := +by simp_rw [s_opp_side, hf.w_opp_side_map_iff, mem_map_iff_mem_of_injective hf] + +@[simp] lemma _root_.affine_equiv.w_opp_side_map_iff {s : affine_subspace R P} {x y : P} + (f : P ≃ᵃ[R] P') : (s.map ↑f).w_opp_side (f x) (f y) ↔ s.w_opp_side x y := +(show function.injective f.to_affine_map, from f.injective).w_opp_side_map_iff + +@[simp] lemma _root_.affine_equiv.s_opp_side_map_iff {s : affine_subspace R P} {x y : P} + (f : P ≃ᵃ[R] P') : (s.map ↑f).s_opp_side (f x) (f y) ↔ s.s_opp_side x y := +(show function.injective f.to_affine_map, from f.injective).s_opp_side_map_iff + +omit V' + +lemma w_same_side.nonempty {s : affine_subspace R P} {x y : P} (h : s.w_same_side x y) : + (s : set P).nonempty := +⟨h.some, h.some_spec.some⟩ + +lemma s_same_side.nonempty {s : affine_subspace R P} {x y : P} (h : s.s_same_side x y) : + (s : set P).nonempty := +⟨h.1.some, h.1.some_spec.some⟩ + +lemma w_opp_side.nonempty {s : affine_subspace R P} {x y : P} (h : s.w_opp_side x y) : + (s : set P).nonempty := +⟨h.some, h.some_spec.some⟩ + +lemma s_opp_side.nonempty {s : affine_subspace R P} {x y : P} (h : s.s_opp_side x y) : + (s : set P).nonempty := +⟨h.1.some, h.1.some_spec.some⟩ + +lemma s_same_side.w_same_side {s : affine_subspace R P} {x y : P} (h : s.s_same_side x y) : + s.w_same_side x y := +h.1 + +lemma s_same_side.left_not_mem {s : affine_subspace R P} {x y : P} (h : s.s_same_side x y) : + x ∉ s := +h.2.1 + +lemma s_same_side.right_not_mem {s : affine_subspace R P} {x y : P} (h : s.s_same_side x y) : + y ∉ s := +h.2.2 + +lemma s_opp_side.w_opp_side {s : affine_subspace R P} {x y : P} (h : s.s_opp_side x y) : + s.w_opp_side x y := +h.1 + +lemma s_opp_side.left_not_mem {s : affine_subspace R P} {x y : P} (h : s.s_opp_side x y) : + x ∉ s := +h.2.1 + +lemma s_opp_side.right_not_mem {s : affine_subspace R P} {x y : P} (h : s.s_opp_side x y) : + y ∉ s := +h.2.2 + +lemma w_same_side_comm {s : affine_subspace R P} {x y : P} : + s.w_same_side x y ↔ s.w_same_side y x := +⟨λ ⟨p₁, hp₁, p₂, hp₂, h⟩, ⟨p₂, hp₂, p₁, hp₁, h.symm⟩, + λ ⟨p₁, hp₁, p₂, hp₂, h⟩, ⟨p₂, hp₂, p₁, hp₁, h.symm⟩⟩ + +alias w_same_side_comm ↔ w_same_side.symm _ + +lemma s_same_side_comm {s : affine_subspace R P} {x y : P} : + s.s_same_side x y ↔ s.s_same_side y x := +by rw [s_same_side, s_same_side, w_same_side_comm, and_comm (x ∉ s)] + +alias s_same_side_comm ↔ s_same_side.symm _ + +lemma w_opp_side_comm {s : affine_subspace R P} {x y : P} : + s.w_opp_side x y ↔ s.w_opp_side y x := +begin + split, + { rintro ⟨p₁, hp₁, p₂, hp₂, h⟩, + refine ⟨p₂, hp₂, p₁, hp₁, _⟩, + rwa [same_ray_comm, ←same_ray_neg_iff, neg_vsub_eq_vsub_rev, neg_vsub_eq_vsub_rev] }, + { rintro ⟨p₁, hp₁, p₂, hp₂, h⟩, + refine ⟨p₂, hp₂, p₁, hp₁, _⟩, + rwa [same_ray_comm, ←same_ray_neg_iff, neg_vsub_eq_vsub_rev, neg_vsub_eq_vsub_rev] } +end + +alias w_opp_side_comm ↔ w_opp_side.symm _ + +lemma s_opp_side_comm {s : affine_subspace R P} {x y : P} : + s.s_opp_side x y ↔ s.s_opp_side y x := +by rw [s_opp_side, s_opp_side, w_opp_side_comm, and_comm (x ∉ s)] + +alias s_opp_side_comm ↔ s_opp_side.symm _ + +lemma not_w_same_side_bot (x y : P) : ¬ (⊥ : affine_subspace R P).w_same_side x y := +by simp [w_same_side, not_mem_bot] + +lemma not_s_same_side_bot (x y : P) : ¬ (⊥ : affine_subspace R P).s_same_side x y := +λ h, not_w_same_side_bot x y h.w_same_side + +lemma not_w_opp_side_bot (x y : P) : ¬ (⊥ : affine_subspace R P).w_opp_side x y := +by simp [w_opp_side, not_mem_bot] + +lemma not_s_opp_side_bot (x y : P) : ¬ (⊥ : affine_subspace R P).s_opp_side x y := +λ h, not_w_opp_side_bot x y h.w_opp_side + +@[simp] lemma w_same_side_self_iff {s : affine_subspace R P} {x : P} : + s.w_same_side x x ↔ (s : set P).nonempty := +⟨λ h, h.nonempty, λ ⟨p, hp⟩, ⟨p, hp, p, hp, same_ray.rfl⟩⟩ + +lemma s_same_side_self_iff {s : affine_subspace R P} {x : P} : + s.s_same_side x x ↔ (s : set P).nonempty ∧ x ∉ s := +⟨λ ⟨h, hx, _⟩, ⟨w_same_side_self_iff.1 h, hx⟩, λ ⟨h, hx⟩, ⟨w_same_side_self_iff.2 h, hx, hx⟩⟩ + +lemma w_same_side_of_left_mem {s : affine_subspace R P} {x : P} (y : P) (hx : x ∈ s) : + s.w_same_side x y := +begin + refine ⟨x, hx, x, hx, _⟩, + simp +end + +lemma w_same_side_of_right_mem {s : affine_subspace R P} (x : P) {y : P} (hy : y ∈ s) : + s.w_same_side x y := +(w_same_side_of_left_mem x hy).symm + +lemma w_opp_side_of_left_mem {s : affine_subspace R P} {x : P} (y : P) (hx : x ∈ s) : + s.w_opp_side x y := +begin + refine ⟨x, hx, x, hx, _⟩, + simp +end + +lemma w_opp_side_of_right_mem {s : affine_subspace R P} (x : P) {y : P} (hy : y ∈ s) : + s.w_opp_side x y := +(w_opp_side_of_left_mem x hy).symm + +lemma w_same_side_vadd_left_iff {s : affine_subspace R P} {x y : P} {v : V} + (hv : v ∈ s.direction) : s.w_same_side (v +ᵥ x) y ↔ s.w_same_side x y := +begin + split, + { rintro ⟨p₁, hp₁, p₂, hp₂, h⟩, + refine ⟨-v +ᵥ p₁, + affine_subspace.vadd_mem_of_mem_direction (submodule.neg_mem _ hv) hp₁, p₂, hp₂, _⟩, + rwa [vsub_vadd_eq_vsub_sub, sub_neg_eq_add, add_comm, ←vadd_vsub_assoc] }, + { rintro ⟨p₁, hp₁, p₂, hp₂, h⟩, + refine ⟨v +ᵥ p₁, + affine_subspace.vadd_mem_of_mem_direction hv hp₁, p₂, hp₂, _⟩, + rwa vadd_vsub_vadd_cancel_left } +end + +lemma w_same_side_vadd_right_iff {s : affine_subspace R P} {x y : P} {v : V} + (hv : v ∈ s.direction) : s.w_same_side x (v +ᵥ y) ↔ s.w_same_side x y := +by rw [w_same_side_comm, w_same_side_vadd_left_iff hv, w_same_side_comm] + +lemma s_same_side_vadd_left_iff {s : affine_subspace R P} {x y : P} {v : V} + (hv : v ∈ s.direction) : s.s_same_side (v +ᵥ x) y ↔ s.s_same_side x y := +by rw [s_same_side, s_same_side, w_same_side_vadd_left_iff hv, + vadd_mem_iff_mem_of_mem_direction hv] + +lemma s_same_side_vadd_right_iff {s : affine_subspace R P} {x y : P} {v : V} + (hv : v ∈ s.direction) : s.s_same_side x (v +ᵥ y) ↔ s.s_same_side x y := +by rw [s_same_side_comm, s_same_side_vadd_left_iff hv, s_same_side_comm] + +lemma w_opp_side_vadd_left_iff {s : affine_subspace R P} {x y : P} {v : V} + (hv : v ∈ s.direction) : s.w_opp_side (v +ᵥ x) y ↔ s.w_opp_side x y := +begin + split, + { rintro ⟨p₁, hp₁, p₂, hp₂, h⟩, + refine ⟨-v +ᵥ p₁, + affine_subspace.vadd_mem_of_mem_direction (submodule.neg_mem _ hv) hp₁, p₂, hp₂, _⟩, + rwa [vsub_vadd_eq_vsub_sub, sub_neg_eq_add, add_comm, ←vadd_vsub_assoc] }, + { rintro ⟨p₁, hp₁, p₂, hp₂, h⟩, + refine ⟨v +ᵥ p₁, + affine_subspace.vadd_mem_of_mem_direction hv hp₁, p₂, hp₂, _⟩, + rwa vadd_vsub_vadd_cancel_left } +end + +lemma w_opp_side_vadd_right_iff {s : affine_subspace R P} {x y : P} {v : V} + (hv : v ∈ s.direction) : s.w_opp_side x (v +ᵥ y) ↔ s.w_opp_side x y := +by rw [w_opp_side_comm, w_opp_side_vadd_left_iff hv, w_opp_side_comm] + +lemma s_opp_side_vadd_left_iff {s : affine_subspace R P} {x y : P} {v : V} + (hv : v ∈ s.direction) : s.s_opp_side (v +ᵥ x) y ↔ s.s_opp_side x y := +by rw [s_opp_side, s_opp_side, w_opp_side_vadd_left_iff hv, + vadd_mem_iff_mem_of_mem_direction hv] + +lemma s_opp_side_vadd_right_iff {s : affine_subspace R P} {x y : P} {v : V} + (hv : v ∈ s.direction) : s.s_opp_side x (v +ᵥ y) ↔ s.s_opp_side x y := +by rw [s_opp_side_comm, s_opp_side_vadd_left_iff hv, s_opp_side_comm] + +lemma w_same_side_smul_vsub_vadd_left {s : affine_subspace R P} {p₁ p₂ : P} (x : P) + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) {t : R} (ht : 0 ≤ t) : s.w_same_side (t • (x -ᵥ p₁) +ᵥ p₂) x := +begin + refine ⟨p₂, hp₂, p₁, hp₁, _⟩, + rw vadd_vsub, + exact same_ray_nonneg_smul_left _ ht +end + +lemma w_same_side_smul_vsub_vadd_right {s : affine_subspace R P} {p₁ p₂ : P} (x : P) + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) {t : R} (ht : 0 ≤ t) : s.w_same_side x (t • (x -ᵥ p₁) +ᵥ p₂) := +(w_same_side_smul_vsub_vadd_left x hp₁ hp₂ ht).symm + +lemma w_same_side_line_map_left {s : affine_subspace R P} {x : P} (y : P) (h : x ∈ s) {t : R} + (ht : 0 ≤ t) : s.w_same_side (line_map x y t) y := +w_same_side_smul_vsub_vadd_left y h h ht + +lemma w_same_side_line_map_right {s : affine_subspace R P} {x : P} (y : P) (h : x ∈ s) {t : R} + (ht : 0 ≤ t) : s.w_same_side y (line_map x y t) := +(w_same_side_line_map_left y h ht).symm + +lemma w_opp_side_smul_vsub_vadd_left {s : affine_subspace R P} {p₁ p₂ : P} (x : P) + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) {t : R} (ht : t ≤ 0) : s.w_opp_side (t • (x -ᵥ p₁) +ᵥ p₂) x := +begin + refine ⟨p₂, hp₂, p₁, hp₁, _⟩, + rw [vadd_vsub, ←neg_neg t, neg_smul, ←smul_neg, neg_vsub_eq_vsub_rev], + exact same_ray_nonneg_smul_left _ (neg_nonneg.2 ht) +end + +lemma w_opp_side_smul_vsub_vadd_right {s : affine_subspace R P} {p₁ p₂ : P} (x : P) + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) {t : R} (ht : t ≤ 0) : s.w_opp_side x (t • (x -ᵥ p₁) +ᵥ p₂) := +(w_opp_side_smul_vsub_vadd_left x hp₁ hp₂ ht).symm + +lemma w_opp_side_line_map_left {s : affine_subspace R P} {x : P} (y : P) (h : x ∈ s) {t : R} + (ht : t ≤ 0) : s.w_opp_side (line_map x y t) y := +w_opp_side_smul_vsub_vadd_left y h h ht + +lemma w_opp_side_line_map_right {s : affine_subspace R P} {x : P} (y : P) (h : x ∈ s) {t : R} + (ht : t ≤ 0) : s.w_opp_side y (line_map x y t) := +(w_opp_side_line_map_left y h ht).symm + +lemma _root_.wbtw.w_same_side₂₃ {s : affine_subspace R P} {x y z : P} (h : wbtw R x y z) + (hx : x ∈ s) : s.w_same_side y z := +begin + rcases h with ⟨t, ⟨ht0, -⟩, rfl⟩, + exact w_same_side_line_map_left z hx ht0 +end + +lemma _root_.wbtw.w_same_side₃₂ {s : affine_subspace R P} {x y z : P} (h : wbtw R x y z) + (hx : x ∈ s) : s.w_same_side z y := +(h.w_same_side₂₃ hx).symm + +lemma _root_.wbtw.w_same_side₁₂ {s : affine_subspace R P} {x y z : P} (h : wbtw R x y z) + (hz : z ∈ s) : s.w_same_side x y := +h.symm.w_same_side₃₂ hz + +lemma _root_.wbtw.w_same_side₂₁ {s : affine_subspace R P} {x y z : P} (h : wbtw R x y z) + (hz : z ∈ s) : s.w_same_side y x := +h.symm.w_same_side₂₃ hz + +lemma _root_.wbtw.w_opp_side₁₃ {s : affine_subspace R P} {x y z : P} (h : wbtw R x y z) + (hy : y ∈ s) : s.w_opp_side x z := +begin + rcases h with ⟨t, ⟨ht0, ht1⟩, rfl⟩, + refine ⟨_, hy, _, hy, _⟩, + rcases ht1.lt_or_eq with ht1' | rfl, swap, { simp }, + rcases ht0.lt_or_eq with ht0' | rfl, swap, { simp }, + refine or.inr (or.inr ⟨1 - t, t, sub_pos.2 ht1', ht0', _⟩), + simp_rw [line_map_apply, vadd_vsub_assoc, vsub_vadd_eq_vsub_sub, ←neg_vsub_eq_vsub_rev z x, + vsub_self, zero_sub, ←neg_one_smul R (z -ᵥ x), ←add_smul, smul_neg, ←neg_smul, + smul_smul], + ring_nf +end + +lemma _root_.wbtw.w_opp_side₃₁ {s : affine_subspace R P} {x y z : P} (h : wbtw R x y z) + (hy : y ∈ s) : s.w_opp_side z x := +h.symm.w_opp_side₁₃ hy + +end strict_ordered_comm_ring + +section linear_ordered_field + +variables [linear_ordered_field R] [add_comm_group V] [module R V] [add_torsor V P] +variables [add_comm_group V'] [module R V'] [add_torsor V' P'] + +include V + +variables {R} + +@[simp] lemma w_opp_side_self_iff {s : affine_subspace R P} {x : P} : s.w_opp_side x x ↔ x ∈ s := +begin + split, + { rintro ⟨p₁, hp₁, p₂, hp₂, h⟩, + obtain ⟨a, -, -, -, -, h₁, -⟩ := h.exists_eq_smul_add, + rw [add_comm, vsub_add_vsub_cancel, ←eq_vadd_iff_vsub_eq] at h₁, + rw h₁, + exact s.smul_vsub_vadd_mem a hp₂ hp₁ hp₁ }, + { exact λ h, ⟨x, h, x, h, same_ray.rfl⟩ } +end + +lemma not_s_opp_side_self (s : affine_subspace R P) (x : P) : ¬s.s_opp_side x x := +by simp [s_opp_side] + +lemma w_same_side_iff_exists_left {s : affine_subspace R P} {x y p₁ : P} (h : p₁ ∈ s) : + s.w_same_side x y ↔ x ∈ s ∨ ∃ p₂ ∈ s, same_ray R (x -ᵥ p₁) (y -ᵥ p₂) := +begin + split, + { rintro ⟨p₁', hp₁', p₂', hp₂', h0 | h0 | ⟨r₁, r₂, hr₁, hr₂, hr⟩⟩, + { rw vsub_eq_zero_iff_eq at h0, + rw h0, + exact or.inl hp₁' }, + { refine or.inr ⟨p₂', hp₂', _⟩, + rw h0, + exact same_ray.zero_right _ }, + { refine or.inr ⟨(r₁ / r₂) • (p₁ -ᵥ p₁') +ᵥ p₂', s.smul_vsub_vadd_mem _ h hp₁' hp₂', + or.inr (or.inr ⟨r₁, r₂, hr₁, hr₂, _⟩)⟩, + rw [vsub_vadd_eq_vsub_sub, smul_sub, ←hr, smul_smul, mul_div_cancel' _ hr₂.ne.symm, + ←smul_sub, vsub_sub_vsub_cancel_right] } }, + { rintro (h' | h'), + { exact w_same_side_of_left_mem y h' }, + { exact ⟨p₁, h, h'⟩ } } +end + +lemma w_same_side_iff_exists_right {s : affine_subspace R P} {x y p₂ : P} (h : p₂ ∈ s) : + s.w_same_side x y ↔ y ∈ s ∨ ∃ p₁ ∈ s, same_ray R (x -ᵥ p₁) (y -ᵥ p₂) := +begin + rw [w_same_side_comm, w_same_side_iff_exists_left h], + simp_rw same_ray_comm +end + +lemma s_same_side_iff_exists_left {s : affine_subspace R P} {x y p₁ : P} (h : p₁ ∈ s) : + s.s_same_side x y ↔ x ∉ s ∧ y ∉ s ∧ ∃ p₂ ∈ s, same_ray R (x -ᵥ p₁) (y -ᵥ p₂) := +begin + rw [s_same_side, and_comm, w_same_side_iff_exists_left h, and_assoc, and.congr_right_iff], + intro hx, + rw or_iff_right hx +end + +lemma s_same_side_iff_exists_right {s : affine_subspace R P} {x y p₂ : P} (h : p₂ ∈ s) : + s.s_same_side x y ↔ x ∉ s ∧ y ∉ s ∧ ∃ p₁ ∈ s, same_ray R (x -ᵥ p₁) (y -ᵥ p₂) := +begin + rw [s_same_side_comm, s_same_side_iff_exists_left h, ←and_assoc, and_comm (y ∉ s), and_assoc], + simp_rw same_ray_comm +end + +lemma w_opp_side_iff_exists_left {s : affine_subspace R P} {x y p₁ : P} (h : p₁ ∈ s) : + s.w_opp_side x y ↔ x ∈ s ∨ ∃ p₂ ∈ s, same_ray R (x -ᵥ p₁) (p₂ -ᵥ y) := +begin + split, + { rintro ⟨p₁', hp₁', p₂', hp₂', h0 | h0 | ⟨r₁, r₂, hr₁, hr₂, hr⟩⟩, + { rw vsub_eq_zero_iff_eq at h0, + rw h0, + exact or.inl hp₁' }, + { refine or.inr ⟨p₂', hp₂', _⟩, + rw h0, + exact same_ray.zero_right _ }, + { refine or.inr ⟨(-r₁ / r₂) • (p₁ -ᵥ p₁') +ᵥ p₂', s.smul_vsub_vadd_mem _ h hp₁' hp₂', + or.inr (or.inr ⟨r₁, r₂, hr₁, hr₂, _⟩)⟩, + rw [vadd_vsub_assoc, smul_add, ←hr, smul_smul, neg_div, mul_neg, + mul_div_cancel' _ hr₂.ne.symm, neg_smul, neg_add_eq_sub, ←smul_sub, + vsub_sub_vsub_cancel_right] } }, + { rintro (h' | h'), + { exact w_opp_side_of_left_mem y h' }, + { exact ⟨p₁, h, h'⟩ } } +end + +lemma w_opp_side_iff_exists_right {s : affine_subspace R P} {x y p₂ : P} (h : p₂ ∈ s) : + s.w_opp_side x y ↔ y ∈ s ∨ ∃ p₁ ∈ s, same_ray R (x -ᵥ p₁) (p₂ -ᵥ y) := +begin + rw [w_opp_side_comm, w_opp_side_iff_exists_left h], + split, + { rintro (hy | ⟨p, hp, hr⟩), { exact or.inl hy }, + refine or.inr ⟨p, hp, _⟩, + rwa [same_ray_comm, ←same_ray_neg_iff, neg_vsub_eq_vsub_rev, neg_vsub_eq_vsub_rev] }, + { rintro (hy | ⟨p, hp, hr⟩), { exact or.inl hy }, + refine or.inr ⟨p, hp, _⟩, + rwa [same_ray_comm, ←same_ray_neg_iff, neg_vsub_eq_vsub_rev, neg_vsub_eq_vsub_rev] } +end + +lemma s_opp_side_iff_exists_left {s : affine_subspace R P} {x y p₁ : P} (h : p₁ ∈ s) : + s.s_opp_side x y ↔ x ∉ s ∧ y ∉ s ∧ ∃ p₂ ∈ s, same_ray R (x -ᵥ p₁) (p₂ -ᵥ y) := +begin + rw [s_opp_side, and_comm, w_opp_side_iff_exists_left h, and_assoc, and.congr_right_iff], + intro hx, + rw or_iff_right hx +end + +lemma s_opp_side_iff_exists_right {s : affine_subspace R P} {x y p₂ : P} (h : p₂ ∈ s) : + s.s_opp_side x y ↔ x ∉ s ∧ y ∉ s ∧ ∃ p₁ ∈ s, same_ray R (x -ᵥ p₁) (p₂ -ᵥ y) := +begin + rw [s_opp_side, and_comm, w_opp_side_iff_exists_right h, and_assoc, and.congr_right_iff, + and.congr_right_iff], + rintro hx hy, + rw or_iff_right hy +end + +lemma w_same_side.trans {s : affine_subspace R P} {x y z : P} (hxy : s.w_same_side x y) + (hyz : s.w_same_side y z) (hy : y ∉ s) : s.w_same_side x z := +begin + rcases hxy with ⟨p₁, hp₁, p₂, hp₂, hxy⟩, + rw [w_same_side_iff_exists_left hp₂, or_iff_right hy] at hyz, + rcases hyz with ⟨p₃, hp₃, hyz⟩, + refine ⟨p₁, hp₁, p₃, hp₃, hxy.trans hyz _⟩, + refine λ h, false.elim _, + rw vsub_eq_zero_iff_eq at h, + exact hy (h.symm ▸ hp₂) +end + +lemma w_same_side.trans_s_same_side {s : affine_subspace R P} {x y z : P} + (hxy : s.w_same_side x y) (hyz : s.s_same_side y z) : s.w_same_side x z := +hxy.trans hyz.1 hyz.2.1 + +lemma w_same_side.trans_w_opp_side {s : affine_subspace R P} {x y z : P} (hxy : s.w_same_side x y) + (hyz : s.w_opp_side y z) (hy : y ∉ s) : s.w_opp_side x z := +begin + rcases hxy with ⟨p₁, hp₁, p₂, hp₂, hxy⟩, + rw [w_opp_side_iff_exists_left hp₂, or_iff_right hy] at hyz, + rcases hyz with ⟨p₃, hp₃, hyz⟩, + refine ⟨p₁, hp₁, p₃, hp₃, hxy.trans hyz _⟩, + refine λ h, false.elim _, + rw vsub_eq_zero_iff_eq at h, + exact hy (h.symm ▸ hp₂) +end + +lemma w_same_side.trans_s_opp_side {s : affine_subspace R P} {x y z : P} (hxy : s.w_same_side x y) + (hyz : s.s_opp_side y z) : s.w_opp_side x z := +hxy.trans_w_opp_side hyz.1 hyz.2.1 + +lemma s_same_side.trans_w_same_side {s : affine_subspace R P} {x y z : P} + (hxy : s.s_same_side x y) (hyz : s.w_same_side y z) : s.w_same_side x z := +(hyz.symm.trans_s_same_side hxy.symm).symm + +lemma s_same_side.trans {s : affine_subspace R P} {x y z : P} (hxy : s.s_same_side x y) + (hyz : s.s_same_side y z) : s.s_same_side x z := +⟨hxy.w_same_side.trans_s_same_side hyz, hxy.2.1, hyz.2.2⟩ + +lemma s_same_side.trans_w_opp_side {s : affine_subspace R P} {x y z : P} (hxy : s.s_same_side x y) + (hyz : s.w_opp_side y z) : s.w_opp_side x z := +hxy.w_same_side.trans_w_opp_side hyz hxy.2.2 + +lemma s_same_side.trans_s_opp_side {s : affine_subspace R P} {x y z : P} (hxy : s.s_same_side x y) + (hyz : s.s_opp_side y z) : s.s_opp_side x z := +⟨hxy.trans_w_opp_side hyz.1, hxy.2.1, hyz.2.2⟩ + +lemma w_opp_side.trans_w_same_side {s : affine_subspace R P} {x y z : P} (hxy : s.w_opp_side x y) + (hyz : s.w_same_side y z) (hy : y ∉ s) : s.w_opp_side x z := +(hyz.symm.trans_w_opp_side hxy.symm hy).symm + +lemma w_opp_side.trans_s_same_side {s : affine_subspace R P} {x y z : P} (hxy : s.w_opp_side x y) + (hyz : s.s_same_side y z) : s.w_opp_side x z := +hxy.trans_w_same_side hyz.1 hyz.2.1 + +lemma w_opp_side.trans {s : affine_subspace R P} {x y z : P} (hxy : s.w_opp_side x y) + (hyz : s.w_opp_side y z) (hy : y ∉ s) : s.w_same_side x z := +begin + rcases hxy with ⟨p₁, hp₁, p₂, hp₂, hxy⟩, + rw [w_opp_side_iff_exists_left hp₂, or_iff_right hy] at hyz, + rcases hyz with ⟨p₃, hp₃, hyz⟩, + rw [←same_ray_neg_iff, neg_vsub_eq_vsub_rev, neg_vsub_eq_vsub_rev] at hyz, + refine ⟨p₁, hp₁, p₃, hp₃, hxy.trans hyz _⟩, + refine λ h, false.elim _, + rw vsub_eq_zero_iff_eq at h, + exact hy (h ▸ hp₂) +end + +lemma w_opp_side.trans_s_opp_side {s : affine_subspace R P} {x y z : P} (hxy : s.w_opp_side x y) + (hyz : s.s_opp_side y z) : s.w_same_side x z := +hxy.trans hyz.1 hyz.2.1 + +lemma s_opp_side.trans_w_same_side {s : affine_subspace R P} {x y z : P} (hxy : s.s_opp_side x y) + (hyz : s.w_same_side y z) : s.w_opp_side x z := +(hyz.symm.trans_s_opp_side hxy.symm).symm + +lemma s_opp_side.trans_s_same_side {s : affine_subspace R P} {x y z : P} (hxy : s.s_opp_side x y) + (hyz : s.s_same_side y z) : s.s_opp_side x z := +(hyz.symm.trans_s_opp_side hxy.symm).symm + +lemma s_opp_side.trans_w_opp_side {s : affine_subspace R P} {x y z : P} (hxy : s.s_opp_side x y) + (hyz : s.w_opp_side y z) : s.w_same_side x z := +(hyz.symm.trans_s_opp_side hxy.symm).symm + +lemma s_opp_side.trans {s : affine_subspace R P} {x y z : P} (hxy : s.s_opp_side x y) + (hyz : s.s_opp_side y z) : s.s_same_side x z := +⟨hxy.trans_w_opp_side hyz.1, hxy.2.1, hyz.2.2⟩ + +lemma w_same_side_and_w_opp_side_iff {s : affine_subspace R P} {x y : P} : + (s.w_same_side x y ∧ s.w_opp_side x y) ↔ (x ∈ s ∨ y ∈ s) := +begin + split, + { rintro ⟨hs, ho⟩, + rw w_opp_side_comm at ho, + by_contra h, + rw not_or_distrib at h, + exact h.1 (w_opp_side_self_iff.1 (hs.trans_w_opp_side ho h.2)) }, + { rintro (h | h), + { exact ⟨w_same_side_of_left_mem y h, w_opp_side_of_left_mem y h⟩ }, + { exact ⟨w_same_side_of_right_mem x h, w_opp_side_of_right_mem x h⟩ } } +end + +lemma w_same_side.not_s_opp_side {s : affine_subspace R P} {x y : P} (h : s.w_same_side x y) : + ¬s.s_opp_side x y := +begin + intro ho, + have hxy := w_same_side_and_w_opp_side_iff.1 ⟨h, ho.1⟩, + rcases hxy with hx | hy, + { exact ho.2.1 hx }, + { exact ho.2.2 hy } +end + +lemma s_same_side.not_w_opp_side {s : affine_subspace R P} {x y : P} (h : s.s_same_side x y) : + ¬s.w_opp_side x y := +begin + intro ho, + have hxy := w_same_side_and_w_opp_side_iff.1 ⟨h.1, ho⟩, + rcases hxy with hx | hy, + { exact h.2.1 hx }, + { exact h.2.2 hy } +end + +lemma s_same_side.not_s_opp_side {s : affine_subspace R P} {x y : P} (h : s.s_same_side x y) : + ¬s.s_opp_side x y := +λ ho, h.not_w_opp_side ho.1 + +lemma w_opp_side.not_s_same_side {s : affine_subspace R P} {x y : P} (h : s.w_opp_side x y) : + ¬s.s_same_side x y := +λ hs, hs.not_w_opp_side h + +lemma s_opp_side.not_w_same_side {s : affine_subspace R P} {x y : P} (h : s.s_opp_side x y) : + ¬s.w_same_side x y := +λ hs, hs.not_s_opp_side h + +lemma s_opp_side.not_s_same_side {s : affine_subspace R P} {x y : P} (h : s.s_opp_side x y) : + ¬s.s_same_side x y := +λ hs, h.not_w_same_side hs.1 + +lemma w_opp_side_iff_exists_wbtw {s : affine_subspace R P} {x y : P} : + s.w_opp_side x y ↔ ∃ p ∈ s, wbtw R x p y := +begin + refine ⟨λ h, _, λ ⟨p, hp, h⟩, h.w_opp_side₁₃ hp⟩, + rcases h with ⟨p₁, hp₁, p₂, hp₂, (h | h | ⟨r₁, r₂, hr₁, hr₂, h⟩)⟩, + { rw vsub_eq_zero_iff_eq at h, + rw h, + exact ⟨p₁, hp₁, wbtw_self_left _ _ _⟩ }, + { rw vsub_eq_zero_iff_eq at h, + rw ←h, + exact ⟨p₂, hp₂, wbtw_self_right _ _ _⟩ }, + { refine ⟨line_map x y (r₂ / (r₁ + r₂)), _, _⟩, + { rw [line_map_apply, ←vsub_vadd x p₁, ←vsub_vadd y p₂, vsub_vadd_eq_vsub_sub, + vadd_vsub_assoc, ←vadd_assoc, vadd_eq_add], + convert s.smul_vsub_vadd_mem (r₂ / (r₁ + r₂)) hp₂ hp₁ hp₁, + rw [add_comm (y -ᵥ p₂), smul_sub, smul_add, add_sub_assoc, add_assoc, add_right_eq_self, + div_eq_inv_mul, ←neg_vsub_eq_vsub_rev, smul_neg, ←smul_smul, ←h, smul_smul, + ←neg_smul, ←sub_smul, ←div_eq_inv_mul, ←div_eq_inv_mul, ←neg_div, ←sub_div, + sub_eq_add_neg, ←neg_add, neg_div, div_self (left.add_pos hr₁ hr₂).ne.symm, + neg_one_smul, neg_add_self] }, + { exact set.mem_image_of_mem _ ⟨div_nonneg hr₂.le (left.add_pos hr₁ hr₂).le, + div_le_one_of_le (le_add_of_nonneg_left hr₁.le) + (left.add_pos hr₁ hr₂).le⟩ } } +end + +lemma s_opp_side.exists_sbtw {s : affine_subspace R P} {x y : P} (h : s.s_opp_side x y) : + ∃ p ∈ s, sbtw R x p y := +begin + obtain ⟨p, hp, hw⟩ := w_opp_side_iff_exists_wbtw.1 h.w_opp_side, + refine ⟨p, hp, hw, _, _⟩, + { rintro rfl, + exact h.2.1 hp }, + { rintro rfl, + exact h.2.2 hp }, +end + +lemma _root_.sbtw.s_opp_side_of_not_mem_of_mem {s : affine_subspace R P} {x y z : P} + (h : sbtw R x y z) (hx : x ∉ s) (hy : y ∈ s) : s.s_opp_side x z := +begin + refine ⟨h.wbtw.w_opp_side₁₃ hy, hx, λ hz, hx _⟩, + rcases h with ⟨⟨t, ⟨ht0, ht1⟩, rfl⟩, hyx, hyz⟩, + rw line_map_apply at hy, + have ht : t ≠ 1, { rintro rfl, simpa [line_map_apply] using hyz }, + have hy' := vsub_mem_direction hy hz, + rw [vadd_vsub_assoc, ←neg_vsub_eq_vsub_rev z, ←neg_one_smul R (z -ᵥ x), ←add_smul, + ←sub_eq_add_neg, s.direction.smul_mem_iff (sub_ne_zero_of_ne ht)] at hy', + rwa vadd_mem_iff_mem_of_mem_direction (submodule.smul_mem _ _ hy') at hy +end + +lemma s_same_side_smul_vsub_vadd_left {s : affine_subspace R P} {x p₁ p₂ : P} (hx : x ∉ s) + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) {t : R} (ht : 0 < t) : s.s_same_side (t • (x -ᵥ p₁) +ᵥ p₂) x := +begin + refine ⟨w_same_side_smul_vsub_vadd_left x hp₁ hp₂ ht.le, λ h, hx _, hx⟩, + rwa [vadd_mem_iff_mem_direction _ hp₂, s.direction.smul_mem_iff ht.ne.symm, + vsub_right_mem_direction_iff_mem hp₁] at h +end + +lemma s_same_side_smul_vsub_vadd_right {s : affine_subspace R P} {x p₁ p₂ : P} (hx : x ∉ s) + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) {t : R} (ht : 0 < t) : s.s_same_side x (t • (x -ᵥ p₁) +ᵥ p₂) := +(s_same_side_smul_vsub_vadd_left hx hp₁ hp₂ ht).symm + +lemma s_same_side_line_map_left {s : affine_subspace R P} {x y : P} (hx : x ∈ s) (hy : y ∉ s) + {t : R} (ht : 0 < t) : s.s_same_side (line_map x y t) y := +s_same_side_smul_vsub_vadd_left hy hx hx ht + +lemma s_same_side_line_map_right {s : affine_subspace R P} {x y : P} (hx : x ∈ s) (hy : y ∉ s) + {t : R} (ht : 0 < t) : s.s_same_side y (line_map x y t) := +(s_same_side_line_map_left hx hy ht).symm + +lemma s_opp_side_smul_vsub_vadd_left {s : affine_subspace R P} {x p₁ p₂ : P} (hx : x ∉ s) + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) {t : R} (ht : t < 0) : s.s_opp_side (t • (x -ᵥ p₁) +ᵥ p₂) x := +begin + refine ⟨w_opp_side_smul_vsub_vadd_left x hp₁ hp₂ ht.le, λ h, hx _, hx⟩, + rwa [vadd_mem_iff_mem_direction _ hp₂, s.direction.smul_mem_iff ht.ne, + vsub_right_mem_direction_iff_mem hp₁] at h +end + +lemma s_opp_side_smul_vsub_vadd_right {s : affine_subspace R P} {x p₁ p₂ : P} (hx : x ∉ s) + (hp₁ : p₁ ∈ s) (hp₂ : p₂ ∈ s) {t : R} (ht : t < 0) : s.s_opp_side x (t • (x -ᵥ p₁) +ᵥ p₂) := +(s_opp_side_smul_vsub_vadd_left hx hp₁ hp₂ ht).symm + +lemma s_opp_side_line_map_left {s : affine_subspace R P} {x y : P} (hx : x ∈ s) (hy : y ∉ s) + {t : R} (ht : t < 0) : s.s_opp_side (line_map x y t) y := +s_opp_side_smul_vsub_vadd_left hy hx hx ht + +lemma s_opp_side_line_map_right {s : affine_subspace R P} {x y : P} (hx : x ∈ s) (hy : y ∉ s) + {t : R} (ht : t < 0) : s.s_opp_side y (line_map x y t) := +(s_opp_side_line_map_left hx hy ht).symm + +lemma set_of_w_same_side_eq_image2 {s : affine_subspace R P} {x p : P} (hx : x ∉ s) (hp : p ∈ s) : + {y | s.w_same_side x y} = set.image2 (λ (t : R) q, t • (x -ᵥ p) +ᵥ q) (set.Ici 0) s := +begin + ext y, + simp_rw [set.mem_set_of, set.mem_image2, set.mem_Ici, mem_coe], + split, + { rw [w_same_side_iff_exists_left hp, or_iff_right hx], + rintro ⟨p₂, hp₂, h | h | ⟨r₁, r₂, hr₁, hr₂, h⟩⟩, + { rw vsub_eq_zero_iff_eq at h, + exact false.elim (hx (h.symm ▸ hp)) }, + { rw vsub_eq_zero_iff_eq at h, + refine ⟨0, p₂, le_refl _, hp₂, _⟩, + simp [h] }, + { refine ⟨r₁ / r₂, p₂, (div_pos hr₁ hr₂).le, hp₂, _⟩, + rw [div_eq_inv_mul, ←smul_smul, h, smul_smul, inv_mul_cancel hr₂.ne.symm, one_smul, + vsub_vadd] } }, + { rintro ⟨t, p', ht, hp', rfl⟩, + exact w_same_side_smul_vsub_vadd_right x hp hp' ht } +end + +lemma set_of_s_same_side_eq_image2 {s : affine_subspace R P} {x p : P} (hx : x ∉ s) (hp : p ∈ s) : + {y | s.s_same_side x y} = set.image2 (λ (t : R) q, t • (x -ᵥ p) +ᵥ q) (set.Ioi 0) s := +begin + ext y, + simp_rw [set.mem_set_of, set.mem_image2, set.mem_Ioi, mem_coe], + split, + { rw s_same_side_iff_exists_left hp, + rintro ⟨-, hy, p₂, hp₂, h | h | ⟨r₁, r₂, hr₁, hr₂, h⟩⟩, + { rw vsub_eq_zero_iff_eq at h, + exact false.elim (hx (h.symm ▸ hp)) }, + { rw vsub_eq_zero_iff_eq at h, + exact false.elim (hy (h.symm ▸ hp₂)) }, + { refine ⟨r₁ / r₂, p₂, div_pos hr₁ hr₂, hp₂, _⟩, + rw [div_eq_inv_mul, ←smul_smul, h, smul_smul, inv_mul_cancel hr₂.ne.symm, one_smul, + vsub_vadd] } }, + { rintro ⟨t, p', ht, hp', rfl⟩, + exact s_same_side_smul_vsub_vadd_right hx hp hp' ht } +end + +lemma set_of_w_opp_side_eq_image2 {s : affine_subspace R P} {x p : P} (hx : x ∉ s) (hp : p ∈ s) : + {y | s.w_opp_side x y} = set.image2 (λ (t : R) q, t • (x -ᵥ p) +ᵥ q) (set.Iic 0) s := +begin + ext y, + simp_rw [set.mem_set_of, set.mem_image2, set.mem_Iic, mem_coe], + split, + { rw [w_opp_side_iff_exists_left hp, or_iff_right hx], + rintro ⟨p₂, hp₂, h | h | ⟨r₁, r₂, hr₁, hr₂, h⟩⟩, + { rw vsub_eq_zero_iff_eq at h, + exact false.elim (hx (h.symm ▸ hp)) }, + { rw vsub_eq_zero_iff_eq at h, + refine ⟨0, p₂, le_refl _, hp₂, _⟩, + simp [h] }, + { refine ⟨-r₁ / r₂, p₂, (div_neg_of_neg_of_pos (left.neg_neg_iff.2 hr₁) hr₂).le, hp₂, _⟩, + rw [div_eq_inv_mul, ←smul_smul, neg_smul, h, smul_neg, smul_smul, + inv_mul_cancel hr₂.ne.symm, one_smul, neg_vsub_eq_vsub_rev, vsub_vadd] } }, + { rintro ⟨t, p', ht, hp', rfl⟩, + exact w_opp_side_smul_vsub_vadd_right x hp hp' ht } +end + +lemma set_of_s_opp_side_eq_image2 {s : affine_subspace R P} {x p : P} (hx : x ∉ s) (hp : p ∈ s) : + {y | s.s_opp_side x y} = set.image2 (λ (t : R) q, t • (x -ᵥ p) +ᵥ q) (set.Iio 0) s := +begin + ext y, + simp_rw [set.mem_set_of, set.mem_image2, set.mem_Iio, mem_coe], + split, + { rw s_opp_side_iff_exists_left hp, + rintro ⟨-, hy, p₂, hp₂, h | h | ⟨r₁, r₂, hr₁, hr₂, h⟩⟩, + { rw vsub_eq_zero_iff_eq at h, + exact false.elim (hx (h.symm ▸ hp)) }, + { rw vsub_eq_zero_iff_eq at h, + exact false.elim (hy (h ▸ hp₂)) }, + { refine ⟨-r₁ / r₂, p₂, div_neg_of_neg_of_pos (left.neg_neg_iff.2 hr₁) hr₂, hp₂, _⟩, + rw [div_eq_inv_mul, ←smul_smul, neg_smul, h, smul_neg, smul_smul, + inv_mul_cancel hr₂.ne.symm, one_smul, neg_vsub_eq_vsub_rev, vsub_vadd] } }, + { rintro ⟨t, p', ht, hp', rfl⟩, + exact s_opp_side_smul_vsub_vadd_right hx hp hp' ht } +end + +end linear_ordered_field + +end affine_subspace From 65fbcb3366f775241c26ceec4ce86a26a04d24ff Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Wed, 12 Oct 2022 16:16:47 +0000 Subject: [PATCH 728/828] feat(number_theory/number_field): add map_mem_ring_of_integers (#16754) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We add `number_field.map_mem_ring_of_integers`, the image of an element of the ring of integers under `f : K →ₐ[ℚ] L` lies in the ring of integer. In particular, the Galois group preserve the ring of integers. From flt-regular --- src/data/polynomial/algebra_map.lean | 13 ++++---- src/field_theory/adjoin.lean | 2 +- src/field_theory/galois.lean | 2 +- .../is_alg_closed/algebraic_closure.lean | 2 +- src/field_theory/minpoly.lean | 2 +- src/field_theory/normal.lean | 19 +++++------ src/field_theory/separable.lean | 2 +- src/field_theory/splitting_field.lean | 2 +- src/linear_algebra/charpoly/basic.lean | 2 +- src/number_theory/number_field/basic.lean | 4 +++ .../number_field/embeddings.lean | 2 +- src/ring_theory/integral_closure.lean | 32 ++++++++++++------- src/ring_theory/norm.lean | 2 +- src/ring_theory/trace.lean | 2 +- 14 files changed, 51 insertions(+), 37 deletions(-) diff --git a/src/data/polynomial/algebra_map.lean b/src/data/polynomial/algebra_map.lean index aea595e742ab4..fcf99d2387bdc 100644 --- a/src/data/polynomial/algebra_map.lean +++ b/src/data/polynomial/algebra_map.lean @@ -217,17 +217,16 @@ theorem eval_unique (φ : R[X] →ₐ[R] A) (p) : φ p = eval₂ (algebra_map R A) (φ X) p := by rw [← aeval_def, aeval_alg_hom, aeval_X_left, alg_hom.comp_id] -theorem aeval_alg_hom_apply (f : A →ₐ[R] B) (x : A) (p : R[X]) : +theorem aeval_alg_hom_apply {F : Type*} [alg_hom_class F R A B] (f : F) (x : A) (p : R[X]) : aeval (f x) p = f (aeval x p) := -alg_hom.ext_iff.1 (aeval_alg_hom f x) p +begin + refine polynomial.induction_on p (by simp) (λ p q hp hq, _) (by simp), + rw [map_add, hp, hq, ← map_add, ← map_add] +end theorem aeval_alg_equiv (f : A ≃ₐ[R] B) (x : A) : aeval (f x) = (f : A →ₐ[R] B).comp (aeval x) := aeval_alg_hom (f : A →ₐ[R] B) x -theorem aeval_alg_equiv_apply (f : A ≃ₐ[R] B) (x : A) (p : R[X]) : - aeval (f x) p = f (aeval x p) := -aeval_alg_hom_apply (f : A →ₐ[R] B) x p - lemma aeval_algebra_map_apply_eq_algebra_map_eval (x : R) (p : R[X]) : aeval (algebra_map R A x) p = algebra_map R A (p.eval x) := aeval_alg_hom_apply (algebra.of_id R A) x p @@ -237,7 +236,7 @@ aeval_alg_hom_apply (algebra.of_id R A) x p @[simp] lemma aeval_fn_apply {X : Type*} (g : R[X]) (f : X → R) (x : X) : ((aeval f) g) x = aeval (f x) g := -(aeval_alg_hom_apply (pi.eval_alg_hom _ _ x) f g).symm +(aeval_alg_hom_apply (pi.eval_alg_hom R (λ _, R) x) f g).symm @[norm_cast] lemma aeval_subalgebra_coe (g : R[X]) {A : Type*} [semiring A] [algebra R A] (s : subalgebra R A) (f : s) : diff --git a/src/field_theory/adjoin.lean b/src/field_theory/adjoin.lean index 88ef68c1027d7..a4eac4651624a 100644 --- a/src/field_theory/adjoin.lean +++ b/src/field_theory/adjoin.lean @@ -886,7 +886,7 @@ end⟩ /-- Extend a lift `x : lifts F E K` to an element `s : E` whose conjugates are all in `K` -/ noncomputable def lifts.lift_of_splits (x : lifts F E K) {s : E} (h1 : is_integral F s) (h2 : (minpoly F s).splits (algebra_map F K)) : lifts F E K := -let h3 : is_integral x.1 s := is_integral_of_is_scalar_tower s h1 in +let h3 : is_integral x.1 s := is_integral_of_is_scalar_tower h1 in let key : (minpoly x.1 s).splits x.2.to_ring_hom := splits_of_splits_of_dvd _ (map_ne_zero (minpoly.ne_zero h1)) ((splits_map_iff _ _).mpr (by {convert h2, exact ring_hom.ext (λ y, x.2.commutes y)})) diff --git a/src/field_theory/galois.lean b/src/field_theory/galois.lean index 132caae2c2b46..6672ba57090be 100644 --- a/src/field_theory/galois.lean +++ b/src/field_theory/galois.lean @@ -339,7 +339,7 @@ lemma of_separable_splitting_field_aux fintype.card (K⟮x⟯.restrict_scalars F →ₐ[F] E) = fintype.card (K →ₐ[F] E) * finrank K K⟮x⟯ := begin - have h : is_integral K x := is_integral_of_is_scalar_tower x + have h : is_integral K x := is_integral_of_is_scalar_tower (is_integral_of_noetherian (is_noetherian.iff_fg.2 hFE) x), have h1 : p ≠ 0 := λ hp, by rwa [hp, polynomial.map_zero, polynomial.roots_zero] at hx, have h2 : (minpoly K x) ∣ p.map (algebra_map F K), diff --git a/src/field_theory/is_alg_closed/algebraic_closure.lean b/src/field_theory/is_alg_closed/algebraic_closure.lean index 94f4671ff1e03..cb76e9dcb4067 100644 --- a/src/field_theory/is_alg_closed/algebraic_closure.lean +++ b/src/field_theory/is_alg_closed/algebraic_closure.lean @@ -257,7 +257,7 @@ def of_step_hom (n) : step k n →ₐ[k] algebraic_closure k := theorem is_algebraic : algebra.is_algebraic k (algebraic_closure k) := λ z, is_algebraic_iff_is_integral.2 $ let ⟨n, x, hx⟩ := exists_of_step k z in -hx ▸ is_integral_alg_hom (of_step_hom k n) (step.is_integral k n x) +hx ▸ map_is_integral (of_step_hom k n) (step.is_integral k n x) instance : is_alg_closure k (algebraic_closure k) := ⟨algebraic_closure.is_alg_closed k, is_algebraic k⟩ diff --git a/src/field_theory/minpoly.lean b/src/field_theory/minpoly.lean index 0d9f3e85d189a..56265f431de63 100644 --- a/src/field_theory/minpoly.lean +++ b/src/field_theory/minpoly.lean @@ -454,7 +454,7 @@ begin let L := fraction_ring S, rw [← gcd_domain_eq_field_fractions K L hs], refine minpoly.eq_of_algebra_map_eq (is_fraction_ring.injective S L) - (is_integral_of_is_scalar_tower _ hs) rfl + (is_integral_of_is_scalar_tower hs) rfl end variable [no_zero_smul_divisors R S] diff --git a/src/field_theory/normal.lean b/src/field_theory/normal.lean index e84972970bbd9..954c7b0cff9b9 100644 --- a/src/field_theory/normal.lean +++ b/src/field_theory/normal.lean @@ -86,7 +86,7 @@ lemma normal.tower_top_of_normal [h : normal F E] : normal K E := normal_iff.2 $ λ x, begin cases h.out x with hx hhx, rw algebra_map_eq F K E at hhx, - exact ⟨is_integral_of_is_scalar_tower x hx, polynomial.splits_of_splits_of_dvd (algebra_map K E) + exact ⟨is_integral_of_is_scalar_tower hx, polynomial.splits_of_splits_of_dvd (algebra_map K E) (polynomial.map_ne_zero (minpoly.ne_zero hx)) ((polynomial.splits_map_iff (algebra_map F K) (algebra_map K E)).mpr hhx) (minpoly.dvd_map_of_is_scalar_tower F K x)⟩, @@ -97,7 +97,7 @@ lemma alg_hom.normal_bijective [h : normal F E] (ϕ : E →ₐ[F] K) : function. { letI : algebra E K := ϕ.to_ring_hom.to_algebra, obtain ⟨h1, h2⟩ := h.out (algebra_map K E x), cases minpoly.mem_range_of_degree_eq_one E x (h2.def.resolve_left (minpoly.ne_zero h1) - (minpoly.irreducible (is_integral_of_is_scalar_tower x + (minpoly.irreducible (is_integral_of_is_scalar_tower ((is_integral_algebra_map_iff (algebra_map K E).injective).mp h1))) (minpoly.dvd E x ((algebra_map K E).injective (by { rw [ring_hom.map_zero, aeval_map_algebra_map, ← aeval_algebra_map_apply], @@ -109,7 +109,7 @@ variables {F} {E} {E' : Type*} [field E'] [algebra F E'] lemma normal.of_alg_equiv [h : normal F E] (f : E ≃ₐ[F] E') : normal F E' := normal_iff.2 $ λ x, begin cases h.out (f.symm x) with hx hhx, - have H := is_integral_alg_hom f.to_alg_hom hx, + have H := map_is_integral f.to_alg_hom hx, rw [alg_equiv.to_alg_hom_eq_coe, alg_equiv.coe_alg_hom, alg_equiv.apply_symm_apply] at H, use H, apply polynomial.splits_of_splits_of_dvd (algebra_map F E') (minpoly.ne_zero hx), @@ -245,11 +245,12 @@ def alg_hom.restrict_normal_aux [h : normal F E] : rintros x ⟨y, ⟨z, hy⟩, hx⟩, rw [←hx, ←hy], apply minpoly.mem_range_of_degree_eq_one E, - exact or.resolve_left (h.splits z).def (minpoly.ne_zero (h.is_integral z)) - (minpoly.irreducible $ is_integral_of_is_scalar_tower _ $ - is_integral_alg_hom ϕ $ is_integral_alg_hom _ $ h.is_integral z) - (minpoly.dvd E _ $ by rw [aeval_map_algebra_map, aeval_alg_hom_apply, aeval_alg_hom_apply, - minpoly.aeval, alg_hom.map_zero, alg_hom.map_zero]) }⟩, + refine or.resolve_left (h.splits z).def (minpoly.ne_zero (h.is_integral z)) + (minpoly.irreducible _) (minpoly.dvd E _ (by simp [aeval_alg_hom_apply])), + simp only [alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom], + suffices : is_integral F _, + { exact is_integral_of_is_scalar_tower this }, + exact map_is_integral ϕ (map_is_integral (to_alg_hom F E K₁) (h.is_integral z)) }⟩, map_zero' := subtype.ext ϕ.map_zero, map_one' := subtype.ext ϕ.map_one, map_add' := λ x y, subtype.ext (ϕ.map_add x y), @@ -330,7 +331,7 @@ noncomputable def alg_hom.lift_normal [h : normal F E] : E →ₐ[F] E := @intermediate_field.alg_hom_mk_adjoin_splits' _ _ _ _ _ _ _ ((is_scalar_tower.to_alg_hom F K₂ E).comp ϕ).to_ring_hom.to_algebra _ (intermediate_field.adjoin_univ _ _) - (λ x hx, ⟨is_integral_of_is_scalar_tower x (h.out x).1, + (λ x hx, ⟨is_integral_of_is_scalar_tower (h.out x).1, splits_of_splits_of_dvd _ (map_ne_zero (minpoly.ne_zero (h.out x).1)) (by { rw [splits_map_iff, ←is_scalar_tower.algebra_map_eq], exact (h.out x).2 }) (minpoly.dvd_map_of_is_scalar_tower F K₁ x)⟩) diff --git a/src/field_theory/separable.lean b/src/field_theory/separable.lean index 1b889d1f75e7e..59a727bd64889 100644 --- a/src/field_theory/separable.lean +++ b/src/field_theory/separable.lean @@ -486,7 +486,7 @@ variables (F K E : Type*) [field F] [field K] [field E] [algebra F K] [algebra F [algebra K E] [is_scalar_tower F K E] lemma is_separable_tower_top_of_is_separable [is_separable F E] : is_separable K E := -⟨λ x, is_integral_of_is_scalar_tower x (is_separable.is_integral F x), +⟨λ x, is_integral_of_is_scalar_tower (is_separable.is_integral F x), λ x, (is_separable.separable F x).map.of_dvd (minpoly.dvd_map_of_is_scalar_tower _ _ _)⟩ lemma is_separable_tower_bot_of_is_separable [h : is_separable F E] : is_separable F K := diff --git a/src/field_theory/splitting_field.lean b/src/field_theory/splitting_field.lean index 7babfe9e330f2..3a86734c03cd6 100644 --- a/src/field_theory/splitting_field.lean +++ b/src/field_theory/splitting_field.lean @@ -467,7 +467,7 @@ begin (submodule.fg_iff_finite_dimensional _).1 (fg_adjoin_of_finite s.finite_to_set H3)).of_subalgebra_to_submodule, letI := field_of_finite_dimensional F (algebra.adjoin F (↑s : set K)), - have H5 : is_integral (algebra.adjoin F (↑s : set K)) a := is_integral_of_is_scalar_tower a H1, + have H5 : is_integral (algebra.adjoin F (↑s : set K)) a := is_integral_of_is_scalar_tower H1, have H6 : (minpoly (algebra.adjoin F (↑s : set K)) a).splits (algebra_map (algebra.adjoin F (↑s : set K)) L), { refine polynomial.splits_of_splits_of_dvd _ diff --git a/src/linear_algebra/charpoly/basic.lean b/src/linear_algebra/charpoly/basic.lean index 630f07ee36211..fe2808b923912 100644 --- a/src/linear_algebra/charpoly/basic.lean +++ b/src/linear_algebra/charpoly/basic.lean @@ -61,7 +61,7 @@ to the linear map itself, is zero. See `matrix.aeval_self_charpoly` for the equivalent statement about matrices. -/ lemma aeval_self_charpoly : aeval f f.charpoly = 0 := begin - apply (linear_equiv.map_eq_zero_iff (alg_equiv_matrix _).to_linear_equiv).1, + apply (linear_equiv.map_eq_zero_iff (alg_equiv_matrix (choose_basis R M)).to_linear_equiv).1, rw [alg_equiv.to_linear_equiv_apply, ← alg_equiv.coe_alg_hom, ← polynomial.aeval_alg_hom_apply _ _ _, charpoly_def], exact aeval_self_charpoly _, diff --git a/src/number_theory/number_field/basic.lean b/src/number_theory/number_field/basic.lean index aaecadef620b5..f7dca7a2d261e 100644 --- a/src/number_theory/number_field/basic.lean +++ b/src/number_theory/number_field/basic.lean @@ -102,6 +102,10 @@ integral_closure.is_integrally_closed_of_finite_extension ℚ lemma is_integral_coe (x : 𝓞 K) : is_integral ℤ (x : K) := x.2 +lemma map_mem_ring_of_integers {F L : Type*} [field L] [char_zero K] [char_zero L] + [alg_hom_class F ℚ K L] (f : F) (x : 𝓞 K) : f x ∈ 𝓞 L := +(mem_ring_of_integers _ _).2 $ map_is_integral_int f $ ring_of_integers.is_integral_coe x + /-- The ring of integers of `K` are equivalent to any integral closure of `ℤ` in `K` -/ protected noncomputable def equiv (R : Type*) [comm_ring R] [algebra R K] [is_integral_closure R ℤ K] : 𝓞 K ≃+* R := diff --git a/src/number_theory/number_field/embeddings.lean b/src/number_theory/number_field/embeddings.lean index 53e4582bd1cce..0261b284955ef 100644 --- a/src/number_theory/number_field/embeddings.lean +++ b/src/number_theory/number_field/embeddings.lean @@ -95,7 +95,7 @@ begin have h_map_ℚ_minpoly := minpoly.gcd_domain_eq_field_fractions' ℚ hx.1, refine ⟨_, ⟨_, λ i, _⟩, (mem_root_set_iff (minpoly.ne_zero hx.1) x).2 (minpoly.aeval ℤ x)⟩, { rw [← (minpoly.monic hx.1).nat_degree_map (algebra_map ℤ ℚ), ← h_map_ℚ_minpoly], - exact minpoly.nat_degree_le (is_integral_of_is_scalar_tower x hx.1), }, + exact minpoly.nat_degree_le (is_integral_of_is_scalar_tower hx.1) }, rw [mem_Icc, ← abs_le, ← @int.cast_le ℝ], refine (eq.trans_le _ $ coeff_bdd_of_norm_le hx.2 i).trans (nat.le_ceil _), rw [h_map_ℚ_minpoly, coeff_map, eq_int_cast, int.norm_cast_rat, int.norm_eq_abs, int.cast_abs], diff --git a/src/ring_theory/integral_closure.lean b/src/ring_theory/integral_closure.lean index 83a6380292f74..075e5adc32769 100644 --- a/src/ring_theory/integral_closure.lean +++ b/src/ring_theory/integral_closure.lean @@ -114,15 +114,20 @@ variables {R A B S : Type*} variables [comm_ring R] [comm_ring A] [comm_ring B] [comm_ring S] variables [algebra R A] [algebra R B] (f : R →+* S) -theorem is_integral_alg_hom {A B : Type*} [ring A] [ring B] [algebra R A] [algebra R B] - (f : A →ₐ[R] B) {x : A} (hx : is_integral R x) : is_integral R (f x) := -let ⟨p, hp, hpx⟩ := -hx in ⟨p, hp, by rw [← aeval_def, aeval_alg_hom_apply, aeval_def, hpx, f.map_zero]⟩ +lemma map_is_integral {B C F : Type*} [ring B] [ring C] [algebra R B] [algebra A B] + [algebra R C] [is_scalar_tower R A B] [algebra A C] [is_scalar_tower R A C] {b : B} + [alg_hom_class F A B C] (f : F) (hb : is_integral R b) : is_integral R (f b) := +begin + obtain ⟨P, hP⟩ := hb, + refine ⟨P, hP.1, _⟩, + rw [← aeval_def, show (aeval (f b)) P = (aeval (f b)) (P.map (algebra_map R A)), by simp, + aeval_alg_hom_apply, aeval_map_algebra_map, aeval_def, hP.2, _root_.map_zero] +end theorem is_integral_alg_hom_iff {A B : Type*} [ring A] [ring B] [algebra R A] [algebra R B] (f : A →ₐ[R] B) (hf : function.injective f) {x : A} : is_integral R (f x) ↔ is_integral R x := begin - refine ⟨_, is_integral_alg_hom f⟩, + refine ⟨_, map_is_integral f⟩, rintros ⟨p, hp, hx⟩, use [p, hp], rwa [← f.comp_algebra_map, ← alg_hom.coe_to_ring_hom, ← polynomial.hom_eval₂, @@ -132,17 +137,22 @@ end @[simp] theorem is_integral_alg_equiv {A B : Type*} [ring A] [ring B] [algebra R A] [algebra R B] (f : A ≃ₐ[R] B) {x : A} : is_integral R (f x) ↔ is_integral R x := -⟨λ h, by simpa using is_integral_alg_hom f.symm.to_alg_hom h, is_integral_alg_hom f.to_alg_hom⟩ +⟨λ h, by simpa using map_is_integral f.symm.to_alg_hom h, map_is_integral f.to_alg_hom⟩ theorem is_integral_of_is_scalar_tower [algebra A B] [is_scalar_tower R A B] - (x : B) (hx : is_integral R x) : is_integral A x := + {x : B} (hx : is_integral R x) : is_integral A x := let ⟨p, hp, hpx⟩ := hx in ⟨p.map $ algebra_map R A, hp.map _, by rw [← aeval_def, aeval_map_algebra_map, aeval_def, hpx]⟩ +lemma map_is_integral_int {B C F : Type*} [ring B] [ring C] {b : B} + [ring_hom_class F B C] (f : F) (hb : is_integral ℤ b) : + is_integral ℤ (f b) := +map_is_integral (f : B →+* C).to_int_alg_hom hb + theorem is_integral_of_subring {x : A} (T : subring R) (hx : is_integral T x) : is_integral R x := -is_integral_of_is_scalar_tower x hx +is_integral_of_is_scalar_tower hx lemma is_integral.algebra_map [algebra A B] [is_scalar_tower R A B] {x : A} (h : is_integral R x) : @@ -503,9 +513,9 @@ begin rw subalgebra.mem_map, split, { rintros ⟨x, hx, rfl⟩, - exact is_integral_alg_hom f hx }, + exact map_is_integral f hx }, { intro hy, - use [f.symm y, is_integral_alg_hom (f.symm : B →ₐ[R] A) hy], + use [f.symm y, map_is_integral (f.symm : B →ₐ[R] A) hy], simp } end @@ -852,7 +862,7 @@ begin { rw [finset.mem_coe, frange, finset.mem_image] at hx, rcases hx with ⟨i, _, rfl⟩, rw coeff_map, - exact is_integral_alg_hom (is_scalar_tower.to_alg_hom R A B) (A_int _) }, + exact map_is_integral (is_scalar_tower.to_alg_hom R A B) (A_int _) }, { apply fg_adjoin_singleton_of_integral, exact is_integral_trans_aux _ pmonic hp } end diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index a4c3f9778722d..a9ee578b28db7 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -305,7 +305,7 @@ lemma is_integral_norm [algebra S L] [algebra S K] [is_scalar_tower S K L] [is_separable K L] [finite_dimensional K L] {x : L} (hx : _root_.is_integral S x) : _root_.is_integral S (norm K x) := begin - have hx' : _root_.is_integral K x := is_integral_of_is_scalar_tower _ hx, + have hx' : _root_.is_integral K x := is_integral_of_is_scalar_tower hx, rw [← is_integral_algebra_map_iff (algebra_map K (algebraic_closure L)).injective, norm_eq_prod_roots], { refine (is_integral.multiset_prod (λ y hy, _)).pow _, diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index 4ec7c28b8cceb..8aa1462475762 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -286,7 +286,7 @@ open polynomial lemma algebra.is_integral_trace [finite_dimensional L F] {x : F} (hx : _root_.is_integral R x) : _root_.is_integral R (algebra.trace L F x) := begin - have hx' : _root_.is_integral L x := is_integral_of_is_scalar_tower _ hx, + have hx' : _root_.is_integral L x := is_integral_of_is_scalar_tower hx, rw [← is_integral_algebra_map_iff (algebra_map L (algebraic_closure F)).injective, trace_eq_sum_roots], { refine (is_integral.multiset_sum _).nsmul _, From c9f6b6f0f91743b1661758867b960d91e2d4b2b8 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 12 Oct 2022 18:38:10 +0000 Subject: [PATCH 729/828] feat(analysis/calculus/monotone): a monotone function is differentiable almost everywhere (#16549) The proof is an almost direct consequence of results on differentiation of measures that are already in mathlib. --- src/analysis/calculus/monotone.lean | 255 +++++++++++++++++++++ src/linear_algebra/affine_space/slope.lean | 3 + src/measure_theory/covering/one_dim.lean | 64 ++++++ 3 files changed, 322 insertions(+) create mode 100644 src/analysis/calculus/monotone.lean create mode 100644 src/measure_theory/covering/one_dim.lean diff --git a/src/analysis/calculus/monotone.lean b/src/analysis/calculus/monotone.lean new file mode 100644 index 0000000000000..d7f179ff05491 --- /dev/null +++ b/src/analysis/calculus/monotone.lean @@ -0,0 +1,255 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import measure_theory.measure.lebesgue +import analysis.calculus.deriv +import measure_theory.covering.one_dim + +/-! +# Differentiability of monotone functions + +We show that a monotone function `f : ℝ → ℝ` is differentiable almost everywhere, in +`monotone.ae_differentiable_at`. (We also give a version for a function monotone on a set, in +`monotone_on.ae_differentiable_within_at`.) + +If the function `f` is continuous, this follows directly from general differentiation of measure +theorems. Let `μ` be the Stieltjes measure associated to `f`. Then, almost everywhere, +`μ [x, y] / Leb [x, y]` (resp. `μ [y, x] / Leb [y, x]`) converges to the Radon-Nikodym derivative +of `μ` with respect to Lebesgue when `y` tends to `x` in `(x, +∞)` (resp. `(-∞, x)`), by +`vitali_family.ae_tendsto_rn_deriv`. As `μ [x, y] = f y - f x` and `Leb [x, y] = y - x`, this +gives differentiability right away. + +When `f` is only monotone, the same argument works up to small adjustments, as the associated +Stieltjes measure satisfies `μ [x, y] = f (y^+) - f (x^-)` (the right and left limits of `f` at `y` +and `x` respectively). One argues that `f (x^-) = f x` almost everywhere (in fact away from a +countable set), and moreover `f ((y - (y-x)^2)^+) ≤ f y ≤ f (y^+)`. This is enough to deduce the +limit of `(f y - f x) / (y - x)` by a lower and upper approximation argument from the known +behavior of `μ [x, y]`. +-/ + +open set filter function metric measure_theory measure_theory.measure is_doubling_measure +open_locale topological_space + +/-- If `(f y - f x) / (y - x)` converges to a limit as `y` tends to `x`, then the same goes if +`y` is shifted a little bit, i.e., `f (y + (y-x)^2) - f x) / (y - x)` converges to the same limit. +This lemma contains a slightly more general version of this statement (where one considers +convergence along some subfilter, typically `𝓝[<] x` or `𝓝[>] x`) tailored to the application +to almost everywhere differentiability of monotone functions. -/ +lemma tendsto_apply_add_mul_sq_div_sub {f : ℝ → ℝ} {x a c d : ℝ} {l : filter ℝ} (hl : l ≤ 𝓝[≠] x) + (hf : tendsto (λ y, (f y - d) / (y - x)) l (𝓝 a)) + (h' : tendsto (λ y, y + c * (y-x)^2) l l) : + tendsto (λ y, (f (y + c * (y-x)^2) - d) / (y - x)) l (𝓝 a) := +begin + have L : tendsto (λ y, (y + c * (y - x)^2 - x) / (y - x)) l (𝓝 1), + { have : tendsto (λ y, (1 + c * (y - x))) l (𝓝 (1 + c * (x - x))), + { apply tendsto.mono_left _ (hl.trans nhds_within_le_nhds), + exact ((tendsto_id.sub_const x).const_mul c).const_add 1 }, + simp only [_root_.sub_self, add_zero, mul_zero] at this, + apply tendsto.congr' (eventually.filter_mono hl _) this, + filter_upwards [self_mem_nhds_within] with y hy, + field_simp [sub_ne_zero.2 hy], + ring }, + have Z := (hf.comp h').mul L, + rw mul_one at Z, + apply tendsto.congr' _ Z, + have : ∀ᶠ y in l, y + c * (y-x)^2 ≠ x := by apply tendsto.mono_right h' hl self_mem_nhds_within, + filter_upwards [this] with y hy, + field_simp [sub_ne_zero.2 hy], +end + +/-- A Stieltjes function is almost everywhere differentiable, with derivative equal to the +Radon-Nikodym derivative of the associated Stieltjes measure with respect to Lebesgue. -/ +lemma stieltjes_function.ae_has_deriv_at (f : stieltjes_function) : + ∀ᵐ x, has_deriv_at f (rn_deriv f.measure volume x).to_real x := +begin + /- Denote by `μ` the Stieltjes measure associated to `f`. + The general theorem `vitali_family.ae_tendsto_rn_deriv` ensures that `μ [x, y] / (y - x)` tends + to the Radon-Nikodym derivative as `y` tends to `x` from the right. As `μ [x, y] = f y - f (x^-)` + and `f (x^-) = f x` almost everywhere, this gives differentiability on the right. + On the left, `μ [y, x] / (x - y)` again tends to the Radon-Nikodym derivative. + As `μ [y, x] = f x - f (y^-)`, this is not exactly the right result, so one uses a sandwiching + argument to deduce the convergence for `(f x - f y) / (x - y)`. -/ + filter_upwards [ + vitali_family.ae_tendsto_rn_deriv (vitali_family (volume : measure ℝ) 1) f.measure, + rn_deriv_lt_top f.measure volume, f.countable_left_lim_ne.ae_not_mem volume] with x hx h'x h''x, + -- Limit on the right, following from differentiation of measures + have L1 : tendsto (λ y, (f y - f x) / (y - x)) + (𝓝[>] x) (𝓝 ((rn_deriv f.measure volume x).to_real)), + { apply tendsto.congr' _ + ((ennreal.tendsto_to_real h'x.ne).comp (hx.comp (real.tendsto_Icc_vitali_family_right x))), + filter_upwards [self_mem_nhds_within], + rintros y (hxy : x < y), + simp only [comp_app, stieltjes_function.measure_Icc, real.volume_Icc, not_not.1 h''x], + rw [← ennreal.of_real_div_of_pos (sub_pos.2 hxy), ennreal.to_real_of_real], + exact div_nonneg (sub_nonneg.2 (f.mono hxy.le)) (sub_pos.2 hxy).le }, + -- Limit on the left, following from differentiation of measures. Its form is not exactly the one + -- we need, due to the appearance of a left limit. + have L2 : tendsto (λ y, (left_lim f y - f x) / (y - x)) + (𝓝[<] x) (𝓝 ((rn_deriv f.measure volume x).to_real)), + { apply tendsto.congr' _ + ((ennreal.tendsto_to_real h'x.ne).comp (hx.comp (real.tendsto_Icc_vitali_family_left x))), + filter_upwards [self_mem_nhds_within], + rintros y (hxy : y < x), + simp only [comp_app, stieltjes_function.measure_Icc, real.volume_Icc], + rw [← ennreal.of_real_div_of_pos (sub_pos.2 hxy), ennreal.to_real_of_real, ← neg_neg (y - x), + div_neg, neg_div', neg_sub, neg_sub], + exact div_nonneg (sub_nonneg.2 (f.mono.left_lim_le hxy.le)) (sub_pos.2 hxy).le }, + -- Shifting a little bit the limit on the left, by `(y - x)^2`. + have L3 : tendsto (λ y, (left_lim f (y + 1 * (y - x)^2) - f x) / (y - x)) + (𝓝[<] x) (𝓝 ((rn_deriv f.measure volume x).to_real)), + { apply tendsto_apply_add_mul_sq_div_sub (nhds_left'_le_nhds_ne x) L2, + apply tendsto_nhds_within_of_tendsto_nhds_of_eventually_within, + { apply tendsto.mono_left _ nhds_within_le_nhds, + have : tendsto (λ (y : ℝ), y + 1 * (y - x) ^ 2) (𝓝 x) (𝓝 (x + 1 * (x - x)^2)) := + tendsto_id.add (((tendsto_id.sub_const x).pow 2).const_mul 1), + simpa using this }, + { have : Ioo (x - 1) x ∈ 𝓝[<] x, + { apply Ioo_mem_nhds_within_Iio, exact ⟨by linarith, le_refl _⟩ }, + filter_upwards [this], + rintros y ⟨hy : x - 1 < y, h'y : y < x⟩, + rw mem_Iio, + nlinarith } }, + -- Deduce the correct limit on the left, by sandwiching. + have L4 : tendsto (λ y, (f y - f x) / (y - x)) + (𝓝[<] x) (𝓝 ((rn_deriv f.measure volume x).to_real)), + { apply tendsto_of_tendsto_of_tendsto_of_le_of_le' L3 L2, + { filter_upwards [self_mem_nhds_within], + rintros y (hy : y < x), + refine div_le_div_of_nonpos_of_le (by linarith) ((sub_le_sub_iff_right _).2 _), + apply f.mono.le_left_lim, + have : 0 < (x - y)^2 := sq_pos_of_pos (sub_pos.2 hy), + linarith }, + { filter_upwards [self_mem_nhds_within], + rintros y (hy : y < x), + refine div_le_div_of_nonpos_of_le (by linarith) _, + simpa only [sub_le_sub_iff_right] using f.mono.left_lim_le (le_refl y) } }, + -- prove the result by splitting into left and right limits. + rw [has_deriv_at_iff_tendsto_slope, slope_fun_def_field, ← nhds_left'_sup_nhds_right', + tendsto_sup], + exact ⟨L4, L1⟩ +end + +/-- A monotone function is almost everywhere differentiable, with derivative equal to the +Radon-Nikodym derivative of the associated Stieltjes measure with respect to Lebesgue. -/ +lemma monotone.ae_has_deriv_at {f : ℝ → ℝ} (hf : monotone f) : + ∀ᵐ x, has_deriv_at f (rn_deriv hf.stieltjes_function.measure volume x).to_real x := +begin + /- We already know that the Stieltjes function associated to `f` (i.e., `g : x ↦ f (x^+)`) is + differentiable almost everywhere. We reduce to this statement by sandwiching values of `f` with + values of `g`, by shifting with `(y - x)^2` (which has no influence on the relevant + scale `y - x`.)-/ + filter_upwards [hf.stieltjes_function.ae_has_deriv_at, + hf.countable_not_continuous_at.ae_not_mem volume] with x hx h'x, + have A : hf.stieltjes_function x = f x, + { rw [not_not, hf.continuous_at_iff_left_lim_eq_right_lim] at h'x, + apply le_antisymm _ (hf.le_right_lim (le_refl _)), + rw ← h'x, + exact hf.left_lim_le (le_refl _) }, + rw [has_deriv_at_iff_tendsto_slope, (nhds_left'_sup_nhds_right' x).symm, tendsto_sup, + slope_fun_def_field, A] at hx, + -- prove differentiability on the right, by sandwiching with values of `g` + have L1 : tendsto (λ y, (f y - f x) / (y - x)) (𝓝[>] x) + (𝓝 (rn_deriv hf.stieltjes_function.measure volume x).to_real), + { -- limit of a helper function, with a small shift compared to `g` + have : tendsto (λ y, (hf.stieltjes_function (y + (-1) * (y-x)^2) - f x) / (y - x)) (𝓝[>] x) + (𝓝 (rn_deriv hf.stieltjes_function.measure volume x).to_real), + { apply tendsto_apply_add_mul_sq_div_sub (nhds_right'_le_nhds_ne x) hx.2, + apply tendsto_nhds_within_of_tendsto_nhds_of_eventually_within, + { apply tendsto.mono_left _ nhds_within_le_nhds, + have : tendsto (λ (y : ℝ), y + (-1) * (y - x) ^ 2) (𝓝 x) (𝓝 (x + (-1) * (x - x)^2)) := + tendsto_id.add (((tendsto_id.sub_const x).pow 2).const_mul (-1)), + simpa using this }, + { have : Ioo x (x+1) ∈ 𝓝[>] x, + { apply Ioo_mem_nhds_within_Ioi, exact ⟨le_refl _, by linarith⟩ }, + filter_upwards [this], + rintros y ⟨hy : x < y, h'y : y < x + 1⟩, + rw mem_Ioi, + nlinarith } }, + -- apply the sandwiching argument, with the helper function and `g` + apply tendsto_of_tendsto_of_tendsto_of_le_of_le' this hx.2, + { filter_upwards [self_mem_nhds_within], + rintros y (hy : x < y), + have : 0 < (y - x)^2, from sq_pos_of_pos (sub_pos.2 hy), + apply div_le_div_of_le_of_nonneg _ (sub_pos.2 hy).le, + exact (sub_le_sub_iff_right _).2 (hf.right_lim_le (by linarith)) }, + { filter_upwards [self_mem_nhds_within], + rintros y (hy : x < y), + apply div_le_div_of_le_of_nonneg _ (sub_pos.2 hy).le, + exact (sub_le_sub_iff_right _).2 (hf.le_right_lim (le_refl y)) } }, + -- prove differentiability on the left, by sandwiching with values of `g` + have L2 : tendsto (λ y, (f y - f x) / (y - x)) (𝓝[<] x) + (𝓝 (rn_deriv hf.stieltjes_function.measure volume x).to_real), + { -- limit of a helper function, with a small shift compared to `g` + have : tendsto (λ y, (hf.stieltjes_function (y + (-1) * (y-x)^2) - f x) / (y - x)) (𝓝[<] x) + (𝓝 (rn_deriv hf.stieltjes_function.measure volume x).to_real), + { apply tendsto_apply_add_mul_sq_div_sub (nhds_left'_le_nhds_ne x) hx.1, + apply tendsto_nhds_within_of_tendsto_nhds_of_eventually_within, + { apply tendsto.mono_left _ nhds_within_le_nhds, + have : tendsto (λ (y : ℝ), y + (-1) * (y - x) ^ 2) (𝓝 x) (𝓝 (x + (-1) * (x - x)^2)) := + tendsto_id.add (((tendsto_id.sub_const x).pow 2).const_mul (-1)), + simpa using this }, + { have : Ioo (x - 1) x ∈ 𝓝[<] x, + { apply Ioo_mem_nhds_within_Iio, exact ⟨by linarith, le_refl _⟩ }, + filter_upwards [this], + rintros y ⟨hy : x - 1 < y, h'y : y < x⟩, + rw mem_Iio, + nlinarith } }, + -- apply the sandwiching argument, with `g` and the helper function + apply tendsto_of_tendsto_of_tendsto_of_le_of_le' hx.1 this, + { filter_upwards [self_mem_nhds_within], + rintros y (hy : y < x), + apply div_le_div_of_nonpos_of_le (sub_neg.2 hy).le, + exact (sub_le_sub_iff_right _).2 (hf.le_right_lim (le_refl _)) }, + { filter_upwards [self_mem_nhds_within], + rintros y (hy : y < x), + have : 0 < (y - x)^2, from sq_pos_of_neg (sub_neg.2 hy), + apply div_le_div_of_nonpos_of_le (sub_neg.2 hy).le, + exact (sub_le_sub_iff_right _).2 (hf.right_lim_le (by linarith)) } }, + -- conclude global differentiability + rw [has_deriv_at_iff_tendsto_slope, slope_fun_def_field, (nhds_left'_sup_nhds_right' x).symm, + tendsto_sup], + exact ⟨L2, L1⟩ +end + +/-- A monotone real function is differentiable Lebesgue-almost everywhere. -/ +theorem monotone.ae_differentiable_at {f : ℝ → ℝ} (hf : monotone f) : + ∀ᵐ x, differentiable_at ℝ f x := +by filter_upwards [hf.ae_has_deriv_at] with x hx using hx.differentiable_at + +/-- A real function which is monotone on a set is differentiable Lebesgue-almost everywhere on +this set. This version does not assume that `s` is measurable. For a formulation with +`volume.restrict s` assuming that `s` is measurable, see `monotone_on.ae_differentiable_within_at`. +-/ +theorem monotone_on.ae_differentiable_within_at_of_mem + {f : ℝ → ℝ} {s : set ℝ} (hf : monotone_on f s) : + ∀ᵐ x, x ∈ s → differentiable_within_at ℝ f s x := +begin + /- We use a global monotone extension of `f`, and argue that this extension is differentiable + almost everywhere. Such an extension need not exist (think of `1/x` on `(0, +∞)`), but it exists + if one restricts first the function to a compact interval `[a, b]`. -/ + apply ae_of_mem_of_ae_of_mem_inter_Ioo, + assume a b as bs hab, + obtain ⟨g, hg, gf⟩ : ∃ (g : ℝ → ℝ), monotone g ∧ eq_on f g (s ∩ Icc a b) := + monotone_on.exists_monotone_extension (hf.mono (inter_subset_left s (Icc a b))) + ⟨⟨as, ⟨le_rfl, hab.le⟩⟩, λ x hx, hx.2.1⟩ ⟨⟨bs, ⟨hab.le, le_rfl⟩⟩, λ x hx, hx.2.2⟩, + filter_upwards [hg.ae_differentiable_at] with x hx, + assume h'x, + apply hx.differentiable_within_at.congr_of_eventually_eq _ (gf ⟨h'x.1, h'x.2.1.le, h'x.2.2.le⟩), + have : Ioo a b ∈ 𝓝[s] x, from nhds_within_le_nhds (Ioo_mem_nhds h'x.2.1 h'x.2.2), + filter_upwards [self_mem_nhds_within, this] with y hy h'y, + exact gf ⟨hy, h'y.1.le, h'y.2.le⟩, +end + +/-- A real function which is monotone on a set is differentiable Lebesgue-almost everywhere on +this set. This version assumes that `s` is measurable and uses `volume.restrict s`. +For a formulation without measurability assumption, +see `monotone_on.ae_differentiable_within_at_of_mem`. -/ +theorem monotone_on.ae_differentiable_within_at + {f : ℝ → ℝ} {s : set ℝ} (hf : monotone_on f s) (hs : measurable_set s) : + ∀ᵐ x ∂(volume.restrict s), differentiable_within_at ℝ f s x := +begin + rw ae_restrict_iff' hs, + exact hf.ae_differentiable_within_at_of_mem +end diff --git a/src/linear_algebra/affine_space/slope.lean b/src/linear_algebra/affine_space/slope.lean index 9613b06504f49..4a4e6574f8a81 100644 --- a/src/linear_algebra/affine_space/slope.lean +++ b/src/linear_algebra/affine_space/slope.lean @@ -36,6 +36,9 @@ omit E lemma slope_def_field (f : k → k) (a b : k) : slope f a b = (f b - f a) / (b - a) := (div_eq_inv_mul _ _).symm +lemma slope_fun_def_field (f : k → k) (a : k) : slope f a = λ b, (f b - f a) / (b - a) := +(div_eq_inv_mul _ _).symm + @[simp] lemma slope_same (f : k → PE) (a : k) : (slope f a a : E) = 0 := by rw [slope, sub_self, inv_zero, zero_smul] diff --git a/src/measure_theory/covering/one_dim.lean b/src/measure_theory/covering/one_dim.lean new file mode 100644 index 0000000000000..9476116be57c0 --- /dev/null +++ b/src/measure_theory/covering/one_dim.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2022 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import measure_theory.covering.density_theorem +import measure_theory.measure.haar_lebesgue + +/-! +# Covering theorems for Lebesgue measure in one dimension + +We have a general theory of covering theorems for doubling measures, developed notably +in `density_theorems.lean`. In this file, we expand the API for this theory in one dimension, +by showing that intervals belong to the relevant Vitali family. +-/ + +open set measure_theory is_doubling_measure filter +open_locale topological_space + +namespace real + +lemma Icc_mem_vitali_family_at_right {x y : ℝ} (hxy : x < y) : + Icc x y ∈ (vitali_family (volume : measure ℝ) 1).sets_at x := +begin + rw Icc_eq_closed_ball, + refine closed_ball_mem_vitali_family_of_dist_le_mul _ _ (by linarith), + rw [dist_comm, real.dist_eq, abs_of_nonneg]; + linarith, +end + +lemma tendsto_Icc_vitali_family_right (x : ℝ) : + tendsto (λ y, Icc x y) (𝓝[>] x) ((vitali_family (volume : measure ℝ) 1).filter_at x) := +begin + refine (vitali_family.tendsto_filter_at_iff _).2 ⟨_, _⟩, + { filter_upwards [self_mem_nhds_within] with y hy using Icc_mem_vitali_family_at_right hy }, + { assume ε εpos, + have : x ∈ Ico x (x + ε) := ⟨le_refl _, by linarith⟩, + filter_upwards [Icc_mem_nhds_within_Ioi this] with y hy, + rw closed_ball_eq_Icc, + exact Icc_subset_Icc (by linarith) hy.2 } +end + +lemma Icc_mem_vitali_family_at_left {x y : ℝ} (hxy : x < y) : + Icc x y ∈ (vitali_family (volume : measure ℝ) 1).sets_at y := +begin + rw Icc_eq_closed_ball, + refine closed_ball_mem_vitali_family_of_dist_le_mul _ _ (by linarith), + rw [real.dist_eq, abs_of_nonneg]; + linarith, +end + +lemma tendsto_Icc_vitali_family_left (x : ℝ) : + tendsto (λ y, Icc y x) (𝓝[<] x) ((vitali_family (volume : measure ℝ) 1).filter_at x) := +begin + refine (vitali_family.tendsto_filter_at_iff _).2 ⟨_, _⟩, + { filter_upwards [self_mem_nhds_within] with y hy using Icc_mem_vitali_family_at_left hy }, + { assume ε εpos, + have : x ∈ Ioc (x - ε) x := ⟨by linarith, le_refl _⟩, + filter_upwards [Icc_mem_nhds_within_Iio this] with y hy, + rw closed_ball_eq_Icc, + exact Icc_subset_Icc hy.1 (by linarith) } +end + +end real From 4dc9ec772a916d906ba81dca9963760c42f200c7 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Wed, 12 Oct 2022 20:19:08 +0000 Subject: [PATCH 730/828] feat(data/finset/prod): image projections out of finset product (#16936) Co-authored-by: Bhavik Mehta --- src/data/finset/prod.lean | 20 ++++++++++++++++++++ 1 file changed, 20 insertions(+) diff --git a/src/data/finset/prod.lean b/src/data/finset/prod.lean index 58e887e9e010b..534df4236002b 100644 --- a/src/data/finset/prod.lean +++ b/src/data/finset/prod.lean @@ -46,6 +46,18 @@ lemma mk_mem_product (ha : a ∈ s) (hb : b ∈ t) : (a, b) ∈ s ×ˢ t := mem_ (↑(s ×ˢ t) : set (α × β)) = s ×ˢ t := set.ext $ λ x, finset.mem_product +lemma subset_product_image_fst [decidable_eq α] : (s ×ˢ t).image prod.fst ⊆ s := +λ i, by simp [mem_image] {contextual := tt} + +lemma subset_product_image_snd [decidable_eq β] : (s ×ˢ t).image prod.snd ⊆ t := +λ i, by simp [mem_image] {contextual := tt} + +lemma product_image_fst [decidable_eq α] (ht : t.nonempty) : (s ×ˢ t).image prod.fst = s := +by { ext i, simp [mem_image, ht.bex] } + +lemma product_image_snd [decidable_eq β] (ht : s.nonempty) : (s ×ˢ t).image prod.snd = t := +by { ext i, simp [mem_image, ht.bex] } + lemma subset_product [decidable_eq α] [decidable_eq β] {s : finset (α × β)} : s ⊆ s.image prod.fst ×ˢ s.image prod.snd := λ p hp, mem_product.2 ⟨mem_image_of_mem _ hp, mem_image_of_mem _ hp⟩ @@ -82,6 +94,14 @@ lemma filter_product (p : α → Prop) (q : β → Prop) [decidable_pred p] [dec by { ext ⟨a, b⟩, simp only [mem_filter, mem_product], exact and_and_and_comm (a ∈ s) (b ∈ t) (p a) (q b) } +lemma filter_product_left (p : α → Prop) [decidable_pred p] : + (s ×ˢ t).filter (λ (x : α × β), p x.1) = s.filter p ×ˢ t := +by simpa using filter_product p (λ _, true) + +lemma filter_product_right (q : β → Prop) [decidable_pred q] : + (s ×ˢ t).filter (λ (x : α × β), q x.2) = s ×ˢ t.filter q := +by simpa using filter_product (λ _ : α, true) q + lemma filter_product_card (s : finset α) (t : finset β) (p : α → Prop) (q : β → Prop) [decidable_pred p] [decidable_pred q] : ((s ×ˢ t).filter (λ (x : α × β), p x.1 ↔ q x.2)).card = From f45d33795824b768fd642b7e89d9e37e101649db Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 12 Oct 2022 23:21:35 +0000 Subject: [PATCH 731/828] feat(measure_theory/covering/differentiation): Lebesgue differentiation theorem (#16906) Given a doubling measure, then at almost every `x` the average of an integrable function on `closed_ball x r` converges to `f x` as `r` tends to zero (this is Lebesgue differentiation theorem). We give a version of this theorem for general Vitali families, and the concrete application to doubling measures. --- .../covering/density_theorem.lean | 32 +++- .../covering/differentiation.lean | 180 ++++++++++++++++++ 2 files changed, 209 insertions(+), 3 deletions(-) diff --git a/src/measure_theory/covering/density_theorem.lean b/src/measure_theory/covering/density_theorem.lean index cd4078adfa1b8..cd37a11be4cb2 100644 --- a/src/measure_theory/covering/density_theorem.lean +++ b/src/measure_theory/covering/density_theorem.lean @@ -134,12 +134,15 @@ end end -/-- A version of *Lebesgue's density theorem* for a sequence of closed balls whose centres are +section applications +variables [sigma_compact_space α] [borel_space α] [is_locally_finite_measure μ] + {E : Type*} [normed_add_comm_group E] + +/-- A version of *Lebesgue's density theorem* for a sequence of closed balls whose centers are not required to be fixed. See also `besicovitch.ae_tendsto_measure_inter_div`. -/ -lemma ae_tendsto_measure_inter_div - [sigma_compact_space α] [borel_space α] [is_locally_finite_measure μ] (S : set α) (K : ℝ) : +lemma ae_tendsto_measure_inter_div (S : set α) (K : ℝ) : ∀ᵐ x ∂μ.restrict S, ∀ {ι : Type*} {l : filter ι} (w : ι → α) (δ : ι → ℝ) (δlim : tendsto δ l (𝓝[>] 0)) (xmem : ∀ᶠ j in l, x ∈ closed_ball (w j) (K * δ j)), @@ -147,4 +150,27 @@ lemma ae_tendsto_measure_inter_div by filter_upwards [(vitali_family μ K).ae_tendsto_measure_inter_div S] with x hx ι l w δ δlim xmem using hx.comp (tendsto_closed_ball_filter_at μ _ _ δlim xmem) +/-- A version of *Lebesgue differentiation theorem* for a sequence of closed balls whose +centers are not required to be fixed. -/ +lemma ae_tendsto_average_norm_sub {f : α → E} (hf : integrable f μ) (K : ℝ) : + ∀ᵐ x ∂μ, ∀ {ι : Type*} {l : filter ι} (w : ι → α) (δ : ι → ℝ) + (δlim : tendsto δ l (𝓝[>] 0)) + (xmem : ∀ᶠ j in l, x ∈ closed_ball (w j) (K * δ j)), + tendsto (λ j, ⨍ y in closed_ball (w j) (δ j), ∥f y - f x∥ ∂μ) l (𝓝 0) := +by filter_upwards [(vitali_family μ K).ae_tendsto_average_norm_sub hf] with x hx ι l w δ δlim xmem +using hx.comp (tendsto_closed_ball_filter_at μ _ _ δlim xmem) + +/-- A version of *Lebesgue differentiation theorem* for a sequence of closed balls whose +centers are not required to be fixed. -/ +lemma ae_tendsto_average [normed_space ℝ E] [complete_space E] + {f : α → E} (hf : integrable f μ) (K : ℝ) : + ∀ᵐ x ∂μ, ∀ {ι : Type*} {l : filter ι} (w : ι → α) (δ : ι → ℝ) + (δlim : tendsto δ l (𝓝[>] 0)) + (xmem : ∀ᶠ j in l, x ∈ closed_ball (w j) (K * δ j)), + tendsto (λ j, ⨍ y in closed_ball (w j) (δ j), f y ∂μ) l (𝓝 (f x)) := +by filter_upwards [(vitali_family μ K).ae_tendsto_average hf] with x hx ι l w δ δlim xmem +using hx.comp (tendsto_closed_ball_filter_at μ _ _ δlim xmem) + +end applications + end is_doubling_measure diff --git a/src/measure_theory/covering/differentiation.lean b/src/measure_theory/covering/differentiation.lean index c5671d1179530..1eb18552396a9 100644 --- a/src/measure_theory/covering/differentiation.lean +++ b/src/measure_theory/covering/differentiation.lean @@ -8,6 +8,8 @@ import measure_theory.measure.regular import measure_theory.function.ae_measurable_order import measure_theory.integral.lebesgue import measure_theory.decomposition.radon_nikodym +import measure_theory.integral.average +import measure_theory.function.locally_integrable /-! # Differentiation of measures @@ -26,6 +28,13 @@ ratio really makes sense. For concrete applications, one needs concrete instances of Vitali families, as provided for instance by `besicovitch.vitali_family` (for balls) or by `vitali.vitali_family` (for doubling measures). +Specific applications to Lebesgue density points and the Lebesgue differentiation theorem are also +derived: +* `vitali_family.ae_tendsto_measure_inter_div` states that, for almost every point `x ∈ s`, + then `μ (s ∩ a) / μ a` tends to `1` as `a` shrinks to `x` along a Vitali family. +* `vitali_family.ae_tendsto_average_norm_sub` states that, for almost every point `x`, then the + average of `y ↦ ∥f y - f x∥` on `a` tends to `0` as `a` shrinks to `x` along a Vitali family. + ## Sketch of proof Let `v` be a Vitali family for `μ`. Assume for simplicity that `ρ` is absolutely continuous with @@ -60,6 +69,8 @@ local attribute [instance] emetric.second_countable_of_sigma_compact variables {α : Type*} [metric_space α] {m0 : measurable_space α} {μ : measure α} (v : vitali_family μ) +{E : Type*} [normed_add_comm_group E] + include v namespace vitali_family @@ -706,6 +717,8 @@ begin { simp only [Bx, zero_add] } end +/-! ### Lebesgue density points -/ + /-- Given a measurable set `s`, then `μ (s ∩ a) / μ a` converges when `a` shrinks to a typical point `x` along a Vitali family. The limit is `1` for `x ∈ s` and `0` for `x ∉ s`. This shows that almost every point of `s` is a Lebesgue density point for `s`. A version for non-measurable sets @@ -745,6 +758,173 @@ begin exact measure_to_measurable_inter_of_sigma_finite ha _, end +/-! ### Lebesgue differentiation theorem -/ + +lemma ae_tendsto_lintegral_div' {f : α → ℝ≥0∞} (hf : measurable f) (h'f : ∫⁻ y, f y ∂μ ≠ ∞) : + ∀ᵐ x ∂μ, tendsto (λ a, (∫⁻ y in a, f y ∂μ) / μ a) (v.filter_at x) (𝓝 (f x)) := +begin + let ρ := μ.with_density f, + haveI : is_finite_measure ρ, from is_finite_measure_with_density h'f, + filter_upwards [ae_tendsto_rn_deriv v ρ, rn_deriv_with_density μ hf] with x hx h'x, + rw ← h'x, + apply hx.congr' _, + filter_upwards [v.eventually_filter_at_measurable_set] with a ha, + rw ← with_density_apply f ha, +end + +lemma ae_tendsto_lintegral_div {f : α → ℝ≥0∞} (hf : ae_measurable f μ) (h'f : ∫⁻ y, f y ∂μ ≠ ∞) : + ∀ᵐ x ∂μ, tendsto (λ a, (∫⁻ y in a, f y ∂μ) / μ a) (v.filter_at x) (𝓝 (f x)) := +begin + have A : ∫⁻ y, hf.mk f y ∂μ ≠ ∞, + { convert h'f using 1, + apply lintegral_congr_ae, + exact hf.ae_eq_mk.symm }, + filter_upwards [v.ae_tendsto_lintegral_div' hf.measurable_mk A, hf.ae_eq_mk] with x hx h'x, + rw h'x, + convert hx, + ext1 a, + congr' 1, + apply lintegral_congr_ae, + exact ae_restrict_of_ae (hf.ae_eq_mk) +end + +lemma ae_tendsto_lintegral_nnnorm_sub_div' + {f : α → E} (hf : integrable f μ) (h'f : strongly_measurable f) : + ∀ᵐ x ∂μ, tendsto (λ a, (∫⁻ y in a, ∥f y - f x∥₊ ∂μ) / μ a) (v.filter_at x) (𝓝 0) := +begin + /- For every `c`, then `(∫⁻ y in a, ∥f y - c∥₊ ∂μ) / μ a` tends almost everywhere to `∥f x - c∥`. + We apply this to a countable set of `c` which is dense in the range of `f`, to deduce the desired + convergence. + A minor technical inconvenience is that constants are not integrable, so to apply previous lemmas + we need to replace `c` with the restriction of `c` to a finite measure set `A n` in the + above sketch. -/ + let A := measure_theory.measure.finite_spanning_sets_in_open μ, + rcases h'f.is_separable_range with ⟨t, t_count, ht⟩, + have main : ∀ᵐ x ∂μ, ∀ (n : ℕ) (c : E) (hc : c ∈ t), + tendsto (λ a, (∫⁻ y in a, ∥f y - (A.set n).indicator (λ y, c) y∥₊ ∂μ) / μ a) + (v.filter_at x) (𝓝 (∥f x - (A.set n).indicator (λ y, c) x∥₊)), + { simp_rw [ae_all_iff, ae_ball_iff t_count], + assume n c hc, + apply ae_tendsto_lintegral_div', + { refine (h'f.sub _).ennnorm, + exact strongly_measurable_const.indicator (is_open.measurable_set (A.set_mem n)) }, + { apply ne_of_lt, + calc ∫⁻ y, ↑∥f y - (A.set n).indicator (λ (y : α), c) y∥₊ ∂μ + ≤ ∫⁻ y, (∥f y∥₊ + ∥(A.set n).indicator (λ (y : α), c) y∥₊) ∂μ : + begin + apply lintegral_mono, + assume x, + dsimp, + rw ← ennreal.coe_add, + exact ennreal.coe_le_coe.2 (nnnorm_sub_le _ _), + end + ... = ∫⁻ y, ∥f y∥₊ ∂μ + ∫⁻ y, ∥(A.set n).indicator (λ (y : α), c) y∥₊ ∂μ : + lintegral_add_left h'f.ennnorm _ + ... < ∞ + ∞ : + begin + have I : integrable ((A.set n).indicator (λ (y : α), c)) μ, + by simp only [integrable_indicator_iff (is_open.measurable_set (A.set_mem n)), + integrable_on_const, A.finite n, or_true], + exact ennreal.add_lt_add hf.2 I.2, + end } }, + filter_upwards [main, v.ae_eventually_measure_pos] with x hx h'x, + have M : ∀ c ∈ t, tendsto (λ a, (∫⁻ y in a, ∥f y - c∥₊ ∂μ) / μ a) + (v.filter_at x) (𝓝 (∥f x - c∥₊)), + { assume c hc, + obtain ⟨n, xn⟩ : ∃ n, x ∈ A.set n, by simpa [← A.spanning] using mem_univ x, + specialize hx n c hc, + simp only [xn, indicator_of_mem] at hx, + apply hx.congr' _, + filter_upwards [v.eventually_filter_at_subset_of_nhds (is_open.mem_nhds (A.set_mem n) xn), + v.eventually_filter_at_measurable_set] + with a ha h'a, + congr' 1, + apply set_lintegral_congr_fun h'a, + apply eventually_of_forall (λ y, _), + assume hy, + simp only [ha hy, indicator_of_mem] }, + apply ennreal.tendsto_nhds_zero.2 (λ ε εpos, _), + obtain ⟨c, ct, xc⟩ : ∃ c ∈ t, (∥f x - c∥₊ : ℝ≥0∞) < ε / 2, + { simp_rw ← edist_eq_coe_nnnorm_sub, + have : f x ∈ closure t, from ht (mem_range_self _), + exact emetric.mem_closure_iff.1 this (ε / 2) (ennreal.half_pos (ne_of_gt εpos)) }, + filter_upwards [(tendsto_order.1 (M c ct)).2 (ε / 2) xc, h'x, v.eventually_measure_lt_top x] + with a ha h'a h''a, + apply ennreal.div_le_of_le_mul, + calc ∫⁻ y in a, ∥f y - f x∥₊ ∂μ + ≤ ∫⁻ y in a, ∥f y - c∥₊ + ∥f x - c∥₊ ∂μ : + begin + apply lintegral_mono (λ x, _), + simpa only [← edist_eq_coe_nnnorm_sub] using edist_triangle_right _ _ _, + end + ... = ∫⁻ y in a, ∥f y - c∥₊ ∂μ + ∫⁻ y in a, ∥f x - c∥₊ ∂μ : + lintegral_add_right _ measurable_const + ... ≤ ε / 2 * μ a + ε / 2 * μ a : + begin + refine add_le_add _ _, + { rw ennreal.div_lt_iff (or.inl (h'a.ne')) (or.inl (h''a.ne)) at ha, + exact ha.le }, + { simp only [lintegral_const, measure.restrict_apply, measurable_set.univ, univ_inter], + exact mul_le_mul_right' xc.le _ } + end + ... = ε * μ a : by rw [← add_mul, ennreal.add_halves] +end + +lemma ae_tendsto_lintegral_nnnorm_sub_div {f : α → E} (hf : integrable f μ) : + ∀ᵐ x ∂μ, tendsto (λ a, (∫⁻ y in a, ∥f y - f x∥₊ ∂μ) / μ a) (v.filter_at x) (𝓝 0) := +begin + have I : integrable (hf.1.mk f) μ, from hf.congr hf.1.ae_eq_mk, + filter_upwards [v.ae_tendsto_lintegral_nnnorm_sub_div' I hf.1.strongly_measurable_mk, + hf.1.ae_eq_mk] with x hx h'x, + apply hx.congr _, + assume a, + congr' 1, + apply lintegral_congr_ae, + apply ae_restrict_of_ae, + filter_upwards [hf.1.ae_eq_mk] with y hy, + rw [hy, h'x] +end + +/-- *Lebesgue differentiation theorem*: for almost every point `x`, the +average of `∥f y - f x∥` on `a` tends to `0` as `a` shrinks to `x` along a Vitali family.-/ +lemma ae_tendsto_average_norm_sub {f : α → E} (hf : integrable f μ) : + ∀ᵐ x ∂μ, tendsto (λ a, ⨍ y in a, ∥f y - f x∥ ∂μ) (v.filter_at x) (𝓝 0) := +begin + filter_upwards [v.ae_tendsto_lintegral_nnnorm_sub_div hf, v.ae_eventually_measure_pos] + with x hx h'x, + have := (ennreal.tendsto_to_real ennreal.zero_ne_top).comp hx, + simp only [ennreal.zero_to_real] at this, + apply tendsto.congr' _ this, + filter_upwards [h'x, v.eventually_measure_lt_top x] with a ha h'a, + simp only [function.comp_app, ennreal.to_real_div, set_average_eq, div_eq_inv_mul], + have A : integrable_on (λ y, (∥f y - f x∥₊ : ℝ)) a μ, + { simp_rw [coe_nnnorm], + exact (hf.integrable_on.sub (integrable_on_const.2 (or.inr h'a))).norm }, + rw [lintegral_coe_eq_integral _ A, ennreal.to_real_of_real], + { simp_rw [coe_nnnorm], + refl }, + { apply integral_nonneg, + assume x, + exact nnreal.coe_nonneg _ } +end + +/-- *Lebesgue differentiation theorem*: for almost every point `x`, the +average of `f` on `a` tends to `f x` as `a` shrinks to `x` along a Vitali family.-/ +lemma ae_tendsto_average [normed_space ℝ E] [complete_space E] {f : α → E} (hf : integrable f μ) : + ∀ᵐ x ∂μ, tendsto (λ a, ⨍ y in a, f y ∂μ) (v.filter_at x) (𝓝 (f x)) := +begin + filter_upwards [v.ae_tendsto_average_norm_sub hf, v.ae_eventually_measure_pos] with x hx h'x, + rw tendsto_iff_norm_tendsto_zero, + refine squeeze_zero' (eventually_of_forall (λ a, norm_nonneg _)) _ hx, + filter_upwards [h'x, v.eventually_measure_lt_top x] with a ha h'a, + nth_rewrite 0 [← set_average_const ha.ne' h'a.ne (f x)], + simp_rw [set_average_eq'], + rw ← integral_sub, + { exact norm_integral_le_integral_norm _ }, + { exact (integrable_inv_smul_measure ha.ne' h'a.ne).2 hf.integrable_on }, + { exact (integrable_inv_smul_measure ha.ne' h'a.ne).2 (integrable_on_const.2 (or.inr h'a)) } +end + end end vitali_family From 8b7c3f5b9ba4db1af57076b3132f1679d4775f7d Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 13 Oct 2022 05:04:56 +0000 Subject: [PATCH 732/828] feat(analysis/convex/between,analysis/convex/side): betweenness and `point_reflection` (#16937) Add lemmas about betweenness properties involving `point_reflection`. --- src/analysis/convex/between.lean | 18 +++++++++++++++++- src/analysis/convex/side.lean | 13 ++++++++++++- 2 files changed, 29 insertions(+), 2 deletions(-) diff --git a/src/analysis/convex/between.lean b/src/analysis/convex/between.lean index cb21a9e52c901..8e0d36f93d523 100644 --- a/src/analysis/convex/between.lean +++ b/src/analysis/convex/between.lean @@ -22,7 +22,7 @@ This file defines notions of a point in an affine space being between two given variables (R : Type*) {V V' P P' : Type*} -open affine_map +open affine_equiv affine_map section ordered_ring @@ -559,4 +559,20 @@ begin ring } end +variables (R) + +lemma wbtw_point_reflection (x y : P) : wbtw R y x (point_reflection R x y) := +begin + refine ⟨2⁻¹, ⟨by norm_num, by norm_num⟩, _⟩, + rw [line_map_apply, point_reflection_apply, vadd_vsub_assoc, ←two_smul R (x -ᵥ y)], + simp +end + +lemma sbtw_point_reflection_of_ne {x y : P} (h : x ≠ y) : sbtw R y x (point_reflection R x y) := +begin + refine ⟨wbtw_point_reflection _ _ _, h, _⟩, + nth_rewrite 0 [←point_reflection_self R x], + exact (point_reflection_involutive R x).injective.ne h +end + end linear_ordered_field diff --git a/src/analysis/convex/side.lean b/src/analysis/convex/side.lean index 5b90dcc960f69..da71d0b9c3768 100644 --- a/src/analysis/convex/side.lean +++ b/src/analysis/convex/side.lean @@ -25,7 +25,7 @@ This file defines notions of two points being on the same or opposite sides of a variables {R V V' P P' : Type*} -open affine_map +open affine_equiv affine_map namespace affine_subspace @@ -798,6 +798,17 @@ begin exact s_opp_side_smul_vsub_vadd_right hx hp hp' ht } end +lemma w_opp_side_point_reflection {s : affine_subspace R P} {x : P} (y : P) (hx : x ∈ s) : + s.w_opp_side y (point_reflection R x y) := +(wbtw_point_reflection R _ _).w_opp_side₁₃ hx + +lemma s_opp_side_point_reflection {s : affine_subspace R P} {x y : P} (hx : x ∈ s) (hy : y ∉ s) : + s.s_opp_side y (point_reflection R x y) := +begin + refine (sbtw_point_reflection_of_ne R (λ h, hy _)).s_opp_side_of_not_mem_of_mem hy hx, + rwa ←h +end + end linear_ordered_field end affine_subspace From 74cf64c3c22ad8c3df4d0d64e12a3e74273bebf0 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 13 Oct 2022 06:00:58 +0000 Subject: [PATCH 733/828] feat(group_theory/perm): add `equiv.perm.equiv_units_End` (#16900) Also add `monoid_hom.to_hom_perm`. --- src/group_theory/perm/basic.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/group_theory/perm/basic.lean b/src/group_theory/perm/basic.lean index 02f99f3bce424..096dfd17881be 100644 --- a/src/group_theory/perm/basic.lean +++ b/src/group_theory/perm/basic.lean @@ -29,6 +29,22 @@ instance perm_group : group (perm α) := mul_one := refl_trans, mul_left_inv := self_trans_symm } +/-- The permutation of a type is equivalent to the units group of the endomorphisms monoid of this +type. -/ +@[simps] def equiv_units_End : perm α ≃* units (function.End α) := +{ to_fun := λ e, ⟨e, e.symm, e.self_comp_symm, e.symm_comp_self⟩, + inv_fun := λ u, ⟨(u : function.End α), (↑u⁻¹ : function.End α), congr_fun u.inv_val, + congr_fun u.val_inv⟩, + left_inv := λ e, ext $ λ x, rfl, + right_inv := λ u, units.ext rfl, + map_mul' := λ e₁ e₂, rfl } + +/-- Lift a monoid homomorphism `f : G →* function.End α` to a monoid homomorphism +`f : G →* equiv.perm α`. -/ +@[simps] def _root_.monoid_hom.to_hom_perm {G : Type*} [group G] (f : G →* function.End α) : + G →* perm α := +equiv_units_End.symm.to_monoid_hom.comp f.to_hom_units + theorem mul_apply (f g : perm α) (x) : (f * g) x = f (g x) := equiv.trans_apply _ _ _ From fc2e3b003a30ab1ecd851fb2ba6b8dcc6e656278 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 13 Oct 2022 06:00:59 +0000 Subject: [PATCH 734/828] refactor(algebra/free_monoid): add `to_list`/`of_list` (#16926) * add `free_monoid.to_list` and `free_monoid.of_list`; * update API to explicitly use these functions; * fix `control.fold`. --- src/algebra/free_monoid/basic.lean | 128 ++++++++++++++++++++++------- src/control/fold.lean | 114 +++++++++++-------------- 2 files changed, 146 insertions(+), 96 deletions(-) diff --git a/src/algebra/free_monoid/basic.lean b/src/algebra/free_monoid/basic.lean index 53c3e796994ca..fe39381ef6667 100644 --- a/src/algebra/free_monoid/basic.lean +++ b/src/algebra/free_monoid/basic.lean @@ -25,47 +25,106 @@ def free_monoid (α) := list α namespace free_monoid +/-- The identity equivalence between `free_monoid α` and `list α`. -/ +@[to_additive "The identity equivalence between `free_add_monoid α` and `list α`."] +def to_list : free_monoid α ≃ list α := equiv.refl _ + +/-- The identity equivalence between `list α` and `free_monoid α`. -/ +@[to_additive "The identity equivalence between `list α` and `free_add_monoid α`."] +def of_list : list α ≃ free_monoid α := equiv.refl _ + +@[simp, to_additive] lemma to_list_symm : (@to_list α).symm = of_list := rfl +@[simp, to_additive] lemma of_list_symm : (@of_list α).symm = to_list := rfl +@[simp, to_additive] lemma to_list_of_list (l : list α) : to_list (of_list l) = l := rfl +@[simp, to_additive] lemma of_list_to_list (xs : free_monoid α) : of_list (to_list xs) = xs := rfl +@[simp, to_additive] lemma to_list_comp_of_list : @to_list α ∘ of_list = id := rfl +@[simp, to_additive] lemma of_list_comp_to_list : @of_list α ∘ to_list = id := rfl + @[to_additive] -instance : monoid (free_monoid α) := -{ one := [], - mul := λ x y, (x ++ y : list α), - mul_one := by intros; apply list.append_nil, - one_mul := by intros; refl, - mul_assoc := by intros; apply list.append_assoc } +instance : cancel_monoid (free_monoid α) := +{ one := of_list [], + mul := λ x y, of_list (x.to_list ++ y.to_list), + mul_one := list.append_nil, + one_mul := list.nil_append, + mul_assoc := list.append_assoc, + mul_left_cancel := λ _ _ _, list.append_left_cancel, + mul_right_cancel := λ _ _ _, list.append_right_cancel } @[to_additive] instance : inhabited (free_monoid α) := ⟨1⟩ -@[to_additive] -lemma one_def : (1 : free_monoid α) = [] := rfl +@[simp, to_additive] lemma to_list_one : (1 : free_monoid α).to_list = [] := rfl +@[simp, to_additive] lemma of_list_nil : of_list ([] : list α) = 1 := rfl -@[to_additive] -lemma mul_def (xs ys : list α) : (xs * ys : free_monoid α) = (xs ++ ys : list α) := +@[simp, to_additive] +lemma to_list_mul (xs ys : free_monoid α) : (xs * ys).to_list = xs.to_list ++ ys.to_list := rfl + +@[simp, to_additive] +lemma of_list_append (xs ys : list α) : + of_list (xs ++ ys) = of_list xs * of_list ys := rfl -@[to_additive] -lemma prod_eq_join (xs : list (free_monoid α)) : xs.prod = xs.join := -by induction xs; simp [*, mul_def, list.join, one_def] +@[simp, to_additive] +lemma to_list_prod (xs : list (free_monoid α)) : to_list xs.prod = (xs.map to_list).join := +by induction xs; simp [*, list.join] + +@[simp, to_additive] +lemma of_list_join (xs : list (list α)) : of_list xs.join = (xs.map of_list).prod := +to_list.injective $ by simp /-- Embeds an element of `α` into `free_monoid α` as a singleton list. -/ @[to_additive "Embeds an element of `α` into `free_add_monoid α` as a singleton list." ] -def of (x : α) : free_monoid α := [x] +def of (x : α) : free_monoid α := of_list [x] -@[to_additive] -lemma of_def (x : α) : of x = [x] := rfl +@[simp, to_additive] lemma to_list_of (x : α) : to_list (of x) = [x] := rfl +@[to_additive] lemma of_list_singleton (x : α) : of_list [x] = of x := rfl -@[to_additive] -lemma of_injective : function.injective (@of α) := -λ a b, list.head_eq_of_cons_eq +@[simp, to_additive] lemma of_list_cons (x : α) (xs : list α) : + of_list (x :: xs) = of x * of_list xs := +rfl + +@[to_additive] lemma to_list_of_mul (x : α) (xs : free_monoid α) : + to_list (of x * xs) = x :: xs.to_list := +rfl -@[to_additive] lemma of_mul_eq_cons (x : α) (l : free_monoid α) : of x * l = x :: l := rfl +@[to_additive] lemma of_injective : function.injective (@of α) := list.singleton_injective -/-- Recursor for `free_monoid` using `1` and `of x * xs` instead of `[]` and `x :: xs`. -/ +/-- Recursor for `free_monoid` using `1` and `free_monoid.of x * xs` instead of `[]` and +`x :: xs`. -/ @[elab_as_eliminator, to_additive - "Recursor for `free_add_monoid` using `0` and `of x + xs` instead of `[]` and `x :: xs`."] + "Recursor for `free_add_monoid` using `0` and `free_add_monoid.of x + xs` instead of `[]` and + `x :: xs`."] def rec_on {C : free_monoid α → Sort*} (xs : free_monoid α) (h0 : C 1) (ih : Π x xs, C xs → C (of x * xs)) : C xs := list.rec_on xs h0 ih +@[simp, to_additive] lemma rec_on_one {C : free_monoid α → Sort*} (h0 : C 1) + (ih : Π x xs, C xs → C (of x * xs)) : + @rec_on α C 1 h0 ih = h0 := +rfl + +@[simp, to_additive] lemma rec_on_of_mul {C : free_monoid α → Sort*} (x : α) (xs : free_monoid α) + (h0 : C 1) (ih : Π x xs, C xs → C (of x * xs)) : + @rec_on α C (of x * xs) h0 ih = ih x xs (rec_on xs h0 ih) := +rfl + +/-- A version of `list.cases_on` for `free_monoid` using `1` and `free_monoid.of x * xs` instead of +`[]` and `x :: xs`. -/ +@[elab_as_eliminator, to_additive + "A version of `list.cases_on` for `free_add_monoid` using `0` and `free_add_monoid.of x + xs` + instead of `[]` and `x :: xs`."] +def cases_on {C : free_monoid α → Sort*} (xs : free_monoid α) (h0 : C 1) + (ih : Π x xs, C (of x * xs)) : C xs := list.cases_on xs h0 ih + +@[simp, to_additive] lemma cases_on_one {C : free_monoid α → Sort*} (h0 : C 1) + (ih : Π x xs, C (of x * xs)) : + @cases_on α C 1 h0 ih = h0 := +rfl + +@[simp, to_additive] lemma cases_on_of_mul {C : free_monoid α → Sort*} (x : α) (xs : free_monoid α) + (h0 : C 1) (ih : Π x xs, C (of x * xs)) : + @cases_on α C (of x * xs) h0 ih = ih x xs := +rfl + @[ext, to_additive] lemma hom_eq ⦃f g : free_monoid α →* M⦄ (h : ∀ x, f (of x) = g (of x)) : f = g := @@ -76,8 +135,8 @@ monoid_hom.ext $ λ l, rec_on l (f.map_one.trans g.map_one.symm) $ @[to_additive "Equivalence between maps `α → A` and additive monoid homomorphisms `free_add_monoid α →+ A`."] def lift : (α → M) ≃ (free_monoid α →* M) := -{ to_fun := λ f, ⟨λ l, (l.map f).prod, rfl, - λ l₁ l₂, by simp only [mul_def, list.map_append, list.prod_append]⟩, +{ to_fun := λ f, ⟨λ l, (l.to_list.map f).prod, rfl, + λ l₁ l₂, by simp only [to_list_mul, list.map_append, list.prod_append]⟩, inv_fun := λ f x, f (of x), left_inv := λ f, funext $ λ x, one_mul (f x), right_inv := λ f, hom_eq $ λ x, one_mul (f (of x)) } @@ -86,10 +145,9 @@ def lift : (α → M) ≃ (free_monoid α →* M) := lemma lift_symm_apply (f : free_monoid α →* M) : lift.symm f = f ∘ of := rfl @[to_additive] -lemma lift_apply (f : α → M) (l : free_monoid α) : lift f l = (l.map f).prod := rfl +lemma lift_apply (f : α → M) (l : free_monoid α) : lift f l = (l.to_list.map f).prod := rfl -@[to_additive] -lemma lift_comp_of (f : α → M) : (lift f) ∘ of = f := lift.symm_apply_apply f +@[to_additive] lemma lift_comp_of (f : α → M) : lift f ∘ of = f := lift.symm_apply_apply f @[simp, to_additive] lemma lift_eval_of (f : α → M) (x : α) : lift f (of x) = f x := @@ -110,10 +168,14 @@ monoid_hom.ext_iff.1 (comp_lift g f) x /-- Define a multiplicative action of `free_monoid α` on `β`. -/ @[to_additive "Define an additive action of `free_add_monoid α` on `β`."] def mk_mul_action (f : α → β → β) : mul_action (free_monoid α) β := -{ smul := λ l b, list.foldr f b l, +{ smul := λ l b, l.to_list.foldr f b, one_smul := λ x, rfl, mul_smul := λ xs ys b, list.foldr_append _ _ _ _ } +@[to_additive] lemma smul_def (f : α → β → β) (l : free_monoid α) (b : β) : + (by haveI := mk_mul_action f; exact l • b = l.to_list.foldr f b) := +rfl + @[simp, to_additive] lemma of_smul (f : α → β → β) (x : α) (y : β) : (by haveI := mk_mul_action f; exact of x • y) = f x y := rfl @@ -123,12 +185,20 @@ each `of x` to `of (f x)`. -/ @[to_additive "The unique additive monoid homomorphism `free_add_monoid α →+ free_add_monoid β` that sends each `of x` to `of (f x)`."] def map (f : α → β) : free_monoid α →* free_monoid β := -{ to_fun := list.map f, +{ to_fun := λ l, of_list $ l.to_list.map f, map_one' := rfl, map_mul' := λ l₁ l₂, list.map_append _ _ _ } @[simp, to_additive] lemma map_of (f : α → β) (x : α) : map f (of x) = of (f x) := rfl +@[to_additive] lemma to_list_map (f : α → β) (xs : free_monoid α) : + (map f xs).to_list = xs.to_list.map f := +rfl + +@[to_additive] lemma of_list_map (f : α → β) (xs : list α) : + of_list (xs.map f) = map f (of_list xs) := +rfl + @[to_additive] lemma lift_of_comp_eq_map (f : α → β) : lift (λ x, of (f x)) = map f := diff --git a/src/control/fold.lean b/src/control/fold.lean index 2944d754f29cf..f00281fcffdeb 100644 --- a/src/control/fold.lean +++ b/src/control/fold.lean @@ -104,24 +104,18 @@ how the monoid of endofunctions define `foldl`. def foldl.mk (f : α → α) : foldl α := op f def foldl.get (x : foldl α) : α → α := unop x @[simps] def foldl.of_free_monoid (f : β → α → β) : free_monoid α →* monoid.foldl β := -{ to_fun := λ xs, op $ flip (list.foldl f) xs, +{ to_fun := λ xs, op $ flip (list.foldl f) xs.to_list, map_one' := rfl, - map_mul' := by intros; simp only [free_monoid.mul_def, flip, unop_op, + map_mul' := by intros; simp only [free_monoid.to_list_mul, flip, unop_op, list.foldl_append, op_inj]; refl } @[reducible] def foldr (α : Type u) : Type u := End α def foldr.mk (f : α → α) : foldr α := f def foldr.get (x : foldr α) : α → α := x @[simps] def foldr.of_free_monoid (f : α → β → β) : free_monoid α →* monoid.foldr β := -{ to_fun := λ xs, flip (list.foldr f) xs, +{ to_fun := λ xs, flip (list.foldr f) xs.to_list, map_one' := rfl, - map_mul' := - begin - intros, - simp only [free_monoid.mul_def, list.foldr_append, flip], - refl - end } - + map_mul' := λ xs ys, funext $ λ z, list.foldr_append _ _ _ _ } @[reducible] def mfoldl (m : Type u → Type u) [monad m] (α : Type u) : Type u := mul_opposite $ End $ Kleisli.mk m α @@ -129,7 +123,7 @@ def mfoldl.mk (f : α → m α) : mfoldl m α := op f def mfoldl.get (x : mfoldl m α) : α → m α := unop x @[simps] def mfoldl.of_free_monoid [is_lawful_monad m] (f : β → α → m β) : free_monoid α →* monoid.mfoldl m β := -{ to_fun := λ xs, op $ flip (list.mfoldl f) xs, +{ to_fun := λ xs, op $ flip (list.mfoldl f) xs.to_list, map_one' := rfl, map_mul' := by intros; apply unop_injective; ext; apply list.mfoldl_append } @@ -139,7 +133,7 @@ def mfoldr.mk (f : α → m α) : mfoldr m α := f def mfoldr.get (x : mfoldr m α) : α → m α := x @[simps] def mfoldr.of_free_monoid [is_lawful_monad m] (f : α → β → m β) : free_monoid α →* monoid.mfoldr m β := -{ to_fun := λ xs, flip (list.mfoldr f) xs, +{ to_fun := λ xs, flip (list.mfoldr f) xs.to_list, map_one' := rfl, map_mul' := by intros; ext; apply list.mfoldr_append } @@ -204,27 +198,18 @@ def map_fold [monoid α] [monoid β] (f : α →* β) : preserves_seq' := by { intros, simp only [f.map_mul, (<*>)], }, preserves_pure' := by { intros, simp only [f.map_one, pure] } } -def free.mk : α → free_monoid α := list.ret - -def free.map (f : α → β) : free_monoid α →* free_monoid β := -{ to_fun := list.map f, - map_mul' := λ x y, - by simp only [free_monoid.mul_def, list.map_append, free_add_monoid.add_def], - map_one' := by simp only [free_monoid.one_def, list.map, free_add_monoid.zero_def] } - lemma free.map_eq_map (f : α → β) (xs : list α) : - f <$> xs = free.map f xs := rfl + f <$> xs = (free_monoid.map f (free_monoid.of_list xs)).to_list := rfl lemma foldl.unop_of_free_monoid (f : β → α → β) (xs : free_monoid α) (a : β) : - unop (foldl.of_free_monoid f xs) a = list.foldl f a xs := rfl + unop (foldl.of_free_monoid f xs) a = list.foldl f a xs.to_list := rfl variables (m : Type u → Type u) [monad m] [is_lawful_monad m] variables {t : Type u → Type u} [traversable t] [is_lawful_traversable t] open is_lawful_traversable -lemma fold_map_hom - [monoid α] [monoid β] (f : α →* β) +lemma fold_map_hom [monoid α] [monoid β] (f : α →* β) (g : γ → α) (x : t γ) : f (fold_map g x) = fold_map (f ∘ g) x := calc f (fold_map g x) @@ -236,19 +221,9 @@ calc f (fold_map g x) lemma fold_map_hom_free [monoid β] (f : free_monoid α →* β) (x : t α) : - f (fold_map free.mk x) = fold_map (f ∘ free.mk) x := + f (fold_map free_monoid.of x) = fold_map (f ∘ free_monoid.of) x := fold_map_hom f _ x -variable {m} - -lemma fold_mfoldl_cons (f : α → β → m α) (x : β) (y : α) : - list.mfoldl f y (free.mk x) = f y x := -by simp only [free.mk, list.ret, list.mfoldl, bind_pure] - -lemma fold_mfoldr_cons (f : β → α → m α) (x : β) (y : α) : - list.mfoldr f y (free.mk x) = f x y := -by simp only [free.mk, list.ret, list.mfoldr, pure_bind] - end applicative_transformation section equalities @@ -257,37 +232,38 @@ variables {α β γ : Type u} variables {t : Type u → Type u} [traversable t] [is_lawful_traversable t] @[simp] -lemma foldl.of_free_monoid_comp_free_mk (f : α → β → α) : - foldl.of_free_monoid f ∘ free.mk = foldl.mk ∘ flip f := rfl +lemma foldl.of_free_monoid_comp_of (f : α → β → α) : + foldl.of_free_monoid f ∘ free_monoid.of = foldl.mk ∘ flip f := rfl @[simp] -lemma foldr.of_free_monoid_comp_free_mk (f : β → α → α) : - foldr.of_free_monoid f ∘ free.mk = foldr.mk ∘ f := rfl +lemma foldr.of_free_monoid_comp_of (f : β → α → α) : + foldr.of_free_monoid f ∘ free_monoid.of = foldr.mk ∘ f := rfl @[simp] -lemma mfoldl.of_free_monoid_comp_free_mk {m} [monad m] [is_lawful_monad m] (f : α → β → m α) : - mfoldl.of_free_monoid f ∘ free.mk = mfoldl.mk ∘ flip f := -by ext; simp [(∘), mfoldl.of_free_monoid, mfoldl.mk, flip, fold_mfoldl_cons]; refl +lemma mfoldl.of_free_monoid_comp_of {m} [monad m] [is_lawful_monad m] (f : α → β → m α) : + mfoldl.of_free_monoid f ∘ free_monoid.of = mfoldl.mk ∘ flip f := +by { ext1 x, simp [(∘), mfoldl.of_free_monoid, mfoldl.mk, flip], refl } @[simp] -lemma mfoldr.of_free_monoid_comp_free_mk {m} [monad m] [is_lawful_monad m] (f : β → α → m α) : - mfoldr.of_free_monoid f ∘ free.mk = mfoldr.mk ∘ f := -by { ext, simp [(∘), mfoldr.of_free_monoid, mfoldr.mk, flip, fold_mfoldr_cons] } +lemma mfoldr.of_free_monoid_comp_of {m} [monad m] [is_lawful_monad m] (f : β → α → m α) : + mfoldr.of_free_monoid f ∘ free_monoid.of = mfoldr.mk ∘ f := +by { ext, simp [(∘), mfoldr.of_free_monoid, mfoldr.mk, flip] } lemma to_list_spec (xs : t α) : - to_list xs = (fold_map free.mk xs : free_monoid _) := + to_list xs = free_monoid.to_list (fold_map free_monoid.of xs) := eq.symm $ -calc fold_map free.mk xs - = (fold_map free.mk xs).reverse.reverse : by simp only [list.reverse_reverse] -... = (list.foldr cons [] (fold_map free.mk xs).reverse).reverse +calc free_monoid.to_list (fold_map free_monoid.of xs) + = free_monoid.to_list (fold_map free_monoid.of xs).reverse.reverse + : by simp only [list.reverse_reverse] +... = free_monoid.to_list (list.foldr cons [] (fold_map free_monoid.of xs).reverse).reverse : by simp only [list.foldr_eta] -... = (unop (foldl.of_free_monoid (flip cons) (fold_map free.mk xs)) []).reverse - : by simp [flip,list.foldr_reverse,foldl.of_free_monoid, unop_op] +... = (unop (foldl.of_free_monoid (flip cons) (fold_map free_monoid.of xs)) []).reverse + : by simp [flip, list.foldr_reverse, foldl.of_free_monoid, unop_op] ... = to_list xs : begin rw fold_map_hom_free (foldl.of_free_monoid (flip $ @cons α)), - simp only [to_list, foldl, list.reverse_inj, foldl.get, - foldl.of_free_monoid_comp_free_mk], - all_goals { apply_instance } + { simp only [to_list, foldl, list.reverse_inj, foldl.get, + foldl.of_free_monoid_comp_of] }, + { apply_instance } end lemma fold_map_map [monoid γ] (f : α → β) (g : β → γ) (xs : t α) : @@ -297,23 +273,27 @@ by simp only [fold_map,traverse_map] lemma foldl_to_list (f : α → β → α) (xs : t β) (x : α) : foldl f x xs = list.foldl f x (to_list xs) := begin - rw ← foldl.unop_of_free_monoid, + rw [← free_monoid.to_list_of_list (to_list xs), ← foldl.unop_of_free_monoid], simp only [foldl, to_list_spec, fold_map_hom_free, - foldl.of_free_monoid_comp_free_mk, foldl.get] + foldl.of_free_monoid_comp_of, foldl.get, free_monoid.of_list_to_list] end lemma foldr_to_list (f : α → β → β) (xs : t α) (x : β) : foldr f x xs = list.foldr f x (to_list xs) := begin - change _ = foldr.of_free_monoid _ _ _, - simp only [foldr, to_list_spec, fold_map_hom_free, - foldr.of_free_monoid_comp_free_mk, foldr.get] + change _ = foldr.of_free_monoid _ (free_monoid.of_list $ to_list xs) _, + rw [to_list_spec, foldr, foldr.get, free_monoid.of_list_to_list, fold_map_hom_free, + foldr.of_free_monoid_comp_of] end +/- + +-/ + lemma to_list_map (f : α → β) (xs : t α) : - to_list (f <$> xs) = f <$> to_list xs := by -{ simp only [to_list_spec,free.map_eq_map,fold_map_hom (free.map f), fold_map_map]; - refl } + to_list (f <$> xs) = f <$> to_list xs := +by simp only [to_list_spec, free.map_eq_map, fold_map_hom, fold_map_map, + free_monoid.of_list_to_list, free_monoid.map_of, (∘)] @[simp] theorem foldl_map (g : β → γ) (f : α → γ → α) (a : α) (l : t β) : foldl f a (g <$> l) = foldl (λ x y, f x (g y)) a l := @@ -328,7 +308,7 @@ begin simp only [to_list_spec, fold_map, traverse], induction xs, case list.nil { refl }, - case list.cons : _ _ ih { unfold list.traverse list.ret, rw ih, refl } + case list.cons : _ _ ih { conv_rhs { rw [← ih] }, refl } end theorem length_to_list {xs : t α} : length xs = list.length (to_list xs) := @@ -353,17 +333,17 @@ variables {m : Type u → Type u} [monad m] [is_lawful_monad m] lemma mfoldl_to_list {f : α → β → m α} {x : α} {xs : t β} : mfoldl f x xs = list.mfoldl f x (to_list xs) := -calc mfoldl f x xs = unop (mfoldl.of_free_monoid f (to_list xs)) x : +calc mfoldl f x xs = unop (mfoldl.of_free_monoid f (free_monoid.of_list $ to_list xs)) x : by simp only [mfoldl, to_list_spec, fold_map_hom_free (mfoldl.of_free_monoid f), - mfoldl.of_free_monoid_comp_free_mk, mfoldl.get] + mfoldl.of_free_monoid_comp_of, mfoldl.get, free_monoid.of_list_to_list] ... = list.mfoldl f x (to_list xs) : by simp [mfoldl.of_free_monoid, unop_op, flip] lemma mfoldr_to_list (f : α → β → m β) (x : β) (xs : t α) : mfoldr f x xs = list.mfoldr f x (to_list xs) := begin - change _ = mfoldr.of_free_monoid f (to_list xs) x, + change _ = mfoldr.of_free_monoid f (free_monoid.of_list $ to_list xs) x, simp only [mfoldr, to_list_spec, fold_map_hom_free (mfoldr.of_free_monoid f), - mfoldr.of_free_monoid_comp_free_mk, mfoldr.get] + mfoldr.of_free_monoid_comp_of, mfoldr.get, free_monoid.of_list_to_list] end @[simp] theorem mfoldl_map (g : β → γ) (f : α → γ → m α) (a : α) (l : t β) : From 8f82d25e27b93224d123fbf1c73d686ea1cb3666 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 13 Oct 2022 06:01:00 +0000 Subject: [PATCH 735/828] feat(geometry/euclidean/basic): `concyclic` (#16927) Define a set of points to be concyclic if it is cospherical and coplanar. I don't expect that much use of this definition, as opposed to the separate `cospherical` and `coplanar` pieces, in mathlib, but it's intended to allow formal statements of results involving concyclic points that correspond more closely to the informal statements. --- docs/undergrad.yaml | 2 +- src/geometry/euclidean/basic.lean | 22 ++++++++++++++++++++++ 2 files changed, 23 insertions(+), 1 deletion(-) diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index e37577deec32e..eccd67d980bec 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -263,7 +263,7 @@ Affine and Euclidean Geometry: angles between vectors: 'inner_product_geometry.angle' angles between planes: '' inscribed angle theorem: '' - cocyclicity: '' + cocyclicity: 'euclidean_geometry.concyclic' group of isometries stabilizing a subset of the plane or of space: '' regular polygons: '' metric relations in the triangle: '' diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index 7b478e10299b6..9d6dc284237ea 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -1518,4 +1518,26 @@ lemma eq_of_mem_sphere_of_mem_sphere_of_finrank_eq_two [finite_dimensional ℝ V eq_of_dist_eq_of_dist_eq_of_finrank_eq_two hd ((sphere.center_ne_iff_ne_of_mem hps₁ hps₂).2 hs) hp hp₁s₁ hp₂s₁ hps₁ hp₁s₂ hp₂s₂ hps₂ +/-- A set of points is concyclic if it is cospherical and coplanar. (Most results are stated +directly in terms of `cospherical` instead of using `concyclic`.) -/ +structure concyclic (ps : set P) : Prop := +(cospherical : cospherical ps) +(coplanar : coplanar ℝ ps) + +/-- A subset of a concyclic set is concyclic. -/ +lemma concyclic.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (h : concyclic ps₂) : concyclic ps₁ := +⟨cospherical_subset hs h.1, h.2.subset hs⟩ + +/-- The empty set is concyclic. -/ +lemma concyclic_empty : concyclic (∅ : set P) := +⟨cospherical_empty, coplanar_empty ℝ P⟩ + +/-- A single point is concyclic. -/ +lemma concyclic_singleton (p : P) : concyclic ({p} : set P) := +⟨cospherical_singleton p, coplanar_singleton ℝ p⟩ + +/-- Two points are concyclic. -/ +lemma concyclic_pair (p₁ p₂ : P) : concyclic ({p₁, p₂} : set P) := +⟨cospherical_pair p₁ p₂, coplanar_pair ℝ p₁ p₂⟩ + end euclidean_geometry From f2259302f3cbbfb08f5c7e35c5850ede99be6ca3 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Thu, 13 Oct 2022 06:01:02 +0000 Subject: [PATCH 736/828] chore(group_theory/*): Golf using `subgroup.subtype_injective` (#16941) This PR uses the recently added `subgroup.subtype_injective` to golf a few lines. --- src/group_theory/schur_zassenhaus.lean | 2 +- src/group_theory/solvable.lean | 2 +- src/group_theory/subgroup/basic.lean | 2 +- src/group_theory/sylow.lean | 2 +- 4 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/group_theory/schur_zassenhaus.lean b/src/group_theory/schur_zassenhaus.lean index a425f1600e0d7..c07fded4e72dc 100644 --- a/src/group_theory/schur_zassenhaus.lean +++ b/src/group_theory/schur_zassenhaus.lean @@ -206,7 +206,7 @@ begin have key := step2 h1 h2 h3 (K.map N.subtype) K.map_subtype_le, rw ← map_bot N.subtype at key, conv at key { congr, skip, to_rhs, rw [←N.subtype_range, N.subtype.range_eq_map] }, - have inj := map_injective (show function.injective N.subtype, from subtype.coe_injective), + have inj := map_injective N.subtype_injective, rwa [inj.eq_iff, inj.eq_iff] at key, end diff --git a/src/group_theory/solvable.lean b/src/group_theory/solvable.lean index 32a39b4b8e698..d1d29b72ce2f5 100644 --- a/src/group_theory/solvable.lean +++ b/src/group_theory/solvable.lean @@ -136,7 +136,7 @@ lemma solvable_of_solvable_injective (hf : function.injective f) [h : is_solvabl solvable_of_ker_le_range (1 : G' →* G) f ((f.ker_eq_bot_iff.mpr hf).symm ▸ bot_le) instance subgroup_solvable_of_solvable (H : subgroup G) [h : is_solvable G] : is_solvable H := -solvable_of_solvable_injective (show function.injective (subtype H), from subtype.val_injective) +solvable_of_solvable_injective H.subtype_injective lemma solvable_of_surjective (hf : function.surjective f) [h : is_solvable G] : is_solvable G' := diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index b504acc242a35..d3974e05b204c 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2417,7 +2417,7 @@ end @[to_additive] lemma closure_preimage_eq_top (s : set G) : closure ((closure s).subtype ⁻¹' s) = ⊤ := begin - apply map_injective (show function.injective (closure s).subtype, from subtype.coe_injective), + apply map_injective (closure s).subtype_injective, rwa [monoid_hom.map_closure, ←monoid_hom.range_eq_map, subtype_range, set.image_preimage_eq_of_subset], rw [coe_subtype, subtype.range_coe_subtype], diff --git a/src/group_theory/sylow.lean b/src/group_theory/sylow.lean index 5d5c5449025c6..e4bd0213abc51 100644 --- a/src/group_theory/sylow.lean +++ b/src/group_theory/sylow.lean @@ -165,7 +165,7 @@ sylow.fintype_of_ker_is_p_group (is_p_group.ker_is_p_group_of_injective hf) /-- If `H` is a subgroup of `G`, then `fintype (sylow p G)` implies `fintype (sylow p H)`. -/ noncomputable instance (H : subgroup G) [fintype (sylow p G)] : fintype (sylow p H) := -sylow.fintype_of_injective (show function.injective H.subtype, from subtype.coe_injective) +sylow.fintype_of_injective H.subtype_injective /-- If `H` is a subgroup of `G`, then `finite (sylow p G)` implies `finite (sylow p H)`. -/ instance (H : subgroup G) [finite (sylow p G)] : finite (sylow p H) := From f8b900bfc7a6bd6b875d61ec2efbaf24a24a8cc4 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 13 Oct 2022 06:01:03 +0000 Subject: [PATCH 737/828] refactor(geometry/euclidean): split up unoriented angles (#16942) Move the definitions and lemmas about unoriented angles from `geometry.euclidean.basic` (quite a large file) to new `geometry.euclidean.angle.unoriented.basic` and `geometry.euclidean.angle.unoriented.affine`. Also move some lemmas about right-angled triangles from `geometry.euclidean.triangle` to `geometry.euclidean.angle.unoriented.right_angle` (there's not much there at present, but I expect to add lots of lemmas about trigonometric functions of angles in right-angled triangles, which should be enough to justify having a separate file for such content). I expect to split up `geometry.euclidean.oriented_angle` similarly, into files under `geometry.euclidean.angle.oriented`, once Heather's refactor is in. When it becomes convenient to deduce results about unoriented angles from those about oriented angles rather than vice versa, as with circle theorems, I expect to put such groups of results in e.g. `geometry.euclidean.angle.sphere` rather than in separate unoriented and oriented files. Although not done here, dependencies could be reduced further by moving two lemmas about conformal maps from `geometry.euclidean.angle.unoriented.basic` to a separate `geometry.euclidean.angle.unoriented.conformal` (most uses of angles don't need `analysis.calculus.conformal.normed_space`). There are no changes to lemma statements or proofs. --- .../euclidean/angle/unoriented/affine.lean | 349 ++++++++++ .../euclidean/angle/unoriented/basic.lean | 337 +++++++++ .../angle/unoriented/right_angle.lean | 81 +++ src/geometry/euclidean/basic.lean | 648 ------------------ src/geometry/euclidean/circumcenter.lean | 3 - src/geometry/euclidean/default.lean | 1 + src/geometry/euclidean/monge_point.lean | 1 - src/geometry/euclidean/oriented_angle.lean | 2 +- src/geometry/euclidean/sphere.lean | 1 + src/geometry/euclidean/triangle.lean | 38 +- 10 files changed, 771 insertions(+), 690 deletions(-) create mode 100644 src/geometry/euclidean/angle/unoriented/affine.lean create mode 100644 src/geometry/euclidean/angle/unoriented/basic.lean create mode 100644 src/geometry/euclidean/angle/unoriented/right_angle.lean diff --git a/src/geometry/euclidean/angle/unoriented/affine.lean b/src/geometry/euclidean/angle/unoriented/affine.lean new file mode 100644 index 0000000000000..f5ba7e1bc5797 --- /dev/null +++ b/src/geometry/euclidean/angle/unoriented/affine.lean @@ -0,0 +1,349 @@ +/- +Copyright (c) 2020 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers, Manuel Candales +-/ +import analysis.convex.between +import geometry.euclidean.angle.unoriented.basic + +/-! +# Angles between points + +This file defines unoriented angles in Euclidean affine spaces. + +## Main definitions + +* `euclidean_geometry.angle`, with notation `∠`, is the undirected angle determined by three + points. + +-/ + +noncomputable theory +open_locale big_operators +open_locale real +open_locale real_inner_product_space + +namespace euclidean_geometry + +open inner_product_geometry + +variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] + [normed_add_torsor V P] +include V + +/-- The undirected angle at `p2` between the line segments to `p1` and +`p3`. If either of those points equals `p2`, this is π/2. Use +`open_locale euclidean_geometry` to access the `∠ p1 p2 p3` +notation. -/ +def angle (p1 p2 p3 : P) : ℝ := angle (p1 -ᵥ p2 : V) (p3 -ᵥ p2) + +localized "notation (name := angle) `∠` := euclidean_geometry.angle" in euclidean_geometry + +lemma continuous_at_angle {x : P × P × P} (hx12 : x.1 ≠ x.2.1) (hx32 : x.2.2 ≠ x.2.1) : + continuous_at (λ y : P × P × P, ∠ y.1 y.2.1 y.2.2) x := +begin + let f : P × P × P → V × V := λ y, (y.1 -ᵥ y.2.1, y.2.2 -ᵥ y.2.1), + have hf1 : (f x).1 ≠ 0, by simp [hx12], + have hf2 : (f x).2 ≠ 0, by simp [hx32], + exact (inner_product_geometry.continuous_at_angle hf1 hf2).comp + ((continuous_fst.vsub continuous_snd.fst).prod_mk + (continuous_snd.snd.vsub continuous_snd.fst)).continuous_at +end + +@[simp] lemma _root_.affine_isometry.angle_map {V₂ P₂ : Type*} [inner_product_space ℝ V₂] + [metric_space P₂] [normed_add_torsor V₂ P₂] (f : P →ᵃⁱ[ℝ] P₂) (p₁ p₂ p₃ : P) : + ∠ (f p₁) (f p₂) (f p₃) = ∠ p₁ p₂ p₃ := +by simp_rw [angle, ←affine_isometry.map_vsub, linear_isometry.angle_map] + +@[simp, norm_cast] lemma _root_.affine_subspace.angle_coe {s : affine_subspace ℝ P} + (p₁ p₂ p₃ : s) : + by haveI : nonempty s := ⟨p₁⟩; exact ∠ (p₁ : P) (p₂ : P) (p₃ : P) = ∠ p₁ p₂ p₃ := +by haveI : nonempty s := ⟨p₁⟩; exact s.subtypeₐᵢ.angle_map p₁ p₂ p₃ + +/-- The angle at a point does not depend on the order of the other two +points. -/ +lemma angle_comm (p1 p2 p3 : P) : ∠ p1 p2 p3 = ∠ p3 p2 p1 := +angle_comm _ _ + +/-- The angle at a point is nonnegative. -/ +lemma angle_nonneg (p1 p2 p3 : P) : 0 ≤ ∠ p1 p2 p3 := +angle_nonneg _ _ + +/-- The angle at a point is at most π. -/ +lemma angle_le_pi (p1 p2 p3 : P) : ∠ p1 p2 p3 ≤ π := +angle_le_pi _ _ + +/-- The angle ∠AAB at a point. -/ +lemma angle_eq_left (p1 p2 : P) : ∠ p1 p1 p2 = π / 2 := +begin + unfold angle, + rw vsub_self, + exact angle_zero_left _ +end + +/-- The angle ∠ABB at a point. -/ +lemma angle_eq_right (p1 p2 : P) : ∠ p1 p2 p2 = π / 2 := +by rw [angle_comm, angle_eq_left] + +/-- The angle ∠ABA at a point. -/ +lemma angle_eq_of_ne {p1 p2 : P} (h : p1 ≠ p2) : ∠ p1 p2 p1 = 0 := +angle_self (λ he, h (vsub_eq_zero_iff_eq.1 he)) + +/-- If the angle ∠ABC at a point is π, the angle ∠BAC is 0. -/ +lemma angle_eq_zero_of_angle_eq_pi_left {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : + ∠ p2 p1 p3 = 0 := +begin + unfold angle at h, + rw angle_eq_pi_iff at h, + rcases h with ⟨hp1p2, ⟨r, ⟨hr, hpr⟩⟩⟩, + unfold angle, + rw angle_eq_zero_iff, + rw [←neg_vsub_eq_vsub_rev, neg_ne_zero] at hp1p2, + use [hp1p2, -r + 1, add_pos (neg_pos_of_neg hr) zero_lt_one], + rw [add_smul, ←neg_vsub_eq_vsub_rev p1 p2, smul_neg], + simp [←hpr] +end + +/-- If the angle ∠ABC at a point is π, the angle ∠BCA is 0. -/ +lemma angle_eq_zero_of_angle_eq_pi_right {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : + ∠ p2 p3 p1 = 0 := +begin + rw angle_comm at h, + exact angle_eq_zero_of_angle_eq_pi_left h +end + +/-- If ∠BCD = π, then ∠ABC = ∠ABD. -/ +lemma angle_eq_angle_of_angle_eq_pi (p1 : P) {p2 p3 p4 : P} (h : ∠ p2 p3 p4 = π) : + ∠ p1 p2 p3 = ∠ p1 p2 p4 := +begin + unfold angle at *, + rcases angle_eq_pi_iff.1 h with ⟨hp2p3, ⟨r, ⟨hr, hpr⟩⟩⟩, + rw [eq_comm], + convert angle_smul_right_of_pos (p1 -ᵥ p2) (p3 -ᵥ p2) (add_pos (neg_pos_of_neg hr) zero_lt_one), + rw [add_smul, ← neg_vsub_eq_vsub_rev p2 p3, smul_neg, neg_smul, ← hpr], + simp +end + +/-- If ∠BCD = π, then ∠ACB + ∠ACD = π. -/ +lemma angle_add_angle_eq_pi_of_angle_eq_pi (p1 : P) {p2 p3 p4 : P} (h : ∠ p2 p3 p4 = π) : + ∠ p1 p3 p2 + ∠ p1 p3 p4 = π := +begin + unfold angle at h, + rw [angle_comm p1 p3 p2, angle_comm p1 p3 p4], + unfold angle, + exact angle_add_angle_eq_pi_of_angle_eq_pi _ h +end + +/-- Vertical Angles Theorem: angles opposite each other, formed by two intersecting straight +lines, are equal. -/ +lemma angle_eq_angle_of_angle_eq_pi_of_angle_eq_pi {p1 p2 p3 p4 p5 : P} + (hapc : ∠ p1 p5 p3 = π) (hbpd : ∠ p2 p5 p4 = π) : ∠ p1 p5 p2 = ∠ p3 p5 p4 := +by linarith [angle_add_angle_eq_pi_of_angle_eq_pi p1 hbpd, angle_comm p4 p5 p1, + angle_add_angle_eq_pi_of_angle_eq_pi p4 hapc, angle_comm p4 p5 p3] + +/-- If ∠ABC = π then dist A B ≠ 0. -/ +lemma left_dist_ne_zero_of_angle_eq_pi {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : dist p1 p2 ≠ 0 := +begin + by_contra heq, + rw [dist_eq_zero] at heq, + rw [heq, angle_eq_left] at h, + exact real.pi_ne_zero (by linarith), +end + +/-- If ∠ABC = π then dist C B ≠ 0. -/ +lemma right_dist_ne_zero_of_angle_eq_pi {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : dist p3 p2 ≠ 0 := +left_dist_ne_zero_of_angle_eq_pi $ (angle_comm _ _ _).trans h + +/-- If ∠ABC = π, then (dist A C) = (dist A B) + (dist B C). -/ +lemma dist_eq_add_dist_of_angle_eq_pi {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : + dist p1 p3 = dist p1 p2 + dist p3 p2 := +begin + rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, ← vsub_sub_vsub_cancel_right], + exact norm_sub_eq_add_norm_of_angle_eq_pi h, +end + +/-- If A ≠ B and C ≠ B then ∠ABC = π if and only if (dist A C) = (dist A B) + (dist B C). -/ +lemma dist_eq_add_dist_iff_angle_eq_pi {p1 p2 p3 : P} (hp1p2 : p1 ≠ p2) (hp3p2 : p3 ≠ p2) : + dist p1 p3 = dist p1 p2 + dist p3 p2 ↔ ∠ p1 p2 p3 = π := +begin + rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, ← vsub_sub_vsub_cancel_right], + exact norm_sub_eq_add_norm_iff_angle_eq_pi + ((λ he, hp1p2 (vsub_eq_zero_iff_eq.1 he))) (λ he, hp3p2 (vsub_eq_zero_iff_eq.1 he)), +end + +/-- If ∠ABC = 0, then (dist A C) = abs ((dist A B) - (dist B C)). -/ +lemma dist_eq_abs_sub_dist_of_angle_eq_zero {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = 0) : + (dist p1 p3) = |(dist p1 p2) - (dist p3 p2)| := +begin + rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, ← vsub_sub_vsub_cancel_right], + exact norm_sub_eq_abs_sub_norm_of_angle_eq_zero h, +end + +/-- If A ≠ B and C ≠ B then ∠ABC = 0 if and only if (dist A C) = abs ((dist A B) - (dist B C)). -/ +lemma dist_eq_abs_sub_dist_iff_angle_eq_zero {p1 p2 p3 : P} (hp1p2 : p1 ≠ p2) (hp3p2 : p3 ≠ p2) : + (dist p1 p3) = |(dist p1 p2) - (dist p3 p2)| ↔ ∠ p1 p2 p3 = 0 := +begin + rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, ← vsub_sub_vsub_cancel_right], + exact norm_sub_eq_abs_sub_norm_iff_angle_eq_zero + ((λ he, hp1p2 (vsub_eq_zero_iff_eq.1 he))) (λ he, hp3p2 (vsub_eq_zero_iff_eq.1 he)), +end + +/-- If M is the midpoint of the segment AB, then ∠AMB = π. -/ +lemma angle_midpoint_eq_pi (p1 p2 : P) (hp1p2 : p1 ≠ p2) : ∠ p1 (midpoint ℝ p1 p2) p2 = π := +have p2 -ᵥ midpoint ℝ p1 p2 = -(p1 -ᵥ midpoint ℝ p1 p2), by { rw neg_vsub_eq_vsub_rev, simp }, +by simp [angle, this, hp1p2, -zero_lt_one] + +/-- If M is the midpoint of the segment AB and C is the same distance from A as it is from B +then ∠CMA = π / 2. -/ +lemma angle_left_midpoint_eq_pi_div_two_of_dist_eq {p1 p2 p3 : P} (h : dist p3 p1 = dist p3 p2) : + ∠ p3 (midpoint ℝ p1 p2) p1 = π / 2 := +begin + let m : P := midpoint ℝ p1 p2, + have h1 : p3 -ᵥ p1 = (p3 -ᵥ m) - (p1 -ᵥ m) := (vsub_sub_vsub_cancel_right p3 p1 m).symm, + have h2 : p3 -ᵥ p2 = (p3 -ᵥ m) + (p1 -ᵥ m), + { rw [left_vsub_midpoint, ← midpoint_vsub_right, vsub_add_vsub_cancel] }, + rw [dist_eq_norm_vsub V p3 p1, dist_eq_norm_vsub V p3 p2, h1, h2] at h, + exact (norm_add_eq_norm_sub_iff_angle_eq_pi_div_two (p3 -ᵥ m) (p1 -ᵥ m)).mp h.symm, +end + +/-- If M is the midpoint of the segment AB and C is the same distance from A as it is from B +then ∠CMB = π / 2. -/ +lemma angle_right_midpoint_eq_pi_div_two_of_dist_eq {p1 p2 p3 : P} (h : dist p3 p1 = dist p3 p2) : + ∠ p3 (midpoint ℝ p1 p2) p2 = π / 2 := +by rw [midpoint_comm p1 p2, angle_left_midpoint_eq_pi_div_two_of_dist_eq h.symm] + +/-- If the second of three points is strictly between the other two, the angle at that point +is π. -/ +lemma _root_.sbtw.angle₁₂₃_eq_pi {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₁ p₂ p₃ = π := +begin + rw [angle, angle_eq_pi_iff], + rcases h with ⟨⟨r, ⟨hr0, hr1⟩, hp₂⟩, hp₂p₁, hp₂p₃⟩, + refine ⟨vsub_ne_zero.2 hp₂p₁.symm, -(1 - r) / r, _⟩, + have hr0' : r ≠ 0, + { rintro rfl, + rw ←hp₂ at hp₂p₁, + simpa using hp₂p₁ }, + have hr1' : r ≠ 1, + { rintro rfl, + rw ←hp₂ at hp₂p₃, + simpa using hp₂p₃ }, + replace hr0 := hr0.lt_of_ne hr0'.symm, + replace hr1 := hr1.lt_of_ne hr1', + refine ⟨div_neg_of_neg_of_pos (left.neg_neg_iff.2 (sub_pos.2 hr1)) hr0, _⟩, + rw [←hp₂, affine_map.line_map_apply, vsub_vadd_eq_vsub_sub, vsub_vadd_eq_vsub_sub, vsub_self, + zero_sub, smul_neg, smul_smul, div_mul_cancel _ hr0', neg_smul, neg_neg, sub_eq_iff_eq_add, + ←add_smul, sub_add_cancel, one_smul] +end + +/-- If the second of three points is strictly between the other two, the angle at that point +(reversed) is π. -/ +lemma _root_.sbtw.angle₃₂₁_eq_pi {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₃ p₂ p₁ = π := +by rw [←h.angle₁₂₃_eq_pi, angle_comm] + +/-- The angle between three points is π if and only if the second point is strictly between the +other two. -/ +lemma angle_eq_pi_iff_sbtw {p₁ p₂ p₃ : P} : ∠ p₁ p₂ p₃ = π ↔ sbtw ℝ p₁ p₂ p₃ := +begin + refine ⟨_, λ h, h.angle₁₂₃_eq_pi⟩, + rw [angle, angle_eq_pi_iff], + rintro ⟨hp₁p₂, r, hr, hp₃p₂⟩, + refine ⟨⟨1 / (1 - r), + ⟨div_nonneg zero_le_one (sub_nonneg.2 (hr.le.trans zero_le_one)), + (div_le_one (sub_pos.2 (hr.trans zero_lt_one))).2 ((le_sub_self_iff 1).2 hr.le)⟩, _⟩, + (vsub_ne_zero.1 hp₁p₂).symm, _⟩, + { rw ←eq_vadd_iff_vsub_eq at hp₃p₂, + rw [affine_map.line_map_apply, hp₃p₂, vadd_vsub_assoc, ←neg_vsub_eq_vsub_rev p₂ p₁, + smul_neg, ←neg_smul, smul_add, smul_smul, ←add_smul, eq_comm, eq_vadd_iff_vsub_eq], + convert (one_smul ℝ (p₂ -ᵥ p₁)).symm, + field_simp [(sub_pos.2 (hr.trans zero_lt_one)).ne.symm], + abel }, + { rw [ne_comm, ←@vsub_ne_zero V, hp₃p₂, smul_ne_zero_iff], + exact ⟨hr.ne, hp₁p₂⟩ } +end + +/-- If the second of three points is weakly between the other two, and not equal to the first, +the angle at the first point is zero. -/ +lemma _root_.wbtw.angle₂₁₃_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₁ : p₂ ≠ p₁) : + ∠ p₂ p₁ p₃ = 0 := +begin + rw [angle, angle_eq_zero_iff], + rcases h with ⟨r, ⟨hr0, hr1⟩, rfl⟩, + have hr0' : r ≠ 0, { rintro rfl, simpa using hp₂p₁ }, + replace hr0 := hr0.lt_of_ne hr0'.symm, + refine ⟨vsub_ne_zero.2 hp₂p₁, r⁻¹, inv_pos.2 hr0, _⟩, + rw [affine_map.line_map_apply, vadd_vsub_assoc, vsub_self, add_zero, smul_smul, + inv_mul_cancel hr0', one_smul] +end + +/-- If the second of three points is strictly between the other two, the angle at the first point +is zero. -/ +lemma _root_.sbtw.angle₂₁₃_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₂ p₁ p₃ = 0 := +h.wbtw.angle₂₁₃_eq_zero_of_ne h.ne_left + +/-- If the second of three points is weakly between the other two, and not equal to the first, +the angle at the first point (reversed) is zero. -/ +lemma _root_.wbtw.angle₃₁₂_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₁ : p₂ ≠ p₁) : + ∠ p₃ p₁ p₂ = 0 := +by rw [←h.angle₂₁₃_eq_zero_of_ne hp₂p₁, angle_comm] + +/-- If the second of three points is strictly between the other two, the angle at the first point +(reversed) is zero. -/ +lemma _root_.sbtw.angle₃₁₂_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₃ p₁ p₂ = 0 := +h.wbtw.angle₃₁₂_eq_zero_of_ne h.ne_left + +/-- If the second of three points is weakly between the other two, and not equal to the third, +the angle at the third point is zero. -/ +lemma _root_.wbtw.angle₂₃₁_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₃ : p₂ ≠ p₃) : + ∠ p₂ p₃ p₁ = 0 := +h.symm.angle₂₁₃_eq_zero_of_ne hp₂p₃ + +/-- If the second of three points is strictly between the other two, the angle at the third point +is zero. -/ +lemma _root_.sbtw.angle₂₃₁_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₂ p₃ p₁ = 0 := +h.wbtw.angle₂₃₁_eq_zero_of_ne h.ne_right + +/-- If the second of three points is weakly between the other two, and not equal to the third, +the angle at the third point (reversed) is zero. -/ +lemma _root_.wbtw.angle₁₃₂_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₃ : p₂ ≠ p₃) : + ∠ p₁ p₃ p₂ = 0 := +h.symm.angle₃₁₂_eq_zero_of_ne hp₂p₃ + +/-- If the second of three points is strictly between the other two, the angle at the third point +(reversed) is zero. -/ +lemma _root_.sbtw.angle₁₃₂_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₁ p₃ p₂ = 0 := +h.wbtw.angle₁₃₂_eq_zero_of_ne h.ne_right + +/-- The angle between three points is zero if and only if one of the first and third points is +weakly between the other two, and not equal to the second. -/ +lemma angle_eq_zero_iff_ne_and_wbtw {p₁ p₂ p₃ : P} : + ∠ p₁ p₂ p₃ = 0 ↔ (p₁ ≠ p₂ ∧ wbtw ℝ p₂ p₁ p₃) ∨ (p₃ ≠ p₂ ∧ wbtw ℝ p₂ p₃ p₁) := +begin + split, + { rw [angle, angle_eq_zero_iff], + rintro ⟨hp₁p₂, r, hr0, hp₃p₂⟩, + rcases le_or_lt 1 r with hr1 | hr1, + { refine or.inl ⟨vsub_ne_zero.1 hp₁p₂, r⁻¹, ⟨(inv_pos.2 hr0).le, inv_le_one hr1⟩, _⟩, + rw [affine_map.line_map_apply, hp₃p₂, smul_smul, inv_mul_cancel hr0.ne.symm, one_smul, + vsub_vadd] }, + { refine or.inr ⟨_, r, ⟨hr0.le, hr1.le⟩, _⟩, + { rw [←@vsub_ne_zero V, hp₃p₂, smul_ne_zero_iff], + exact ⟨hr0.ne.symm, hp₁p₂⟩ }, + { rw [affine_map.line_map_apply, ←hp₃p₂, vsub_vadd] } } }, + { rintro (⟨hp₁p₂, h⟩ | ⟨hp₃p₂, h⟩), + { exact h.angle₂₁₃_eq_zero_of_ne hp₁p₂ }, + { exact h.angle₃₁₂_eq_zero_of_ne hp₃p₂ } } +end + +/-- The angle between three points is zero if and only if one of the first and third points is +strictly between the other two, or those two points are equal but not equal to the second. -/ +lemma angle_eq_zero_iff_eq_and_ne_or_sbtw {p₁ p₂ p₃ : P} : + ∠ p₁ p₂ p₃ = 0 ↔ (p₁ = p₃ ∧ p₁ ≠ p₂) ∨ sbtw ℝ p₂ p₁ p₃ ∨ sbtw ℝ p₂ p₃ p₁ := +begin + rw angle_eq_zero_iff_ne_and_wbtw, + by_cases hp₁p₂ : p₁ = p₂, { simp [hp₁p₂] }, + by_cases hp₁p₃ : p₁ = p₃, { simp [hp₁p₃] }, + by_cases hp₃p₂ : p₃ = p₂, { simp [hp₃p₂] }, + simp [hp₁p₂, hp₁p₃, ne.symm hp₁p₃, sbtw, hp₃p₂] +end + +end euclidean_geometry diff --git a/src/geometry/euclidean/angle/unoriented/basic.lean b/src/geometry/euclidean/angle/unoriented/basic.lean new file mode 100644 index 0000000000000..cade9bc78797b --- /dev/null +++ b/src/geometry/euclidean/angle/unoriented/basic.lean @@ -0,0 +1,337 @@ +/- +Copyright (c) 2020 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers, Manuel Candales +-/ +import analysis.calculus.conformal.normed_space +import analysis.inner_product_space.basic +import analysis.special_functions.trigonometric.inverse + +/-! +# Angles between vectors + +This file defines unoriented angles in real inner product spaces. + +## Main definitions + +* `inner_product_geometry.angle` is the undirected angle between two vectors. + +-/ + +noncomputable theory +open_locale big_operators +open_locale real +open_locale real_inner_product_space + +namespace inner_product_geometry + +variables {V : Type*} [inner_product_space ℝ V] + +/-- The undirected angle between two vectors. If either vector is 0, +this is π/2. See `orientation.oangle` for the corresponding oriented angle +definition. -/ +def angle (x y : V) : ℝ := real.arccos (inner x y / (∥x∥ * ∥y∥)) + +lemma continuous_at_angle {x : V × V} (hx1 : x.1 ≠ 0) (hx2 : x.2 ≠ 0) : + continuous_at (λ y : V × V, angle y.1 y.2) x := +real.continuous_arccos.continuous_at.comp $ continuous_inner.continuous_at.div + ((continuous_norm.comp continuous_fst).mul (continuous_norm.comp continuous_snd)).continuous_at + (by simp [hx1, hx2]) + +lemma angle_smul_smul {c : ℝ} (hc : c ≠ 0) (x y : V) : + angle (c • x) (c • y) = angle x y := +have c * c ≠ 0, from mul_ne_zero hc hc, +by rw [angle, angle, real_inner_smul_left, inner_smul_right, norm_smul, norm_smul, real.norm_eq_abs, + mul_mul_mul_comm _ (∥x∥), abs_mul_abs_self, ← mul_assoc c c, mul_div_mul_left _ _ this] + +@[simp] lemma _root_.linear_isometry.angle_map {E F : Type*} + [inner_product_space ℝ E] [inner_product_space ℝ F] (f : E →ₗᵢ[ℝ] F) (u v : E) : + angle (f u) (f v) = angle u v := +by rw [angle, angle, f.inner_map_map, f.norm_map, f.norm_map] + +@[simp, norm_cast] lemma _root_.submodule.angle_coe {s : submodule ℝ V} (x y : s) : + angle (x : V) (y : V) = angle x y := +s.subtypeₗᵢ.angle_map x y + +lemma is_conformal_map.preserves_angle {E F : Type*} + [inner_product_space ℝ E] [inner_product_space ℝ F] + {f' : E →L[ℝ] F} (h : is_conformal_map f') (u v : E) : + angle (f' u) (f' v) = angle u v := +begin + obtain ⟨c, hc, li, rfl⟩ := h, + exact (angle_smul_smul hc _ _).trans (li.angle_map _ _) +end + +/-- If a real differentiable map `f` is conformal at a point `x`, + then it preserves the angles at that point. -/ +lemma conformal_at.preserves_angle {E F : Type*} + [inner_product_space ℝ E] [inner_product_space ℝ F] + {f : E → F} {x : E} {f' : E →L[ℝ] F} + (h : has_fderiv_at f f' x) (H : conformal_at f x) (u v : E) : + angle (f' u) (f' v) = angle u v := +let ⟨f₁, h₁, c⟩ := H in h₁.unique h ▸ is_conformal_map.preserves_angle c u v + +/-- The cosine of the angle between two vectors. -/ +lemma cos_angle (x y : V) : real.cos (angle x y) = inner x y / (∥x∥ * ∥y∥) := +real.cos_arccos (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 + (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).2 + +/-- The angle between two vectors does not depend on their order. -/ +lemma angle_comm (x y : V) : angle x y = angle y x := +begin + unfold angle, + rw [real_inner_comm, mul_comm] +end + +/-- The angle between the negation of two vectors. -/ +@[simp] lemma angle_neg_neg (x y : V) : angle (-x) (-y) = angle x y := +begin + unfold angle, + rw [inner_neg_neg, norm_neg, norm_neg] +end + +/-- The angle between two vectors is nonnegative. -/ +lemma angle_nonneg (x y : V) : 0 ≤ angle x y := +real.arccos_nonneg _ + +/-- The angle between two vectors is at most π. -/ +lemma angle_le_pi (x y : V) : angle x y ≤ π := +real.arccos_le_pi _ + +/-- The angle between a vector and the negation of another vector. -/ +lemma angle_neg_right (x y : V) : angle x (-y) = π - angle x y := +begin + unfold angle, + rw [←real.arccos_neg, norm_neg, inner_neg_right, neg_div] +end + +/-- The angle between the negation of a vector and another vector. -/ +lemma angle_neg_left (x y : V) : angle (-x) y = π - angle x y := +by rw [←angle_neg_neg, neg_neg, angle_neg_right] + +/-- The angle between the zero vector and a vector. -/ +@[simp] lemma angle_zero_left (x : V) : angle 0 x = π / 2 := +begin + unfold angle, + rw [inner_zero_left, zero_div, real.arccos_zero] +end + +/-- The angle between a vector and the zero vector. -/ +@[simp] lemma angle_zero_right (x : V) : angle x 0 = π / 2 := +begin + unfold angle, + rw [inner_zero_right, zero_div, real.arccos_zero] +end + +/-- The angle between a nonzero vector and itself. -/ +@[simp] lemma angle_self {x : V} (hx : x ≠ 0) : angle x x = 0 := +begin + unfold angle, + rw [←real_inner_self_eq_norm_mul_norm, div_self (λ h, hx (inner_self_eq_zero.1 h)), + real.arccos_one] +end + +/-- The angle between a nonzero vector and its negation. -/ +@[simp] lemma angle_self_neg_of_nonzero {x : V} (hx : x ≠ 0) : angle x (-x) = π := +by rw [angle_neg_right, angle_self hx, sub_zero] + +/-- The angle between the negation of a nonzero vector and that +vector. -/ +@[simp] lemma angle_neg_self_of_nonzero {x : V} (hx : x ≠ 0) : angle (-x) x = π := +by rw [angle_comm, angle_self_neg_of_nonzero hx] + +/-- The angle between a vector and a positive multiple of a vector. -/ +@[simp] lemma angle_smul_right_of_pos (x y : V) {r : ℝ} (hr : 0 < r) : + angle x (r • y) = angle x y := +begin + unfold angle, + rw [inner_smul_right, norm_smul, real.norm_eq_abs, abs_of_nonneg (le_of_lt hr), ←mul_assoc, + mul_comm _ r, mul_assoc, mul_div_mul_left _ _ (ne_of_gt hr)] +end + +/-- The angle between a positive multiple of a vector and a vector. -/ +@[simp] lemma angle_smul_left_of_pos (x y : V) {r : ℝ} (hr : 0 < r) : + angle (r • x) y = angle x y := +by rw [angle_comm, angle_smul_right_of_pos y x hr, angle_comm] + +/-- The angle between a vector and a negative multiple of a vector. -/ +@[simp] lemma angle_smul_right_of_neg (x y : V) {r : ℝ} (hr : r < 0) : + angle x (r • y) = angle x (-y) := +by rw [←neg_neg r, neg_smul, angle_neg_right, angle_smul_right_of_pos x y (neg_pos_of_neg hr), + angle_neg_right] + +/-- The angle between a negative multiple of a vector and a vector. -/ +@[simp] lemma angle_smul_left_of_neg (x y : V) {r : ℝ} (hr : r < 0) : + angle (r • x) y = angle (-x) y := +by rw [angle_comm, angle_smul_right_of_neg y x hr, angle_comm] + +/-- The cosine of the angle between two vectors, multiplied by the +product of their norms. -/ +lemma cos_angle_mul_norm_mul_norm (x y : V) : real.cos (angle x y) * (∥x∥ * ∥y∥) = inner x y := +begin + rw [cos_angle, div_mul_cancel_of_imp], + simp [or_imp_distrib] { contextual := tt }, +end + +/-- The sine of the angle between two vectors, multiplied by the +product of their norms. -/ +lemma sin_angle_mul_norm_mul_norm (x y : V) : real.sin (angle x y) * (∥x∥ * ∥y∥) = + real.sqrt (inner x x * inner y y - inner x y * inner x y) := +begin + unfold angle, + rw [real.sin_arccos (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 + (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).2, + ←real.sqrt_mul_self (mul_nonneg (norm_nonneg x) (norm_nonneg y)), + ←real.sqrt_mul' _ (mul_self_nonneg _), sq, + real.sqrt_mul_self (mul_nonneg (norm_nonneg x) (norm_nonneg y)), + real_inner_self_eq_norm_mul_norm, + real_inner_self_eq_norm_mul_norm], + by_cases h : (∥x∥ * ∥y∥) = 0, + { rw [(show ∥x∥ * ∥x∥ * (∥y∥ * ∥y∥) = (∥x∥ * ∥y∥) * (∥x∥ * ∥y∥), by ring), h, mul_zero, mul_zero, + zero_sub], + cases eq_zero_or_eq_zero_of_mul_eq_zero h with hx hy, + { rw norm_eq_zero at hx, + rw [hx, inner_zero_left, zero_mul, neg_zero] }, + { rw norm_eq_zero at hy, + rw [hy, inner_zero_right, zero_mul, neg_zero] } }, + { field_simp [h], ring_nf } +end + +/-- The angle between two vectors is zero if and only if they are +nonzero and one is a positive multiple of the other. -/ +lemma angle_eq_zero_iff {x y : V} : angle x y = 0 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), 0 < r ∧ y = r • x) := +begin + rw [angle, ← real_inner_div_norm_mul_norm_eq_one_iff, real.arccos_eq_zero, has_le.le.le_iff_eq, + eq_comm], + exact (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).2 +end + +/-- The angle between two vectors is π if and only if they are nonzero +and one is a negative multiple of the other. -/ +lemma angle_eq_pi_iff {x y : V} : angle x y = π ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r < 0 ∧ y = r • x) := +begin + rw [angle, ← real_inner_div_norm_mul_norm_eq_neg_one_iff, real.arccos_eq_pi, has_le.le.le_iff_eq], + exact (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 +end + +/-- If the angle between two vectors is π, the angles between those +vectors and a third vector add to π. -/ +lemma angle_add_angle_eq_pi_of_angle_eq_pi {x y : V} (z : V) (h : angle x y = π) : + angle x z + angle y z = π := +begin + rcases angle_eq_pi_iff.1 h with ⟨hx, ⟨r, ⟨hr, rfl⟩⟩⟩, + rw [angle_smul_left_of_neg x z hr, angle_neg_left, add_sub_cancel'_right] +end + +/-- Two vectors have inner product 0 if and only if the angle between +them is π/2. -/ +lemma inner_eq_zero_iff_angle_eq_pi_div_two (x y : V) : ⟪x, y⟫ = 0 ↔ angle x y = π / 2 := +iff.symm $ by simp [angle, or_imp_distrib] { contextual := tt } + +/-- If the angle between two vectors is π, the inner product equals the negative product +of the norms. -/ +lemma inner_eq_neg_mul_norm_of_angle_eq_pi {x y : V} (h : angle x y = π) : ⟪x, y⟫ = - (∥x∥ * ∥y∥) := +by simp [← cos_angle_mul_norm_mul_norm, h] + +/-- If the angle between two vectors is 0, the inner product equals the product of the norms. -/ +lemma inner_eq_mul_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : ⟪x, y⟫ = ∥x∥ * ∥y∥ := +by simp [← cos_angle_mul_norm_mul_norm, h] + +/-- The inner product of two non-zero vectors equals the negative product of their norms +if and only if the angle between the two vectors is π. -/ +lemma inner_eq_neg_mul_norm_iff_angle_eq_pi {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + ⟪x, y⟫ = - (∥x∥ * ∥y∥) ↔ angle x y = π := +begin + refine ⟨λ h, _, inner_eq_neg_mul_norm_of_angle_eq_pi⟩, + have h₁ : (∥x∥ * ∥y∥) ≠ 0 := (mul_pos (norm_pos_iff.mpr hx) (norm_pos_iff.mpr hy)).ne', + rw [angle, h, neg_div, div_self h₁, real.arccos_neg_one], +end + +/-- The inner product of two non-zero vectors equals the product of their norms +if and only if the angle between the two vectors is 0. -/ +lemma inner_eq_mul_norm_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + ⟪x, y⟫ = ∥x∥ * ∥y∥ ↔ angle x y = 0 := +begin + refine ⟨λ h, _, inner_eq_mul_norm_of_angle_eq_zero⟩, + have h₁ : (∥x∥ * ∥y∥) ≠ 0 := (mul_pos (norm_pos_iff.mpr hx) (norm_pos_iff.mpr hy)).ne', + rw [angle, h, div_self h₁, real.arccos_one], +end + +/-- If the angle between two vectors is π, the norm of their difference equals +the sum of their norms. -/ +lemma norm_sub_eq_add_norm_of_angle_eq_pi {x y : V} (h : angle x y = π) : ∥x - y∥ = ∥x∥ + ∥y∥ := +begin + rw ← sq_eq_sq (norm_nonneg (x - y)) (add_nonneg (norm_nonneg x) (norm_nonneg y)), + rw [norm_sub_pow_two_real, inner_eq_neg_mul_norm_of_angle_eq_pi h], + ring, +end + +/-- If the angle between two vectors is 0, the norm of their sum equals +the sum of their norms. -/ +lemma norm_add_eq_add_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : ∥x + y∥ = ∥x∥ + ∥y∥ := +begin + rw ← sq_eq_sq (norm_nonneg (x + y)) (add_nonneg (norm_nonneg x) (norm_nonneg y)), + rw [norm_add_pow_two_real, inner_eq_mul_norm_of_angle_eq_zero h], + ring, +end + +/-- If the angle between two vectors is 0, the norm of their difference equals +the absolute value of the difference of their norms. -/ +lemma norm_sub_eq_abs_sub_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : + ∥x - y∥ = |∥x∥ - ∥y∥| := +begin + rw [← sq_eq_sq (norm_nonneg (x - y)) (abs_nonneg (∥x∥ - ∥y∥)), + norm_sub_pow_two_real, inner_eq_mul_norm_of_angle_eq_zero h, sq_abs (∥x∥ - ∥y∥)], + ring, +end + +/-- The norm of the difference of two non-zero vectors equals the sum of their norms +if and only the angle between the two vectors is π. -/ +lemma norm_sub_eq_add_norm_iff_angle_eq_pi {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + ∥x - y∥ = ∥x∥ + ∥y∥ ↔ angle x y = π := +begin + refine ⟨λ h, _, norm_sub_eq_add_norm_of_angle_eq_pi⟩, + rw ← inner_eq_neg_mul_norm_iff_angle_eq_pi hx hy, + obtain ⟨hxy₁, hxy₂⟩ := ⟨norm_nonneg (x - y), add_nonneg (norm_nonneg x) (norm_nonneg y)⟩, + rw [← sq_eq_sq hxy₁ hxy₂, norm_sub_pow_two_real] at h, + calc inner x y = (∥x∥ ^ 2 + ∥y∥ ^ 2 - (∥x∥ + ∥y∥) ^ 2) / 2 : by linarith + ... = -(∥x∥ * ∥y∥) : by ring, +end + +/-- The norm of the sum of two non-zero vectors equals the sum of their norms +if and only the angle between the two vectors is 0. -/ +lemma norm_add_eq_add_norm_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + ∥x + y∥ = ∥x∥ + ∥y∥ ↔ angle x y = 0 := +begin + refine ⟨λ h, _, norm_add_eq_add_norm_of_angle_eq_zero⟩, + rw ← inner_eq_mul_norm_iff_angle_eq_zero hx hy, + obtain ⟨hxy₁, hxy₂⟩ := ⟨norm_nonneg (x + y), add_nonneg (norm_nonneg x) (norm_nonneg y)⟩, + rw [← sq_eq_sq hxy₁ hxy₂, norm_add_pow_two_real] at h, + calc inner x y = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith + ... = ∥x∥ * ∥y∥ : by ring, +end + +/-- The norm of the difference of two non-zero vectors equals the absolute value +of the difference of their norms if and only the angle between the two vectors is 0. -/ +lemma norm_sub_eq_abs_sub_norm_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : + ∥x - y∥ = |∥x∥ - ∥y∥| ↔ angle x y = 0 := +begin + refine ⟨λ h, _, norm_sub_eq_abs_sub_norm_of_angle_eq_zero⟩, + rw ← inner_eq_mul_norm_iff_angle_eq_zero hx hy, + have h1 : ∥x - y∥ ^ 2 = (∥x∥ - ∥y∥) ^ 2, { rw h, exact sq_abs (∥x∥ - ∥y∥) }, + rw norm_sub_pow_two_real at h1, + calc inner x y = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith + ... = ∥x∥ * ∥y∥ : by ring, +end + +/-- The norm of the sum of two vectors equals the norm of their difference if and only if +the angle between them is π/2. -/ +lemma norm_add_eq_norm_sub_iff_angle_eq_pi_div_two (x y : V) : + ∥x + y∥ = ∥x - y∥ ↔ angle x y = π / 2 := +begin + rw [← sq_eq_sq (norm_nonneg (x + y)) (norm_nonneg (x - y)), + ← inner_eq_zero_iff_angle_eq_pi_div_two x y, norm_add_pow_two_real, norm_sub_pow_two_real], + split; intro h; linarith, +end + +end inner_product_geometry diff --git a/src/geometry/euclidean/angle/unoriented/right_angle.lean b/src/geometry/euclidean/angle/unoriented/right_angle.lean new file mode 100644 index 0000000000000..a8b75b85e3a8b --- /dev/null +++ b/src/geometry/euclidean/angle/unoriented/right_angle.lean @@ -0,0 +1,81 @@ +/- +Copyright (c) 2020 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers +-/ +import geometry.euclidean.angle.unoriented.affine + +/-! +# Right-angled triangles + +This file proves basic geometrical results about distances and angles in (possibly degenerate) +right-angled triangles in real inner product spaces and Euclidean affine spaces. + +## Implementation notes + +Results in this file are generally given in a form with only those non-degeneracy conditions +needed for the particular result, rather than requiring affine independence of the points of a +triangle unnecessarily. + +## References + +* https://en.wikipedia.org/wiki/Pythagorean_theorem + +-/ + +noncomputable theory +open_locale big_operators +open_locale euclidean_geometry +open_locale real +open_locale real_inner_product_space + +namespace inner_product_geometry + +variables {V : Type*} [inner_product_space ℝ V] + +/-- Pythagorean theorem, if-and-only-if vector angle form. -/ +lemma norm_add_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two (x y : V) : + ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ ↔ angle x y = π / 2 := +begin + rw norm_add_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero, + exact inner_eq_zero_iff_angle_eq_pi_div_two x y +end + +/-- Pythagorean theorem, vector angle form. -/ +lemma norm_add_sq_eq_norm_sq_add_norm_sq' (x y : V) (h : angle x y = π / 2) : + ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ := +(norm_add_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two x y).2 h + +/-- Pythagorean theorem, subtracting vectors, if-and-only-if vector angle form. -/ +lemma norm_sub_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two (x y : V) : + ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ ↔ angle x y = π / 2 := +begin + rw norm_sub_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero, + exact inner_eq_zero_iff_angle_eq_pi_div_two x y +end + +/-- Pythagorean theorem, subtracting vectors, vector angle form. -/ +lemma norm_sub_sq_eq_norm_sq_add_norm_sq' (x y : V) (h : angle x y = π / 2) : + ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ := +(norm_sub_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two x y).2 h + +end inner_product_geometry + +namespace euclidean_geometry + +open inner_product_geometry + +variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] + [normed_add_torsor V P] +include V + +/-- **Pythagorean theorem**, if-and-only-if angle-at-point form. -/ +lemma dist_sq_eq_dist_sq_add_dist_sq_iff_angle_eq_pi_div_two (p1 p2 p3 : P) : + dist p1 p3 * dist p1 p3 = dist p1 p2 * dist p1 p2 + dist p3 p2 * dist p3 p2 ↔ + ∠ p1 p2 p3 = π / 2 := +by erw [dist_comm p3 p2, dist_eq_norm_vsub V p1 p3, dist_eq_norm_vsub V p1 p2, + dist_eq_norm_vsub V p2 p3, + ←norm_sub_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two, + vsub_sub_vsub_cancel_right p1, ←neg_vsub_eq_vsub_rev p2 p3, norm_neg] + +end euclidean_geometry diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index 9d6dc284237ea..3de9a8c31ed93 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -3,12 +3,8 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Manuel Candales -/ -import analysis.convex.between import analysis.inner_product_space.projection -import analysis.special_functions.trigonometric.inverse import algebra.quadratic_discriminant -import linear_algebra.affine_space.finite_dimensional -import analysis.calculus.conformal.normed_space /-! # Euclidean spaces @@ -22,12 +18,6 @@ proofs or more geometrical content generally go in separate files. ## Main definitions -* `inner_product_geometry.angle` is the undirected angle between two - vectors. - -* `euclidean_geometry.angle`, with notation `∠`, is the undirected - angle determined by three points. - * `euclidean_geometry.orthogonal_projection` is the orthogonal projection of a point onto an affine subspace. @@ -55,330 +45,8 @@ theorems that need it. noncomputable theory open_locale big_operators open_locale classical -open_locale real open_locale real_inner_product_space -namespace inner_product_geometry -/-! -### Geometrical results on real inner product spaces - -This section develops some geometrical definitions and results on real -inner product spaces, where those definitions and results can most -conveniently be developed in terms of vectors and then used to deduce -corresponding results for Euclidean affine spaces. --/ - -variables {V : Type*} [inner_product_space ℝ V] - -/-- The undirected angle between two vectors. If either vector is 0, -this is π/2. See `orientation.oangle` for the corresponding oriented angle -definition. -/ -def angle (x y : V) : ℝ := real.arccos (inner x y / (∥x∥ * ∥y∥)) - -lemma continuous_at_angle {x : V × V} (hx1 : x.1 ≠ 0) (hx2 : x.2 ≠ 0) : - continuous_at (λ y : V × V, angle y.1 y.2) x := -real.continuous_arccos.continuous_at.comp $ continuous_inner.continuous_at.div - ((continuous_norm.comp continuous_fst).mul (continuous_norm.comp continuous_snd)).continuous_at - (by simp [hx1, hx2]) - -lemma angle_smul_smul {c : ℝ} (hc : c ≠ 0) (x y : V) : - angle (c • x) (c • y) = angle x y := -have c * c ≠ 0, from mul_ne_zero hc hc, -by rw [angle, angle, real_inner_smul_left, inner_smul_right, norm_smul, norm_smul, real.norm_eq_abs, - mul_mul_mul_comm _ (∥x∥), abs_mul_abs_self, ← mul_assoc c c, mul_div_mul_left _ _ this] - -@[simp] lemma _root_.linear_isometry.angle_map {E F : Type*} - [inner_product_space ℝ E] [inner_product_space ℝ F] (f : E →ₗᵢ[ℝ] F) (u v : E) : - angle (f u) (f v) = angle u v := -by rw [angle, angle, f.inner_map_map, f.norm_map, f.norm_map] - -@[simp, norm_cast] lemma _root_.submodule.angle_coe {s : submodule ℝ V} (x y : s) : - angle (x : V) (y : V) = angle x y := -s.subtypeₗᵢ.angle_map x y - -lemma is_conformal_map.preserves_angle {E F : Type*} - [inner_product_space ℝ E] [inner_product_space ℝ F] - {f' : E →L[ℝ] F} (h : is_conformal_map f') (u v : E) : - angle (f' u) (f' v) = angle u v := -begin - obtain ⟨c, hc, li, rfl⟩ := h, - exact (angle_smul_smul hc _ _).trans (li.angle_map _ _) -end - -/-- If a real differentiable map `f` is conformal at a point `x`, - then it preserves the angles at that point. -/ -lemma conformal_at.preserves_angle {E F : Type*} - [inner_product_space ℝ E] [inner_product_space ℝ F] - {f : E → F} {x : E} {f' : E →L[ℝ] F} - (h : has_fderiv_at f f' x) (H : conformal_at f x) (u v : E) : - angle (f' u) (f' v) = angle u v := -let ⟨f₁, h₁, c⟩ := H in h₁.unique h ▸ is_conformal_map.preserves_angle c u v - -/-- The cosine of the angle between two vectors. -/ -lemma cos_angle (x y : V) : real.cos (angle x y) = inner x y / (∥x∥ * ∥y∥) := -real.cos_arccos (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 - (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).2 - -/-- The angle between two vectors does not depend on their order. -/ -lemma angle_comm (x y : V) : angle x y = angle y x := -begin - unfold angle, - rw [real_inner_comm, mul_comm] -end - -/-- The angle between the negation of two vectors. -/ -@[simp] lemma angle_neg_neg (x y : V) : angle (-x) (-y) = angle x y := -begin - unfold angle, - rw [inner_neg_neg, norm_neg, norm_neg] -end - -/-- The angle between two vectors is nonnegative. -/ -lemma angle_nonneg (x y : V) : 0 ≤ angle x y := -real.arccos_nonneg _ - -/-- The angle between two vectors is at most π. -/ -lemma angle_le_pi (x y : V) : angle x y ≤ π := -real.arccos_le_pi _ - -/-- The angle between a vector and the negation of another vector. -/ -lemma angle_neg_right (x y : V) : angle x (-y) = π - angle x y := -begin - unfold angle, - rw [←real.arccos_neg, norm_neg, inner_neg_right, neg_div] -end - -/-- The angle between the negation of a vector and another vector. -/ -lemma angle_neg_left (x y : V) : angle (-x) y = π - angle x y := -by rw [←angle_neg_neg, neg_neg, angle_neg_right] - -/-- The angle between the zero vector and a vector. -/ -@[simp] lemma angle_zero_left (x : V) : angle 0 x = π / 2 := -begin - unfold angle, - rw [inner_zero_left, zero_div, real.arccos_zero] -end - -/-- The angle between a vector and the zero vector. -/ -@[simp] lemma angle_zero_right (x : V) : angle x 0 = π / 2 := -begin - unfold angle, - rw [inner_zero_right, zero_div, real.arccos_zero] -end - -/-- The angle between a nonzero vector and itself. -/ -@[simp] lemma angle_self {x : V} (hx : x ≠ 0) : angle x x = 0 := -begin - unfold angle, - rw [←real_inner_self_eq_norm_mul_norm, div_self (λ h, hx (inner_self_eq_zero.1 h)), - real.arccos_one] -end - -/-- The angle between a nonzero vector and its negation. -/ -@[simp] lemma angle_self_neg_of_nonzero {x : V} (hx : x ≠ 0) : angle x (-x) = π := -by rw [angle_neg_right, angle_self hx, sub_zero] - -/-- The angle between the negation of a nonzero vector and that -vector. -/ -@[simp] lemma angle_neg_self_of_nonzero {x : V} (hx : x ≠ 0) : angle (-x) x = π := -by rw [angle_comm, angle_self_neg_of_nonzero hx] - -/-- The angle between a vector and a positive multiple of a vector. -/ -@[simp] lemma angle_smul_right_of_pos (x y : V) {r : ℝ} (hr : 0 < r) : - angle x (r • y) = angle x y := -begin - unfold angle, - rw [inner_smul_right, norm_smul, real.norm_eq_abs, abs_of_nonneg (le_of_lt hr), ←mul_assoc, - mul_comm _ r, mul_assoc, mul_div_mul_left _ _ (ne_of_gt hr)] -end - -/-- The angle between a positive multiple of a vector and a vector. -/ -@[simp] lemma angle_smul_left_of_pos (x y : V) {r : ℝ} (hr : 0 < r) : - angle (r • x) y = angle x y := -by rw [angle_comm, angle_smul_right_of_pos y x hr, angle_comm] - -/-- The angle between a vector and a negative multiple of a vector. -/ -@[simp] lemma angle_smul_right_of_neg (x y : V) {r : ℝ} (hr : r < 0) : - angle x (r • y) = angle x (-y) := -by rw [←neg_neg r, neg_smul, angle_neg_right, angle_smul_right_of_pos x y (neg_pos_of_neg hr), - angle_neg_right] - -/-- The angle between a negative multiple of a vector and a vector. -/ -@[simp] lemma angle_smul_left_of_neg (x y : V) {r : ℝ} (hr : r < 0) : - angle (r • x) y = angle (-x) y := -by rw [angle_comm, angle_smul_right_of_neg y x hr, angle_comm] - -/-- The cosine of the angle between two vectors, multiplied by the -product of their norms. -/ -lemma cos_angle_mul_norm_mul_norm (x y : V) : real.cos (angle x y) * (∥x∥ * ∥y∥) = inner x y := -begin - rw [cos_angle, div_mul_cancel_of_imp], - simp [or_imp_distrib] { contextual := tt }, -end - -/-- The sine of the angle between two vectors, multiplied by the -product of their norms. -/ -lemma sin_angle_mul_norm_mul_norm (x y : V) : real.sin (angle x y) * (∥x∥ * ∥y∥) = - real.sqrt (inner x x * inner y y - inner x y * inner x y) := -begin - unfold angle, - rw [real.sin_arccos (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 - (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).2, - ←real.sqrt_mul_self (mul_nonneg (norm_nonneg x) (norm_nonneg y)), - ←real.sqrt_mul' _ (mul_self_nonneg _), sq, - real.sqrt_mul_self (mul_nonneg (norm_nonneg x) (norm_nonneg y)), - real_inner_self_eq_norm_mul_norm, - real_inner_self_eq_norm_mul_norm], - by_cases h : (∥x∥ * ∥y∥) = 0, - { rw [(show ∥x∥ * ∥x∥ * (∥y∥ * ∥y∥) = (∥x∥ * ∥y∥) * (∥x∥ * ∥y∥), by ring), h, mul_zero, mul_zero, - zero_sub], - cases eq_zero_or_eq_zero_of_mul_eq_zero h with hx hy, - { rw norm_eq_zero at hx, - rw [hx, inner_zero_left, zero_mul, neg_zero] }, - { rw norm_eq_zero at hy, - rw [hy, inner_zero_right, zero_mul, neg_zero] } }, - { field_simp [h], ring_nf } -end - -/-- The angle between two vectors is zero if and only if they are -nonzero and one is a positive multiple of the other. -/ -lemma angle_eq_zero_iff {x y : V} : angle x y = 0 ↔ (x ≠ 0 ∧ ∃ (r : ℝ), 0 < r ∧ y = r • x) := -begin - rw [angle, ← real_inner_div_norm_mul_norm_eq_one_iff, real.arccos_eq_zero, has_le.le.le_iff_eq, - eq_comm], - exact (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).2 -end - -/-- The angle between two vectors is π if and only if they are nonzero -and one is a negative multiple of the other. -/ -lemma angle_eq_pi_iff {x y : V} : angle x y = π ↔ (x ≠ 0 ∧ ∃ (r : ℝ), r < 0 ∧ y = r • x) := -begin - rw [angle, ← real_inner_div_norm_mul_norm_eq_neg_one_iff, real.arccos_eq_pi, has_le.le.le_iff_eq], - exact (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 -end - -/-- If the angle between two vectors is π, the angles between those -vectors and a third vector add to π. -/ -lemma angle_add_angle_eq_pi_of_angle_eq_pi {x y : V} (z : V) (h : angle x y = π) : - angle x z + angle y z = π := -begin - rcases angle_eq_pi_iff.1 h with ⟨hx, ⟨r, ⟨hr, rfl⟩⟩⟩, - rw [angle_smul_left_of_neg x z hr, angle_neg_left, add_sub_cancel'_right] -end - -/-- Two vectors have inner product 0 if and only if the angle between -them is π/2. -/ -lemma inner_eq_zero_iff_angle_eq_pi_div_two (x y : V) : ⟪x, y⟫ = 0 ↔ angle x y = π / 2 := -iff.symm $ by simp [angle, or_imp_distrib] { contextual := tt } - -/-- If the angle between two vectors is π, the inner product equals the negative product -of the norms. -/ -lemma inner_eq_neg_mul_norm_of_angle_eq_pi {x y : V} (h : angle x y = π) : ⟪x, y⟫ = - (∥x∥ * ∥y∥) := -by simp [← cos_angle_mul_norm_mul_norm, h] - -/-- If the angle between two vectors is 0, the inner product equals the product of the norms. -/ -lemma inner_eq_mul_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : ⟪x, y⟫ = ∥x∥ * ∥y∥ := -by simp [← cos_angle_mul_norm_mul_norm, h] - -/-- The inner product of two non-zero vectors equals the negative product of their norms -if and only if the angle between the two vectors is π. -/ -lemma inner_eq_neg_mul_norm_iff_angle_eq_pi {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ⟪x, y⟫ = - (∥x∥ * ∥y∥) ↔ angle x y = π := -begin - refine ⟨λ h, _, inner_eq_neg_mul_norm_of_angle_eq_pi⟩, - have h₁ : (∥x∥ * ∥y∥) ≠ 0 := (mul_pos (norm_pos_iff.mpr hx) (norm_pos_iff.mpr hy)).ne', - rw [angle, h, neg_div, div_self h₁, real.arccos_neg_one], -end - -/-- The inner product of two non-zero vectors equals the product of their norms -if and only if the angle between the two vectors is 0. -/ -lemma inner_eq_mul_norm_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ⟪x, y⟫ = ∥x∥ * ∥y∥ ↔ angle x y = 0 := -begin - refine ⟨λ h, _, inner_eq_mul_norm_of_angle_eq_zero⟩, - have h₁ : (∥x∥ * ∥y∥) ≠ 0 := (mul_pos (norm_pos_iff.mpr hx) (norm_pos_iff.mpr hy)).ne', - rw [angle, h, div_self h₁, real.arccos_one], -end - -/-- If the angle between two vectors is π, the norm of their difference equals -the sum of their norms. -/ -lemma norm_sub_eq_add_norm_of_angle_eq_pi {x y : V} (h : angle x y = π) : ∥x - y∥ = ∥x∥ + ∥y∥ := -begin - rw ← sq_eq_sq (norm_nonneg (x - y)) (add_nonneg (norm_nonneg x) (norm_nonneg y)), - rw [norm_sub_pow_two_real, inner_eq_neg_mul_norm_of_angle_eq_pi h], - ring, -end - -/-- If the angle between two vectors is 0, the norm of their sum equals -the sum of their norms. -/ -lemma norm_add_eq_add_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : ∥x + y∥ = ∥x∥ + ∥y∥ := -begin - rw ← sq_eq_sq (norm_nonneg (x + y)) (add_nonneg (norm_nonneg x) (norm_nonneg y)), - rw [norm_add_pow_two_real, inner_eq_mul_norm_of_angle_eq_zero h], - ring, -end - -/-- If the angle between two vectors is 0, the norm of their difference equals -the absolute value of the difference of their norms. -/ -lemma norm_sub_eq_abs_sub_norm_of_angle_eq_zero {x y : V} (h : angle x y = 0) : - ∥x - y∥ = |∥x∥ - ∥y∥| := -begin - rw [← sq_eq_sq (norm_nonneg (x - y)) (abs_nonneg (∥x∥ - ∥y∥)), - norm_sub_pow_two_real, inner_eq_mul_norm_of_angle_eq_zero h, sq_abs (∥x∥ - ∥y∥)], - ring, -end - -/-- The norm of the difference of two non-zero vectors equals the sum of their norms -if and only the angle between the two vectors is π. -/ -lemma norm_sub_eq_add_norm_iff_angle_eq_pi {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ∥x - y∥ = ∥x∥ + ∥y∥ ↔ angle x y = π := -begin - refine ⟨λ h, _, norm_sub_eq_add_norm_of_angle_eq_pi⟩, - rw ← inner_eq_neg_mul_norm_iff_angle_eq_pi hx hy, - obtain ⟨hxy₁, hxy₂⟩ := ⟨norm_nonneg (x - y), add_nonneg (norm_nonneg x) (norm_nonneg y)⟩, - rw [← sq_eq_sq hxy₁ hxy₂, norm_sub_pow_two_real] at h, - calc inner x y = (∥x∥ ^ 2 + ∥y∥ ^ 2 - (∥x∥ + ∥y∥) ^ 2) / 2 : by linarith - ... = -(∥x∥ * ∥y∥) : by ring, -end - -/-- The norm of the sum of two non-zero vectors equals the sum of their norms -if and only the angle between the two vectors is 0. -/ -lemma norm_add_eq_add_norm_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ∥x + y∥ = ∥x∥ + ∥y∥ ↔ angle x y = 0 := -begin - refine ⟨λ h, _, norm_add_eq_add_norm_of_angle_eq_zero⟩, - rw ← inner_eq_mul_norm_iff_angle_eq_zero hx hy, - obtain ⟨hxy₁, hxy₂⟩ := ⟨norm_nonneg (x + y), add_nonneg (norm_nonneg x) (norm_nonneg y)⟩, - rw [← sq_eq_sq hxy₁ hxy₂, norm_add_pow_two_real] at h, - calc inner x y = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith - ... = ∥x∥ * ∥y∥ : by ring, -end - -/-- The norm of the difference of two non-zero vectors equals the absolute value -of the difference of their norms if and only the angle between the two vectors is 0. -/ -lemma norm_sub_eq_abs_sub_norm_iff_angle_eq_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) : - ∥x - y∥ = |∥x∥ - ∥y∥| ↔ angle x y = 0 := -begin - refine ⟨λ h, _, norm_sub_eq_abs_sub_norm_of_angle_eq_zero⟩, - rw ← inner_eq_mul_norm_iff_angle_eq_zero hx hy, - have h1 : ∥x - y∥ ^ 2 = (∥x∥ - ∥y∥) ^ 2, { rw h, exact sq_abs (∥x∥ - ∥y∥) }, - rw norm_sub_pow_two_real at h1, - calc inner x y = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith - ... = ∥x∥ * ∥y∥ : by ring, -end - -/-- The norm of the sum of two vectors equals the norm of their difference if and only if -the angle between them is π/2. -/ -lemma norm_add_eq_norm_sub_iff_angle_eq_pi_div_two (x y : V) : - ∥x + y∥ = ∥x - y∥ ↔ angle x y = π / 2 := -begin - rw [← sq_eq_sq (norm_nonneg (x + y)) (norm_nonneg (x - y)), - ← inner_eq_zero_iff_angle_eq_pi_div_two x y, norm_add_pow_two_real, norm_sub_pow_two_real], - split; intro h; linarith, -end - -end inner_product_geometry - namespace euclidean_geometry /-! ### Geometrical results on Euclidean affine spaces @@ -386,333 +54,17 @@ namespace euclidean_geometry This section develops some geometrical definitions and results on Euclidean affine spaces. -/ -open inner_product_geometry variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] local notation `⟪`x`, `y`⟫` := @inner ℝ V _ x y include V -/-- The undirected angle at `p2` between the line segments to `p1` and -`p3`. If either of those points equals `p2`, this is π/2. Use -`open_locale euclidean_geometry` to access the `∠ p1 p2 p3` -notation. -/ -def angle (p1 p2 p3 : P) : ℝ := angle (p1 -ᵥ p2 : V) (p3 -ᵥ p2) - -localized "notation (name := angle) `∠` := euclidean_geometry.angle" in euclidean_geometry - -lemma continuous_at_angle {x : P × P × P} (hx12 : x.1 ≠ x.2.1) (hx32 : x.2.2 ≠ x.2.1) : - continuous_at (λ y : P × P × P, ∠ y.1 y.2.1 y.2.2) x := -begin - let f : P × P × P → V × V := λ y, (y.1 -ᵥ y.2.1, y.2.2 -ᵥ y.2.1), - have hf1 : (f x).1 ≠ 0, by simp [hx12], - have hf2 : (f x).2 ≠ 0, by simp [hx32], - exact (inner_product_geometry.continuous_at_angle hf1 hf2).comp - ((continuous_fst.vsub continuous_snd.fst).prod_mk - (continuous_snd.snd.vsub continuous_snd.fst)).continuous_at -end - -@[simp] lemma _root_.affine_isometry.angle_map {V₂ P₂ : Type*} [inner_product_space ℝ V₂] - [metric_space P₂] [normed_add_torsor V₂ P₂] (f : P →ᵃⁱ[ℝ] P₂) (p₁ p₂ p₃ : P) : - ∠ (f p₁) (f p₂) (f p₃) = ∠ p₁ p₂ p₃ := -by simp_rw [angle, ←affine_isometry.map_vsub, linear_isometry.angle_map] - -@[simp, norm_cast] lemma _root_.affine_subspace.angle_coe {s : affine_subspace ℝ P} - (p₁ p₂ p₃ : s) : - by haveI : nonempty s := ⟨p₁⟩; exact ∠ (p₁ : P) (p₂ : P) (p₃ : P) = ∠ p₁ p₂ p₃ := -by haveI : nonempty s := ⟨p₁⟩; exact s.subtypeₐᵢ.angle_map p₁ p₂ p₃ - -/-- The angle at a point does not depend on the order of the other two -points. -/ -lemma angle_comm (p1 p2 p3 : P) : ∠ p1 p2 p3 = ∠ p3 p2 p1 := -angle_comm _ _ - -/-- The angle at a point is nonnegative. -/ -lemma angle_nonneg (p1 p2 p3 : P) : 0 ≤ ∠ p1 p2 p3 := -angle_nonneg _ _ - -/-- The angle at a point is at most π. -/ -lemma angle_le_pi (p1 p2 p3 : P) : ∠ p1 p2 p3 ≤ π := -angle_le_pi _ _ - -/-- The angle ∠AAB at a point. -/ -lemma angle_eq_left (p1 p2 : P) : ∠ p1 p1 p2 = π / 2 := -begin - unfold angle, - rw vsub_self, - exact angle_zero_left _ -end - -/-- The angle ∠ABB at a point. -/ -lemma angle_eq_right (p1 p2 : P) : ∠ p1 p2 p2 = π / 2 := -by rw [angle_comm, angle_eq_left] - -/-- The angle ∠ABA at a point. -/ -lemma angle_eq_of_ne {p1 p2 : P} (h : p1 ≠ p2) : ∠ p1 p2 p1 = 0 := -angle_self (λ he, h (vsub_eq_zero_iff_eq.1 he)) - -/-- If the angle ∠ABC at a point is π, the angle ∠BAC is 0. -/ -lemma angle_eq_zero_of_angle_eq_pi_left {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : - ∠ p2 p1 p3 = 0 := -begin - unfold angle at h, - rw angle_eq_pi_iff at h, - rcases h with ⟨hp1p2, ⟨r, ⟨hr, hpr⟩⟩⟩, - unfold angle, - rw angle_eq_zero_iff, - rw [←neg_vsub_eq_vsub_rev, neg_ne_zero] at hp1p2, - use [hp1p2, -r + 1, add_pos (neg_pos_of_neg hr) zero_lt_one], - rw [add_smul, ←neg_vsub_eq_vsub_rev p1 p2, smul_neg], - simp [←hpr] -end - -/-- If the angle ∠ABC at a point is π, the angle ∠BCA is 0. -/ -lemma angle_eq_zero_of_angle_eq_pi_right {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : - ∠ p2 p3 p1 = 0 := -begin - rw angle_comm at h, - exact angle_eq_zero_of_angle_eq_pi_left h -end - -/-- If ∠BCD = π, then ∠ABC = ∠ABD. -/ -lemma angle_eq_angle_of_angle_eq_pi (p1 : P) {p2 p3 p4 : P} (h : ∠ p2 p3 p4 = π) : - ∠ p1 p2 p3 = ∠ p1 p2 p4 := -begin - unfold angle at *, - rcases angle_eq_pi_iff.1 h with ⟨hp2p3, ⟨r, ⟨hr, hpr⟩⟩⟩, - rw [eq_comm], - convert angle_smul_right_of_pos (p1 -ᵥ p2) (p3 -ᵥ p2) (add_pos (neg_pos_of_neg hr) zero_lt_one), - rw [add_smul, ← neg_vsub_eq_vsub_rev p2 p3, smul_neg, neg_smul, ← hpr], - simp -end - -/-- If ∠BCD = π, then ∠ACB + ∠ACD = π. -/ -lemma angle_add_angle_eq_pi_of_angle_eq_pi (p1 : P) {p2 p3 p4 : P} (h : ∠ p2 p3 p4 = π) : - ∠ p1 p3 p2 + ∠ p1 p3 p4 = π := -begin - unfold angle at h, - rw [angle_comm p1 p3 p2, angle_comm p1 p3 p4], - unfold angle, - exact angle_add_angle_eq_pi_of_angle_eq_pi _ h -end - -/-- Vertical Angles Theorem: angles opposite each other, formed by two intersecting straight -lines, are equal. -/ -lemma angle_eq_angle_of_angle_eq_pi_of_angle_eq_pi {p1 p2 p3 p4 p5 : P} - (hapc : ∠ p1 p5 p3 = π) (hbpd : ∠ p2 p5 p4 = π) : ∠ p1 p5 p2 = ∠ p3 p5 p4 := -by linarith [angle_add_angle_eq_pi_of_angle_eq_pi p1 hbpd, angle_comm p4 p5 p1, - angle_add_angle_eq_pi_of_angle_eq_pi p4 hapc, angle_comm p4 p5 p3] - -/-- If ∠ABC = π then dist A B ≠ 0. -/ -lemma left_dist_ne_zero_of_angle_eq_pi {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : dist p1 p2 ≠ 0 := -begin - by_contra heq, - rw [dist_eq_zero] at heq, - rw [heq, angle_eq_left] at h, - exact real.pi_ne_zero (by linarith), -end - -/-- If ∠ABC = π then dist C B ≠ 0. -/ -lemma right_dist_ne_zero_of_angle_eq_pi {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : dist p3 p2 ≠ 0 := -left_dist_ne_zero_of_angle_eq_pi $ (angle_comm _ _ _).trans h - -/-- If ∠ABC = π, then (dist A C) = (dist A B) + (dist B C). -/ -lemma dist_eq_add_dist_of_angle_eq_pi {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = π) : - dist p1 p3 = dist p1 p2 + dist p3 p2 := -begin - rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, ← vsub_sub_vsub_cancel_right], - exact norm_sub_eq_add_norm_of_angle_eq_pi h, -end - -/-- If A ≠ B and C ≠ B then ∠ABC = π if and only if (dist A C) = (dist A B) + (dist B C). -/ -lemma dist_eq_add_dist_iff_angle_eq_pi {p1 p2 p3 : P} (hp1p2 : p1 ≠ p2) (hp3p2 : p3 ≠ p2) : - dist p1 p3 = dist p1 p2 + dist p3 p2 ↔ ∠ p1 p2 p3 = π := -begin - rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, ← vsub_sub_vsub_cancel_right], - exact norm_sub_eq_add_norm_iff_angle_eq_pi - ((λ he, hp1p2 (vsub_eq_zero_iff_eq.1 he))) (λ he, hp3p2 (vsub_eq_zero_iff_eq.1 he)), -end - -/-- If ∠ABC = 0, then (dist A C) = abs ((dist A B) - (dist B C)). -/ -lemma dist_eq_abs_sub_dist_of_angle_eq_zero {p1 p2 p3 : P} (h : ∠ p1 p2 p3 = 0) : - (dist p1 p3) = |(dist p1 p2) - (dist p3 p2)| := -begin - rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, ← vsub_sub_vsub_cancel_right], - exact norm_sub_eq_abs_sub_norm_of_angle_eq_zero h, -end - -/-- If A ≠ B and C ≠ B then ∠ABC = 0 if and only if (dist A C) = abs ((dist A B) - (dist B C)). -/ -lemma dist_eq_abs_sub_dist_iff_angle_eq_zero {p1 p2 p3 : P} (hp1p2 : p1 ≠ p2) (hp3p2 : p3 ≠ p2) : - (dist p1 p3) = |(dist p1 p2) - (dist p3 p2)| ↔ ∠ p1 p2 p3 = 0 := -begin - rw [dist_eq_norm_vsub V, dist_eq_norm_vsub V, dist_eq_norm_vsub V, ← vsub_sub_vsub_cancel_right], - exact norm_sub_eq_abs_sub_norm_iff_angle_eq_zero - ((λ he, hp1p2 (vsub_eq_zero_iff_eq.1 he))) (λ he, hp3p2 (vsub_eq_zero_iff_eq.1 he)), -end - /-- The midpoint of the segment AB is the same distance from A as it is from B. -/ lemma dist_left_midpoint_eq_dist_right_midpoint (p1 p2 : P) : dist p1 (midpoint ℝ p1 p2) = dist p2 (midpoint ℝ p1 p2) := by rw [dist_left_midpoint p1 p2, dist_right_midpoint p1 p2] -/-- If M is the midpoint of the segment AB, then ∠AMB = π. -/ -lemma angle_midpoint_eq_pi (p1 p2 : P) (hp1p2 : p1 ≠ p2) : ∠ p1 (midpoint ℝ p1 p2) p2 = π := -have p2 -ᵥ midpoint ℝ p1 p2 = -(p1 -ᵥ midpoint ℝ p1 p2), by { rw neg_vsub_eq_vsub_rev, simp }, -by simp [angle, this, hp1p2, -zero_lt_one] - -/-- If M is the midpoint of the segment AB and C is the same distance from A as it is from B -then ∠CMA = π / 2. -/ -lemma angle_left_midpoint_eq_pi_div_two_of_dist_eq {p1 p2 p3 : P} (h : dist p3 p1 = dist p3 p2) : - ∠ p3 (midpoint ℝ p1 p2) p1 = π / 2 := -begin - let m : P := midpoint ℝ p1 p2, - have h1 : p3 -ᵥ p1 = (p3 -ᵥ m) - (p1 -ᵥ m) := (vsub_sub_vsub_cancel_right p3 p1 m).symm, - have h2 : p3 -ᵥ p2 = (p3 -ᵥ m) + (p1 -ᵥ m), - { rw [left_vsub_midpoint, ← midpoint_vsub_right, vsub_add_vsub_cancel] }, - rw [dist_eq_norm_vsub V p3 p1, dist_eq_norm_vsub V p3 p2, h1, h2] at h, - exact (norm_add_eq_norm_sub_iff_angle_eq_pi_div_two (p3 -ᵥ m) (p1 -ᵥ m)).mp h.symm, -end - -/-- If M is the midpoint of the segment AB and C is the same distance from A as it is from B -then ∠CMB = π / 2. -/ -lemma angle_right_midpoint_eq_pi_div_two_of_dist_eq {p1 p2 p3 : P} (h : dist p3 p1 = dist p3 p2) : - ∠ p3 (midpoint ℝ p1 p2) p2 = π / 2 := -by rw [midpoint_comm p1 p2, angle_left_midpoint_eq_pi_div_two_of_dist_eq h.symm] - -/-- If the second of three points is strictly between the other two, the angle at that point -is π. -/ -lemma _root_.sbtw.angle₁₂₃_eq_pi {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₁ p₂ p₃ = π := -begin - rw [angle, angle_eq_pi_iff], - rcases h with ⟨⟨r, ⟨hr0, hr1⟩, hp₂⟩, hp₂p₁, hp₂p₃⟩, - refine ⟨vsub_ne_zero.2 hp₂p₁.symm, -(1 - r) / r, _⟩, - have hr0' : r ≠ 0, - { rintro rfl, - rw ←hp₂ at hp₂p₁, - simpa using hp₂p₁ }, - have hr1' : r ≠ 1, - { rintro rfl, - rw ←hp₂ at hp₂p₃, - simpa using hp₂p₃ }, - replace hr0 := hr0.lt_of_ne hr0'.symm, - replace hr1 := hr1.lt_of_ne hr1', - refine ⟨div_neg_of_neg_of_pos (left.neg_neg_iff.2 (sub_pos.2 hr1)) hr0, _⟩, - rw [←hp₂, affine_map.line_map_apply, vsub_vadd_eq_vsub_sub, vsub_vadd_eq_vsub_sub, vsub_self, - zero_sub, smul_neg, smul_smul, div_mul_cancel _ hr0', neg_smul, neg_neg, sub_eq_iff_eq_add, - ←add_smul, sub_add_cancel, one_smul] -end - -/-- If the second of three points is strictly between the other two, the angle at that point -(reversed) is π. -/ -lemma _root_.sbtw.angle₃₂₁_eq_pi {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₃ p₂ p₁ = π := -by rw [←h.angle₁₂₃_eq_pi, angle_comm] - -/-- The angle between three points is π if and only if the second point is strictly between the -other two. -/ -lemma angle_eq_pi_iff_sbtw {p₁ p₂ p₃ : P} : ∠ p₁ p₂ p₃ = π ↔ sbtw ℝ p₁ p₂ p₃ := -begin - refine ⟨_, λ h, h.angle₁₂₃_eq_pi⟩, - rw [angle, angle_eq_pi_iff], - rintro ⟨hp₁p₂, r, hr, hp₃p₂⟩, - refine ⟨⟨1 / (1 - r), - ⟨div_nonneg zero_le_one (sub_nonneg.2 (hr.le.trans zero_le_one)), - (div_le_one (sub_pos.2 (hr.trans zero_lt_one))).2 ((le_sub_self_iff 1).2 hr.le)⟩, _⟩, - (vsub_ne_zero.1 hp₁p₂).symm, _⟩, - { rw ←eq_vadd_iff_vsub_eq at hp₃p₂, - rw [affine_map.line_map_apply, hp₃p₂, vadd_vsub_assoc, ←neg_vsub_eq_vsub_rev p₂ p₁, - smul_neg, ←neg_smul, smul_add, smul_smul, ←add_smul, eq_comm, eq_vadd_iff_vsub_eq], - convert (one_smul ℝ (p₂ -ᵥ p₁)).symm, - field_simp [(sub_pos.2 (hr.trans zero_lt_one)).ne.symm], - abel }, - { rw [ne_comm, ←@vsub_ne_zero V, hp₃p₂, smul_ne_zero_iff], - exact ⟨hr.ne, hp₁p₂⟩ } -end - -/-- If the second of three points is weakly between the other two, and not equal to the first, -the angle at the first point is zero. -/ -lemma _root_.wbtw.angle₂₁₃_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₁ : p₂ ≠ p₁) : - ∠ p₂ p₁ p₃ = 0 := -begin - rw [angle, angle_eq_zero_iff], - rcases h with ⟨r, ⟨hr0, hr1⟩, rfl⟩, - have hr0' : r ≠ 0, { rintro rfl, simpa using hp₂p₁ }, - replace hr0 := hr0.lt_of_ne hr0'.symm, - refine ⟨vsub_ne_zero.2 hp₂p₁, r⁻¹, inv_pos.2 hr0, _⟩, - rw [affine_map.line_map_apply, vadd_vsub_assoc, vsub_self, add_zero, smul_smul, - inv_mul_cancel hr0', one_smul] -end - -/-- If the second of three points is strictly between the other two, the angle at the first point -is zero. -/ -lemma _root_.sbtw.angle₂₁₃_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₂ p₁ p₃ = 0 := -h.wbtw.angle₂₁₃_eq_zero_of_ne h.ne_left - -/-- If the second of three points is weakly between the other two, and not equal to the first, -the angle at the first point (reversed) is zero. -/ -lemma _root_.wbtw.angle₃₁₂_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₁ : p₂ ≠ p₁) : - ∠ p₃ p₁ p₂ = 0 := -by rw [←h.angle₂₁₃_eq_zero_of_ne hp₂p₁, angle_comm] - -/-- If the second of three points is strictly between the other two, the angle at the first point -(reversed) is zero. -/ -lemma _root_.sbtw.angle₃₁₂_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₃ p₁ p₂ = 0 := -h.wbtw.angle₃₁₂_eq_zero_of_ne h.ne_left - -/-- If the second of three points is weakly between the other two, and not equal to the third, -the angle at the third point is zero. -/ -lemma _root_.wbtw.angle₂₃₁_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₃ : p₂ ≠ p₃) : - ∠ p₂ p₃ p₁ = 0 := -h.symm.angle₂₁₃_eq_zero_of_ne hp₂p₃ - -/-- If the second of three points is strictly between the other two, the angle at the third point -is zero. -/ -lemma _root_.sbtw.angle₂₃₁_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₂ p₃ p₁ = 0 := -h.wbtw.angle₂₃₁_eq_zero_of_ne h.ne_right - -/-- If the second of three points is weakly between the other two, and not equal to the third, -the angle at the third point (reversed) is zero. -/ -lemma _root_.wbtw.angle₁₃₂_eq_zero_of_ne {p₁ p₂ p₃ : P} (h : wbtw ℝ p₁ p₂ p₃) (hp₂p₃ : p₂ ≠ p₃) : - ∠ p₁ p₃ p₂ = 0 := -h.symm.angle₃₁₂_eq_zero_of_ne hp₂p₃ - -/-- If the second of three points is strictly between the other two, the angle at the third point -(reversed) is zero. -/ -lemma _root_.sbtw.angle₁₃₂_eq_zero {p₁ p₂ p₃ : P} (h : sbtw ℝ p₁ p₂ p₃) : ∠ p₁ p₃ p₂ = 0 := -h.wbtw.angle₁₃₂_eq_zero_of_ne h.ne_right - -/-- The angle between three points is zero if and only if one of the first and third points is -weakly between the other two, and not equal to the second. -/ -lemma angle_eq_zero_iff_ne_and_wbtw {p₁ p₂ p₃ : P} : - ∠ p₁ p₂ p₃ = 0 ↔ (p₁ ≠ p₂ ∧ wbtw ℝ p₂ p₁ p₃) ∨ (p₃ ≠ p₂ ∧ wbtw ℝ p₂ p₃ p₁) := -begin - split, - { rw [angle, angle_eq_zero_iff], - rintro ⟨hp₁p₂, r, hr0, hp₃p₂⟩, - rcases le_or_lt 1 r with hr1 | hr1, - { refine or.inl ⟨vsub_ne_zero.1 hp₁p₂, r⁻¹, ⟨(inv_pos.2 hr0).le, inv_le_one hr1⟩, _⟩, - rw [affine_map.line_map_apply, hp₃p₂, smul_smul, inv_mul_cancel hr0.ne.symm, one_smul, - vsub_vadd] }, - { refine or.inr ⟨_, r, ⟨hr0.le, hr1.le⟩, _⟩, - { rw [←@vsub_ne_zero V, hp₃p₂, smul_ne_zero_iff], - exact ⟨hr0.ne.symm, hp₁p₂⟩ }, - { rw [affine_map.line_map_apply, ←hp₃p₂, vsub_vadd] } } }, - { rintro (⟨hp₁p₂, h⟩ | ⟨hp₃p₂, h⟩), - { exact h.angle₂₁₃_eq_zero_of_ne hp₁p₂ }, - { exact h.angle₃₁₂_eq_zero_of_ne hp₃p₂ } } -end - -/-- The angle between three points is zero if and only if one of the first and third points is -strictly between the other two, or those two points are equal but not equal to the second. -/ -lemma angle_eq_zero_iff_eq_and_ne_or_sbtw {p₁ p₂ p₃ : P} : - ∠ p₁ p₂ p₃ = 0 ↔ (p₁ = p₃ ∧ p₁ ≠ p₂) ∨ sbtw ℝ p₂ p₁ p₃ ∨ sbtw ℝ p₂ p₃ p₁ := -begin - rw angle_eq_zero_iff_ne_and_wbtw, - by_cases hp₁p₂ : p₁ = p₂, { simp [hp₁p₂] }, - by_cases hp₁p₃ : p₁ = p₃, { simp [hp₁p₃] }, - by_cases hp₃p₂ : p₃ = p₂, { simp [hp₃p₂] }, - simp [hp₁p₂, hp₁p₃, ne.symm hp₁p₃, sbtw, hp₃p₂] -end - /-- The inner product of two vectors given with `weighted_vsub`, in terms of the pairwise distances. -/ lemma inner_weighted_vsub {ι₁ : Type*} {s₁ : finset ι₁} {w₁ : ι₁ → ℝ} (p₁ : ι₁ → P) diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index 9d1f9f40241c2..14483a778efe6 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -30,13 +30,10 @@ the circumcenter. noncomputable theory open_locale big_operators open_locale classical -open_locale real open_locale real_inner_product_space namespace euclidean_geometry -open inner_product_geometry - variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V diff --git a/src/geometry/euclidean/default.lean b/src/geometry/euclidean/default.lean index a79e5c08e564b..9002867580ffe 100644 --- a/src/geometry/euclidean/default.lean +++ b/src/geometry/euclidean/default.lean @@ -1,3 +1,4 @@ +import geometry.euclidean.angle.unoriented.right_angle import geometry.euclidean.basic import geometry.euclidean.circumcenter import geometry.euclidean.monge_point diff --git a/src/geometry/euclidean/monge_point.lean b/src/geometry/euclidean/monge_point.lean index 2e03285bc3c35..6b919c11600ed 100644 --- a/src/geometry/euclidean/monge_point.lean +++ b/src/geometry/euclidean/monge_point.lean @@ -49,7 +49,6 @@ generalization, the Monge point of a simplex. noncomputable theory open_locale big_operators open_locale classical -open_locale real open_locale real_inner_product_space namespace affine diff --git a/src/geometry/euclidean/oriented_angle.lean b/src/geometry/euclidean/oriented_angle.lean index 2616f17c64193..7a4154c59b70c 100644 --- a/src/geometry/euclidean/oriented_angle.lean +++ b/src/geometry/euclidean/oriented_angle.lean @@ -7,7 +7,7 @@ import analysis.complex.arg import analysis.inner_product_space.orientation import analysis.inner_product_space.pi_L2 import analysis.special_functions.complex.circle -import geometry.euclidean.basic +import geometry.euclidean.angle.unoriented.affine /-! # Oriented angles. diff --git a/src/geometry/euclidean/sphere.lean b/src/geometry/euclidean/sphere.lean index b2495efd70e79..516e092f4e428 100644 --- a/src/geometry/euclidean/sphere.lean +++ b/src/geometry/euclidean/sphere.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Manuel Candales. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Manuel Candales, Benjamin Davidson -/ +import geometry.euclidean.basic import geometry.euclidean.triangle /-! diff --git a/src/geometry/euclidean/triangle.lean b/src/geometry/euclidean/triangle.lean index c27a9f89ddaa9..2b229b7e30f40 100644 --- a/src/geometry/euclidean/triangle.lean +++ b/src/geometry/euclidean/triangle.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Manuel Candales -/ -import geometry.euclidean.basic +import geometry.euclidean.angle.unoriented.affine import tactic.interval_cases /-! @@ -26,7 +26,6 @@ unnecessarily. ## References -* https://en.wikipedia.org/wiki/Pythagorean_theorem * https://en.wikipedia.org/wiki/Law_of_cosines * https://en.wikipedia.org/wiki/Pons_asinorum * https://en.wikipedia.org/wiki/Sum_of_angles_of_a_triangle @@ -52,32 +51,6 @@ deduce corresponding results for Euclidean affine spaces. variables {V : Type*} [inner_product_space ℝ V] -/-- Pythagorean theorem, if-and-only-if vector angle form. -/ -lemma norm_add_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two (x y : V) : - ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ ↔ angle x y = π / 2 := -begin - rw norm_add_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero, - exact inner_eq_zero_iff_angle_eq_pi_div_two x y -end - -/-- Pythagorean theorem, vector angle form. -/ -lemma norm_add_sq_eq_norm_sq_add_norm_sq' (x y : V) (h : angle x y = π / 2) : - ∥x + y∥ * ∥x + y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ := -(norm_add_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two x y).2 h - -/-- Pythagorean theorem, subtracting vectors, if-and-only-if vector angle form. -/ -lemma norm_sub_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two (x y : V) : - ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ ↔ angle x y = π / 2 := -begin - rw norm_sub_sq_eq_norm_sq_add_norm_sq_iff_real_inner_eq_zero, - exact inner_eq_zero_iff_angle_eq_pi_div_two x y -end - -/-- Pythagorean theorem, subtracting vectors, vector angle form. -/ -lemma norm_sub_sq_eq_norm_sq_add_norm_sq' (x y : V) (h : angle x y = π / 2) : - ∥x - y∥ * ∥x - y∥ = ∥x∥ * ∥x∥ + ∥y∥ * ∥y∥ := -(norm_sub_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two x y).2 h - /-- Law of cosines (cosine rule), vector angle form. -/ lemma norm_sub_sq_eq_norm_sq_add_norm_sq_sub_two_mul_norm_mul_norm_mul_cos_angle (x y : V) : @@ -292,15 +265,6 @@ variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] include V -/-- **Pythagorean theorem**, if-and-only-if angle-at-point form. -/ -lemma dist_sq_eq_dist_sq_add_dist_sq_iff_angle_eq_pi_div_two (p1 p2 p3 : P) : - dist p1 p3 * dist p1 p3 = dist p1 p2 * dist p1 p2 + dist p3 p2 * dist p3 p2 ↔ - ∠ p1 p2 p3 = π / 2 := -by erw [dist_comm p3 p2, dist_eq_norm_vsub V p1 p3, dist_eq_norm_vsub V p1 p2, - dist_eq_norm_vsub V p2 p3, - ←norm_sub_sq_eq_norm_sq_add_norm_sq_iff_angle_eq_pi_div_two, - vsub_sub_vsub_cancel_right p1, ←neg_vsub_eq_vsub_rev p2 p3, norm_neg] - /-- **Law of cosines** (cosine rule), angle-at-point form. -/ lemma dist_sq_eq_dist_sq_add_dist_sq_sub_two_mul_dist_mul_dist_mul_cos_angle (p1 p2 p3 : P) : From 46fc73a356be8a6295ee49b38937965b675a36b7 Mon Sep 17 00:00:00 2001 From: Alexander Bentkamp Date: Thu, 13 Oct 2022 08:46:32 +0000 Subject: [PATCH 738/828] feat(linear_algebra/matrix): inverse of block triangular is block triangular (#16150) The inverse of a block triangular matrix is block triangular. I took the opportunity to refactor the proof about determinants of block triangular matrices as well. Co-authored-by: Eric Co-authored-by: Alexander Bentkamp --- src/data/fintype/basic.lean | 31 ++++ src/data/matrix/block.lean | 65 ++++++++- src/linear_algebra/matrix/block.lean | 209 ++++++++++++++++++++------- 3 files changed, 249 insertions(+), 56 deletions(-) diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index ede6df4440323..78315c8cb6568 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -1306,6 +1306,37 @@ instance Prop.fintype : fintype Prop := instance subtype.fintype (p : α → Prop) [decidable_pred p] [fintype α] : fintype {x // p x} := fintype.subtype (univ.filter p) (by simp) +lemma image_subtype_ne_univ_eq_image_erase [fintype α] (k : β) (b : α → β) : + image (λ i : {a // b a ≠ k}, b ↑i) univ = (image b univ).erase k := +begin + apply subset_antisymm, + { rw image_subset_iff, + intros i _, + apply mem_erase_of_ne_of_mem i.2 (mem_image_of_mem _ (mem_univ _)) }, + { intros i hi, + rw mem_image, + rcases mem_image.1 (erase_subset _ _ hi) with ⟨a, _, ha⟩, + subst ha, + exact ⟨⟨a, ne_of_mem_erase hi⟩, mem_univ _, rfl⟩ } +end + +lemma image_subtype_univ_ssubset_image_univ [fintype α] (k : β) (b : α → β) + (hk : k ∈ image b univ) (p : β → Prop) [decidable_pred p] (hp : ¬ p k) : + image (λ i : {a // p (b a)}, b ↑i) univ ⊂ image b univ := +begin + split, + { intros x hx, + rcases mem_image.1 hx with ⟨y, _, hy⟩, + exact hy ▸ mem_image_of_mem b (mem_univ y) }, + { intros h, + rw mem_image at hk, + rcases hk with ⟨k', _, hk'⟩, subst hk', + have := h (mem_image_of_mem b (mem_univ k')), + rw mem_image at this, + rcases this with ⟨j, hj, hj'⟩, + exact hp (hj' ▸ j.2) } +end + @[simp] lemma set.to_finset_eq_empty_iff {s : set α} [fintype s] : s.to_finset = ∅ ↔ s = ∅ := by simp only [ext_iff, set.ext_iff, set.mem_to_finset, not_mem_empty, set.mem_empty_iff_false] diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index 2cf32751d5ad0..d659e5d98878f 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -220,16 +220,50 @@ by { ext i, cases i; simp [vec_mul, dot_product] } variables [decidable_eq l] [decidable_eq m] -@[simp] lemma from_blocks_diagonal [has_zero α] (d₁ : l → α) (d₂ : m → α) : +section has_zero +variables [has_zero α] + +lemma to_block_diagonal_self (d : m → α) (p : m → Prop) : + matrix.to_block (diagonal d) p p = diagonal (λ i : subtype p, d ↑i) := +begin + ext i j, + by_cases i = j, + { simp [h] }, + { simp [has_one.one, h, λ h', h $ subtype.ext h'], } +end + +lemma to_block_diagonal_disjoint (d : m → α) {p q : m → Prop} (hpq : disjoint p q) : + matrix.to_block (diagonal d) p q = 0 := +begin + ext ⟨i, hi⟩ ⟨j, hj⟩, + have : i ≠ j, from λ heq, hpq i ⟨hi, heq.symm ▸ hj⟩, + simp [diagonal_apply_ne d this] +end + +@[simp] lemma from_blocks_diagonal (d₁ : l → α) (d₂ : m → α) : from_blocks (diagonal d₁) 0 0 (diagonal d₂) = diagonal (sum.elim d₁ d₂) := begin ext i j, rcases i; rcases j; simp [diagonal], end -@[simp] lemma from_blocks_one [has_zero α] [has_one α] : +end has_zero + +section has_zero_has_one +variables [has_zero α] [has_one α] + +@[simp] lemma from_blocks_one : from_blocks (1 : matrix l l α) 0 0 (1 : matrix m m α) = 1 := by { ext i j, rcases i; rcases j; simp [one_apply] } +@[simp] lemma to_block_one_self (p : m → Prop) : matrix.to_block (1 : matrix m m α) p p = 1 := +to_block_diagonal_self _ p + +lemma to_block_one_disjoint {p q : m → Prop} (hpq : disjoint p q) : + matrix.to_block (1 : matrix m m α) p q = 0 := +to_block_diagonal_disjoint _ hpq + +end has_zero_has_one + end block_matrices section block_diagonal @@ -662,4 +696,31 @@ rfl end block_diag' +section +variables [comm_ring R] + +lemma to_block_mul_eq_mul {m n k : Type*} [fintype n] (p : m → Prop) (q : k → Prop) + (A : matrix m n R) (B : matrix n k R) : + (A ⬝ B).to_block p q = A.to_block p ⊤ ⬝ B.to_block ⊤ q := +begin + ext i k, + simp only [to_block_apply, mul_apply], + rw finset.sum_subtype, + simp [has_top.top, complete_lattice.top, bounded_order.top], +end + +lemma to_block_mul_eq_add + {m n k : Type*} [fintype n] (p : m → Prop) (q : n → Prop) [decidable_pred q] (r : k → Prop) + (A : matrix m n R) (B : matrix n k R) : + (A ⬝ B).to_block p r = + A.to_block p q ⬝ B.to_block q r + A.to_block p (λ i, ¬ q i) ⬝ B.to_block (λ i, ¬ q i) r := +begin + classical, + ext i k, + simp only [to_block_apply, mul_apply, pi.add_apply], + convert (fintype.sum_subtype_add_sum_subtype q (λ x, A ↑i x * B x ↑k)).symm +end + +end + end matrix diff --git a/src/linear_algebra/matrix/block.lean b/src/linear_algebra/matrix/block.lean index de0d47e2d99bd..c8bba64c869ee 100644 --- a/src/linear_algebra/matrix/block.lean +++ b/src/linear_algebra/matrix/block.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Patrick Massot, Casper Putz, Anne Baanen -/ import linear_algebra.matrix.determinant +import linear_algebra.matrix.nonsingular_inverse import tactic.fin_cases /-! @@ -35,8 +36,8 @@ open_locale big_operators matrix universes v -variables {α m n : Type*} -variables {R : Type v} [comm_ring R] {M : matrix m m R} {b : m → α} +variables {α β m n o : Type*} {m' n' : α → Type*} +variables {R : Type v} [comm_ring R] {M N : matrix m m R} {b : m → α} namespace matrix @@ -67,28 +68,66 @@ protected lemma block_triangular.transpose : @[simp] protected lemma block_triangular_transpose_iff {b : m → αᵒᵈ} : Mᵀ.block_triangular b ↔ M.block_triangular (of_dual ∘ b) := forall_swap +@[simp] lemma block_triangular_zero : block_triangular (0 : matrix m m R) b := λ i j h, rfl + +protected lemma block_triangular.neg (hM : block_triangular M b) : block_triangular (-M) b := +λ i j h, neg_eq_zero.2 $ hM h + +lemma block_triangular.add (hM : block_triangular M b) (hN : block_triangular N b) : + block_triangular (M + N) b := +λ i j h, by simp_rw [pi.add_apply, hM h, hN h, zero_add] + +lemma block_triangular.sub (hM : block_triangular M b) (hN : block_triangular N b) : + block_triangular (M - N) b := +λ i j h, by simp_rw [pi.sub_apply, hM h, hN h, sub_zero] + end has_lt +section preorder +variables [preorder α] + +lemma block_triangular_diagonal [decidable_eq m] (d : m → R) : + block_triangular (diagonal d) b := +λ i j h, diagonal_apply_ne' d (λ h', ne_of_lt h (congr_arg _ h')) + +lemma block_triangular_block_diagonal' [decidable_eq α] (d : Π (i : α), matrix (m' i) (m' i) R) : + block_triangular (block_diagonal' d) sigma.fst := +begin + rintros ⟨i, i'⟩ ⟨j, j'⟩ h, + apply block_diagonal'_apply_ne d i' j' (λ h', ne_of_lt h h'.symm), +end + +lemma block_triangular_block_diagonal [decidable_eq α] (d : α → matrix m m R) : + block_triangular (block_diagonal d) prod.snd := +begin + rintros ⟨i, i'⟩ ⟨j, j'⟩ h, + rw [block_diagonal'_eq_block_diagonal, block_triangular_block_diagonal'], + exact h +end + +end preorder + +section linear_order +variables [linear_order α] + +lemma block_triangular.mul [fintype m] + {M N : matrix m m R} (hM : block_triangular M b) (hN : block_triangular N b): + block_triangular (M * N) b := +begin + intros i j hij, + apply finset.sum_eq_zero, + intros k hk, + by_cases hki : b k < b i, + { simp_rw [hM hki, zero_mul] }, + { simp_rw [hN (lt_of_lt_of_le hij (le_of_not_lt hki)), mul_zero] }, +end + +end linear_order + lemma upper_two_block_triangular [preorder α] (A : matrix m m R) (B : matrix m n R) (D : matrix n n R) {a b : α} (hab : a < b) : block_triangular (from_blocks A B 0 D) (sum.elim (λ i, a) (λ j, b)) := -begin - intros k1 k2 hk12, - have hor : ∀ (k : m ⊕ n), sum.elim (λ i, a) (λ j, b) k = a ∨ sum.elim (λ i, a) (λ j, b) k = b, - { simp }, - have hne : a ≠ b, from λ h, lt_irrefl _ (lt_of_lt_of_eq hab h.symm), - have ha : ∀ (k : m ⊕ n), sum.elim (λ i, a) (λ j, b) k = a → ∃ i, k = sum.inl i, - { simp [hne.symm] }, - have hb : ∀ (k : m ⊕ n), sum.elim (λ i, a) (λ j, b) k = b → ∃ j, k = sum.inr j, - { simp [hne] }, - cases (hor k1) with hk1 hk1; cases (hor k2) with hk2 hk2; rw [hk1, hk2] at hk12, - { exact false.elim (lt_irrefl a hk12), }, - { exact false.elim (lt_irrefl _ (lt_trans hab hk12)) }, - { obtain ⟨i, hi⟩ := hb k1 hk1, - obtain ⟨j, hj⟩ := ha k2 hk2, - rw [hi, hj], simp }, - { exact absurd hk12 (irrefl b) } -end +by rintro (c | c) (d | d) hcd; simpa [hab.not_lt] using hcd <|> simp /-! ### Determinant -/ @@ -148,45 +187,31 @@ begin unfreezingI { induction hs : univ.image b using finset.strong_induction with s ih generalizing m }, subst hs, - by_cases h : univ.image b = ∅, - { haveI := univ_eq_empty_iff.1 (image_eq_empty.1 h), - simp [h] }, - { let k := (univ.image b).max' (nonempty_of_ne_empty h), - rw two_block_triangular_det' M (λ i, b i = k), - { have : univ.image b = insert k ((univ.image b).erase k), - { rw insert_erase, apply max'_mem }, - rw [this, prod_insert (not_mem_erase _ _)], - refine congr_arg _ _, - let b' := λ i : {a // b a ≠ k}, b ↑i, - have h' : block_triangular (M.to_square_block_prop (λ (i : m), b i ≠ k)) b', - { intros i j, apply hM }, - have hb' : image b' univ = (image b univ).erase k, - { apply subset_antisymm, - { rw image_subset_iff, - intros i _, - apply mem_erase_of_ne_of_mem i.2 (mem_image_of_mem _ (mem_univ _)) }, - { intros i hi, - rw mem_image, - rcases mem_image.1 (erase_subset _ _ hi) with ⟨a, _, ha⟩, - subst ha, - exact ⟨⟨a, ne_of_mem_erase hi⟩, mem_univ _, rfl⟩ } }, - rw ih ((univ.image b).erase k) (erase_ssubset (max'_mem _ _)) h' hb', - apply finset.prod_congr rfl, - intros l hl, - let he : {a // b' a = l} ≃ {a // b a = l}, - { have hc : ∀ (i : m), (λ a, b a = l) i → (λ a, b a ≠ k) i, - { intros i hbi, rw hbi, exact ne_of_mem_erase hl }, - exact equiv.subtype_subtype_equiv_subtype hc }, - simp only [to_square_block_def], - rw ← matrix.det_reindex_self he.symm (λ (i j : {a // b a = l}), M ↑i ↑j), - refine congr_arg _ _, - ext, - simp [to_square_block_def, to_square_block_prop_def] }, + casesI is_empty_or_nonempty m, + { simp }, + let k := (univ.image b).max' (univ_nonempty.image _), + rw two_block_triangular_det' M (λ i, b i = k), + { have : univ.image b = insert k ((univ.image b).erase k), + { rw insert_erase, apply max'_mem }, + rw [this, prod_insert (not_mem_erase _ _)], + refine congr_arg _ _, + let b' := λ i : {a // b a ≠ k}, b ↑i, + have h' : block_triangular (M.to_square_block_prop (λ i, b i ≠ k)) b' := hM.submatrix, + have hb' : image b' univ = (image b univ).erase k, + { convert image_subtype_ne_univ_eq_image_erase k b }, + rw ih _ (erase_ssubset $ max'_mem _ _) h' hb', + refine finset.prod_congr rfl (λ l hl, _), + let he : {a // b' a = l} ≃ {a // b a = l}, + { have hc : ∀ i, b i = l → b i ≠ k := λ i hi, ne_of_eq_of_ne hi (ne_of_mem_erase hl), + exact equiv.subtype_subtype_equiv_subtype hc }, + simp only [to_square_block_def], + rw ← matrix.det_reindex_self he.symm (λ (i j : {a // b a = l}), M ↑i ↑j), + refl }, { intros i hi j hj, apply hM, rw hi, apply lt_of_le_of_ne _ hj, - exact finset.le_max' (univ.image b) _ (mem_image_of_mem _ (mem_univ _)) } } + exact finset.le_max' (univ.image b) _ (mem_image_of_mem _ (mem_univ _)) } end lemma block_triangular.det_fintype [decidable_eq α] [fintype α] [linear_order α] @@ -209,4 +234,80 @@ lemma det_of_lower_triangular [linear_order m] (M : matrix m m R) (h : M.block_t M.det = ∏ i : m, M i i := by { rw ←det_transpose, exact det_of_upper_triangular h.transpose } +/-! ### Invertible -/ + +lemma block_triangular.to_block_inverse_mul_to_block_eq_one [linear_order α] [invertible M] + (hM : block_triangular M b) (k : α) : + M⁻¹.to_block (λ i, b i < k) (λ i, b i < k) ⬝ M.to_block (λ i, b i < k) (λ i, b i < k) = 1 := +begin + let p := (λ i, b i < k), + have h_sum : M⁻¹.to_block p p ⬝ M.to_block p p + + M⁻¹.to_block p (λ i, ¬ p i) ⬝ M.to_block (λ i, ¬ p i) p = 1, + by rw [←to_block_mul_eq_add, inv_mul_of_invertible M, to_block_one_self], + have h_zero : M.to_block (λ i, ¬ p i) p = 0, + { ext i j, + simpa using hM (lt_of_lt_of_le j.2 (le_of_not_lt i.2)) }, + simpa [h_zero] using h_sum +end + +/-- The inverse of an upper-left subblock of a block-triangular matrix `M` is the upper-left +subblock of `M⁻¹`. -/ +lemma block_triangular.inv_to_block [linear_order α] [invertible M] + (hM : block_triangular M b) (k : α) : + (M.to_block (λ i, b i < k) (λ i, b i < k))⁻¹ = M⁻¹.to_block (λ i, b i < k) (λ i, b i < k) := +inv_eq_left_inv $ hM.to_block_inverse_mul_to_block_eq_one k + +/-- An upper-left subblock of an invertible block-triangular matrix is invertible. -/ +def block_triangular.invertible_to_block [linear_order α] [invertible M] + (hM : block_triangular M b) (k : α) : + invertible (M.to_block (λ i, b i < k) (λ i, b i < k)) := +invertible_of_left_inverse _ ((⅟M).to_block (λ i, b i < k) (λ i, b i < k)) $ + by simpa only [inv_of_eq_nonsing_inv] using hM.to_block_inverse_mul_to_block_eq_one k + +/-- A lower-left subblock of the inverse of a block-triangular matrix is zero. This is a first step +towards `block_triangular.inv_to_block` below. -/ +lemma to_block_inverse_eq_zero [linear_order α] [invertible M] (hM : block_triangular M b) (k : α) : + M⁻¹.to_block (λ i, k ≤ b i) (λ i, b i < k) = 0 := +begin + let p := λ i, b i < k, + let q := λ i, ¬ b i < k, + have h_sum : M⁻¹.to_block q p ⬝ M.to_block p p + M⁻¹.to_block q q ⬝ M.to_block q p = 0, + { rw [←to_block_mul_eq_add, inv_mul_of_invertible M, to_block_one_disjoint], + exact λ i h, h.1 h.2 }, + have h_zero : M.to_block q p = 0, + { ext i j, + simpa using hM (lt_of_lt_of_le j.2 $ le_of_not_lt i.2) }, + have h_mul_eq_zero : M⁻¹.to_block q p ⬝ M.to_block p p = 0 := by simpa [h_zero] using h_sum, + haveI : invertible (M.to_block p p) := hM.invertible_to_block k, + have : (λ i, k ≤ b i) = q := by { ext, exact not_lt.symm }, + rw [this, ← matrix.zero_mul (M.to_block p p)⁻¹, ← h_mul_eq_zero, + mul_inv_cancel_right_of_invertible], +end + +/-- The inverse of a block-triangular matrix is block-triangular. -/ +lemma block_triangular_inv_of_block_triangular [linear_order α] [invertible M] + (hM : block_triangular M b) : + block_triangular M⁻¹ b := +begin + unfreezingI { induction hs : univ.image b using finset.strong_induction + with s ih generalizing m }, + subst hs, + intros i j hij, + haveI : inhabited m := ⟨i⟩, + let k := (univ.image b).max' (univ_nonempty.image _), + let b' := λ i : {a // b a < k}, b ↑i, + let A := M.to_block (λ i, b i < k) (λ j, b j < k), + obtain hbi | hi : b i = k ∨ _ := (le_max' _ (b i) $ mem_image_of_mem _ $ mem_univ _).eq_or_lt, + { have : M⁻¹.to_block (λ i, k ≤ b i) (λ i, b i < k) ⟨i, hbi.ge⟩ ⟨j, hbi ▸ hij⟩ = 0, + { simp only [to_block_inverse_eq_zero hM k, pi.zero_apply] }, + simp [this.symm] }, + haveI : invertible A := hM.invertible_to_block _, + have hA : A.block_triangular b' := hM.submatrix, + have hb' : image b' univ ⊂ image b univ, + { convert image_subtype_univ_ssubset_image_univ k b _ (λ a, a < k) (lt_irrefl _), + convert max'_mem _ _ }, + have hij' : b' ⟨j, hij.trans hi⟩ < b' ⟨i, hi⟩, by simp_rw [b', subtype.coe_mk, hij], + simp [hM.inv_to_block k, (ih (image b' univ) hb' hA rfl hij').symm], +end + end matrix From 1675aa423ff09df03113d8a7e019f911bc0ec774 Mon Sep 17 00:00:00 2001 From: Alain Chavarri Date: Thu, 13 Oct 2022 08:46:33 +0000 Subject: [PATCH 739/828] feat(data/int/gcd): add lemmas on gcd with negatives (#16922) some simp lemmas for dealing with gcds of negated integers --- src/data/int/gcd.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/data/int/gcd.lean b/src/data/int/gcd.lean index ef153ac08f22d..eb38cc5061d55 100644 --- a/src/data/int/gcd.lean +++ b/src/data/int/gcd.lean @@ -199,6 +199,12 @@ theorem gcd_assoc (i j k : ℤ) : gcd (gcd i j) k = gcd i (gcd j k) := nat.gcd_a @[simp] theorem gcd_one_right (i : ℤ) : gcd i 1 = 1 := nat.gcd_one_right _ +@[simp] lemma gcd_neg_right {x y : ℤ} : gcd x (-y) = gcd x y := +by rw [int.gcd, int.gcd, nat_abs_neg] + +@[simp] lemma gcd_neg_left {x y : ℤ} : gcd (-x) y = gcd x y := +by rw [int.gcd, int.gcd, nat_abs_neg] + theorem gcd_mul_left (i j k : ℤ) : gcd (i * j) (i * k) = nat_abs i * gcd j k := by { rw [int.gcd, int.gcd, nat_abs_mul, nat_abs_mul], apply nat.gcd_mul_left } From 1a170053e192f9750c4eeb535b2809a267ae395a Mon Sep 17 00:00:00 2001 From: Wrenna Robson Date: Thu, 13 Oct 2022 08:46:34 +0000 Subject: [PATCH 740/828] feat(data/set/finite): Add `off_diag` finitary instance and lemmas (#16923) Adds a `fintype` instance for `off_diag`, finite lemmas for it, and adds a related missing prod lemma. Co-authored-by: Wrenna Robson <34025592+linesthatinterlace@users.noreply.github.com> --- src/data/fintype/basic.lean | 3 +++ src/data/set/finite.lean | 12 ++++++++++++ 2 files changed, 15 insertions(+) diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 78315c8cb6568..869603c72e00c 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -693,6 +693,9 @@ lemma to_finset_prod (s : set α) (t : set β) [fintype s] [fintype t] [fintype (s ×ˢ t).to_finset = s.to_finset ×ˢ t.to_finset := by { ext, simp } +lemma to_finset_off_diag {s : set α} [decidable_eq α] [fintype s] [fintype s.off_diag] : +s.off_diag.to_finset = s.to_finset.off_diag := finset.ext $ by simp + /- TODO Without the coercion arrow (`↥`) there is an elaboration bug; it essentially infers `fintype.{v} (set.univ.{u} : set α)` with `v` and `u` distinct. Reported in leanprover-community/lean#672 -/ diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index a96ebf110fa03..050cb239e9d28 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -312,6 +312,9 @@ instance fintype_prod (s : set α) (t : set β) [fintype s] [fintype t] : fintype (s ×ˢ t : set (α × β)) := fintype.of_finset (s.to_finset ×ˢ t.to_finset) $ by simp +instance fintype_off_diag [decidable_eq α] (s : set α) [fintype s] : fintype s.off_diag := +fintype.of_finset s.to_finset.off_diag $ by simp + /-- `image2 f s t` is `fintype` if `s` and `t` are. -/ instance fintype_image2 [decidable_eq γ] (f : α → β → γ) (s : set α) (t : set β) [hs : fintype s] [ht : fintype t] : fintype (image2 f s t : set γ) := @@ -580,6 +583,9 @@ lemma finite.prod {s : set α} {t : set β} (hs : s.finite) (ht : t.finite) : (s ×ˢ t : set (α × β)).finite := by { casesI hs, casesI ht, apply to_finite } +lemma finite.off_diag {s : set α} (hs : s.finite) : s.off_diag.finite := +by { classical, casesI hs, apply set.to_finite } + lemma finite.image2 (f : α → β → γ) {s : set α} {t : set β} (hs : s.finite) (ht : t.finite) : (image2 f s t).finite := by { casesI hs, casesI ht, apply to_finite } @@ -668,6 +674,12 @@ lemma finite.to_finset_insert' [decidable_eq α] {a : α} {s : set α} (hs : s.f (hs.insert a).to_finset = insert a hs.to_finset := finite.to_finset_insert _ +lemma finite.to_finset_prod {s : set α} {t : set β} (hs : s.finite) (ht : t.finite) : +hs.to_finset ×ˢ ht.to_finset = (hs.prod ht).to_finset := finset.ext $ by simp + +lemma finite.to_finset_off_diag {s : set α} [decidable_eq α] (hs : s.finite) : +hs.off_diag.to_finset = hs.to_finset.off_diag := finset.ext $ by simp + lemma finite.fin_embedding {s : set α} (h : s.finite) : ∃ (n : ℕ) (f : fin n ↪ α), range f = s := ⟨_, (fintype.equiv_fin (h.to_finset : set α)).symm.as_embedding, by simp⟩ From 080cbe2f5f002d1f44f4b3fc443d2ae19f7d7a3b Mon Sep 17 00:00:00 2001 From: Wrenna Robson Date: Thu, 13 Oct 2022 08:46:35 +0000 Subject: [PATCH 741/828] feat(data/set/fintype): Add `to_finset_nonempty`. (#16925) We have `finset.nonempty_coe` and `finite.nonempty_to_finset`, but we don't have this for fintypes. This PR adds this lemma. --- src/data/fintype/basic.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 869603c72e00c..f9532e261f42a 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -655,6 +655,9 @@ multiset.card_map subtype.val finset.univ.val @[simp] theorem coe_to_finset (s : set α) [fintype s] : (↑s.to_finset : set α) = s := set.ext $ λ _, mem_to_finset +@[simp] lemma to_finset_nonempty {s : set α} [fintype s] : s.to_finset.nonempty ↔ s.nonempty := +by rw [←finset.coe_nonempty, coe_to_finset] + @[simp] theorem to_finset_inj {s t : set α} [fintype s] [fintype t] : s.to_finset = t.to_finset ↔ s = t := ⟨λ h, by rw [←s.coe_to_finset, h, t.coe_to_finset], λ h, by simp [h]; congr⟩ From 8b80cafb12a244b7ce7178673051c2e626d2a8ac Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Thu, 13 Oct 2022 11:47:24 +0000 Subject: [PATCH 742/828] feat(topology/sheaves): Localization of a sheaf. (#15331) --- src/algebra/category/Ring/instances.lean | 12 +++ src/topology/sheaves/operations.lean | 121 +++++++++++++++++++++++ 2 files changed, 133 insertions(+) create mode 100644 src/topology/sheaves/operations.lean diff --git a/src/algebra/category/Ring/instances.lean b/src/algebra/category/Ring/instances.lean index 55c663a16477d..8e7e967a1fddc 100644 --- a/src/algebra/category/Ring/instances.lean +++ b/src/algebra/category/Ring/instances.lean @@ -19,3 +19,15 @@ is_iso.of_iso (is_localization.at_one R (localization.away (1 : R))).to_ring_equ instance localization_unit_is_iso' (R : CommRing) : @is_iso CommRing _ R _ (CommRing.of_hom $ algebra_map R (localization.away (1 : R))) := by { cases R, exact localization_unit_is_iso _ } + +lemma is_localization.epi {R : Type*} [comm_ring R] (M : submonoid R) (S : Type*) [comm_ring S] + [algebra R S] [is_localization M S] : epi (CommRing.of_hom $ algebra_map R S) := +⟨λ T f₁ f₂, @is_localization.ring_hom_ext R _ M S _ _ T _ _ _ _⟩ + +instance localization.epi {R : Type*} [comm_ring R] (M : submonoid R) : epi + (CommRing.of_hom $ algebra_map R $ localization M) := +is_localization.epi M _ + +instance localization.epi' {R : CommRing} (M : submonoid R) : @epi CommRing _ R _ + (CommRing.of_hom $ algebra_map R $ localization M : _) := +by { cases R, exact is_localization.epi M _ } diff --git a/src/topology/sheaves/operations.lean b/src/topology/sheaves/operations.lean new file mode 100644 index 0000000000000..ecf93ebc1111c --- /dev/null +++ b/src/topology/sheaves/operations.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2022 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ +import category_theory.sites.compatible_sheafification +import algebra.category.Ring.instances +import algebra.category.Ring.filtered_colimits +import algebra.category.Group.limits +import ring_theory.localization.basic +import topology.sheaves.sheafify +import topology.sheaves.sheaf_condition.sites + +/-! + +# Operations on sheaves + +## Main definition + +- `submonoid_presheaf` : A subpresheaf with a submonoid structure on each of the components. +- `localization_presheaf` : The localization of a presheaf of commrings at a `submonoid_presheaf`. +- `total_quotient_presheaf` : The presheaf of total quotient rings. + +-/ + +open_locale non_zero_divisors + +open topological_space opposite category_theory + +universes v u w + +namespace Top + +namespace presheaf + +variables {X : Top.{w}} {C : Type u} [category.{v} C] [concrete_category C] + +local attribute [instance] concrete_category.has_coe_to_sort + +/-- A subpresheaf with a submonoid structure on each of the components. -/ +structure submonoid_presheaf [∀ X : C, mul_one_class X] + [∀ (X Y : C), monoid_hom_class (X ⟶ Y) X Y] (F : X.presheaf C) := +(obj : ∀ U, submonoid (F.obj U)) +(map : ∀ {U V : (opens X)ᵒᵖ} (i : U ⟶ V), (obj U) ≤ (obj V).comap (F.map i)) + +variables {F : X.presheaf CommRing.{w}} (G : F.submonoid_presheaf) + +/-- The localization of a presheaf of `CommRing`s with respect to a `submonoid_presheaf`. -/ +protected noncomputable +def submonoid_presheaf.localization_presheaf : + X.presheaf CommRing := +{ obj := λ U, CommRing.of $ localization (G.obj U), + map := λ U V i, CommRing.of_hom $ is_localization.map _ (F.map i) (G.map i), + map_id' := λ U, begin + apply is_localization.ring_hom_ext (G.obj U), + any_goals { dsimp, apply_instance }, + refine (is_localization.map_comp _).trans _, + rw F.map_id, + refl, + end, + map_comp' := λ U V W i j, by { refine eq.trans _ (is_localization.map_comp_map _ _).symm, + ext, dsimp, congr, rw F.map_comp, refl } } + +/-- The map into the localization presheaf. -/ +def submonoid_presheaf.to_localization_presheaf : + F ⟶ G.localization_presheaf := +{ app := λ U, CommRing.of_hom $ algebra_map (F.obj U) (localization $ G.obj U), + naturality' := λ U V i, (is_localization.map_comp (G.map i)).symm } + +instance : epi G.to_localization_presheaf := +@@nat_trans.epi_of_epi_app _ _ G.to_localization_presheaf (λ U, localization.epi' (G.obj U)) + +variable (F) + +/-- Given a submonoid at each of the stalks, we may define a submonoid presheaf consisting of +sections whose restriction onto each stalk falls in the given submonoid. -/ +@[simps] noncomputable +def submonoid_presheaf_of_stalk (S : ∀ x : X, submonoid (F.stalk x)) : + F.submonoid_presheaf := +{ obj := λ U, ⨅ x : (unop U), submonoid.comap (F.germ x) (S x), + map := λ U V i, begin + intros s hs, + simp only [submonoid.mem_comap, submonoid.mem_infi] at ⊢ hs, + intro x, + change (F.map i.unop.op ≫ F.germ x) s ∈ _, + rw F.germ_res, + exact hs _, + end } + +noncomputable +instance : inhabited F.submonoid_presheaf := ⟨F.submonoid_presheaf_of_stalk (λ _, ⊥)⟩ + +/-- The localization of a presheaf of `CommRing`s at locally non-zero-divisor sections. -/ +noncomputable +def total_quotient_presheaf : X.presheaf CommRing.{w} := +(F.submonoid_presheaf_of_stalk (λ x, (F.stalk x)⁰)).localization_presheaf + +/-- The map into the presheaf of total quotient rings -/ +@[derive epi] noncomputable +def to_total_quotient_presheaf : F ⟶ F.total_quotient_presheaf := +submonoid_presheaf.to_localization_presheaf _ + +instance (F : X.sheaf CommRing.{w}) : mono F.presheaf.to_total_quotient_presheaf := +begin + apply_with nat_trans.mono_of_mono_app { instances := ff }, + intro U, + apply concrete_category.mono_of_injective, + apply is_localization.injective _, + swap 3, { exact localization.is_localization }, + intros s hs t e, + apply section_ext F (unop U), + intro x, + rw map_zero, + apply submonoid.mem_infi.mp hs x, + rw [← map_mul, e, map_zero] +end + + +end presheaf + +end Top From 856cfd48476696148040e15ccb6cc1557cdd070b Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Thu, 13 Oct 2022 11:47:25 +0000 Subject: [PATCH 743/828] feat(order/liminf_limsup): define `filter.blimsup`, `filter.bliminf` (#16819) Also characterise `cofinite.limsup` and `cofinite.liminf` for the lattice of sets. --- src/analysis/analytic/radius_liminf.lean | 2 +- src/analysis/normed_space/spectrum.lean | 5 +- .../constructions/borel_space.lean | 12 +- src/measure_theory/function/lp_space.lean | 14 +- src/measure_theory/integral/lebesgue.lean | 18 +- .../finite_measure_weak_convergence.lean | 4 +- src/measure_theory/measure/hausdorff.lean | 14 +- src/measure_theory/measure/measure_space.lean | 4 +- src/order/liminf_limsup.lean | 320 ++++++++++++++---- .../martingale/borel_cantelli.lean | 2 +- src/probability/martingale/convergence.lean | 2 +- src/topology/algebra/order/liminf_limsup.lean | 14 +- src/topology/instances/ennreal.lean | 12 +- .../metric_space/hausdorff_dimension.lean | 4 +- 14 files changed, 302 insertions(+), 125 deletions(-) diff --git a/src/analysis/analytic/radius_liminf.lean b/src/analysis/analytic/radius_liminf.lean index 429adbf102d1a..d3f1220e00408 100644 --- a/src/analysis/analytic/radius_liminf.lean +++ b/src/analysis/analytic/radius_liminf.lean @@ -27,7 +27,7 @@ variables (p : formal_multilinear_series 𝕜 E F) /-- The radius of a formal multilinear series is equal to $\liminf_{n\to\infty} \frac{1}{\sqrt[n]{∥p n∥}}$. The actual statement uses `ℝ≥0` and some coercions. -/ -lemma radius_eq_liminf : p.radius = liminf at_top (λ n, 1/((∥p n∥₊) ^ (1 / (n : ℝ)) : ℝ≥0)) := +lemma radius_eq_liminf : p.radius = liminf (λ n, 1/((∥p n∥₊) ^ (1 / (n : ℝ)) : ℝ≥0)) at_top := begin have : ∀ (r : ℝ≥0) {n : ℕ}, 0 < n → ((r : ℝ≥0∞) ≤ 1 / ↑(∥p n∥₊ ^ (1 / (n : ℝ))) ↔ ∥p n∥₊ * r ^ n ≤ 1), diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 367d2fcbb12c6..cd7c42335f53d 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -312,7 +312,7 @@ variables /-- The `limsup` relationship for the spectral radius used to prove `spectrum.gelfand_formula`. -/ lemma limsup_pow_nnnorm_pow_one_div_le_spectral_radius (a : A) : - limsup at_top (λ n : ℕ, ↑∥a ^ n∥₊ ^ (1 / n : ℝ)) ≤ spectral_radius ℂ a := + limsup (λ n : ℕ, ↑∥a ^ n∥₊ ^ (1 / n : ℝ)) at_top ≤ spectral_radius ℂ a := begin refine ennreal.inv_le_inv.mp (le_of_forall_pos_nnreal_lt (λ r r_pos r_lt, _)), simp_rw [inv_limsup, ←one_div], @@ -321,7 +321,8 @@ begin suffices h : (r : ℝ≥0∞) ≤ p.radius, { convert h, simp only [p.radius_eq_liminf, ←norm_to_nnreal, norm_mk_pi_field], - refine congr_arg _ (funext (λ n, congr_arg _ _)), + congr, + ext n, rw [norm_to_nnreal, ennreal.coe_rpow_def (∥a ^ n∥₊) (1 / n : ℝ), if_neg], exact λ ha, by linarith [ha.2, (one_div_nonneg.mpr n.cast_nonneg : 0 ≤ (1 / n : ℝ))], }, { have H₁ := (differentiable_on_inverse_one_sub_smul r_lt).has_fpower_series_on_ball r_pos, diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 58fb2f7ca3b3d..341d5e46b0317 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -1243,7 +1243,7 @@ end -/ lemma measurable_liminf' {ι ι'} {f : ι → δ → α} {u : filter ι} (hf : ∀ i, measurable (f i)) {p : ι' → Prop} {s : ι' → set ι} (hu : u.has_countable_basis p s) (hs : ∀ i, (s i).countable) : - measurable (λ x, liminf u (λ i, f i x)) := + measurable (λ x, liminf (λ i, f i x) u) := begin simp_rw [hu.to_has_basis.liminf_eq_supr_infi], refine measurable_bsupr _ hu.countable _, @@ -1254,7 +1254,7 @@ end -/ lemma measurable_limsup' {ι ι'} {f : ι → δ → α} {u : filter ι} (hf : ∀ i, measurable (f i)) {p : ι' → Prop} {s : ι' → set ι} (hu : u.has_countable_basis p s) (hs : ∀ i, (s i).countable) : - measurable (λ x, limsup u (λ i, f i x)) := + measurable (λ x, limsup (λ i, f i x) u) := begin simp_rw [hu.to_has_basis.limsup_eq_infi_supr], refine measurable_binfi _ hu.countable _, @@ -1265,14 +1265,14 @@ end -/ @[measurability] lemma measurable_liminf {f : ℕ → δ → α} (hf : ∀ i, measurable (f i)) : - measurable (λ x, liminf at_top (λ i, f i x)) := + measurable (λ x, liminf (λ i, f i x) at_top) := measurable_liminf' hf at_top_countable_basis (λ i, to_countable _) /-- `limsup` over `ℕ` is measurable. See `measurable_limsup'` for a version with a general filter. -/ @[measurability] lemma measurable_limsup {f : ℕ → δ → α} (hf : ∀ i, measurable (f i)) : - measurable (λ x, limsup at_top (λ i, f i x)) := + measurable (λ x, limsup (λ i, f i x) at_top) := measurable_limsup' hf at_top_countable_basis (λ i, to_countable _) end complete_linear_order @@ -1900,10 +1900,10 @@ lemma measurable_of_tendsto_ennreal' {ι} {f : ι → α → ℝ≥0∞} {g : α begin rcases u.exists_seq_tendsto with ⟨x, hx⟩, rw [tendsto_pi_nhds] at lim, - have : (λ y, liminf at_top (λ n, (f (x n) y : ℝ≥0∞))) = g := + have : (λ y, liminf (λ n, (f (x n) y : ℝ≥0∞)) at_top) = g := by { ext1 y, exact ((lim y).comp hx).liminf_eq, }, rw ← this, - show measurable (λ y, liminf at_top (λ n, (f (x n) y : ℝ≥0∞))), + show measurable (λ y, liminf (λ n, (f (x n) y : ℝ≥0∞)) at_top), exact measurable_liminf (λ n, hf (x n)), end diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 5d38b024c5ccf..bd581d859eecb 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -1433,7 +1433,7 @@ variables [measurable_space E] [opens_measurable_space E] {R : ℝ≥0} lemma ae_bdd_liminf_at_top_rpow_of_snorm_bdd {p : ℝ≥0∞} {f : ℕ → α → E} (hfmeas : ∀ n, measurable (f n)) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : - ∀ᵐ x ∂μ, liminf at_top (λ n, (∥f n x∥₊ ^ p.to_real : ℝ≥0∞)) < ∞ := + ∀ᵐ x ∂μ, liminf (λ n, (∥f n x∥₊ ^ p.to_real : ℝ≥0∞)) at_top < ∞ := begin by_cases hp0 : p.to_real = 0, { simp only [hp0, ennreal.rpow_zero], @@ -1456,7 +1456,7 @@ end lemma ae_bdd_liminf_at_top_of_snorm_bdd {p : ℝ≥0∞} (hp : p ≠ 0) {f : ℕ → α → E} (hfmeas : ∀ n, measurable (f n)) (hbdd : ∀ n, snorm (f n) p μ ≤ R) : - ∀ᵐ x ∂μ, liminf at_top (λ n, (∥f n x∥₊ : ℝ≥0∞)) < ∞ := + ∀ᵐ x ∂μ, liminf (λ n, (∥f n x∥₊ : ℝ≥0∞)) at_top < ∞ := begin by_cases hp' : p = ∞, { subst hp', @@ -1470,14 +1470,14 @@ begin (ennreal.add_lt_top.2 ⟨ennreal.coe_lt_top, ennreal.one_lt_top⟩) }, filter_upwards [ae_bdd_liminf_at_top_rpow_of_snorm_bdd hfmeas hbdd] with x hx, have hppos : 0 < p.to_real := ennreal.to_real_pos hp hp', - have : liminf at_top (λ n, (∥f n x∥₊ ^ p.to_real : ℝ≥0∞)) = - liminf at_top (λ n, (∥f n x∥₊ : ℝ≥0∞)) ^ p.to_real, - { change liminf at_top (λ n, ennreal.order_iso_rpow p.to_real hppos (∥f n x∥₊ : ℝ≥0∞)) = - ennreal.order_iso_rpow p.to_real hppos (liminf at_top (λ n, (∥f n x∥₊ : ℝ≥0∞))), + have : liminf (λ n, (∥f n x∥₊ ^ p.to_real : ℝ≥0∞)) at_top = + (liminf (λ n, (∥f n x∥₊ : ℝ≥0∞)) at_top)^ p.to_real, + { change liminf (λ n, ennreal.order_iso_rpow p.to_real hppos (∥f n x∥₊ : ℝ≥0∞)) at_top = + ennreal.order_iso_rpow p.to_real hppos (liminf (λ n, (∥f n x∥₊ : ℝ≥0∞)) at_top), refine (order_iso.liminf_apply (ennreal.order_iso_rpow p.to_real _) _ _ _ _).symm; is_bounded_default }, rw this at hx, - rw [← ennreal.rpow_one (liminf at_top (λ n, ∥f n x∥₊)), ← mul_inv_cancel hppos.ne.symm, + rw [← ennreal.rpow_one (liminf (λ n, ∥f n x∥₊) at_top), ← mul_inv_cancel hppos.ne.symm, ennreal.rpow_mul], exact ennreal.rpow_lt_top_of_nonneg (inv_nonneg.2 hppos.le) hx.ne, end diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index c60e524d6800d..f660c99b69315 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1866,9 +1866,9 @@ lintegral_infi_ae h_meas (λ n, ae_of_all _ $ h_anti n.le_succ) h_fin /-- Known as Fatou's lemma, version with `ae_measurable` functions -/ lemma lintegral_liminf_le' {f : ℕ → α → ℝ≥0∞} (h_meas : ∀n, ae_measurable (f n) μ) : - ∫⁻ a, liminf at_top (λ n, f n a) ∂μ ≤ liminf at_top (λ n, ∫⁻ a, f n a ∂μ) := + ∫⁻ a, liminf (λ n, f n a) at_top ∂μ ≤ liminf (λ n, ∫⁻ a, f n a ∂μ) at_top := calc - ∫⁻ a, liminf at_top (λ n, f n a) ∂μ = ∫⁻ a, ⨆n:ℕ, ⨅i≥n, f i a ∂μ : + ∫⁻ a, liminf (λ n, f n a) at_top ∂μ = ∫⁻ a, ⨆n:ℕ, ⨅i≥n, f i a ∂μ : by simp only [liminf_eq_supr_infi_of_nat] ... = ⨆n:ℕ, ∫⁻ a, ⨅i≥n, f i a ∂μ : lintegral_supr' @@ -1880,14 +1880,14 @@ calc /-- Known as Fatou's lemma -/ lemma lintegral_liminf_le {f : ℕ → α → ℝ≥0∞} (h_meas : ∀n, measurable (f n)) : - ∫⁻ a, liminf at_top (λ n, f n a) ∂μ ≤ liminf at_top (λ n, ∫⁻ a, f n a ∂μ) := + ∫⁻ a, liminf (λ n, f n a) at_top ∂μ ≤ liminf (λ n, ∫⁻ a, f n a ∂μ) at_top := lintegral_liminf_le' (λ n, (h_meas n).ae_measurable) lemma limsup_lintegral_le {f : ℕ → α → ℝ≥0∞} {g : α → ℝ≥0∞} (hf_meas : ∀ n, measurable (f n)) (h_bound : ∀n, f n ≤ᵐ[μ] g) (h_fin : ∫⁻ a, g a ∂μ ≠ ∞) : - limsup at_top (λn, ∫⁻ a, f n a ∂μ) ≤ ∫⁻ a, limsup at_top (λn, f n a) ∂μ := + limsup (λn, ∫⁻ a, f n a ∂μ) at_top ≤ ∫⁻ a, limsup (λn, f n a) at_top ∂μ := calc - limsup at_top (λn, ∫⁻ a, f n a ∂μ) = ⨅n:ℕ, ⨆i≥n, ∫⁻ a, f i a ∂μ : + limsup (λn, ∫⁻ a, f n a ∂μ) at_top = ⨅n:ℕ, ⨆i≥n, ∫⁻ a, f i a ∂μ : limsup_eq_infi_supr_of_nat ... ≤ ⨅n:ℕ, ∫⁻ a, ⨆i≥n, f i a ∂μ : infi_mono $ assume n, supr₂_lintegral_le _ @@ -1900,7 +1900,7 @@ calc refine (ae_all_iff.2 h_bound).mono (λ n hn, _), exact supr_le (λ i, supr_le $ λ hi, hn i) } end - ... = ∫⁻ a, limsup at_top (λn, f n a) ∂μ : + ... = ∫⁻ a, limsup (λn, f n a) at_top ∂μ : by simp only [limsup_eq_infi_supr_of_nat] /-- Dominated convergence theorem for nonnegative functions -/ @@ -1911,10 +1911,10 @@ lemma tendsto_lintegral_of_dominated_convergence (h_lim : ∀ᵐ a ∂μ, tendsto (λ n, F n a) at_top (𝓝 (f a))) : tendsto (λn, ∫⁻ a, F n a ∂μ) at_top (𝓝 (∫⁻ a, f a ∂μ)) := tendsto_of_le_liminf_of_limsup_le -(calc ∫⁻ a, f a ∂μ = ∫⁻ a, liminf at_top (λ (n : ℕ), F n a) ∂μ : +(calc ∫⁻ a, f a ∂μ = ∫⁻ a, liminf (λ (n : ℕ), F n a) at_top ∂μ : lintegral_congr_ae $ h_lim.mono $ assume a h, h.liminf_eq.symm - ... ≤ liminf at_top (λ n, ∫⁻ a, F n a ∂μ) : lintegral_liminf_le hF_meas) -(calc limsup at_top (λ (n : ℕ), ∫⁻ a, F n a ∂μ) ≤ ∫⁻ a, limsup at_top (λn, F n a) ∂μ : + ... ≤ liminf (λ n, ∫⁻ a, F n a ∂μ) at_top : lintegral_liminf_le hF_meas) +(calc limsup (λ (n : ℕ), ∫⁻ a, F n a ∂μ) at_top ≤ ∫⁻ a, limsup (λn, F n a) at_top ∂μ : limsup_lintegral_le hF_meas h_bound h_fin ... = ∫⁻ a, f a ∂μ : lintegral_congr_ae $ h_lim.mono $ λ a h, h.limsup_eq) diff --git a/src/measure_theory/measure/finite_measure_weak_convergence.lean b/src/measure_theory/measure/finite_measure_weak_convergence.lean index 7f16b36478e8a..849b1b8ff57bd 100644 --- a/src/measure_theory/measure/finite_measure_weak_convergence.lean +++ b/src/measure_theory/measure/finite_measure_weak_convergence.lean @@ -1063,7 +1063,7 @@ lemma le_measure_compl_liminf_of_limsup_measure_le begin by_cases L_bot : L = ⊥, { simp only [L_bot, le_top, - (show liminf ⊥ (λ i, μs i Eᶜ) = ⊤, by simp only [liminf, filter.map_bot, Liminf_bot])], }, + (show liminf (λ i, μs i Eᶜ) ⊥ = ⊤, by simp only [liminf, filter.map_bot, Liminf_bot])], }, haveI : L.ne_bot, from {ne' := L_bot}, have meas_Ec : μ Eᶜ = 1 - μ E, { simpa only [measure_univ] using measure_compl E_mble (measure_lt_top μ E).ne, }, @@ -1094,7 +1094,7 @@ lemma limsup_measure_compl_le_of_le_liminf_measure begin by_cases L_bot : L = ⊥, { simp only [L_bot, bot_le, - (show limsup ⊥ (λ i, μs i Eᶜ) = ⊥, by simp only [limsup, filter.map_bot, Limsup_bot])], }, + (show limsup (λ i, μs i Eᶜ) ⊥ = ⊥, by simp only [limsup, filter.map_bot, Limsup_bot])], }, haveI : L.ne_bot, from {ne' := L_bot}, have meas_Ec : μ Eᶜ = 1 - μ E, { simpa only [measure_univ] using measure_compl E_mble (measure_lt_top μ E).ne, }, diff --git a/src/measure_theory/measure/hausdorff.lean b/src/measure_theory/measure/hausdorff.lean index fb4aa3800f7ec..da0be5b93bb46 100644 --- a/src/measure_theory/measure/hausdorff.lean +++ b/src/measure_theory/measure/hausdorff.lean @@ -514,7 +514,7 @@ lemma mk_metric_le_liminf_tsum {β : Type*} {ι : β → Type*} [∀ n, countabl {l : filter β} (r : β → ℝ≥0∞) (hr : tendsto r l (𝓝 0)) (t : Π (n : β), ι n → set X) (ht : ∀ᶠ n in l, ∀ i, diam (t n i) ≤ r n) (hst : ∀ᶠ n in l, s ⊆ ⋃ i, t n i) (m : ℝ≥0∞ → ℝ≥0∞) : - mk_metric m s ≤ liminf l (λ n, ∑' i, m (diam (t n i))) := + mk_metric m s ≤ liminf (λ n, ∑' i, m (diam (t n i))) l := begin haveI : Π n, encodable (ι n) := λ n, encodable.of_countable _, simp only [mk_metric_apply], @@ -541,7 +541,7 @@ lemma mk_metric_le_liminf_sum {β : Type*} {ι : β → Type*} [hι : ∀ n, fin {l : filter β} (r : β → ℝ≥0∞) (hr : tendsto r l (𝓝 0)) (t : Π (n : β), ι n → set X) (ht : ∀ᶠ n in l, ∀ i, diam (t n i) ≤ r n) (hst : ∀ᶠ n in l, s ⊆ ⋃ i, t n i) (m : ℝ≥0∞ → ℝ≥0∞) : - mk_metric m s ≤ liminf l (λ n, ∑ i, m (diam (t n i))) := + mk_metric m s ≤ liminf (λ n, ∑ i, m (diam (t n i))) l := by simpa only [tsum_fintype] using mk_metric_le_liminf_tsum s r hr t ht hst m /-! @@ -571,7 +571,7 @@ lemma hausdorff_measure_le_liminf_tsum {β : Type*} {ι : β → Type*} [hι : (d : ℝ) (s : set X) {l : filter β} (r : β → ℝ≥0∞) (hr : tendsto r l (𝓝 0)) (t : Π (n : β), ι n → set X) (ht : ∀ᶠ n in l, ∀ i, diam (t n i) ≤ r n) (hst : ∀ᶠ n in l, s ⊆ ⋃ i, t n i) : - μH[d] s ≤ liminf l (λ n, ∑' i, diam (t n i) ^ d) := + μH[d] s ≤ liminf (λ n, ∑' i, diam (t n i) ^ d) l := mk_metric_le_liminf_tsum s r hr t ht hst _ /-- To bound the Hausdorff measure of a set, one may use coverings with maximum diameter tending @@ -580,7 +580,7 @@ lemma hausdorff_measure_le_liminf_sum {β : Type*} {ι : β → Type*} [hι : (d : ℝ) (s : set X) {l : filter β} (r : β → ℝ≥0∞) (hr : tendsto r l (𝓝 0)) (t : Π (n : β), ι n → set X) (ht : ∀ᶠ n in l, ∀ i, diam (t n i) ≤ r n) (hst : ∀ᶠ n in l, s ⊆ ⋃ i, t n i) : - μH[d] s ≤ liminf l (λ n, ∑ i, diam (t n i) ^ d) := + μH[d] s ≤ liminf (λ n, ∑ i, diam (t n i) ^ d) l := mk_metric_le_liminf_sum s r hr t ht hst _ /-- If `d₁ < d₂`, then for any set `s` we have either `μH[d₂] s = 0`, or `μH[d₁] s = ∞`. -/ @@ -753,10 +753,10 @@ begin ... ≤ (a i : ℝ) + (⌊(x i - a i) * n⌋₊ + 1) / n : add_le_add le_rfl ((div_le_div_right npos).2 (nat.lt_floor_add_one _).le) } }, calc μH[fintype.card ι] (set.pi univ (λ (i : ι), Ioo (a i : ℝ) (b i))) - ≤ liminf at_top (λ (n : ℕ), ∑ (i : γ n), diam (t n i) ^ ↑(fintype.card ι)) : + ≤ liminf (λ (n : ℕ), ∑ (i : γ n), diam (t n i) ^ ↑(fintype.card ι)) at_top : hausdorff_measure_le_liminf_sum _ (set.pi univ (λ i, Ioo (a i : ℝ) (b i))) (λ (n : ℕ), 1/(n : ℝ≥0∞)) A t B C - ... ≤ liminf at_top (λ (n : ℕ), ∑ (i : γ n), (1/n) ^ (fintype.card ι)) : + ... ≤ liminf (λ (n : ℕ), ∑ (i : γ n), (1/n) ^ (fintype.card ι)) at_top : begin refine liminf_le_liminf _ (by is_bounded_default), filter_upwards [B] with _ hn, @@ -764,7 +764,7 @@ begin rw ennreal.rpow_nat_cast, exact pow_le_pow_of_le_left' (hn i) _, end - ... = liminf at_top (λ (n : ℕ), ∏ (i : ι), (⌈((b i : ℝ) - a i) * n⌉₊ : ℝ≥0∞) / n) : + ... = liminf (λ (n : ℕ), ∏ (i : ι), (⌈((b i : ℝ) - a i) * n⌉₊ : ℝ≥0∞) / n) at_top : begin simp only [finset.card_univ, nat.cast_prod, one_mul, fintype.card_fin, finset.sum_const, nsmul_eq_mul, fintype.card_pi, div_eq_mul_inv, finset.prod_mul_distrib, diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index eaf74338f9bc9..3b978df81e1fd 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -503,13 +503,13 @@ end /-- One direction of the **Borel-Cantelli lemma**: if (sᵢ) is a sequence of sets such that `∑ μ sᵢ` is finite, then the limit superior of the `sᵢ` is a null set. -/ -lemma measure_limsup_eq_zero {s : ℕ → set α} (hs : ∑' i, μ (s i) ≠ ∞) : μ (limsup at_top s) = 0 := +lemma measure_limsup_eq_zero {s : ℕ → set α} (hs : ∑' i, μ (s i) ≠ ∞) : μ (limsup s at_top) = 0 := begin -- First we replace the sequence `sₙ` with a sequence of measurable sets `tₙ ⊇ sₙ` of the same -- measure. set t : ℕ → set α := λ n, to_measurable μ (s n), have ht : ∑' i, μ (t i) ≠ ∞, by simpa only [t, measure_to_measurable] using hs, - suffices : μ (limsup at_top t) = 0, + suffices : μ (limsup t at_top) = 0, { have A : s ≤ t := λ n, subset_to_measurable μ (s n), -- TODO default args fail exact measure_mono_null (limsup_le_limsup (eventually_of_forall (pi.le_def.mp A)) diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index d6caaef0e8ee0..279f595ef96de 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -11,9 +11,9 @@ import order.filter.cofinite Defines the Liminf/Limsup of a function taking values in a conditionally complete lattice, with respect to an arbitrary filter. -We define `f.Limsup` (`f.Liminf`) where `f` is a filter taking values in a conditionally complete -lattice. `f.Limsup` is the smallest element `a` such that, eventually, `u ≤ a` (and vice versa for -`f.Liminf`). To work with the Limsup along a function `u` use `(f.map u).Limsup`. +We define `Limsup f` (`Liminf f`) where `f` is a filter taking values in a conditionally complete +lattice. `Limsup f` is the smallest element `a` such that, eventually, `u ≤ a` (and vice versa for +`Liminf f`). To work with the Limsup along a function `u` use `Limsup (map u f)`. Usually, one defines the Limsup as `Inf (Sup s)` where the Inf is taken over all sets in the filter. For instance, in ℕ along a function `u`, this is `Inf_n (Sup_{k ≥ n} u k)` (and the latter quantity @@ -48,7 +48,7 @@ def is_bounded (r : α → α → Prop) (f : filter α) := ∃ b, ∀ᶠ x in f, /-- `f.is_bounded_under (≺) u`: the image of the filter `f` under `u` is eventually bounded w.r.t. the relation `≺`, i.e. eventually, it is bounded by some uniform bound. -/ -def is_bounded_under (r : α → α → Prop) (f : filter β) (u : β → α) := (f.map u).is_bounded r +def is_bounded_under (r : α → α → Prop) (f : filter β) (u : β → α) := (map u f).is_bounded r variables {r : α → α → Prop} {f g : filter α} @@ -154,7 +154,7 @@ def is_cobounded (r : α → α → Prop) (f : filter α) := ∃b, ∀a, (∀ᶠ /-- `is_cobounded_under (≺) f u` states that the image of the filter `f` under the map `u` does not tend to infinity w.r.t. `≺`. This is also called frequently bounded. Will be usually instantiated with `≤` or `≥`. -/ -def is_cobounded_under (r : α → α → Prop) (f : filter β) (u : β → α) := (f.map u).is_cobounded r +def is_cobounded_under (r : α → α → Prop) (f : filter β) (u : β → α) := (map u f).is_cobounded r /-- To check that a filter is frequently bounded, it suffices to have a witness which bounds `f` at some point for every admissible set. @@ -273,124 +273,165 @@ def Liminf (f : filter α) : α := Sup { a | ∀ᶠ n in f, a ≤ n } /-- The `limsup` of a function `u` along a filter `f` is the infimum of the `a` such that, eventually for `f`, holds `u x ≤ a`. -/ -def limsup (f : filter β) (u : β → α) : α := (f.map u).Limsup +def limsup (u : β → α) (f : filter β) : α := Limsup (map u f) /-- The `liminf` of a function `u` along a filter `f` is the supremum of the `a` such that, eventually for `f`, holds `u x ≥ a`. -/ -def liminf (f : filter β) (u : β → α) : α := (f.map u).Liminf +def liminf (u : β → α) (f : filter β) : α := Liminf (map u f) + +/-- The `blimsup` of a function `u` along a filter `f`, bounded by a predicate `p`, is the infimum +of the `a` such that, eventually for `f`, `u x ≤ a` whenever `p x` holds. -/ +def blimsup (u : β → α) (f : filter β) (p : β → Prop) := +Inf { a | ∀ᶠ x in f, p x → u x ≤ a } + +/-- The `bliminf` of a function `u` along a filter `f`, bounded by a predicate `p`, is the supremum +of the `a` such that, eventually for `f`, `a ≤ u x` whenever `p x` holds. -/ +def bliminf (u : β → α) (f : filter β) (p : β → Prop) := +Sup { a | ∀ᶠ x in f, p x → a ≤ u x } section -variables {f : filter β} {u : β → α} -theorem limsup_eq : f.limsup u = Inf { a | ∀ᶠ n in f, u n ≤ a } := rfl -theorem liminf_eq : f.liminf u = Sup { a | ∀ᶠ n in f, a ≤ u n } := rfl + +variables {f : filter β} {u : β → α} {p : β → Prop} + +theorem limsup_eq : limsup u f = Inf { a | ∀ᶠ n in f, u n ≤ a } := rfl + +theorem liminf_eq : liminf u f = Sup { a | ∀ᶠ n in f, a ≤ u n } := rfl + +theorem blimsup_eq : blimsup u f p = Inf { a | ∀ᶠ x in f, p x → u x ≤ a } := rfl + +theorem bliminf_eq : bliminf u f p = Sup { a | ∀ᶠ x in f, p x → a ≤ u x } := rfl + +end + +@[simp] lemma blimsup_true (f : filter β) (u : β → α) : + blimsup u f (λ x, true) = limsup u f := +by simp [blimsup_eq, limsup_eq] + +@[simp] lemma bliminf_true (f : filter β) (u : β → α) : + bliminf u f (λ x, true) = liminf u f := +by simp [bliminf_eq, liminf_eq] + +lemma blimsup_eq_limsup_subtype {f : filter β} {u : β → α} {p : β → Prop} : + blimsup u f p = limsup (u ∘ (coe : {x | p x} → β)) (comap coe f) := +begin + simp only [blimsup_eq, limsup_eq, function.comp_app, eventually_comap, set_coe.forall, + subtype.coe_mk, mem_set_of_eq], + congr, + ext a, + exact eventually_congr (eventually_of_forall + (λ x, ⟨λ hx y hy hxy, hxy.symm ▸ (hx (hxy ▸ hy)), λ hx hx', hx x hx' rfl⟩)), end +lemma bliminf_eq_liminf_subtype {f : filter β} {u : β → α} {p : β → Prop} : + bliminf u f p = liminf (u ∘ (coe : {x | p x} → β)) (comap coe f) := +@blimsup_eq_limsup_subtype αᵒᵈ β _ f u p + theorem Limsup_le_of_le {f : filter α} {a} - (hf : f.is_cobounded (≤) . is_bounded_default) (h : ∀ᶠ n in f, n ≤ a) : f.Limsup ≤ a := + (hf : f.is_cobounded (≤) . is_bounded_default) (h : ∀ᶠ n in f, n ≤ a) : Limsup f ≤ a := cInf_le hf h theorem le_Liminf_of_le {f : filter α} {a} - (hf : f.is_cobounded (≥) . is_bounded_default) (h : ∀ᶠ n in f, a ≤ n) : a ≤ f.Liminf := + (hf : f.is_cobounded (≥) . is_bounded_default) (h : ∀ᶠ n in f, a ≤ n) : a ≤ Liminf f := le_cSup hf h theorem limsup_le_of_le {f : filter β} {u : β → α} {a} (hf : f.is_cobounded_under (≤) u . is_bounded_default) (h : ∀ᶠ n in f, u n ≤ a) : - f.limsup u ≤ a := + limsup u f ≤ a := cInf_le hf h theorem le_liminf_of_le {f : filter β} {u : β → α} {a} (hf : f.is_cobounded_under (≥) u . is_bounded_default) (h : ∀ᶠ n in f, a ≤ u n) : - a ≤ f.liminf u := + a ≤ liminf u f := le_cSup hf h theorem le_Limsup_of_le {f : filter α} {a} (hf : f.is_bounded (≤) . is_bounded_default) (h : ∀ b, (∀ᶠ n in f, n ≤ b) → a ≤ b) : - a ≤ f.Limsup := + a ≤ Limsup f := le_cInf hf h theorem Liminf_le_of_le {f : filter α} {a} (hf : f.is_bounded (≥) . is_bounded_default) (h : ∀ b, (∀ᶠ n in f, b ≤ n) → b ≤ a) : - f.Liminf ≤ a := + Liminf f ≤ a := cSup_le hf h theorem le_limsup_of_le {f : filter β} {u : β → α} {a} (hf : f.is_bounded_under (≤) u . is_bounded_default) (h : ∀ b, (∀ᶠ n in f, u n ≤ b) → a ≤ b) : - a ≤ f.limsup u := + a ≤ limsup u f := le_cInf hf h theorem liminf_le_of_le {f : filter β} {u : β → α} {a} (hf : f.is_bounded_under (≥) u . is_bounded_default) (h : ∀ b, (∀ᶠ n in f, b ≤ u n) → b ≤ a) : - f.liminf u ≤ a := + liminf u f ≤ a := cSup_le hf h theorem Liminf_le_Limsup {f : filter α} [ne_bot f] (h₁ : f.is_bounded (≤) . is_bounded_default) (h₂ : f.is_bounded (≥) . is_bounded_default) : - f.Liminf ≤ f.Limsup := + Liminf f ≤ Limsup f := Liminf_le_of_le h₂ $ assume a₀ ha₀, le_Limsup_of_le h₁ $ assume a₁ ha₁, show a₀ ≤ a₁, from let ⟨b, hb₀, hb₁⟩ := (ha₀.and ha₁).exists in le_trans hb₀ hb₁ lemma liminf_le_limsup {f : filter β} [ne_bot f] {u : β → α} (h : f.is_bounded_under (≤) u . is_bounded_default) (h' : f.is_bounded_under (≥) u . is_bounded_default) : - liminf f u ≤ limsup f u := + liminf u f ≤ limsup u f := Liminf_le_Limsup h h' lemma Limsup_le_Limsup {f g : filter α} (hf : f.is_cobounded (≤) . is_bounded_default) (hg : g.is_bounded (≤) . is_bounded_default) - (h : ∀ a, (∀ᶠ n in g, n ≤ a) → ∀ᶠ n in f, n ≤ a) : f.Limsup ≤ g.Limsup := + (h : ∀ a, (∀ᶠ n in g, n ≤ a) → ∀ᶠ n in f, n ≤ a) : Limsup f ≤ Limsup g := cInf_le_cInf hf hg h lemma Liminf_le_Liminf {f g : filter α} (hf : f.is_bounded (≥) . is_bounded_default) (hg : g.is_cobounded (≥) . is_bounded_default) - (h : ∀ a, (∀ᶠ n in f, a ≤ n) → ∀ᶠ n in g, a ≤ n) : f.Liminf ≤ g.Liminf := + (h : ∀ a, (∀ᶠ n in f, a ≤ n) → ∀ᶠ n in g, a ≤ n) : Liminf f ≤ Liminf g := cSup_le_cSup hg hf h lemma limsup_le_limsup {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} (h : u ≤ᶠ[f] v) (hu : f.is_cobounded_under (≤) u . is_bounded_default) (hv : f.is_bounded_under (≤) v . is_bounded_default) : - f.limsup u ≤ f.limsup v := + limsup u f ≤ limsup v f := Limsup_le_Limsup hu hv $ assume b, h.trans lemma liminf_le_liminf {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} (h : ∀ᶠ a in f, u a ≤ v a) (hu : f.is_bounded_under (≥) u . is_bounded_default) (hv : f.is_cobounded_under (≥) v . is_bounded_default) : - f.liminf u ≤ f.liminf v := + liminf u f ≤ liminf v f := @limsup_le_limsup βᵒᵈ α _ _ _ _ h hv hu lemma Limsup_le_Limsup_of_le {f g : filter α} (h : f ≤ g) (hf : f.is_cobounded (≤) . is_bounded_default) (hg : g.is_bounded (≤) . is_bounded_default) : - f.Limsup ≤ g.Limsup := + Limsup f ≤ Limsup g := Limsup_le_Limsup hf hg (assume a ha, h ha) lemma Liminf_le_Liminf_of_le {f g : filter α} (h : g ≤ f) (hf : f.is_bounded (≥) . is_bounded_default) (hg : g.is_cobounded (≥) . is_bounded_default) : - f.Liminf ≤ g.Liminf := + Liminf f ≤ Liminf g := Liminf_le_Liminf hf hg (assume a ha, h ha) lemma limsup_le_limsup_of_le {α β} [conditionally_complete_lattice β] {f g : filter α} (h : f ≤ g) {u : α → β} (hf : f.is_cobounded_under (≤) u . is_bounded_default) (hg : g.is_bounded_under (≤) u . is_bounded_default) : - f.limsup u ≤ g.limsup u := + limsup u f ≤ limsup u g := Limsup_le_Limsup_of_le (map_mono h) hf hg lemma liminf_le_liminf_of_le {α β} [conditionally_complete_lattice β] {f g : filter α} (h : g ≤ f) {u : α → β} (hf : f.is_bounded_under (≥) u . is_bounded_default) (hg : g.is_cobounded_under (≥) u . is_bounded_default) : - f.liminf u ≤ g.liminf u := + liminf u f ≤ liminf u g := Liminf_le_Liminf_of_le (map_mono h) hf hg theorem Limsup_principal {s : set α} (h : bdd_above s) (hs : s.nonempty) : - (𝓟 s).Limsup = Sup s := + Limsup (𝓟 s) = Sup s := by simp [Limsup]; exact cInf_upper_bounds_eq_cSup h hs theorem Liminf_principal {s : set α} (h : bdd_below s) (hs : s.nonempty) : - (𝓟 s).Liminf = Inf s := + Liminf (𝓟 s) = Inf s := @Limsup_principal αᵒᵈ _ s h hs lemma limsup_congr {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} - (h : ∀ᶠ a in f, u a = v a) : limsup f u = limsup f v := + (h : ∀ᶠ a in f, u a = v a) : limsup u f = limsup v f := begin rw limsup_eq, congr' with b, @@ -398,15 +439,15 @@ begin end lemma liminf_congr {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} - (h : ∀ᶠ a in f, u a = v a) : liminf f u = liminf f v := + (h : ∀ᶠ a in f, u a = v a) : liminf u f = liminf v f := @limsup_congr βᵒᵈ _ _ _ _ _ h lemma limsup_const {α : Type*} [conditionally_complete_lattice β] {f : filter α} [ne_bot f] - (b : β) : limsup f (λ x, b) = b := + (b : β) : limsup (λ x, b) f = b := by simpa only [limsup_eq, eventually_const] using cInf_Ici lemma liminf_const {α : Type*} [conditionally_complete_lattice β] {f : filter α} [ne_bot f] - (b : β) : liminf f (λ x, b) = b := + (b : β) : liminf (λ x, b) f = b := @limsup_const βᵒᵈ α _ f _ b @@ -415,88 +456,114 @@ end conditionally_complete_lattice section complete_lattice variables [complete_lattice α] -@[simp] theorem Limsup_bot : (⊥ : filter α).Limsup = ⊥ := +@[simp] theorem Limsup_bot : Limsup (⊥ : filter α) = ⊥ := bot_unique $ Inf_le $ by simp -@[simp] theorem Liminf_bot : (⊥ : filter α).Liminf = ⊤ := +@[simp] theorem Liminf_bot : Liminf (⊥ : filter α) = ⊤ := top_unique $ le_Sup $ by simp -@[simp] theorem Limsup_top : (⊤ : filter α).Limsup = ⊤ := +@[simp] theorem Limsup_top : Limsup (⊤ : filter α) = ⊤ := top_unique $ le_Inf $ by simp [eq_univ_iff_forall]; exact assume b hb, (top_unique $ hb _) -@[simp] theorem Liminf_top : (⊤ : filter α).Liminf = ⊥ := +@[simp] theorem Liminf_top : Liminf (⊤ : filter α) = ⊥ := bot_unique $ Sup_le $ by simp [eq_univ_iff_forall]; exact assume b hb, (bot_unique $ hb _) +@[simp] lemma blimsup_false {f : filter β} {u : β → α} : + blimsup u f (λ x, false) = ⊥ := +by simp [blimsup_eq] + +@[simp] lemma bliminf_false {f : filter β} {u : β → α} : + bliminf u f (λ x, false) = ⊤ := +by simp [bliminf_eq] + /-- Same as limsup_const applied to `⊥` but without the `ne_bot f` assumption -/ -lemma limsup_const_bot {f : filter β} : limsup f (λ x : β, (⊥ : α)) = (⊥ : α) := +lemma limsup_const_bot {f : filter β} : limsup (λ x : β, (⊥ : α)) f = (⊥ : α) := begin rw [limsup_eq, eq_bot_iff], exact Inf_le (eventually_of_forall (λ x, le_rfl)), end /-- Same as limsup_const applied to `⊤` but without the `ne_bot f` assumption -/ -lemma liminf_const_top {f : filter β} : liminf f (λ x : β, (⊤ : α)) = (⊤ : α) := +lemma liminf_const_top {f : filter β} : liminf (λ x : β, (⊤ : α)) f = (⊤ : α) := @limsup_const_bot αᵒᵈ β _ _ theorem has_basis.Limsup_eq_infi_Sup {ι} {p : ι → Prop} {s} {f : filter α} (h : f.has_basis p s) : - f.Limsup = ⨅ i (hi : p i), Sup (s i) := + Limsup f = ⨅ i (hi : p i), Sup (s i) := le_antisymm (le_infi₂ $ λ i hi, Inf_le $ h.eventually_iff.2 ⟨i, hi, λ x, le_Sup⟩) (le_Inf $ assume a ha, let ⟨i, hi, ha⟩ := h.eventually_iff.1 ha in infi₂_le_of_le _ hi $ Sup_le ha) theorem has_basis.Liminf_eq_supr_Inf {p : ι → Prop} {s : ι → set α} {f : filter α} - (h : f.has_basis p s) : f.Liminf = ⨆ i (hi : p i), Inf (s i) := + (h : f.has_basis p s) : Liminf f = ⨆ i (hi : p i), Inf (s i) := @has_basis.Limsup_eq_infi_Sup αᵒᵈ _ _ _ _ _ h -theorem Limsup_eq_infi_Sup {f : filter α} : f.Limsup = ⨅ s ∈ f, Sup s := +theorem Limsup_eq_infi_Sup {f : filter α} : Limsup f = ⨅ s ∈ f, Sup s := f.basis_sets.Limsup_eq_infi_Sup -theorem Liminf_eq_supr_Inf {f : filter α} : f.Liminf = ⨆ s ∈ f, Inf s := +theorem Liminf_eq_supr_Inf {f : filter α} : Liminf f = ⨆ s ∈ f, Inf s := @Limsup_eq_infi_Sup αᵒᵈ _ _ -theorem limsup_le_supr {f : filter β} {u : β → α} : f.limsup u ≤ ⨆ n, u n := +theorem limsup_le_supr {f : filter β} {u : β → α} : limsup u f ≤ ⨆ n, u n := limsup_le_of_le (by is_bounded_default) (eventually_of_forall (le_supr u)) -theorem infi_le_liminf {f : filter β} {u : β → α} : (⨅ n, u n) ≤ f.liminf u := +theorem infi_le_liminf {f : filter β} {u : β → α} : (⨅ n, u n) ≤ liminf u f := le_liminf_of_le (by is_bounded_default) (eventually_of_forall (infi_le u)) /-- In a complete lattice, the limsup of a function is the infimum over sets `s` in the filter of the supremum of the function over `s` -/ -theorem limsup_eq_infi_supr {f : filter β} {u : β → α} : f.limsup u = ⨅ s ∈ f, ⨆ a ∈ s, u a := +theorem limsup_eq_infi_supr {f : filter β} {u : β → α} : limsup u f = ⨅ s ∈ f, ⨆ a ∈ s, u a := (f.basis_sets.map u).Limsup_eq_infi_Sup.trans $ by simp only [Sup_image, id] -lemma limsup_eq_infi_supr_of_nat {u : ℕ → α} : limsup at_top u = ⨅ n : ℕ, ⨆ i ≥ n, u i := +lemma limsup_eq_infi_supr_of_nat {u : ℕ → α} : limsup u at_top = ⨅ n : ℕ, ⨆ i ≥ n, u i := (at_top_basis.map u).Limsup_eq_infi_Sup.trans $ by simp only [Sup_image, infi_const]; refl -lemma limsup_eq_infi_supr_of_nat' {u : ℕ → α} : limsup at_top u = ⨅ n : ℕ, ⨆ i : ℕ, u (i + n) := +lemma limsup_eq_infi_supr_of_nat' {u : ℕ → α} : limsup u at_top = ⨅ n : ℕ, ⨆ i : ℕ, u (i + n) := by simp only [limsup_eq_infi_supr_of_nat, supr_ge_eq_supr_nat_add] theorem has_basis.limsup_eq_infi_supr {p : ι → Prop} {s : ι → set β} {f : filter β} {u : β → α} - (h : f.has_basis p s) : f.limsup u = ⨅ i (hi : p i), ⨆ a ∈ s i, u a := + (h : f.has_basis p s) : limsup u f = ⨅ i (hi : p i), ⨆ a ∈ s i, u a := (h.map u).Limsup_eq_infi_Sup.trans $ by simp only [Sup_image, id] +lemma blimsup_eq_infi_bsupr {f : filter β} {p : β → Prop} {u : β → α} : + blimsup u f p = ⨅ s ∈ f, ⨆ b (hb : p b ∧ b ∈ s), u b := +begin + refine le_antisymm (Inf_le_Inf _) (infi_le_iff.mpr $ λ a ha, le_Inf_iff.mpr $ λ a' ha', _), + { rintros - ⟨s, rfl⟩, + simp only [mem_set_of_eq, le_infi_iff], + conv { congr, funext, rw imp.swap, }, + refine eventually_imp_distrib_left.mpr (λ h, eventually_iff_exists_mem.2 ⟨s, h, λ x h₁ h₂, _⟩), + exact @le_supr₂ α β (λ b, p b ∧ b ∈ s) _ (λ b hb, u b) x ⟨h₂, h₁⟩, }, + { obtain ⟨s, hs, hs'⟩ := eventually_iff_exists_mem.mp ha', + simp_rw imp.swap at hs', + exact (le_infi_iff.mp (ha s) hs).trans (by simpa only [supr₂_le_iff, and_imp]), }, +end + /-- In a complete lattice, the liminf of a function is the infimum over sets `s` in the filter of the supremum of the function over `s` -/ -theorem liminf_eq_supr_infi {f : filter β} {u : β → α} : f.liminf u = ⨆ s ∈ f, ⨅ a ∈ s, u a := +theorem liminf_eq_supr_infi {f : filter β} {u : β → α} : liminf u f = ⨆ s ∈ f, ⨅ a ∈ s, u a := @limsup_eq_infi_supr αᵒᵈ β _ _ _ -lemma liminf_eq_supr_infi_of_nat {u : ℕ → α} : liminf at_top u = ⨆ n : ℕ, ⨅ i ≥ n, u i := +lemma liminf_eq_supr_infi_of_nat {u : ℕ → α} : liminf u at_top = ⨆ n : ℕ, ⨅ i ≥ n, u i := @limsup_eq_infi_supr_of_nat αᵒᵈ _ u -lemma liminf_eq_supr_infi_of_nat' {u : ℕ → α} : liminf at_top u = ⨆ n : ℕ, ⨅ i : ℕ, u (i + n) := +lemma liminf_eq_supr_infi_of_nat' {u : ℕ → α} : liminf u at_top = ⨆ n : ℕ, ⨅ i : ℕ, u (i + n) := @limsup_eq_infi_supr_of_nat' αᵒᵈ _ _ theorem has_basis.liminf_eq_supr_infi {p : ι → Prop} {s : ι → set β} {f : filter β} {u : β → α} - (h : f.has_basis p s) : f.liminf u = ⨆ i (hi : p i), ⨅ a ∈ s i, u a := + (h : f.has_basis p s) : liminf u f = ⨆ i (hi : p i), ⨅ a ∈ s i, u a := @has_basis.limsup_eq_infi_supr αᵒᵈ _ _ _ _ _ _ _ h +lemma bliminf_eq_supr_binfi {f : filter β} {p : β → Prop} {u : β → α} : + bliminf u f p = ⨆ s ∈ f, ⨅ b (hb : p b ∧ b ∈ s), u b := +@blimsup_eq_infi_bsupr αᵒᵈ β _ f p u + lemma limsup_eq_Inf_Sup {ι R : Type*} (F : filter ι) [complete_lattice R] (a : ι → R) : - F.limsup a = Inf ((λ I, Sup (a '' I)) '' F.sets) := + limsup a F = Inf ((λ I, Sup (a '' I)) '' F.sets) := begin refine le_antisymm _ _, { rw limsup_eq, @@ -511,20 +578,20 @@ begin end lemma liminf_eq_Sup_Inf {ι R : Type*} (F : filter ι) [complete_lattice R] (a : ι → R) : - F.liminf a = Sup ((λ I, Inf (a '' I)) '' F.sets) := + liminf a F = Sup ((λ I, Inf (a '' I)) '' F.sets) := @filter.limsup_eq_Inf_Sup ι (order_dual R) _ _ a @[simp] lemma liminf_nat_add (f : ℕ → α) (k : ℕ) : - at_top.liminf (λ i, f (i + k)) = at_top.liminf f := + liminf (λ i, f (i + k)) at_top = liminf f at_top := by { simp_rw liminf_eq_supr_infi_of_nat, exact supr_infi_ge_nat_add f k } @[simp] lemma limsup_nat_add (f : ℕ → α) (k : ℕ) : - at_top.limsup (λ i, f (i + k)) = at_top.limsup f := + limsup (λ i, f (i + k)) at_top = limsup f at_top := @liminf_nat_add αᵒᵈ _ f k lemma liminf_le_of_frequently_le' {α β} [complete_lattice β] {f : filter α} {u : α → β} {x : β} (h : ∃ᶠ a in f, u a ≤ x) : - f.liminf u ≤ x := + liminf u f ≤ x := begin rw liminf_eq, refine Sup_le (λ b hb, _), @@ -537,15 +604,124 @@ end lemma le_limsup_of_frequently_le' {α β} [complete_lattice β] {f : filter α} {u : α → β} {x : β} (h : ∃ᶠ a in f, x ≤ u a) : - x ≤ f.limsup u := + x ≤ limsup u f := @liminf_le_of_frequently_le' _ βᵒᵈ _ _ _ _ h +variables {f g : filter β} {p q : β → Prop} {u v : β → α} + +lemma blimsup_mono (h : ∀ x, p x → q x) : + blimsup u f p ≤ blimsup u f q := +Inf_le_Inf $ λ a ha, ha.mono $ by tauto + +lemma bliminf_antitone (h : ∀ x, p x → q x) : + bliminf u f q ≤ bliminf u f p := +Sup_le_Sup $ λ a ha, ha.mono $ by tauto + +lemma mono_blimsup' (h : ∀ᶠ x in f, u x ≤ v x) : + blimsup u f p ≤ blimsup v f p := +Inf_le_Inf $ λ a ha, (ha.and h).mono $ λ x hx hx', hx.2.trans (hx.1 hx') + +lemma mono_blimsup (h : ∀ x, u x ≤ v x) : + blimsup u f p ≤ blimsup v f p := +mono_blimsup' $ eventually_of_forall h + +lemma mono_bliminf' (h : ∀ᶠ x in f, u x ≤ v x) : + bliminf u f p ≤ bliminf v f p := +Sup_le_Sup $ λ a ha, (ha.and h).mono $ λ x hx hx', (hx.1 hx').trans hx.2 + +lemma mono_bliminf (h : ∀ x, u x ≤ v x) : + bliminf u f p ≤ bliminf v f p := +mono_bliminf' $ eventually_of_forall h + +lemma bliminf_antitone_filter (h : f ≤ g) : + bliminf u g p ≤ bliminf u f p := +Sup_le_Sup $ λ a ha, ha.filter_mono h + +lemma blimsup_monotone_filter (h : f ≤ g) : + blimsup u f p ≤ blimsup u g p := +Inf_le_Inf $ λ a ha, ha.filter_mono h + +@[simp] lemma blimsup_and_le_inf : + blimsup u f (λ x, p x ∧ q x) ≤ blimsup u f p ⊓ blimsup u f q := +le_inf (blimsup_mono $ by tauto) (blimsup_mono $ by tauto) + +@[simp] lemma bliminf_sup_le_and : + bliminf u f p ⊔ bliminf u f q ≤ bliminf u f (λ x, p x ∧ q x) := +@blimsup_and_le_inf αᵒᵈ β _ f p q u + +/-- See also `filter.blimsup_or_eq_sup`. -/ +@[simp] lemma blimsup_sup_le_or : + blimsup u f p ⊔ blimsup u f q ≤ blimsup u f (λ x, p x ∨ q x) := +sup_le (blimsup_mono $ by tauto) (blimsup_mono $ by tauto) + +/-- See also `filter.bliminf_or_eq_inf`. -/ +@[simp] lemma bliminf_or_le_inf : + bliminf u f (λ x, p x ∨ q x) ≤ bliminf u f p ⊓ bliminf u f q := +@blimsup_sup_le_or αᵒᵈ β _ f p q u + end complete_lattice +section complete_distrib_lattice + +variables [complete_distrib_lattice α] {f : filter β} {p q : β → Prop} {u : β → α} + +@[simp] lemma blimsup_or_eq_sup : + blimsup u f (λ x, p x ∨ q x) = blimsup u f p ⊔ blimsup u f q := +begin + refine le_antisymm _ blimsup_sup_le_or, + simp only [blimsup_eq, Inf_sup_eq, sup_Inf_eq, le_infi₂_iff, mem_set_of_eq], + refine λ a' ha' a ha, Inf_le ((ha.and ha').mono $ λ b h hb, _), + exact or.elim hb (λ hb, le_sup_of_le_left $ h.1 hb) (λ hb, le_sup_of_le_right $ h.2 hb), +end + +@[simp] lemma bliminf_or_eq_inf : + bliminf u f (λ x, p x ∨ q x) = bliminf u f p ⊓ bliminf u f q := +@blimsup_or_eq_sup αᵒᵈ β _ f p q u + +end complete_distrib_lattice + +section set_lattice + +variables {p : ι → Prop} {s : ι → set α} + +lemma cofinite.blimsup_set_eq : + blimsup s cofinite p = { x | { n | p n ∧ x ∈ s n }.infinite } := +begin + simp only [blimsup_eq, le_eq_subset, eventually_cofinite, not_forall, Inf_eq_sInter, exists_prop], + ext x, + refine ⟨λ h, _, λ hx t h, _⟩; + contrapose! h, + { simp only [mem_sInter, mem_set_of_eq, not_forall, exists_prop], + exact ⟨{x}ᶜ, by simpa using h, by simp⟩, }, + { exact hx.mono (λ i hi, ⟨hi.1, λ hit, h (hit hi.2)⟩), }, +end + +lemma cofinite.bliminf_set_eq : + bliminf s cofinite p = { x | { n | p n ∧ x ∉ s n }.finite } := +begin + rw ← compl_inj_iff, + simpa only [bliminf_eq_supr_binfi, compl_infi, compl_supr, ← blimsup_eq_infi_bsupr, + cofinite.blimsup_set_eq], +end + +/-- In other words, `limsup cofinite s` is the set of elements lying inside the family `s` +infinitely often. -/ +lemma cofinite.limsup_set_eq : + limsup s cofinite = { x | { n | x ∈ s n }.infinite } := +by simp only [← cofinite.blimsup_true s, cofinite.blimsup_set_eq, true_and] + +/-- In other words, `liminf cofinite s` is the set of elements lying outside the family `s` +finitely often. -/ +lemma cofinite.liminf_set_eq : + liminf s cofinite = { x | { n | x ∉ s n }.finite } := +by simp only [← cofinite.bliminf_true s, cofinite.bliminf_set_eq, true_and] + +end set_lattice + section conditionally_complete_linear_order lemma frequently_lt_of_lt_Limsup {f : filter α} [conditionally_complete_linear_order α] {a : α} - (hf : f.is_cobounded (≤) . is_bounded_default) (h : a < f.Limsup) : ∃ᶠ n in f, a < n := + (hf : f.is_cobounded (≤) . is_bounded_default) (h : a < Limsup f) : ∃ᶠ n in f, a < n := begin contrapose! h, simp only [not_frequently, not_lt] at h, @@ -553,11 +729,11 @@ begin end lemma frequently_lt_of_Liminf_lt {f : filter α} [conditionally_complete_linear_order α] {a : α} - (hf : f.is_cobounded (≥) . is_bounded_default) (h : f.Liminf < a) : ∃ᶠ n in f, n < a := + (hf : f.is_cobounded (≥) . is_bounded_default) (h : Liminf f < a) : ∃ᶠ n in f, n < a := @frequently_lt_of_lt_Limsup (order_dual α) f _ a hf h lemma eventually_lt_of_lt_liminf {f : filter α} [conditionally_complete_linear_order β] - {u : α → β} {b : β} (h : b < liminf f u) (hu : f.is_bounded_under (≥) u . is_bounded_default) : + {u : α → β} {b : β} (h : b < liminf u f) (hu : f.is_bounded_under (≥) u . is_bounded_default) : ∀ᶠ a in f, b < u a := begin obtain ⟨c, hc, hbc⟩ : ∃ (c : β) (hc : c ∈ {c : β | ∀ᶠ (n : α) in f, c ≤ u n}), b < c := @@ -566,14 +742,14 @@ begin end lemma eventually_lt_of_limsup_lt {f : filter α} [conditionally_complete_linear_order β] - {u : α → β} {b : β} (h : limsup f u < b) (hu : f.is_bounded_under (≤) u . is_bounded_default) : + {u : α → β} {b : β} (h : limsup u f < b) (hu : f.is_bounded_under (≤) u . is_bounded_default) : ∀ᶠ a in f, u a < b := @eventually_lt_of_lt_liminf _ βᵒᵈ _ _ _ _ h hu lemma le_limsup_of_frequently_le {α β} [conditionally_complete_linear_order β] {f : filter α} {u : α → β} {b : β} (hu_le : ∃ᶠ x in f, b ≤ u x) (hu : f.is_bounded_under (≤) u . is_bounded_default) : - b ≤ f.limsup u := + b ≤ limsup u f := begin revert hu_le, rw [←not_imp_not, not_frequently], @@ -584,12 +760,12 @@ end lemma liminf_le_of_frequently_le {α β} [conditionally_complete_linear_order β] {f : filter α} {u : α → β} {b : β} (hu_le : ∃ᶠ x in f, u x ≤ b) (hu : f.is_bounded_under (≥) u . is_bounded_default) : - f.liminf u ≤ b := + liminf u f ≤ b := @le_limsup_of_frequently_le _ βᵒᵈ _ f u b hu_le hu lemma frequently_lt_of_lt_limsup {α β} [conditionally_complete_linear_order β] {f : filter α} {u : α → β} {b : β} - (hu : f.is_cobounded_under (≤) u . is_bounded_default) (h : b < f.limsup u) : + (hu : f.is_cobounded_under (≤) u . is_bounded_default) (h : b < limsup u f) : ∃ᶠ x in f, b < u x := begin contrapose! h, @@ -599,7 +775,7 @@ end lemma frequently_lt_of_liminf_lt {α β} [conditionally_complete_linear_order β] {f : filter α} {u : α → β} {b : β} - (hu : f.is_cobounded_under (≥) u . is_bounded_default) (h : f.liminf u < b) : + (hu : f.is_cobounded_under (≥) u . is_bounded_default) (h : liminf u f < b) : ∃ᶠ x in f, u x < b := @frequently_lt_of_lt_limsup _ βᵒᵈ _ f u b hu h @@ -644,7 +820,7 @@ lemma galois_connection.l_limsup_le [conditionally_complete_lattice β] {l : β → γ} {u : γ → β} (gc : galois_connection l u) (hlv : f.is_bounded_under (≤) (λ x, l (v x)) . is_bounded_default) (hv_co : f.is_cobounded_under (≤) v . is_bounded_default) : - l (f.limsup v) ≤ f.limsup (λ x, l (v x)) := + l (limsup v f) ≤ limsup (λ x, l (v x)) f := begin refine le_Limsup_of_le hlv (λ c hc, _), rw filter.eventually_map at hc, @@ -658,10 +834,10 @@ lemma order_iso.limsup_apply {γ} [conditionally_complete_lattice β] (hu_co : f.is_cobounded_under (≤) u . is_bounded_default) (hgu : f.is_bounded_under (≤) (λ x, g (u x)) . is_bounded_default) (hgu_co : f.is_cobounded_under (≤) (λ x, g (u x)) . is_bounded_default) : - g (f.limsup u) = f.limsup (λ x, g (u x)) := + g (limsup u f) = limsup (λ x, g (u x)) f := begin refine le_antisymm (g.to_galois_connection.l_limsup_le hgu hu_co) _, - rw [←(g.symm.symm_apply_apply (f.limsup (λ (x : α), g (u x)))), g.symm_symm], + rw [←(g.symm.symm_apply_apply $ limsup (λ x, g (u x)) f), g.symm_symm], refine g.monotone _, have hf : u = λ i, g.symm (g (u i)), from funext (λ i, (g.symm_apply_apply (u i)).symm), nth_rewrite 0 hf, @@ -676,7 +852,7 @@ lemma order_iso.liminf_apply {γ} [conditionally_complete_lattice β] (hu_co : f.is_cobounded_under (≥) u . is_bounded_default) (hgu : f.is_bounded_under (≥) (λ x, g (u x)) . is_bounded_default) (hgu_co : f.is_cobounded_under (≥) (λ x, g (u x)) . is_bounded_default) : - g (f.liminf u) = f.liminf (λ x, g (u x)) := + g (liminf u f) = liminf (λ x, g (u x)) f := @order_iso.limsup_apply α βᵒᵈ γᵒᵈ _ _ f u g.dual hu hu_co hgu hgu_co end order diff --git a/src/probability/martingale/borel_cantelli.lean b/src/probability/martingale/borel_cantelli.lean index ba26caf71b6ad..d2c1a9641fdbc 100644 --- a/src/probability/martingale/borel_cantelli.lean +++ b/src/probability/martingale/borel_cantelli.lean @@ -389,7 +389,7 @@ filtration `ℱ` such that for all `n`, `s n` is `ℱ n`-measurable, `at_top.lim everywhere equal to the set for which `∑ k, ℙ(s (k + 1) | ℱ k) = ∞`. -/ theorem ae_mem_limsup_at_top_iff [is_finite_measure μ] {s : ℕ → set Ω} (hs : ∀ n, measurable_set[ℱ n] (s n)) : - ∀ᵐ ω ∂μ, ω ∈ limsup at_top s ↔ + ∀ᵐ ω ∂μ, ω ∈ limsup s at_top ↔ tendsto (λ n, ∑ k in finset.range n, μ[(s (k + 1)).indicator (1 : Ω → ℝ) | ℱ k] ω) at_top at_top := (limsup_eq_tendsto_sum_indicator_at_top ℝ s).symm ▸ tendsto_sum_indicator_at_top_iff' hs diff --git a/src/probability/martingale/convergence.lean b/src/probability/martingale/convergence.lean index d29bc0750b90a..012be9ecce518 100644 --- a/src/probability/martingale/convergence.lean +++ b/src/probability/martingale/convergence.lean @@ -136,7 +136,7 @@ convergent. We use the spelling `< ∞` instead of the standard `≠ ∞` in the assumptions since it is not as easy to change `<` to `≠` under binders. -/ lemma tendsto_of_uncrossing_lt_top - (hf₁ : liminf at_top (λ n, (∥f n ω∥₊ : ℝ≥0∞)) < ∞) + (hf₁ : liminf (λ n, (∥f n ω∥₊ : ℝ≥0∞)) at_top < ∞) (hf₂ : ∀ a b : ℚ, a < b → upcrossings a b f ω < ∞) : ∃ c, tendsto (λ n, f n ω) at_top (𝓝 c) := begin diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index a50d7c1808ed9..5200a03484d2f 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -151,18 +151,18 @@ theorem Limsup_eq_of_le_nhds : ∀ {f : filter α} {a : α} [ne_bot f], f ≤ /-- If a function has a limit, then its limsup coincides with its limit. -/ theorem filter.tendsto.limsup_eq {f : filter β} {u : β → α} {a : α} [ne_bot f] - (h : tendsto u f (𝓝 a)) : limsup f u = a := + (h : tendsto u f (𝓝 a)) : limsup u f = a := Limsup_eq_of_le_nhds h /-- If a function has a limit, then its liminf coincides with its limit. -/ theorem filter.tendsto.liminf_eq {f : filter β} {u : β → α} {a : α} [ne_bot f] - (h : tendsto u f (𝓝 a)) : liminf f u = a := + (h : tendsto u f (𝓝 a)) : liminf u f = a := Liminf_eq_of_le_nhds h /-- If the liminf and the limsup of a function coincide, then the limit of the function exists and has the same value -/ theorem tendsto_of_liminf_eq_limsup {f : filter β} {u : β → α} {a : α} - (hinf : liminf f u = a) (hsup : limsup f u = a) + (hinf : liminf u f = a) (hsup : limsup u f = a) (h : f.is_bounded_under (≤) u . is_bounded_default) (h' : f.is_bounded_under (≥) u . is_bounded_default) : tendsto u f (𝓝 a) := @@ -171,7 +171,7 @@ le_nhds_of_Limsup_eq_Liminf h h' hsup hinf /-- If a number `a` is less than or equal to the `liminf` of a function `f` at some filter and is greater than or equal to the `limsup` of `f`, then `f` tends to `a` along this filter. -/ theorem tendsto_of_le_liminf_of_limsup_le {f : filter β} {u : β → α} {a : α} - (hinf : a ≤ liminf f u) (hsup : limsup f u ≤ a) + (hinf : a ≤ liminf u f) (hsup : limsup u f ≤ a) (h : f.is_bounded_under (≤) u . is_bounded_default) (h' : f.is_bounded_under (≥) u . is_bounded_default) : tendsto u f (𝓝 a) := @@ -193,7 +193,7 @@ lemma tendsto_of_no_upcrossings [densely_ordered α] begin by_cases hbot : f = ⊥, { rw hbot, exact ⟨Inf ∅, tendsto_bot⟩ }, haveI : ne_bot f := ⟨hbot⟩, - refine ⟨limsup f u, _⟩, + refine ⟨limsup u f, _⟩, apply tendsto_of_le_liminf_of_limsup_le _ le_rfl h h', by_contra' hlt, obtain ⟨a, ⟨⟨la, au⟩, as⟩⟩ : ∃ a, (f.liminf u < a ∧ a < f.limsup u) ∧ a ∈ s := @@ -324,7 +324,7 @@ section indicator open_locale big_operators lemma limsup_eq_tendsto_sum_indicator_nat_at_top (s : ℕ → set α) : - limsup at_top s = + limsup s at_top = {ω | tendsto (λ n, ∑ k in finset.range n, (s (k + 1)).indicator (1 : α → ℕ) ω) at_top at_top} := begin ext ω, @@ -383,7 +383,7 @@ end lemma limsup_eq_tendsto_sum_indicator_at_top (R : Type*) [strict_ordered_semiring R] [nontrivial R] [archimedean R] (s : ℕ → set α) : - limsup at_top s = + limsup s at_top = {ω | tendsto (λ n, ∑ k in finset.range n, (s (k + 1)).indicator (1 : α → R) ω) at_top at_top} := begin rw limsup_eq_tendsto_sum_indicator_nat_at_top s, diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 89a7b6f46ff5f..eda048d0cf97a 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -488,11 +488,11 @@ lemma inv_map_supr {ι : Sort*} {x : ι → ℝ≥0∞} : order_iso.inv_ennreal.map_supr x lemma inv_limsup {ι : Sort*} {x : ι → ℝ≥0∞} {l : filter ι} : - (l.limsup x)⁻¹ = l.liminf (λ i, (x i)⁻¹) := + (limsup x l)⁻¹ = liminf (λ i, (x i)⁻¹) l := by simp only [limsup_eq_infi_supr, inv_map_infi, inv_map_supr, liminf_eq_supr_infi] lemma inv_liminf {ι : Sort*} {x : ι → ℝ≥0∞} {l : filter ι} : - (l.liminf x)⁻¹ = l.limsup (λ i, (x i)⁻¹) := + (liminf x l)⁻¹ = limsup (λ i, (x i)⁻¹) l := by simp only [limsup_eq_infi_supr, inv_map_infi, inv_map_supr, liminf_eq_supr_infi] instance : has_continuous_inv ℝ≥0∞ := ⟨order_iso.inv_ennreal.continuous⟩ @@ -654,7 +654,7 @@ end topological_space section liminf lemma exists_frequently_lt_of_liminf_ne_top - {ι : Type*} {l : filter ι} {x : ι → ℝ} (hx : liminf l (λ n, (∥x n∥₊ : ℝ≥0∞)) ≠ ∞) : + {ι : Type*} {l : filter ι} {x : ι → ℝ} (hx : liminf (λ n, (∥x n∥₊ : ℝ≥0∞)) l ≠ ∞) : ∃ R, ∃ᶠ n in l, x n < R := begin by_contra h, @@ -665,7 +665,7 @@ begin end lemma exists_frequently_lt_of_liminf_ne_top' - {ι : Type*} {l : filter ι} {x : ι → ℝ} (hx : liminf l (λ n, (∥x n∥₊ : ℝ≥0∞)) ≠ ∞) : + {ι : Type*} {l : filter ι} {x : ι → ℝ} (hx : liminf (λ n, (∥x n∥₊ : ℝ≥0∞)) l ≠ ∞) : ∃ R, ∃ᶠ n in l, R < x n := begin by_contra h, @@ -677,7 +677,7 @@ end lemma exists_upcrossings_of_not_bounded_under {ι : Type*} {l : filter ι} {x : ι → ℝ} - (hf : liminf l (λ i, (∥x i∥₊ : ℝ≥0∞)) ≠ ∞) + (hf : liminf (λ i, (∥x i∥₊ : ℝ≥0∞)) l ≠ ∞) (hbdd : ¬ is_bounded_under (≤) l (λ i, |x i|)) : ∃ a b : ℚ, a < b ∧ (∃ᶠ i in l, x i < a) ∧ (∃ᶠ i in l, ↑b < x i) := begin @@ -783,7 +783,7 @@ protected lemma tsum_eq_supr_nat {f : ℕ → ℝ≥0∞} : ennreal.tsum_eq_supr_sum' _ finset.exists_nat_subset_range protected lemma tsum_eq_liminf_sum_nat {f : ℕ → ℝ≥0∞} : - ∑' i, f i = filter.at_top.liminf (λ n, ∑ i in finset.range n, f i) := + ∑' i, f i = liminf (λ n, ∑ i in finset.range n, f i) at_top := begin rw [ennreal.tsum_eq_supr_nat, filter.liminf_eq_supr_infi_of_nat], congr, diff --git a/src/topology/metric_space/hausdorff_dimension.lean b/src/topology/metric_space/hausdorff_dimension.lean index fcfcd22bc4213..976e0f419be16 100644 --- a/src/topology/metric_space/hausdorff_dimension.lean +++ b/src/topology/metric_space/hausdorff_dimension.lean @@ -226,7 +226,7 @@ end /-- In an (extended) metric space with second countable topology, the Hausdorff dimension of a set `s` is the supremum over `x ∈ s` of the limit superiors of `dimH t` along `(𝓝[s] x).small_sets`. -/ -lemma bsupr_limsup_dimH (s : set X) : (⨆ x ∈ s, limsup (𝓝[s] x).small_sets dimH) = dimH s := +lemma bsupr_limsup_dimH (s : set X) : (⨆ x ∈ s, limsup dimH (𝓝[s] x).small_sets) = dimH s := begin refine le_antisymm (supr₂_le $ λ x hx, _) _, { refine Limsup_le_of_le (by apply_auto_param) (eventually_map.2 _), @@ -241,7 +241,7 @@ end /-- In an (extended) metric space with second countable topology, the Hausdorff dimension of a set `s` is the supremum over all `x` of the limit superiors of `dimH t` along `(𝓝[s] x).small_sets`. -/ -lemma supr_limsup_dimH (s : set X) : (⨆ x, limsup (𝓝[s] x).small_sets dimH) = dimH s := +lemma supr_limsup_dimH (s : set X) : (⨆ x, limsup dimH (𝓝[s] x).small_sets) = dimH s := begin refine le_antisymm (supr_le $ λ x, _) _, { refine Limsup_le_of_le (by apply_auto_param) (eventually_map.2 _), From 938f1f1b89b04dc15659cca28163f250d201d6a1 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 13 Oct 2022 11:47:26 +0000 Subject: [PATCH 744/828] chore(tactic/core): fix bug in `resolve_attribute_expr_list` (#16827) See https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/namespace.20affects.20behaviour.20of.20.60apply_list_expr.60 Co-authored-by: Alex J Best Co-authored-by: Alex J Best --- src/tactic/core.lean | 3 ++- test/apply_rules.lean | 12 ++++++++++-- 2 files changed, 12 insertions(+), 3 deletions(-) diff --git a/src/tactic/core.lean b/src/tactic/core.lean index ee8b157133eb7..845f0254d8fc3 100644 --- a/src/tactic/core.lean +++ b/src/tactic/core.lean @@ -1009,7 +1009,8 @@ application of `i_to_expr_for_apply` to a declaration with that attribute. -/ meta def resolve_attribute_expr_list (attr_name : name) : tactic (list (tactic expr)) := do l ← attribute.get_instances attr_name, - list.map i_to_expr_for_apply <$> list.reverse <$> l.mmap resolve_name + list.map i_to_expr_for_apply <$> list.reverse + <$> l.mmap (λ n, do c ← (mk_const n), return (pexpr.of_expr c)) /--`apply_rules args attrs n`: apply the lists of rules `args` (given as pexprs) and `attrs` (given diff --git a/test/apply_rules.lean b/test/apply_rules.lean index ca0595b9b710e..9e1a078289f7f 100644 --- a/test/apply_rules.lean +++ b/test/apply_rules.lean @@ -1,5 +1,4 @@ - -import data.nat.basic +import data.int.basic open nat @@ -49,3 +48,12 @@ axiom p_rules : P 0 example : P 0 := by success_if_fail {apply_rules with p_rules}; apply_rules [p_rules] example : P 10 := by apply_rules with p_rules 60 + +attribute [p_rules] pow_lt_pow_of_lt_left + +open nat + +-- This tests for the following bug: +-- https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/namespace.20affects.20behaviour.20of.20.60apply_list_expr.60 +example {x y : ℤ} (n : ℕ) (h1 : x < y) (h2 : 0 ≤ x) (h3 : 0 < n) : x ^ n < y ^ n := +by apply_rules with p_rules From 02b85de47704de46eaff748bd4c71aaae5245451 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 13 Oct 2022 11:47:27 +0000 Subject: [PATCH 745/828] chore(algebra/order/sub): split file (#16868) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit - [x] depends on: #16861 [![Open in Gitpod](https://gitpod.io/button/open-in-gitpod.svg)](https://gitpod.io/from-referrer/) Co-authored-by: Yaël Dillies --- .../order/{group.lean => group/basic.lean} | 25 +- src/algebra/order/group/type_tags.lean | 27 ++ src/algebra/order/lattice_group.lean | 2 +- src/algebra/order/monoid/canonical.lean | 2 +- src/algebra/order/ring.lean | 3 +- src/algebra/order/sub/basic.lean | 396 ++++++++++++++++ .../order/{sub.lean => sub/canonical.lean} | 423 +----------------- src/algebra/order/sub/with_top.lean | 46 ++ src/algebra/order/with_zero.lean | 2 +- src/data/nat/enat.lean | 1 + src/data/nat/with_bot.lean | 2 +- src/data/real/ennreal.lean | 1 + src/data/set/intervals/basic.lean | 2 +- 13 files changed, 486 insertions(+), 446 deletions(-) rename src/algebra/order/{group.lean => group/basic.lean} (98%) create mode 100644 src/algebra/order/group/type_tags.lean create mode 100644 src/algebra/order/sub/basic.lean rename src/algebra/order/{sub.lean => sub/canonical.lean} (52%) create mode 100644 src/algebra/order/sub/with_top.lean diff --git a/src/algebra/order/group.lean b/src/algebra/order/group/basic.lean similarity index 98% rename from src/algebra/order/group.lean rename to src/algebra/order/group/basic.lean index 9aa93a76c2d33..6d7854c018612 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group/basic.lean @@ -4,11 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import algebra.abs -import algebra.order.sub +import algebra.order.sub.basic import algebra.order.monoid.min_max import algebra.order.monoid.prod -import algebra.order.monoid.type_tags import algebra.order.monoid.units +import algebra.order.monoid.order_dual +import algebra.order.monoid.with_top /-! # Ordered groups @@ -1337,26 +1338,6 @@ instance [ordered_comm_group G] [ordered_comm_group H] : end prod -section type_tags - -instance [ordered_add_comm_group α] : ordered_comm_group (multiplicative α) := -{ ..multiplicative.comm_group, - ..multiplicative.ordered_comm_monoid } - -instance [ordered_comm_group α] : ordered_add_comm_group (additive α) := -{ ..additive.add_comm_group, - ..additive.ordered_add_comm_monoid } - -instance [linear_ordered_add_comm_group α] : linear_ordered_comm_group (multiplicative α) := -{ ..multiplicative.linear_order, - ..multiplicative.ordered_comm_group } - -instance [linear_ordered_comm_group α] : linear_ordered_add_comm_group (additive α) := -{ ..additive.linear_order, - ..additive.ordered_add_comm_group } - -end type_tags - section norm_num_lemmas /- The following lemmas are stated so that the `norm_num` tactic can use them with the expected signatures. -/ diff --git a/src/algebra/order/group/type_tags.lean b/src/algebra/order/group/type_tags.lean new file mode 100644 index 0000000000000..88310593826b7 --- /dev/null +++ b/src/algebra/order/group/type_tags.lean @@ -0,0 +1,27 @@ +/- +Copyright (c) 2016 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl +-/ +import algebra.order.group.basic +import algebra.order.monoid.type_tags + +/-! # Ordered group structures on `multiplicative α` and `additive α`. -/ + +variables {α : Type*} + +instance [ordered_add_comm_group α] : ordered_comm_group (multiplicative α) := +{ ..multiplicative.comm_group, + ..multiplicative.ordered_comm_monoid } + +instance [ordered_comm_group α] : ordered_add_comm_group (additive α) := +{ ..additive.add_comm_group, + ..additive.ordered_add_comm_monoid } + +instance [linear_ordered_add_comm_group α] : linear_ordered_comm_group (multiplicative α) := +{ ..multiplicative.linear_order, + ..multiplicative.ordered_comm_group } + +instance [linear_ordered_comm_group α] : linear_ordered_add_comm_group (additive α) := +{ ..additive.linear_order, + ..additive.ordered_add_comm_group } diff --git a/src/algebra/order/lattice_group.lean b/src/algebra/order/lattice_group.lean index 4ddd79c0c4d14..05fc3f40c9749 100644 --- a/src/algebra/order/lattice_group.lean +++ b/src/algebra/order/lattice_group.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Christopher Hoskin -/ import algebra.group_power.basic -- Needed for squares -import algebra.order.group +import algebra.order.group.basic import tactic.nth_rewrite /-! diff --git a/src/algebra/order/monoid/canonical.lean b/src/algebra/order/monoid/canonical.lean index 88e4ec9edfb57..9931cafb67783 100644 --- a/src/algebra/order/monoid/canonical.lean +++ b/src/algebra/order/monoid/canonical.lean @@ -7,7 +7,7 @@ import algebra.order.monoid.basic import order.min_max /-! -# Canonical ordered monoids +# Canonically ordered monoids -/ universe u diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index 6bec664ae1ad2..a1d66ceab2751 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -6,7 +6,8 @@ Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro import algebra.char_zero.defs import algebra.ring.divisibility import algebra.hom.ring -import algebra.order.group +import algebra.order.group.basic +import algebra.order.sub.canonical import algebra.order.ring_lemmas /-! diff --git a/src/algebra/order/sub/basic.lean b/src/algebra/order/sub/basic.lean new file mode 100644 index 0000000000000..306d4427a8276 --- /dev/null +++ b/src/algebra/order/sub/basic.lean @@ -0,0 +1,396 @@ +/- +Copyright (c) 2021 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn +-/ +import order.hom.basic +import algebra.hom.equiv +import algebra.ring.basic + +/-! +# Ordered Subtraction + +This file proves lemmas relating (truncated) subtraction with an order. We provide a class +`has_ordered_sub` stating that `a - b ≤ c ↔ a ≤ c + b`. + +The subtraction discussed here could both be normal subtraction in an additive group or truncated +subtraction on a canonically ordered monoid (`ℕ`, `multiset`, `part_enat`, `ennreal`, ...) + +## Implementation details + +`has_ordered_sub` is a mixin type-class, so that we can use the results in this file even in cases +where we don't have a `canonically_ordered_add_monoid` instance +(even though that is our main focus). Conversely, this means we can use +`canonically_ordered_add_monoid` without necessarily having to define a subtraction. + +The results in this file are ordered by the type-class assumption needed to prove it. +This means that similar results might not be close to each other. Furthermore, we don't prove +implications if a bi-implication can be proven under the same assumptions. + +Lemmas using this class are named using `tsub` instead of `sub` (short for "truncated subtraction"). +This is to avoid naming conflicts with similar lemmas about ordered groups. + +We provide a second version of most results that require `[contravariant_class α α (+) (≤)]`. In the +second version we replace this type-class assumption by explicit `add_le_cancellable` assumptions. + +TODO: maybe we should make a multiplicative version of this, so that we can replace some identical +lemmas about subtraction/division in `ordered_[add_]comm_group` with these. + +TODO: generalize `nat.le_of_le_of_sub_le_sub_right`, `nat.sub_le_sub_right_iff`, + `nat.mul_self_sub_mul_self_eq` +-/ + +variables {α β : Type*} + +/-- `has_ordered_sub α` means that `α` has a subtraction characterized by `a - b ≤ c ↔ a ≤ c + b`. +In other words, `a - b` is the least `c` such that `a ≤ b + c`. + +This is satisfied both by the subtraction in additive ordered groups and by truncated subtraction +in canonically ordered monoids on many specific types. +-/ +class has_ordered_sub (α : Type*) [has_le α] [has_add α] [has_sub α] := +(tsub_le_iff_right : ∀ a b c : α, a - b ≤ c ↔ a ≤ c + b) + +section has_add + +variables [preorder α] [has_add α] [has_sub α] [has_ordered_sub α] {a b c d : α} + +@[simp] lemma tsub_le_iff_right : a - b ≤ c ↔ a ≤ c + b := +has_ordered_sub.tsub_le_iff_right a b c + +/-- See `add_tsub_cancel_right` for the equality if `contravariant_class α α (+) (≤)`. -/ +lemma add_tsub_le_right : a + b - b ≤ a := +tsub_le_iff_right.mpr le_rfl + +lemma le_tsub_add : b ≤ (b - a) + a := +tsub_le_iff_right.mp le_rfl + +lemma add_hom.le_map_tsub [preorder β] [has_add β] [has_sub β] [has_ordered_sub β] + (f : add_hom α β) (hf : monotone f) (a b : α) : + f a - f b ≤ f (a - b) := +by { rw [tsub_le_iff_right, ← f.map_add], exact hf le_tsub_add } + +lemma le_mul_tsub {R : Type*} [distrib R] [preorder R] [has_sub R] [has_ordered_sub R] + [covariant_class R R (*) (≤)] {a b c : R} : + a * b - a * c ≤ a * (b - c) := +(add_hom.mul_left a).le_map_tsub (monotone_id.const_mul' a) _ _ + +lemma le_tsub_mul {R : Type*} [comm_semiring R] [preorder R] [has_sub R] [has_ordered_sub R] + [covariant_class R R (*) (≤)] {a b c : R} : + a * c - b * c ≤ (a - b) * c := +by simpa only [mul_comm _ c] using le_mul_tsub + +end has_add + +/-- An order isomorphism between types with ordered subtraction preserves subtraction provided that +it preserves addition. -/ +lemma order_iso.map_tsub {M N : Type*} [preorder M] [has_add M] [has_sub M] [has_ordered_sub M] + [partial_order N] [has_add N] [has_sub N] [has_ordered_sub N] (e : M ≃o N) + (h_add : ∀ a b, e (a + b) = e a + e b) (a b : M) : + e (a - b) = e a - e b := +begin + set e_add : M ≃+ N := { map_add' := h_add, .. e }, + refine le_antisymm _ (e_add.to_add_hom.le_map_tsub e.monotone a b), + suffices : e (e.symm (e a) - e.symm (e b)) ≤ e (e.symm (e a - e b)), by simpa, + exact e.monotone (e_add.symm.to_add_hom.le_map_tsub e.symm.monotone _ _) +end + +/-! ### Preorder -/ + +section ordered_add_comm_semigroup + +section preorder +variables [preorder α] + +section add_comm_semigroup +variables [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] {a b c d : α} + +lemma tsub_le_iff_left : a - b ≤ c ↔ a ≤ b + c := +by rw [tsub_le_iff_right, add_comm] + +lemma le_add_tsub : a ≤ b + (a - b) := +tsub_le_iff_left.mp le_rfl + +/-- See `add_tsub_cancel_left` for the equality if `contravariant_class α α (+) (≤)`. -/ +lemma add_tsub_le_left : a + b - a ≤ b := +tsub_le_iff_left.mpr le_rfl + +lemma tsub_le_tsub_right (h : a ≤ b) (c : α) : a - c ≤ b - c := +tsub_le_iff_left.mpr $ h.trans le_add_tsub + +lemma tsub_le_iff_tsub_le : a - b ≤ c ↔ a - c ≤ b := +by rw [tsub_le_iff_left, tsub_le_iff_right] + +/-- See `tsub_tsub_cancel_of_le` for the equality. -/ +lemma tsub_tsub_le : b - (b - a) ≤ a := +tsub_le_iff_right.mpr le_add_tsub + +section cov +variable [covariant_class α α (+) (≤)] + +lemma tsub_le_tsub_left (h : a ≤ b) (c : α) : c - b ≤ c - a := +tsub_le_iff_left.mpr $ le_add_tsub.trans $ add_le_add_right h _ + +lemma tsub_le_tsub (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c := +(tsub_le_tsub_right hab _).trans $ tsub_le_tsub_left hcd _ + +lemma antitone_const_tsub : antitone (λ x, c - x) := +λ x y hxy, tsub_le_tsub rfl.le hxy + +/-- See `add_tsub_assoc_of_le` for the equality. -/ +lemma add_tsub_le_assoc : a + b - c ≤ a + (b - c) := +by { rw [tsub_le_iff_left, add_left_comm], exact add_le_add_left le_add_tsub a } + +/-- See `tsub_add_eq_add_tsub` for the equality. -/ +lemma add_tsub_le_tsub_add : a + b - c ≤ a - c + b := +by { rw [add_comm, add_comm _ b], exact add_tsub_le_assoc } + +lemma add_le_add_add_tsub : a + b ≤ (a + c) + (b - c) := +by { rw [add_assoc], exact add_le_add_left le_add_tsub a } + +lemma le_tsub_add_add : a + b ≤ (a - c) + (b + c) := +by { rw [add_comm a, add_comm (a - c)], exact add_le_add_add_tsub } + +lemma tsub_le_tsub_add_tsub : a - c ≤ (a - b) + (b - c) := +begin + rw [tsub_le_iff_left, ← add_assoc, add_right_comm], + exact le_add_tsub.trans (add_le_add_right le_add_tsub _), +end + +lemma tsub_tsub_tsub_le_tsub : (c - a) - (c - b) ≤ b - a := +begin + rw [tsub_le_iff_left, tsub_le_iff_left, add_left_comm], + exact le_tsub_add.trans (add_le_add_left le_add_tsub _), +end + +lemma tsub_tsub_le_tsub_add {a b c : α} : a - (b - c) ≤ a - b + c := +tsub_le_iff_right.2 $ calc + a ≤ a - b + b : le_tsub_add + ... ≤ a - b + (c + (b - c)) : add_le_add_left le_add_tsub _ + ... = a - b + c + (b - c) : (add_assoc _ _ _).symm + +/-- See `tsub_add_tsub_comm` for the equality. -/ +lemma add_tsub_add_le_tsub_add_tsub : a + b - (c + d) ≤ a - c + (b - d) := +begin + rw [add_comm c, tsub_le_iff_left, add_assoc, ←tsub_le_iff_left, ←tsub_le_iff_left], + refine (tsub_le_tsub_right add_tsub_le_assoc c).trans _, + rw [add_comm a, add_comm (a - c)], + exact add_tsub_le_assoc, +end + +/-- See `add_tsub_add_eq_tsub_left` for the equality. -/ +lemma add_tsub_add_le_tsub_left : a + b - (a + c) ≤ b - c := +by { rw [tsub_le_iff_left, add_assoc], exact add_le_add_left le_add_tsub _ } + +/-- See `add_tsub_add_eq_tsub_right` for the equality. -/ +lemma add_tsub_add_le_tsub_right : a + c - (b + c) ≤ a - b := +by { rw [tsub_le_iff_left, add_right_comm], exact add_le_add_right le_add_tsub c } + +end cov + +/-! #### Lemmas that assume that an element is `add_le_cancellable` -/ + +namespace add_le_cancellable + +protected lemma le_add_tsub_swap (hb : add_le_cancellable b) : a ≤ b + a - b := hb le_add_tsub + +protected lemma le_add_tsub (hb : add_le_cancellable b) : a ≤ a + b - b := +by { rw add_comm, exact hb.le_add_tsub_swap } + +protected lemma le_tsub_of_add_le_left (ha : add_le_cancellable a) (h : a + b ≤ c) : b ≤ c - a := +ha $ h.trans le_add_tsub + +protected lemma le_tsub_of_add_le_right (hb : add_le_cancellable b) (h : a + b ≤ c) : a ≤ c - b := +hb.le_tsub_of_add_le_left $ by rwa add_comm + +end add_le_cancellable + +/-! ### Lemmas where addition is order-reflecting -/ + +section contra +variable [contravariant_class α α (+) (≤)] + +lemma le_add_tsub_swap : a ≤ b + a - b := contravariant.add_le_cancellable.le_add_tsub_swap + +lemma le_add_tsub' : a ≤ a + b - b := contravariant.add_le_cancellable.le_add_tsub + +lemma le_tsub_of_add_le_left (h : a + b ≤ c) : b ≤ c - a := +contravariant.add_le_cancellable.le_tsub_of_add_le_left h + +lemma le_tsub_of_add_le_right (h : a + b ≤ c) : a ≤ c - b := +contravariant.add_le_cancellable.le_tsub_of_add_le_right h + +end contra + +end add_comm_semigroup + +variables [add_comm_monoid α] [has_sub α] [has_ordered_sub α] {a b c d : α} + +lemma tsub_nonpos : a - b ≤ 0 ↔ a ≤ b := by rw [tsub_le_iff_left, add_zero] + +alias tsub_nonpos ↔ _ tsub_nonpos_of_le + +lemma add_monoid_hom.le_map_tsub [preorder β] [add_comm_monoid β] [has_sub β] + [has_ordered_sub β] (f : α →+ β) (hf : monotone f) (a b : α) : + f a - f b ≤ f (a - b) := +f.to_add_hom.le_map_tsub hf a b + +end preorder + +/-! ### Partial order -/ + +variables [partial_order α] [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] {a b c d : α} + +lemma tsub_tsub (b a c : α) : b - a - c = b - (a + c) := +begin + apply le_antisymm, + { rw [tsub_le_iff_left, tsub_le_iff_left, ← add_assoc, ← tsub_le_iff_left] }, + { rw [tsub_le_iff_left, add_assoc, ← tsub_le_iff_left, ← tsub_le_iff_left] } +end + +lemma tsub_add_eq_tsub_tsub (a b c : α) : a - (b + c) = a - b - c := (tsub_tsub _ _ _).symm + +lemma tsub_add_eq_tsub_tsub_swap (a b c : α) : a - (b + c) = a - c - b := +by { rw [add_comm], apply tsub_add_eq_tsub_tsub } + +lemma tsub_right_comm : a - b - c = a - c - b := +by simp_rw [← tsub_add_eq_tsub_tsub, add_comm] + +/-! ### Lemmas that assume that an element is `add_le_cancellable`. -/ + +namespace add_le_cancellable + +protected lemma tsub_eq_of_eq_add (hb : add_le_cancellable b) (h : a = c + b) : a - b = c := +le_antisymm (tsub_le_iff_right.mpr h.le) $ + by { rw h, exact hb.le_add_tsub } + +protected lemma eq_tsub_of_add_eq (hc : add_le_cancellable c) (h : a + c = b) : a = b - c := +(hc.tsub_eq_of_eq_add h.symm).symm + +protected theorem tsub_eq_of_eq_add_rev (hb : add_le_cancellable b) (h : a = b + c) : a - b = c := +hb.tsub_eq_of_eq_add $ by rw [add_comm, h] + +@[simp] +protected lemma add_tsub_cancel_right (hb : add_le_cancellable b) : a + b - b = a := +hb.tsub_eq_of_eq_add $ by rw [add_comm] + +@[simp] +protected lemma add_tsub_cancel_left (ha : add_le_cancellable a) : a + b - a = b := +ha.tsub_eq_of_eq_add $ add_comm a b + +protected lemma lt_add_of_tsub_lt_left (hb : add_le_cancellable b) (h : a - b < c) : a < b + c := +begin + rw [lt_iff_le_and_ne, ← tsub_le_iff_left], + refine ⟨h.le, _⟩, + rintro rfl, + simpa [hb] using h, +end + +protected lemma lt_add_of_tsub_lt_right (hc : add_le_cancellable c) (h : a - c < b) : a < b + c := +begin + rw [lt_iff_le_and_ne, ← tsub_le_iff_right], + refine ⟨h.le, _⟩, + rintro rfl, + simpa [hc] using h, +end + +protected lemma lt_tsub_of_add_lt_right (hc : add_le_cancellable c) (h : a + c < b) : a < b - c := +(hc.le_tsub_of_add_le_right h.le).lt_of_ne $ by { rintro rfl, exact h.not_le le_tsub_add } + +protected lemma lt_tsub_of_add_lt_left (ha : add_le_cancellable a) (h : a + c < b) : c < b - a := +ha.lt_tsub_of_add_lt_right $ by rwa add_comm + +end add_le_cancellable + +/-! #### Lemmas where addition is order-reflecting. -/ + +section contra +variable [contravariant_class α α (+) (≤)] + +lemma tsub_eq_of_eq_add (h : a = c + b) : a - b = c := +contravariant.add_le_cancellable.tsub_eq_of_eq_add h + +lemma eq_tsub_of_add_eq (h : a + c = b) : a = b - c := +contravariant.add_le_cancellable.eq_tsub_of_add_eq h + +lemma tsub_eq_of_eq_add_rev (h : a = b + c) : a - b = c := +contravariant.add_le_cancellable.tsub_eq_of_eq_add_rev h + +@[simp] +lemma add_tsub_cancel_right (a b : α) : a + b - b = a := +contravariant.add_le_cancellable.add_tsub_cancel_right + +@[simp] +lemma add_tsub_cancel_left (a b : α) : a + b - a = b := +contravariant.add_le_cancellable.add_tsub_cancel_left + +lemma lt_add_of_tsub_lt_left (h : a - b < c) : a < b + c := +contravariant.add_le_cancellable.lt_add_of_tsub_lt_left h + +lemma lt_add_of_tsub_lt_right (h : a - c < b) : a < b + c := +contravariant.add_le_cancellable.lt_add_of_tsub_lt_right h + +/-- This lemma (and some of its corollaries) also holds for `ennreal`, but this proof doesn't work +for it. Maybe we should add this lemma as field to `has_ordered_sub`? -/ +lemma lt_tsub_of_add_lt_left : a + c < b → c < b - a := +contravariant.add_le_cancellable.lt_tsub_of_add_lt_left + +lemma lt_tsub_of_add_lt_right : a + c < b → a < b - c := +contravariant.add_le_cancellable.lt_tsub_of_add_lt_right + +end contra + +section both +variables [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤)] + +lemma add_tsub_add_eq_tsub_right (a c b : α) : (a + c) - (b + c) = a - b := +begin + refine add_tsub_add_le_tsub_right.antisymm (tsub_le_iff_right.2 $ le_of_add_le_add_right _), swap, + rw add_assoc, + exact le_tsub_add, +end + +lemma add_tsub_add_eq_tsub_left (a b c : α) : (a + b) - (a + c) = b - c := +by rw [add_comm a b, add_comm a c, add_tsub_add_eq_tsub_right] + +end both + +end ordered_add_comm_semigroup + +/-! ### Lemmas in a linearly ordered monoid. -/ +section linear_order +variables {a b c d : α} [linear_order α] [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] + +/-- See `lt_of_tsub_lt_tsub_right_of_le` for a weaker statement in a partial order. -/ +lemma lt_of_tsub_lt_tsub_right (h : a - c < b - c) : a < b := +lt_imp_lt_of_le_imp_le (λ h, tsub_le_tsub_right h c) h + +/-- See `lt_tsub_iff_right_of_le` for a weaker statement in a partial order. -/ +lemma lt_tsub_iff_right : a < b - c ↔ a + c < b := +lt_iff_lt_of_le_iff_le tsub_le_iff_right + +/-- See `lt_tsub_iff_left_of_le` for a weaker statement in a partial order. -/ +lemma lt_tsub_iff_left : a < b - c ↔ c + a < b := +lt_iff_lt_of_le_iff_le tsub_le_iff_left + +lemma lt_tsub_comm : a < b - c ↔ c < b - a := +lt_tsub_iff_left.trans lt_tsub_iff_right.symm + +section cov +variable [covariant_class α α (+) (≤)] + +/-- See `lt_of_tsub_lt_tsub_left_of_le` for a weaker statement in a partial order. -/ +lemma lt_of_tsub_lt_tsub_left (h : a - b < a - c) : c < b := +lt_imp_lt_of_le_imp_le (λ h, tsub_le_tsub_left h a) h + +end cov + +end linear_order + +section ordered_add_comm_monoid +variables [partial_order α] [add_comm_monoid α] [has_sub α] [has_ordered_sub α] + +@[simp] lemma tsub_zero (a : α) : a - 0 = a := +add_le_cancellable.tsub_eq_of_eq_add add_le_cancellable_zero (add_zero _).symm + +end ordered_add_comm_monoid diff --git a/src/algebra/order/sub.lean b/src/algebra/order/sub/canonical.lean similarity index 52% rename from src/algebra/order/sub.lean rename to src/algebra/order/sub/canonical.lean index c9cc427b90096..5d3f58c8983b7 100644 --- a/src/algebra/order/sub.lean +++ b/src/algebra/order/sub/canonical.lean @@ -3,394 +3,14 @@ Copyright (c) 2021 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ -import algebra.order.monoid.with_top -/-! -# Ordered Subtraction - -This file proves lemmas relating (truncated) subtraction with an order. We provide a class -`has_ordered_sub` stating that `a - b ≤ c ↔ a ≤ c + b`. - -The subtraction discussed here could both be normal subtraction in an additive group or truncated -subtraction on a canonically ordered monoid (`ℕ`, `multiset`, `part_enat`, `ennreal`, ...) - -## Implementation details - -`has_ordered_sub` is a mixin type-class, so that we can use the results in this file even in cases -where we don't have a `canonically_ordered_add_monoid` instance -(even though that is our main focus). Conversely, this means we can use -`canonically_ordered_add_monoid` without necessarily having to define a subtraction. - -The results in this file are ordered by the type-class assumption needed to prove it. -This means that similar results might not be close to each other. Furthermore, we don't prove -implications if a bi-implication can be proven under the same assumptions. - -Lemmas using this class are named using `tsub` instead of `sub` (short for "truncated subtraction"). -This is to avoid naming conflicts with similar lemmas about ordered groups. - -We provide a second version of most results that require `[contravariant_class α α (+) (≤)]`. In the -second version we replace this type-class assumption by explicit `add_le_cancellable` assumptions. - -TODO: maybe we should make a multiplicative version of this, so that we can replace some identical -lemmas about subtraction/division in `ordered_[add_]comm_group` with these. - -TODO: generalize `nat.le_of_le_of_sub_le_sub_right`, `nat.sub_le_sub_right_iff`, - `nat.mul_self_sub_mul_self_eq` --/ - -variables {α β : Type*} - -/-- `has_ordered_sub α` means that `α` has a subtraction characterized by `a - b ≤ c ↔ a ≤ c + b`. -In other words, `a - b` is the least `c` such that `a ≤ b + c`. +import algebra.order.monoid.canonical +import algebra.order.sub.basic -This is satisfied both by the subtraction in additive ordered groups and by truncated subtraction -in canonically ordered monoids on many specific types. +/-! +# Lemmas about subtraction in canonically ordered monoids -/ -class has_ordered_sub (α : Type*) [has_le α] [has_add α] [has_sub α] := -(tsub_le_iff_right : ∀ a b c : α, a - b ≤ c ↔ a ≤ c + b) - -section has_add - -variables [preorder α] [has_add α] [has_sub α] [has_ordered_sub α] {a b c d : α} - -@[simp] lemma tsub_le_iff_right : a - b ≤ c ↔ a ≤ c + b := -has_ordered_sub.tsub_le_iff_right a b c - -/-- See `add_tsub_cancel_right` for the equality if `contravariant_class α α (+) (≤)`. -/ -lemma add_tsub_le_right : a + b - b ≤ a := -tsub_le_iff_right.mpr le_rfl - -lemma le_tsub_add : b ≤ (b - a) + a := -tsub_le_iff_right.mp le_rfl - -lemma add_hom.le_map_tsub [preorder β] [has_add β] [has_sub β] [has_ordered_sub β] - (f : add_hom α β) (hf : monotone f) (a b : α) : - f a - f b ≤ f (a - b) := -by { rw [tsub_le_iff_right, ← f.map_add], exact hf le_tsub_add } - -lemma le_mul_tsub {R : Type*} [distrib R] [preorder R] [has_sub R] [has_ordered_sub R] - [covariant_class R R (*) (≤)] {a b c : R} : - a * b - a * c ≤ a * (b - c) := -(add_hom.mul_left a).le_map_tsub (monotone_id.const_mul' a) _ _ - -lemma le_tsub_mul {R : Type*} [comm_semiring R] [preorder R] [has_sub R] [has_ordered_sub R] - [covariant_class R R (*) (≤)] {a b c : R} : - a * c - b * c ≤ (a - b) * c := -by simpa only [mul_comm _ c] using le_mul_tsub - -end has_add - -/-- An order isomorphism between types with ordered subtraction preserves subtraction provided that -it preserves addition. -/ -lemma order_iso.map_tsub {M N : Type*} [preorder M] [has_add M] [has_sub M] [has_ordered_sub M] - [partial_order N] [has_add N] [has_sub N] [has_ordered_sub N] (e : M ≃o N) - (h_add : ∀ a b, e (a + b) = e a + e b) (a b : M) : - e (a - b) = e a - e b := -begin - set e_add : M ≃+ N := { map_add' := h_add, .. e }, - refine le_antisymm _ (e_add.to_add_hom.le_map_tsub e.monotone a b), - suffices : e (e.symm (e a) - e.symm (e b)) ≤ e (e.symm (e a - e b)), by simpa, - exact e.monotone (e_add.symm.to_add_hom.le_map_tsub e.symm.monotone _ _) -end - -/-! ### Preorder -/ - -section ordered_add_comm_semigroup - -section preorder -variables [preorder α] - -section add_comm_semigroup -variables [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] {a b c d : α} - -lemma tsub_le_iff_left : a - b ≤ c ↔ a ≤ b + c := -by rw [tsub_le_iff_right, add_comm] - -lemma le_add_tsub : a ≤ b + (a - b) := -tsub_le_iff_left.mp le_rfl - -/-- See `add_tsub_cancel_left` for the equality if `contravariant_class α α (+) (≤)`. -/ -lemma add_tsub_le_left : a + b - a ≤ b := -tsub_le_iff_left.mpr le_rfl - -lemma tsub_le_tsub_right (h : a ≤ b) (c : α) : a - c ≤ b - c := -tsub_le_iff_left.mpr $ h.trans le_add_tsub - -lemma tsub_le_iff_tsub_le : a - b ≤ c ↔ a - c ≤ b := -by rw [tsub_le_iff_left, tsub_le_iff_right] - -/-- See `tsub_tsub_cancel_of_le` for the equality. -/ -lemma tsub_tsub_le : b - (b - a) ≤ a := -tsub_le_iff_right.mpr le_add_tsub - -section cov -variable [covariant_class α α (+) (≤)] - -lemma tsub_le_tsub_left (h : a ≤ b) (c : α) : c - b ≤ c - a := -tsub_le_iff_left.mpr $ le_add_tsub.trans $ add_le_add_right h _ - -lemma tsub_le_tsub (hab : a ≤ b) (hcd : c ≤ d) : a - d ≤ b - c := -(tsub_le_tsub_right hab _).trans $ tsub_le_tsub_left hcd _ - -lemma antitone_const_tsub : antitone (λ x, c - x) := -λ x y hxy, tsub_le_tsub rfl.le hxy - -/-- See `add_tsub_assoc_of_le` for the equality. -/ -lemma add_tsub_le_assoc : a + b - c ≤ a + (b - c) := -by { rw [tsub_le_iff_left, add_left_comm], exact add_le_add_left le_add_tsub a } - -/-- See `tsub_add_eq_add_tsub` for the equality. -/ -lemma add_tsub_le_tsub_add : a + b - c ≤ a - c + b := -by { rw [add_comm, add_comm _ b], exact add_tsub_le_assoc } - -lemma add_le_add_add_tsub : a + b ≤ (a + c) + (b - c) := -by { rw [add_assoc], exact add_le_add_left le_add_tsub a } - -lemma le_tsub_add_add : a + b ≤ (a - c) + (b + c) := -by { rw [add_comm a, add_comm (a - c)], exact add_le_add_add_tsub } - -lemma tsub_le_tsub_add_tsub : a - c ≤ (a - b) + (b - c) := -begin - rw [tsub_le_iff_left, ← add_assoc, add_right_comm], - exact le_add_tsub.trans (add_le_add_right le_add_tsub _), -end - -lemma tsub_tsub_tsub_le_tsub : (c - a) - (c - b) ≤ b - a := -begin - rw [tsub_le_iff_left, tsub_le_iff_left, add_left_comm], - exact le_tsub_add.trans (add_le_add_left le_add_tsub _), -end - -lemma tsub_tsub_le_tsub_add {a b c : α} : a - (b - c) ≤ a - b + c := -tsub_le_iff_right.2 $ calc - a ≤ a - b + b : le_tsub_add - ... ≤ a - b + (c + (b - c)) : add_le_add_left le_add_tsub _ - ... = a - b + c + (b - c) : (add_assoc _ _ _).symm - -/-- See `tsub_add_tsub_comm` for the equality. -/ -lemma add_tsub_add_le_tsub_add_tsub : a + b - (c + d) ≤ a - c + (b - d) := -begin - rw [add_comm c, tsub_le_iff_left, add_assoc, ←tsub_le_iff_left, ←tsub_le_iff_left], - refine (tsub_le_tsub_right add_tsub_le_assoc c).trans _, - rw [add_comm a, add_comm (a - c)], - exact add_tsub_le_assoc, -end - -/-- See `add_tsub_add_eq_tsub_left` for the equality. -/ -lemma add_tsub_add_le_tsub_left : a + b - (a + c) ≤ b - c := -by { rw [tsub_le_iff_left, add_assoc], exact add_le_add_left le_add_tsub _ } - -/-- See `add_tsub_add_eq_tsub_right` for the equality. -/ -lemma add_tsub_add_le_tsub_right : a + c - (b + c) ≤ a - b := -by { rw [tsub_le_iff_left, add_right_comm], exact add_le_add_right le_add_tsub c } - -end cov - -/-! #### Lemmas that assume that an element is `add_le_cancellable` -/ - -namespace add_le_cancellable - -protected lemma le_add_tsub_swap (hb : add_le_cancellable b) : a ≤ b + a - b := hb le_add_tsub - -protected lemma le_add_tsub (hb : add_le_cancellable b) : a ≤ a + b - b := -by { rw add_comm, exact hb.le_add_tsub_swap } - -protected lemma le_tsub_of_add_le_left (ha : add_le_cancellable a) (h : a + b ≤ c) : b ≤ c - a := -ha $ h.trans le_add_tsub - -protected lemma le_tsub_of_add_le_right (hb : add_le_cancellable b) (h : a + b ≤ c) : a ≤ c - b := -hb.le_tsub_of_add_le_left $ by rwa add_comm - -end add_le_cancellable - -/-! ### Lemmas where addition is order-reflecting -/ - -section contra -variable [contravariant_class α α (+) (≤)] - -lemma le_add_tsub_swap : a ≤ b + a - b := contravariant.add_le_cancellable.le_add_tsub_swap - -lemma le_add_tsub' : a ≤ a + b - b := contravariant.add_le_cancellable.le_add_tsub - -lemma le_tsub_of_add_le_left (h : a + b ≤ c) : b ≤ c - a := -contravariant.add_le_cancellable.le_tsub_of_add_le_left h - -lemma le_tsub_of_add_le_right (h : a + b ≤ c) : a ≤ c - b := -contravariant.add_le_cancellable.le_tsub_of_add_le_right h - -end contra - -end add_comm_semigroup - -variables [add_comm_monoid α] [has_sub α] [has_ordered_sub α] {a b c d : α} -lemma tsub_nonpos : a - b ≤ 0 ↔ a ≤ b := by rw [tsub_le_iff_left, add_zero] - -alias tsub_nonpos ↔ _ tsub_nonpos_of_le - -lemma add_monoid_hom.le_map_tsub [preorder β] [add_comm_monoid β] [has_sub β] - [has_ordered_sub β] (f : α →+ β) (hf : monotone f) (a b : α) : - f a - f b ≤ f (a - b) := -f.to_add_hom.le_map_tsub hf a b - -end preorder - -/-! ### Partial order -/ - -variables [partial_order α] [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] {a b c d : α} - -lemma tsub_tsub (b a c : α) : b - a - c = b - (a + c) := -begin - apply le_antisymm, - { rw [tsub_le_iff_left, tsub_le_iff_left, ← add_assoc, ← tsub_le_iff_left] }, - { rw [tsub_le_iff_left, add_assoc, ← tsub_le_iff_left, ← tsub_le_iff_left] } -end - -lemma tsub_add_eq_tsub_tsub (a b c : α) : a - (b + c) = a - b - c := (tsub_tsub _ _ _).symm - -lemma tsub_add_eq_tsub_tsub_swap (a b c : α) : a - (b + c) = a - c - b := -by { rw [add_comm], apply tsub_add_eq_tsub_tsub } - -lemma tsub_right_comm : a - b - c = a - c - b := -by simp_rw [← tsub_add_eq_tsub_tsub, add_comm] - -/-! ### Lemmas that assume that an element is `add_le_cancellable`. -/ - -namespace add_le_cancellable - -protected lemma tsub_eq_of_eq_add (hb : add_le_cancellable b) (h : a = c + b) : a - b = c := -le_antisymm (tsub_le_iff_right.mpr h.le) $ - by { rw h, exact hb.le_add_tsub } - -protected lemma eq_tsub_of_add_eq (hc : add_le_cancellable c) (h : a + c = b) : a = b - c := -(hc.tsub_eq_of_eq_add h.symm).symm - -protected theorem tsub_eq_of_eq_add_rev (hb : add_le_cancellable b) (h : a = b + c) : a - b = c := -hb.tsub_eq_of_eq_add $ by rw [add_comm, h] - -@[simp] -protected lemma add_tsub_cancel_right (hb : add_le_cancellable b) : a + b - b = a := -hb.tsub_eq_of_eq_add $ by rw [add_comm] - -@[simp] -protected lemma add_tsub_cancel_left (ha : add_le_cancellable a) : a + b - a = b := -ha.tsub_eq_of_eq_add $ add_comm a b - -protected lemma lt_add_of_tsub_lt_left (hb : add_le_cancellable b) (h : a - b < c) : a < b + c := -begin - rw [lt_iff_le_and_ne, ← tsub_le_iff_left], - refine ⟨h.le, _⟩, - rintro rfl, - simpa [hb] using h, -end - -protected lemma lt_add_of_tsub_lt_right (hc : add_le_cancellable c) (h : a - c < b) : a < b + c := -begin - rw [lt_iff_le_and_ne, ← tsub_le_iff_right], - refine ⟨h.le, _⟩, - rintro rfl, - simpa [hc] using h, -end - -protected lemma lt_tsub_of_add_lt_right (hc : add_le_cancellable c) (h : a + c < b) : a < b - c := -(hc.le_tsub_of_add_le_right h.le).lt_of_ne $ by { rintro rfl, exact h.not_le le_tsub_add } - -protected lemma lt_tsub_of_add_lt_left (ha : add_le_cancellable a) (h : a + c < b) : c < b - a := -ha.lt_tsub_of_add_lt_right $ by rwa add_comm - -end add_le_cancellable - -/-! #### Lemmas where addition is order-reflecting. -/ - -section contra -variable [contravariant_class α α (+) (≤)] - -lemma tsub_eq_of_eq_add (h : a = c + b) : a - b = c := -contravariant.add_le_cancellable.tsub_eq_of_eq_add h - -lemma eq_tsub_of_add_eq (h : a + c = b) : a = b - c := -contravariant.add_le_cancellable.eq_tsub_of_add_eq h - -lemma tsub_eq_of_eq_add_rev (h : a = b + c) : a - b = c := -contravariant.add_le_cancellable.tsub_eq_of_eq_add_rev h - -@[simp] -lemma add_tsub_cancel_right (a b : α) : a + b - b = a := -contravariant.add_le_cancellable.add_tsub_cancel_right - -@[simp] -lemma add_tsub_cancel_left (a b : α) : a + b - a = b := -contravariant.add_le_cancellable.add_tsub_cancel_left - -lemma lt_add_of_tsub_lt_left (h : a - b < c) : a < b + c := -contravariant.add_le_cancellable.lt_add_of_tsub_lt_left h - -lemma lt_add_of_tsub_lt_right (h : a - c < b) : a < b + c := -contravariant.add_le_cancellable.lt_add_of_tsub_lt_right h - -/-- This lemma (and some of its corollaries) also holds for `ennreal`, but this proof doesn't work -for it. Maybe we should add this lemma as field to `has_ordered_sub`? -/ -lemma lt_tsub_of_add_lt_left : a + c < b → c < b - a := -contravariant.add_le_cancellable.lt_tsub_of_add_lt_left - -lemma lt_tsub_of_add_lt_right : a + c < b → a < b - c := -contravariant.add_le_cancellable.lt_tsub_of_add_lt_right - -end contra - -section both -variables [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤)] - -lemma add_tsub_add_eq_tsub_right (a c b : α) : (a + c) - (b + c) = a - b := -begin - refine add_tsub_add_le_tsub_right.antisymm (tsub_le_iff_right.2 $ le_of_add_le_add_right _), swap, - rw add_assoc, - exact le_tsub_add, -end - -lemma add_tsub_add_eq_tsub_left (a b c : α) : (a + b) - (a + c) = b - c := -by rw [add_comm a b, add_comm a c, add_tsub_add_eq_tsub_right] - -end both - -end ordered_add_comm_semigroup - -/-! ### Lemmas in a linearly ordered monoid. -/ -section linear_order -variables {a b c d : α} [linear_order α] [add_comm_semigroup α] [has_sub α] [has_ordered_sub α] - -/-- See `lt_of_tsub_lt_tsub_right_of_le` for a weaker statement in a partial order. -/ -lemma lt_of_tsub_lt_tsub_right (h : a - c < b - c) : a < b := -lt_imp_lt_of_le_imp_le (λ h, tsub_le_tsub_right h c) h - -/-- See `lt_tsub_iff_right_of_le` for a weaker statement in a partial order. -/ -lemma lt_tsub_iff_right : a < b - c ↔ a + c < b := -lt_iff_lt_of_le_iff_le tsub_le_iff_right - -/-- See `lt_tsub_iff_left_of_le` for a weaker statement in a partial order. -/ -lemma lt_tsub_iff_left : a < b - c ↔ c + a < b := -lt_iff_lt_of_le_iff_le tsub_le_iff_left - -lemma lt_tsub_comm : a < b - c ↔ c < b - a := -lt_tsub_iff_left.trans lt_tsub_iff_right.symm - -section cov -variable [covariant_class α α (+) (≤)] - -/-- See `lt_of_tsub_lt_tsub_left_of_le` for a weaker statement in a partial order. -/ -lemma lt_of_tsub_lt_tsub_left (h : a - b < a - c) : c < b := -lt_imp_lt_of_le_imp_le (λ h, tsub_le_tsub_left h a) h - -end cov - -end linear_order - -section ordered_add_comm_monoid -variables [partial_order α] [add_comm_monoid α] [has_sub α] [has_ordered_sub α] - -@[simp] lemma tsub_zero (a : α) : a - 0 = a := -add_le_cancellable.tsub_eq_of_eq_add add_le_cancellable_zero (add_zero _).symm - -end ordered_add_comm_monoid +variables {α : Type*} section has_exists_add_of_le variables [add_comm_semigroup α] [partial_order α] [has_exists_add_of_le α] @@ -788,36 +408,3 @@ lemma tsub_add_min : a - b + min a b = a := by { rw [← tsub_min, tsub_add_cancel_of_le], apply min_le_left } end canonically_linear_ordered_add_monoid - -namespace with_top - -section -variables [has_sub α] [has_zero α] - -/-- If `α` has subtraction and `0`, we can extend the subtraction to `with_top α`. -/ -protected def sub : Π (a b : with_top α), with_top α -| _ ⊤ := 0 -| ⊤ (x : α) := ⊤ -| (x : α) (y : α) := (x - y : α) - -instance : has_sub (with_top α) := -⟨with_top.sub⟩ - -@[simp, norm_cast] lemma coe_sub {a b : α} : (↑(a - b) : with_top α) = ↑a - ↑b := rfl -@[simp] lemma top_sub_coe {a : α} : (⊤ : with_top α) - a = ⊤ := rfl -@[simp] lemma sub_top {a : with_top α} : a - ⊤ = 0 := by { cases a; refl } - -end - -variables [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] -instance : has_ordered_sub (with_top α) := -begin - constructor, - rintro x y z, - induction y using with_top.rec_top_coe, { simp }, - induction x using with_top.rec_top_coe, { simp }, - induction z using with_top.rec_top_coe, { simp }, - norm_cast, exact tsub_le_iff_right -end - -end with_top diff --git a/src/algebra/order/sub/with_top.lean b/src/algebra/order/sub/with_top.lean new file mode 100644 index 0000000000000..e93225209b9c6 --- /dev/null +++ b/src/algebra/order/sub/with_top.lean @@ -0,0 +1,46 @@ +/- +Copyright (c) 2021 Floris van Doorn. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn +-/ +import algebra.order.sub.basic +import algebra.order.monoid.with_top + +/-! +# Lemma about subtraction in ordered monoids with a top element adjoined. +-/ + +variables {α : Type*} + +namespace with_top + +section +variables [has_sub α] [has_zero α] + +/-- If `α` has subtraction and `0`, we can extend the subtraction to `with_top α`. -/ +protected def sub : Π (a b : with_top α), with_top α +| _ ⊤ := 0 +| ⊤ (x : α) := ⊤ +| (x : α) (y : α) := (x - y : α) + +instance : has_sub (with_top α) := +⟨with_top.sub⟩ + +@[simp, norm_cast] lemma coe_sub {a b : α} : (↑(a - b) : with_top α) = ↑a - ↑b := rfl +@[simp] lemma top_sub_coe {a : α} : (⊤ : with_top α) - a = ⊤ := rfl +@[simp] lemma sub_top {a : with_top α} : a - ⊤ = 0 := by { cases a; refl } + +end + +variables [canonically_ordered_add_monoid α] [has_sub α] [has_ordered_sub α] +instance : has_ordered_sub (with_top α) := +begin + constructor, + rintro x y z, + induction y using with_top.rec_top_coe, { simp }, + induction x using with_top.rec_top_coe, { simp }, + induction z using with_top.rec_top_coe, { simp }, + norm_cast, exact tsub_le_iff_right +end + +end with_top diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index 0173838fee2d8..306173382fae1 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Johan Commelin, Patrick Massot -/ -import algebra.order.group +import algebra.order.group.type_tags import algebra.order.monoid.with_zero /-! diff --git a/src/data/nat/enat.lean b/src/data/nat/enat.lean index b46f12583089a..909cdbca77d30 100644 --- a/src/data/nat/enat.lean +++ b/src/data/nat/enat.lean @@ -5,6 +5,7 @@ Authors: Yury Kudryashov -/ import data.nat.lattice import data.nat.succ_pred +import algebra.order.sub.with_top /-! # Definition and basic properties of extended natural numbers diff --git a/src/data/nat/with_bot.lean b/src/data/nat/with_bot.lean index 46172cf2b3647..d1bf69d5de077 100644 --- a/src/data/nat/with_bot.lean +++ b/src/data/nat/with_bot.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import data.nat.basic -import algebra.order.group +import algebra.order.group.basic /-! # `with_bot ℕ` diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 3aa0589e4c1a6..d4e164a2cea8c 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ import data.real.nnreal +import algebra.order.sub.with_top /-! # Extended non-negative reals diff --git a/src/data/set/intervals/basic.lean b/src/data/set/intervals/basic.lean index 7f8ead612f168..d5b3fc664d2dd 100644 --- a/src/data/set/intervals/basic.lean +++ b/src/data/set/intervals/basic.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, Patrick Massot, Yury Kudryashov, Rémy Degenne -/ -import algebra.order.group +import algebra.order.group.basic import order.rel_iso /-! From 9644d9683398c98f703652b9cb1889ae0b71607c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Thu, 13 Oct 2022 11:47:29 +0000 Subject: [PATCH 746/828] feat(algebra/group_with_zero/units): add ring.inverse lemmas (#16889) --- src/algebra/group_with_zero/units.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/algebra/group_with_zero/units.lean b/src/algebra/group_with_zero/units.lean index 00589fb05a5f1..b075720ba2f9b 100644 --- a/src/algebra/group_with_zero/units.lean +++ b/src/algebra/group_with_zero/units.lean @@ -97,6 +97,15 @@ by rw [← mul_assoc, mul_inverse_cancel x h, one_mul] lemma inverse_mul_cancel_left (x y : M₀) (h : is_unit x) : inverse x * (x * y) = y := by rw [← mul_assoc, inverse_mul_cancel x h, one_mul] +lemma inverse_mul_eq_iff_eq_mul (x y z : M₀) (h : is_unit x) : + inverse x * y = z ↔ y = x * z := +⟨λ h1, by rw [← h1, mul_inverse_cancel_left _ _ h], λ h1, by rw [h1, inverse_mul_cancel_left _ _ h]⟩ + +lemma eq_mul_inverse_iff_mul_eq (x y z : M₀) (h : is_unit z) : + x = y * inverse z ↔ x * z = y := +⟨λ h1, by rw [h1, inverse_mul_cancel_right _ _ h], + λ h1, by rw [← h1, mul_inverse_cancel_right _ _ h]⟩ + variables (M₀) @[simp] lemma inverse_one : inverse (1 : M₀) = 1 := From 5147123efc3d090ffc28eca0165ffe63ae74b428 Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Thu, 13 Oct 2022 11:47:30 +0000 Subject: [PATCH 747/828] feat(ring_theory/dedekind_domain/ideal): height_one_spectrum is equivalent to maximal_spectrum (#16920) - [x] depends on: #16905 [define maximal spectrum] --- src/ring_theory/dedekind_domain/ideal.lean | 43 ++++++++++++---------- 1 file changed, 24 insertions(+), 19 deletions(-) diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index ebad5844b9405..c905a1aad74ad 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenji Nakagawa, Anne Baanen, Filippo A. E. Nuccio -/ import algebra.algebra.subalgebra.pointwise +import algebraic_geometry.prime_spectrum.maximal import algebraic_geometry.prime_spectrum.noetherian import order.hom.basic import ring_theory.dedekind_domain.basic @@ -830,14 +831,13 @@ end end is_dedekind_domain -section height_one_spectrum - /-! ### Height one spectrum of a Dedekind domain If `R` is a Dedekind domain of Krull dimension 1, the maximal ideals of `R` are exactly its nonzero prime ideals. We define `height_one_spectrum` and provide lemmas to recover the facts that prime ideals of height -one are prime and irreducible. -/ +one are prime and irreducible. +-/ namespace is_dedekind_domain @@ -849,31 +849,36 @@ variables [is_domain R] [is_dedekind_domain R] structure height_one_spectrum := (as_ideal : ideal R) (is_prime : as_ideal.is_prime) -(ne_bot : as_ideal ≠ ⊥) +(ne_bot : as_ideal ≠ ⊥) + +attribute [instance] height_one_spectrum.is_prime variables (v : height_one_spectrum R) {R} -lemma height_one_spectrum.prime (v : height_one_spectrum R) : prime v.as_ideal := -ideal.prime_of_is_prime v.ne_bot v.is_prime +namespace height_one_spectrum -lemma height_one_spectrum.irreducible (v : height_one_spectrum R) : - irreducible v.as_ideal := -begin - rw [unique_factorization_monoid.irreducible_iff_prime], - apply v.prime, -end +instance is_maximal : v.as_ideal.is_maximal := dimension_le_one v.as_ideal v.ne_bot v.is_prime -lemma height_one_spectrum.associates_irreducible (v : height_one_spectrum R) : - irreducible (associates.mk v.as_ideal) := -begin - rw [associates.irreducible_mk _], - apply v.irreducible, -end +lemma prime : prime v.as_ideal := ideal.prime_of_is_prime v.ne_bot v.is_prime -end is_dedekind_domain +lemma irreducible : irreducible v.as_ideal := +unique_factorization_monoid.irreducible_iff_prime.mpr v.prime + +lemma associates_irreducible : _root_.irreducible $ associates.mk v.as_ideal := +(associates.irreducible_mk _).mpr v.irreducible + +/-- An equivalence between the height one and maximal spectra for rings of Krull dimension 1. -/ +def equiv_maximal_spectrum (hR : ¬is_field R) : height_one_spectrum R ≃ maximal_spectrum R := +{ to_fun := λ v, ⟨v.as_ideal, dimension_le_one v.as_ideal v.ne_bot v.is_prime⟩, + inv_fun := λ v, + ⟨v.as_ideal, v.is_maximal.is_prime, ring.ne_bot_of_is_maximal_of_not_is_field v.is_maximal hR⟩, + left_inv := λ ⟨_, _, _⟩, rfl, + right_inv := λ ⟨_, _⟩, rfl } end height_one_spectrum +end is_dedekind_domain + section open ideal From 297610838471f6ea3368bf26d2642e63a159fbcf Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Thu, 13 Oct 2022 11:47:31 +0000 Subject: [PATCH 748/828] chore(number_theory/number_field/basic): fix a name (#16943) This lemma is in the `ring_of_integers` namespace, so the name was redundant. --- src/number_theory/number_field/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/number_theory/number_field/basic.lean b/src/number_theory/number_field/basic.lean index f7dca7a2d261e..5fc4093435653 100644 --- a/src/number_theory/number_field/basic.lean +++ b/src/number_theory/number_field/basic.lean @@ -102,7 +102,7 @@ integral_closure.is_integrally_closed_of_finite_extension ℚ lemma is_integral_coe (x : 𝓞 K) : is_integral ℤ (x : K) := x.2 -lemma map_mem_ring_of_integers {F L : Type*} [field L] [char_zero K] [char_zero L] +lemma map_mem {F L : Type*} [field L] [char_zero K] [char_zero L] [alg_hom_class F ℚ K L] (f : F) (x : 𝓞 K) : f x ∈ 𝓞 L := (mem_ring_of_integers _ _).2 $ map_is_integral_int f $ ring_of_integers.is_integral_coe x From 91b5bdde5c5078b74a44c52e886f9c5904ce6590 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 13 Oct 2022 15:02:23 +0000 Subject: [PATCH 749/828] feat(analysis/normed_space/lp_equiv): equivalences between `lp` and `pi_Lp`, `bounded_continuous_function` (#15872) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This adds a new file for equivalences between various L^p spaces. We do this in a new file instead of `analysis/normed_space/lp_space` in order to minimize imports. We begin by establishing the equivalence between `lp` and `pi_Lp` when the index type is a fintype, and then proceed to recognize the equivalence between `lp` (for `p = ∞`) and `bounded_continuous_function` when the codomain has various algebraic structures. - [x] depends on: #15833 - [x] depends on: #15852 --- src/analysis/normed_space/lp_equiv.lean | 172 ++++++++++++++++++++++++ src/analysis/normed_space/lp_space.lean | 5 - 2 files changed, 172 insertions(+), 5 deletions(-) create mode 100644 src/analysis/normed_space/lp_equiv.lean diff --git a/src/analysis/normed_space/lp_equiv.lean b/src/analysis/normed_space/lp_equiv.lean new file mode 100644 index 0000000000000..8a761c9788197 --- /dev/null +++ b/src/analysis/normed_space/lp_equiv.lean @@ -0,0 +1,172 @@ +/- +Copyright (c) 2022 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import analysis.normed_space.lp_space +import analysis.normed_space.pi_Lp +import topology.continuous_function.bounded + +/-! +# Equivalences among $$L^p$$ spaces + +In this file we collect a variety of equivalences among various $$L^p$$ spaces. In particular, +when `α` is a `fintype`, given `E : α → Type u` and `p : ℝ≥0∞`, there is a natural linear isometric +equivalence `lp_pi_Lpₗᵢ : lp E p ≃ₗᵢ pi_Lp p E`. In addition, when `α` is a discrete topological +space, the bounded continuous functions `α →ᵇ β` correspond exactly to `lp (λ _, β) ∞`. Here there +can be more structure, including ring and algebra structures, and we implement these equivalences +accordingly as well. + +We keep this as a separate file so that the various $$L^p$$ space files don't import the others. + +Recall that `pi_Lp` is just a type synonym for `Π i, E i` but given a different metric and norm +structure, although the topological, uniform and bornological structures coincide definitionally. +These structures are only defined on `pi_Lp` for `fintype α`, so there are no issues of convergence +to consider. + +While `pre_lp` is also a type synonym for `Π i, E i`, it allows for infinite index types. On this +type there is a predicate `mem_ℓp` which says that the relevant `p`-norm is finite and `lp E p` is +the subtype of `pre_lp` satisfying `mem_ℓp`. + +## TODO + +* Equivalence between `lp` and `measure_theory.Lp`, for `f : α → E` (i.e., functions rather than + pi-types) and the counting measure on `α` + +-/ + +open_locale ennreal + +section lp_pi_Lp + +variables {α : Type*} {E : α → Type*} [Π i, normed_add_comm_group (E i)] {p : ℝ≥0∞} + +/-- When `α` is `finite`, every `f : pre_lp E p` satisfies `mem_ℓp f p`. -/ +lemma mem_ℓp.all [finite α] (f : Π i, E i) : mem_ℓp f p := +begin + rcases p.trichotomy with (rfl | rfl | h), + { exact mem_ℓp_zero_iff.mpr {i : α | f i ≠ 0}.to_finite, }, + { exact mem_ℓp_infty_iff.mpr (set.finite.bdd_above (set.range (λ (i : α), ∥f i∥)).to_finite) }, + { casesI nonempty_fintype α, exact mem_ℓp_gen ⟨finset.univ.sum _, has_sum_fintype _⟩ } +end + +variables [fintype α] + +/-- The canonical `equiv` between `lp E p ≃ pi_Lp p E` when `E : α → Type u` with `[fintype α]`. -/ +def equiv.lp_pi_Lp : lp E p ≃ pi_Lp p E := +{ to_fun := λ f, f, + inv_fun := λ f, ⟨f, mem_ℓp.all f⟩, + left_inv := λ f, lp.ext $ funext $ λ x, rfl, + right_inv := λ f, funext $ λ x, rfl } + +lemma coe_equiv_lp_pi_Lp (f : lp E p) : equiv.lp_pi_Lp f = f := rfl +lemma coe_equiv_lp_pi_Lp_symm (f : pi_Lp p E) : (equiv.lp_pi_Lp.symm f : Π i, E i) = f := rfl + +lemma equiv_lp_pi_Lp_norm (f : lp E p) : ∥equiv.lp_pi_Lp f∥ = ∥f∥ := +begin + unfreezingI { rcases p.trichotomy with (rfl | rfl | h) }, + { rw [pi_Lp.norm_eq_card, lp.norm_eq_card_dsupport], refl }, + { rw [pi_Lp.norm_eq_csupr, lp.norm_eq_csupr], refl }, + { rw [pi_Lp.norm_eq_sum h, lp.norm_eq_tsum_rpow h, tsum_fintype], refl }, +end + +/-- The canonical `add_equiv` between `lp E p` and `pi_Lp p E` when `E : α → Type u` with +`[fintype α]` and `[fact (1 ≤ p)]`. -/ +def add_equiv.lp_pi_Lp [fact (1 ≤ p)] : lp E p ≃+ pi_Lp p E := +{ map_add' := λ f g, rfl, + .. equiv.lp_pi_Lp } + +lemma coe_add_equiv_lp_pi_Lp [fact (1 ≤ p)] (f : lp E p) : + add_equiv.lp_pi_Lp f = f := rfl +lemma coe_add_equiv_lp_pi_Lp_symm [fact (1 ≤ p)] (f : pi_Lp p E) : + (add_equiv.lp_pi_Lp.symm f : Π i, E i) = f := rfl + +section equivₗᵢ +variables (𝕜 : Type*) [nontrivially_normed_field 𝕜] [Π i, normed_space 𝕜 (E i)] + +/-- The canonical `linear_isometry_equiv` between `lp E p` and `pi_Lp p E` when `E : α → Type u` +with `[fintype α]` and `[fact (1 ≤ p)]`. -/ +noncomputable def lp_pi_Lpₗᵢ [fact (1 ≤ p)] : lp E p ≃ₗᵢ[𝕜] pi_Lp p E := +{ map_smul' := λ k f, rfl, + norm_map' := equiv_lp_pi_Lp_norm, + .. add_equiv.lp_pi_Lp } + +variables {𝕜} + +lemma coe_lp_pi_Lpₗᵢ [fact (1 ≤ p)] (f : lp E p) : + lp_pi_Lpₗᵢ 𝕜 f = f := rfl +lemma coe_lp_pi_Lpₗᵢ_symm [fact (1 ≤ p)] (f : pi_Lp p E) : + ((lp_pi_Lpₗᵢ 𝕜).symm f : Π i, E i) = f := rfl + +end equivₗᵢ + +end lp_pi_Lp + +section lp_bcf + +open_locale bounded_continuous_function +open bounded_continuous_function + +-- note: `R` and `A` are explicit because otherwise Lean has elaboration problems +variables {α E : Type*} (R A 𝕜 : Type*) [topological_space α] [discrete_topology α] +variables [normed_ring A] [norm_one_class A] [nontrivially_normed_field 𝕜] [normed_algebra 𝕜 A] +variables [normed_add_comm_group E] [normed_space 𝕜 E] [non_unital_normed_ring R] + + +section normed_add_comm_group + +/-- The canonical map between `lp (λ (_ : α), E) ∞` and `α →ᵇ E` as an `add_equiv`. -/ +noncomputable def add_equiv.lp_bcf : + lp (λ (_ : α), E) ∞ ≃+ (α →ᵇ E) := +{ to_fun := λ f, of_normed_add_comm_group_discrete f (∥f∥) $ le_csupr (mem_ℓp_infty_iff.mp f.prop), + inv_fun := λ f, ⟨f, f.bdd_above_range_norm_comp⟩, + left_inv := λ f, lp.ext rfl, + right_inv := λ f, ext $ λ x, rfl, + map_add' := λ f g, ext $ λ x, rfl } + +lemma coe_add_equiv_lp_bcf (f : lp (λ (_ : α), E) ∞) : + (add_equiv.lp_bcf f : α → E) = f := rfl +lemma coe_add_equiv_lp_bcf_symm (f : α →ᵇ E) : (add_equiv.lp_bcf.symm f : α → E) = f := rfl + +/-- The canonical map between `lp (λ (_ : α), E) ∞` and `α →ᵇ E` as a `linear_isometry_equiv`. -/ +noncomputable def lp_bcfₗᵢ : lp (λ (_ : α), E) ∞ ≃ₗᵢ[𝕜] (α →ᵇ E) := +{ map_smul' := λ k f, rfl, + norm_map' := λ f, by { simp only [norm_eq_supr_norm, lp.norm_eq_csupr], refl }, + .. add_equiv.lp_bcf } + +variables {𝕜} + +lemma coe_lp_bcfₗᵢ (f : lp (λ (_ : α), E) ∞) : (lp_bcfₗᵢ 𝕜 f : α → E) = f := rfl +lemma coe_lp_bcfₗᵢ_symm (f : α →ᵇ E) : ((lp_bcfₗᵢ 𝕜).symm f : α → E) = f := rfl + +end normed_add_comm_group + +section ring_algebra + +/-- The canonical map between `lp (λ (_ : α), R) ∞` and `α →ᵇ R` as a `ring_equiv`. -/ +noncomputable def ring_equiv.lp_bcf : lp (λ (_ : α), R) ∞ ≃+* (α →ᵇ R) := +{ map_mul' := λ f g, ext $ λ x, rfl, .. @add_equiv.lp_bcf _ R _ _ _ } + +variables {R} +lemma coe_ring_equiv_lp_bcf (f : lp (λ (_ : α), R) ∞) : + (ring_equiv.lp_bcf R f : α → R) = f := rfl +lemma coe_ring_equiv_lp_bcf_symm (f : α →ᵇ R) : + ((ring_equiv.lp_bcf R).symm f : α → R) = f := rfl + +variables (α) -- even `α` needs to be explicit here for elaboration + +-- the `norm_one_class A` shouldn't really be necessary, but currently it is for +-- `one_mem_ℓp_infty` to get the `ring` instance on `lp`. +/-- The canonical map between `lp (λ (_ : α), A) ∞` and `α →ᵇ A` as an `alg_equiv`. -/ +noncomputable def alg_equiv.lp_bcf : lp (λ (_ : α), A) ∞ ≃ₐ[𝕜] (α →ᵇ A) := +{ commutes' := λ k, rfl, .. ring_equiv.lp_bcf A } + +variables {α A 𝕜} +lemma coe_alg_equiv_lp_bcf (f : lp (λ (_ : α), A) ∞) : + (alg_equiv.lp_bcf α A 𝕜 f : α → A) = f := rfl +lemma coe_alg_equiv_lp_bcf_symm (f : α →ᵇ A) : + ((alg_equiv.lp_bcf α A 𝕜).symm f : α → A) = f := rfl + +end ring_algebra + +end lp_bcf diff --git a/src/analysis/normed_space/lp_space.lean b/src/analysis/normed_space/lp_space.lean index 6fe3fb7086574..74f5523c0a237 100644 --- a/src/analysis/normed_space/lp_space.lean +++ b/src/analysis/normed_space/lp_space.lean @@ -51,11 +51,6 @@ say that `∥-f∥ = ∥f∥`, instead of the non-working `f.norm_neg`. * More versions of Hölder's inequality (for example: the case `p = 1`, `q = ∞`; a version for normed rings which has `∥∑' i, f i * g i∥` rather than `∑' i, ∥f i∥ * g i∥` on the RHS; a version for three exponents satisfying `1 / r = 1 / p + 1 / q`) -* Equivalence with `pi_Lp`, for `α` finite -* Equivalence with `measure_theory.Lp`, for `f : α → E` (i.e., functions rather than pi-types) and - the counting measure on `α` -* Equivalence with `bounded_continuous_function`, for `f : α → E` (i.e., functions rather than - pi-types) and `p = ∞`, and the discrete topology on `α` -/ From 414df82235c39247fe17bcb7349955348665bb1e Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 13 Oct 2022 15:02:24 +0000 Subject: [PATCH 750/828] feat(analysis/seminorm): seminorms are a `conditionally_complete_lattice` (#16582) --- src/analysis/seminorm.lean | 103 ++++++++++++++- src/order/conditionally_complete_lattice.lean | 120 ++++++++++++++++++ 2 files changed, 221 insertions(+), 2 deletions(-) diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index c10ddad892510..d55a5221d9ae7 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -192,8 +192,6 @@ instance [semiring R] [module R ℝ] [has_smul R ℝ≥0] [is_scalar_tower R ℝ module R (seminorm 𝕜 E) := (coe_fn_add_monoid_hom_injective 𝕜 E).module R _ coe_smul --- TODO: define `has_Sup` too, from the skeleton at --- https://github.com/leanprover-community/mathlib/pull/11329#issuecomment-1008915345 instance : has_sup (seminorm 𝕜 E) := { sup := λ p q, { to_fun := p ⊔ q, @@ -385,6 +383,107 @@ begin smul_eq_mul, real.mul_infi_of_nonneg (subtype.prop _), mul_add], end +section classical + +open_locale classical + +/-- We define the supremum of an arbitrary subset of `seminorm 𝕜 E` as follows: +* if `s` is `bdd_above` *as a set of functions `E → ℝ`* (that is, if `s` is pointwise bounded +above), we take the pointwise supremum of all elements of `s`, and we prove that it is indeed a +seminorm. +* otherwise, we take the zero seminorm `⊥`. + +There are two things worth mentionning here: +* First, it is not trivial at first that `s` being bounded above *by a function* implies +being bounded above *as a seminorm*. We show this in `seminorm.bdd_above_iff` by using +that the `Sup s` as defined here is then a bounding seminorm for `s`. So it is important to make +the case disjunction on `bdd_above (coe_fn '' s : set (E → ℝ))` and not `bdd_above s`. +* Since the pointwise `Sup` already gives `0` at points where a family of functions is +not bounded above, one could hope that just using the pointwise `Sup` would work here, without the +need for an additional case disjunction. As discussed on Zulip, this doesn't work because this can +give a function which does *not* satisfy the seminorm axioms (typically sub-additivity). +-/ +noncomputable instance : has_Sup (seminorm 𝕜 E) := +{ Sup := λ s, if h : bdd_above (coe_fn '' s : set (E → ℝ)) then + { to_fun := ⨆ p : s, ((p : seminorm 𝕜 E) : E → ℝ), + map_zero' := + begin + rw [supr_apply, ← @real.csupr_const_zero s], + congrm ⨆ i, _, + exact map_zero i.1 + end, + add_le' := λ x y, + begin + rcases h with ⟨q, hq⟩, + obtain rfl | h := s.eq_empty_or_nonempty, + { simp [real.csupr_empty] }, + haveI : nonempty ↥s := nonempty_coe_sort.mpr h, + simp only [supr_apply], + refine csupr_le (λ i, ((i : seminorm 𝕜 E).add_le' x y).trans $ + add_le_add (le_csupr ⟨q x, _⟩ i) (le_csupr ⟨q y, _⟩ i)); + rw [mem_upper_bounds, forall_range_iff]; + exact λ j, hq (mem_image_of_mem _ j.2) _, + end, + neg' := λ x, + begin + simp only [supr_apply], + congrm ⨆ i, _, + exact i.1.neg' _ + end, + smul' := λ a x, + begin + simp only [supr_apply], + rw [← smul_eq_mul, real.smul_supr_of_nonneg (norm_nonneg a) (λ i : s, (i : seminorm 𝕜 E) x)], + congrm ⨆ i, _, + exact i.1.smul' a x + end } + else ⊥ } + +protected lemma coe_Sup_eq' {s : set $ seminorm 𝕜 E} (hs : bdd_above (coe_fn '' s : set (E → ℝ))) : + coe_fn (Sup s) = ⨆ p : s, p := +congr_arg _ (dif_pos hs) + +protected lemma bdd_above_iff {s : set $ seminorm 𝕜 E} : + bdd_above s ↔ bdd_above (coe_fn '' s : set (E → ℝ)) := +⟨λ ⟨q, hq⟩, ⟨q, ball_image_of_ball $ λ p hp, hq hp⟩, + λ H, ⟨Sup s, λ p hp x, + begin + rw [seminorm.coe_Sup_eq' H, supr_apply], + rcases H with ⟨q, hq⟩, + exact le_csupr ⟨q x, forall_range_iff.mpr $ λ i : s, hq (mem_image_of_mem _ i.2) x⟩ ⟨p, hp⟩ + end ⟩⟩ + +protected lemma coe_Sup_eq {s : set $ seminorm 𝕜 E} (hs : bdd_above s) : + coe_fn (Sup s) = ⨆ p : s, p := +seminorm.coe_Sup_eq' (seminorm.bdd_above_iff.mp hs) + +protected lemma coe_supr_eq {ι : Type*} {p : ι → seminorm 𝕜 E} (hp : bdd_above (range p)) : + coe_fn (⨆ i, p i) = ⨆ i, p i := +by rw [← Sup_range, seminorm.coe_Sup_eq hp]; exact supr_range' (coe_fn : seminorm 𝕜 E → E → ℝ) p + +private lemma seminorm.is_lub_Sup (s : set (seminorm 𝕜 E)) (hs₁ : bdd_above s) (hs₂ : s.nonempty) : + is_lub s (Sup s) := +begin + refine ⟨λ p hp x, _, λ p hp x, _⟩; + haveI : nonempty ↥s := nonempty_coe_sort.mpr hs₂; + rw [seminorm.coe_Sup_eq hs₁, supr_apply], + { rcases hs₁ with ⟨q, hq⟩, + exact le_csupr ⟨q x, forall_range_iff.mpr $ λ i : s, hq i.2 x⟩ ⟨p, hp⟩ }, + { exact csupr_le (λ q, hp q.2 x) } +end + +/-- `seminorm 𝕜 E` is a conditionally complete lattice. + +Note that, while `inf`, `sup` and `Sup` have good definitional properties (corresponding to +`seminorm.has_inf`, `seminorm.has_sup` and `seminorm.has_Sup` respectively), `Inf s` is just +defined as the supremum of the lower bounds of `s`, which is not really useful in practice. If you +need to use `Inf` on seminorms, then you should probably provide a more workable definition first, +but this is unlikely to happen so we keep the "bad" definition for now. -/ +noncomputable instance : conditionally_complete_lattice (seminorm 𝕜 E) := +conditionally_complete_lattice_of_lattice_of_Sup (seminorm 𝕜 E) seminorm.is_lub_Sup + +end classical + end normed_field /-! ### Seminorm ball -/ diff --git a/src/order/conditionally_complete_lattice.lean b/src/order/conditionally_complete_lattice.lean index 679c79d02b05b..c345094f16e9a 100644 --- a/src/order/conditionally_complete_lattice.lean +++ b/src/order/conditionally_complete_lattice.lean @@ -235,6 +235,126 @@ instance (α : Type*) [conditionally_complete_linear_order α] : end order_dual +/-- Create a `conditionally_complete_lattice` from a `partial_order` and `Sup` function +that returns the least upper bound of a nonempty set which is bounded above. Usually this +constructor provides poor definitional equalities. If other fields are known explicitly, they +should be provided; for example, if `inf` is known explicitly, construct the +`conditionally_complete_lattice` instance as +``` +instance : conditionally_complete_lattice my_T := +{ inf := better_inf, + le_inf := ..., + inf_le_right := ..., + inf_le_left := ... + -- don't care to fix sup, Inf + ..conditionally_complete_lattice_of_Sup my_T _ } +``` +-/ +def conditionally_complete_lattice_of_Sup (α : Type*) [H1 : partial_order α] + [H2 : has_Sup α] + (bdd_above_pair : ∀ a b : α, bdd_above ({a, b} : set α)) + (bdd_below_pair : ∀ a b : α, bdd_below ({a, b} : set α)) + (is_lub_Sup : ∀ s : set α, bdd_above s → s.nonempty → is_lub s (Sup s)) : + conditionally_complete_lattice α := +{ sup := λ a b, Sup {a, b}, + le_sup_left := λ a b, (is_lub_Sup {a, b} (bdd_above_pair a b) + (insert_nonempty _ _)).1 (mem_insert _ _), + le_sup_right := λ a b, (is_lub_Sup {a, b} (bdd_above_pair a b) + (insert_nonempty _ _)).1 (mem_insert_of_mem _ (mem_singleton _)), + sup_le := λ a b c hac hbc, (is_lub_Sup {a, b} (bdd_above_pair a b) + (insert_nonempty _ _)).2 (forall_insert_of_forall (forall_eq.mpr hbc) hac), + inf := λ a b, Sup (lower_bounds {a, b}), + inf_le_left := λ a b, (is_lub_Sup (lower_bounds {a, b}) + (nonempty.bdd_above_lower_bounds ⟨a, mem_insert _ _⟩) (bdd_below_pair a b)).2 + (λ c hc, hc $ mem_insert _ _), + inf_le_right := λ a b, (is_lub_Sup (lower_bounds {a, b}) + (nonempty.bdd_above_lower_bounds ⟨a, mem_insert _ _⟩) + (bdd_below_pair a b)).2 (λ c hc, hc $ mem_insert_of_mem _ (mem_singleton _)), + le_inf := λ c a b hca hcb, (is_lub_Sup (lower_bounds {a, b}) + (nonempty.bdd_above_lower_bounds ⟨a, mem_insert _ _⟩) + ⟨c, forall_insert_of_forall (forall_eq.mpr hcb) hca⟩).1 + (forall_insert_of_forall (forall_eq.mpr hcb) hca), + Inf := λ s, Sup (lower_bounds s), + cSup_le := λ s a hs ha, (is_lub_Sup s ⟨a, ha⟩ hs).2 ha, + le_cSup := λ s a hs ha, (is_lub_Sup s hs ⟨a, ha⟩).1 ha, + cInf_le := λ s a hs ha, (is_lub_Sup (lower_bounds s) + (nonempty.bdd_above_lower_bounds ⟨a, ha⟩) hs).2 (λ b hb, hb ha), + le_cInf := λ s a hs ha, (is_lub_Sup (lower_bounds s) hs.bdd_above_lower_bounds ⟨a, ha⟩).1 ha, + .. H1, .. H2 } + +/-- Create a `conditionally_complete_lattice_of_Inf` from a `partial_order` and `Inf` function +that returns the greatest lower bound of a nonempty set which is bounded below. Usually this +constructor provides poor definitional equalities. If other fields are known explicitly, they +should be provided; for example, if `inf` is known explicitly, construct the +`conditionally_complete_lattice` instance as +``` +instance : conditionally_complete_lattice my_T := +{ inf := better_inf, + le_inf := ..., + inf_le_right := ..., + inf_le_left := ... + -- don't care to fix sup, Sup + ..conditionally_complete_lattice_of_Inf my_T _ } +``` +-/ +def conditionally_complete_lattice_of_Inf (α : Type*) [H1 : partial_order α] + [H2 : has_Inf α] + (bdd_above_pair : ∀ a b : α, bdd_above ({a, b} : set α)) + (bdd_below_pair : ∀ a b : α, bdd_below ({a, b} : set α)) + (is_glb_Inf : ∀ s : set α, bdd_below s → s.nonempty → is_glb s (Inf s)) : + conditionally_complete_lattice α := +{ inf := λ a b, Inf {a, b}, + inf_le_left := λ a b, (is_glb_Inf {a, b} (bdd_below_pair a b) + (insert_nonempty _ _)).1 (mem_insert _ _), + inf_le_right := λ a b, (is_glb_Inf {a, b} (bdd_below_pair a b) + (insert_nonempty _ _)).1 (mem_insert_of_mem _ (mem_singleton _)), + le_inf := λ c a b hca hcb, (is_glb_Inf {a, b} (bdd_below_pair a b) + (insert_nonempty _ _)).2 (forall_insert_of_forall (forall_eq.mpr hcb) hca), + sup := λ a b, Inf (upper_bounds {a, b}), + le_sup_left := λ a b, (is_glb_Inf (upper_bounds {a, b}) + (nonempty.bdd_below_upper_bounds ⟨a, mem_insert _ _⟩) (bdd_above_pair a b)).2 + (λ c hc, hc $ mem_insert _ _), + le_sup_right := λ a b, (is_glb_Inf (upper_bounds {a, b}) + (nonempty.bdd_below_upper_bounds ⟨a, mem_insert _ _⟩) + (bdd_above_pair a b)).2 (λ c hc, hc $ mem_insert_of_mem _ (mem_singleton _)), + sup_le := λ a b c hac hbc, (is_glb_Inf (upper_bounds {a, b}) + (nonempty.bdd_below_upper_bounds ⟨a, mem_insert _ _⟩) + ⟨c, forall_insert_of_forall (forall_eq.mpr hbc) hac⟩).1 + (forall_insert_of_forall (forall_eq.mpr hbc) hac), + Sup := λ s, Inf (upper_bounds s), + le_cInf := λ s a hs ha, (is_glb_Inf s ⟨a, ha⟩ hs).2 ha, + cInf_le := λ s a hs ha, (is_glb_Inf s hs ⟨a, ha⟩).1 ha, + le_cSup := λ s a hs ha, (is_glb_Inf (upper_bounds s) + (nonempty.bdd_below_upper_bounds ⟨a, ha⟩) hs).2 (λ b hb, hb ha), + cSup_le := λ s a hs ha, (is_glb_Inf (upper_bounds s) hs.bdd_below_upper_bounds ⟨a, ha⟩).1 ha, + .. H1, .. H2 } + +/-- +A version of `conditionally_complete_lattice_of_Sup` when we already know that `α` is a lattice. + +This should only be used when it is both hard and unnecessary to provide `Inf` explicitly. -/ +def conditionally_complete_lattice_of_lattice_of_Sup (α : Type*) [H1 : lattice α] + [H2 : has_Sup α] + (is_lub_Sup : ∀ s : set α, bdd_above s → s.nonempty → is_lub s (Sup s)) : + conditionally_complete_lattice α := +{ ..H1, ..conditionally_complete_lattice_of_Sup α + (λ a b, ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) + (λ a b, ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) + is_lub_Sup } + +/-- +A version of `conditionally_complete_lattice_of_Inf` when we already know that `α` is a lattice. + +This should only be used when it is both hard and unnecessary to provide `Sup` explicitly. -/ +def conditionally_complete_lattice_of_lattice_of_Inf (α : Type*) [H1 : lattice α] + [H2 : has_Inf α] + (is_glb_Inf : ∀ s : set α, bdd_below s → s.nonempty → is_glb s (Inf s)) : + conditionally_complete_lattice α := +{ ..H1, ..conditionally_complete_lattice_of_Inf α + (λ a b, ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) + (λ a b, ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) + is_glb_Inf } + section conditionally_complete_lattice variables [conditionally_complete_lattice α] {s t : set α} {a b : α} From b68e4a899f41b91923e70404a43ce275da5ffa91 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 13 Oct 2022 15:02:26 +0000 Subject: [PATCH 751/828] refactor(topology/continuous_function/{algebra, compact}): move `continuous_map.comp_right_alg_hom` (#16875) This moves `continuous_function.comp_right_alg_hom` to a better home. Along the way, we generalize it significantly. We also provide term-mode proofs for some very slow `tidy`s, and remove some unnecessary hypotheses in `topology/continuous_function/algebra` --- src/topology/continuous_function/algebra.lean | 18 +++++++- src/topology/continuous_function/compact.lean | 43 +++++-------------- .../continuous_function/polynomial.lean | 2 +- .../continuous_function/weierstrass.lean | 2 +- 4 files changed, 29 insertions(+), 36 deletions(-) diff --git a/src/topology/continuous_function/algebra.lean b/src/topology/continuous_function/algebra.lean index f4ac6d39dbd4d..185f38f7625c6 100644 --- a/src/topology/continuous_function/algebra.lean +++ b/src/topology/continuous_function/algebra.lean @@ -575,8 +575,6 @@ def continuous_map.C : R →+* C(α, A) := @[simp] lemma continuous_map.C_apply (r : R) (a : α) : continuous_map.C r a = algebra_map R A r := rfl -variables [has_continuous_const_smul R A] [has_continuous_const_smul R A₂] - instance continuous_map.algebra : algebra R C(α, A) := { to_ring_hom := continuous_map.C, commutes' := λ c f, by ext x; exact algebra.commutes' _ _, @@ -592,6 +590,22 @@ variables (R) { commutes' := λ c, continuous_map.ext $ λ _, g.commutes' _, .. g.to_ring_hom.comp_left_continuous α hg } +variables (A) + +/-- +Precomposition of functions into a normed ring by a continuous map is an algebra homomorphism. +-/ +@[simps] def continuous_map.comp_right_alg_hom {α β : Type*} [topological_space α] + [topological_space β] (f : C(α, β)) : C(β, A) →ₐ[R] C(α, A) := +{ to_fun := λ g, g.comp f, + map_zero' := by { ext, refl, }, + map_add' := λ g₁ g₂, by { ext, refl, }, + map_one' := by { ext, refl, }, + map_mul' := λ g₁ g₂, by { ext, refl, }, + commutes' := λ r, by { ext, refl, }, } + +variables {A} + /-- Coercion to a function as an `alg_hom`. -/ @[simps] def continuous_map.coe_fn_alg_hom : C(α, A) →ₐ[R] (α → A) := diff --git a/src/topology/continuous_function/compact.lean b/src/topology/continuous_function/compact.lean index 6d94f0acc0e9b..00929bfc10e36 100644 --- a/src/topology/continuous_function/compact.lean +++ b/src/topology/continuous_function/compact.lean @@ -363,7 +363,7 @@ section comp_right Precomposition by a continuous map is itself a continuous map between spaces of continuous maps. -/ def comp_right_continuous_map {X Y : Type*} (T : Type*) [topological_space X] [compact_space X] - [topological_space Y] [compact_space Y] [normed_add_comm_group T] + [topological_space Y] [compact_space Y] [metric_space T] (f : C(X, Y)) : C(C(Y, T), C(X, T)) := { to_fun := λ g, g.comp f, continuous_to_fun := @@ -376,7 +376,7 @@ def comp_right_continuous_map {X Y : Type*} (T : Type*) [topological_space X] [c end } @[simp] lemma comp_right_continuous_map_apply {X Y : Type*} (T : Type*) [topological_space X] - [compact_space X] [topological_space Y] [compact_space Y] [normed_add_comm_group T] + [compact_space X] [topological_space Y] [compact_space Y] [metric_space T] (f : C(X, Y)) (g : C(Y, T)) : (comp_right_continuous_map T f) g = g.comp f := rfl @@ -385,39 +385,18 @@ rfl Precomposition by a homeomorphism is itself a homeomorphism between spaces of continuous maps. -/ def comp_right_homeomorph {X Y : Type*} (T : Type*) [topological_space X] [compact_space X] - [topological_space Y] [compact_space Y] [normed_add_comm_group T] + [topological_space Y] [compact_space Y] [metric_space T] (f : X ≃ₜ Y) : C(Y, T) ≃ₜ C(X, T) := { to_fun := comp_right_continuous_map T f.to_continuous_map, inv_fun := comp_right_continuous_map T f.symm.to_continuous_map, - left_inv := by tidy, - right_inv := by tidy, } - -/-- -Precomposition of functions into a normed ring by continuous map is an algebra homomorphism. --/ -def comp_right_alg_hom {X Y : Type*} (R : Type*) - [topological_space X] [topological_space Y] [normed_comm_ring R] (f : C(X, Y)) : - C(Y, R) →ₐ[R] C(X, R) := -{ to_fun := λ g, g.comp f, - map_zero' := by { ext, simp, }, - map_add' := λ g₁ g₂, by { ext, simp, }, - map_one' := by { ext, simp, }, - map_mul' := λ g₁ g₂, by { ext, simp, }, - commutes' := λ r, by { ext, simp, }, } - -@[simp] lemma comp_right_alg_hom_apply {X Y : Type*} (R : Type*) - [topological_space X] [topological_space Y] [normed_comm_ring R] (f : C(X, Y)) (g : C(Y, R)) : - (comp_right_alg_hom R f) g = g.comp f := -rfl - -lemma comp_right_alg_hom_continuous {X Y : Type*} (R : Type*) - [topological_space X] [compact_space X] [topological_space Y] [compact_space Y] - [normed_comm_ring R] (f : C(X, Y)) : - continuous (comp_right_alg_hom R f) := -begin - change continuous (comp_right_continuous_map R f), - continuity, -end + left_inv := λ g, ext $ λ _, congr_arg g (f.apply_symm_apply _), + right_inv := λ g, ext $ λ _, congr_arg g (f.symm_apply_apply _) } + +lemma comp_right_alg_hom_continuous {X Y : Type*} (R A : Type*) + [topological_space X] [compact_space X] [topological_space Y] [compact_space Y] [comm_semiring R] + [semiring A] [metric_space A] [topological_semiring A] [algebra R A] (f : C(X, Y)) : + continuous (comp_right_alg_hom R A f) := +map_continuous (comp_right_continuous_map A f) end comp_right diff --git a/src/topology/continuous_function/polynomial.lean b/src/topology/continuous_function/polynomial.lean index 04ae3993512da..e03d45c5d6e79 100644 --- a/src/topology/continuous_function/polynomial.lean +++ b/src/topology/continuous_function/polynomial.lean @@ -138,7 +138,7 @@ open continuous_map is the polynomials on `[a,b]`. -/ lemma polynomial_functions.comap_comp_right_alg_hom_Icc_homeo_I (a b : ℝ) (h : a < b) : (polynomial_functions I).comap - (comp_right_alg_hom ℝ (Icc_homeo_I a b h).symm.to_continuous_map) = + (comp_right_alg_hom ℝ ℝ (Icc_homeo_I a b h).symm.to_continuous_map) = polynomial_functions (set.Icc a b) := begin ext f, diff --git a/src/topology/continuous_function/weierstrass.lean b/src/topology/continuous_function/weierstrass.lean index 6cf001aa8dfed..f2189eb8f3659 100644 --- a/src/topology/continuous_function/weierstrass.lean +++ b/src/topology/continuous_function/weierstrass.lean @@ -58,7 +58,7 @@ begin { -- We can pullback continuous functions on `[a,b]` to continuous functions on `[0,1]`, -- by precomposing with an affine map. let W : C(set.Icc a b, ℝ) →ₐ[ℝ] C(I, ℝ) := - comp_right_alg_hom ℝ (Icc_homeo_I a b h).symm.to_continuous_map, + comp_right_alg_hom ℝ ℝ (Icc_homeo_I a b h).symm.to_continuous_map, -- This operation is itself a homeomorphism -- (with respect to the norm topologies on continuous functions). let W' : C(set.Icc a b, ℝ) ≃ₜ C(I, ℝ) := comp_right_homeomorph ℝ (Icc_homeo_I a b h).symm, From 234ddfeaa5572bc13716dd215c6444410a679a8e Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Thu, 13 Oct 2022 15:02:27 +0000 Subject: [PATCH 752/828] chore(category_theory): put biproduct results on preadditive categories into the correct file (#16924) --- src/algebra/category/Group/biproducts.lean | 2 +- src/category_theory/abelian/basic.lean | 2 +- .../limits/preserves/shapes/biproducts.lean | 216 +---- .../limits/shapes/biproducts.lean | 613 +------------ .../preadditive/additive_functor.lean | 2 +- .../preadditive/biproducts.lean | 813 +++++++++++++++++- .../preadditive/hom_orthogonal.lean | 2 +- 7 files changed, 818 insertions(+), 832 deletions(-) diff --git a/src/algebra/category/Group/biproducts.lean b/src/algebra/category/Group/biproducts.lean index 7b2e208d29549..7616b1378093f 100644 --- a/src/algebra/category/Group/biproducts.lean +++ b/src/algebra/category/Group/biproducts.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison -/ import algebra.group.pi import algebra.category.Group.preadditive -import category_theory.limits.shapes.biproducts +import category_theory.preadditive.biproducts import algebra.category.Group.limits /-! diff --git a/src/category_theory/abelian/basic.lean b/src/category_theory/abelian/basic.lean index 1d0d3e8df9ed4..28d79e49f15a8 100644 --- a/src/category_theory/abelian/basic.lean +++ b/src/category_theory/abelian/basic.lean @@ -5,7 +5,7 @@ Authors: Markus Himmel, Johan Commelin, Scott Morrison -/ import category_theory.limits.constructions.pullbacks -import category_theory.limits.shapes.biproducts +import category_theory.preadditive.biproducts import category_theory.limits.shapes.images import category_theory.limits.constructions.limits_of_products_and_equalizers import category_theory.limits.constructions.epi_mono diff --git a/src/category_theory/limits/preserves/shapes/biproducts.lean b/src/category_theory/limits/preserves/shapes/biproducts.lean index 5727f969e1b36..c897b208c5347 100644 --- a/src/category_theory/limits/preserves/shapes/biproducts.lean +++ b/src/category_theory/limits/preserves/shapes/biproducts.lean @@ -18,9 +18,7 @@ classes `preserves_biproduct` and `preserves_binary_biproduct`. We then * construct the comparison morphisms between the image of a biproduct and the biproduct of the images and show that the biproduct is preserved if one of them is an isomorphism, * give the canonical isomorphism between the image of a biproduct and the biproduct of the images - in case that the biproduct is preserved, -* show that in a preadditive category, a functor preserves a biproduct if and only if it preserves - the corresponding product if and only if it preserves the corresponding coproduct. + in case that the biproduct is preserved. -/ @@ -380,216 +378,4 @@ end limits end has_zero_morphisms -open category_theory.functor - -section preadditive -variables [preadditive C] [preadditive D] (F : C ⥤ D) [preserves_zero_morphisms F] - -namespace limits - -section fintype -variables {J : Type} [fintype J] - -local attribute [tidy] tactic.discrete_cases - -/-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts - preserves finite products. -/ -def preserves_product_of_preserves_biproduct {f : J → C} [preserves_biproduct f F] : - preserves_limit (discrete.functor f) F := -{ preserves := λ c hc, is_limit.of_iso_limit - ((is_limit.postcompose_inv_equiv (discrete.comp_nat_iso_discrete _ _) _).symm - (is_bilimit_of_preserves F (bicone_is_bilimit_of_limit_cone_of_is_limit hc)).is_limit) $ - cones.ext (iso.refl _) (by tidy) } - -section -local attribute [instance] preserves_product_of_preserves_biproduct - -/-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts - preserves finite products. -/ -def preserves_products_of_shape_of_preserves_biproducts_of_shape - [preserves_biproducts_of_shape J F] : preserves_limits_of_shape (discrete J) F := -{ preserves_limit := λ f, preserves_limit_of_iso_diagram _ discrete.nat_iso_functor.symm } - -end - -/-- A functor between preadditive categories that preserves (zero morphisms and) finite products - preserves finite biproducts. -/ -def preserves_biproduct_of_preserves_product {f : J → C} [preserves_limit (discrete.functor f) F] : - preserves_biproduct f F := -{ preserves := λ b hb, is_bilimit_of_is_limit _ $ - is_limit.of_iso_limit ((is_limit.postcompose_hom_equiv (discrete.comp_nat_iso_discrete _ _) - (F.map_cone b.to_cone)).symm (is_limit_of_preserves F hb.is_limit)) $ - cones.ext (iso.refl _) (by tidy) } - -/-- If the (product-like) biproduct comparison for `F` and `f` is a monomorphism, then `F` - preserves the biproduct of `f`. For the converse, see `map_biproduct`. -/ -def preserves_biproduct_of_mono_biproduct_comparison {f : J → C} [has_biproduct f] - [has_biproduct (F.obj ∘ f)] [mono (biproduct_comparison F f)] : preserves_biproduct f F := -begin - have : pi_comparison F f = (F.map_iso (biproduct.iso_product f)).inv ≫ - biproduct_comparison F f ≫ (biproduct.iso_product _).hom, - { ext, convert pi_comparison_comp_π F f j.as; simp [← functor.map_comp] }, - haveI : is_iso (biproduct_comparison F f) := is_iso_of_mono_of_is_split_epi _, - haveI : is_iso (pi_comparison F f) := by { rw this, apply_instance }, - haveI := preserves_product.of_iso_comparison F f, - apply preserves_biproduct_of_preserves_product -end - -/-- If the (coproduct-like) biproduct comparison for `F` and `f` is an epimorphism, then `F` - preserves the biproduct of `F` and `f`. For the converse, see `map_biproduct`. -/ -def preserves_biproduct_of_epi_biproduct_comparison' {f : J → C} [has_biproduct f] - [has_biproduct (F.obj ∘ f)] [epi (biproduct_comparison' F f)] : preserves_biproduct f F := -begin - haveI : epi ((split_epi_biproduct_comparison F f).section_) := by simpa, - haveI : is_iso (biproduct_comparison F f) := is_iso.of_epi_section' - (split_epi_biproduct_comparison F f), - apply preserves_biproduct_of_mono_biproduct_comparison -end - -/-- A functor between preadditive categories that preserves (zero morphisms and) finite products - preserves finite biproducts. -/ -def preserves_biproducts_of_shape_of_preserves_products_of_shape - [preserves_limits_of_shape (discrete J) F] : preserves_biproducts_of_shape J F := -{ preserves := λ f, preserves_biproduct_of_preserves_product F } - -/-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts - preserves finite coproducts. -/ -def preserves_coproduct_of_preserves_biproduct {f : J → C} [preserves_biproduct f F] : - preserves_colimit (discrete.functor f) F := -{ preserves := λ c hc, is_colimit.of_iso_colimit - ((is_colimit.precompose_hom_equiv (discrete.comp_nat_iso_discrete _ _) _).symm - (is_bilimit_of_preserves F - (bicone_is_bilimit_of_colimit_cocone_of_is_colimit hc)).is_colimit) $ - cocones.ext (iso.refl _) (by tidy) } - -section -local attribute [instance] preserves_coproduct_of_preserves_biproduct - -/-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts - preserves finite coproducts. -/ -def preserves_coproducts_of_shape_of_preserves_biproducts_of_shape - [preserves_biproducts_of_shape J F] : preserves_colimits_of_shape (discrete J) F := -{ preserves_colimit := λ f, preserves_colimit_of_iso_diagram _ discrete.nat_iso_functor.symm } - -end - -/-- A functor between preadditive categories that preserves (zero morphisms and) finite coproducts - preserves finite biproducts. -/ -def preserves_biproduct_of_preserves_coproduct {f : J → C} - [preserves_colimit (discrete.functor f) F] : preserves_biproduct f F := -{ preserves := λ b hb, is_bilimit_of_is_colimit _ $ - is_colimit.of_iso_colimit ((is_colimit.precompose_inv_equiv (discrete.comp_nat_iso_discrete _ _) - (F.map_cocone b.to_cocone)).symm (is_colimit_of_preserves F hb.is_colimit)) $ - cocones.ext (iso.refl _) (by tidy) } - -/-- A functor between preadditive categories that preserves (zero morphisms and) finite coproducts - preserves finite biproducts. -/ -def preserves_biproducts_of_shape_of_preserves_coproducts_of_shape - [preserves_colimits_of_shape (discrete J) F] : preserves_biproducts_of_shape J F := -{ preserves := λ f, preserves_biproduct_of_preserves_coproduct F } - -end fintype - -/-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts - preserves binary products. -/ -def preserves_binary_product_of_preserves_binary_biproduct {X Y : C} - [preserves_binary_biproduct X Y F] : preserves_limit (pair X Y) F := -{ preserves := λ c hc, is_limit.of_iso_limit - ((is_limit.postcompose_inv_equiv (by exact diagram_iso_pair _) _).symm - (is_binary_bilimit_of_preserves F - (binary_bicone_is_bilimit_of_limit_cone_of_is_limit hc)).is_limit) $ - cones.ext (iso.refl _) (λ j, by { rcases j with ⟨⟨⟩⟩, tidy }) } - -section -local attribute [instance] preserves_binary_product_of_preserves_binary_biproduct - -/-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts - preserves binary products. -/ -def preserves_binary_products_of_preserves_binary_biproducts - [preserves_binary_biproducts F] : preserves_limits_of_shape (discrete walking_pair) F := -{ preserves_limit := λ K, preserves_limit_of_iso_diagram _ (diagram_iso_pair _).symm } - -end - -/-- A functor between preadditive categories that preserves (zero morphisms and) binary products - preserves binary biproducts. -/ -def preserves_binary_biproduct_of_preserves_binary_product {X Y : C} - [preserves_limit (pair X Y) F] : preserves_binary_biproduct X Y F := -{ preserves := λ b hb, is_binary_bilimit_of_is_limit _ $ - is_limit.of_iso_limit ((is_limit.postcompose_hom_equiv (by exact diagram_iso_pair _) - (F.map_cone b.to_cone)).symm (is_limit_of_preserves F hb.is_limit)) $ - cones.ext (iso.refl _) (λ j, by { rcases j with ⟨⟨⟩⟩, tidy }) } - -/-- If the (product-like) biproduct comparison for `F`, `X` and `Y` is a monomorphism, then - `F` preserves the biproduct of `X` and `Y`. For the converse, see `map_biprod`. -/ -def preserves_binary_biproduct_of_mono_biprod_comparison {X Y : C} [has_binary_biproduct X Y] - [has_binary_biproduct (F.obj X) (F.obj Y)] [mono (biprod_comparison F X Y)] : - preserves_binary_biproduct X Y F := -begin - have : prod_comparison F X Y = (F.map_iso (biprod.iso_prod X Y)).inv ≫ - biprod_comparison F X Y ≫ (biprod.iso_prod _ _).hom := by { ext; simp [← functor.map_comp] }, - haveI : is_iso (biprod_comparison F X Y) := is_iso_of_mono_of_is_split_epi _, - haveI : is_iso (prod_comparison F X Y) := by { rw this, apply_instance }, - haveI := preserves_limit_pair.of_iso_prod_comparison F X Y, - apply preserves_binary_biproduct_of_preserves_binary_product -end - -/-- If the (coproduct-like) biproduct comparison for `F`, `X` and `Y` is an epimorphism, then - `F` preserves the biproduct of `X` and `Y`. For the converse, see `map_biprod`. -/ -def preserves_binary_biproduct_of_epi_biprod_comparison' {X Y : C} [has_binary_biproduct X Y] - [has_binary_biproduct (F.obj X) (F.obj Y)] [epi (biprod_comparison' F X Y)] : - preserves_binary_biproduct X Y F := -begin - haveI : epi ((split_epi_biprod_comparison F X Y).section_) := by simpa, - haveI : is_iso (biprod_comparison F X Y) := is_iso.of_epi_section' - (split_epi_biprod_comparison F X Y), - apply preserves_binary_biproduct_of_mono_biprod_comparison -end - -/-- A functor between preadditive categories that preserves (zero morphisms and) binary products - preserves binary biproducts. -/ -def preserves_binary_biproducts_of_preserves_binary_products - [preserves_limits_of_shape (discrete walking_pair) F] : preserves_binary_biproducts F := -{ preserves := λ X Y, preserves_binary_biproduct_of_preserves_binary_product F } - -/-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts - preserves binary coproducts. -/ -def preserves_binary_coproduct_of_preserves_binary_biproduct {X Y : C} - [preserves_binary_biproduct X Y F] : preserves_colimit (pair X Y) F := -{ preserves := λ c hc, is_colimit.of_iso_colimit - ((is_colimit.precompose_hom_equiv (by exact diagram_iso_pair _) _).symm - (is_binary_bilimit_of_preserves F - (binary_bicone_is_bilimit_of_colimit_cocone_of_is_colimit hc)).is_colimit) $ - cocones.ext (iso.refl _) (λ j, by { rcases j with ⟨⟨⟩⟩, tidy }) } - -section -local attribute [instance] preserves_binary_coproduct_of_preserves_binary_biproduct - -/-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts - preserves binary coproducts. -/ -def preserves_binary_coproducts_of_preserves_binary_biproducts - [preserves_binary_biproducts F] : preserves_colimits_of_shape (discrete walking_pair) F := -{ preserves_colimit := λ K, preserves_colimit_of_iso_diagram _ (diagram_iso_pair _).symm } - -end - -/-- A functor between preadditive categories that preserves (zero morphisms and) binary coproducts - preserves binary biproducts. -/ -def preserves_binary_biproduct_of_preserves_binary_coproduct {X Y : C} - [preserves_colimit (pair X Y) F] : preserves_binary_biproduct X Y F := -{ preserves := λ b hb, is_binary_bilimit_of_is_colimit _ $ - is_colimit.of_iso_colimit ((is_colimit.precompose_inv_equiv (by exact diagram_iso_pair _) - (F.map_cocone b.to_cocone)).symm (is_colimit_of_preserves F hb.is_colimit)) $ - cocones.ext (iso.refl _) (λ j, by { rcases j with ⟨⟨⟩⟩, tidy }) } - -/-- A functor between preadditive categories that preserves (zero morphisms and) binary coproducts - preserves binary biproducts. -/ -def preserves_binary_biproducts_of_preserves_binary_coproducts - [preserves_colimits_of_shape (discrete walking_pair) F] : preserves_binary_biproducts F := -{ preserves := λ X Y, preserves_binary_biproduct_of_preserves_binary_coproduct F } - -end limits - -end preadditive - end category_theory diff --git a/src/category_theory/limits/shapes/biproducts.lean b/src/category_theory/limits/shapes/biproducts.lean index bc4b5f2cca98c..ec6556ca17f15 100644 --- a/src/category_theory/limits/shapes/biproducts.lean +++ b/src/category_theory/limits/shapes/biproducts.lean @@ -18,8 +18,8 @@ These are slightly unusual relative to the other shapes in the library, as they are simultaneously limits and colimits. (Zero objects are similar; they are "biterminal".) -We treat first the case of a general category with zero morphisms, -and subsequently the case of a preadditive category. +For results about biproducts in preadditive categories see +`category_theory.preadditive.biproducts`. In a category with zero morphisms, we model the (binary) biproduct of `P Q : C` using a `binary_bicone`, which has a cone point `X`, @@ -28,20 +28,10 @@ such that `inl ≫ fst = 𝟙 P`, `inl ≫ snd = 0`, `inr ≫ fst = 0`, and `inr Such a `binary_bicone` is a biproduct if the cone is a limit cone, and the cocone is a colimit cocone. -In a preadditive category, -* any `binary_biproduct` satisfies `total : fst ≫ inl + snd ≫ inr = 𝟙 X` -* any `binary_product` is a `binary_biproduct` -* any `binary_coproduct` is a `binary_biproduct` - For biproducts indexed by a `fintype J`, a `bicone` again consists of a cone point `X` and morphisms `π j : X ⟶ F j` and `ι j : F j ⟶ X` for each `j`, such that `ι j ≫ π j'` is the identity when `j = j'` and zero otherwise. -In a preadditive category, -* any `biproduct` satisfies `total : ∑ j : J, biproduct.π f j ≫ biproduct.ι f j = 𝟙 (⨁ f)` -* any `product` is a `biproduct` -* any `coproduct` is a `biproduct` - ## Notation As `⊕` is already taken for the sum of types, we introduce the notation `X ⊞ Y` for a binary biproduct. We introduce `⨁ f` for the indexed biproduct. @@ -1332,14 +1322,11 @@ def biprod.unique_up_to_iso (X Y : C) [has_binary_biproduct X Y] {b : binary_bic inv_hom_id' := by rw [← biprod.cone_point_unique_up_to_iso_hom X Y hb, ← biprod.cone_point_unique_up_to_iso_inv X Y hb, iso.inv_hom_id] } -section -variables (X Y : C) [has_binary_biproduct X Y] - -- There are three further variations, -- about `is_iso biprod.inr`, `is_iso biprod.fst` and `is_iso biprod.snd`, -- but any one suffices to prove `indecomposable_of_simple` -- and they are likely not separately useful. -lemma biprod.is_iso_inl_iff_id_eq_fst_comp_inl : +lemma biprod.is_iso_inl_iff_id_eq_fst_comp_inl (X Y : C) [has_binary_biproduct X Y] : is_iso (biprod.inl : X ⟶ X ⊞ Y) ↔ 𝟙 (X ⊞ Y) = biprod.fst ≫ biprod.inl := begin split, @@ -1350,8 +1337,6 @@ begin { intro h, exact ⟨⟨biprod.fst, biprod.inl_fst, h.symm⟩⟩, }, end -end - section biprod_kernel section binary_bicone @@ -1577,570 +1562,58 @@ by simp end --- TODO: --- If someone is interested, they could provide the constructions: --- has_binary_biproducts ↔ has_finite_biproducts - end limits -namespace limits - -section preadditive -variables {C : Type u} [category.{v} C] [preadditive C] -variables {J : Type} [fintype J] - -open category_theory.preadditive -open_locale big_operators - -/-- -In a preadditive category, we can construct a biproduct for `f : J → C` from -any bicone `b` for `f` satisfying `total : ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X`. - -(That is, such a bicone is a limit cone and a colimit cocone.) --/ -def is_bilimit_of_total {f : J → C} (b : bicone f) (total : ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X) : - b.is_bilimit := -{ is_limit := - { lift := λ s, ∑ (j : J), s.π.app ⟨j⟩ ≫ b.ι j, - uniq' := λ s m h, - begin - erw [←category.comp_id m, ←total, comp_sum], - apply finset.sum_congr rfl, - intros j m, - erw [reassoc_of (h ⟨j⟩)], - end, - fac' := λ s j, - begin - cases j, - simp only [sum_comp, category.assoc, bicone.to_cone_π_app, b.ι_π, comp_dite], - -- See note [dsimp, simp]. - dsimp, simp, - end }, - is_colimit := - { desc := λ s, ∑ (j : J), b.π j ≫ s.ι.app ⟨j⟩, - uniq' := λ s m h, - begin - erw [←category.id_comp m, ←total, sum_comp], - apply finset.sum_congr rfl, - intros j m, - erw [category.assoc, h ⟨j⟩], - end, - fac' := λ s j, - begin - cases j, - simp only [comp_sum, ←category.assoc, bicone.to_cocone_ι_app, b.ι_π, dite_comp], - dsimp, simp, - end } } - -lemma is_bilimit.total {f : J → C} {b : bicone f} (i : b.is_bilimit) : - ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X := -i.is_limit.hom_ext (λ j, by { cases j, simp [sum_comp, b.ι_π, comp_dite] }) - -/-- -In a preadditive category, we can construct a biproduct for `f : J → C` from -any bicone `b` for `f` satisfying `total : ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X`. - -(That is, such a bicone is a limit cone and a colimit cocone.) --/ -lemma has_biproduct_of_total {f : J → C} (b : bicone f) (total : ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X) : - has_biproduct f := -has_biproduct.mk -{ bicone := b, - is_bilimit := is_bilimit_of_total b total } - -/-- In a preadditive category, any finite bicone which is a limit cone is in fact a bilimit - bicone. -/ -def is_bilimit_of_is_limit {f : J → C} (t : bicone f) (ht : is_limit t.to_cone) : t.is_bilimit := -is_bilimit_of_total _ $ ht.hom_ext $ - λ j, by { cases j, simp [sum_comp, t.ι_π, dite_comp, comp_dite] } - -/-- We can turn any limit cone over a pair into a bilimit bicone. -/ -def bicone_is_bilimit_of_limit_cone_of_is_limit {f : J → C} {t : cone (discrete.functor f)} - (ht : is_limit t) : (bicone.of_limit_cone ht).is_bilimit := -is_bilimit_of_is_limit _ $ - is_limit.of_iso_limit ht $ cones.ext (iso.refl _) (by { rintro ⟨j⟩, tidy }) - -/-- In a preadditive category, if the product over `f : J → C` exists, - then the biproduct over `f` exists. -/ -lemma has_biproduct.of_has_product {J : Type} [finite J] (f : J → C) [has_product f] : - has_biproduct f := -by casesI nonempty_fintype J; exact -has_biproduct.mk -{ bicone := _, - is_bilimit := bicone_is_bilimit_of_limit_cone_of_is_limit (limit.is_limit _) } - -/-- In a preadditive category, any finite bicone which is a colimit cocone is in fact a bilimit - bicone. -/ -def is_bilimit_of_is_colimit {f : J → C} (t : bicone f) (ht : is_colimit t.to_cocone) : - t.is_bilimit := -is_bilimit_of_total _ $ ht.hom_ext $ λ j, begin - cases j, - simp_rw [bicone.to_cocone_ι_app, comp_sum, ← category.assoc, t.ι_π, dite_comp], - tidy -end - -/-- We can turn any limit cone over a pair into a bilimit bicone. -/ -def bicone_is_bilimit_of_colimit_cocone_of_is_colimit {f : J → C} {t : cocone (discrete.functor f)} - (ht : is_colimit t) : (bicone.of_colimit_cocone ht).is_bilimit := -is_bilimit_of_is_colimit _ $ - is_colimit.of_iso_colimit ht $ cocones.ext (iso.refl _) (by { rintro ⟨j⟩, tidy }) - -/-- In a preadditive category, if the coproduct over `f : J → C` exists, - then the biproduct over `f` exists. -/ -lemma has_biproduct.of_has_coproduct {J : Type} [finite J] (f : J → C) [has_coproduct f] : - has_biproduct f := -by casesI nonempty_fintype J; exact -has_biproduct.mk -{ bicone := _, - is_bilimit := bicone_is_bilimit_of_colimit_cocone_of_is_colimit (colimit.is_colimit _) } - -/-- A preadditive category with finite products has finite biproducts. -/ -lemma has_finite_biproducts.of_has_finite_products [has_finite_products C] : - has_finite_biproducts C := -⟨λ J _, { has_biproduct := λ F, by exactI has_biproduct.of_has_product _ }⟩ - -/-- A preadditive category with finite coproducts has finite biproducts. -/ -lemma has_finite_biproducts.of_has_finite_coproducts [has_finite_coproducts C] : - has_finite_biproducts C := -⟨λ J _, { has_biproduct := λ F, by exactI has_biproduct.of_has_coproduct _ }⟩ - -section -variables {f : J → C} [has_biproduct f] - -/-- -In any preadditive category, any biproduct satsifies -`∑ j : J, biproduct.π f j ≫ biproduct.ι f j = 𝟙 (⨁ f)` --/ -@[simp] lemma biproduct.total : ∑ j : J, biproduct.π f j ≫ biproduct.ι f j = 𝟙 (⨁ f) := -is_bilimit.total (biproduct.is_bilimit _) - -lemma biproduct.lift_eq {T : C} {g : Π j, T ⟶ f j} : - biproduct.lift g = ∑ j, g j ≫ biproduct.ι f j := -begin - ext j, - simp [sum_comp, biproduct.ι_π, comp_dite], -end - -lemma biproduct.desc_eq {T : C} {g : Π j, f j ⟶ T} : - biproduct.desc g = ∑ j, biproduct.π f j ≫ g j := -begin - ext j, - simp [comp_sum, biproduct.ι_π_assoc, dite_comp], -end - -@[simp, reassoc] lemma 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 simp [biproduct.lift_eq, biproduct.desc_eq, comp_sum, sum_comp, biproduct.ι_π_assoc, - comp_dite, dite_comp] - -lemma biproduct.map_eq [has_finite_biproducts C] {f g : J → C} {h : Π j, f j ⟶ g j} : - biproduct.map h = ∑ j : J, biproduct.π f j ≫ h j ≫ biproduct.ι g j := -begin - ext, - simp [biproduct.ι_π, biproduct.ι_π_assoc, comp_sum, sum_comp, comp_dite, dite_comp], -end - -@[simp, reassoc] -lemma biproduct.matrix_desc - {K : Type} [fintype K] [has_finite_biproducts C] - {f : J → C} {g : K → C} (m : Π j k, f j ⟶ g k) {P} (x : Π k, g k ⟶ P) : - biproduct.matrix m ≫ biproduct.desc x = biproduct.desc (λ j, ∑ k, m j k ≫ x k) := -by { ext, simp, } - -@[simp, reassoc] -lemma biproduct.lift_matrix - {K : Type} [fintype K] [has_finite_biproducts C] - {f : J → C} {g : K → C} {P} (x : Π j, P ⟶ f j) (m : Π j k, f j ⟶ g k) : - biproduct.lift x ≫ biproduct.matrix m = biproduct.lift (λ k, ∑ j, x j ≫ m j k) := -by { ext, simp, } - -@[reassoc] -lemma biproduct.matrix_map - {K : Type} [fintype K] [has_finite_biproducts C] - {f : J → C} {g : K → C} {h : K → C} (m : Π j k, f j ⟶ g k) (n : Π k, g k ⟶ h k) : - biproduct.matrix m ≫ biproduct.map n = biproduct.matrix (λ j k, m j k ≫ n k) := -by { ext, simp, } - -@[reassoc] -lemma biproduct.map_matrix - {K : Type} [fintype K] [has_finite_biproducts C] - {f : J → C} {g : J → C} {h : K → C} (m : Π k, f k ⟶ g k) (n : Π j k, g j ⟶ h k) : - biproduct.map m ≫ biproduct.matrix n = biproduct.matrix (λ j k, m j ≫ n j k) := -by { ext, simp, } - -end - -/-- Reindex a categorical biproduct via an equivalence of the index types. -/ -@[simps] -def biproduct.reindex {β γ : Type} [fintype β] [decidable_eq β] [decidable_eq γ] - (ε : β ≃ γ) (f : γ → C) [has_biproduct f] [has_biproduct (f ∘ ε)] : (⨁ (f ∘ ε)) ≅ (⨁ f) := -{ hom := biproduct.desc (λ b, biproduct.ι f (ε b)), - inv := biproduct.lift (λ b, biproduct.π f (ε b)), - hom_inv_id' := by { ext b b', by_cases h : b = b', { subst h, simp, }, { simp [h], }, }, - inv_hom_id' := begin - ext g g', - by_cases h : g = g'; - simp [preadditive.sum_comp, preadditive.comp_sum, biproduct.ι_π, biproduct.ι_π_assoc, comp_dite, - equiv.apply_eq_iff_eq_symm_apply, finset.sum_dite_eq' finset.univ (ε.symm g') _, h], - end, } - -/-- -In a preadditive category, we can construct a binary biproduct for `X Y : C` from -any binary bicone `b` satisfying `total : b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X`. - -(That is, such a bicone is a limit cone and a colimit cocone.) --/ -def is_binary_bilimit_of_total {X Y : C} (b : binary_bicone X Y) - (total : b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X) : b.is_bilimit := -{ is_limit := - { lift := λ s, binary_fan.fst s ≫ b.inl + - binary_fan.snd s ≫ b.inr, - uniq' := λ s m h, by erw [←category.comp_id m, ←total, - comp_add, reassoc_of (h ⟨walking_pair.left⟩), reassoc_of (h ⟨walking_pair.right⟩)], - fac' := λ s j, by rcases j with ⟨⟨⟩⟩; simp, }, - is_colimit := - { desc := λ s, b.fst ≫ binary_cofan.inl s + - b.snd ≫ binary_cofan.inr s, - uniq' := λ s m h, by erw [←category.id_comp m, ←total, - add_comp, category.assoc, category.assoc, h ⟨walking_pair.left⟩, h ⟨walking_pair.right⟩], - fac' := λ s j, by rcases j with ⟨⟨⟩⟩; simp, } } - -lemma is_bilimit.binary_total {X Y : C} {b : binary_bicone X Y} (i : b.is_bilimit) : - b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X := -i.is_limit.hom_ext (λ j, by { rcases j with ⟨⟨⟩⟩; simp, }) - -/-- -In a preadditive category, we can construct a binary biproduct for `X Y : C` from -any binary bicone `b` satisfying `total : b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X`. - -(That is, such a bicone is a limit cone and a colimit cocone.) --/ -lemma has_binary_biproduct_of_total {X Y : C} (b : binary_bicone X Y) - (total : b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X) : has_binary_biproduct X Y := -has_binary_biproduct.mk -{ bicone := b, - is_bilimit := is_binary_bilimit_of_total b total } - -/-- We can turn any limit cone over a pair into a bicone. -/ -@[simps] -def binary_bicone.of_limit_cone {X Y : C} {t : cone (pair X Y)} (ht : is_limit t) : - binary_bicone X Y := -{ X := t.X, - fst := t.π.app ⟨walking_pair.left⟩, - snd := t.π.app ⟨walking_pair.right⟩, - inl := ht.lift (binary_fan.mk (𝟙 X) 0), - inr := ht.lift (binary_fan.mk 0 (𝟙 Y)) } - -lemma inl_of_is_limit {X Y : C} {t : binary_bicone X Y} (ht : is_limit t.to_cone) : - t.inl = ht.lift (binary_fan.mk (𝟙 X) 0) := -by apply ht.uniq (binary_fan.mk (𝟙 X) 0); rintro ⟨⟨⟩⟩; dsimp; simp - -lemma inr_of_is_limit {X Y : C} {t : binary_bicone X Y} (ht : is_limit t.to_cone) : - t.inr = ht.lift (binary_fan.mk 0 (𝟙 Y)) := -by apply ht.uniq (binary_fan.mk 0 (𝟙 Y)); rintro ⟨⟨⟩⟩; dsimp; simp - -/-- In a preadditive category, any binary bicone which is a limit cone is in fact a bilimit - bicone. -/ -def is_binary_bilimit_of_is_limit {X Y : C} (t : binary_bicone X Y) (ht : is_limit t.to_cone) : - t.is_bilimit := -is_binary_bilimit_of_total _ (by refine binary_fan.is_limit.hom_ext ht _ _; simp) - -/-- We can turn any limit cone over a pair into a bilimit bicone. -/ -def binary_bicone_is_bilimit_of_limit_cone_of_is_limit {X Y : C} {t : cone (pair X Y)} - (ht : is_limit t) : (binary_bicone.of_limit_cone ht).is_bilimit := -is_binary_bilimit_of_total _ $ binary_fan.is_limit.hom_ext ht (by simp) (by simp) - -/-- In a preadditive category, if the product of `X` and `Y` exists, then the - binary biproduct of `X` and `Y` exists. -/ -lemma has_binary_biproduct.of_has_binary_product (X Y : C) [has_binary_product X Y] : - has_binary_biproduct X Y := -has_binary_biproduct.mk -{ bicone := _, - is_bilimit := binary_bicone_is_bilimit_of_limit_cone_of_is_limit (limit.is_limit _) } - -/-- In a preadditive category, if all binary products exist, then all binary biproducts exist. -/ -lemma has_binary_biproducts.of_has_binary_products [has_binary_products C] : - has_binary_biproducts C := -{ has_binary_biproduct := λ X Y, has_binary_biproduct.of_has_binary_product X Y, } - -/-- We can turn any colimit cocone over a pair into a bicone. -/ -@[simps] -def binary_bicone.of_colimit_cocone {X Y : C} {t : cocone (pair X Y)} (ht : is_colimit t) : - binary_bicone X Y := -{ X := t.X, - fst := ht.desc (binary_cofan.mk (𝟙 X) 0), - snd := ht.desc (binary_cofan.mk 0 (𝟙 Y)), - inl := t.ι.app ⟨walking_pair.left⟩, - inr := t.ι.app ⟨walking_pair.right⟩ } - -lemma fst_of_is_colimit {X Y : C} {t : binary_bicone X Y} (ht : is_colimit t.to_cocone) : - t.fst = ht.desc (binary_cofan.mk (𝟙 X) 0) := -begin - apply ht.uniq (binary_cofan.mk (𝟙 X) 0), - rintro ⟨⟨⟩⟩; dsimp; simp -end - -lemma snd_of_is_colimit {X Y : C} {t : binary_bicone X Y} (ht : is_colimit t.to_cocone) : - t.snd = ht.desc (binary_cofan.mk 0 (𝟙 Y)) := -begin - apply ht.uniq (binary_cofan.mk 0 (𝟙 Y)), - rintro ⟨⟨⟩⟩; dsimp; simp -end - -/-- In a preadditive category, any binary bicone which is a colimit cocone is in fact a - bilimit bicone. -/ -def is_binary_bilimit_of_is_colimit {X Y : C} (t : binary_bicone X Y) - (ht : is_colimit t.to_cocone) : t.is_bilimit := -is_binary_bilimit_of_total _ -begin - refine binary_cofan.is_colimit.hom_ext ht _ _; simp, - { rw [category.comp_id t.inl] }, - { rw [category.comp_id t.inr] } -end +open category_theory.limits -/-- We can turn any colimit cocone over a pair into a bilimit bicone. -/ -def binary_bicone_is_bilimit_of_colimit_cocone_of_is_colimit {X Y : C} {t : cocone (pair X Y)} - (ht : is_colimit t) : (binary_bicone.of_colimit_cocone ht).is_bilimit := -is_binary_bilimit_of_is_colimit (binary_bicone.of_colimit_cocone ht) $ - is_colimit.of_iso_colimit ht $ cocones.ext (iso.refl _) $ λ j, by { rcases j with ⟨⟨⟩⟩, tidy } - -/-- In a preadditive category, if the coproduct of `X` and `Y` exists, then the - binary biproduct of `X` and `Y` exists. -/ -lemma has_binary_biproduct.of_has_binary_coproduct (X Y : C) [has_binary_coproduct X Y] : - has_binary_biproduct X Y := -has_binary_biproduct.mk -{ bicone := _, - is_bilimit := binary_bicone_is_bilimit_of_colimit_cocone_of_is_colimit (colimit.is_colimit _) } - -/-- In a preadditive category, if all binary coproducts exist, then all binary biproducts exist. -/ -lemma has_binary_biproducts.of_has_binary_coproducts [has_binary_coproducts C] : - has_binary_biproducts C := -{ has_binary_biproduct := λ X Y, has_binary_biproduct.of_has_binary_coproduct X Y, } +-- TODO: +-- If someone is interested, they could provide the constructions: +-- has_binary_biproducts ↔ has_finite_biproducts +variables {C : Type.{u}} [category.{v} C] [has_zero_morphisms C] [has_binary_biproducts C] -section -variables {X Y : C} [has_binary_biproduct X Y] +/-- An object is indecomposable if it cannot be written as the biproduct of two nonzero objects. -/ +def indecomposable (X : C) : Prop := ¬ is_zero X ∧ ∀ Y Z, (X ≅ Y ⊞ Z) → is_zero Y ∨ is_zero Z /-- -In any preadditive category, any binary biproduct satsifies -`biprod.fst ≫ biprod.inl + biprod.snd ≫ biprod.inr = 𝟙 (X ⊞ Y)`. +If +``` +(f 0) +(0 g) +``` +is invertible, then `f` is invertible. -/ -@[simp] lemma biprod.total : biprod.fst ≫ biprod.inl + biprod.snd ≫ biprod.inr = 𝟙 (X ⊞ Y) := -begin - ext; simp [add_comp], -end - -lemma biprod.lift_eq {T : C} {f : T ⟶ X} {g : T ⟶ Y} : - biprod.lift f g = f ≫ biprod.inl + g ≫ biprod.inr := -begin - ext; simp [add_comp], -end - -lemma biprod.desc_eq {T : C} {f : X ⟶ T} {g : Y ⟶ T} : - biprod.desc f g = biprod.fst ≫ f + biprod.snd ≫ g := -begin - ext; simp [add_comp], -end - -@[simp, reassoc] lemma biprod.lift_desc {T U : C} {f : T ⟶ X} {g : T ⟶ Y} {h : X ⟶ U} {i : Y ⟶ U} : - biprod.lift f g ≫ biprod.desc h i = f ≫ h + g ≫ i := -by simp [biprod.lift_eq, biprod.desc_eq] - -lemma biprod.map_eq [has_binary_biproducts C] {W X Y Z : C} {f : W ⟶ Y} {g : X ⟶ Z} : - biprod.map f g = biprod.fst ≫ f ≫ biprod.inl + biprod.snd ≫ g ≫ biprod.inr := -by apply biprod.hom_ext; apply biprod.hom_ext'; simp - -/-- -Every split mono `f` with a cokernel induces a binary bicone with `f` as its `inl` and -the cokernel map as its `snd`. -We will show in `is_bilimit_binary_bicone_of_split_mono_of_cokernel` that this binary bicone is in -fact already a biproduct. -/ -@[simps] -def binary_bicone_of_is_split_mono_of_cokernel {X Y : C} {f : X ⟶ Y} [is_split_mono f] - {c : cokernel_cofork f} (i : is_colimit c) : binary_bicone X c.X := -{ X := Y, - fst := retraction f, - snd := c.π, - inl := f, - inr := - let c' : cokernel_cofork (𝟙 Y - (𝟙 Y - retraction f ≫ f)) := - cokernel_cofork.of_π (cofork.π c) (by simp) in - let i' : is_colimit c' := is_cokernel_epi_comp i (retraction f) (by simp) in - let i'' := is_colimit_cofork_of_cokernel_cofork i' in - (split_epi_of_idempotent_of_is_colimit_cofork C (by simp) i'').section_, - inl_fst' := by simp, - inl_snd' := by simp, - inr_fst' := - begin - dsimp only, - rw [split_epi_of_idempotent_of_is_colimit_cofork_section_, - is_colimit_cofork_of_cokernel_cofork_desc, is_cokernel_epi_comp_desc], - dsimp only [cokernel_cofork_of_cofork_of_π], - letI := epi_of_is_colimit_cofork i, - apply zero_of_epi_comp c.π, - simp only [sub_comp, comp_sub, category.comp_id, category.assoc, is_split_mono.id, sub_self, - cofork.is_colimit.π_desc_assoc, cokernel_cofork.π_of_π, is_split_mono.id_assoc], - apply sub_eq_zero_of_eq, - apply category.id_comp +lemma is_iso_left_of_is_iso_biprod_map + {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [is_iso (biprod.map f g)] : is_iso f := +⟨⟨biprod.inl ≫ inv (biprod.map f g) ≫ biprod.fst, + ⟨begin + have t := congr_arg (λ p : W ⊞ X ⟶ W ⊞ X, biprod.inl ≫ p ≫ biprod.fst) + (is_iso.hom_inv_id (biprod.map f g)), + simp only [category.id_comp, category.assoc, biprod.inl_map_assoc] at t, + simp [t], end, - inr_snd' := by apply split_epi.id } - -/-- The bicone constructed in `binary_bicone_of_split_mono_of_cokernel` is a bilimit. -This is a version of the splitting lemma that holds in all preadditive categories. -/ -def is_bilimit_binary_bicone_of_is_split_mono_of_cokernel {X Y : C} {f : X ⟶ Y} [is_split_mono f] - {c : cokernel_cofork f} (i : is_colimit c) : - (binary_bicone_of_is_split_mono_of_cokernel i).is_bilimit := -is_binary_bilimit_of_total _ -begin - simp only [binary_bicone_of_is_split_mono_of_cokernel_fst, - binary_bicone_of_is_split_mono_of_cokernel_inr, binary_bicone_of_is_split_mono_of_cokernel_snd, - split_epi_of_idempotent_of_is_colimit_cofork_section_], - dsimp only [binary_bicone_of_is_split_mono_of_cokernel_X], - rw [is_colimit_cofork_of_cokernel_cofork_desc, is_cokernel_epi_comp_desc], - simp only [binary_bicone_of_is_split_mono_of_cokernel_inl, cofork.is_colimit.π_desc, - cokernel_cofork_of_cofork_π, cofork.π_of_π, add_sub_cancel'_right] -end - -/-- If `b` is a binary bicone such that `b.inl` is a kernel of `b.snd`, then `b` is a bilimit - bicone. -/ -def binary_bicone.is_bilimit_of_kernel_inl {X Y : C} (b : binary_bicone X Y) - (hb : is_limit b.snd_kernel_fork) : b.is_bilimit := -is_binary_bilimit_of_is_limit _ $ binary_fan.is_limit.mk _ - (λ T f g, f ≫ b.inl + g ≫ b.inr) (λ T f g, by simp) (λ T f g, by simp) $ λ T f g m h₁ h₂, - begin - have h₁' : (m - (f ≫ b.inl + g ≫ b.inr)) ≫ b.fst = 0 := by simpa using sub_eq_zero.2 h₁, - have h₂' : (m - (f ≫ b.inl + g ≫ b.inr)) ≫ b.snd = 0 := by simpa using sub_eq_zero.2 h₂, - obtain ⟨q : T ⟶ X, hq : q ≫ b.inl = m - (f ≫ b.inl + g ≫ b.inr)⟩ := - kernel_fork.is_limit.lift' hb _ h₂', - rw [←sub_eq_zero, ←hq, ←category.comp_id q, ←b.inl_fst, ←category.assoc, hq, h₁', zero_comp] - end - -/-- If `b` is a binary bicone such that `b.inr` is a kernel of `b.fst`, then `b` is a bilimit - bicone. -/ -def binary_bicone.is_bilimit_of_kernel_inr {X Y : C} (b : binary_bicone X Y) - (hb : is_limit b.fst_kernel_fork) : b.is_bilimit := -is_binary_bilimit_of_is_limit _ $ binary_fan.is_limit.mk _ - (λ T f g, f ≫ b.inl + g ≫ b.inr) (λ t f g, by simp) (λ t f g, by simp) $ λ T f g m h₁ h₂, begin - have h₁' : (m - (f ≫ b.inl + g ≫ b.inr)) ≫ b.fst = 0 := by simpa using sub_eq_zero.2 h₁, - have h₂' : (m - (f ≫ b.inl + g ≫ b.inr)) ≫ b.snd = 0 := by simpa using sub_eq_zero.2 h₂, - obtain ⟨q : T ⟶ Y, hq : q ≫ b.inr = m - (f ≫ b.inl + g ≫ b.inr)⟩ := - kernel_fork.is_limit.lift' hb _ h₁', - rw [←sub_eq_zero, ←hq, ←category.comp_id q, ←b.inr_snd, ←category.assoc, hq, h₂', zero_comp] - end - -/-- If `b` is a binary bicone such that `b.fst` is a cokernel of `b.inr`, then `b` is a bilimit - bicone. -/ -def binary_bicone.is_bilimit_of_cokernel_fst {X Y : C} (b : binary_bicone X Y) - (hb : is_colimit b.inr_cokernel_cofork) : b.is_bilimit := -is_binary_bilimit_of_is_colimit _ $ binary_cofan.is_colimit.mk _ - (λ T f g, b.fst ≫ f + b.snd ≫ g) (λ T f g, by simp) (λ T f g, by simp) $ λ T f g m h₁ h₂, - begin - have h₁' : b.inl ≫ (m - (b.fst ≫ f + b.snd ≫ g)) = 0 := by simpa using sub_eq_zero.2 h₁, - have h₂' : b.inr ≫ (m - (b.fst ≫ f + b.snd ≫ g)) = 0 := by simpa using sub_eq_zero.2 h₂, - obtain ⟨q : X ⟶ T, hq : b.fst ≫ q = m - (b.fst ≫ f + b.snd ≫ g)⟩ := - cokernel_cofork.is_colimit.desc' hb _ h₂', - rw [←sub_eq_zero, ←hq, ←category.id_comp q, ←b.inl_fst, category.assoc, hq, h₁', comp_zero] - end - -/-- If `b` is a binary bicone such that `b.snd` is a cokernel of `b.inl`, then `b` is a bilimit - bicone. -/ -def binary_bicone.is_bilimit_of_cokernel_snd {X Y : C} (b : binary_bicone X Y) - (hb : is_colimit b.inl_cokernel_cofork) : b.is_bilimit := -is_binary_bilimit_of_is_colimit _ $ binary_cofan.is_colimit.mk _ - (λ T f g, b.fst ≫ f + b.snd ≫ g) (λ T f g, by simp) (λ T f g, by simp) $ λ T f g m h₁ h₂, - begin - have h₁' : b.inl ≫ (m - (b.fst ≫ f + b.snd ≫ g)) = 0 := by simpa using sub_eq_zero.2 h₁, - have h₂' : b.inr ≫ (m - (b.fst ≫ f + b.snd ≫ g)) = 0 := by simpa using sub_eq_zero.2 h₂, - obtain ⟨q : Y ⟶ T, hq : b.snd ≫ q = m - (b.fst ≫ f + b.snd ≫ g)⟩ := - cokernel_cofork.is_colimit.desc' hb _ h₁', - rw [←sub_eq_zero, ←hq, ←category.id_comp q, ←b.inr_snd, category.assoc, hq, h₂', comp_zero] - end + have t := congr_arg (λ p : Y ⊞ Z ⟶ Y ⊞ Z, biprod.inl ≫ p ≫ biprod.fst) + (is_iso.inv_hom_id (biprod.map f g)), + simp only [category.id_comp, category.assoc, biprod.map_fst] at t, + simp only [category.assoc], + simp [t], + end⟩⟩⟩ /-- -Every split epi `f` with a kernel induces a binary bicone with `f` as its `snd` and -the kernel map as its `inl`. -We will show in `binary_bicone_of_is_split_mono_of_cokernel` that this binary bicone is in fact -already a biproduct. -/ -@[simps] -def binary_bicone_of_is_split_epi_of_kernel {X Y : C} {f : X ⟶ Y} [is_split_epi f] - {c : kernel_fork f} (i : is_limit c) : binary_bicone c.X Y := -{ X := X, - fst := - let c' : kernel_fork (𝟙 X - (𝟙 X - f ≫ section_ f)) := - kernel_fork.of_ι (fork.ι c) (by simp) in - let i' : is_limit c' := is_kernel_comp_mono i (section_ f) (by simp) in - let i'' := is_limit_fork_of_kernel_fork i' in - (split_mono_of_idempotent_of_is_limit_fork C (by simp) i'').retraction, - snd := f, - inl := c.ι, - inr := section_ f, - inl_fst' := by apply split_mono.id, - inl_snd' := by simp, - inr_fst' := - begin - dsimp only, - rw [split_mono_of_idempotent_of_is_limit_fork_retraction, - is_limit_fork_of_kernel_fork_lift, is_kernel_comp_mono_lift], - dsimp only [kernel_fork_of_fork_ι], - letI := mono_of_is_limit_fork i, - apply zero_of_comp_mono c.ι, - simp only [comp_sub, category.comp_id, category.assoc, sub_self, fork.is_limit.lift_ι, - fork.ι_of_ι, is_split_epi.id_assoc] - end, - inr_snd' := by simp } - -/-- The bicone constructed in `binary_bicone_of_is_split_epi_of_kernel` is a bilimit. -This is a version of the splitting lemma that holds in all preadditive categories. -/ -def is_bilimit_binary_bicone_of_is_split_epi_of_kernel {X Y : C} {f : X ⟶ Y} [is_split_epi f] - {c : kernel_fork f} (i : is_limit c) : - (binary_bicone_of_is_split_epi_of_kernel i).is_bilimit := -binary_bicone.is_bilimit_of_kernel_inl _ $ i.of_iso_limit $ fork.ext (iso.refl _) (by simp) - -end - -section -variables {X Y : C} (f g : X ⟶ Y) - -/-- The existence of binary biproducts implies that there is at most one preadditive structure. -/ -lemma biprod.add_eq_lift_id_desc [has_binary_biproduct X X] : - f + g = biprod.lift (𝟙 X) (𝟙 X) ≫ biprod.desc f g := -by simp - -/-- The existence of binary biproducts implies that there is at most one preadditive structure. -/ -lemma biprod.add_eq_lift_desc_id [has_binary_biproduct Y Y] : - f + g = biprod.lift f g ≫ biprod.desc (𝟙 Y) (𝟙 Y) := -by simp - -end - -end preadditive - -end limits - -open category_theory.limits - -section -local attribute [ext] preadditive - -/-- The existence of binary biproducts implies that there is at most one preadditive structure. -/ -instance subsingleton_preadditive_of_has_binary_biproducts {C : Type u} [category.{v} C] - [has_zero_morphisms C] [has_binary_biproducts C] : subsingleton (preadditive C) := -subsingleton.intro $ λ a b, +If +``` +(f 0) +(0 g) +``` +is invertible, then `g` is invertible. +-/ +lemma is_iso_right_of_is_iso_biprod_map + {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [is_iso (biprod.map f g)] : is_iso g := begin - ext X Y f g, - have h₁ := @biprod.add_eq_lift_id_desc _ _ a _ _ f g - (by convert (infer_instance : has_binary_biproduct X X)), - have h₂ := @biprod.add_eq_lift_id_desc _ _ b _ _ f g - (by convert (infer_instance : has_binary_biproduct X X)), - refine h₁.trans (eq.trans _ h₂.symm), - congr' 2; - exact subsingleton.elim _ _ -end + letI : is_iso (biprod.map g f) := by + { rw [←biprod.braiding_map_braiding], + apply_instance, }, + exact is_iso_left_of_is_iso_biprod_map g f, end -variables {C : Type u} [category.{v} C] [has_zero_morphisms C] [has_binary_biproducts C] - -/-- An object is indecomposable if it cannot be written as the biproduct of two nonzero objects. -/ -def indecomposable (X : C) : Prop := ¬ is_zero X ∧ ∀ Y Z, (X ≅ Y ⊞ Z) → is_zero Y ∨ is_zero Z - end category_theory diff --git a/src/category_theory/preadditive/additive_functor.lean b/src/category_theory/preadditive/additive_functor.lean index da442099950a0..05903b8bb0c89 100644 --- a/src/category_theory/preadditive/additive_functor.lean +++ b/src/category_theory/preadditive/additive_functor.lean @@ -5,7 +5,7 @@ Authors: Adam Topaz, Scott Morrison -/ import category_theory.limits.exact_functor import category_theory.limits.preserves.finite -import category_theory.limits.preserves.shapes.biproducts +import category_theory.preadditive.biproducts import category_theory.preadditive.functor_category /-! diff --git a/src/category_theory/preadditive/biproducts.lean b/src/category_theory/preadditive/biproducts.lean index eb707be966f12..65dbb3bbacda3 100644 --- a/src/category_theory/preadditive/biproducts.lean +++ b/src/category_theory/preadditive/biproducts.lean @@ -5,16 +5,23 @@ Authors: Scott Morrison -/ import tactic.abel import category_theory.limits.shapes.biproducts +import category_theory.limits.preserves.shapes.biproducts import category_theory.preadditive /-! -# Basic facts about morphisms between biproducts in preadditive categories. +# Basic facts about biproducts in preadditive categories. + +In (or between) preadditive categories, + +* Any biproduct satisfies the equality + `total : ∑ j : J, biproduct.π f j ≫ biproduct.ι f j = 𝟙 (⨁ f)`, + or, in the binary case, `total : fst ≫ inl + snd ≫ inr = 𝟙 X`. + +* Any (binary) `product` or (binary) `coproduct` is a (binary) `biproduct`. * In any category (with zero morphisms), if `biprod.map f g` is an isomorphism, then both `f` and `g` are isomorphisms. -The remaining lemmas hold in any preadditive category. - * If `f` is a morphism `X₁ ⊞ X₂ ⟶ Y₁ ⊞ Y₂` whose `X₁ ⟶ Y₁` entry is an isomorphism, then we can construct isomorphisms `L : X₁ ⊞ X₂ ≅ X₁ ⊞ X₂` and `R : Y₁ ⊞ Y₂ ≅ Y₁ ⊞ Y₂` so that `L.hom ≫ g ≫ R.hom` is diagonal (with `X₁ ⟶ Y₁` component still `f`), @@ -30,68 +37,579 @@ The remaining lemmas hold in any preadditive category. * If `f : ⨁ S ⟶ ⨁ T` is an isomorphism, then every column (corresponding to a nonzero summand in the domain) has some nonzero matrix entry. + +* A functor preserves a biproduct if and only if it preserves + the corresponding product if and only if it preserves the corresponding coproduct. -/ open category_theory open category_theory.preadditive open category_theory.limits +open category_theory.functor +open category_theory.preadditive -universes v u +open_locale classical +open_locale big_operators + +universes v v' u u' noncomputable theory namespace category_theory -variables {C : Type u} [category.{v} C] +variables {C : Type u} [category.{v} C] [preadditive C] + +namespace limits + +variables {J : Type} [fintype J] + +/-- +In a preadditive category, we can construct a biproduct for `f : J → C` from +any bicone `b` for `f` satisfying `total : ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X`. + +(That is, such a bicone is a limit cone and a colimit cocone.) +-/ +def is_bilimit_of_total {f : J → C} (b : bicone f) (total : ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X) : + b.is_bilimit := +{ is_limit := + { lift := λ s, ∑ (j : J), s.π.app ⟨j⟩ ≫ b.ι j, + uniq' := λ s m h, + begin + erw [←category.comp_id m, ←total, comp_sum], + apply finset.sum_congr rfl, + intros j m, + erw [reassoc_of (h ⟨j⟩)], + end, + fac' := λ s j, + begin + cases j, + simp only [sum_comp, category.assoc, bicone.to_cone_π_app, b.ι_π, comp_dite], + -- See note [dsimp, simp]. + dsimp, simp, + end }, + is_colimit := + { desc := λ s, ∑ (j : J), b.π j ≫ s.ι.app ⟨j⟩, + uniq' := λ s m h, + begin + erw [←category.id_comp m, ←total, sum_comp], + apply finset.sum_congr rfl, + intros j m, + erw [category.assoc, h ⟨j⟩], + end, + fac' := λ s j, + begin + cases j, + simp only [comp_sum, ←category.assoc, bicone.to_cocone_ι_app, b.ι_π, dite_comp], + dsimp, simp, + end } } + +lemma is_bilimit.total {f : J → C} {b : bicone f} (i : b.is_bilimit) : + ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X := +i.is_limit.hom_ext (λ j, by { cases j, simp [sum_comp, b.ι_π, comp_dite] }) + +/-- +In a preadditive category, we can construct a biproduct for `f : J → C` from +any bicone `b` for `f` satisfying `total : ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X`. + +(That is, such a bicone is a limit cone and a colimit cocone.) +-/ +lemma has_biproduct_of_total {f : J → C} (b : bicone f) (total : ∑ j : J, b.π j ≫ b.ι j = 𝟙 b.X) : + has_biproduct f := +has_biproduct.mk +{ bicone := b, + is_bilimit := is_bilimit_of_total b total } + +/-- In a preadditive category, any finite bicone which is a limit cone is in fact a bilimit + bicone. -/ +def is_bilimit_of_is_limit {f : J → C} (t : bicone f) (ht : is_limit t.to_cone) : t.is_bilimit := +is_bilimit_of_total _ $ ht.hom_ext $ + λ j, by { cases j, simp [sum_comp, t.ι_π, dite_comp, comp_dite] } + +/-- We can turn any limit cone over a pair into a bilimit bicone. -/ +def bicone_is_bilimit_of_limit_cone_of_is_limit {f : J → C} {t : cone (discrete.functor f)} + (ht : is_limit t) : (bicone.of_limit_cone ht).is_bilimit := +is_bilimit_of_is_limit _ $ + is_limit.of_iso_limit ht $ cones.ext (iso.refl _) (by { rintro ⟨j⟩, tidy }) + +/-- In a preadditive category, if the product over `f : J → C` exists, + then the biproduct over `f` exists. -/ +lemma has_biproduct.of_has_product {J : Type} [finite J] (f : J → C) [has_product f] : + has_biproduct f := +by casesI nonempty_fintype J; exact +has_biproduct.mk +{ bicone := _, + is_bilimit := bicone_is_bilimit_of_limit_cone_of_is_limit (limit.is_limit _) } + +/-- In a preadditive category, any finite bicone which is a colimit cocone is in fact a bilimit + bicone. -/ +def is_bilimit_of_is_colimit {f : J → C} (t : bicone f) (ht : is_colimit t.to_cocone) : + t.is_bilimit := +is_bilimit_of_total _ $ ht.hom_ext $ λ j, begin + cases j, + simp_rw [bicone.to_cocone_ι_app, comp_sum, ← category.assoc, t.ι_π, dite_comp], + tidy +end + +/-- We can turn any limit cone over a pair into a bilimit bicone. -/ +def bicone_is_bilimit_of_colimit_cocone_of_is_colimit {f : J → C} {t : cocone (discrete.functor f)} + (ht : is_colimit t) : (bicone.of_colimit_cocone ht).is_bilimit := +is_bilimit_of_is_colimit _ $ + is_colimit.of_iso_colimit ht $ cocones.ext (iso.refl _) (by { rintro ⟨j⟩, tidy }) + +/-- In a preadditive category, if the coproduct over `f : J → C` exists, + then the biproduct over `f` exists. -/ +lemma has_biproduct.of_has_coproduct {J : Type} [finite J] (f : J → C) [has_coproduct f] : + has_biproduct f := +by casesI nonempty_fintype J; exact +has_biproduct.mk +{ bicone := _, + is_bilimit := bicone_is_bilimit_of_colimit_cocone_of_is_colimit (colimit.is_colimit _) } + +/-- A preadditive category with finite products has finite biproducts. -/ +lemma has_finite_biproducts.of_has_finite_products [has_finite_products C] : + has_finite_biproducts C := +⟨λ J _, { has_biproduct := λ F, by exactI has_biproduct.of_has_product _ }⟩ + +/-- A preadditive category with finite coproducts has finite biproducts. -/ +lemma has_finite_biproducts.of_has_finite_coproducts [has_finite_coproducts C] : + has_finite_biproducts C := +⟨λ J _, { has_biproduct := λ F, by exactI has_biproduct.of_has_coproduct _ }⟩ + section -variables [has_zero_morphisms.{v} C] [has_binary_biproducts.{v} C] +variables {f : J → C} [has_biproduct f] /-- -If -``` -(f 0) -(0 g) -``` -is invertible, then `f` is invertible. +In any preadditive category, any biproduct satsifies +`∑ j : J, biproduct.π f j ≫ biproduct.ι f j = 𝟙 (⨁ f)` -/ -lemma is_iso_left_of_is_iso_biprod_map - {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [is_iso (biprod.map f g)] : is_iso f := -⟨⟨biprod.inl ≫ inv (biprod.map f g) ≫ biprod.fst, - ⟨begin - have t := congr_arg (λ p : W ⊞ X ⟶ W ⊞ X, biprod.inl ≫ p ≫ biprod.fst) - (is_iso.hom_inv_id (biprod.map f g)), - simp only [category.id_comp, category.assoc, biprod.inl_map_assoc] at t, - simp [t], - end, - begin - have t := congr_arg (λ p : Y ⊞ Z ⟶ Y ⊞ Z, biprod.inl ≫ p ≫ biprod.fst) - (is_iso.inv_hom_id (biprod.map f g)), - simp only [category.id_comp, category.assoc, biprod.map_fst] at t, - simp only [category.assoc], - simp [t], - end⟩⟩⟩ +@[simp] lemma biproduct.total : ∑ j : J, biproduct.π f j ≫ biproduct.ι f j = 𝟙 (⨁ f) := +is_bilimit.total (biproduct.is_bilimit _) + +lemma biproduct.lift_eq {T : C} {g : Π j, T ⟶ f j} : + biproduct.lift g = ∑ j, g j ≫ biproduct.ι f j := +begin + ext j, + simp only [sum_comp, biproduct.ι_π, comp_dite, biproduct.lift_π, category.assoc, comp_zero, + finset.sum_dite_eq', finset.mem_univ, eq_to_hom_refl, category.comp_id, if_true], +end + +lemma biproduct.desc_eq {T : C} {g : Π j, f j ⟶ T} : + biproduct.desc g = ∑ j, biproduct.π f j ≫ g j := +begin + ext j, + simp [comp_sum, biproduct.ι_π_assoc, dite_comp], +end + +@[simp, reassoc] lemma 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 simp [biproduct.lift_eq, biproduct.desc_eq, comp_sum, sum_comp, biproduct.ι_π_assoc, + comp_dite, dite_comp] + +lemma biproduct.map_eq [has_finite_biproducts C] {f g : J → C} {h : Π j, f j ⟶ g j} : + biproduct.map h = ∑ j : J, biproduct.π f j ≫ h j ≫ biproduct.ι g j := +begin + ext, + simp [biproduct.ι_π, biproduct.ι_π_assoc, comp_sum, sum_comp, comp_dite, dite_comp], +end + +@[simp, reassoc] +lemma biproduct.matrix_desc + {K : Type} [fintype K] [has_finite_biproducts C] + {f : J → C} {g : K → C} (m : Π j k, f j ⟶ g k) {P} (x : Π k, g k ⟶ P) : + biproduct.matrix m ≫ biproduct.desc x = biproduct.desc (λ j, ∑ k, m j k ≫ x k) := +by { ext, simp, } + +@[simp, reassoc] +lemma biproduct.lift_matrix + {K : Type} [fintype K] [has_finite_biproducts C] + {f : J → C} {g : K → C} {P} (x : Π j, P ⟶ f j) (m : Π j k, f j ⟶ g k) : + biproduct.lift x ≫ biproduct.matrix m = biproduct.lift (λ k, ∑ j, x j ≫ m j k) := +by { ext, simp, } + +@[reassoc] +lemma biproduct.matrix_map + {K : Type} [fintype K] [has_finite_biproducts C] + {f : J → C} {g : K → C} {h : K → C} (m : Π j k, f j ⟶ g k) (n : Π k, g k ⟶ h k) : + biproduct.matrix m ≫ biproduct.map n = biproduct.matrix (λ j k, m j k ≫ n k) := +by { ext, simp, } + +@[reassoc] +lemma biproduct.map_matrix + {K : Type} [fintype K] [has_finite_biproducts C] + {f : J → C} {g : J → C} {h : K → C} (m : Π k, f k ⟶ g k) (n : Π j k, g j ⟶ h k) : + biproduct.map m ≫ biproduct.matrix n = biproduct.matrix (λ j k, m j ≫ n j k) := +by { ext, simp, } + +end + +/-- Reindex a categorical biproduct via an equivalence of the index types. -/ +@[simps] +def biproduct.reindex {β γ : Type} [fintype β] [decidable_eq β] [decidable_eq γ] + (ε : β ≃ γ) (f : γ → C) [has_biproduct f] [has_biproduct (f ∘ ε)] : (⨁ (f ∘ ε)) ≅ (⨁ f) := +{ hom := biproduct.desc (λ b, biproduct.ι f (ε b)), + inv := biproduct.lift (λ b, biproduct.π f (ε b)), + hom_inv_id' := by { ext b b', by_cases h : b = b', { subst h, simp, }, { simp [h], }, }, + inv_hom_id' := begin + ext g g', + by_cases h : g = g'; + simp [preadditive.sum_comp, preadditive.comp_sum, biproduct.ι_π, biproduct.ι_π_assoc, comp_dite, + equiv.apply_eq_iff_eq_symm_apply, finset.sum_dite_eq' finset.univ (ε.symm g') _, h], + end, } /-- -If -``` -(f 0) -(0 g) -``` -is invertible, then `g` is invertible. +In a preadditive category, we can construct a binary biproduct for `X Y : C` from +any binary bicone `b` satisfying `total : b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X`. + +(That is, such a bicone is a limit cone and a colimit cocone.) -/ -lemma is_iso_right_of_is_iso_biprod_map - {W X Y Z : C} (f : W ⟶ Y) (g : X ⟶ Z) [is_iso (biprod.map f g)] : is_iso g := +def is_binary_bilimit_of_total {X Y : C} (b : binary_bicone X Y) + (total : b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X) : b.is_bilimit := +{ is_limit := + { lift := λ s, binary_fan.fst s ≫ b.inl + + binary_fan.snd s ≫ b.inr, + uniq' := λ s m h, by erw [←category.comp_id m, ←total, + comp_add, reassoc_of (h ⟨walking_pair.left⟩), reassoc_of (h ⟨walking_pair.right⟩)], + fac' := λ s j, by rcases j with ⟨⟨⟩⟩; simp, }, + is_colimit := + { desc := λ s, b.fst ≫ binary_cofan.inl s + + b.snd ≫ binary_cofan.inr s, + uniq' := λ s m h, by erw [←category.id_comp m, ←total, + add_comp, category.assoc, category.assoc, h ⟨walking_pair.left⟩, h ⟨walking_pair.right⟩], + fac' := λ s j, by rcases j with ⟨⟨⟩⟩; simp, } } + +lemma is_bilimit.binary_total {X Y : C} {b : binary_bicone X Y} (i : b.is_bilimit) : + b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X := +i.is_limit.hom_ext (λ j, by { rcases j with ⟨⟨⟩⟩; simp, }) + +/-- +In a preadditive category, we can construct a binary biproduct for `X Y : C` from +any binary bicone `b` satisfying `total : b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X`. + +(That is, such a bicone is a limit cone and a colimit cocone.) +-/ +lemma has_binary_biproduct_of_total {X Y : C} (b : binary_bicone X Y) + (total : b.fst ≫ b.inl + b.snd ≫ b.inr = 𝟙 b.X) : has_binary_biproduct X Y := +has_binary_biproduct.mk +{ bicone := b, + is_bilimit := is_binary_bilimit_of_total b total } + +/-- We can turn any limit cone over a pair into a bicone. -/ +@[simps] +def binary_bicone.of_limit_cone {X Y : C} {t : cone (pair X Y)} (ht : is_limit t) : + binary_bicone X Y := +{ X := t.X, + fst := t.π.app ⟨walking_pair.left⟩, + snd := t.π.app ⟨walking_pair.right⟩, + inl := ht.lift (binary_fan.mk (𝟙 X) 0), + inr := ht.lift (binary_fan.mk 0 (𝟙 Y)) } + +lemma inl_of_is_limit {X Y : C} {t : binary_bicone X Y} (ht : is_limit t.to_cone) : + t.inl = ht.lift (binary_fan.mk (𝟙 X) 0) := +by apply ht.uniq (binary_fan.mk (𝟙 X) 0); rintro ⟨⟨⟩⟩; dsimp; simp + +lemma inr_of_is_limit {X Y : C} {t : binary_bicone X Y} (ht : is_limit t.to_cone) : + t.inr = ht.lift (binary_fan.mk 0 (𝟙 Y)) := +by apply ht.uniq (binary_fan.mk 0 (𝟙 Y)); rintro ⟨⟨⟩⟩; dsimp; simp + +/-- In a preadditive category, any binary bicone which is a limit cone is in fact a bilimit + bicone. -/ +def is_binary_bilimit_of_is_limit {X Y : C} (t : binary_bicone X Y) (ht : is_limit t.to_cone) : + t.is_bilimit := +is_binary_bilimit_of_total _ (by refine binary_fan.is_limit.hom_ext ht _ _; simp) + +/-- We can turn any limit cone over a pair into a bilimit bicone. -/ +def binary_bicone_is_bilimit_of_limit_cone_of_is_limit {X Y : C} {t : cone (pair X Y)} + (ht : is_limit t) : (binary_bicone.of_limit_cone ht).is_bilimit := +is_binary_bilimit_of_total _ $ binary_fan.is_limit.hom_ext ht (by simp) (by simp) + +/-- In a preadditive category, if the product of `X` and `Y` exists, then the + binary biproduct of `X` and `Y` exists. -/ +lemma has_binary_biproduct.of_has_binary_product (X Y : C) [has_binary_product X Y] : + has_binary_biproduct X Y := +has_binary_biproduct.mk +{ bicone := _, + is_bilimit := binary_bicone_is_bilimit_of_limit_cone_of_is_limit (limit.is_limit _) } + +/-- In a preadditive category, if all binary products exist, then all binary biproducts exist. -/ +lemma has_binary_biproducts.of_has_binary_products [has_binary_products C] : + has_binary_biproducts C := +{ has_binary_biproduct := λ X Y, has_binary_biproduct.of_has_binary_product X Y, } + +/-- We can turn any colimit cocone over a pair into a bicone. -/ +@[simps] +def binary_bicone.of_colimit_cocone {X Y : C} {t : cocone (pair X Y)} (ht : is_colimit t) : + binary_bicone X Y := +{ X := t.X, + fst := ht.desc (binary_cofan.mk (𝟙 X) 0), + snd := ht.desc (binary_cofan.mk 0 (𝟙 Y)), + inl := t.ι.app ⟨walking_pair.left⟩, + inr := t.ι.app ⟨walking_pair.right⟩ } + +lemma fst_of_is_colimit {X Y : C} {t : binary_bicone X Y} (ht : is_colimit t.to_cocone) : + t.fst = ht.desc (binary_cofan.mk (𝟙 X) 0) := +begin + apply ht.uniq (binary_cofan.mk (𝟙 X) 0), + rintro ⟨⟨⟩⟩; dsimp; simp +end + +lemma snd_of_is_colimit {X Y : C} {t : binary_bicone X Y} (ht : is_colimit t.to_cocone) : + t.snd = ht.desc (binary_cofan.mk 0 (𝟙 Y)) := begin - letI : is_iso (biprod.map g f) := by - { rw [←biprod.braiding_map_braiding], - apply_instance, }, - exact is_iso_left_of_is_iso_biprod_map g f, + apply ht.uniq (binary_cofan.mk 0 (𝟙 Y)), + rintro ⟨⟨⟩⟩; dsimp; simp end +/-- In a preadditive category, any binary bicone which is a colimit cocone is in fact a + bilimit bicone. -/ +def is_binary_bilimit_of_is_colimit {X Y : C} (t : binary_bicone X Y) + (ht : is_colimit t.to_cocone) : t.is_bilimit := +is_binary_bilimit_of_total _ +begin + refine binary_cofan.is_colimit.hom_ext ht _ _; simp, + { rw [category.comp_id t.inl] }, + { rw [category.comp_id t.inr] } end +/-- We can turn any colimit cocone over a pair into a bilimit bicone. -/ +def binary_bicone_is_bilimit_of_colimit_cocone_of_is_colimit {X Y : C} {t : cocone (pair X Y)} + (ht : is_colimit t) : (binary_bicone.of_colimit_cocone ht).is_bilimit := +is_binary_bilimit_of_is_colimit (binary_bicone.of_colimit_cocone ht) $ + is_colimit.of_iso_colimit ht $ cocones.ext (iso.refl _) $ λ j, by { rcases j with ⟨⟨⟩⟩, tidy } + +/-- In a preadditive category, if the coproduct of `X` and `Y` exists, then the + binary biproduct of `X` and `Y` exists. -/ +lemma has_binary_biproduct.of_has_binary_coproduct (X Y : C) [has_binary_coproduct X Y] : + has_binary_biproduct X Y := +has_binary_biproduct.mk +{ bicone := _, + is_bilimit := binary_bicone_is_bilimit_of_colimit_cocone_of_is_colimit (colimit.is_colimit _) } + +/-- In a preadditive category, if all binary coproducts exist, then all binary biproducts exist. -/ +lemma has_binary_biproducts.of_has_binary_coproducts [has_binary_coproducts C] : + has_binary_biproducts C := +{ has_binary_biproduct := λ X Y, has_binary_biproduct.of_has_binary_coproduct X Y, } + section -variables [preadditive.{v} C] [has_binary_biproducts.{v} C] +variables {X Y : C} [has_binary_biproduct X Y] + +/-- +In any preadditive category, any binary biproduct satsifies +`biprod.fst ≫ biprod.inl + biprod.snd ≫ biprod.inr = 𝟙 (X ⊞ Y)`. +-/ +@[simp] lemma biprod.total : biprod.fst ≫ biprod.inl + biprod.snd ≫ biprod.inr = 𝟙 (X ⊞ Y) := +begin + ext; simp [add_comp], +end + +lemma biprod.lift_eq {T : C} {f : T ⟶ X} {g : T ⟶ Y} : + biprod.lift f g = f ≫ biprod.inl + g ≫ biprod.inr := +begin + ext; simp [add_comp], +end + +lemma biprod.desc_eq {T : C} {f : X ⟶ T} {g : Y ⟶ T} : + biprod.desc f g = biprod.fst ≫ f + biprod.snd ≫ g := +begin + ext; simp [add_comp], +end + +@[simp, reassoc] lemma biprod.lift_desc {T U : C} {f : T ⟶ X} {g : T ⟶ Y} {h : X ⟶ U} {i : Y ⟶ U} : + biprod.lift f g ≫ biprod.desc h i = f ≫ h + g ≫ i := +by simp [biprod.lift_eq, biprod.desc_eq] + +lemma biprod.map_eq [has_binary_biproducts C] {W X Y Z : C} {f : W ⟶ Y} {g : X ⟶ Z} : + biprod.map f g = biprod.fst ≫ f ≫ biprod.inl + biprod.snd ≫ g ≫ biprod.inr := +by apply biprod.hom_ext; apply biprod.hom_ext'; simp + +/-- +Every split mono `f` with a cokernel induces a binary bicone with `f` as its `inl` and +the cokernel map as its `snd`. +We will show in `is_bilimit_binary_bicone_of_split_mono_of_cokernel` that this binary bicone is in +fact already a biproduct. -/ +@[simps] +def binary_bicone_of_is_split_mono_of_cokernel {X Y : C} {f : X ⟶ Y} [is_split_mono f] + {c : cokernel_cofork f} (i : is_colimit c) : binary_bicone X c.X := +{ X := Y, + fst := retraction f, + snd := c.π, + inl := f, + inr := + let c' : cokernel_cofork (𝟙 Y - (𝟙 Y - retraction f ≫ f)) := + cokernel_cofork.of_π (cofork.π c) (by simp) in + let i' : is_colimit c' := is_cokernel_epi_comp i (retraction f) (by simp) in + let i'' := is_colimit_cofork_of_cokernel_cofork i' in + (split_epi_of_idempotent_of_is_colimit_cofork C (by simp) i'').section_, + inl_fst' := by simp, + inl_snd' := by simp, + inr_fst' := + begin + dsimp only, + rw [split_epi_of_idempotent_of_is_colimit_cofork_section_, + is_colimit_cofork_of_cokernel_cofork_desc, is_cokernel_epi_comp_desc], + dsimp only [cokernel_cofork_of_cofork_of_π], + letI := epi_of_is_colimit_cofork i, + apply zero_of_epi_comp c.π, + simp only [sub_comp, comp_sub, category.comp_id, category.assoc, is_split_mono.id, sub_self, + cofork.is_colimit.π_desc_assoc, cokernel_cofork.π_of_π, is_split_mono.id_assoc], + apply sub_eq_zero_of_eq, + apply category.id_comp + end, + inr_snd' := by apply split_epi.id } + +/-- The bicone constructed in `binary_bicone_of_split_mono_of_cokernel` is a bilimit. +This is a version of the splitting lemma that holds in all preadditive categories. -/ +def is_bilimit_binary_bicone_of_is_split_mono_of_cokernel {X Y : C} {f : X ⟶ Y} [is_split_mono f] + {c : cokernel_cofork f} (i : is_colimit c) : + (binary_bicone_of_is_split_mono_of_cokernel i).is_bilimit := +is_binary_bilimit_of_total _ +begin + simp only [binary_bicone_of_is_split_mono_of_cokernel_fst, + binary_bicone_of_is_split_mono_of_cokernel_inr, binary_bicone_of_is_split_mono_of_cokernel_snd, + split_epi_of_idempotent_of_is_colimit_cofork_section_], + dsimp only [binary_bicone_of_is_split_mono_of_cokernel_X], + rw [is_colimit_cofork_of_cokernel_cofork_desc, is_cokernel_epi_comp_desc], + simp only [binary_bicone_of_is_split_mono_of_cokernel_inl, cofork.is_colimit.π_desc, + cokernel_cofork_of_cofork_π, cofork.π_of_π, add_sub_cancel'_right] +end + +/-- If `b` is a binary bicone such that `b.inl` is a kernel of `b.snd`, then `b` is a bilimit + bicone. -/ +def binary_bicone.is_bilimit_of_kernel_inl {X Y : C} (b : binary_bicone X Y) + (hb : is_limit b.snd_kernel_fork) : b.is_bilimit := +is_binary_bilimit_of_is_limit _ $ binary_fan.is_limit.mk _ + (λ T f g, f ≫ b.inl + g ≫ b.inr) (λ T f g, by simp) (λ T f g, by simp) $ λ T f g m h₁ h₂, + begin + have h₁' : (m - (f ≫ b.inl + g ≫ b.inr)) ≫ b.fst = 0 := by simpa using sub_eq_zero.2 h₁, + have h₂' : (m - (f ≫ b.inl + g ≫ b.inr)) ≫ b.snd = 0 := by simpa using sub_eq_zero.2 h₂, + obtain ⟨q : T ⟶ X, hq : q ≫ b.inl = m - (f ≫ b.inl + g ≫ b.inr)⟩ := + kernel_fork.is_limit.lift' hb _ h₂', + rw [←sub_eq_zero, ←hq, ←category.comp_id q, ←b.inl_fst, ←category.assoc, hq, h₁', zero_comp] + end + +/-- If `b` is a binary bicone such that `b.inr` is a kernel of `b.fst`, then `b` is a bilimit + bicone. -/ +def binary_bicone.is_bilimit_of_kernel_inr {X Y : C} (b : binary_bicone X Y) + (hb : is_limit b.fst_kernel_fork) : b.is_bilimit := +is_binary_bilimit_of_is_limit _ $ binary_fan.is_limit.mk _ + (λ T f g, f ≫ b.inl + g ≫ b.inr) (λ t f g, by simp) (λ t f g, by simp) $ λ T f g m h₁ h₂, + begin + have h₁' : (m - (f ≫ b.inl + g ≫ b.inr)) ≫ b.fst = 0 := by simpa using sub_eq_zero.2 h₁, + have h₂' : (m - (f ≫ b.inl + g ≫ b.inr)) ≫ b.snd = 0 := by simpa using sub_eq_zero.2 h₂, + obtain ⟨q : T ⟶ Y, hq : q ≫ b.inr = m - (f ≫ b.inl + g ≫ b.inr)⟩ := + kernel_fork.is_limit.lift' hb _ h₁', + rw [←sub_eq_zero, ←hq, ←category.comp_id q, ←b.inr_snd, ←category.assoc, hq, h₂', zero_comp] + end + +/-- If `b` is a binary bicone such that `b.fst` is a cokernel of `b.inr`, then `b` is a bilimit + bicone. -/ +def binary_bicone.is_bilimit_of_cokernel_fst {X Y : C} (b : binary_bicone X Y) + (hb : is_colimit b.inr_cokernel_cofork) : b.is_bilimit := +is_binary_bilimit_of_is_colimit _ $ binary_cofan.is_colimit.mk _ + (λ T f g, b.fst ≫ f + b.snd ≫ g) (λ T f g, by simp) (λ T f g, by simp) $ λ T f g m h₁ h₂, + begin + have h₁' : b.inl ≫ (m - (b.fst ≫ f + b.snd ≫ g)) = 0 := by simpa using sub_eq_zero.2 h₁, + have h₂' : b.inr ≫ (m - (b.fst ≫ f + b.snd ≫ g)) = 0 := by simpa using sub_eq_zero.2 h₂, + obtain ⟨q : X ⟶ T, hq : b.fst ≫ q = m - (b.fst ≫ f + b.snd ≫ g)⟩ := + cokernel_cofork.is_colimit.desc' hb _ h₂', + rw [←sub_eq_zero, ←hq, ←category.id_comp q, ←b.inl_fst, category.assoc, hq, h₁', comp_zero] + end + +/-- If `b` is a binary bicone such that `b.snd` is a cokernel of `b.inl`, then `b` is a bilimit + bicone. -/ +def binary_bicone.is_bilimit_of_cokernel_snd {X Y : C} (b : binary_bicone X Y) + (hb : is_colimit b.inl_cokernel_cofork) : b.is_bilimit := +is_binary_bilimit_of_is_colimit _ $ binary_cofan.is_colimit.mk _ + (λ T f g, b.fst ≫ f + b.snd ≫ g) (λ T f g, by simp) (λ T f g, by simp) $ λ T f g m h₁ h₂, + begin + have h₁' : b.inl ≫ (m - (b.fst ≫ f + b.snd ≫ g)) = 0 := by simpa using sub_eq_zero.2 h₁, + have h₂' : b.inr ≫ (m - (b.fst ≫ f + b.snd ≫ g)) = 0 := by simpa using sub_eq_zero.2 h₂, + obtain ⟨q : Y ⟶ T, hq : b.snd ≫ q = m - (b.fst ≫ f + b.snd ≫ g)⟩ := + cokernel_cofork.is_colimit.desc' hb _ h₁', + rw [←sub_eq_zero, ←hq, ←category.id_comp q, ←b.inr_snd, category.assoc, hq, h₂', comp_zero] + end + +/-- +Every split epi `f` with a kernel induces a binary bicone with `f` as its `snd` and +the kernel map as its `inl`. +We will show in `binary_bicone_of_is_split_mono_of_cokernel` that this binary bicone is in fact +already a biproduct. -/ +@[simps] +def binary_bicone_of_is_split_epi_of_kernel {X Y : C} {f : X ⟶ Y} [is_split_epi f] + {c : kernel_fork f} (i : is_limit c) : binary_bicone c.X Y := +{ X := X, + fst := + let c' : kernel_fork (𝟙 X - (𝟙 X - f ≫ section_ f)) := + kernel_fork.of_ι (fork.ι c) (by simp) in + let i' : is_limit c' := is_kernel_comp_mono i (section_ f) (by simp) in + let i'' := is_limit_fork_of_kernel_fork i' in + (split_mono_of_idempotent_of_is_limit_fork C (by simp) i'').retraction, + snd := f, + inl := c.ι, + inr := section_ f, + inl_fst' := by apply split_mono.id, + inl_snd' := by simp, + inr_fst' := + begin + dsimp only, + rw [split_mono_of_idempotent_of_is_limit_fork_retraction, + is_limit_fork_of_kernel_fork_lift, is_kernel_comp_mono_lift], + dsimp only [kernel_fork_of_fork_ι], + letI := mono_of_is_limit_fork i, + apply zero_of_comp_mono c.ι, + simp only [comp_sub, category.comp_id, category.assoc, sub_self, fork.is_limit.lift_ι, + fork.ι_of_ι, is_split_epi.id_assoc] + end, + inr_snd' := by simp } + +/-- The bicone constructed in `binary_bicone_of_is_split_epi_of_kernel` is a bilimit. +This is a version of the splitting lemma that holds in all preadditive categories. -/ +def is_bilimit_binary_bicone_of_is_split_epi_of_kernel {X Y : C} {f : X ⟶ Y} [is_split_epi f] + {c : kernel_fork f} (i : is_limit c) : + (binary_bicone_of_is_split_epi_of_kernel i).is_bilimit := +binary_bicone.is_bilimit_of_kernel_inl _ $ i.of_iso_limit $ fork.ext (iso.refl _) (by simp) + +end + +section +variables {X Y : C} (f g : X ⟶ Y) + +/-- The existence of binary biproducts implies that there is at most one preadditive structure. -/ +lemma biprod.add_eq_lift_id_desc [has_binary_biproduct X X] : + f + g = biprod.lift (𝟙 X) (𝟙 X) ≫ biprod.desc f g := +by simp + +/-- The existence of binary biproducts implies that there is at most one preadditive structure. -/ +lemma biprod.add_eq_lift_desc_id [has_binary_biproduct Y Y] : + f + g = biprod.lift f g ≫ biprod.desc (𝟙 Y) (𝟙 Y) := +by simp + +end + +end limits + +open category_theory.limits + +section +local attribute [ext] preadditive + +/-- The existence of binary biproducts implies that there is at most one preadditive structure. -/ +instance subsingleton_preadditive_of_has_binary_biproducts {C : Type u} [category.{v} C] + [has_zero_morphisms C] [has_binary_biproducts C] : subsingleton (preadditive C) := +subsingleton.intro $ λ a b, +begin + ext X Y f g, + have h₁ := @biprod.add_eq_lift_id_desc _ _ a _ _ f g + (by convert (infer_instance : has_binary_biproduct X X)), + have h₂ := @biprod.add_eq_lift_id_desc _ _ b _ _ f g + (by convert (infer_instance : has_binary_biproduct X X)), + refine h₁.trans (eq.trans _ h₂.symm), + congr' 2; + exact subsingleton.elim _ _ +end +end + +section +variables [has_binary_biproducts.{v} C] variables {X₁ X₂ Y₁ Y₂ : C} variables (f₁₁ : X₁ ⟶ Y₁) (f₁₂ : X₁ ⟶ Y₂) (f₂₁ : X₂ ⟶ Y₁) (f₂₂ : X₂ ⟶ Y₂) @@ -270,8 +788,6 @@ end end -variables [preadditive.{v} C] - lemma biproduct.column_nonzero_of_iso' {σ τ : Type} [finite τ] {S : σ → C} [has_biproduct S] {T : τ → C} [has_biproduct T] @@ -310,4 +826,215 @@ begin exact nz (t h) end +section preadditive +variables {D : Type.{u'}} [category.{v'} D] [preadditive.{v'} D] +variables (F : C ⥤ D) [preserves_zero_morphisms F] + +namespace limits + +section fintype +variables {J : Type} [fintype J] + +local attribute [tidy] tactic.discrete_cases + +/-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts + preserves finite products. -/ +def preserves_product_of_preserves_biproduct {f : J → C} [preserves_biproduct f F] : + preserves_limit (discrete.functor f) F := +{ preserves := λ c hc, is_limit.of_iso_limit + ((is_limit.postcompose_inv_equiv (discrete.comp_nat_iso_discrete _ _) _).symm + (is_bilimit_of_preserves F (bicone_is_bilimit_of_limit_cone_of_is_limit hc)).is_limit) $ + cones.ext (iso.refl _) (by tidy) } + +section +local attribute [instance] preserves_product_of_preserves_biproduct + +/-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts + preserves finite products. -/ +def preserves_products_of_shape_of_preserves_biproducts_of_shape + [preserves_biproducts_of_shape J F] : preserves_limits_of_shape (discrete J) F := +{ preserves_limit := λ f, preserves_limit_of_iso_diagram _ discrete.nat_iso_functor.symm } + +end + +/-- A functor between preadditive categories that preserves (zero morphisms and) finite products + preserves finite biproducts. -/ +def preserves_biproduct_of_preserves_product {f : J → C} [preserves_limit (discrete.functor f) F] : + preserves_biproduct f F := +{ preserves := λ b hb, is_bilimit_of_is_limit _ $ + is_limit.of_iso_limit ((is_limit.postcompose_hom_equiv (discrete.comp_nat_iso_discrete _ _) + (F.map_cone b.to_cone)).symm (is_limit_of_preserves F hb.is_limit)) $ + cones.ext (iso.refl _) (by tidy) } + +/-- If the (product-like) biproduct comparison for `F` and `f` is a monomorphism, then `F` + preserves the biproduct of `f`. For the converse, see `map_biproduct`. -/ +def preserves_biproduct_of_mono_biproduct_comparison {f : J → C} [has_biproduct f] + [has_biproduct (F.obj ∘ f)] [mono (biproduct_comparison F f)] : preserves_biproduct f F := +begin + have : pi_comparison F f = (F.map_iso (biproduct.iso_product f)).inv ≫ + biproduct_comparison F f ≫ (biproduct.iso_product _).hom, + { ext, convert pi_comparison_comp_π F f j.as; simp [← functor.map_comp] }, + haveI : is_iso (biproduct_comparison F f) := is_iso_of_mono_of_is_split_epi _, + haveI : is_iso (pi_comparison F f) := by { rw this, apply_instance }, + haveI := preserves_product.of_iso_comparison F f, + apply preserves_biproduct_of_preserves_product +end + +/-- If the (coproduct-like) biproduct comparison for `F` and `f` is an epimorphism, then `F` + preserves the biproduct of `F` and `f`. For the converse, see `map_biproduct`. -/ +def preserves_biproduct_of_epi_biproduct_comparison' {f : J → C} [has_biproduct f] + [has_biproduct (F.obj ∘ f)] [epi (biproduct_comparison' F f)] : preserves_biproduct f F := +begin + haveI : epi ((split_epi_biproduct_comparison F f).section_) := by simpa, + haveI : is_iso (biproduct_comparison F f) := is_iso.of_epi_section' + (split_epi_biproduct_comparison F f), + apply preserves_biproduct_of_mono_biproduct_comparison +end + +/-- A functor between preadditive categories that preserves (zero morphisms and) finite products + preserves finite biproducts. -/ +def preserves_biproducts_of_shape_of_preserves_products_of_shape + [preserves_limits_of_shape (discrete J) F] : preserves_biproducts_of_shape J F := +{ preserves := λ f, preserves_biproduct_of_preserves_product F } + +/-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts + preserves finite coproducts. -/ +def preserves_coproduct_of_preserves_biproduct {f : J → C} [preserves_biproduct f F] : + preserves_colimit (discrete.functor f) F := +{ preserves := λ c hc, is_colimit.of_iso_colimit + ((is_colimit.precompose_hom_equiv (discrete.comp_nat_iso_discrete _ _) _).symm + (is_bilimit_of_preserves F + (bicone_is_bilimit_of_colimit_cocone_of_is_colimit hc)).is_colimit) $ + cocones.ext (iso.refl _) (by tidy) } + +section +local attribute [instance] preserves_coproduct_of_preserves_biproduct + +/-- A functor between preadditive categories that preserves (zero morphisms and) finite biproducts + preserves finite coproducts. -/ +def preserves_coproducts_of_shape_of_preserves_biproducts_of_shape + [preserves_biproducts_of_shape J F] : preserves_colimits_of_shape (discrete J) F := +{ preserves_colimit := λ f, preserves_colimit_of_iso_diagram _ discrete.nat_iso_functor.symm } + +end + +/-- A functor between preadditive categories that preserves (zero morphisms and) finite coproducts + preserves finite biproducts. -/ +def preserves_biproduct_of_preserves_coproduct {f : J → C} + [preserves_colimit (discrete.functor f) F] : preserves_biproduct f F := +{ preserves := λ b hb, is_bilimit_of_is_colimit _ $ + is_colimit.of_iso_colimit ((is_colimit.precompose_inv_equiv (discrete.comp_nat_iso_discrete _ _) + (F.map_cocone b.to_cocone)).symm (is_colimit_of_preserves F hb.is_colimit)) $ + cocones.ext (iso.refl _) (by tidy) } + +/-- A functor between preadditive categories that preserves (zero morphisms and) finite coproducts + preserves finite biproducts. -/ +def preserves_biproducts_of_shape_of_preserves_coproducts_of_shape + [preserves_colimits_of_shape (discrete J) F] : preserves_biproducts_of_shape J F := +{ preserves := λ f, preserves_biproduct_of_preserves_coproduct F } + +end fintype + +/-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts + preserves binary products. -/ +def preserves_binary_product_of_preserves_binary_biproduct {X Y : C} + [preserves_binary_biproduct X Y F] : preserves_limit (pair X Y) F := +{ preserves := λ c hc, is_limit.of_iso_limit + ((is_limit.postcompose_inv_equiv (by exact diagram_iso_pair _) _).symm + (is_binary_bilimit_of_preserves F + (binary_bicone_is_bilimit_of_limit_cone_of_is_limit hc)).is_limit) $ + cones.ext (iso.refl _) (λ j, by { rcases j with ⟨⟨⟩⟩, tidy }) } + +section +local attribute [instance] preserves_binary_product_of_preserves_binary_biproduct + +/-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts + preserves binary products. -/ +def preserves_binary_products_of_preserves_binary_biproducts + [preserves_binary_biproducts F] : preserves_limits_of_shape (discrete walking_pair) F := +{ preserves_limit := λ K, preserves_limit_of_iso_diagram _ (diagram_iso_pair _).symm } + +end + +/-- A functor between preadditive categories that preserves (zero morphisms and) binary products + preserves binary biproducts. -/ +def preserves_binary_biproduct_of_preserves_binary_product {X Y : C} + [preserves_limit (pair X Y) F] : preserves_binary_biproduct X Y F := +{ preserves := λ b hb, is_binary_bilimit_of_is_limit _ $ + is_limit.of_iso_limit ((is_limit.postcompose_hom_equiv (by exact diagram_iso_pair _) + (F.map_cone b.to_cone)).symm (is_limit_of_preserves F hb.is_limit)) $ + cones.ext (iso.refl _) (λ j, by { rcases j with ⟨⟨⟩⟩, tidy }) } + +/-- If the (product-like) biproduct comparison for `F`, `X` and `Y` is a monomorphism, then + `F` preserves the biproduct of `X` and `Y`. For the converse, see `map_biprod`. -/ +def preserves_binary_biproduct_of_mono_biprod_comparison {X Y : C} [has_binary_biproduct X Y] + [has_binary_biproduct (F.obj X) (F.obj Y)] [mono (biprod_comparison F X Y)] : + preserves_binary_biproduct X Y F := +begin + have : prod_comparison F X Y = (F.map_iso (biprod.iso_prod X Y)).inv ≫ + biprod_comparison F X Y ≫ (biprod.iso_prod _ _).hom := by { ext; simp [← functor.map_comp] }, + haveI : is_iso (biprod_comparison F X Y) := is_iso_of_mono_of_is_split_epi _, + haveI : is_iso (prod_comparison F X Y) := by { rw this, apply_instance }, + haveI := preserves_limit_pair.of_iso_prod_comparison F X Y, + apply preserves_binary_biproduct_of_preserves_binary_product +end + +/-- If the (coproduct-like) biproduct comparison for `F`, `X` and `Y` is an epimorphism, then + `F` preserves the biproduct of `X` and `Y`. For the converse, see `map_biprod`. -/ +def preserves_binary_biproduct_of_epi_biprod_comparison' {X Y : C} [has_binary_biproduct X Y] + [has_binary_biproduct (F.obj X) (F.obj Y)] [epi (biprod_comparison' F X Y)] : + preserves_binary_biproduct X Y F := +begin + haveI : epi ((split_epi_biprod_comparison F X Y).section_) := by simpa, + haveI : is_iso (biprod_comparison F X Y) := is_iso.of_epi_section' + (split_epi_biprod_comparison F X Y), + apply preserves_binary_biproduct_of_mono_biprod_comparison +end + +/-- A functor between preadditive categories that preserves (zero morphisms and) binary products + preserves binary biproducts. -/ +def preserves_binary_biproducts_of_preserves_binary_products + [preserves_limits_of_shape (discrete walking_pair) F] : preserves_binary_biproducts F := +{ preserves := λ X Y, preserves_binary_biproduct_of_preserves_binary_product F } + +/-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts + preserves binary coproducts. -/ +def preserves_binary_coproduct_of_preserves_binary_biproduct {X Y : C} + [preserves_binary_biproduct X Y F] : preserves_colimit (pair X Y) F := +{ preserves := λ c hc, is_colimit.of_iso_colimit + ((is_colimit.precompose_hom_equiv (by exact diagram_iso_pair _) _).symm + (is_binary_bilimit_of_preserves F + (binary_bicone_is_bilimit_of_colimit_cocone_of_is_colimit hc)).is_colimit) $ + cocones.ext (iso.refl _) (λ j, by { rcases j with ⟨⟨⟩⟩, tidy }) } + +section +local attribute [instance] preserves_binary_coproduct_of_preserves_binary_biproduct + +/-- A functor between preadditive categories that preserves (zero morphisms and) binary biproducts + preserves binary coproducts. -/ +def preserves_binary_coproducts_of_preserves_binary_biproducts + [preserves_binary_biproducts F] : preserves_colimits_of_shape (discrete walking_pair) F := +{ preserves_colimit := λ K, preserves_colimit_of_iso_diagram _ (diagram_iso_pair _).symm } + +end + +/-- A functor between preadditive categories that preserves (zero morphisms and) binary coproducts + preserves binary biproducts. -/ +def preserves_binary_biproduct_of_preserves_binary_coproduct {X Y : C} + [preserves_colimit (pair X Y) F] : preserves_binary_biproduct X Y F := +{ preserves := λ b hb, is_binary_bilimit_of_is_colimit _ $ + is_colimit.of_iso_colimit ((is_colimit.precompose_inv_equiv (by exact diagram_iso_pair _) + (F.map_cocone b.to_cocone)).symm (is_colimit_of_preserves F hb.is_colimit)) $ + cocones.ext (iso.refl _) (λ j, by { rcases j with ⟨⟨⟩⟩, tidy }) } + +/-- A functor between preadditive categories that preserves (zero morphisms and) binary coproducts + preserves binary biproducts. -/ +def preserves_binary_biproducts_of_preserves_binary_coproducts + [preserves_colimits_of_shape (discrete walking_pair) F] : preserves_binary_biproducts F := +{ preserves := λ X Y, preserves_binary_biproduct_of_preserves_binary_coproduct F } + +end limits + +end preadditive + end category_theory diff --git a/src/category_theory/preadditive/hom_orthogonal.lean b/src/category_theory/preadditive/hom_orthogonal.lean index 656a49c4c6c64..e153c5f7f2bc1 100644 --- a/src/category_theory/preadditive/hom_orthogonal.lean +++ b/src/category_theory/preadditive/hom_orthogonal.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import category_theory.linear -import category_theory.limits.shapes.biproducts +import category_theory.preadditive.biproducts import linear_algebra.matrix.invariant_basis_number /-! From 95d4f6586d313c8c28e00f36621d2a6a66893aa6 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Thu, 13 Oct 2022 15:02:28 +0000 Subject: [PATCH 753/828] refactor(topology/sets/compacts): rename projection to is_compact (#16949) In the structures `compacts`, `nonempty_compacts`, `compact_opens`, `positive_compacts` rename the field `compact` to `is_compact`. --- src/measure_theory/measure/content.lean | 4 +- src/measure_theory/measure/haar.lean | 24 +++++------ src/measure_theory/measure/haar_lebesgue.lean | 4 +- src/measure_theory/measure/haar_quotient.lean | 2 +- src/topology/algebra/group.lean | 2 +- src/topology/metric_space/baire.lean | 2 +- src/topology/metric_space/closeds.lean | 8 ++-- .../metric_space/gromov_hausdorff.lean | 8 ++-- src/topology/metric_space/kuratowski.lean | 2 +- src/topology/sets/compacts.lean | 42 +++++++++---------- 10 files changed, 49 insertions(+), 49 deletions(-) diff --git a/src/measure_theory/measure/content.lean b/src/measure_theory/measure/content.lean index db25dc4c278bd..1bd299745cfd6 100644 --- a/src/measure_theory/measure/content.lean +++ b/src/measure_theory/measure/content.lean @@ -157,9 +157,9 @@ begin { intros n s hn ih, rw [finset.sup_insert, finset.sum_insert hn], exact le_trans (μ.sup_le _ _) (add_le_add_left ih _) }}, refine supr₂_le (λ K hK, _), - obtain ⟨t, ht⟩ := K.compact.elim_finite_subcover _ (λ i, (U i).prop) _, swap, + obtain ⟨t, ht⟩ := K.is_compact.elim_finite_subcover _ (λ i, (U i).prop) _, swap, { convert hK, rw [opens.supr_def, subtype.coe_mk] }, - rcases K.compact.finite_compact_cover t (coe ∘ U) (λ i _, (U _).prop) (by simp only [ht]) + rcases K.is_compact.finite_compact_cover t (coe ∘ U) (λ i _, (U _).prop) (by simp only [ht]) with ⟨K', h1K', h2K', h3K'⟩, let L : ℕ → compacts G := λ n, ⟨K' n, h1K' n⟩, convert le_trans (h3 t L) _, diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar.lean index f2d5b992b74f4..50b143bef2667 100644 --- a/src/measure_theory/measure/haar.lean +++ b/src/measure_theory/measure/haar.lean @@ -150,8 +150,8 @@ lemma le_index_mul (K₀ : positive_compacts G) (K : compacts G) {V : set G} (hV : (interior V).nonempty) : index (K : set G) V ≤ index (K : set G) K₀ * index (K₀ : set G) V := begin - obtain ⟨s, h1s, h2s⟩ := index_elim K.compact K₀.interior_nonempty, - obtain ⟨t, h1t, h2t⟩ := index_elim K₀.compact hV, + obtain ⟨s, h1s, h2s⟩ := index_elim K.is_compact K₀.interior_nonempty, + obtain ⟨t, h1t, h2t⟩ := index_elim K₀.is_compact hV, rw [← h2s, ← h2t, mul_comm], refine le_trans _ finset.card_mul_le, apply nat.Inf_le, refine ⟨_, _, rfl⟩, rw [mem_set_of_eq], refine subset.trans h1s _, @@ -169,7 +169,7 @@ begin { rintro ⟨t, h1t, h2t⟩, rw [finset.card_eq_zero] at h2t, subst h2t, obtain ⟨g, hg⟩ := K.interior_nonempty, show g ∈ (∅ : set G), convert h1t (interior_subset hg), symmetry, apply bUnion_empty }, - { exact index_defined K.compact hV } + { exact index_defined K.is_compact hV } end @[to_additive add_index_mono] @@ -300,7 +300,7 @@ by { simp only [prehaar], rw [div_add_div_same], congr', exact_mod_cast index_un lemma is_left_invariant_prehaar {K₀ : positive_compacts G} {U : set G} (hU : (interior U).nonempty) (g : G) (K : compacts G) : prehaar (K₀ : set G) U (K.map _ $ continuous_mul_left g) = prehaar (K₀ : set G) U K := -by simp only [prehaar, compacts.coe_map, is_left_invariant_index K.compact _ hU] +by simp only [prehaar, compacts.coe_map, is_left_invariant_index K.is_compact _ hU] /-! ### Lemmas about `haar_product` @@ -540,9 +540,9 @@ end lemma haar_measure_self {K₀ : positive_compacts G} : haar_measure K₀ K₀ = 1 := begin haveI : locally_compact_space G := K₀.locally_compact_space_of_group, - rw [haar_measure_apply K₀.compact.measurable_set, ennreal.div_self], + rw [haar_measure_apply K₀.is_compact.measurable_set, ennreal.div_self], { rw [← pos_iff_ne_zero], exact haar_content_outer_measure_self_pos }, - { exact (content.outer_measure_lt_top_of_is_compact _ K₀.compact).ne } + { exact (content.outer_measure_lt_top_of_is_compact _ K₀.is_compact).ne } end /-- The Haar measure is regular. -/ @@ -569,7 +569,7 @@ gives finite mass to compact sets and positive mass to nonempty open sets."] instance is_haar_measure_haar_measure (K₀ : positive_compacts G) : is_haar_measure (haar_measure K₀) := begin - apply is_haar_measure_of_is_compact_nonempty_interior (haar_measure K₀) K₀ K₀.compact + apply is_haar_measure_of_is_compact_nonempty_interior (haar_measure K₀) K₀ K₀.is_compact K₀.interior_nonempty, { simp only [haar_measure_self], exact one_ne_zero }, { simp only [haar_measure_self], exact ennreal.coe_ne_top }, @@ -593,9 +593,9 @@ left invariant measure is a scalar multiple of the additive Haar measure. This i than assuming that `μ` is an additive Haar measure (in particular we don't require `μ ≠ 0`)."] theorem haar_measure_unique (μ : measure G) [sigma_finite μ] [is_mul_left_invariant μ] (K₀ : positive_compacts G) : μ = μ K₀ • haar_measure K₀ := -(measure_eq_div_smul μ (haar_measure K₀) K₀.compact.measurable_set +(measure_eq_div_smul μ (haar_measure K₀) K₀.is_compact.measurable_set (measure_pos_of_nonempty_interior _ K₀.interior_nonempty).ne' - K₀.compact.measure_lt_top.ne).trans (by rw [haar_measure_self, div_one]) + K₀.is_compact.measure_lt_top.ne).trans (by rw [haar_measure_self, div_one]) example [locally_compact_space G] (μ : measure G) [is_haar_measure μ] (K₀ : positive_compacts G) : μ = μ K₀.1 • haar_measure K₀ := @@ -617,11 +617,11 @@ theorem is_haar_measure_eq_smul_is_haar_measure begin have K : positive_compacts G := classical.arbitrary _, have νpos : 0 < ν K := measure_pos_of_nonempty_interior _ K.interior_nonempty, - have νne : ν K ≠ ∞ := K.compact.measure_lt_top.ne, + have νne : ν K ≠ ∞ := K.is_compact.measure_lt_top.ne, refine ⟨μ K / ν K, _, _, _⟩, { simp only [νne, (μ.measure_pos_of_nonempty_interior K.interior_nonempty).ne', ne.def, ennreal.div_zero_iff, not_false_iff, or_self] }, - { simp only [div_eq_mul_inv, νpos.ne', (K.compact.measure_lt_top).ne, or_self, + { simp only [div_eq_mul_inv, νpos.ne', (K.is_compact.measure_lt_top).ne, or_self, ennreal.inv_eq_top, with_top.mul_eq_top_iff, ne.def, not_false_iff, and_false, false_and] }, { calc μ = μ K • haar_measure K : haar_measure_unique μ K @@ -721,7 +721,7 @@ begin simp, }, have : c^2 = 1^2 := (ennreal.mul_eq_mul_right (measure_pos_of_nonempty_interior _ K.interior_nonempty).ne' - K.compact.measure_lt_top.ne).1 this, + K.is_compact.measure_lt_top.ne).1 this, have : c = 1 := (ennreal.pow_strict_mono two_ne_zero).injective this, rw [hc, this, one_smul] end diff --git a/src/measure_theory/measure/haar_lebesgue.lean b/src/measure_theory/measure/haar_lebesgue.lean index c84c51cb624f7..1453aef367053 100644 --- a/src/measure_theory/measure/haar_lebesgue.lean +++ b/src/measure_theory/measure/haar_lebesgue.lean @@ -42,7 +42,7 @@ open_locale ennreal pointwise topological_space nnreal /-- The interval `[0,1]` as a compact set with non-empty interior. -/ def topological_space.positive_compacts.Icc01 : positive_compacts ℝ := { carrier := Icc 0 1, - compact' := is_compact_Icc, + is_compact' := is_compact_Icc, interior_nonempty' := by simp_rw [interior_Icc, nonempty_Ioo, zero_lt_one] } universe u @@ -51,7 +51,7 @@ universe u def topological_space.positive_compacts.pi_Icc01 (ι : Type*) [fintype ι] : positive_compacts (ι → ℝ) := { carrier := pi univ (λ i, Icc 0 1), - compact' := is_compact_univ_pi (λ i, is_compact_Icc), + is_compact' := is_compact_univ_pi (λ i, is_compact_Icc), interior_nonempty' := by simp only [interior_pi_set, set.to_finite, interior_Icc, univ_pi_nonempty_iff, nonempty_Ioo, implies_true_iff, zero_lt_one] } diff --git a/src/measure_theory/measure/haar_quotient.lean b/src/measure_theory/measure/haar_quotient.lean index 688188c0760e0..aa703d77da1d7 100644 --- a/src/measure_theory/measure/haar_quotient.lean +++ b/src/measure_theory/measure/haar_quotient.lean @@ -156,7 +156,7 @@ begin h𝓕.is_mul_left_invariant_map, rw [measure.haar_measure_unique (measure.map (quotient_group.mk' Γ) (μ.restrict 𝓕)) K, measure.map_apply meas_π, measure.restrict_apply₀' 𝓕meas, inter_comm], - exact K.compact.measurable_set, + exact K.is_compact.measurable_set, end /-- Given a normal subgroup `Γ` of a topological group `G` with Haar measure `μ`, which is also diff --git a/src/topology/algebra/group.lean b/src/topology/algebra/group.lean index 449af7ee8a905..2b6e8e2453d50 100644 --- a/src/topology/algebra/group.lean +++ b/src/topology/algebra/group.lean @@ -1174,7 +1174,7 @@ begin refine locally_compact_of_compact_nhds (λ x, _), obtain ⟨y, hy⟩ := K.interior_nonempty, let F := homeomorph.mul_left (x * y⁻¹), - refine ⟨F '' K, _, K.compact.image F.continuous⟩, + refine ⟨F '' K, _, K.is_compact.image F.continuous⟩, suffices : F.symm ⁻¹' K ∈ 𝓝 x, by { convert this, apply equiv.image_eq_preimage }, apply continuous_at.preimage_mem_nhds F.symm.continuous.continuous_at, have : F.symm x = y, by simp [F, homeomorph.mul_left_symm], diff --git a/src/topology/metric_space/baire.lean b/src/topology/metric_space/baire.lean index 543c96184f490..cc0417baedb4e 100644 --- a/src/topology/metric_space/baire.lean +++ b/src/topology/metric_space/baire.lean @@ -171,7 +171,7 @@ begin have hK_nonempty : (⋂ n, K n : set α).nonempty, from is_compact.nonempty_Inter_of_sequence_nonempty_compact_closed _ (λ n, (hK_decreasing n).trans (inter_subset_right _ _)) - (λ n, (K n).nonempty) (K 0).compact (λ n, (K n).compact.is_closed), + (λ n, (K n).nonempty) (K 0).is_compact (λ n, (K n).is_compact.is_closed), exact hK_nonempty.mono hK_subset end diff --git a/src/topology/metric_space/closeds.lean b/src/topology/metric_space/closeds.lean index 2655f95e0c8a8..8b4154d495cc3 100644 --- a/src/topology/metric_space/closeds.lean +++ b/src/topology/metric_space/closeds.lean @@ -237,7 +237,7 @@ instance nonempty_compacts.emetric_space : emetric_space (nonempty_compacts α) edist_triangle := λs t u, Hausdorff_edist_triangle, eq_of_edist_eq_zero := λ s t h, nonempty_compacts.ext $ begin have : closure (s : set α) = closure t := Hausdorff_edist_zero_iff_closure_eq_closure.1 h, - rwa [s.compact.is_closed.closure_eq, t.compact.is_closed.closure_eq] at this, + rwa [s.is_compact.is_closed.closure_eq, t.is_compact.is_closed.closure_eq] at this, end } /-- `nonempty_compacts.to_closeds` is a uniform embedding (as it is an isometry) -/ @@ -254,7 +254,7 @@ begin { ext s, refine ⟨_, λ h, ⟨⟨⟨s, h.2⟩, h.1⟩, closeds.ext rfl⟩⟩, rintro ⟨s, hs, rfl⟩, - exact ⟨s.nonempty, s.compact⟩ }, + exact ⟨s.nonempty, s.is_compact⟩ }, rw this, refine is_closed_of_closure_subset (λs hs, ⟨_, _⟩), { -- take a set set t which is nonempty and at a finite distance of s @@ -330,7 +330,7 @@ begin have Fspec : ∀x, F x ∈ s ∧ edist x (F x) < δ/2 := λx, some_spec (Exy x), -- cover `t` with finitely many balls. Their centers form a set `a` - have : totally_bounded (t : set α) := t.compact.totally_bounded, + have : totally_bounded (t : set α) := t.is_compact.totally_bounded, rcases totally_bounded_iff.1 this (δ/2) δpos' with ⟨a, af, ta⟩, -- a : set α, af : a.finite, ta : t ⊆ ⋃ (y : α) (H : y ∈ a), eball y (δ / 2) -- replace each center by a nearby approximation in `s`, giving a new set `b` @@ -392,7 +392,7 @@ variables {α : Type u} [metric_space α] edistance between two such sets is finite. -/ instance nonempty_compacts.metric_space : metric_space (nonempty_compacts α) := emetric_space.to_metric_space $ λ x y, Hausdorff_edist_ne_top_of_nonempty_of_bounded - x.nonempty y.nonempty x.compact.bounded y.compact.bounded + x.nonempty y.nonempty x.is_compact.bounded y.is_compact.bounded /-- The distance on `nonempty_compacts α` is the Hausdorff distance, by construction -/ lemma nonempty_compacts.dist_eq {x y : nonempty_compacts α} : diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index 80d6d2b3311bc..01f8b3a330e73 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -249,7 +249,7 @@ begin have : Φ xX ∈ ↑p := Φrange.subst (mem_range_self _), exact exists_dist_lt_of_Hausdorff_dist_lt this bound (Hausdorff_edist_ne_top_of_nonempty_of_bounded p.nonempty q.nonempty - p.compact.bounded q.compact.bounded) }, + p.is_compact.bounded q.is_compact.bounded) }, rcases this with ⟨y, hy, dy⟩, rcases mem_range.1 hy with ⟨z, hzy⟩, rw ← hzy at dy, @@ -288,7 +288,7 @@ begin { apply mem_union_right, apply mem_range_self } }, refine dist_le_diam_of_mem _ (A _) (A _), rw [Φrange, Ψrange], - exact (p ⊔ q).compact.bounded, + exact (p ⊔ q).is_compact.bounded, end ... ≤ 2 * diam (univ : set X) + 1 + 2 * diam (univ : set Y) : I } }, let Fb := candidates_b_of_candidates F Fgood, @@ -300,7 +300,7 @@ begin have : f (inl x) ∈ ↑p := Φrange.subst (mem_range_self _), rcases exists_dist_lt_of_Hausdorff_dist_lt this hr (Hausdorff_edist_ne_top_of_nonempty_of_bounded p.nonempty q.nonempty - p.compact.bounded q.compact.bounded) + p.is_compact.bounded q.is_compact.bounded) with ⟨z, zq, hz⟩, have : z ∈ range Ψ, by rwa [← Ψrange] at zq, rcases mem_range.1 this with ⟨y, hy⟩, @@ -314,7 +314,7 @@ begin have : f (inr y) ∈ ↑q := Ψrange.subst (mem_range_self _), rcases exists_dist_lt_of_Hausdorff_dist_lt' this hr (Hausdorff_edist_ne_top_of_nonempty_of_bounded p.nonempty q.nonempty - p.compact.bounded q.compact.bounded) + p.is_compact.bounded q.is_compact.bounded) with ⟨z, zq, hz⟩, have : z ∈ range Φ, by rwa [← Φrange] at zq, rcases mem_range.1 this with ⟨x, hx⟩, diff --git a/src/topology/metric_space/kuratowski.lean b/src/topology/metric_space/kuratowski.lean index b7107bedd1658..99fd9607cca37 100644 --- a/src/topology/metric_space/kuratowski.lean +++ b/src/topology/metric_space/kuratowski.lean @@ -113,5 +113,5 @@ def nonempty_compacts.Kuratowski_embedding (α : Type u) [metric_space α] [comp [nonempty α] : nonempty_compacts ℓ_infty_ℝ := { carrier := range (Kuratowski_embedding α), - compact' := is_compact_range (Kuratowski_embedding.isometry α).continuous, + is_compact' := is_compact_range (Kuratowski_embedding.isometry α).continuous, nonempty' := range_nonempty _ } diff --git a/src/topology/sets/compacts.lean b/src/topology/sets/compacts.lean index f4f37e3489b03..0524ac1ac956c 100644 --- a/src/topology/sets/compacts.lean +++ b/src/topology/sets/compacts.lean @@ -32,7 +32,7 @@ namespace topological_space /-- The type of compact sets of a topological space. -/ structure compacts (α : Type*) [topological_space α] := (carrier : set α) -(compact' : is_compact carrier) +(is_compact' : is_compact carrier) namespace compacts variables {α} @@ -41,9 +41,9 @@ instance : set_like (compacts α) α := { coe := compacts.carrier, coe_injective' := λ s t h, by { cases s, cases t, congr' } } -lemma compact (s : compacts α) : is_compact (s : set α) := s.compact' +protected lemma is_compact (s : compacts α) : is_compact (s : set α) := s.is_compact' -instance (K : compacts α) : compact_space K := is_compact_iff_compact_space.1 K.compact +instance (K : compacts α) : compact_space K := is_compact_iff_compact_space.1 K.is_compact instance : can_lift (set α) (compacts α) coe is_compact := { prf := λ K hK, ⟨⟨K, hK⟩, rfl⟩ } @@ -54,8 +54,8 @@ instance : can_lift (set α) (compacts α) coe is_compact := @[simp] lemma carrier_eq_coe (s : compacts α) : s.carrier = s := rfl -instance : has_sup (compacts α) := ⟨λ s t, ⟨s ∪ t, s.compact.union t.compact⟩⟩ -instance [t2_space α] : has_inf (compacts α) := ⟨λ s t, ⟨s ∩ t, s.compact.inter t.compact⟩⟩ +instance : has_sup (compacts α) := ⟨λ s t, ⟨s ∪ t, s.is_compact.union t.is_compact⟩⟩ +instance [t2_space α] : has_inf (compacts α) := ⟨λ s t, ⟨s ∩ t, s.is_compact.inter t.is_compact⟩⟩ instance [compact_space α] : has_top (compacts α) := ⟨⟨univ, compact_univ⟩⟩ instance : has_bot (compacts α) := ⟨⟨∅, is_compact_empty⟩⟩ @@ -108,7 +108,7 @@ congr_fun (image_eq_preimage_of_inverse f.left_inv f.right_inv) K.1 /-- The product of two `compacts`, as a `compacts` in the product space. -/ protected def prod (K : compacts α) (L : compacts β) : compacts (α × β) := { carrier := K ×ˢ L, - compact' := is_compact.prod K.2 L.2 } + is_compact' := is_compact.prod K.2 L.2 } @[simp] lemma coe_prod (K : compacts α) (L : compacts β) : (K.prod L : set (α × β)) = K ×ˢ L := rfl @@ -126,11 +126,11 @@ instance : set_like (nonempty_compacts α) α := { coe := λ s, s.carrier, coe_injective' := λ s t h, by { obtain ⟨⟨_, _⟩, _⟩ := s, obtain ⟨⟨_, _⟩, _⟩ := t, congr' } } -lemma compact (s : nonempty_compacts α) : is_compact (s : set α) := s.compact' +protected lemma is_compact (s : nonempty_compacts α) : is_compact (s : set α) := s.is_compact' protected lemma nonempty (s : nonempty_compacts α) : (s : set α).nonempty := s.nonempty' /-- Reinterpret a nonempty compact as a closed set. -/ -def to_closeds [t2_space α] (s : nonempty_compacts α) : closeds α := ⟨s, s.compact.is_closed⟩ +def to_closeds [t2_space α] (s : nonempty_compacts α) : closeds α := ⟨s, s.is_compact.is_closed⟩ @[ext] protected lemma ext {s t : nonempty_compacts α} (h : (s : set α) = t) : s = t := set_like.ext' h @@ -156,10 +156,10 @@ order_top.lift (coe : _ → set α) (λ _ _, id) rfl /-- In an inhabited space, the type of nonempty compact subsets is also inhabited, with default element the singleton set containing the default element. -/ instance [inhabited α] : inhabited (nonempty_compacts α) := -⟨{ carrier := {default}, compact' := is_compact_singleton, nonempty' := singleton_nonempty _ }⟩ +⟨{ carrier := {default}, is_compact' := is_compact_singleton, nonempty' := singleton_nonempty _ }⟩ instance to_compact_space {s : nonempty_compacts α} : compact_space s := -is_compact_iff_compact_space.1 s.compact +is_compact_iff_compact_space.1 s.is_compact instance to_nonempty {s : nonempty_compacts α} : nonempty s := s.nonempty.to_subtype @@ -187,7 +187,7 @@ instance : set_like (positive_compacts α) α := { coe := λ s, s.carrier, coe_injective' := λ s t h, by { obtain ⟨⟨_, _⟩, _⟩ := s, obtain ⟨⟨_, _⟩, _⟩ := t, congr' } } -lemma compact (s : positive_compacts α) : is_compact (s : set α) := s.compact' +protected lemma is_compact (s : positive_compacts α) : is_compact (s : set α) := s.is_compact' lemma interior_nonempty (s : positive_compacts α) : (interior (s : set α)).nonempty := s.interior_nonempty' @@ -252,7 +252,7 @@ end positive_compacts /-- The type of compact open sets of a topological space. This is useful in non Hausdorff contexts, in particular spectral spaces. -/ structure compact_opens (α : Type*) [topological_space α] extends compacts α := -(open' : is_open carrier) +(is_open' : is_open carrier) namespace compact_opens @@ -260,22 +260,22 @@ instance : set_like (compact_opens α) α := { coe := λ s, s.carrier, coe_injective' := λ s t h, by { obtain ⟨⟨_, _⟩, _⟩ := s, obtain ⟨⟨_, _⟩, _⟩ := t, congr' } } -lemma compact (s : compact_opens α) : is_compact (s : set α) := s.compact' -lemma «open» (s : compact_opens α) : is_open (s : set α) := s.open' +protected lemma is_compact (s : compact_opens α) : is_compact (s : set α) := s.is_compact' +protected lemma is_open (s : compact_opens α) : is_open (s : set α) := s.is_open' /-- Reinterpret a compact open as an open. -/ -@[simps] def to_opens (s : compact_opens α) : opens α := ⟨s, s.open⟩ +@[simps] def to_opens (s : compact_opens α) : opens α := ⟨s, s.is_open⟩ /-- Reinterpret a compact open as a clopen. -/ @[simps] def to_clopens [t2_space α] (s : compact_opens α) : clopens α := -⟨s, s.open, s.compact.is_closed⟩ +⟨s, s.is_open, s.is_compact.is_closed⟩ @[ext] protected lemma ext {s t : compact_opens α} (h : (s : set α) = t) : s = t := set_like.ext' h @[simp] lemma coe_mk (s : compacts α) (h) : (mk s h : set α) = s := rfl instance : has_sup (compact_opens α) := -⟨λ s t, ⟨s.to_compacts ⊔ t.to_compacts, s.open.union t.open⟩⟩ +⟨λ s t, ⟨s.to_compacts ⊔ t.to_compacts, s.is_open.union t.is_open⟩⟩ instance [quasi_separated_space α] : has_inf (compact_opens α) := ⟨λ U V, ⟨⟨(U : set α) ∩ (V : set α), @@ -286,9 +286,9 @@ set_like.coe_injective.semilattice_inf _ (λ _ _, rfl) instance [compact_space α] : has_top (compact_opens α) := ⟨⟨⊤, is_open_univ⟩⟩ instance : has_bot (compact_opens α) := ⟨⟨⊥, is_open_empty⟩⟩ instance [t2_space α] : has_sdiff (compact_opens α) := -⟨λ s t, ⟨⟨s \ t, s.compact.diff t.open⟩, s.open.sdiff t.compact.is_closed⟩⟩ +⟨λ s t, ⟨⟨s \ t, s.is_compact.diff t.is_open⟩, s.is_open.sdiff t.is_compact.is_closed⟩⟩ instance [t2_space α] [compact_space α] : has_compl (compact_opens α) := -⟨λ s, ⟨⟨sᶜ, s.open.is_closed_compl.is_compact⟩, s.compact.is_closed.is_open_compl⟩⟩ +⟨λ s, ⟨⟨sᶜ, s.is_open.is_closed_compl.is_compact⟩, s.is_compact.is_closed.is_open_compl⟩⟩ instance : semilattice_sup (compact_opens α) := set_like.coe_injective.semilattice_sup _ (λ _ _, rfl) @@ -317,7 +317,7 @@ instance : inhabited (compact_opens α) := ⟨⊥⟩ /-- The image of a compact open under a continuous open map. -/ @[simps] def map (f : α → β) (hf : continuous f) (hf' : is_open_map f) (s : compact_opens α) : compact_opens β := -⟨s.to_compacts.map f hf, hf' _ s.open⟩ +⟨s.to_compacts.map f hf, hf' _ s.is_open⟩ @[simp] lemma coe_map {f : α → β} (hf : continuous f) (hf' : is_open_map f) (s : compact_opens α) : (s.map f hf hf' : set β) = f '' s := rfl @@ -325,7 +325,7 @@ instance : inhabited (compact_opens α) := ⟨⊥⟩ /-- The product of two `compact_opens`, as a `compact_opens` in the product space. -/ protected def prod (K : compact_opens α) (L : compact_opens β) : compact_opens (α × β) := -{ open' := K.open.prod L.open, +{ is_open' := K.is_open.prod L.is_open, .. K.to_compacts.prod L.to_compacts } @[simp] lemma coe_prod (K : compact_opens α) (L : compact_opens β) : From 9540b8cfd33b831695561cd9f436730f18e333c8 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Thu, 13 Oct 2022 18:03:29 +0000 Subject: [PATCH 754/828] feat(ring_theory/ideals/operations): add induction lemma (#16945) Add `submodule.smul_induction_on'` which is a dependent version of `submodule.smul_induction_on` Co-authored-by: Antoine Chambert-Loir --- src/ring_theory/ideal/operations.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index 81495a3e4a4ee..c2a4a27542b04 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -95,6 +95,21 @@ begin exact Hb _ hi _ hj, end +/-- Dependent version of `submodule.smul_induction_on`. -/ +@[elab_as_eliminator] +theorem smul_induction_on' {x : M} (hx : x ∈ I • N) + {p : Π x, x ∈ I • N → Prop} + (Hb : ∀ (r : R) (hr : r ∈ I) (n : M) (hn : n ∈ N), + p (r • n) (smul_mem_smul hr hn)) + (H1 : ∀ x hx y hy, p x hx → p y hy → p (x + y) (submodule.add_mem _ ‹_› ‹_›)) : + p x hx := +begin + refine exists.elim _ (λ (h : x ∈ I • N) (H : p x h), H), + exact smul_induction_on hx + (λ a ha x hx, ⟨_, Hb _ ha _ hx⟩) + (λ x y ⟨_, hx⟩ ⟨_, hy⟩, ⟨_, H1 _ _ _ _ hx hy⟩), +end + theorem mem_smul_span_singleton {I : ideal R} {m : M} {x : M} : x ∈ I • span R ({m} : set M) ↔ ∃ y ∈ I, y • m = x := ⟨λ hx, smul_induction_on hx From ec247d43814751ffceb33b758e8820df2372bf6f Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Thu, 13 Oct 2022 20:27:51 +0000 Subject: [PATCH 755/828] feat(measure_theory/group/prod): add properties for right-invariant measures (#16913) * Use `measure_preserving` more in `measure_theory.group.measure`. * Add instance `is_haar_measure.is_inv_invariant` for abelian groups; remove two (unused) lemmas that are a consequence of this. * Improve docstrings in `measure_theory.group.prod` * Sort the file `measure_theory.group.prod` now in sections `left_invariant`, `right_invariant`, `quasi_measure_preserving` * Continuation of #16668 * Part of #16706 --- src/analysis/convolution.lean | 3 +- src/measure_theory/constructions/prod.lean | 12 +- src/measure_theory/group/integration.lean | 11 +- src/measure_theory/group/measure.lean | 30 ++- src/measure_theory/group/prod.lean | 204 ++++++++++++++++----- src/measure_theory/measure/haar.lean | 42 ++--- 6 files changed, 215 insertions(+), 87 deletions(-) diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 88cdb7da1a1ea..1b560154cb2a5 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -223,7 +223,8 @@ variables [sigma_finite μ] [is_add_right_invariant μ] lemma measure_theory.ae_strongly_measurable.convolution_integrand (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) : ae_strongly_measurable (λ p : G × G, L (f p.2) (g (p.1 - p.2))) (μ.prod μ) := -hf.convolution_integrand' L $ hg.mono' (quasi_measure_preserving_sub μ).absolutely_continuous +hf.convolution_integrand' L $ hg.mono' + (quasi_measure_preserving_sub_of_right_invariant μ μ).absolutely_continuous lemma measure_theory.integrable.convolution_integrand (hf : integrable f μ) (hg : integrable g μ) : integrable (λ p : G × G, L (f p.2) (g (p.1 - p.2))) (μ.prod μ) := diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod.lean index a5cb2528f285e..1028df450013a 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod.lean @@ -89,7 +89,7 @@ end variables [measurable_space α] [measurable_space α'] [measurable_space β] [measurable_space β'] variables [measurable_space γ] -variables {μ : measure α} {ν : measure β} {τ : measure γ} +variables {μ μ' : measure α} {ν ν' : measure β} {τ : measure γ} variables [normed_add_comm_group E] /-! ### Measurability @@ -325,7 +325,7 @@ bind μ $ λ x : α, map (prod.mk x) ν instance prod.measure_space {α β} [measure_space α] [measure_space β] : measure_space (α × β) := { volume := volume.prod volume } -variables {μ ν} [sigma_finite ν] +variables [sigma_finite ν] lemma volume_eq_prod (α β) [measure_space α] [measure_space β] : (volume : measure (α × β)) = (volume : measure α).prod (volume : measure β) := @@ -449,6 +449,14 @@ begin eventually_of_forall $ λ x, zero_le _⟩ end +lemma absolutely_continuous.prod [sigma_finite ν'] (h1 : μ ≪ μ') (h2 : ν ≪ ν') : + μ.prod ν ≪ μ'.prod ν' := +begin + refine absolutely_continuous.mk (λ s hs h2s, _), + simp_rw [measure_prod_null hs] at h2s ⊢, + exact (h2s.filter_mono h1.ae_le).mono (λ _ h, h2 h) +end + /-- Note: the converse is not true. For a counterexample, see Walter Rudin *Real and Complex Analysis*, example (c) in section 8.9. -/ lemma ae_ae_of_ae_prod {p : α × β → Prop} (h : ∀ᵐ z ∂μ.prod ν, p z) : diff --git a/src/measure_theory/group/integration.lean b/src/measure_theory/group/integration.lean index fee69426abdf3..04d574f004c59 100644 --- a/src/measure_theory/group/integration.lean +++ b/src/measure_theory/group/integration.lean @@ -140,15 +140,7 @@ variables [has_measurable_inv G] lemma integrable.comp_div_left {f : G → F} [is_inv_invariant μ] [is_mul_left_invariant μ] (hf : integrable f μ) (g : G) : integrable (λ t, f (g / t)) μ := -begin - rw [← map_mul_right_inv_eq_self μ g⁻¹, integrable_map_measure, function.comp], - { simp_rw [div_inv_eq_mul, mul_inv_cancel_left], exact hf }, - { refine ae_strongly_measurable.comp_measurable _ (measurable_id.const_div g), - simp_rw [map_map (measurable_id'.const_div g) (measurable_id'.const_mul g⁻¹).inv, - function.comp, div_inv_eq_mul, mul_inv_cancel_left, map_id'], - exact hf.ae_strongly_measurable }, - { exact (measurable_id'.const_mul g⁻¹).inv.ae_measurable } -end +((measure_preserving_div_left μ g).integrable_comp hf.ae_strongly_measurable).mpr hf @[simp, to_additive] lemma integrable_comp_div_left (f : G → F) @@ -183,7 +175,6 @@ end end smul - section topological_group variables [topological_space G] [group G] [topological_group G] [borel_space G] diff --git a/src/measure_theory/group/measure.lean b/src/measure_theory/group/measure.lean index 0c573bc97d2ea..bd9ac61686f9e 100644 --- a/src/measure_theory/group/measure.lean +++ b/src/measure_theory/group/measure.lean @@ -229,6 +229,13 @@ is_inv_invariant.inv_eq_self lemma map_inv_eq_self (μ : measure G) [is_inv_invariant μ] : map has_inv.inv μ = μ := is_inv_invariant.inv_eq_self +variables [has_measurable_inv G] + +@[to_additive] +lemma measure_preserving_inv (μ : measure G) [is_inv_invariant μ] : + measure_preserving has_inv.inv μ μ := +⟨measurable_inv, map_inv_eq_self μ⟩ + end inv section has_involutive_inv @@ -282,21 +289,28 @@ begin end @[to_additive] -lemma map_div_left_eq_self (μ : measure G) [is_inv_invariant μ] [is_mul_left_invariant μ] (g : G) : - map (λ t, g / t) μ = μ := +lemma measure_preserving_div_left (μ : measure G) [is_inv_invariant μ] [is_mul_left_invariant μ] + (g : G) : measure_preserving (λ t, g / t) μ μ := begin simp_rw [div_eq_mul_inv], - conv_rhs { rw [← map_mul_left_eq_self μ g, ← map_inv_eq_self μ] }, - exact (map_map (measurable_const_mul g) measurable_inv).symm + exact (measure_preserving_mul_left μ g).comp (measure_preserving_inv μ) end +@[to_additive] +lemma map_div_left_eq_self (μ : measure G) [is_inv_invariant μ] [is_mul_left_invariant μ] (g : G) : + map (λ t, g / t) μ = μ := +(measure_preserving_div_left μ g).map_eq + +@[to_additive] +lemma measure_preserving_mul_right_inv (μ : measure G) + [is_inv_invariant μ] [is_mul_left_invariant μ] (g : G) : + measure_preserving (λ t, (g * t)⁻¹) μ μ := +(measure_preserving_inv μ).comp $ measure_preserving_mul_left μ g + @[to_additive] lemma map_mul_right_inv_eq_self (μ : measure G) [is_inv_invariant μ] [is_mul_left_invariant μ] (g : G) : map (λ t, (g * t)⁻¹) μ = μ := -begin - conv_rhs { rw [← map_inv_eq_self μ, ← map_mul_left_eq_self μ g] }, - exact (map_map measurable_inv (measurable_const_mul g)).symm -end +(measure_preserving_mul_right_inv μ g).map_eq @[to_additive] lemma map_div_left_ae (μ : measure G) [is_mul_left_invariant μ] [is_inv_invariant μ] (x : G) : diff --git a/src/measure_theory/group/prod.lean b/src/measure_theory/group/prod.lean index 351eb20e19929..c2ad58d818513 100644 --- a/src/measure_theory/group/prod.lean +++ b/src/measure_theory/group/prod.lean @@ -18,7 +18,7 @@ The idea of the proof is to use the translation invariance of measures to prove for two sets `s` and `t`, where `c` is a constant that does not depend on `μ`. Let `e` and `f` be the characteristic functions of `s` and `t`. Assume that `μ` and `ν` are left-invariant measures. Then the map `(x, y) ↦ (y * x, x⁻¹)` -preserves the measure `μ.prod ν`, which means that +preserves the measure `μ × ν`, which means that ``` ∫ x, ∫ y, h x y ∂ν ∂μ = ∫ x, ∫ y, h (y * x) x⁻¹ ∂ν ∂μ ``` @@ -40,25 +40,34 @@ variables (G : Type*) [measurable_space G] variables [group G] [has_measurable_mul₂ G] variables (μ ν : measure G) [sigma_finite ν] [sigma_finite μ] {s : set G} -/-- The map `(x, y) ↦ (x, xy)` as a `measurable_equiv`. This is a shear mapping. -/ -@[to_additive "The map `(x, y) ↦ (x, x + y)` as a `measurable_equiv`. -This is a shear mapping."] +/-- The map `(x, y) ↦ (x, xy)` as a `measurable_equiv`. -/ +@[to_additive "The map `(x, y) ↦ (x, x + y)` as a `measurable_equiv`."] protected def measurable_equiv.shear_mul_right [has_measurable_inv G] : G × G ≃ᵐ G × G := { measurable_to_fun := measurable_fst.prod_mk measurable_mul, measurable_inv_fun := measurable_fst.prod_mk $ measurable_fst.inv.mul measurable_snd, .. equiv.prod_shear (equiv.refl _) equiv.mul_left } +/-- The map `(x, y) ↦ (x, y / x)` as a `measurable_equiv` with as inverse `(x, y) ↦ (x, yx)` -/ +@[to_additive + "The map `(x, y) ↦ (x, y - x)` as a `measurable_equiv` with as inverse `(x, y) ↦ (x, y + x)`."] +protected def measurable_equiv.shear_div_right [has_measurable_inv G] : G × G ≃ᵐ G × G := +{ measurable_to_fun := measurable_fst.prod_mk $ measurable_snd.div measurable_fst, + measurable_inv_fun := measurable_fst.prod_mk $ measurable_snd.mul measurable_fst, + .. equiv.prod_shear (equiv.refl _) (equiv.div_right) } + variables {G} namespace measure_theory open measure +section left_invariant + /-- The multiplicative shear mapping `(x, y) ↦ (x, xy)` preserves the measure `μ × ν`. This condition is part of the definition of a measurable group in [Halmos, §59]. There, the map in this lemma is called `S`. -/ @[to_additive measure_preserving_prod_add - /-" The shear mapping `(x, y) ↦ (x, x + y)` preserves the measure `μ.prod ν`. "-/] + /-" The shear mapping `(x, y) ↦ (x, x + y)` preserves the measure `μ × ν`. "-/] lemma measure_preserving_prod_mul [is_mul_left_invariant ν] : measure_preserving (λ z : G × G, (z.1, z.1 * z.2)) (μ.prod ν) (μ.prod ν) := (measure_preserving.id μ).skew_product measurable_mul $ @@ -66,7 +75,7 @@ lemma measure_preserving_prod_mul [is_mul_left_invariant ν] : /-- The map `(x, y) ↦ (y, yx)` sends the measure `μ × ν` to `ν × μ`. This is the map `SR` in [Halmos, §59]. -`S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ +`S` is the map `(x, y) ↦ (x, xy)` and `R` is `prod.swap`. -/ @[to_additive measure_preserving_prod_add_swap /-" The map `(x, y) ↦ (y, y + x)` sends the measure `μ × ν` to `ν × μ`. "-/] lemma measure_preserving_prod_mul_swap [is_mul_left_invariant μ] : @@ -88,26 +97,18 @@ variables [has_measurable_inv G] /-- The map `(x, y) ↦ (x, x⁻¹y)` is measure-preserving. This is the function `S⁻¹` in [Halmos, §59], -where `S` is the map in `measure_preserving_prod_mul`. -/ +where `S` is the map `(x, y) ↦ (x, xy)`. -/ @[to_additive measure_preserving_prod_neg_add "The map `(x, y) ↦ (x, - x + y)` is measure-preserving."] lemma measure_preserving_prod_inv_mul [is_mul_left_invariant ν] : measure_preserving (λ z : G × G, (z.1, z.1⁻¹ * z.2)) (μ.prod ν) (μ.prod ν) := (measure_preserving_prod_mul μ ν).symm $ measurable_equiv.shear_mul_right G -@[to_additive] -lemma quasi_measure_preserving_div [is_mul_right_invariant μ] : - quasi_measure_preserving (λ (p : G × G), p.1 / p.2) (μ.prod μ) μ := -begin - refine quasi_measure_preserving.prod_of_left measurable_div (eventually_of_forall $ λ y, _), - exact (measure_preserving_div_right μ y).quasi_measure_preserving -end - variables [is_mul_left_invariant μ] /-- The map `(x, y) ↦ (y, y⁻¹x)` sends `μ × ν` to `ν × μ`. This is the function `S⁻¹R` in [Halmos, §59], -where `S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ +where `S` is the map `(x, y) ↦ (x, xy)` and `R` is `prod.swap`. -/ @[to_additive measure_preserving_prod_neg_add_swap "The map `(x, y) ↦ (y, - y + x)` sends `μ × ν` to `ν × μ`."] lemma measure_preserving_prod_inv_mul_swap : @@ -116,7 +117,7 @@ lemma measure_preserving_prod_inv_mul_swap : /-- The map `(x, y) ↦ (yx, x⁻¹)` is measure-preserving. This is the function `S⁻¹RSR` in [Halmos, §59], -where `S` is the map in `map_prod_mul_eq` and `R` is `prod.swap`. -/ +where `S` is the map `(x, y) ↦ (x, xy)` and `R` is `prod.swap`. -/ @[to_additive measure_preserving_add_prod_neg "The map `(x, y) ↦ (y + x, - x)` is measure-preserving."] lemma measure_preserving_mul_prod_inv [is_mul_left_invariant ν] : @@ -147,15 +148,19 @@ end lemma measure_inv_null : μ s⁻¹ = 0 ↔ μ s = 0 := begin refine ⟨λ hs, _, (quasi_measure_preserving_inv μ).preimage_null⟩, - convert (quasi_measure_preserving_inv μ).preimage_null hs, - exact (inv_inv _).symm + rw [← inv_inv s], + exact (quasi_measure_preserving_inv μ).preimage_null hs end @[to_additive] -lemma absolutely_continuous_map_inv : μ ≪ map has_inv.inv μ := +lemma inv_absolutely_continuous : μ.inv ≪ μ := +(quasi_measure_preserving_inv μ).absolutely_continuous + +@[to_additive] +lemma absolutely_continuous_inv : μ ≪ μ.inv := begin refine absolutely_continuous.mk (λ s hs, _), - simp_rw [map_apply measurable_inv hs, inv_preimage, measure_inv_null, imp_self] + simp_rw [inv_apply μ s, measure_inv_null, imp_self] end @[to_additive] @@ -186,17 +191,6 @@ lemma measure_mul_right_ne_zero (h2s : μ s ≠ 0) (y : G) : μ ((λ x, x * y) ⁻¹' s) ≠ 0 := (not_iff_not_of_iff (measure_mul_right_null μ y)).mpr h2s -/-- A *left*-invariant measure is quasi-preserved by *right*-multiplication. -This should not be confused with `(measure_preserving_mul_right μ g).quasi_measure_preserving`. -/ -@[to_additive /-"A *left*-invariant measure is quasi-preserved by *right*-addition. -This should not be confused with `(measure_preserving_add_right μ g).quasi_measure_preserving`. "-/] -lemma quasi_measure_preserving_mul_right (g : G) : - quasi_measure_preserving (λ h : G, h * g) μ μ := -begin - refine ⟨measurable_mul_const g, absolutely_continuous.mk $ λ s hs, _⟩, - rw [map_apply (measurable_mul_const g) hs, measure_mul_right_null], exact id, -end - @[to_additive] lemma absolutely_continuous_map_mul_right (g : G) : μ ≪ map (* g) μ := begin @@ -204,21 +198,13 @@ begin rw [map_apply (measurable_mul_const g) hs, measure_mul_right_null], exact id end -@[to_additive] lemma quasi_measure_preserving_div_left (g : G) : - quasi_measure_preserving (λ h : G, g / h) μ μ := -begin - simp_rw [div_eq_mul_inv], - exact (measure_preserving_mul_left μ g).quasi_measure_preserving.comp - (quasi_measure_preserving_inv μ) -end - @[to_additive] lemma absolutely_continuous_map_div_left (g : G) : μ ≪ map (λ h, g / h) μ := begin simp_rw [div_eq_mul_inv], rw [← map_map (measurable_const_mul g) measurable_inv], conv_lhs { rw [← map_mul_left_eq_self μ g] }, - exact (absolutely_continuous_map_inv μ).map (measurable_const_mul g) + exact (absolutely_continuous_inv μ).map (measurable_const_mul g) end /-- This is the computation performed in the proof of [Halmos, §60 Th. A]. -/ @@ -339,4 +325,140 @@ begin measure_mul_measure_eq μ ν hs ht h2s h3s, mul_div_assoc, ennreal.mul_div_cancel' h2s h3s] end +end left_invariant + +section right_invariant + +@[to_additive measure_preserving_prod_add_right] +lemma measure_preserving_prod_mul_right [is_mul_right_invariant ν] : + measure_preserving (λ z : G × G, (z.1, z.2 * z.1)) (μ.prod ν) (μ.prod ν) := +(measure_preserving.id μ).skew_product (by exact measurable_snd.mul measurable_fst) $ + filter.eventually_of_forall $ map_mul_right_eq_self ν + +/-- The map `(x, y) ↦ (y, xy)` sends the measure `μ × ν` to `ν × μ`. -/ +@[to_additive measure_preserving_prod_add_swap_right + /-" The map `(x, y) ↦ (y, x + y)` sends the measure `μ × ν` to `ν × μ`. "-/] +lemma measure_preserving_prod_mul_swap_right [is_mul_right_invariant μ] : + measure_preserving (λ z : G × G, (z.2, z.1 * z.2)) (μ.prod ν) (ν.prod μ) := +(measure_preserving_prod_mul_right ν μ).comp measure_preserving_swap + +/-- The map `(x, y) ↦ (xy, y)` preserves the measure `μ × ν`. -/ +@[to_additive measure_preserving_add_prod + /-" The map `(x, y) ↦ (x + y, y)` preserves the measure `μ × ν`. "-/] +lemma measure_preserving_mul_prod [is_mul_right_invariant μ] : + measure_preserving (λ z : G × G, (z.1 * z.2, z.2)) (μ.prod ν) (μ.prod ν) := +measure_preserving_swap.comp $ by apply measure_preserving_prod_mul_swap_right μ ν + +variables [has_measurable_inv G] + +/-- The map `(x, y) ↦ (x, y / x)` is measure-preserving. -/ +@[to_additive measure_preserving_prod_sub + "The map `(x, y) ↦ (x, y - x)` is measure-preserving."] +lemma measure_preserving_prod_div [is_mul_right_invariant ν] : + measure_preserving (λ z : G × G, (z.1, z.2 / z.1)) (μ.prod ν) (μ.prod ν) := +(measure_preserving_prod_mul_right μ ν).symm (measurable_equiv.shear_div_right G).symm + +/-- The map `(x, y) ↦ (y, x / y)` sends `μ × ν` to `ν × μ`. -/ +@[to_additive measure_preserving_prod_sub_swap + "The map `(x, y) ↦ (y, x - y)` sends `μ × ν` to `ν × μ`."] +lemma measure_preserving_prod_div_swap [is_mul_right_invariant μ] : + measure_preserving (λ z : G × G, (z.2, z.1 / z.2)) (μ.prod ν) (ν.prod μ) := +(measure_preserving_prod_div ν μ).comp measure_preserving_swap + +/-- The map `(x, y) ↦ (x / y, y)` preserves the measure `μ × ν`. -/ +@[to_additive measure_preserving_sub_prod + /-" The map `(x, y) ↦ (x - y, y)` preserves the measure `μ × ν`. "-/] +lemma measure_preserving_div_prod [is_mul_right_invariant μ] : + measure_preserving (λ z : G × G, (z.1 / z.2, z.2)) (μ.prod ν) (μ.prod ν) := +measure_preserving_swap.comp $ by apply measure_preserving_prod_div_swap μ ν + +/-- The map `(x, y) ↦ (xy, x⁻¹)` is measure-preserving. -/ +@[to_additive measure_preserving_add_prod_neg_right + "The map `(x, y) ↦ (x + y, - x)` is measure-preserving."] +lemma measure_preserving_mul_prod_inv_right [is_mul_right_invariant μ] [is_mul_right_invariant ν] : + measure_preserving (λ z : G × G, (z.1 * z.2, z.1⁻¹)) (μ.prod ν) (μ.prod ν) := +begin + convert (measure_preserving_prod_div_swap ν μ).comp + (measure_preserving_prod_mul_swap_right μ ν), + ext1 ⟨x, y⟩, + simp_rw [function.comp_apply, div_mul_eq_div_div_swap, div_self', one_div] +end + +end right_invariant + +section quasi_measure_preserving + +variables [has_measurable_inv G] + +@[to_additive] +lemma quasi_measure_preserving_inv_of_right_invariant [is_mul_right_invariant μ] : + quasi_measure_preserving (has_inv.inv : G → G) μ μ := +begin + rw [← μ.inv_inv], + exact (quasi_measure_preserving_inv μ.inv).mono + (inv_absolutely_continuous μ.inv) (absolutely_continuous_inv μ.inv) +end + +@[to_additive] +lemma quasi_measure_preserving_div_left [is_mul_left_invariant μ] (g : G) : + quasi_measure_preserving (λ h : G, g / h) μ μ := +begin + simp_rw [div_eq_mul_inv], + exact (measure_preserving_mul_left μ g).quasi_measure_preserving.comp + (quasi_measure_preserving_inv μ) +end + +@[to_additive] +lemma quasi_measure_preserving_div_left_of_right_invariant [is_mul_right_invariant μ] (g : G) : + quasi_measure_preserving (λ h : G, g / h) μ μ := +begin + rw [← μ.inv_inv], + exact (quasi_measure_preserving_div_left μ.inv g).mono + (inv_absolutely_continuous μ.inv) (absolutely_continuous_inv μ.inv) +end + +@[to_additive] +lemma quasi_measure_preserving_div_of_right_invariant [is_mul_right_invariant μ] : + quasi_measure_preserving (λ (p : G × G), p.1 / p.2) (μ.prod ν) μ := +begin + refine quasi_measure_preserving.prod_of_left measurable_div (eventually_of_forall $ λ y, _), + exact (measure_preserving_div_right μ y).quasi_measure_preserving +end + +@[to_additive] +lemma quasi_measure_preserving_div [is_mul_left_invariant μ] : + quasi_measure_preserving (λ (p : G × G), p.1 / p.2) (μ.prod ν) μ := +(quasi_measure_preserving_div_of_right_invariant μ.inv ν).mono + ((absolutely_continuous_inv μ).prod absolutely_continuous.rfl) + (inv_absolutely_continuous μ) + +/-- A *left*-invariant measure is quasi-preserved by *right*-multiplication. +This should not be confused with `(measure_preserving_mul_right μ g).quasi_measure_preserving`. -/ +@[to_additive /-"A *left*-invariant measure is quasi-preserved by *right*-addition. +This should not be confused with `(measure_preserving_add_right μ g).quasi_measure_preserving`. "-/] +lemma quasi_measure_preserving_mul_right [is_mul_left_invariant μ] (g : G) : + quasi_measure_preserving (λ h : G, h * g) μ μ := +begin + refine ⟨measurable_mul_const g, absolutely_continuous.mk $ λ s hs, _⟩, + rw [map_apply (measurable_mul_const g) hs, measure_mul_right_null], exact id, +end + +/-- A *right*-invariant measure is quasi-preserved by *left*-multiplication. +This should not be confused with `(measure_preserving_mul_left μ g).quasi_measure_preserving`. -/ +@[to_additive /-"A *right*-invariant measure is quasi-preserved by *left*-addition. +This should not be confused with `(measure_preserving_add_left μ g).quasi_measure_preserving`. "-/] +lemma quasi_measure_preserving_mul_left [is_mul_right_invariant μ] (g : G) : + quasi_measure_preserving (λ h : G, g * h) μ μ := +begin + have := (quasi_measure_preserving_mul_right μ.inv g⁻¹).mono + (inv_absolutely_continuous μ.inv) (absolutely_continuous_inv μ.inv), + rw [μ.inv_inv] at this, + have := (quasi_measure_preserving_inv_of_right_invariant μ).comp + (this.comp (quasi_measure_preserving_inv_of_right_invariant μ)), + simp_rw [function.comp, mul_inv_rev, inv_inv] at this, + exact this +end + +end quasi_measure_preserving + end measure_theory diff --git a/src/measure_theory/measure/haar.lean b/src/measure_theory/measure/haar.lean index 50b143bef2667..c87dc30ad4235 100644 --- a/src/measure_theory/measure/haar.lean +++ b/src/measure_theory/measure/haar.lean @@ -62,11 +62,12 @@ noncomputable theory open set has_inv function topological_space measurable_space open_locale nnreal classical ennreal pointwise topological_space -variables {G : Type*} [group G] - namespace measure_theory namespace measure +section group +variables {G : Type*} [group G] + /-! We put the internal functions in the construction of the Haar measure in a namespace, so that the chosen names don't clash with other declarations. We first define a couple of the functions before proving the properties (that require that `G` @@ -694,17 +695,23 @@ end end second_countable -/-- Any Haar measure is invariant under inversion in a commutative group. -/ -@[to_additive "Any additive Haar measure is invariant under negation in a commutative group."] -lemma map_haar_inv - {G : Type*} [comm_group G] [topological_space G] [topological_group G] [t2_space G] +end group + +section comm_group + +variables {G : Type*} [comm_group G] [topological_space G] [topological_group G] [t2_space G] [measurable_space G] [borel_space G] [locally_compact_space G] [second_countable_topology G] - (μ : measure G) [is_haar_measure μ] : - measure.map has_inv.inv μ = μ := + +/-- Any Haar measure is invariant under inversion in an abelian group. -/ +@[priority 100, to_additive + "Any additive Haar measure is invariant under negation in an abelian group."] +instance is_haar_measure.is_inv_invariant (μ : measure G) [is_haar_measure μ] : + is_inv_invariant μ := begin -- the image measure is a Haar measure. By uniqueness up to multiplication, it is of the form -- `c μ`. Applying again inversion, one gets the measure `c^2 μ`. But since inversion is an -- involution, this is also `μ`. Hence, `c^2 = 1`, which implies `c = 1`. + constructor, haveI : is_haar_measure (measure.map has_inv.inv μ) := is_haar_measure_map μ (mul_equiv.inv G) continuous_inv continuous_inv, obtain ⟨c, cpos, clt, hc⟩ : ∃ (c : ℝ≥0∞), (c ≠ 0) ∧ (c ≠ ∞) ∧ (measure.map has_inv.inv μ = c • μ) @@ -723,25 +730,10 @@ begin (ennreal.mul_eq_mul_right (measure_pos_of_nonempty_interior _ K.interior_nonempty).ne' K.is_compact.measure_lt_top.ne).1 this, have : c = 1 := (ennreal.pow_strict_mono two_ne_zero).injective this, - rw [hc, this, one_smul] + rw [measure.inv, hc, this, one_smul] end -@[simp, to_additive] lemma haar_preimage_inv - {G : Type*} [comm_group G] [topological_space G] [topological_group G] [t2_space G] - [measurable_space G] [borel_space G] [locally_compact_space G] [second_countable_topology G] - (μ : measure G) [is_haar_measure μ] (s : set G) : - μ (s⁻¹) = μ s := -calc μ (s⁻¹) = measure.map (has_inv.inv) μ s : - ((homeomorph.inv G).to_measurable_equiv.map_apply s).symm -... = μ s : by rw map_haar_inv - -@[to_additive] -lemma measure_preserving_inv - {G : Type*} [comm_group G] [topological_space G] [topological_group G] [t2_space G] - [measurable_space G] [borel_space G] [locally_compact_space G] [second_countable_topology G] - (μ : measure G) [is_haar_measure μ] : - measure_preserving has_inv.inv μ μ := -⟨measurable_inv, map_haar_inv μ⟩ +end comm_group end measure end measure_theory From a08db7e9ea8f4ca1efea31b91960cef03e29a71a Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 13 Oct 2022 20:27:52 +0000 Subject: [PATCH 756/828] feat(analysis/inner_product_space/two_dim): new file (#16928) This file defines constructions specific to the geometry of an oriented two-dimensional real inner product space. Main declarations: * `orientation.area_form` * `orientation.right_angle_rotation` * `orientation.kahler` (renaming suggestions are welcome) [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/adding.20angles/near/299839950) --- src/analysis/inner_product_space/two_dim.lean | 454 ++++++++++++++++++ .../normed_space/linear_isometry.lean | 18 + 2 files changed, 472 insertions(+) create mode 100644 src/analysis/inner_product_space/two_dim.lean diff --git a/src/analysis/inner_product_space/two_dim.lean b/src/analysis/inner_product_space/two_dim.lean new file mode 100644 index 0000000000000..ea78c1a6ff3cf --- /dev/null +++ b/src/analysis/inner_product_space/two_dim.lean @@ -0,0 +1,454 @@ +/- +Copyright (c) 2022 Heather Macbeth. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Heather Macbeth +-/ +import analysis.inner_product_space.dual +import analysis.inner_product_space.orientation +import tactic.linear_combination + +/-! +# Oriented two-dimensional real inner product spaces + +This file defines constructions specific to the geometry of an oriented two-dimensional real inner +product space `E`. + +## Main declarations + +* `orientation.area_form`: an antisymmetric bilinear form `E →ₗ[ℝ] E →ₗ[ℝ] ℝ` (usual notation `ω`). + Morally, when `ω` is evaluated on two vectors, it gives the oriented area of the parallelogram + they span. (But mathlib does not yet have a construction of oriented area, and in fact the + construction of oriented area should pass through `ω`.) + +* `orientation.right_angle_rotation`: an isometric automorphism `E ≃ₗᵢ[ℝ] E` (usual notation `J`). + This automorphism squares to -1. TODO: this agrees with the rotation by 90 degrees defined in + `orientation.rotation`. + +* `orientation.basis_right_angle_rotation`: for a nonzero vector `x` in `E`, the basis `![x, J x]` + for `E`. + +* `orientation.kahler`: a complex-valued real-bilinear map `E →ₗ[ℝ] E →ₗ[ℝ] ℂ`. Its real part is the + inner product and its imaginary part is `orientation.area_form`. For vectors `x` and `y` in `E`, + the complex number `o.kahler x y` has modulus `∥x∥ * ∥y∥`. TODO: the argument of `o.kahler x y` is + the oriented angle (`orientation.oangle`) from `x` to `y`. + +## Implementation notes + +Notation `ω` for `orientation.area_form` and `J` for `orientation.right_angle_rotation` should be +defined locally in each file which uses them, since otherwise one would need a more cumbersome +notation which mentions the orientation explicitly (something like `ω[o]`). Write + +``` +local notation `ω` := o.area_form +local notation `J` := o.right_angle_rotation +``` + +-/ + +noncomputable theory + +open_locale real_inner_product_space complex_conjugate +open finite_dimensional + +local attribute [instance] fact_finite_dimensional_of_finrank_eq_succ + +variables {E : Type*} [inner_product_space ℝ E] [fact (finrank ℝ E = 2)] + (o : orientation ℝ E (fin 2)) + +namespace orientation + +include o + +/-- An antisymmetric bilinear form on an oriented real inner product space of dimension 2 (usual +notation `ω`). When evaluated on two vectors, it gives the oriented area of the parallelogram they +span. -/ +@[irreducible] def area_form : E →ₗ[ℝ] E →ₗ[ℝ] ℝ := +begin + let z : alternating_map ℝ E ℝ (fin 0) ≃ₗ[ℝ] ℝ := + alternating_map.const_linear_equiv_of_is_empty.symm, + let y : alternating_map ℝ E ℝ (fin 1) →ₗ[ℝ] E →ₗ[ℝ] ℝ := + (linear_map.llcomp ℝ E (alternating_map ℝ E ℝ (fin 0)) ℝ z) ∘ₗ + alternating_map.curry_left_linear_map, + exact y ∘ₗ (alternating_map.curry_left_linear_map o.volume_form), +end + +omit o + +local notation `ω` := o.area_form + +lemma area_form_to_volume_form (x y : E) : ω x y = o.volume_form ![x, y] := by simp [area_form] + +@[simp] lemma area_form_apply_self (x : E) : ω x x = 0 := +begin + rw area_form_to_volume_form, + refine o.volume_form.map_eq_zero_of_eq ![x, x] _ (_ : (0 : fin 2) ≠ 1), + { simp }, + { norm_num } +end + +lemma area_form_swap (x y : E) : ω x y = - ω y x := +begin + simp only [area_form_to_volume_form], + convert o.volume_form.map_swap ![y, x] (_ : (0 : fin 2) ≠ 1), + { ext i, + fin_cases i; refl }, + { norm_num } +end + +/-- Continuous linear map version of `orientation.area_form`, useful for calculus. -/ +def area_form' : E →L[ℝ] (E →L[ℝ] ℝ) := +((↑(linear_map.to_continuous_linear_map : (E →ₗ[ℝ] ℝ) ≃ₗ[ℝ] (E →L[ℝ] ℝ))) + ∘ₗ o.area_form).to_continuous_linear_map + +@[simp] lemma area_form'_apply (x : E) : + o.area_form' x = (o.area_form x).to_continuous_linear_map := +rfl + +lemma abs_area_form_le (x y : E) : |ω x y| ≤ ∥x∥ * ∥y∥ := +by simpa [area_form_to_volume_form, fin.prod_univ_succ] using o.abs_volume_form_apply_le ![x, y] + +lemma area_form_le (x y : E) : ω x y ≤ ∥x∥ * ∥y∥ := +by simpa [area_form_to_volume_form, fin.prod_univ_succ] using o.volume_form_apply_le ![x, y] + +lemma abs_area_form_of_orthogonal {x y : E} (h : ⟪x, y⟫ = 0) : |ω x y| = ∥x∥ * ∥y∥ := +begin + rw [o.area_form_to_volume_form, o.abs_volume_form_apply_of_pairwise_orthogonal], + { simp [fin.prod_univ_succ] }, + intros i j hij, + fin_cases i; fin_cases j, + { simpa }, + { simpa using h }, + { simpa [real_inner_comm] using h }, + { simpa } +end + +/-- Auxiliary construction for `orientation.right_angle_rotation`, rotation by 90 degrees in an +oriented real inner product space of dimension 2. -/ +@[irreducible] def right_angle_rotation_aux₁ : E →ₗ[ℝ] E := +let to_dual : E ≃ₗ[ℝ] (E →ₗ[ℝ] ℝ) := + (inner_product_space.to_dual ℝ E).to_linear_equiv ≪≫ₗ linear_map.to_continuous_linear_map.symm in +↑to_dual.symm ∘ₗ ω + +@[simp] lemma inner_right_angle_rotation_aux₁_left (x y : E) : + ⟪o.right_angle_rotation_aux₁ x, y⟫ = ω x y := +by simp [right_angle_rotation_aux₁] + +@[simp] lemma inner_right_angle_rotation_aux₁_right (x y : E) : + ⟪x, o.right_angle_rotation_aux₁ y⟫ = - ω x y := +begin + rw real_inner_comm, + simp [o.area_form_swap y x], +end + +/-- Auxiliary construction for `orientation.right_angle_rotation`, rotation by 90 degrees in an +oriented real inner product space of dimension 2. -/ +def right_angle_rotation_aux₂ : E →ₗᵢ[ℝ] E := +{ norm_map' := λ x, begin + dsimp, + refine le_antisymm _ _, + { cases eq_or_lt_of_le (norm_nonneg (o.right_angle_rotation_aux₁ x)) with h h, + { rw ← h, + positivity }, + refine le_of_mul_le_mul_right _ h, + rw [← real_inner_self_eq_norm_mul_norm, o.inner_right_angle_rotation_aux₁_left], + exact o.area_form_le x (o.right_angle_rotation_aux₁ x) }, + { let K : submodule ℝ E := ℝ ∙ x, + haveI : nontrivial Kᗮ, + { apply @finite_dimensional.nontrivial_of_finrank_pos ℝ, + have : finrank ℝ K ≤ finset.card {x}, + { rw ← set.to_finset_singleton, + exact finrank_span_le_card ({x} : set E) }, + have : finset.card {x} = 1 := finset.card_singleton x, + have : finrank ℝ K + finrank ℝ Kᗮ = finrank ℝ E := K.finrank_add_finrank_orthogonal, + have : finrank ℝ E = 2 := fact.out _, + linarith }, + obtain ⟨w, hw₀⟩ : ∃ w : Kᗮ, w ≠ 0 := exists_ne 0, + have hw' : ⟪x, (w:E)⟫ = 0 := inner_right_of_mem_orthogonal_singleton x w.2, -- hw'₀, + have hw : (w:E) ≠ 0 := λ h, hw₀ (submodule.coe_eq_zero.mp h), + refine le_of_mul_le_mul_right _ (by rwa norm_pos_iff : 0 < ∥(w:E)∥), + rw ← o.abs_area_form_of_orthogonal hw', + rw ← o.inner_right_angle_rotation_aux₁_left x w, + exact abs_real_inner_le_norm (o.right_angle_rotation_aux₁ x) w }, + end, + .. o.right_angle_rotation_aux₁ } + +@[simp] lemma right_angle_rotation_aux₁_right_angle_rotation_aux₁ (x : E) : + o.right_angle_rotation_aux₁ (o.right_angle_rotation_aux₁ x) = - x := +begin + apply ext_inner_left ℝ, + intros y, + have : ⟪o.right_angle_rotation_aux₁ y, o.right_angle_rotation_aux₁ x⟫ = ⟪y, x⟫ := + linear_isometry.inner_map_map o.right_angle_rotation_aux₂ y x, + rw [o.inner_right_angle_rotation_aux₁_right, ← o.inner_right_angle_rotation_aux₁_left, this, + inner_neg_right], +end + +/-- An isometric automorphism of an oriented real inner product space of dimension 2 (usual notation +`J`). This automorphism squares to -1. We will define rotations in such a way that this +automorphism is equal to rotation by 90 degrees. -/ +@[irreducible] def right_angle_rotation : E ≃ₗᵢ[ℝ] E := +linear_isometry_equiv.of_linear_isometry + o.right_angle_rotation_aux₂ + (-o.right_angle_rotation_aux₁) + (by ext; simp [right_angle_rotation_aux₂]) + (by ext; simp [right_angle_rotation_aux₂]) + +local notation `J` := o.right_angle_rotation + +@[simp] lemma inner_right_angle_rotation_left (x y : E) : ⟪J x, y⟫ = ω x y := +begin + rw right_angle_rotation, + exact o.inner_right_angle_rotation_aux₁_left x y +end + +@[simp] lemma inner_right_angle_rotation_right (x y : E) : ⟪x, J y⟫ = - ω x y := +begin + rw right_angle_rotation, + exact o.inner_right_angle_rotation_aux₁_right x y +end + +@[simp] lemma right_angle_rotation_right_angle_rotation (x : E) : J (J x) = - x := +begin + rw right_angle_rotation, + exact o.right_angle_rotation_aux₁_right_angle_rotation_aux₁ x +end + +@[simp] lemma right_angle_rotation_symm : + linear_isometry_equiv.symm J = linear_isometry_equiv.trans J (linear_isometry_equiv.neg ℝ) := +begin + rw right_angle_rotation, + exact linear_isometry_equiv.to_linear_isometry_injective rfl +end + +@[simp] lemma inner_right_angle_rotation_self (x : E) : ⟪J x, x⟫ = 0 := by simp + +lemma inner_right_angle_rotation_swap (x y : E) : ⟪x, J y⟫ = - ⟪J x, y⟫ := by simp + +lemma inner_right_angle_rotation_swap' (x y : E) : ⟪J x, y⟫ = - ⟪x, J y⟫ := +by simp [o.inner_right_angle_rotation_swap x y] + +lemma inner_comp_right_angle_rotation (x y : E) : ⟪J x, J y⟫ = ⟪x, y⟫ := +linear_isometry_equiv.inner_map_map J x y + +@[simp] lemma area_form_right_angle_rotation_left (x y : E) : ω (J x) y = - ⟪x, y⟫ := +by rw [← o.inner_comp_right_angle_rotation, o.inner_right_angle_rotation_right, neg_neg] + +@[simp] lemma area_form_right_angle_rotation_right (x y : E) : ω x (J y) = ⟪x, y⟫ := +by rw [← o.inner_right_angle_rotation_left, o.inner_comp_right_angle_rotation] + +@[simp] lemma area_form_comp_right_angle_rotation (x y : E) : ω (J x) (J y) = ω x y := +by simp + +@[simp] lemma right_angle_rotation_trans_right_angle_rotation : + linear_isometry_equiv.trans J J = linear_isometry_equiv.neg ℝ := +by ext; simp + +/-- For a nonzero vector `x` in an oriented two-dimensional real inner product space `E`, +`![x, J x]` forms an (orthogonal) basis for `E`. -/ +def basis_right_angle_rotation (x : E) (hx : x ≠ 0) : basis (fin 2) ℝ E := +@basis_of_linear_independent_of_card_eq_finrank _ _ _ _ _ _ _ _ ![x, J x] +(linear_independent_of_ne_zero_of_inner_eq_zero (λ i, by { fin_cases i; simp [hx] }) + begin + intros i j hij, + fin_cases i; fin_cases j, + { simpa }, + { simp }, + { simp }, + { simpa } + end) +(fact.out (finrank ℝ E = 2)).symm + +@[simp] lemma coe_basis_right_angle_rotation (x : E) (hx : x ≠ 0) : + ⇑(o.basis_right_angle_rotation x hx) = ![x, J x] := +coe_basis_of_linear_independent_of_card_eq_finrank _ _ + +/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ∥a∥ ^ 2 * ⟪x, y⟫`. (See +`orientation.inner_mul_inner_add_area_form_mul_area_form` for the "applied" form.)-/ +lemma inner_mul_inner_add_area_form_mul_area_form' (a x : E) : + ⟪a, x⟫ • @innerₛₗ ℝ _ _ _ a + ω a x • ω a = ∥a∥ ^ 2 • @innerₛₗ ℝ _ _ _ x := +begin + by_cases ha : a = 0, + { simp [ha] }, + apply (o.basis_right_angle_rotation a ha).ext, + intros i, + fin_cases i, + { simp only [real_inner_self_eq_norm_sq, algebra.id.smul_eq_mul, innerₛₗ_apply, + linear_map.smul_apply, linear_map.add_apply, matrix.cons_val_zero, + o.coe_basis_right_angle_rotation, o.area_form_apply_self, real_inner_comm], + ring }, + { simp only [real_inner_self_eq_norm_sq, algebra.id.smul_eq_mul, innerₛₗ_apply, + linear_map.smul_apply, neg_inj, linear_map.add_apply, matrix.cons_val_one, matrix.head_cons, + o.coe_basis_right_angle_rotation, o.area_form_right_angle_rotation_right, + o.area_form_apply_self, o.inner_right_angle_rotation_right], + rw o.area_form_swap, + ring, } +end + +/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ∥a∥ ^ 2 * ⟪x, y⟫`. -/ +lemma inner_mul_inner_add_area_form_mul_area_form (a x y : E) : + ⟪a, x⟫ * ⟪a, y⟫ + ω a x * ω a y = ∥a∥ ^ 2 * ⟪x, y⟫ := +congr_arg (λ f : E →ₗ[ℝ] ℝ, f y) (o.inner_mul_inner_add_area_form_mul_area_form' a x) + +lemma inner_sq_add_area_form_sq (a b : E) : ⟪a, b⟫ ^ 2 + ω a b ^ 2 = ∥a∥ ^ 2 * ∥b∥ ^ 2 := +by simpa [sq, real_inner_self_eq_norm_sq] using o.inner_mul_inner_add_area_form_mul_area_form a b b + +/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ∥a∥ ^ 2 * ω x y`. (See +`orientation.inner_mul_area_form_sub` for the "applied" form.)-/ +lemma inner_mul_area_form_sub' (a x : E) : + ⟪a, x⟫ • ω a - ω a x • @innerₛₗ ℝ _ _ _ a = ∥a∥ ^ 2 • ω x := +begin + by_cases ha : a = 0, + { simp [ha] }, + apply (o.basis_right_angle_rotation a ha).ext, + intros i, + fin_cases i, + { simp only [o.coe_basis_right_angle_rotation, o.area_form_apply_self, o.area_form_swap a x, + real_inner_self_eq_norm_sq, algebra.id.smul_eq_mul, innerₛₗ_apply, linear_map.sub_apply, + linear_map.smul_apply, matrix.cons_val_zero], + ring }, + { simp only [o.area_form_right_angle_rotation_right, o.area_form_apply_self, + o.coe_basis_right_angle_rotation, o.inner_right_angle_rotation_right, + real_inner_self_eq_norm_sq, real_inner_comm, algebra.id.smul_eq_mul, innerₛₗ_apply, + linear_map.smul_apply, linear_map.sub_apply, matrix.cons_val_one, matrix.head_cons], + ring}, +end + +/-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ∥a∥ ^ 2 * ω x y`. -/ +lemma inner_mul_area_form_sub (a x y : E) : ⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ∥a∥ ^ 2 * ω x y := +congr_arg (λ f : E →ₗ[ℝ] ℝ, f y) (o.inner_mul_area_form_sub' a x) + +lemma nonneg_inner_and_area_form_eq_zero_iff_same_ray (x y : E) : + 0 ≤ ⟪x, y⟫ ∧ ω x y = 0 ↔ same_ray ℝ x y := +begin + by_cases hx : x = 0, + { simp [hx] }, + split, + { let a : ℝ := (o.basis_right_angle_rotation x hx).repr y 0, + let b : ℝ := (o.basis_right_angle_rotation x hx).repr y 1, + suffices : 0 ≤ a * ∥x∥ ^ 2 ∧ b * ∥x∥ ^ 2 = 0 → same_ray ℝ x (a • x + b • J x), + { -- TODO trace the `dsimp` lemmas in this block to make a single `simp only` + rw ← (o.basis_right_angle_rotation x hx).sum_repr y, + simp only [fin.sum_univ_succ, coe_basis_right_angle_rotation], + dsimp, + simp only [o.area_form_apply_self, map_smul, map_add, map_zero, inner_smul_left, + inner_smul_right, inner_add_left, inner_add_right, inner_zero_right, linear_map.add_apply, + matrix.cons_val_one], + dsimp, + simp only [o.area_form_right_angle_rotation_right, mul_zero, add_zero, zero_add, neg_zero, + o.inner_right_angle_rotation_right, o.area_form_apply_self, real_inner_self_eq_norm_sq], + exact this }, + rintros ⟨ha, hb⟩, + have hx' : 0 < ∥x∥ := by simpa using hx, + have ha' : 0 ≤ a := nonneg_of_mul_nonneg_left ha (by positivity), + have hb' : b = 0 := eq_zero_of_ne_zero_of_mul_right_eq_zero (pow_ne_zero 2 hx'.ne') hb, + simpa [hb'] using same_ray_nonneg_smul_right x ha' }, + { intros h, + obtain ⟨r, hr, rfl⟩ := h.exists_nonneg_left hx, + simp only [inner_smul_right, real_inner_self_eq_norm_sq, linear_map.map_smulₛₗ, + area_form_apply_self, algebra.id.smul_eq_mul, mul_zero, eq_self_iff_true, and_true], + positivity }, +end + +/-- A complex-valued real-bilinear map on an oriented real inner product space of dimension 2. Its +real part is the inner product and its imaginary part is `orientation.area_form`. + +TODO On `ℂ` with the standard orientation, `kahler w z = conj w * z`. -/ +def kahler : E →ₗ[ℝ] E →ₗ[ℝ] ℂ := +(linear_map.llcomp ℝ E ℝ ℂ complex.of_real_clm) ∘ₗ (@innerₛₗ ℝ E _ _) ++ (linear_map.llcomp ℝ E ℝ ℂ ((linear_map.lsmul ℝ ℂ).flip complex.I)) ∘ₗ ω + +lemma kahler_apply_apply (x y : E) : o.kahler x y = ⟪x, y⟫ + ω x y • complex.I := rfl + +lemma kahler_swap (x y : E) : o.kahler x y = conj (o.kahler y x) := +begin + simp only [kahler_apply_apply], + rw [real_inner_comm, area_form_swap], + simp, +end + +@[simp] lemma kahler_apply_self (x : E) : o.kahler x x = ∥x∥ ^ 2 := +by simp [kahler_apply_apply, real_inner_self_eq_norm_sq] + +@[simp] lemma kahler_right_angle_rotation_left (x y : E) : + o.kahler (J x) y = - complex.I * o.kahler x y := +begin + simp only [o.area_form_right_angle_rotation_left, o.inner_right_angle_rotation_left, + o.kahler_apply_apply, complex.of_real_neg, complex.real_smul], + linear_combination ω x y * complex.I_sq, +end + +@[simp] lemma kahler_right_angle_rotation_right (x y : E) : + o.kahler x (J y) = complex.I * o.kahler x y := +begin + simp only [o.area_form_right_angle_rotation_right, o.inner_right_angle_rotation_right, + o.kahler_apply_apply, complex.of_real_neg, complex.real_smul], + linear_combination - ω x y * complex.I_sq, +end + +@[simp] lemma kahler_comp_right_angle_rotation (x y : E) : o.kahler (J x) (J y) = o.kahler x y := +begin + simp only [kahler_right_angle_rotation_left, kahler_right_angle_rotation_right], + linear_combination - o.kahler x y * complex.I_sq, +end + +lemma kahler_mul (a x y : E) : o.kahler x a * o.kahler a y = ∥a∥ ^ 2 * o.kahler x y := +begin + transitivity (↑(∥a∥ ^ 2) : ℂ) * o.kahler x y, + { ext, + { simp only [o.kahler_apply_apply, complex.add_im, complex.add_re, complex.I_im, complex.I_re, + complex.mul_im, complex.mul_re, complex.of_real_im, complex.of_real_re, complex.real_smul], + rw [real_inner_comm a x, o.area_form_swap x a], + linear_combination o.inner_mul_inner_add_area_form_mul_area_form a x y }, + { simp only [o.kahler_apply_apply, complex.add_im, complex.add_re, complex.I_im, complex.I_re, + complex.mul_im, complex.mul_re, complex.of_real_im, complex.of_real_re, complex.real_smul], + rw [real_inner_comm a x, o.area_form_swap x a], + linear_combination o.inner_mul_area_form_sub a x y } }, + { norm_cast }, +end + +lemma norm_sq_kahler (x y : E) : complex.norm_sq (o.kahler x y) = ∥x∥ ^ 2 * ∥y∥ ^ 2 := +by simpa [kahler_apply_apply, complex.norm_sq, sq] using o.inner_sq_add_area_form_sq x y + +lemma abs_kahler (x y : E) : complex.abs (o.kahler x y) = ∥x∥ * ∥y∥ := +begin + rw [← sq_eq_sq, complex.sq_abs], + { linear_combination o.norm_sq_kahler x y }, + { positivity }, + { positivity } +end + +lemma norm_kahler (x y : E) : ∥o.kahler x y∥ = ∥x∥ * ∥y∥ := by simpa using o.abs_kahler x y + +lemma eq_zero_or_eq_zero_of_kahler_eq_zero {x y : E} (hx : o.kahler x y = 0) : x = 0 ∨ y = 0 := +begin + have : ∥x∥ * ∥y∥ = 0 := by simpa [hx] using (o.norm_kahler x y).symm, + cases eq_zero_or_eq_zero_of_mul_eq_zero this with h h, + { left, + simpa using h }, + { right, + simpa using h }, +end + +lemma kahler_eq_zero_iff (x y : E) : o.kahler x y = 0 ↔ x = 0 ∨ y = 0 := +begin + refine ⟨o.eq_zero_or_eq_zero_of_kahler_eq_zero, _⟩, + rintros (rfl | rfl); + simp, +end + +lemma kahler_ne_zero {x y : E} (hx : x ≠ 0) (hy : y ≠ 0) : o.kahler x y ≠ 0 := +begin + apply mt o.eq_zero_or_eq_zero_of_kahler_eq_zero, + tauto, +end + +lemma kahler_ne_zero_iff (x y : E) : o.kahler x y ≠ 0 ↔ x ≠ 0 ∧ y ≠ 0 := +begin + refine ⟨_, λ h, o.kahler_ne_zero h.1 h.2⟩, + contrapose, + simp only [not_and_distrib, not_not, kahler_apply_apply, complex.real_smul], + rintros (rfl | rfl); + simp, +end + +end orientation diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index 862806468477f..799ca8e08e90a 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -726,6 +726,24 @@ noncomputable def of_surjective (f : F →ₛₗᵢ[σ₁₂] E₂) ⇑(linear_isometry_equiv.of_surjective f hfr) = f := by { ext, refl } +/-- If a linear isometry has an inverse, it is a linear isometric equivalence. -/ +def of_linear_isometry (f : E →ₛₗᵢ[σ₁₂] E₂) (g : E₂ →ₛₗ[σ₂₁] E) + (h₁ : f.to_linear_map.comp g = linear_map.id) (h₂ : g.comp f.to_linear_map = linear_map.id) : + E ≃ₛₗᵢ[σ₁₂] E₂ := +{ norm_map' := λ x, f.norm_map x, + .. linear_equiv.of_linear f.to_linear_map g h₁ h₂ } + +@[simp] lemma coe_of_linear_isometry (f : E →ₛₗᵢ[σ₁₂] E₂) (g : E₂ →ₛₗ[σ₂₁] E) + (h₁ : f.to_linear_map.comp g = linear_map.id) (h₂ : g.comp f.to_linear_map = linear_map.id) : + (of_linear_isometry f g h₁ h₂ : E → E₂) = (f : E → E₂) := +rfl + +@[simp] lemma coe_of_linear_isometry_symm (f : E →ₛₗᵢ[σ₁₂] E₂) + (g : E₂ →ₛₗ[σ₂₁] E) (h₁ : f.to_linear_map.comp g = linear_map.id) + (h₂ : g.comp f.to_linear_map = linear_map.id) : + ((of_linear_isometry f g h₁ h₂).symm : E₂ → E) = (g : E₂ → E) := +rfl + omit σ₂₁ variables (R) From 1e9b3df5b266b0823ae183932f0f94d09ca19187 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Thu, 13 Oct 2022 20:27:53 +0000 Subject: [PATCH 757/828] refactor(geometry/euclidean/basic): rename `cospherical_subset` (#16947) Rename `cospherical_subset` to `cospherical.subset` to allow use of dot notation. --- src/geometry/euclidean/basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index 3de9a8c31ed93..14862db9dd226 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -757,7 +757,7 @@ lemma sphere.cospherical (s : sphere P) : cospherical (s : set P) := cospherical_iff_exists_sphere.2 ⟨s, set.subset.rfl⟩ /-- A subset of a cospherical set is cospherical. -/ -lemma cospherical_subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (hc : cospherical ps₂) : +lemma cospherical.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (hc : cospherical ps₂) : cospherical ps₁ := begin rcases hc with ⟨c, r, hcr⟩, @@ -878,7 +878,7 @@ structure concyclic (ps : set P) : Prop := /-- A subset of a concyclic set is concyclic. -/ lemma concyclic.subset {ps₁ ps₂ : set P} (hs : ps₁ ⊆ ps₂) (h : concyclic ps₂) : concyclic ps₁ := -⟨cospherical_subset hs h.1, h.2.subset hs⟩ +⟨h.1.subset hs, h.2.subset hs⟩ /-- The empty set is concyclic. -/ lemma concyclic_empty : concyclic (∅ : set P) := From 8f9520eb24f83056021ccea9d1181508a999e46e Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Fri, 14 Oct 2022 05:28:43 +0000 Subject: [PATCH 758/828] chore(scripts): update nolints.txt (#16968) I am happy to remove some nolints for you! --- scripts/nolints.txt | 2 -- 1 file changed, 2 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index aade010d56cee..80d18571f1557 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -67,8 +67,6 @@ apply_nolint monoid.mfoldr.of_free_monoid doc_blame apply_nolint traversable.fold_map doc_blame apply_nolint traversable.foldl doc_blame apply_nolint traversable.foldr doc_blame -apply_nolint traversable.free.map doc_blame -apply_nolint traversable.free.mk doc_blame apply_nolint traversable.length doc_blame apply_nolint traversable.map_fold doc_blame apply_nolint traversable.mfoldl doc_blame From 7a52b9ce97045dd692d8dce3760a9f3b89cc6f5a Mon Sep 17 00:00:00 2001 From: bottine Date: Fri, 14 Oct 2022 06:45:36 +0000 Subject: [PATCH 759/828] feat(category_theory/groupoid/subgroupoid): subgroupoids and basic properties (#16742) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add a definition of (bundled) subgroupoid and prove some of their basic properties. Co-authored-by: Rémi Bottinelli Co-authored-by: bottine --- src/category_theory/endomorphism.lean | 2 + src/category_theory/groupoid/subgroupoid.lean | 354 ++++++++++++++++++ .../groupoid/vertex_group.lean | 81 ++++ 3 files changed, 437 insertions(+) create mode 100644 src/category_theory/groupoid/subgroupoid.lean create mode 100644 src/category_theory/groupoid/vertex_group.lean diff --git a/src/category_theory/endomorphism.lean b/src/category_theory/endomorphism.lean index 55057caac0b4c..c12f58e70f7f0 100644 --- a/src/category_theory/endomorphism.lean +++ b/src/category_theory/endomorphism.lean @@ -117,6 +117,8 @@ simp [flip, (*), monoid.mul, mul_one_class.mul, mul_one_class.one, has_one.one, lemma Aut_mul_def (f g : Aut X) : f * g = g.trans f := rfl +lemma Aut_inv_def (f : Aut X) : f ⁻¹ = f.symm := rfl + /-- Units in the monoid of endomorphisms of an object are (multiplicatively) equivalent to automorphisms of that object. diff --git a/src/category_theory/groupoid/subgroupoid.lean b/src/category_theory/groupoid/subgroupoid.lean new file mode 100644 index 0000000000000..b0dc1e7ae9ff0 --- /dev/null +++ b/src/category_theory/groupoid/subgroupoid.lean @@ -0,0 +1,354 @@ +/- +Copyright (c) 2022 Rémi Bottinelli. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémi Bottinelli, Junyan Xu +-/ +import category_theory.groupoid.vertex_group +import category_theory.groupoid +import algebra.group.defs +import algebra.hom.group +import algebra.hom.equiv +import data.set.lattice +import combinatorics.quiver.connected_component +import group_theory.subgroup.basic +/-! +# Subgroupoid + +This file defines subgroupoids as `structure`s containing the subsets of arrows and their +stability under composition and inversion. +Also defined are + +* containment of subgroupoids is a complete lattice; +* images and preimages of subgroupoids under a functor; +* the notion of normality of subgroupoids and its stability under intersection and preimage; +* compatibility of the above with `groupoid.vertex_group`. + + +## Main definitions + +Given a type `C` with associated `groupoid C` instance. + +* `subgroupoid C` is the type of subgroupoids of `C` +* `subgroupoid.is_normal` is the property that the subgroupoid is stable under conjugation + by arbitrary arrows, _and_ that all identity arrows are contained in the subgroupoid. +* `subgroupoid.comap` is the "preimage" map of subgroupoids along a functor. +* `subgroupoid.map` is the "image" map of subgroupoids along a functor _injective on objects_. +* `subgroupoid.vertex_subgroup` is the subgroup of the `vertex group` at a given vertex `v`, + assuming `v` is contained in the `subgroupoid` (meaning, by definition, that the arrow `𝟙 v` + is contained in the subgroupoid). + +## Implementation details + +The structure of this file is copied from/inspired by `group_theory.subgroup.basic` +and `combinatorics.simple_graph.subgraph`. + +## TODO + +* Equivalent inductive characterization of generated (normal) subgroupoids. +* Characterization of normal subgroupoids as kernels. + +## Tags + +subgroupoid + +-/ + +namespace category_theory + +open set groupoid + +local attribute [protected] category_theory.inv + +universes u v + +variables {C : Type u} [groupoid C] + +/-- +A sugroupoid of `C` consists of a choice of arrows for each pair of vertices, closed +under composition and inverses. +-/ +@[ext] structure subgroupoid (C : Type u) [groupoid C] := +(arrows : ∀ (c d : C), set (c ⟶ d)) +(inv : ∀ {c d} {p : c ⟶ d} (hp : p ∈ arrows c d), + inv p ∈ arrows d c) +(mul : ∀ {c d e} {p} (hp : p ∈ arrows c d) {q} (hq : q ∈ arrows d e), + p ≫ q ∈ arrows c e) + +attribute [protected] subgroupoid.inv subgroupoid.mul + +namespace subgroupoid + +variable (S : subgroupoid C) + +/-- The vertices of `C` on which `S` has non-trivial isotropy -/ +def objs : set C := {c : C | (S.arrows c c).nonempty} + +lemma id_mem_of_nonempty_isotropy (c : C) : + c ∈ objs S → 𝟙 c ∈ S.arrows c c := +begin + rintro ⟨γ,hγ⟩, + convert S.mul hγ (S.inv hγ), + simp only [inv_eq_inv, is_iso.hom_inv_id], +end + +/-- A subgroupoid seen as a quiver on vertex set `C` -/ +def as_wide_quiver : quiver C := ⟨λ c d, subtype $ S.arrows c d⟩ + +/-- The coercion of a subgroupoid as a groupoid -/ +instance coe : groupoid S.objs := +{ hom := λ a b, S.arrows a.val b.val, + id := λ a, ⟨𝟙 a.val, id_mem_of_nonempty_isotropy S a.val a.prop⟩, + comp := λ a b c p q, ⟨p.val ≫ q.val, S.mul p.prop q.prop⟩, + id_comp' := λ a b ⟨p,hp⟩, by simp only [category.id_comp], + comp_id' := λ a b ⟨p,hp⟩, by simp only [category.comp_id], + assoc' := λ a b c d ⟨p,hp⟩ ⟨q,hq⟩ ⟨r,hr⟩, by simp only [category.assoc], + inv := λ a b p, ⟨inv p.val, S.inv p.prop⟩, + inv_comp' := λ a b ⟨p,hp⟩, by simp only [inv_comp], + comp_inv' := λ a b ⟨p,hp⟩, by simp only [comp_inv] } + +/-- The embedding of the coerced subgroupoid to its parent-/ +def hom : S.objs ⥤ C := +{ obj := λ c, c.val, + map := λ c d f, f.val, + map_id' := λ c, rfl, + map_comp' := λ c d e f g, rfl } + +lemma hom.inj_on_objects : function.injective (hom S).obj := +by { rintros ⟨c,hc⟩ ⟨d,hd⟩ hcd, simp only [subtype.mk_eq_mk], exact hcd } + +lemma hom.faithful : + ∀ c d, function.injective (λ (f : c ⟶ d), (hom S).map f) := +by { rintros ⟨c,hc⟩ ⟨d,hd⟩ ⟨f,hf⟩ ⟨g,hg⟩ hfg, simp only [subtype.mk_eq_mk], exact hfg, } + +/-- The subgroup of the vertex group at `c` given by the subgroupoid -/ +def vertex_subgroup {c : C} (hc : c ∈ S.objs) : subgroup (c ⟶ c) := +{ carrier := S.arrows c c, + mul_mem' := λ f g hf hg, S.mul hf hg, + one_mem' := id_mem_of_nonempty_isotropy _ _ hc, + inv_mem' := λ f hf, S.inv hf } + +instance : set_like (subgroupoid C) (Σ (c d : C), c ⟶ d) := +{ coe := λ S, {F | F.2.2 ∈ S.arrows F.1 F.2.1}, + coe_injective' := λ ⟨S, _, _⟩ ⟨T, _, _⟩ h, by { ext c d f, apply set.ext_iff.1 h ⟨c, d, f⟩ } } + +lemma mem_iff (S : subgroupoid C) (F : Σ c d, c ⟶ d) : + F ∈ S ↔ F.2.2 ∈ S.arrows F.1 F.2.1 := iff.rfl + +lemma le_iff (S T : subgroupoid C) : (S ≤ T) ↔ (∀ {c d}, (S.arrows c d) ⊆ (T.arrows c d)) := +by { rw [set_like.le_def, sigma.forall], exact forall_congr (λ c, sigma.forall) } + +instance : has_top (subgroupoid C) := +⟨ { arrows := (λ _ _, set.univ), + mul := by { rintros, trivial, }, + inv := by { rintros, trivial, } } ⟩ + +instance : has_bot (subgroupoid C) := +⟨ { arrows := (λ _ _, ∅), + mul := λ _ _ _ _, false.elim, + inv := λ _ _ _, false.elim } ⟩ + +instance : inhabited (subgroupoid C) := ⟨⊤⟩ + +instance : has_inf (subgroupoid C) := +⟨ λ S T, + { arrows := (λ c d, (S.arrows c d) ∩ (T.arrows c d)), + inv := by { rintros, exact ⟨S.inv hp.1, T.inv hp.2⟩, }, + mul := by { rintros, exact ⟨S.mul hp.1 hq.1, T.mul hp.2 hq.2⟩, } } ⟩ + +instance : has_Inf (subgroupoid C) := +⟨ λ s, + { arrows := λ c d, ⋂ S ∈ s, (subgroupoid.arrows S c d), + inv := by { intros, rw mem_Inter₂ at hp ⊢, exact λ S hS, S.inv (hp S hS) }, + mul := by { intros, rw mem_Inter₂ at hp hq ⊢,exact λ S hS, S.mul (hp S hS) (hq S hS) } } ⟩ + +instance : complete_lattice (subgroupoid C) := +{ bot := (⊥), + bot_le := λ S, empty_subset _, + top := (⊤), + le_top := λ S, subset_univ _, + inf := (⊓), + le_inf := λ R S T RS RT _ pR, ⟨RS pR, RT pR⟩, + inf_le_left := λ R S _, and.left, + inf_le_right := λ R S _, and.right, + .. complete_lattice_of_Inf (subgroupoid C) + begin + refine (λ s, ⟨λ S Ss F, _, λ T Tl F fT, _⟩); + simp only [Inf, mem_iff, mem_Inter], + exacts [λ hp, hp S Ss, λ S Ss, Tl Ss fT], + end } + +lemma le_objs {S T : subgroupoid C} (h : S ≤ T) : S.objs ⊆ T.objs := +λ s ⟨γ, hγ⟩, ⟨γ, @h ⟨s, s, γ⟩ hγ⟩ + +/-- The functor associated to the embedding of subgroupoids -/ +def inclusion {S T : subgroupoid C} (h : S ≤ T) : S.objs ⥤ T.objs := +{ obj := λ s, ⟨s.val, le_objs h s.prop⟩, + map := λ s t f, ⟨f.val, @h ⟨s, t, f.val⟩ f.prop⟩, + map_id' := λ _, rfl, + map_comp' := λ _ _ _ _ _, rfl } + +lemma inclusion_inj_on_objects {S T : subgroupoid C} (h : S ≤ T) : + function.injective (inclusion h).obj := +λ ⟨s,hs⟩ ⟨t,ht⟩, by simpa only [inclusion, subtype.mk_eq_mk] using id + +lemma inclusion_faithful {S T : subgroupoid C} (h : S ≤ T) (s t : S.objs): + function.injective (λ (f : s ⟶ t), (inclusion h).map f) := +λ ⟨f,hf⟩ ⟨g,hg⟩, by { dsimp only [inclusion], simpa only [subtype.mk_eq_mk] using id } + +lemma inclusion_refl {S : subgroupoid C} : inclusion (le_refl S) = 𝟭 S.objs := +functor.hext (λ ⟨s,hs⟩, rfl) (λ ⟨s,hs⟩ ⟨t,ht⟩ ⟨f,hf⟩, heq_of_eq rfl) + +lemma inclusion_trans {R S T : subgroupoid C} (k : R ≤ S) (h : S ≤ T) : + inclusion (k.trans h) = (inclusion k) ⋙ (inclusion h) := rfl + +lemma inclusion_comp_embedding {S T : subgroupoid C} (h : S ≤ T) : + (inclusion h) ⋙ T.hom = S.hom := rfl + +/-- The family of arrows of the discrete groupoid -/ +inductive discrete.arrows : Π (c d : C), (c ⟶ d) → Prop +| id (c : C) : discrete.arrows c c (𝟙 c) + +/-- The only arrows of the discrete groupoid are the identity arrows. -/ +def discrete : subgroupoid C := +{ arrows := discrete.arrows, + inv := by { rintros _ _ _ ⟨⟩, simp only [inv_eq_inv, is_iso.inv_id], split, }, + mul := by { rintros _ _ _ _ ⟨⟩ _ ⟨⟩, rw category.comp_id, split, } } + +lemma mem_discrete_iff {c d : C} (f : c ⟶ d): + (f ∈ (discrete).arrows c d) ↔ (∃ (h : c = d), f = eq_to_hom h) := +⟨by { rintro ⟨⟩, exact ⟨rfl, rfl⟩ }, by { rintro ⟨rfl, rfl⟩, split }⟩ + +/-- A subgroupoid is normal if it is “wide” (meaning that its carrier set is all of `C`) + and satisfies the expected stability under conjugacy. -/ +structure is_normal : Prop := +(wide : ∀ c, (𝟙 c) ∈ (S.arrows c c)) +(conj : ∀ {c d} (p : c ⟶ d) {γ : c ⟶ c} (hs : γ ∈ S.arrows c c), + ((inv p) ≫ γ ≫ p) ∈ (S.arrows d d)) + +lemma is_normal.conj' {S : subgroupoid C} (Sn : is_normal S) : + ∀ {c d} (p : d ⟶ c) {γ : c ⟶ c} (hs : γ ∈ S.arrows c c), (p ≫ γ ≫ (inv p)) ∈ (S.arrows d d) := +λ c d p γ hs, by { convert Sn.conj (inv p) hs, simp, } + +lemma is_normal.conjugation_bij (Sn : is_normal S) {c d} (p : c ⟶ d) : + set.bij_on (λ γ : c ⟶ c, (inv p) ≫ γ ≫ p) (S.arrows c c) (S.arrows d d) := +begin + refine ⟨λ γ γS, Sn.conj p γS, λ γ₁ γ₁S γ₂ γ₂S h, _, λ δ δS, ⟨p ≫ δ ≫ (inv p), Sn.conj' p δS, _⟩⟩, + { simpa only [inv_eq_inv, category.assoc, is_iso.hom_inv_id, + category.comp_id, is_iso.hom_inv_id_assoc] using p ≫= h =≫ inv p }, + { simp only [inv_eq_inv, category.assoc, is_iso.inv_hom_id, + category.comp_id, is_iso.inv_hom_id_assoc] }, +end + +lemma top_is_normal : is_normal (⊤ : subgroupoid C) := +{ wide := (λ c, trivial), + conj := (λ a b c d e, trivial) } + +lemma Inf_is_normal (s : set $ subgroupoid C) (sn : ∀ S ∈ s, is_normal S) : is_normal (Inf s) := +{ wide := by { simp_rw [Inf, mem_Inter₂], exact λ c S Ss, (sn S Ss).wide c }, + conj := by { simp_rw [Inf, mem_Inter₂], exact λ c d p γ hγ S Ss, (sn S Ss).conj p (hγ S Ss) } } + +lemma is_normal.vertex_subgroup (Sn : is_normal S) (c : C) (cS : c ∈ S.objs) : + (S.vertex_subgroup cS).normal := +{ conj_mem := λ x hx y, by { rw mul_assoc, exact Sn.conj' y hx } } + +section generated_subgroupoid + +-- TODO: proof that generated is just "words in X" and generated_normal is similarly +variable (X : ∀ c d : C, set (c ⟶ d)) + +/-- The subgropoid generated by the set of arrows `X` -/ +def generated : subgroupoid C := +Inf {S : subgroupoid C | ∀ c d, X c d ⊆ S.arrows c d} + +/-- The normal sugroupoid generated by the set of arrows `X` -/ +def generated_normal : subgroupoid C := +Inf {S : subgroupoid C | (∀ c d, X c d ⊆ S.arrows c d) ∧ S.is_normal} + +lemma generated_normal_is_normal : (generated_normal X).is_normal := +Inf_is_normal _ (λ S h, h.right) + +end generated_subgroupoid + +section hom + +variables {D : Type*} [groupoid D] (φ : C ⥤ D) + +/-- +A functor between groupoid defines a map of subgroupoids in the reverse direction +by taking preimages. + -/ +def comap (S : subgroupoid D) : subgroupoid C := +{ arrows := λ c d, {f : c ⟶ d | φ.map f ∈ S.arrows (φ.obj c) (φ.obj d)}, + inv := λ c d p hp, by { rw [mem_set_of, inv_eq_inv, φ.map_inv p, ← inv_eq_inv], exact S.inv hp }, + mul := begin + rintros, + simp only [mem_set_of, functor.map_comp], + apply S.mul; assumption, + end } + +lemma comap_mono (S T : subgroupoid D) : + S ≤ T → comap φ S ≤ comap φ T := λ ST ⟨c,d,p⟩, @ST ⟨_,_,_⟩ + +lemma is_normal_comap {S : subgroupoid D} (Sn : is_normal S) : is_normal (comap φ S) := +{ wide := λ c, by { rw [comap, mem_set_of, functor.map_id], apply Sn.wide, }, + conj := λ c d f γ hγ, begin + simp only [comap, mem_set_of, functor.map_comp, functor.map_inv, inv_eq_inv], + rw [←inv_eq_inv], + exact Sn.conj _ hγ, + end } + +/-- The kernel of a functor between subgroupoid is the preimage. -/ +def ker : subgroupoid C := comap φ discrete + +lemma mem_ker_iff {c d : C} (f : c ⟶ d) : + f ∈ (ker φ).arrows c d ↔ ∃ (h : φ.obj c = φ.obj d), φ.map f = eq_to_hom h := +mem_discrete_iff (φ.map f) + +/-- The family of arrows of the image of a subgroupoid under a functor injective on objects -/ +inductive map.arrows (hφ : function.injective φ.obj) (S : subgroupoid C) : + Π (c d : D), (c ⟶ d) → Prop +| im {c d : C} (f : c ⟶ d) (hf : f ∈ S.arrows c d) : map.arrows (φ.obj c) (φ.obj d) (φ.map f) + +lemma map.mem_arrows_iff (hφ : function.injective φ.obj) (S : subgroupoid C) {c d : D} (f : c ⟶ d): + map.arrows φ hφ S c d f ↔ + ∃ (a b : C) (g : a ⟶ b) (ha : φ.obj a = c) (hb : φ.obj b = d) (hg : g ∈ S.arrows a b), + f = (eq_to_hom ha.symm) ≫ φ.map g ≫ (eq_to_hom hb) := +begin + split, + { rintro ⟨a,b,g,hg⟩, exact ⟨a,b,g,rfl,rfl,hg, eq_conj_eq_to_hom _⟩ }, + { rintro ⟨a,b,g,rfl,rfl,hg,rfl⟩, rw ← eq_conj_eq_to_hom, split, exact hg }, +end + +/-- The "forward" image of a subgroupoid under a functor injective on objects -/ +def map (hφ : function.injective φ.obj) (S : subgroupoid C) : subgroupoid D := +{ arrows := map.arrows φ hφ S, + inv := begin + rintro _ _ _ ⟨⟩, + rw [inv_eq_inv, ←functor.map_inv, ←inv_eq_inv], + split, apply S.inv, assumption, + end, + mul := begin + rintro _ _ _ _ ⟨c₁,c₂,f,hf⟩ q hq, + obtain ⟨c₃,c₄,g,he,rfl,hg,gq⟩ := (map.mem_arrows_iff φ hφ S q).mp hq, + cases hφ he, rw [gq, ← eq_conj_eq_to_hom, ← φ.map_comp], + split, exact S.mul hf hg, + end } + +lemma map_mono (hφ : function.injective φ.obj) (S T : subgroupoid C) : + S ≤ T → map φ hφ S ≤ map φ hφ T := +by { rintros ST ⟨c,d,f⟩ ⟨_,_,_,h⟩, split, exact @ST ⟨_,_,_⟩ h } + +/-- The image of a functor injective on objects -/ +def im (hφ : function.injective φ.obj) := map φ hφ (⊤) + +lemma mem_im_iff (hφ : function.injective φ.obj) {c d : D} (f : c ⟶ d) : + f ∈ (im φ hφ).arrows c d ↔ + ∃ (a b : C) (g : a ⟶ b) (ha : φ.obj a = c) (hb : φ.obj b = d), + f = (eq_to_hom ha.symm) ≫ φ.map g ≫ (eq_to_hom hb) := +by { convert map.mem_arrows_iff φ hφ ⊤ f, simp only [has_top.top, mem_univ, exists_true_left] } + +end hom + +end subgroupoid + +end category_theory diff --git a/src/category_theory/groupoid/vertex_group.lean b/src/category_theory/groupoid/vertex_group.lean new file mode 100644 index 0000000000000..a65b09cc7ea28 --- /dev/null +++ b/src/category_theory/groupoid/vertex_group.lean @@ -0,0 +1,81 @@ +/- +Copyright (c) 2022 Rémi Bottinelli. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Rémi Bottinelli +-/ +import category_theory.groupoid +import category_theory.path_category +import algebra.group.defs +import algebra.hom.group +import algebra.hom.equiv +import combinatorics.quiver.path + +/-! +# Vertex group + +This file defines the vertex group (*aka* isotropy group) of a groupoid at a vertex. + +## Implementation notes + +* The instance is defined "manually", instead of relying on `category_theory.Aut.group` or + using `category_theory.inv`. +* The multiplication order therefore matches the categorical one : `x * y = x ≫ y`. +* The inverse is directly defined in terms of the groupoidal inverse : `x ⁻¹ = groupoid.inv x`. + +## Tags + +isotropy, vertex group, groupoid +-/ + +namespace category_theory + +namespace groupoid + +universes u v + +variables {C : Type u} [groupoid C] + +/-- The vertex group at `c`. -/ +@[simps] instance vertex_group (c : C): group (c ⟶ c) := +{ mul := λ (x y : c ⟶ c), x ≫ y, + mul_assoc := category.assoc, + one := 𝟙 c, + one_mul := category.id_comp, + mul_one := category.comp_id, + inv := groupoid.inv, + mul_left_inv := inv_comp } + +/-- The inverse in the group is equal to the inverse given by `category_theory.inv`. -/ +lemma vertex_group.inv_eq_inv (c : C) (γ : c ⟶ c) : + γ ⁻¹ = category_theory.inv γ := groupoid.inv_eq_inv γ + +/-- +An arrow in the groupoid defines, by conjugation, an isomorphism of groups between +its endpoints. +-/ +@[simps] def vertex_group_isom_of_map {c d : C} (f : c ⟶ d) : (c ⟶ c) ≃* (d ⟶ d) := +{ to_fun := λ γ, inv f ≫ γ ≫ f, + inv_fun := λ δ, f ≫ δ ≫ inv f, + left_inv := λ γ, by simp_rw [category.assoc, comp_inv, category.comp_id, + ←category.assoc, comp_inv, category.id_comp], + right_inv := λ δ, by simp_rw [category.assoc, inv_comp, ←category.assoc, + inv_comp, category.id_comp, category.comp_id], + map_mul' := λ γ₁ γ₂, by simp only [vertex_group_mul, inv_eq_inv, + category.assoc, is_iso.hom_inv_id_assoc] } + +/-- +A path in the groupoid defines an isomorphism between its endpoints. +-/ +def vertex_group_isom_of_path {c d : C} (p : quiver.path c d) : (c ⟶ c) ≃* (d ⟶ d) := +vertex_group_isom_of_map (compose_path p) + +/-- A functor defines a morphism of vertex group. -/ +@[simps] def _root_.category_theory.functor.map_vertex_group {D : Type v} [groupoid D] + (φ : C ⥤ D) (c : C) : (c ⟶ c) →* (φ.obj c ⟶ φ.obj c) := +{ to_fun := φ.map, + map_one' := φ.map_id c, + map_mul' := φ.map_comp } + +end groupoid + +end category_theory From 71429ce7c6fa5b0f069ea89c05c0bac8b073a5a6 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 14 Oct 2022 08:49:54 +0000 Subject: [PATCH 760/828] feat(analysis/bounded_variation): functions of bounded variation are ae differentiable (#16680) We define functions of bounded variation. We show that such functions are a difference of monotone functions, and deduce that they are differentiable almost everywhere. This applies in particular to Lipschitz functions. --- src/analysis/bounded_variation.lean | 108 ++++++++++++++++++++++++++-- 1 file changed, 104 insertions(+), 4 deletions(-) diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index a9a24529d5bad..c59bf89d54592 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -5,12 +5,15 @@ Authors: Sébastien Gouëzel -/ import data.set.intervals.monotone import measure_theory.measure.lebesgue +import analysis.calculus.monotone /-! # Functions of bounded variation We study functions of bounded variation. In particular, we show that a bounded variation function -is a difference of monotone functions, and that Lipschitz functions have bounded variation +is a difference of monotone functions, and differentiable almost everywhere. This implies that +Lipschitz functions from the real line into finite-dimensional vector space are also differentiable +almost everywhere. ## Main definitions and results @@ -25,6 +28,9 @@ is a difference of monotone functions, and that Lipschitz functions have bounded with locally bounded variation is the difference of two monotone functions. * `lipschitz_with.has_locally_bounded_variation_on` shows that a Lipschitz function has locally bounded variation. +* `has_locally_bounded_variation_on.ae_differentiable_within_at` shows that a bounded variation + function into a finite dimensional real vector space is differentiable almost everywhere. +* `lipschitz_on_with.ae_differentiable_within_at` is the same result for Lipschitz functions. We also give several variations around these results. @@ -240,7 +246,7 @@ begin if_false, h], }, { have A : ¬(i ≤ n), by linarith, have B : ¬(i + 1 ≤ n), by linarith, - simp [A, B] } }, + simp only [A, B, if_false]} }, refine ⟨v, n+2, hv, vs, (mem_image _ _ _).2 ⟨n+1, _, _⟩, _⟩, { rw mem_Iio, exact nat.lt_succ_self (n+1) }, { have : ¬(n + 1 ≤ n), by linarith, @@ -268,10 +274,10 @@ begin exacts [us _, hx, us _] }, have hw : monotone w, { apply monotone_nat_of_le_succ (λ i, _), - dsimp [w], + dsimp only [w], rcases lt_trichotomy (i + 1) N with hi|hi|hi, { have : i < N, by linarith, - simp [hi, this], + simp only [hi, this, if_true], exact hu (nat.le_succ _) }, { have A : i < N, by linarith, have B : ¬(i + 1 < N), by linarith, @@ -651,3 +657,97 @@ hf.comp_has_locally_bounded_variation_on (maps_to_id _) lemma lipschitz_with.has_locally_bounded_variation_on {f : ℝ → E} {C : ℝ≥0} (hf : lipschitz_with C f) (s : set ℝ) : has_locally_bounded_variation_on f s := (hf.lipschitz_on_with s).has_locally_bounded_variation_on + + +/-! ## Almost everywhere differentiability of functions with locally bounded variation -/ + +namespace has_locally_bounded_variation_on + +/-- A bounded variation function into `ℝ` is differentiable almost everywhere. Superseded by +`ae_differentiable_within_at_of_mem`. -/ +theorem ae_differentiable_within_at_of_mem_real + {f : ℝ → ℝ} {s : set ℝ} (h : has_locally_bounded_variation_on f s) : + ∀ᵐ x, x ∈ s → differentiable_within_at ℝ f s x := +begin + obtain ⟨p, q, hp, hq, fpq⟩ : ∃ p q, monotone_on p s ∧ monotone_on q s ∧ f = p - q, + from h.exists_monotone_on_sub_monotone_on, + filter_upwards [hp.ae_differentiable_within_at_of_mem, hq.ae_differentiable_within_at_of_mem] + with x hxp hxq xs, + have fpq : ∀ x, f x = p x - q x, by simp [fpq], + refine ((hxp xs).sub (hxq xs)).congr (λ y hy, fpq y) (fpq x), +end + +/-- A bounded variation function into a finite dimensional product vector space is differentiable +almost everywhere. Superseded by `ae_differentiable_within_at_of_mem`. -/ +theorem ae_differentiable_within_at_of_mem_pi {ι : Type*} [fintype ι] + {f : ℝ → (ι → ℝ)} {s : set ℝ} (h : has_locally_bounded_variation_on f s) : + ∀ᵐ x, x ∈ s → differentiable_within_at ℝ f s x := +begin + have A : ∀ (i : ι), lipschitz_with 1 (λ (x : ι → ℝ), x i) := λ i, lipschitz_with.eval i, + have : ∀ (i : ι), ∀ᵐ x, x ∈ s → differentiable_within_at ℝ (λ (x : ℝ), f x i) s x, + { assume i, + apply ae_differentiable_within_at_of_mem_real, + exact lipschitz_with.comp_has_locally_bounded_variation_on (A i) h }, + filter_upwards [ae_all_iff.2 this] with x hx xs, + exact differentiable_within_at_pi.2 (λ i, hx i xs), +end + +/-- A real function into a finite dimensional real vector space with bounded variation on a set +is differentiable almost everywhere in this set. -/ +theorem ae_differentiable_within_at_of_mem + {f : ℝ → V} {s : set ℝ} (h : has_locally_bounded_variation_on f s) : + ∀ᵐ x, x ∈ s → differentiable_within_at ℝ f s x := +begin + let A := (basis.of_vector_space ℝ V).equiv_fun.to_continuous_linear_equiv, + suffices H : ∀ᵐ x, x ∈ s → differentiable_within_at ℝ (A ∘ f) s x, + { filter_upwards [H] with x hx xs, + have : f = (A.symm ∘ A) ∘ f, + by simp only [continuous_linear_equiv.symm_comp_self, function.comp.left_id], + rw this, + exact A.symm.differentiable_at.comp_differentiable_within_at x (hx xs) }, + apply ae_differentiable_within_at_of_mem_pi, + exact A.lipschitz.comp_has_locally_bounded_variation_on h, +end + +/-- A real function into a finite dimensional real vector space with bounded variation on a set +is differentiable almost everywhere in this set. -/ +theorem ae_differentiable_within_at + {f : ℝ → V} {s : set ℝ} (h : has_locally_bounded_variation_on f s) (hs : measurable_set s) : + ∀ᵐ x ∂(volume.restrict s), differentiable_within_at ℝ f s x := +begin + rw ae_restrict_iff' hs, + exact h.ae_differentiable_within_at_of_mem +end + +/-- A real function into a finite dimensional real vector space with bounded variation +is differentiable almost everywhere. -/ +theorem ae_differentiable_at {f : ℝ → V} (h : has_locally_bounded_variation_on f univ) : + ∀ᵐ x, differentiable_at ℝ f x := +begin + filter_upwards [h.ae_differentiable_within_at_of_mem] with x hx, + rw differentiable_within_at_univ at hx, + exact hx (mem_univ _), +end + +end has_locally_bounded_variation_on + +/-- A real function into a finite dimensional real vector space which is Lipschitz on a set +is differentiable almost everywhere in this set . -/ +lemma lipschitz_on_with.ae_differentiable_within_at_of_mem + {C : ℝ≥0} {f : ℝ → V} {s : set ℝ} (h : lipschitz_on_with C f s) : + ∀ᵐ x, x ∈ s → differentiable_within_at ℝ f s x := +h.has_locally_bounded_variation_on.ae_differentiable_within_at_of_mem + +/-- A real function into a finite dimensional real vector space which is Lipschitz on a set +is differentiable almost everywhere in this set. -/ +lemma lipschitz_on_with.ae_differentiable_within_at + {C : ℝ≥0} {f : ℝ → V} {s : set ℝ} (h : lipschitz_on_with C f s) (hs : measurable_set s) : + ∀ᵐ x ∂(volume.restrict s), differentiable_within_at ℝ f s x := +h.has_locally_bounded_variation_on.ae_differentiable_within_at hs + +/-- A real Lipschitz function into a finite dimensional real vector space is differentiable +almost everywhere. -/ +lemma lipschitz_with.ae_differentiable_at + {C : ℝ≥0} {f : ℝ → V} (h : lipschitz_with C f) : + ∀ᵐ x, differentiable_at ℝ f x := +(h.has_locally_bounded_variation_on univ).ae_differentiable_at From e0ce21862273a48e47abe598035d4d043e3e3ecb Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Fri, 14 Oct 2022 10:34:10 +0000 Subject: [PATCH 761/828] =?UTF-8?q?feat(analysis/normed=5Fspace/star/gelfa?= =?UTF-8?q?nd=5Fduality):=20functoriality=20of=20`character=5Fspace=20?= =?UTF-8?q?=E2=84=82`=20(#16835)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is one of the two contravariant functors involved in Gelfand duality. --- .../normed_space/star/gelfand_duality.lean | 42 +++++++++++++++++++ 1 file changed, 42 insertions(+) diff --git a/src/analysis/normed_space/star/gelfand_duality.lean b/src/analysis/normed_space/star/gelfand_duality.lean index 501f3d88a7b3a..ac95ef0534b3d 100644 --- a/src/analysis/normed_space/star/gelfand_duality.lean +++ b/src/analysis/normed_space/star/gelfand_duality.lean @@ -24,6 +24,8 @@ and even an equivalence between C⋆-algebras. * `ideal.to_character_space` : constructs an element of the character space from a maximal ideal in a commutative complex Banach algebra +* `weak_dual.character_space.comp_continuous_map`: The functorial map taking `ψ : A →⋆ₐ[ℂ] B` to a + continuous function `character_space ℂ B → character_space ℂ A` given by pre-composition with `ψ`. ## Main statements @@ -164,3 +166,43 @@ begin end end complex_cstar_algebra + +section functoriality + +namespace weak_dual + +namespace character_space + +variables {A B C : Type*} +variables [normed_ring A] [normed_algebra ℂ A] [complete_space A] [star_ring A] +variables [normed_ring B] [normed_algebra ℂ B] [complete_space B] [star_ring B] +variables [normed_ring C] [normed_algebra ℂ C] [complete_space C] [star_ring C] + +/-- The functorial map taking `ψ : A →⋆ₐ[ℂ] B` to a continuous function +`character_space ℂ B → character_space ℂ A` obtained by pre-composition with `ψ`. -/ +@[simps] +noncomputable def comp_continuous_map (ψ : A →⋆ₐ[ℂ] B) : + C(character_space ℂ B, character_space ℂ A) := +{ to_fun := λ φ, equiv_alg_hom.symm ((equiv_alg_hom φ).comp (ψ.to_alg_hom)), + continuous_to_fun := continuous.subtype_mk (continuous_of_continuous_eval $ + λ a, map_continuous $ gelfand_transform ℂ B (ψ a)) _ } + +variables (A) + +/-- `weak_dual.character_space.comp_continuous_map` sends the identity to the identity. -/ +@[simp] lemma comp_continuous_map_id : + comp_continuous_map (star_alg_hom.id ℂ A) = continuous_map.id (character_space ℂ A) := +continuous_map.ext $ λ a, ext $ λ x, rfl + +variables {A} + +/-- `weak_dual.character_space.comp_continuous_map` is functorial. -/ +@[simp] lemma comp_continuous_map_comp (ψ₂ : B →⋆ₐ[ℂ] C) (ψ₁ : A →⋆ₐ[ℂ] B) : + comp_continuous_map (ψ₂.comp ψ₁) = (comp_continuous_map ψ₁).comp (comp_continuous_map ψ₂) := +continuous_map.ext $ λ a, ext $ λ x, rfl + +end character_space + +end weak_dual + +end functoriality From 4c2452e8bfeb9b7151686cdf6d653ea41cd774f4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 14 Oct 2022 12:07:45 +0000 Subject: [PATCH 762/828] feat(topology/metric_space/emetric_space): `nhds_within` lemmas (#16952) A few lemmas about emetric spaces and `nhds_within`. --- src/topology/metric_space/emetric_space.lean | 33 +++++++++++++++++++- 1 file changed, 32 insertions(+), 1 deletion(-) diff --git a/src/topology/metric_space/emetric_space.lean b/src/topology/metric_space/emetric_space.lean index 3f9fc8b553efa..f61625720ab7a 100644 --- a/src/topology/metric_space/emetric_space.lean +++ b/src/topology/metric_space/emetric_space.lean @@ -493,7 +493,7 @@ finset.sup_le_iff.trans $ by simp only [finset.mem_univ, forall_const] end pi namespace emetric -variables {x y z : α} {ε ε₁ ε₂ : ℝ≥0∞} {s : set α} +variables {x y z : α} {ε ε₁ ε₂ : ℝ≥0∞} {s t : set α} /-- `emetric.ball x ε` is the set of all points `y` with `edist y x < ε` -/ def ball (x : α) (ε : ℝ≥0∞) : set α := {y | edist y x < ε} @@ -582,14 +582,45 @@ by rw [emetric.ball_eq_empty_iff] theorem nhds_basis_eball : (𝓝 x).has_basis (λ ε:ℝ≥0∞, 0 < ε) (ball x) := nhds_basis_uniformity uniformity_basis_edist +lemma nhds_within_basis_eball : (𝓝[s] x).has_basis (λ ε : ℝ≥0∞, 0 < ε) (λ ε, ball x ε ∩ s) := +nhds_within_has_basis nhds_basis_eball s + theorem nhds_basis_closed_eball : (𝓝 x).has_basis (λ ε:ℝ≥0∞, 0 < ε) (closed_ball x) := nhds_basis_uniformity uniformity_basis_edist_le +lemma nhds_within_basis_closed_eball : + (𝓝[s] x).has_basis (λ ε : ℝ≥0∞, 0 < ε) (λ ε, closed_ball x ε ∩ s) := +nhds_within_has_basis nhds_basis_closed_eball s + theorem nhds_eq : 𝓝 x = (⨅ε>0, 𝓟 (ball x ε)) := nhds_basis_eball.eq_binfi theorem mem_nhds_iff : s ∈ 𝓝 x ↔ ∃ε>0, ball x ε ⊆ s := nhds_basis_eball.mem_iff +lemma mem_nhds_within_iff : s ∈ 𝓝[t] x ↔ ∃ ε > 0, ball x ε ∩ t ⊆ s := +nhds_within_basis_eball.mem_iff + +section +variables [pseudo_emetric_space β] {f : α → β} + +lemma tendsto_nhds_within_nhds_within {t : set β} {a b} : + tendsto f (𝓝[s] a) (𝓝[t] b) ↔ + ∀ ε > 0, ∃ δ > 0, ∀ ⦃x⦄, x ∈ s → edist x a < δ → f x ∈ t ∧ edist (f x) b < ε := +(nhds_within_basis_eball.tendsto_iff nhds_within_basis_eball).trans $ + forall₂_congr $ λ ε hε, exists₂_congr $ λ δ hδ, + forall_congr $ λ x, by simp; itauto + +lemma tendsto_nhds_within_nhds {a b} : + tendsto f (𝓝[s] a) (𝓝 b) ↔ + ∀ ε > 0, ∃ δ > 0, ∀{x:α}, x ∈ s → edist x a < δ → edist (f x) b < ε := +by { rw [← nhds_within_univ b, tendsto_nhds_within_nhds_within], simp only [mem_univ, true_and] } + +lemma tendsto_nhds_nhds {a b} : + tendsto f (𝓝 a) (𝓝 b) ↔ ∀ ε > 0, ∃ δ > 0, ∀ ⦃x⦄, edist x a < δ → edist (f x) b < ε := +nhds_basis_eball.tendsto_iff nhds_basis_eball + +end + theorem is_open_iff : is_open s ↔ ∀x∈s, ∃ε>0, ball x ε ⊆ s := by simp [is_open_iff_nhds, mem_nhds_iff] From 34f53b6c3c21677ae52305667051211ea474d026 Mon Sep 17 00:00:00 2001 From: jakelev Date: Fri, 14 Oct 2022 14:12:20 +0000 Subject: [PATCH 763/828] refactor(combinatorics/young_tableau): move young diagram modules (#16954) Rename `combinatorics/young_tableaux` to `combinatorics/young` and move `combinatorics/young_diagram.lean` to `combinatorics/young/young_diagram.lean`. This is to group all topics related to Young tableaux together. Future additions will be in this directory. --- .../semistandard.lean => young/semistandard_tableau.lean} | 2 +- src/combinatorics/{ => young}/young_diagram.lean | 0 2 files changed, 1 insertion(+), 1 deletion(-) rename src/combinatorics/{young_tableaux/semistandard.lean => young/semistandard_tableau.lean} (99%) rename src/combinatorics/{ => young}/young_diagram.lean (100%) diff --git a/src/combinatorics/young_tableaux/semistandard.lean b/src/combinatorics/young/semistandard_tableau.lean similarity index 99% rename from src/combinatorics/young_tableaux/semistandard.lean rename to src/combinatorics/young/semistandard_tableau.lean index e0784ecf00138..1046ea9bb9189 100644 --- a/src/combinatorics/young_tableaux/semistandard.lean +++ b/src/combinatorics/young/semistandard_tableau.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jake Levinson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jake Levinson -/ -import combinatorics.young_diagram +import combinatorics.young.young_diagram /-! # Semistandard Young tableaux diff --git a/src/combinatorics/young_diagram.lean b/src/combinatorics/young/young_diagram.lean similarity index 100% rename from src/combinatorics/young_diagram.lean rename to src/combinatorics/young/young_diagram.lean From e623a7e43f627d9939114c1b87b38d452911a66a Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Fri, 14 Oct 2022 14:12:22 +0000 Subject: [PATCH 764/828] fix(topology/metric_space/lipschitz): typo (#16962) --- src/topology/metric_space/lipschitz.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/topology/metric_space/lipschitz.lean b/src/topology/metric_space/lipschitz.lean index b0279dedeb874..345d52c330fb0 100644 --- a/src/topology/metric_space/lipschitz.lean +++ b/src/topology/metric_space/lipschitz.lean @@ -223,7 +223,7 @@ end protected lemma prod_mk_left (a : α) : lipschitz_with 1 (prod.mk a : β → α × β) := by simpa only [max_eq_right zero_le_one] using (lipschitz_with.const a).prod lipschitz_with.id -protected lemma prod_mk_right (b : α) : lipschitz_with 1 (λ a : α, (a, b)) := +protected lemma prod_mk_right (b : β) : lipschitz_with 1 (λ a : α, (a, b)) := by simpa only [max_eq_left zero_le_one] using lipschitz_with.id.prod (lipschitz_with.const b) protected lemma uncurry {f : α → β → γ} {Kα Kβ : ℝ≥0} (hα : ∀ b, lipschitz_with Kα (λ a, f a b)) From f6b127639ee2cc43534b85fe015bffd942e635ce Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Fri, 14 Oct 2022 16:19:13 +0000 Subject: [PATCH 765/828] feat(topology/algebra/infinite_sum): generalize `tsum_le_of_sum_range_le` (#16951) --- src/topology/algebra/infinite_sum.lean | 11 +++++++++++ src/topology/instances/ennreal.lean | 11 +++++++---- 2 files changed, 18 insertions(+), 4 deletions(-) diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 1d59f5836bdd5..140ecce54409a 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -1395,6 +1395,17 @@ end end topological_group +section preorder + +variables {E : Type*} [preorder E] [add_comm_monoid E] + [topological_space E] [order_closed_topology E] [t2_space E] + +lemma tsum_le_of_sum_range_le {f : ℕ → E} {c : E} (hsum : summable f) + (h : ∀ n, ∑ i in finset.range n, f i ≤ c) : ∑' n, f n ≤ c := +let ⟨l, hl⟩ := hsum in hl.tsum_eq.symm ▸ le_of_tendsto' hl.tendsto_sum_nat h + +end preorder + section linear_order /-! For infinite sums taking values in a linearly ordered monoid, the existence of a least upper diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index eda048d0cf97a..51630bbfe861c 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -1055,7 +1055,7 @@ end lemma tsum_le_of_sum_range_le {f : ℕ → ℝ≥0} {c : ℝ≥0} (h : ∀ n, ∑ i in finset.range n, f i ≤ c) : ∑' n, f n ≤ c := -le_of_tendsto' (has_sum_iff_tendsto_nat.1 (summable_of_sum_range_le h).has_sum) h +tsum_le_of_sum_range_le (summable_of_sum_range_le h) h lemma tsum_comp_le_tsum_of_inj {β : Type*} {f : α → ℝ≥0} (hf : summable f) {i : β → α} (hi : function.injective i) : ∑' x, f (i x) ≤ ∑' x, f x := @@ -1145,6 +1145,10 @@ begin exact_mod_cast nnreal.tendsto_sum_nat_add f end +lemma tsum_le_of_sum_range_le {f : ℕ → ℝ≥0∞} {c : ℝ≥0∞} + (h : ∀ n, ∑ i in finset.range n, f i ≤ c) : ∑' n, f n ≤ c := +tsum_le_of_sum_range_le ennreal.summable h + end ennreal lemma tsum_comp_le_tsum_of_inj {β : Type*} {f : α → ℝ} (hf : summable f) (hn : ∀ a, 0 ≤ f a) @@ -1209,10 +1213,9 @@ begin exact lt_irrefl _ (hn.trans_le (h n)), end -lemma tsum_le_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) +lemma real.tsum_le_of_sum_range_le {f : ℕ → ℝ} {c : ℝ} (hf : ∀ n, 0 ≤ f n) (h : ∀ n, ∑ i in finset.range n, f i ≤ c) : ∑' n, f n ≤ c := -le_of_tendsto' ((has_sum_iff_tendsto_nat_of_nonneg hf _).1 - (summable_of_sum_range_le hf h).has_sum) h +tsum_le_of_sum_range_le (summable_of_sum_range_le hf h) h /-- If a sequence `f` with non-negative terms is dominated by a sequence `g` with summable series and at least one term of `f` is strictly smaller than the corresponding term in `g`, From e6f28f56c2017b994deef67890aad06410d86301 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 14 Oct 2022 18:32:12 +0000 Subject: [PATCH 766/828] feat(topology/algebra/order/basic): `nhds_within` versions (#16956) Version of `exists_seq_strict_mono_tendsto`/`exists_seq_strict_anti_tendsto` strengthened to `nhds_within`. --- src/topology/algebra/order/basic.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index 4ca8225e2a0d2..e4890a1185205 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -2278,6 +2278,12 @@ begin exact ⟨u, hu_mono, λ n, (hu_mem n).2, hux⟩ end +lemma exists_seq_strict_mono_tendsto_nhds_within [densely_ordered α] [no_min_order α] + [first_countable_topology α] (x : α) : + ∃ u : ℕ → α, strict_mono u ∧ (∀ n, u n < x) ∧ tendsto u at_top (𝓝[<] x) := +let ⟨u, hu, hx, h⟩ := exists_seq_strict_mono_tendsto x in ⟨u, hu, hx, + tendsto_nhds_within_mono_right (range_subset_iff.2 hx) $ tendsto_nhds_within_range.2 h⟩ + lemma exists_seq_tendsto_Sup {α : Type*} [conditionally_complete_linear_order α] [topological_space α] [order_topology α] [first_countable_topology α] {S : set α} (hS : S.nonempty) (hS' : bdd_above S) : @@ -2309,6 +2315,11 @@ lemma exists_seq_strict_anti_tendsto [densely_ordered α] [no_max_order α] ∃ u : ℕ → α, strict_anti u ∧ (∀ n, x < u n) ∧ tendsto u at_top (𝓝 x) := @exists_seq_strict_mono_tendsto αᵒᵈ _ _ _ _ _ _ x +lemma exists_seq_strict_anti_tendsto_nhds_within [densely_ordered α] [no_max_order α] + [first_countable_topology α] (x : α) : + ∃ u : ℕ → α, strict_anti u ∧ (∀ n, x < u n) ∧ tendsto u at_top (𝓝[>] x) := +@exists_seq_strict_mono_tendsto_nhds_within αᵒᵈ _ _ _ _ _ _ _ + lemma exists_seq_strict_anti_strict_mono_tendsto [densely_ordered α] [first_countable_topology α] {x y : α} (h : x < y) : ∃ (u v : ℕ → α), strict_anti u ∧ strict_mono v ∧ (∀ k, u k ∈ Ioo x y) ∧ (∀ l, v l ∈ Ioo x y) ∧ From da4010c2a5121b5284c21c732a6ed43633713336 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Fri, 14 Oct 2022 20:35:49 +0000 Subject: [PATCH 767/828] feat(probability/independence): Kolmogorov's 0-1 law (#16648) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove that any event in the tail σ-algebra of an independent sequence of sub-σ-algebras has probability 0 or 1. Co-authored-by: Rémy Degenne Co-authored-by: sgouezel --- docs/undergrad.yaml | 2 +- src/algebra/big_operators/basic.lean | 5 + src/data/set/basic.lean | 18 + src/measure_theory/measurable_space_def.lean | 23 +- src/measure_theory/pi_system.lean | 37 ++ src/probability/independence.lean | 339 ++++++++++++++++++- 6 files changed, 401 insertions(+), 23 deletions(-) diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index eccd67d980bec..037346590d914 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -536,7 +536,7 @@ Probability Theory: independent events: 'probability_theory.Indep_set' sigma-algebra: 'measurable_space' independent sigma-algebra: 'probability_theory.Indep' - $0$-$1$ law: '' + $0$-$1$ law: 'probability_theory.measure_zero_or_one_of_measurable_set_limsup_at_top' Borel-Cantelli lemma (easy direction): 'measure_theory.measure_limsup_eq_zero' Borel-Cantelli lemma (difficult direction): '' conditional probability: 'probability_theory.cond' diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index a99b0f85fcf5f..689369cc6c7f2 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -821,6 +821,11 @@ lemma prod_extend_by_one [decidable_eq α] (s : finset α) (f : α → β) : ∏ i in s, (if i ∈ s then f i else 1) = ∏ i in s, f i := prod_congr rfl $ λ i hi, if_pos hi +@[simp, to_additive] +lemma prod_ite_mem [decidable_eq α] (s t : finset α) (f : α → β) : + ∏ i in s, (if i ∈ t then f i else 1) = ∏ i in (s ∩ t), f i := +by rw [← finset.prod_filter, finset.filter_mem_eq_inter] + @[simp, to_additive] lemma prod_dite_eq [decidable_eq α] (s : finset α) (a : α) (b : Π x : α, a = x → β) : (∏ x in s, (if h : a = x then b x h else 1)) = ite (a ∈ s) (b a rfl) 1 := diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 8503fe1dfff55..730d083b0ad55 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -1232,6 +1232,24 @@ ext $ λ s, subset_empty_iff @[simp] theorem powerset_univ : 𝒫 (univ : set α) = univ := eq_univ_of_forall subset_univ +/-! ### Sets defined as an if-then-else -/ + +lemma mem_dite_univ_right (p : Prop) [decidable p] (t : p → set α) (x : α) : + (x ∈ if h : p then t h else univ) ↔ (∀ h : p, x ∈ t h) := +by split_ifs; simp [h] + +@[simp] lemma mem_ite_univ_right (p : Prop) [decidable p] (t : set α) (x : α) : + x ∈ ite p t set.univ ↔ (p → x ∈ t) := +mem_dite_univ_right p (λ _, t) x + +lemma mem_dite_univ_left (p : Prop) [decidable p] (t : ¬ p → set α) (x : α) : + (x ∈ if h : p then univ else t h) ↔ (∀ h : ¬ p, x ∈ t h) := +by split_ifs; simp [h] + +@[simp] lemma mem_ite_univ_left (p : Prop) [decidable p] (t : set α) (x : α) : + x ∈ ite p set.univ t ↔ (¬ p → x ∈ t) := +mem_dite_univ_left p (λ _, t) x + /-! ### If-then-else for sets -/ /-- `ite` for sets: `set.ite t s s' ∩ t = s ∩ t`, `set.ite t s s' ∩ tᶜ = s' ∩ tᶜ`. diff --git a/src/measure_theory/measurable_space_def.lean b/src/measure_theory/measurable_space_def.lean index 65c9533c049ac..ee1ec99dd1573 100644 --- a/src/measure_theory/measurable_space_def.lean +++ b/src/measure_theory/measurable_space_def.lean @@ -182,6 +182,11 @@ h₁.inter h₂.compl measurable_set (t.ite s₁ s₂) := (h₁.inter ht).union (h₂.diff ht) +lemma measurable_set.ite' {s t : set α} {p : Prop} + (hs : p → measurable_set s) (ht : ¬ p → measurable_set t) : + measurable_set (ite p s t) := +by { split_ifs, exacts [hs h, ht h], } + @[simp] lemma measurable_set.cond {s₁ s₂ : set α} (h₁ : measurable_set s₁) (h₂ : measurable_set s₂) {i : bool} : measurable_set (cond i s₁ s₂) := by { cases i, exacts [h₂, h₁] } @@ -371,6 +376,12 @@ begin { exact measurable_set_generate_from ht, }, end +@[simp] lemma generate_from_singleton_empty : generate_from {∅} = (⊥ : measurable_space α) := +by { rw [eq_bot_iff, generate_from_le_iff], simp, } + +@[simp] lemma generate_from_singleton_univ : generate_from {set.univ} = (⊥ : measurable_space α) := +by { rw [eq_bot_iff, generate_from_le_iff], simp, } + lemma measurable_set_bot_iff {s : set α} : @measurable_set α ⊥ s ↔ (s = ∅ ∨ s = univ) := let b : measurable_space α := { measurable_set' := λ s, s = ∅ ∨ s = univ, @@ -419,12 +430,12 @@ theorem measurable_set_supr {ι} {m : ι → measurable_space α} {s : set α} : by simp only [supr, measurable_set_Sup, exists_range_iff] lemma measurable_space_supr_eq (m : ι → measurable_space α) : - (⨆ n, m n) = measurable_space.generate_from {s | ∃ n, measurable_set[m n] s} := -begin - ext s, - rw measurable_space.measurable_set_supr, - refl, -end + (⨆ n, m n) = generate_from {s | ∃ n, measurable_set[m n] s} := +by { ext s, rw measurable_set_supr, refl, } + +lemma generate_from_Union_measurable_set (m : ι → measurable_space α) : + generate_from (⋃ n, {t | measurable_set[m n] t}) = ⨆ n, m n := +(@gi_generate_from α).l_supr_u m end complete_lattice diff --git a/src/measure_theory/pi_system.lean b/src/measure_theory/pi_system.lean index 78675f24a168c..b533be024faef 100644 --- a/src/measure_theory/pi_system.lean +++ b/src/measure_theory/pi_system.lean @@ -113,6 +113,23 @@ begin simpa using hst, end +lemma is_pi_system_Union_of_directed_le {α ι} (p : ι → set (set α)) + (hp_pi : ∀ n, is_pi_system (p n)) (hp_directed : directed (≤) p) : + is_pi_system (⋃ n, p n) := +begin + intros t1 ht1 t2 ht2 h, + rw set.mem_Union at ht1 ht2 ⊢, + cases ht1 with n ht1, + cases ht2 with m ht2, + obtain ⟨k, hpnk, hpmk⟩ : ∃ k, p n ≤ p k ∧ p m ≤ p k := hp_directed n m, + exact ⟨k, hp_pi k t1 (hpnk ht1) t2 (hpmk ht2) h⟩, +end + +lemma is_pi_system_Union_of_monotone {α ι} [semilattice_sup ι] (p : ι → set (set α)) + (hp_pi : ∀ n, is_pi_system (p n)) (hp_mono : monotone p) : + is_pi_system (⋃ n, p n) := +is_pi_system_Union_of_directed_le p hp_pi (monotone.directed_le hp_mono) + section order variables {α : Type*} {ι ι' : Sort*} [linear_order α] @@ -435,6 +452,26 @@ begin exact generate_from_mono (mem_pi_Union_Inter_of_measurable_set m hpS hpi), }, end +lemma generate_from_pi_Union_Inter_subsets {α ι} (m : ι → measurable_space α) (S : set ι) : + generate_from (pi_Union_Inter (λ i, {t | measurable_set[m i] t}) {t : finset ι | ↑t ⊆ S}) + = ⨆ i ∈ S, m i := +begin + rw generate_from_pi_Union_Inter_measurable_space, + simp only [set.mem_set_of_eq, exists_prop, supr_exists], + congr' 1, + ext1 i, + by_cases hiS : i ∈ S, + { simp only [hiS, csupr_pos], + refine le_antisymm (supr₂_le (λ t ht, le_rfl)) _, + rw le_supr_iff, + intros m' hm', + specialize hm' {i}, + simpa only [hiS, finset.coe_singleton, set.singleton_subset_iff, finset.mem_singleton, + eq_self_iff_true, and_self, csupr_pos] using hm', }, + { simp only [hiS, supr_false, supr₂_eq_bot, and_imp], + exact λ t htS hit, absurd (htS hit) hiS, }, +end + end Union_Inter namespace measurable_space diff --git a/src/probability/independence.lean b/src/probability/independence.lean index 4cc0e046c3894..5f8b789c0d436 100644 --- a/src/probability/independence.lean +++ b/src/probability/independence.lean @@ -27,8 +27,11 @@ import measure_theory.constructions.pi ## Main statements * `Indep_sets.Indep`: if π-systems are independent as sets of sets, then the -measurable space structures they generate are independent. + measurable space structures they generate are independent. * `indep_sets.indep`: variant with two π-systems. +* `measure_zero_or_one_of_measurable_set_limsup_at_top`: Kolmogorov's 0-1 law. Any set which is + measurable with respect to the tail σ-algebra `limsup s at_top` of an independent sequence of + σ-algebras `s` has probability 0 or 1. ## Implementation notes @@ -63,7 +66,7 @@ Part A, Chapter 4. -/ open measure_theory measurable_space -open_locale big_operators classical measure_theory +open_locale big_operators measure_theory ennreal namespace probability_theory @@ -155,6 +158,16 @@ lemma indep_bot_left (m' : measurable_space Ω) {m : measurable_space Ω} indep ⊥ m' μ := (indep_bot_right m').symm +lemma indep_set_empty_right {m : measurable_space Ω} {μ : measure Ω} [is_probability_measure μ] + (s : set Ω) : + indep_set s ∅ μ := +by { simp only [indep_set, generate_from_singleton_empty], exact indep_bot_right _, } + +lemma indep_set_empty_left {m : measurable_space Ω} {μ : measure Ω} [is_probability_measure μ] + (s : set Ω) : + indep_set ∅ s μ := +(indep_set_empty_right s).symm + lemma indep_sets_of_indep_sets_of_le_left {s₁ s₂ s₃: set (set Ω)} [measurable_space Ω] {μ : measure Ω} (h_indep : indep_sets s₁ s₂ μ) (h31 : s₃ ⊆ s₁) : indep_sets s₃ s₂ μ := @@ -226,6 +239,7 @@ lemma Indep_sets.indep_sets {s : ι → set (set Ω)} [measurable_space Ω] {μ (h_indep : Indep_sets s μ) {i j : ι} (hij : i ≠ j) : indep_sets (s i) (s j) μ := begin + classical, intros t₁ t₂ ht₁ ht₂, have hf_m : ∀ (x : ι), x ∈ {i, j} → (ite (x=i) t₁ t₂) ∈ s x, { intros x hx, @@ -244,7 +258,7 @@ begin rw h1, nth_rewrite 1 h2, nth_rewrite 3 h2, - rw [←h_inter, ←h_prod, h_indep {i, j} hf_m], + rw [← h_inter, ← h_prod, h_indep {i, j} hf_m], end lemma Indep.indep {m : ι → measurable_space Ω} [measurable_space Ω] {μ : measure Ω} @@ -300,7 +314,7 @@ begin have h_univ : μ_inter set.univ = ν set.univ, by rw [measure.restrict_apply_univ, measure.smul_apply, smul_eq_mul, measure_univ, mul_one], haveI : is_finite_measure μ_inter := @restrict.is_finite_measure Ω _ t1 μ ⟨measure_lt_top μ t1⟩, - rw [set.inter_comm, ←@measure.restrict_apply Ω _ μ t1 t2 (h2 t2 ht2m)], + rw [set.inter_comm, ← measure.restrict_apply (h2 t2 ht2m)], refine ext_on_measurable_space_of_generate_finite m p2 (λ t ht, _) h2 hpm2 hp2 h_univ ht2m, have ht2 : measurable_set[m] t, { refine h2 _ _, @@ -322,7 +336,7 @@ begin have h_univ : μ_inter set.univ = ν set.univ, by rw [measure.restrict_apply_univ, measure.smul_apply, smul_eq_mul, measure_univ, mul_one], haveI : is_finite_measure μ_inter := @restrict.is_finite_measure Ω _ t2 μ ⟨measure_lt_top μ t2⟩, - rw [mul_comm, ←@measure.restrict_apply Ω _ μ t2 t1 (h1 t1 ht1)], + rw [mul_comm, ← measure.restrict_apply (h1 t1 ht1)], refine ext_on_measurable_space_of_generate_finite m p1 (λ t ht, _) h1 hpm1 hp1 h_univ ht1, have ht1 : measurable_set[m] t, { refine h1 _ _, @@ -334,6 +348,100 @@ end variables {m0 : measurable_space Ω} {μ : measure Ω} +lemma indep_sets_pi_Union_Inter_of_disjoint [decidable_eq ι] [is_probability_measure μ] + {s : ι → set (set Ω)} {S T : set (finset ι)} + (h_indep : Indep_sets s μ) (hST : ∀ u v, u ∈ S → v ∈ T → disjoint u v) : + indep_sets (pi_Union_Inter s S) (pi_Union_Inter s T) μ := +begin + rintros t1 t2 ⟨p1, hp1, f1, ht1_m, ht1_eq⟩ ⟨p2, hp2, f2, ht2_m, ht2_eq⟩, + let g := λ i, ite (i ∈ p1) (f1 i) set.univ ∩ ite (i ∈ p2) (f2 i) set.univ, + have h_P_inter : μ (t1 ∩ t2) = ∏ n in p1 ∪ p2, μ (g n), + { have hgm : ∀ i ∈ p1 ∪ p2, g i ∈ s i, + { intros i hi_mem_union, + rw finset.mem_union at hi_mem_union, + cases hi_mem_union with hi1 hi2, + { have hi2 : i ∉ p2 := finset.disjoint_left.mp (hST p1 p2 hp1 hp2) hi1, + simp_rw [g, if_pos hi1, if_neg hi2, set.inter_univ], + exact ht1_m i hi1, }, + { have hi1 : i ∉ p1 := finset.disjoint_right.mp (hST p1 p2 hp1 hp2) hi2, + simp_rw [g, if_neg hi1, if_pos hi2, set.univ_inter], + exact ht2_m i hi2, }, }, + have h_p1_inter_p2 : ((⋂ x ∈ p1, f1 x) ∩ ⋂ x ∈ p2, f2 x) + = ⋂ i ∈ p1 ∪ p2, (ite (i ∈ p1) (f1 i) set.univ ∩ ite (i ∈ p2) (f2 i) set.univ), + { ext1 x, + simp only [set.mem_ite_univ_right, set.mem_inter_iff, set.mem_Inter, finset.mem_union], + exact ⟨λ h i _, ⟨h.1 i, h.2 i⟩, + λ h, ⟨λ i hi, (h i (or.inl hi)).1 hi, λ i hi, (h i (or.inr hi)).2 hi⟩⟩, }, + rw [ht1_eq, ht2_eq, h_p1_inter_p2, ← h_indep _ hgm], }, + have h_μg : ∀ n, μ (g n) = (ite (n ∈ p1) (μ (f1 n)) 1) * (ite (n ∈ p2) (μ (f2 n)) 1), + { intro n, + simp_rw g, + split_ifs, + { exact absurd rfl (finset.disjoint_iff_ne.mp (hST p1 p2 hp1 hp2) n h n h_1), }, + all_goals { simp only [measure_univ, one_mul, mul_one, set.inter_univ, set.univ_inter], }, }, + simp_rw [h_P_inter, h_μg, finset.prod_mul_distrib, + finset.prod_ite_mem (p1 ∪ p2) p1 (λ x, μ (f1 x)), + finset.union_inter_cancel_left, finset.prod_ite_mem (p1 ∪ p2) p2 (λ x, μ (f2 x)), + finset.union_inter_cancel_right, ht1_eq, ← h_indep p1 ht1_m, ht2_eq, ← h_indep p2 ht2_m], +end + +lemma indep_supr_of_disjoint [is_probability_measure μ] {m : ι → measurable_space Ω} + (h_le : ∀ i, m i ≤ m0) (h_indep : Indep m μ) {S T : set ι} (hST : disjoint S T) : + indep (⨆ i ∈ S, m i) (⨆ i ∈ T, m i) μ := +begin + refine indep_sets.indep (supr₂_le (λ i _, h_le i)) (supr₂_le (λ i _, h_le i)) _ _ + (generate_from_pi_Union_Inter_subsets m S).symm + (generate_from_pi_Union_Inter_subsets m T).symm _, + { refine is_pi_system_pi_Union_Inter _ (λ n, @is_pi_system_measurable_set Ω (m n)) _ _, + intros s t hs ht, + simp only [finset.sup_eq_union, set.mem_set_of_eq, finset.coe_union, set.union_subset_iff], + exact ⟨hs, ht⟩, }, + { refine is_pi_system_pi_Union_Inter _ (λ n, @is_pi_system_measurable_set Ω (m n)) _ _, + intros s t hs ht, + simp only [finset.sup_eq_union, set.mem_set_of_eq, finset.coe_union, set.union_subset_iff], + exact ⟨hs, ht⟩, }, + { classical, + refine indep_sets_pi_Union_Inter_of_disjoint h_indep (λ s t hs ht, _), + rw finset.disjoint_iff_ne, + exact λ i his j hjt, set.disjoint_iff_forall_ne.mp hST i (hs his) j (ht hjt), }, +end + +lemma indep_supr_of_directed_le {Ω} {m : ι → measurable_space Ω} + {m' m0 : measurable_space Ω} {μ : measure Ω} [is_probability_measure μ] + (h_indep : ∀ i, indep (m i) m' μ) (h_le : ∀ i, m i ≤ m0) (h_le' : m' ≤ m0) + (hm : directed (≤) m) : + indep (⨆ i, m i) m' μ := +begin + let p : ι → set (set Ω) := λ n, {t | measurable_set[m n] t}, + have hp : ∀ n, is_pi_system (p n) := λ n, @is_pi_system_measurable_set Ω (m n), + have h_gen_n : ∀ n, m n = generate_from (p n), + from λ n, (@generate_from_measurable_set Ω (m n)).symm, + have hp_supr_pi : is_pi_system (⋃ n, p n) := is_pi_system_Union_of_directed_le p hp hm, + let p' := {t : set Ω | measurable_set[m'] t}, + have hp'_pi : is_pi_system p' := @is_pi_system_measurable_set Ω m', + have h_gen' : m' = generate_from p' := (@generate_from_measurable_set Ω m').symm, + -- the π-systems defined are independent + have h_pi_system_indep : indep_sets (⋃ n, p n) p' μ, + { refine indep_sets.Union _, + simp_rw [h_gen_n, h_gen'] at h_indep, + exact λ n, (h_indep n).indep_sets, }, + -- now go from π-systems to σ-algebras + refine indep_sets.indep (supr_le h_le) h_le' hp_supr_pi hp'_pi _ h_gen' h_pi_system_indep, + exact (generate_from_Union_measurable_set _).symm, +end + +lemma indep_supr_of_monotone [semilattice_sup ι] {Ω} {m : ι → measurable_space Ω} + {m' m0 : measurable_space Ω} {μ : measure Ω} [is_probability_measure μ] + (h_indep : ∀ i, indep (m i) m' μ) (h_le : ∀ i, m i ≤ m0) (h_le' : m' ≤ m0) (hm : monotone m) : + indep (⨆ i, m i) m' μ := +indep_supr_of_directed_le h_indep h_le h_le' (monotone.directed_le hm) + +lemma indep_supr_of_antitone [semilattice_inf ι] {Ω} {m : ι → measurable_space Ω} + {m' m0 : measurable_space Ω} {μ : measure Ω} [is_probability_measure μ] + (h_indep : ∀ i, indep (m i) m' μ) (h_le : ∀ i, m i ≤ m0) (h_le' : m' ≤ m0) (hm : antitone m) : + indep (⨆ i, m i) m' μ := +indep_supr_of_directed_le h_indep h_le h_le' (directed_of_inf hm) + lemma Indep_sets.pi_Union_Inter_singleton {π : ι → set (set Ω)} {a : ι} {S : finset ι} (hp_ind : Indep_sets π μ) (haS : a ∉ S) : indep_sets (pi_Union_Inter π {S}) (π a) μ := @@ -341,6 +449,7 @@ begin rintros t1 t2 ⟨s, hs_mem, ft1, hft1_mem, ht1_eq⟩ ht2_mem_pia, rw set.mem_singleton_iff at hs_mem, subst hs_mem, + classical, let f := λ n, ite (n = a) t2 (ite (n ∈ s) (ft1 n) set.univ), have h_f_mem : ∀ n ∈ insert a s, f n ∈ π n, { intros n hn_mem_insert, @@ -359,12 +468,12 @@ begin intros n hnS, have hn_ne_a : n ≠ a, by { rintro rfl, exact haS hnS, }, simp_rw [f, if_pos hnS, if_neg hn_ne_a], }, - have h_μ_t1 : μ t1 = ∏ n in s, μ (f n), by rw [h_t1, ←hp_ind s h_f_mem_pi], + have h_μ_t1 : μ t1 = ∏ n in s, μ (f n), by rw [h_t1, ← hp_ind s h_f_mem_pi], have h_t2 : t2 = f a, by { simp_rw [f], simp, }, have h_μ_inter : μ (t1 ∩ t2) = ∏ n in insert a s, μ (f n), { have h_t1_inter_t2 : t1 ∩ t2 = ⋂ n ∈ insert a s, f n, by rw [h_t1, h_t2, finset.set_bInter_insert, set.inter_comm], - rw [h_t1_inter_t2, ←hp_ind (insert a s) h_f_mem], }, + rw [h_t1_inter_t2, ← hp_ind (insert a s) h_f_mem], }, rw [h_μ_inter, finset.prod_insert haS, h_t2, mul_comm, h_μ_t1], end @@ -375,10 +484,11 @@ theorem Indep_sets.Indep_aux [is_probability_measure μ] (m : ι → measurable_ (h_ind : Indep_sets π μ) : Indep m μ := begin + classical, refine finset.induction (by simp [measure_univ]) _, intros a S ha_notin_S h_rec f hf_m, have hf_m_S : ∀ x ∈ S, measurable_set[m x] (f x) := λ x hx, hf_m x (by simp [hx]), - rw [finset.set_bInter_insert, finset.prod_insert ha_notin_S, ←h_rec hf_m_S], + rw [finset.set_bInter_insert, finset.prod_insert ha_notin_S, ← h_rec hf_m_S], let p := pi_Union_Inter π {S}, set m_p := generate_from p with hS_eq_generate, have h_indep : indep m_p (m a) μ, @@ -413,6 +523,7 @@ begin rw [h_generate i, generate_from_insert_univ (π i)], }, have h_ind' : Indep_sets π' μ, { intros S f hfπ', + classical, let S' := finset.filter (λ i, f i ≠ set.univ) S, have h_mem : ∀ i ∈ S', f i ∈ π i, { intros i hi, @@ -455,9 +566,10 @@ We prove the following equivalences on `indep_set`, for measurable sets `s, t`. * `indep_set s t μ ↔ indep_sets {s} {t} μ`. -/ -variables [measurable_space Ω] {s t : set Ω} (S T : set (set Ω)) +variables {s t : set Ω} (S T : set (set Ω)) -lemma indep_set_iff_indep_sets_singleton (hs_meas : measurable_set s) (ht_meas : measurable_set t) +lemma indep_set_iff_indep_sets_singleton {m0 : measurable_space Ω} + (hs_meas : measurable_set s) (ht_meas : measurable_set t) (μ : measure Ω . volume_tac) [is_probability_measure μ] : indep_set s t μ ↔ indep_sets {s} {t} μ := ⟨indep.indep_sets, λ h, indep_sets.indep @@ -465,17 +577,42 @@ lemma indep_set_iff_indep_sets_singleton (hs_meas : measurable_set s) (ht_meas : (generate_from_le (λ u hu, by rwa set.mem_singleton_iff.mp hu)) (is_pi_system.singleton s) (is_pi_system.singleton t) rfl rfl h⟩ -lemma indep_set_iff_measure_inter_eq_mul (hs_meas : measurable_set s) (ht_meas : measurable_set t) +lemma indep_set_iff_measure_inter_eq_mul {m0 : measurable_space Ω} + (hs_meas : measurable_set s) (ht_meas : measurable_set t) (μ : measure Ω . volume_tac) [is_probability_measure μ] : indep_set s t μ ↔ μ (s ∩ t) = μ s * μ t := (indep_set_iff_indep_sets_singleton hs_meas ht_meas μ).trans indep_sets_singleton_iff -lemma indep_sets.indep_set_of_mem (hs : s ∈ S) (ht : t ∈ T) (hs_meas : measurable_set s) - (ht_meas : measurable_set t) (μ : measure Ω . volume_tac) [is_probability_measure μ] - (h_indep : indep_sets S T μ) : +lemma indep_sets.indep_set_of_mem {m0 : measurable_space Ω} (hs : s ∈ S) (ht : t ∈ T) + (hs_meas : measurable_set s) (ht_meas : measurable_set t) (μ : measure Ω . volume_tac) + [is_probability_measure μ] (h_indep : indep_sets S T μ) : indep_set s t μ := (indep_set_iff_measure_inter_eq_mul hs_meas ht_meas μ).mpr (h_indep s t hs ht) +lemma indep.indep_set_of_measurable_set {m₁ m₂ m0 : measurable_space Ω} {μ : measure Ω} + (h_indep : indep m₁ m₂ μ) {s t : set Ω} (hs : measurable_set[m₁] s) (ht : measurable_set[m₂] t) : + indep_set s t μ := +begin + refine λ s' t' hs' ht', h_indep s' t' _ _, + { refine generate_from_induction (λ u, measurable_set[m₁] u) {s} _ _ _ _ hs', + { simp only [hs, set.mem_singleton_iff, set.mem_set_of_eq, forall_eq], }, + { exact @measurable_set.empty _ m₁, }, + { exact λ u hu, hu.compl, }, + { exact λ f hf, measurable_set.Union hf, }, }, + { refine generate_from_induction (λ u, measurable_set[m₂] u) {t} _ _ _ _ ht', + { simp only [ht, set.mem_singleton_iff, set.mem_set_of_eq, forall_eq], }, + { exact @measurable_set.empty _ m₂, }, + { exact λ u hu, hu.compl, }, + { exact λ f hf, measurable_set.Union hf, },}, +end + +lemma indep_iff_forall_indep_set (m₁ m₂ : measurable_space Ω) {m0 : measurable_space Ω} + (μ : measure Ω) : + indep m₁ m₂ μ ↔ ∀ s t, measurable_set[m₁] s → measurable_set[m₂] t → indep_set s t μ := +⟨λ h, λ s t hs ht, h.indep_set_of_measurable_set hs ht, + λ h s t hs ht, h s t hs ht s t (measurable_set_generate_from (set.mem_singleton s)) + (measurable_set_generate_from (set.mem_singleton t))⟩ + end indep_set section indep_fun @@ -505,6 +642,7 @@ lemma Indep_fun_iff_measure_inter_preimage_eq_mul {ι : Type*} {β : ι → Type begin refine ⟨λ h S sets h_meas, h _ (λ i hi_mem, ⟨sets i, h_meas i hi_mem, rfl⟩), _⟩, intros h S setsΩ h_meas, + classical, let setsβ : (Π i : ι, set (β i)) := λ i, dite (i ∈ S) (λ hi_mem, (h_meas i hi_mem).some) (λ _, set.univ), have h_measβ : ∀ i ∈ S, measurable_set[m i] (setsβ i), @@ -549,7 +687,7 @@ begin rintro _ _ ⟨A, hA, rfl⟩ ⟨B, hB, rfl⟩, have h1 : f ⁻¹' A =ᵐ[μ] f' ⁻¹' A := hf.fun_comp A, have h2 : g ⁻¹' B =ᵐ[μ] g' ⁻¹' B := hg.fun_comp B, - rw [←measure_congr h1, ←measure_congr h2, ←measure_congr (h1.inter h2)], + rw [← measure_congr h1, ← measure_congr h2, ← measure_congr (h1.inter h2)], exact hfg _ _ ⟨_, hA, rfl⟩ ⟨_, hB, rfl⟩ end @@ -568,7 +706,7 @@ end two disjoint finite index sets, then the tuple formed by `f i` for `i ∈ S` is independent of the tuple `(f i)_i` for `i ∈ T`. -/ lemma Indep_fun.indep_fun_finset [is_probability_measure μ] - {ι : Type*} {β : ι → Type*} {m : Π i, measurable_space (β i)} + {ι : Type*} [decidable_eq ι] {β : ι → Type*} {m : Π i, measurable_space (β i)} {f : Π i, Ω → β i} (S T : finset ι) (hST : disjoint S T) (hf_Indep : Indep_fun m f μ) (hf_meas : ∀ i, measurable (f i)) : indep_fun (λ a (i : S), f i a) (λ a (i : T), f i a) μ := @@ -733,4 +871,173 @@ hf_Indep.indep_fun_finset_prod_of_not_mem hf_meas finset.not_mem_range_self end indep_fun + +/-! ### Kolmogorov's 0-1 law + +Let `s : ι → measurable_space Ω` be an independent sequence of sub-σ-algebras. Then any set which +is measurable with respect to the tail σ-algebra `limsup s at_top` has probability 0 or 1. +-/ + +section zero_one_law + +variables {m m0 : measurable_space Ω} {μ : measure Ω} + +lemma measure_eq_zero_or_one_or_top_of_indep_set_self {t : set Ω} (h_indep : indep_set t t μ) : + μ t = 0 ∨ μ t = 1 ∨ μ t = ∞ := +begin + specialize h_indep t t (measurable_set_generate_from (set.mem_singleton t)) + (measurable_set_generate_from (set.mem_singleton t)), + by_cases h0 : μ t = 0, + { exact or.inl h0, }, + by_cases h_top : μ t = ∞, + { exact or.inr (or.inr h_top), }, + rw [← one_mul (μ (t ∩ t)), set.inter_self, ennreal.mul_eq_mul_right h0 h_top] at h_indep, + exact or.inr (or.inl h_indep.symm), +end + +lemma measure_eq_zero_or_one_of_indep_set_self [is_finite_measure μ] {t : set Ω} + (h_indep : indep_set t t μ) : + μ t = 0 ∨ μ t = 1 := +begin + have h_0_1_top := measure_eq_zero_or_one_or_top_of_indep_set_self h_indep, + simpa [measure_ne_top μ] using h_0_1_top, +end + +variables [is_probability_measure μ] {s : ι → measurable_space Ω} + +open filter + +lemma indep_bsupr_compl (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) (t : set ι) : + indep (⨆ n ∈ t, s n) (⨆ n ∈ tᶜ, s n) μ := +indep_supr_of_disjoint h_le h_indep disjoint_compl_right + +section abstract +variables {α : Type*} {p : set ι → Prop} {f : filter ι} {ns : α → set ι} + +/-! We prove a version of Kolmogorov's 0-1 law for the σ-algebra `limsup s f` where `f` is a filter +for which we can define the following two functions: +* `p : set ι → Prop` such that for a set `t`, `p t → tᶜ ∈ f`, +* `ns : α → set ι` a directed sequence of sets which all verify `p` and such that + `⋃ a, ns a = set.univ`. + +For the example of `f = at_top`, we can take `p = bdd_above` and `ns : ι → set ι := λ i, set.Iic i`. +-/ + +lemma indep_bsupr_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) + (hf : ∀ t, p t → tᶜ ∈ f) {t : set ι} (ht : p t) : + indep (⨆ n ∈ t, s n) (limsup s f) μ := +begin + refine indep_of_indep_of_le_right (indep_bsupr_compl h_le h_indep t) _, + refine Limsup_le_of_le (by is_bounded_default) _, + simp only [set.mem_compl_iff, eventually_map], + exact eventually_of_mem (hf t ht) le_supr₂, +end + +lemma indep_supr_directed_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) + (hf : ∀ t, p t → tᶜ ∈ f) (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) : + indep (⨆ a, ⨆ n ∈ (ns a), s n) (limsup s f) μ := +begin + refine indep_supr_of_directed_le _ _ _ _, + { exact λ a, indep_bsupr_limsup h_le h_indep hf (hnsp a), }, + { exact λ a, supr₂_le (λ n hn, h_le n), }, + { exact limsup_le_supr.trans (supr_le h_le), }, + { intros a b, + obtain ⟨c, hc⟩ := hns a b, + refine ⟨c, _, _⟩; refine supr_mono (λ n, supr_mono' (λ hn, ⟨_, le_rfl⟩)), + { exact hc.1 hn, }, + { exact hc.2 hn, }, }, +end + +lemma indep_supr_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) (hf : ∀ t, p t → tᶜ ∈ f) + (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : + indep (⨆ n, s n) (limsup s f) μ := +begin + suffices : (⨆ a, ⨆ n ∈ (ns a), s n) = ⨆ n, s n, + { rw ← this, + exact indep_supr_directed_limsup h_le h_indep hf hns hnsp, }, + rw supr_comm, + refine supr_congr (λ n, _), + have : (⨆ (i : α) (H : n ∈ ns i), s n) = (⨆ (h : ∃ i, n ∈ ns i), s n), by rw supr_exists, + haveI : nonempty (∃ (i : α), n ∈ ns i) := ⟨hns_univ n⟩, + rw [this, supr_const], +end + +lemma indep_limsup_self (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) (hf : ∀ t, p t → tᶜ ∈ f) + (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) (hns_univ : ∀ n, ∃ a, n ∈ ns a) : + indep (limsup s f) (limsup s f) μ := +indep_of_indep_of_le_left (indep_supr_limsup h_le h_indep hf hns hnsp hns_univ) limsup_le_supr + +theorem measure_zero_or_one_of_measurable_set_limsup (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) + (hf : ∀ t, p t → tᶜ ∈ f) (hns : directed (≤) ns) (hnsp : ∀ a, p (ns a)) + (hns_univ : ∀ n, ∃ a, n ∈ ns a) {t : set Ω} (ht_tail : measurable_set[limsup s f] t) : + μ t = 0 ∨ μ t = 1 := +measure_eq_zero_or_one_of_indep_set_self + ((indep_limsup_self h_le h_indep hf hns hnsp hns_univ).indep_set_of_measurable_set + ht_tail ht_tail) + +end abstract + +section at_top +variables [semilattice_sup ι] [no_max_order ι] [nonempty ι] + +lemma indep_limsup_at_top_self (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) : + indep (limsup s at_top) (limsup s at_top) μ := +begin + let ns : ι → set ι := set.Iic, + have hnsp : ∀ i, bdd_above (ns i) := λ i, bdd_above_Iic, + refine indep_limsup_self h_le h_indep _ _ hnsp _, + { simp only [mem_at_top_sets, ge_iff_le, set.mem_compl_iff, bdd_above, upper_bounds, + set.nonempty], + rintros t ⟨a, ha⟩, + obtain ⟨b, hb⟩ : ∃ b, a < b := exists_gt a, + refine ⟨b, λ c hc hct, _⟩, + suffices : ∀ i ∈ t, i < c, from lt_irrefl c (this c hct), + exact λ i hi, (ha hi).trans_lt (hb.trans_le hc), }, + { exact monotone.directed_le (λ i j hij k hki, le_trans hki hij), }, + { exact λ n, ⟨n, le_rfl⟩, }, +end + +/-- **Kolmogorov's 0-1 law** : any event in the tail σ-algebra of an independent sequence of +sub-σ-algebras has probability 0 or 1. +The tail σ-algebra `limsup s at_top` is the same as `⋂ n, ⋃ i ≥ n, s i`. -/ +theorem measure_zero_or_one_of_measurable_set_limsup_at_top (h_le : ∀ n, s n ≤ m0) + (h_indep : Indep s μ) {t : set Ω} (ht_tail : measurable_set[limsup s at_top] t) : + μ t = 0 ∨ μ t = 1 := +measure_eq_zero_or_one_of_indep_set_self + ((indep_limsup_at_top_self h_le h_indep).indep_set_of_measurable_set ht_tail ht_tail) + +end at_top + +section at_bot +variables [semilattice_inf ι] [no_min_order ι] [nonempty ι] + +lemma indep_limsup_at_bot_self (h_le : ∀ n, s n ≤ m0) (h_indep : Indep s μ) : + indep (limsup s at_bot) (limsup s at_bot) μ := +begin + let ns : ι → set ι := set.Ici, + have hnsp : ∀ i, bdd_below (ns i) := λ i, bdd_below_Ici, + refine indep_limsup_self h_le h_indep _ _ hnsp _, + { simp only [mem_at_bot_sets, ge_iff_le, set.mem_compl_iff, bdd_below, lower_bounds, + set.nonempty], + rintros t ⟨a, ha⟩, + obtain ⟨b, hb⟩ : ∃ b, b < a := exists_lt a, + refine ⟨b, λ c hc hct, _⟩, + suffices : ∀ i ∈ t, c < i, from lt_irrefl c (this c hct), + exact λ i hi, hc.trans_lt (hb.trans_le (ha hi)), }, + { exact directed_of_inf (λ i j hij k hki, hij.trans hki), }, + { exact λ n, ⟨n, le_rfl⟩, }, +end + +/-- **Kolmogorov's 0-1 law** : any event in the tail σ-algebra of an independent sequence of +sub-σ-algebras has probability 0 or 1. -/ +theorem measure_zero_or_one_of_measurable_set_limsup_at_bot (h_le : ∀ n, s n ≤ m0) + (h_indep : Indep s μ) {t : set Ω} (ht_tail : measurable_set[limsup s at_bot] t) : + μ t = 0 ∨ μ t = 1 := +measure_eq_zero_or_one_of_indep_set_self + ((indep_limsup_at_bot_self h_le h_indep).indep_set_of_measurable_set ht_tail ht_tail) + +end at_bot + +end zero_one_law + end probability_theory From 5c19b83b7da2964a07475ae7c3182368921381c0 Mon Sep 17 00:00:00 2001 From: Vincent Beffara Date: Fri, 14 Oct 2022 20:35:50 +0000 Subject: [PATCH 768/828] feat(analysis/complex/open_mapping): the open mapping thm for holomorphic functions (#16780) Co-authored-by: sgouezel Co-authored-by: Johan Commelin --- src/analysis/analytic/basic.lean | 4 + src/analysis/analytic/isolated_zeros.lean | 8 + src/analysis/calculus/deriv.lean | 26 ++-- src/analysis/complex/abs_max.lean | 12 ++ src/analysis/complex/cauchy_integral.lean | 4 + src/analysis/complex/open_mapping.lean | 173 ++++++++++++++++++++++ src/analysis/normed/group/basic.lean | 4 + src/topology/algebra/field.lean | 12 ++ src/topology/algebra/order/compact.lean | 22 +++ src/topology/metric_space/basic.lean | 11 ++ 10 files changed, 263 insertions(+), 13 deletions(-) create mode 100644 src/analysis/complex/open_mapping.lean diff --git a/src/analysis/analytic/basic.lean b/src/analysis/analytic/basic.lean index ac10a5351b276..54c76bfbbb20a 100644 --- a/src/analysis/analytic/basic.lean +++ b/src/analysis/analytic/basic.lean @@ -516,6 +516,10 @@ lemma analytic_at.sub (hf : analytic_at 𝕜 f x) (hg : analytic_at 𝕜 g x) : analytic_at 𝕜 (f - g) x := by simpa only [sub_eq_add_neg] using hf.add hg.neg +lemma analytic_on.mono {s t : set E} (hf : analytic_on 𝕜 f t) (hst : s ⊆ t) : + analytic_on 𝕜 f s := +λ z hz, hf z (hst hz) + lemma analytic_on.add {s : set E} (hf : analytic_on 𝕜 f s) (hg : analytic_on 𝕜 g s) : analytic_on 𝕜 (f + g) s := λ z hz, (hf z hz).add (hg z hz) diff --git a/src/analysis/analytic/isolated_zeros.lean b/src/analysis/analytic/isolated_zeros.lean index 0a4d40c6a4363..61d0fd16336c1 100644 --- a/src/analysis/analytic/isolated_zeros.lean +++ b/src/analysis/analytic/isolated_zeros.lean @@ -141,11 +141,19 @@ begin { exact or.inr (hp.locally_ne_zero h) } end +lemma eventually_eq_or_eventually_ne (hf : analytic_at 𝕜 f z₀) (hg : analytic_at 𝕜 g z₀) : + (∀ᶠ z in 𝓝 z₀, f z = g z) ∨ (∀ᶠ z in 𝓝[≠] z₀, f z ≠ g z) := +by simpa [sub_eq_zero] using (hf.sub hg).eventually_eq_zero_or_eventually_ne_zero + lemma frequently_zero_iff_eventually_zero {f : 𝕜 → E} {w : 𝕜} (hf : analytic_at 𝕜 f w) : (∃ᶠ z in 𝓝[≠] w, f z = 0) ↔ (∀ᶠ z in 𝓝 w, f z = 0) := ⟨hf.eventually_eq_zero_or_eventually_ne_zero.resolve_right, λ h, (h.filter_mono nhds_within_le_nhds).frequently⟩ +lemma frequently_eq_iff_eventually_eq (hf : analytic_at 𝕜 f z₀) (hg : analytic_at 𝕜 g z₀) : + (∃ᶠ z in 𝓝[≠] z₀, f z = g z) ↔ (∀ᶠ z in 𝓝 z₀, f z = g z) := +by simpa [sub_eq_zero] using frequently_zero_iff_eventually_zero (hf.sub hg) + end analytic_at namespace analytic_on diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 15b565b864e92..055fa55a3c09d 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -1520,7 +1520,7 @@ begin exact fderiv_inv end -variables {c : 𝕜 → 𝕜} {c' : 𝕜} +variables {c : 𝕜 → 𝕜} {h : E → 𝕜} {c' : 𝕜} {z : E} {S : set E} lemma has_deriv_within_at.inv (hc : has_deriv_within_at c c' s x) (hx : c x ≠ 0) : @@ -1537,21 +1537,21 @@ begin exact hc.inv hx end -lemma differentiable_within_at.inv (hc : differentiable_within_at 𝕜 c s x) (hx : c x ≠ 0) : - differentiable_within_at 𝕜 (λx, (c x)⁻¹) s x := -(hc.has_deriv_within_at.inv hx).differentiable_within_at +lemma differentiable_within_at.inv (hf : differentiable_within_at 𝕜 h S z) (hz : h z ≠ 0) : + differentiable_within_at 𝕜 (λx, (h x)⁻¹) S z := +(differentiable_at_inv.mpr hz).comp_differentiable_within_at z hf -@[simp] lemma differentiable_at.inv (hc : differentiable_at 𝕜 c x) (hx : c x ≠ 0) : - differentiable_at 𝕜 (λx, (c x)⁻¹) x := -(hc.has_deriv_at.inv hx).differentiable_at +@[simp] lemma differentiable_at.inv (hf : differentiable_at 𝕜 h z) (hz : h z ≠ 0) : + differentiable_at 𝕜 (λx, (h x)⁻¹) z := +(differentiable_at_inv.mpr hz).comp z hf -lemma differentiable_on.inv (hc : differentiable_on 𝕜 c s) (hx : ∀ x ∈ s, c x ≠ 0) : - differentiable_on 𝕜 (λx, (c x)⁻¹) s := -λx h, (hc x h).inv (hx x h) +lemma differentiable_on.inv (hf : differentiable_on 𝕜 h S) (hz : ∀ x ∈ S, h x ≠ 0) : + differentiable_on 𝕜 (λx, (h x)⁻¹) S := +λx h, (hf x h).inv (hz x h) -@[simp] lemma differentiable.inv (hc : differentiable 𝕜 c) (hx : ∀ x, c x ≠ 0) : - differentiable 𝕜 (λx, (c x)⁻¹) := -λx, (hc x).inv (hx x) +@[simp] lemma differentiable.inv (hf : differentiable 𝕜 h) (hz : ∀ x, h x ≠ 0) : + differentiable 𝕜 (λx, (h x)⁻¹) := +λx, (hf x).inv (hz x) lemma deriv_within_inv' (hc : differentiable_within_at 𝕜 c s x) (hx : c x ≠ 0) (hxs : unique_diff_within_at 𝕜 s x) : diff --git a/src/analysis/complex/abs_max.lean b/src/analysis/complex/abs_max.lean index 151176781c868..86195f7f44d50 100644 --- a/src/analysis/complex/abs_max.lean +++ b/src/analysis/complex/abs_max.lean @@ -348,6 +348,18 @@ begin (λ x hx, (hr $ ball_subset_closed_ball hx).2)⟩ end +lemma eventually_eq_or_eq_zero_of_is_local_min_norm {f : E → ℂ} {c : E} + (hf : ∀ᶠ z in 𝓝 c, differentiable_at ℂ f z) (hc : is_local_min (norm ∘ f) c) : + (∀ᶠ z in 𝓝 c, f z = f c) ∨ (f c = 0) := +begin + refine or_iff_not_imp_right.mpr (λ h, _), + have h1 : ∀ᶠ z in 𝓝 c, f z ≠ 0 := hf.self_of_nhds.continuous_at.eventually_ne h, + have h2 : is_local_max (norm ∘ f)⁻¹ c := hc.inv (h1.mono (λ z, norm_pos_iff.mpr)), + have h3 : is_local_max (norm ∘ f⁻¹) c := by { refine h2.congr (eventually_of_forall _); simp }, + have h4 : ∀ᶠ z in 𝓝 c, differentiable_at ℂ f⁻¹ z, by filter_upwards [hf, h1] with z h using h.inv, + filter_upwards [eventually_eq_of_is_local_max_norm h4 h3] with z using inv_inj.mp +end + end strict_convex /-! diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index 777b13bd4428d..2f7410cca4c6b 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -570,6 +570,10 @@ begin exact ((hd.mono hRs).has_fpower_series_on_ball hR0).analytic_at end +lemma _root_.differentiable_on.analytic_on {s : set ℂ} {f : ℂ → E} (hd : differentiable_on ℂ f s) + (hs : is_open s) : analytic_on ℂ f s := +λ z hz, hd.analytic_at (hs.mem_nhds hz) + /-- A complex differentiable function `f : ℂ → E` is analytic at every point. -/ protected lemma _root_.differentiable.analytic_at {f : ℂ → E} (hf : differentiable ℂ f) (z : ℂ) : analytic_at ℂ f z := diff --git a/src/analysis/complex/open_mapping.lean b/src/analysis/complex/open_mapping.lean new file mode 100644 index 0000000000000..34d0abf9501c1 --- /dev/null +++ b/src/analysis/complex/open_mapping.lean @@ -0,0 +1,173 @@ +/- +Copyright (c) 2022 Vincent Beffara. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Vincent Beffara +-/ +import analysis.analytic.isolated_zeros +import analysis.complex.cauchy_integral +import analysis.complex.abs_max +import topology.algebra.field + +/-! +# The open mapping theorem for holomorphic functions + +This file proves the open mapping theorem for holomorphic functions, namely that an analytic +function on a preconnected set of the complex plane is either constant or open. The main step is to +show a local version of the theorem that states that if `f` is analytic at a point `z₀`, then either +it is constant in a neighborhood of `z₀` or it maps any neighborhood of `z₀` to a neighborhood of +its image `f z₀`. The results extend in higher dimension to `g : E → ℂ`. + +The proof of the local version on `ℂ` goes through two main steps: first, assuming that the function +is not constant around `z₀`, use the isolated zero principle to show that `∥f z∥` is bounded below +on a small `sphere z₀ r` around `z₀`, and then use the maximum principle applied to the auxiliary +function `(λ z, ∥f z - v∥)` to show that any `v` close enough to `f z₀` is in `f '' ball z₀ r`. That +second step is implemented in `diff_cont_on_cl.ball_subset_image_closed_ball`. + +## Main results + +* `analytic_at.eventually_constant_or_nhds_le_map_nhds` is the local version of the open mapping + theorem around a point; +* `analytic_on.is_constant_or_is_open` is the open mapping theorem on a connected open set. +-/ + +open set filter metric complex +open_locale topological_space + +variables {E : Type*} [normed_add_comm_group E] [normed_space ℂ E] {U : set E} + {f : ℂ → ℂ} {g : E → ℂ} {z₀ w : ℂ} {ε r m : ℝ} + +/-- If the modulus of a holomorphic function `f` is bounded below by `ε` on a circle, then its range +contains a disk of radius `ε / 2`. -/ +lemma diff_cont_on_cl.ball_subset_image_closed_ball (h : diff_cont_on_cl ℂ f (ball z₀ r)) + (hr : 0 < r) (hf : ∀ z ∈ sphere z₀ r, ε ≤ ∥f z - f z₀∥) (hz₀ : ∃ᶠ z in 𝓝 z₀, f z ≠ f z₀) : + ball (f z₀) (ε / 2) ⊆ f '' closed_ball z₀ r := +begin + /- This is a direct application of the maximum principle. Pick `v` close to `f z₀`, and look at + the function `λ z, ∥f z - v∥`: it is bounded below on the circle, and takes a small value at `z₀` + so it is not constant on the disk, which implies that its infimum is equal to `0` and hence that + `v` is in the range of `f`. -/ + rintro v hv, + have h1 : diff_cont_on_cl ℂ (λ z, f z - v) (ball z₀ r) := h.sub_const v, + have h2 : continuous_on (λ z, ∥f z - v∥) (closed_ball z₀ r), + from continuous_norm.comp_continuous_on (closure_ball z₀ hr.ne.symm ▸ h1.continuous_on), + have h3 : analytic_on ℂ f (ball z₀ r) := h.differentiable_on.analytic_on is_open_ball, + have h4 : ∀ z ∈ sphere z₀ r, ε / 2 ≤ ∥f z - v∥, + from λ z hz, by linarith [hf z hz, (show ∥v - f z₀∥ < ε / 2, from mem_ball.mp hv), + norm_sub_sub_norm_sub_le_norm_sub (f z) v (f z₀)], + have h5 : ∥f z₀ - v∥ < ε / 2 := by simpa [← dist_eq_norm, dist_comm] using mem_ball.mp hv, + obtain ⟨z, hz1, hz2⟩ : ∃ z ∈ ball z₀ r, is_local_min (λ z, ∥f z - v∥) z, + from exists_local_min_mem_ball h2 (mem_closed_ball_self hr.le) (λ z hz, h5.trans_le (h4 z hz)), + refine ⟨z, ball_subset_closed_ball hz1, sub_eq_zero.mp _⟩, + have h6 := h1.differentiable_on.eventually_differentiable_at (is_open_ball.mem_nhds hz1), + refine (eventually_eq_or_eq_zero_of_is_local_min_norm h6 hz2).resolve_left (λ key, _), + have h7 : ∀ᶠ w in 𝓝 z, f w = f z := by { filter_upwards [key] with h; field_simp }, + replace h7 : ∃ᶠ w in 𝓝[≠] z, f w = f z := (h7.filter_mono nhds_within_le_nhds).frequently, + have h8 : is_preconnected (ball z₀ r) := (convex_ball z₀ r).is_preconnected, + have h9 := h3.eq_on_of_preconnected_of_frequently_eq analytic_on_const h8 hz1 h7, + have h10 : f z = f z₀ := (h9 (mem_ball_self hr)).symm, + exact not_eventually.mpr hz₀ (mem_of_superset (ball_mem_nhds z₀ hr) (h10 ▸ h9)) +end + +/-- A function `f : ℂ → ℂ` which is analytic at a point `z₀` is either constant in a neighborhood +of `z₀`, or behaves locally like an open function (in the sense that the image of every neighborhood +of `z₀` is a neighborhood of `f z₀`, as in `is_open_map_iff_nhds_le`). For a function `f : E → ℂ` +the same result holds, see `analytic_at.eventually_constant_or_nhds_le_map_nhds`. -/ +lemma analytic_at.eventually_constant_or_nhds_le_map_nhds_aux (hf : analytic_at ℂ f z₀) : + (∀ᶠ z in 𝓝 z₀, f z = f z₀) ∨ (𝓝 (f z₀) ≤ map f (𝓝 z₀)) := +begin + /- The function `f` is analytic in a neighborhood of `z₀`; by the isolated zeros principle, if `f` + is not constant in a neighborhood of `z₀`, then it is nonzero, and therefore bounded below, on + every small enough circle around `z₀` and then `diff_cont_on_cl.ball_subset_image_closed_ball` + provides an explicit ball centered at `f z₀` contained in the range of `f`. -/ + refine or_iff_not_imp_left.mpr (λ h, _), + refine (nhds_basis_ball.le_basis_iff (nhds_basis_closed_ball.map f)).mpr (λ R hR, _), + have h1 := (hf.eventually_eq_or_eventually_ne analytic_at_const).resolve_left h, + have h2 : ∀ᶠ z in 𝓝 z₀, analytic_at ℂ f z := (is_open_analytic_at ℂ f).eventually_mem hf, + obtain ⟨ρ, hρ, h3, h4⟩ : ∃ ρ > 0, analytic_on ℂ f (closed_ball z₀ ρ) ∧ + ∀ z ∈ closed_ball z₀ ρ, z ≠ z₀ → f z ≠ f z₀, + by simpa only [set_of_and, subset_inter_iff] using + nhds_basis_closed_ball.mem_iff.mp (h2.and (eventually_nhds_within_iff.mp h1)), + replace h3 : diff_cont_on_cl ℂ f (ball z₀ ρ), + from ⟨h3.differentiable_on.mono ball_subset_closed_ball, + (closure_ball z₀ hρ.lt.ne.symm).symm ▸ h3.continuous_on⟩, + let r := ρ ⊓ R, + have hr : 0 < r := lt_inf_iff.mpr ⟨hρ, hR⟩, + have h5 : closed_ball z₀ r ⊆ closed_ball z₀ ρ := closed_ball_subset_closed_ball inf_le_left, + have h6 : diff_cont_on_cl ℂ f (ball z₀ r) := h3.mono (ball_subset_ball inf_le_left), + have h7 : ∀ z ∈ sphere z₀ r, f z ≠ f z₀, + from λ z hz, h4 z (h5 (sphere_subset_closed_ball hz)) (ne_of_mem_sphere hz hr.ne.symm), + have h8 : (sphere z₀ r).nonempty := normed_space.sphere_nonempty.mpr hr.le, + have h9 : continuous_on (λ x, ∥f x - f z₀∥) (sphere z₀ r), + from continuous_norm.comp_continuous_on + ((h6.sub_const (f z₀)).continuous_on_ball.mono sphere_subset_closed_ball), + obtain ⟨x, hx, hfx⟩ := (is_compact_sphere z₀ r).exists_forall_le h8 h9, + refine ⟨∥f x - f z₀∥ / 2, half_pos (norm_sub_pos_iff.mpr (h7 x hx)), _⟩, + exact (h6.ball_subset_image_closed_ball hr (λ z hz, hfx z hz) (not_eventually.mp h)).trans + (image_subset f (closed_ball_subset_closed_ball inf_le_right)) +end + +/-- The *open mapping theorem* for holomorphic functions, local version: is a function `g : E → ℂ` +is analytic at a point `z₀`, then either it is constant in a neighborhood of `z₀`, or it maps every +neighborhood of `z₀` to a neighborhood of `z₀`. For the particular case of a holomorphic function on +`ℂ`, see `analytic_at.eventually_constant_or_nhds_le_map_nhds_aux`. -/ +lemma analytic_at.eventually_constant_or_nhds_le_map_nhds {z₀ : E} (hg : analytic_at ℂ g z₀) : + (∀ᶠ z in 𝓝 z₀, g z = g z₀) ∨ (𝓝 (g z₀) ≤ map g (𝓝 z₀)) := +begin + /- The idea of the proof is to use the one-dimensional version applied to the restriction of `g` + to lines going through `z₀` (indexed by `sphere (0 : E) 1`). If the restriction is eventually + constant along each of these lines, then the identity theorem implies that `g` is constant on any + ball centered at `z₀` on which it is analytic, and in particular `g` is eventually constant. If on + the other hand there is one line along which `g` is not eventually constant, then the + one-dimensional version of the open mapping theorem can be used to conclude. -/ + let ray : E → ℂ → E := λ z t, z₀ + t • z, + let gray : E → ℂ → ℂ := λ z, (g ∘ ray z), + obtain ⟨r, hr, hgr⟩ := is_open_iff.mp (is_open_analytic_at ℂ g) z₀ hg, + have h1 : ∀ z ∈ sphere (0 : E) 1, analytic_on ℂ (gray z) (ball 0 r), + { refine λ z hz t ht, analytic_at.comp _ _, + { exact hgr (by simpa [ray, norm_smul, mem_sphere_zero_iff_norm.mp hz] using ht) }, + { exact analytic_at_const.add + ((continuous_linear_map.smul_right (continuous_linear_map.id ℂ ℂ) z).analytic_at t) } }, + by_cases (∀ z ∈ sphere (0 : E) 1, ∀ᶠ t in 𝓝 0, gray z t = gray z 0), + { left, -- If g is eventually constant along every direction, then it is eventually constant + refine eventually_of_mem (ball_mem_nhds z₀ hr) (λ z hz, _), + refine (eq_or_ne z z₀).cases_on (congr_arg g) (λ h', _), + replace h' : ∥z - z₀∥ ≠ 0 := by simpa only [ne.def, norm_eq_zero, sub_eq_zero], + let w : E := ∥z - z₀∥⁻¹ • (z - z₀), + have h3 : ∀ t ∈ ball (0 : ℂ) r, gray w t = g z₀, + { have e1 : is_preconnected (ball (0 : ℂ) r) := (convex_ball 0 r).is_preconnected, + have e2 : w ∈ sphere (0 : E) 1 := by simp [w, norm_smul, h'], + specialize h1 w e2, + apply h1.eq_on_of_preconnected_of_eventually_eq analytic_on_const e1 (mem_ball_self hr), + simpa [gray, ray] using h w e2 }, + have h4 : ∥z - z₀∥ < r := by simpa [dist_eq_norm] using mem_ball.mp hz, + replace h4 : ↑∥z - z₀∥ ∈ ball (0 : ℂ) r := by simpa only [mem_ball_zero_iff, norm_eq_abs, + abs_of_real, abs_norm_eq_norm], + simpa only [gray, ray, smul_smul, mul_inv_cancel h', one_smul, add_sub_cancel'_right, + function.comp_app, coe_smul] using h3 ↑∥z - z₀∥ h4 }, + { right, -- Otherwise, it is open along at least one direction and that implies the result + push_neg at h, + obtain ⟨z, hz, hrz⟩ := h, + specialize h1 z hz 0 (mem_ball_self hr), + have h7 := h1.eventually_constant_or_nhds_le_map_nhds_aux.resolve_left hrz, + rw [show gray z 0 = g z₀, by simp [gray, ray], ← map_compose] at h7, + refine h7.trans (map_mono _), + have h10 : continuous (λ (t : ℂ), z₀ + t • z), + from continuous_const.add (continuous_id'.smul continuous_const), + simpa using h10.tendsto 0 } +end + +/-- The *open mapping theorem* for holomorphic functions, global version: if a function `g : E → ℂ` +is analytic on a connected set `U`, then either it is constant on `U`, or it is open on `U` (in the +sense that it maps any open set contained in `U` to an open set in `ℂ`). -/ +theorem analytic_on.is_constant_or_is_open (hg : analytic_on ℂ g U) (hU : is_preconnected U) : + (∃ w, ∀ z ∈ U, g z = w) ∨ (∀ s ⊆ U, is_open s → is_open (g '' s)) := +begin + by_cases ∃ z₀ ∈ U, ∀ᶠ z in 𝓝 z₀, g z = g z₀, + { obtain ⟨z₀, hz₀, h⟩ := h, + exact or.inl ⟨g z₀, hg.eq_on_of_preconnected_of_eventually_eq analytic_on_const hU hz₀ h⟩ }, + { push_neg at h, + refine or.inr (λ s hs1 hs2, is_open_iff_mem_nhds.mpr _), + rintro z ⟨w, hw1, rfl⟩, + exact (hg w (hs1 hw1)).eventually_constant_or_nhds_le_map_nhds.resolve_left (h w (hs1 hw1)) + (image_mem_map (hs2.mem_nhds hw1)) } +end diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index 2ec4b28c52b73..ae023ed55045d 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -447,6 +447,10 @@ norm_le_of_mem_closed_ball' lemma norm_lt_of_mem_ball' (h : b ∈ ball a r) : ∥b∥ < ∥a∥ + r := (norm_le_norm_add_norm_div' _ _).trans_lt $ add_lt_add_left (by rwa ←dist_eq_norm_div) _ +@[to_additive] +lemma norm_div_sub_norm_div_le_norm_div (u v w : E) : ∥u / w∥ - ∥v / w∥ ≤ ∥u / v∥ := +by simpa only [div_div_div_cancel_right'] using norm_sub_norm_le' (u / w) (v / w) + @[to_additive bounded_iff_forall_norm_le] lemma bounded_iff_forall_norm_le' : bounded s ↔ ∃ C, ∀ x ∈ s, ∥x∥ ≤ C := by simpa only [set.subset_def, mem_closed_ball_one_iff] using bounded_iff_subset_ball (1 : E) diff --git a/src/topology/algebra/field.lean b/src/topology/algebra/field.lean index 98745ad7625cc..3764d0a56b597 100644 --- a/src/topology/algebra/field.lean +++ b/src/topology/algebra/field.lean @@ -5,6 +5,7 @@ Authors: Patrick Massot, Scott Morrison -/ import topology.algebra.ring import topology.algebra.group_with_zero +import topology.local_extr /-! # Topological fields @@ -141,3 +142,14 @@ def affine_homeomorph (a b : 𝕜) (h : a ≠ 0) : 𝕜 ≃ₜ 𝕜 := right_inv := λ y, by { simp [mul_div_cancel' _ h], }, } end affine_homeomorph + +section local_extr + +variables {α β : Type*} [topological_space α] [linear_ordered_semifield β] {a : α} +open_locale topological_space + +lemma is_local_min.inv {f : α → β} {a : α} (h1 : is_local_min f a) (h2 : ∀ᶠ z in 𝓝 a, 0 < f z) : + is_local_max f⁻¹ a := +by filter_upwards [h1, h2] with z h3 h4 using (inv_le_inv h4 h2.self_of_nhds).mpr h3 + +end local_extr diff --git a/src/topology/algebra/order/compact.lean b/src/topology/algebra/order/compact.lean index e4a19908650b9..2286eeec194fd 100644 --- a/src/topology/algebra/order/compact.lean +++ b/src/topology/algebra/order/compact.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Yury Kudryashov -/ import topology.algebra.order.intermediate_value +import topology.local_extr /-! # Compactness of a closed interval @@ -388,3 +389,24 @@ begin end end continuous_on + +lemma is_compact.exists_local_min_on_mem_subset {f : β → α} {s t : set β} {z : β} + (ht : is_compact t) (hf : continuous_on f t) (hz : z ∈ t) (hfz : ∀ z' ∈ t \ s, f z < f z') : + ∃ x ∈ s, is_local_min_on f t x := +begin + obtain ⟨x, hx, hfx⟩ : ∃ x ∈ t, ∀ y ∈ t, f x ≤ f y := ht.exists_forall_le ⟨z, hz⟩ hf, + have key : ∀ ⦃y⦄, y ∈ t → (∀ z' ∈ t \ s, f y < f z') → y ∈ s := λ y hy hfy, + by { by_contra; simpa using ((hfy y ((mem_diff y).mpr ⟨hy,h⟩))) }, + have h1 : ∀ z' ∈ t \ s, f x < f z' := λ z' hz', (hfx z hz).trans_lt (hfz z' hz'), + have h2 : x ∈ s := key hx h1, + refine ⟨x, h2, eventually_nhds_within_of_forall hfx⟩ +end + +lemma is_compact.exists_local_min_mem_open {f : β → α} {s t : set β} {z : β} (ht : is_compact t) + (hst : s ⊆ t) (hf : continuous_on f t) (hz : z ∈ t) (hfz : ∀ z' ∈ t \ s, f z < f z') + (hs : is_open s) : + ∃ x ∈ s, is_local_min f x := +begin + obtain ⟨x, hx, hfx⟩ := ht.exists_local_min_on_mem_subset hf hz hfz, + exact ⟨x, hx, hfx.is_local_min (filter.mem_of_superset (hs.mem_nhds hx) hst)⟩ +end diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 30fd0843e0972..d150d66ad7e1a 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -2493,6 +2493,17 @@ lemma nonempty_Inter_of_nonempty_bInter [complete_space α] {s : ℕ → set α} end diam +lemma exists_local_min_mem_ball [proper_space α] [topological_space β] + [conditionally_complete_linear_order β] [order_topology β] + {f : α → β} {a z : α} {r : ℝ} (hf : continuous_on f (closed_ball a r)) + (hz : z ∈ closed_ball a r) (hf1 : ∀ z' ∈ sphere a r, f z < f z') : + ∃ z ∈ ball a r, is_local_min f z := +begin + simp_rw [← closed_ball_diff_ball] at hf1, + exact (is_compact_closed_ball a r).exists_local_min_mem_open ball_subset_closed_ball hf hz hf1 + is_open_ball, +end + end metric namespace tactic From 7121691d5bab8da8857a43f83c67123bf033b1ef Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 14 Oct 2022 20:35:52 +0000 Subject: [PATCH 769/828] feat(data/multiset/basic): multiset inclusion is decidable (#16994) --- src/data/list/perm.lean | 3 +++ src/data/multiset/basic.lean | 3 +++ 2 files changed, 6 insertions(+) diff --git a/src/data/list/perm.lean b/src/data/list/perm.lean index 1aaddb4d71628..763deafcfd68a 100644 --- a/src/data/list/perm.lean +++ b/src/data/list/perm.lean @@ -767,6 +767,9 @@ begin convert (subperm_append_right _).mpr nil_subperm using 1 end +instance decidable_subperm : decidable_rel ((<+~) : list α → list α → Prop) := +λ l₁ l₂, decidable_of_iff _ list.subperm_ext_iff.symm + @[simp] lemma subperm_singleton_iff {α} {l : list α} {a : α} : [a] <+~ l ↔ a ∈ l := ⟨λ ⟨s, hla, h⟩, by rwa [perm_singleton.mp hla, singleton_sublist] at h, λ h, ⟨[a], perm.refl _, singleton_sublist.mpr h⟩⟩ diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 040a00767e83a..7c52815e8b238 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -331,6 +331,9 @@ instance : partial_order (multiset α) := le_trans := by rintros ⟨l₁⟩ ⟨l₂⟩ ⟨l₃⟩; exact @subperm.trans _ _ _ _, le_antisymm := by rintros ⟨l₁⟩ ⟨l₂⟩ h₁ h₂; exact quot.sound (subperm.antisymm h₁ h₂) } +instance decidable_le [decidable_eq α] : decidable_rel ((≤) : multiset α → multiset α → Prop) := +λ s t, quotient.rec_on_subsingleton₂ s t list.decidable_subperm + section variables {s t : multiset α} {a : α} From 31263840fccb6bc5ea3d3d49676a1d16d596cfc0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Fri, 14 Oct 2022 23:58:26 +0000 Subject: [PATCH 770/828] refactor(data/finset/noncomm_prod): Use `set.pairwise` (#16859) Redefine the various `noncomm_prod`s using `set.pairwise`. This has the advantage of having a solid API around monotonicity in the set or in the relation and avoiding checking the trivial `i = j` case. --- src/algebra/group/commute.lean | 6 + src/analysis/normed_space/exponential.lean | 13 +- .../normed_space/matrix_exponential.lean | 4 +- src/data/finset/basic.lean | 3 + src/data/finset/noncomm_prod.lean | 232 +++++++----------- src/data/finset/pairwise.lean | 19 +- src/data/list/nodup.lean | 14 +- src/group_theory/noncomm_pi_coprod.lean | 5 +- src/group_theory/perm/cycle/basic.lean | 68 ++--- src/group_theory/perm/cycle/type.lean | 6 +- src/group_theory/perm/support.lean | 2 + src/group_theory/subgroup/basic.lean | 18 +- src/group_theory/submonoid/membership.lean | 7 +- 13 files changed, 160 insertions(+), 237 deletions(-) diff --git a/src/algebra/group/commute.lean b/src/algebra/group/commute.lean index 0a52ee04dedee..80133fc0323cf 100644 --- a/src/algebra/group/commute.lean +++ b/src/algebra/group/commute.lean @@ -54,6 +54,12 @@ protected lemma symm {a b : S} (h : commute a b) : commute b a := eq.symm h protected theorem symm_iff {a b : S} : commute a b ↔ commute b a := ⟨commute.symm, commute.symm⟩ +@[to_additive] instance : is_refl S commute := ⟨commute.refl⟩ + +-- This instance is useful for `finset.noncomm_prod` +@[to_additive] instance on_is_refl {f : G → S} : is_refl G (λ a b, commute (f a) (f b)) := +⟨λ _, commute.refl _⟩ + end has_mul section semigroup diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 81581cc4ea78e..882cf9621fcde 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -452,18 +452,17 @@ end /-- In a Banach-algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ`, if a family of elements `f i` mutually commute then `exp 𝕂 (∑ i, f i) = ∏ i, exp 𝕂 (f i)`. -/ lemma exp_sum_of_commute {ι} (s : finset ι) (f : ι → 𝔸) - (h : ∀ (i ∈ s) (j ∈ s), commute (f i) (f j)) : + (h : (s : set ι).pairwise $ λ i j, commute (f i) (f j)) : exp 𝕂 (∑ i in s, f i) = s.noncomm_prod (λ i, exp 𝕂 (f i)) - (λ i hi j hj, (h i hi j hj).exp 𝕂) := + (λ i hi j hj _, (h.of_refl hi hj).exp 𝕂) := begin classical, induction s using finset.induction_on with a s ha ih, { simp }, rw [finset.noncomm_prod_insert_of_not_mem _ _ _ _ ha, finset.sum_insert ha, - exp_add_of_commute, ih], - refine commute.sum_right _ _ _ _, - intros i hi, - exact h _ (finset.mem_insert_self _ _) _ (finset.mem_insert_of_mem hi), + exp_add_of_commute, ih (h.mono $ finset.subset_insert _ _)], + refine commute.sum_right _ _ _ (λ i hi, _), + exact h.of_refl (finset.mem_insert_self _ _) (finset.mem_insert_of_mem hi), end lemma exp_nsmul (n : ℕ) (x : 𝔸) : @@ -591,7 +590,7 @@ lemma exp_sum {ι} (s : finset ι) (f : ι → 𝔸) : exp 𝕂 (∑ i in s, f i) = ∏ i in s, exp 𝕂 (f i) := begin rw [exp_sum_of_commute, finset.noncomm_prod_eq_prod], - exact λ i hi j hj, commute.all _ _, + exact λ i hi j hj _, commute.all _ _, end end comm_algebra diff --git a/src/analysis/normed_space/matrix_exponential.lean b/src/analysis/normed_space/matrix_exponential.lean index 1a1aea1e5afa5..423ad6dfbd9ea 100644 --- a/src/analysis/normed_space/matrix_exponential.lean +++ b/src/analysis/normed_space/matrix_exponential.lean @@ -149,9 +149,9 @@ begin end lemma exp_sum_of_commute {ι} (s : finset ι) (f : ι → matrix m m 𝔸) - (h : ∀ (i ∈ s) (j ∈ s), commute (f i) (f j)) : + (h : (s : set ι).pairwise $ λ i j, commute (f i) (f j)) : exp 𝕂 (∑ i in s, f i) = s.noncomm_prod (λ i, exp 𝕂 (f i)) - (λ i hi j hj, (h i hi j hj).exp 𝕂) := + (λ i hi j hj _, (h.of_refl hi hj).exp 𝕂) := begin letI : semi_normed_ring (matrix m m 𝔸) := matrix.linfty_op_semi_normed_ring, letI : normed_ring (matrix m m 𝔸) := matrix.linfty_op_normed_ring, diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 36c919d376986..7aefff28180fa 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -2025,6 +2025,9 @@ def to_finset (l : list α) : finset α := multiset.to_finset l lemma to_finset_eq (n : nodup l) : @finset.mk α l n = l.to_finset := multiset.to_finset_eq n @[simp] lemma mem_to_finset : a ∈ l.to_finset ↔ a ∈ l := mem_dedup +@[simp, norm_cast] lemma coe_to_finset (l : list α) : (l.to_finset : set α) = {a | a ∈ l} := +set.ext $ λ _, list.mem_to_finset + @[simp] lemma to_finset_nil : to_finset (@nil α) = ∅ := rfl @[simp] lemma to_finset_cons : to_finset (a :: l) = insert a (to_finset l) := diff --git a/src/data/finset/noncomm_prod.lean b/src/data/finset/noncomm_prod.lean index 5cef6169bb93b..be920fd0a5a3d 100644 --- a/src/data/finset/noncomm_prod.lean +++ b/src/data/finset/noncomm_prod.lean @@ -27,18 +27,21 @@ the two must be equal. -/ -variables {α β γ : Type*} (f : α → β → β) (op : α → α → α) +variables {F ι α β γ : Type*} (f : α → β → β) (op : α → α → α) namespace multiset /-- Fold of a `s : multiset α` with `f : α → β → β`, given a proof that `left_commutative f` on all elements `x ∈ s`. -/ -def noncomm_foldr (s : multiset α) (comm : ∀ (x ∈ s) (y ∈ s) b, f x (f y b) = f y (f x b)) (b : β) : +def noncomm_foldr (s : multiset α) + (comm : {x | x ∈ s}.pairwise $ λ x y, ∀ b, f x (f y b) = f y (f x b)) (b : β) : β := -s.attach.foldr (f ∘ subtype.val) (λ ⟨x, hx⟩ ⟨y, hy⟩, comm x hx y hy) b +s.attach.foldr (f ∘ subtype.val) (λ ⟨x, hx⟩ ⟨y, hy⟩, begin + haveI : is_refl α (λ x y, ∀ b, f x (f y b) = f y (f x b)) := ⟨λ x b, rfl⟩, + exact comm.of_refl hx hy, + end) b -@[simp] lemma noncomm_foldr_coe (l : list α) - (comm : ∀ (x ∈ (l : multiset α)) (y ∈ (l : multiset α)) b, f x (f y b) = f y (f x b)) (b : β) : +@[simp] lemma noncomm_foldr_coe (l : list α) (comm) (b : β) : noncomm_foldr f (l : multiset α) comm b = l.foldr f b := begin simp only [noncomm_foldr, coe_foldr, coe_attach, list.attach], @@ -46,14 +49,9 @@ begin simp [list.map_pmap, list.pmap_eq_map] end -@[simp] lemma noncomm_foldr_empty - (h : ∀ (x ∈ (0 : multiset α)) (y ∈ (0 : multiset α)) b, f x (f y b) = f y (f x b)) (b : β) : - noncomm_foldr f (0 : multiset α) h b = b := rfl +@[simp] lemma noncomm_foldr_empty (h) (b : β) : noncomm_foldr f (0 : multiset α) h b = b := rfl -lemma noncomm_foldr_cons (s : multiset α) (a : α) - (h : ∀ (x ∈ (a ::ₘ s)) (y ∈ (a ::ₘ s)) b, f x (f y b) = f y (f x b)) - (h' : ∀ (x ∈ s) (y ∈ s) b, f x (f y b) = f y (f x b)) - (b : β) : +lemma noncomm_foldr_cons (s : multiset α) (a : α) (h h') (b : β) : noncomm_foldr f (a ::ₘ s) h b = f a (noncomm_foldr f s h' b) := begin induction s using quotient.induction_on, @@ -61,7 +59,7 @@ begin end lemma noncomm_foldr_eq_foldr (s : multiset α) (h : left_commutative f) (b : β) : - noncomm_foldr f s (λ x _ y _, h x y) b = foldr f h b s := + noncomm_foldr f s (λ x _ y _ _, h x y) b = foldr f h b s := begin induction s using quotient.induction_on, simp @@ -72,34 +70,24 @@ include assoc /-- Fold of a `s : multiset α` with an associative `op : α → α → α`, given a proofs that `op` is commutative on all elements `x ∈ s`. -/ -def noncomm_fold (s : multiset α) - (comm : ∀ (x ∈ s) (y ∈ s), op x y = op y x) - (a : α) : α := -noncomm_foldr op s (λ x hx y hy b, by rw [←assoc.assoc, comm _ hx _ hy, assoc.assoc]) a - -@[simp] lemma noncomm_fold_coe (l : list α) - (comm : ∀ (x ∈ (l : multiset α)) (y ∈ (l : multiset α)), op x y = op y x) - (a : α) : +def noncomm_fold (s : multiset α) (comm : {x | x ∈ s}.pairwise $ λ x y, op x y = op y x) : α → α := +noncomm_foldr op s (λ x hx y hy h b, by rw [←assoc.assoc, comm hx hy h, assoc.assoc]) + +@[simp] lemma noncomm_fold_coe (l : list α) (comm) (a : α) : noncomm_fold op (l : multiset α) comm a = l.foldr op a := by simp [noncomm_fold] -@[simp] lemma noncomm_fold_empty - (h : ∀ (x ∈ (0 : multiset α)) (y ∈ (0 : multiset α)), op x y = op y x) (a : α) : - noncomm_fold op (0 : multiset α) h a = a := rfl +@[simp] lemma noncomm_fold_empty (h) (a : α) : noncomm_fold op (0 : multiset α) h a = a := rfl -lemma noncomm_fold_cons (s : multiset α) (a : α) - (h : ∀ (x ∈ a ::ₘ s) (y ∈ a ::ₘ s), op x y = op y x) - (h' : ∀ (x ∈ s) (y ∈ s), op x y = op y x) - (x : α) : +lemma noncomm_fold_cons (s : multiset α) (a : α) (h h') (x : α) : noncomm_fold op (a ::ₘ s) h x = op a (noncomm_fold op s h' x) := begin induction s using quotient.induction_on, simp end -lemma noncomm_fold_eq_fold (s : multiset α) [is_commutative α op] - (a : α) : - noncomm_fold op s (λ x _ y _, is_commutative.comm x y) a = fold op a s := +lemma noncomm_fold_eq_fold (s : multiset α) [is_commutative α op] (a : α) : + noncomm_fold op s (λ x _ y _ _, is_commutative.comm x y) a = fold op a s := begin induction s using quotient.induction_on, simp @@ -112,11 +100,10 @@ variables [monoid α] [monoid β] on all elements `x ∈ s`. -/ @[to_additive "Sum of a `s : multiset α` with `[add_monoid α]`, given a proof that `+` commutes on all elements `x ∈ s`." ] -def noncomm_prod (s : multiset α) (comm : ∀ (x ∈ s) (y ∈ s), commute x y) : α := +def noncomm_prod (s : multiset α) (comm : {x | x ∈ s}.pairwise commute) : α := s.noncomm_fold (*) comm 1 -@[simp, to_additive] lemma noncomm_prod_coe (l : list α) - (comm : ∀ (x ∈ (l : multiset α)) (y ∈ (l : multiset α)), commute x y) : +@[simp, to_additive] lemma noncomm_prod_coe (l : list α) (comm) : noncomm_prod (l : multiset α) comm = l.prod := begin rw [noncomm_prod], @@ -125,26 +112,17 @@ begin { simp }, { rw [list.prod_cons, list.foldr, hl], intros x hx y hy, - exact comm x (list.mem_cons_of_mem _ hx) y (list.mem_cons_of_mem _ hy) } + exact comm (list.mem_cons_of_mem _ hx) (list.mem_cons_of_mem _ hy) } end -@[simp, to_additive] lemma noncomm_prod_empty - (h : ∀ (x ∈ (0 : multiset α)) (y ∈ (0 : multiset α)), commute x y) : - noncomm_prod (0 : multiset α) h = 1 := rfl +@[simp, to_additive] lemma noncomm_prod_empty (h) : noncomm_prod (0 : multiset α) h = 1 := rfl -@[simp, to_additive] lemma noncomm_prod_cons (s : multiset α) (a : α) - (comm : ∀ (x ∈ a ::ₘ s) (y ∈ a ::ₘ s), commute x y) : - noncomm_prod (a ::ₘ s) comm = a * noncomm_prod s - (λ x hx y hy, comm _ (mem_cons_of_mem hx) _ (mem_cons_of_mem hy)) := -begin - induction s using quotient.induction_on, - simp -end +@[simp, to_additive] lemma noncomm_prod_cons (s : multiset α) (a : α) (comm) : + noncomm_prod (a ::ₘ s) comm = a * noncomm_prod s (comm.mono $ λ _, mem_cons_of_mem) := +by { induction s using quotient.induction_on, simp } -@[to_additive] lemma noncomm_prod_cons' (s : multiset α) (a : α) - (comm : ∀ (x ∈ a ::ₘ s) (y ∈ a ::ₘ s), commute x y) : - noncomm_prod (a ::ₘ s) comm = noncomm_prod s - (λ x hx y hy, comm _ (mem_cons_of_mem hx) _ (mem_cons_of_mem hy)) * a := +@[to_additive] lemma noncomm_prod_cons' (s : multiset α) (a : α) (comm) : + noncomm_prod (a ::ₘ s) comm = noncomm_prod s (comm.mono $ λ _, mem_cons_of_mem) * a := begin induction s using quotient.induction_on with s, simp only [quot_mk_to_coe, cons_coe, noncomm_prod_coe, list.prod_cons], @@ -152,7 +130,7 @@ begin { simp }, { rw [list.prod_cons, mul_assoc, ←IH, ←mul_assoc, ←mul_assoc], { congr' 1, - apply comm; + apply comm.of_refl; simp }, { intros x hx y hy, simp only [quot_mk_to_coe, list.mem_cons_iff, mem_coe, cons_coe] at hx hy, @@ -164,28 +142,26 @@ begin end @[protected, to_additive] -lemma nocomm_prod_map_aux (s : multiset α) - (comm : ∀ (x ∈ s) (y ∈ s), commute x y) - {F : Type*} [monoid_hom_class F α β] (f : F) : - ∀ (x ∈ s.map f) (y ∈ s.map f), commute x y := +lemma noncomm_prod_map_aux [monoid_hom_class F α β] (s : multiset α) + (comm : {x | x ∈ s}.pairwise commute) (f : F) : + {x | x ∈ s.map f}.pairwise commute := begin simp only [multiset.mem_map], - rintros _ ⟨x, hx, rfl⟩ _ ⟨y, hy, rfl⟩, - exact (comm _ hx _ hy).map f, + rintros _ ⟨x, hx, rfl⟩ _ ⟨y, hy, rfl⟩ _, + exact (comm.of_refl hx hy).map f, end @[to_additive] -lemma noncomm_prod_map (s : multiset α) (comm : ∀ (x ∈ s) (y ∈ s), commute x y) - {F : Type*} [monoid_hom_class F α β] (f : F) : - f (s.noncomm_prod comm) = (s.map f).noncomm_prod (nocomm_prod_map_aux s comm f) := +lemma noncomm_prod_map [monoid_hom_class F α β] (s : multiset α) (comm) (f : F) : + f (s.noncomm_prod comm) = (s.map f).noncomm_prod (noncomm_prod_map_aux s comm f) := begin induction s using quotient.induction_on, simpa using map_list_prod f _, end @[to_additive noncomm_sum_eq_card_nsmul] -lemma noncomm_prod_eq_pow_card (s : multiset α) (comm : ∀ (x ∈ s) (y ∈ s), commute x y) - (m : α) (h : ∀ (x ∈ s), x = m) : s.noncomm_prod comm = m ^ s.card := +lemma noncomm_prod_eq_pow_card (s : multiset α) (comm) (m : α) (h : ∀ x ∈ s, x = m) : + s.noncomm_prod comm = m ^ s.card := begin induction s using quotient.induction_on, simp only [quot_mk_to_coe, noncomm_prod_coe, coe_card, mem_coe] at *, @@ -193,16 +169,15 @@ begin end @[to_additive] lemma noncomm_prod_eq_prod {α : Type*} [comm_monoid α] (s : multiset α) : - noncomm_prod s (λ _ _ _ _, commute.all _ _) = prod s := + noncomm_prod s (λ _ _ _ _ _, commute.all _ _) = prod s := begin induction s using quotient.induction_on, simp end @[to_additive noncomm_sum_add_commute] -lemma noncomm_prod_commute (s : multiset α) - (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute x y) - (y : α) (h : ∀ (x : α), x ∈ s → commute y x) : commute y (s.noncomm_prod comm) := +lemma noncomm_prod_commute (s : multiset α) (comm) (y : α) (h : ∀ x ∈ s, commute y x) : + commute y (s.noncomm_prod comm) := begin induction s using quotient.induction_on, simp only [quot_mk_to_coe, noncomm_prod_coe], @@ -219,19 +194,19 @@ variables [monoid β] [monoid γ] given a proof that `*` commutes on all elements `f x` for `x ∈ s`. -/ @[to_additive "Sum of a `s : finset α` mapped with `f : α → β` with `[add_monoid β]`, given a proof that `+` commutes on all elements `f x` for `x ∈ s`."] -def noncomm_prod (s : finset α) (f : α → β) (comm : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y)) : β := -(s.1.map f).noncomm_prod (by simpa [multiset.mem_map, ←finset.mem_def] using comm) +def noncomm_prod (s : finset α) (f : α → β) + (comm : (s : set α).pairwise (λ a b, commute (f a) (f b))) : β := +(s.1.map f).noncomm_prod $ + by { simp_rw multiset.mem_map, rintro _ ⟨a, ha, rfl⟩ _ ⟨b, hb, rfl⟩ _, exact comm.of_refl ha hb } @[congr, to_additive] -lemma noncomm_prod_congr - {s₁ s₂ : finset α} {f g : α → β} (h₁ : s₁ = s₂) (h₂ : ∀ (x ∈ s₂), f x = g x) - (comm : ∀ (x ∈ s₁) (y ∈ s₁), commute (f x) (f y)) : +lemma noncomm_prod_congr {s₁ s₂ : finset α} {f g : α → β} (h₁ : s₁ = s₂) + (h₂ : ∀ (x ∈ s₂), f x = g x) (comm) : noncomm_prod s₁ f comm = noncomm_prod s₂ g - (λ x hx y hy, h₂ x hx ▸ h₂ y hy ▸ comm x (h₁.symm ▸ hx) y (h₁.symm ▸ hy)) := + (λ x hx y hy h, by { rw [←h₂ _ hx, ←h₂ _ hy], subst h₁, exact comm hx hy h }) := by simp_rw [noncomm_prod, multiset.map_congr (congr_arg _ h₁) h₂] -@[simp, to_additive] lemma noncomm_prod_to_finset [decidable_eq α] (l : list α) (f : α → β) - (comm : ∀ (x ∈ l.to_finset) (y ∈ l.to_finset), commute (f x) (f y)) +@[simp, to_additive] lemma noncomm_prod_to_finset [decidable_eq α] (l : list α) (f : α → β) (comm) (hl : l.nodup) : noncomm_prod l.to_finset f comm = (l.map f).prod := begin @@ -239,43 +214,32 @@ begin simp [noncomm_prod, hl] end -@[simp, to_additive] lemma noncomm_prod_empty (f : α → β) - (h : ∀ (x ∈ (∅ : finset α)) (y ∈ (∅ : finset α)), commute (f x) (f y)) : +@[simp, to_additive] lemma noncomm_prod_empty (f : α → β) (h) : noncomm_prod (∅ : finset α) f h = 1 := rfl @[simp, to_additive] lemma noncomm_prod_insert_of_not_mem [decidable_eq α] (s : finset α) (a : α) - (f : α → β) - (comm : ∀ (x ∈ insert a s) (y ∈ insert a s), commute (f x) (f y)) - (ha : a ∉ s) : - noncomm_prod (insert a s) f comm = f a * noncomm_prod s f - (λ x hx y hy, comm _ (mem_insert_of_mem hx) _ (mem_insert_of_mem hy)) := + (f : α → β) (comm) (ha : a ∉ s) : + noncomm_prod (insert a s) f comm = f a * noncomm_prod s f (comm.mono $ λ _, mem_insert_of_mem) := by simp [insert_val_of_not_mem ha, noncomm_prod] @[to_additive] lemma noncomm_prod_insert_of_not_mem' [decidable_eq α] (s : finset α) (a : α) - (f : α → β) - (comm : ∀ (x ∈ insert a s) (y ∈ insert a s), commute (f x) (f y)) - (ha : a ∉ s) : - noncomm_prod (insert a s) f comm = noncomm_prod s f - (λ x hx y hy, comm _ (mem_insert_of_mem hx) _ (mem_insert_of_mem hy)) * f a := + (f : α → β) (comm) (ha : a ∉ s) : + noncomm_prod (insert a s) f comm = noncomm_prod s f (comm.mono $ λ _, mem_insert_of_mem) * f a := by simp [noncomm_prod, insert_val_of_not_mem ha, multiset.noncomm_prod_cons'] @[simp, to_additive] lemma noncomm_prod_singleton (a : α) (f : α → β) : - noncomm_prod ({a} : finset α) f - (λ x hx y hy, by rw [mem_singleton.mp hx, mem_singleton.mp hy]) = f a := + noncomm_prod ({a} : finset α) f (by { norm_cast, exact set.pairwise_singleton _ _ }) = f a := by simp [noncomm_prod, ←multiset.cons_zero] @[to_additive] -lemma noncomm_prod_map (s : finset α) (f : α → β) - (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - {F : Type*} [monoid_hom_class F β γ] (g : F) : +lemma noncomm_prod_map [monoid_hom_class F β γ] (s : finset α) (f : α → β) (comm) (g : F) : g (s.noncomm_prod f comm) = s.noncomm_prod (λ i, g (f i)) - (λ x hx y hy, (comm x hx y hy).map g) := + (λ x hx y hy h, (comm.of_refl hx hy).map g) := by simp [noncomm_prod, multiset.noncomm_prod_map] @[to_additive noncomm_sum_eq_card_nsmul] -lemma noncomm_prod_eq_pow_card (s : finset α) (f : α → β) - (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - (m : β) (h : ∀ (x : α), x ∈ s → f x = m) : s.noncomm_prod f comm = m ^ s.card := +lemma noncomm_prod_eq_pow_card (s : finset α) (f : α → β) (comm) (m : β) (h : ∀ x ∈ s, f x = m) : + s.noncomm_prod f comm = m ^ s.card := begin rw [noncomm_prod, multiset.noncomm_prod_eq_pow_card _ _ m], simp only [finset.card_def, multiset.card_map], @@ -283,9 +247,8 @@ begin end @[to_additive noncomm_sum_add_commute] -lemma noncomm_prod_commute (s : finset α) (f : α → β) - (comm : ∀ (x : α), x ∈ s → ∀ (y : α), y ∈ s → commute (f x) (f y)) - (y : β) (h : ∀ (x : α), x ∈ s → commute y (f x)) : commute y (s.noncomm_prod f comm) := +lemma noncomm_prod_commute (s : finset α) (f : α → β) (comm) (y : β) + (h : ∀ x ∈ s, commute y (f x)) : commute y (s.noncomm_prod f comm) := begin apply multiset.noncomm_prod_commute, intro y, @@ -295,7 +258,7 @@ begin end @[to_additive] lemma noncomm_prod_eq_prod {β : Type*} [comm_monoid β] (s : finset α) (f : α → β) : - noncomm_prod s f (λ _ _ _ _, commute.all _ _) = s.prod f := + noncomm_prod s f (λ _ _ _ _ _, commute.all _ _) = s.prod f := begin classical, induction s using finset.induction_on with a s ha IH, @@ -303,16 +266,12 @@ begin { simp [ha, IH] } end -/- The non-commutative version of `finset.prod_union` -/ +/-- The non-commutative version of `finset.prod_union` -/ @[to_additive "The non-commutative version of `finset.sum_union`"] lemma noncomm_prod_union_of_disjoint [decidable_eq α] {s t : finset α} - (h : disjoint s t) (f : α → β) - (comm : ∀ (x ∈ s ∪ t) (y ∈ s ∪ t), commute (f x) (f y)) - (scomm : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y) := - λ _ hx _ hy, comm _ (mem_union_left _ hx) _ (mem_union_left _ hy)) - (tcomm : ∀ (x ∈ t) (y ∈ t), commute (f x) (f y) := - λ _ hx _ hy, comm _ (mem_union_right _ hx) _ (mem_union_right _ hy)) : - noncomm_prod (s ∪ t) f comm = noncomm_prod s f scomm * noncomm_prod t f tcomm := + (h : disjoint s t) (f : α → β) (comm : {x | x ∈ s ∪ t}.pairwise $ λ a b, commute (f a) (f b)) : + noncomm_prod (s ∪ t) f comm = noncomm_prod s f (comm.mono $ coe_subset.2 $ subset_union_left _ _) + * noncomm_prod t f (comm.mono $ coe_subset.2 $ subset_union_right _ _) := begin obtain ⟨sl, sl', rfl⟩ := exists_list_nodup_eq s, obtain ⟨tl, tl', rfl⟩ := exists_list_nodup_eq t, @@ -323,56 +282,46 @@ end @[protected, to_additive] lemma noncomm_prod_mul_distrib_aux {s : finset α} {f : α → β} {g : α → β} - (comm_ff : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y)) - (comm_gg : ∀ (x ∈ s) (y ∈ s), commute (g x) (g y)) - (comm_gf : ∀ (x ∈ s) (y ∈ s), x ≠ y → commute (g x) (f y)) : - (∀ (x ∈ s) (y ∈ s), commute ((f * g) x) ((f * g) y)) := + (comm_ff : (s : set α).pairwise $ λ x y, commute (f x) (f y)) + (comm_gg : (s : set α).pairwise $ λ x y, commute (g x) (g y)) + (comm_gf : (s : set α).pairwise $ λ x y, commute (g x) (f y)) : + (s : set α).pairwise $ λ x y, commute ((f * g) x) ((f * g) y) := begin - intros x hx y hy, - by_cases h : x = y, { subst h }, + intros x hx y hy h, apply commute.mul_left; apply commute.mul_right, - { exact comm_ff x hx y hy }, - { exact (comm_gf y hy x hx (ne.symm h)).symm }, - { exact comm_gf x hx y hy h }, - { exact comm_gg x hx y hy }, + { exact comm_ff.of_refl hx hy }, + { exact (comm_gf hy hx h.symm).symm }, + { exact comm_gf hx hy h }, + { exact comm_gg.of_refl hx hy } end /-- The non-commutative version of `finset.prod_mul_distrib` -/ @[to_additive "The non-commutative version of `finset.sum_add_distrib`"] -lemma noncomm_prod_mul_distrib {s : finset α} (f : α → β) (g : α → β) - (comm_ff : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y)) - (comm_gg : ∀ (x ∈ s) (y ∈ s), commute (g x) (g y)) - (comm_gf : ∀ (x ∈ s) (y ∈ s), x ≠ y → commute (g x) (f y)) : +lemma noncomm_prod_mul_distrib {s : finset α} (f : α → β) (g : α → β) (comm_ff comm_gg comm_gf) : noncomm_prod s (f * g) (noncomm_prod_mul_distrib_aux comm_ff comm_gg comm_gf) = noncomm_prod s f comm_ff * noncomm_prod s g comm_gg := begin classical, induction s using finset.induction_on with x s hnmem ih, { simp, }, - { simp only [finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem], - specialize ih - (λ x hx y hy, comm_ff x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) - (λ x hx y hy, comm_gg x (mem_insert_of_mem hx) y (mem_insert_of_mem hy)) - (λ x hx y hy hne, comm_gf x (mem_insert_of_mem hx) y (mem_insert_of_mem hy) hne), - rw [ih, pi.mul_apply], - simp only [mul_assoc], - congr' 1, - simp only [← mul_assoc], - congr' 1, - apply noncomm_prod_commute, - intros y hy, - have : x ≠ y, by {rintro rfl, contradiction}, - exact comm_gf x (mem_insert_self x s) y (mem_insert_of_mem hy) this, } + simp only [finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem], + specialize ih (comm_ff.mono $ λ _, mem_insert_of_mem) (comm_gg.mono $ λ _, mem_insert_of_mem) + (comm_gf.mono $ λ _, mem_insert_of_mem), + rw [ih, pi.mul_apply], + simp only [mul_assoc], + congr' 1, + simp only [← mul_assoc], + congr' 1, + refine noncomm_prod_commute _ _ _ _ (λ y hy, _), + exact comm_gf (mem_insert_self x s) (mem_insert_of_mem hy) (ne_of_mem_of_not_mem hy hnmem).symm, end section finite_pi - -variables {ι : Type*} [fintype ι] [decidable_eq ι] {M : ι → Type*} [∀ i, monoid (M i)] -variables (x : Π i, M i) +variables {M : ι → Type*} [Π i, monoid (M i)] @[to_additive] -lemma noncomm_prod_mul_single : - univ.noncomm_prod (λ i, pi.mul_single i (x i)) (λ i _ j _, pi.mul_single_apply_commute x i j) +lemma noncomm_prod_mul_single [fintype ι] [decidable_eq ι] (x : Π i, M i) : + univ.noncomm_prod (λ i, pi.mul_single i (x i)) (λ i _ j _ _, pi.mul_single_apply_commute x i j) = x := begin ext i, @@ -387,10 +336,11 @@ begin end @[to_additive] -lemma _root_.monoid_hom.pi_ext {f g : (Π i, M i) →* γ} +lemma _root_.monoid_hom.pi_ext [finite ι] [decidable_eq ι] {f g : (Π i, M i) →* γ} (h : ∀ i x, f (pi.mul_single i x) = g (pi.mul_single i x)) : f = g := begin + casesI nonempty_fintype ι, ext x, rw [← noncomm_prod_mul_single x, univ.noncomm_prod_map, univ.noncomm_prod_map], congr' 1 with i, exact h i (x i), diff --git a/src/data/finset/pairwise.lean b/src/data/finset/pairwise.lean index 47e14a48da52e..b9346127ed985 100644 --- a/src/data/finset/pairwise.lean +++ b/src/data/finset/pairwise.lean @@ -70,26 +70,11 @@ variables {β : Type*} [decidable_eq α] {r : α → α → Prop} {l : list α} lemma pairwise_of_coe_to_finset_pairwise (hl : (l.to_finset : set α).pairwise r) (hn : l.nodup) : l.pairwise r := -begin - induction l with hd tl IH, - { simp }, - simp only [set.pairwise_insert, pairwise_cons, to_finset_cons, finset.coe_insert, - finset.mem_coe, mem_to_finset, ne.def, nodup_cons] at hl hn ⊢, - refine ⟨λ x hx, (hl.right x hx _).left, IH hl.left hn.right⟩, - rintro rfl, - exact hn.left hx -end +by { rw coe_to_finset at hl, exact hn.pairwise_of_set_pairwise hl } lemma pairwise_iff_coe_to_finset_pairwise (hn : l.nodup) (hs : symmetric r) : (l.to_finset : set α).pairwise r ↔ l.pairwise r := -begin - refine ⟨λ h, pairwise_of_coe_to_finset_pairwise h hn, λ h, _⟩, - induction l with hd tl IH, - { simp }, - simp only [set.pairwise_insert, to_finset_cons, finset.coe_insert, finset.mem_coe, - mem_to_finset, ne.def, pairwise_cons, nodup_cons] at hn h ⊢, - exact ⟨IH hn.right h.right, λ x hx hne, ⟨h.left _ hx, hs (h.left _ hx)⟩⟩ -end +by { rw [coe_to_finset, hn.pairwise_coe], exact ⟨hs⟩ } lemma pairwise_disjoint_of_coe_to_finset_pairwise_disjoint {α ι} [semilattice_inf α] [order_bot α] [decidable_eq ι] {l : list ι} {f : ι → α} diff --git a/src/data/list/nodup.lean b/src/data/list/nodup.lean index 1a77d853637ca..b15a56f6bd412 100644 --- a/src/data/list/nodup.lean +++ b/src/data/list/nodup.lean @@ -18,7 +18,7 @@ universes u v open nat function -variables {α : Type u} {β : Type v} {l l₁ l₂ : list α} {a b : α} +variables {α : Type u} {β : Type v} {l l₁ l₂ : list α} {r : α → α → Prop} {a b : α} namespace list @@ -352,6 +352,18 @@ lemma nodup.pairwise_of_set_pairwise {l : list α} {r : α → α → Prop} (hl : l.nodup) (h : {x | x ∈ l}.pairwise r) : l.pairwise r := hl.pairwise_of_forall_ne h +@[simp] lemma nodup.pairwise_coe [is_symm α r] (hl : l.nodup) : + {a | a ∈ l}.pairwise r ↔ l.pairwise r := +begin + induction l with a l ih, + { simp }, + rw list.nodup_cons at hl, + have : ∀ b ∈ l, ¬a = b → r a b ↔ r a b := + λ b hb, imp_iff_right (ne_of_mem_of_not_mem hb hl.1).symm, + simp [set.set_of_or, set.pairwise_insert_of_symmetric (@symm_of _ r _), ih hl.2, and_comm, + forall₂_congr this], +end + end list theorem option.to_list_nodup {α} : ∀ o : option α, o.to_list.nodup diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 3ed0b04fc1417..0c7b7e4d871e2 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -55,7 +55,7 @@ variables {N : ι → Type*} [∀ i, monoid (N i)] variables (ϕ : Π (i : ι), N i →* M) -- We assume that the elements of different morphism commute -variables (hcomm : ∀ (i j : ι), i ≠ j → ∀ (x : N i) (y : N j), commute (ϕ i x) (ϕ j y)) +variables (hcomm : pairwise $ λ i j, ∀ x y, commute (ϕ i x) (ϕ j y)) include hcomm -- We use `f` and `g` to denote elements of `Π (i : ι), N i` @@ -68,8 +68,7 @@ namespace monoid_hom See also `linear_map.lsum` for a linear version without the commutativity assumption."] def noncomm_pi_coprod : (Π (i : ι), N i) →* M := -{ to_fun := λ f, finset.univ.noncomm_prod (λ i, ϕ i (f i)) $ - by { rintros i - j -, by_cases h : i = j, { subst h }, { exact hcomm _ _ h _ _ } }, +{ to_fun := λ f, finset.univ.noncomm_prod (λ i, ϕ i (f i)) $ λ i _ j _ h, hcomm _ _ h _ _, map_one' := by {apply (finset.noncomm_prod_eq_pow_card _ _ _ _ _).trans (one_pow _), simp}, map_mul' := λ f g, begin diff --git a/src/group_theory/perm/cycle/basic.lean b/src/group_theory/perm/cycle/basic.lean index ad15315bd79e6..f7e276194a9c0 100644 --- a/src/group_theory/perm/cycle/basic.lean +++ b/src/group_theory/perm/cycle/basic.lean @@ -929,46 +929,27 @@ end lemma cycle_factors_finset_eq_finset {σ : perm α} {s : finset (perm α)} : σ.cycle_factors_finset = s ↔ (∀ f : perm α, f ∈ s → f.is_cycle) ∧ - (∃ h : (∀ (a ∈ s) (b ∈ s), a ≠ b → disjoint a b), s.noncomm_prod id - (λ a ha b hb, (em (a = b)).by_cases (λ h, h ▸ commute.refl a) - (set.pairwise.mono' (λ _ _, disjoint.commute) h ha hb)) = σ) := + ∃ h : (s : set (perm α)).pairwise disjoint, + s.noncomm_prod id (h.mono' $ λ _ _, disjoint.commute) = σ := begin obtain ⟨l, hl, rfl⟩ := s.exists_list_nodup_eq, - rw cycle_factors_finset_eq_list_to_finset hl, - simp only [noncomm_prod_to_finset, hl, exists_prop, list.mem_to_finset, and.congr_left_iff, - and.congr_right_iff, list.map_id, ne.def], - intros, - exact ⟨list.pairwise.forall disjoint.symmetric, hl.pairwise_of_forall_ne⟩ + simp [cycle_factors_finset_eq_list_to_finset, hl], end -lemma cycle_factors_finset_pairwise_disjoint (p : perm α) (hp : p ∈ cycle_factors_finset f) - (q : perm α) (hq : q ∈ cycle_factors_finset f) (h : p ≠ q) : - disjoint p q := -begin - have : f.cycle_factors_finset = f.cycle_factors_finset := rfl, - obtain ⟨-, hd, -⟩ := cycle_factors_finset_eq_finset.mp this, - exact hd p hp q hq h -end +lemma cycle_factors_finset_pairwise_disjoint : + (cycle_factors_finset f : set (perm α)).pairwise disjoint := +(cycle_factors_finset_eq_finset.mp rfl).2.some -lemma cycle_factors_finset_mem_commute (p : perm α) (hp : p ∈ cycle_factors_finset f) - (q : perm α) (hq : q ∈ cycle_factors_finset f) : - _root_.commute p q := -begin - by_cases h : p = q, - { exact h ▸ commute.refl _ }, - { exact (cycle_factors_finset_pairwise_disjoint _ _ hp _ hq h).commute } -end +lemma cycle_factors_finset_mem_commute : + (cycle_factors_finset f : set (perm α)).pairwise commute := +(cycle_factors_finset_pairwise_disjoint _).mono' $ λ _ _, disjoint.commute /-- The product of cycle factors is equal to the original `f : perm α`. -/ lemma cycle_factors_finset_noncomm_prod - (comm : ∀ (g ∈ f.cycle_factors_finset) (h ∈ f.cycle_factors_finset), - commute (id g) (id h) := cycle_factors_finset_mem_commute f) : - f.cycle_factors_finset.noncomm_prod id (comm) = f := -begin - have : f.cycle_factors_finset = f.cycle_factors_finset := rfl, - obtain ⟨-, hd, hp⟩ := cycle_factors_finset_eq_finset.mp this, - exact hp -end + (comm : (cycle_factors_finset f : set (perm α)).pairwise commute := + cycle_factors_finset_mem_commute f) : + f.cycle_factors_finset.noncomm_prod id comm = f := +(cycle_factors_finset_eq_finset.mp rfl).2.some_spec lemma mem_cycle_factors_finset_iff {f p : perm α} : p ∈ cycle_factors_finset f ↔ p.is_cycle ∧ ∀ (a ∈ p.support), p a = f a := @@ -1056,21 +1037,14 @@ lemma disjoint.cycle_factors_finset_mul_eq_union {f g : perm α} (h : disjoint f cycle_factors_finset (f * g) = cycle_factors_finset f ∪ cycle_factors_finset g := begin rw cycle_factors_finset_eq_finset, - split, - { simp only [mem_cycle_factors_finset_iff, mem_union], - rintro _ (⟨h, -⟩ | ⟨h, -⟩); - exact h }, - { refine ⟨_, _⟩, - { simp_rw mem_union, - rintros x (hx | hx) y (hy | hy) hxy, - { exact cycle_factors_finset_pairwise_disjoint _ _ hx _ hy hxy }, - { exact h.mono (mem_cycle_factors_finset_support_le hx) - (mem_cycle_factors_finset_support_le hy) }, - { exact h.symm.mono (mem_cycle_factors_finset_support_le hx) - (mem_cycle_factors_finset_support_le hy) }, - { exact cycle_factors_finset_pairwise_disjoint _ _ hx _ hy hxy } }, - { rw noncomm_prod_union_of_disjoint h.disjoint_cycle_factors_finset, - rw [cycle_factors_finset_noncomm_prod, cycle_factors_finset_noncomm_prod] } } + refine ⟨_, _, _⟩, + { simp [or_imp_distrib, mem_cycle_factors_finset_iff, forall_swap] }, + { rw [coe_union, set.pairwise_union_of_symmetric disjoint.symmetric], + exact ⟨cycle_factors_finset_pairwise_disjoint _, cycle_factors_finset_pairwise_disjoint _, + λ x hx y hy hxy, h.mono (mem_cycle_factors_finset_support_le hx) + (mem_cycle_factors_finset_support_le hy)⟩ }, + { rw noncomm_prod_union_of_disjoint h.disjoint_cycle_factors_finset, + rw [cycle_factors_finset_noncomm_prod, cycle_factors_finset_noncomm_prod] } end lemma disjoint_mul_inv_of_mem_cycle_factors_finset {f g : perm α} (h : f ∈ cycle_factors_finset g) : diff --git a/src/group_theory/perm/cycle/type.lean b/src/group_theory/perm/cycle/type.lean index b00aac036a120..afc73b7b01826 100644 --- a/src/group_theory/perm/cycle/type.lean +++ b/src/group_theory/perm/cycle/type.lean @@ -47,10 +47,8 @@ lemma cycle_type_def (σ : perm α) : σ.cycle_type = σ.cycle_factors_finset.1.map (finset.card ∘ support) := rfl lemma cycle_type_eq' {σ : perm α} (s : finset (perm α)) - (h1 : ∀ f : perm α, f ∈ s → f.is_cycle) (h2 : ∀ (a ∈ s) (b ∈ s), a ≠ b → disjoint a b) - (h0 : s.noncomm_prod id - (λ a ha b hb, (em (a = b)).by_cases (λ h, h ▸ commute.refl a) - (set.pairwise.mono' (λ _ _, disjoint.commute) h2 ha hb)) = σ) : + (h1 : ∀ f : perm α, f ∈ s → f.is_cycle) (h2 : (s : set (perm α)).pairwise disjoint) + (h0 : s.noncomm_prod id (h2.imp $ λ _ _, disjoint.commute) = σ) : σ.cycle_type = s.1.map (finset.card ∘ support) := begin rw cycle_type_def, diff --git a/src/group_theory/perm/support.lean b/src/group_theory/perm/support.lean index 602a4d84feb29..df00a3cc77659 100644 --- a/src/group_theory/perm/support.lean +++ b/src/group_theory/perm/support.lean @@ -41,6 +41,8 @@ by simp only [disjoint, or.comm, imp_self] lemma disjoint.symmetric : symmetric (@disjoint α) := λ _ _, disjoint.symm +instance : is_symm (perm α) disjoint := ⟨disjoint.symmetric⟩ + lemma disjoint_comm : disjoint f g ↔ disjoint g f := ⟨disjoint.symm, disjoint.symm⟩ diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index d3974e05b204c..f3f9d7db5d446 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -477,8 +477,7 @@ protected lemma multiset_prod_mem {G} [comm_group G] (K : subgroup G) (g : multi (∀ a ∈ g, a ∈ K) → g.prod ∈ K := multiset_prod_mem g @[to_additive] -lemma multiset_noncomm_prod_mem (K : subgroup G) (g : multiset G) - (comm : ∀ (x ∈ g) (y ∈ g), commute x y) : +lemma multiset_noncomm_prod_mem (K : subgroup G) (g : multiset G) (comm) : (∀ a ∈ g, a ∈ K) → g.noncomm_prod comm ∈ K := K.to_submonoid.multiset_noncomm_prod_mem g comm @@ -492,8 +491,7 @@ protected lemma prod_mem {G : Type*} [comm_group G] (K : subgroup G) prod_mem h @[to_additive] -lemma noncomm_prod_mem (K : subgroup G) - {ι : Type*} {t : finset ι} {f : ι → G} (comm : ∀ (x ∈ t) (y ∈ t), commute (f x) (f y)) : +lemma noncomm_prod_mem (K : subgroup G) {ι : Type*} {t : finset ι} {f : ι → G} (comm) : (∀ c ∈ t, f c ∈ K) → t.noncomm_prod f comm ∈ K := K.to_submonoid.noncomm_prod_mem t f comm @@ -3127,8 +3125,7 @@ disjoint_def'.trans ⟨λ h x y hx hy hxy, generalizes (one direction of) `subgroup.disjoint_iff_mul_eq_one`. -/ @[to_additive "`finset.noncomm_sum` is “injective” in `f` if `f` maps into independent subgroups. This generalizes (one direction of) `add_subgroup.disjoint_iff_add_eq_zero`. "] -lemma eq_one_of_noncomm_prod_eq_one_of_independent {ι : Type*} - (s : finset ι) (f : ι → G) (comm : ∀ (x ∈ s) (y ∈ s), commute (f x) (f y)) +lemma eq_one_of_noncomm_prod_eq_one_of_independent {ι : Type*} (s : finset ι) (f : ι → G) (comm) (K : ι → subgroup G) (hind : complete_lattice.independent K) (hmem : ∀ (x ∈ s), f x ∈ K x) (heq1 : s.noncomm_prod f comm = 1) : ∀ (i ∈ s), f i = 1 := begin @@ -3136,9 +3133,9 @@ begin revert heq1, induction s using finset.induction_on with i s hnmem ih, { simp, }, - { simp only [finset.forall_mem_insert] at comm hmem, - specialize ih (λ x hx, (comm.2 x hx).2) hmem.2, - have hmem_bsupr: s.noncomm_prod f (λ x hx, (comm.2 x hx).2) ∈ ⨆ (i ∈ (s : set ι)), K i, + { have hcomm := comm.mono (finset.coe_subset.2 $ finset.subset_insert _ _), + simp only [finset.forall_mem_insert] at hmem, + have hmem_bsupr: s.noncomm_prod f hcomm ∈ ⨆ (i ∈ (s : set ι)), K i, { refine subgroup.noncomm_prod_mem _ _ _, intros x hx, have : K x ≤ ⨆ (i ∈ (s : set ι)), K i := le_supr₂ x hx, @@ -3148,12 +3145,11 @@ begin have hnmem' : i ∉ (s : set ι), by simpa, obtain ⟨heq1i : f i = 1, heq1S : s.noncomm_prod f _ = 1⟩ := subgroup.disjoint_iff_mul_eq_one.mp (hind.disjoint_bsupr hnmem') hmem.1 hmem_bsupr heq1, - specialize ih heq1S, intros i h, simp only [finset.mem_insert] at h, rcases h with ⟨rfl | _⟩, { exact heq1i }, - { exact (ih _ h), } } + { exact ih hcomm hmem.2 heq1S _ h } } end end subgroup diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index d70d06f8b1aa5..8a0e34c4a2119 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -109,8 +109,7 @@ lemma multiset_prod_mem {M} [comm_monoid M] (S : submonoid M) (m : multiset M) by { lift m to multiset S using hm, rw ← coe_multiset_prod, exact m.prod.coe_prop } @[to_additive] -lemma multiset_noncomm_prod_mem (S : submonoid M) (m : multiset M) - (comm : ∀ (x ∈ m) (y ∈ m), commute x y) (h : ∀ (x ∈ m), x ∈ S) : +lemma multiset_noncomm_prod_mem (S : submonoid M) (m : multiset M) (comm) (h : ∀ (x ∈ m), x ∈ S) : m.noncomm_prod comm ∈ S := begin induction m using quotient.induction_on with l, @@ -128,8 +127,8 @@ lemma prod_mem {M : Type*} [comm_monoid M] (S : submonoid M) S.multiset_prod_mem (t.1.map f) $ λ x hx, let ⟨i, hi, hix⟩ := multiset.mem_map.1 hx in hix ▸ h i hi @[to_additive] -lemma noncomm_prod_mem (S : submonoid M) {ι : Type*} (t : finset ι) (f : ι → M) - (comm : ∀ (x ∈ t) (y ∈ t), commute (f x) (f y)) (h : ∀ c ∈ t, f c ∈ S) : +lemma noncomm_prod_mem (S : submonoid M) {ι : Type*} (t : finset ι) (f : ι → M) (comm) + (h : ∀ c ∈ t, f c ∈ S) : t.noncomm_prod f comm ∈ S := begin apply multiset_noncomm_prod_mem, From 5d7ab6e9d99fe27a31fceb586ade608af44645a5 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Fri, 14 Oct 2022 23:58:27 +0000 Subject: [PATCH 771/828] feat(analysis/inner_product_space/two_dim): volume/area form of negated orientation (#16939) --- .../inner_product_space/orientation.lean | 60 +++++++++++++++++-- src/analysis/inner_product_space/two_dim.lean | 22 +++++++ 2 files changed, 76 insertions(+), 6 deletions(-) diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index b0daf2a16f765..b31f28dbcdd84 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -57,6 +57,17 @@ begin linarith, end +/-- The change-of-basis matrix between two orthonormal bases with the opposite orientations has +determinant -1. -/ +lemma det_to_matrix_orthonormal_basis_of_opposite_orientation + (h : e.to_basis.orientation ≠ f.to_basis.orientation) : + e.to_basis.det f = -1 := +begin + contrapose! h, + simp [e.to_basis.orientation_eq_iff_det_pos, + (e.det_to_matrix_orthonormal_basis_real f).resolve_right h], +end + variables {e f} /-- Two orthonormal bases with the same orientation determine the same "determinant" top-dimensional @@ -73,7 +84,17 @@ begin simp [e.det_to_matrix_orthonormal_basis_of_same_orientation f h], }, end -variables (e) +variables (e f) + +/-- Two orthonormal bases with opposite orientations determine opposite "determinant" +top-dimensional forms on `E`. -/ +lemma det_eq_neg_det_of_opposite_orientation + (h : e.to_basis.orientation ≠ f.to_basis.orientation) : + e.to_basis.det = -f.to_basis.det := +begin + rw e.to_basis.det.eq_smul_basis_det f.to_basis, + simp [e.det_to_matrix_orthonormal_basis_of_opposite_orientation f h], +end section adjust_to_orientation include ne @@ -161,8 +182,7 @@ begin classical, unfreezingI { cases n }, { let opos : alternating_map ℝ E ℝ (fin 0) := alternating_map.const_of_is_empty ℝ E (1:ℝ), - let oneg : alternating_map ℝ E ℝ (fin 0) := alternating_map.const_of_is_empty ℝ E (-1:ℝ), - exact o.eq_or_eq_neg_of_is_empty.by_cases (λ _, opos) (λ _, oneg) }, + exact o.eq_or_eq_neg_of_is_empty.by_cases (λ _, opos) (λ _, -opos) }, { exact (o.fin_orthonormal_basis n.succ_pos _i.out).to_basis.det } end @@ -171,11 +191,11 @@ omit _i o @[simp] lemma volume_form_zero_pos [_i : fact (finrank ℝ E = 0)] : orientation.volume_form (positive_orientation : orientation ℝ E (fin 0)) = alternating_map.const_linear_equiv_of_is_empty 1 := -by simp [volume_form, or.by_cases, dif_pos] +by simp [volume_form, or.by_cases, if_pos] -@[simp] lemma volume_form_zero_neg [_i : fact (finrank ℝ E = 0)] : +lemma volume_form_zero_neg [_i : fact (finrank ℝ E = 0)] : orientation.volume_form (-positive_orientation : orientation ℝ E (fin 0)) - = alternating_map.const_linear_equiv_of_is_empty (-1) := + = - alternating_map.const_linear_equiv_of_is_empty 1 := begin dsimp [volume_form, or.by_cases, positive_orientation], apply if_neg, @@ -200,6 +220,34 @@ begin exact o.fin_orthonormal_basis_orientation _ _ }, end +/-- The volume form on an oriented real inner product space can be evaluated as the determinant with +respect to any orthonormal basis of the space compatible with the orientation. -/ +lemma volume_form_robust_neg (b : orthonormal_basis (fin n) ℝ E) + (hb : b.to_basis.orientation ≠ o) : + o.volume_form = - b.to_basis.det := +begin + unfreezingI { cases n }, + { have : positive_orientation ≠ o := by rwa b.to_basis.orientation_is_empty at hb, + simp [volume_form, or.by_cases, dif_neg this.symm] }, + let e : orthonormal_basis (fin n.succ) ℝ E := o.fin_orthonormal_basis n.succ_pos (fact.out _), + dsimp [volume_form], + apply e.det_eq_neg_det_of_opposite_orientation b, + convert hb.symm, + exact o.fin_orthonormal_basis_orientation _ _, +end + +@[simp] lemma volume_form_neg_orientation : (-o).volume_form = - o.volume_form := +begin + unfreezingI { cases n }, + { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp [volume_form_zero_neg] }, + let e : orthonormal_basis (fin n.succ) ℝ E := o.fin_orthonormal_basis n.succ_pos (fact.out _), + have h₁ : e.to_basis.orientation = o := o.fin_orthonormal_basis_orientation _ _, + have h₂ : e.to_basis.orientation ≠ -o, + { symmetry, + rw [e.to_basis.orientation_ne_iff_eq_neg, h₁] }, + rw [o.volume_form_robust e h₁, (-o).volume_form_robust_neg e h₂], +end + lemma volume_form_robust' (b : orthonormal_basis (fin n) ℝ E) (v : fin n → E) : |o.volume_form v| = |b.to_basis.det v| := begin diff --git a/src/analysis/inner_product_space/two_dim.lean b/src/analysis/inner_product_space/two_dim.lean index ea78c1a6ff3cf..483f6ed609686 100644 --- a/src/analysis/inner_product_space/two_dim.lean +++ b/src/analysis/inner_product_space/two_dim.lean @@ -95,6 +95,12 @@ begin { norm_num } end +@[simp] lemma area_form_neg_orientation : (-o).area_form = -o.area_form := +begin + ext x y, + simp [area_form_to_volume_form] +end + /-- Continuous linear map version of `orientation.area_form`, useful for calculus. -/ def area_form' : E →L[ℝ] (E →L[ℝ] ℝ) := ((↑(linear_map.to_continuous_linear_map : (E →ₗ[ℝ] ℝ) ≃ₗ[ℝ] (E →L[ℝ] ℝ))) @@ -243,6 +249,19 @@ by simp linear_isometry_equiv.trans J J = linear_isometry_equiv.neg ℝ := by ext; simp +lemma right_angle_rotation_neg_orientation (x : E) : + (-o).right_angle_rotation x = - o.right_angle_rotation x := +begin + apply ext_inner_right ℝ, + intros y, + rw inner_right_angle_rotation_left, + simp +end + +@[simp] lemma right_angle_rotation_trans_neg_orientation : + (-o).right_angle_rotation = o.right_angle_rotation.trans (linear_isometry_equiv.neg ℝ) := +linear_isometry_equiv.ext $ o.right_angle_rotation_neg_orientation + /-- For a nonzero vector `x` in an oriented two-dimensional real inner product space `E`, `![x, J x]` forms an (orthogonal) basis for `E`. -/ def basis_right_angle_rotation (x : E) (hx : x ≠ 0) : basis (fin 2) ℝ E := @@ -391,6 +410,9 @@ begin linear_combination - o.kahler x y * complex.I_sq, end +@[simp] lemma kahler_neg_orientation (x y : E) : (-o).kahler x y = conj (o.kahler x y) := +by simp [kahler_apply_apply] + lemma kahler_mul (a x y : E) : o.kahler x a * o.kahler a y = ∥a∥ ^ 2 * o.kahler x y := begin transitivity (↑(∥a∥ ^ 2) : ℂ) * o.kahler x y, From 372310d0bfda6bf7e3406fe1bf57d154ca3191a0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 15 Oct 2022 03:47:53 +0000 Subject: [PATCH 772/828] chore(algebra/group/*): attribute copyright properly, correct instance names (#16967) #16847 copied over the copyright from `algebra.group.basic` without actually attributing things properly. All the material in `algebra.group.commutator` is from #12309 and the material is mostly from #16122 with some dating back from #8564 Also fix `order_dual.has_pow` being additivized as `order_dual.has_nsmul` rather than `order_dual.has_smul`. --- src/algebra/group/commutator.lean | 4 ++-- src/algebra/group/order_synonym.lean | 21 +++++++++++++-------- 2 files changed, 15 insertions(+), 10 deletions(-) diff --git a/src/algebra/group/commutator.lean b/src/algebra/group/commutator.lean index 2b25672d50d3e..dcec0c29ecffd 100644 --- a/src/algebra/group/commutator.lean +++ b/src/algebra/group/commutator.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Copyright (c) 2022 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro +Authors: Thomas Browning -/ import algebra.group.defs diff --git a/src/algebra/group/order_synonym.lean b/src/algebra/group/order_synonym.lean index 5020ec4fc9ed1..9f2400047cde8 100644 --- a/src/algebra/group/order_synonym.lean +++ b/src/algebra/group/order_synonym.lean @@ -1,26 +1,31 @@ /- -Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro +Authors: Yury Kudryashov, Yaël Dillies -/ import algebra.group.defs import order.synonym -/-! # Group structure on the type synonyms -/ +/-! +# Group structure on the order type synonyms -variables {α β : Type*} - -/-! ### order_dual -/ +Transfer algebraic instances from `α` to `αᵒᵈ` and `lex α`. +-/ open order_dual +variables {α β : Type*} + +/-! ### `order_dual` -/ + @[to_additive] instance [h : has_one α] : has_one αᵒᵈ := h @[to_additive] instance [h : has_mul α] : has_mul αᵒᵈ := h @[to_additive] instance [h : has_inv α] : has_inv αᵒᵈ := h @[to_additive] instance [h : has_div α] : has_div αᵒᵈ := h @[to_additive] instance [h : has_smul α β] : has_smul α βᵒᵈ := h -@[to_additive] instance order_dual.has_pow [h : has_pow α β] : has_pow αᵒᵈ β := h +@[to_additive order_dual.has_smul] +instance order_dual.has_pow [h : has_pow α β] : has_pow αᵒᵈ β := h @[to_additive] instance [h : semigroup α] : semigroup αᵒᵈ := h @[to_additive] instance [h : comm_semigroup α] : comm_semigroup αᵒᵈ := h @[to_additive] instance [h : left_cancel_semigroup α] : left_cancel_semigroup αᵒᵈ := h @@ -71,7 +76,7 @@ lemma of_dual_pow [has_pow α β] (a : αᵒᵈ) (b : β) : of_dual (a ^ b) = of @[to_additive] instance [h : has_inv α] : has_inv (lex α) := h @[to_additive] instance [h : has_div α] : has_div (lex α) := h @[to_additive] instance [h : has_smul α β] : has_smul α (lex β) := h -@[to_additive] instance lex.has_pow [h : has_pow α β] : has_pow (lex α) β := h +@[to_additive lex.has_smul] instance lex.has_pow [h : has_pow α β] : has_pow (lex α) β := h @[to_additive] instance [h : semigroup α] : semigroup (lex α) := h @[to_additive] instance [h : comm_semigroup α] : comm_semigroup (lex α) := h @[to_additive] instance [h : left_cancel_semigroup α] : left_cancel_semigroup (lex α) := h From 5316314b553dcf8c6716541851517c1a9715e22b Mon Sep 17 00:00:00 2001 From: Wrenna Robson Date: Sat, 15 Oct 2022 07:59:10 +0000 Subject: [PATCH 773/828] feat(topology/metric_space/infsep): Add "infimum separation". (#15689) Informally, this is the "minimum distance", though "minimum" makes sense mainly only in the finite context. This is analogous to diam in some sense. We provide finset and set versions for both the emetric and metric cases. We generally try to have the analogous lemmas that diam has, and any other useful utility lemmas to allow for the use of the definition without unpacking it. --- src/topology/metric_space/infsep.lean | 478 ++++++++++++++++++++++++++ 1 file changed, 478 insertions(+) create mode 100644 src/topology/metric_space/infsep.lean diff --git a/src/topology/metric_space/infsep.lean b/src/topology/metric_space/infsep.lean new file mode 100644 index 0000000000000..3de55c17a31e9 --- /dev/null +++ b/src/topology/metric_space/infsep.lean @@ -0,0 +1,478 @@ +/- +Copyright (c) 2022 Wrenna Robson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Wrenna Robson +-/ +import topology.metric_space.basic + +/-! +# Infimum separation + +This file defines the extended infimum separation of a set. This is approximately dual to the +diameter of a set, but where the extended diameter of a set is the supremum of the extended distance +between elements of the set, the extended infimum separation is the infimum of the (extended) +distance between *distinct* elements in the set. + +We also define the infimum separation as the cast of the extended infimum separation to the reals. +This is the infimum of the distance between distinct elements of the set when in a pseudometric +space. + +All lemmas and definitions are in the `set` namespace to give access to dot notation. + +## Main definitions +* `set.einfsep`: Extended infimum separation of a set. +* `set.infsep`: Infimum separation of a set (when in a pseudometric space). + +!-/ + +variables {α β : Type*} + +namespace set + +section einfsep +open_locale ennreal +open function + +/-- The "extended infimum separation" of a set with an edist function. -/ +noncomputable def einfsep [has_edist α] (s : set α) : ℝ≥0∞ := +⨅ (x ∈ s) (y ∈ s) (hxy : x ≠ y), edist x y + +section has_edist +variables [has_edist α] {x y : α} {s t : set α} + +lemma le_einfsep_iff {d} : d ≤ s.einfsep ↔ ∀ (x y ∈ s) (hxy : x ≠ y), d ≤ edist x y := +by simp_rw [einfsep, le_infi_iff] + +theorem einfsep_zero : + s.einfsep = 0 ↔ ∀ C (hC : 0 < C), ∃ (x y ∈ s) (hxy : x ≠ y), edist x y < C := +by simp_rw [einfsep, ← bot_eq_zero, infi_eq_bot, infi_lt_iff] + +theorem einfsep_pos : + 0 < s.einfsep ↔ ∃ C (hC : 0 < C), ∀ (x y ∈ s) (hxy : x ≠ y), C ≤ edist x y := +by { rw [pos_iff_ne_zero, ne.def, einfsep_zero], simp only [not_forall, not_exists, not_lt] } + +lemma einfsep_top : s.einfsep = ∞ ↔ ∀ (x y ∈ s) (hxy : x ≠ y), edist x y = ∞ := +by simp_rw [einfsep, infi_eq_top] + +lemma einfsep_lt_top : s.einfsep < ∞ ↔ ∃ (x y ∈ s) (hxy : x ≠ y), edist x y < ∞ := +by simp_rw [einfsep, infi_lt_iff] + +lemma einfsep_ne_top : s.einfsep ≠ ∞ ↔ ∃ (x y ∈ s) (hxy : x ≠ y), edist x y ≠ ∞ := +by simp_rw [←lt_top_iff_ne_top, einfsep_lt_top] + +lemma einfsep_lt_iff {d} : s.einfsep < d ↔ ∃ (x y ∈ s) (h : x ≠ y), edist x y < d := +by simp_rw [einfsep, infi_lt_iff] + +lemma nontrivial_of_einfsep_lt_top (hs : s.einfsep < ∞) : s.nontrivial := +by { rcases einfsep_lt_top.1 hs with ⟨_, hx, _, hy, hxy, _⟩, exact ⟨_, hx, _, hy, hxy⟩ } + +lemma nontrivial_of_einfsep_ne_top (hs : s.einfsep ≠ ∞) : s.nontrivial := +nontrivial_of_einfsep_lt_top (lt_top_iff_ne_top.mpr hs) + +lemma subsingleton.einfsep (hs : s.subsingleton) : s.einfsep = ∞ := +by { rw einfsep_top, exact λ _ hx _ hy hxy, (hxy $ hs hx hy).elim } + +lemma le_einfsep_image_iff {d} {f : β → α} {s : set β} : + d ≤ einfsep (f '' s) ↔ ∀ x y ∈ s, f x ≠ f y → d ≤ edist (f x) (f y) := +by simp_rw [le_einfsep_iff, ball_image_iff] + +lemma le_edist_of_le_einfsep {d x} (hx : x ∈ s) {y} (hy : y ∈ s) (hxy : x ≠ y) + (hd : d ≤ s.einfsep) : d ≤ edist x y := le_einfsep_iff.1 hd x hx y hy hxy + +lemma einfsep_le_edist_of_mem {x} (hx : x ∈ s) {y} (hy : y ∈ s) (hxy : x ≠ y) : + s.einfsep ≤ edist x y := le_edist_of_le_einfsep hx hy hxy le_rfl + +lemma einfsep_le_of_mem_of_edist_le {d x} (hx : x ∈ s) {y} (hy : y ∈ s) (hxy : x ≠ y) + (hxy' : edist x y ≤ d) : s.einfsep ≤ d := le_trans (einfsep_le_edist_of_mem hx hy hxy) hxy' + +lemma le_einfsep {d} (h : ∀ (x y ∈ s) (hxy : x ≠ y), d ≤ edist x y) : + d ≤ s.einfsep := le_einfsep_iff.2 h + +@[simp] lemma einfsep_empty : (∅ : set α).einfsep = ∞ := subsingleton_empty.einfsep + +@[simp] lemma einfsep_singleton : ({x} : set α).einfsep = ∞ := subsingleton_singleton.einfsep + +lemma einfsep_Union_mem_option {ι : Type*} (o : option ι) (s : ι → set α) : + (⋃ i ∈ o, s i).einfsep = ⨅ i ∈ o, (s i).einfsep := by cases o; simp + +lemma einfsep_anti (hst : s ⊆ t) : t.einfsep ≤ s.einfsep := +le_einfsep $ λ x hx y hy, einfsep_le_edist_of_mem (hst hx) (hst hy) + +lemma einfsep_insert_le : (insert x s).einfsep ≤ ⨅ (y ∈ s) (hxy : x ≠ y), edist x y := +begin + simp_rw le_infi_iff, + refine λ _ hy hxy, einfsep_le_edist_of_mem (mem_insert _ _) (mem_insert_of_mem _ hy) hxy +end + +lemma le_einfsep_pair : edist x y ⊓ edist y x ≤ ({x, y} : set α).einfsep := +begin + simp_rw [le_einfsep_iff, inf_le_iff, mem_insert_iff, mem_singleton_iff], + rintros a (rfl | rfl) b (rfl | rfl) hab; finish +end + +lemma einfsep_pair_le_left (hxy : x ≠ y) : ({x, y} : set α).einfsep ≤ edist x y := +einfsep_le_edist_of_mem (mem_insert _ _) (mem_insert_of_mem _ (mem_singleton _)) hxy + +lemma einfsep_pair_le_right (hxy : x ≠ y) : ({x, y} : set α).einfsep ≤ edist y x := +by rw pair_comm; exact einfsep_pair_le_left hxy.symm + +lemma einfsep_pair_eq_inf (hxy : x ≠ y) : ({x, y} : set α).einfsep = (edist x y) ⊓ (edist y x) := +le_antisymm (le_inf (einfsep_pair_le_left hxy) (einfsep_pair_le_right hxy)) le_einfsep_pair + +lemma einfsep_eq_infi : s.einfsep = ⨅ d : s.off_diag, (uncurry edist) (d : α × α) := +begin + refine eq_of_forall_le_iff (λ _, _), + simp_rw [le_einfsep_iff, le_infi_iff, imp_forall_iff, set_coe.forall, subtype.coe_mk, + mem_off_diag, prod.forall, uncurry_apply_pair, and_imp] +end + +lemma einfsep_of_fintype [decidable_eq α] [fintype s] : + s.einfsep = s.off_diag.to_finset.inf (uncurry edist) := +begin + refine eq_of_forall_le_iff (λ _, _), + simp_rw [le_einfsep_iff, imp_forall_iff, finset.le_inf_iff, mem_to_finset, mem_off_diag, + prod.forall, uncurry_apply_pair, and_imp] +end + +lemma finite.einfsep (hs : s.finite) : + s.einfsep = hs.off_diag.to_finset.inf (uncurry edist) := +begin + refine eq_of_forall_le_iff (λ _, _), + simp_rw [le_einfsep_iff, imp_forall_iff, finset.le_inf_iff, finite.mem_to_finset, mem_off_diag, + prod.forall, uncurry_apply_pair, and_imp] +end + +lemma finset.coe_einfsep [decidable_eq α] {s : finset α} : + (s : set α).einfsep = s.off_diag.inf (uncurry edist) := +by simp_rw [einfsep_of_fintype, ← finset.coe_off_diag, finset.to_finset_coe] + +lemma nontrivial.einfsep_exists_of_finite [finite s] (hs : s.nontrivial) : + ∃ (x y ∈ s) (hxy : x ≠ y), s.einfsep = edist x y := +begin + classical, + casesI nonempty_fintype s, + simp_rw einfsep_of_fintype, + rcases @finset.exists_mem_eq_inf _ _ _ _ (s.off_diag.to_finset) (by simpa) (uncurry edist) + with ⟨_, hxy, hed⟩, + simp_rw mem_to_finset at hxy, + refine ⟨w.fst, hxy.1, w.snd, hxy.2.1, hxy.2.2, hed⟩ +end + +lemma finite.einfsep_exists_of_nontrivial (hsf : s.finite) (hs : s.nontrivial) : + ∃ (x y ∈ s) (hxy : x ≠ y), s.einfsep = edist x y := +by { letI := hsf.fintype, exact hs.einfsep_exists_of_finite } + +end has_edist + +section pseudo_emetric_space +variables [pseudo_emetric_space α] {x y z : α} {s t : set α} + +lemma einfsep_pair (hxy : x ≠ y) : ({x, y} : set α).einfsep = edist x y := +begin + nth_rewrite 0 [← min_self (edist x y)], + convert einfsep_pair_eq_inf hxy using 2, + rw edist_comm +end + +lemma einfsep_insert : + einfsep (insert x s) = (⨅ (y ∈ s) (hxy : x ≠ y), edist x y) ⊓ (s.einfsep) := +begin + refine le_antisymm (le_min einfsep_insert_le (einfsep_anti (subset_insert _ _))) _, + simp_rw [le_einfsep_iff, inf_le_iff, mem_insert_iff], + rintros y (rfl | hy) z (rfl | hz) hyz, + { exact false.elim (hyz rfl) }, + { exact or.inl (infi_le_of_le _ (infi₂_le hz hyz)) }, + { rw edist_comm, exact or.inl (infi_le_of_le _ (infi₂_le hy hyz.symm)) }, + { exact or.inr (einfsep_le_edist_of_mem hy hz hyz) } +end + +lemma einfsep_triple (hxy : x ≠ y) (hyz : y ≠ z) (hxz : x ≠ z) : + einfsep ({x, y, z} : set α) = edist x y ⊓ edist x z ⊓ edist y z := +by simp_rw [einfsep_insert, infi_insert, infi_singleton, einfsep_singleton, + inf_top_eq, cinfi_pos hxy, cinfi_pos hyz, cinfi_pos hxz] + +lemma le_einfsep_pi_of_le {π : β → Type*} [fintype β] [∀ b, pseudo_emetric_space (π b)] + {s : Π (b : β), set (π b)} {c : ℝ≥0∞} (h : ∀ b, c ≤ einfsep (s b) ) : + c ≤ einfsep (set.pi univ s) := +begin + refine le_einfsep (λ x hx y hy hxy, _), + rw mem_univ_pi at hx hy, + rcases function.ne_iff.mp hxy with ⟨i, hi⟩, + exact le_trans (le_einfsep_iff.1 (h i) _ (hx _) _ (hy _) hi) (edist_le_pi_edist _ _ i) +end + +end pseudo_emetric_space + +section pseudo_metric_space +variables [pseudo_metric_space α] {s : set α} + +theorem subsingleton_of_einfsep_eq_top (hs : s.einfsep = ∞) : s.subsingleton := +begin + rw einfsep_top at hs, + exact λ _ hx _ hy, of_not_not (λ hxy, edist_ne_top _ _ (hs _ hx _ hy hxy)) +end + +theorem einfsep_eq_top_iff : s.einfsep = ∞ ↔ s.subsingleton := +⟨subsingleton_of_einfsep_eq_top, subsingleton.einfsep⟩ + +theorem nontrivial.einfsep_ne_top (hs : s.nontrivial) : s.einfsep ≠ ∞ := +by { contrapose! hs, rw not_nontrivial_iff, exact subsingleton_of_einfsep_eq_top hs } + +theorem nontrivial.einfsep_lt_top (hs : s.nontrivial) : s.einfsep < ∞ := +by { rw lt_top_iff_ne_top, exact hs.einfsep_ne_top } + +theorem einfsep_lt_top_iff : s.einfsep < ∞ ↔ s.nontrivial := +⟨nontrivial_of_einfsep_lt_top, nontrivial.einfsep_lt_top⟩ + +theorem einfsep_ne_top_iff : s.einfsep ≠ ∞ ↔ s.nontrivial := +⟨nontrivial_of_einfsep_ne_top, nontrivial.einfsep_ne_top⟩ + +lemma le_einfsep_of_forall_dist_le {d} (h : ∀ (x y ∈ s) (hxy : x ≠ y), d ≤ dist x y) : + ennreal.of_real d ≤ s.einfsep := +le_einfsep $ +λ x hx y hy hxy, (edist_dist x y).symm ▸ ennreal.of_real_le_of_real (h x hx y hy hxy) + +end pseudo_metric_space + +section emetric_space +variables [emetric_space α] {x y z : α} {s t : set α} {C : ℝ≥0∞} {sC : set ℝ≥0∞} + +lemma einfsep_pos_of_finite [finite s] : 0 < s.einfsep := +begin + casesI nonempty_fintype s, + by_cases hs : s.nontrivial, + { rcases hs.einfsep_exists_of_finite with ⟨x, hx, y, hy, hxy, hxy'⟩, + exact hxy'.symm ▸ edist_pos.2 hxy }, + { rw not_nontrivial_iff at hs, + exact hs.einfsep.symm ▸ with_top.zero_lt_top } +end + +lemma relatively_discrete_of_finite [finite s] : + ∃ C (hC : 0 < C), ∀ (x y ∈ s) (hxy : x ≠ y), C ≤ edist x y := +by { rw ← einfsep_pos, exact einfsep_pos_of_finite } + +lemma finite.einfsep_pos (hs : s.finite) : 0 < s.einfsep := +by { letI := hs.fintype, exact einfsep_pos_of_finite } + +lemma finite.relatively_discrete (hs : s.finite) : + ∃ C (hC : 0 < C), ∀ (x y ∈ s) (hxy : x ≠ y), C ≤ edist x y := +by { letI := hs.fintype, exact relatively_discrete_of_finite } + +end emetric_space + +end einfsep + +section infsep +open_locale ennreal +open set function + +/-- The "infimum separation" of a set with an edist function. -/ +noncomputable def infsep [has_edist α] (s : set α) : ℝ := ennreal.to_real (s.einfsep) + +section has_edist +variables [has_edist α] {x y : α} {s : set α} + +lemma infsep_zero : s.infsep = 0 ↔ s.einfsep = 0 ∨ s.einfsep = ∞ := +by rw [infsep, ennreal.to_real_eq_zero_iff] + +lemma infsep_nonneg : 0 ≤ s.infsep := ennreal.to_real_nonneg + +lemma infsep_pos : 0 < s.infsep ↔ 0 < s.einfsep ∧ s.einfsep < ∞ := +by simp_rw [infsep, ennreal.to_real_pos_iff] + +lemma subsingleton.infsep_zero (hs : s.subsingleton) : s.infsep = 0 := +by { rw [infsep_zero, hs.einfsep], right, refl } + +lemma nontrivial_of_infsep_pos (hs : 0 < s.infsep) : s.nontrivial := +by { contrapose hs, rw not_nontrivial_iff at hs, exact hs.infsep_zero ▸ lt_irrefl _ } + +lemma infsep_empty : (∅ : set α).infsep = 0 := +subsingleton_empty.infsep_zero + +lemma infsep_singleton : ({x} : set α).infsep = 0 := +subsingleton_singleton.infsep_zero + +lemma infsep_pair_le_to_real_inf (hxy : x ≠ y) : + ({x, y} : set α).infsep ≤ (edist x y ⊓ edist y x).to_real := +by simp_rw [infsep, einfsep_pair_eq_inf hxy] + +end has_edist + +section pseudo_emetric_space +variables [pseudo_emetric_space α] {x y : α} {s : set α} + +lemma infsep_pair_eq_to_real : ({x, y} : set α).infsep = (edist x y).to_real := +begin + by_cases hxy : x = y, + { rw hxy, simp only [infsep_singleton, pair_eq_singleton, edist_self, ennreal.zero_to_real] }, + { rw [infsep, einfsep_pair hxy] } +end + +end pseudo_emetric_space + +section pseudo_metric_space + +variables [pseudo_metric_space α] {x y z: α} {s t : set α} + +lemma nontrivial.le_infsep_iff {d} (hs : s.nontrivial) : + d ≤ s.infsep ↔ ∀ (x y ∈ s) (hxy : x ≠ y), d ≤ dist x y := +by simp_rw [infsep, ← ennreal.of_real_le_iff_le_to_real (hs.einfsep_ne_top), le_einfsep_iff, + edist_dist, ennreal.of_real_le_of_real_iff (dist_nonneg)] + +lemma nontrivial.infsep_lt_iff {d} (hs : s.nontrivial) : + s.infsep < d ↔ ∃ (x y ∈ s) (hxy : x ≠ y), dist x y < d := +by { rw ← not_iff_not, push_neg, exact hs.le_infsep_iff } + +lemma nontrivial.le_infsep {d} (hs : s.nontrivial) (h : ∀ (x y ∈ s) (hxy : x ≠ y), d ≤ dist x y) : + d ≤ s.infsep := hs.le_infsep_iff.2 h + +lemma le_edist_of_le_infsep {d x} (hx : x ∈ s) {y} (hy : y ∈ s) + (hxy : x ≠ y) (hd : d ≤ s.infsep) : d ≤ dist x y := +begin + by_cases hs : s.nontrivial, + { exact hs.le_infsep_iff.1 hd x hx y hy hxy }, + { rw not_nontrivial_iff at hs, + rw hs.infsep_zero at hd, + exact le_trans hd dist_nonneg } +end + +lemma infsep_le_dist_of_mem (hx : x ∈ s) (hy : y ∈ s) (hxy : x ≠ y) : s.infsep ≤ dist x y := +le_edist_of_le_infsep hx hy hxy le_rfl + +lemma infsep_le_of_mem_of_edist_le {d x} (hx : x ∈ s) {y} (hy : y ∈ s) (hxy : x ≠ y) + (hxy' : dist x y ≤ d) : s.infsep ≤ d := le_trans (infsep_le_dist_of_mem hx hy hxy) hxy' + +lemma infsep_pair : ({x, y} : set α).infsep = dist x y := +by { rw [infsep_pair_eq_to_real, edist_dist], exact ennreal.to_real_of_real (dist_nonneg) } + +lemma infsep_triple (hxy : x ≠ y) (hyz : y ≠ z) (hxz : x ≠ z) : + ({x, y, z} : set α).infsep = dist x y ⊓ dist x z ⊓ dist y z := +by simp only [infsep, einfsep_triple hxy hyz hxz, ennreal.to_real_inf, edist_ne_top x y, + edist_ne_top x z, edist_ne_top y z, dist_edist, ne.def, inf_eq_top_iff, + and_self, not_false_iff] + + +lemma nontrivial.infsep_anti (hs : s.nontrivial) (hst : s ⊆ t) : t.infsep ≤ s.infsep := +ennreal.to_real_mono hs.einfsep_ne_top (einfsep_anti hst) + +lemma infsep_eq_infi [decidable s.nontrivial] : + s.infsep = if s.nontrivial then ⨅ d : s.off_diag, (uncurry dist) (d : α × α) else 0 := +begin + split_ifs with hs, + { have hb : bdd_below (uncurry dist '' s.off_diag), + { refine ⟨0, λ d h, _⟩, + simp_rw [mem_image, prod.exists, uncurry_apply_pair] at h, + rcases h with ⟨_, _, _, rfl⟩, + exact dist_nonneg }, + refine eq_of_forall_le_iff (λ _, _), + simp_rw [hs.le_infsep_iff, le_cinfi_set_iff (off_diag_nonempty.mpr hs) hb, imp_forall_iff, + mem_off_diag, prod.forall, uncurry_apply_pair, and_imp] }, + { exact ((not_nontrivial_iff).mp hs).infsep_zero } +end + +lemma nontrivial.infsep_eq_infi (hs : s.nontrivial) + : s.infsep = ⨅ d : s.off_diag, (uncurry dist) (d : α × α) := +by { classical, rw [infsep_eq_infi, if_pos hs] } + +lemma infsep_of_fintype [decidable s.nontrivial] [decidable_eq α] [fintype s] : + s.infsep = if hs : s.nontrivial then s.off_diag.to_finset.inf' (by simpa) (uncurry dist) else 0 := +begin + split_ifs with hs, + { refine eq_of_forall_le_iff (λ _, _), + simp_rw [hs.le_infsep_iff, imp_forall_iff, finset.le_inf'_iff, mem_to_finset, mem_off_diag, + prod.forall, uncurry_apply_pair, and_imp] }, + { rw not_nontrivial_iff at hs, exact hs.infsep_zero } +end + +lemma nontrivial.infsep_of_fintype [decidable_eq α] [fintype s] (hs : s.nontrivial) : + s.infsep = s.off_diag.to_finset.inf' (by simpa) (uncurry dist) := +by { classical, rw [infsep_of_fintype, dif_pos hs] } + +lemma finite.infsep [decidable s.nontrivial] (hsf : s.finite) : + s.infsep = if hs : s.nontrivial then hsf.off_diag.to_finset.inf' (by simpa) (uncurry dist) + else 0 := +begin + split_ifs with hs, + { refine eq_of_forall_le_iff (λ _, _), + simp_rw [hs.le_infsep_iff, imp_forall_iff, finset.le_inf'_iff, finite.mem_to_finset, + mem_off_diag, prod.forall, uncurry_apply_pair, and_imp] }, + { rw not_nontrivial_iff at hs, exact hs.infsep_zero } +end + +lemma finite.infsep_of_nontrivial (hsf : s.finite) (hs : s.nontrivial) : + s.infsep = hsf.off_diag.to_finset.inf' (by simpa) (uncurry dist) := + by { classical, simp_rw [hsf.infsep, dif_pos hs] } + +lemma _root_.finset.coe_infsep [decidable_eq α] (s : finset α) : + (s : set α).infsep = if hs : s.off_diag.nonempty then s.off_diag.inf' hs (uncurry dist) + else 0 := +begin + have H : (s : set α).nontrivial ↔ s.off_diag.nonempty, + by rwa [← set.off_diag_nonempty, ← finset.coe_off_diag, finset.coe_nonempty], + split_ifs with hs, + { simp_rw [(H.mpr hs).infsep_of_fintype, ← finset.coe_off_diag, finset.to_finset_coe] }, + { exact ((not_nontrivial_iff).mp (H.mp.mt hs)).infsep_zero } +end + +lemma _root_.finset.coe_infsep_of_off_diag_nonempty [decidable_eq α] {s : finset α} + (hs : s.off_diag.nonempty) : (s : set α).infsep = s.off_diag.inf' hs (uncurry dist) := +by rw [finset.coe_infsep, dif_pos hs] + +lemma _root_.finset.coe_infsep_of_off_diag_empty [decidable_eq α] {s : finset α} + (hs : s.off_diag = ∅) : (s : set α).infsep = 0 := +by { rw ← finset.not_nonempty_iff_eq_empty at hs, rw [finset.coe_infsep, dif_neg hs] } + +lemma nontrivial.infsep_exists_of_finite [finite s] (hs : s.nontrivial) : + ∃ (x y ∈ s) (hxy : x ≠ y), s.infsep = dist x y := +begin + classical, + casesI nonempty_fintype s, + simp_rw hs.infsep_of_fintype, + rcases @finset.exists_mem_eq_inf' _ _ _ (s.off_diag.to_finset) (by simpa) (uncurry dist) + with ⟨_, hxy, hed⟩, + simp_rw mem_to_finset at hxy, + exact ⟨w.fst, hxy.1, w.snd, hxy.2.1, hxy.2.2, hed⟩ +end + +lemma finite.infsep_exists_of_nontrivial (hsf : s.finite) (hs : s.nontrivial) : + ∃ (x y ∈ s) (hxy : x ≠ y), s.infsep = dist x y := +by { letI := hsf.fintype, exact hs.infsep_exists_of_finite } + +end pseudo_metric_space + +section metric_space +variables [metric_space α] {s : set α} + +lemma infsep_zero_iff_subsingleton_of_finite [finite s] : + s.infsep = 0 ↔ s.subsingleton := +begin + rw [infsep_zero, einfsep_eq_top_iff, or_iff_right_iff_imp], + exact λ H, (einfsep_pos_of_finite.ne' H).elim +end + +lemma infsep_pos_iff_nontrivial_of_finite [finite s] : + 0 < s.infsep ↔ s.nontrivial := +begin + rw [infsep_pos, einfsep_lt_top_iff, and_iff_right_iff_imp], + exact λ _, einfsep_pos_of_finite +end + +lemma finite.infsep_zero_iff_subsingleton (hs : s.finite) : + s.infsep = 0 ↔ s.subsingleton := +by { letI := hs.fintype, exact infsep_zero_iff_subsingleton_of_finite } + +lemma finite.infsep_pos_iff_nontrivial (hs : s.finite) : + 0 < s.infsep ↔ s.nontrivial := +by { letI := hs.fintype, exact infsep_pos_iff_nontrivial_of_finite } + +lemma _root_.finset.infsep_zero_iff_subsingleton (s : finset α) : + (s : set α).infsep = 0 ↔ (s : set α).subsingleton := infsep_zero_iff_subsingleton_of_finite + +lemma _root_.finset.infsep_pos_iff_nontrivial (s : finset α) : + 0 < (s : set α).infsep ↔ (s : set α).nontrivial := infsep_pos_iff_nontrivial_of_finite + +end metric_space + +end infsep + +end set From 73c9121e11837237af7b28fde9dd7ed53b32a676 Mon Sep 17 00:00:00 2001 From: mkaratarakis Date: Sat, 15 Oct 2022 07:59:11 +0000 Subject: [PATCH 774/828] feat(algebra/ring/aut): add apply_mul_semiring_action & mul_semiring_action.to_ring_aut (#16857) Co-authored-by: mkaratarakis <40603357+mkaratarakis@users.noreply.github.com> --- src/algebra/ring/aut.lean | 38 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 38 insertions(+) diff --git a/src/algebra/ring/aut.lean b/src/algebra/ring/aut.lean index 642a865bbe7c9..834b9092e6b75 100644 --- a/src/algebra/ring/aut.lean +++ b/src/algebra/ring/aut.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ +import algebra.group_ring_action import algebra.hom.aut import algebra.ring.equiv @@ -29,6 +30,9 @@ ring_aut @[reducible] def ring_aut (R : Type*) [has_mul R] [has_add R] := ring_equiv R R namespace ring_aut + +section mul_add + variables (R : Type*) [has_mul R] [has_add R] /-- @@ -60,4 +64,38 @@ by refine_struct { to_fun := ring_equiv.to_mul_equiv }; intros; refl def to_perm : ring_aut R →* equiv.perm R := by refine_struct { to_fun := ring_equiv.to_equiv }; intros; refl +end mul_add + +section semiring + +variables {G R : Type*} [group G] [semiring R] + +/-- The tautological action by the group of automorphism of a ring `R` on `R`.-/ +instance apply_mul_semiring_action : mul_semiring_action (ring_aut R) R := +{ smul := ($), + smul_zero := ring_equiv.map_zero, + smul_add := ring_equiv.map_add, + smul_one := ring_equiv.map_one, + smul_mul := ring_equiv.map_mul, + one_smul := λ _, rfl, + mul_smul := λ _ _ _, rfl } + +@[simp] +protected lemma smul_def (f : ring_aut R) (r : R) : f • r = f r := rfl + +instance apply_has_faithful_smul : has_faithful_smul (ring_aut R) R := ⟨λ _ _, ring_equiv.ext⟩ + +variables (G R) + +/-- Each element of the group defines a ring automorphism. + +This is a stronger version of `distrib_mul_action.to_add_aut` and +`mul_distrib_mul_action.to_mul_aut`. -/ +@[simps] def _root_.mul_semiring_action.to_ring_aut [mul_semiring_action G R] : G →* ring_aut R := +{ to_fun := mul_semiring_action.to_ring_equiv G R, + map_mul' := λ g h, ring_equiv.ext $ mul_smul g h, + map_one' := ring_equiv.ext $ one_smul _, } + +end semiring + end ring_aut From 0a375900aa08bede6ff0e20e927e4941fd93a789 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 15 Oct 2022 07:59:12 +0000 Subject: [PATCH 775/828] feat(algebra/star/subalgebra): develop some API for `star_subalgebra` (#16896) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Some of this is necessary to define the (closed) C⋆-algebra generated by a single element, which is in turn necessary to define the continuous functional calculus. --- src/algebra/star/basic.lean | 17 ++++ src/algebra/star/subalgebra.lean | 167 ++++++++++++++++++++++++++++++- 2 files changed, 180 insertions(+), 4 deletions(-) diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index c3550fa6430b9..b61603e6a7106 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -7,6 +7,7 @@ import algebra.ring.aut import algebra.ring.comp_typeclasses import data.rat.cast import group_theory.group_action.opposite +import data.set_like.basic /-! # Star monoids, rings, and modules @@ -57,6 +58,22 @@ A star operation (e.g. complex conjugate). -/ add_decl_doc star +/-- `star_mem_class S G` states `S` is a type of subsets `s ⊆ G` closed under star. -/ +class star_mem_class (S R : Type*) [has_star R] [set_like S R] := +(star_mem : ∀ {s : S} {r : R}, r ∈ s → star r ∈ s) + +export star_mem_class (star_mem) + +namespace star_mem_class + +variables {S : Type u} [has_star R] [set_like S R] [hS : star_mem_class S R] (s : S) +include hS + +instance : has_star s := +{ star := λ r, ⟨star (r : R), star_mem r.prop⟩ } + +end star_mem_class + /-- Typeclass for a star operation with is involutive. diff --git a/src/algebra/star/subalgebra.lean b/src/algebra/star/subalgebra.lean index 7bf003398f68e..5f1331e95b3bb 100644 --- a/src/algebra/star/subalgebra.lean +++ b/src/algebra/star/subalgebra.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import algebra.star.basic +import algebra.star.star_alg_hom import algebra.algebra.subalgebra.basic /-! @@ -30,8 +30,10 @@ Forgetting that a *-subalgebra is closed under *. -/ add_decl_doc star_subalgebra.to_subalgebra -variables (R : Type u) (A : Type v) [comm_semiring R] [star_ring R] - [semiring A] [star_ring A] [algebra R A] [star_module R A] +variables {R A B C : Type*} [comm_semiring R] [star_ring R] +variables [semiring A] [star_ring A] [algebra R A] [star_module R A] +variables [semiring B] [star_ring B] [algebra R B] [star_module R B] +variables [semiring C] [star_ring C] [algebra R C] [star_module R C] instance : set_like (star_subalgebra R A) A := ⟨star_subalgebra.carrier, λ p q h, by cases p; cases q; congr'⟩ @@ -41,8 +43,165 @@ instance : has_top (star_subalgebra R A) := instance : inhabited (star_subalgebra R A) := ⟨⊤⟩ +instance : star_mem_class (star_subalgebra R A) A := +{ star_mem := λ s a, s.star_mem' } + +instance : subsemiring_class (star_subalgebra R A) A := +{ add_mem := add_mem', + mul_mem := mul_mem', + one_mem := one_mem', + zero_mem := zero_mem' } + +-- this uses the `has_star` instance `s` inherits from `star_mem_class (star_subalgebra R A) A` +instance (s : star_subalgebra R A) : star_ring s := +{ star := star, + star_involutive := λ r, subtype.ext (star_star r), + star_mul := λ r₁ r₂, subtype.ext (star_mul r₁ r₂), + star_add := λ r₁ r₂, subtype.ext (star_add r₁ r₂) } + +instance (s : star_subalgebra R A) : algebra R s := s.to_subalgebra.algebra' + +instance (s : star_subalgebra R A) : star_module R s := +{ star_smul := λ r a, subtype.ext (star_smul r a) } + +@[simp] +lemma mem_carrier {s : star_subalgebra R A} {x : A} : x ∈ s.carrier ↔ x ∈ s := iff.rfl + +@[ext] theorem ext {S T : star_subalgebra R A} (h : ∀ x : A, x ∈ S ↔ x ∈ T) : S = T := +set_like.ext h + +@[simp] lemma mem_to_subalgebra {S : star_subalgebra R A} {x} : x ∈ S.to_subalgebra ↔ x ∈ S := +iff.rfl + +@[simp] lemma coe_to_subalgebra (S : star_subalgebra R A) : (S.to_subalgebra : set A) = S := rfl + +theorem to_subalgebra_injective : + function.injective (to_subalgebra : star_subalgebra R A → subalgebra R A) := +λ S T h, ext $ λ x, by rw [← mem_to_subalgebra, ← mem_to_subalgebra, h] + +theorem to_subalgebra_inj {S U : star_subalgebra R A} : S.to_subalgebra = U.to_subalgebra ↔ S = U := +to_subalgebra_injective.eq_iff + +/-- Copy of a star subalgebra with a new `carrier` equal to the old one. Useful to fix definitional +equalities. -/ +protected def copy (S : star_subalgebra R A) (s : set A) (hs : s = ↑S) : star_subalgebra R A := +{ carrier := s, + add_mem' := λ _ _, hs.symm ▸ S.add_mem', + mul_mem' := λ _ _, hs.symm ▸ S.mul_mem', + algebra_map_mem' := hs.symm ▸ S.algebra_map_mem', + star_mem' := λ _, hs.symm ▸ S.star_mem' } + +@[simp] lemma coe_copy (S : star_subalgebra R A) (s : set A) (hs : s = ↑S) : + (S.copy s hs : set A) = s := rfl + +lemma copy_eq (S : star_subalgebra R A) (s : set A) (hs : s = ↑S) : S.copy s hs = S := +set_like.coe_injective hs + +variables (S : star_subalgebra R A) + +theorem algebra_map_mem (r : R) : algebra_map R A r ∈ S := +S.algebra_map_mem' r + +theorem srange_le : (algebra_map R A).srange ≤ S.to_subalgebra.to_subsemiring := +λ x ⟨r, hr⟩, hr ▸ S.algebra_map_mem r + +theorem range_subset : set.range (algebra_map R A) ⊆ S := +λ x ⟨r, hr⟩, hr ▸ S.algebra_map_mem r + +theorem range_le : set.range (algebra_map R A) ≤ S := +S.range_subset + +protected theorem smul_mem {x : A} (hx : x ∈ S) (r : R) : r • x ∈ S := +(algebra.smul_def r x).symm ▸ mul_mem (S.algebra_map_mem r) hx + +/-- Embedding of a subalgebra into the algebra. -/ +def subtype : S →⋆ₐ[R] A := +by refine_struct { to_fun := (coe : S → A) }; intros; refl + +@[simp] lemma coe_subtype : (S.subtype : S → A) = coe := rfl + +lemma subtype_apply (x : S) : S.subtype x = (x : A) := rfl + +@[simp] lemma to_subalgebra_subtype : S.to_subalgebra.val = S.subtype.to_alg_hom := +rfl + +section map + +/-- Transport a star subalgebra via a star algebra homomorphism. -/ +def map (f : A →⋆ₐ[R] B) (S : star_subalgebra R A) : star_subalgebra R B := +{ star_mem' := + begin + rintro _ ⟨a, ha, rfl⟩, + exact map_star f a ▸ set.mem_image_of_mem _ (S.star_mem' ha), + end, + .. S.to_subalgebra.map f.to_alg_hom } + +lemma map_mono {S₁ S₂ : star_subalgebra R A} {f : A →⋆ₐ[R] B} : + S₁ ≤ S₂ → S₁.map f ≤ S₂.map f := +set.image_subset f + +lemma map_injective {f : A →⋆ₐ[R] B} (hf : function.injective f) : + function.injective (map f) := +λ S₁ S₂ ih, ext $ set.ext_iff.1 $ set.image_injective.2 hf $ set.ext $ set_like.ext_iff.mp ih + +@[simp] lemma map_id (S : star_subalgebra R A) : S.map (star_alg_hom.id R A) = S := +set_like.coe_injective $ set.image_id _ + +lemma map_map (S : star_subalgebra R A) (g : B →⋆ₐ[R] C) (f : A →⋆ₐ[R] B) : + (S.map f).map g = S.map (g.comp f) := +set_like.coe_injective $ set.image_image _ _ _ + +lemma mem_map {S : star_subalgebra R A} {f : A →⋆ₐ[R] B} {y : B} : + y ∈ map f S ↔ ∃ x ∈ S, f x = y := +subsemiring.mem_map + +lemma map_to_subalgebra {S : star_subalgebra R A} {f : A →⋆ₐ[R] B} : + (S.map f).to_subalgebra = S.to_subalgebra.map f.to_alg_hom := +set_like.coe_injective rfl + +@[simp] lemma coe_map (S : star_subalgebra R A) (f : A →⋆ₐ[R] B) : + (S.map f : set B) = f '' S := +rfl + +/-- Preimage of a star subalgebra under an star algebra homomorphism. -/ +def comap (f : A →⋆ₐ[R] B) (S : star_subalgebra R B) : star_subalgebra R A := +{ star_mem' := λ a ha, show f (star a) ∈ S, from (map_star f a).symm ▸ star_mem ha, + .. S.to_subalgebra.comap f.to_alg_hom } + +theorem map_le_iff_le_comap {S : star_subalgebra R A} {f : A →⋆ₐ[R] B} {U : star_subalgebra R B} : + map f S ≤ U ↔ S ≤ comap f U := +set.image_subset_iff + +lemma gc_map_comap (f : A →⋆ₐ[R] B) : galois_connection (map f) (comap f) := +λ S U, map_le_iff_le_comap + +lemma comap_mono {S₁ S₂ : star_subalgebra R B} {f : A →⋆ₐ[R] B} : + S₁ ≤ S₂ → S₁.comap f ≤ S₂.comap f := +set.preimage_mono + +lemma comap_injective {f : A →⋆ₐ[R] B} (hf : function.surjective f) : + function.injective (comap f) := +λ S₁ S₂ h, ext $ λ b, let ⟨x, hx⟩ := hf b in let this := set_like.ext_iff.1 h x in hx ▸ this + +@[simp] lemma comap_id (S : star_subalgebra R A) : S.comap (star_alg_hom.id R A) = S := +set_like.coe_injective $ set.preimage_id + +lemma comap_comap (S : star_subalgebra R C) (g : B →⋆ₐ[R] C) (f : A →⋆ₐ[R] B) : + (S.comap g).comap f = S.comap (g.comp f) := +set_like.coe_injective $ set.preimage_preimage + +@[simp] lemma mem_comap (S : star_subalgebra R B) (f : A →⋆ₐ[R] B) (x : A) : + x ∈ S.comap f ↔ f x ∈ S := +iff.rfl + +@[simp, norm_cast] lemma coe_comap (S : star_subalgebra R B) (f : A →⋆ₐ[R] B) : + (S.comap f : set A) = f ⁻¹' (S : set B) := +rfl + +end map + section centralizer -variables {A} +variables (R) {A} /-- The centralizer, or commutant, of a *-closed set as star subalgebra. -/ def centralizer From 4a4d740d7099034d076168755cb4e45ce64dd160 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 15 Oct 2022 07:59:13 +0000 Subject: [PATCH 776/828] =?UTF-8?q?feat(logic/lemmas):=20Aliases=20for=20`?= =?UTF-8?q?=3D`-`=E2=89=A0`=20transitivity=20lemmas=20(#16959)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/logic/lemmas.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/src/logic/lemmas.lean b/src/logic/lemmas.lean index c91dd5cc78d0b..f61df64f5fc50 100644 --- a/src/logic/lemmas.lean +++ b/src/logic/lemmas.lean @@ -22,6 +22,9 @@ alias heq_iff_eq ↔ heq.eq eq.heq attribute [protected] heq.eq eq.heq +alias ne_of_eq_of_ne ← eq.trans_ne +alias ne_of_ne_of_eq ← ne.trans_eq + variables {α : Sort*} {p q r : Prop} [decidable p] [decidable q] {a b c : α} lemma dite_dite_distrib_left {a : p → α} {b : ¬ p → q → α} {c : ¬ p → ¬ q → α} : From e05f1302d785c39d03e3a1276348f087922e0511 Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Sat, 15 Oct 2022 10:55:30 +0000 Subject: [PATCH 777/828] feat(topology/algebra/module/basic): a hyperplane is either closed or dense (#16782) --- src/topology/algebra/module/basic.lean | 24 ++++++++++++++++++- .../algebra/module/finite_dimension.lean | 13 ++++++++++ 2 files changed, 36 insertions(+), 1 deletion(-) diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index a1467fa48599f..06a28e425d605 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -12,6 +12,7 @@ import algebra.algebra.basic import linear_algebra.projection import linear_algebra.pi import linear_algebra.determinant +import ring_theory.simple_module /-! # Theory of topological modules and continuous linear maps. @@ -167,10 +168,13 @@ S.to_add_subgroup.topological_add_group end submodule section closure -variables {R : Type u} {M : Type v} +variables {R R' : Type u} {M M' : Type v} [semiring R] [topological_space R] +[ring R'] [topological_space R'] [topological_space M] [add_comm_monoid M] +[topological_space M'] [add_comm_group M'] [module R M] [has_continuous_smul R M] +[module R' M'] [has_continuous_smul R' M'] lemma submodule.closure_smul_self_subset (s : submodule R M) : (λ p : R × M, p.1 • p.2) '' (set.univ ×ˢ closure s) ⊆ closure s := @@ -238,6 +242,24 @@ instance {M' : Type*} [add_comm_monoid M'] [module R M'] [uniform_space M'] complete_space U.topological_closure := is_closed_closure.complete_space_coe +/-- A maximal proper subspace of a topological module (i.e a `submodule` satisfying `is_coatom`) +is either closed or dense. -/ +lemma submodule.is_closed_or_dense_of_is_coatom (s : submodule R M) (hs : is_coatom s) : + is_closed (s : set M) ∨ dense (s : set M) := +(hs.le_iff.mp s.submodule_topological_closure).swap.imp (is_closed_of_closure_subset ∘ eq.le) + submodule.dense_iff_topological_closure_eq_top.mpr + +lemma linear_map.is_closed_or_dense_ker [has_continuous_add M'] [is_simple_module R' R'] + (l : M' →ₗ[R'] R') : + is_closed (l.ker : set M') ∨ dense (l.ker : set M') := +begin + rcases l.surjective_or_eq_zero with (hl|rfl), + { refine l.ker.is_closed_or_dense_of_is_coatom (linear_map.is_coatom_ker_of_surjective hl) }, + { rw linear_map.ker_zero, + left, + exact is_closed_univ }, +end + end closure /-- Continuous linear maps between modules. We only put the type classes that are necessary for the diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index 5ff331d315321..db94793b01f22 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -197,6 +197,19 @@ lemma linear_map.continuous_iff_is_closed_ker (l : E →ₗ[𝕜] 𝕜) : continuous l ↔ is_closed (l.ker : set E) := ⟨λ h, is_closed_singleton.preimage h, l.continuous_of_is_closed_ker⟩ +/-- Over a nontrivially normed field, any linear form which is nonzero on a nonempty open set is + automatically continuous. -/ +lemma linear_map.continuous_of_nonzero_on_open (l : E →ₗ[𝕜] 𝕜) (s : set E) (hs₁ : is_open s) + (hs₂ : s.nonempty) (hs₃ : ∀ x ∈ s, l x ≠ 0) : continuous l := +begin + refine l.continuous_of_is_closed_ker (l.is_closed_or_dense_ker.resolve_right $ λ hl, _), + rcases hs₂ with ⟨x, hx⟩, + have : x ∈ interior (l.ker : set E)ᶜ, + { rw mem_interior_iff_mem_nhds, + exact mem_of_superset (hs₁.mem_nhds hx) hs₃ }, + rwa hl.interior_compl at this +end + variables [complete_space 𝕜] /-- This version imposes `ι` and `E` to live in the same universe, so you should instead use From 3605313c4201bcc7d5cfd62dfb40ab3439b8ce83 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Sat, 15 Oct 2022 10:55:31 +0000 Subject: [PATCH 778/828] feat(data/polynomial): a lemma relating `nth_roots 2` to `is_square` (#16990) --- src/data/polynomial/ring_division.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index 608c51c1090f9..fc6cc0503c5e1 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -575,6 +575,11 @@ then if h : (X : R[X]) ^ n - C a = 0 else by rw [← with_bot.coe_le_coe, ← degree_X_pow_sub_C (nat.pos_of_ne_zero hn) a]; exact card_roots (X_pow_sub_C_ne_zero (nat.pos_of_ne_zero hn) a) +@[simp] +lemma nth_roots_two_eq_zero_iff {r : R} : nth_roots 2 r = 0 ↔ ¬ is_square r := +by simp_rw [is_square_iff_exists_sq, eq_zero_iff_forall_not_mem, + mem_nth_roots (by norm_num : 0 < 2), ← not_exists, eq_comm] + /-- The multiset `nth_roots ↑n (1 : R)` as a finset. -/ def nth_roots_finset (n : ℕ) (R : Type*) [comm_ring R] [is_domain R] : finset R := multiset.to_finset (nth_roots n (1 : R)) From 583cae777bea4e5188198ee8e367ab2ea8200434 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Sat, 15 Oct 2022 13:04:14 +0000 Subject: [PATCH 779/828] chore(analysis/normed_space/operator_norm): fix naming of `continuous_linear_map.lmul` and similar (#15968) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This renames the continuous linear (bilinear) map given by multiplication in a non-unital algebra to match the non-continuous version (`linear_map.mul`). The changes are as follows: 1. `continuous_linear_map.lmul` → `continuous_linear_map.mul` (the "l" was for "linear", which is redundant) 2. `continuous_linear_map.lmul_right` is deleted. This is literally just the `continuous_linear_map.flip` of the map in (1) above. It's only use in mathlib is to define the map in (3) below, and creating it as a separate def seems unnecessary (and actively obscures the `.flip`) 3. `continuous_linear_map.lmul_left_right` → `continuous_linear_map.mul_left_right`. This matches `linear_map.mul_left_right` (although the latter is uncurried). Docstrings have also been updated, both to reflect the non-unital generalization from #15868, but also because the old docstrings mentioned "left multiplication", but this should just be called "multiplication". --- src/analysis/calculus/cont_diff.lean | 6 +- src/analysis/calculus/fderiv.lean | 24 +++--- src/analysis/convolution.lean | 10 +-- .../normed_space/bounded_linear_maps.lean | 9 +- src/analysis/normed_space/operator_norm.lean | 83 +++++++------------ 5 files changed, 55 insertions(+), 77 deletions(-) diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 04f2c3399ab00..2c3b1f22ecdbc 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -2582,7 +2582,7 @@ variables {𝔸 𝔸' ι 𝕜' : Type*} [normed_ring 𝔸] [normed_algebra 𝕜 /- The product is smooth. -/ lemma cont_diff_mul : cont_diff 𝕜 n (λ p : 𝔸 × 𝔸, p.1 * p.2) := -(continuous_linear_map.lmul 𝕜 𝔸).is_bounded_bilinear_map.cont_diff +(continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.cont_diff /-- The product of two `C^n` functions within a set at a point is `C^n` within this set at this point. -/ @@ -2881,12 +2881,12 @@ begin exact (inverse_continuous_at x').continuous_within_at }, { simp [ftaylor_series_within] } } }, { apply cont_diff_at_succ_iff_has_fderiv_at.mpr, - refine ⟨λ (x : R), - lmul_left_right 𝕜 R (inverse x) (inverse x), _, _⟩, + refine ⟨λ (x : R), - mul_left_right 𝕜 R (inverse x) (inverse x), _, _⟩, { refine ⟨{y : R | is_unit y}, x.nhds, _⟩, rintros _ ⟨y, rfl⟩, rw [inverse_unit], exact has_fderiv_at_ring_inverse y }, - { convert (lmul_left_right_is_bounded_bilinear 𝕜 R).cont_diff.neg.comp_cont_diff_at + { convert (mul_left_right_is_bounded_bilinear 𝕜 R).cont_diff.neg.comp_cont_diff_at (x : R) (IH.prod IH) } }, { exact cont_diff_at_top.mpr Itop } end diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index 320afbfbf658a..7e103ee6a0757 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -2371,7 +2371,7 @@ variables {𝔸 𝔸' : Type*} [normed_ring 𝔸] [normed_comm_ring 𝔸'] [norm theorem has_strict_fderiv_at.mul' {x : E} (ha : has_strict_fderiv_at a a' x) (hb : has_strict_fderiv_at b b' x) : has_strict_fderiv_at (λ y, a y * b y) (a x • b' + a'.smul_right (b x)) x := -((continuous_linear_map.lmul 𝕜 𝔸).is_bounded_bilinear_map.has_strict_fderiv_at (a x, b x)).comp x +((continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.has_strict_fderiv_at (a x, b x)).comp x (ha.prod hb) theorem has_strict_fderiv_at.mul @@ -2382,7 +2382,7 @@ by { convert hc.mul' hd, ext z, apply mul_comm } theorem has_fderiv_within_at.mul' (ha : has_fderiv_within_at a a' s x) (hb : has_fderiv_within_at b b' s x) : has_fderiv_within_at (λ y, a y * b y) (a x • b' + a'.smul_right (b x)) s x := -((continuous_linear_map.lmul 𝕜 𝔸).is_bounded_bilinear_map.has_fderiv_at +((continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.has_fderiv_at (a x, b x)).comp_has_fderiv_within_at x (ha.prod hb) theorem has_fderiv_within_at.mul @@ -2393,7 +2393,7 @@ by { convert hc.mul' hd, ext z, apply mul_comm } theorem has_fderiv_at.mul' (ha : has_fderiv_at a a' x) (hb : has_fderiv_at b b' x) : has_fderiv_at (λ y, a y * b y) (a x • b' + a'.smul_right (b x)) x := -((continuous_linear_map.lmul 𝕜 𝔸).is_bounded_bilinear_map.has_fderiv_at (a x, b x)).comp x +((continuous_linear_map.mul 𝕜 𝔸).is_bounded_bilinear_map.has_fderiv_at (a x, b x)).comp x (ha.prod hb) theorem has_fderiv_at.mul (hc : has_fderiv_at c c' x) (hd : has_fderiv_at d d' x) : @@ -2458,7 +2458,7 @@ lemma fderiv_mul (hc : differentiable_at 𝕜 c x) (hd : differentiable_at 𝕜 theorem has_strict_fderiv_at.mul_const' (ha : has_strict_fderiv_at a a' x) (b : 𝔸) : has_strict_fderiv_at (λ y, a y * b) (a'.smul_right b) x := -(((continuous_linear_map.lmul 𝕜 𝔸).flip b).has_strict_fderiv_at).comp x ha +(((continuous_linear_map.mul 𝕜 𝔸).flip b).has_strict_fderiv_at).comp x ha theorem has_strict_fderiv_at.mul_const (hc : has_strict_fderiv_at c c' x) (d : 𝔸') : has_strict_fderiv_at (λ y, c y * d) (d • c') x := @@ -2466,7 +2466,7 @@ by { convert hc.mul_const' d, ext z, apply mul_comm } theorem has_fderiv_within_at.mul_const' (ha : has_fderiv_within_at a a' s x) (b : 𝔸) : has_fderiv_within_at (λ y, a y * b) (a'.smul_right b) s x := -(((continuous_linear_map.lmul 𝕜 𝔸).flip b).has_fderiv_at).comp_has_fderiv_within_at x ha +(((continuous_linear_map.mul 𝕜 𝔸).flip b).has_fderiv_at).comp_has_fderiv_within_at x ha theorem has_fderiv_within_at.mul_const (hc : has_fderiv_within_at c c' s x) (d : 𝔸') : has_fderiv_within_at (λ y, c y * d) (d • c') s x := @@ -2474,7 +2474,7 @@ by { convert hc.mul_const' d, ext z, apply mul_comm } theorem has_fderiv_at.mul_const' (ha : has_fderiv_at a a' x) (b : 𝔸) : has_fderiv_at (λ y, a y * b) (a'.smul_right b) x := -(((continuous_linear_map.lmul 𝕜 𝔸).flip b).has_fderiv_at).comp x ha +(((continuous_linear_map.mul 𝕜 𝔸).flip b).has_fderiv_at).comp x ha theorem has_fderiv_at.mul_const (hc : has_fderiv_at c c' x) (d : 𝔸') : has_fderiv_at (λ y, c y * d) (d • c') x := @@ -2517,16 +2517,16 @@ lemma fderiv_mul_const (hc : differentiable_at 𝕜 c x) (d : 𝔸') : theorem has_strict_fderiv_at.const_mul (ha : has_strict_fderiv_at a a' x) (b : 𝔸) : has_strict_fderiv_at (λ y, b * a y) (b • a') x := -(((continuous_linear_map.lmul 𝕜 𝔸) b).has_strict_fderiv_at).comp x ha +(((continuous_linear_map.mul 𝕜 𝔸) b).has_strict_fderiv_at).comp x ha theorem has_fderiv_within_at.const_mul (ha : has_fderiv_within_at a a' s x) (b : 𝔸) : has_fderiv_within_at (λ y, b * a y) (b • a') s x := -(((continuous_linear_map.lmul 𝕜 𝔸) b).has_fderiv_at).comp_has_fderiv_within_at x ha +(((continuous_linear_map.mul 𝕜 𝔸) b).has_fderiv_at).comp_has_fderiv_within_at x ha theorem has_fderiv_at.const_mul (ha : has_fderiv_at a a' x) (b : 𝔸) : has_fderiv_at (λ y, b * a y) (b • a') x := -(((continuous_linear_map.lmul 𝕜 𝔸) b).has_fderiv_at).comp x ha +(((continuous_linear_map.mul 𝕜 𝔸) b).has_fderiv_at).comp x ha lemma differentiable_within_at.const_mul (ha : differentiable_within_at 𝕜 a s x) (b : 𝔸) : @@ -2563,7 +2563,7 @@ open normed_ring continuous_linear_map ring /-- At an invertible element `x` of a normed algebra `R`, the Fréchet derivative of the inversion operation is the linear map `λ t, - x⁻¹ * t * x⁻¹`. -/ lemma has_fderiv_at_ring_inverse (x : Rˣ) : - has_fderiv_at ring.inverse (-lmul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹) x := + has_fderiv_at ring.inverse (-mul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹) x := begin have h_is_o : (λ (t : R), inverse (↑x + t) - ↑x⁻¹ + ↑x⁻¹ * t * ↑x⁻¹) =o[𝓝 0] (λ (t : R), t), { refine (inverse_add_norm_diff_second_order x).trans_is_o ((is_o_norm_norm).mp _), @@ -2577,7 +2577,7 @@ begin simp only [has_fderiv_at, has_fderiv_at_filter], convert h_is_o.comp_tendsto h_lim, ext y, - simp only [coe_comp', function.comp_app, lmul_left_right_apply, neg_apply, inverse_unit x, + simp only [coe_comp', function.comp_app, mul_left_right_apply, neg_apply, inverse_unit x, units.inv_mul, add_sub_cancel'_right, mul_sub, sub_mul, one_mul, sub_neg_eq_add] end @@ -2585,7 +2585,7 @@ lemma differentiable_at_inverse (x : Rˣ) : differentiable_at 𝕜 (@ring.invers (has_fderiv_at_ring_inverse x).differentiable_at lemma fderiv_inverse (x : Rˣ) : - fderiv 𝕜 (@ring.inverse R _) x = - lmul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹ := + fderiv 𝕜 (@ring.inverse R _) x = - mul_left_right 𝕜 R ↑x⁻¹ ↑x⁻¹ := (has_fderiv_at_ring_inverse x).fderiv end algebra_inverse diff --git a/src/analysis/convolution.lean b/src/analysis/convolution.lean index 1b560154cb2a5..0ed115d6fc8cd 100644 --- a/src/analysis/convolution.lean +++ b/src/analysis/convolution.lean @@ -18,7 +18,7 @@ group as domain. We use a continuous bilinear operation `L` on these function va "multiplication". The domain must be equipped with a Haar measure `μ` (though many individual results have weaker conditions on `μ`). -For many applications we can take `L = lsmul ℝ ℝ` or `L = lmul ℝ ℝ`. +For many applications we can take `L = lsmul ℝ ℝ` or `L = mul ℝ ℝ`. We also define `convolution_exists` and `convolution_exists_at` to state that the convolution is well-defined (everywhere or at a single point). These conditions are needed for pointwise @@ -375,8 +375,8 @@ lemma convolution_lsmul [has_sub G] {f : G → 𝕜} {g : G → F} : (f ⋆[lsmul 𝕜 𝕜, μ] g : G → F) x = ∫ t, f t • g (x - t) ∂μ := rfl /-- The definition of convolution where the bilinear operator is multiplication. -/ -lemma convolution_lmul [has_sub G] [normed_space ℝ 𝕜] [complete_space 𝕜] {f : G → 𝕜} {g : G → 𝕜} : - (f ⋆[lmul 𝕜 𝕜, μ] g) x = ∫ t, f t * g (x - t) ∂μ := rfl +lemma convolution_mul [has_sub G] [normed_space ℝ 𝕜] [complete_space 𝕜] {f : G → 𝕜} {g : G → 𝕜} : + (f ⋆[mul 𝕜 𝕜, μ] g) x = ∫ t, f t * g (x - t) ∂μ := rfl section group @@ -547,8 +547,8 @@ lemma convolution_lsmul_swap {f : G → 𝕜} {g : G → F}: convolution_eq_swap _ /-- The symmetric definition of convolution where the bilinear operator is multiplication. -/ -lemma convolution_lmul_swap [normed_space ℝ 𝕜] [complete_space 𝕜] {f : G → 𝕜} {g : G → 𝕜} : - (f ⋆[lmul 𝕜 𝕜, μ] g) x = ∫ t, f (x - t) * g t ∂μ := +lemma convolution_mul_swap [normed_space ℝ 𝕜] [complete_space 𝕜] {f : G → 𝕜} {g : G → 𝕜} : + (f ⋆[mul 𝕜 𝕜, μ] g) x = ∫ t, f (x - t) * g t ∂μ := convolution_eq_swap _ end measurable diff --git a/src/analysis/normed_space/bounded_linear_maps.lean b/src/analysis/normed_space/bounded_linear_maps.lean index 3e53c007d2d19..2c9f6ffffa049 100644 --- a/src/analysis/normed_space/bounded_linear_maps.lean +++ b/src/analysis/normed_space/bounded_linear_maps.lean @@ -488,11 +488,12 @@ end variables (𝕜) -/-- The function `lmul_left_right : 𝕜' × 𝕜' → (𝕜' →L[𝕜] 𝕜')` is a bounded bilinear map. -/ -lemma continuous_linear_map.lmul_left_right_is_bounded_bilinear +/-- The function `continuous_linear_map.mul_left_right : 𝕜' × 𝕜' → (𝕜' →L[𝕜] 𝕜')` is a bounded +bilinear map. -/ +lemma continuous_linear_map.mul_left_right_is_bounded_bilinear (𝕜' : Type*) [normed_ring 𝕜'] [normed_algebra 𝕜 𝕜'] : - is_bounded_bilinear_map 𝕜 (λ p : 𝕜' × 𝕜', continuous_linear_map.lmul_left_right 𝕜 𝕜' p.1 p.2) := -(continuous_linear_map.lmul_left_right 𝕜 𝕜').is_bounded_bilinear_map + is_bounded_bilinear_map 𝕜 (λ p : 𝕜' × 𝕜', continuous_linear_map.mul_left_right 𝕜 𝕜' p.1 p.2) := +(continuous_linear_map.mul_left_right 𝕜 𝕜').is_bounded_bilinear_map variables {𝕜} diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index f15433ba581cd..44fcf1655e801 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -828,77 +828,56 @@ section non_unital variables (𝕜) (𝕜' : Type*) [non_unital_semi_normed_ring 𝕜'] [normed_space 𝕜 𝕜'] [is_scalar_tower 𝕜 𝕜' 𝕜'] [smul_comm_class 𝕜 𝕜' 𝕜'] -/-- Left multiplication in a normed algebra as a continuous bilinear map. -/ -def lmul : 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' := -(linear_map.mul 𝕜 𝕜').mk_continuous₂ 1 $ +/-- Multiplication in a non-unital normed algebra as a continuous bilinear map. -/ +def mul : 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' := (linear_map.mul 𝕜 𝕜').mk_continuous₂ 1 $ λ x y, by simpa using norm_mul_le x y -@[simp] lemma lmul_apply (x y : 𝕜') : lmul 𝕜 𝕜' x y = x * y := rfl +@[simp] lemma mul_apply' (x y : 𝕜') : mul 𝕜 𝕜' x y = x * y := rfl -@[simp] lemma op_norm_lmul_apply_le (x : 𝕜') : ∥lmul 𝕜 𝕜' x∥ ≤ ∥x∥ := +@[simp] lemma op_norm_mul_apply_le (x : 𝕜') : ∥mul 𝕜 𝕜' x∥ ≤ ∥x∥ := (op_norm_le_bound _ (norm_nonneg x) (norm_mul_le x)) -/-- Right-multiplication in a normed algebra, considered as a continuous linear map. -/ -def lmul_right : 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' := (lmul 𝕜 𝕜').flip +/-- Simultaneous left- and right-multiplication in a non-unital normed algebra, considered as a +continuous trilinear map. This is akin to its non-continuous version `linear_map.mul_left_right`, +but there is a minor difference: `linear_map.mul_left_right` is uncurried. -/ +def mul_left_right : 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' := +((compL 𝕜 𝕜' 𝕜' 𝕜').comp (mul 𝕜 𝕜').flip).flip.comp (mul 𝕜 𝕜') -@[simp] lemma lmul_right_apply (x y : 𝕜') : lmul_right 𝕜 𝕜' x y = y * x := rfl +@[simp] lemma mul_left_right_apply (x y z : 𝕜') : + mul_left_right 𝕜 𝕜' x y z = x * z * y := rfl -@[simp] lemma op_norm_lmul_right_apply_le (x : 𝕜') : ∥lmul_right 𝕜 𝕜' x∥ ≤ ∥x∥ := -op_norm_le_bound _ (norm_nonneg x) (λ y, (norm_mul_le y x).trans_eq (mul_comm _ _)) - -/-- Simultaneous left- and right-multiplication in a normed algebra, considered as a continuous -trilinear map. -/ -def lmul_left_right : 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' →L[𝕜] 𝕜' := -((compL 𝕜 𝕜' 𝕜' 𝕜').comp (lmul_right 𝕜 𝕜')).flip.comp (lmul 𝕜 𝕜') - -@[simp] lemma lmul_left_right_apply (x y z : 𝕜') : - lmul_left_right 𝕜 𝕜' x y z = x * z * y := rfl - -lemma op_norm_lmul_left_right_apply_apply_le (x y : 𝕜') : - ∥lmul_left_right 𝕜 𝕜' x y∥ ≤ ∥x∥ * ∥y∥ := +lemma op_norm_mul_left_right_apply_apply_le (x y : 𝕜') : + ∥mul_left_right 𝕜 𝕜' x y∥ ≤ ∥x∥ * ∥y∥ := (op_norm_comp_le _ _).trans $ (mul_comm _ _).trans_le $ - mul_le_mul (op_norm_lmul_apply_le _ _ _) (op_norm_lmul_right_apply_le _ _ _) + mul_le_mul (op_norm_mul_apply_le _ _ _) + (op_norm_le_bound _ (norm_nonneg _) (λ _, (norm_mul_le _ _).trans_eq (mul_comm _ _))) (norm_nonneg _) (norm_nonneg _) -lemma op_norm_lmul_left_right_apply_le (x : 𝕜') : - ∥lmul_left_right 𝕜 𝕜' x∥ ≤ ∥x∥ := -op_norm_le_bound _ (norm_nonneg x) (op_norm_lmul_left_right_apply_apply_le 𝕜 𝕜' x) +lemma op_norm_mul_left_right_apply_le (x : 𝕜') : + ∥mul_left_right 𝕜 𝕜' x∥ ≤ ∥x∥ := +op_norm_le_bound _ (norm_nonneg x) (op_norm_mul_left_right_apply_apply_le 𝕜 𝕜' x) -lemma op_norm_lmul_left_right_le : - ∥lmul_left_right 𝕜 𝕜'∥ ≤ 1 := -op_norm_le_bound _ zero_le_one (λ x, (one_mul ∥x∥).symm ▸ op_norm_lmul_left_right_apply_le 𝕜 𝕜' x) +lemma op_norm_mul_left_right_le : + ∥mul_left_right 𝕜 𝕜'∥ ≤ 1 := +op_norm_le_bound _ zero_le_one (λ x, (one_mul ∥x∥).symm ▸ op_norm_mul_left_right_apply_le 𝕜 𝕜' x) end non_unital section unital variables (𝕜) (𝕜' : Type*) [semi_normed_ring 𝕜'] [normed_algebra 𝕜 𝕜'] [norm_one_class 𝕜'] -/-- Left multiplication in a normed algebra as a linear isometry to the space of +/-- Multiplication in a normed algebra as a linear isometry to the space of continuous linear maps. -/ -def lmulₗᵢ : 𝕜' →ₗᵢ[𝕜] 𝕜' →L[𝕜] 𝕜' := -{ to_linear_map := lmul 𝕜 𝕜', - norm_map' := λ x, le_antisymm (op_norm_lmul_apply_le _ _ _) +def mulₗᵢ : 𝕜' →ₗᵢ[𝕜] 𝕜' →L[𝕜] 𝕜' := +{ to_linear_map := mul 𝕜 𝕜', + norm_map' := λ x, le_antisymm (op_norm_mul_apply_le _ _ _) (by { convert ratio_le_op_norm _ (1 : 𝕜'), simp [norm_one], apply_instance }) } -@[simp] lemma coe_lmulₗᵢ : ⇑(lmulₗᵢ 𝕜 𝕜') = lmul 𝕜 𝕜' := rfl - -@[simp] lemma op_norm_lmul_apply (x : 𝕜') : ∥lmul 𝕜 𝕜' x∥ = ∥x∥ := -(lmulₗᵢ 𝕜 𝕜').norm_map x - -@[simp] lemma op_norm_lmul_right_apply (x : 𝕜') : ∥lmul_right 𝕜 𝕜' x∥ = ∥x∥ := -le_antisymm - (op_norm_lmul_right_apply_le _ _ _) - (by { convert ratio_le_op_norm _ (1 : 𝕜'), simp [norm_one], - apply_instance }) - -/-- Right-multiplication in a normed algebra, considered as a linear isometry to the space of -continuous linear maps. -/ -def lmul_rightₗᵢ : 𝕜' →ₗᵢ[𝕜] 𝕜' →L[𝕜] 𝕜' := -{ to_linear_map := lmul_right 𝕜 𝕜', - norm_map' := op_norm_lmul_right_apply 𝕜 𝕜' } +@[simp] lemma coe_mulₗᵢ : ⇑(mulₗᵢ 𝕜 𝕜') = mul 𝕜 𝕜' := rfl -@[simp] lemma coe_lmul_rightₗᵢ : ⇑(lmul_rightₗᵢ 𝕜 𝕜') = lmul_right 𝕜 𝕜' := rfl +@[simp] lemma op_norm_mul_apply (x : 𝕜') : ∥mul 𝕜 𝕜' x∥ = ∥x∥ := +(mulₗᵢ 𝕜 𝕜').norm_map x end unital @@ -1659,11 +1638,9 @@ variables (𝕜) (𝕜' : Type*) section variables [normed_ring 𝕜'] [normed_algebra 𝕜 𝕜'] -@[simp] lemma op_norm_lmul [norm_one_class 𝕜'] : ∥lmul 𝕜 𝕜'∥ = 1 := -by haveI := norm_one_class.nontrivial 𝕜'; exact (lmulₗᵢ 𝕜 𝕜').norm_to_continuous_linear_map +@[simp] lemma op_norm_mul [norm_one_class 𝕜'] : ∥mul 𝕜 𝕜'∥ = 1 := +by haveI := norm_one_class.nontrivial 𝕜'; exact (mulₗᵢ 𝕜 𝕜').norm_to_continuous_linear_map -@[simp] lemma op_norm_lmul_right [norm_one_class 𝕜'] : ∥lmul_right 𝕜 𝕜'∥ = 1 := -(op_norm_flip (lmul 𝕜 𝕜')).trans (op_norm_lmul _ _) end /-- The norm of `lsmul` equals 1 in any nontrivial normed group. From 86bf83abdd6a11a6a52ec03782ef749f8fd595c9 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Sat, 15 Oct 2022 13:04:15 +0000 Subject: [PATCH 780/828] feat(algebra/ring): mark a useful lemma simp and generalize unit neg lemmas (#16993) Mark the lemma that states `is_unit (-x) \iff is_unit x` as `simp`, and while we are here generalize a couple of typeclasses to a general monoid with some distributive negation. --- src/algebra/ring/basic.lean | 17 ++++++++++++++--- 1 file changed, 14 insertions(+), 3 deletions(-) diff --git a/src/algebra/ring/basic.lean b/src/algebra/ring/basic.lean index afef8dabffe13..0460017430973 100644 --- a/src/algebra/ring/basic.lean +++ b/src/algebra/ring/basic.lean @@ -784,7 +784,9 @@ protected def function.surjective.ring end ring namespace units -variables [ring α] {a b : α} + +section has_distrib_neg +variables [monoid α] [has_distrib_neg α] {a b : α} /-- Each element of the group of units of a ring has an additive inverse. -/ instance : has_neg αˣ := ⟨λu, ⟨-↑u, -↑u⁻¹, by simp, by simp⟩ ⟩ @@ -800,6 +802,12 @@ instance : has_distrib_neg αˣ := units.ext.has_distrib_neg _ units.coe_neg uni @[field_simps] lemma neg_divp (a : α) (u : αˣ) : -(a /ₚ u) = (-a) /ₚ u := by simp only [divp, neg_mul] +end has_distrib_neg + +section ring + +variables [ring α] {a b : α} + @[field_simps] lemma divp_add_divp_same (a b : α) (u : αˣ) : a /ₚ u + b /ₚ u = (a + b) /ₚ u := by simp only [divp, add_mul] @@ -823,12 +831,15 @@ begin assoc_rw [units.mul_inv, mul_one], end +end ring + end units -lemma is_unit.neg [ring α] {a : α} : is_unit a → is_unit (-a) +lemma is_unit.neg [monoid α] [has_distrib_neg α] {a : α} : is_unit a → is_unit (-a) | ⟨x, hx⟩ := hx ▸ (-x).is_unit -lemma is_unit.neg_iff [ring α] (a : α) : is_unit (-a) ↔ is_unit a := +@[simp] +lemma is_unit.neg_iff [monoid α] [has_distrib_neg α] (a : α) : is_unit (-a) ↔ is_unit a := ⟨λ h, neg_neg a ▸ h.neg, is_unit.neg⟩ lemma is_unit.sub_iff [ring α] {x y : α} : From 11c53f174270aa43140c0b26dabce5fc4a253e80 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Dupuis?= Date: Sat, 15 Oct 2022 15:59:18 +0000 Subject: [PATCH 781/828] fix(*): use `old_structure_cmd` for all morphism classes (#16242) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR changes all morphism classes to use the old structure command, as briefly discussed [here](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/old_structure_cmd.20in.20morphism.20type.20classes). Co-authored-by: Frédéric Dupuis <31101893+dupuisf@users.noreply.github.com> Co-authored-by: Jireh Loreaux --- src/algebra/order/hom/monoid.lean | 21 +++++++++++++++-- src/data/fun_like/embedding.lean | 12 ++++++++++ src/data/fun_like/equiv.lean | 2 +- src/model_theory/elementary_maps.lean | 3 ++- src/order/hom/basic.lean | 5 ++++ src/order/hom/bounded.lean | 16 ++++++++++--- src/order/hom/complete_lattice.lean | 23 +++++++++++++++---- src/order/hom/lattice.lean | 20 ++++++++++++---- src/order/rel_iso.lean | 5 ++++ .../algebra/continuous_monoid_hom.lean | 8 +++++++ src/topology/bornology/hom.lean | 5 ++++ src/topology/continuous_function/bounded.lean | 5 ++++ .../continuous_function/cocompact_map.lean | 5 ++++ .../continuous_function/zero_at_infty.lean | 8 ++++++- src/topology/hom/open.lean | 5 ++++ src/topology/homotopy/basic.lean | 5 ++++ src/topology/order/hom/basic.lean | 5 ++++ src/topology/order/hom/esakia.lean | 14 ++++++++--- src/topology/spectral/hom.lean | 5 ++++ 19 files changed, 151 insertions(+), 21 deletions(-) diff --git a/src/algebra/order/hom/monoid.lean b/src/algebra/order/hom/monoid.lean index ad3d404337a62..eff8a3e4b97fc 100644 --- a/src/algebra/order/hom/monoid.lean +++ b/src/algebra/order/hom/monoid.lean @@ -70,14 +70,20 @@ structure order_add_monoid_hom (α β : Type*) [preorder α] [preorder β] [add_ infixr ` →+o `:25 := order_add_monoid_hom +section +set_option old_structure_cmd true + /-- `order_add_monoid_hom_class F α β` states that `F` is a type of ordered monoid homomorphisms. You should also extend this typeclass when you extend `order_add_monoid_hom`. -/ +@[ancestor add_monoid_hom_class] class order_add_monoid_hom_class (F : Type*) (α β : out_param $ Type*) [preorder α] [preorder β] [add_zero_class α] [add_zero_class β] extends add_monoid_hom_class F α β := (monotone (f : F) : monotone f) +end + -- Instances and lemmas are defined below through `@[to_additive]`. end add_monoid @@ -101,19 +107,25 @@ structure order_monoid_hom (α β : Type*) [preorder α] [preorder β] [mul_one_ infixr ` →*o `:25 := order_monoid_hom +section +set_option old_structure_cmd true + /-- `order_monoid_hom_class F α β` states that `F` is a type of ordered monoid homomorphisms. You should also extend this typeclass when you extend `order_monoid_hom`. -/ -@[to_additive] +@[ancestor monoid_hom_class, to_additive] class order_monoid_hom_class (F : Type*) (α β : out_param $ Type*) [preorder α] [preorder β] [mul_one_class α] [mul_one_class β] extends monoid_hom_class F α β := (monotone (f : F) : monotone f) +end + @[priority 100, to_additive] -- See note [lower instance priority] instance order_monoid_hom_class.to_order_hom_class [order_monoid_hom_class F α β] : order_hom_class F α β := -{ map_rel := order_monoid_hom_class.monotone } +{ map_rel := order_monoid_hom_class.monotone, + .. ‹order_monoid_hom_class F α β› } @[to_additive] instance [order_monoid_hom_class F α β] : has_coe_t F (α →*o β) := @@ -141,6 +153,9 @@ structure order_monoid_with_zero_hom (α β : Type*) [preorder α] [preorder β] infixr ` →*₀o `:25 := order_monoid_with_zero_hom +section +set_option old_structure_cmd true + /-- `order_monoid_with_zero_hom_class F α β` states that `F` is a type of ordered monoid with zero homomorphisms. @@ -150,6 +165,8 @@ class order_monoid_with_zero_hom_class (F : Type*) (α β : out_param $ Type*) extends monoid_with_zero_hom_class F α β := (monotone (f : F) : monotone f) +end + @[priority 100] -- See note [lower instance priority] instance order_monoid_with_zero_hom_class.to_order_monoid_hom_class [order_monoid_with_zero_hom_class F α β] : order_monoid_hom_class F α β := diff --git a/src/data/fun_like/embedding.lean b/src/data/fun_like/embedding.lean index ac3107e07fd46..f5a6e824a3232 100644 --- a/src/data/fun_like/embedding.lean +++ b/src/data/fun_like/embedding.lean @@ -59,12 +59,17 @@ the axioms of your new type of morphisms. Continuing the example above: ``` +section +set_option old_structure_cmd true + /-- `my_embedding_class F A B` states that `F` is a type of `my_class.op`-preserving embeddings. You should extend this class when you extend `my_embedding`. -/ class my_embedding_class (F : Type*) (A B : out_param $ Type*) [my_class A] [my_class B] extends embedding_like F A B := (map_op : ∀ (f : F) (x y : A), f (my_class.op x y) = my_class.op (f x) (f y)) +end + @[simp] lemma map_op {F A B : Type*} [my_class A] [my_class B] [my_embedding_class F A B] (f : F) (x y : A) : f (my_class.op x y) = my_class.op (f x) (f y) := my_embedding_class.map_op @@ -88,10 +93,15 @@ structure cooler_embedding (A B : Type*) [cool_class A] [cool_class B] extends my_embedding A B := (map_cool' : to_fun cool_class.cool = cool_class.cool) +section +set_option old_structure_cmd true + class cooler_embedding_class (F : Type*) (A B : out_param $ Type*) [cool_class A] [cool_class B] extends my_embedding_class F A B := (map_cool : ∀ (f : F), f cool_class.cool = cool_class.cool) +end + @[simp] lemma map_cool {F A B : Type*} [cool_class A] [cool_class B] [cooler_embedding_class F A B] (f : F) : f cool_class.cool = cool_class.cool := my_embedding_class.map_op @@ -120,6 +130,8 @@ instead of linearly increasing the work per `my_embedding`-related declaration. -/ +set_option old_structure_cmd true + /-- The class `embedding_like F α β` expresses that terms of type `F` have an injective coercion to injective functions `α ↪ β`. -/ diff --git a/src/data/fun_like/equiv.lean b/src/data/fun_like/equiv.lean index 7327b35a5416d..772d7a72cc9ad 100644 --- a/src/data/fun_like/equiv.lean +++ b/src/data/fun_like/equiv.lean @@ -151,7 +151,7 @@ lemma inv_injective : function.injective (equiv_like.inv : E → (β → α)) := @[priority 100] instance to_embedding_like : embedding_like E α β := -{ coe := coe, +{ coe := (coe : E → α → β), coe_injective' := λ e g h, coe_injective' e g h ((left_inv e).eq_right_inverse (h.symm ▸ right_inv g)), injective' := λ e, (left_inv e).injective } diff --git a/src/model_theory/elementary_maps.lean b/src/model_theory/elementary_maps.lean index ada6f2c3ce971..d62105d2b95d0 100644 --- a/src/model_theory/elementary_maps.lean +++ b/src/model_theory/elementary_maps.lean @@ -107,7 +107,8 @@ begin end instance embedding_like : embedding_like (M ↪ₑ[L] N) M N := -{ injective' := injective } +{ injective' := injective, + .. show fun_like (M ↪ₑ[L] N) M (λ _, N), from infer_instance } @[simp] lemma map_fun (φ : M ↪ₑ[L] N) {n : ℕ} (f : L.functions n) (x : fin n → M) : φ (fun_map f x) = fun_map f (φ ∘ x) := diff --git a/src/order/hom/basic.lean b/src/order/hom/basic.lean index e66c9933a9205..d27ae18c68fca 100644 --- a/src/order/hom/basic.lean +++ b/src/order/hom/basic.lean @@ -90,6 +90,9 @@ abbreviation order_iso (α β : Type*) [has_le α] [has_le β] := @rel_iso α β infix ` ≃o `:25 := order_iso +section +set_option old_structure_cmd true + /-- `order_hom_class F α b` asserts that `F` is a type of `≤`-preserving morphisms. -/ abbreviation order_hom_class (F : Type*) (α β : out_param Type*) [has_le α] [has_le β] := rel_hom_class F ((≤) : α → α → Prop) ((≤) : β → β → Prop) @@ -101,6 +104,8 @@ class order_iso_class (F : Type*) (α β : out_param Type*) [has_le α] [has_le extends equiv_like F α β := (map_le_map_iff (f : F) {a b : α} : f a ≤ f b ↔ a ≤ b) +end + export order_iso_class (map_le_map_iff) attribute [simp] map_le_map_iff diff --git a/src/order/hom/bounded.lean b/src/order/hom/bounded.lean index 87a1e61bd35f4..4e6c16a053b42 100644 --- a/src/order/hom/bounded.lean +++ b/src/order/hom/bounded.lean @@ -48,6 +48,9 @@ structure bounded_order_hom (α β : Type*) [preorder α] [preorder β] [bounded (map_top' : to_fun ⊤ = ⊤) (map_bot' : to_fun ⊥ = ⊥) +section +set_option old_structure_cmd true + /-- `top_hom_class F α β` states that `F` is a type of `⊤`-preserving morphisms. You should extend this class when you extend `top_hom`. -/ @@ -71,6 +74,8 @@ class bounded_order_hom_class (F : Type*) (α β : out_param $ Type*) [has_le α (map_top (f : F) : f ⊤ = ⊤) (map_bot (f : F) : f ⊥ = ⊥) +end + export top_hom_class (map_top) bot_hom_class (map_bot) attribute [simp] map_top map_bot @@ -91,19 +96,24 @@ instance bounded_order_hom_class.to_bot_hom_class [has_le α] [has_le β] instance order_iso_class.to_top_hom_class [has_le α] [order_top α] [partial_order β] [order_top β] [order_iso_class F α β] : top_hom_class F α β := -⟨λ f, top_le_iff.1 $ (map_inv_le_iff f).1 le_top⟩ +{ map_top := λ f, top_le_iff.1 $ (map_inv_le_iff f).1 le_top, + .. show order_hom_class F α β, from infer_instance } @[priority 100] -- See note [lower instance priority] instance order_iso_class.to_bot_hom_class [has_le α] [order_bot α] [partial_order β] [order_bot β] [order_iso_class F α β] : bot_hom_class F α β := -⟨λ f, le_bot_iff.1 $ (le_map_inv_iff f).1 bot_le⟩ +--⟨λ f, le_bot_iff.1 $ (le_map_inv_iff f).1 bot_le⟩ +{ map_bot := λ f, le_bot_iff.1 $ (le_map_inv_iff f).1 bot_le, + .. show order_hom_class F α β, from infer_instance } @[priority 100] -- See note [lower instance priority] instance order_iso_class.to_bounded_order_hom_class [has_le α] [bounded_order α] [partial_order β] [bounded_order β] [order_iso_class F α β] : bounded_order_hom_class F α β := -{ ..order_iso_class.to_top_hom_class, ..order_iso_class.to_bot_hom_class } +{ ..show order_hom_class F α β, from infer_instance, + ..order_iso_class.to_top_hom_class, + ..order_iso_class.to_bot_hom_class } @[simp] lemma map_eq_top_iff [has_le α] [order_top α] [partial_order β] [order_top β] [order_iso_class F α β] (f : F) {a : α} : f a = ⊤ ↔ a = ⊤ := diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index dc5e26e305dc7..eb8524db44056 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -62,6 +62,9 @@ structure complete_lattice_hom (α β : Type*) [complete_lattice α] [complete_l extends Inf_hom α β := (map_Sup' (s : set α) : to_fun (Sup s) = Sup (to_fun '' s)) +section +set_option old_structure_cmd true + /-- `Sup_hom_class F α β` states that `F` is a type of `⨆`-preserving morphisms. You should extend this class when you extend `Sup_hom`. -/ @@ -90,6 +93,8 @@ class complete_lattice_hom_class (F : Type*) (α β : out_param $ Type*) [comple [complete_lattice β] extends Inf_hom_class F α β := (map_Sup (f : F) (s : set α) : f (Sup s) = Sup (f '' s)) +end + export Sup_hom_class (map_Sup) export Inf_hom_class (map_Inf) @@ -116,14 +121,16 @@ instance Sup_hom_class.to_sup_bot_hom_class [complete_lattice α] [complete_latt [Sup_hom_class F α β] : sup_bot_hom_class F α β := { map_sup := λ f a b, by rw [←Sup_pair, map_Sup, set.image_pair, Sup_pair], - map_bot := λ f, by rw [←Sup_empty, map_Sup, set.image_empty, Sup_empty] } + map_bot := λ f, by rw [←Sup_empty, map_Sup, set.image_empty, Sup_empty], + ..‹Sup_hom_class F α β› } @[priority 100] -- See note [lower instance priority] instance Inf_hom_class.to_inf_top_hom_class [complete_lattice α] [complete_lattice β] [Inf_hom_class F α β] : inf_top_hom_class F α β := { map_inf := λ f a b, by rw [←Inf_pair, map_Inf, set.image_pair, Inf_pair], - map_top := λ f, by rw [←Inf_empty, map_Inf, set.image_empty, Inf_empty] } + map_top := λ f, by rw [←Inf_empty, map_Inf, set.image_empty, Inf_empty], + ..‹Inf_hom_class F α β› } @[priority 100] -- See note [lower instance priority] instance frame_hom_class.to_Sup_hom_class [complete_lattice α] [complete_lattice β] @@ -153,19 +160,25 @@ instance complete_lattice_hom_class.to_bounded_lattice_hom_class [complete_latti instance order_iso_class.to_Sup_hom_class [complete_lattice α] [complete_lattice β] [order_iso_class F α β] : Sup_hom_class F α β := -⟨λ f s, eq_of_forall_ge_iff $ λ c, by simp only [←le_map_inv_iff, Sup_le_iff, set.ball_image_iff]⟩ +{ map_Sup := λ f s, eq_of_forall_ge_iff $ + λ c, by simp only [←le_map_inv_iff, Sup_le_iff, set.ball_image_iff], + .. show order_hom_class F α β, from infer_instance } @[priority 100] -- See note [lower instance priority] instance order_iso_class.to_Inf_hom_class [complete_lattice α] [complete_lattice β] [order_iso_class F α β] : Inf_hom_class F α β := -⟨λ f s, eq_of_forall_le_iff $ λ c, by simp only [←map_inv_le_iff, le_Inf_iff, set.ball_image_iff]⟩ +{ map_Inf := λ f s, eq_of_forall_le_iff $ + λ c, by simp only [←map_inv_le_iff, le_Inf_iff, set.ball_image_iff], + .. show order_hom_class F α β, from infer_instance } @[priority 100] -- See note [lower instance priority] instance order_iso_class.to_complete_lattice_hom_class [complete_lattice α] [complete_lattice β] [order_iso_class F α β] : complete_lattice_hom_class F α β := -{ ..order_iso_class.to_Sup_hom_class, ..order_iso_class.to_lattice_hom_class } +{ ..order_iso_class.to_Sup_hom_class, + ..order_iso_class.to_lattice_hom_class, + .. show Inf_hom_class F α β, from infer_instance } instance [has_Sup α] [has_Sup β] [Sup_hom_class F α β] : has_coe_t F (Sup_hom α β) := ⟨λ f, ⟨f, map_Sup f⟩⟩ diff --git a/src/order/hom/lattice.lean b/src/order/hom/lattice.lean index 42fce255b6019..0e125b40b0f6b 100644 --- a/src/order/hom/lattice.lean +++ b/src/order/hom/lattice.lean @@ -73,6 +73,9 @@ structure bounded_lattice_hom (α β : Type*) [lattice α] [lattice β] [bounded (map_top' : to_fun ⊤ = ⊤) (map_bot' : to_fun ⊥ = ⊥) +section +set_option old_structure_cmd true + /-- `sup_hom_class F α β` states that `F` is a type of `⊔`-preserving morphisms. You should extend this class when you extend `sup_hom`. -/ @@ -117,6 +120,8 @@ class bounded_lattice_hom_class (F : Type*) (α β : out_param $ Type*) [lattice (map_top (f : F) : f ⊤ = ⊤) (map_bot (f : F) : f ⊥ = ⊥) +end + export sup_hom_class (map_sup) export inf_hom_class (map_inf) @@ -126,12 +131,14 @@ attribute [simp] map_top map_bot map_sup map_inf instance sup_hom_class.to_order_hom_class [semilattice_sup α] [semilattice_sup β] [sup_hom_class F α β] : order_hom_class F α β := -⟨λ f a b h, by rw [←sup_eq_right, ←map_sup, sup_eq_right.2 h]⟩ +{ map_rel := λ f a b h, by rw [←sup_eq_right, ←map_sup, sup_eq_right.2 h], + ..‹sup_hom_class F α β› } @[priority 100] -- See note [lower instance priority] instance inf_hom_class.to_order_hom_class [semilattice_inf α] [semilattice_inf β] [inf_hom_class F α β] : order_hom_class F α β := -⟨λ f a b h, by rw [←inf_eq_left, ←map_inf, inf_eq_left.2 h]⟩ +{ map_rel := λ f a b h, by rw [←inf_eq_left, ←map_inf, inf_eq_left.2 h] + ..‹inf_hom_class F α β› } @[priority 100] -- See note [lower instance priority] instance sup_bot_hom_class.to_bot_hom_class [has_sup α] [has_sup β] [has_bot α] [has_bot β] @@ -166,19 +173,22 @@ instance bounded_lattice_hom_class.to_inf_top_hom_class [lattice α] [lattice β instance bounded_lattice_hom_class.to_bounded_order_hom_class [lattice α] [lattice β] [bounded_order α] [bounded_order β] [bounded_lattice_hom_class F α β] : bounded_order_hom_class F α β := -{ .. ‹bounded_lattice_hom_class F α β› } +{ .. show order_hom_class F α β, from infer_instance, + .. ‹bounded_lattice_hom_class F α β› } @[priority 100] -- See note [lower instance priority] instance order_iso_class.to_sup_hom_class [semilattice_sup α] [semilattice_sup β] [order_iso_class F α β] : sup_hom_class F α β := -⟨λ f a b, eq_of_forall_ge_iff $ λ c, by simp only [←le_map_inv_iff, sup_le_iff]⟩ +{ map_sup := λ f a b, eq_of_forall_ge_iff $ λ c, by simp only [←le_map_inv_iff, sup_le_iff], + .. show order_hom_class F α β, from infer_instance } @[priority 100] -- See note [lower instance priority] instance order_iso_class.to_inf_hom_class [semilattice_inf α] [semilattice_inf β] [order_iso_class F α β] : inf_hom_class F α β := -⟨λ f a b, eq_of_forall_le_iff $ λ c, by simp only [←map_inv_le_iff, le_inf_iff]⟩ +{ map_inf := λ f a b, eq_of_forall_le_iff $ λ c, by simp only [←map_inv_le_iff, le_inf_iff], + .. show order_hom_class F α β, from infer_instance } @[priority 100] -- See note [lower instance priority] instance order_iso_class.to_sup_bot_hom_class [semilattice_sup α] [order_bot α] [semilattice_sup β] diff --git a/src/order/rel_iso.lean b/src/order/rel_iso.lean index aa48961a89d7f..fd73fc92939ac 100644 --- a/src/order/rel_iso.lean +++ b/src/order/rel_iso.lean @@ -48,6 +48,9 @@ structure rel_hom {α β : Type*} (r : α → α → Prop) (s : β → β → Pr infix ` →r `:25 := rel_hom +section +set_option old_structure_cmd true + /-- `rel_hom_class F r s` asserts that `F` is a type of functions such that all `f : F` satisfy `r a b → s (f a) (f b)`. @@ -63,6 +66,8 @@ export rel_hom_class (map_rel) -- The free parameters `r` and `s` are `out_param`s so this is not dangerous. attribute [nolint dangerous_instance] rel_hom_class.to_fun_like +end + namespace rel_hom_class variables {F : Type*} diff --git a/src/topology/algebra/continuous_monoid_hom.lean b/src/topology/algebra/continuous_monoid_hom.lean index e95e500b9095e..fd27981dc4649 100644 --- a/src/topology/algebra/continuous_monoid_hom.lean +++ b/src/topology/algebra/continuous_monoid_hom.lean @@ -46,6 +46,9 @@ When you extend this structure, make sure to extend `continuous_add_monoid_hom_c structure continuous_monoid_hom extends A →* B := (continuous_to_fun : continuous to_fun) +section +set_option old_structure_cmd true + /-- `continuous_add_monoid_hom_class F A B` states that `F` is a type of continuous additive monoid homomorphisms. @@ -62,6 +65,11 @@ You should also extend this typeclass when you extend `continuous_monoid_hom`. - class continuous_monoid_hom_class extends monoid_hom_class F A B := (map_continuous (f : F) : continuous f) +attribute [to_additive continuous_add_monoid_hom_class.to_add_monoid_hom_class] + continuous_monoid_hom_class.to_monoid_hom_class + +end + /-- Reinterpret a `continuous_monoid_hom` as a `monoid_hom`. -/ add_decl_doc continuous_monoid_hom.to_monoid_hom diff --git a/src/topology/bornology/hom.lean b/src/topology/bornology/hom.lean index aeca18629976d..8f2a7d7e84b1d 100644 --- a/src/topology/bornology/hom.lean +++ b/src/topology/bornology/hom.lean @@ -31,6 +31,9 @@ structure locally_bounded_map (α β : Type*) [bornology α] [bornology β] := (to_fun : α → β) (comap_cobounded_le' : (cobounded β).comap to_fun ≤ cobounded α) +section +set_option old_structure_cmd true + /-- `locally_bounded_map_class F α β` states that `F` is a type of bounded maps. You should extend this class when you extend `locally_bounded_map`. -/ @@ -39,6 +42,8 @@ class locally_bounded_map_class (F : Type*) (α β : out_param $ Type*) [bornolo extends fun_like F α (λ _, β) := (comap_cobounded_le (f : F) : (cobounded β).comap f ≤ cobounded α) +end + export locally_bounded_map_class (comap_cobounded_le) lemma is_bounded.image [bornology α] [bornology β] [locally_bounded_map_class F α β] {f : F} diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 30eaf3bf12450..fc932b9fb97c6 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -40,6 +40,9 @@ structure bounded_continuous_function (α : Type u) (β : Type v) localized "infixr (name := bounded_continuous_function) ` →ᵇ `:25 := bounded_continuous_function" in bounded_continuous_function +section +set_option old_structure_cmd true + /-- `bounded_continuous_map_class F α β` states that `F` is a type of bounded continuous maps. You should also extend this typeclass when you extend `bounded_continuous_function`. -/ @@ -47,6 +50,8 @@ class bounded_continuous_map_class (F α β : Type*) [topological_space α] [pse extends continuous_map_class F α β := (map_bounded (f : F) : ∃ C, ∀ x y, dist (f x) (f y) ≤ C) +end + export bounded_continuous_map_class (map_bounded) namespace bounded_continuous_function diff --git a/src/topology/continuous_function/cocompact_map.lean b/src/topology/continuous_function/cocompact_map.lean index 5ed716fb37a9b..557e382e16337 100644 --- a/src/topology/continuous_function/cocompact_map.lean +++ b/src/topology/continuous_function/cocompact_map.lean @@ -27,6 +27,9 @@ structure cocompact_map (α : Type u) (β : Type v) [topological_space α] [topo extends continuous_map α β : Type (max u v) := (cocompact_tendsto' : tendsto to_fun (cocompact α) (cocompact β)) +section +set_option old_structure_cmd true + /-- `cocompact_map_class F α β` states that `F` is a type of cocompact continuous maps. You should also extend this typeclass when you extend `cocompact_map`. -/ @@ -34,6 +37,8 @@ class cocompact_map_class (F : Type*) (α β : out_param $ Type*) [topological_s [topological_space β] extends continuous_map_class F α β := (cocompact_tendsto (f : F) : tendsto f (cocompact α) (cocompact β)) +end + namespace cocompact_map_class variables {F α β : Type*} [topological_space α] [topological_space β] diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index c41cd5e5196c4..9911e7e6e6c80 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -44,6 +44,9 @@ localized "notation [priority 2000] (name := zero_at_infty_continuous_map) localized "notation (name := zero_at_infty_continuous_map.arrow) α ` →C₀ ` β := zero_at_infty_continuous_map α β" in zero_at_infty +section +set_option old_structure_cmd true + /-- `zero_at_infty_continuous_map_class F α β` states that `F` is a type of continuous maps which vanish at infinity. @@ -52,6 +55,8 @@ class zero_at_infty_continuous_map_class (F : Type*) (α β : out_param $ Type*) [has_zero β] [topological_space β] extends continuous_map_class F α β := (zero_at_infty (f : F) : tendsto f (cocompact α) (𝓝 0)) +end + export zero_at_infty_continuous_map_class (zero_at_infty) namespace zero_at_infty_continuous_map @@ -315,7 +320,8 @@ f.bounded_range.mono $ image_subset_range _ _ @[priority 100] instance : bounded_continuous_map_class F α β := -{ map_bounded := λ f, zero_at_infty_continuous_map.bounded f } +{ map_bounded := λ f, zero_at_infty_continuous_map.bounded f, + ..‹zero_at_infty_continuous_map_class F α β› } /-- Construct a bounded continuous function from a continuous function vanishing at infinity. -/ @[simps] diff --git a/src/topology/hom/open.lean b/src/topology/hom/open.lean index 67ad4a5d027fd..c6097171f8860 100644 --- a/src/topology/hom/open.lean +++ b/src/topology/hom/open.lean @@ -33,6 +33,9 @@ structure continuous_open_map (α β : Type*) [topological_space α] [topologica infixr ` →CO `:25 := continuous_open_map +section +set_option old_structure_cmd true + /-- `continuous_open_map_class F α β` states that `F` is a type of continuous open maps. You should extend this class when you extend `continuous_open_map`. -/ @@ -40,6 +43,8 @@ class continuous_open_map_class (F : Type*) (α β : out_param $ Type*) [topolog [topological_space β] extends continuous_map_class F α β := (map_open (f : F) : is_open_map f) +end + export continuous_open_map_class (map_open) instance [topological_space α] [topological_space β] [continuous_open_map_class F α β] : diff --git a/src/topology/homotopy/basic.lean b/src/topology/homotopy/basic.lean index ed41f986392f3..7deeee09350bc 100644 --- a/src/topology/homotopy/basic.lean +++ b/src/topology/homotopy/basic.lean @@ -74,6 +74,9 @@ structure homotopy (f₀ f₁ : C(X, Y)) extends C(I × X, Y) := (map_zero_left' : ∀ x, to_fun (0, x) = f₀ x) (map_one_left' : ∀ x, to_fun (1, x) = f₁ x) +section +set_option old_structure_cmd true + /-- `continuous_map.homotopy_like F f₀ f₁` states that `F` is a type of homotopies between `f₀` and `f₁`. @@ -83,6 +86,8 @@ class homotopy_like (F : Type*) (f₀ f₁ : out_param $ C(X, Y)) (map_zero_left (f : F) : ∀ x, f (0, x) = f₀ x) (map_one_left (f : F) : ∀ x, f (1, x) = f₁ x) +end + -- `f₀` and `f₁` are `out_param` so this is not dangerous attribute [nolint dangerous_instance] homotopy_like.to_continuous_map_class diff --git a/src/topology/order/hom/basic.lean b/src/topology/order/hom/basic.lean index 885b93478fa7d..7d4c283e88f99 100644 --- a/src/topology/order/hom/basic.lean +++ b/src/topology/order/hom/basic.lean @@ -37,6 +37,9 @@ structure continuous_order_hom (α β : Type*) [preorder α] [preorder β] [topo infixr ` →Co `:25 := continuous_order_hom +section +set_option old_structure_cmd true + /-- `continuous_order_hom_class F α β` states that `F` is a type of continuous monotone maps. You should extend this class when you extend `continuous_order_hom`. -/ @@ -45,6 +48,8 @@ class continuous_order_hom_class (F : Type*) (α β : out_param $ Type*) [preord extends rel_hom_class F ((≤) : α → α → Prop) ((≤) : β → β → Prop) := (map_continuous (f : F) : continuous f) +end + @[priority 100] -- See note [lower instance priority] instance continuous_order_hom_class.to_continuous_map_class [preorder α] [preorder β] [topological_space α] [topological_space β] [continuous_order_hom_class F α β] : diff --git a/src/topology/order/hom/esakia.lean b/src/topology/order/hom/esakia.lean index 067ada35f1621..930950984f105 100644 --- a/src/topology/order/hom/esakia.lean +++ b/src/topology/order/hom/esakia.lean @@ -44,6 +44,9 @@ structure esakia_hom (α β : Type*) [topological_space α] [preorder α] [topol [preorder β] extends α →Co β := (exists_map_eq_of_map_le' ⦃a : α⦄ ⦃b : β⦄ : to_fun a ≤ b → ∃ c, a ≤ c ∧ to_fun c = b) +section +set_option old_structure_cmd true + /-- `pseudo_epimorphism_class F α β` states that `F` is a type of `⊔`-preserving morphisms. You should extend this class when you extend `pseudo_epimorphism`. -/ @@ -59,18 +62,23 @@ class esakia_hom_class (F : Type*) (α β : out_param $ Type*) [topological_spac extends continuous_order_hom_class F α β := (exists_map_eq_of_map_le (f : F) ⦃a : α⦄ ⦃b : β⦄ : f a ≤ b → ∃ c, a ≤ c ∧ f c = b) +end + export pseudo_epimorphism_class (exists_map_eq_of_map_le) @[priority 100] -- See note [lower instance priority] instance pseudo_epimorphism_class.to_top_hom_class [partial_order α] [order_top α] [preorder β] [order_top β] [pseudo_epimorphism_class F α β] : top_hom_class F α β := -⟨λ f, let ⟨b, h⟩ := exists_map_eq_of_map_le f (@le_top _ _ _ $ f ⊤) in - by rw [←top_le_iff.1 h.1, h.2]⟩ +{ map_top := λ f, let ⟨b, h⟩ := exists_map_eq_of_map_le f (@le_top _ _ _ $ f ⊤) in + by rw [←top_le_iff.1 h.1, h.2] + .. ‹pseudo_epimorphism_class F α β› } @[priority 100] -- See note [lower instance priority] instance order_iso_class.to_pseudo_epimorphism_class [preorder α] [preorder β] [order_iso_class F α β] : pseudo_epimorphism_class F α β := -⟨λ f a b h, ⟨equiv_like.inv f b, (le_map_inv_iff f).2 h, equiv_like.right_inv _ _⟩⟩ +{ exists_map_eq_of_map_le := + λ f a b h, ⟨equiv_like.inv f b, (le_map_inv_iff f).2 h, equiv_like.right_inv _ _⟩, + .. order_iso_class.to_order_hom_class } @[priority 100] -- See note [lower instance priority] instance esakia_hom_class.to_pseudo_epimorphism_class [topological_space α] [preorder α] diff --git a/src/topology/spectral/hom.lean b/src/topology/spectral/hom.lean index 9a939f0bc670c..0b44d6ad54dae 100644 --- a/src/topology/spectral/hom.lean +++ b/src/topology/spectral/hom.lean @@ -56,6 +56,9 @@ structure spectral_map (α β : Type*) [topological_space α] [topological_space (to_fun : α → β) (spectral' : is_spectral_map to_fun) +section +set_option old_structure_cmd true + /-- `spectral_map_class F α β` states that `F` is a type of spectral maps. You should extend this class when you extend `spectral_map`. -/ @@ -64,6 +67,8 @@ class spectral_map_class (F : Type*) (α β : out_param $ Type*) [topological_sp extends fun_like F α (λ _, β) := (map_spectral (f : F) : is_spectral_map f) +end + export spectral_map_class (map_spectral) attribute [simp] map_spectral From 0fecc32c1f68d73bf34b699b4f67425b396688f8 Mon Sep 17 00:00:00 2001 From: mkaratarakis Date: Sat, 15 Oct 2022 15:59:19 +0000 Subject: [PATCH 782/828] feat(algebra/group_ring_action): add `mul_semiring_action.comp_hom` (#16850) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A multiplicative action of a monoid `M` on a semiring `R` and a monoid homomorphism `N →* M` induce a multiplicative action of a monoid `N` on `R`. I need this for my study of the decomposition and inertia groups Co-authored-by: mkaratarakis <40603357+mkaratarakis@users.noreply.github.com> --- src/algebra/group_ring_action.lean | 15 ++++++++++++++- 1 file changed, 14 insertions(+), 1 deletion(-) diff --git a/src/algebra/group_ring_action.lean b/src/algebra/group_ring_action.lean index 520c11f19a127..ff741043f6a74 100644 --- a/src/algebra/group_ring_action.lean +++ b/src/algebra/group_ring_action.lean @@ -40,7 +40,7 @@ class mul_semiring_action (M : Type u) (R : Type v) [monoid M] [semiring R] section semiring -variables (M G : Type u) [monoid M] [group G] +variables (M N G : Type*) [monoid M] [monoid N] [group G] variables (A R S F : Type v) [add_monoid A] [semiring R] [comm_semiring S] [division_ring F] -- note we could not use `extends` since these typeclasses are made with `old_structure_cmd` @@ -65,6 +65,19 @@ def mul_semiring_action.to_ring_equiv [mul_semiring_action G R] (x : G) : R ≃+ { .. distrib_mul_action.to_add_equiv R x, .. mul_semiring_action.to_ring_hom G R x } +section +variables {M N} + +/-- Compose a `mul_semiring_action` with a `monoid_hom`, with action `f r' • m`. +See note [reducible non-instances]. -/ +@[reducible] def mul_semiring_action.comp_hom (f : N →* M) [mul_semiring_action M R] : + mul_semiring_action N R := +{ smul := has_smul.comp.smul f, + ..distrib_mul_action.comp_hom R f, + ..mul_distrib_mul_action.comp_hom R f } + +end + section variables {M G R} From 5ea225ecd1dbd0dda5f60a9da3cd295e7f15a42d Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 15 Oct 2022 19:03:11 +0000 Subject: [PATCH 783/828] feat(data/set/function): add `set.eq_on_union` (#16984) Also add `set.eq_on.union` and turn `set.comp_eq_of_eq_on_range` into an `iff` lemma `set.eq_on_range`. --- src/data/set/function.lean | 16 ++++++++++++---- 1 file changed, 12 insertions(+), 4 deletions(-) diff --git a/src/data/set/function.lean b/src/data/set/function.lean index 6e7b7284f964e..a2d9c62abda54 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -166,11 +166,19 @@ ext $ λ x, and.congr_right_iff.2 $ λ hx, by rw [mem_preimage, mem_preimage, he lemma eq_on.mono (hs : s₁ ⊆ s₂) (hf : eq_on f₁ f₂ s₂) : eq_on f₁ f₂ s₁ := λ x hx, hf (hs hx) +@[simp] lemma eq_on_union : eq_on f₁ f₂ (s₁ ∪ s₂) ↔ eq_on f₁ f₂ s₁ ∧ eq_on f₁ f₂ s₂ := +ball_or_left_distrib + +lemma eq_on.union (h₁ : eq_on f₁ f₂ s₁) (h₂ : eq_on f₁ f₂ s₂) : eq_on f₁ f₂ (s₁ ∪ s₂) := +eq_on_union.2 ⟨h₁, h₂⟩ + lemma eq_on.comp_left (h : s.eq_on f₁ f₂) : s.eq_on (g ∘ f₁) (g ∘ f₂) := λ a ha, congr_arg _ $ h ha -lemma comp_eq_of_eq_on_range {ι : Sort*} {f : ι → α} {g₁ g₂ : α → β} (h : eq_on g₁ g₂ (range f)) : - g₁ ∘ f = g₂ ∘ f := -funext $ λ x, h $ mem_range_self _ +@[simp] lemma eq_on_range {ι : Sort*} {f : ι → α} {g₁ g₂ : α → β} : + eq_on g₁ g₂ (range f) ↔ g₁ ∘ f = g₂ ∘ f := +forall_range_iff.trans $ funext_iff.symm + +alias eq_on_range ↔ eq_on.comp_eq _ /-! ### Congruence lemmas -/ @@ -987,7 +995,7 @@ funext $ λ x, if hx : x ∈ s then by simp [hx] else by simp [hx] @[simp] lemma piecewise_range_comp {ι : Sort*} (f : ι → α) [Π j, decidable (j ∈ range f)] (g₁ g₂ : α → β) : (range f).piecewise g₁ g₂ ∘ f = g₁ ∘ f := -comp_eq_of_eq_on_range $ piecewise_eq_on _ _ _ +eq_on.comp_eq $ piecewise_eq_on _ _ _ theorem maps_to.piecewise_ite {s s₁ s₂ : set α} {t t₁ t₂ : set β} {f₁ f₂ : α → β} [∀ i, decidable (i ∈ s)] From 4d8fe495f692830e951e4aa409e614ba48739f39 Mon Sep 17 00:00:00 2001 From: Patrick Massot Date: Sat, 15 Oct 2022 19:03:12 +0000 Subject: [PATCH 784/828] docs(*): Missing undergrad and overview items (#17003) --- docs/overview.yaml | 3 +++ docs/undergrad.yaml | 4 ++-- 2 files changed, 5 insertions(+), 2 deletions(-) diff --git a/docs/overview.yaml b/docs/overview.yaml index fae0637b299f0..6e0f6640e4cce 100644 --- a/docs/overview.yaml +++ b/docs/overview.yaml @@ -339,6 +339,8 @@ Analysis: Cauchy integral formula: 'diff_cont_on_cl.circle_integral_sub_inv_smul' Liouville theorem: 'differentiable.apply_eq_apply_of_bounded' maximum modulus principle: 'complex.eventually_eq_of_is_local_max_norm' + principle of isolated zeros: 'analytic_at.eventually_eq_zero_or_eventually_ne_zero' + principle of analytic continuation: 'analytic_on.eq_on_of_preconnected_of_frequently_eq' analyticity of holomorphic functions: 'differentiable_on.analytic_at' Schwarz lemma: 'complex.abs_le_abs_of_maps_to_ball_self' removable singularity: 'complex.differentiable_on_update_lim_insert_of_is_o' @@ -360,6 +362,7 @@ Probability Theory: probability density function: 'measure_theory.has_pdf' variance of a real-valued random variable: 'probability_theory.variance' independence of random variables: 'probability_theory.Indep_fun' + Kolmogorov's $0$-$1$ law: 'probability_theory.measure_zero_or_one_of_measurable_set_limsup_at_top' mean of product of independent random variables: 'probability_theory.indep_fun.integral_mul_of_integrable' moment of a random variable: 'probability_theory.moment' Bernoulli law: 'pmf.bernoulli' diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index 037346590d914..9c0f4d78a13bc 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -379,8 +379,8 @@ Single Variable Complex Analysis: winding number of a closed curve in $\C$ with respect to a point: '' Cauchy formulas: '' analyticity of a holomorphic function: 'differentiable_on.analytic_at' - principle of isolated zeros: '' - principle of analytic continuation: '' + principle of isolated zeros: 'analytic_at.eventually_eq_zero_or_eventually_ne_zero' + principle of analytic continuation: 'analytic_on.eq_on_of_preconnected_of_frequently_eq' maximum principle: 'complex.eventually_eq_of_is_local_max_norm' isolated singularities: '' Laurent series: '' From 71a759b2a0139506ff209da2536198395aa67ce3 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Sun, 16 Oct 2022 01:04:45 +0000 Subject: [PATCH 785/828] feat(group_theory/subgroup/basic): Range and kernel of `monoid_hom.restrict` (#16988) This PRs adds simp-lemmas for the range and kernel of `monoid_hom.restrict`. --- src/group_theory/subgroup/basic.lean | 7 +++++++ src/group_theory/submonoid/operations.lean | 8 ++++++++ 2 files changed, 15 insertions(+) diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index f3f9d7db5d446..100188847a13d 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -2051,6 +2051,10 @@ iff.rfl @[to_additive] lemma range_eq_map (f : G →* N) : f.range = (⊤ : subgroup G).map f := by ext; simp +@[simp, to_additive] lemma restrict_range (f : G →* N) : (f.restrict K).range = K.map f := +by simp_rw [set_like.ext_iff, mem_range, mem_map, restrict_apply, set_like.exists, subtype.coe_mk, + iff_self, forall_const] + /-- The canonical surjective group homomorphism `G →* f(G)` induced by a group homomorphism `G →* N`. -/ @[to_additive "The canonical surjective `add_group` homomorphism `G →+ f(G)` induced by a group @@ -2177,6 +2181,9 @@ lemma comap_ker (g : N →* P) (f : G →* N) : g.ker.comap f = (g.comp f).ker : @[simp, to_additive] lemma comap_bot (f : G →* N) : (⊥ : subgroup N).comap f = f.ker := rfl +@[simp, to_additive] lemma restrict_ker (f : G →* N) : (f.restrict K).ker = f.ker.subgroup_of K := +rfl + @[to_additive] lemma range_restrict_ker (f : G →* N) : ker (range_restrict f) = ker f := begin ext, diff --git a/src/group_theory/submonoid/operations.lean b/src/group_theory/submonoid/operations.lean index 2aa27c94f47dc..9646c9c246b50 100644 --- a/src/group_theory/submonoid/operations.lean +++ b/src/group_theory/submonoid/operations.lean @@ -831,6 +831,10 @@ lemma restrict_apply {N S : Type*} [mul_one_class N] [set_like S M] [submonoid_c (f : M →* N) (s : S) (x : s) : f.restrict s x = f x := rfl +@[simp, to_additive] lemma restrict_mrange (f : M →* N) : (f.restrict S).mrange = S.map f := +by simp_rw [set_like.ext_iff, mem_mrange, mem_map, restrict_apply, set_like.exists, subtype.coe_mk, + iff_self, forall_const] + /-- Restriction of a monoid hom to a submonoid of the codomain. -/ @[to_additive "Restriction of an `add_monoid` hom to an `add_submonoid` of the codomain.", simps apply] @@ -881,6 +885,10 @@ include mc (⊥ : submonoid N).comap f = mker f := rfl omit mc +@[simp, to_additive] +lemma restrict_mker (f : M →* N) : (f.restrict S).mker = f.mker.comap S.subtype := +rfl + @[to_additive] lemma range_restrict_mker (f : M →* N) : mker (mrange_restrict f) = mker f := begin ext, From 6568f2f2b41d6c6a92d6dfe762222743fb198953 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sun, 16 Oct 2022 05:42:13 +0000 Subject: [PATCH 786/828] chore(scripts): update nolints.txt (#17008) I am happy to remove some nolints for you! --- scripts/nolints.txt | 5 ----- 1 file changed, 5 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 80d18571f1557..67cd9ddfb58c3 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -137,11 +137,6 @@ apply_nolint tactic.interactive.traverse_constructor unused_arguments apply_nolint tactic.interactive.traverse_field unused_arguments apply_nolint tactic.interactive.with_prefix doc_blame --- data/finset/noncomm_prod.lean -apply_nolint add_monoid_hom.pi_ext fintype_finite -apply_nolint finset.noncomm_prod_union_of_disjoint to_additive_doc -apply_nolint monoid_hom.pi_ext fintype_finite - -- data/finsupp/pwo.lean apply_nolint finsupp.is_pwo fintype_finite From 4698e35ca56a0d4fa53aa5639c3364e0a77f4eba Mon Sep 17 00:00:00 2001 From: Oleksandr Manzyuk Date: Sun, 16 Oct 2022 11:47:35 +0000 Subject: [PATCH 787/828] feat(category_theory/monoidal): define the bicategory of algebras and bimodules internal to some monoidal category (#14402) We generalise the [well-known construction](https://ncatlab.org/nlab/show/bimodule#AsMorphismsInA2Category) of the 2-category of rings, bimodules, and intertwiners. Given a monoidal category `C` that has coequalizers and in which left and right tensor products preserve colimits, we define a bicategory whose - objects are monoids in C; - 1-morphisms are bimodules; - 2-morphisms are bimodule homomorphisms. The composition of 1-morphisms is given by the tensor product of bimodules over the middle monoid. Co-authored-by: Scott Morrison --- .../limits/preserves/shapes/equalizers.lean | 41 + .../limits/shapes/equalizers.lean | 7 + src/category_theory/monoidal/Bimod.lean | 1017 +++++++++++++++++ 3 files changed, 1065 insertions(+) create mode 100644 src/category_theory/monoidal/Bimod.lean diff --git a/src/category_theory/limits/preserves/shapes/equalizers.lean b/src/category_theory/limits/preserves/shapes/equalizers.lean index a02122d528bfb..8aaee3c376d9d 100644 --- a/src/category_theory/limits/preserves/shapes/equalizers.lean +++ b/src/category_theory/limits/preserves/shapes/equalizers.lean @@ -183,6 +183,47 @@ begin apply_instance end +instance map_π_epi : epi (G.map (coequalizer.π f g)) := +⟨λ W h k, by { rw ←ι_comp_coequalizer_comparison, apply (cancel_epi _).1, apply epi_comp }⟩ + +@[reassoc] +lemma map_π_preserves_coequalizer_inv : + G.map (coequalizer.π f g) ≫ (preserves_coequalizer.iso G f g).inv = + coequalizer.π (G.map f) (G.map g) := +begin + rw [←ι_comp_coequalizer_comparison_assoc, ←preserves_coequalizer.iso_hom, + iso.hom_inv_id, comp_id], +end + +@[reassoc] +lemma map_π_preserves_coequalizer_inv_desc + {W : D} (k : G.obj Y ⟶ W) (wk : G.map f ≫ k = G.map g ≫ k) : + G.map (coequalizer.π f g) ≫ (preserves_coequalizer.iso G f g).inv ≫ coequalizer.desc k wk = k := +by rw [←category.assoc, map_π_preserves_coequalizer_inv, coequalizer.π_desc] + +@[reassoc] +lemma map_π_preserves_coequalizer_inv_colim_map + {X' Y' : D} (f' g' : X' ⟶ Y') [has_coequalizer f' g'] (p : G.obj X ⟶ X') (q : G.obj Y ⟶ Y') + (wf : (G.map f) ≫ q = p ≫ f') (wg : (G.map g) ≫ q = p ≫ g') : + G.map (coequalizer.π f g) ≫ (preserves_coequalizer.iso G f g).inv ≫ + colim_map (parallel_pair_hom (G.map f) (G.map g) f' g' p q wf wg) = + q ≫ coequalizer.π f' g' := +by rw [←category.assoc, map_π_preserves_coequalizer_inv, ι_colim_map, parallel_pair_hom_app_one] + +@[reassoc] +lemma map_π_preserves_coequalizer_inv_colim_map_desc + {X' Y' : D} (f' g' : X' ⟶ Y') [has_coequalizer f' g'] (p : G.obj X ⟶ X') (q : G.obj Y ⟶ Y') + (wf : (G.map f) ≫ q = p ≫ f') (wg : (G.map g) ≫ q = p ≫ g') + {Z' : D} (h : Y' ⟶ Z') (wh : f' ≫ h = g' ≫ h) : + G.map (coequalizer.π f g) ≫ (preserves_coequalizer.iso G f g).inv ≫ + colim_map (parallel_pair_hom (G.map f) (G.map g) f' g' p q wf wg) ≫ + coequalizer.desc h wh = + q ≫ h := +begin + slice_lhs 1 3 { rw map_π_preserves_coequalizer_inv_colim_map }, + slice_lhs 2 3 { rw coequalizer.π_desc }, +end + /-- Any functor preserves coequalizers of split pairs. -/ @[priority 1] instance preserves_split_coequalizers (f g : X ⟶ Y) [has_split_coequalizer f g] : diff --git a/src/category_theory/limits/shapes/equalizers.lean b/src/category_theory/limits/shapes/equalizers.lean index 0e534c0271773..750dc84ea4f73 100644 --- a/src/category_theory/limits/shapes/equalizers.lean +++ b/src/category_theory/limits/shapes/equalizers.lean @@ -751,6 +751,13 @@ lemma coequalizer.π_desc {W : C} (k : Y ⟶ W) (h : f ≫ k = g ≫ k) : coequalizer.π f g ≫ coequalizer.desc k h = k := colimit.ι_desc _ _ +lemma coequalizer.π_colim_map_desc {X' Y' Z : C} (f' g' : X' ⟶ Y') [has_coequalizer f' g'] + (p : X ⟶ X') (q : Y ⟶ Y') (wf : f ≫ q = p ≫ f') (wg : g ≫ q = p ≫ g') + (h : Y' ⟶ Z) (wh : f' ≫ h = g' ≫ h) : + coequalizer.π f g ≫ colim_map (parallel_pair_hom f g f' g' p q wf wg) ≫ coequalizer.desc h wh = + q ≫ h := +by rw [ι_colim_map_assoc, parallel_pair_hom_app_one, coequalizer.π_desc] + /-- Any morphism `k : Y ⟶ W` satisfying `f ≫ k = g ≫ k` induces a morphism `l : coequalizer f g ⟶ W` satisfying `coequalizer.π ≫ g = l`. -/ def coequalizer.desc' {W : C} (k : Y ⟶ W) (h : f ≫ k = g ≫ k) : diff --git a/src/category_theory/monoidal/Bimod.lean b/src/category_theory/monoidal/Bimod.lean new file mode 100644 index 0000000000000..7ff041dcfd99e --- /dev/null +++ b/src/category_theory/monoidal/Bimod.lean @@ -0,0 +1,1017 @@ +/- +Copyright (c) 2022 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison, Oleksandr Manzyuk +-/ +import category_theory.bicategory.basic +import category_theory.monoidal.Mon_ +import category_theory.limits.preserves.shapes.equalizers + +/-! +# The category of bimodule objects over a pair of monoid objects. +-/ + +universes v₁ v₂ u₁ u₂ + +open category_theory +open category_theory.monoidal_category + +variables {C : Type u₁} [category.{v₁} C] [monoidal_category.{v₁} C] + +section + +open category_theory.limits + +variables [has_coequalizers C] + +section + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_left X)] + +lemma id_tensor_π_preserves_coequalizer_inv_desc + {W X Y Z : C} (f g : X ⟶ Y) + (h : Z ⊗ Y ⟶ W) (wh : (𝟙 Z ⊗ f) ≫ h = (𝟙 Z ⊗ g) ≫ h) : + (𝟙 Z ⊗ coequalizer.π f g) ≫ (preserves_coequalizer.iso (tensor_left Z) f g).inv ≫ + coequalizer.desc h wh = h := +map_π_preserves_coequalizer_inv_desc (tensor_left Z) f g h wh + +lemma id_tensor_π_preserves_coequalizer_inv_colim_map_desc + {X Y Z X' Y' Z' : C} (f g : X ⟶ Y) (f' g' : X' ⟶ Y') (p : Z ⊗ X ⟶ X') (q : Z ⊗ Y ⟶ Y') + (wf : (𝟙 Z ⊗ f) ≫ q = p ≫ f') (wg : (𝟙 Z ⊗ g) ≫ q = p ≫ g') + (h : Y' ⟶ Z') (wh : f' ≫ h = g' ≫ h) : + (𝟙 Z ⊗ coequalizer.π f g) ≫ (preserves_coequalizer.iso (tensor_left Z) f g).inv ≫ + colim_map (parallel_pair_hom (𝟙 Z ⊗ f) (𝟙 Z ⊗ g) f' g' p q wf wg) ≫ + coequalizer.desc h wh = + q ≫ h := +map_π_preserves_coequalizer_inv_colim_map_desc (tensor_left Z) f g f' g' p q wf wg h wh + +end + +section + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_right X)] + +lemma π_tensor_id_preserves_coequalizer_inv_desc + {W X Y Z : C} (f g : X ⟶ Y) + (h : Y ⊗ Z ⟶ W) (wh : (f ⊗ 𝟙 Z) ≫ h = (g ⊗ 𝟙 Z) ≫ h) : + (coequalizer.π f g ⊗ 𝟙 Z) ≫ (preserves_coequalizer.iso (tensor_right Z) f g).inv ≫ + coequalizer.desc h wh = h := +map_π_preserves_coequalizer_inv_desc (tensor_right Z) f g h wh + +lemma π_tensor_id_preserves_coequalizer_inv_colim_map_desc + {X Y Z X' Y' Z' : C} (f g : X ⟶ Y) (f' g' : X' ⟶ Y') (p : X ⊗ Z ⟶ X') (q : Y ⊗ Z ⟶ Y') + (wf : (f ⊗ 𝟙 Z) ≫ q = p ≫ f') (wg : (g ⊗ 𝟙 Z) ≫ q = p ≫ g') + (h : Y' ⟶ Z') (wh : f' ≫ h = g' ≫ h) : + (coequalizer.π f g ⊗ 𝟙 Z) ≫ (preserves_coequalizer.iso (tensor_right Z) f g).inv ≫ + colim_map (parallel_pair_hom (f ⊗ 𝟙 Z) (g ⊗ 𝟙 Z) f' g' p q wf wg) ≫ + coequalizer.desc h wh = + q ≫ h := +map_π_preserves_coequalizer_inv_colim_map_desc (tensor_right Z) f g f' g' p q wf wg h wh + +end + +end + +/-- A bimodule object for a pair of monoid objects, all internal to some monoidal category. -/ +structure Bimod (A B : Mon_ C) := +(X : C) +(act_left : A.X ⊗ X ⟶ X) +(one_act_left' : (A.one ⊗ 𝟙 X) ≫ act_left = (λ_ X).hom . obviously) +(left_assoc' : + (A.mul ⊗ 𝟙 X) ≫ act_left = (α_ A.X A.X X).hom ≫ (𝟙 A.X ⊗ act_left) ≫ act_left . obviously) +(act_right : X ⊗ B.X ⟶ X) +(act_right_one' : (𝟙 X ⊗ B.one) ≫ act_right = (ρ_ X).hom . obviously) +(right_assoc' : + (𝟙 X ⊗ B.mul) ≫ act_right = (α_ X B.X B.X).inv ≫ (act_right ⊗ 𝟙 B.X) ≫ act_right . obviously) +(middle_assoc' : + (act_left ⊗ 𝟙 B.X) ≫ act_right = (α_ A.X X B.X).hom ≫ (𝟙 A.X ⊗ act_right) ≫ act_left . obviously) + +restate_axiom Bimod.one_act_left' +restate_axiom Bimod.act_right_one' +restate_axiom Bimod.left_assoc' +restate_axiom Bimod.right_assoc' +restate_axiom Bimod.middle_assoc' +attribute [simp, reassoc] +Bimod.one_act_left Bimod.act_right_one Bimod.left_assoc Bimod.right_assoc Bimod.middle_assoc + +namespace Bimod + +variables {A B : Mon_ C} (M : Bimod A B) + +/-- A morphism of bimodule objects. -/ +@[ext] +structure hom (M N : Bimod A B) := +(hom : M.X ⟶ N.X) +(left_act_hom' : M.act_left ≫ hom = (𝟙 A.X ⊗ hom) ≫ N.act_left . obviously) +(right_act_hom' : M.act_right ≫ hom = (hom ⊗ 𝟙 B.X) ≫ N.act_right . obviously) + +restate_axiom hom.left_act_hom' +restate_axiom hom.right_act_hom' +attribute [simp, reassoc] hom.left_act_hom hom.right_act_hom + +/-- The identity morphism on a bimodule object. -/ +@[simps] +def id' (M : Bimod A B) : hom M M := +{ hom := 𝟙 M.X, } + +instance hom_inhabited (M : Bimod A B) : inhabited (hom M M) := ⟨id' M⟩ + +/-- Composition of bimodule object morphisms. -/ +@[simps] +def comp {M N O : Bimod A B} (f : hom M N) (g : hom N O) : hom M O := +{ hom := f.hom ≫ g.hom, } + +instance : category (Bimod A B) := +{ hom := λ M N, hom M N, + id := id', + comp := λ M N O f g, comp f g, } + +@[simp] lemma id_hom' (M : Bimod A B) : (𝟙 M : hom M M).hom = 𝟙 M.X := rfl +@[simp] lemma comp_hom' {M N K : Bimod A B} (f : M ⟶ N) (g : N ⟶ K) : + (f ≫ g : hom M K).hom = f.hom ≫ g.hom := rfl + +/-- +Construct an isomorphism of bimodules by giving an isomorphism between the underlying objects +and checking compatibility with left and right actions only in the forward direction. +-/ +@[simps] +def iso_of_iso {X Y : Mon_ C} {P Q : Bimod X Y} + (f : P.X ≅ Q.X) + (f_left_act_hom : P.act_left ≫ f.hom = (𝟙 X.X ⊗ f.hom) ≫ Q.act_left) + (f_right_act_hom : P.act_right ≫ f.hom = (f.hom ⊗ 𝟙 Y.X) ≫ Q.act_right) : + P ≅ Q := +{ hom := ⟨f.hom⟩, + inv := + { hom := f.inv, + left_act_hom' := begin + rw [←(cancel_mono f.hom), category.assoc, category.assoc, iso.inv_hom_id, category.comp_id, + f_left_act_hom, ←category.assoc, ←id_tensor_comp, iso.inv_hom_id, + monoidal_category.tensor_id, category.id_comp], + end, + right_act_hom' := begin + rw [←(cancel_mono f.hom), category.assoc, category.assoc, iso.inv_hom_id, category.comp_id, + f_right_act_hom, ←category.assoc, ←comp_tensor_id, iso.inv_hom_id, + monoidal_category.tensor_id, category.id_comp], + end }, + hom_inv_id' := begin + ext, dsimp, rw iso.hom_inv_id, + end, + inv_hom_id' := begin + ext, dsimp, rw iso.inv_hom_id, + end } + +variables (A) + +/-- A monoid object as a bimodule over itself. -/ +@[simps] +def regular : Bimod A A := +{ X := A.X, + act_left := A.mul, + act_right := A.mul, } + +instance : inhabited (Bimod A A) := ⟨regular A⟩ + +/-- The forgetful functor from bimodule objects to the ambient category. -/ +def forget : Bimod A B ⥤ C := +{ obj := λ A, A.X, + map := λ A B f, f.hom, } + +open category_theory.limits + +variables [has_coequalizers C] + +namespace tensor_Bimod +variables {R S T : Mon_ C} (P : Bimod R S) (Q : Bimod S T) + +/-- The underlying object of the tensor product of two bimodules. -/ +noncomputable +def X : C := coequalizer (P.act_right ⊗ 𝟙 Q.X) ((α_ _ _ _).hom ≫ (𝟙 P.X ⊗ Q.act_left)) + +section + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_left X)] + +/-- Left action for the tensor product of two bimodules. -/ +noncomputable +def act_left : R.X ⊗ X P Q ⟶ X P Q := +(preserves_coequalizer.iso (tensor_left R.X) _ _).inv ≫ +colim_map + (parallel_pair_hom _ _ _ _ + ((𝟙 _ ⊗ (α_ _ _ _).hom) ≫ (α_ _ _ _).inv ≫ (P.act_left ⊗ 𝟙 S.X ⊗ 𝟙 Q.X) ≫ (α_ _ _ _).inv) + ((α_ _ _ _).inv ≫ (P.act_left ⊗ 𝟙 Q.X)) + begin + dsimp, + slice_lhs 1 2 { rw associator_inv_naturality }, + slice_rhs 3 4 { rw associator_inv_naturality }, + slice_rhs 4 5 { rw [←tensor_comp, middle_assoc, tensor_comp, comp_tensor_id] }, + coherence, + end + begin + dsimp, + slice_lhs 1 1 { rw id_tensor_comp }, + slice_lhs 2 3 { rw associator_inv_naturality }, + slice_lhs 3 4 { rw [tensor_id, id_tensor_comp_tensor_id] }, + slice_rhs 4 6 { rw iso.inv_hom_id_assoc }, + slice_rhs 3 4 { rw [tensor_id, tensor_id_comp_id_tensor] }, + end) + +lemma id_tensor_π_act_left : + (𝟙 R.X ⊗ coequalizer.π _ _) ≫ act_left P Q = + (α_ _ _ _).inv ≫ (P.act_left ⊗ 𝟙 Q.X) ≫ coequalizer.π _ _ := +begin + erw map_π_preserves_coequalizer_inv_colim_map (tensor_left _), + simp only [category.assoc], +end + +lemma one_act_left' : (R.one ⊗ 𝟙 _) ≫ act_left P Q = (λ_ _).hom := +begin + refine (cancel_epi ((tensor_left _).map (coequalizer.π _ _))).1 _, + dsimp [X], + slice_lhs 1 2 { rw [id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor] }, + slice_lhs 2 3 { rw id_tensor_π_act_left }, + slice_lhs 1 2 { rw [←monoidal_category.tensor_id, associator_inv_naturality] }, + slice_lhs 2 3 { rw [←comp_tensor_id, one_act_left] }, + slice_rhs 1 2 { rw left_unitor_naturality }, + coherence, +end + +lemma left_assoc' : + (R.mul ⊗ 𝟙 _) ≫ act_left P Q = + (α_ R.X R.X _).hom ≫ (𝟙 R.X ⊗ act_left P Q) ≫ act_left P Q := +begin + refine (cancel_epi ((tensor_left _).map (coequalizer.π _ _))).1 _, + dsimp [X], + slice_lhs 1 2 { rw [id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor] }, + slice_lhs 2 3 { rw id_tensor_π_act_left }, + slice_lhs 1 2 { rw [←monoidal_category.tensor_id, associator_inv_naturality] }, + slice_lhs 2 3 { rw [←comp_tensor_id, left_assoc, comp_tensor_id, comp_tensor_id] }, + slice_rhs 1 2 { rw [←monoidal_category.tensor_id, associator_naturality] }, + slice_rhs 2 3 { rw [←id_tensor_comp, id_tensor_π_act_left, id_tensor_comp, id_tensor_comp] }, + slice_rhs 4 5 { rw id_tensor_π_act_left }, + slice_rhs 3 4 { rw associator_inv_naturality }, + coherence, +end + +end + +section + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_right X)] + +/-- Right action for the tensor product of two bimodules. -/ +noncomputable +def act_right : X P Q ⊗ T.X ⟶ X P Q := +(preserves_coequalizer.iso (tensor_right T.X) _ _).inv ≫ +colim_map + (parallel_pair_hom _ _ _ _ + ((α_ _ _ _).hom ≫ (α_ _ _ _).hom ≫ (𝟙 P.X ⊗ 𝟙 S.X ⊗ Q.act_right) ≫ (α_ _ _ _).inv) + ((α_ _ _ _).hom ≫ (𝟙 P.X ⊗ Q.act_right)) + begin + dsimp, + slice_lhs 1 2 { rw associator_naturality }, + slice_lhs 2 3 { rw [tensor_id, tensor_id_comp_id_tensor] }, + slice_rhs 3 4 { rw associator_inv_naturality }, + slice_rhs 2 4 { rw iso.hom_inv_id_assoc }, + slice_rhs 2 3 { rw [tensor_id, id_tensor_comp_tensor_id] }, + end + begin + dsimp, + slice_lhs 1 1 { rw comp_tensor_id }, + slice_lhs 2 3 { rw associator_naturality }, + slice_lhs 3 4 { rw [←id_tensor_comp, middle_assoc, id_tensor_comp] }, + slice_rhs 4 6 { rw iso.inv_hom_id_assoc }, + slice_rhs 3 4 { rw ←id_tensor_comp }, + coherence, + end) + +lemma π_tensor_id_act_right : + (coequalizer.π _ _ ⊗ 𝟙 T.X) ≫ act_right P Q = + (α_ _ _ _).hom ≫ (𝟙 P.X ⊗ Q.act_right) ≫ coequalizer.π _ _ := +begin + erw map_π_preserves_coequalizer_inv_colim_map (tensor_right _), + simp only [category.assoc], +end + +lemma act_right_one' : (𝟙 _ ⊗ T.one) ≫ act_right P Q = (ρ_ _).hom := +begin + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp [X], + slice_lhs 1 2 { rw [tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id] }, + slice_lhs 2 3 { rw π_tensor_id_act_right }, + slice_lhs 1 2 { rw [←monoidal_category.tensor_id, associator_naturality] }, + slice_lhs 2 3 { rw [←id_tensor_comp, act_right_one] }, + slice_rhs 1 2 { rw right_unitor_naturality }, + coherence, +end + +lemma right_assoc' : + (𝟙 _ ⊗ T.mul) ≫ act_right P Q = + (α_ _ T.X T.X).inv ≫ (act_right P Q ⊗ 𝟙 T.X) ≫ act_right P Q := +begin + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp [X], + slice_lhs 1 2 { rw [tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id] }, + slice_lhs 2 3 { rw π_tensor_id_act_right }, + slice_lhs 1 2 { rw [←monoidal_category.tensor_id, associator_naturality] }, + slice_lhs 2 3 { rw [←id_tensor_comp, right_assoc, id_tensor_comp, id_tensor_comp] }, + slice_rhs 1 2 { rw [←monoidal_category.tensor_id, associator_inv_naturality] }, + slice_rhs 2 3 { rw [←comp_tensor_id, π_tensor_id_act_right, comp_tensor_id, comp_tensor_id] }, + slice_rhs 4 5 { rw π_tensor_id_act_right }, + slice_rhs 3 4 { rw associator_naturality }, + coherence, +end + +end + +section + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_left X)] +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_right X)] + +lemma middle_assoc' : + (act_left P Q ⊗ 𝟙 T.X) ≫ act_right P Q = + (α_ R.X _ T.X).hom ≫ (𝟙 R.X ⊗ act_right P Q) ≫ act_left P Q := +begin + refine (cancel_epi ((tensor_left _ ⋙ tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp [X], + slice_lhs 1 2 { rw [←comp_tensor_id, id_tensor_π_act_left, comp_tensor_id, comp_tensor_id] }, + slice_lhs 3 4 { rw π_tensor_id_act_right }, + slice_lhs 2 3 { rw associator_naturality }, + slice_lhs 3 4 { rw [monoidal_category.tensor_id, tensor_id_comp_id_tensor] }, + slice_rhs 1 2 { rw associator_naturality }, + slice_rhs 2 3 { rw [←id_tensor_comp, π_tensor_id_act_right, id_tensor_comp, id_tensor_comp] }, + slice_rhs 4 5 { rw id_tensor_π_act_left }, + slice_rhs 3 4 { rw associator_inv_naturality }, + slice_rhs 4 5 { rw [monoidal_category.tensor_id, id_tensor_comp_tensor_id] }, + coherence, +end + +end + +end tensor_Bimod + +section + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_left X)] +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_right X)] + +/-- Tensor product of two bimodule objects as a bimodule object. -/ +@[simps] +noncomputable +def tensor_Bimod {X Y Z : Mon_ C} (M : Bimod X Y) (N : Bimod Y Z) : Bimod X Z := +{ X := tensor_Bimod.X M N, + act_left := tensor_Bimod.act_left M N, + act_right := tensor_Bimod.act_right M N, + one_act_left' := tensor_Bimod.one_act_left' M N, + act_right_one' := tensor_Bimod.act_right_one' M N, + left_assoc' := tensor_Bimod.left_assoc' M N, + right_assoc' := tensor_Bimod.right_assoc' M N, + middle_assoc' := tensor_Bimod.middle_assoc' M N, } + +/-- Tensor product of two morphisms of bimodule objects. -/ +@[simps] +noncomputable +def tensor_hom {X Y Z : Mon_ C} {M₁ M₂ : Bimod X Y} {N₁ N₂ : Bimod Y Z} + (f : M₁ ⟶ M₂) (g : N₁ ⟶ N₂) : M₁.tensor_Bimod N₁ ⟶ M₂.tensor_Bimod N₂ := +{ hom := + colim_map + (parallel_pair_hom _ _ _ _ ((f.hom ⊗ 𝟙 Y.X) ⊗ g.hom) (f.hom ⊗ g.hom) + (by rw [←tensor_comp, ←tensor_comp, hom.right_act_hom, category.id_comp, category.comp_id]) + begin + slice_lhs 2 3 { rw [←tensor_comp, hom.left_act_hom, category.id_comp] }, + slice_rhs 1 2 { rw associator_naturality }, + slice_rhs 2 3 { rw [←tensor_comp, category.comp_id] }, + end), + left_act_hom' := begin + refine (cancel_epi ((tensor_left _).map (coequalizer.π _ _))).1 _, + dsimp, + slice_lhs 1 2 { rw tensor_Bimod.id_tensor_π_act_left }, + slice_lhs 3 4 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_lhs 2 3 { rw [←tensor_comp, hom.left_act_hom, category.id_comp] }, + slice_rhs 1 2 { rw [←id_tensor_comp, ι_colim_map, parallel_pair_hom_app_one, id_tensor_comp] }, + slice_rhs 2 3 { rw tensor_Bimod.id_tensor_π_act_left }, + slice_rhs 1 2 { rw associator_inv_naturality }, + slice_rhs 2 3 { rw [←tensor_comp, category.comp_id] }, + end, + right_act_hom' := begin + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp, + slice_lhs 1 2 { rw tensor_Bimod.π_tensor_id_act_right }, + slice_lhs 3 4 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_lhs 2 3 { rw [←tensor_comp, category.id_comp, hom.right_act_hom] }, + slice_rhs 1 2 { rw [←comp_tensor_id, ι_colim_map, parallel_pair_hom_app_one, comp_tensor_id] }, + slice_rhs 2 3 { rw tensor_Bimod.π_tensor_id_act_right }, + slice_rhs 1 2 { rw associator_naturality }, + slice_rhs 2 3 { rw [←tensor_comp, category.comp_id] }, + end } + +lemma tensor_id {X Y Z : Mon_ C} {M : Bimod X Y} {N : Bimod Y Z} : + tensor_hom (𝟙 M) (𝟙 N) = 𝟙 (M.tensor_Bimod N) := +begin + ext, + simp only [id_hom', tensor_id, tensor_hom_hom, ι_colim_map, parallel_pair_hom_app_one], + dsimp, dunfold tensor_Bimod.X, + simp only [category.id_comp, category.comp_id], +end + +lemma tensor_comp {X Y Z : Mon_ C} {M₁ M₂ M₃ : Bimod X Y} {N₁ N₂ N₃ : Bimod Y Z} + (f₁ : M₁ ⟶ M₂) (f₂ : M₂ ⟶ M₃) (g₁ : N₁ ⟶ N₂) (g₂ : N₂ ⟶ N₃) : + tensor_hom (f₁ ≫ f₂) (g₁ ≫ g₂) = tensor_hom f₁ g₁ ≫ tensor_hom f₂ g₂ := +begin + ext, + simp only [comp_hom', tensor_comp, tensor_hom_hom, ι_colim_map, parallel_pair_hom_app_one, + category.assoc, ι_colim_map_assoc] +end + +end + +namespace associator_Bimod + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_left X)] +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_right X)] + +variables {R S T U : Mon_ C} (P : Bimod R S) (Q : Bimod S T) (L : Bimod T U) + +/-- An auxiliary morphism for the definition of the underlying morphism of the forward component of +the associator isomorphism. -/ +noncomputable +def hom_aux : (P.tensor_Bimod Q).X ⊗ L.X ⟶ (P.tensor_Bimod (Q.tensor_Bimod L)).X := +(preserves_coequalizer.iso (tensor_right L.X) _ _).inv ≫ +coequalizer.desc + ((α_ _ _ _).hom ≫ (𝟙 P.X ⊗ (coequalizer.π _ _)) ≫ (coequalizer.π _ _)) + begin + dsimp, dsimp [tensor_Bimod.X], + slice_lhs 1 2 { rw associator_naturality }, + slice_lhs 2 3 { rw [monoidal_category.tensor_id, + tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id] }, + slice_lhs 3 4 { rw coequalizer.condition }, + slice_lhs 2 3 { rw [←monoidal_category.tensor_id, associator_naturality] }, + slice_lhs 3 4 { rw [←id_tensor_comp, tensor_Bimod.id_tensor_π_act_left, id_tensor_comp] }, + slice_rhs 1 1 { rw comp_tensor_id }, + slice_rhs 2 3 { rw associator_naturality }, + slice_rhs 3 4 { rw ←id_tensor_comp }, + coherence, + end + +/-- The underlying morphism of the forward component of the associator isomorphism. -/ +noncomputable +def hom : ((P.tensor_Bimod Q).tensor_Bimod L).X ⟶ (P.tensor_Bimod (Q.tensor_Bimod L)).X := +coequalizer.desc + (hom_aux P Q L) + begin + dsimp [hom_aux], + refine (cancel_epi ((tensor_right _ ⋙ tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp [tensor_Bimod.X], + slice_lhs 1 2 { rw [←comp_tensor_id, + tensor_Bimod.π_tensor_id_act_right, + comp_tensor_id, comp_tensor_id] }, + slice_lhs 3 5 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_lhs 2 3 { rw associator_naturality }, + slice_lhs 3 4 { rw [←id_tensor_comp, coequalizer.condition, id_tensor_comp, id_tensor_comp] }, + slice_rhs 1 2 { rw associator_naturality }, + slice_rhs 2 3 { rw [monoidal_category.tensor_id, + tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id] }, + slice_rhs 3 5 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_rhs 2 3 { rw [←monoidal_category.tensor_id, associator_naturality] }, + coherence, + end + +lemma hom_left_act_hom' : + ((P.tensor_Bimod Q).tensor_Bimod L).act_left ≫ hom P Q L = + (𝟙 R.X ⊗ hom P Q L) ≫ (P.tensor_Bimod (Q.tensor_Bimod L)).act_left := +begin + dsimp, dsimp [hom, hom_aux], + refine (cancel_epi ((tensor_left _).map (coequalizer.π _ _))).1 _, + rw tensor_left_map, + slice_lhs 1 2 { rw tensor_Bimod.id_tensor_π_act_left }, + slice_lhs 3 4 { rw coequalizer.π_desc }, + slice_rhs 1 2 { rw [←id_tensor_comp, coequalizer.π_desc, id_tensor_comp] }, + refine (cancel_epi ((tensor_right _ ⋙ tensor_left _).map (coequalizer.π _ _))).1 _, + dsimp, dsimp [tensor_Bimod.X], + slice_lhs 1 2 { rw associator_inv_naturality }, + slice_lhs 2 3 { rw [←comp_tensor_id, + tensor_Bimod.id_tensor_π_act_left, + comp_tensor_id, comp_tensor_id] }, + slice_lhs 4 6 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_lhs 3 4 { rw associator_naturality }, + slice_lhs 4 5 { rw [monoidal_category.tensor_id, tensor_id_comp_id_tensor] }, + slice_rhs 1 3 { rw [←id_tensor_comp, ←id_tensor_comp, + π_tensor_id_preserves_coequalizer_inv_desc, + id_tensor_comp, id_tensor_comp] }, + slice_rhs 3 4 { erw tensor_Bimod.id_tensor_π_act_left P (Q.tensor_Bimod L) }, + slice_rhs 2 3 { erw associator_inv_naturality }, + slice_rhs 3 4 { erw [monoidal_category.tensor_id, id_tensor_comp_tensor_id] }, + coherence, +end + +lemma hom_right_act_hom' : + ((P.tensor_Bimod Q).tensor_Bimod L).act_right ≫ hom P Q L = + (hom P Q L ⊗ 𝟙 U.X) ≫ (P.tensor_Bimod (Q.tensor_Bimod L)).act_right := +begin + dsimp, dsimp [hom, hom_aux], + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + rw tensor_right_map, + slice_lhs 1 2 { rw tensor_Bimod.π_tensor_id_act_right }, + slice_lhs 3 4 { rw coequalizer.π_desc }, + slice_rhs 1 2 { rw [←comp_tensor_id, coequalizer.π_desc, comp_tensor_id] }, + refine (cancel_epi ((tensor_right _ ⋙ tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp, dsimp [tensor_Bimod.X], + slice_lhs 1 2 { rw associator_naturality }, + slice_lhs 2 3 { rw [monoidal_category.tensor_id, + tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id] }, + slice_lhs 3 5 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_lhs 2 3 { rw [←monoidal_category.tensor_id, + associator_naturality] }, + slice_rhs 1 3 { rw [←comp_tensor_id, ←comp_tensor_id, + π_tensor_id_preserves_coequalizer_inv_desc, + comp_tensor_id, comp_tensor_id] }, + slice_rhs 3 4 { erw tensor_Bimod.π_tensor_id_act_right P (Q.tensor_Bimod L) }, + slice_rhs 2 3 { erw associator_naturality }, + dsimp, + slice_rhs 3 4 { rw [←id_tensor_comp, + tensor_Bimod.π_tensor_id_act_right, + id_tensor_comp, id_tensor_comp] }, + coherence, +end + +/-- An auxiliary morphism for the definition of the underlying morphism of the inverse component of +the associator isomorphism. -/ +noncomputable +def inv_aux : P.X ⊗ (Q.tensor_Bimod L).X ⟶ ((P.tensor_Bimod Q).tensor_Bimod L).X := +(preserves_coequalizer.iso (tensor_left P.X) _ _).inv ≫ +coequalizer.desc + ((α_ _ _ _).inv ≫ ((coequalizer.π _ _) ⊗ 𝟙 L.X) ≫ (coequalizer.π _ _)) + begin + dsimp, dsimp [tensor_Bimod.X], + slice_lhs 1 2 { rw associator_inv_naturality }, + rw [←(iso.inv_hom_id_assoc (α_ _ _ _) (𝟙 P.X ⊗ Q.act_right)), comp_tensor_id], + slice_lhs 3 4 { rw [←comp_tensor_id, category.assoc, ←tensor_Bimod.π_tensor_id_act_right, + comp_tensor_id] }, + slice_lhs 4 5 { rw coequalizer.condition }, + slice_lhs 3 4 { rw associator_naturality }, + slice_lhs 4 5 { rw [monoidal_category.tensor_id, tensor_id_comp_id_tensor] }, + slice_rhs 1 2 { rw id_tensor_comp }, + slice_rhs 2 3 { rw associator_inv_naturality }, + slice_rhs 3 4 { rw [monoidal_category.tensor_id, id_tensor_comp_tensor_id] }, + coherence, + end + +/-- The underlying morphism of the inverse component of the associator isomorphism. -/ +noncomputable +def inv : (P.tensor_Bimod (Q.tensor_Bimod L)).X ⟶ ((P.tensor_Bimod Q).tensor_Bimod L).X := +coequalizer.desc + (inv_aux P Q L) + begin + dsimp [inv_aux], + refine (cancel_epi ((tensor_left _).map (coequalizer.π _ _))).1 _, + dsimp [tensor_Bimod.X], + slice_lhs 1 2 { rw [id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor] }, + slice_lhs 2 4 { rw id_tensor_π_preserves_coequalizer_inv_desc }, + slice_lhs 1 2 { rw [←monoidal_category.tensor_id, associator_inv_naturality] }, + slice_lhs 2 3 { rw [←comp_tensor_id, coequalizer.condition, comp_tensor_id, comp_tensor_id] }, + slice_rhs 1 2 { rw [←monoidal_category.tensor_id, associator_naturality] }, + slice_rhs 2 3 { rw [←id_tensor_comp, + tensor_Bimod.id_tensor_π_act_left, + id_tensor_comp, id_tensor_comp] }, + slice_rhs 4 6 { rw id_tensor_π_preserves_coequalizer_inv_desc }, + slice_rhs 3 4 { rw associator_inv_naturality }, + coherence, + end + +lemma hom_inv_id : hom P Q L ≫ inv P Q L = 𝟙 _ := +begin + dsimp [hom, hom_aux, inv, inv_aux], + ext, + slice_lhs 1 2 { rw coequalizer.π_desc }, + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + rw tensor_right_map, + slice_lhs 1 3 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_lhs 3 4 { rw coequalizer.π_desc }, + slice_lhs 2 4 { rw id_tensor_π_preserves_coequalizer_inv_desc }, + slice_lhs 1 3 { rw iso.hom_inv_id_assoc }, + dunfold tensor_Bimod.X, + slice_rhs 2 3 { rw category.comp_id }, + refl, +end + +lemma inv_hom_id : inv P Q L ≫ hom P Q L = 𝟙 _ := +begin + dsimp [hom, hom_aux, inv, inv_aux], + ext, + slice_lhs 1 2 { rw coequalizer.π_desc }, + refine (cancel_epi ((tensor_left _).map (coequalizer.π _ _))).1 _, + rw tensor_left_map, + slice_lhs 1 3 { rw id_tensor_π_preserves_coequalizer_inv_desc }, + slice_lhs 3 4 { rw coequalizer.π_desc }, + slice_lhs 2 4 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_lhs 1 3 { rw iso.inv_hom_id_assoc }, + dunfold tensor_Bimod.X, + slice_rhs 2 3 { rw category.comp_id }, + refl, +end + +end associator_Bimod + +namespace left_unitor_Bimod +variables {R S : Mon_ C} (P : Bimod R S) + +/-- The underlying morphism of the forward component of the left unitor isomorphism. -/ +noncomputable +def hom : tensor_Bimod.X (regular R) P ⟶ P.X := +coequalizer.desc P.act_left (by { dsimp, rw [category.assoc, left_assoc] }) + +/-- The underlying morphism of the inverse component of the left unitor isomorphism. -/ +noncomputable +def inv : P.X ⟶ tensor_Bimod.X (regular R) P := +(λ_ P.X).inv ≫ (R.one ⊗ 𝟙 _) ≫ coequalizer.π _ _ + +lemma hom_inv_id : hom P ≫ inv P = 𝟙 _ := +begin + dunfold hom inv tensor_Bimod.X, + ext, dsimp, + slice_lhs 1 2 { rw coequalizer.π_desc }, + slice_lhs 1 2 { rw left_unitor_inv_naturality }, + slice_lhs 2 3 { rw [id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor] }, + slice_lhs 3 3 { rw ←(iso.inv_hom_id_assoc (α_ R.X R.X P.X) (𝟙 R.X ⊗ P.act_left)) }, + slice_lhs 4 6 { rw [←category.assoc, ←coequalizer.condition] }, + slice_lhs 2 3 { rw [←monoidal_category.tensor_id, associator_inv_naturality] }, + slice_lhs 3 4 { rw [←comp_tensor_id, Mon_.one_mul] }, + slice_rhs 1 2 { rw category.comp_id }, + coherence, +end + +lemma inv_hom_id : inv P ≫ hom P = 𝟙 _ := +begin + dsimp [hom, inv], + slice_lhs 3 4 { rw coequalizer.π_desc }, + rw [one_act_left, iso.inv_hom_id], +end + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_left X)] +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_right X)] + +lemma hom_left_act_hom' : + ((regular R).tensor_Bimod P).act_left ≫ hom P = (𝟙 R.X ⊗ hom P) ≫ P.act_left := +begin + dsimp, dsimp [hom, tensor_Bimod.act_left, regular], + refine (cancel_epi ((tensor_left _).map (coequalizer.π _ _))).1 _, + dsimp, + slice_lhs 1 4 { rw id_tensor_π_preserves_coequalizer_inv_colim_map_desc }, + slice_lhs 2 3 { rw left_assoc }, + slice_rhs 1 2 { rw [←id_tensor_comp, coequalizer.π_desc] }, + rw iso.inv_hom_id_assoc, +end + +lemma hom_right_act_hom' : + ((regular R).tensor_Bimod P).act_right ≫ hom P = (hom P ⊗ 𝟙 S.X) ≫ P.act_right := +begin + dsimp, dsimp [hom, tensor_Bimod.act_right, regular], + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp, + slice_lhs 1 4 { rw π_tensor_id_preserves_coequalizer_inv_colim_map_desc }, + slice_rhs 1 2 { rw [←comp_tensor_id, coequalizer.π_desc] }, + slice_rhs 1 2 { rw middle_assoc }, + simp only [category.assoc], +end + +end left_unitor_Bimod + +namespace right_unitor_Bimod +variables {R S : Mon_ C} (P : Bimod R S) + +/-- The underlying morphism of the forward component of the right unitor isomorphism. -/ +noncomputable +def hom : tensor_Bimod.X P (regular S) ⟶ P.X := +coequalizer.desc P.act_right + (by { dsimp, rw [category.assoc, right_assoc, iso.hom_inv_id_assoc] }) + +/-- The underlying morphism of the inverse component of the right unitor isomorphism. -/ +noncomputable +def inv : P.X ⟶ tensor_Bimod.X P (regular S) := +(ρ_ P.X).inv ≫ (𝟙 _ ⊗ S.one) ≫ coequalizer.π _ _ + +lemma hom_inv_id : hom P ≫ inv P = 𝟙 _ := +begin + dunfold hom inv tensor_Bimod.X, + ext, dsimp, + slice_lhs 1 2 { rw coequalizer.π_desc }, + slice_lhs 1 2 { rw right_unitor_inv_naturality }, + slice_lhs 2 3 { rw [tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id] }, + slice_lhs 3 4 { rw coequalizer.condition }, + slice_lhs 2 3 { rw [←monoidal_category.tensor_id, associator_naturality] }, + slice_lhs 3 4 { rw [←id_tensor_comp, Mon_.mul_one] }, + slice_rhs 1 2 { rw category.comp_id }, + coherence, +end + +lemma inv_hom_id : inv P ≫ hom P = 𝟙 _ := +begin + dsimp [hom, inv], + slice_lhs 3 4 { rw coequalizer.π_desc }, + rw [act_right_one, iso.inv_hom_id], +end + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_left X)] +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_right X)] + +lemma hom_left_act_hom' : + (P.tensor_Bimod (regular S)).act_left ≫ hom P = (𝟙 R.X ⊗ hom P) ≫ P.act_left := +begin + dsimp, dsimp [hom, tensor_Bimod.act_left, regular], + refine (cancel_epi ((tensor_left _).map (coequalizer.π _ _))).1 _, + dsimp, + slice_lhs 1 4 { rw id_tensor_π_preserves_coequalizer_inv_colim_map_desc }, + slice_lhs 2 3 { rw middle_assoc }, + slice_rhs 1 2 { rw [←id_tensor_comp, coequalizer.π_desc] }, + rw iso.inv_hom_id_assoc, +end + +lemma hom_right_act_hom' : + (P.tensor_Bimod (regular S)).act_right ≫ hom P = (hom P ⊗ 𝟙 S.X) ≫ P.act_right := +begin + dsimp, dsimp [hom, tensor_Bimod.act_right, regular], + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp, + slice_lhs 1 4 { rw π_tensor_id_preserves_coequalizer_inv_colim_map_desc }, + slice_lhs 2 3 { rw right_assoc }, + slice_rhs 1 2 { rw [←comp_tensor_id, coequalizer.π_desc] }, + rw iso.hom_inv_id_assoc, +end + +end right_unitor_Bimod + +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_left X)] +variables [∀ X : C, preserves_colimits_of_size.{0 0} (tensor_right X)] + +/-- The associator as a bimodule isomorphism. -/ +noncomputable +def associator_Bimod {W X Y Z : Mon_ C} (L : Bimod W X) (M : Bimod X Y) (N : Bimod Y Z) : + (L.tensor_Bimod M).tensor_Bimod N ≅ L.tensor_Bimod (M.tensor_Bimod N) := +iso_of_iso + { hom := associator_Bimod.hom L M N, + inv := associator_Bimod.inv L M N, + hom_inv_id' := associator_Bimod.hom_inv_id L M N, + inv_hom_id' := associator_Bimod.inv_hom_id L M N } + (associator_Bimod.hom_left_act_hom' L M N) + (associator_Bimod.hom_right_act_hom' L M N) + +/-- The left unitor as a bimodule isomorphism. -/ +noncomputable +def left_unitor_Bimod {X Y : Mon_ C} (M : Bimod X Y) : (regular X).tensor_Bimod M ≅ M := +iso_of_iso + { hom := left_unitor_Bimod.hom M, + inv := left_unitor_Bimod.inv M, + hom_inv_id' := left_unitor_Bimod.hom_inv_id M, + inv_hom_id' := left_unitor_Bimod.inv_hom_id M } + (left_unitor_Bimod.hom_left_act_hom' M) + (left_unitor_Bimod.hom_right_act_hom' M) + +/-- The right unitor as a bimodule isomorphism. -/ +noncomputable +def right_unitor_Bimod {X Y : Mon_ C} (M : Bimod X Y) : M.tensor_Bimod (regular Y) ≅ M := +iso_of_iso + { hom := right_unitor_Bimod.hom M, + inv := right_unitor_Bimod.inv M, + hom_inv_id' := right_unitor_Bimod.hom_inv_id M, + inv_hom_id' := right_unitor_Bimod.inv_hom_id M } + (right_unitor_Bimod.hom_left_act_hom' M) + (right_unitor_Bimod.hom_right_act_hom' M) + +lemma whisker_left_comp_Bimod {X Y Z : Mon_ C} + (M : Bimod X Y) {N P Q : Bimod Y Z} (f : N ⟶ P) (g : P ⟶ Q) : + tensor_hom (𝟙 M) (f ≫ g) = tensor_hom (𝟙 M) f ≫ tensor_hom (𝟙 M) g := +by rw [←tensor_comp, category.comp_id] + +lemma id_whisker_left_Bimod {X Y : Mon_ C} {M N : Bimod X Y} (f : M ⟶ N) : + tensor_hom (𝟙 (regular X)) f = (left_unitor_Bimod M).hom ≫ f ≫ (left_unitor_Bimod N).inv := +begin + dsimp [tensor_hom, regular, left_unitor_Bimod], + ext, dsimp, + slice_lhs 1 2 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + dsimp [left_unitor_Bimod.hom], + slice_rhs 1 2 { rw coequalizer.π_desc }, + dsimp [left_unitor_Bimod.inv], + slice_rhs 1 2 { rw hom.left_act_hom }, + slice_rhs 2 3 { rw left_unitor_inv_naturality }, + slice_rhs 3 4 { rw [id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor] }, + slice_rhs 4 4 { rw ←(iso.inv_hom_id_assoc (α_ X.X X.X N.X) (𝟙 X.X ⊗ N.act_left)) }, + slice_rhs 5 7 { rw [←category.assoc, ←coequalizer.condition] }, + slice_rhs 3 4 { rw [←monoidal_category.tensor_id, associator_inv_naturality] }, + slice_rhs 4 5 { rw [←comp_tensor_id, Mon_.one_mul] }, + have : + (λ_ (X.X ⊗ N.X)).inv ≫ (α_ (𝟙_ C) X.X N.X).inv ≫ ((λ_ X.X).hom ⊗ 𝟙 N.X) = 𝟙 _ := + by pure_coherence, + slice_rhs 2 4 { rw this }, + slice_rhs 1 2 { rw category.comp_id }, +end + +lemma comp_whisker_left_Bimod {W X Y Z : Mon_ C} + (M : Bimod W X) (N : Bimod X Y) {P P' : Bimod Y Z} (f : P ⟶ P') : + tensor_hom (𝟙 (M.tensor_Bimod N)) f = + (associator_Bimod M N P).hom ≫ tensor_hom (𝟙 M) (tensor_hom (𝟙 N) f) ≫ + (associator_Bimod M N P').inv := +begin + dsimp [tensor_hom, tensor_Bimod, associator_Bimod], + ext, dsimp, + slice_lhs 1 2 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + dsimp [tensor_Bimod.X, associator_Bimod.hom], + slice_rhs 1 2 { rw coequalizer.π_desc }, + dsimp [associator_Bimod.hom_aux, associator_Bimod.inv], + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + rw tensor_right_map, + slice_rhs 1 3 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_rhs 3 4 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_rhs 2 3 { rw [←id_tensor_comp, ι_colim_map, parallel_pair_hom_app_one] }, + slice_rhs 3 4 { rw coequalizer.π_desc }, + dsimp [associator_Bimod.inv_aux], + slice_rhs 2 2 { rw id_tensor_comp }, + slice_rhs 3 5 { rw id_tensor_π_preserves_coequalizer_inv_desc }, + slice_rhs 2 3 { rw associator_inv_naturality }, + slice_rhs 1 3 { rw [iso.hom_inv_id_assoc, monoidal_category.tensor_id] }, + slice_lhs 1 2 { rw [tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id] }, + dunfold tensor_Bimod.X, + simp only [category.assoc], +end + +lemma comp_whisker_right_Bimod {X Y Z : Mon_ C} + {M N P : Bimod X Y} (f : M ⟶ N) (g : N ⟶ P) (Q : Bimod Y Z) : + tensor_hom (f ≫ g) (𝟙 Q) = tensor_hom f (𝟙 Q) ≫ tensor_hom g (𝟙 Q) := +by rw [←tensor_comp, category.comp_id] + +lemma whisker_right_id_Bimod {X Y : Mon_ C} {M N : Bimod X Y} (f : M ⟶ N) : + tensor_hom f (𝟙 (regular Y)) = (right_unitor_Bimod M).hom ≫ f ≫ (right_unitor_Bimod N).inv := +begin + dsimp [tensor_hom, regular, right_unitor_Bimod], + ext, dsimp, + slice_lhs 1 2 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + dsimp [right_unitor_Bimod.hom], + slice_rhs 1 2 { rw coequalizer.π_desc }, + dsimp [right_unitor_Bimod.inv], + slice_rhs 1 2 { rw hom.right_act_hom }, + slice_rhs 2 3 { rw right_unitor_inv_naturality }, + slice_rhs 3 4 { rw [tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id] }, + slice_rhs 4 5 { rw coequalizer.condition }, + slice_rhs 3 4 { rw [←monoidal_category.tensor_id, associator_naturality] }, + slice_rhs 4 5 { rw [←id_tensor_comp, Mon_.mul_one] }, + have : + (ρ_ (N.X ⊗ Y.X)).inv ≫ (α_ N.X Y.X (𝟙_ C)).hom ≫ (𝟙 N.X ⊗ (ρ_ Y.X).hom) = 𝟙 _ := + by pure_coherence, + slice_rhs 2 4 { rw this }, + slice_rhs 1 2 { rw category.comp_id }, +end + +lemma whisker_right_comp_Bimod {W X Y Z : Mon_ C} + {M M' : Bimod W X} (f : M ⟶ M') (N : Bimod X Y) (P : Bimod Y Z) : + tensor_hom f (𝟙 (N.tensor_Bimod P)) = + (associator_Bimod M N P).inv ≫ tensor_hom (tensor_hom f (𝟙 N)) (𝟙 P) ≫ + (associator_Bimod M' N P).hom := +begin + dsimp [tensor_hom, tensor_Bimod, associator_Bimod], + ext, dsimp, + slice_lhs 1 2 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + dsimp [tensor_Bimod.X, associator_Bimod.inv], + slice_rhs 1 2 { rw coequalizer.π_desc }, + dsimp [associator_Bimod.inv_aux, associator_Bimod.hom], + refine (cancel_epi ((tensor_left _).map (coequalizer.π _ _))).1 _, + rw tensor_left_map, + slice_rhs 1 3 { rw id_tensor_π_preserves_coequalizer_inv_desc }, + slice_rhs 3 4 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_rhs 2 3 { rw [←comp_tensor_id, ι_colim_map, parallel_pair_hom_app_one] }, + slice_rhs 3 4 { rw coequalizer.π_desc }, + dsimp [associator_Bimod.hom_aux], + slice_rhs 2 2 { rw comp_tensor_id }, + slice_rhs 3 5 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_rhs 2 3 { rw associator_naturality }, + slice_rhs 1 3 { rw [iso.inv_hom_id_assoc, monoidal_category.tensor_id] }, + slice_lhs 1 2 { rw [id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor] }, + dunfold tensor_Bimod.X, + simp only [category.assoc], +end + +lemma whisker_assoc_Bimod {W X Y Z : Mon_ C} + (M : Bimod W X) {N N' : Bimod X Y} (f : N ⟶ N') (P : Bimod Y Z) : + tensor_hom (tensor_hom (𝟙 M) f) (𝟙 P) = + (associator_Bimod M N P).hom ≫ tensor_hom (𝟙 M) (tensor_hom f (𝟙 P)) ≫ + (associator_Bimod M N' P).inv := +begin + dsimp [tensor_hom, tensor_Bimod, associator_Bimod], + ext, dsimp, + slice_lhs 1 2 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + dsimp [associator_Bimod.hom], + slice_rhs 1 2 { rw coequalizer.π_desc }, + dsimp [associator_Bimod.hom_aux], + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + rw tensor_right_map, + slice_lhs 1 2 { rw [←comp_tensor_id, ι_colim_map, parallel_pair_hom_app_one] }, + slice_rhs 1 3 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_rhs 3 4 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_rhs 2 3 { rw [←id_tensor_comp, ι_colim_map, parallel_pair_hom_app_one] }, + dsimp [associator_Bimod.inv], + slice_rhs 3 4 { rw coequalizer.π_desc }, + dsimp [associator_Bimod.inv_aux], + slice_rhs 2 2 { rw id_tensor_comp }, + slice_rhs 3 5 { rw id_tensor_π_preserves_coequalizer_inv_desc }, + slice_rhs 2 3 { rw associator_inv_naturality }, + slice_rhs 1 3 { rw iso.hom_inv_id_assoc }, + slice_lhs 1 1 { rw comp_tensor_id }, +end + +lemma whisker_exchange_Bimod {X Y Z : Mon_ C} + {M N : Bimod X Y} {P Q : Bimod Y Z} (f : M ⟶ N) (g : P ⟶ Q) : + tensor_hom (𝟙 M) g ≫ tensor_hom f (𝟙 Q) = tensor_hom f (𝟙 P) ≫ tensor_hom (𝟙 N) g := +begin + dsimp [tensor_hom], + ext, dsimp, + slice_lhs 1 2 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_lhs 2 3 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_lhs 1 2 { rw id_tensor_comp_tensor_id }, + slice_rhs 1 2 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_rhs 2 3 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_rhs 1 2 { rw tensor_id_comp_id_tensor }, +end + +lemma pentagon_Bimod {V W X Y Z : Mon_ C} + (M : Bimod V W) (N : Bimod W X) (P : Bimod X Y) (Q : Bimod Y Z) : + tensor_hom (associator_Bimod M N P).hom (𝟙 Q) ≫ (associator_Bimod M (N.tensor_Bimod P) Q).hom ≫ + tensor_hom (𝟙 M) (associator_Bimod N P Q).hom = + (associator_Bimod (M.tensor_Bimod N) P Q).hom ≫ (associator_Bimod M N (P.tensor_Bimod Q)).hom := +begin + dsimp [tensor_hom, associator_Bimod], ext, dsimp, + dunfold associator_Bimod.hom, + slice_lhs 1 2 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_lhs 2 3 { rw coequalizer.π_desc }, + slice_rhs 1 2 { rw coequalizer.π_desc }, + dsimp [associator_Bimod.hom_aux], + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp, + slice_lhs 1 2 { rw [←comp_tensor_id, coequalizer.π_desc] }, + slice_rhs 1 3 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_rhs 3 4 { rw coequalizer.π_desc }, + refine (cancel_epi ((tensor_right _ ⋙ tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp, + slice_lhs 1 2 { rw [←comp_tensor_id, + π_tensor_id_preserves_coequalizer_inv_desc, + comp_tensor_id, comp_tensor_id ]}, + slice_lhs 3 5 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + dunfold tensor_Bimod.X, + slice_lhs 2 3 { rw associator_naturality }, + slice_lhs 5 6 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + slice_lhs 4 5 { rw [←id_tensor_comp, coequalizer.π_desc] }, + slice_lhs 3 4 { rw [←id_tensor_comp, + π_tensor_id_preserves_coequalizer_inv_desc, + id_tensor_comp, id_tensor_comp] }, + slice_rhs 1 2 { rw associator_naturality }, + slice_rhs 2 3 { rw [monoidal_category.tensor_id, + tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id] }, + slice_rhs 3 5 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_rhs 2 3 { rw [←monoidal_category.tensor_id, associator_naturality] }, + coherence, +end + +lemma triangle_Bimod {X Y Z : Mon_ C} (M : Bimod X Y) (N : Bimod Y Z) : + (associator_Bimod M (regular Y) N).hom ≫ tensor_hom (𝟙 M) (left_unitor_Bimod N).hom = + tensor_hom (right_unitor_Bimod M).hom (𝟙 N) := +begin + dsimp [tensor_hom, associator_Bimod, left_unitor_Bimod, right_unitor_Bimod], + ext, dsimp, + dsimp [associator_Bimod.hom], + slice_lhs 1 2 { rw coequalizer.π_desc }, + dsimp [associator_Bimod.hom_aux], + slice_rhs 1 2 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + dsimp [right_unitor_Bimod.hom], + refine (cancel_epi ((tensor_right _).map (coequalizer.π _ _))).1 _, + dsimp [regular], + slice_lhs 1 3 { rw π_tensor_id_preserves_coequalizer_inv_desc }, + slice_lhs 3 4 { rw [ι_colim_map, parallel_pair_hom_app_one] }, + dsimp [left_unitor_Bimod.hom], + slice_lhs 2 3 { rw [←id_tensor_comp, coequalizer.π_desc] }, + slice_rhs 1 2 { rw [←comp_tensor_id, coequalizer.π_desc] }, + slice_rhs 1 2 { rw coequalizer.condition }, + simp only [category.assoc], +end + +/-- The bicategory of algebras (monoids) and bimodules, all internal to some monoidal category. -/ +noncomputable +def Mon_bicategory : bicategory (Mon_ C) := +{ hom := λ X Y, Bimod X Y, + id := λ X, regular X, + comp := λ _ _ _ M N, tensor_Bimod M N, + whisker_left := λ _ _ _ L _ _ f, tensor_hom (𝟙 L) f, + whisker_right := λ _ _ _ _ _ f N, tensor_hom f (𝟙 N), + associator := λ _ _ _ _ L M N, associator_Bimod L M N, + left_unitor := λ _ _ M, left_unitor_Bimod M, + right_unitor := λ _ _ M, right_unitor_Bimod M, + whisker_left_id' := λ _ _ _ _ _, tensor_id, + whisker_left_comp' := λ _ _ _ M _ _ _ f g, whisker_left_comp_Bimod M f g, + id_whisker_left' := λ _ _ _ _ f, id_whisker_left_Bimod f, + comp_whisker_left' := λ _ _ _ _ M N _ _ f, comp_whisker_left_Bimod M N f, + id_whisker_right' := λ _ _ _ _ _, tensor_id, + comp_whisker_right' := λ _ _ _ _ _ _ f g Q, comp_whisker_right_Bimod f g Q, + whisker_right_id' := λ _ _ _ _ f, whisker_right_id_Bimod f, + whisker_right_comp' := λ _ _ _ _ _ _ f N P, whisker_right_comp_Bimod f N P, + whisker_assoc' := λ _ _ _ _ M _ _ f P, whisker_assoc_Bimod M f P, + whisker_exchange' := λ _ _ _ _ _ _ _ f g, whisker_exchange_Bimod f g, + pentagon' := λ _ _ _ _ _ M N P Q, pentagon_Bimod M N P Q, + triangle' := λ _ _ _ M N, triangle_Bimod M N } + +end Bimod From 64505769e222dbe47158b9f147c0f0b46714b4ca Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Sun, 16 Oct 2022 15:52:00 +0000 Subject: [PATCH 788/828] feat(algebra/order/floor): several round and fract lemmas (#17001) Lemmas about fract of negations, and about rounds under various operations --- src/algebra/order/floor.lean | 67 ++++++++++++++++++++++++++++++++++++ 1 file changed, 67 insertions(+) diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index 53bd2d11f81a1..cdcfe97b3b240 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -519,12 +519,21 @@ ext $ λ x, floor_eq_iff @[simp] lemma fract_add_int (a : α) (m : ℤ) : fract (a + m) = fract a := by { rw fract, simp } +@[simp] lemma fract_add_nat (a : α) (m : ℕ) : fract (a + m) = fract a := +by { rw fract, simp } + @[simp] lemma fract_sub_int (a : α) (m : ℤ) : fract (a - m) = fract a := by { rw fract, simp } @[simp] lemma fract_int_add (m : ℤ) (a : α) : fract (↑m + a) = fract a := by rw [add_comm, fract_add_int] +@[simp] lemma fract_sub_nat (a : α) (n : ℕ) : fract (a - n) = fract a := +by { rw fract, simp } + +@[simp] lemma fract_int_nat (n : ℕ) (a : α) : fract (↑n + a) = fract a := +by rw [add_comm, fract_add_nat] + @[simp] lemma self_sub_fract (a : α) : a - fract a = ⌊a⌋ := sub_sub_cancel _ _ @[simp] lemma fract_sub_self (a : α) : fract a - a = -⌊a⌋ := sub_sub_cancel_left _ _ @@ -582,6 +591,26 @@ fract_eq_self.2 ⟨fract_nonneg _, fract_lt_one _⟩ lemma fract_add (a b : α) : ∃ z : ℤ, fract (a + b) - fract a - fract b = z := ⟨⌊a⌋ + ⌊b⌋ - ⌊a + b⌋, by { unfold fract, simp [sub_eq_add_neg], abel }⟩ +lemma fract_neg {x : α} (hx : fract x ≠ 0) : + fract (-x) = 1 - fract x := +begin + rw fract_eq_iff, + split, + { rw [le_sub_iff_add_le, zero_add], + exact (fract_lt_one x).le, }, + refine ⟨sub_lt_self _ (lt_of_le_of_ne' (fract_nonneg x) hx), -⌊x⌋ - 1, _⟩, + simp only [sub_sub_eq_add_sub, cast_sub, cast_neg, cast_one, sub_left_inj], + conv in (-x) {rw ← floor_add_fract x}, + simp [-floor_add_fract], +end + +@[simp] +lemma fract_neg_eq_zero {x : α} : fract (-x) = 0 ↔ fract x = 0 := +begin + simp only [fract_eq_iff, le_refl, zero_lt_one, tsub_zero, true_and], + split; rintros ⟨z, hz⟩; use [-z]; simp [← hz], +end + lemma fract_mul_nat (a : α) (b : ℕ) : ∃ z : ℤ, fract a * b - fract (a * b) = z := begin induction b with c hc, @@ -669,12 +698,18 @@ lemma ceil_mono : monotone (ceil : α → ℤ) := gc_ceil_coe.monotone_l @[simp] lemma ceil_add_int (a : α) (z : ℤ) : ⌈a + z⌉ = ⌈a⌉ + z := by rw [←neg_inj, neg_add', ←floor_neg, ←floor_neg, neg_add', floor_sub_int] +@[simp] lemma ceil_add_nat (a : α) (n : ℕ) : ⌈a + n⌉ = ⌈a⌉ + n := +by rw [← int.cast_coe_nat, ceil_add_int] + @[simp] lemma ceil_add_one (a : α) : ⌈a + 1⌉ = ⌈a⌉ + 1 := by { convert ceil_add_int a (1 : ℤ), exact cast_one.symm } @[simp] lemma ceil_sub_int (a : α) (z : ℤ) : ⌈a - z⌉ = ⌈a⌉ - z := eq.trans (by rw [int.cast_neg, sub_eq_add_neg]) (ceil_add_int _ _) +@[simp] lemma ceil_sub_nat (a : α) (n : ℕ) : ⌈a - n⌉ = ⌈a⌉ - n := +by convert ceil_sub_int a n using 1; simp + @[simp] lemma ceil_sub_one (a : α) : ⌈a - 1⌉ = ⌈a⌉ - 1 := by rw [eq_sub_iff_add_eq, ← ceil_add_one, sub_add_cancel] @@ -774,6 +809,38 @@ def round (x : α) : ℤ := if 2 * fract x < 1 then ⌊x⌋ else ⌈x⌉ @[simp] lemma round_int_cast (n : ℤ) : round (n : α) = n := by simp [round] +@[simp] +lemma round_add_int (x : α) (y : ℤ) : round (x + y) = round x + y := +by rw [round, round, int.fract_add_int, int.floor_add_int, int.ceil_add_int, ← apply_ite2, if_t_t] + +@[simp] +lemma round_add_one (a : α) : round (a + 1) = round a + 1 := +by { convert round_add_int a 1, exact int.cast_one.symm } + +@[simp] +lemma round_sub_int (x : α) (y : ℤ) : round (x - y) = round x - y := +by { rw [sub_eq_add_neg], norm_cast, rw [round_add_int, sub_eq_add_neg] } + +@[simp] +lemma round_sub_one (a : α) : round (a - 1) = round a - 1 := +by { convert round_sub_int a 1, exact int.cast_one.symm } + +@[simp] +lemma round_add_nat (x : α) (y : ℕ) : round (x + y) = round x + y := +by rw [round, round, fract_add_nat, int.floor_add_nat, int.ceil_add_nat, ← apply_ite2, if_t_t] + +@[simp] +lemma round_sub_nat (x : α) (y : ℕ) : round (x - y) = round x - y := +by { rw [sub_eq_add_neg, ← int.cast_coe_nat], norm_cast, rw [round_add_int, sub_eq_add_neg] } + +@[simp] +lemma round_int_add (x : α) (y : ℤ) : round ((y : α) + x) = y + round x := +by { rw [add_comm, round_add_int, add_comm] } + +@[simp] +lemma round_nat_add (x : α) (y : ℕ) : round ((y : α) + x) = y + round x := +by { rw [add_comm, round_add_nat, add_comm] } + lemma abs_sub_round_eq_min (x : α) : |x - round x| = min (fract x) (1 - fract x) := begin simp_rw [round, min_def', two_mul, ← lt_tsub_iff_left], From e905eeab3a6b9160abb392784bf2a7ea4da23f2e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 16 Oct 2022 18:10:53 +0000 Subject: [PATCH 789/828] feat(order/locally_finite): add lemmas expanding `finset.Icc` (#16996) This avoids the need to unfold `finset.Icc` to access the component-wise definitions. I only really needed this for `prod`, but added a handful of lemmas elsewhere too. --- src/data/dfinsupp/interval.lean | 2 ++ src/data/finsupp/interval.lean | 2 ++ src/data/pi/interval.lean | 2 ++ src/order/locally_finite.lean | 17 +++++++++++++++++ 4 files changed, 23 insertions(+) diff --git a/src/data/dfinsupp/interval.lean b/src/data/dfinsupp/interval.lean index 2516db9bbc19a..97bfab8275d70 100644 --- a/src/data/dfinsupp/interval.lean +++ b/src/data/dfinsupp/interval.lean @@ -152,6 +152,8 @@ locally_finite_order.of_Icc (Π₀ i, α i) variables (f g : Π₀ i, α i) +lemma Icc_eq : Icc f g = (f.support ∪ g.support).dfinsupp (f.range_Icc g) := rfl + lemma card_Icc : (Icc f g).card = ∏ i in f.support ∪ g.support, (Icc (f i) (g i)).card := card_dfinsupp _ _ diff --git a/src/data/finsupp/interval.lean b/src/data/finsupp/interval.lean index e82f01c069ac3..6a35ac884ef7e 100644 --- a/src/data/finsupp/interval.lean +++ b/src/data/finsupp/interval.lean @@ -75,6 +75,8 @@ locally_finite_order.of_Icc (ι →₀ α) exact forall_and_distrib, end) +lemma Icc_eq : Icc f g = (f.support ∪ g.support).finsupp (f.range_Icc g) := rfl + lemma card_Icc : (Icc f g).card = ∏ i in f.support ∪ g.support, (Icc (f i) (g i)).card := card_finsupp _ _ diff --git a/src/data/pi/interval.lean b/src/data/pi/interval.lean index 4b4aace20a8b4..e65aeea6bc910 100644 --- a/src/data/pi/interval.lean +++ b/src/data/pi/interval.lean @@ -28,6 +28,8 @@ locally_finite_order.of_Icc _ variables (a b : Π i, α i) +lemma Icc_eq : Icc a b = pi_finset (λ i, Icc (a i) (b i)) := rfl + lemma card_Icc : (Icc a b).card = ∏ i, (Icc (a i) (b i)).card := card_pi_finset _ lemma card_Ico : (Ico a b).card = (∏ i, (Icc (a i) (b i)).card) - 1 := diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index b9df19815ebb0..aab3206456d4f 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -656,6 +656,8 @@ lemma Iio_of_dual (a : αᵒᵈ) : Iio (of_dual a) = (Ioi a).map of_dual.to_embe end locally_finite_order_top +namespace prod + instance [locally_finite_order α] [locally_finite_order β] [decidable_rel ((≤) : α × β → α × β → Prop)] : locally_finite_order (α × β) := @@ -675,6 +677,21 @@ instance [locally_finite_order_bot α] [locally_finite_order_bot β] locally_finite_order_bot.of_Iic' (α × β) (λ a, Iic a.fst ×ˢ Iic a.snd) (λ a x, by { rw [mem_product, mem_Iic, mem_Iic], refl }) +lemma Icc_eq [locally_finite_order α] [locally_finite_order β] + [decidable_rel ((≤) : α × β → α × β → Prop)] (p q : α × β) : + finset.Icc p q = finset.Icc p.1 q.1 ×ˢ finset.Icc p.2 q.2 := rfl + +@[simp] lemma Icc_mk_mk [locally_finite_order α] [locally_finite_order β] + [decidable_rel ((≤) : α × β → α × β → Prop)] (a₁ a₂ : α) (b₁ b₂ : β) : + finset.Icc (a₁, b₁) (a₂, b₂) = finset.Icc a₁ a₂ ×ˢ finset.Icc b₁ b₂ := rfl + +lemma card_Icc [locally_finite_order α] [locally_finite_order β] + [decidable_rel ((≤) : α × β → α × β → Prop)] (p q : α × β) : + (finset.Icc p q).card = (finset.Icc p.1 q.1).card * (finset.Icc p.2 q.2).card := +finset.card_product _ _ + +end prod + end preorder /-! From 2f5533c9a08e6711f058a7e6aa88e75d58f75255 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 16 Oct 2022 20:42:49 +0000 Subject: [PATCH 790/828] feat(ring_theory/finiteness): Kernel of surjective morphism between finitely presented algebras is fg. (#15969) --- src/algebra/algebra/basic.lean | 2 +- src/ring_theory/finiteness.lean | 74 +++++++++++++++++++++++++++++++++ 2 files changed, 75 insertions(+), 1 deletion(-) diff --git a/src/algebra/algebra/basic.lean b/src/algebra/algebra/basic.lean index c4eda7542b773..e249ea34b8b37 100644 --- a/src/algebra/algebra/basic.lean +++ b/src/algebra/algebra/basic.lean @@ -771,7 +771,7 @@ def comp (φ₁ : B →ₐ[R] C) (φ₂ : A →ₐ[R] B) : A →ₐ[R] C := lemma comp_apply (φ₁ : B →ₐ[R] C) (φ₂ : A →ₐ[R] B) (p : A) : φ₁.comp φ₂ p = φ₁ (φ₂ p) := rfl lemma comp_to_ring_hom (φ₁ : B →ₐ[R] C) (φ₂ : A →ₐ[R] B) : - ⇑(φ₁.comp φ₂ : A →+* C) = (φ₁ : B →+* C).comp ↑φ₂ := rfl + (φ₁.comp φ₂ : A →+* C) = (φ₁ : B →+* C).comp ↑φ₂ := rfl @[simp] theorem comp_id : φ.comp (alg_hom.id R A) = φ := ext $ λ x, rfl diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index 375ca0041060a..3ad61a258c627 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -422,6 +422,80 @@ begin exact equiv ((mv_polynomial_of_finite_presentation hfpA _).quotient hfg) (e.restrict_scalars R) end +open mv_polynomial + +/-- This is used to prove the strictly stronger `ker_fg_of_surjective`. Use it instead. -/ +lemma ker_fg_of_mv_polynomial {n : ℕ} (f : mv_polynomial (fin n) R →ₐ[R] A) + (hf : function.surjective f) (hfp : finite_presentation R A) : f.to_ring_hom.ker.fg := +begin + obtain ⟨m, f', hf', s, hs⟩ := hfp, + let RXn := mv_polynomial (fin n) R, let RXm := mv_polynomial (fin m) R, + have := λ (i : fin n), hf' (f $ X i), + choose g hg, + have := λ (i : fin m), hf (f' $ X i), + choose h hh, + let aeval_h : RXm →ₐ[R] RXn := aeval h, + let g' : fin n → RXn := λ i, X i - aeval_h (g i), + refine ⟨finset.univ.image g' ∪ s.image aeval_h, _⟩, + simp only [finset.coe_image, finset.coe_union, finset.coe_univ, set.image_univ], + have hh' : ∀ x, f (aeval_h x) = f' x, + { intro x, rw [← f.coe_to_ring_hom, map_aeval], simp_rw [alg_hom.coe_to_ring_hom, hh], + rw [alg_hom.comp_algebra_map, ← aeval_eq_eval₂_hom, ← aeval_unique] }, + let s' := set.range g' ∪ aeval_h '' s, + have leI : ideal.span s' ≤ f.to_ring_hom.ker, + { rw ideal.span_le, + rintros _ (⟨i, rfl⟩|⟨x, hx, rfl⟩), + { change f (g' i) = 0, rw [map_sub, ← hg, hh', sub_self], }, + { change f (aeval_h x) = 0, + rw hh', + change x ∈ f'.to_ring_hom.ker, + rw ← hs, + exact ideal.subset_span hx } }, + apply leI.antisymm, + intros x hx, + have : x ∈ aeval_h.range.to_add_submonoid ⊔ (ideal.span s').to_add_submonoid, + { have : x ∈ adjoin R (set.range X : set RXn) := by { rw [adjoin_range_X], trivial }, + apply adjoin_induction this, + { rintros _ ⟨i, rfl⟩, + rw [← sub_add_cancel (X i) (aeval h (g i)), add_comm], + apply add_submonoid.add_mem_sup, + { exact set.mem_range_self _ }, + { apply submodule.subset_span, + apply set.mem_union_left, + exact set.mem_range_self _ } }, + { intros r, apply add_submonoid.mem_sup_left, exact ⟨C r, aeval_C _ _⟩ }, + { intros _ _ h₁ h₂, exact add_mem h₁ h₂ }, + { intros p₁ p₂ h₁ h₂, + obtain ⟨_, ⟨x₁, rfl⟩, y₁, hy₁, rfl⟩ := add_submonoid.mem_sup.mp h₁, + obtain ⟨_, ⟨x₂, rfl⟩, y₂, hy₂, rfl⟩ := add_submonoid.mem_sup.mp h₂, + rw [mul_add, add_mul, add_assoc, ← map_mul], + apply add_submonoid.add_mem_sup, + { exact set.mem_range_self _ }, + { exact add_mem (ideal.mul_mem_right _ _ hy₁) (ideal.mul_mem_left _ _ hy₂) } } }, + obtain ⟨_, ⟨x, rfl⟩, y, hy, rfl⟩ := add_submonoid.mem_sup.mp this, + refine add_mem _ hy, + simp only [ring_hom.mem_ker, alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom, map_add, + (show f y = 0, from leI hy), add_zero, hh'] at hx, + suffices : ideal.span (s : set RXm) ≤ (ideal.span s').comap aeval_h, + { apply this, rwa hs }, + rw ideal.span_le, + intros x hx, + apply submodule.subset_span, + apply set.mem_union_right, + exact set.mem_image_of_mem _ hx +end + +/-- If `f : A →ₐ[R] B` is a sujection between finitely-presented `R`-algebras, then the kernel of +`f` is finitely generated. -/ +lemma ker_fg_of_surjective (f : A →ₐ[R] B) (hf : function.surjective f) + (hRA : finite_presentation R A) (hRB : finite_presentation R B) : f.to_ring_hom.ker.fg := +begin + obtain ⟨n, g, hg, hg'⟩ := hRA, + convert (ker_fg_of_mv_polynomial (f.comp g) (hf.comp hg) hRB).map g.to_ring_hom, + simp_rw [ring_hom.ker_eq_comap_bot, alg_hom.to_ring_hom_eq_coe, alg_hom.comp_to_ring_hom], + rw [← ideal.comap_comap, ideal.map_comap_of_surjective (g : mv_polynomial (fin n) R →+* A) hg], +end + end finite_presentation end algebra From 0c171f291c5e590f368974dc3a293c7b50de44a4 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Mon, 17 Oct 2022 04:40:50 +0000 Subject: [PATCH 791/828] feat(order/upper_lower): Product of upper sets (#17022) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the product of two upper/lower sets as an upper/lower set. Co-authored-by: Yaël Dillies --- src/order/upper_lower.lean | 63 +++++++++++++++++++++++++++++++++++++- 1 file changed, 62 insertions(+), 1 deletion(-) diff --git a/src/order/upper_lower.lean b/src/order/upper_lower.lean index e89c2245b90c0..559ae43dc7afd 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower.lean @@ -39,7 +39,7 @@ Lattice structure on antichains. Order equivalence between upper/lower sets and open order_dual set -variables {α : Type*} {ι : Sort*} {κ : ι → Sort*} +variables {α β : Type*} {ι : Sort*} {κ : ι → Sort*} /-! ### Unbundled upper/lower sets -/ @@ -689,3 +689,64 @@ begin end end closure + +/-! ### Product -/ + +section preorder +variables [preorder α] [preorder β] {s : set α} {t : set β} {x : α × β} + +lemma is_upper_set.prod (hs : is_upper_set s) (ht : is_upper_set t) : is_upper_set (s ×ˢ t) := +λ a b h ha, ⟨hs h.1 ha.1, ht h.2 ha.2⟩ + +lemma is_lower_set.prod (hs : is_lower_set s) (ht : is_lower_set t) : is_lower_set (s ×ˢ t) := +λ a b h ha, ⟨hs h.1 ha.1, ht h.2 ha.2⟩ + +namespace upper_set + +/-- The product of two upper sets as an upper set. -/ +def prod (s : upper_set α) (t : upper_set β) : upper_set (α × β) := ⟨s ×ˢ t, s.2.prod t.2⟩ + +@[simp] lemma coe_prod (s : upper_set α) (t : upper_set β) : (s.prod t : set (α × β)) = s ×ˢ t := +rfl + +@[simp] lemma mem_prod {s : upper_set α} {t : upper_set β} : x ∈ s.prod t ↔ x.1 ∈ s ∧ x.2 ∈ t := +iff.rfl + +lemma Ici_prod (x : α × β) : Ici x = (Ici x.1).prod (Ici x.2) := rfl +@[simp] lemma Ici_prod_Ici (a : α) (b : β) : (Ici a).prod (Ici b) = Ici (a, b) := rfl + +@[simp] lemma bot_prod_bot : (⊥ : upper_set α).prod (⊥ : upper_set β) = ⊥ := ext univ_prod_univ +@[simp] lemma prod_top (s : upper_set α) : s.prod (⊤ : upper_set β) = ⊤ := ext prod_empty +@[simp] lemma top_prod (t : upper_set β) : (⊤ : upper_set α).prod t = ⊤ := ext empty_prod + +end upper_set + +namespace lower_set + +/-- The product of two lower sets as a lower set. -/ +def prod (s : lower_set α) (t : lower_set β) : lower_set (α × β) := ⟨s ×ˢ t, s.2.prod t.2⟩ + +@[simp] lemma coe_prod (s : lower_set α) (t : lower_set β) : (s.prod t : set (α × β)) = s ×ˢ t := +rfl + +@[simp] lemma mem_prod {s : lower_set α} {t : lower_set β} : x ∈ s.prod t ↔ x.1 ∈ s ∧ x.2 ∈ t := +iff.rfl + +lemma Iic_prod (x : α × β) : Iic x = (Iic x.1).prod (Iic x.2) := rfl +@[simp] lemma Ici_prod_Ici (a : α) (b : β) : (Iic a).prod (Iic b) = Iic (a, b) := rfl + +@[simp] lemma prod_bot (s : lower_set α) : s.prod (⊥ : lower_set β) = ⊥ := ext prod_empty +@[simp] lemma bot_prod (t : lower_set β) : (⊥ : lower_set α).prod t = ⊥ := ext empty_prod +@[simp] lemma top_prod_top : (⊤ : lower_set α).prod (⊤ : lower_set β) = ⊤ := ext univ_prod_univ + +end lower_set + +@[simp] lemma upper_closure_prod (s : set α) (t : set β) : + upper_closure (s ×ˢ t) = (upper_closure s).prod (upper_closure t) := +by { ext, simp [prod.le_def, and_and_and_comm _ (_ ∈ t)] } + +@[simp] lemma lower_closure_prod (s : set α) (t : set β) : + lower_closure (s ×ˢ t) = (lower_closure s).prod (lower_closure t) := +by { ext, simp [prod.le_def, and_and_and_comm _ (_ ∈ t)] } + +end preorder From 158724c8970a2d5b3c6b82fff7c82cfdbe9e2b32 Mon Sep 17 00:00:00 2001 From: Thomas Browning Date: Mon, 17 Oct 2022 08:42:08 +0000 Subject: [PATCH 792/828] feat(group_theory/subgroup/basic): Multiplication of disjoint subgroups is injective (#16966) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR proves that if `H₁` and `H₂` are disjoint subgroups of `G`, then the multiplication map `H₁ × H₂ → G` is injective. This is useful for a couple results about complementary subgroups. --- src/group_theory/complement.lean | 21 +++++++++++---------- src/group_theory/subgroup/basic.lean | 18 ++++++++++++++---- 2 files changed, 25 insertions(+), 14 deletions(-) diff --git a/src/group_theory/complement.lean b/src/group_theory/complement.lean index a9f02a2192708..a51fe9db2844f 100644 --- a/src/group_theory/complement.lean +++ b/src/group_theory/complement.lean @@ -27,7 +27,7 @@ In this file we define the complement of a subgroup. - `is_complement_of_coprime` : Subgroups of coprime order are complements. -/ -open_locale big_operators +open_locale big_operators pointwise namespace subgroup @@ -368,18 +368,19 @@ lemma is_complement'.card_mul [fintype G] [fintype H] [fintype K] (h : is_comple fintype.card H * fintype.card K = fintype.card G := h.card_mul +lemma is_complement'_of_disjoint_and_mul_eq_univ + (h1 : disjoint H K) (h2 : ↑H * ↑K = (set.univ : set G)) : is_complement' H K := +begin + refine ⟨mul_injective_of_disjoint h1, λ g, _⟩, + obtain ⟨h, k, hh, hk, hg⟩ := set.eq_univ_iff_forall.mp h2 g, + exact ⟨(⟨h, hh⟩, ⟨k, hk⟩), hg⟩, +end + lemma is_complement'_of_card_mul_and_disjoint [fintype G] [fintype H] [fintype K] (h1 : fintype.card H * fintype.card K = fintype.card G) (h2 : disjoint H K) : is_complement' H K := -begin - refine (fintype.bijective_iff_injective_and_card _).mpr - ⟨λ x y h, _, (fintype.card_prod H K).trans h1⟩, - rw [←eq_inv_mul_iff_mul_eq, ←mul_assoc, ←mul_inv_eq_iff_eq_mul] at h, - change ↑(x.2 * y.2⁻¹) = ↑(x.1⁻¹ * y.1) at h, - rw [prod.ext_iff, ←@inv_mul_eq_one H _ x.1 y.1, ←@mul_inv_eq_one K _ x.2 y.2, subtype.ext_iff, - subtype.ext_iff, coe_one, coe_one, h, and_self, ←mem_bot, ←h2.eq_bot, mem_inf], - exact ⟨subtype.mem ((x.1)⁻¹ * (y.1)), (congr_arg (∈ K) h).mp (subtype.mem (x.2 * (y.2)⁻¹))⟩, -end +(fintype.bijective_iff_injective_and_card _).mpr + ⟨mul_injective_of_disjoint h2, (fintype.card_prod H K).trans h1⟩ lemma is_complement'_iff_card_mul_and_disjoint [fintype G] [fintype H] [fintype K] : is_complement' H K ↔ diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 100188847a13d..b96e900790a6a 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -3112,13 +3112,12 @@ end subgroup_normal @[to_additive] lemma disjoint_def {H₁ H₂ : subgroup G} : disjoint H₁ H₂ ↔ ∀ {x : G}, x ∈ H₁ → x ∈ H₂ → x = 1 := -show (∀ x, x ∈ H₁ ∧ x ∈ H₂ → x ∈ ({1} : set G)) ↔ _, by simp +by simp only [disjoint, set_like.le_def, mem_inf, mem_bot, and_imp] @[to_additive] lemma disjoint_def' {H₁ H₂ : subgroup G} : disjoint H₁ H₂ ↔ ∀ {x y : G}, x ∈ H₁ → y ∈ H₂ → x = y → x = 1 := -disjoint_def.trans ⟨λ h x y hx hy hxy, h hx $ hxy.symm ▸ hy, - λ h x hx hx', h hx hx' rfl⟩ +disjoint_def.trans ⟨λ h x y hx hy hxy, h hx $ hxy.symm ▸ hy, λ h x hx hx', h hx hx' rfl⟩ @[to_additive] lemma disjoint_iff_mul_eq_one {H₁ H₂ : subgroup G} : @@ -3126,7 +3125,18 @@ lemma disjoint_iff_mul_eq_one {H₁ H₂ : subgroup G} : disjoint_def'.trans ⟨λ h x y hx hy hxy, let hx1 : x = 1 := h hx (H₂.inv_mem hy) (eq_inv_iff_mul_eq_one.mpr hxy) in ⟨hx1, by simpa [hx1] using hxy⟩, - λ h x y hx hy hxy, (h hx (H₂.inv_mem hy) (mul_inv_eq_one.mpr hxy)).1 ⟩ + λ h x y hx hy hxy, (h hx (H₂.inv_mem hy) (mul_inv_eq_one.mpr hxy)).1⟩ + +@[to_additive] +lemma mul_injective_of_disjoint {H₁ H₂ : subgroup G} (h : disjoint H₁ H₂) : + function.injective (λ g, g.1 * g.2 : H₁ × H₂ → G) := +begin + intros x y hxy, + rw [←inv_mul_eq_iff_eq_mul, ←mul_assoc, ←mul_inv_eq_one, mul_assoc] at hxy, + replace hxy := disjoint_iff_mul_eq_one.mp h (y.1⁻¹ * x.1).prop (x.2 * y.2⁻¹).prop hxy, + rwa [coe_mul, coe_mul, coe_inv, coe_inv, inv_mul_eq_one, mul_inv_eq_one, + ←subtype.ext_iff, ←subtype.ext_iff, eq_comm, ←prod.ext_iff] at hxy, +end /-- `finset.noncomm_prod` is “injective” in `f` if `f` maps into independent subgroups. This generalizes (one direction of) `subgroup.disjoint_iff_mul_eq_one`. -/ From 5f17f4dabbab9da2673e7711e6cea33ff190819f Mon Sep 17 00:00:00 2001 From: David Kurniadi Angdinata Date: Mon, 17 Oct 2022 11:19:56 +0000 Subject: [PATCH 793/828] feat(algebraic_geometry/prime_spectrum/basic): intersection of localisations (#16860) - [x] depends on: #16905 [define maximal spectrum] - [x] depends on: #16920 [refactor height one spectrum] --- .../prime_spectrum/maximal.lean | 57 +++++++++++++++++++ src/ring_theory/dedekind_domain/ideal.lean | 22 +++++++ src/ring_theory/localization/at_prime.lean | 6 +- .../valuation/valuation_subring.lean | 3 +- 4 files changed, 84 insertions(+), 4 deletions(-) diff --git a/src/algebraic_geometry/prime_spectrum/maximal.lean b/src/algebraic_geometry/prime_spectrum/maximal.lean index 4b2ad313271b9..07d8465365e49 100644 --- a/src/algebraic_geometry/prime_spectrum/maximal.lean +++ b/src/algebraic_geometry/prime_spectrum/maximal.lean @@ -5,6 +5,7 @@ Authors: David Kurniadi Angdinata -/ import algebraic_geometry.prime_spectrum.basic +import ring_theory.localization.as_subring /-! # Maximal spectrum of a commutative ring @@ -73,4 +74,60 @@ instance : t1_space $ maximal_spectrum R := lemma to_prime_spectrum_continuous : continuous $ @to_prime_spectrum R _ := continuous_induced_dom +variables (R) [is_domain R] (K : Type v) [field K] [algebra R K] [is_fraction_ring R K] + +/-- An integral domain is equal to the intersection of its localizations at all its maximal ideals +viewed as subalgebras of its field of fractions. -/ +theorem infi_localization_eq_bot : + (⨅ v : maximal_spectrum R, + localization.subalgebra.of_field K _ v.as_ideal.prime_compl_le_non_zero_divisors) = ⊥ := +begin + ext x, + rw [algebra.mem_bot, algebra.mem_infi], + split, + { apply imp_of_not_imp_not, + intros hrange hlocal, + let denom : ideal R := (submodule.span R {1} : submodule R K).colon (submodule.span R {x}), + have hdenom : (1 : R) ∉ denom := + begin + intro hdenom, + rcases submodule.mem_span_singleton.mp + (submodule.mem_colon.mp hdenom x $ submodule.mem_span_singleton_self x) with ⟨y, hy⟩, + exact hrange ⟨y, by rw [← mul_one $ algebra_map R K y, ← algebra.smul_def, hy, one_smul]⟩ + end, + rcases denom.exists_le_maximal (λ h, (h ▸ hdenom) submodule.mem_top) with ⟨max, hmax, hle⟩, + rcases hlocal ⟨max, hmax⟩ with ⟨n, d, hd, rfl⟩, + apply hd (hle $ submodule.mem_colon.mpr $ λ _ hy, _), + rcases submodule.mem_span_singleton.mp hy with ⟨y, rfl⟩, + exact submodule.mem_span_singleton.mpr + ⟨y * n, by rw [algebra.smul_def, mul_one, map_mul, smul_comm, algebra.smul_def, + algebra.smul_def, mul_comm $ algebra_map R K d, inv_mul_cancel_right₀ $ + (map_ne_zero_iff _ $ no_zero_smul_divisors.algebra_map_injective R K).mpr $ + λ h, (h ▸ hd) max.zero_mem]⟩ }, + { rintro ⟨y, rfl⟩ ⟨v, hv⟩, + exact ⟨y, 1, v.ne_top_iff_one.mp hv.ne_top, by rw [map_one, inv_one, mul_one]⟩ } +end + end maximal_spectrum + +namespace prime_spectrum + +variables (R) [is_domain R] (K : Type v) [field K] [algebra R K] [is_fraction_ring R K] + +/-- An integral domain is equal to the intersection of its localizations at all its prime ideals +viewed as subalgebras of its field of fractions. -/ +theorem infi_localization_eq_bot : + (⨅ v : prime_spectrum R, + localization.subalgebra.of_field K _ $ v.as_ideal.prime_compl_le_non_zero_divisors) = ⊥ := +begin + ext x, + rw [algebra.mem_infi], + split, + { rw [← maximal_spectrum.infi_localization_eq_bot, algebra.mem_infi], + exact λ hx ⟨v, hv⟩, hx ⟨v, hv.is_prime⟩ }, + { rw [algebra.mem_bot], + rintro ⟨y, rfl⟩ ⟨v, hv⟩, + exact ⟨y, 1, v.ne_top_iff_one.mp hv.ne_top, by rw [map_one, inv_one, mul_one]⟩ } +end + +end prime_spectrum diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index c905a1aad74ad..70b2eb9b778e7 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -875,6 +875,28 @@ def equiv_maximal_spectrum (hR : ¬is_field R) : height_one_spectrum R ≃ maxim left_inv := λ ⟨_, _, _⟩, rfl, right_inv := λ ⟨_, _⟩, rfl } +variables (R K) + +/-- A Dedekind domain is equal to the intersection of its localizations at all its height one +non-zero prime ideals viewed as subalgebras of its field of fractions. -/ +theorem infi_localization_eq_bot [algebra R K] [hK : is_fraction_ring R K] : + (⨅ v : height_one_spectrum R, + localization.subalgebra.of_field K _ v.as_ideal.prime_compl_le_non_zero_divisors) = ⊥ := +begin + ext x, + rw [algebra.mem_infi], + split, + by_cases hR : is_field R, + { rcases function.bijective_iff_has_inverse.mp + (is_field.localization_map_bijective (flip non_zero_divisors.ne_zero rfl : 0 ∉ R⁰) hR) + with ⟨algebra_map_inv, _, algebra_map_right_inv⟩, + exact λ _, algebra.mem_bot.mpr ⟨algebra_map_inv x, algebra_map_right_inv x⟩, + exact hK }, + all_goals { rw [← maximal_spectrum.infi_localization_eq_bot, algebra.mem_infi] }, + { exact λ hx ⟨v, hv⟩, hx ((equiv_maximal_spectrum hR).symm ⟨v, hv⟩) }, + { exact λ hx ⟨v, hv, hbot⟩, hx ⟨v, dimension_le_one v hbot hv⟩ } +end + end height_one_spectrum end is_dedekind_domain diff --git a/src/ring_theory/localization/at_prime.lean b/src/ring_theory/localization/at_prime.lean index 49ec3ca61a51f..62cd3aaab6ebd 100644 --- a/src/ring_theory/localization/at_prime.lean +++ b/src/ring_theory/localization/at_prime.lean @@ -44,6 +44,9 @@ def prime_compl : one_mem' := by convert I.ne_top_iff_one.1 hp.1; refl, mul_mem' := λ x y hnx hny hxy, or.cases_on (hp.mem_or_mem hxy) hnx hny } +lemma prime_compl_le_non_zero_divisors [no_zero_divisors R] : I.prime_compl ≤ non_zero_divisors R := +le_non_zero_divisors_of_no_zero_divisors $ not_not_intro I.zero_mem + end ideal variables (S) @@ -115,8 +118,7 @@ The localization of an integral domain at the complement of a prime ideal is an -/ instance is_domain_of_local_at_prime {P : ideal A} (hp : P.is_prime) : is_domain (localization.at_prime P) := -is_domain_localization (le_non_zero_divisors_of_no_zero_divisors - (not_not_intro P.zero_mem)) +is_domain_localization P.prime_compl_le_non_zero_divisors namespace at_prime diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index 94acc78a0b841..63bf36c387b1f 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -211,8 +211,7 @@ instance prime_ideal_of_le (R S : valuation_subring K) (h : R ≤ S) : /-- The coarsening of a valuation ring associated to a prime ideal. -/ def of_prime (A : valuation_subring K) (P : ideal A) [P.is_prime] : valuation_subring K := -of_le A (localization.subalgebra.of_field K P.prime_compl $ - le_non_zero_divisors_of_no_zero_divisors $ not_not_intro P.zero_mem).to_subring $ +of_le A (localization.subalgebra.of_field K _ P.prime_compl_le_non_zero_divisors).to_subring $ λ a ha, subalgebra.algebra_map_mem _ (⟨a, ha⟩ : A) instance of_prime_algebra (A : valuation_subring K) (P : ideal A) [P.is_prime] : From 3c95a7274409e91bd7497010502941f1ee6c3a52 Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Mon, 17 Oct 2022 11:19:57 +0000 Subject: [PATCH 794/828] style(geometry/euclidean): consistently use inner product notation (#16961) Make `geometry.euclidean` files consistently use notation for the inner product, rather than mixing notation with non-notation calls to `inner`. Also always get that notation from `open_locale real_inner_product_space` instead of a local `notation` declaration. --- .../euclidean/angle/unoriented/basic.lean | 20 +++++++++---------- src/geometry/euclidean/basic.lean | 5 ++--- src/geometry/euclidean/triangle.lean | 12 +++++------ 3 files changed, 18 insertions(+), 19 deletions(-) diff --git a/src/geometry/euclidean/angle/unoriented/basic.lean b/src/geometry/euclidean/angle/unoriented/basic.lean index cade9bc78797b..547e019006ef4 100644 --- a/src/geometry/euclidean/angle/unoriented/basic.lean +++ b/src/geometry/euclidean/angle/unoriented/basic.lean @@ -30,7 +30,7 @@ variables {V : Type*} [inner_product_space ℝ V] /-- The undirected angle between two vectors. If either vector is 0, this is π/2. See `orientation.oangle` for the corresponding oriented angle definition. -/ -def angle (x y : V) : ℝ := real.arccos (inner x y / (∥x∥ * ∥y∥)) +def angle (x y : V) : ℝ := real.arccos (⟪x, y⟫ / (∥x∥ * ∥y∥)) lemma continuous_at_angle {x : V × V} (hx1 : x.1 ≠ 0) (hx2 : x.2 ≠ 0) : continuous_at (λ y : V × V, angle y.1 y.2) x := @@ -72,7 +72,7 @@ lemma conformal_at.preserves_angle {E F : Type*} let ⟨f₁, h₁, c⟩ := H in h₁.unique h ▸ is_conformal_map.preserves_angle c u v /-- The cosine of the angle between two vectors. -/ -lemma cos_angle (x y : V) : real.cos (angle x y) = inner x y / (∥x∥ * ∥y∥) := +lemma cos_angle (x y : V) : real.cos (angle x y) = ⟪x, y⟫ / (∥x∥ * ∥y∥) := real.cos_arccos (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).2 @@ -167,7 +167,7 @@ by rw [angle_comm, angle_smul_right_of_neg y x hr, angle_comm] /-- The cosine of the angle between two vectors, multiplied by the product of their norms. -/ -lemma cos_angle_mul_norm_mul_norm (x y : V) : real.cos (angle x y) * (∥x∥ * ∥y∥) = inner x y := +lemma cos_angle_mul_norm_mul_norm (x y : V) : real.cos (angle x y) * (∥x∥ * ∥y∥) = ⟪x, y⟫ := begin rw [cos_angle, div_mul_cancel_of_imp], simp [or_imp_distrib] { contextual := tt }, @@ -176,7 +176,7 @@ end /-- The sine of the angle between two vectors, multiplied by the product of their norms. -/ lemma sin_angle_mul_norm_mul_norm (x y : V) : real.sin (angle x y) * (∥x∥ * ∥y∥) = - real.sqrt (inner x x * inner y y - inner x y * inner x y) := + real.sqrt (⟪x, x⟫ * ⟪y, y⟫ - ⟪x, y⟫ * ⟪x, y⟫) := begin unfold angle, rw [real.sin_arccos (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 @@ -294,8 +294,8 @@ begin rw ← inner_eq_neg_mul_norm_iff_angle_eq_pi hx hy, obtain ⟨hxy₁, hxy₂⟩ := ⟨norm_nonneg (x - y), add_nonneg (norm_nonneg x) (norm_nonneg y)⟩, rw [← sq_eq_sq hxy₁ hxy₂, norm_sub_pow_two_real] at h, - calc inner x y = (∥x∥ ^ 2 + ∥y∥ ^ 2 - (∥x∥ + ∥y∥) ^ 2) / 2 : by linarith - ... = -(∥x∥ * ∥y∥) : by ring, + calc ⟪x, y⟫ = (∥x∥ ^ 2 + ∥y∥ ^ 2 - (∥x∥ + ∥y∥) ^ 2) / 2 : by linarith + ... = -(∥x∥ * ∥y∥) : by ring, end /-- The norm of the sum of two non-zero vectors equals the sum of their norms @@ -307,8 +307,8 @@ begin rw ← inner_eq_mul_norm_iff_angle_eq_zero hx hy, obtain ⟨hxy₁, hxy₂⟩ := ⟨norm_nonneg (x + y), add_nonneg (norm_nonneg x) (norm_nonneg y)⟩, rw [← sq_eq_sq hxy₁ hxy₂, norm_add_pow_two_real] at h, - calc inner x y = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith - ... = ∥x∥ * ∥y∥ : by ring, + calc ⟪x, y⟫ = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith + ... = ∥x∥ * ∥y∥ : by ring, end /-- The norm of the difference of two non-zero vectors equals the absolute value @@ -320,8 +320,8 @@ begin rw ← inner_eq_mul_norm_iff_angle_eq_zero hx hy, have h1 : ∥x - y∥ ^ 2 = (∥x∥ - ∥y∥) ^ 2, { rw h, exact sq_abs (∥x∥ - ∥y∥) }, rw norm_sub_pow_two_real at h1, - calc inner x y = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith - ... = ∥x∥ * ∥y∥ : by ring, + calc ⟪x, y⟫ = ((∥x∥ + ∥y∥) ^ 2 - ∥x∥ ^ 2 - ∥y∥ ^ 2)/ 2 : by linarith + ... = ∥x∥ * ∥y∥ : by ring, end /-- The norm of the sum of two vectors equals the norm of their difference if and only if diff --git a/src/geometry/euclidean/basic.lean b/src/geometry/euclidean/basic.lean index 14862db9dd226..db47bad9c1e2f 100644 --- a/src/geometry/euclidean/basic.lean +++ b/src/geometry/euclidean/basic.lean @@ -57,7 +57,6 @@ Euclidean affine spaces. variables {V : Type*} {P : Type*} [inner_product_space ℝ V] [metric_space P] [normed_add_torsor V P] -local notation `⟪`x`, `y`⟫` := @inner ℝ V _ x y include V /-- The midpoint of the segment AB is the same distance from A as it is from B. -/ @@ -70,7 +69,7 @@ terms of the pairwise distances. -/ lemma inner_weighted_vsub {ι₁ : Type*} {s₁ : finset ι₁} {w₁ : ι₁ → ℝ} (p₁ : ι₁ → P) (h₁ : ∑ i in s₁, w₁ i = 0) {ι₂ : Type*} {s₂ : finset ι₂} {w₂ : ι₂ → ℝ} (p₂ : ι₂ → P) (h₂ : ∑ i in s₂, w₂ i = 0) : - inner (s₁.weighted_vsub p₁ w₁) (s₂.weighted_vsub p₂ w₂) = + ⟪s₁.weighted_vsub p₁ w₁, s₂.weighted_vsub p₂ w₂⟫ = (-∑ i₁ in s₁, ∑ i₂ in s₂, w₁ i₁ * w₂ i₂ * (dist (p₁ i₁) (p₂ i₂) * dist (p₁ i₁) (p₂ i₂))) / 2 := begin @@ -138,7 +137,7 @@ begin ←real_inner_self_eq_norm_mul_norm, sub_self] }, have hvi : ⟪v, v⟫ ≠ 0, by simpa using hv, have hd : discrim ⟪v, v⟫ (2 * ⟪v, p₁ -ᵥ p₂⟫) 0 = - (2 * inner v (p₁ -ᵥ p₂)) * (2 * inner v (p₁ -ᵥ p₂)), + (2 * ⟪v, p₁ -ᵥ p₂⟫) * (2 * ⟪v, p₁ -ᵥ p₂⟫), { rw discrim, ring }, rw [quadratic_eq_zero_iff hvi hd, add_left_neg, zero_div, neg_mul_eq_neg_mul, ←mul_sub_right_distrib, sub_eq_add_neg, ←mul_two, mul_assoc, mul_div_assoc, diff --git a/src/geometry/euclidean/triangle.lean b/src/geometry/euclidean/triangle.lean index 2b229b7e30f40..741ed3f0d7126 100644 --- a/src/geometry/euclidean/triangle.lean +++ b/src/geometry/euclidean/triangle.lean @@ -123,12 +123,12 @@ begin ∥x∥ * ∥y∥ * ∥x - y∥ * ∥x - y∥ = (real.sin (angle x (x - y)) * (∥x∥ * ∥x - y∥)) * (real.sin (angle y (y - x)) * (∥y∥ * ∥x - y∥)), { ring }, - have H2 : ⟪x, x⟫ * (inner x x - inner x y - (inner x y - inner y y)) - - (inner x x - inner x y) * (inner x x - inner x y) = - inner x x * inner y y - inner x y * inner x y, { ring }, - have H3 : ⟪y, y⟫ * (inner y y - inner x y - (inner x y - inner x x)) - - (inner y y - inner x y) * (inner y y - inner x y) = - inner x x * inner y y - inner x y * inner x y, { ring }, + have H2 : ⟪x, x⟫ * (⟪x, x⟫ - ⟪x, y⟫ - (⟪x, y⟫ - ⟪y, y⟫)) - + (⟪x, x⟫ - ⟪x, y⟫) * (⟪x, x⟫ - ⟪x, y⟫) = + ⟪x, x⟫ * ⟪y, y⟫ - ⟪x, y⟫ * ⟪x, y⟫, { ring }, + have H3 : ⟪y, y⟫ * (⟪y, y⟫ - ⟪x, y⟫ - (⟪x, y⟫ - ⟪x, x⟫)) - + (⟪y, y⟫ - ⟪x, y⟫) * (⟪y, y⟫ - ⟪x, y⟫) = + ⟪x, x⟫ * ⟪y, y⟫ - ⟪x, y⟫ * ⟪x, y⟫, { ring }, rw [mul_sub_right_distrib, mul_sub_right_distrib, mul_sub_right_distrib, mul_sub_right_distrib, H1, sin_angle_mul_norm_mul_norm, norm_sub_rev x y, sin_angle_mul_norm_mul_norm, norm_sub_rev y x, inner_sub_left, inner_sub_left, From 2b69cd5eca1cc3c5f4a754abe7dab71111959a9d Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Mon, 17 Oct 2022 11:19:58 +0000 Subject: [PATCH 795/828] refactor(geometry/manifold/cont_mdiff): split everything about mfderiv (#16982) * Remove import `mfderiv` from `cont_mdiff` * Separate everything about `mfderiv` and `mdifferentiable` to a separate file * The reason is that many properties about `mfderiv` are nicely proven once we already know about things like `cont_mdiff_fst` (unless we want to prove `mdifferentiable_fst` separately?) * We could also try to separate out `mdifferentiable` from `mfderiv`. However, this is harder, since proofs about `mdifferentiable` are given using `has_mfderiv_at` * Needed (sort-of) to port more results from the sphere eversion project. --- src/geometry/manifold/cont_mdiff.lean | 603 +--------------- src/geometry/manifold/cont_mdiff_map.lean | 2 +- src/geometry/manifold/cont_mdiff_mfderiv.lean | 643 ++++++++++++++++++ src/geometry/manifold/mfderiv.lean | 5 - .../smooth_manifold_with_corners.lean | 9 +- 5 files changed, 653 insertions(+), 609 deletions(-) create mode 100644 src/geometry/manifold/cont_mdiff_mfderiv.lean diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index b9202c2bced10..b0d7d42e4d820 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -3,8 +3,7 @@ Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ - -import geometry.manifold.mfderiv +import geometry.manifold.smooth_manifold_with_corners import geometry.manifold.local_invariant_properties /-! @@ -24,10 +23,6 @@ Let `M ` and `M'` be two smooth manifolds, with respect to model with corners `I * `cont_mdiff_on I I' n f s` states that the function `f` is `Cⁿ` on the set `s` * `cont_mdiff I I' n f` states that the function `f` is `Cⁿ`. * `cont_mdiff_on.comp` gives the invariance of the `Cⁿ` property under composition -* `cont_mdiff_on.cont_mdiff_on_tangent_map_within` states that the bundled derivative - of a `Cⁿ` function in a domain is `Cᵐ` when `m + 1 ≤ n`. -* `cont_mdiff.cont_mdiff_tangent_map` states that the bundled derivative - of a `Cⁿ` function is `Cᵐ` when `m + 1 ≤ n`. * `cont_mdiff_iff_cont_diff` states that, for functions between vector spaces, manifold-smoothness is equivalent to usual smoothness. @@ -641,43 +636,6 @@ lemma cont_mdiff.continuous (hf : cont_mdiff I I' n f) : continuous f := continuous_iff_continuous_at.2 $ λ x, (hf x).continuous_at -/-! ### Deducing differentiability from smoothness -/ - -lemma cont_mdiff_within_at.mdifferentiable_within_at - (hf : cont_mdiff_within_at I I' n f s x) (hn : 1 ≤ n) : - mdifferentiable_within_at I I' f s x := -begin - suffices h : mdifferentiable_within_at I I' f (s ∩ (f ⁻¹' (ext_chart_at I' (f x)).source)) x, - { rwa mdifferentiable_within_at_inter' at h, - apply (hf.1).preimage_mem_nhds_within, - exact is_open.mem_nhds (ext_chart_at_open_source I' (f x)) (mem_ext_chart_source I' (f x)) }, - rw mdifferentiable_within_at_iff, - exact ⟨hf.1.mono (inter_subset_left _ _), - (hf.2.differentiable_within_at hn).mono (by mfld_set_tac)⟩, -end - -lemma cont_mdiff_at.mdifferentiable_at (hf : cont_mdiff_at I I' n f x) (hn : 1 ≤ n) : - mdifferentiable_at I I' f x := -mdifferentiable_within_at_univ.1 $ cont_mdiff_within_at.mdifferentiable_within_at hf hn - -lemma cont_mdiff_on.mdifferentiable_on (hf : cont_mdiff_on I I' n f s) (hn : 1 ≤ n) : - mdifferentiable_on I I' f s := -λ x hx, (hf x hx).mdifferentiable_within_at hn - -lemma cont_mdiff.mdifferentiable (hf : cont_mdiff I I' n f) (hn : 1 ≤ n) : - mdifferentiable I I' f := -λ x, (hf x).mdifferentiable_at hn - -lemma smooth.mdifferentiable (hf : smooth I I' f) : mdifferentiable I I' f := -cont_mdiff.mdifferentiable hf le_top - -lemma smooth.mdifferentiable_at (hf : smooth I I' f) : mdifferentiable_at I I' f x := -hf.mdifferentiable x - -lemma smooth.mdifferentiable_within_at (hf : smooth I I' f) : - mdifferentiable_within_at I I' f s x := -hf.mdifferentiable_at.mdifferentiable_within_at - /-! ### `C^∞` smoothness -/ lemma cont_mdiff_within_at_top : @@ -1201,565 +1159,6 @@ alias cont_mdiff_iff_cont_diff ↔ end module -/-! ### The tangent map of a smooth function is smooth -/ - -section tangent_map - -/-- If a function is `C^n` with `1 ≤ n` on a domain with unique derivatives, then its bundled -derivative is continuous. In this auxiliary lemma, we prove this fact when the source and target -space are model spaces in models with corners. The general fact is proved in -`cont_mdiff_on.continuous_on_tangent_map_within`-/ -lemma cont_mdiff_on.continuous_on_tangent_map_within_aux - {f : H → H'} {s : set H} - (hf : cont_mdiff_on I I' n f s) (hn : 1 ≤ n) (hs : unique_mdiff_on I s) : - continuous_on (tangent_map_within I I' f s) ((tangent_bundle.proj I H) ⁻¹' s) := -begin - suffices h : continuous_on (λ (p : H × E), (f p.fst, - (fderiv_within 𝕜 (written_in_ext_chart_at I I' p.fst f) (I.symm ⁻¹' s ∩ range I) - ((ext_chart_at I p.fst) p.fst) : E →L[𝕜] E') p.snd)) (prod.fst ⁻¹' s), - { have A := (tangent_bundle_model_space_homeomorph H I).continuous, - rw continuous_iff_continuous_on_univ at A, - have B := ((tangent_bundle_model_space_homeomorph H' I').symm.continuous.comp_continuous_on h) - .comp' A, - have : (univ ∩ ⇑(tangent_bundle_model_space_homeomorph H I) ⁻¹' (prod.fst ⁻¹' s)) = - tangent_bundle.proj I H ⁻¹' s, - by { ext ⟨x, v⟩, simp only with mfld_simps }, - rw this at B, - apply B.congr, - rintros ⟨x, v⟩ hx, - dsimp [tangent_map_within], - ext, { refl }, - simp only with mfld_simps, - apply congr_fun, - apply congr_arg, - rw mdifferentiable_within_at.mfderiv_within (hf.mdifferentiable_on hn x hx), - refl }, - suffices h : continuous_on (λ (p : H × E), (fderiv_within 𝕜 (I' ∘ f ∘ I.symm) - (I.symm ⁻¹' s ∩ range I) (I p.fst) : E →L[𝕜] E') p.snd) (prod.fst ⁻¹' s), - { dsimp [written_in_ext_chart_at, ext_chart_at], - apply continuous_on.prod - (continuous_on.comp hf.continuous_on continuous_fst.continuous_on (subset.refl _)), - apply h.congr, - assume p hp, - refl }, - suffices h : continuous_on (fderiv_within 𝕜 (I' ∘ f ∘ I.symm) - (I.symm ⁻¹' s ∩ range I)) (I '' s), - { have C := continuous_on.comp h I.continuous_to_fun.continuous_on (subset.refl _), - have A : continuous (λq : (E →L[𝕜] E') × E, q.1 q.2) := - is_bounded_bilinear_map_apply.continuous, - have B : continuous_on (λp : H × E, - (fderiv_within 𝕜 (I' ∘ f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) - (I p.1), p.2)) (prod.fst ⁻¹' s), - { apply continuous_on.prod _ continuous_snd.continuous_on, - refine (continuous_on.comp C continuous_fst.continuous_on _ : _), - exact preimage_mono (subset_preimage_image _ _) }, - exact A.comp_continuous_on B }, - rw cont_mdiff_on_iff at hf, - let x : H := I.symm (0 : E), - let y : H' := I'.symm (0 : E'), - have A := hf.2 x y, - simp only [I.image_eq, inter_comm] with mfld_simps at A ⊢, - apply A.continuous_on_fderiv_within _ hn, - convert hs.unique_diff_on_target_inter x using 1, - simp only [inter_comm] with mfld_simps -end - -/-- If a function is `C^n` on a domain with unique derivatives, then its bundled derivative is -`C^m` when `m+1 ≤ n`. In this auxiliary lemma, we prove this fact when the source and target space -are model spaces in models with corners. The general fact is proved in -`cont_mdiff_on.cont_mdiff_on_tangent_map_within` -/ -lemma cont_mdiff_on.cont_mdiff_on_tangent_map_within_aux - {f : H → H'} {s : set H} - (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : - cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) - ((tangent_bundle.proj I H) ⁻¹' s) := -begin - have m_le_n : m ≤ n, - { apply le_trans _ hmn, - have : m + 0 ≤ m + 1 := add_le_add_left (zero_le _) _, - simpa only [add_zero] using this }, - have one_le_n : 1 ≤ n, - { apply le_trans _ hmn, - change 0 + 1 ≤ m + 1, - exact add_le_add_right (zero_le _) _ }, - have U': unique_diff_on 𝕜 (range I ∩ I.symm ⁻¹' s), - { assume y hy, - simpa only [unique_mdiff_on, unique_mdiff_within_at, hy.1, inter_comm] with mfld_simps - using hs (I.symm y) hy.2 }, - rw cont_mdiff_on_iff, - refine ⟨hf.continuous_on_tangent_map_within_aux one_le_n hs, λp q, _⟩, - have A : range I ×ˢ univ ∩ - ((equiv.sigma_equiv_prod H E).symm ∘ λ (p : E × E), ((I.symm) p.fst, p.snd)) ⁻¹' - (tangent_bundle.proj I H ⁻¹' s) - = (range I ∩ I.symm ⁻¹' s) ×ˢ univ, - by { ext ⟨x, v⟩, simp only with mfld_simps }, - suffices h : cont_diff_on 𝕜 m (((λ (p : H' × E'), (I' p.fst, p.snd)) ∘ - (equiv.sigma_equiv_prod H' E')) ∘ tangent_map_within I I' f s ∘ - ((equiv.sigma_equiv_prod H E).symm) ∘ λ (p : E × E), (I.symm p.fst, p.snd)) - ((range ⇑I ∩ ⇑(I.symm) ⁻¹' s) ×ˢ univ), - by simpa [A] using h, - change cont_diff_on 𝕜 m (λ (p : E × E), - ((I' (f (I.symm p.fst)), ((mfderiv_within I I' f s (I.symm p.fst)) : E → E') p.snd) : E' × E')) - ((range I ∩ I.symm ⁻¹' s) ×ˢ univ), - -- check that all bits in this formula are `C^n` - have hf' := cont_mdiff_on_iff.1 hf, - have A : cont_diff_on 𝕜 m (I' ∘ f ∘ I.symm) (range I ∩ I.symm ⁻¹' s) := - by simpa only with mfld_simps using (hf'.2 (I.symm 0) (I'.symm 0)).of_le m_le_n, - have B : cont_diff_on 𝕜 m ((I' ∘ f ∘ I.symm) ∘ prod.fst) - ((range I ∩ I.symm ⁻¹' s) ×ˢ univ) := - A.comp (cont_diff_fst.cont_diff_on) (prod_subset_preimage_fst _ _), - suffices C : cont_diff_on 𝕜 m (λ (p : E × E), - ((fderiv_within 𝕜 (I' ∘ f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) p.1 : _) p.2)) - ((range I ∩ I.symm ⁻¹' s) ×ˢ univ), - { apply cont_diff_on.prod B _, - apply C.congr (λp hp, _), - simp only with mfld_simps at hp, - simp only [mfderiv_within, hf.mdifferentiable_on one_le_n _ hp.2, hp.1, dif_pos] - with mfld_simps }, - have D : cont_diff_on 𝕜 m (λ x, - (fderiv_within 𝕜 (I' ∘ f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) x)) - (range I ∩ I.symm ⁻¹' s), - { have : cont_diff_on 𝕜 n (I' ∘ f ∘ I.symm) (range I ∩ I.symm ⁻¹' s) := - by simpa only with mfld_simps using (hf'.2 (I.symm 0) (I'.symm 0)), - simpa only [inter_comm] using this.fderiv_within U' hmn }, - have := D.comp (cont_diff_fst.cont_diff_on) (prod_subset_preimage_fst _ _), - have := cont_diff_on.prod this (cont_diff_snd.cont_diff_on), - exact is_bounded_bilinear_map_apply.cont_diff.comp_cont_diff_on this, -end - -include Is I's - -/-- If a function is `C^n` on a domain with unique derivatives, then its bundled derivative -is `C^m` when `m+1 ≤ n`. -/ -theorem cont_mdiff_on.cont_mdiff_on_tangent_map_within - (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : - cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) - ((tangent_bundle.proj I M) ⁻¹' s) := -begin - /- The strategy of the proof is to avoid unfolding the definitions, and reduce by functoriality - to the case of functions on the model spaces, where we have already proved the result. - Let `l` and `r` be the charts to the left and to the right, so that we have - ``` - l^{-1} f r - H --------> M ---> M' ---> H' - ``` - Then the tangent map `T(r ∘ f ∘ l)` is smooth by a previous result. Consider the composition - ``` - Tl T(r ∘ f ∘ l^{-1}) Tr^{-1} - TM -----> TH -------------------> TH' ---------> TM' - ``` - where `Tr^{-1}` and `Tl` are the tangent maps of `r^{-1}` and `l`. Writing `Tl` and `Tr^{-1}` as - composition of charts (called `Dl` and `il` for `l` and `Dr` and `ir` in the proof below), it - follows that they are smooth. The composition of all these maps is `Tf`, and is therefore smooth - as a composition of smooth maps. - -/ - have m_le_n : m ≤ n, - { apply le_trans _ hmn, - have : m + 0 ≤ m + 1 := add_le_add_left (zero_le _) _, - simpa only [add_zero] }, - have one_le_n : 1 ≤ n, - { apply le_trans _ hmn, - change 0 + 1 ≤ m + 1, - exact add_le_add_right (zero_le _) _ }, - /- First step: local reduction on the space, to a set `s'` which is contained in chart domains. -/ - refine cont_mdiff_on_of_locally_cont_mdiff_on (λp hp, _), - have hf' := cont_mdiff_on_iff.1 hf, - simp [tangent_bundle.proj] at hp, - let l := chart_at H p.1, - set Dl := chart_at (model_prod H E) p with hDl, - let r := chart_at H' (f p.1), - let Dr := chart_at (model_prod H' E') (tangent_map_within I I' f s p), - let il := chart_at (model_prod H E) (tangent_map I I l p), - let ir := chart_at (model_prod H' E') (tangent_map I I' (r ∘ f) p), - let s' := f ⁻¹' r.source ∩ s ∩ l.source, - let s'_lift := (tangent_bundle.proj I M)⁻¹' s', - let s'l := l.target ∩ l.symm ⁻¹' s', - let s'l_lift := (tangent_bundle.proj I H) ⁻¹' s'l, - rcases continuous_on_iff'.1 hf'.1 r.source r.open_source with ⟨o, o_open, ho⟩, - suffices h : cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) s'_lift, - { refine ⟨(tangent_bundle.proj I M)⁻¹' (o ∩ l.source), _, _, _⟩, - show is_open ((tangent_bundle.proj I M)⁻¹' (o ∩ l.source)), from - (is_open.inter o_open l.open_source).preimage (tangent_bundle_proj_continuous _ _) , - show p ∈ tangent_bundle.proj I M ⁻¹' (o ∩ l.source), - { simp [tangent_bundle.proj] at ⊢, - have : p.1 ∈ f ⁻¹' r.source ∩ s, by simp [hp], - rw ho at this, - exact this.1 }, - { have : tangent_bundle.proj I M ⁻¹' s ∩ tangent_bundle.proj I M ⁻¹' (o ∩ l.source) = s'_lift, - { dsimp only [s'_lift, s'], rw [ho], mfld_set_tac }, - rw this, - exact h } }, - /- Second step: check that all functions are smooth, and use the chain rule to write the bundled - derivative as a composition of a function between model spaces and of charts. - Convention: statements about the differentiability of `a ∘ b ∘ c` are named `diff_abc`. Statements - about differentiability in the bundle have a `_lift` suffix. -/ - have U' : unique_mdiff_on I s', - { apply unique_mdiff_on.inter _ l.open_source, - rw [ho, inter_comm], - exact hs.inter o_open }, - have U'l : unique_mdiff_on I s'l := - U'.unique_mdiff_on_preimage (mdifferentiable_chart _ _), - have diff_f : cont_mdiff_on I I' n f s' := - hf.mono (by mfld_set_tac), - have diff_r : cont_mdiff_on I' I' n r r.source := - cont_mdiff_on_chart, - have diff_rf : cont_mdiff_on I I' n (r ∘ f) s', - { apply cont_mdiff_on.comp diff_r diff_f (λx hx, _), - simp only [s'] with mfld_simps at hx, simp only [hx] with mfld_simps }, - have diff_l : cont_mdiff_on I I n l.symm s'l, - { have A : cont_mdiff_on I I n l.symm l.target := - cont_mdiff_on_chart_symm, - exact A.mono (by mfld_set_tac) }, - have diff_rfl : cont_mdiff_on I I' n (r ∘ f ∘ l.symm) s'l, - { apply cont_mdiff_on.comp diff_rf diff_l, - mfld_set_tac }, - have diff_rfl_lift : cont_mdiff_on I.tangent I'.tangent m - (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l) s'l_lift := - diff_rfl.cont_mdiff_on_tangent_map_within_aux hmn U'l, - have diff_irrfl_lift : cont_mdiff_on I.tangent I'.tangent m - (ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l)) s'l_lift, - { have A : cont_mdiff_on I'.tangent I'.tangent m ir ir.source := cont_mdiff_on_chart, - exact cont_mdiff_on.comp A diff_rfl_lift (λp hp, by simp only [ir] with mfld_simps) }, - have diff_Drirrfl_lift : cont_mdiff_on I.tangent I'.tangent m - (Dr.symm ∘ (ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l))) s'l_lift, - { have A : cont_mdiff_on I'.tangent I'.tangent m Dr.symm Dr.target := - cont_mdiff_on_chart_symm, - apply cont_mdiff_on.comp A diff_irrfl_lift (λp hp, _), - simp only [s'l_lift, tangent_bundle.proj] with mfld_simps at hp, - simp only [ir, @local_equiv.refl_coe (model_prod H' E'), hp] with mfld_simps }, - -- conclusion of this step: the composition of all the maps above is smooth - have diff_DrirrflilDl : cont_mdiff_on I.tangent I'.tangent m - (Dr.symm ∘ (ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l)) ∘ - (il.symm ∘ Dl)) s'_lift, - { have A : cont_mdiff_on I.tangent I.tangent m Dl Dl.source := cont_mdiff_on_chart, - have A' : cont_mdiff_on I.tangent I.tangent m Dl s'_lift, - { apply A.mono (λp hp, _), - simp only [s'_lift, tangent_bundle.proj] with mfld_simps at hp, - simp only [Dl, hp] with mfld_simps }, - have B : cont_mdiff_on I.tangent I.tangent m il.symm il.target := - cont_mdiff_on_chart_symm, - have C : cont_mdiff_on I.tangent I.tangent m (il.symm ∘ Dl) s'_lift := - cont_mdiff_on.comp B A' (λp hp, by simp only [il] with mfld_simps), - apply cont_mdiff_on.comp diff_Drirrfl_lift C (λp hp, _), - simp only [s'_lift, tangent_bundle.proj] with mfld_simps at hp, - simp only [il, s'l_lift, hp, tangent_bundle.proj] with mfld_simps }, - /- Third step: check that the composition of all the maps indeed coincides with the derivative we - are looking for -/ - have eq_comp : ∀q ∈ s'_lift, tangent_map_within I I' f s q = - (Dr.symm ∘ ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l) ∘ - (il.symm ∘ Dl)) q, - { assume q hq, - simp only [s'_lift, tangent_bundle.proj] with mfld_simps at hq, - have U'q : unique_mdiff_within_at I s' q.1, - by { apply U', simp only [hq, s'] with mfld_simps }, - have U'lq : unique_mdiff_within_at I s'l (Dl q).1, - by { apply U'l, simp only [hq, s'l] with mfld_simps }, - have A : tangent_map_within I I' ((r ∘ f) ∘ l.symm) s'l (il.symm (Dl q)) = - tangent_map_within I I' (r ∘ f) s' (tangent_map_within I I l.symm s'l (il.symm (Dl q))), - { refine tangent_map_within_comp_at (il.symm (Dl q)) _ _ (λp hp, _) U'lq, - { apply diff_rf.mdifferentiable_on one_le_n, - simp only [hq] with mfld_simps }, - { apply diff_l.mdifferentiable_on one_le_n, - simp only [s'l, hq] with mfld_simps }, - { simp only with mfld_simps at hp, simp only [hp] with mfld_simps } }, - have B : tangent_map_within I I l.symm s'l (il.symm (Dl q)) = q, - { have : tangent_map_within I I l.symm s'l (il.symm (Dl q)) - = tangent_map I I l.symm (il.symm (Dl q)), - { refine tangent_map_within_eq_tangent_map U'lq _, - refine mdifferentiable_at_atlas_symm _ (chart_mem_atlas _ _) _, - simp only [hq] with mfld_simps }, - rw [this, tangent_map_chart_symm, hDl], - { simp only [hq] with mfld_simps, - have : q ∈ (chart_at (model_prod H E) p).source, by simp only [hq] with mfld_simps, - exact (chart_at (model_prod H E) p).left_inv this }, - { simp only [hq] with mfld_simps } }, - have C : tangent_map_within I I' (r ∘ f) s' q - = tangent_map_within I' I' r r.source (tangent_map_within I I' f s' q), - { refine tangent_map_within_comp_at q _ _ (λr hr, _) U'q, - { apply diff_r.mdifferentiable_on one_le_n, - simp only [hq] with mfld_simps }, - { apply diff_f.mdifferentiable_on one_le_n, - simp only [hq] with mfld_simps }, - { simp only [s'] with mfld_simps at hr, - simp only [hr] with mfld_simps } }, - have D : Dr.symm (ir (tangent_map_within I' I' r r.source (tangent_map_within I I' f s' q))) - = tangent_map_within I I' f s' q, - { have A : tangent_map_within I' I' r r.source (tangent_map_within I I' f s' q) = - tangent_map I' I' r (tangent_map_within I I' f s' q), - { apply tangent_map_within_eq_tangent_map, - { apply is_open.unique_mdiff_within_at _ r.open_source, simp [hq] }, - { refine mdifferentiable_at_atlas _ (chart_mem_atlas _ _) _, - simp only [hq] with mfld_simps } }, - have : f p.1 = (tangent_map_within I I' f s p).1 := rfl, - rw [A], - dsimp [r, Dr], - rw [this, tangent_map_chart], - { simp only [hq] with mfld_simps, - have : tangent_map_within I I' f s' q ∈ - (chart_at (model_prod H' E') (tangent_map_within I I' f s p)).source, - by simp only [hq] with mfld_simps, - exact (chart_at (model_prod H' E') (tangent_map_within I I' f s p)).left_inv this }, - { simp only [hq] with mfld_simps } }, - have E : tangent_map_within I I' f s' q = tangent_map_within I I' f s q, - { refine tangent_map_within_subset (by mfld_set_tac) U'q _, - apply hf.mdifferentiable_on one_le_n, - simp only [hq] with mfld_simps }, - simp only [(∘), A, B, C, D, E.symm] }, - exact diff_DrirrflilDl.congr eq_comp, -end - -/-- If a function is `C^n` on a domain with unique derivatives, with `1 ≤ n`, then its bundled -derivative is continuous there. -/ -theorem cont_mdiff_on.continuous_on_tangent_map_within - (hf : cont_mdiff_on I I' n f s) (hmn : 1 ≤ n) (hs : unique_mdiff_on I s) : - continuous_on (tangent_map_within I I' f s) ((tangent_bundle.proj I M) ⁻¹' s) := -begin - have : cont_mdiff_on I.tangent I'.tangent 0 (tangent_map_within I I' f s) - ((tangent_bundle.proj I M) ⁻¹' s) := - hf.cont_mdiff_on_tangent_map_within hmn hs, - exact this.continuous_on -end - -/-- If a function is `C^n`, then its bundled derivative is `C^m` when `m+1 ≤ n`. -/ -theorem cont_mdiff.cont_mdiff_tangent_map - (hf : cont_mdiff I I' n f) (hmn : m + 1 ≤ n) : - cont_mdiff I.tangent I'.tangent m (tangent_map I I' f) := -begin - rw ← cont_mdiff_on_univ at hf ⊢, - convert hf.cont_mdiff_on_tangent_map_within hmn unique_mdiff_on_univ, - rw tangent_map_within_univ -end - -/-- If a function is `C^n`, with `1 ≤ n`, then its bundled derivative is continuous. -/ -theorem cont_mdiff.continuous_tangent_map - (hf : cont_mdiff I I' n f) (hmn : 1 ≤ n) : - continuous (tangent_map I I' f) := -begin - rw ← cont_mdiff_on_univ at hf, - rw continuous_iff_continuous_on_univ, - convert hf.continuous_on_tangent_map_within hmn unique_mdiff_on_univ, - rw tangent_map_within_univ -end - -end tangent_map - -/-! ### Smoothness of the projection in a basic smooth bundle -/ - -namespace basic_smooth_vector_bundle_core - -variables (Z : basic_smooth_vector_bundle_core I M E') - -/-- A version of `cont_mdiff_at_iff_target` when the codomain is the total space of - a `basic_smooth_vector_bundle_core`. The continuity condition in the RHS is weaker. -/ -lemma cont_mdiff_at_iff_target {f : N → Z.to_topological_vector_bundle_core.total_space} - {x : N} {n : ℕ∞} : - cont_mdiff_at J (I.prod 𝓘(𝕜, E')) n f x ↔ continuous_at (bundle.total_space.proj ∘ f) x ∧ - cont_mdiff_at J 𝓘(𝕜, E × E') n (ext_chart_at (I.prod 𝓘(𝕜, E')) (f x) ∘ f) x := -begin - let Z' := Z.to_topological_vector_bundle_core, - rw [cont_mdiff_at_iff_target, and.congr_left_iff], - refine λ hf, ⟨λ h, Z'.continuous_proj.continuous_at.comp h, λ h, _⟩, - exact (Z'.local_triv ⟨chart_at _ (f x).1, chart_mem_atlas _ _⟩).to_fiber_bundle_trivialization - .continuous_at_of_comp_left h (mem_chart_source _ _) (h.prod hf.continuous_at.snd) -end - -lemma cont_mdiff_proj : - cont_mdiff (I.prod 𝓘(𝕜, E')) I n Z.to_topological_vector_bundle_core.proj := -begin - assume x, - rw [cont_mdiff_at, cont_mdiff_within_at_iff'], - refine ⟨Z.to_topological_vector_bundle_core.continuous_proj.continuous_within_at, _⟩, - simp only [(∘), chart_at, chart] with mfld_simps, - apply cont_diff_within_at_fst.congr, - { rintros ⟨a, b⟩ hab, - simp only with mfld_simps at hab, - simp only [hab] with mfld_simps }, - { simp only with mfld_simps } -end - -lemma smooth_proj : - smooth (I.prod 𝓘(𝕜, E')) I Z.to_topological_vector_bundle_core.proj := -cont_mdiff_proj Z - -lemma cont_mdiff_on_proj {s : set (Z.to_topological_vector_bundle_core.total_space)} : - cont_mdiff_on (I.prod 𝓘(𝕜, E')) I n - Z.to_topological_vector_bundle_core.proj s := -Z.cont_mdiff_proj.cont_mdiff_on - -lemma smooth_on_proj {s : set (Z.to_topological_vector_bundle_core.total_space)} : - smooth_on (I.prod 𝓘(𝕜, E')) I Z.to_topological_vector_bundle_core.proj s := -cont_mdiff_on_proj Z - -lemma cont_mdiff_at_proj {p : Z.to_topological_vector_bundle_core.total_space} : - cont_mdiff_at (I.prod 𝓘(𝕜, E')) I n - Z.to_topological_vector_bundle_core.proj p := -Z.cont_mdiff_proj.cont_mdiff_at - -lemma smooth_at_proj {p : Z.to_topological_vector_bundle_core.total_space} : - smooth_at (I.prod 𝓘(𝕜, E')) I Z.to_topological_vector_bundle_core.proj p := -Z.cont_mdiff_at_proj - -lemma cont_mdiff_within_at_proj - {s : set (Z.to_topological_vector_bundle_core.total_space)} - {p : Z.to_topological_vector_bundle_core.total_space} : - cont_mdiff_within_at (I.prod 𝓘(𝕜, E')) I n - Z.to_topological_vector_bundle_core.proj s p := -Z.cont_mdiff_at_proj.cont_mdiff_within_at - -lemma smooth_within_at_proj - {s : set (Z.to_topological_vector_bundle_core.total_space)} - {p : Z.to_topological_vector_bundle_core.total_space} : - smooth_within_at (I.prod 𝓘(𝕜, E')) I - Z.to_topological_vector_bundle_core.proj s p := -Z.cont_mdiff_within_at_proj - -/-- If an element of `E'` is invariant under all coordinate changes, then one can define a -corresponding section of the fiber bundle, which is smooth. This applies in particular to the -zero section of a vector bundle. Another example (not yet defined) would be the identity -section of the endomorphism bundle of a vector bundle. -/ -lemma smooth_const_section (v : E') - (h : ∀ (i j : atlas H M), ∀ x ∈ i.1.source ∩ j.1.source, Z.coord_change i j (i.1 x) v = v) : - smooth I (I.prod 𝓘(𝕜, E')) - (show M → Z.to_topological_vector_bundle_core.total_space, from λ x, ⟨x, v⟩) := -begin - assume x, - rw [cont_mdiff_at, cont_mdiff_within_at_iff'], - split, - { apply continuous.continuous_within_at, - apply topological_fiber_bundle_core.continuous_const_section, - assume i j y hy, - exact h _ _ _ hy }, - { have : cont_diff 𝕜 ⊤ (λ (y : E), (y, v)) := cont_diff_id.prod cont_diff_const, - apply this.cont_diff_within_at.congr, - { assume y hy, - simp only with mfld_simps at hy, - simp only [chart, hy, chart_at, prod.mk.inj_iff, to_topological_vector_bundle_core] - with mfld_simps, - apply h, - simp only [hy, subtype.val_eq_coe] with mfld_simps, - exact mem_chart_source H (((chart_at H x).symm) ((model_with_corners.symm I) y)) }, - { simp only [chart, chart_at, prod.mk.inj_iff, to_topological_vector_bundle_core] - with mfld_simps, - apply h, - simp only [subtype.val_eq_coe] with mfld_simps, - exact mem_chart_source H x, } } -end - -end basic_smooth_vector_bundle_core - -/-! ### Smoothness of the tangent bundle projection -/ - -namespace tangent_bundle - -include Is - -lemma cont_mdiff_proj : - cont_mdiff I.tangent I n (proj I M) := -basic_smooth_vector_bundle_core.cont_mdiff_proj _ - -lemma smooth_proj : smooth I.tangent I (proj I M) := -basic_smooth_vector_bundle_core.smooth_proj _ - -lemma cont_mdiff_on_proj {s : set (tangent_bundle I M)} : - cont_mdiff_on I.tangent I n (proj I M) s := -basic_smooth_vector_bundle_core.cont_mdiff_on_proj _ - -lemma smooth_on_proj {s : set (tangent_bundle I M)} : - smooth_on I.tangent I (proj I M) s := -basic_smooth_vector_bundle_core.smooth_on_proj _ - -lemma cont_mdiff_at_proj {p : tangent_bundle I M} : - cont_mdiff_at I.tangent I n - (proj I M) p := -basic_smooth_vector_bundle_core.cont_mdiff_at_proj _ - -lemma smooth_at_proj {p : tangent_bundle I M} : - smooth_at I.tangent I (proj I M) p := -basic_smooth_vector_bundle_core.smooth_at_proj _ - -lemma cont_mdiff_within_at_proj - {s : set (tangent_bundle I M)} {p : tangent_bundle I M} : - cont_mdiff_within_at I.tangent I n - (proj I M) s p := -basic_smooth_vector_bundle_core.cont_mdiff_within_at_proj _ - -lemma smooth_within_at_proj - {s : set (tangent_bundle I M)} {p : tangent_bundle I M} : - smooth_within_at I.tangent I - (proj I M) s p := -basic_smooth_vector_bundle_core.smooth_within_at_proj _ - -variables (I M) -/-- The zero section of the tangent bundle -/ -def zero_section : M → tangent_bundle I M := λ x, ⟨x, 0⟩ -variables {I M} - -lemma smooth_zero_section : smooth I I.tangent (zero_section I M) := -begin - apply basic_smooth_vector_bundle_core.smooth_const_section (tangent_bundle_core I M) 0, - assume i j x hx, - simp only [tangent_bundle_core, continuous_linear_map.map_zero, continuous_linear_map.coe_coe] - with mfld_simps, -end - -open bundle - -/-- The derivative of the zero section of the tangent bundle maps `⟨x, v⟩` to `⟨⟨x, 0⟩, ⟨v, 0⟩⟩`. - -Note that, as currently framed, this is a statement in coordinates, thus reliant on the choice -of the coordinate system we use on the tangent bundle. - -However, the result itself is coordinate-dependent only to the extent that the coordinates -determine a splitting of the tangent bundle. Moreover, there is a canonical splitting at each -point of the zero section (since there is a canonical horizontal space there, the tangent space -to the zero section, in addition to the canonical vertical space which is the kernel of the -derivative of the projection), and this canonical splitting is also the one that comes from the -coordinates on the tangent bundle in our definitions. So this statement is not as crazy as it -may seem. - -TODO define splittings of vector bundles; state this result invariantly. -/ -lemma tangent_map_tangent_bundle_pure (p : tangent_bundle I M) : - tangent_map I I.tangent (tangent_bundle.zero_section I M) p = ⟨⟨p.1, 0⟩, ⟨p.2, 0⟩⟩ := -begin - rcases p with ⟨x, v⟩, - have N : I.symm ⁻¹' (chart_at H x).target ∈ 𝓝 (I ((chart_at H x) x)), - { apply is_open.mem_nhds, - apply (local_homeomorph.open_target _).preimage I.continuous_inv_fun, - simp only with mfld_simps }, - have A : mdifferentiable_at I I.tangent (λ x, @total_space_mk M (tangent_space I) x 0) x := - tangent_bundle.smooth_zero_section.mdifferentiable_at, - have B : fderiv_within 𝕜 (λ (x_1 : E), (x_1, (0 : E))) (set.range ⇑I) (I ((chart_at H x) x)) v - = (v, 0), - { rw [fderiv_within_eq_fderiv, differentiable_at.fderiv_prod], - { simp }, - { exact differentiable_at_id' }, - { exact differentiable_at_const _ }, - { exact model_with_corners.unique_diff_at_image I }, - { exact differentiable_at_id'.prod (differentiable_at_const _) } }, - simp only [tangent_bundle.zero_section, tangent_map, mfderiv, - A, dif_pos, chart_at, basic_smooth_vector_bundle_core.chart, - basic_smooth_vector_bundle_core.to_topological_vector_bundle_core, tangent_bundle_core, - function.comp, continuous_linear_map.map_zero] with mfld_simps, - rw ← fderiv_within_inter N (I.unique_diff (I ((chart_at H x) x)) (set.mem_range_self _)) at B, - rw [← fderiv_within_inter N (I.unique_diff (I ((chart_at H x) x)) (set.mem_range_self _)), ← B], - congr' 2, - apply fderiv_within_congr _ (λ y hy, _), - { simp only [prod.mk.inj_iff] with mfld_simps, - exact ((tangent_bundle_core I M).to_topological_vector_bundle_core.coord_change - ((tangent_bundle_core I M).to_topological_vector_bundle_core.index_at (((chart_at H x).symm) - (I.symm (I ((chart_at H x) x))))) ⟨chart_at H x, _⟩ (((chart_at H x).symm) - (I.symm (I ((chart_at H x) x))))).map_zero, }, - { apply unique_diff_within_at.inter (I.unique_diff _ _) N, - simp only with mfld_simps }, - { simp only with mfld_simps at hy, - simp only [hy, prod.mk.inj_iff] with mfld_simps, - exact ((tangent_bundle_core I M).to_topological_vector_bundle_core.coord_change - ((tangent_bundle_core I M).to_topological_vector_bundle_core.index_at (((chart_at H x).symm) - (I.symm y))) ⟨chart_at H x, _⟩ (((chart_at H x).symm) (I.symm y))).map_zero, }, -end - -end tangent_bundle - /-! ### Smoothness of standard maps associated to the product of manifolds -/ section prod_mk diff --git a/src/geometry/manifold/cont_mdiff_map.lean b/src/geometry/manifold/cont_mdiff_map.lean index 3f168262d5c0c..73697d2b7a292 100644 --- a/src/geometry/manifold/cont_mdiff_map.lean +++ b/src/geometry/manifold/cont_mdiff_map.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Nicolò Cavalleri -/ -import geometry.manifold.cont_mdiff +import geometry.manifold.cont_mdiff_mfderiv import topology.continuous_function.basic /-! diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean new file mode 100644 index 0000000000000..c663fc78db4eb --- /dev/null +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -0,0 +1,643 @@ +/- +Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import geometry.manifold.cont_mdiff +import geometry.manifold.mfderiv + +/-! +### Interactions between differentiability, smoothness and manifold derivatives + +We give the relation between `mdifferentiable`, `cont_mdiff`, `mfderiv`, `tangent_map` +and related notions. + +## Main statements + +* `cont_mdiff_on.cont_mdiff_on_tangent_map_within` states that the bundled derivative + of a `Cⁿ` function in a domain is `Cᵐ` when `m + 1 ≤ n`. +* `cont_mdiff.cont_mdiff_tangent_map` states that the bundled derivative + of a `Cⁿ` function is `Cᵐ` when `m + 1 ≤ n`. +-/ + +open set function filter charted_space smooth_manifold_with_corners +open_locale topological_space manifold + +/-! ### Definition of smooth functions between manifolds -/ + +variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] +-- declare a smooth manifold `M` over the pair `(E, H)`. +{E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] +{H : Type*} [topological_space H] {I : model_with_corners 𝕜 E H} +{M : Type*} [topological_space M] [charted_space H M] [Is : smooth_manifold_with_corners I M] +-- declare a smooth manifold `M'` over the pair `(E', H')`. +{E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] +{H' : Type*} [topological_space H'] {I' : model_with_corners 𝕜 E' H'} +{M' : Type*} [topological_space M'] [charted_space H' M'] [I's : smooth_manifold_with_corners I' M'] +-- declare a smooth manifold `N` over the pair `(F, G)`. +{F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] +{G : Type*} [topological_space G] {J : model_with_corners 𝕜 F G} +{N : Type*} [topological_space N] [charted_space G N] [Js : smooth_manifold_with_corners J N] +-- declare a smooth manifold `N'` over the pair `(F', G')`. +{F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜 F'] +{G' : Type*} [topological_space G'] {J' : model_with_corners 𝕜 F' G'} +{N' : Type*} [topological_space N'] [charted_space G' N'] [J's : smooth_manifold_with_corners J' N'] +-- declare functions, sets, points and smoothness indices +{f f₁ : M → M'} {s s₁ t : set M} {x : M} {m n : ℕ∞} + +/-! ### Deducing differentiability from smoothness -/ + +lemma cont_mdiff_within_at.mdifferentiable_within_at + (hf : cont_mdiff_within_at I I' n f s x) (hn : 1 ≤ n) : + mdifferentiable_within_at I I' f s x := +begin + suffices h : mdifferentiable_within_at I I' f (s ∩ (f ⁻¹' (ext_chart_at I' (f x)).source)) x, + { rwa mdifferentiable_within_at_inter' at h, + apply (hf.1).preimage_mem_nhds_within, + exact is_open.mem_nhds (ext_chart_at_open_source I' (f x)) (mem_ext_chart_source I' (f x)) }, + rw mdifferentiable_within_at_iff, + exact ⟨hf.1.mono (inter_subset_left _ _), + (hf.2.differentiable_within_at hn).mono (by mfld_set_tac)⟩, +end + +lemma cont_mdiff_at.mdifferentiable_at (hf : cont_mdiff_at I I' n f x) (hn : 1 ≤ n) : + mdifferentiable_at I I' f x := +mdifferentiable_within_at_univ.1 $ cont_mdiff_within_at.mdifferentiable_within_at hf hn + +lemma cont_mdiff_on.mdifferentiable_on (hf : cont_mdiff_on I I' n f s) (hn : 1 ≤ n) : + mdifferentiable_on I I' f s := +λ x hx, (hf x hx).mdifferentiable_within_at hn + +lemma cont_mdiff.mdifferentiable (hf : cont_mdiff I I' n f) (hn : 1 ≤ n) : + mdifferentiable I I' f := +λ x, (hf x).mdifferentiable_at hn + +lemma smooth.mdifferentiable (hf : smooth I I' f) : mdifferentiable I I' f := +cont_mdiff.mdifferentiable hf le_top + +lemma smooth.mdifferentiable_at (hf : smooth I I' f) : mdifferentiable_at I I' f x := +hf.mdifferentiable x + +lemma smooth.mdifferentiable_within_at (hf : smooth I I' f) : + mdifferentiable_within_at I I' f s x := +hf.mdifferentiable_at.mdifferentiable_within_at + + +/-! ### The tangent map of a smooth function is smooth -/ + +section tangent_map + +/-- If a function is `C^n` with `1 ≤ n` on a domain with unique derivatives, then its bundled +derivative is continuous. In this auxiliary lemma, we prove this fact when the source and target +space are model spaces in models with corners. The general fact is proved in +`cont_mdiff_on.continuous_on_tangent_map_within`-/ +lemma cont_mdiff_on.continuous_on_tangent_map_within_aux + {f : H → H'} {s : set H} + (hf : cont_mdiff_on I I' n f s) (hn : 1 ≤ n) (hs : unique_mdiff_on I s) : + continuous_on (tangent_map_within I I' f s) ((tangent_bundle.proj I H) ⁻¹' s) := +begin + suffices h : continuous_on (λ (p : H × E), (f p.fst, + (fderiv_within 𝕜 (written_in_ext_chart_at I I' p.fst f) (I.symm ⁻¹' s ∩ range I) + ((ext_chart_at I p.fst) p.fst) : E →L[𝕜] E') p.snd)) (prod.fst ⁻¹' s), + { have A := (tangent_bundle_model_space_homeomorph H I).continuous, + rw continuous_iff_continuous_on_univ at A, + have B := ((tangent_bundle_model_space_homeomorph H' I').symm.continuous.comp_continuous_on h) + .comp' A, + have : (univ ∩ ⇑(tangent_bundle_model_space_homeomorph H I) ⁻¹' (prod.fst ⁻¹' s)) = + tangent_bundle.proj I H ⁻¹' s, + by { ext ⟨x, v⟩, simp only with mfld_simps }, + rw this at B, + apply B.congr, + rintros ⟨x, v⟩ hx, + dsimp [tangent_map_within], + ext, { refl }, + simp only with mfld_simps, + apply congr_fun, + apply congr_arg, + rw mdifferentiable_within_at.mfderiv_within (hf.mdifferentiable_on hn x hx), + refl }, + suffices h : continuous_on (λ (p : H × E), (fderiv_within 𝕜 (I' ∘ f ∘ I.symm) + (I.symm ⁻¹' s ∩ range I) (I p.fst) : E →L[𝕜] E') p.snd) (prod.fst ⁻¹' s), + { dsimp [written_in_ext_chart_at, ext_chart_at], + apply continuous_on.prod + (continuous_on.comp hf.continuous_on continuous_fst.continuous_on (subset.refl _)), + apply h.congr, + assume p hp, + refl }, + suffices h : continuous_on (fderiv_within 𝕜 (I' ∘ f ∘ I.symm) + (I.symm ⁻¹' s ∩ range I)) (I '' s), + { have C := continuous_on.comp h I.continuous_to_fun.continuous_on (subset.refl _), + have A : continuous (λq : (E →L[𝕜] E') × E, q.1 q.2) := + is_bounded_bilinear_map_apply.continuous, + have B : continuous_on (λp : H × E, + (fderiv_within 𝕜 (I' ∘ f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) + (I p.1), p.2)) (prod.fst ⁻¹' s), + { apply continuous_on.prod _ continuous_snd.continuous_on, + refine (continuous_on.comp C continuous_fst.continuous_on _ : _), + exact preimage_mono (subset_preimage_image _ _) }, + exact A.comp_continuous_on B }, + rw cont_mdiff_on_iff at hf, + let x : H := I.symm (0 : E), + let y : H' := I'.symm (0 : E'), + have A := hf.2 x y, + simp only [I.image_eq, inter_comm] with mfld_simps at A ⊢, + apply A.continuous_on_fderiv_within _ hn, + convert hs.unique_diff_on_target_inter x using 1, + simp only [inter_comm] with mfld_simps +end + +/-- If a function is `C^n` on a domain with unique derivatives, then its bundled derivative is +`C^m` when `m+1 ≤ n`. In this auxiliary lemma, we prove this fact when the source and target space +are model spaces in models with corners. The general fact is proved in +`cont_mdiff_on.cont_mdiff_on_tangent_map_within` -/ +lemma cont_mdiff_on.cont_mdiff_on_tangent_map_within_aux + {f : H → H'} {s : set H} + (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : + cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) + ((tangent_bundle.proj I H) ⁻¹' s) := +begin + have m_le_n : m ≤ n, + { apply le_trans _ hmn, + have : m + 0 ≤ m + 1 := add_le_add_left (zero_le _) _, + simpa only [add_zero] using this }, + have one_le_n : 1 ≤ n, + { apply le_trans _ hmn, + change 0 + 1 ≤ m + 1, + exact add_le_add_right (zero_le _) _ }, + have U': unique_diff_on 𝕜 (range I ∩ I.symm ⁻¹' s), + { assume y hy, + simpa only [unique_mdiff_on, unique_mdiff_within_at, hy.1, inter_comm] with mfld_simps + using hs (I.symm y) hy.2 }, + rw cont_mdiff_on_iff, + refine ⟨hf.continuous_on_tangent_map_within_aux one_le_n hs, λp q, _⟩, + have A : range I ×ˢ univ ∩ + ((equiv.sigma_equiv_prod H E).symm ∘ λ (p : E × E), ((I.symm) p.fst, p.snd)) ⁻¹' + (tangent_bundle.proj I H ⁻¹' s) + = (range I ∩ I.symm ⁻¹' s) ×ˢ univ, + by { ext ⟨x, v⟩, simp only with mfld_simps }, + suffices h : cont_diff_on 𝕜 m (((λ (p : H' × E'), (I' p.fst, p.snd)) ∘ + (equiv.sigma_equiv_prod H' E')) ∘ tangent_map_within I I' f s ∘ + ((equiv.sigma_equiv_prod H E).symm) ∘ λ (p : E × E), (I.symm p.fst, p.snd)) + ((range ⇑I ∩ ⇑(I.symm) ⁻¹' s) ×ˢ univ), + by simpa [A] using h, + change cont_diff_on 𝕜 m (λ (p : E × E), + ((I' (f (I.symm p.fst)), ((mfderiv_within I I' f s (I.symm p.fst)) : E → E') p.snd) : E' × E')) + ((range I ∩ I.symm ⁻¹' s) ×ˢ univ), + -- check that all bits in this formula are `C^n` + have hf' := cont_mdiff_on_iff.1 hf, + have A : cont_diff_on 𝕜 m (I' ∘ f ∘ I.symm) (range I ∩ I.symm ⁻¹' s) := + by simpa only with mfld_simps using (hf'.2 (I.symm 0) (I'.symm 0)).of_le m_le_n, + have B : cont_diff_on 𝕜 m ((I' ∘ f ∘ I.symm) ∘ prod.fst) + ((range I ∩ I.symm ⁻¹' s) ×ˢ univ) := + A.comp (cont_diff_fst.cont_diff_on) (prod_subset_preimage_fst _ _), + suffices C : cont_diff_on 𝕜 m (λ (p : E × E), + ((fderiv_within 𝕜 (I' ∘ f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) p.1 : _) p.2)) + ((range I ∩ I.symm ⁻¹' s) ×ˢ univ), + { apply cont_diff_on.prod B _, + apply C.congr (λp hp, _), + simp only with mfld_simps at hp, + simp only [mfderiv_within, hf.mdifferentiable_on one_le_n _ hp.2, hp.1, dif_pos] + with mfld_simps }, + have D : cont_diff_on 𝕜 m (λ x, + (fderiv_within 𝕜 (I' ∘ f ∘ I.symm) (I.symm ⁻¹' s ∩ range I) x)) + (range I ∩ I.symm ⁻¹' s), + { have : cont_diff_on 𝕜 n (I' ∘ f ∘ I.symm) (range I ∩ I.symm ⁻¹' s) := + by simpa only with mfld_simps using (hf'.2 (I.symm 0) (I'.symm 0)), + simpa only [inter_comm] using this.fderiv_within U' hmn }, + have := D.comp (cont_diff_fst.cont_diff_on) (prod_subset_preimage_fst _ _), + have := cont_diff_on.prod this (cont_diff_snd.cont_diff_on), + exact is_bounded_bilinear_map_apply.cont_diff.comp_cont_diff_on this, +end + +include Is I's + +/-- If a function is `C^n` on a domain with unique derivatives, then its bundled derivative +is `C^m` when `m+1 ≤ n`. -/ +theorem cont_mdiff_on.cont_mdiff_on_tangent_map_within + (hf : cont_mdiff_on I I' n f s) (hmn : m + 1 ≤ n) (hs : unique_mdiff_on I s) : + cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) + ((tangent_bundle.proj I M) ⁻¹' s) := +begin + /- The strategy of the proof is to avoid unfolding the definitions, and reduce by functoriality + to the case of functions on the model spaces, where we have already proved the result. + Let `l` and `r` be the charts to the left and to the right, so that we have + ``` + l^{-1} f r + H --------> M ---> M' ---> H' + ``` + Then the tangent map `T(r ∘ f ∘ l)` is smooth by a previous result. Consider the composition + ``` + Tl T(r ∘ f ∘ l^{-1}) Tr^{-1} + TM -----> TH -------------------> TH' ---------> TM' + ``` + where `Tr^{-1}` and `Tl` are the tangent maps of `r^{-1}` and `l`. Writing `Tl` and `Tr^{-1}` as + composition of charts (called `Dl` and `il` for `l` and `Dr` and `ir` in the proof below), it + follows that they are smooth. The composition of all these maps is `Tf`, and is therefore smooth + as a composition of smooth maps. + -/ + have m_le_n : m ≤ n, + { apply le_trans _ hmn, + have : m + 0 ≤ m + 1 := add_le_add_left (zero_le _) _, + simpa only [add_zero] }, + have one_le_n : 1 ≤ n, + { apply le_trans _ hmn, + change 0 + 1 ≤ m + 1, + exact add_le_add_right (zero_le _) _ }, + /- First step: local reduction on the space, to a set `s'` which is contained in chart domains. -/ + refine cont_mdiff_on_of_locally_cont_mdiff_on (λp hp, _), + have hf' := cont_mdiff_on_iff.1 hf, + simp [tangent_bundle.proj] at hp, + let l := chart_at H p.1, + set Dl := chart_at (model_prod H E) p with hDl, + let r := chart_at H' (f p.1), + let Dr := chart_at (model_prod H' E') (tangent_map_within I I' f s p), + let il := chart_at (model_prod H E) (tangent_map I I l p), + let ir := chart_at (model_prod H' E') (tangent_map I I' (r ∘ f) p), + let s' := f ⁻¹' r.source ∩ s ∩ l.source, + let s'_lift := (tangent_bundle.proj I M)⁻¹' s', + let s'l := l.target ∩ l.symm ⁻¹' s', + let s'l_lift := (tangent_bundle.proj I H) ⁻¹' s'l, + rcases continuous_on_iff'.1 hf'.1 r.source r.open_source with ⟨o, o_open, ho⟩, + suffices h : cont_mdiff_on I.tangent I'.tangent m (tangent_map_within I I' f s) s'_lift, + { refine ⟨(tangent_bundle.proj I M)⁻¹' (o ∩ l.source), _, _, _⟩, + show is_open ((tangent_bundle.proj I M)⁻¹' (o ∩ l.source)), from + (is_open.inter o_open l.open_source).preimage (tangent_bundle_proj_continuous _ _) , + show p ∈ tangent_bundle.proj I M ⁻¹' (o ∩ l.source), + { simp [tangent_bundle.proj] at ⊢, + have : p.1 ∈ f ⁻¹' r.source ∩ s, by simp [hp], + rw ho at this, + exact this.1 }, + { have : tangent_bundle.proj I M ⁻¹' s ∩ tangent_bundle.proj I M ⁻¹' (o ∩ l.source) = s'_lift, + { dsimp only [s'_lift, s'], rw [ho], mfld_set_tac }, + rw this, + exact h } }, + /- Second step: check that all functions are smooth, and use the chain rule to write the bundled + derivative as a composition of a function between model spaces and of charts. + Convention: statements about the differentiability of `a ∘ b ∘ c` are named `diff_abc`. Statements + about differentiability in the bundle have a `_lift` suffix. -/ + have U' : unique_mdiff_on I s', + { apply unique_mdiff_on.inter _ l.open_source, + rw [ho, inter_comm], + exact hs.inter o_open }, + have U'l : unique_mdiff_on I s'l := + U'.unique_mdiff_on_preimage (mdifferentiable_chart _ _), + have diff_f : cont_mdiff_on I I' n f s' := + hf.mono (by mfld_set_tac), + have diff_r : cont_mdiff_on I' I' n r r.source := + cont_mdiff_on_chart, + have diff_rf : cont_mdiff_on I I' n (r ∘ f) s', + { apply cont_mdiff_on.comp diff_r diff_f (λx hx, _), + simp only [s'] with mfld_simps at hx, simp only [hx] with mfld_simps }, + have diff_l : cont_mdiff_on I I n l.symm s'l, + { have A : cont_mdiff_on I I n l.symm l.target := + cont_mdiff_on_chart_symm, + exact A.mono (by mfld_set_tac) }, + have diff_rfl : cont_mdiff_on I I' n (r ∘ f ∘ l.symm) s'l, + { apply cont_mdiff_on.comp diff_rf diff_l, + mfld_set_tac }, + have diff_rfl_lift : cont_mdiff_on I.tangent I'.tangent m + (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l) s'l_lift := + diff_rfl.cont_mdiff_on_tangent_map_within_aux hmn U'l, + have diff_irrfl_lift : cont_mdiff_on I.tangent I'.tangent m + (ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l)) s'l_lift, + { have A : cont_mdiff_on I'.tangent I'.tangent m ir ir.source := cont_mdiff_on_chart, + exact cont_mdiff_on.comp A diff_rfl_lift (λp hp, by simp only [ir] with mfld_simps) }, + have diff_Drirrfl_lift : cont_mdiff_on I.tangent I'.tangent m + (Dr.symm ∘ (ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l))) s'l_lift, + { have A : cont_mdiff_on I'.tangent I'.tangent m Dr.symm Dr.target := + cont_mdiff_on_chart_symm, + apply cont_mdiff_on.comp A diff_irrfl_lift (λp hp, _), + simp only [s'l_lift, tangent_bundle.proj] with mfld_simps at hp, + simp only [ir, @local_equiv.refl_coe (model_prod H' E'), hp] with mfld_simps }, + -- conclusion of this step: the composition of all the maps above is smooth + have diff_DrirrflilDl : cont_mdiff_on I.tangent I'.tangent m + (Dr.symm ∘ (ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l)) ∘ + (il.symm ∘ Dl)) s'_lift, + { have A : cont_mdiff_on I.tangent I.tangent m Dl Dl.source := cont_mdiff_on_chart, + have A' : cont_mdiff_on I.tangent I.tangent m Dl s'_lift, + { apply A.mono (λp hp, _), + simp only [s'_lift, tangent_bundle.proj] with mfld_simps at hp, + simp only [Dl, hp] with mfld_simps }, + have B : cont_mdiff_on I.tangent I.tangent m il.symm il.target := + cont_mdiff_on_chart_symm, + have C : cont_mdiff_on I.tangent I.tangent m (il.symm ∘ Dl) s'_lift := + cont_mdiff_on.comp B A' (λp hp, by simp only [il] with mfld_simps), + apply cont_mdiff_on.comp diff_Drirrfl_lift C (λp hp, _), + simp only [s'_lift, tangent_bundle.proj] with mfld_simps at hp, + simp only [il, s'l_lift, hp, tangent_bundle.proj] with mfld_simps }, + /- Third step: check that the composition of all the maps indeed coincides with the derivative we + are looking for -/ + have eq_comp : ∀q ∈ s'_lift, tangent_map_within I I' f s q = + (Dr.symm ∘ ir ∘ (tangent_map_within I I' (r ∘ f ∘ l.symm) s'l) ∘ + (il.symm ∘ Dl)) q, + { assume q hq, + simp only [s'_lift, tangent_bundle.proj] with mfld_simps at hq, + have U'q : unique_mdiff_within_at I s' q.1, + by { apply U', simp only [hq, s'] with mfld_simps }, + have U'lq : unique_mdiff_within_at I s'l (Dl q).1, + by { apply U'l, simp only [hq, s'l] with mfld_simps }, + have A : tangent_map_within I I' ((r ∘ f) ∘ l.symm) s'l (il.symm (Dl q)) = + tangent_map_within I I' (r ∘ f) s' (tangent_map_within I I l.symm s'l (il.symm (Dl q))), + { refine tangent_map_within_comp_at (il.symm (Dl q)) _ _ (λp hp, _) U'lq, + { apply diff_rf.mdifferentiable_on one_le_n, + simp only [hq] with mfld_simps }, + { apply diff_l.mdifferentiable_on one_le_n, + simp only [s'l, hq] with mfld_simps }, + { simp only with mfld_simps at hp, simp only [hp] with mfld_simps } }, + have B : tangent_map_within I I l.symm s'l (il.symm (Dl q)) = q, + { have : tangent_map_within I I l.symm s'l (il.symm (Dl q)) + = tangent_map I I l.symm (il.symm (Dl q)), + { refine tangent_map_within_eq_tangent_map U'lq _, + refine mdifferentiable_at_atlas_symm _ (chart_mem_atlas _ _) _, + simp only [hq] with mfld_simps }, + rw [this, tangent_map_chart_symm, hDl], + { simp only [hq] with mfld_simps, + have : q ∈ (chart_at (model_prod H E) p).source, by simp only [hq] with mfld_simps, + exact (chart_at (model_prod H E) p).left_inv this }, + { simp only [hq] with mfld_simps } }, + have C : tangent_map_within I I' (r ∘ f) s' q + = tangent_map_within I' I' r r.source (tangent_map_within I I' f s' q), + { refine tangent_map_within_comp_at q _ _ (λr hr, _) U'q, + { apply diff_r.mdifferentiable_on one_le_n, + simp only [hq] with mfld_simps }, + { apply diff_f.mdifferentiable_on one_le_n, + simp only [hq] with mfld_simps }, + { simp only [s'] with mfld_simps at hr, + simp only [hr] with mfld_simps } }, + have D : Dr.symm (ir (tangent_map_within I' I' r r.source (tangent_map_within I I' f s' q))) + = tangent_map_within I I' f s' q, + { have A : tangent_map_within I' I' r r.source (tangent_map_within I I' f s' q) = + tangent_map I' I' r (tangent_map_within I I' f s' q), + { apply tangent_map_within_eq_tangent_map, + { apply is_open.unique_mdiff_within_at _ r.open_source, simp [hq] }, + { refine mdifferentiable_at_atlas _ (chart_mem_atlas _ _) _, + simp only [hq] with mfld_simps } }, + have : f p.1 = (tangent_map_within I I' f s p).1 := rfl, + rw [A], + dsimp [r, Dr], + rw [this, tangent_map_chart], + { simp only [hq] with mfld_simps, + have : tangent_map_within I I' f s' q ∈ + (chart_at (model_prod H' E') (tangent_map_within I I' f s p)).source, + by simp only [hq] with mfld_simps, + exact (chart_at (model_prod H' E') (tangent_map_within I I' f s p)).left_inv this }, + { simp only [hq] with mfld_simps } }, + have E : tangent_map_within I I' f s' q = tangent_map_within I I' f s q, + { refine tangent_map_within_subset (by mfld_set_tac) U'q _, + apply hf.mdifferentiable_on one_le_n, + simp only [hq] with mfld_simps }, + simp only [(∘), A, B, C, D, E.symm] }, + exact diff_DrirrflilDl.congr eq_comp, +end + +/-- If a function is `C^n` on a domain with unique derivatives, with `1 ≤ n`, then its bundled +derivative is continuous there. -/ +theorem cont_mdiff_on.continuous_on_tangent_map_within + (hf : cont_mdiff_on I I' n f s) (hmn : 1 ≤ n) (hs : unique_mdiff_on I s) : + continuous_on (tangent_map_within I I' f s) ((tangent_bundle.proj I M) ⁻¹' s) := +begin + have : cont_mdiff_on I.tangent I'.tangent 0 (tangent_map_within I I' f s) + ((tangent_bundle.proj I M) ⁻¹' s) := + hf.cont_mdiff_on_tangent_map_within hmn hs, + exact this.continuous_on +end + +/-- If a function is `C^n`, then its bundled derivative is `C^m` when `m+1 ≤ n`. -/ +theorem cont_mdiff.cont_mdiff_tangent_map + (hf : cont_mdiff I I' n f) (hmn : m + 1 ≤ n) : + cont_mdiff I.tangent I'.tangent m (tangent_map I I' f) := +begin + rw ← cont_mdiff_on_univ at hf ⊢, + convert hf.cont_mdiff_on_tangent_map_within hmn unique_mdiff_on_univ, + rw tangent_map_within_univ +end + +/-- If a function is `C^n`, with `1 ≤ n`, then its bundled derivative is continuous. -/ +theorem cont_mdiff.continuous_tangent_map + (hf : cont_mdiff I I' n f) (hmn : 1 ≤ n) : + continuous (tangent_map I I' f) := +begin + rw ← cont_mdiff_on_univ at hf, + rw continuous_iff_continuous_on_univ, + convert hf.continuous_on_tangent_map_within hmn unique_mdiff_on_univ, + rw tangent_map_within_univ +end + +end tangent_map + +/-! ### Smoothness of the projection in a basic smooth bundle -/ + +namespace basic_smooth_vector_bundle_core + +variables (Z : basic_smooth_vector_bundle_core I M E') + +/-- A version of `cont_mdiff_at_iff_target` when the codomain is the total space of + a `basic_smooth_vector_bundle_core`. The continuity condition in the RHS is weaker. -/ +lemma cont_mdiff_at_iff_target {f : N → Z.to_topological_vector_bundle_core.total_space} + {x : N} {n : ℕ∞} : + cont_mdiff_at J (I.prod 𝓘(𝕜, E')) n f x ↔ continuous_at (bundle.total_space.proj ∘ f) x ∧ + cont_mdiff_at J 𝓘(𝕜, E × E') n (ext_chart_at (I.prod 𝓘(𝕜, E')) (f x) ∘ f) x := +begin + let Z' := Z.to_topological_vector_bundle_core, + rw [cont_mdiff_at_iff_target, and.congr_left_iff], + refine λ hf, ⟨λ h, Z'.continuous_proj.continuous_at.comp h, λ h, _⟩, + exact (Z'.local_triv ⟨chart_at _ (f x).1, chart_mem_atlas _ _⟩).to_fiber_bundle_trivialization + .continuous_at_of_comp_left h (mem_chart_source _ _) (h.prod hf.continuous_at.snd) +end + +lemma cont_mdiff_proj : + cont_mdiff (I.prod 𝓘(𝕜, E')) I n Z.to_topological_vector_bundle_core.proj := +begin + assume x, + rw [cont_mdiff_at, cont_mdiff_within_at_iff'], + refine ⟨Z.to_topological_vector_bundle_core.continuous_proj.continuous_within_at, _⟩, + simp only [(∘), chart_at, chart] with mfld_simps, + apply cont_diff_within_at_fst.congr, + { rintros ⟨a, b⟩ hab, + simp only with mfld_simps at hab, + simp only [hab] with mfld_simps }, + { simp only with mfld_simps } +end + +lemma smooth_proj : + smooth (I.prod 𝓘(𝕜, E')) I Z.to_topological_vector_bundle_core.proj := +cont_mdiff_proj Z + +lemma cont_mdiff_on_proj {s : set (Z.to_topological_vector_bundle_core.total_space)} : + cont_mdiff_on (I.prod 𝓘(𝕜, E')) I n + Z.to_topological_vector_bundle_core.proj s := +Z.cont_mdiff_proj.cont_mdiff_on + +lemma smooth_on_proj {s : set (Z.to_topological_vector_bundle_core.total_space)} : + smooth_on (I.prod 𝓘(𝕜, E')) I Z.to_topological_vector_bundle_core.proj s := +cont_mdiff_on_proj Z + +lemma cont_mdiff_at_proj {p : Z.to_topological_vector_bundle_core.total_space} : + cont_mdiff_at (I.prod 𝓘(𝕜, E')) I n + Z.to_topological_vector_bundle_core.proj p := +Z.cont_mdiff_proj.cont_mdiff_at + +lemma smooth_at_proj {p : Z.to_topological_vector_bundle_core.total_space} : + smooth_at (I.prod 𝓘(𝕜, E')) I Z.to_topological_vector_bundle_core.proj p := +Z.cont_mdiff_at_proj + +lemma cont_mdiff_within_at_proj + {s : set (Z.to_topological_vector_bundle_core.total_space)} + {p : Z.to_topological_vector_bundle_core.total_space} : + cont_mdiff_within_at (I.prod 𝓘(𝕜, E')) I n + Z.to_topological_vector_bundle_core.proj s p := +Z.cont_mdiff_at_proj.cont_mdiff_within_at + +lemma smooth_within_at_proj + {s : set (Z.to_topological_vector_bundle_core.total_space)} + {p : Z.to_topological_vector_bundle_core.total_space} : + smooth_within_at (I.prod 𝓘(𝕜, E')) I + Z.to_topological_vector_bundle_core.proj s p := +Z.cont_mdiff_within_at_proj + +/-- If an element of `E'` is invariant under all coordinate changes, then one can define a +corresponding section of the fiber bundle, which is smooth. This applies in particular to the +zero section of a vector bundle. Another example (not yet defined) would be the identity +section of the endomorphism bundle of a vector bundle. -/ +lemma smooth_const_section (v : E') + (h : ∀ (i j : atlas H M), ∀ x ∈ i.1.source ∩ j.1.source, Z.coord_change i j (i.1 x) v = v) : + smooth I (I.prod 𝓘(𝕜, E')) + (show M → Z.to_topological_vector_bundle_core.total_space, from λ x, ⟨x, v⟩) := +begin + assume x, + rw [cont_mdiff_at, cont_mdiff_within_at_iff'], + split, + { apply continuous.continuous_within_at, + apply topological_fiber_bundle_core.continuous_const_section, + assume i j y hy, + exact h _ _ _ hy }, + { have : cont_diff 𝕜 ⊤ (λ (y : E), (y, v)) := cont_diff_id.prod cont_diff_const, + apply this.cont_diff_within_at.congr, + { assume y hy, + simp only with mfld_simps at hy, + simp only [chart, hy, chart_at, prod.mk.inj_iff, to_topological_vector_bundle_core] + with mfld_simps, + apply h, + simp only [hy, subtype.val_eq_coe] with mfld_simps, + exact mem_chart_source H (((chart_at H x).symm) ((model_with_corners.symm I) y)) }, + { simp only [chart, chart_at, prod.mk.inj_iff, to_topological_vector_bundle_core] + with mfld_simps, + apply h, + simp only [subtype.val_eq_coe] with mfld_simps, + exact mem_chart_source H x, } } +end + +end basic_smooth_vector_bundle_core + +/-! ### Smoothness of the tangent bundle projection -/ + +namespace tangent_bundle + +include Is + +lemma cont_mdiff_proj : + cont_mdiff I.tangent I n (proj I M) := +basic_smooth_vector_bundle_core.cont_mdiff_proj _ + +lemma smooth_proj : smooth I.tangent I (proj I M) := +basic_smooth_vector_bundle_core.smooth_proj _ + +lemma cont_mdiff_on_proj {s : set (tangent_bundle I M)} : + cont_mdiff_on I.tangent I n (proj I M) s := +basic_smooth_vector_bundle_core.cont_mdiff_on_proj _ + +lemma smooth_on_proj {s : set (tangent_bundle I M)} : + smooth_on I.tangent I (proj I M) s := +basic_smooth_vector_bundle_core.smooth_on_proj _ + +lemma cont_mdiff_at_proj {p : tangent_bundle I M} : + cont_mdiff_at I.tangent I n + (proj I M) p := +basic_smooth_vector_bundle_core.cont_mdiff_at_proj _ + +lemma smooth_at_proj {p : tangent_bundle I M} : + smooth_at I.tangent I (proj I M) p := +basic_smooth_vector_bundle_core.smooth_at_proj _ + +lemma cont_mdiff_within_at_proj + {s : set (tangent_bundle I M)} {p : tangent_bundle I M} : + cont_mdiff_within_at I.tangent I n + (proj I M) s p := +basic_smooth_vector_bundle_core.cont_mdiff_within_at_proj _ + +lemma smooth_within_at_proj + {s : set (tangent_bundle I M)} {p : tangent_bundle I M} : + smooth_within_at I.tangent I + (proj I M) s p := +basic_smooth_vector_bundle_core.smooth_within_at_proj _ + +variables (I M) +/-- The zero section of the tangent bundle -/ +def zero_section : M → tangent_bundle I M := λ x, ⟨x, 0⟩ +variables {I M} + +lemma smooth_zero_section : smooth I I.tangent (zero_section I M) := +begin + apply basic_smooth_vector_bundle_core.smooth_const_section (tangent_bundle_core I M) 0, + assume i j x hx, + simp only [tangent_bundle_core, continuous_linear_map.map_zero, continuous_linear_map.coe_coe] + with mfld_simps, +end + +open bundle + +/-- The derivative of the zero section of the tangent bundle maps `⟨x, v⟩` to `⟨⟨x, 0⟩, ⟨v, 0⟩⟩`. + +Note that, as currently framed, this is a statement in coordinates, thus reliant on the choice +of the coordinate system we use on the tangent bundle. + +However, the result itself is coordinate-dependent only to the extent that the coordinates +determine a splitting of the tangent bundle. Moreover, there is a canonical splitting at each +point of the zero section (since there is a canonical horizontal space there, the tangent space +to the zero section, in addition to the canonical vertical space which is the kernel of the +derivative of the projection), and this canonical splitting is also the one that comes from the +coordinates on the tangent bundle in our definitions. So this statement is not as crazy as it +may seem. + +TODO define splittings of vector bundles; state this result invariantly. -/ +lemma tangent_map_tangent_bundle_pure (p : tangent_bundle I M) : + tangent_map I I.tangent (tangent_bundle.zero_section I M) p = ⟨⟨p.1, 0⟩, ⟨p.2, 0⟩⟩ := +begin + rcases p with ⟨x, v⟩, + have N : I.symm ⁻¹' (chart_at H x).target ∈ 𝓝 (I ((chart_at H x) x)), + { apply is_open.mem_nhds, + apply (local_homeomorph.open_target _).preimage I.continuous_inv_fun, + simp only with mfld_simps }, + have A : mdifferentiable_at I I.tangent (λ x, @total_space_mk M (tangent_space I) x 0) x := + tangent_bundle.smooth_zero_section.mdifferentiable_at, + have B : fderiv_within 𝕜 (λ (x_1 : E), (x_1, (0 : E))) (set.range ⇑I) (I ((chart_at H x) x)) v + = (v, 0), + { rw [fderiv_within_eq_fderiv, differentiable_at.fderiv_prod], + { simp }, + { exact differentiable_at_id' }, + { exact differentiable_at_const _ }, + { exact model_with_corners.unique_diff_at_image I }, + { exact differentiable_at_id'.prod (differentiable_at_const _) } }, + simp only [tangent_bundle.zero_section, tangent_map, mfderiv, + A, dif_pos, chart_at, basic_smooth_vector_bundle_core.chart, + basic_smooth_vector_bundle_core.to_topological_vector_bundle_core, tangent_bundle_core, + function.comp, continuous_linear_map.map_zero] with mfld_simps, + rw ← fderiv_within_inter N (I.unique_diff (I ((chart_at H x) x)) (set.mem_range_self _)) at B, + rw [← fderiv_within_inter N (I.unique_diff (I ((chart_at H x) x)) (set.mem_range_self _)), ← B], + congr' 2, + apply fderiv_within_congr _ (λ y hy, _), + { simp only [prod.mk.inj_iff] with mfld_simps, + exact ((tangent_bundle_core I M).to_topological_vector_bundle_core.coord_change + ((tangent_bundle_core I M).to_topological_vector_bundle_core.index_at (((chart_at H x).symm) + (I.symm (I ((chart_at H x) x))))) ⟨chart_at H x, _⟩ (((chart_at H x).symm) + (I.symm (I ((chart_at H x) x))))).map_zero, }, + { apply unique_diff_within_at.inter (I.unique_diff _ _) N, + simp only with mfld_simps }, + { simp only with mfld_simps at hy, + simp only [hy, prod.mk.inj_iff] with mfld_simps, + exact ((tangent_bundle_core I M).to_topological_vector_bundle_core.coord_change + ((tangent_bundle_core I M).to_topological_vector_bundle_core.index_at (((chart_at H x).symm) + (I.symm y))) ⟨chart_at H x, _⟩ (((chart_at H x).symm) (I.symm y))).map_zero, }, +end + +end tangent_bundle diff --git a/src/geometry/manifold/mfderiv.lean b/src/geometry/manifold/mfderiv.lean index 1616efda1224a..b812169762b9b 100644 --- a/src/geometry/manifold/mfderiv.lean +++ b/src/geometry/manifold/mfderiv.lean @@ -186,11 +186,6 @@ unique_diff_within_at 𝕜 ((ext_chart_at I x).symm ⁻¹' s ∩ range I) ((ext_ def unique_mdiff_on (s : set M) := ∀x∈s, unique_mdiff_within_at I s x -/-- Conjugating a function to write it in the preferred charts around `x`. The manifold derivative -of `f` will just be the derivative of this conjugated function. -/ -@[simp, mfld_simps] def written_in_ext_chart_at (x : M) (f : M → M') : E → E' := -(ext_chart_at I' (f x)) ∘ f ∘ (ext_chart_at I x).symm - /-- `mdifferentiable_within_at I I' f s x` indicates that the function `f` between manifolds has a derivative at the point `x` within the set `s`. This is a generalization of `differentiable_within_at` to manifolds. diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index a6c6ba9631dab..7300bbe09e949 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -680,6 +680,9 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E : Type*} [normed_add_comm_group E] [normed_space 𝕜 E] {H : Type*} [topological_space H] (I : model_with_corners 𝕜 E H) {M : Type*} [topological_space M] [charted_space H M] + {E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] + {H' : Type*} [topological_space H'] (I' : model_with_corners 𝕜 E' H') + {M' : Type*} [topological_space M'] [charted_space H' M'] (x : M) {s t : set M} /-! @@ -935,6 +938,11 @@ begin exact I.image_mem_nhds_within ((local_homeomorph.open_source _).mem_nhds hz) end +/-- Conjugating a function to write it in the preferred charts around `x`. +The manifold derivative of `f` will just be the derivative of this conjugated function. -/ +@[simp, mfld_simps] def written_in_ext_chart_at (x : M) (f : M → M') : E → E' := +ext_chart_at I' (f x) ∘ f ∘ (ext_chart_at I x).symm + variable (𝕜) lemma ext_chart_self_eq {x : H} : ⇑(ext_chart_at I x) = I := rfl @@ -947,5 +955,4 @@ by simp only with mfld_simps lemma ext_chart_model_space_apply {x y : E} : ext_chart_at 𝓘(𝕜, E) x y = y := rfl - end extended_charts From 6d94ff7bc65cc7895934e102664dd889d1f07522 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 17 Oct 2022 11:19:59 +0000 Subject: [PATCH 796/828] feat(topology/subset_properties, noetherian_space): Noetherian spaces have finite irreducible components. (#17015) The previous statement was too weak. Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/topology/noetherian_space.lean | 19 +++++++-------- src/topology/subset_properties.lean | 32 ++++++++++++++++++++++--- src/topology/uniform_space/compact.lean | 4 ++-- 3 files changed, 39 insertions(+), 16 deletions(-) diff --git a/src/topology/noetherian_space.lean b/src/topology/noetherian_space.lean index 1cf5f3f65e33c..79625b2a697d5 100644 --- a/src/topology/noetherian_space.lean +++ b/src/topology/noetherian_space.lean @@ -206,28 +206,25 @@ begin end lemma noetherian_space.finite_irreducible_components [noetherian_space α] : - (set.range irreducible_component : set (set α)).finite := + (irreducible_components α).finite := begin classical, obtain ⟨S, hS₁, hS₂⟩ := noetherian_space.exists_finset_irreducible (⊤ : closeds α), - suffices : ∀ x : α, ∃ s : S, irreducible_component x = s, - { choose f hf, - rw [show irreducible_component = coe ∘ f, from funext hf, set.range_comp], - exact (set.finite.intro infer_instance).image _ }, - intro x, - obtain ⟨z, hz, hz'⟩ : ∃ (z : set α) (H : z ∈ finset.image coe S), irreducible_component x ⊆ z, + suffices : irreducible_components α ⊆ coe '' (S : set $ closeds α), + { exact set.finite.subset ((set.finite.intro infer_instance).image _) this }, + intros K hK, + obtain ⟨z, hz, hz'⟩ : ∃ (z : set α) (H : z ∈ finset.image coe S), K ⊆ z, { convert is_irreducible_iff_sUnion_closed.mp - is_irreducible_irreducible_component (S.image coe) _ _, - { apply_instance }, + hK.1 (S.image coe) _ _, { simp only [finset.mem_image, exists_prop, forall_exists_index, and_imp], rintro _ z hz rfl, exact z.2 }, { exact (set.subset_univ _).trans ((congr_arg coe hS₂).trans $ by simp).subset } }, obtain ⟨s, hs, e⟩ := finset.mem_image.mp hz, rw ← e at hz', - use ⟨s, hs⟩, + refine ⟨s, hs, _⟩, symmetry, - apply eq_irreducible_component (hS₁ _).2, + suffices : K ≤ s, { exact this.antisymm (hK.2 (hS₁ ⟨s, hs⟩) this) }, simpa, end diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index b8706ea8dc33c..70aace824f04a 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -9,6 +9,7 @@ import data.finset.order import data.set.accumulate import tactic.tfae import topology.bornology.basic +import order.minimal /-! # Properties of subsets of topological spaces @@ -1580,6 +1581,29 @@ let ⟨m, hm, hsm, hmm⟩ := zorn_subset_nonempty {t : set α | is_preirreducibl λ x hxc, subset_sUnion_of_mem hxc⟩) s H in ⟨m, hm, hsm, λ u hu hmu, hmm _ hu hmu⟩ +/-- The set of irreducible components of a topological space. -/ +def irreducible_components (α : Type*) [topological_space α] : set (set α) := +maximals (≤) { s : set α | is_irreducible s } + +lemma is_closed_of_mem_irreducible_components (s ∈ irreducible_components α) : + is_closed s := +begin + rw [← closure_eq_iff_is_closed, eq_comm], + exact subset_closure.antisymm (H.2 H.1.closure subset_closure), +end + +lemma irreducible_components_eq_maximals_closed (α : Type*) [topological_space α] : + irreducible_components α = maximals (≤) { s : set α | is_closed s ∧ is_irreducible s } := +begin + ext s, + split, + { intro H, exact ⟨⟨is_closed_of_mem_irreducible_components _ H, H.1⟩, λ x h e, H.2 h.2 e⟩ }, + { intro H, refine ⟨H.1.2, λ x h e, _⟩, + have : closure x ≤ s, + { exact H.2 ⟨is_closed_closure, h.closure⟩ (e.trans subset_closure) }, + exact le_trans subset_closure this } +end + /-- A maximal irreducible set that contains a given point. -/ def irreducible_component (x : α) : set α := classical.some (exists_preirreducible {x} is_irreducible_singleton.is_preirreducible) @@ -1599,11 +1623,13 @@ theorem eq_irreducible_component {x : α} : ∀ {s : set α}, is_preirreducible s → irreducible_component x ⊆ s → s = irreducible_component x := (irreducible_component_property x).2.2 +lemma irreducible_component_mem_irreducible_components (x : α) : + irreducible_component x ∈ irreducible_components α := +⟨is_irreducible_irreducible_component, λ s h₁ h₂,(eq_irreducible_component h₁.2 h₂).le⟩ + theorem is_closed_irreducible_component {x : α} : is_closed (irreducible_component x) := -closure_eq_iff_is_closed.1 $ eq_irreducible_component - is_irreducible_irreducible_component.is_preirreducible.closure - subset_closure +is_closed_of_mem_irreducible_components _ (irreducible_component_mem_irreducible_components x) /-- A preirreducible space is one where there is no non-trivial pair of disjoint opens. -/ class preirreducible_space (α : Type u) [topological_space α] : Prop := diff --git a/src/topology/uniform_space/compact.lean b/src/topology/uniform_space/compact.lean index 9fddded951809..bd87660593be7 100644 --- a/src/topology/uniform_space/compact.lean +++ b/src/topology/uniform_space/compact.lean @@ -80,7 +80,7 @@ def uniform_space_of_compact_t2 [topological_space γ] [compact_space γ] [t2_sp end, symm := begin refine le_of_eq _, - rw map_supr, + rw filter.map_supr, congr' with x : 1, erw [nhds_prod_eq, ← prod_comm], end, @@ -184,7 +184,7 @@ lemma compact_space.uniform_continuous_of_continuous [compact_space α] {f : α → β} (h : continuous f) : uniform_continuous f := calc map (prod.map f f) (𝓤 α) = map (prod.map f f) (⨆ x, 𝓝 (x, x)) : by rw compact_space_uniformity - ... = ⨆ x, map (prod.map f f) (𝓝 (x, x)) : by rw map_supr + ... = ⨆ x, map (prod.map f f) (𝓝 (x, x)) : by rw filter.map_supr ... ≤ ⨆ x, 𝓝 (f x, f x) : supr_mono (λ x, (h.prod_map h).continuous_at) ... ≤ ⨆ y, 𝓝 (y, y) : supr_comp_le (λ y, 𝓝 (y, y)) f ... ≤ 𝓤 β : supr_nhds_le_uniformity From cf94743b4edb5834e28b962ee5a5e2dd4dd5faa1 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 17 Oct 2022 13:11:14 +0000 Subject: [PATCH 797/828] feat(ring_theory/finiteness): A morphism is finitely presented if the composition with a finite morphism is. (#15949) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- src/ring_theory/adjoin/basic.lean | 10 ++++ src/ring_theory/finiteness.lean | 91 +++++++++++++++++++++++++++++++ 2 files changed, 101 insertions(+) diff --git a/src/ring_theory/adjoin/basic.lean b/src/ring_theory/adjoin/basic.lean index 63a64bcbc86e7..fe6d707d084df 100644 --- a/src/ring_theory/adjoin/basic.lean +++ b/src/ring_theory/adjoin/basic.lean @@ -270,6 +270,16 @@ begin congr' 1 with z, simp [submonoid.closure_union, submonoid.mem_sup, set.mem_mul] end +lemma adjoin_adjoin_of_tower [semiring B] [algebra R B] [algebra A B] [is_scalar_tower R A B] + (s : set B) : adjoin A (adjoin R s : set B) = adjoin A s := +begin + apply le_antisymm (adjoin_le _), + { exact adjoin_mono subset_adjoin }, + { change adjoin R s ≤ (adjoin A s).restrict_scalars R, + refine adjoin_le _, + exact subset_adjoin } +end + variable {R} lemma pow_smul_mem_of_smul_subset_of_mem_adjoin diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index 3ad61a258c627..6bce12e8193b4 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -6,6 +6,7 @@ Authors: Johan Commelin import group_theory.finiteness import ring_theory.algebra_tower +import ring_theory.mv_polynomial.tower import ring_theory.ideal.quotient import ring_theory.noetherian @@ -424,7 +425,86 @@ end open mv_polynomial +-- We follow the proof of https://stacks.math.columbia.edu/tag/0561 +-- TODO: extract out helper lemmas and tidy proof. +lemma of_restrict_scalars_finite_presentation [algebra A B] [is_scalar_tower R A B] + (hRB : finite_presentation R B) [hRA : finite_type R A] : finite_presentation A B := +begin + obtain ⟨n, f, hf, s, hs⟩ := hRB, + let RX := mv_polynomial (fin n) R, let AX := mv_polynomial (fin n) A, + refine ⟨n, mv_polynomial.aeval (f ∘ X), _, _⟩, + { rw [← algebra.range_top_iff_surjective, ← algebra.adjoin_range_eq_range_aeval, set.range_comp, + _root_.eq_top_iff, ← @adjoin_adjoin_of_tower R A B, adjoin_image, + adjoin_range_X, algebra.map_top, (algebra.range_top_iff_surjective _).mpr hf], + exact subset_adjoin }, + { obtain ⟨t, ht⟩ := hRA.out, + have := λ i : t, hf (algebra_map A B i), + choose t' ht', + have ht'' : algebra.adjoin R ((algebra_map A AX) '' t ∪ set.range (X : _ → AX)) = ⊤, + { rw [adjoin_union_eq_adjoin_adjoin, ← subalgebra.restrict_scalars_top R], + congr' 1, + swap, { exact subalgebra.is_scalar_tower_mid _ }, + rw [adjoin_algebra_map, ht], + apply subalgebra.restrict_scalars_injective R, + rw [← adjoin_restrict_scalars, adjoin_range_X, subalgebra.restrict_scalars_top, + subalgebra.restrict_scalars_top] }, + let g : t → AX := λ x, C (x : A) - map (algebra_map R A) (t' x), + refine ⟨s.image (map (algebra_map R A)) ∪ t.attach.image g, _⟩, + rw [finset.coe_union, finset.coe_image, finset.coe_image, finset.attach_eq_univ, + finset.coe_univ, set.image_univ], + let s₀ := _, let I := _, change ideal.span s₀ = I, + have leI : ideal.span s₀ ≤ I, + { rw [ideal.span_le], + rintros _ (⟨x, hx, rfl⟩|⟨⟨x, hx⟩, rfl⟩), + all_goals + { dsimp [g], rw [ring_hom.mem_ker, alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom] }, + { rw [mv_polynomial.aeval_map_algebra_map, ← aeval_unique], + have := ideal.subset_span hx, + rwa hs at this }, + { rw [map_sub, mv_polynomial.aeval_map_algebra_map, ← aeval_unique, + aeval_C, ht', subtype.coe_mk, sub_self] } }, + apply leI.antisymm, + intros x hx, + rw [ring_hom.mem_ker, alg_hom.to_ring_hom_eq_coe, alg_hom.coe_to_ring_hom] at hx, + let s₀ := _, change x ∈ ideal.span s₀, + have : x ∈ (map (algebra_map R A) : _ →+* AX).srange.to_add_submonoid ⊔ + (ideal.span s₀).to_add_submonoid, + { have : x ∈ (⊤ : subalgebra R AX) := trivial, + rw ← ht'' at this, + apply adjoin_induction this, + { rintros _ (⟨x, hx, rfl⟩|⟨i, rfl⟩), + { rw [algebra_map_eq, ← sub_add_cancel (C x) (map (algebra_map R A) (t' ⟨x, hx⟩)), + add_comm], + apply add_submonoid.add_mem_sup, + { exact set.mem_range_self _ }, + { apply ideal.subset_span, + apply set.mem_union_right, + exact set.mem_range_self ⟨x, hx⟩ } }, + { apply add_submonoid.mem_sup_left, + exact ⟨X i, map_X _ _⟩ } }, + { intro r, apply add_submonoid.mem_sup_left, exact ⟨C r, map_C _ _⟩ }, + { intros _ _ h₁ h₂, exact add_mem h₁ h₂ }, + { intros x₁ x₂ h₁ h₂, + obtain ⟨_, ⟨p₁, rfl⟩, q₁, hq₁, rfl⟩ := add_submonoid.mem_sup.mp h₁, + obtain ⟨_, ⟨p₂, rfl⟩, q₂, hq₂, rfl⟩ := add_submonoid.mem_sup.mp h₂, + rw [add_mul, mul_add, add_assoc, ← map_mul], + apply add_submonoid.add_mem_sup, + { exact set.mem_range_self _ }, + { refine add_mem (ideal.mul_mem_left _ _ hq₂) (ideal.mul_mem_right _ _ hq₁) } } }, + obtain ⟨_, ⟨p, rfl⟩, q, hq, rfl⟩ := add_submonoid.mem_sup.mp this, + rw [map_add, aeval_map_algebra_map, ← aeval_unique, (show aeval (f ∘ X) q = 0, from leI hq), + add_zero] at hx, + suffices : ideal.span (s : set RX) ≤ (ideal.span s₀).comap (map $ algebra_map R A), + { refine add_mem _ hq, rw hs at this, exact this hx }, + rw ideal.span_le, + intros x hx, + apply ideal.subset_span, + apply set.mem_union_left, + exact set.mem_image_of_mem _ hx } +end + /-- This is used to prove the strictly stronger `ker_fg_of_surjective`. Use it instead. -/ +-- TODO: extract out helper lemmas and tidy proof. lemma ker_fg_of_mv_polynomial {n : ℕ} (f : mv_polynomial (fin n) R →ₐ[R] A) (hf : function.surjective f) (hfp : finite_presentation R A) : f.to_ring_hom.ker.fg := begin @@ -631,6 +711,13 @@ lemma comp {g : B →+* C} {f : A →+* B} (hg : g.finite_presentation) (hf : f. end } hf hg +lemma of_comp_finite_type (f : A →+* B) {g : B →+* C} (hg : (g.comp f).finite_presentation) + (hf : f.finite_type) : g.finite_presentation := +@@algebra.finite_presentation.of_restrict_scalars_finite_presentation _ _ f.to_algebra _ + (g.comp f).to_algebra g.to_algebra + (@@is_scalar_tower.of_algebra_map_eq' _ _ _ f.to_algebra g.to_algebra (g.comp f).to_algebra rfl) + hg hf + end finite_presentation end ring_hom @@ -727,6 +814,10 @@ lemma of_finite_type [is_noetherian_ring A] {f : A →ₐ[R] B} : f.finite_type ↔ f.finite_presentation := ring_hom.finite_presentation.of_finite_type +lemma of_comp_finite_type (f : A →ₐ[R] B) {g : B →ₐ[R] C} (h : (g.comp f).finite_presentation) + (h' : f.finite_type) : g.finite_presentation := +h.of_comp_finite_type _ h' + end finite_presentation end alg_hom From d8dd620ccc133d5810714eb1156a2ac680a80ff3 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Mon, 17 Oct 2022 13:11:15 +0000 Subject: [PATCH 798/828] feat(data/polynomial): a couple of lemmas on nat degree and roots of maps (#16992) The nat_degree of sub lemmas complete the (nat/degree)-(add/sub) square. And the map lemma similar to those above lets the user provide injectivity of the map to avoid proving the image is non-zero which is often more convenient. --- src/data/polynomial/degree/definitions.lean | 8 ++++++++ src/data/polynomial/ring_division.lean | 7 +++++++ 2 files changed, 15 insertions(+) diff --git a/src/data/polynomial/degree/definitions.lean b/src/data/polynomial/degree/definitions.lean index d2d843cfeaa80..be1ff5eaf804d 100644 --- a/src/data/polynomial/degree/definitions.lean +++ b/src/data/polynomial/degree/definitions.lean @@ -1073,6 +1073,14 @@ by { rw ← degree_neg q at h, rw [sub_eq_add_neg, degree_add_eq_left_of_degree_ lemma degree_sub_eq_right_of_degree_lt (h : degree p < degree q) : degree (p - q) = degree q := by { rw ← degree_neg q at h, rw [sub_eq_add_neg, degree_add_eq_right_of_degree_lt h, degree_neg] } +lemma nat_degree_sub_eq_left_of_nat_degree_lt (h : nat_degree q < nat_degree p) : + nat_degree (p - q) = nat_degree p := +nat_degree_eq_of_degree_eq (degree_sub_eq_left_of_degree_lt (degree_lt_degree h)) + +lemma nat_degree_sub_eq_right_of_nat_degree_lt (h : nat_degree p < nat_degree q) : + nat_degree (p - q) = nat_degree q := +nat_degree_eq_of_degree_eq (degree_sub_eq_right_of_degree_lt (degree_lt_degree h)) + end ring section nonzero_ring diff --git a/src/data/polynomial/ring_division.lean b/src/data/polynomial/ring_division.lean index fc6cc0503c5e1..ba02c01fba34b 100644 --- a/src/data/polynomial/ring_division.lean +++ b/src/data/polynomial/ring_division.lean @@ -923,6 +923,13 @@ lemma card_roots_le_map [is_domain A] [is_domain B] {p : A[X]} {f : A →+* B} ( p.roots.card ≤ (p.map f).roots.card := by { rw ← p.roots.card_map f, exact multiset.card_le_of_le (map_roots_le h) } +lemma card_roots_le_map_of_injective [is_domain A] [is_domain B] {p : A[X]} {f : A →+* B} + (hf : function.injective f) : p.roots.card ≤ (p.map f).roots.card := +begin + by_cases hp0 : p = 0, { simp only [hp0, roots_zero, polynomial.map_zero, multiset.card_zero], }, + exact card_roots_le_map ((polynomial.map_ne_zero_iff hf).mpr hp0), +end + lemma roots_map_of_injective_of_card_eq_nat_degree [is_domain A] [is_domain B] {p : A[X]} {f : A →+* B} (hf : function.injective f) (hroots : p.roots.card = p.nat_degree) : p.roots.map f = (p.map f).roots := From 8f774b71ea62d466fa0e0112fc23627a56615840 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 18 Oct 2022 04:47:57 +0000 Subject: [PATCH 799/828] feat(combinatorics/simple_graph/regularity/bound): `positivity` extension (#16639) Add a `positivity` extension for `szemeredi_regularity.initial_bound` and `szemeredi_regularity.bound`. --- .../simple_graph/regularity/bound.lean | 26 ++++++++++++++----- 1 file changed, 20 insertions(+), 6 deletions(-) diff --git a/src/combinatorics/simple_graph/regularity/bound.lean b/src/combinatorics/simple_graph/regularity/bound.lean index 4771c319aacf1..b90bf7744f082 100644 --- a/src/combinatorics/simple_graph/regularity/bound.lean +++ b/src/combinatorics/simple_graph/regularity/bound.lean @@ -33,8 +33,7 @@ lemma le_step_bound : id ≤ step_bound := λ n, nat.le_mul_of_pos_right $ pow_p lemma step_bound_mono : monotone step_bound := λ a b h, nat.mul_le_mul h $ nat.pow_le_pow_of_le_right (by norm_num) h -lemma step_bound_pos_iff {n : ℕ} : 0 < step_bound n ↔ 0 < n := -zero_lt_mul_right $ pow_pos (by norm_num) _ +lemma step_bound_pos_iff {n : ℕ} : 0 < step_bound n ↔ 0 < n := zero_lt_mul_right $ by positivity alias step_bound_pos_iff ↔ _ step_bound_pos @@ -62,12 +61,10 @@ pos_of_mul_pos_right ((by norm_num : (0 : ℝ) < 100).trans_le hPε) $ pow_nonne lemma eps_pos (hPε : 100 ≤ 4^P.parts.card * ε^5) : 0 < ε := pow_bit1_pos_iff.1 $ eps_pow_five_pos hPε -lemma four_pow_pos {n : ℕ} : 0 < (4 : ℝ)^n := pow_pos (by norm_num) n - lemma hundred_div_ε_pow_five_le_m [nonempty α] (hPα : P.parts.card * 16^P.parts.card ≤ card α) (hPε : 100 ≤ 4^P.parts.card * ε^5) : 100 / ε^5 ≤ m := -(div_le_of_nonneg_of_le_mul (eps_pow_five_pos hPε).le four_pow_pos.le hPε).trans +(div_le_of_nonneg_of_le_mul (eps_pow_five_pos hPε).le (by positivity) hPε).trans begin norm_cast, rwa [nat.le_div_iff_mul_le'(step_bound_pos (P.parts_nonempty $ univ_nonempty.ne_empty).card_pos), @@ -141,9 +138,26 @@ noncomputable def bound : ℕ := (step_bound^[⌊4 / ε^5⌋₊] $ initial_bound ε l) * 16 ^ (step_bound^[⌊4 / ε^5⌋₊] $ initial_bound ε l) lemma initial_bound_le_bound : initial_bound ε l ≤ bound ε l := -(id_le_iterate_of_id_le le_step_bound _ _).trans $ nat.le_mul_of_pos_right $ pow_pos (by norm_num) _ +(id_le_iterate_of_id_le le_step_bound _ _).trans $ nat.le_mul_of_pos_right $ by positivity lemma le_bound : l ≤ bound ε l := (le_initial_bound ε l).trans $ initial_bound_le_bound ε l lemma bound_pos : 0 < bound ε l := (initial_bound_pos ε l).trans_le $ initial_bound_le_bound ε l end szemeredi_regularity + +namespace tactic +open positivity szemeredi_regularity + +/-- Extension for the `positivity` tactic: `szemeredi_regularity.initial_bound` and +`szemeredi_regularity.bound` are always positive. -/ +@[positivity] +meta def positivity_szemeredi_regularity_bound : expr → tactic strictness +| `(szemeredi_regularity.initial_bound %%ε %%l) := positive <$> mk_app ``initial_bound_pos [ε, l] +| `(szemeredi_regularity.bound %%ε %%l) := positive <$> mk_app ``bound_pos [ε, l] +| e := pp e >>= fail ∘ format.bracket "The expression `" + "` isn't of the form `szemeredi_regularity.initial_bound ε l` nor `szemeredi_regularity.bound ε l`" + +example (ε : ℝ) (l : ℕ) : 0 < szemeredi_regularity.initial_bound ε l := by positivity +example (ε : ℝ) (l : ℕ) : 0 < szemeredi_regularity.bound ε l := by positivity + +end tactic From 38992f943d67b4f200f25686cfcd9371749c8d85 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 18 Oct 2022 05:41:29 +0000 Subject: [PATCH 800/828] feat(analysis/calculus/cont_diff): strengthen lemma (#16933) * The new `cont_diff_within_at_succ_iff_has_fderiv_within_at'` generalizes `cont_diff_within_at.has_fderiv_within_at_nhds` and `cont_diff_within_at_succ_iff_has_fderiv_within_at_of_mem` * Prove that `has_fderiv_within_at` and `cont_diff_within_at` are not dependent on adding a single point to the set * Add some more lemmas needed for a follow-up PR * From the sphere eversion project --- src/analysis/asymptotics/asymptotics.lean | 24 ++++++++ src/analysis/calculus/cont_diff.lean | 70 ++++++++++++++--------- src/analysis/calculus/fderiv.lean | 23 ++++++++ src/data/set/basic.lean | 3 + src/data/set/function.lean | 3 + src/data/set/prod.lean | 5 ++ src/topology/separation.lean | 23 ++++++++ 7 files changed, 124 insertions(+), 27 deletions(-) diff --git a/src/analysis/asymptotics/asymptotics.lean b/src/analysis/asymptotics/asymptotics.lean index be4c66e886f89..970336579d7cd 100644 --- a/src/analysis/asymptotics/asymptotics.lean +++ b/src/analysis/asymptotics/asymptotics.lean @@ -467,6 +467,30 @@ is_o.of_is_O_with $ λ c cpos, (h.forall_is_O_with cpos).sup (h'.forall_is_O_wit @[simp] lemma is_o_sup : f =o[l ⊔ l'] g ↔ f =o[l] g ∧ f =o[l'] g := ⟨λ h, ⟨h.mono le_sup_left, h.mono le_sup_right⟩, λ h, h.1.sup h.2⟩ +lemma is_O_with_insert [topological_space α] {x : α} {s : set α} {C : ℝ} {g : α → E} {g' : α → F} + (h : ∥g x∥ ≤ C * ∥g' x∥) : + is_O_with C (𝓝[insert x s] x) g g' ↔ is_O_with C (𝓝[s] x) g g' := +by simp_rw [is_O_with, nhds_within_insert, eventually_sup, eventually_pure, h, true_and] + +lemma is_O_with.insert [topological_space α] {x : α} {s : set α} {C : ℝ} {g : α → E} {g' : α → F} + (h1 : is_O_with C (𝓝[s] x) g g') (h2 : ∥g x∥ ≤ C * ∥g' x∥) : + is_O_with C (𝓝[insert x s] x) g g' := +(is_O_with_insert h2).mpr h1 + +lemma is_o_insert [topological_space α] {x : α} {s : set α} {g : α → E'} {g' : α → F'} + (h : g x = 0) : g =o[𝓝[insert x s] x] g' ↔ g =o[𝓝[s] x] g' := +begin + simp_rw [is_o], + refine forall_congr (λ c, forall_congr (λ hc, _)), + rw [is_O_with_insert], + rw [h, norm_zero], + exact mul_nonneg hc.le (norm_nonneg _) +end + +lemma is_o.insert [topological_space α] {x : α} {s : set α} {g : α → E'} {g' : α → F'} + (h1 : g =o[𝓝[s] x] g') (h2 : g x = 0) : g =o[𝓝[insert x s] x] g' := +(is_o_insert h2).mpr h1 + /-! ### Simplification : norm, abs -/ section norm_abs diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index 2c3b1f22ecdbc..8d625c2c65085 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -511,6 +511,21 @@ lemma cont_diff_within_at_inter (h : t ∈ 𝓝 x) : cont_diff_within_at 𝕜 n f (s ∩ t) x ↔ cont_diff_within_at 𝕜 n f s x := cont_diff_within_at_inter' (mem_nhds_within_of_mem_nhds h) +lemma cont_diff_within_at_insert {y : E} : + cont_diff_within_at 𝕜 n f (insert y s) x ↔ cont_diff_within_at 𝕜 n f s x := +begin + simp_rw [cont_diff_within_at], + rcases eq_or_ne x y with rfl|h, + { simp_rw [insert_eq_of_mem (mem_insert _ _)] }, + simp_rw [insert_comm x y, nhds_within_insert_of_ne h] +end + +alias cont_diff_within_at_insert ↔ cont_diff_within_at.of_insert cont_diff_within_at.insert' + +lemma cont_diff_within_at.insert (h : cont_diff_within_at 𝕜 n f s x) : + cont_diff_within_at 𝕜 n f (insert x s) x := +h.insert' + /-- If a function is `C^n` within a set at a point, with `n ≥ 1`, then it is differentiable within this set at this point. -/ lemma cont_diff_within_at.differentiable_within_at' @@ -576,36 +591,27 @@ begin rw [snoc_last, init_snoc] } } } } end -/-- One direction of `cont_diff_within_at_succ_iff_has_fderiv_within_at`, but where all derivatives - are taken within the same set. -/ -lemma cont_diff_within_at.has_fderiv_within_at_nhds {n : ℕ} - (hf : cont_diff_within_at 𝕜 (n + 1 : ℕ) f s x) : - ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ∃ f' : E → E →L[𝕜] F, - (∀ x ∈ u, has_fderiv_within_at f (f' x) s x) ∧ cont_diff_within_at 𝕜 n f' s x := -begin - obtain ⟨u, hu, f', huf', hf'⟩ := cont_diff_within_at_succ_iff_has_fderiv_within_at.mp hf, - obtain ⟨w, hw, hxw, hwu⟩ := mem_nhds_within.mp hu, - rw [inter_comm] at hwu, - refine ⟨insert x s ∩ w, inter_mem_nhds_within _ (hw.mem_nhds hxw), inter_subset_left _ _, - f', λ y hy, _, _⟩, - { refine ((huf' y $ hwu hy).mono hwu).mono_of_mem _, - refine mem_of_superset _ (inter_subset_inter_left _ (subset_insert _ _)), - refine inter_mem_nhds_within _ (hw.mem_nhds hy.2) }, - { exact hf'.mono_of_mem (nhds_within_mono _ (subset_insert _ _) hu) } -end - /-- A version of `cont_diff_within_at_succ_iff_has_fderiv_within_at` where all derivatives - are taken within the same set. This lemma assumes `x ∈ s`. -/ -lemma cont_diff_within_at_succ_iff_has_fderiv_within_at_of_mem {n : ℕ} (hx : x ∈ s) : + are taken within the same set. -/ +lemma cont_diff_within_at_succ_iff_has_fderiv_within_at' {n : ℕ} : cont_diff_within_at 𝕜 (n + 1 : ℕ) f s x - ↔ ∃ u ∈ 𝓝[s] x, u ⊆ s ∧ ∃ f' : E → E →L[𝕜] F, + ↔ ∃ u ∈ 𝓝[insert x s] x, u ⊆ insert x s ∧ ∃ f' : E → E →L[𝕜] F, (∀ x ∈ u, has_fderiv_within_at f (f' x) s x) ∧ cont_diff_within_at 𝕜 n f' s x := begin - split, - { intro hf, simpa only [insert_eq_of_mem hx] using hf.has_fderiv_within_at_nhds }, - rw [cont_diff_within_at_succ_iff_has_fderiv_within_at, insert_eq_of_mem hx], - rintro ⟨u, hu, hus, f', huf', hf'⟩, - exact ⟨u, hu, f', λ y hy, (huf' y hy).mono hus, hf'.mono hus⟩ + refine ⟨λ hf, _, _⟩, + { obtain ⟨u, hu, f', huf', hf'⟩ := cont_diff_within_at_succ_iff_has_fderiv_within_at.mp hf, + obtain ⟨w, hw, hxw, hwu⟩ := mem_nhds_within.mp hu, + rw [inter_comm] at hwu, + refine ⟨insert x s ∩ w, inter_mem_nhds_within _ (hw.mem_nhds hxw), inter_subset_left _ _, + f', λ y hy, _, _⟩, + { refine ((huf' y $ hwu hy).mono hwu).mono_of_mem _, + refine mem_of_superset _ (inter_subset_inter_left _ (subset_insert _ _)), + refine inter_mem_nhds_within _ (hw.mem_nhds hy.2) }, + { exact hf'.mono_of_mem (nhds_within_mono _ (subset_insert _ _) hu) } }, + { rw [← cont_diff_within_at_insert, cont_diff_within_at_succ_iff_has_fderiv_within_at, + insert_eq_of_mem (mem_insert _ _)], + rintro ⟨u, hu, hus, f', huf', hf'⟩, + refine ⟨u, hu, f', λ y hy, (huf' y hy).insert'.mono hus, hf'.insert.mono hus⟩ } end /-! ### Smooth functions within a set -/ @@ -1145,7 +1151,8 @@ lemma cont_diff_within_at.fderiv_within' begin have : ∀ k : ℕ, (k + 1 : ℕ∞) ≤ n → cont_diff_within_at 𝕜 k (fderiv_within 𝕜 f s) s x, { intros k hkn, - obtain ⟨v, hv, -, f', hvf', hf'⟩ := (hf.of_le hkn).has_fderiv_within_at_nhds, + obtain ⟨v, hv, -, f', hvf', hf'⟩ := + cont_diff_within_at_succ_iff_has_fderiv_within_at'.mp (hf.of_le hkn), apply hf'.congr_of_eventually_eq_insert, filter_upwards [hv, hs], exact λ y hy h2y, (hvf' y hy).fderiv_within h2y }, @@ -2100,6 +2107,15 @@ begin rwa [insert_eq_of_mem xmem, this] at Z, end +/-- The composition of `C^n` functions at points in domains is `C^n`, + with a weaker condition on `s` and `t`. -/ +lemma cont_diff_within_at.comp_of_mem + {s : set E} {t : set F} {g : F → G} {f : E → F} (x : E) + (hg : cont_diff_within_at 𝕜 n g t (f x)) + (hf : cont_diff_within_at 𝕜 n f s x) (hs : t ∈ 𝓝[f '' s] f x) : + cont_diff_within_at 𝕜 n (g ∘ f) s x := +(hg.mono_of_mem hs).comp x hf (subset_preimage_image f s) + /-- The composition of `C^n` functions at points in domains is `C^n`. -/ lemma cont_diff_within_at.comp' {s : set E} {t : set F} {g : F → G} {f : E → F} (x : E) diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index 7e103ee6a0757..4b4525da1bebc 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -346,6 +346,23 @@ lemma has_fderiv_at.differentiable_at (h : has_fderiv_at f f' x) : differentiabl has_fderiv_within_at f f' univ x ↔ has_fderiv_at f f' x := by { simp only [has_fderiv_within_at, nhds_within_univ], refl } +lemma has_fderiv_within_at_insert {y : E} {g' : E →L[𝕜] F} : + has_fderiv_within_at g g' (insert y s) x ↔ has_fderiv_within_at g g' s x := +begin + rcases eq_or_ne x y with rfl|h, + { simp_rw [has_fderiv_within_at, has_fderiv_at_filter], + apply asymptotics.is_o_insert, + simp only [sub_self, g'.map_zero] }, + refine ⟨λ h, h.mono $ subset_insert y s, λ hg, hg.mono_of_mem _⟩, + simp_rw [nhds_within_insert_of_ne h, self_mem_nhds_within] +end + +alias has_fderiv_within_at_insert ↔ has_fderiv_within_at.of_insert has_fderiv_within_at.insert' + +lemma has_fderiv_within_at.insert {g' : E →L[𝕜] F} (h : has_fderiv_within_at g g' s x) : + has_fderiv_within_at g g' (insert x s) x := +h.insert' + lemma has_strict_fderiv_at.is_O_sub (hf : has_strict_fderiv_at f f' x) : (λ p : E × E, f p.1 - f p.2) =O[𝓝 (x, x)] (λ p : E × E, p.1 - p.2) := hf.is_O.congr_of_sub.2 (f'.is_O_comp _ _) @@ -1086,6 +1103,12 @@ theorem has_fderiv_at.comp_has_fderiv_within_at {g : F → G} {g' : F →L[𝕜] has_fderiv_within_at (g ∘ f) (g'.comp f') s x := hg.comp x hf hf.continuous_within_at +theorem has_fderiv_within_at.comp_of_mem {g : F → G} {g' : F →L[𝕜] G} {t : set F} + (hg : has_fderiv_within_at g g' t (f x)) (hf : has_fderiv_within_at f f' s x) + (hst : tendsto f (𝓝[s] x) (𝓝[t] f x)) : + has_fderiv_within_at (g ∘ f) (g'.comp f') s x := +has_fderiv_at_filter.comp x hg hf hst + /-- The chain rule. -/ theorem has_fderiv_at.comp {g : F → G} {g' : F →L[𝕜] G} (hg : has_fderiv_at g g' (f x)) (hf : has_fderiv_at f f' x) : diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 730d083b0ad55..270116ff4841d 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -997,6 +997,9 @@ theorem diff_eq (s t : set α) : s \ t = s ∩ tᶜ := rfl theorem mem_diff_of_mem {s t : set α} {x : α} (h1 : x ∈ s) (h2 : x ∉ t) : x ∈ s \ t := ⟨h1, h2⟩ +lemma not_mem_diff_of_mem {s t : set α} {x : α} (hx : x ∈ t) : x ∉ s \ t := +λ h, h.2 hx + theorem mem_of_mem_diff {s t : set α} {x : α} (h : x ∈ s \ t) : x ∈ s := h.left diff --git a/src/data/set/function.lean b/src/data/set/function.lean index a2d9c62abda54..8eae815d3b7e0 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -291,6 +291,9 @@ lemma maps_to_iff_exists_map_subtype : maps_to f s t ↔ ∃ g : s → t, ∀ x theorem maps_to' : maps_to f s t ↔ f '' s ⊆ t := image_subset_iff.symm +lemma maps_to.subset_preimage {f : α → β} {s : set α} {t : set β} (hf : maps_to f s t) : + s ⊆ f ⁻¹' t := hf + @[simp] theorem maps_to_singleton {x : α} : maps_to f {x} t ↔ f x ∈ t := singleton_subset_iff theorem maps_to_empty (f : α → β) (t : set β) : maps_to f ∅ t := empty_subset _ diff --git a/src/data/set/prod.lean b/src/data/set/prod.lean index ca10b1cf8880b..6123fe3f0f57c 100644 --- a/src/data/set/prod.lean +++ b/src/data/set/prod.lean @@ -192,6 +192,11 @@ lemma prod_sub_preimage_iff {W : set γ} {f : α × β → γ} : s ×ˢ t ⊆ f ⁻¹' W ↔ ∀ a b, a ∈ s → b ∈ t → f (a, b) ∈ W := by simp [subset_def] + +lemma image_prod_mk_subset_prod {f : α → β} {g : α → γ} {s : set α} : + (λ x, (f x, g x)) '' s ⊆ (f '' s) ×ˢ (g '' s) := +by { rintros _ ⟨x, hx, rfl⟩, exact mk_mem_prod (mem_image_of_mem f hx) (mem_image_of_mem g hx) } + lemma image_prod_mk_subset_prod_left (hb : b ∈ t) : (λ a, (a, b)) '' s ⊆ s ×ˢ t := by { rintro _ ⟨a, ha, rfl⟩, exact ⟨ha, hb⟩ } diff --git a/src/topology/separation.lean b/src/topology/separation.lean index 5d40e4cbbec0c..8d2d90fc6f76d 100644 --- a/src/topology/separation.lean +++ b/src/topology/separation.lean @@ -543,6 +543,29 @@ lemma is_closed_map_const {α β} [topological_space α] [topological_space β] is_closed_map (function.const α y) := is_closed_map.of_nonempty $ λ s hs h2s, by simp_rw [h2s.image_const, is_closed_singleton] +lemma nhds_within_insert_of_ne [t1_space α] {x y : α} {s : set α} (hxy : x ≠ y) : + 𝓝[insert y s] x = 𝓝[s] x := +begin + refine le_antisymm (λ t ht, _) (nhds_within_mono x $ subset_insert y s), + obtain ⟨o, ho, hxo, host⟩ := mem_nhds_within.mp ht, + refine mem_nhds_within.mpr ⟨o \ {y}, ho.sdiff is_closed_singleton, ⟨hxo, hxy⟩, _⟩, + rw [inter_insert_of_not_mem $ not_mem_diff_of_mem (mem_singleton y)], + exact (inter_subset_inter (diff_subset _ _) subset.rfl).trans host +end + +/-- If `t` is a subset of `s`, except for one point, +then `insert x s` is a neighborhood of `x` within `t`. -/ +lemma insert_mem_nhds_within_of_subset_insert [t1_space α] {x y : α} {s t : set α} + (hu : t ⊆ insert y s) : + insert x s ∈ 𝓝[t] x := +begin + rcases eq_or_ne x y with rfl|h, + { exact mem_of_superset self_mem_nhds_within hu }, + refine nhds_within_mono x hu _, + rw [nhds_within_insert_of_ne h], + exact mem_of_superset self_mem_nhds_within (subset_insert x s) +end + lemma bInter_basis_nhds [t1_space α] {ι : Sort*} {p : ι → Prop} {s : ι → set α} {x : α} (h : (𝓝 x).has_basis p s) : (⋂ i (h : p i), s i) = {x} := begin From a5347cb91abfcc702e3b92dc72a33f9e8bd4a3e7 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Tue, 18 Oct 2022 08:40:57 +0000 Subject: [PATCH 801/828] feat(*): add various order-related lemmas (#16918) These all came up while preparing to introduce the definition of ergodic maps but it seems better to PR them separately. Co-authored-by: Johan Commelin --- src/data/nat/basic.lean | 2 + src/data/set/basic.lean | 14 +++- src/dynamics/fixed_points/basic.lean | 4 + src/measure_theory/measurable_space.lean | 15 +++- src/measure_theory/measure/measure_space.lean | 80 ++++++++++++++++++ .../measure/measure_space_def.lean | 12 +++ src/order/complete_lattice.lean | 35 ++++++-- src/order/liminf_limsup.lean | 84 +++++++++++++++++++ 8 files changed, 239 insertions(+), 7 deletions(-) diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index 5a97adb64eb0a..42322740b9539 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -150,6 +150,8 @@ open set theorem zero_union_range_succ : {0} ∪ range succ = univ := by { ext n, cases n; simp } +@[simp] protected lemma range_succ : range succ = {i | 0 < i} := by ext (_ | i); simp [succ_pos] + variables {α : Type*} theorem range_of_succ (f : ℕ → α) : {f 0} ∪ range (f ∘ succ) = range f := diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 270116ff4841d..8fbc8e1318458 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ import order.symm_diff +import logic.function.iterate /-! # Basic properties of sets @@ -1359,7 +1360,9 @@ rfl @[simp] theorem preimage_set_of_eq {p : α → Prop} {f : β → α} : f ⁻¹' {a | p a} = {a | p (f a)} := rfl -@[simp] theorem preimage_id {s : set α} : id ⁻¹' s = s := rfl +@[simp] lemma preimage_id_eq : preimage (id : α → α) = id := rfl + +theorem preimage_id {s : set α} : id ⁻¹' s = s := rfl @[simp] theorem preimage_id' {s : set α} : (λ x, x) ⁻¹' s = s := rfl @@ -1377,6 +1380,15 @@ by { split_ifs with hb hb, exacts [preimage_const_of_mem hb, preimage_const_of_n theorem preimage_comp {s : set γ} : (g ∘ f) ⁻¹' s = f ⁻¹' (g ⁻¹' s) := rfl +lemma preimage_comp_eq : preimage (g ∘ f) = preimage f ∘ preimage g := rfl + +@[simp] lemma preimage_iterate_eq {f : α → α} {n : ℕ} : + set.preimage (f^[n]) = ((set.preimage f)^[n]) := +begin + induction n with n ih, { simp, }, + rw [iterate_succ, iterate_succ', set.preimage_comp_eq, ih], +end + lemma preimage_preimage {g : β → γ} {f : α → β} {s : set γ} : f ⁻¹' (g ⁻¹' s) = (λ x, g (f x)) ⁻¹' s := preimage_comp.symm diff --git a/src/dynamics/fixed_points/basic.lean b/src/dynamics/fixed_points/basic.lean index 27623adab8802..5799eab98318e 100644 --- a/src/dynamics/fixed_points/basic.lean +++ b/src/dynamics/fixed_points/basic.lean @@ -71,6 +71,10 @@ calc fb (g x) = g (fa x) : (h.eq x).symm protected lemma apply {x : α} (hx : is_fixed_pt f x) : is_fixed_pt f (f x) := by convert hx +lemma preimage_iterate {s : set α} (h : is_fixed_pt (set.preimage f) s) (n : ℕ) : + is_fixed_pt (set.preimage (f^[n])) s := +by { rw set.preimage_iterate_eq, exact h.iterate n, } + end is_fixed_pt @[simp] lemma injective.is_fixed_pt_apply_iff (hf : injective f) {x : α} : diff --git a/src/measure_theory/measurable_space.lean b/src/measure_theory/measurable_space.lean index 0c4493e754935..11acb85e19cb9 100644 --- a/src/measure_theory/measurable_space.lean +++ b/src/measure_theory/measurable_space.lean @@ -267,7 +267,6 @@ begin exact (hf ht).inter h.measurable_set.of_compl, end - end measurable_functions section constructions @@ -1462,4 +1461,18 @@ instance : boolean_algebra (subtype (measurable_set : set α → Prop)) := .. measurable_set.subtype.bounded_order, .. measurable_set.subtype.distrib_lattice } +@[measurability] lemma measurable_set_limsup {s : ℕ → set α} (hs : ∀ n, measurable_set $ s n) : + measurable_set $ filter.limsup s filter.at_top := +begin + simp only [filter.limsup_eq_infi_supr_of_nat', supr_eq_Union, infi_eq_Inter], + exact measurable_set.Inter (λ n, measurable_set.Union $ λ m, hs $ m + n), +end + +@[measurability] lemma measurable_set_liminf {s : ℕ → set α} (hs : ∀ n, measurable_set $ s n) : + measurable_set $ filter.liminf s filter.at_top := +begin + simp only [filter.liminf_eq_supr_infi_of_nat', supr_eq_Union, infi_eq_Inter], + exact measurable_set.Union (λ n, measurable_set.Inter $ λ m, hs $ m + n), +end + end measurable_set diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 3b978df81e1fd..e8f41673d5561 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -527,6 +527,42 @@ begin exact λ ⟨i, hi⟩, ⟨i + (m - n), by simpa only [add_assoc, tsub_add_cancel_of_le hnm] using hi⟩ end +lemma measure_liminf_eq_zero {s : ℕ → set α} (h : ∑' i, μ (s i) ≠ ⊤) : μ (liminf s at_top) = 0 := +begin + rw ← le_zero_iff, + have : liminf s at_top ≤ limsup s at_top := + liminf_le_limsup (by is_bounded_default) (by is_bounded_default), + exact (μ.mono this).trans (by simp [measure_limsup_eq_zero h]), +end + +lemma limsup_ae_eq_of_forall_ae_eq (s : ℕ → set α) {t : set α} (h : ∀ n, s n =ᵐ[μ] t) : + -- Need `@` below because of diamond; see gh issue #16932 + @limsup (set α) ℕ _ s at_top =ᵐ[μ] t := +begin + simp_rw ae_eq_set at h ⊢, + split, + { rw at_top.limsup_sdiff s t, + apply measure_limsup_eq_zero, + simp [h], }, + { rw at_top.sdiff_limsup s t, + apply measure_liminf_eq_zero, + simp [h], }, +end + +lemma liminf_ae_eq_of_forall_ae_eq (s : ℕ → set α) {t : set α} (h : ∀ n, s n =ᵐ[μ] t) : + -- Need `@` below because of diamond; see gh issue #16932 + @liminf (set α) ℕ _ s at_top =ᵐ[μ] t := +begin + simp_rw ae_eq_set at h ⊢, + split, + { rw at_top.liminf_sdiff s t, + apply measure_liminf_eq_zero, + simp [h], }, + { rw at_top.sdiff_liminf s t, + apply measure_limsup_eq_zero, + simp [h], }, +end + lemma measure_if {x : β} {t : set β} {s : set α} : μ (if x ∈ t then s else ∅) = indicator t (λ _, μ s) x := by { split_ifs; simp [h] } @@ -2006,6 +2042,32 @@ lemma preimage_null (h : quasi_measure_preserving f μa μb) {s : set β} (hs : μa (f ⁻¹' s) = 0 := preimage_null_of_map_null h.ae_measurable (h.2 hs) +lemma limsup_preimage_iterate_ae_eq {f : α → α} (hf : quasi_measure_preserving f μ μ) + (hs : f⁻¹' s =ᵐ[μ] s) : + -- Need `@` below because of diamond; see gh issue #16932 + @limsup (set α) ℕ _ (λ n, (preimage f)^[n] s) at_top =ᵐ[μ] s := +begin + have : ∀ n, (preimage f)^[n] s =ᵐ[μ] s, + { intros n, + induction n with n ih, { simp, }, + simpa only [iterate_succ', comp_app] using ae_eq_trans (hf.ae_eq ih) hs, }, + exact (limsup_ae_eq_of_forall_ae_eq (λ n, (preimage f)^[n] s) this).trans (ae_eq_refl _), +end + +lemma liminf_preimage_iterate_ae_eq {f : α → α} (hf : quasi_measure_preserving f μ μ) + (hs : f⁻¹' s =ᵐ[μ] s) : + -- Need `@` below because of diamond; see gh issue #16932 + @liminf (set α) ℕ _ (λ n, (preimage f)^[n] s) at_top =ᵐ[μ] s := +begin + -- Need `@` below because of diamond; see gh issue #16932 + rw [← ae_eq_set_compl_compl, @filter.liminf_compl (set α)], + rw [← ae_eq_set_compl_compl, ← preimage_compl] at hs, + convert hf.limsup_preimage_iterate_ae_eq hs, + ext1 n, + simp only [← set.preimage_iterate_eq, comp_app, preimage_compl], +end + + end quasi_measure_preserving /-! ### The `cofinite` filter -/ @@ -2530,6 +2592,24 @@ lemma is_probability_measure_map [is_probability_measure μ] {f : α → β} (hf is_probability_measure (map f μ) := ⟨by simp [map_apply_of_ae_measurable, hf]⟩ +@[simp] lemma one_le_prob_iff [is_probability_measure μ] : 1 ≤ μ s ↔ μ s = 1 := +⟨λ h, le_antisymm prob_le_one h, λ h, h ▸ le_refl _⟩ + +/-- Note that this is not quite as useful as it looks because the measure takes values in `ℝ≥0∞`. +Thus the subtraction appearing is the truncated subtraction of `ℝ≥0∞`, rather than the +better-behaved subtraction of `ℝ`. -/ +lemma prob_compl_eq_one_sub [is_probability_measure μ] (hs : measurable_set s) : + μ sᶜ = 1 - μ s := +by simpa only [measure_univ] using measure_compl hs (measure_lt_top μ s).ne + +@[simp] lemma prob_compl_eq_zero_iff [is_probability_measure μ] (hs : measurable_set s) : + μ sᶜ = 0 ↔ μ s = 1 := +by simp only [prob_compl_eq_one_sub hs, tsub_eq_zero_iff_le, one_le_prob_iff] + +@[simp] lemma prob_compl_eq_one_iff [is_probability_measure μ] (hs : measurable_set s) : + μ sᶜ = 1 ↔ μ s = 0 := +by rwa [← prob_compl_eq_zero_iff hs.compl, compl_compl] + end is_probability_measure section no_atoms diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index d406543fa6104..ca841de9c6bf9 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -376,6 +376,18 @@ lemma ae_eq_set {s t : set α} : s =ᵐ[μ] t ↔ μ (s \ t) = 0 ∧ μ (t \ s) = 0 := by simp [eventually_le_antisymm_iff, ae_le_set] +@[simp] lemma measure_symm_diff_eq_zero_iff {s t : set α} : + μ (s ∆ t) = 0 ↔ s =ᵐ[μ] t := +by simp [ae_eq_set, symm_diff_def] + +@[simp] lemma ae_eq_set_compl_compl {s t : set α} : + sᶜ =ᵐ[μ] tᶜ ↔ s =ᵐ[μ] t := +by simp only [← measure_symm_diff_eq_zero_iff, compl_symm_diff_compl] + +lemma ae_eq_set_compl {s t : set α} : + sᶜ =ᵐ[μ] t ↔ s =ᵐ[μ] tᶜ := +by rw [← ae_eq_set_compl_compl, compl_compl] + lemma ae_eq_set_inter {s' t' : set α} (h : s =ᵐ[μ] t) (h' : s' =ᵐ[μ] t') : (s ∩ s' : set α) =ᵐ[μ] (t ∩ t' : set α) := h.inter h' diff --git a/src/order/complete_lattice.lean b/src/order/complete_lattice.lean index 565b68593117f..5c17ea84ba601 100644 --- a/src/order/complete_lattice.lean +++ b/src/order/complete_lattice.lean @@ -818,6 +818,15 @@ supr_subtype lemma infi_subtype'' {ι} (s : set ι) (f : ι → α) : (⨅ i : s, f i) = ⨅ (t : ι) (H : t ∈ s), f t := infi_subtype +lemma bsupr_const {ι : Sort*} {a : α} {s : set ι} (hs : s.nonempty) : (⨆ i ∈ s, a) = a := +begin + haveI : nonempty s := set.nonempty_coe_sort.mpr hs, + rw [← supr_subtype'', supr_const], +end + +lemma binfi_const {ι : Sort*} {a : α} {s : set ι} (hs : s.nonempty) : (⨅ i ∈ s, a) = a := +@bsupr_const αᵒᵈ _ ι _ s hs + theorem supr_sup_eq : (⨆ x, f x ⊔ g x) = (⨆ x, f x) ⊔ (⨆ x, g x) := le_antisymm (supr_le $ λ i, sup_le_sup (le_supr _ _) $ le_supr _ _) @@ -846,14 +855,22 @@ by rw [supr_sup_eq, supr_const] lemma inf_infi [nonempty ι] {f : ι → α} {a : α} : a ⊓ (⨅ x, f x) = ⨅ x, a ⊓ f x := by rw [infi_inf_eq, infi_const] -lemma binfi_inf {p : ι → Prop} {f : Π i (hi : p i), α} {a : α} (h : ∃ i, p i) : - (⨅ i (h : p i), f i h) ⊓ a = ⨅ i (h : p i), f i h ⊓ a := +lemma bsupr_sup {p : ι → Prop} {f : Π i, p i → α} {a : α} (h : ∃ i, p i) : + (⨆ i (h : p i), f i h) ⊔ a = ⨆ i (h : p i), f i h ⊔ a := by haveI : nonempty {i // p i} := (let ⟨i, hi⟩ := h in ⟨⟨i, hi⟩⟩); - rw [infi_subtype', infi_subtype', infi_inf] + rw [supr_subtype', supr_subtype', supr_sup] -lemma inf_binfi {p : ι → Prop} {f : Π i (hi : p i), α} {a : α} (h : ∃ i, p i) : +lemma sup_bsupr {p : ι → Prop} {f : Π i, p i → α} {a : α} (h : ∃ i, p i) : + a ⊔ (⨆ i (h : p i), f i h) = ⨆ i (h : p i), a ⊔ f i h := +by simpa only [sup_comm] using bsupr_sup h + +lemma binfi_inf {p : ι → Prop} {f : Π i, p i → α} {a : α} (h : ∃ i, p i) : + (⨅ i (h : p i), f i h) ⊓ a = ⨅ i (h : p i), f i h ⊓ a := +@bsupr_sup αᵒᵈ ι _ p f _ h + +lemma inf_binfi {p : ι → Prop} {f : Π i, p i → α} {a : α} (h : ∃ i, p i) : a ⊓ (⨅ i (h : p i), f i h) = ⨅ i (h : p i), a ⊓ f i h := -by simpa only [inf_comm] using binfi_inf h +@sup_bsupr αᵒᵈ ι _ p f _ h /-! ### `supr` and `infi` under `Prop` -/ @@ -1161,6 +1178,14 @@ end lemma inf_infi_nat_succ (u : ℕ → α) : u 0 ⊓ (⨅ i, u (i + 1)) = ⨅ i, u i := @sup_supr_nat_succ αᵒᵈ _ u +lemma infi_nat_gt_zero_eq (f : ℕ → α) : + (⨅ (i > 0), f i) = ⨅ i, f (i + 1) := +by simpa only [(by simp : ∀ (i : ℕ), i > 0 ↔ i ∈ range nat.succ)] using infi_range + +lemma supr_nat_gt_zero_eq (f : ℕ → α) : + (⨆ (i > 0), f i) = ⨆ i, f (i + 1) := +@infi_nat_gt_zero_eq αᵒᵈ _ f + end section complete_linear_order diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index 279f595ef96de..ff6dd40b3ec18 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel, Johannes Hölzl, Rémy Degenne -/ import order.filter.cofinite +import order.hom.complete_lattice /-! # liminfs and limsups of functions and filters @@ -607,6 +608,25 @@ lemma le_limsup_of_frequently_le' {α β} [complete_lattice β] x ≤ limsup u f := @liminf_le_of_frequently_le' _ βᵒᵈ _ _ _ _ h +/-- If `f : α → α` is a morphism of complete lattices, then the limsup of its iterates of any +`a : α` is a fixed point. -/ +@[simp] lemma complete_lattice_hom.apply_limsup_iterate (f : complete_lattice_hom α α) (a : α) : + f (limsup (λ n, f^[n] a) at_top) = limsup (λ n, f^[n] a) at_top := +begin + rw [limsup_eq_infi_supr_of_nat', map_infi], + simp_rw [_root_.map_supr, ← function.comp_apply f, ← function.iterate_succ' f, ← nat.add_succ], + conv_rhs { rw infi_split _ ((<) (0 : ℕ)), }, + simp only [not_lt, le_zero_iff, infi_infi_eq_left, add_zero, infi_nat_gt_zero_eq, left_eq_inf], + refine (infi_le (λ i, ⨆ j, (f^[j + (i + 1)]) a) 0).trans _, + simp only [zero_add, function.comp_app, supr_le_iff], + exact λ i, le_supr (λ i, (f^[i] a)) (i + 1), +end + +/-- If `f : α → α` is a morphism of complete lattices, then the liminf of its iterates of any +`a : α` is a fixed point. -/ +lemma complete_lattice_hom.apply_liminf_iterate (f : complete_lattice_hom α α) (a : α) : + f (liminf (λ n, f^[n] a) at_top) = liminf (λ n, f^[n] a) at_top := +(complete_lattice_hom.dual f).apply_limsup_iterate _ variables {f g : filter β} {p q : β → Prop} {u v : β → α} lemma blimsup_mono (h : ∀ x, p x → q x) : @@ -678,8 +698,72 @@ end bliminf u f (λ x, p x ∨ q x) = bliminf u f p ⊓ bliminf u f q := @blimsup_or_eq_sup αᵒᵈ β _ f p q u +lemma sup_limsup [ne_bot f] (a : α) : + a ⊔ limsup u f = limsup (λ x, a ⊔ u x) f := +begin + simp only [limsup_eq_infi_supr, supr_sup_eq, sup_binfi_eq], + congr, ext s, congr, ext hs, congr, + exact (bsupr_const (nonempty_of_mem hs)).symm, +end + +lemma inf_liminf [ne_bot f] (a : α) : + a ⊓ liminf u f = liminf (λ x, a ⊓ u x) f := +@sup_limsup αᵒᵈ β _ f _ _ _ + +lemma sup_liminf (a : α) : + a ⊔ liminf u f = liminf (λ x, a ⊔ u x) f := +begin + simp only [liminf_eq_supr_infi], + rw [sup_comm, bsupr_sup (⟨univ, univ_mem⟩ : ∃ (i : set β), i ∈ f)], + simp_rw [binfi_sup_eq, @sup_comm _ _ a], +end + +lemma inf_limsup (a : α) : + a ⊓ limsup u f = limsup (λ x, a ⊓ u x) f := +@sup_liminf αᵒᵈ β _ f _ _ + end complete_distrib_lattice +section complete_boolean_algebra + +variables [complete_boolean_algebra α] (f : filter β) (u : β → α) + +lemma limsup_compl : + (limsup u f)ᶜ = liminf (compl ∘ u) f := +by simp only [limsup_eq_infi_supr, liminf_eq_supr_infi, compl_infi, compl_supr] + +lemma liminf_compl : + (liminf u f)ᶜ = limsup (compl ∘ u) f := +by simp only [limsup_eq_infi_supr, liminf_eq_supr_infi, compl_infi, compl_supr] + +lemma limsup_sdiff (a : α) : + (limsup u f) \ a = limsup (λ b, (u b) \ a) f := +begin + simp only [limsup_eq_infi_supr, sdiff_eq], + rw binfi_inf (⟨univ, univ_mem⟩ : ∃ (i : set β), i ∈ f), + simp_rw [inf_comm, inf_bsupr_eq, inf_comm], +end + +lemma liminf_sdiff [ne_bot f] (a : α) : + (liminf u f) \ a = liminf (λ b, (u b) \ a) f := +by simp only [sdiff_eq, @inf_comm _ _ _ aᶜ, inf_liminf] + +lemma sdiff_limsup [ne_bot f] (a : α) : + a \ limsup u f = liminf (λ b, a \ u b) f := +begin + rw ← compl_inj_iff, + simp only [sdiff_eq, liminf_compl, (∘), compl_inf, compl_compl, sup_limsup], +end + +lemma sdiff_liminf (a : α) : + a \ liminf u f = limsup (λ b, a \ u b) f := +begin + rw ← compl_inj_iff, + simp only [sdiff_eq, limsup_compl, (∘), compl_inf, compl_compl, sup_liminf], +end + +end complete_boolean_algebra + section set_lattice variables {p : ι → Prop} {s : ι → set α} From 6d4f61075550afb3e4c3d969637320d42ec2a196 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 18 Oct 2022 08:40:58 +0000 Subject: [PATCH 802/828] feat(scripts/port_status): script to show currently portable files (#17009) This was written by Patrick, and I made a few tweaks. Co-authored-by: Scott Morrison --- .gitignore | 1 + scripts/port_status.py | 65 ++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 66 insertions(+) create mode 100755 scripts/port_status.py diff --git a/.gitignore b/.gitignore index 889a3819ca4e8..ed583e63ad380 100644 --- a/.gitignore +++ b/.gitignore @@ -9,3 +9,4 @@ all.lean *~ .DS_Store *.lock +port_status.yaml diff --git a/scripts/port_status.py b/scripts/port_status.py new file mode 100755 index 0000000000000..b36981b642b0a --- /dev/null +++ b/scripts/port_status.py @@ -0,0 +1,65 @@ +#!/usr/bin/env python3 +import os +import re +import yaml +import networkx as nx +from urllib.request import urlopen + +from pathlib import Path + +import_re = re.compile(r"^import ([^ ]*)") + +def mk_label(path: Path) -> str: + rel = path.relative_to(Path('src')) + return str(rel.with_suffix('')).replace(os.sep, '.') + +graph = nx.DiGraph() + +for path in Path('src').glob('**/*.lean'): + if path.parts[1] in ['tactic', 'meta']: + continue + graph.add_node(mk_label(path)) + +for path in Path('src').glob('**/*.lean'): + if path.parts[1] in ['tactic', 'meta']: + continue + for line in path.read_text().split('\n'): + if (m := import_re.match(line)): + imported = m.group(1) + if imported.startswith('tactic.') or imported.startswith('meta.'): + continue + if imported not in graph.nodes: + if imported + '.default' in graph.nodes: + imported = imported + '.default' + else: + imported = 'lean_core.' + imported + graph.add_edge(imported, mk_label(path)) + +data = yaml.safe_load(urlopen('https://raw.githubusercontent.com/wiki/leanprover-community/mathlib/mathlib4-port-status.md')) + +# First make sure all nodes exists in the data set +for node in graph.nodes: + data.setdefault(node, 'No') +yaml.dump(data, Path('port_status.yaml').open('w')) + +allDone = dict() +parentsDone = dict() +for node in graph.nodes: + ancestors = nx.ancestors(graph, node) + if all(data[imported] == 'Yes' for imported in ancestors) and data[node] == 'No': + allDone[node] = len(nx.descendants(graph, node)) + else: + if all(data[imported] == 'Yes' for imported in graph.predecessors(node)) and data[node] == 'No': + parentsDone[node] = len(nx.descendants(graph, node)) + +print('# The following files have all dependencies ported already, and should be ready to port:') +print('# Earlier items in the list are required in more places in mathlib.') +allDone = dict(sorted(allDone.items(), key=lambda item: -item[1])) +for k in allDone: + print(k) + +print() +print('# The following files have their immediate dependencies ported already, and may be ready to port:') +parentsDone = dict(sorted(parentsDone.items(), key=lambda item: -item[1])) +for k in parentsDone: + print(k) From 3812c1cec157b0dee3c3d4c0fe9b9e70bfc065bd Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 18 Oct 2022 14:35:38 +0000 Subject: [PATCH 803/828] chore(order/lattice): add lemmas about `sup` and `inf` on `prod` (#17040) The lemmas are inspired by the corresponding lemmas for add and mul --- src/order/lattice.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/src/order/lattice.lean b/src/order/lattice.lean index 63ae03cc0a8a4..0310c324c4278 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -912,6 +912,24 @@ variables (α β) instance [has_sup α] [has_sup β] : has_sup (α × β) := ⟨λp q, ⟨p.1 ⊔ q.1, p.2 ⊔ q.2⟩⟩ instance [has_inf α] [has_inf β] : has_inf (α × β) := ⟨λp q, ⟨p.1 ⊓ q.1, p.2 ⊓ q.2⟩⟩ +@[simp] lemma mk_sup_mk [has_sup α] [has_sup β] (a₁ a₂ : α) (b₁ b₂ : β) : + (a₁, b₁) ⊔ (a₂, b₂) = (a₁ ⊔ a₂, b₁ ⊔ b₂) := rfl + +@[simp] lemma mk_inf_mk [has_inf α] [has_inf β] (a₁ a₂ : α) (b₁ b₂ : β) : + (a₁, b₁) ⊓ (a₂, b₂) = (a₁ ⊓ a₂, b₁ ⊓ b₂) := rfl + +@[simp] lemma fst_sup [has_sup α] [has_sup β] (p q : α × β) : (p ⊔ q).fst = p.fst ⊔ q.fst := rfl +@[simp] lemma fst_inf [has_inf α] [has_inf β] (p q : α × β) : (p ⊓ q).fst = p.fst ⊓ q.fst := rfl + +@[simp] lemma snd_sup [has_sup α] [has_sup β] (p q : α × β) : (p ⊔ q).snd = p.snd ⊔ q.snd := rfl +@[simp] lemma snd_inf [has_inf α] [has_inf β] (p q : α × β) : (p ⊓ q).snd = p.snd ⊓ q.snd := rfl + +@[simp] lemma swap_sup [has_sup α] [has_sup β] (p q : α × β) : (p ⊔ q).swap = p.swap ⊔ q.swap := rfl +@[simp] lemma swap_inf [has_inf α] [has_inf β] (p q : α × β) : (p ⊓ q).swap = p.swap ⊓ q.swap := rfl + +lemma sup_def [has_sup α] [has_sup β] (p q : α × β) : p ⊔ q = (p.fst ⊔ q.fst, p.snd ⊔ q.snd) := rfl +lemma inf_def [has_inf α] [has_inf β] (p q : α × β) : p ⊓ q = (p.fst ⊓ q.fst, p.snd ⊓ q.snd) := rfl + instance [semilattice_sup α] [semilattice_sup β] : semilattice_sup (α × β) := { sup_le := assume a b c h₁ h₂, ⟨sup_le h₁.1 h₂.1, sup_le h₁.2 h₂.2⟩, le_sup_left := assume a b, ⟨le_sup_left, le_sup_left⟩, From 0b480ff9797f715d9fd9e27c051e818e88a607fa Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Tue, 18 Oct 2022 17:37:28 +0000 Subject: [PATCH 804/828] feat(measure_theory/measure): `measure_Union_null_iff` for Props (#17029) from #2819, this enables the simplifier to deal with proving measure zero-ness of a Union indexed by elements satisfying some condition --- src/measure_theory/measure/measure_space_def.lean | 7 +++++++ src/measure_theory/measure/outer_measure.lean | 7 +++++++ 2 files changed, 14 insertions(+) diff --git a/src/measure_theory/measure/measure_space_def.lean b/src/measure_theory/measure/measure_space_def.lean index ca841de9c6bf9..82dfa24b09c8b 100644 --- a/src/measure_theory/measure/measure_space_def.lean +++ b/src/measure_theory/measure/measure_space_def.lean @@ -225,6 +225,13 @@ lemma measure_Union_null [countable β] {s : β → set α} : (∀ i, μ (s i) = μ (⋃ i, s i) = 0 ↔ ∀ i, μ (s i) = 0 := μ.to_outer_measure.Union_null_iff +/-- A version of `measure_Union_null_iff` for unions indexed by Props +TODO: in the long run it would be better to combine this with `measure_Union_null_iff` by +generalising to `Sort`. -/ +@[simp] lemma measure_Union_null_iff' {ι : Prop} {s : ι → set α} : + μ (⋃ i, s i) = 0 ↔ ∀ i, μ (s i) = 0 := +μ.to_outer_measure.Union_null_iff' + lemma measure_bUnion_null_iff {s : set ι} (hs : s.countable) {t : ι → set α} : μ (⋃ i ∈ s, t i) = 0 ↔ ∀ i ∈ s, μ (t i) = 0 := μ.to_outer_measure.bUnion_null_iff hs diff --git a/src/measure_theory/measure/outer_measure.lean b/src/measure_theory/measure/outer_measure.lean index 577bd61ffb9c6..c599c8dae5320 100644 --- a/src/measure_theory/measure/outer_measure.lean +++ b/src/measure_theory/measure/outer_measure.lean @@ -99,6 +99,13 @@ by simpa [h] using m.Union s m (⋃ i, s i) = 0 ↔ ∀ i, m (s i) = 0 := ⟨λ h i, m.mono_null (subset_Union _ _) h, m.Union_null⟩ +/-- A version of `Union_null_iff` for unions indexed by Props. +TODO: in the long run it would be better to combine this with `Union_null_iff` by +generalising to `Sort`. -/ +@[simp] lemma Union_null_iff' (m : outer_measure α) {ι : Prop} {s : ι → set α} : + m (⋃ i, s i) = 0 ↔ ∀ i, m (s i) = 0 := +by by_cases i : ι; simp [i] + lemma bUnion_null_iff (m : outer_measure α) {s : set β} (hs : s.countable) {t : β → set α} : m (⋃ i ∈ s, t i) = 0 ↔ ∀ i ∈ s, m (t i) = 0 := by { haveI := hs.to_encodable, rw [bUnion_eq_Union, Union_null_iff, set_coe.forall'] } From 3d197c697faef85110c1df1b5681ad99ec9f0555 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 18 Oct 2022 17:37:29 +0000 Subject: [PATCH 805/828] fix(topology/metric_space/gromov_hausdorff): squeeze all simps (#17049) --- .../metric_space/gromov_hausdorff.lean | 57 +++++++++++-------- 1 file changed, 33 insertions(+), 24 deletions(-) diff --git a/src/topology/metric_space/gromov_hausdorff.lean b/src/topology/metric_space/gromov_hausdorff.lean index 01f8b3a330e73..c91e2bde6a02c 100644 --- a/src/topology/metric_space/gromov_hausdorff.lean +++ b/src/topology/metric_space/gromov_hausdorff.lean @@ -213,8 +213,12 @@ begin have BY : ⟦B⟧ = to_GH_space Y, { rw eq_to_GH_space_iff, exact ⟨λ x, F (Ψ' x), (Kuratowski_embedding.isometry _).comp IΨ', range_comp _ _⟩ }, - refine cInf_le ⟨0, - begin simp [lower_bounds], assume t _ _ _ _ ht, rw ← ht, exact Hausdorff_dist_nonneg end⟩ _, + refine cInf_le ⟨0, _⟩ _, + { simp only [lower_bounds, mem_image, mem_prod, mem_set_of_eq, prod.exists, and_imp, + forall_exists_index], + assume t _ _ _ _ ht, + rw ← ht, + exact Hausdorff_dist_nonneg }, apply (mem_image _ _ _).2, existsi (⟨A, B⟩ : nonempty_compacts ℓ_infty_ℝ × nonempty_compacts ℓ_infty_ℝ), simp [AX, BY], @@ -305,7 +309,7 @@ begin have : z ∈ range Ψ, by rwa [← Ψrange] at zq, rcases mem_range.1 this with ⟨y, hy⟩, calc (⨅ y, Fb (inl x, inr y)) ≤ Fb (inl x, inr y) : - cinfi_le (by simpa using HD_below_aux1 0) y + cinfi_le (by simpa only [add_zero] using HD_below_aux1 0) y ... = dist (Φ x) (Ψ y) : rfl ... = dist (f (inl x)) z : by rw hy ... ≤ r : le_of_lt hz }, @@ -319,11 +323,11 @@ begin have : z ∈ range Φ, by rwa [← Φrange] at zq, rcases mem_range.1 this with ⟨x, hx⟩, calc (⨅ x, Fb (inl x, inr y)) ≤ Fb (inl x, inr y) : - cinfi_le (by simpa using HD_below_aux2 0) x + cinfi_le (by simpa only [add_zero] using HD_below_aux2 0) x ... = dist (Φ x) (Ψ y) : rfl ... = dist z (f (inr y)) : by rw hx ... ≤ r : le_of_lt hz }, - simp [HD, csupr_le I1, csupr_le I2] }, + simp only [HD, csupr_le I1, csupr_le I2, max_le_iff, and_self] }, /- Get the same inequality for any coupling. If the coupling is quite good, the desired inequality has been proved above. If it is bad, then the inequality is obvious. -/ have B : ∀ p q : nonempty_compacts ℓ_infty_ℝ, ⟦p⟧ = to_GH_space X → ⟦q⟧ = to_GH_space Y → @@ -372,7 +376,9 @@ instance : metric_space GH_space := refine le_antisymm _ _, { apply cInf_le, { exact ⟨0, by { rintro b ⟨⟨u, v⟩, ⟨hu, hv⟩, rfl⟩, exact Hausdorff_dist_nonneg } ⟩}, - { simp, existsi [y, y], simpa } }, + { simp only [mem_image, mem_prod, mem_set_of_eq, prod.exists], + existsi [y, y], + simpa only [and_self, Hausdorff_dist_self_zero, eq_self_iff_true, and_true]} }, { apply le_cInf, { exact (nonempty.prod ⟨y, hy⟩ ⟨y, hy⟩).image _ }, { rintro b ⟨⟨u, v⟩, ⟨hu, hv⟩, rfl⟩, exact Hausdorff_dist_nonneg } }, @@ -383,8 +389,8 @@ instance : metric_space GH_space := ({a | ⟦a⟧ = x} ×ˢ {b | ⟦b⟧ = y}) = ((λ (p : nonempty_compacts ℓ_infty_ℝ × nonempty_compacts ℓ_infty_ℝ), Hausdorff_dist (p.1 : set ℓ_infty_ℝ) p.2) ∘ prod.swap) '' - ({a | ⟦a⟧ = x} ×ˢ {b | ⟦b⟧ = y}) := - by { congr, funext, simp, rw Hausdorff_dist_comm }, + ({a | ⟦a⟧ = x} ×ˢ {b | ⟦b⟧ = y}), + { congr, funext, simp only [comp_app, prod.fst_swap, prod.snd_swap], rw Hausdorff_dist_comm }, simp only [dist, A, image_comp, image_swap_prod], end, eq_of_dist_eq_zero := λ x y hxy, begin @@ -562,7 +568,7 @@ begin { rw [← image_univ, Hausdorff_dist_image Il], have : 0 ≤ ε₁ := le_trans dist_nonneg Dxs, refine Hausdorff_dist_le_of_mem_dist this (λ x hx, hs x) - (λ x hx, ⟨x, mem_univ _, by simpa⟩) }, + (λ x hx, ⟨x, mem_univ _, by simpa only [dist_self]⟩) }, have : Hausdorff_dist (Fl '' s) (Fr '' (range Φ)) ≤ ε₂/2 + δ, { refine Hausdorff_dist_le_of_mem_dist (by linarith) _ _, { assume x' hx', @@ -581,7 +587,8 @@ begin rcases exists_mem_of_nonempty Y with ⟨xY, _⟩, rcases hs' xY with ⟨xs', Dxs'⟩, have : 0 ≤ ε₃ := le_trans dist_nonneg Dxs', - refine Hausdorff_dist_le_of_mem_dist this (λ x hx, ⟨x, mem_univ _, by simpa⟩) (λ x _, _), + refine Hausdorff_dist_le_of_mem_dist this (λ x hx, ⟨x, mem_univ _, by simpa only [dist_self]⟩) + (λ x _, _), rcases hs' x with ⟨y, Dy⟩, exact ⟨Φ y, mem_range_self _, Dy⟩ }, linarith @@ -595,7 +602,8 @@ begin let ε := (2/5) * δ, have εpos : 0 < ε := mul_pos (by norm_num) δpos, have : ∀ p:GH_space, ∃ s : set p.rep, s.finite ∧ (univ ⊆ (⋃x∈s, ball x ε)) := - λ p, by simpa using finite_cover_balls_of_compact (@compact_univ p.rep _ _) εpos, + λ p, by simpa only [subset_univ, exists_true_left] + using finite_cover_balls_of_compact compact_univ εpos, -- for each `p`, `s p` is a finite `ε`-dense subset of `p` (or rather the metric space -- `p.rep` representing `p`) choose s hs using this, @@ -649,9 +657,9 @@ begin have C1 : (E p) z = ⟨i, hip⟩ := (E p).apply_symm_apply ⟨i, hip⟩, have C2 : fin.cast Npq ⟨i, hip⟩ = ⟨i, hi⟩ := rfl, have C3 : (E q).symm ⟨i, hi⟩ = ⟨y, ys⟩, - by { rw ihi_eq, exact (E q).symm_apply_apply ⟨y, ys⟩ }, - have : Φ z = y := - by { simp only [Φ, Ψ], rw [C1, C2, C3], refl }, + { rw ihi_eq, exact (E q).symm_apply_apply ⟨y, ys⟩ }, + have : Φ z = y, + { simp only [Φ, Ψ], rw [C1, C2, C3], refl }, rw this, exact le_of_lt hy }, show ∀ x y : s p, |dist x y - dist (Φ x) (Φ y)| ≤ ε, @@ -665,12 +673,13 @@ begin let i : ℕ := E p x, have hip : i < N p := ((E p) x).2, have hiq : i < N q, by rwa Npq at hip, - have i' : i = ((E q) (Ψ x)), by { simp [Ψ] }, + have i' : i = ((E q) (Ψ x)), by { simp only [equiv.apply_symm_apply, fin.coe_cast] }, -- introduce `j`, that codes both `y` and `Φ y` in `fin (N p) = fin (N q)` let j : ℕ := E p y, have hjp : j < N p := ((E p) y).2, have hjq : j < N q, by rwa Npq at hjp, - have j' : j = ((E q) (Ψ y)).1, by { simp [Ψ] }, + have j' : j = ((E q) (Ψ y)).1, + { simp only [equiv.apply_symm_apply, fin.val_eq_coe, fin.coe_cast] }, -- Express `dist x y` in terms of `F p` have : (F p).2 ((E p) x) ((E p) y) = floor (ε⁻¹ * dist x y), by simp only [F, (E p).symm_apply_apply], @@ -710,7 +719,7 @@ begin ... = ε : mul_one _ } }, calc dist p q = GH_dist p.rep (q.rep) : dist_GH_dist p q ... ≤ ε + ε/2 + ε : main - ... = δ : by { simp [ε], ring } + ... = δ : by { simp only [ε], ring } end /-- Compactness criterion: a closed set of compact metric spaces is compact if the spaces have @@ -743,7 +752,7 @@ begin { assume p, by_cases hp : p ∉ t, { have : nonempty (equiv (∅ : set p.rep) (fin 0)), - { rw ← fintype.card_eq, simp }, + { rw ← fintype.card_eq, simp only [empty_card', fintype.card_fin] }, use [∅, 0, bot_le, choice (this)] }, { rcases hcov _ (set.not_not_mem.1 hp) n with ⟨s, ⟨scard, scover⟩⟩, rcases cardinal.lt_aleph_0.1 (lt_of_le_of_lt scard (cardinal.nat_lt_aleph_0 _)) with ⟨N, hN⟩, @@ -751,7 +760,7 @@ begin have : cardinal.mk s = cardinal.mk (fin N), by rw [hN, cardinal.mk_fin], cases quotient.exact this with E, use [s, N, scard, E], - simp [hp, scover] } }, + simp only [scover, implies_true_iff] } }, choose s N hN E hs using this, -- Define a function `F` taking values in a finite type and associating to `p` enough data -- to reconstruct it up to `ε`, namely the (discretized) distances between elements of `s p`. @@ -808,12 +817,12 @@ begin let i : ℕ := E p x, have hip : i < N p := ((E p) x).2, have hiq : i < N q, by rwa Npq at hip, - have i' : i = ((E q) (Ψ x)), by { simp [Ψ] }, + have i' : i = ((E q) (Ψ x)), by { simp only [equiv.apply_symm_apply, fin.coe_cast] }, -- introduce `j`, that codes both `y` and `Φ y` in `fin (N p) = fin (N q)` let j : ℕ := E p y, have hjp : j < N p := ((E p) y).2, have hjq : j < N q, by rwa Npq at hjp, - have j' : j = ((E q) (Ψ y)), by { simp [Ψ] }, + have j' : j = ((E q) (Ψ y)), by { simp only [equiv.apply_symm_apply, fin.coe_cast] }, -- Express `dist x y` in terms of `F p` have Ap : ((F p).2 ⟨i, hip⟩ ⟨j, hjp⟩).1 = ⌊ε⁻¹ * dist x y⌋₊ := calc ((F p).2 ⟨i, hip⟩ ⟨j, hjp⟩).1 = ((F p).2 ((E p) x) ((E p) y)).1 : @@ -878,7 +887,7 @@ begin ... = ε : mul_one _ } }, calc dist p q = GH_dist p.rep (q.rep) : dist_GH_dist p q ... ≤ ε + ε/2 + ε : main - ... = δ/2 : by { simp [ε], ring } + ... = δ/2 : by { simp only [ε, one_div], ring } ... < δ : half_lt_self δpos end @@ -943,7 +952,7 @@ begin letI : ∀ n, metric_space (Y n).space := λ n, (Y n).metric, have E : ∀ n : ℕ, glue_space (Y n).isom (isometry_optimal_GH_injl (X n) (X n.succ)) = (Y n.succ).space := - λ n, by { simp [Y, aux_gluing], refl }, + λ n, by { simp only [Y, aux_gluing], refl }, let c := λ n, cast (E n), have ic : ∀ n, isometry (c n) := λ n x y, rfl, -- there is a canonical embedding of `Y n` in `Y (n+1)`, by construction @@ -1011,7 +1020,7 @@ begin constructor, convert (isom n).isometric_on_range.symm, }, -- Finally, we have proved the convergence of `u n` - exact ⟨L.to_GH_space, by simpa [this] using M⟩ + exact ⟨L.to_GH_space, by simpa only [this] using M⟩ end end complete--section From 4af7699b69c11c9d1f7256bea607707456b1a0a2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 18 Oct 2022 18:51:23 +0000 Subject: [PATCH 806/828] feat(order/upper_lower): Maps of upper sets (#17007) An order isomorphism of preorders induces an order isomorphisms of their upper sets. --- src/data/set/lattice.lean | 20 +++- src/order/hom/basic.lean | 5 + src/order/upper_lower.lean | 213 +++++++++++++++++++++++++++++++++++-- 3 files changed, 224 insertions(+), 14 deletions(-) diff --git a/src/data/set/lattice.lean b/src/data/set/lattice.lean index 69f34fce9cbdf..7e145fd807d59 100644 --- a/src/data/set/lattice.lean +++ b/src/data/set/lattice.lean @@ -511,9 +511,8 @@ end /-! ### Unions and intersections indexed by `Prop` -/ -@[simp] theorem Inter_false {s : false → set α} : Inter s = univ := infi_false - -@[simp] theorem Union_false {s : false → set α} : Union s = ∅ := supr_false +theorem Inter_false {s : false → set α} : Inter s = univ := infi_false +theorem Union_false {s : false → set α} : Union s = ∅ := supr_false @[simp] theorem Inter_true {s : true → set α} : Inter s = s trivial := infi_true @@ -982,6 +981,9 @@ by simp only [←sUnion_range, subtype.range_coe] lemma sInter_eq_Inter {s : set (set α)} : (⋂₀ s) = (⋂ (i : s), i) := by simp only [←sInter_range, subtype.range_coe] +@[simp] lemma Union_of_empty [is_empty ι] (s : ι → set α) : (⋃ i, s i) = ∅ := supr_of_empty _ +@[simp] lemma Inter_of_empty [is_empty ι] (s : ι → set α) : (⋂ i, s i) = univ := infi_of_empty _ + lemma union_eq_Union {s₁ s₂ : set α} : s₁ ∪ s₂ = ⋃ b : bool, cond b s₁ s₂ := sup_eq_supr s₁ s₂ @@ -1190,6 +1192,18 @@ begin simpa only [Union, supr_subtype'] using h end +lemma image_Inter {f : α → β} (hf : bijective f) (s : ι → set α) : + f '' (⋂ i, s i) = ⋂ i, f '' s i := +begin + casesI is_empty_or_nonempty ι, + { simp_rw [Inter_of_empty, image_univ_of_surjective hf.surjective] }, + { exact (hf.injective.inj_on _).image_Inter_eq } +end + +lemma image_Inter₂ {f : α → β} (hf : bijective f) (s : Π i, κ i → set α) : + f '' (⋂ i j, s i j) = ⋂ i j, f '' s i j := +by simp_rw image_Inter hf + lemma inj_on_Union_of_directed {s : ι → set α} (hs : directed (⊆) s) {f : α → β} (hf : ∀ i, inj_on f (s i)) : inj_on f (⋃ i, s i) := diff --git a/src/order/hom/basic.lean b/src/order/hom/basic.lean index d27ae18c68fca..ff1bbee6744f2 100644 --- a/src/order/hom/basic.lean +++ b/src/order/hom/basic.lean @@ -590,6 +590,11 @@ e.to_equiv.preimage_image s @[simp] lemma trans_refl (e : α ≃o β) : e.trans (refl β) = e := by { ext x, refl } +@[simp] lemma symm_trans_apply (e₁ : α ≃o β) (e₂ : β ≃o γ) (c : γ) : + (e₁.trans e₂).symm c = e₁.symm (e₂.symm c) := rfl + +lemma symm_trans (e₁ : α ≃o β) (e₂ : β ≃o γ) : (e₁.trans e₂).symm = e₂.symm.trans e₁.symm := rfl + /-- `prod.swap` as an `order_iso`. -/ def prod_comm : (α × β) ≃o (β × α) := { to_equiv := equiv.prod_comm α β, diff --git a/src/order/upper_lower.lean b/src/order/upper_lower.lean index 559ae43dc7afd..0ceca9db42920 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower.lean @@ -39,12 +39,12 @@ Lattice structure on antichains. Order equivalence between upper/lower sets and open order_dual set -variables {α β : Type*} {ι : Sort*} {κ : ι → Sort*} +variables {α β γ : Type*} {ι : Sort*} {κ : ι → Sort*} /-! ### Unbundled upper/lower sets -/ section has_le -variables [has_le α] {s t : set α} +variables [has_le α] [has_le β] {s t : set α} /-- An upper set in an order `α` is a set such that any element greater than one of its members is also a member. Also called up-set, upward-closed set. -/ @@ -136,7 +136,7 @@ alias is_upper_set_preimage_to_dual_iff ↔ _ is_lower_set.to_dual end has_le section preorder -variables [preorder α] {s : set α} (a : α) +variables [preorder α] [preorder β] {s : set α} (a : α) lemma is_upper_set_Ici : is_upper_set (Ici a) := λ _ _, ge_trans lemma is_lower_set_Iic : is_lower_set (Iic a) := λ _ _, le_trans @@ -158,6 +158,22 @@ lemma is_upper_set.ord_connected (h : is_upper_set s) : s.ord_connected := lemma is_lower_set.ord_connected (h : is_lower_set s) : s.ord_connected := ⟨λ a _ b hb, Icc_subset_Iic_self.trans $ h.Iic_subset hb⟩ +lemma is_upper_set.preimage (hs : is_upper_set s) {f : β → α} (hf : monotone f) : + is_upper_set (f ⁻¹' s : set β) := +λ x y hxy, hs $ hf hxy + +lemma is_lower_set.preimage (hs : is_lower_set s) {f : β → α} (hf : monotone f) : + is_lower_set (f ⁻¹' s : set β) := +λ x y hxy, hs $ hf hxy + +lemma is_upper_set.image (hs : is_upper_set s) (f : α ≃o β) : is_upper_set (f '' s : set β) := +by { change is_upper_set ((f : α ≃ β) '' s), rw set.image_equiv_eq_preimage_symm, + exact hs.preimage f.symm.monotone } + +lemma is_lower_set.image (hs : is_lower_set s) (f : α ≃o β) : is_lower_set (f '' s : set β) := +by { change is_lower_set ((f : α ≃ β) '' s), rw set.image_equiv_eq_preimage_symm, + exact hs.preimage f.symm.monotone } + section order_top variables [order_top α] @@ -217,6 +233,26 @@ lemma not_bdd_below_Iio : ¬ bdd_below (Iio a) := (is_lower_set_Iio _).not_bdd_b end no_min_order end preorder +section partial_order +variables [partial_order α] {s : set α} + +lemma is_upper_set_iff_forall_lt : is_upper_set s ↔ ∀ ⦃a b : α⦄, a < b → a ∈ s → b ∈ s := +forall_congr $ λ a, by simp [le_iff_eq_or_lt, or_imp_distrib, forall_and_distrib] + +lemma is_lower_set_iff_forall_lt : is_lower_set s ↔ ∀ ⦃a b : α⦄, b < a → a ∈ s → b ∈ s := +forall_congr $ λ a, by simp [le_iff_eq_or_lt, or_imp_distrib, forall_and_distrib] + +lemma is_upper_set_iff_Ioi_subset : is_upper_set s ↔ ∀ ⦃a⦄, a ∈ s → Ioi a ⊆ s := +by simp [is_upper_set_iff_forall_lt, subset_def, @forall_swap (_ ∈ s)] + +lemma is_lower_set_iff_Iio_subset : is_lower_set s ↔ ∀ ⦃a⦄, a ∈ s → Iio a ⊆ s := +by simp [is_lower_set_iff_forall_lt, subset_def, @forall_swap (_ ∈ s)] + +alias is_upper_set_iff_Ioi_subset ↔ is_upper_set.Ioi_subset _ +alias is_lower_set_iff_Iio_subset ↔ is_lower_set.Iio_subset _ + +end partial_order + /-! ### Bundled upper/lower sets -/ section has_le @@ -244,6 +280,9 @@ instance : set_like (upper_set α) α := protected lemma upper (s : upper_set α) : is_upper_set (s : set α) := s.upper' +@[simp] lemma mem_mk (carrier : set α) (upper') {a : α} : a ∈ mk carrier upper' ↔ a ∈ carrier := +iff.rfl + end upper_set namespace lower_set @@ -258,6 +297,9 @@ instance : set_like (lower_set α) α := protected lemma lower (s : lower_set α) : is_lower_set (s : set α) := s.lower' +@[simp] lemma mem_mk (carrier : set α) (lower') {a : α} : a ∈ mk carrier lower' ↔ a ∈ carrier := +iff.rfl + end lower_set /-! #### Order -/ @@ -454,11 +496,124 @@ end lower_set end has_le +/-! #### Map -/ + +section +variables [preorder α] [preorder β] [preorder γ] + +namespace upper_set +variables {f : α ≃o β} {s t : upper_set α} {a : α} {b : β} + +/-- An order isomorphism of preorders induces an order isomorphism of their upper sets. -/ +def map (f : α ≃o β) : upper_set α ≃o upper_set β := +{ to_fun := λ s, ⟨f '' s, s.upper.image f⟩, + inv_fun := λ t, ⟨f ⁻¹' t, t.upper.preimage f.monotone⟩, + left_inv := λ _, ext $ f.preimage_image _, + right_inv := λ _, ext $ f.image_preimage _, + map_rel_iff' := λ s t, image_subset_image_iff f.injective } + +@[simp] lemma symm_map (f : α ≃o β) : (map f).symm = map f.symm := +fun_like.ext _ _ $ λ s, ext $ set.preimage_equiv_eq_image_symm _ _ + +@[simp] lemma mem_map : b ∈ map f s ↔ f.symm b ∈ s := +by { rw [←f.symm_symm, ←symm_map, f.symm_symm], refl } + +@[simp] lemma map_refl : map (order_iso.refl α) = order_iso.refl _ := by { ext, simp } + +@[simp] lemma map_map (g : β ≃o γ) (f : α ≃o β) : map g (map f s) = map (f.trans g) s := +by { ext, simp } + +variables (f s t) + +@[simp, norm_cast] lemma coe_map : (map f s : set β) = f '' s := rfl + +@[simp] protected lemma map_sup : map f (s ⊔ t) = map f s ⊔ map f t := +ext $ (image_inter f.injective).symm +@[simp] protected lemma map_inf : map f (s ⊓ t) = map f s ⊓ map f t := ext $ image_union _ _ _ +@[simp] protected lemma map_top : map f ⊤ = ⊤ := ext $ image_empty _ +@[simp] protected lemma map_bot : map f ⊥ = ⊥ := ext $ image_univ_of_surjective f.surjective +@[simp] protected lemma map_Sup (S : set (upper_set α)) : map f (Sup S) = ⨆ s ∈ S, map f s := +ext $ by { push_cast, exact image_Inter₂ f.bijective _ } + +@[simp] protected lemma map_Inf (S : set (upper_set α)) : map f (Inf S) = ⨅ s ∈ S, map f s := +ext $ by { push_cast, exact image_Union₂ _ _ } + +@[simp] protected lemma map_supr (g : ι → upper_set α) : map f (⨆ i, g i) = ⨆ i, map f (g i) := +ext $ by { push_cast, exact image_Inter f.bijective _ } + +@[simp] protected lemma map_infi (g : ι → upper_set α) : map f (⨅ i, g i) = ⨅ i, map f (g i) := +ext $ by { push_cast, exact image_Union } + +end upper_set + +namespace lower_set +variables {f : α ≃o β} {s t : lower_set α} {a : α} {b : β} + +/-- An order isomorphism of preorders induces an order isomorphism of their lower sets. -/ +def map (f : α ≃o β) : lower_set α ≃o lower_set β := +{ to_fun := λ s, ⟨f '' s, s.lower.image f⟩, + inv_fun := λ t, ⟨f ⁻¹' t, t.lower.preimage f.monotone⟩, + left_inv := λ _, set_like.coe_injective $ f.preimage_image _, + right_inv := λ _, set_like.coe_injective $ f.image_preimage _, + map_rel_iff' := λ s t, image_subset_image_iff f.injective } + +@[simp] lemma symm_map (f : α ≃o β) : (map f).symm = map f.symm := +fun_like.ext _ _ $ λ s, set_like.coe_injective $ set.preimage_equiv_eq_image_symm _ _ + +@[simp] lemma mem_map {f : α ≃o β} {b : β} : b ∈ map f s ↔ f.symm b ∈ s := +by { rw [←f.symm_symm, ←symm_map, f.symm_symm], refl } + +@[simp] lemma map_refl : map (order_iso.refl α) = order_iso.refl _ := by { ext, simp } + +@[simp] lemma map_map (g : β ≃o γ) (f : α ≃o β) : map g (map f s) = map (f.trans g) s := +by { ext, simp } + +variables (f s t) + +@[simp, norm_cast] lemma coe_map : (map f s : set β) = f '' s := rfl + +@[simp] protected lemma map_sup : map f (s ⊔ t) = map f s ⊔ map f t := ext $ image_union _ _ _ +@[simp] protected lemma map_inf : map f (s ⊓ t) = map f s ⊓ map f t := +ext $ (image_inter f.injective).symm +@[simp] protected lemma map_top : map f ⊤ = ⊤ := ext $ image_univ_of_surjective f.surjective +@[simp] protected lemma map_bot : map f ⊥ = ⊥ := ext $ image_empty _ +@[simp] protected lemma map_Sup (S : set (lower_set α)) : map f (Sup S) = ⨆ s ∈ S, map f s := +ext $ by { push_cast, exact image_Union₂ _ _ } + +protected lemma map_Inf (S : set (lower_set α)) : map f (Inf S) = ⨅ s ∈ S, map f s := +ext $ by { push_cast, exact image_Inter₂ f.bijective _ } + +protected lemma map_supr (g : ι → lower_set α) : map f (⨆ i, g i) = ⨆ i, map f (g i) := +ext $ by { push_cast, exact image_Union } + +protected lemma map_infi (g : ι → lower_set α) : map f (⨅ i, g i) = ⨅ i, map f (g i) := +ext $ by { push_cast, exact image_Inter f.bijective _ } + +end lower_set + +namespace upper_set + +@[simp] lemma compl_map (f : α ≃o β) (s : upper_set α) : + (map f s).compl = lower_set.map f s.compl := +set_like.coe_injective (set.image_compl_eq f.bijective).symm + +end upper_set + +namespace lower_set + +@[simp] lemma compl_map (f : α ≃o β) (s : lower_set α) : + (map f s).compl = upper_set.map f s.compl := +set_like.coe_injective (set.image_compl_eq f.bijective).symm + +end lower_set + +end + /-! #### Principal sets -/ namespace upper_set section preorder -variables [preorder α] {a b : α} +variables [preorder α] [preorder β] {s : upper_set α} {a b : α} /-- The smallest upper set containing a given element. -/ def Ici (a : α) : upper_set α := ⟨Ici a, is_upper_set_Ici a⟩ @@ -470,6 +625,8 @@ def Ioi (a : α) : upper_set α := ⟨Ioi a, is_upper_set_Ioi a⟩ @[simp] lemma coe_Ioi (a : α) : ↑(Ioi a) = set.Ioi a := rfl @[simp] lemma mem_Ici_iff : b ∈ Ici a ↔ a ≤ b := iff.rfl @[simp] lemma mem_Ioi_iff : b ∈ Ioi a ↔ a < b := iff.rfl +@[simp] lemma map_Ici (f : α ≃o β) (a : α) : map f (Ici a) = Ici (f a) := by { ext, simp } +@[simp] lemma map_Ioi (f : α ≃o β) (a : α) : map f (Ioi a) = Ioi (f a) := by { ext, simp } lemma Ici_le_Ioi (a : α) : Ici a ≤ Ioi a := Ioi_subset_Ici_self @@ -512,7 +669,7 @@ end upper_set namespace lower_set section preorder -variables [preorder α] {a b : α} +variables [preorder α] [preorder β] {s : lower_set α} {a b : α} /-- Principal lower set. `set.Iic` as a lower set. The smallest lower set containing a given element. -/ @@ -525,6 +682,8 @@ def Iio (a : α) : lower_set α := ⟨Iio a, is_lower_set_Iio a⟩ @[simp] lemma coe_Iio (a : α) : ↑(Iio a) = set.Iio a := rfl @[simp] lemma mem_Iic_iff : b ∈ Iic a ↔ b ≤ a := iff.rfl @[simp] lemma mem_Iio_iff : b ∈ Iio a ↔ b < a := iff.rfl +@[simp] lemma map_Iic (f : α ≃o β) (a : α) : map f (Iic a) = Iic (f a) := by { ext, simp } +@[simp] lemma map_Iio (f : α ≃o β) (a : α) : map f (Iio a) = Iio (f a) := by { ext, simp } lemma Ioi_le_Ici (a : α) : Ioi a ≤ Ici a := Ioi_subset_Ici_self @@ -569,7 +728,7 @@ end complete_lattice end lower_set section closure -variables [preorder α] {s t : set α} {x : α} +variables [preorder α] [preorder β] {s t : set α} {x : α} /-- The greatest upper set containing a given set. -/ def upper_closure (s : set α) : upper_set α := @@ -579,11 +738,9 @@ def upper_closure (s : set α) : upper_set α := def lower_closure (s : set α) : lower_set α := ⟨{x | ∃ a ∈ s, x ≤ a}, λ x y h, Exists₂.imp $ λ a _, h.trans⟩ -@[simp, norm_cast] lemma coe_upper_closure (s : set α) : - ↑(upper_closure s) = {x | ∃ a ∈ s, a ≤ x} := rfl - -@[simp, norm_cast] lemma coe_lower_closure (s : set α) : - ↑(lower_closure s) = {x | ∃ a ∈ s, x ≤ a} := rfl +-- We do not tag those two as `simp` to respect the abstraction. +@[norm_cast] lemma coe_upper_closure (s : set α) : ↑(upper_closure s) = {x | ∃ a ∈ s, a ≤ x} := rfl +@[norm_cast] lemma coe_lower_closure (s : set α) : ↑(lower_closure s) = {x | ∃ a ∈ s, x ≤ a} := rfl @[simp] lemma mem_upper_closure : x ∈ upper_closure s ↔ ∃ a ∈ s, a ≤ x := iff.rfl @[simp] lemma mem_lower_closure : x ∈ lower_closure s ↔ ∃ a ∈ s, x ≤ a := iff.rfl @@ -597,6 +754,34 @@ lemma upper_closure_min (h : s ⊆ t) (ht : is_upper_set t) : ↑(upper_closure lemma lower_closure_min (h : s ⊆ t) (ht : is_lower_set t) : ↑(lower_closure s) ⊆ t := λ a ⟨b, hb, hab⟩, ht hab $ h hb +protected lemma is_upper_set.upper_closure (hs : is_upper_set s) : ↑(upper_closure s) = s := +(upper_closure_min subset.rfl hs).antisymm subset_upper_closure + +protected lemma is_lower_set.lower_closure (hs : is_lower_set s) : ↑(lower_closure s) = s := +(lower_closure_min subset.rfl hs).antisymm subset_lower_closure + +@[simp] protected lemma upper_set.upper_closure (s : upper_set α) : upper_closure (s : set α) = s := +set_like.coe_injective s.2.upper_closure + +@[simp] protected lemma lower_set.lower_closure (s : lower_set α) : lower_closure (s : set α) = s := +set_like.coe_injective s.2.lower_closure + +@[simp] lemma upper_closure_image (f : α ≃o β) : + upper_closure (f '' s) = upper_set.map f (upper_closure s) := +begin + rw [←f.symm_symm, ←upper_set.symm_map, f.symm_symm], + ext, + simp [-upper_set.symm_map, upper_set.map, order_iso.symm, ←f.le_symm_apply], +end + +@[simp] lemma lower_closure_image (f : α ≃o β) : + lower_closure (f '' s) = lower_set.map f (lower_closure s) := +begin + rw [←f.symm_symm, ←lower_set.symm_map, f.symm_symm], + ext, + simp [-lower_set.symm_map, lower_set.map, order_iso.symm, ←f.symm_apply_le], +end + @[simp] lemma upper_set.infi_Ici (s : set α) : (⨅ a ∈ s, upper_set.Ici a) = upper_closure s := by { ext, simp } @@ -637,6 +822,12 @@ gc_lower_closure_coe.monotone_l @[simp] lemma upper_closure_empty : upper_closure (∅ : set α) = ⊤ := by { ext, simp } @[simp] lemma lower_closure_empty : lower_closure (∅ : set α) = ⊥ := by { ext, simp } +@[simp] lemma upper_closure_singleton (a : α) : upper_closure ({a} : set α) = upper_set.Ici a := +by { ext, simp } + +@[simp] lemma lower_closure_singleton (a : α) : lower_closure ({a} : set α) = lower_set.Iic a := +by { ext, simp } + @[simp] lemma upper_closure_univ : upper_closure (univ : set α) = ⊥ := le_bot_iff.1 subset_upper_closure From 8c9b005780be967e44ba1b761c70b91c509034a8 Mon Sep 17 00:00:00 2001 From: Andrew Yang <36414270+erdOne@users.noreply.github.com> Date: Tue, 18 Oct 2022 18:51:24 +0000 Subject: [PATCH 807/828] fix(data/set/basic): Fix `set.lt_iff_ssubset`. (#17047) --- src/data/set/basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index 8fbc8e1318458..b2dee2c64a91e 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -106,7 +106,7 @@ instance : has_inter (set α) := ⟨(⊓)⟩ @[simp] lemma lt_eq_ssubset : ((<) : set α → set α → Prop) = (⊂) := rfl lemma le_iff_subset : s ≤ t ↔ s ⊆ t := iff.rfl -lemma lt_iff_ssubset : s ≤ t ↔ s ⊆ t := iff.rfl +lemma lt_iff_ssubset : s < t ↔ s ⊂ t := iff.rfl alias le_iff_subset ↔ _root_.has_le.le.subset _root_.has_subset.subset.le alias lt_iff_ssubset ↔ _root_.has_lt.lt.ssubset _root_.has_ssubset.ssubset.lt From 9407b03373c8cd201df99d6bc5514fc2db44054f Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 18 Oct 2022 21:51:47 +0000 Subject: [PATCH 808/828] chore(logic/equiv/basic): split into two files (#17038) A fairly easy split of `logic.equiv.basic` into `logic.equiv.defs` (containing the stuff with very little prerequisites, and which suffices for many other files), and `logic.equiv.basic` (the stuff that requires other imports from around `logic.*`). `logic.equiv.basic` is still a ~1500 line file, so I'd like to split it further shortly. Co-authored-by: Scott Morrison --- src/algebra/free.lean | 2 +- src/algebra/group/type_tags.lean | 2 +- src/algebra/group/with_one.lean | 2 +- src/algebra/opposites.lean | 2 +- .../functor/fully_faithful.lean | 2 +- src/category_theory/types.lean | 2 +- src/combinatorics/derangements/basic.lean | 2 +- src/control/equiv_functor.lean | 2 +- src/control/monad/basic.lean | 2 +- src/control/monad/writer.lean | 2 +- src/control/traversable/equiv.lean | 2 +- src/data/erased.lean | 2 +- src/data/opposite.lean | 2 +- src/data/part.lean | 2 +- src/logic/embedding.lean | 2 +- src/logic/equiv/basic.lean | 822 +----------------- src/logic/equiv/defs.lean | 798 +++++++++++++++++ src/logic/equiv/fin.lean | 2 +- src/logic/equiv/fintype.lean | 2 +- src/logic/equiv/functor.lean | 2 +- src/logic/equiv/local_equiv.lean | 2 +- src/logic/equiv/set.lean | 2 +- src/logic/equiv/transfer_instance.lean | 2 +- src/order/bounded_order.lean | 1 + src/order/synonym.lean | 2 +- src/tactic/equiv_rw.lean | 2 +- src/tactic/norm_swap.lean | 2 +- test/norm_swap.lean | 2 +- test/simp_result.lean | 2 +- 29 files changed, 853 insertions(+), 820 deletions(-) create mode 100644 src/logic/equiv/defs.lean diff --git a/src/algebra/free.lean b/src/algebra/free.lean index e62e005b0c2c9..4e94ee31c8e30 100644 --- a/src/algebra/free.lean +++ b/src/algebra/free.lean @@ -6,7 +6,7 @@ Authors: Kenny Lau import algebra.hom.group import control.applicative import control.traversable.basic -import logic.equiv.basic +import logic.equiv.defs /-! # Free constructions diff --git a/src/algebra/group/type_tags.lean b/src/algebra/group/type_tags.lean index 68fcbe0feda94..d03b666e93413 100644 --- a/src/algebra/group/type_tags.lean +++ b/src/algebra/group/type_tags.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import algebra.hom.group -import logic.equiv.basic +import logic.equiv.defs import data.finite.defs /-! # Type tags that turn additive structures into multiplicative, and vice versa diff --git a/src/algebra/group/with_one.lean b/src/algebra/group/with_one.lean index 0b06f4b89f4bc..9efaeab5b0ef8 100644 --- a/src/algebra/group/with_one.lean +++ b/src/algebra/group/with_one.lean @@ -5,7 +5,7 @@ Authors: Mario Carneiro, Johan Commelin -/ import algebra.hom.equiv import algebra.ring.basic -import logic.equiv.basic +import logic.equiv.defs import logic.equiv.option /-! diff --git a/src/algebra/opposites.lean b/src/algebra/opposites.lean index 8ed49d0dff820..1dacfd60a5186 100644 --- a/src/algebra/opposites.lean +++ b/src/algebra/opposites.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import algebra.group.defs -import logic.equiv.basic +import logic.equiv.defs import logic.nontrivial /-! diff --git a/src/category_theory/functor/fully_faithful.lean b/src/category_theory/functor/fully_faithful.lean index 91635c9022566..6d24d20794f3e 100644 --- a/src/category_theory/functor/fully_faithful.lean +++ b/src/category_theory/functor/fully_faithful.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import category_theory.natural_isomorphism -import logic.equiv.basic +import logic.equiv.defs /-! # Full and faithful functors diff --git a/src/category_theory/types.lean b/src/category_theory/types.lean index 1f16fcbe2ad27..389c586fe537b 100644 --- a/src/category_theory/types.lean +++ b/src/category_theory/types.lean @@ -5,7 +5,7 @@ Authors: Stephen Morgan, Scott Morrison, Johannes Hölzl -/ import category_theory.epi_mono import category_theory.functor.fully_faithful -import logic.equiv.basic +import logic.equiv.defs /-! # The category `Type`. diff --git a/src/combinatorics/derangements/basic.lean b/src/combinatorics/derangements/basic.lean index 81abc44a5891b..6d0cce36db9ec 100644 --- a/src/combinatorics/derangements/basic.lean +++ b/src/combinatorics/derangements/basic.lean @@ -5,7 +5,7 @@ Authors: Henry Swanson -/ import dynamics.fixed_points.basic import group_theory.perm.option -import logic.equiv.basic +import logic.equiv.defs import logic.equiv.option /-! diff --git a/src/control/equiv_functor.lean b/src/control/equiv_functor.lean index 5305c7e0136e6..9f9459fecef12 100644 --- a/src/control/equiv_functor.lean +++ b/src/control/equiv_functor.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import logic.equiv.basic +import logic.equiv.defs /-! # Functions functorial with respect to equivalences diff --git a/src/control/monad/basic.lean b/src/control/monad/basic.lean index 28ba0a0117a64..6255d7b8ec089 100644 --- a/src/control/monad/basic.lean +++ b/src/control/monad/basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon -/ -import logic.equiv.basic +import logic.equiv.defs import tactic.basic /-! diff --git a/src/control/monad/writer.lean b/src/control/monad/writer.lean index 837d27543a48c..2b5fb0135e872 100644 --- a/src/control/monad/writer.lean +++ b/src/control/monad/writer.lean @@ -6,7 +6,7 @@ Authors: Simon Hudon The writer monad transformer for passing immutable state. -/ import algebra.group.defs -import logic.equiv.basic +import logic.equiv.defs universes u v w u₀ u₁ v₀ v₁ diff --git a/src/control/traversable/equiv.lean b/src/control/traversable/equiv.lean index 160d2e5cdcc22..ea9d1588094a3 100644 --- a/src/control/traversable/equiv.lean +++ b/src/control/traversable/equiv.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon -/ import control.traversable.lemmas -import logic.equiv.basic +import logic.equiv.defs /-! # Transferring `traversable` instances along isomorphisms diff --git a/src/data/erased.lean b/src/data/erased.lean index 7ccca5b054b77..b9ffd5bbde2eb 100644 --- a/src/data/erased.lean +++ b/src/data/erased.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import logic.equiv.basic +import logic.equiv.defs /-! # A type for VM-erased data diff --git a/src/data/opposite.lean b/src/data/opposite.lean index 985150bd4faf6..0d66e07b27685 100644 --- a/src/data/opposite.lean +++ b/src/data/opposite.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Reid Barton, Simon Hudon, Kenny Lau -/ -import logic.equiv.basic +import logic.equiv.defs /-! # Opposites diff --git a/src/data/part.lean b/src/data/part.lean index b956cca2b72b9..2e8a86de8be66 100644 --- a/src/data/part.lean +++ b/src/data/part.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Jeremy Avigad, Simon Hudon -/ import data.set.basic -import logic.equiv.basic +import logic.equiv.defs /-! # Partial values of a type diff --git a/src/logic/embedding.lean b/src/logic/embedding.lean index 8e7063d287391..d15adbf122f60 100644 --- a/src/logic/embedding.lean +++ b/src/logic/embedding.lean @@ -7,7 +7,7 @@ import data.fun_like.embedding import data.prod.pprod import data.set.basic import data.sigma.basic -import logic.equiv.basic +import logic.equiv.defs /-! # Injective functions diff --git a/src/logic/equiv/basic.lean b/src/logic/equiv/basic.lean index 207261d9339ea..87865c7c93b99 100644 --- a/src/logic/equiv/basic.lean +++ b/src/logic/equiv/basic.lean @@ -3,35 +3,20 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro -/ -import data.fun_like.equiv -import data.option.basic +import logic.equiv.defs import data.prod.basic -import data.quot import data.sigma.basic import data.subtype import data.sum.basic import logic.function.conjugate -import logic.unique -import tactic.norm_cast -import tactic.simps /-! # Equivalence between types -In this file we define two types: - -* `equiv α β` a.k.a. `α ≃ β`: a bijective map `α → β` bundled with its inverse map; we use this (and - not equality!) to express that various `Type`s or `Sort`s are equivalent. - -* `equiv.perm α`: the group of permutations `α ≃ α`. More lemmas about `equiv.perm` can be found in - `group_theory/perm`. - -Then we define +In this file we continue the work on equivalences begun in `logic/equiv/defs.lean`, defining * canonical isomorphisms between various types: e.g., - - `equiv.refl α` is the identity map interpreted as `α ≃ α`; - - `equiv.sum_equiv_sigma_bool` is the canonical equivalence between the sum of two types `α ⊕ β` and the sigma-type `Σ b : bool, cond b α β`; @@ -40,21 +25,9 @@ Then we define * operations on equivalences: e.g., - - `equiv.symm e : β ≃ α` is the inverse of `e : α ≃ β`; - - - `equiv.trans e₁ e₂ : α ≃ γ` is the composition of `e₁ : α ≃ β` and `e₂ : β ≃ γ` (note the order - of the arguments!); - - `equiv.prod_congr ea eb : α₁ × β₁ ≃ α₂ × β₂`: combine two equivalences `ea : α₁ ≃ α₂` and `eb : β₁ ≃ β₂` using `prod.map`. -* definitions that transfer some instances along an equivalence. By convention, we transfer - instances from right to left. - - - `equiv.inhabited` takes `e : α ≃ β` and `[inhabited β]` and returns `inhabited α`; - - `equiv.unique` takes `e : α ≃ β` and `[unique β]` and returns `unique α`; - - `equiv.decidable_eq` takes `e : α ≃ β` and `[decidable_eq β]` and returns `decidable_eq α`. - More definitions of this kind can be found in other files. E.g., `data/equiv/transfer_instance` does it for many algebraic type classes like `group`, `module`, etc. @@ -68,325 +41,8 @@ open function universes u v w z variables {α : Sort u} {β : Sort v} {γ : Sort w} -/-- `α ≃ β` is the type of functions from `α → β` with a two-sided inverse. -/ -structure equiv (α : Sort*) (β : Sort*) := -(to_fun : α → β) -(inv_fun : β → α) -(left_inv : left_inverse inv_fun to_fun) -(right_inv : right_inverse inv_fun to_fun) - -infix ` ≃ `:25 := equiv - -instance {F} [equiv_like F α β] : has_coe_t F (α ≃ β) := -⟨λ f, { to_fun := f, inv_fun := equiv_like.inv f, left_inv := equiv_like.left_inv f, - right_inv := equiv_like.right_inv f }⟩ - -/-- `perm α` is the type of bijections from `α` to itself. -/ -@[reducible] def equiv.perm (α : Sort*) := equiv α α - namespace equiv -instance : equiv_like (α ≃ β) α β := -{ coe := to_fun, inv := inv_fun, left_inv := left_inv, right_inv := right_inv, - coe_injective' := λ e₁ e₂ h₁ h₂, by { cases e₁, cases e₂, congr' } } - -instance : has_coe_to_fun (α ≃ β) (λ _, α → β) := ⟨to_fun⟩ - -@[simp] theorem coe_fn_mk (f : α → β) (g l r) : (equiv.mk f g l r : α → β) = f := -rfl - -/-- The map `coe_fn : (r ≃ s) → (r → s)` is injective. -/ -theorem coe_fn_injective : @function.injective (α ≃ β) (α → β) coe_fn := fun_like.coe_injective -protected lemma coe_inj {e₁ e₂ : α ≃ β} : (e₁ : α → β) = e₂ ↔ e₁ = e₂ := fun_like.coe_fn_eq -@[ext] lemma ext {f g : equiv α β} (H : ∀ x, f x = g x) : f = g := fun_like.ext f g H -protected lemma congr_arg {f : equiv α β} {x x' : α} : x = x' → f x = f x' := fun_like.congr_arg f -protected lemma congr_fun {f g : equiv α β} (h : f = g) (x : α) : f x = g x := -fun_like.congr_fun h x -lemma ext_iff {f g : equiv α β} : f = g ↔ ∀ x, f x = g x := fun_like.ext_iff - -@[ext] lemma perm.ext {σ τ : equiv.perm α} (H : ∀ x, σ x = τ x) : σ = τ := -equiv.ext H - -protected lemma perm.congr_arg {f : equiv.perm α} {x x' : α} : x = x' → f x = f x' := -equiv.congr_arg - -protected lemma perm.congr_fun {f g : equiv.perm α} (h : f = g) (x : α) : f x = g x := -equiv.congr_fun h x - -lemma perm.ext_iff {σ τ : equiv.perm α} : σ = τ ↔ ∀ x, σ x = τ x := -ext_iff - -/-- Any type is equivalent to itself. -/ -@[refl] protected def refl (α : Sort*) : α ≃ α := ⟨id, id, λ x, rfl, λ x, rfl⟩ - -instance inhabited' : inhabited (α ≃ α) := ⟨equiv.refl α⟩ - -/-- Inverse of an equivalence `e : α ≃ β`. -/ -@[symm] protected def symm (e : α ≃ β) : β ≃ α := ⟨e.inv_fun, e.to_fun, e.right_inv, e.left_inv⟩ - -/-- See Note [custom simps projection] -/ -def simps.symm_apply (e : α ≃ β) : β → α := e.symm - -initialize_simps_projections equiv (to_fun → apply, inv_fun → symm_apply) - -/-- Composition of equivalences `e₁ : α ≃ β` and `e₂ : β ≃ γ`. -/ -@[trans] protected def trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : α ≃ γ := -⟨e₂ ∘ e₁, e₁.symm ∘ e₂.symm, e₂.left_inv.comp e₁.left_inv, e₂.right_inv.comp e₁.right_inv⟩ - -@[simp] -lemma to_fun_as_coe (e : α ≃ β) : e.to_fun = e := rfl - -@[simp] -lemma inv_fun_as_coe (e : α ≃ β) : e.inv_fun = e.symm := rfl - -protected theorem injective (e : α ≃ β) : injective e := equiv_like.injective e -protected theorem surjective (e : α ≃ β) : surjective e := equiv_like.surjective e -protected theorem bijective (e : α ≃ β) : bijective e := equiv_like.bijective e - -protected theorem subsingleton (e : α ≃ β) [subsingleton β] : subsingleton α := -e.injective.subsingleton - -protected theorem subsingleton.symm (e : α ≃ β) [subsingleton α] : subsingleton β := -e.symm.injective.subsingleton - -lemma subsingleton_congr (e : α ≃ β) : subsingleton α ↔ subsingleton β := -⟨λ h, by exactI e.symm.subsingleton, λ h, by exactI e.subsingleton⟩ - -instance equiv_subsingleton_cod [subsingleton β] : - subsingleton (α ≃ β) := -⟨λ f g, equiv.ext $ λ x, subsingleton.elim _ _⟩ - -instance equiv_subsingleton_dom [subsingleton α] : - subsingleton (α ≃ β) := -⟨λ f g, equiv.ext $ λ x, @subsingleton.elim _ (equiv.subsingleton.symm f) _ _⟩ - -instance perm_unique [subsingleton α] : unique (perm α) := -unique_of_subsingleton (equiv.refl α) - -lemma perm.subsingleton_eq_refl [subsingleton α] (e : perm α) : - e = equiv.refl α := subsingleton.elim _ _ - -/-- Transfer `decidable_eq` across an equivalence. -/ -protected def decidable_eq (e : α ≃ β) [decidable_eq β] : decidable_eq α := -e.injective.decidable_eq - -lemma nonempty_congr (e : α ≃ β) : nonempty α ↔ nonempty β := -nonempty.congr e e.symm - -protected lemma nonempty (e : α ≃ β) [nonempty β] : nonempty α := -e.nonempty_congr.mpr ‹_› - -/-- If `α ≃ β` and `β` is inhabited, then so is `α`. -/ -protected def inhabited [inhabited β] (e : α ≃ β) : inhabited α := -⟨e.symm default⟩ - -/-- If `α ≃ β` and `β` is a singleton type, then so is `α`. -/ -protected def unique [unique β] (e : α ≃ β) : unique α := -e.symm.surjective.unique - -/-- Equivalence between equal types. -/ -protected def cast {α β : Sort*} (h : α = β) : α ≃ β := -⟨cast h, cast h.symm, λ x, by { cases h, refl }, λ x, by { cases h, refl }⟩ - -@[simp] theorem coe_fn_symm_mk (f : α → β) (g l r) : ((equiv.mk f g l r).symm : β → α) = g := -rfl - -@[simp] theorem coe_refl : ⇑(equiv.refl α) = id := rfl - -/-- This cannot be a `simp` lemmas as it incorrectly matches against `e : α ≃ synonym α`, when -`synonym α` is semireducible. This makes a mess of `multiplicative.of_add` etc. -/ -theorem perm.coe_subsingleton {α : Type*} [subsingleton α] (e : perm α) : ⇑(e) = id := -by rw [perm.subsingleton_eq_refl e, coe_refl] - -theorem refl_apply (x : α) : equiv.refl α x = x := rfl - -@[simp] theorem coe_trans (f : α ≃ β) (g : β ≃ γ) : ⇑(f.trans g) = g ∘ f := rfl - -theorem trans_apply (f : α ≃ β) (g : β ≃ γ) (a : α) : (f.trans g) a = g (f a) := rfl - -@[simp] theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := -e.right_inv x - -@[simp] theorem symm_apply_apply (e : α ≃ β) (x : α) : e.symm (e x) = x := -e.left_inv x - -@[simp] theorem symm_comp_self (e : α ≃ β) : e.symm ∘ e = id := funext e.symm_apply_apply - -@[simp] theorem self_comp_symm (e : α ≃ β) : e ∘ e.symm = id := funext e.apply_symm_apply - -@[simp] lemma symm_trans_apply (f : α ≃ β) (g : β ≃ γ) (a : γ) : - (f.trans g).symm a = f.symm (g.symm a) := rfl - --- The `simp` attribute is needed to make this a `dsimp` lemma. --- `simp` will always rewrite with `equiv.symm_symm` before this has a chance to fire. -@[simp, nolint simp_nf] theorem symm_symm_apply (f : α ≃ β) (b : α) : f.symm.symm b = f b := rfl - -theorem apply_eq_iff_eq (f : α ≃ β) {x y : α} : f x = f y ↔ x = y := equiv_like.apply_eq_iff_eq f - -theorem apply_eq_iff_eq_symm_apply {α β : Sort*} (f : α ≃ β) {x : α} {y : β} : - f x = y ↔ x = f.symm y := -begin - conv_lhs { rw ←apply_symm_apply f y, }, - rw apply_eq_iff_eq, -end - -@[simp] theorem cast_apply {α β} (h : α = β) (x : α) : equiv.cast h x = cast h x := rfl - -@[simp] theorem cast_symm {α β} (h : α = β) : (equiv.cast h).symm = equiv.cast h.symm := rfl - -@[simp] theorem cast_refl {α} (h : α = α := rfl) : equiv.cast h = equiv.refl α := rfl - -@[simp] theorem cast_trans {α β γ} (h : α = β) (h2 : β = γ) : - (equiv.cast h).trans (equiv.cast h2) = equiv.cast (h.trans h2) := -ext $ λ x, by { substs h h2, refl } - -lemma cast_eq_iff_heq {α β} (h : α = β) {a : α} {b : β} : equiv.cast h a = b ↔ a == b := -by { subst h, simp } - -lemma symm_apply_eq {α β} (e : α ≃ β) {x y} : e.symm x = y ↔ x = e y := -⟨λ H, by simp [H.symm], λ H, by simp [H]⟩ - -lemma eq_symm_apply {α β} (e : α ≃ β) {x y} : y = e.symm x ↔ e y = x := -(eq_comm.trans e.symm_apply_eq).trans eq_comm - -@[simp] theorem symm_symm (e : α ≃ β) : e.symm.symm = e := by { cases e, refl } - -@[simp] theorem trans_refl (e : α ≃ β) : e.trans (equiv.refl β) = e := by { cases e, refl } - -@[simp] theorem refl_symm : (equiv.refl α).symm = equiv.refl α := rfl - -@[simp] theorem refl_trans (e : α ≃ β) : (equiv.refl α).trans e = e := by { cases e, refl } - -@[simp] theorem symm_trans_self (e : α ≃ β) : e.symm.trans e = equiv.refl β := ext (by simp) - -@[simp] theorem self_trans_symm (e : α ≃ β) : e.trans e.symm = equiv.refl α := ext (by simp) - -lemma trans_assoc {δ} (ab : α ≃ β) (bc : β ≃ γ) (cd : γ ≃ δ) : - (ab.trans bc).trans cd = ab.trans (bc.trans cd) := -equiv.ext $ assume a, rfl - -theorem left_inverse_symm (f : equiv α β) : left_inverse f.symm f := f.left_inv - -theorem right_inverse_symm (f : equiv α β) : function.right_inverse f.symm f := f.right_inv - -lemma injective_comp (e : α ≃ β) (f : β → γ) : injective (f ∘ e) ↔ injective f := -equiv_like.injective_comp e f - -lemma comp_injective (f : α → β) (e : β ≃ γ) : injective (e ∘ f) ↔ injective f := -equiv_like.comp_injective f e - -lemma surjective_comp (e : α ≃ β) (f : β → γ) : surjective (f ∘ e) ↔ surjective f := -equiv_like.surjective_comp e f - -lemma comp_surjective (f : α → β) (e : β ≃ γ) : surjective (e ∘ f) ↔ surjective f := -equiv_like.comp_surjective f e - -lemma bijective_comp (e : α ≃ β) (f : β → γ) : bijective (f ∘ e) ↔ bijective f := -equiv_like.bijective_comp e f - -lemma comp_bijective (f : α → β) (e : β ≃ γ) : bijective (e ∘ f) ↔ bijective f := -equiv_like.comp_bijective f e - -/-- If `α` is equivalent to `β` and `γ` is equivalent to `δ`, then the type of equivalences `α ≃ γ` -is equivalent to the type of equivalences `β ≃ δ`. -/ -def equiv_congr {δ} (ab : α ≃ β) (cd : γ ≃ δ) : (α ≃ γ) ≃ (β ≃ δ) := -⟨ λac, (ab.symm.trans ac).trans cd, λbd, ab.trans $ bd.trans $ cd.symm, - assume ac, by { ext x, simp }, assume ac, by { ext x, simp } ⟩ - -@[simp] lemma equiv_congr_refl {α β} : - (equiv.refl α).equiv_congr (equiv.refl β) = equiv.refl (α ≃ β) := by { ext, refl } - -@[simp] lemma equiv_congr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : - (ab.equiv_congr cd).symm = ab.symm.equiv_congr cd.symm := by { ext, refl } - -@[simp] lemma equiv_congr_trans {δ ε ζ} (ab : α ≃ β) (de : δ ≃ ε) (bc : β ≃ γ) (ef : ε ≃ ζ) : - (ab.equiv_congr de).trans (bc.equiv_congr ef) = (ab.trans bc).equiv_congr (de.trans ef) := -by { ext, refl } - -@[simp] lemma equiv_congr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : - (equiv.refl α).equiv_congr bg e = e.trans bg := rfl - -@[simp] lemma equiv_congr_refl_right {α β} (ab e : α ≃ β) : - ab.equiv_congr (equiv.refl β) e = ab.symm.trans e := rfl - -@[simp] lemma equiv_congr_apply_apply {δ} (ab : α ≃ β) (cd : γ ≃ δ) (e : α ≃ γ) (x) : - ab.equiv_congr cd e x = cd (e (ab.symm x)) := rfl - -section perm_congr - -variables {α' β' : Type*} (e : α' ≃ β') - -/-- If `α` is equivalent to `β`, then `perm α` is equivalent to `perm β`. -/ -def perm_congr : perm α' ≃ perm β' := -equiv_congr e e - -lemma perm_congr_def (p : equiv.perm α') : - e.perm_congr p = (e.symm.trans p).trans e := rfl - -@[simp] lemma perm_congr_refl : - e.perm_congr (equiv.refl _) = equiv.refl _ := -by simp [perm_congr_def] - -@[simp] lemma perm_congr_symm : - e.perm_congr.symm = e.symm.perm_congr := rfl - -@[simp] lemma perm_congr_apply (p : equiv.perm α') (x) : - e.perm_congr p x = e (p (e.symm x)) := rfl - -lemma perm_congr_symm_apply (p : equiv.perm β') (x) : - e.perm_congr.symm p x = e.symm (p (e x)) := rfl - -lemma perm_congr_trans (p p' : equiv.perm α') : - (e.perm_congr p).trans (e.perm_congr p') = e.perm_congr (p.trans p') := -by { ext, simp } - -end perm_congr - -/-- Two empty types are equivalent. -/ -def equiv_of_is_empty (α β : Sort*) [is_empty α] [is_empty β] : α ≃ β := -⟨is_empty_elim, is_empty_elim, is_empty_elim, is_empty_elim⟩ - -/-- If `α` is an empty type, then it is equivalent to the `empty` type. -/ -def equiv_empty (α : Sort u) [is_empty α] : α ≃ empty := -equiv_of_is_empty α _ - -/-- If `α` is an empty type, then it is equivalent to the `pempty` type in any universe. -/ -def equiv_pempty (α : Sort v) [is_empty α] : α ≃ pempty.{u} := -equiv_of_is_empty α _ - -/-- `α` is equivalent to an empty type iff `α` is empty. -/ -def equiv_empty_equiv (α : Sort u) : (α ≃ empty) ≃ is_empty α := -⟨λ e, function.is_empty e, @equiv_empty α, λ e, ext $ λ x, (e x).elim, λ p, rfl⟩ - -/-- The `Sort` of proofs of a false proposition is equivalent to `pempty`. -/ -def prop_equiv_pempty {p : Prop} (h : ¬p) : p ≃ pempty := -@equiv_pempty p $ is_empty.prop_iff.2 h - -/-- If both `α` and `β` have a unique element, then `α ≃ β`. -/ -def equiv_of_unique (α β : Sort*) [unique α] [unique β] : α ≃ β := -{ to_fun := default, - inv_fun := default, - left_inv := λ _, subsingleton.elim _ _, - right_inv := λ _, subsingleton.elim _ _ } - -/-- If `α` has a unique element, then it is equivalent to any `punit`. -/ -def equiv_punit (α : Sort*) [unique α] : α ≃ punit.{v} := -equiv_of_unique α _ - -/-- The `Sort` of proofs of a true proposition is equivalent to `punit`. -/ -def prop_equiv_punit {p : Prop} (h : p) : p ≃ punit := -@equiv_punit p $ unique_prop h - -/-- `ulift α` is equivalent to `α`. -/ -@[simps apply symm_apply {fully_applied := ff}] -protected def ulift {α : Type v} : ulift.{u} α ≃ α := -⟨ulift.down, ulift.up, ulift.up_down, λ a, rfl⟩ - -/-- `plift α` is equivalent to `α`. -/ -@[simps apply symm_apply {fully_applied := ff}] -protected def plift : plift α ≃ α := -⟨plift.down, plift.up, plift.up_down, plift.down_up⟩ - /-- `pprod α β` is equivalent to `α × β` -/ @[simps apply symm_apply] def pprod_equiv_prod {α β : Type*} : pprod α β ≃ α × β := @@ -421,171 +77,6 @@ def prod_pprod {α₁ β₁ : Type*} {α₂ β₂ : Sort*} (ea : α₁ ≃ α₂ def pprod_equiv_prod_plift {α β : Sort*} : pprod α β ≃ plift α × plift β := equiv.plift.symm.pprod_prod equiv.plift.symm -/-- equivalence of propositions is the same as iff -/ -def of_iff {P Q : Prop} (h : P ↔ Q) : P ≃ Q := -{ to_fun := h.mp, - inv_fun := h.mpr, - left_inv := λ x, rfl, - right_inv := λ y, rfl } - -/-- If `α₁` is equivalent to `α₂` and `β₁` is equivalent to `β₂`, then the type of maps `α₁ → β₁` -is equivalent to the type of maps `α₂ → β₂`. -/ -@[congr, simps apply] def arrow_congr {α₁ β₁ α₂ β₂ : Sort*} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : - (α₁ → β₁) ≃ (α₂ → β₂) := -{ to_fun := λ f, e₂ ∘ f ∘ e₁.symm, - inv_fun := λ f, e₂.symm ∘ f ∘ e₁, - left_inv := λ f, funext $ λ x, by simp, - right_inv := λ f, funext $ λ x, by simp } - -lemma arrow_congr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort*} - (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) (f : α₁ → β₁) (g : β₁ → γ₁) : - arrow_congr ea ec (g ∘ f) = (arrow_congr eb ec g) ∘ (arrow_congr ea eb f) := -by { ext, simp only [comp, arrow_congr_apply, eb.symm_apply_apply] } - -@[simp] lemma arrow_congr_refl {α β : Sort*} : - arrow_congr (equiv.refl α) (equiv.refl β) = equiv.refl (α → β) := rfl - -@[simp] lemma arrow_congr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort*} - (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : - arrow_congr (e₁.trans e₂) (e₁'.trans e₂') = (arrow_congr e₁ e₁').trans (arrow_congr e₂ e₂') := -rfl - -@[simp] lemma arrow_congr_symm {α₁ β₁ α₂ β₂ : Sort*} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : - (arrow_congr e₁ e₂).symm = arrow_congr e₁.symm e₂.symm := -rfl - -/-- -A version of `equiv.arrow_congr` in `Type`, rather than `Sort`. - -The `equiv_rw` tactic is not able to use the default `Sort` level `equiv.arrow_congr`, -because Lean's universe rules will not unify `?l_1` with `imax (1 ?m_1)`. --/ -@[congr, simps apply] -def arrow_congr' {α₁ β₁ α₂ β₂ : Type*} (hα : α₁ ≃ α₂) (hβ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) := -equiv.arrow_congr hα hβ - -@[simp] lemma arrow_congr'_refl {α β : Type*} : - arrow_congr' (equiv.refl α) (equiv.refl β) = equiv.refl (α → β) := rfl - -@[simp] lemma arrow_congr'_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Type*} - (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : - arrow_congr' (e₁.trans e₂) (e₁'.trans e₂') = (arrow_congr' e₁ e₁').trans (arrow_congr' e₂ e₂') := -rfl - -@[simp] lemma arrow_congr'_symm {α₁ β₁ α₂ β₂ : Type*} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : - (arrow_congr' e₁ e₂).symm = arrow_congr' e₁.symm e₂.symm := -rfl - -/-- Conjugate a map `f : α → α` by an equivalence `α ≃ β`. -/ -@[simps apply] -def conj (e : α ≃ β) : (α → α) ≃ (β → β) := arrow_congr e e - -@[simp] lemma conj_refl : conj (equiv.refl α) = equiv.refl (α → α) := rfl - -@[simp] lemma conj_symm (e : α ≃ β) : e.conj.symm = e.symm.conj := rfl - -@[simp] lemma conj_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : - (e₁.trans e₂).conj = e₁.conj.trans e₂.conj := -rfl - --- This should not be a simp lemma as long as `(∘)` is reducible: --- when `(∘)` is reducible, Lean can unify `f₁ ∘ f₂` with any `g` using --- `f₁ := g` and `f₂ := λ x, x`. This causes nontermination. -lemma conj_comp (e : α ≃ β) (f₁ f₂ : α → α) : - e.conj (f₁ ∘ f₂) = (e.conj f₁) ∘ (e.conj f₂) := -by apply arrow_congr_comp - -lemma eq_comp_symm {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : - f = g ∘ e.symm ↔ f ∘ e = g := -(e.arrow_congr (equiv.refl γ)).symm_apply_eq.symm - -lemma comp_symm_eq {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : - g ∘ e.symm = f ↔ g = f ∘ e := -(e.arrow_congr (equiv.refl γ)).eq_symm_apply.symm - -lemma eq_symm_comp {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) : - f = e.symm ∘ g ↔ e ∘ f = g := -((equiv.refl γ).arrow_congr e).eq_symm_apply - -lemma symm_comp_eq {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) : - e.symm ∘ g = f ↔ g = e ∘ f := -((equiv.refl γ).arrow_congr e).symm_apply_eq - -section binary_op - -variables {α₁ β₁ : Type*} (e : α₁ ≃ β₁) (f : α₁ → α₁ → α₁) - -lemma semiconj_conj (f : α₁ → α₁) : semiconj e f (e.conj f) := λ x, by simp - -lemma semiconj₂_conj : semiconj₂ e f (e.arrow_congr e.conj f) := λ x y, by simp - -instance [is_associative α₁ f] : - is_associative β₁ (e.arrow_congr (e.arrow_congr e) f) := -(e.semiconj₂_conj f).is_associative_right e.surjective - -instance [is_idempotent α₁ f] : - is_idempotent β₁ (e.arrow_congr (e.arrow_congr e) f) := -(e.semiconj₂_conj f).is_idempotent_right e.surjective - -instance [is_left_cancel α₁ f] : - is_left_cancel β₁ (e.arrow_congr (e.arrow_congr e) f) := -⟨e.surjective.forall₃.2 $ λ x y z, by simpa using @is_left_cancel.left_cancel _ f _ x y z⟩ - -instance [is_right_cancel α₁ f] : - is_right_cancel β₁ (e.arrow_congr (e.arrow_congr e) f) := -⟨e.surjective.forall₃.2 $ λ x y z, by simpa using @is_right_cancel.right_cancel _ f _ x y z⟩ - -end binary_op - -/-- `punit` sorts in any two universes are equivalent. -/ -def punit_equiv_punit : punit.{v} ≃ punit.{w} := -⟨λ _, punit.star, λ _, punit.star, λ u, by { cases u, refl }, λ u, by { cases u, reflexivity }⟩ - -section -/-- The sort of maps to `punit.{v}` is equivalent to `punit.{w}`. -/ -def arrow_punit_equiv_punit (α : Sort*) : (α → punit.{v}) ≃ punit.{w} := -⟨λ f, punit.star, λ u f, punit.star, - λ f, by { funext x, cases f x, refl }, λ u, by { cases u, reflexivity }⟩ - -/-- If `α` is `subsingleton` and `a : α`, then the type of dependent functions `Π (i : α), β -i` is equivalent to `β i`. -/ -@[simps] -def Pi_subsingleton {α} (β : α → Sort*) [subsingleton α] (a : α) : (Π a', β a') ≃ β a := -{ to_fun := eval a, - inv_fun := λ x b, cast (congr_arg β $ subsingleton.elim a b) x, - left_inv := λ f, funext $ λ b, by { rw subsingleton.elim b a, reflexivity }, - right_inv := λ b, rfl } - -/-- If `α` has a unique term, then the type of function `α → β` is equivalent to `β`. -/ -@[simps { fully_applied := ff }] def fun_unique (α β) [unique α] : (α → β) ≃ β := -Pi_subsingleton _ default - -/-- The sort of maps from `punit` is equivalent to the codomain. -/ -def punit_arrow_equiv (α : Sort*) : (punit.{u} → α) ≃ α := -fun_unique _ _ - -/-- The sort of maps from `true` is equivalent to the codomain. -/ -def true_arrow_equiv (α : Sort*) : (true → α) ≃ α := -fun_unique _ _ - -/-- The sort of maps from a type that `is_empty` is equivalent to `punit`. -/ -def arrow_punit_of_is_empty (α β : Sort*) [is_empty α] : (α → β) ≃ punit.{u} := -⟨λ f, punit.star, λ u, is_empty_elim, λ f, funext is_empty_elim, λ u, by { cases u, refl }⟩ - -/-- The sort of maps from `empty` is equivalent to `punit`. -/ -def empty_arrow_equiv_punit (α : Sort*) : (empty → α) ≃ punit.{u} := -arrow_punit_of_is_empty _ _ - -/-- The sort of maps from `pempty` is equivalent to `punit`. -/ -def pempty_arrow_equiv_punit (α : Sort*) : (pempty → α) ≃ punit.{u} := -arrow_punit_of_is_empty _ _ - -/-- The sort of maps from `false` is equivalent to `punit`. -/ -def false_arrow_equiv_punit (α : Sort*) : (false → α) ≃ punit.{u} := -arrow_punit_of_is_empty _ _ - -end - /-- Product of two equivalences. If `α₁ ≃ α₂` and `β₁ ≃ β₂`, then `α₁ × β₁ ≃ α₂ × β₂`. This is `prod.map` as an equivalence. -/ @[congr, simps apply] @@ -743,11 +234,6 @@ def bool_equiv_punit_sum_punit : bool ≃ punit.{u+1} ⊕ punit.{v+1} := λ b, by cases b; refl, λ s, by rcases s with ⟨⟨⟩⟩ | ⟨⟨⟩⟩; refl⟩ -/-- `Prop` is noncomputably equivalent to `bool`. -/ -noncomputable def Prop_equiv_bool : Prop ≃ bool := -⟨λ p, @to_bool p (classical.prop_decidable _), - λ b, b, λ p, by simp, λ b, by simp⟩ - /-- Sum of types is commutative up to an equivalence. This is `sum.swap` as an equivalence. -/ @[simps apply {fully_applied := ff}] def sum_comm (α β : Type*) : α ⊕ β ≃ β ⊕ α := @@ -1011,146 +497,6 @@ def Pi_curry {α} {β : α → Sort*} (γ : Π a, β a → Sort*) : end -section - -/-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ -@[simps apply symm_apply] def psigma_equiv_sigma {α} (β : α → Type*) : (Σ' i, β i) ≃ Σ i, β i := -⟨λ a, ⟨a.1, a.2⟩, λ a, ⟨a.1, a.2⟩, λ ⟨a, b⟩, rfl, λ ⟨a, b⟩, rfl⟩ - -/-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ -@[simps apply symm_apply] def psigma_equiv_sigma_plift {α} (β : α → Sort*) : - (Σ' i, β i) ≃ Σ i : plift α, plift (β i.down) := -⟨λ a, ⟨plift.up a.1, plift.up a.2⟩, λ a, ⟨a.1.down, a.2.down⟩, λ ⟨a, b⟩, rfl, λ ⟨⟨a⟩, ⟨b⟩⟩, rfl⟩ - -/-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ' a, β₁ a` and -`Σ' a, β₂ a`. -/ -@[simps apply] -def psigma_congr_right {α} {β₁ β₂ : α → Sort*} (F : Π a, β₁ a ≃ β₂ a) : (Σ' a, β₁ a) ≃ Σ' a, β₂ a := -⟨λ a, ⟨a.1, F a.1 a.2⟩, λ a, ⟨a.1, (F a.1).symm a.2⟩, - λ ⟨a, b⟩, congr_arg (psigma.mk a) $ symm_apply_apply (F a) b, - λ ⟨a, b⟩, congr_arg (psigma.mk a) $ apply_symm_apply (F a) b⟩ - -@[simp] lemma psigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Sort*} - (F : Π a, β₁ a ≃ β₂ a) (G : Π a, β₂ a ≃ β₃ a) : - (psigma_congr_right F).trans (psigma_congr_right G) = - psigma_congr_right (λ a, (F a).trans (G a)) := -by { ext1 x, cases x, refl } - -@[simp] lemma psigma_congr_right_symm {α} {β₁ β₂ : α → Sort*} (F : Π a, β₁ a ≃ β₂ a) : - (psigma_congr_right F).symm = psigma_congr_right (λ a, (F a).symm) := -by { ext1 x, cases x, refl } - -@[simp] lemma psigma_congr_right_refl {α} {β : α → Sort*} : - (psigma_congr_right (λ a, equiv.refl (β a))) = equiv.refl (Σ' a, β a) := -by { ext1 x, cases x, refl } - -/-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ a, β₁ a` and -`Σ a, β₂ a`. -/ -@[simps apply] -def sigma_congr_right {α} {β₁ β₂ : α → Type*} (F : Π a, β₁ a ≃ β₂ a) : (Σ a, β₁ a) ≃ Σ a, β₂ a := -⟨λ a, ⟨a.1, F a.1 a.2⟩, λ a, ⟨a.1, (F a.1).symm a.2⟩, - λ ⟨a, b⟩, congr_arg (sigma.mk a) $ symm_apply_apply (F a) b, - λ ⟨a, b⟩, congr_arg (sigma.mk a) $ apply_symm_apply (F a) b⟩ - -@[simp] lemma sigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Type*} - (F : Π a, β₁ a ≃ β₂ a) (G : Π a, β₂ a ≃ β₃ a) : - (sigma_congr_right F).trans (sigma_congr_right G) = sigma_congr_right (λ a, (F a).trans (G a)) := -by { ext1 x, cases x, refl } - -@[simp] lemma sigma_congr_right_symm {α} {β₁ β₂ : α → Type*} (F : Π a, β₁ a ≃ β₂ a) : - (sigma_congr_right F).symm = sigma_congr_right (λ a, (F a).symm) := -by { ext1 x, cases x, refl } - -@[simp] lemma sigma_congr_right_refl {α} {β : α → Type*} : - (sigma_congr_right (λ a, equiv.refl (β a))) = equiv.refl (Σ a, β a) := -by { ext1 x, cases x, refl } - -/-- A `psigma` with `Prop` fibers is equivalent to the subtype. -/ -def psigma_equiv_subtype {α : Type v} (P : α → Prop) : - (Σ' i, P i) ≃ subtype P := -{ to_fun := λ x, ⟨x.1, x.2⟩, - inv_fun := λ x, ⟨x.1, x.2⟩, - left_inv := λ x, by { cases x, refl, }, - right_inv := λ x, by { cases x, refl, }, } - -/-- A `sigma` with `plift` fibers is equivalent to the subtype. -/ -def sigma_plift_equiv_subtype {α : Type v} (P : α → Prop) : - (Σ i, plift (P i)) ≃ subtype P := -((psigma_equiv_sigma _).symm.trans (psigma_congr_right (λ a, equiv.plift))).trans - (psigma_equiv_subtype P) - -/-- -A `sigma` with `λ i, ulift (plift (P i))` fibers is equivalent to `{ x // P x }`. -Variant of `sigma_plift_equiv_subtype`. --/ -def sigma_ulift_plift_equiv_subtype {α : Type v} (P : α → Prop) : - (Σ i, ulift (plift (P i))) ≃ subtype P := -(sigma_congr_right (λ a, equiv.ulift)).trans (sigma_plift_equiv_subtype P) - -namespace perm - -/-- A family of permutations `Π a, perm (β a)` generates a permuation `perm (Σ a, β₁ a)`. -/ -@[reducible] -def sigma_congr_right {α} {β : α → Sort*} (F : Π a, perm (β a)) : perm (Σ a, β a) := -equiv.sigma_congr_right F - -@[simp] lemma sigma_congr_right_trans {α} {β : α → Sort*} - (F : Π a, perm (β a)) (G : Π a, perm (β a)) : - (sigma_congr_right F).trans (sigma_congr_right G) = sigma_congr_right (λ a, (F a).trans (G a)) := -equiv.sigma_congr_right_trans F G - -@[simp] lemma sigma_congr_right_symm {α} {β : α → Sort*} (F : Π a, perm (β a)) : - (sigma_congr_right F).symm = sigma_congr_right (λ a, (F a).symm) := -equiv.sigma_congr_right_symm F - -@[simp] lemma sigma_congr_right_refl {α} {β : α → Sort*} : - (sigma_congr_right (λ a, equiv.refl (β a))) = equiv.refl (Σ a, β a) := -equiv.sigma_congr_right_refl - -end perm - -/-- An equivalence `f : α₁ ≃ α₂` generates an equivalence between `Σ a, β (f a)` and `Σ a, β a`. -/ -@[simps apply] -def sigma_congr_left {α₁ α₂} {β : α₂ → Sort*} (e : α₁ ≃ α₂) : (Σ a:α₁, β (e a)) ≃ (Σ a:α₂, β a) := -⟨λ a, ⟨e a.1, a.2⟩, λ a, ⟨e.symm a.1, @@eq.rec β a.2 (e.right_inv a.1).symm⟩, - λ ⟨a, b⟩, match e.symm (e a), e.left_inv a : ∀ a' (h : a' = a), - @sigma.mk _ (β ∘ e) _ (@@eq.rec β b (congr_arg e h.symm)) = ⟨a, b⟩ with - | _, rfl := rfl end, - λ ⟨a, b⟩, match e (e.symm a), _ : ∀ a' (h : a' = a), - sigma.mk a' (@@eq.rec β b h.symm) = ⟨a, b⟩ with - | _, rfl := rfl end⟩ - -/-- Transporting a sigma type through an equivalence of the base -/ -def sigma_congr_left' {α₁ α₂} {β : α₁ → Sort*} (f : α₁ ≃ α₂) : - (Σ a:α₁, β a) ≃ (Σ a:α₂, β (f.symm a)) := -(sigma_congr_left f.symm).symm - -/-- Transporting a sigma type through an equivalence of the base and a family of equivalences -of matching fibers -/ -def sigma_congr {α₁ α₂} {β₁ : α₁ → Sort*} {β₂ : α₂ → Sort*} (f : α₁ ≃ α₂) - (F : ∀ a, β₁ a ≃ β₂ (f a)) : - sigma β₁ ≃ sigma β₂ := -(sigma_congr_right F).trans (sigma_congr_left f) - -/-- `sigma` type with a constant fiber is equivalent to the product. -/ -@[simps apply symm_apply] def sigma_equiv_prod (α β : Type*) : (Σ_:α, β) ≃ α × β := -⟨λ a, ⟨a.1, a.2⟩, λ a, ⟨a.1, a.2⟩, λ ⟨a, b⟩, rfl, λ ⟨a, b⟩, rfl⟩ - -/-- If each fiber of a `sigma` type is equivalent to a fixed type, then the sigma type -is equivalent to the product. -/ -def sigma_equiv_prod_of_equiv {α β} {β₁ : α → Sort*} (F : Π a, β₁ a ≃ β) : sigma β₁ ≃ α × β := -(sigma_congr_right F).trans (sigma_equiv_prod α β) - -/-- Dependent product of types is associative up to an equivalence. -/ -def sigma_assoc {α : Type*} {β : α → Type*} (γ : Π (a : α), β a → Type*) : - (Σ (ab : Σ (a : α), β a), γ ab.1 ab.2) ≃ Σ (a : α), (Σ (b : β a), γ a b) := -{ to_fun := λ x, ⟨x.1.1, ⟨x.1.2, x.2⟩⟩, - inv_fun := λ x, ⟨⟨x.1, x.2.1⟩, x.2.2⟩, - left_inv := λ ⟨⟨a, b⟩, c⟩, rfl, - right_inv := λ ⟨a, ⟨b, c⟩⟩, rfl } - -end - section prod_congr variables {α₁ β₁ β₂ : Type*} (e : α₁ → β₁ ≃ β₂) @@ -1909,89 +1255,6 @@ end namespace equiv -protected lemma exists_unique_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) - (h : ∀{x}, p x ↔ q (f x)) : (∃! x, p x) ↔ ∃! y, q y := -begin - split, - { rintro ⟨a, ha₁, ha₂⟩, - exact ⟨f a, h.1 ha₁, λ b hb, f.symm_apply_eq.1 (ha₂ (f.symm b) (h.2 (by simpa using hb)))⟩ }, - { rintro ⟨b, hb₁, hb₂⟩, - exact ⟨f.symm b, h.2 (by simpa using hb₁), λ y hy, (eq_symm_apply f).2 (hb₂ _ (h.1 hy))⟩ } -end - -protected lemma exists_unique_congr_left' {p : α → Prop} (f : α ≃ β) : - (∃! x, p x) ↔ (∃! y, p (f.symm y)) := -equiv.exists_unique_congr f (λx, by simp) - -protected lemma exists_unique_congr_left {p : β → Prop} (f : α ≃ β) : - (∃! x, p (f x)) ↔ (∃! y, p y) := -(equiv.exists_unique_congr_left' f.symm).symm - -protected lemma forall_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) - (h : ∀{x}, p x ↔ q (f x)) : (∀x, p x) ↔ (∀y, q y) := -begin - split; intros h₂ x, - { rw [←f.right_inv x], apply h.mp, apply h₂ }, - apply h.mpr, apply h₂ -end -protected lemma forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ β) - (h : ∀{x}, p (f.symm x) ↔ q x) : (∀x, p x) ↔ (∀y, q y) := -(equiv.forall_congr f.symm (λ x, h.symm)).symm - --- We next build some higher arity versions of `equiv.forall_congr`. --- Although they appear to just be repeated applications of `equiv.forall_congr`, --- unification of metavariables works better with these versions. --- In particular, they are necessary in `equiv_rw`. --- (Stopping at ternary functions seems reasonable: at least in 1-categorical mathematics, --- it's rare to have axioms involving more than 3 elements at once.) -universes ua1 ua2 ub1 ub2 ug1 ug2 -variables {α₁ : Sort ua1} {α₂ : Sort ua2} - {β₁ : Sort ub1} {β₂ : Sort ub2} - {γ₁ : Sort ug1} {γ₂ : Sort ug2} - -protected lemma forall₂_congr {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) - (eβ : β₁ ≃ β₂) (h : ∀{x y}, p x y ↔ q (eα x) (eβ y)) : - (∀x y, p x y) ↔ (∀x y, q x y) := -begin - apply equiv.forall_congr, - intros, - apply equiv.forall_congr, - intros, - apply h, -end -protected lemma forall₂_congr' {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) - (eβ : β₁ ≃ β₂) (h : ∀{x y}, p (eα.symm x) (eβ.symm y) ↔ q x y) : - (∀x y, p x y) ↔ (∀x y, q x y) := -(equiv.forall₂_congr eα.symm eβ.symm (λ x y, h.symm)).symm - -protected lemma forall₃_congr {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} - (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) - (h : ∀{x y z}, p x y z ↔ q (eα x) (eβ y) (eγ z)) : (∀x y z, p x y z) ↔ (∀x y z, q x y z) := -begin - apply equiv.forall₂_congr, - intros, - apply equiv.forall_congr, - intros, - apply h, -end -protected lemma forall₃_congr' {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} - (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) - (h : ∀{x y z}, p (eα.symm x) (eβ.symm y) (eγ.symm z) ↔ q x y z) : - (∀x y z, p x y z) ↔ (∀x y z, q x y z) := -(equiv.forall₃_congr eα.symm eβ.symm eγ.symm (λ x y z, h.symm)).symm - -protected lemma forall_congr_left' {p : α → Prop} (f : α ≃ β) : - (∀x, p x) ↔ (∀y, p (f.symm y)) := -equiv.forall_congr f (λx, by simp) - -protected lemma forall_congr_left {p : β → Prop} (f : α ≃ β) : - (∀x, p (f x)) ↔ (∀y, p y) := -(equiv.forall_congr_left' f.symm).symm - -protected lemma exists_congr_left {α β} (f : α ≃ β) {p : α → Prop} : - (∃ a, p a) ↔ (∃ b, p (f.symm b)) := -⟨λ ⟨a, h⟩, ⟨f a, by simpa using h⟩, λ ⟨b, h⟩, ⟨_, h⟩⟩ - section variables (P : α → Sort w) (e : α ≃ β) @@ -2085,6 +1348,32 @@ end end +section binary_op + +variables {α₁ β₁ : Type*} (e : α₁ ≃ β₁) (f : α₁ → α₁ → α₁) + +lemma semiconj_conj (f : α₁ → α₁) : semiconj e f (e.conj f) := λ x, by simp + +lemma semiconj₂_conj : semiconj₂ e f (e.arrow_congr e.conj f) := λ x y, by simp + +instance [is_associative α₁ f] : + is_associative β₁ (e.arrow_congr (e.arrow_congr e) f) := +(e.semiconj₂_conj f).is_associative_right e.surjective + +instance [is_idempotent α₁ f] : + is_idempotent β₁ (e.arrow_congr (e.arrow_congr e) f) := +(e.semiconj₂_conj f).is_idempotent_right e.surjective + +instance [is_left_cancel α₁ f] : + is_left_cancel β₁ (e.arrow_congr (e.arrow_congr e) f) := +⟨e.surjective.forall₃.2 $ λ x y z, by simpa using @is_left_cancel.left_cancel _ f _ x y z⟩ + +instance [is_right_cancel α₁ f] : + is_right_cancel β₁ (e.arrow_congr (e.arrow_congr e) f) := +⟨e.surjective.forall₃.2 $ λ x y z, by simpa using @is_right_cancel.right_cancel _ f _ x y z⟩ + +end binary_op + end equiv lemma function.injective.swap_apply [decidable_eq α] [decidable_eq β] {f : α → β} @@ -2129,61 +1418,6 @@ def unique_unique_equiv : unique (unique α) ≃ unique α := equiv_of_subsingleton_of_subsingleton (λ h, h.default) (λ h, { default := h, uniq := λ _, subsingleton.elim _ _ }) -namespace quot - -/-- An equivalence `e : α ≃ β` generates an equivalence between quotient spaces, -if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ -protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) - (eq : ∀a₁ a₂, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) : - quot ra ≃ quot rb := -{ to_fun := quot.map e (assume a₁ a₂, (eq a₁ a₂).1), - inv_fun := quot.map e.symm - (assume b₁ b₂ h, - (eq (e.symm b₁) (e.symm b₂)).2 - ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h)), - left_inv := by { rintros ⟨a⟩, dunfold quot.map, simp only [equiv.symm_apply_apply] }, - right_inv := by { rintros ⟨a⟩, dunfold quot.map, simp only [equiv.apply_symm_apply] } } - -@[simp] -lemma congr_mk {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) - (eq : ∀ (a₁ a₂ : α), ra a₁ a₂ ↔ rb (e a₁) (e a₂)) (a : α) : - quot.congr e eq (quot.mk ra a) = quot.mk rb (e a) := rfl - -/-- Quotients are congruent on equivalences under equality of their relation. -An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ -protected def congr_right {r r' : α → α → Prop} (eq : ∀a₁ a₂, r a₁ a₂ ↔ r' a₁ a₂) : - quot r ≃ quot r' := -quot.congr (equiv.refl α) eq - -/-- An equivalence `e : α ≃ β` generates an equivalence between the quotient space of `α` -by a relation `ra` and the quotient space of `β` by the image of this relation under `e`. -/ -protected def congr_left {r : α → α → Prop} (e : α ≃ β) : - quot r ≃ quot (λ b b', r (e.symm b) (e.symm b')) := -@quot.congr α β r (λ b b', r (e.symm b) (e.symm b')) e (λ a₁ a₂, by simp only [e.symm_apply_apply]) - -end quot - -namespace quotient -/-- An equivalence `e : α ≃ β` generates an equivalence between quotient spaces, -if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ -protected def congr {ra : setoid α} {rb : setoid β} (e : α ≃ β) - (eq : ∀a₁ a₂, @setoid.r α ra a₁ a₂ ↔ @setoid.r β rb (e a₁) (e a₂)) : - quotient ra ≃ quotient rb := -quot.congr e eq - -@[simp] -lemma congr_mk {ra : setoid α} {rb : setoid β} (e : α ≃ β) - (eq : ∀ (a₁ a₂ : α), setoid.r a₁ a₂ ↔ setoid.r (e a₁) (e a₂)) (a : α): - quotient.congr e eq (quotient.mk a) = quotient.mk (e a) := -rfl - -/-- Quotients are congruent on equivalences under equality of their relation. -An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ -protected def congr_right {r r' : setoid α} - (eq : ∀a₁ a₂, @setoid.r α r a₁ a₂ ↔ @setoid.r α r' a₁ a₂) : quotient r ≃ quotient r' := -quot.congr_right eq -end quotient - namespace function lemma update_comp_equiv {α β α' : Sort*} [decidable_eq α'] [decidable_eq α] (f : α → β) (g : α' ≃ α) diff --git a/src/logic/equiv/defs.lean b/src/logic/equiv/defs.lean new file mode 100644 index 0000000000000..bef85d5dad93a --- /dev/null +++ b/src/logic/equiv/defs.lean @@ -0,0 +1,798 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Mario Carneiro +-/ +import data.fun_like.equiv +import logic.unique + +/-! +# Equivalence between types + +In this file we define two types: + +* `equiv α β` a.k.a. `α ≃ β`: a bijective map `α → β` bundled with its inverse map; we use this (and + not equality!) to express that various `Type`s or `Sort`s are equivalent. + +* `equiv.perm α`: the group of permutations `α ≃ α`. More lemmas about `equiv.perm` can be found in + `group_theory/perm`. + +Then we define + +* canonical isomorphisms between various types: e.g., + + - `equiv.refl α` is the identity map interpreted as `α ≃ α`; + +* operations on equivalences: e.g., + + - `equiv.symm e : β ≃ α` is the inverse of `e : α ≃ β`; + + - `equiv.trans e₁ e₂ : α ≃ γ` is the composition of `e₁ : α ≃ β` and `e₂ : β ≃ γ` (note the order + of the arguments!); + +* definitions that transfer some instances along an equivalence. By convention, we transfer + instances from right to left. + + - `equiv.inhabited` takes `e : α ≃ β` and `[inhabited β]` and returns `inhabited α`; + - `equiv.unique` takes `e : α ≃ β` and `[unique β]` and returns `unique α`; + - `equiv.decidable_eq` takes `e : α ≃ β` and `[decidable_eq β]` and returns `decidable_eq α`. + + More definitions of this kind can be found in other files. E.g., `data/equiv/transfer_instance` + does it for many algebraic type classes like `group`, `module`, etc. + +Many more such isomorphisms and operations are defined in `logic/equiv/basic`. + +## Tags + +equivalence, congruence, bijective map +-/ + +open function + +universes u v w z +variables {α : Sort u} {β : Sort v} {γ : Sort w} + +/-- `α ≃ β` is the type of functions from `α → β` with a two-sided inverse. -/ +structure equiv (α : Sort*) (β : Sort*) := +(to_fun : α → β) +(inv_fun : β → α) +(left_inv : left_inverse inv_fun to_fun) +(right_inv : right_inverse inv_fun to_fun) + +infix ` ≃ `:25 := equiv + +instance {F} [equiv_like F α β] : has_coe_t F (α ≃ β) := +⟨λ f, { to_fun := f, inv_fun := equiv_like.inv f, left_inv := equiv_like.left_inv f, + right_inv := equiv_like.right_inv f }⟩ + +/-- `perm α` is the type of bijections from `α` to itself. -/ +@[reducible] def equiv.perm (α : Sort*) := equiv α α + +namespace equiv + +instance : equiv_like (α ≃ β) α β := +{ coe := to_fun, inv := inv_fun, left_inv := left_inv, right_inv := right_inv, + coe_injective' := λ e₁ e₂ h₁ h₂, by { cases e₁, cases e₂, congr' } } + +instance : has_coe_to_fun (α ≃ β) (λ _, α → β) := ⟨to_fun⟩ + +@[simp] theorem coe_fn_mk (f : α → β) (g l r) : (equiv.mk f g l r : α → β) = f := +rfl + +/-- The map `coe_fn : (r ≃ s) → (r → s)` is injective. -/ +theorem coe_fn_injective : @function.injective (α ≃ β) (α → β) coe_fn := fun_like.coe_injective +protected lemma coe_inj {e₁ e₂ : α ≃ β} : (e₁ : α → β) = e₂ ↔ e₁ = e₂ := fun_like.coe_fn_eq +@[ext] lemma ext {f g : equiv α β} (H : ∀ x, f x = g x) : f = g := fun_like.ext f g H +protected lemma congr_arg {f : equiv α β} {x x' : α} : x = x' → f x = f x' := fun_like.congr_arg f +protected lemma congr_fun {f g : equiv α β} (h : f = g) (x : α) : f x = g x := +fun_like.congr_fun h x +lemma ext_iff {f g : equiv α β} : f = g ↔ ∀ x, f x = g x := fun_like.ext_iff + +@[ext] lemma perm.ext {σ τ : equiv.perm α} (H : ∀ x, σ x = τ x) : σ = τ := +equiv.ext H + +protected lemma perm.congr_arg {f : equiv.perm α} {x x' : α} : x = x' → f x = f x' := +equiv.congr_arg + +protected lemma perm.congr_fun {f g : equiv.perm α} (h : f = g) (x : α) : f x = g x := +equiv.congr_fun h x + +lemma perm.ext_iff {σ τ : equiv.perm α} : σ = τ ↔ ∀ x, σ x = τ x := +ext_iff + +/-- Any type is equivalent to itself. -/ +@[refl] protected def refl (α : Sort*) : α ≃ α := ⟨id, id, λ x, rfl, λ x, rfl⟩ + +instance inhabited' : inhabited (α ≃ α) := ⟨equiv.refl α⟩ + +/-- Inverse of an equivalence `e : α ≃ β`. -/ +@[symm] protected def symm (e : α ≃ β) : β ≃ α := ⟨e.inv_fun, e.to_fun, e.right_inv, e.left_inv⟩ + +/-- See Note [custom simps projection] -/ +def simps.symm_apply (e : α ≃ β) : β → α := e.symm + +initialize_simps_projections equiv (to_fun → apply, inv_fun → symm_apply) + +/-- Composition of equivalences `e₁ : α ≃ β` and `e₂ : β ≃ γ`. -/ +@[trans] protected def trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : α ≃ γ := +⟨e₂ ∘ e₁, e₁.symm ∘ e₂.symm, e₂.left_inv.comp e₁.left_inv, e₂.right_inv.comp e₁.right_inv⟩ + +@[simp] +lemma to_fun_as_coe (e : α ≃ β) : e.to_fun = e := rfl + +@[simp] +lemma inv_fun_as_coe (e : α ≃ β) : e.inv_fun = e.symm := rfl + +protected theorem injective (e : α ≃ β) : injective e := equiv_like.injective e +protected theorem surjective (e : α ≃ β) : surjective e := equiv_like.surjective e +protected theorem bijective (e : α ≃ β) : bijective e := equiv_like.bijective e + +protected theorem subsingleton (e : α ≃ β) [subsingleton β] : subsingleton α := +e.injective.subsingleton + +protected theorem subsingleton.symm (e : α ≃ β) [subsingleton α] : subsingleton β := +e.symm.injective.subsingleton + +lemma subsingleton_congr (e : α ≃ β) : subsingleton α ↔ subsingleton β := +⟨λ h, by exactI e.symm.subsingleton, λ h, by exactI e.subsingleton⟩ + +instance equiv_subsingleton_cod [subsingleton β] : + subsingleton (α ≃ β) := +⟨λ f g, equiv.ext $ λ x, subsingleton.elim _ _⟩ + +instance equiv_subsingleton_dom [subsingleton α] : + subsingleton (α ≃ β) := +⟨λ f g, equiv.ext $ λ x, @subsingleton.elim _ (equiv.subsingleton.symm f) _ _⟩ + +instance perm_unique [subsingleton α] : unique (perm α) := +unique_of_subsingleton (equiv.refl α) + +lemma perm.subsingleton_eq_refl [subsingleton α] (e : perm α) : + e = equiv.refl α := subsingleton.elim _ _ + +/-- Transfer `decidable_eq` across an equivalence. -/ +protected def decidable_eq (e : α ≃ β) [decidable_eq β] : decidable_eq α := +e.injective.decidable_eq + +lemma nonempty_congr (e : α ≃ β) : nonempty α ↔ nonempty β := +nonempty.congr e e.symm + +protected lemma nonempty (e : α ≃ β) [nonempty β] : nonempty α := +e.nonempty_congr.mpr ‹_› + +/-- If `α ≃ β` and `β` is inhabited, then so is `α`. -/ +protected def inhabited [inhabited β] (e : α ≃ β) : inhabited α := +⟨e.symm default⟩ + +/-- If `α ≃ β` and `β` is a singleton type, then so is `α`. -/ +protected def unique [unique β] (e : α ≃ β) : unique α := +e.symm.surjective.unique + +/-- Equivalence between equal types. -/ +protected def cast {α β : Sort*} (h : α = β) : α ≃ β := +⟨cast h, cast h.symm, λ x, by { cases h, refl }, λ x, by { cases h, refl }⟩ + +@[simp] theorem coe_fn_symm_mk (f : α → β) (g l r) : ((equiv.mk f g l r).symm : β → α) = g := +rfl + +@[simp] theorem coe_refl : ⇑(equiv.refl α) = id := rfl + +/-- This cannot be a `simp` lemmas as it incorrectly matches against `e : α ≃ synonym α`, when +`synonym α` is semireducible. This makes a mess of `multiplicative.of_add` etc. -/ +theorem perm.coe_subsingleton {α : Type*} [subsingleton α] (e : perm α) : ⇑(e) = id := +by rw [perm.subsingleton_eq_refl e, coe_refl] + +theorem refl_apply (x : α) : equiv.refl α x = x := rfl + +@[simp] theorem coe_trans (f : α ≃ β) (g : β ≃ γ) : ⇑(f.trans g) = g ∘ f := rfl + +theorem trans_apply (f : α ≃ β) (g : β ≃ γ) (a : α) : (f.trans g) a = g (f a) := rfl + +@[simp] theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := +e.right_inv x + +@[simp] theorem symm_apply_apply (e : α ≃ β) (x : α) : e.symm (e x) = x := +e.left_inv x + +@[simp] theorem symm_comp_self (e : α ≃ β) : e.symm ∘ e = id := funext e.symm_apply_apply + +@[simp] theorem self_comp_symm (e : α ≃ β) : e ∘ e.symm = id := funext e.apply_symm_apply + +@[simp] lemma symm_trans_apply (f : α ≃ β) (g : β ≃ γ) (a : γ) : + (f.trans g).symm a = f.symm (g.symm a) := rfl + +-- The `simp` attribute is needed to make this a `dsimp` lemma. +-- `simp` will always rewrite with `equiv.symm_symm` before this has a chance to fire. +@[simp, nolint simp_nf] theorem symm_symm_apply (f : α ≃ β) (b : α) : f.symm.symm b = f b := rfl + +theorem apply_eq_iff_eq (f : α ≃ β) {x y : α} : f x = f y ↔ x = y := equiv_like.apply_eq_iff_eq f + +theorem apply_eq_iff_eq_symm_apply {α β : Sort*} (f : α ≃ β) {x : α} {y : β} : + f x = y ↔ x = f.symm y := +begin + conv_lhs { rw ←apply_symm_apply f y, }, + rw apply_eq_iff_eq, +end + +@[simp] theorem cast_apply {α β} (h : α = β) (x : α) : equiv.cast h x = cast h x := rfl + +@[simp] theorem cast_symm {α β} (h : α = β) : (equiv.cast h).symm = equiv.cast h.symm := rfl + +@[simp] theorem cast_refl {α} (h : α = α := rfl) : equiv.cast h = equiv.refl α := rfl + +@[simp] theorem cast_trans {α β γ} (h : α = β) (h2 : β = γ) : + (equiv.cast h).trans (equiv.cast h2) = equiv.cast (h.trans h2) := +ext $ λ x, by { substs h h2, refl } + +lemma cast_eq_iff_heq {α β} (h : α = β) {a : α} {b : β} : equiv.cast h a = b ↔ a == b := +by { subst h, simp } + +lemma symm_apply_eq {α β} (e : α ≃ β) {x y} : e.symm x = y ↔ x = e y := +⟨λ H, by simp [H.symm], λ H, by simp [H]⟩ + +lemma eq_symm_apply {α β} (e : α ≃ β) {x y} : y = e.symm x ↔ e y = x := +(eq_comm.trans e.symm_apply_eq).trans eq_comm + +@[simp] theorem symm_symm (e : α ≃ β) : e.symm.symm = e := by { cases e, refl } + +@[simp] theorem trans_refl (e : α ≃ β) : e.trans (equiv.refl β) = e := by { cases e, refl } + +@[simp] theorem refl_symm : (equiv.refl α).symm = equiv.refl α := rfl + +@[simp] theorem refl_trans (e : α ≃ β) : (equiv.refl α).trans e = e := by { cases e, refl } + +@[simp] theorem symm_trans_self (e : α ≃ β) : e.symm.trans e = equiv.refl β := ext (by simp) + +@[simp] theorem self_trans_symm (e : α ≃ β) : e.trans e.symm = equiv.refl α := ext (by simp) + +lemma trans_assoc {δ} (ab : α ≃ β) (bc : β ≃ γ) (cd : γ ≃ δ) : + (ab.trans bc).trans cd = ab.trans (bc.trans cd) := +equiv.ext $ assume a, rfl + +theorem left_inverse_symm (f : equiv α β) : left_inverse f.symm f := f.left_inv + +theorem right_inverse_symm (f : equiv α β) : function.right_inverse f.symm f := f.right_inv + +lemma injective_comp (e : α ≃ β) (f : β → γ) : injective (f ∘ e) ↔ injective f := +equiv_like.injective_comp e f + +lemma comp_injective (f : α → β) (e : β ≃ γ) : injective (e ∘ f) ↔ injective f := +equiv_like.comp_injective f e + +lemma surjective_comp (e : α ≃ β) (f : β → γ) : surjective (f ∘ e) ↔ surjective f := +equiv_like.surjective_comp e f + +lemma comp_surjective (f : α → β) (e : β ≃ γ) : surjective (e ∘ f) ↔ surjective f := +equiv_like.comp_surjective f e + +lemma bijective_comp (e : α ≃ β) (f : β → γ) : bijective (f ∘ e) ↔ bijective f := +equiv_like.bijective_comp e f + +lemma comp_bijective (f : α → β) (e : β ≃ γ) : bijective (e ∘ f) ↔ bijective f := +equiv_like.comp_bijective f e + +/-- If `α` is equivalent to `β` and `γ` is equivalent to `δ`, then the type of equivalences `α ≃ γ` +is equivalent to the type of equivalences `β ≃ δ`. -/ +def equiv_congr {δ} (ab : α ≃ β) (cd : γ ≃ δ) : (α ≃ γ) ≃ (β ≃ δ) := +⟨ λac, (ab.symm.trans ac).trans cd, λbd, ab.trans $ bd.trans $ cd.symm, + assume ac, by { ext x, simp }, assume ac, by { ext x, simp } ⟩ + +@[simp] lemma equiv_congr_refl {α β} : + (equiv.refl α).equiv_congr (equiv.refl β) = equiv.refl (α ≃ β) := by { ext, refl } + +@[simp] lemma equiv_congr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : + (ab.equiv_congr cd).symm = ab.symm.equiv_congr cd.symm := by { ext, refl } + +@[simp] lemma equiv_congr_trans {δ ε ζ} (ab : α ≃ β) (de : δ ≃ ε) (bc : β ≃ γ) (ef : ε ≃ ζ) : + (ab.equiv_congr de).trans (bc.equiv_congr ef) = (ab.trans bc).equiv_congr (de.trans ef) := +by { ext, refl } + +@[simp] lemma equiv_congr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : + (equiv.refl α).equiv_congr bg e = e.trans bg := rfl + +@[simp] lemma equiv_congr_refl_right {α β} (ab e : α ≃ β) : + ab.equiv_congr (equiv.refl β) e = ab.symm.trans e := rfl + +@[simp] lemma equiv_congr_apply_apply {δ} (ab : α ≃ β) (cd : γ ≃ δ) (e : α ≃ γ) (x) : + ab.equiv_congr cd e x = cd (e (ab.symm x)) := rfl + +section perm_congr + +variables {α' β' : Type*} (e : α' ≃ β') + +/-- If `α` is equivalent to `β`, then `perm α` is equivalent to `perm β`. -/ +def perm_congr : perm α' ≃ perm β' := +equiv_congr e e + +lemma perm_congr_def (p : equiv.perm α') : + e.perm_congr p = (e.symm.trans p).trans e := rfl + +@[simp] lemma perm_congr_refl : + e.perm_congr (equiv.refl _) = equiv.refl _ := +by simp [perm_congr_def] + +@[simp] lemma perm_congr_symm : + e.perm_congr.symm = e.symm.perm_congr := rfl + +@[simp] lemma perm_congr_apply (p : equiv.perm α') (x) : + e.perm_congr p x = e (p (e.symm x)) := rfl + +lemma perm_congr_symm_apply (p : equiv.perm β') (x) : + e.perm_congr.symm p x = e.symm (p (e x)) := rfl + +lemma perm_congr_trans (p p' : equiv.perm α') : + (e.perm_congr p).trans (e.perm_congr p') = e.perm_congr (p.trans p') := +by { ext, simp } + +end perm_congr + +/-- Two empty types are equivalent. -/ +def equiv_of_is_empty (α β : Sort*) [is_empty α] [is_empty β] : α ≃ β := +⟨is_empty_elim, is_empty_elim, is_empty_elim, is_empty_elim⟩ + +/-- If `α` is an empty type, then it is equivalent to the `empty` type. -/ +def equiv_empty (α : Sort u) [is_empty α] : α ≃ empty := +equiv_of_is_empty α _ + +/-- If `α` is an empty type, then it is equivalent to the `pempty` type in any universe. -/ +def equiv_pempty (α : Sort v) [is_empty α] : α ≃ pempty.{u} := +equiv_of_is_empty α _ + +/-- `α` is equivalent to an empty type iff `α` is empty. -/ +def equiv_empty_equiv (α : Sort u) : (α ≃ empty) ≃ is_empty α := +⟨λ e, function.is_empty e, @equiv_empty α, λ e, ext $ λ x, (e x).elim, λ p, rfl⟩ + +/-- The `Sort` of proofs of a false proposition is equivalent to `pempty`. -/ +def prop_equiv_pempty {p : Prop} (h : ¬p) : p ≃ pempty := +@equiv_pempty p $ is_empty.prop_iff.2 h + +/-- If both `α` and `β` have a unique element, then `α ≃ β`. -/ +def equiv_of_unique (α β : Sort*) [unique α] [unique β] : α ≃ β := +{ to_fun := default, + inv_fun := default, + left_inv := λ _, subsingleton.elim _ _, + right_inv := λ _, subsingleton.elim _ _ } + +/-- If `α` has a unique element, then it is equivalent to any `punit`. -/ +def equiv_punit (α : Sort*) [unique α] : α ≃ punit.{v} := +equiv_of_unique α _ + +/-- The `Sort` of proofs of a true proposition is equivalent to `punit`. -/ +def prop_equiv_punit {p : Prop} (h : p) : p ≃ punit := +@equiv_punit p $ unique_prop h + +/-- `ulift α` is equivalent to `α`. -/ +@[simps apply symm_apply {fully_applied := ff}] +protected def ulift {α : Type v} : ulift.{u} α ≃ α := +⟨ulift.down, ulift.up, ulift.up_down, λ a, rfl⟩ + +/-- `plift α` is equivalent to `α`. -/ +@[simps apply symm_apply {fully_applied := ff}] +protected def plift : plift α ≃ α := +⟨plift.down, plift.up, plift.up_down, plift.down_up⟩ + +/-- equivalence of propositions is the same as iff -/ +def of_iff {P Q : Prop} (h : P ↔ Q) : P ≃ Q := +{ to_fun := h.mp, + inv_fun := h.mpr, + left_inv := λ x, rfl, + right_inv := λ y, rfl } + +/-- If `α₁` is equivalent to `α₂` and `β₁` is equivalent to `β₂`, then the type of maps `α₁ → β₁` +is equivalent to the type of maps `α₂ → β₂`. -/ +@[congr, simps apply] def arrow_congr {α₁ β₁ α₂ β₂ : Sort*} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : + (α₁ → β₁) ≃ (α₂ → β₂) := +{ to_fun := λ f, e₂ ∘ f ∘ e₁.symm, + inv_fun := λ f, e₂.symm ∘ f ∘ e₁, + left_inv := λ f, funext $ λ x, by simp, + right_inv := λ f, funext $ λ x, by simp } + +lemma arrow_congr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort*} + (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) (f : α₁ → β₁) (g : β₁ → γ₁) : + arrow_congr ea ec (g ∘ f) = (arrow_congr eb ec g) ∘ (arrow_congr ea eb f) := +by { ext, simp only [comp, arrow_congr_apply, eb.symm_apply_apply] } + +@[simp] lemma arrow_congr_refl {α β : Sort*} : + arrow_congr (equiv.refl α) (equiv.refl β) = equiv.refl (α → β) := rfl + +@[simp] lemma arrow_congr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort*} + (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : + arrow_congr (e₁.trans e₂) (e₁'.trans e₂') = (arrow_congr e₁ e₁').trans (arrow_congr e₂ e₂') := +rfl + +@[simp] lemma arrow_congr_symm {α₁ β₁ α₂ β₂ : Sort*} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : + (arrow_congr e₁ e₂).symm = arrow_congr e₁.symm e₂.symm := +rfl + +/-- +A version of `equiv.arrow_congr` in `Type`, rather than `Sort`. + +The `equiv_rw` tactic is not able to use the default `Sort` level `equiv.arrow_congr`, +because Lean's universe rules will not unify `?l_1` with `imax (1 ?m_1)`. +-/ +@[congr, simps apply] +def arrow_congr' {α₁ β₁ α₂ β₂ : Type*} (hα : α₁ ≃ α₂) (hβ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) := +equiv.arrow_congr hα hβ + +@[simp] lemma arrow_congr'_refl {α β : Type*} : + arrow_congr' (equiv.refl α) (equiv.refl β) = equiv.refl (α → β) := rfl + +@[simp] lemma arrow_congr'_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Type*} + (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : + arrow_congr' (e₁.trans e₂) (e₁'.trans e₂') = (arrow_congr' e₁ e₁').trans (arrow_congr' e₂ e₂') := +rfl + +@[simp] lemma arrow_congr'_symm {α₁ β₁ α₂ β₂ : Type*} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : + (arrow_congr' e₁ e₂).symm = arrow_congr' e₁.symm e₂.symm := +rfl + +/-- Conjugate a map `f : α → α` by an equivalence `α ≃ β`. -/ +@[simps apply] +def conj (e : α ≃ β) : (α → α) ≃ (β → β) := arrow_congr e e + +@[simp] lemma conj_refl : conj (equiv.refl α) = equiv.refl (α → α) := rfl + +@[simp] lemma conj_symm (e : α ≃ β) : e.conj.symm = e.symm.conj := rfl + +@[simp] lemma conj_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : + (e₁.trans e₂).conj = e₁.conj.trans e₂.conj := +rfl + +-- This should not be a simp lemma as long as `(∘)` is reducible: +-- when `(∘)` is reducible, Lean can unify `f₁ ∘ f₂` with any `g` using +-- `f₁ := g` and `f₂ := λ x, x`. This causes nontermination. +lemma conj_comp (e : α ≃ β) (f₁ f₂ : α → α) : + e.conj (f₁ ∘ f₂) = (e.conj f₁) ∘ (e.conj f₂) := +by apply arrow_congr_comp + +lemma eq_comp_symm {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : + f = g ∘ e.symm ↔ f ∘ e = g := +(e.arrow_congr (equiv.refl γ)).symm_apply_eq.symm + +lemma comp_symm_eq {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : + g ∘ e.symm = f ↔ g = f ∘ e := +(e.arrow_congr (equiv.refl γ)).eq_symm_apply.symm + +lemma eq_symm_comp {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) : + f = e.symm ∘ g ↔ e ∘ f = g := +((equiv.refl γ).arrow_congr e).eq_symm_apply + +lemma symm_comp_eq {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) : + e.symm ∘ g = f ↔ g = e ∘ f := +((equiv.refl γ).arrow_congr e).symm_apply_eq + +/-- `punit` sorts in any two universes are equivalent. -/ +def punit_equiv_punit : punit.{v} ≃ punit.{w} := +⟨λ _, punit.star, λ _, punit.star, λ u, by { cases u, refl }, λ u, by { cases u, reflexivity }⟩ + +/-- `Prop` is noncomputably equivalent to `bool`. -/ +noncomputable def Prop_equiv_bool : Prop ≃ bool := +⟨λ p, @to_bool p (classical.prop_decidable _), + λ b, b, λ p, by simp, λ b, by simp⟩ + +section +/-- The sort of maps to `punit.{v}` is equivalent to `punit.{w}`. -/ +def arrow_punit_equiv_punit (α : Sort*) : (α → punit.{v}) ≃ punit.{w} := +⟨λ f, punit.star, λ u f, punit.star, + λ f, by { funext x, cases f x, refl }, λ u, by { cases u, reflexivity }⟩ + +/-- If `α` is `subsingleton` and `a : α`, then the type of dependent functions `Π (i : α), β +i` is equivalent to `β i`. -/ +@[simps] +def Pi_subsingleton {α} (β : α → Sort*) [subsingleton α] (a : α) : (Π a', β a') ≃ β a := +{ to_fun := eval a, + inv_fun := λ x b, cast (congr_arg β $ subsingleton.elim a b) x, + left_inv := λ f, funext $ λ b, by { rw subsingleton.elim b a, reflexivity }, + right_inv := λ b, rfl } + +/-- If `α` has a unique term, then the type of function `α → β` is equivalent to `β`. -/ +@[simps { fully_applied := ff }] def fun_unique (α β) [unique α] : (α → β) ≃ β := +Pi_subsingleton _ default + +/-- The sort of maps from `punit` is equivalent to the codomain. -/ +def punit_arrow_equiv (α : Sort*) : (punit.{u} → α) ≃ α := +fun_unique _ _ + +/-- The sort of maps from `true` is equivalent to the codomain. -/ +def true_arrow_equiv (α : Sort*) : (true → α) ≃ α := +fun_unique _ _ + +/-- The sort of maps from a type that `is_empty` is equivalent to `punit`. -/ +def arrow_punit_of_is_empty (α β : Sort*) [is_empty α] : (α → β) ≃ punit.{u} := +⟨λ f, punit.star, λ u, is_empty_elim, λ f, funext is_empty_elim, λ u, by { cases u, refl }⟩ + +/-- The sort of maps from `empty` is equivalent to `punit`. -/ +def empty_arrow_equiv_punit (α : Sort*) : (empty → α) ≃ punit.{u} := +arrow_punit_of_is_empty _ _ + +/-- The sort of maps from `pempty` is equivalent to `punit`. -/ +def pempty_arrow_equiv_punit (α : Sort*) : (pempty → α) ≃ punit.{u} := +arrow_punit_of_is_empty _ _ + +/-- The sort of maps from `false` is equivalent to `punit`. -/ +def false_arrow_equiv_punit (α : Sort*) : (false → α) ≃ punit.{u} := +arrow_punit_of_is_empty _ _ + +end + +section + +/-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ +@[simps apply symm_apply] def psigma_equiv_sigma {α} (β : α → Type*) : (Σ' i, β i) ≃ Σ i, β i := +⟨λ a, ⟨a.1, a.2⟩, λ a, ⟨a.1, a.2⟩, λ ⟨a, b⟩, rfl, λ ⟨a, b⟩, rfl⟩ + +/-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ +@[simps apply symm_apply] def psigma_equiv_sigma_plift {α} (β : α → Sort*) : + (Σ' i, β i) ≃ Σ i : plift α, plift (β i.down) := +⟨λ a, ⟨plift.up a.1, plift.up a.2⟩, λ a, ⟨a.1.down, a.2.down⟩, λ ⟨a, b⟩, rfl, λ ⟨⟨a⟩, ⟨b⟩⟩, rfl⟩ + +/-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ' a, β₁ a` and +`Σ' a, β₂ a`. -/ +@[simps apply] +def psigma_congr_right {α} {β₁ β₂ : α → Sort*} (F : Π a, β₁ a ≃ β₂ a) : (Σ' a, β₁ a) ≃ Σ' a, β₂ a := +⟨λ a, ⟨a.1, F a.1 a.2⟩, λ a, ⟨a.1, (F a.1).symm a.2⟩, + λ ⟨a, b⟩, congr_arg (psigma.mk a) $ symm_apply_apply (F a) b, + λ ⟨a, b⟩, congr_arg (psigma.mk a) $ apply_symm_apply (F a) b⟩ + +@[simp] lemma psigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Sort*} + (F : Π a, β₁ a ≃ β₂ a) (G : Π a, β₂ a ≃ β₃ a) : + (psigma_congr_right F).trans (psigma_congr_right G) = + psigma_congr_right (λ a, (F a).trans (G a)) := +by { ext1 x, cases x, refl } + +@[simp] lemma psigma_congr_right_symm {α} {β₁ β₂ : α → Sort*} (F : Π a, β₁ a ≃ β₂ a) : + (psigma_congr_right F).symm = psigma_congr_right (λ a, (F a).symm) := +by { ext1 x, cases x, refl } + +@[simp] lemma psigma_congr_right_refl {α} {β : α → Sort*} : + (psigma_congr_right (λ a, equiv.refl (β a))) = equiv.refl (Σ' a, β a) := +by { ext1 x, cases x, refl } + +/-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ a, β₁ a` and +`Σ a, β₂ a`. -/ +@[simps apply] +def sigma_congr_right {α} {β₁ β₂ : α → Type*} (F : Π a, β₁ a ≃ β₂ a) : (Σ a, β₁ a) ≃ Σ a, β₂ a := +⟨λ a, ⟨a.1, F a.1 a.2⟩, λ a, ⟨a.1, (F a.1).symm a.2⟩, + λ ⟨a, b⟩, congr_arg (sigma.mk a) $ symm_apply_apply (F a) b, + λ ⟨a, b⟩, congr_arg (sigma.mk a) $ apply_symm_apply (F a) b⟩ + +@[simp] lemma sigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Type*} + (F : Π a, β₁ a ≃ β₂ a) (G : Π a, β₂ a ≃ β₃ a) : + (sigma_congr_right F).trans (sigma_congr_right G) = sigma_congr_right (λ a, (F a).trans (G a)) := +by { ext1 x, cases x, refl } + +@[simp] lemma sigma_congr_right_symm {α} {β₁ β₂ : α → Type*} (F : Π a, β₁ a ≃ β₂ a) : + (sigma_congr_right F).symm = sigma_congr_right (λ a, (F a).symm) := +by { ext1 x, cases x, refl } + +@[simp] lemma sigma_congr_right_refl {α} {β : α → Type*} : + (sigma_congr_right (λ a, equiv.refl (β a))) = equiv.refl (Σ a, β a) := +by { ext1 x, cases x, refl } + +/-- A `psigma` with `Prop` fibers is equivalent to the subtype. -/ +def psigma_equiv_subtype {α : Type v} (P : α → Prop) : + (Σ' i, P i) ≃ subtype P := +{ to_fun := λ x, ⟨x.1, x.2⟩, + inv_fun := λ x, ⟨x.1, x.2⟩, + left_inv := λ x, by { cases x, refl, }, + right_inv := λ x, by { cases x, refl, }, } + +/-- A `sigma` with `plift` fibers is equivalent to the subtype. -/ +def sigma_plift_equiv_subtype {α : Type v} (P : α → Prop) : + (Σ i, plift (P i)) ≃ subtype P := +((psigma_equiv_sigma _).symm.trans (psigma_congr_right (λ a, equiv.plift))).trans + (psigma_equiv_subtype P) + +/-- +A `sigma` with `λ i, ulift (plift (P i))` fibers is equivalent to `{ x // P x }`. +Variant of `sigma_plift_equiv_subtype`. +-/ +def sigma_ulift_plift_equiv_subtype {α : Type v} (P : α → Prop) : + (Σ i, ulift (plift (P i))) ≃ subtype P := +(sigma_congr_right (λ a, equiv.ulift)).trans (sigma_plift_equiv_subtype P) + +namespace perm + +/-- A family of permutations `Π a, perm (β a)` generates a permuation `perm (Σ a, β₁ a)`. -/ +@[reducible] +def sigma_congr_right {α} {β : α → Sort*} (F : Π a, perm (β a)) : perm (Σ a, β a) := +equiv.sigma_congr_right F + +@[simp] lemma sigma_congr_right_trans {α} {β : α → Sort*} + (F : Π a, perm (β a)) (G : Π a, perm (β a)) : + (sigma_congr_right F).trans (sigma_congr_right G) = sigma_congr_right (λ a, (F a).trans (G a)) := +equiv.sigma_congr_right_trans F G + +@[simp] lemma sigma_congr_right_symm {α} {β : α → Sort*} (F : Π a, perm (β a)) : + (sigma_congr_right F).symm = sigma_congr_right (λ a, (F a).symm) := +equiv.sigma_congr_right_symm F + +@[simp] lemma sigma_congr_right_refl {α} {β : α → Sort*} : + (sigma_congr_right (λ a, equiv.refl (β a))) = equiv.refl (Σ a, β a) := +equiv.sigma_congr_right_refl + +end perm + +/-- An equivalence `f : α₁ ≃ α₂` generates an equivalence between `Σ a, β (f a)` and `Σ a, β a`. -/ +@[simps apply] +def sigma_congr_left {α₁ α₂} {β : α₂ → Sort*} (e : α₁ ≃ α₂) : (Σ a:α₁, β (e a)) ≃ (Σ a:α₂, β a) := +⟨λ a, ⟨e a.1, a.2⟩, λ a, ⟨e.symm a.1, @@eq.rec β a.2 (e.right_inv a.1).symm⟩, + λ ⟨a, b⟩, match e.symm (e a), e.left_inv a : ∀ a' (h : a' = a), + @sigma.mk _ (β ∘ e) _ (@@eq.rec β b (congr_arg e h.symm)) = ⟨a, b⟩ with + | _, rfl := rfl end, + λ ⟨a, b⟩, match e (e.symm a), _ : ∀ a' (h : a' = a), + sigma.mk a' (@@eq.rec β b h.symm) = ⟨a, b⟩ with + | _, rfl := rfl end⟩ + +/-- Transporting a sigma type through an equivalence of the base -/ +def sigma_congr_left' {α₁ α₂} {β : α₁ → Sort*} (f : α₁ ≃ α₂) : + (Σ a:α₁, β a) ≃ (Σ a:α₂, β (f.symm a)) := +(sigma_congr_left f.symm).symm + +/-- Transporting a sigma type through an equivalence of the base and a family of equivalences +of matching fibers -/ +def sigma_congr {α₁ α₂} {β₁ : α₁ → Sort*} {β₂ : α₂ → Sort*} (f : α₁ ≃ α₂) + (F : ∀ a, β₁ a ≃ β₂ (f a)) : + sigma β₁ ≃ sigma β₂ := +(sigma_congr_right F).trans (sigma_congr_left f) + +/-- `sigma` type with a constant fiber is equivalent to the product. -/ +@[simps apply symm_apply] def sigma_equiv_prod (α β : Type*) : (Σ_:α, β) ≃ α × β := +⟨λ a, ⟨a.1, a.2⟩, λ a, ⟨a.1, a.2⟩, λ ⟨a, b⟩, rfl, λ ⟨a, b⟩, rfl⟩ + +/-- If each fiber of a `sigma` type is equivalent to a fixed type, then the sigma type +is equivalent to the product. -/ +def sigma_equiv_prod_of_equiv {α β} {β₁ : α → Sort*} (F : Π a, β₁ a ≃ β) : sigma β₁ ≃ α × β := +(sigma_congr_right F).trans (sigma_equiv_prod α β) + +/-- Dependent product of types is associative up to an equivalence. -/ +def sigma_assoc {α : Type*} {β : α → Type*} (γ : Π (a : α), β a → Type*) : + (Σ (ab : Σ (a : α), β a), γ ab.1 ab.2) ≃ Σ (a : α), (Σ (b : β a), γ a b) := +{ to_fun := λ x, ⟨x.1.1, ⟨x.1.2, x.2⟩⟩, + inv_fun := λ x, ⟨⟨x.1, x.2.1⟩, x.2.2⟩, + left_inv := λ ⟨⟨a, b⟩, c⟩, rfl, + right_inv := λ ⟨a, ⟨b, c⟩⟩, rfl } + +end + +protected lemma exists_unique_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) + (h : ∀{x}, p x ↔ q (f x)) : (∃! x, p x) ↔ ∃! y, q y := +begin + split, + { rintro ⟨a, ha₁, ha₂⟩, + exact ⟨f a, h.1 ha₁, λ b hb, f.symm_apply_eq.1 (ha₂ (f.symm b) (h.2 (by simpa using hb)))⟩ }, + { rintro ⟨b, hb₁, hb₂⟩, + exact ⟨f.symm b, h.2 (by simpa using hb₁), λ y hy, (eq_symm_apply f).2 (hb₂ _ (h.1 hy))⟩ } +end + +protected lemma exists_unique_congr_left' {p : α → Prop} (f : α ≃ β) : + (∃! x, p x) ↔ (∃! y, p (f.symm y)) := +equiv.exists_unique_congr f (λx, by simp) + +protected lemma exists_unique_congr_left {p : β → Prop} (f : α ≃ β) : + (∃! x, p (f x)) ↔ (∃! y, p y) := +(equiv.exists_unique_congr_left' f.symm).symm + +protected lemma forall_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) + (h : ∀{x}, p x ↔ q (f x)) : (∀x, p x) ↔ (∀y, q y) := +begin + split; intros h₂ x, + { rw [←f.right_inv x], apply h.mp, apply h₂ }, + apply h.mpr, apply h₂ +end +protected lemma forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ β) + (h : ∀{x}, p (f.symm x) ↔ q x) : (∀x, p x) ↔ (∀y, q y) := +(equiv.forall_congr f.symm (λ x, h.symm)).symm + +-- We next build some higher arity versions of `equiv.forall_congr`. +-- Although they appear to just be repeated applications of `equiv.forall_congr`, +-- unification of metavariables works better with these versions. +-- In particular, they are necessary in `equiv_rw`. +-- (Stopping at ternary functions seems reasonable: at least in 1-categorical mathematics, +-- it's rare to have axioms involving more than 3 elements at once.) +universes ua1 ua2 ub1 ub2 ug1 ug2 +variables {α₁ : Sort ua1} {α₂ : Sort ua2} + {β₁ : Sort ub1} {β₂ : Sort ub2} + {γ₁ : Sort ug1} {γ₂ : Sort ug2} + +protected lemma forall₂_congr {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) + (eβ : β₁ ≃ β₂) (h : ∀{x y}, p x y ↔ q (eα x) (eβ y)) : + (∀x y, p x y) ↔ (∀x y, q x y) := +begin + apply equiv.forall_congr, + intros, + apply equiv.forall_congr, + intros, + apply h, +end +protected lemma forall₂_congr' {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) + (eβ : β₁ ≃ β₂) (h : ∀{x y}, p (eα.symm x) (eβ.symm y) ↔ q x y) : + (∀x y, p x y) ↔ (∀x y, q x y) := +(equiv.forall₂_congr eα.symm eβ.symm (λ x y, h.symm)).symm + +protected lemma forall₃_congr {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} + (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) + (h : ∀{x y z}, p x y z ↔ q (eα x) (eβ y) (eγ z)) : (∀x y z, p x y z) ↔ (∀x y z, q x y z) := +begin + apply equiv.forall₂_congr, + intros, + apply equiv.forall_congr, + intros, + apply h, +end +protected lemma forall₃_congr' {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} + (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) + (h : ∀{x y z}, p (eα.symm x) (eβ.symm y) (eγ.symm z) ↔ q x y z) : + (∀x y z, p x y z) ↔ (∀x y z, q x y z) := +(equiv.forall₃_congr eα.symm eβ.symm eγ.symm (λ x y z, h.symm)).symm + +protected lemma forall_congr_left' {p : α → Prop} (f : α ≃ β) : + (∀x, p x) ↔ (∀y, p (f.symm y)) := +equiv.forall_congr f (λx, by simp) + +protected lemma forall_congr_left {p : β → Prop} (f : α ≃ β) : + (∀x, p (f x)) ↔ (∀y, p y) := +(equiv.forall_congr_left' f.symm).symm + +protected lemma exists_congr_left {α β} (f : α ≃ β) {p : α → Prop} : + (∃ a, p a) ↔ (∃ b, p (f.symm b)) := +⟨λ ⟨a, h⟩, ⟨f a, by simpa using h⟩, λ ⟨b, h⟩, ⟨_, h⟩⟩ + +end equiv + + +namespace quot + +/-- An equivalence `e : α ≃ β` generates an equivalence between quotient spaces, +if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ +protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) + (eq : ∀a₁ a₂, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) : + quot ra ≃ quot rb := +{ to_fun := quot.map e (assume a₁ a₂, (eq a₁ a₂).1), + inv_fun := quot.map e.symm + (assume b₁ b₂ h, + (eq (e.symm b₁) (e.symm b₂)).2 + ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h)), + left_inv := by { rintros ⟨a⟩, dunfold quot.map, simp only [equiv.symm_apply_apply] }, + right_inv := by { rintros ⟨a⟩, dunfold quot.map, simp only [equiv.apply_symm_apply] } } + +@[simp] +lemma congr_mk {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) + (eq : ∀ (a₁ a₂ : α), ra a₁ a₂ ↔ rb (e a₁) (e a₂)) (a : α) : + quot.congr e eq (quot.mk ra a) = quot.mk rb (e a) := rfl + +/-- Quotients are congruent on equivalences under equality of their relation. +An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ +protected def congr_right {r r' : α → α → Prop} (eq : ∀a₁ a₂, r a₁ a₂ ↔ r' a₁ a₂) : + quot r ≃ quot r' := +quot.congr (equiv.refl α) eq + +/-- An equivalence `e : α ≃ β` generates an equivalence between the quotient space of `α` +by a relation `ra` and the quotient space of `β` by the image of this relation under `e`. -/ +protected def congr_left {r : α → α → Prop} (e : α ≃ β) : + quot r ≃ quot (λ b b', r (e.symm b) (e.symm b')) := +@quot.congr α β r (λ b b', r (e.symm b) (e.symm b')) e (λ a₁ a₂, by simp only [e.symm_apply_apply]) + +end quot + +namespace quotient +/-- An equivalence `e : α ≃ β` generates an equivalence between quotient spaces, +if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ +protected def congr {ra : setoid α} {rb : setoid β} (e : α ≃ β) + (eq : ∀a₁ a₂, @setoid.r α ra a₁ a₂ ↔ @setoid.r β rb (e a₁) (e a₂)) : + quotient ra ≃ quotient rb := +quot.congr e eq + +@[simp] +lemma congr_mk {ra : setoid α} {rb : setoid β} (e : α ≃ β) + (eq : ∀ (a₁ a₂ : α), setoid.r a₁ a₂ ↔ setoid.r (e a₁) (e a₂)) (a : α): + quotient.congr e eq (quotient.mk a) = quotient.mk (e a) := +rfl + +/-- Quotients are congruent on equivalences under equality of their relation. +An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ +protected def congr_right {r r' : setoid α} + (eq : ∀a₁ a₂, @setoid.r α r a₁ a₂ ↔ @setoid.r α r' a₁ a₂) : quotient r ≃ quotient r' := +quot.congr_right eq + +end quotient diff --git a/src/logic/equiv/fin.lean b/src/logic/equiv/fin.lean index 6beacf7bd6b2b..a5b14633c8397 100644 --- a/src/logic/equiv/fin.lean +++ b/src/logic/equiv/fin.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import data.fin.vec_notation -import logic.equiv.basic +import logic.equiv.defs /-! # Equivalences for `fin n` diff --git a/src/logic/equiv/fintype.lean b/src/logic/equiv/fintype.lean index a04aa3628a1a4..97b99b7ec5ff2 100644 --- a/src/logic/equiv/fintype.lean +++ b/src/logic/equiv/fintype.lean @@ -6,7 +6,7 @@ Authors: Yakov Pechersky import data.fintype.basic import group_theory.perm.sign -import logic.equiv.basic +import logic.equiv.defs /-! # Equivalence between fintypes diff --git a/src/logic/equiv/functor.lean b/src/logic/equiv/functor.lean index 3ecd01ad4c5de..35b625f2b8f35 100644 --- a/src/logic/equiv/functor.lean +++ b/src/logic/equiv/functor.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Simon Hudon, Scott Morrison -/ import control.bifunctor -import logic.equiv.basic +import logic.equiv.defs /-! # Functor and bifunctors can be applied to `equiv`s. diff --git a/src/logic/equiv/local_equiv.lean b/src/logic/equiv/local_equiv.lean index 558e1f788c99f..6b7ea9aef3b8c 100644 --- a/src/logic/equiv/local_equiv.lean +++ b/src/logic/equiv/local_equiv.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import data.set.function -import logic.equiv.basic +import logic.equiv.defs /-! # Local equivalences diff --git a/src/logic/equiv/set.lean b/src/logic/equiv/set.lean index 4d6c03580206a..76666b80078ca 100644 --- a/src/logic/equiv/set.lean +++ b/src/logic/equiv/set.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro -/ import data.set.function -import logic.equiv.basic +import logic.equiv.defs /-! # Equivalences and sets diff --git a/src/logic/equiv/transfer_instance.lean b/src/logic/equiv/transfer_instance.lean index c5cdb3b84e8bc..1c3bdd8361686 100644 --- a/src/logic/equiv/transfer_instance.lean +++ b/src/logic/equiv/transfer_instance.lean @@ -6,7 +6,7 @@ Authors: Johannes Hölzl import algebra.algebra.basic import algebra.field.basic import algebra.group.type_tags -import logic.equiv.basic +import logic.equiv.defs import ring_theory.ideal.local_ring /-! diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index a6e39ab08d3da..d67d4b03812ce 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import order.lattice +import logic.equiv.basic /-! # ⊤ and ⊥, bounded lattices and variants diff --git a/src/order/synonym.lean b/src/order/synonym.lean index 05a20a865f6dc..acf8d6b2272cb 100644 --- a/src/order/synonym.lean +++ b/src/order/synonym.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Johan Commelin, Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Damiano Testa, Yaël Dillies -/ -import logic.equiv.basic +import logic.equiv.defs import logic.nontrivial import order.basic diff --git a/src/tactic/equiv_rw.lean b/src/tactic/equiv_rw.lean index d23a1d86db529..766703bf8924f 100644 --- a/src/tactic/equiv_rw.lean +++ b/src/tactic/equiv_rw.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import logic.equiv.basic +import logic.equiv.defs import tactic.clear import tactic.simp_result import tactic.apply diff --git a/src/tactic/norm_swap.lean b/src/tactic/norm_swap.lean index f6d29ae9991f9..d6dd9aa727c7e 100644 --- a/src/tactic/norm_swap.lean +++ b/src/tactic/norm_swap.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Yakov Pechersky All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yakov Pechersky -/ -import logic.equiv.basic +import logic.equiv.defs import tactic.norm_fin /-! diff --git a/test/norm_swap.lean b/test/norm_swap.lean index 104f0d8988203..b89621b1b9ce5 100644 --- a/test/norm_swap.lean +++ b/test/norm_swap.lean @@ -1,4 +1,4 @@ -import logic.equiv.basic +import logic.equiv.defs import tactic.norm_swap diff --git a/test/simp_result.lean b/test/simp_result.lean index 1b04491e4b0d8..42fd22cf8b083 100644 --- a/test/simp_result.lean +++ b/test/simp_result.lean @@ -1,4 +1,4 @@ -import logic.equiv.basic +import logic.equiv.defs import tactic.simp_result open tactic From edfb391a574b045d67be5994e06cb6f6b490106e Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 19 Oct 2022 01:53:33 +0000 Subject: [PATCH 809/828] chore(algebra/module/basic): reduce imports (#17042) Avoids importing `algebra.big_operators.basic` in a few files that don't need it. Co-authored-by: Scott Morrison --- src/algebra/module/basic.lean | 3 --- src/data/dfinsupp/basic.lean | 1 + 2 files changed, 1 insertion(+), 3 deletions(-) diff --git a/src/algebra/module/basic.lean b/src/algebra/module/basic.lean index 2ea00c13212df..a4b3a3ab0e125 100644 --- a/src/algebra/module/basic.lean +++ b/src/algebra/module/basic.lean @@ -3,10 +3,7 @@ Copyright (c) 2015 Nathaniel Thomas. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro -/ -import algebra.big_operators.basic import algebra.smul_with_zero -import data.rat.cast -import group_theory.group_action.big_operators import group_theory.group_action.group import tactic.abel diff --git a/src/data/dfinsupp/basic.lean b/src/data/dfinsupp/basic.lean index bfb3adf4aab11..68aff64e3e7cc 100644 --- a/src/data/dfinsupp/basic.lean +++ b/src/data/dfinsupp/basic.lean @@ -8,6 +8,7 @@ import algebra.module.linear_map import algebra.big_operators.basic import data.set.finite import group_theory.submonoid.membership +import group_theory.group_action.big_operators import data.finset.preimage /-! From f01de4e1f3a0bf32fd07c381a7b361d007875e1b Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Wed, 19 Oct 2022 04:48:07 +0000 Subject: [PATCH 810/828] =?UTF-8?q?feat(topology/algebra/uniform=5Fconverg?= =?UTF-8?q?ence):=20criterion=20for=20a=20vector=20subspace=20of=20`=CE=B1?= =?UTF-8?q?=20=E2=86=92=20E`=20to=20be=20a=20TVS=20for=20the=20topology=20?= =?UTF-8?q?of=20=F0=9D=94=96-convergence=20(#14857)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The main theorem is `uniform_convergence_on.has_continuous_smul_induced_of_image_bounded`. As explained in the module docstring, one could get rid of requiring `𝔖` to be nonempty and directed, but the easiest way to get that is to wait until we know that replacing `𝔖` by its ***noncovering*** bornology (i.e ***not*** what `bornology` currently refers to in mathlib) doesn't change the topology. This will allow to prove that strong topologies on the space of continuous linear maps between two TVSs are also TVS topologies --- src/topology/algebra/uniform_convergence.lean | 154 +++++++++++++++++- 1 file changed, 148 insertions(+), 6 deletions(-) diff --git a/src/topology/algebra/uniform_convergence.lean b/src/topology/algebra/uniform_convergence.lean index 461c976b2f718..767333f59a91c 100644 --- a/src/topology/algebra/uniform_convergence.lean +++ b/src/topology/algebra/uniform_convergence.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anatole Dedecker -/ import topology.uniform_space.uniform_convergence_topology -import topology.algebra.uniform_group +import analysis.locally_convex.bounded +import topology.algebra.filter_basis /-! # Algebraic facts about the topology of uniform convergence @@ -18,17 +19,23 @@ space of continuous linear maps between two topological vector spaces. * `uniform_convergence.uniform_group` : if `G` is a uniform group, then the uniform structure of uniform convergence makes `α → G` a uniform group * `uniform_convergence_on.uniform_group` : if `G` is a uniform group, then the uniform structure of - `𝔖`-convergence, for any `𝔖 : set (set α)`, makes `α → G` a uniform group + `𝔖`-convergence, for any `𝔖 : set (set α)`, makes `α → G` a uniform group. +* `uniform_convergence_on.has_continuous_smul_of_image_bounded` : let `E` be a TVS, + `𝔖 : set (set α)` and `H` a submodule of `α → E`. If the image of any `S ∈ 𝔖` by any `u ∈ H` is + bounded (in the sense of `bornology.is_vonN_bounded`), then `H`, equipped with the topology of + `𝔖`-convergence, is a TVS. ## TODO -* Let `E` be a TVS, `𝔖 : set (set α)` and `H` a submodule of `α → E`. If the image of any `S ∈ 𝔖` - by any `u ∈ H` is bounded (in the sense of `bornology.is_vonN_bounded`), then `H`, equipped with - the topology of `𝔖`-convergence, is a TVS. +* `uniform_convergence_on.has_continuous_smul_of_image_bounded` unnecessarily asks for `𝔖` to be + nonempty and directed. This will be easy to solve once we know that replacing `𝔖` by its + ***noncovering*** bornology (i.e ***not*** what `bornology` currently refers to in mathlib) + doesn't change the topology. ## References * [N. Bourbaki, *General Topology, Chapter X*][bourbaki1966] +* [N. Bourbaki, *Topological Vector Spaces*][bourbaki1987] ## Tags @@ -36,11 +43,15 @@ uniform convergence, strong dual -/ +open filter +open_locale topological_space pointwise + section group -variables {α G : Type*} [group G] [uniform_space G] [uniform_group G] {𝔖 : set $ set α} +variables {α G ι : Type*} [group G] [uniform_space G] [uniform_group G] {𝔖 : set $ set α} local attribute [-instance] Pi.uniform_space +local attribute [-instance] Pi.topological_space /-- If `G` is a uniform group, then the uniform structure of uniform convergence makes `α → G` a uniform group as well. -/ @@ -60,6 +71,26 @@ begin uniform_convergence.uniform_equiv_prod_arrow.symm.uniform_continuous⟩ end +@[to_additive] +protected lemma uniform_convergence.has_basis_nhds_one_of_basis {p : ι → Prop} + {b : ι → set G} (h : (𝓝 1 : filter G).has_basis p b) : + (@nhds (α → G) (uniform_convergence.topological_space α G) 1).has_basis p + (λ i, {f : α → G | ∀ x, f x ∈ b i}) := +begin + have := h.comap (λ p : G × G, p.2 / p.1), + rw ← uniformity_eq_comap_nhds_one at this, + convert uniform_convergence.has_basis_nhds_of_basis α _ 1 this, + ext i f, + simp [uniform_convergence.gen] +end + +@[to_additive] +protected lemma uniform_convergence.has_basis_nhds_one : + (@nhds (α → G) (uniform_convergence.topological_space α G) 1).has_basis + (λ V : set G, V ∈ (𝓝 1 : filter G)) + (λ V, {f : α → G | ∀ x, f x ∈ V}) := +uniform_convergence.has_basis_nhds_one_of_basis (basis_sets _) + /-- Let `𝔖 : set (set α)`. If `G` is a uniform group, then the uniform structure of `𝔖`-convergence makes `α → G` a uniform group as well. -/ @[to_additive "Let `𝔖 : set (set α)`. If `G` is a uniform additive group, then the uniform @@ -78,4 +109,115 @@ begin uniform_convergence_on.uniform_equiv_prod_arrow.symm.uniform_continuous⟩ end +@[to_additive] +protected lemma uniform_convergence_on.has_basis_nhds_one_of_basis (𝔖 : set $ set α) + (h𝔖₁ : 𝔖.nonempty) (h𝔖₂ : directed_on (⊆) 𝔖) {p : ι → Prop} + {b : ι → set G} (h : (𝓝 1 : filter G).has_basis p b) : + (@nhds (α → G) (uniform_convergence_on.topological_space α G 𝔖) 1).has_basis + (λ Si : set α × ι, Si.1 ∈ 𝔖 ∧ p Si.2) + (λ Si, {f : α → G | ∀ x ∈ Si.1, f x ∈ b Si.2}) := +begin + have := h.comap (λ p : G × G, p.1 / p.2), + rw ← uniformity_eq_comap_nhds_one_swapped at this, + convert uniform_convergence_on.has_basis_nhds_of_basis α _ 𝔖 1 h𝔖₁ h𝔖₂ this, + ext i f, + simp [uniform_convergence_on.gen] +end + +@[to_additive] +protected lemma uniform_convergence_on.has_basis_nhds_one (𝔖 : set $ set α) + (h𝔖₁ : 𝔖.nonempty) (h𝔖₂ : directed_on (⊆) 𝔖) : + (@nhds (α → G) (uniform_convergence_on.topological_space α G 𝔖) 1).has_basis + (λ SV : set α × set G, SV.1 ∈ 𝔖 ∧ SV.2 ∈ (𝓝 1 : filter G)) + (λ SV, {f : α → G | ∀ x ∈ SV.1, f x ∈ SV.2}) := +uniform_convergence_on.has_basis_nhds_one_of_basis 𝔖 h𝔖₁ h𝔖₂ (basis_sets _) + end group + +section module + +variables (𝕜 α E H : Type*) {hom : Type*} [normed_field 𝕜] [add_comm_group H] [module 𝕜 H] + [add_comm_group E] [module 𝕜 E] [linear_map_class hom 𝕜 H (α → E)] [topological_space H] + [uniform_space E] [uniform_add_group E] [has_continuous_smul 𝕜 E] {𝔖 : set $ set α} + +local attribute [-instance] Pi.uniform_space +local attribute [-instance] Pi.topological_space + +/-- Let `E` be a TVS, `𝔖 : set (set α)` and `H` a submodule of `α → E`. If the image of any `S ∈ 𝔖` +by any `u ∈ H` is bounded (in the sense of `bornology.is_vonN_bounded`), then `H`, equipped with +the topology of `𝔖`-convergence, is a TVS. + +For convenience, we don't literally ask for `H : submodule (α → E)`. Instead, we prove the result +for any vector space `H` equipped with a linear inducing to `α → E`, which is often easier to use. +We also state the `submodule` version as +`uniform_convergence_on.has_continuous_smul_submodule_of_image_bounded`. -/ +lemma uniform_convergence_on.has_continuous_smul_induced_of_image_bounded + (h𝔖₁ : 𝔖.nonempty) (h𝔖₂ : directed_on (⊆) 𝔖) + (φ : hom) (hφ : @inducing _ _ _ (uniform_convergence_on.topological_space α E 𝔖) φ) + (h : ∀ u : H, ∀ s ∈ 𝔖, bornology.is_vonN_bounded 𝕜 ((φ u : α → E) '' s)) : + has_continuous_smul 𝕜 H := +begin + letI : uniform_space (α → E) := uniform_convergence_on.uniform_space α E 𝔖, + haveI : uniform_add_group (α → E) := uniform_convergence_on.uniform_add_group, + haveI : topological_add_group H, + { rw hφ.induced, + exact topological_add_group_induced φ }, + have : (𝓝 0 : filter H).has_basis _ _, + { rw [hφ.induced, nhds_induced, map_zero], + exact ((uniform_convergence_on.has_basis_nhds_zero 𝔖 h𝔖₁ h𝔖₂).comap φ) }, + refine has_continuous_smul.of_basis_zero this _ _ _, + { rintros ⟨S, V⟩ ⟨hS, hV⟩, + have : tendsto (λ kx : (𝕜 × E), kx.1 • kx.2) (𝓝 (0, 0)) (𝓝 $ (0 : 𝕜) • 0) := + continuous_smul.tendsto (0 : 𝕜 × E), + rw [zero_smul, nhds_prod_eq] at this, + have := this hV, + rw [mem_map, mem_prod_iff] at this, + rcases this with ⟨U, hU, W, hW, hUW⟩, + refine ⟨U, hU, ⟨S, W⟩, ⟨hS, hW⟩, _⟩, + rw set.smul_subset_iff, + intros a ha u hu x hx, + rw smul_hom_class.map_smul, + exact hUW (⟨ha, hu x hx⟩ : (a, φ u x) ∈ U ×ˢ W) }, + { rintros a ⟨S, V⟩ ⟨hS, hV⟩, + have : tendsto (λ x : E, a • x) (𝓝 0) (𝓝 $ a • 0) := tendsto_id.const_smul a, + rw [smul_zero] at this, + refine ⟨⟨S, ((•) a) ⁻¹' V⟩, ⟨hS, this hV⟩, λ f hf x hx, _⟩, + rw [smul_hom_class.map_smul], + exact hf x hx }, + { rintros u ⟨S, V⟩ ⟨hS, hV⟩, + rcases h u S hS hV with ⟨r, hrpos, hr⟩, + rw metric.eventually_nhds_iff_ball, + refine ⟨r⁻¹, inv_pos.mpr hrpos, λ a ha x hx, _⟩, + by_cases ha0 : a = 0, + { rw ha0, + simp [mem_of_mem_nhds hV] }, + { rw mem_ball_zero_iff at ha, + rw [smul_hom_class.map_smul, pi.smul_apply], + have : φ u x ∈ a⁻¹ • V, + { have ha0 : 0<∥a∥ := norm_pos_iff.mpr ha0, + refine (hr a⁻¹ _) (set.mem_image_of_mem (φ u) hx), + rw [norm_inv, le_inv hrpos ha0], + exact ha.le }, + rwa set.mem_inv_smul_set_iff₀ ha0 at this } } +end + +/-- Let `E` be a TVS, `𝔖 : set (set α)` and `H` a submodule of `α → E`. If the image of any `S ∈ 𝔖` +by any `u ∈ H` is bounded (in the sense of `bornology.is_vonN_bounded`), then `H`, equipped with +the topology of `𝔖`-convergence, is a TVS. + +If you have a hard time using this lemma, try the one above instead. -/ +lemma uniform_convergence_on.has_continuous_smul_submodule_of_image_bounded + (h𝔖₁ : 𝔖.nonempty) (h𝔖₂ : directed_on (⊆) 𝔖) (H : submodule 𝕜 (α → E)) + (h : ∀ u ∈ H, ∀ s ∈ 𝔖, bornology.is_vonN_bounded 𝕜 (u '' s)) : + @has_continuous_smul 𝕜 H _ _ + ((uniform_convergence_on.topological_space α E 𝔖).induced (coe : H → α → E)) := +begin + letI : uniform_space (α → E) := uniform_convergence_on.uniform_space α E 𝔖, + haveI : uniform_add_group (α → E) := uniform_convergence_on.uniform_add_group, + haveI : topological_add_group H := topological_add_group_induced + (linear_map.id.dom_restrict H : H →ₗ[𝕜] α → E), + exact uniform_convergence_on.has_continuous_smul_induced_of_image_bounded 𝕜 α E H h𝔖₁ h𝔖₂ + (linear_map.id.dom_restrict H : H →ₗ[𝕜] α → E) inducing_coe (λ ⟨u, hu⟩, h u hu) +end + +end module From 689184f748be74e4e7757e98d78838b47ec36313 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Wed, 19 Oct 2022 04:48:08 +0000 Subject: [PATCH 811/828] feat(algebraic_geometry/prime_spectrum/noetherian): The spectrum of a noetherian ring is noetherian. (#17016) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> Co-authored-by: Junyan Xu --- .../prime_spectrum/basic.lean | 23 ++++++++++++++++++- .../prime_spectrum/noetherian.lean | 13 ++++++++++- 2 files changed, 34 insertions(+), 2 deletions(-) diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index 54352ff14e92e..457091180a66c 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -9,7 +9,7 @@ import ring_theory.nilpotent import ring_theory.localization.away import ring_theory.ideal.prod import ring_theory.ideal.over -import topology.sets.opens +import topology.sets.closeds import topology.sober /-! @@ -405,6 +405,27 @@ lemma vanishing_ideal_closure (t : set (prime_spectrum R)) : vanishing_ideal (closure t) = vanishing_ideal t := zero_locus_vanishing_ideal_eq_closure t ▸ (gc R).u_l_u_eq_u t +lemma vanishing_ideal_anti_mono_iff {s t : set (prime_spectrum R)} + (ht : is_closed t) : s ⊆ t ↔ vanishing_ideal t ≤ vanishing_ideal s := +⟨vanishing_ideal_anti_mono, λ h, +begin + rw [← ht.closure_subset_iff, ← ht.closure_eq], + convert ← zero_locus_anti_mono_ideal h; + apply zero_locus_vanishing_ideal_eq_closure, +end⟩ + +lemma vanishing_ideal_strict_anti_mono_iff {s t : set (prime_spectrum R)} + (hs : is_closed s) (ht : is_closed t) : + s ⊂ t ↔ vanishing_ideal t < vanishing_ideal s := +by rw [set.ssubset_def, vanishing_ideal_anti_mono_iff hs, + vanishing_ideal_anti_mono_iff ht, lt_iff_le_not_le] + +/-- The antitone order embedding of closed subsets of `Spec R` into ideals of `R`. -/ +def closeds_embedding (R : Type*) [comm_ring R] : + (topological_space.closeds $ prime_spectrum R)ᵒᵈ ↪o ideal R := +order_embedding.of_map_le_iff (λ s, vanishing_ideal s.of_dual) + (λ s t, (vanishing_ideal_anti_mono_iff s.2).symm) + lemma t1_space_iff_is_field [is_domain R] : t1_space (prime_spectrum R) ↔ is_field R := begin diff --git a/src/algebraic_geometry/prime_spectrum/noetherian.lean b/src/algebraic_geometry/prime_spectrum/noetherian.lean index c9abaabee3c72..9da9ef197e557 100644 --- a/src/algebraic_geometry/prime_spectrum/noetherian.lean +++ b/src/algebraic_geometry/prime_spectrum/noetherian.lean @@ -1,9 +1,10 @@ /- Copyright (c) 2020 Filippo A. E. Nuccio. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Filippo A. E. Nuccio +Authors: Filippo A. E. Nuccio, Andrew Yang -/ import algebraic_geometry.prime_spectrum.basic +import topology.noetherian_space /-! This file proves additional properties of the prime spectrum a ring is Noetherian. -/ @@ -90,4 +91,14 @@ begin { rintro (hx | hy); contradiction }, end +open topological_space + +instance : noetherian_space (prime_spectrum R) := +begin + rw (noetherian_space_tfae $ prime_spectrum R).out 0 1, + have H := ‹is_noetherian_ring R›, + rw [is_noetherian_ring_iff, is_noetherian_iff_well_founded] at H, + exact (closeds_embedding R).dual.well_founded H +end + end prime_spectrum From 056a7a1a7e91c32c982097098d31f21ee6a89c3f Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 19 Oct 2022 04:48:09 +0000 Subject: [PATCH 812/828] feat(data/multiset/interval): locally finite order on multisets (#17028) --- src/data/dfinsupp/multiset.lean | 86 +++++++++++++++++++++++++++ src/data/multiset/basic.lean | 23 +++++-- src/data/multiset/interval.lean | 67 +++++++++++++++++++++ src/data/multiset/locally_finite.lean | 2 + 4 files changed, 172 insertions(+), 6 deletions(-) create mode 100644 src/data/dfinsupp/multiset.lean create mode 100644 src/data/multiset/interval.lean diff --git a/src/data/dfinsupp/multiset.lean b/src/data/dfinsupp/multiset.lean new file mode 100644 index 0000000000000..3d1da56331209 --- /dev/null +++ b/src/data/dfinsupp/multiset.lean @@ -0,0 +1,86 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import data.dfinsupp.order + +/-! +# Equivalence between `multiset` and `ℕ`-valued finitely supported functions + +This defines `dfinsupp.to_multiset` the equivalence between `Π₀ a : α, ℕ` and `multiset α`, along +with `multiset.to_dfinsupp` the reverse equivalence. + +Note that this provides a computable alternative to `finsupp.to_multiset`. +-/ + +variables {α : Type*} {β : α → Type*} + +namespace dfinsupp + +/-- Non-dependent special case of `dfinsupp.add_zero_class` to help typeclass search. -/ +instance add_zero_class' {β} [add_zero_class β] : add_zero_class (Π₀ a : α, β) := +@dfinsupp.add_zero_class α (λ _, β) _ + +variables [decidable_eq α] + +/-- A computable version of `finsupp.to_multiset`. -/ +def to_multiset : (Π₀ a : α, ℕ) →+ multiset α := +dfinsupp.sum_add_hom (λ a : α, multiset.repeat_add_monoid_hom a) + +@[simp] lemma to_multiset_single (a : α) (n : ℕ) : + to_multiset (dfinsupp.single a n) = multiset.repeat a n := +dfinsupp.sum_add_hom_single _ _ _ + +end dfinsupp + +namespace multiset +variables [decidable_eq α] + +/-- A computable version of `multiset.to_finsupp` -/ +def to_dfinsupp : multiset α →+ Π₀ a : α, ℕ := +{ to_fun := λ s, + { to_fun := λ n, s.count n, + support' := trunc.mk ⟨s, λ i, (em (i ∈ s)).imp_right multiset.count_eq_zero_of_not_mem⟩ }, + map_zero' := rfl, + map_add' := λ s t, dfinsupp.ext $ λ _, multiset.count_add _ _ _ } + +@[simp] lemma to_dfinsupp_apply (s : multiset α) (a : α) : + s.to_dfinsupp a = s.count a := rfl + +@[simp] lemma to_dfinsupp_support (s : multiset α) : + s.to_dfinsupp.support = s.to_finset := +(finset.filter_eq_self _).mpr (λ x hx, count_ne_zero.mpr $ multiset.mem_to_finset.1 hx) + +@[simp] lemma to_dfinsupp_repeat (a : α) (n : ℕ) : + to_dfinsupp (multiset.repeat a n) = dfinsupp.single a n := +begin + ext i, + dsimp [to_dfinsupp], + simp [count_repeat, eq_comm], +end + +@[simp] lemma to_dfinsupp_singleton (a : α) : to_dfinsupp {a} = dfinsupp.single a 1 := +by rw [←repeat_one, to_dfinsupp_repeat] + +/-- `multiset.to_dfinsupp` as an `add_equiv`. -/ +@[simps apply symm_apply] +def equiv_dfinsupp : multiset α ≃+ Π₀ a : α, ℕ := +add_monoid_hom.to_add_equiv + multiset.to_dfinsupp + dfinsupp.to_multiset + (by { ext x : 1, simp }) + (by { refine @dfinsupp.add_hom_ext α (λ _, ℕ) _ _ _ _ _ _ (λ i hi, _), simp, }) + +@[simp] lemma to_dfinsupp_to_multiset (s : multiset α) : s.to_dfinsupp.to_multiset = s := +equiv_dfinsupp.symm_apply_apply s + +@[simp] lemma to_dfinsupp_le_to_dfinsupp (s t : multiset α) : + to_dfinsupp s ≤ to_dfinsupp t ↔ s ≤ t := +by simp [multiset.le_iff_count, dfinsupp.le_def] + +end multiset + +@[simp] lemma dfinsupp.to_multiset_to_dfinsupp [decidable_eq α] (f : Π₀ a : α, ℕ) : + f.to_multiset.to_dfinsupp = f := +multiset.equiv_dfinsupp.apply_symm_apply f diff --git a/src/data/multiset/basic.lean b/src/data/multiset/basic.lean index 7c52815e8b238..3b168b4532024 100644 --- a/src/data/multiset/basic.lean +++ b/src/data/multiset/basic.lean @@ -616,6 +616,13 @@ def repeat (a : α) (n : ℕ) : multiset α := repeat a n @[simp] lemma repeat_succ (a : α) (n) : repeat a (n+1) = a ::ₘ repeat a n := by simp [repeat] +lemma repeat_add (a : α) (m n : ℕ) : repeat a (m + n) = repeat a m + repeat a n := +congr_arg _ $ list.repeat_add _ _ _ + +/-- `multiset.repeat` as an `add_monoid_hom`. -/ +@[simps] def repeat_add_monoid_hom (a : α) : ℕ →+ multiset α := +{ to_fun := repeat a, map_zero' := repeat_zero a, map_add' := repeat_add a } + @[simp] lemma repeat_one (a : α) : repeat a 1 = {a} := by simp only [repeat_succ, ←cons_zero, eq_self_iff_true, repeat_zero, cons_inj_right] @@ -657,12 +664,7 @@ begin end lemma nsmul_repeat {a : α} (n m : ℕ) : n • (repeat a m) = repeat a (n * m) := -begin - rw eq_repeat, - split, - { rw [card_nsmul, card_repeat] }, - { exact λ b hb, eq_of_mem_repeat (mem_of_mem_nsmul hb) }, -end +((repeat_add_monoid_hom a).map_nsmul _ _).symm /-! ### Erasing one copy of an element -/ section erase @@ -1911,6 +1913,15 @@ by rw [inter_comm, repeat_inter, min_comm] end +@[ext] +lemma add_hom_ext [add_zero_class β] ⦃f g : multiset α →+ β⦄ (h : ∀ x, f {x} = g {x}) : f = g := +begin + ext s, + induction s using multiset.induction_on with a s ih, + { simp only [_root_.map_zero] }, + { simp only [←singleton_add, _root_.map_add, ih, h] } +end + section embedding @[simp] lemma map_le_map_iff {f : α → β} (hf : function.injective f) {s t : multiset α} : diff --git a/src/data/multiset/interval.lean b/src/data/multiset/interval.lean new file mode 100644 index 0000000000000..af501ed5fd1bd --- /dev/null +++ b/src/data/multiset/interval.lean @@ -0,0 +1,67 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import data.finset.locally_finite +import data.dfinsupp.interval +import data.dfinsupp.multiset +import data.nat.interval + +/-! +# Finite intervals of multisets + +This file provides the `locally_finite_order` instance for `multiset α` and calculates the +cardinality of its finite intervals. + +## Implementation notes + +We implement the intervals via the intervals on `dfinsupp`, rather than via filtering +`multiset.powerset`; this is because `(multiset.repeat x n).powerset` has `2^n` entries not `n+1` +entries as it contains duplicates. We do not go via `finsupp` as this would be noncomputable, and +multisets are typically used computationally. + +-/ + +open finset dfinsupp function +open_locale big_operators pointwise + +variables {α : Type*} {β : α → Type*} + +namespace multiset + +variables [decidable_eq α] (f g : multiset α) + +instance : locally_finite_order (multiset α) := +locally_finite_order.of_Icc (multiset α) + (λ f g, (finset.Icc f.to_dfinsupp g.to_dfinsupp).map + (multiset.equiv_dfinsupp.to_equiv.symm.to_embedding)) + (λ f g x, by simp) + +lemma Icc_eq : + finset.Icc f g = + (finset.Icc f.to_dfinsupp g.to_dfinsupp).map + (multiset.equiv_dfinsupp.to_equiv.symm.to_embedding) := rfl + +lemma card_Icc : + (finset.Icc f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) := +by simp_rw [Icc_eq, finset.card_map, dfinsupp.card_Icc, nat.card_Icc, multiset.to_dfinsupp_apply, + to_dfinsupp_support] + +lemma card_Ico : + (finset.Ico f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) - 1 := +by rw [card_Ico_eq_card_Icc_sub_one, card_Icc] + +lemma card_Ioc : + (finset.Ioc f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) - 1 := +by rw [card_Ioc_eq_card_Icc_sub_one, card_Icc] + +lemma card_Ioo : + (finset.Ioo f g).card = ∏ i in f.to_finset ∪ g.to_finset, (g.count i + 1 - f.count i) - 2 := +by rw [card_Ioo_eq_card_Icc_sub_two, card_Icc] + +lemma card_Iic : + (finset.Iic f).card = ∏ i in f.to_finset, (f.count i + 1) := +by simp_rw [Iic_eq_Icc, card_Icc, bot_eq_zero, to_finset_zero, empty_union, count_zero, tsub_zero] + +end multiset diff --git a/src/data/multiset/locally_finite.lean b/src/data/multiset/locally_finite.lean index 8873513dbf223..988094e14a917 100644 --- a/src/data/multiset/locally_finite.lean +++ b/src/data/multiset/locally_finite.lean @@ -10,6 +10,8 @@ import data.finset.locally_finite This file provides basic results about all the `multiset.Ixx`, which are defined in `order.locally_finite`. + +Note that intervals of multisets themselves (`multiset.locally_finite_order`) are defined elsewhere. -/ variables {α : Type*} From 22e33e9de03c12e9dcd56a2010b76836c122fdb3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 19 Oct 2022 04:48:10 +0000 Subject: [PATCH 813/828] refactor(data/finsupp/ne_locus): Match `dist` API (#17036) Rename `ne_locus` lemmas so they match the `dist` one. This is the name of them both being binary functions `f` such that `f a b = g (a - b)` for some `g`. --- src/data/finsupp/ne_locus.lean | 45 ++++++++++++++++++++-------------- 1 file changed, 27 insertions(+), 18 deletions(-) diff --git a/src/data/finsupp/ne_locus.lean b/src/data/finsupp/ne_locus.lean index 2dfb2a4d6d19b..39b84a7643a10 100644 --- a/src/data/finsupp/ne_locus.lean +++ b/src/data/finsupp/ne_locus.lean @@ -88,35 +88,44 @@ by { ext, simpa only [mem_ne_locus] using hF.ne_iff } end ne_locus_and_maps -section cancel_and_group variables [decidable_eq N] -@[simp] lemma add_ne_locus_add_eq_left [add_left_cancel_monoid N] (f g h : α →₀ N) : - (f + g).ne_locus (f + h) = g.ne_locus h := +@[simp] lemma ne_locus_add_left [add_left_cancel_monoid N] (f g h : α →₀ N) : + (f + g).ne_locus (f + h) = g.ne_locus h := zip_with_ne_locus_eq_left _ _ _ _ add_right_injective -@[simp] lemma add_ne_locus_add_eq_right [add_right_cancel_monoid N] (f g h : α →₀ N) : - (f + h).ne_locus (g + h) = f.ne_locus g := +@[simp] lemma ne_locus_add_right [add_right_cancel_monoid N] (f g h : α →₀ N) : + (f + h).ne_locus (g + h) = f.ne_locus g := zip_with_ne_locus_eq_right _ _ _ _ add_left_injective -@[simp] lemma neg_ne_locus_neg [add_group N] (f g : α →₀ N) : (- f).ne_locus (- g) = f.ne_locus g := +section add_group +variables [add_group N] (f f₁ f₂ g g₁ g₂ : α →₀ N) + +@[simp] lemma ne_locus_neg_neg : ne_locus (-f) (-g) = f.ne_locus g := map_range_ne_locus_eq _ _ neg_zero neg_injective -lemma ne_locus_neg [add_group N] (f g : α →₀ N) : (- f).ne_locus g = f.ne_locus (- g) := -by rw [← neg_ne_locus_neg _ _, neg_neg] +lemma ne_locus_neg : ne_locus (-f) g = f.ne_locus (-g) := by rw [←ne_locus_neg_neg, neg_neg] + +lemma ne_locus_eq_support_sub : f.ne_locus g = (f - g).support := +by rw [←ne_locus_add_right _ _ (-g), add_right_neg, ne_locus_zero_right, sub_eq_add_neg] + +@[simp] lemma ne_locus_sub_left : ne_locus (f - g₁) (f - g₂) = ne_locus g₁ g₂ := +by simp only [sub_eq_add_neg, ne_locus_add_left, ne_locus_neg_neg] + +@[simp] lemma ne_locus_sub_right : ne_locus (f₁ - g) (f₂ - g) = ne_locus f₁ f₂ := +by simpa only [sub_eq_add_neg] using ne_locus_add_right _ _ _ -lemma ne_locus_eq_support_sub [add_group N] (f g : α →₀ N) : - f.ne_locus g = (f - g).support := -by rw [← add_ne_locus_add_eq_right _ _ (- g), add_right_neg, ne_locus_zero_right, sub_eq_add_neg] +@[simp] lemma ne_locus_self_add_right : ne_locus f (f + g) = g.support := +by rw [←ne_locus_zero_left, ←ne_locus_add_left f 0 g, add_zero] -@[simp] lemma sub_ne_locus_sub_eq_left [add_group N] (f g h : α →₀ N) : - (f - g).ne_locus (f - h) = g.ne_locus h := -zip_with_ne_locus_eq_left _ _ _ _ (λ fn, sub_right_injective) +@[simp] lemma ne_locus_self_add_left : ne_locus (f + g) f = g.support := +by rw [ne_locus_comm, ne_locus_self_add_right] -@[simp] lemma sub_ne_locus_sub_eq_right [add_group N] (f g h : α →₀ N) : - (f - h).ne_locus (g - h) = f.ne_locus g := -zip_with_ne_locus_eq_right _ _ _ _ (λ hn, sub_left_injective) +@[simp] lemma ne_locus_self_sub_right : ne_locus f (f - g) = g.support := +by rw [sub_eq_add_neg, ne_locus_self_add_right, support_neg] -end cancel_and_group +@[simp] lemma ne_locus_self_sub_left : ne_locus (f - g) f = g.support := +by rw [ne_locus_comm, ne_locus_self_sub_right] +end add_group end finsupp From 0b045c730613b5ef68d4d3799d771e5885c2c8de Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?R=C3=A9my=20Degenne?= Date: Wed, 19 Oct 2022 04:48:12 +0000 Subject: [PATCH 814/828] chore(probability/independence): simplify the definition of pi_Union_Inter (#17043) Due to the simplification, `sup_closed` is no longer useful. Since it was introduced for this application and is not used anywhere else, I removed it. --- src/data/set/lattice.lean | 32 --------- src/measure_theory/pi_system.lean | 96 +++++++++---------------- src/probability/independence.lean | 112 ++++++++---------------------- 3 files changed, 59 insertions(+), 181 deletions(-) diff --git a/src/data/set/lattice.lean b/src/data/set/lattice.lean index 7e145fd807d59..736947280448f 100644 --- a/src/data/set/lattice.lean +++ b/src/data/set/lattice.lean @@ -1835,38 +1835,6 @@ inf_infi_nat_succ u end set -section sup_closed - -/-- A set `s` is sup-closed if for all `x₁, x₂ ∈ s`, `x₁ ⊔ x₂ ∈ s`. -/ -def sup_closed [has_sup α] (s : set α) : Prop := ∀ x1 x2, x1 ∈ s → x2 ∈ s → x1 ⊔ x2 ∈ s - -lemma sup_closed_singleton [semilattice_sup α] (x : α) : sup_closed ({x} : set α) := -λ _ _ y1_mem y2_mem, by { rw set.mem_singleton_iff at *, rw [y1_mem, y2_mem, sup_idem], } - -lemma sup_closed.inter [semilattice_sup α] {s t : set α} (hs : sup_closed s) - (ht : sup_closed t) : - sup_closed (s ∩ t) := -begin - intros x y hx hy, - rw set.mem_inter_iff at hx hy ⊢, - exact ⟨hs x y hx.left hy.left, ht x y hx.right hy.right⟩, -end - -lemma sup_closed_of_totally_ordered [semilattice_sup α] (s : set α) - (hs : ∀ x y : α, x ∈ s → y ∈ s → y ≤ x ∨ x ≤ y) : - sup_closed s := -begin - intros x y hxs hys, - cases hs x y hxs hys, - { rwa (sup_eq_left.mpr h), }, - { rwa (sup_eq_right.mpr h), }, -end - -lemma sup_closed_of_linear_order [linear_order α] (s : set α) : sup_closed s := -sup_closed_of_totally_ordered s (λ x y hxs hys, le_total y x) - -end sup_closed - open set variables [complete_lattice β] diff --git a/src/measure_theory/pi_system.lean b/src/measure_theory/pi_system.lean index b533be024faef..d5879d5894a2d 100644 --- a/src/measure_theory/pi_system.lean +++ b/src/measure_theory/pi_system.lean @@ -43,10 +43,8 @@ import measure_theory.measurable_space_def represented as the intersection of a finite number of elements from these sets. * `pi_Union_Inter` defines a new π-system from a family of π-systems `π : ι → set (set α)` and a - set of finsets `S : set (finset α)`. `pi_Union_Inter π S` is the set of sets that can be written - as `⋂ x ∈ t, f x` for some `t ∈ S` and sets `f x ∈ π x`. If `S` is union-closed, then it is a - π-system. The π-systems used to prove Kolmogorov's 0-1 law will be defined using this mechanism - (TODO). + set of indices `S : set ι`. `pi_Union_Inter π S` is the set of sets that can be written + as `⋂ x ∈ t, f x` for some finset `t ∈ S` and sets `f x ∈ π x`. ## Implementation details @@ -329,25 +327,25 @@ section Union_Inter /-! ### π-system generated by finite intersections of sets of a π-system family -/ -/-- From a set of finsets `S : set (finset ι)` and a family of sets of sets `π : ι → set (set α)`, -define the set of sets that can be written as `⋂ x ∈ t, f x` for some `t ∈ S` and sets `f x ∈ π x`. - -If `S` is union-closed and `π` is a family of π-systems, then it is a π-system. -The π-systems used to prove Kolmogorov's 0-1 law are of that form. -/ -def pi_Union_Inter {α ι} (π : ι → set (set α)) (S : set (finset ι)) : set (set α) := -{s : set α | ∃ (t : finset ι) (htS : t ∈ S) (f : ι → set α) (hf : ∀ x, x ∈ t → f x ∈ π x), +/-- From a set of indices `S : set ι` and a family of sets of sets `π : ι → set (set α)`, +define the set of sets that can be written as `⋂ x ∈ t, f x` for some finset `t ⊆ S` and sets +`f x ∈ π x`. If `π` is a family of π-systems, then it is a π-system. -/ +def pi_Union_Inter {α ι} (π : ι → set (set α)) (S : set ι) : set (set α) := +{s : set α | ∃ (t : finset ι) (htS : ↑t ⊆ S) (f : ι → set α) (hf : ∀ x, x ∈ t → f x ∈ π x), s = ⋂ x ∈ t, f x} /-- If `S` is union-closed and `π` is a family of π-systems, then `pi_Union_Inter π S` is a π-system. -/ lemma is_pi_system_pi_Union_Inter {α ι} (π : ι → set (set α)) - (hpi : ∀ x, is_pi_system (π x)) (S : set (finset ι)) (h_sup : sup_closed S) : + (hpi : ∀ x, is_pi_system (π x)) (S : set ι) : is_pi_system (pi_Union_Inter π S) := begin rintros t1 ⟨p1, hp1S, f1, hf1m, ht1_eq⟩ t2 ⟨p2, hp2S, f2, hf2m, ht2_eq⟩ h_nonempty, simp_rw [pi_Union_Inter, set.mem_set_of_eq], let g := λ n, (ite (n ∈ p1) (f1 n) set.univ) ∩ (ite (n ∈ p2) (f2 n) set.univ), - use [p1 ∪ p2, h_sup p1 p2 hp1S hp2S, g], + have hp_union_ss : ↑(p1 ∪ p2) ⊆ S, + { simp only [hp1S, hp2S, finset.coe_union, union_subset_iff, and_self], }, + use [p1 ∪ p2, hp_union_ss, g], have h_inter_eq : t1 ∩ t2 = ⋂ i ∈ p1 ∪ p2, g i, { rw [ht1_eq, ht2_eq], simp_rw [← set.inf_eq_inter, g], @@ -380,7 +378,7 @@ begin end lemma pi_Union_Inter_mono_left {α ι} {π π' : ι → set (set α)} (h_le : ∀ i, π i ⊆ π' i) - (S : set (finset ι)) : + (S : set ι) : pi_Union_Inter π S ⊆ pi_Union_Inter π' S := begin rintros s ⟨t, ht_mem, ft, hft_mem_pi, rfl⟩, @@ -388,7 +386,7 @@ begin end lemma generate_from_pi_Union_Inter_le {α ι} {m : measurable_space α} - (π : ι → set (set α)) (h : ∀ n, generate_from (π n) ≤ m) (S : set (finset ι)) : + (π : ι → set (set α)) (h : ∀ n, generate_from (π n) ≤ m) (S : set ι) : generate_from (pi_Union_Inter π S) ≤ m := begin refine generate_from_le _, @@ -397,79 +395,47 @@ begin exact measurable_set_generate_from (hft_mem_pi x hx_mem), end -lemma subset_pi_Union_Inter {α ι} {π : ι → set (set α)} {S : set (finset ι)} - (h_univ : ∀ i, set.univ ∈ π i) {i : ι} {s : finset ι} (hsS : s ∈ S) (his : i ∈ s) : +lemma subset_pi_Union_Inter {α ι} {π : ι → set (set α)} {S : set ι} {i : ι} (his : i ∈ S) : π i ⊆ pi_Union_Inter π S := begin - refine λ t ht_pii, ⟨s, hsS, (λ j, ite (j = i) t set.univ), ⟨λ m h_pm, _, _⟩⟩, - { split_ifs, - { rwa h, }, - { exact h_univ m, }, }, - { ext1 x, - simp_rw set.mem_Inter, - split; intro hx, - { intros j h_p_j, - split_ifs, - { exact hx, }, - { exact set.mem_univ _, }, }, - { simpa using hx i his, }, }, + refine λ t ht_pii, ⟨{i}, _, (λ j, t), ⟨λ m h_pm, _, _⟩⟩, + { simp only [his, finset.coe_singleton, singleton_subset_iff], }, + { rw finset.mem_singleton at h_pm, + rwa h_pm, }, + { simp only [finset.mem_singleton, Inter_Inter_eq_left], }, end lemma mem_pi_Union_Inter_of_measurable_set {α ι} (m : ι → measurable_space α) - {S : set (finset ι)} {i : ι} {t : finset ι} (htS : t ∈ S) (hit : i ∈ t) (s : set α) + {S : set ι} {i : ι} (hiS : i ∈ S) (s : set α) (hs : measurable_set[m i] s) : s ∈ pi_Union_Inter (λ n, {s | measurable_set[m n] s}) S := -subset_pi_Union_Inter (λ i, measurable_set.univ) htS hit hs +subset_pi_Union_Inter hiS hs lemma le_generate_from_pi_Union_Inter {α ι} {π : ι → set (set α)} - (S : set (finset ι)) (h_univ : ∀ n, set.univ ∈ π n) {x : ι} - {t : finset ι} (htS : t ∈ S) (hxt : x ∈ t) : + (S : set ι) {x : ι} (hxS : x ∈ S) : generate_from (π x) ≤ generate_from (pi_Union_Inter π S) := -generate_from_mono (subset_pi_Union_Inter h_univ htS hxt) +generate_from_mono (subset_pi_Union_Inter hxS) lemma measurable_set_supr_of_mem_pi_Union_Inter {α ι} (m : ι → measurable_space α) - (S : set (finset ι)) (t : set α) (ht : t ∈ pi_Union_Inter (λ n, {s | measurable_set[m n] s}) S) : - measurable_set[⨆ i (hi : ∃ s ∈ S, i ∈ s), m i] t := + (S : set ι) (t : set α) (ht : t ∈ pi_Union_Inter (λ n, {s | measurable_set[m n] s}) S) : + measurable_set[⨆ i ∈ S, m i] t := begin rcases ht with ⟨pt, hpt, ft, ht_m, rfl⟩, refine pt.measurable_set_bInter (λ i hi, _), - suffices h_le : m i ≤ (⨆ i (hi : ∃ s ∈ S, i ∈ s), m i), from h_le (ft i) (ht_m i hi), - have hi' : ∃ s ∈ S, i ∈ s, from ⟨pt, hpt, hi⟩, + suffices h_le : m i ≤ (⨆ i ∈ S, m i), from h_le (ft i) (ht_m i hi), + have hi' : i ∈ S := hpt hi, exact le_supr₂ i hi', end -lemma generate_from_pi_Union_Inter_measurable_space {α ι} (m : ι → measurable_space α) - (S : set (finset ι)) : - generate_from (pi_Union_Inter (λ n, {s | measurable_set[m n] s}) S) - = ⨆ i (hi : ∃ p ∈ S, i ∈ p), m i := +lemma generate_from_pi_Union_Inter_measurable_set {α ι} (m : ι → measurable_space α) (S : set ι) : + generate_from (pi_Union_Inter (λ n, {s | measurable_set[m n] s}) S) = ⨆ i ∈ S, m i := begin refine le_antisymm _ _, - { rw ← @generate_from_measurable_set α (⨆ i (hi : ∃ p ∈ S, i ∈ p), m i), + { rw ← @generate_from_measurable_set α (⨆ i ∈ S, m i), exact generate_from_mono (measurable_set_supr_of_mem_pi_Union_Inter m S), }, { refine supr₂_le (λ i hi, _), - rcases hi with ⟨p, hpS, hpi⟩, rw ← @generate_from_measurable_set α (m i), - exact generate_from_mono (mem_pi_Union_Inter_of_measurable_set m hpS hpi), }, -end - -lemma generate_from_pi_Union_Inter_subsets {α ι} (m : ι → measurable_space α) (S : set ι) : - generate_from (pi_Union_Inter (λ i, {t | measurable_set[m i] t}) {t : finset ι | ↑t ⊆ S}) - = ⨆ i ∈ S, m i := -begin - rw generate_from_pi_Union_Inter_measurable_space, - simp only [set.mem_set_of_eq, exists_prop, supr_exists], - congr' 1, - ext1 i, - by_cases hiS : i ∈ S, - { simp only [hiS, csupr_pos], - refine le_antisymm (supr₂_le (λ t ht, le_rfl)) _, - rw le_supr_iff, - intros m' hm', - specialize hm' {i}, - simpa only [hiS, finset.coe_singleton, set.singleton_subset_iff, finset.mem_singleton, - eq_self_iff_true, and_self, csupr_pos] using hm', }, - { simp only [hiS, supr_false, supr₂_eq_bot, and_imp], - exact λ t htS hit, absurd (htS hit) hiS, }, + exact generate_from_mono (mem_pi_Union_Inter_of_measurable_set m hi), }, end end Union_Inter diff --git a/src/probability/independence.lean b/src/probability/independence.lean index 5f8b789c0d436..08ce45aeb11f9 100644 --- a/src/probability/independence.lean +++ b/src/probability/independence.lean @@ -348,22 +348,23 @@ end variables {m0 : measurable_space Ω} {μ : measure Ω} -lemma indep_sets_pi_Union_Inter_of_disjoint [decidable_eq ι] [is_probability_measure μ] - {s : ι → set (set Ω)} {S T : set (finset ι)} - (h_indep : Indep_sets s μ) (hST : ∀ u v, u ∈ S → v ∈ T → disjoint u v) : +lemma indep_sets_pi_Union_Inter_of_disjoint [is_probability_measure μ] + {s : ι → set (set Ω)} {S T : set ι} + (h_indep : Indep_sets s μ) (hST : disjoint S T) : indep_sets (pi_Union_Inter s S) (pi_Union_Inter s T) μ := begin rintros t1 t2 ⟨p1, hp1, f1, ht1_m, ht1_eq⟩ ⟨p2, hp2, f2, ht2_m, ht2_eq⟩, + classical, let g := λ i, ite (i ∈ p1) (f1 i) set.univ ∩ ite (i ∈ p2) (f2 i) set.univ, have h_P_inter : μ (t1 ∩ t2) = ∏ n in p1 ∪ p2, μ (g n), { have hgm : ∀ i ∈ p1 ∪ p2, g i ∈ s i, { intros i hi_mem_union, rw finset.mem_union at hi_mem_union, cases hi_mem_union with hi1 hi2, - { have hi2 : i ∉ p2 := finset.disjoint_left.mp (hST p1 p2 hp1 hp2) hi1, + { have hi2 : i ∉ p2 := λ hip2, set.disjoint_left.mp hST (hp1 hi1) (hp2 hip2), simp_rw [g, if_pos hi1, if_neg hi2, set.inter_univ], exact ht1_m i hi1, }, - { have hi1 : i ∉ p1 := finset.disjoint_right.mp (hST p1 p2 hp1 hp2) hi2, + { have hi1 : i ∉ p1 := λ hip1, set.disjoint_right.mp hST (hp2 hi2) (hp1 hip1), simp_rw [g, if_neg hi1, if_pos hi2, set.univ_inter], exact ht2_m i hi2, }, }, have h_p1_inter_p2 : ((⋂ x ∈ p1, f1 x) ∩ ⋂ x ∈ p2, f2 x) @@ -377,7 +378,7 @@ begin { intro n, simp_rw g, split_ifs, - { exact absurd rfl (finset.disjoint_iff_ne.mp (hST p1 p2 hp1 hp2) n h n h_1), }, + { exact absurd rfl (set.disjoint_iff_forall_ne.mp hST _ (hp1 h) _ (hp2 h_1)), }, all_goals { simp only [measure_univ, one_mul, mul_one, set.inter_univ, set.univ_inter], }, }, simp_rw [h_P_inter, h_μg, finset.prod_mul_distrib, finset.prod_ite_mem (p1 ∪ p2) p1 (λ x, μ (f1 x)), @@ -390,20 +391,12 @@ lemma indep_supr_of_disjoint [is_probability_measure μ] {m : ι → measurable_ indep (⨆ i ∈ S, m i) (⨆ i ∈ T, m i) μ := begin refine indep_sets.indep (supr₂_le (λ i _, h_le i)) (supr₂_le (λ i _, h_le i)) _ _ - (generate_from_pi_Union_Inter_subsets m S).symm - (generate_from_pi_Union_Inter_subsets m T).symm _, - { refine is_pi_system_pi_Union_Inter _ (λ n, @is_pi_system_measurable_set Ω (m n)) _ _, - intros s t hs ht, - simp only [finset.sup_eq_union, set.mem_set_of_eq, finset.coe_union, set.union_subset_iff], - exact ⟨hs, ht⟩, }, - { refine is_pi_system_pi_Union_Inter _ (λ n, @is_pi_system_measurable_set Ω (m n)) _ _, - intros s t hs ht, - simp only [finset.sup_eq_union, set.mem_set_of_eq, finset.coe_union, set.union_subset_iff], - exact ⟨hs, ht⟩, }, + (generate_from_pi_Union_Inter_measurable_set m S).symm + (generate_from_pi_Union_Inter_measurable_set m T).symm _, + { exact is_pi_system_pi_Union_Inter _ (λ n, @is_pi_system_measurable_set Ω (m n)) _, }, + { exact is_pi_system_pi_Union_Inter _ (λ n, @is_pi_system_measurable_set Ω (m n)) _ , }, { classical, - refine indep_sets_pi_Union_Inter_of_disjoint h_indep (λ s t hs ht, _), - rw finset.disjoint_iff_ne, - exact λ i his j hjt, set.disjoint_iff_forall_ne.mp hST i (hs his) j (ht hjt), }, + exact indep_sets_pi_Union_Inter_of_disjoint h_indep hST, }, end lemma indep_supr_of_directed_le {Ω} {m : ι → measurable_space Ω} @@ -444,11 +437,10 @@ indep_supr_of_directed_le h_indep h_le h_le' (directed_of_inf hm) lemma Indep_sets.pi_Union_Inter_singleton {π : ι → set (set Ω)} {a : ι} {S : finset ι} (hp_ind : Indep_sets π μ) (haS : a ∉ S) : - indep_sets (pi_Union_Inter π {S}) (π a) μ := + indep_sets (pi_Union_Inter π S) (π a) μ := begin rintros t1 t2 ⟨s, hs_mem, ft1, hft1_mem, ht1_eq⟩ ht2_mem_pia, - rw set.mem_singleton_iff at hs_mem, - subst hs_mem, + rw [finset.coe_subset] at hs_mem, classical, let f := λ n, ite (n = a) t2 (ite (n ∈ s) (ft1 n) set.univ), have h_f_mem : ∀ n ∈ insert a s, f n ∈ π n, @@ -456,7 +448,7 @@ begin simp_rw f, cases (finset.mem_insert.mp hn_mem_insert) with hn_mem hn_mem, { simp [hn_mem, ht2_mem_pia], }, - { have hn_ne_a : n ≠ a, by { rintro rfl, exact haS hn_mem, }, + { have hn_ne_a : n ≠ a, by { rintro rfl, exact haS (hs_mem hn_mem), }, simp [hn_ne_a, hn_mem, hft1_mem n hn_mem], }, }, have h_f_mem_pi : ∀ n ∈ s, f n ∈ π n, from λ x hxS, h_f_mem x (by simp [hxS]), have h_t1 : t1 = ⋂ n ∈ s, f n, @@ -466,7 +458,7 @@ begin congr' with hns y, simp only [(h_forall n hns).symm], }, intros n hnS, - have hn_ne_a : n ≠ a, by { rintro rfl, exact haS hnS, }, + have hn_ne_a : n ≠ a, by { rintro rfl, exact haS (hs_mem hnS), }, simp_rw [f, if_pos hnS, if_neg hn_ne_a], }, have h_μ_t1 : μ t1 = ∏ n in s, μ (f n), by rw [h_t1, ← hp_ind s h_f_mem_pi], have h_t2 : t2 = f a, by { simp_rw [f], simp, }, @@ -474,88 +466,40 @@ begin { have h_t1_inter_t2 : t1 ∩ t2 = ⋂ n ∈ insert a s, f n, by rw [h_t1, h_t2, finset.set_bInter_insert, set.inter_comm], rw [h_t1_inter_t2, ← hp_ind (insert a s) h_f_mem], }, - rw [h_μ_inter, finset.prod_insert haS, h_t2, mul_comm, h_μ_t1], + have has : a ∉ s := λ has_mem, haS (hs_mem has_mem), + rw [h_μ_inter, finset.prod_insert has, h_t2, mul_comm, h_μ_t1], end -/-- Auxiliary lemma for `Indep_sets.Indep`. -/ -theorem Indep_sets.Indep_aux [is_probability_measure μ] (m : ι → measurable_space Ω) +/-- The measurable space structures generated by independent pi-systems are independent. -/ +theorem Indep_sets.Indep [is_probability_measure μ] (m : ι → measurable_space Ω) (h_le : ∀ i, m i ≤ m0) (π : ι → set (set Ω)) (h_pi : ∀ n, is_pi_system (π n)) - (hp_univ : ∀ i, set.univ ∈ π i) (h_generate : ∀ i, m i = generate_from (π i)) - (h_ind : Indep_sets π μ) : + (h_generate : ∀ i, m i = generate_from (π i)) (h_ind : Indep_sets π μ) : Indep m μ := begin classical, - refine finset.induction (by simp [measure_univ]) _, + refine finset.induction _ _, + { simp only [measure_univ, implies_true_iff, set.Inter_false, set.Inter_univ, finset.prod_empty, + eq_self_iff_true], }, intros a S ha_notin_S h_rec f hf_m, have hf_m_S : ∀ x ∈ S, measurable_set[m x] (f x) := λ x hx, hf_m x (by simp [hx]), rw [finset.set_bInter_insert, finset.prod_insert ha_notin_S, ← h_rec hf_m_S], - let p := pi_Union_Inter π {S}, + let p := pi_Union_Inter π S, set m_p := generate_from p with hS_eq_generate, have h_indep : indep m_p (m a) μ, - { have hp : is_pi_system p := is_pi_system_pi_Union_Inter π h_pi {S} (sup_closed_singleton S), + { have hp : is_pi_system p := is_pi_system_pi_Union_Inter π h_pi S, have h_le' : ∀ i, generate_from (π i) ≤ m0 := λ i, (h_generate i).symm.trans_le (h_le i), - have hm_p : m_p ≤ m0 := generate_from_pi_Union_Inter_le π h_le' {S}, + have hm_p : m_p ≤ m0 := generate_from_pi_Union_Inter_le π h_le' S, exact indep_sets.indep hm_p (h_le a) hp (h_pi a) hS_eq_generate (h_generate a) (h_ind.pi_Union_Inter_singleton ha_notin_S), }, refine h_indep.symm (f a) (⋂ n ∈ S, f n) (hf_m a (finset.mem_insert_self a S)) _, have h_le_p : ∀ i ∈ S, m i ≤ m_p, { intros n hn, rw [hS_eq_generate, h_generate n], - exact le_generate_from_pi_Union_Inter {S} hp_univ (set.mem_singleton _) hn, }, + exact le_generate_from_pi_Union_Inter S hn, }, have h_S_f : ∀ i ∈ S, measurable_set[m_p] (f i) := λ i hi, (h_le_p i hi) (f i) (hf_m_S i hi), exact S.measurable_set_bInter h_S_f, end -/-- The measurable space structures generated by independent pi-systems are independent. -/ -theorem Indep_sets.Indep [is_probability_measure μ] (m : ι → measurable_space Ω) - (h_le : ∀ i, m i ≤ m0) (π : ι → set (set Ω)) (h_pi : ∀ n, is_pi_system (π n)) - (h_generate : ∀ i, m i = generate_from (π i)) (h_ind : Indep_sets π μ) : - Indep m μ := -begin - -- We want to apply `Indep_sets.Indep_aux`, but `π i` does not contain `univ`, hence we replace - -- `π` with a new augmented pi-system `π'`, and prove all hypotheses for that pi-system. - let π' := λ i, insert set.univ (π i), - have h_subset : ∀ i, π i ⊆ π' i := λ i, set.subset_insert _ _, - have h_pi' : ∀ n, is_pi_system (π' n) := λ n, (h_pi n).insert_univ, - have h_univ' : ∀ i, set.univ ∈ π' i, from λ i, set.mem_insert _ _, - have h_gen' : ∀ i, m i = generate_from (π' i), - { intros i, - rw [h_generate i, generate_from_insert_univ (π i)], }, - have h_ind' : Indep_sets π' μ, - { intros S f hfπ', - classical, - let S' := finset.filter (λ i, f i ≠ set.univ) S, - have h_mem : ∀ i ∈ S', f i ∈ π i, - { intros i hi, - simp_rw [S', finset.mem_filter] at hi, - cases hfπ' i hi.1, - { exact absurd h hi.2, }, - { exact h, }, }, - have h_left : (⋂ i ∈ S, f i) = ⋂ i ∈ S', f i, - { ext1 x, - simp only [set.mem_Inter, finset.mem_filter, ne.def, and_imp], - split, - { exact λ h i hiS hif, h i hiS, }, - { intros h i hiS, - by_cases hfi_univ : f i = set.univ, - { rw hfi_univ, exact set.mem_univ _, }, - { exact h i hiS hfi_univ, }, }, }, - have h_right : ∏ i in S, μ (f i) = ∏ i in S', μ (f i), - { rw ← finset.prod_filter_mul_prod_filter_not S (λ i, f i ≠ set.univ), - simp only [ne.def, finset.filter_congr_decidable, not_not], - suffices : ∏ x in finset.filter (λ x, f x = set.univ) S, μ (f x) = 1, - { rw [this, mul_one], }, - calc ∏ x in finset.filter (λ x, f x = set.univ) S, μ (f x) - = ∏ x in finset.filter (λ x, f x = set.univ) S, μ set.univ : - finset.prod_congr rfl (λ x hx, by { rw finset.mem_filter at hx, rw hx.2, }) - ... = ∏ x in finset.filter (λ x, f x = set.univ) S, 1 : - finset.prod_congr rfl (λ _ _, measure_univ) - ... = 1 : finset.prod_const_one, }, - rw [h_left, h_right], - exact h_ind S' h_mem, }, - exact Indep_sets.Indep_aux m h_le π' h_pi' h_univ' h_gen' h_ind', -end - end from_pi_systems_to_measurable_spaces section indep_set From 813725f8aa036240cc9dd1c4b05e700ba97f8f0d Mon Sep 17 00:00:00 2001 From: Paul Lezeau <72892199+Paul-Lez@users.noreply.github.com> Date: Wed, 19 Oct 2022 07:41:22 +0000 Subject: [PATCH 815/828] feat(number_theory/kummer_dedekind): define conductor and add basic API plus theorem (#16833) In this PR we define the conductor ideal and prove some basic results about it. In a later PR (#16870), these will be used to prove the Dedekind-Kummer theorem in full generality. Co-authored-by: Paul Lezeau --- src/number_theory/kummer_dedekind.lean | 157 +++++++++++++++++++++++-- src/ring_theory/ideal/operations.lean | 12 ++ src/ring_theory/ideal/quotient.lean | 9 ++ 3 files changed, 170 insertions(+), 8 deletions(-) diff --git a/src/number_theory/kummer_dedekind.lean b/src/number_theory/kummer_dedekind.lean index 789eb779ba264..04577ebe05292 100644 --- a/src/number_theory/kummer_dedekind.lean +++ b/src/number_theory/kummer_dedekind.lean @@ -35,7 +35,7 @@ with a formula). ## TODO - * Define the conductor ideal and prove the Kummer-Dedekind theorem in full generality. + * Prove the Kummer-Dedekind theorem in full generality. * Prove the converse of `ideal.irreducible_map_of_irreducible_minpoly`. @@ -52,20 +52,161 @@ with a formula). kummer, dedekind, kummer dedekind, dedekind-kummer, dedekind kummer -/ +variables (R : Type*) {S : Type*} [comm_ring R] [comm_ring S] [algebra R S] + +open ideal polynomial double_quot unique_factorization_monoid algebra ring_hom + +local notation R`<`:std.prec.max_plus x `>` := adjoin R ({x} : set S) + +/-- Let `S / R` be a ring extension and `x : S`, then the conductor of `R` is the + biggest ideal of `S` contained in `R`. -/ +def conductor (x : S) : ideal S := +{ carrier := {a | ∀ (b : S), a * b ∈ R }, + zero_mem' := λ b, by simpa only [zero_mul] using subalgebra.zero_mem _, + add_mem' := λ a b ha hb c, by simpa only [add_mul] using subalgebra.add_mem _ (ha c) (hb c), + smul_mem' := λ c a ha b, by simpa only [smul_eq_mul, mul_left_comm, mul_assoc] using ha (c * b) } + +variables {R} {x : S} + +lemma conductor_eq_of_eq {y : S} (h : (R : set S) = R): + conductor R x = conductor R y := +ideal.ext $ λ a, forall_congr $ λ b, set.ext_iff.mp h _ + +lemma conductor_subset_adjoin : (conductor R x : set S) ⊆ R := +λ y hy, by simpa only [mul_one] using hy 1 + +lemma mem_conductor_iff {y : S} : y ∈ conductor R x ↔ ∀ (b : S), y * b ∈ R := +⟨λ h, h, λ h, h⟩ + +variables {I : ideal R} + +/-- This technical lemma tell us that if `C` is the conductor of `R` and `I` is an ideal of `R` + then `p * (I * S) ⊆ I * R` for any `p` in `C ∩ R` -/ +lemma prod_mem_ideal_map_of_mem_conductor {p : R} {z : S} + (hp : p ∈ ideal.comap (algebra_map R S) (conductor R x)) (hz' : z ∈ (I.map (algebra_map R S))) : + (algebra_map R S p) * z ∈ + algebra_map R S '' ↑(I.map (algebra_map R R)) := +begin + rw [ideal.map, ideal.span, finsupp.mem_span_image_iff_total] at hz', + obtain ⟨l, H, H'⟩ := hz', + rw finsupp.total_apply at H', + rw [← H', mul_comm, finsupp.sum_mul], + have lem : ∀ {a : R}, a ∈ I → (l a • (algebra_map R S a) * (algebra_map R S p)) ∈ + (algebra_map R S) '' (I.map (algebra_map R R)), + { intros a ha, + rw [algebra.id.smul_eq_mul, mul_assoc, mul_comm, mul_assoc, set.mem_image], + refine exists.intro (algebra_map R R a * ⟨l a * algebra_map R S p, + show l a * algebra_map R S p ∈ R, from _ ⟩) _, + { rw mul_comm, + exact mem_conductor_iff.mp (ideal.mem_comap.mp hp) _ }, + refine ⟨_, by simpa only [ring_hom.map_mul, mul_comm (algebra_map R S p) (l a)]⟩, + rw mul_comm, + apply ideal.mul_mem_left (I.map (algebra_map R R)) _ + (ideal.mem_map_of_mem _ ha) }, + refine finset.sum_induction _ (λ u, u ∈ (algebra_map R S) '' + (I.map (algebra_map R R))) + (λ a b, _) _ _, + rintro ⟨z, hz, rfl⟩ ⟨y, hy, rfl⟩, + rw [← ring_hom.map_add], + exact ⟨z + y, ideal.add_mem _ (set_like.mem_coe.mp hz) hy, rfl⟩, + { refine ⟨0, set_like.mem_coe.mpr $ ideal.zero_mem _, ring_hom.map_zero _⟩ }, + { intros y hy, + exact lem ((finsupp.mem_supported _ l).mp H hy) }, +end + +/-- A technical result telling us that `(I * S) ∩ R = I * R` for any ideal `I` of `R`. -/ +lemma comap_map_eq_map_adjoin_of_coprime_conductor + (hx : (conductor R x).comap (algebra_map R S) ⊔ I = ⊤) + (h_alg : function.injective (algebra_map R S)): + (I.map (algebra_map R S)).comap (algebra_map R S) = I.map (algebra_map R R) := +begin + apply le_antisymm, + { -- This is adapted from [Neukirch1992]. Let `C = (conductor R x)`. The idea of the proof + -- is that since `I` and `C ∩ R` are coprime, we have + -- `(I * S) ∩ R ⊆ (I + C) * ((I * S) ∩ R) ⊆ I * R + I * C * S ⊆ I * R`. + intros y hy, + obtain ⟨z, hz⟩ := y, + obtain ⟨p, hp, q, hq, hpq⟩ := submodule.mem_sup.mp ((ideal.eq_top_iff_one _).mp hx), + have temp : (algebra_map R S p)*z + (algebra_map R S q)*z = z, + { simp only [←add_mul, ←ring_hom.map_add (algebra_map R S), hpq, map_one, one_mul] }, + suffices : z ∈ algebra_map R S '' (I.map (algebra_map R R)) ↔ (⟨z, hz⟩ : R) ∈ + I.map (algebra_map R R), + { rw [← this, ← temp], + obtain ⟨a, ha⟩ := (set.mem_image _ _ _).mp (prod_mem_ideal_map_of_mem_conductor hp + (show z ∈ I.map (algebra_map R S), by rwa ideal.mem_comap at hy )), + use a + (algebra_map R R q) * ⟨z, hz⟩, + refine ⟨ ideal.add_mem (I.map (algebra_map R R)) ha.left _, + by simpa only [ha.right, map_add, alg_hom.map_mul, add_right_inj] ⟩, + rw mul_comm, + exact ideal.mul_mem_left (I.map (algebra_map R R)) _ (ideal.mem_map_of_mem _ hq) }, + refine ⟨ λ h, _, λ h, (set.mem_image _ _ _).mpr (exists.intro ⟨z, hz⟩ ⟨by simp [h], rfl⟩ ) ⟩, + { obtain ⟨x₁, hx₁, hx₂⟩ := (set.mem_image _ _ _).mp h, + have : x₁ = ⟨z, hz⟩, + { apply h_alg, + simpa [hx₂], }, + rwa ← this } }, + + { -- The converse inclusion is trivial + have : algebra_map R S = (algebra_map _ S).comp (algebra_map R R) := by { ext, refl }, + rw [this, ← ideal.map_map], + apply ideal.le_comap_map } +end + +/-- The canonical morphism of rings from `R ⧸ (I*R)` to `S ⧸ (I*S)` is an isomorphism + when `I` and `(conductor R x) ∩ R` are coprime. -/ +noncomputable def quot_adjoin_equiv_quot_map (hx : (conductor R x).comap (algebra_map R S) ⊔ I = ⊤) + (h_alg : function.injective (algebra_map R S)) : + R ⧸ (I.map (algebra_map R R)) ≃+* S ⧸ (I.map (algebra_map R S)) := +ring_equiv.of_bijective (ideal.quotient.lift (I.map (algebra_map R R)) + (((I.map (algebra_map R S))^.quotient.mk).comp (algebra_map R S )) (λ r hr, + begin + have : algebra_map R S = (algebra_map R S).comp + (algebra_map R R) := by { ext, refl }, + rw [ring_hom.comp_apply, ideal.quotient.eq_zero_iff_mem, this, ← ideal.map_map], + exact ideal.mem_map_of_mem _ hr + end)) +begin + split, + { --the kernel of the map is clearly `(I * S) ∩ R`. To get injectivity, we need to show that + --this is contained in `I * R`, which is the content of the previous lemma. + refine ring_hom.lift_injective_of_ker_le_ideal _ _ (λ u hu, _), + rwa [ring_hom.mem_ker, ring_hom.comp_apply, ideal.quotient.eq_zero_iff_mem, + ← ideal.mem_comap, comap_map_eq_map_adjoin_of_coprime_conductor hx h_alg] at hu }, + { -- Surjectivity follows from the surjectivity of the canonical map `R → S ⧸ (I * S)`, + -- which in turn follows from the fact that `I * S + (conductor R x) = S`. + refine ideal.quotient.lift_surjective_of_surjective _ _ (λ y, _), + obtain ⟨z, hz⟩ := ideal.quotient.mk_surjective y, + have : z ∈ conductor R x ⊔ (I.map (algebra_map R S)), + { suffices : conductor R x ⊔ (I.map (algebra_map R S)) = ⊤, + { simp only [this] }, + rw ideal.eq_top_iff_one at hx ⊢, + replace hx := ideal.mem_map_of_mem (algebra_map R S) hx, + rw [ideal.map_sup, ring_hom.map_one] at hx, + exact (sup_le_sup (show ((conductor R x).comap (algebra_map R S)).map (algebra_map R S) ≤ + conductor R x, from ideal.map_comap_le) (le_refl (I.map (algebra_map R S)))) hx }, + rw [← ideal.mem_quotient_iff_mem_sup, hz, ideal.mem_map_iff_of_surjective] at this, + obtain ⟨u, hu, hu'⟩ := this, + use ⟨u, conductor_subset_adjoin hu⟩, + simpa only [← hu'], + { exact ideal.quotient.mk_surjective } } +end + +@[simp] +lemma quot_adjoin_equiv_quot_map_apply_mk (hx : (conductor R x).comap (algebra_map R S) ⊔ I = ⊤) + (h_alg : function.injective (algebra_map R S)) (a : R) : + quot_adjoin_equiv_quot_map hx h_alg ((I.map (algebra_map R R))^.quotient.mk a) + = (I.map (algebra_map R S))^.quotient.mk ↑a := +rfl + namespace kummer_dedekind open_locale big_operators polynomial classical -open ideal polynomial double_quot unique_factorization_monoid - -variables {R : Type*} [comm_ring R] -variables {S : Type*} [comm_ring S] [is_domain S] [is_dedekind_domain S] [algebra R S] -variables (pb : power_basis R S) {I : ideal R} +variables [is_domain R] [is_domain S] [is_dedekind_domain S] +variables (pb : power_basis R S) local attribute [instance] ideal.quotient.field -variables [is_domain R] - /-- The first half of the **Kummer-Dedekind Theorem** in the monogenic case, stating that the prime factors of `I*S` are in bijection with those of the minimal polynomial of the generator of `S` over `R`, taken `mod I`.-/ diff --git a/src/ring_theory/ideal/operations.lean b/src/ring_theory/ideal/operations.lean index c2a4a27542b04..178191d1facdb 100644 --- a/src/ring_theory/ideal/operations.lean +++ b/src/ring_theory/ideal/operations.lean @@ -1598,6 +1598,18 @@ assume a b, quotient.induction_on₂' a b $ assume a b (h : f a = f b), ideal.quotient.eq.2 $ show a - b ∈ ker f, by rw [mem_ker, map_sub, h, sub_self] +lemma lift_injective_of_ker_le_ideal (I : ideal R) {f : R →+* S} + (H : ∀ (a : R), a ∈ I → f a = 0) (hI : f.ker ≤ I) : + function.injective (ideal.quotient.lift I f H) := +begin + rw [ring_hom.injective_iff_ker_eq_bot, ring_hom.ker_eq_bot_iff_eq_zero], + intros u hu, + obtain ⟨v, rfl⟩ := ideal.quotient.mk_surjective u, + rw ideal.quotient.lift_mk at hu, + rw [ideal.quotient.eq_zero_iff_mem], + exact hI ((ring_hom.mem_ker f).mpr hu), +end + variable {f} /-- The **first isomorphism theorem** for commutative rings, computable version. -/ diff --git a/src/ring_theory/ideal/quotient.lean b/src/ring_theory/ideal/quotient.lean index 8bff80f2e77d5..61e8b459e2d82 100644 --- a/src/ring_theory/ideal/quotient.lean +++ b/src/ring_theory/ideal/quotient.lean @@ -210,6 +210,15 @@ def lift (I : ideal R) (f : R →+* S) (H : ∀ (a : R), a ∈ I → f a = 0) : @[simp] lemma lift_mk (I : ideal R) (f : R →+* S) (H : ∀ (a : R), a ∈ I → f a = 0) : lift I f H (mk I a) = f a := rfl +lemma lift_surjective_of_surjective (I : ideal R) {f : R →+* S} (H : ∀ (a : R), a ∈ I → f a = 0) + (hf : function.surjective f) : function.surjective (ideal.quotient.lift I f H) := +begin + intro y, + obtain ⟨x, rfl⟩ := hf y, + use ideal.quotient.mk I x, + simp only [ideal.quotient.lift_mk], +end + /-- The ring homomorphism from the quotient by a smaller ideal to the quotient by a larger ideal. This is the `ideal.quotient` version of `quot.factor` -/ From 8bf24777564e29dc7c30170f49f810dde95ac147 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 19 Oct 2022 09:54:19 +0000 Subject: [PATCH 816/828] feat(analysis/asymptotics): prove that a function has derivative 0 using asymptotics (#16934) * From the sphere eversion project * Needed for computing derivatives of the parametric interval integral Co-authored by: Patrick Massot [patrickmassot@free.fr](patrickmassot@free.fr) --- src/analysis/asymptotics/asymptotics.lean | 33 +++++++++++++++++++++++ src/analysis/calculus/fderiv.lean | 32 +++++++++++++++++++--- src/analysis/normed/field/basic.lean | 2 ++ 3 files changed, 64 insertions(+), 3 deletions(-) diff --git a/src/analysis/asymptotics/asymptotics.lean b/src/analysis/asymptotics/asymptotics.lean index 970336579d7cd..3543859939d16 100644 --- a/src/analysis/asymptotics/asymptotics.lean +++ b/src/analysis/asymptotics/asymptotics.lean @@ -101,6 +101,9 @@ by simp only [is_O, is_O_with] lemma is_O.of_bound (c : ℝ) (h : ∀ᶠ x in l, ∥f x∥ ≤ c * ∥g x∥) : f =O[l] g := is_O_iff.2 ⟨c, h⟩ +lemma is_O.of_bound' (h : ∀ᶠ x in l, ∥f x∥ ≤ ∥g x∥) : f =O[l] g := +is_O.of_bound 1 $ by { simp_rw one_mul, exact h } + lemma is_O.bound : f =O[l] g → ∃ c : ℝ, ∀ᶠ x in l, ∥f x∥ ≤ c * ∥g x∥ := is_O_iff.1 /-- The Landau notation `f =o[l] g` where `f` and `g` are two functions on a type `α` and `l` is @@ -382,6 +385,14 @@ let ⟨c, cpos, hc⟩ := hfg.exists_pos in hc.trans_is_o hgk cpos (hgk : g =o[l] k) : f =o[l] k := hfg.trans_is_O_with hgk.is_O_with one_pos +lemma _root_.filter.eventually.trans_is_O {f : α → E} {g : α → F'} {k : α → G} + (hfg : ∀ᶠ x in l, ∥f x∥ ≤ ∥g x∥) (hgk : g =O[l] k) : f =O[l] k := +(is_O.of_bound' hfg).trans hgk + +lemma _root_.filter.eventually.is_O {f : α → E} {g : α → ℝ} {l : filter α} + (hfg : ∀ᶠ x in l, ∥f x∥ ≤ g x) : f =O[l] g := +is_O.of_bound' $ hfg.mono $ λ x hx, hx.trans $ real.le_norm_self _ + section variable (l) @@ -1547,6 +1558,28 @@ theorem is_o_norm_pow_id {n : ℕ} (h : 1 < n) : (λ x : E', ∥x∥^n) =o[𝓝 0] (λ x, x) := by simpa only [pow_one, is_o_norm_right] using @is_o_norm_pow_norm_pow E' _ _ _ h +lemma is_O.eq_zero_of_norm_pow_within {f : E'' → F''} {s : set E''} {x₀ : E''} {n : ℕ} + (h : f =O[𝓝[s] x₀] λ x, ∥x - x₀∥ ^ n) (hx₀ : x₀ ∈ s) (hn : 0 < n) : f x₀ = 0 := +mem_of_mem_nhds_within hx₀ h.eq_zero_imp $ by simp_rw [sub_self, norm_zero, zero_pow hn] + +lemma is_O.eq_zero_of_norm_pow {f : E'' → F''} {x₀ : E''} {n : ℕ} + (h : f =O[𝓝 x₀] λ x, ∥x - x₀∥ ^ n) (hn : 0 < n) : f x₀ = 0 := +by { rw [← nhds_within_univ] at h, exact h.eq_zero_of_norm_pow_within (mem_univ _) hn } + +lemma is_o_pow_sub_pow_sub (x₀ : E') {n m : ℕ} (h : n < m) : + (λ x, ∥x - x₀∥ ^ m) =o[𝓝 x₀] λ x, ∥x - x₀∥^n := +begin + have : tendsto (λ x, ∥x - x₀∥) (𝓝 x₀) (𝓝 0), + { apply tendsto_norm_zero.comp, + rw ← sub_self x₀, + exact tendsto_id.sub tendsto_const_nhds }, + exact (is_o_pow_pow h).comp_tendsto this +end + +lemma is_o_pow_sub_sub (x₀ : E') {m : ℕ} (h : 1 < m) : + (λ x, ∥x - x₀∥^m) =o[𝓝 x₀] λ x, x - x₀ := +by simpa only [is_o_norm_right, pow_one] using is_o_pow_sub_pow_sub x₀ h + theorem is_O_with.right_le_sub_of_lt_1 {f₁ f₂ : α → E'} (h : is_O_with c l f₁ f₂) (hc : c < 1) : is_O_with (1 / (1 - c)) l f₂ (λx, f₂ x - f₁ x) := is_O_with.of_bound $ mem_of_superset h.bound $ λ x hx, diff --git a/src/analysis/calculus/fderiv.lean b/src/analysis/calculus/fderiv.lean index 4b4525da1bebc..90fa617b7b12f 100644 --- a/src/analysis/calculus/fderiv.lean +++ b/src/analysis/calculus/fderiv.lean @@ -346,6 +346,8 @@ lemma has_fderiv_at.differentiable_at (h : has_fderiv_at f f' x) : differentiabl has_fderiv_within_at f f' univ x ↔ has_fderiv_at f f' x := by { simp only [has_fderiv_within_at, nhds_within_univ], refl } +alias has_fderiv_within_at_univ ↔ has_fderiv_within_at.has_fderiv_at_of_univ _ + lemma has_fderiv_within_at_insert {y : E} {g' : E →L[𝕜] F} : has_fderiv_within_at g g' (insert y s) x ↔ has_fderiv_within_at g g' s x := begin @@ -406,7 +408,7 @@ lemma has_fderiv_at.lim (hf : has_fderiv_at f f' x) (v : E) {α : Type*} {c : α {l : filter α} (hc : tendsto (λ n, ∥c n∥) l at_top) : tendsto (λ n, (c n) • (f (x + (c n)⁻¹ • v) - f x)) l (𝓝 (f' v)) := begin - refine (has_fderiv_within_at_univ.2 hf).lim _ (univ_mem' (λ _, trivial)) hc _, + refine (has_fderiv_within_at_univ.2 hf).lim _ univ_mem hc _, assume U hU, refine (eventually_ne_of_tendsto_norm_at_top hc (0:𝕜)).mono (λ y hy, _), convert mem_of_mem_nhds hU, @@ -646,6 +648,30 @@ lemma fderiv_within_mem_iff {f : E → F} {t : set E} {s : set (E →L[𝕜] F)} by by_cases hx : differentiable_within_at 𝕜 f t x; simp [fderiv_within_zero_of_not_differentiable_within_at, *] +lemma asymptotics.is_O.has_fderiv_within_at {s : set E} {x₀ : E} {n : ℕ} + (h : f =O[𝓝[s] x₀] λ x, ∥x - x₀∥^n) (hx₀ : x₀ ∈ s) (hn : 1 < n) : + has_fderiv_within_at f (0 : E →L[𝕜] F) s x₀ := +by simp_rw [has_fderiv_within_at, has_fderiv_at_filter, + h.eq_zero_of_norm_pow_within hx₀ $ zero_lt_one.trans hn, zero_apply, sub_zero, + h.trans_is_o ((is_o_pow_sub_sub x₀ hn).mono nhds_within_le_nhds)] + +lemma asymptotics.is_O.has_fderiv_at {x₀ : E} {n : ℕ} + (h : f =O[𝓝 x₀] λ x, ∥x - x₀∥^n) (hn : 1 < n) : + has_fderiv_at f (0 : E →L[𝕜] F) x₀ := +begin + rw [← nhds_within_univ] at h, + exact (h.has_fderiv_within_at (mem_univ _) hn).has_fderiv_at_of_univ +end + +lemma has_fderiv_within_at.is_O {f : E → F} {s : set E} {x₀ : E} {f' : E →L[𝕜] F} + (h : has_fderiv_within_at f f' s x₀) : + (λ x, f x - f x₀) =O[𝓝[s] x₀] λ x, x - x₀ := +by simpa only [sub_add_cancel] using h.is_O.add (is_O_sub f' (𝓝[s] x₀) x₀) + +lemma has_fderiv_at.is_O {f : E → F} {x₀ : E} {f' : E →L[𝕜] F} (h : has_fderiv_at f f' x₀) : + (λ x, f x - f x₀) =O[𝓝 x₀] λ x, x - x₀ := +by simpa only [sub_add_cancel] using h.is_O.add (is_O_sub f' (𝓝 x₀) x₀) + end fderiv_properties section continuous @@ -2699,7 +2725,7 @@ end lemma comp_has_fderiv_at_iff {f : G → E} {x : G} {f' : G →L[𝕜] E} : has_fderiv_at (iso ∘ f) ((iso : E →L[𝕜] F).comp f') x ↔ has_fderiv_at f f' x := -by rw [← has_fderiv_within_at_univ, ← has_fderiv_within_at_univ, iso.comp_has_fderiv_within_at_iff] +by simp_rw [← has_fderiv_within_at_univ, iso.comp_has_fderiv_within_at_iff] lemma comp_has_fderiv_within_at_iff' {f : G → E} {s : set G} {x : G} {f' : G →L[𝕜] F} : @@ -2710,7 +2736,7 @@ by rw [← iso.comp_has_fderiv_within_at_iff, ← continuous_linear_map.comp_ass lemma comp_has_fderiv_at_iff' {f : G → E} {x : G} {f' : G →L[𝕜] F} : has_fderiv_at (iso ∘ f) f' x ↔ has_fderiv_at f ((iso.symm : F →L[𝕜] E).comp f') x := -by rw [← has_fderiv_within_at_univ, ← has_fderiv_within_at_univ, iso.comp_has_fderiv_within_at_iff'] +by simp_rw [← has_fderiv_within_at_univ, iso.comp_has_fderiv_within_at_iff'] lemma comp_fderiv_within {f : G → E} {s : set G} {x : G} (hxs : unique_diff_within_at 𝕜 s x) : diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index fd035c2cbe88e..7920b8e0fb064 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -638,6 +638,8 @@ abs_of_nonneg hx lemma norm_of_nonpos {x : ℝ} (hx : x ≤ 0) : ∥x∥ = -x := abs_of_nonpos hx +lemma le_norm_self (r : ℝ) : r ≤ ∥r∥ := le_abs_self r + @[simp] lemma norm_coe_nat (n : ℕ) : ∥(n : ℝ)∥ = n := abs_of_nonneg n.cast_nonneg @[simp] lemma nnnorm_coe_nat (n : ℕ) : ∥(n : ℝ)∥₊ = n := nnreal.eq $ by simp From 91b330171dee505a29a0820fe924e2883b031b60 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Wed, 19 Oct 2022 09:54:21 +0000 Subject: [PATCH 817/828] feat(geometry/manifold/cont_mdiff): add various properties (#16980) * More lemmas about `smooth`, more basic lemmas, and some interactions with `cont_diff` and continuous linear maps * Also add various useful lemmas throughout the manifold directory * Add me as author to two files (in `local_invariant_properties` for earlier refactoring PRs) * From the sphere eversion project --- src/geometry/manifold/charted_space.lean | 7 + src/geometry/manifold/cont_mdiff.lean | 130 +++++++++++++++++- src/geometry/manifold/cont_mdiff_mfderiv.lean | 16 +++ .../manifold/local_invariant_properties.lean | 2 +- .../smooth_manifold_with_corners.lean | 17 ++- src/geometry/manifold/tangent_bundle.lean | 38 +++++ src/logic/equiv/local_equiv.lean | 4 + src/topology/local_homeomorph.lean | 5 + 8 files changed, 212 insertions(+), 7 deletions(-) diff --git a/src/geometry/manifold/charted_space.lean b/src/geometry/manifold/charted_space.lean index 10d6c976af42a..c324375db2680 100644 --- a/src/geometry/manifold/charted_space.lean +++ b/src/geometry/manifold/charted_space.lean @@ -456,6 +456,7 @@ The model space is written as an explicit parameter as there can be several mode given topological space. For instance, a complex manifold (modelled over `ℂ^n`) will also be seen sometimes as a real manifold over `ℝ^(2n)`. -/ +@[ext] class charted_space (H : Type*) [topological_space H] (M : Type*) [topological_space M] := (atlas [] : set (local_homeomorph M H)) (chart_at [] : M → local_homeomorph M H) @@ -509,6 +510,9 @@ lemma coe_achart (x : M) : (achart H x : local_homeomorph M H) = chart_at H x := @[simp, mfld_simps] lemma achart_val (x : M) : (achart H x).1 = chart_at H x := rfl +lemma mem_achart_source (x : M) : x ∈ (achart H x).1.source := +mem_chart_source H x + open topological_space lemma charted_space.second_countable_of_countable_cover [second_countable_topology H] @@ -646,6 +650,9 @@ variables [topological_space H] [topological_space M] [charted_space H M] @[simp, mfld_simps] lemma prod_charted_space_chart_at : (chart_at (model_prod H H') x) = (chart_at H x.fst).prod (chart_at H' x.snd) := rfl +lemma charted_space_self_prod : prod_charted_space H H H' H' = charted_space_self (H × H') := +by { ext1, { simp [prod_charted_space, atlas] }, { ext1, simp [chart_at_self_eq], refl } } + end prod_charted_space /-- The product of a finite family of charted spaces is naturally a charted space, with the diff --git a/src/geometry/manifold/cont_mdiff.lean b/src/geometry/manifold/cont_mdiff.lean index b0d7d42e4d820..66b0da15f1e57 100644 --- a/src/geometry/manifold/cont_mdiff.lean +++ b/src/geometry/manifold/cont_mdiff.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Floris van Doorn -/ import geometry.manifold.smooth_manifold_with_corners import geometry.manifold.local_invariant_properties @@ -56,6 +56,10 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {E' : Type*} [normed_add_comm_group E'] [normed_space 𝕜 E'] {H' : Type*} [topological_space H'] (I' : model_with_corners 𝕜 E' H') {M' : Type*} [topological_space M'] [charted_space H' M'] [I's : smooth_manifold_with_corners I' M'] +-- declare a manifold `M''` over the pair `(E'', H'')`. +{E'' : Type*} [normed_add_comm_group E''] [normed_space 𝕜 E''] +{H'' : Type*} [topological_space H''] {I'' : model_with_corners 𝕜 E'' H''} +{M'' : Type*} [topological_space M''] [charted_space H'' M''] -- declare a smooth manifold `N` over the pair `(F, G)`. {F : Type*} [normed_add_comm_group F] [normed_space 𝕜 F] {G : Type*} [topological_space G] {J : model_with_corners 𝕜 F G} @@ -64,6 +68,8 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] {F' : Type*} [normed_add_comm_group F'] [normed_space 𝕜 F'] {G' : Type*} [topological_space G'] {J' : model_with_corners 𝕜 F' G'} {N' : Type*} [topological_space N'] [charted_space G' N'] [J's : smooth_manifold_with_corners J' N'] +-- F'' is a normed space +{F'' : Type*} [normed_add_comm_group F''] [normed_space 𝕜 F''] -- declare functions, sets, points and smoothness indices {f f₁ : M → M'} {s s₁ t : set M} {x : M} {m n : ℕ∞} @@ -876,10 +882,6 @@ lemma cont_mdiff_of_locally_cont_mdiff_on section composition -variables {E'' : Type*} [normed_add_comm_group E''] [normed_space 𝕜 E''] -{H'' : Type*} [topological_space H''] {I'' : model_with_corners 𝕜 E'' H''} -{M'' : Type*} [topological_space M''] [charted_space H'' M''] - /-- The composition of `C^n` functions within domains at points is `C^n`. -/ lemma cont_mdiff_within_at.comp {t : set M'} {g : M' → M''} (x : M) (hg : cont_mdiff_within_at I' I'' n g t (f x)) @@ -914,18 +916,37 @@ begin { simp only [written_in_ext_chart_at, (∘), mem_ext_chart_source, e.left_inv, e'.left_inv] } end +/-- The composition of `C^∞` functions within domains at points is `C^∞`. -/ +lemma smooth_within_at.comp {t : set M'} {g : M' → M''} (x : M) + (hg : smooth_within_at I' I'' g t (f x)) + (hf : smooth_within_at I I' f s x) + (st : maps_to f s t) : smooth_within_at I I'' (g ∘ f) s x := +hg.comp x hf st + /-- The composition of `C^n` functions on domains is `C^n`. -/ lemma cont_mdiff_on.comp {t : set M'} {g : M' → M''} (hg : cont_mdiff_on I' I'' n g t) (hf : cont_mdiff_on I I' n f s) (st : s ⊆ f ⁻¹' t) : cont_mdiff_on I I'' n (g ∘ f) s := λ x hx, (hg _ (st hx)).comp x (hf x hx) st +/-- The composition of `C^∞` functions on domains is `C^∞`. -/ +lemma smooth_on.comp {t : set M'} {g : M' → M''} + (hg : smooth_on I' I'' g t) (hf : smooth_on I I' f s) + (st : s ⊆ f ⁻¹' t) : smooth_on I I'' (g ∘ f) s := +hg.comp hf st + /-- The composition of `C^n` functions on domains is `C^n`. -/ lemma cont_mdiff_on.comp' {t : set M'} {g : M' → M''} (hg : cont_mdiff_on I' I'' n g t) (hf : cont_mdiff_on I I' n f s) : cont_mdiff_on I I'' n (g ∘ f) (s ∩ f ⁻¹' t) := hg.comp (hf.mono (inter_subset_left _ _)) (inter_subset_right _ _) +/-- The composition of `C^∞` functions is `C^∞`. -/ +lemma smooth_on.comp' {t : set M'} {g : M' → M''} + (hg : smooth_on I' I'' g t) (hf : smooth_on I I' f s) : + smooth_on I I'' (g ∘ f) (s ∩ f ⁻¹' t) := +hg.comp' hf + /-- The composition of `C^n` functions is `C^n`. -/ lemma cont_mdiff.comp {g : M' → M''} (hg : cont_mdiff I' I'' n g) (hf : cont_mdiff I I' n f) : @@ -935,6 +956,11 @@ begin exact hg.comp hf subset_preimage_univ, end +/-- The composition of `C^∞` functions is `C^∞`. -/ +lemma smooth.comp {g : M' → M''} (hg : smooth I' I'' g) (hf : smooth I I' f) : + smooth I I'' (g ∘ f) := +hg.comp hf + /-- The composition of `C^n` functions within domains at points is `C^n`. -/ lemma cont_mdiff_within_at.comp' {t : set M'} {g : M' → M''} (x : M) (hg : cont_mdiff_within_at I' I'' n g t (f x)) @@ -942,6 +968,13 @@ lemma cont_mdiff_within_at.comp' {t : set M'} {g : M' → M''} (x : M) cont_mdiff_within_at I I'' n (g ∘ f) (s ∩ f⁻¹' t) x := hg.comp x (hf.mono (inter_subset_left _ _)) (inter_subset_right _ _) +/-- The composition of `C^∞` functions within domains at points is `C^∞`. -/ +lemma smooth_within_at.comp' {t : set M'} {g : M' → M''} (x : M) + (hg : smooth_within_at I' I'' g t (f x)) + (hf : smooth_within_at I I' f s x) : + smooth_within_at I I'' (g ∘ f) (s ∩ f⁻¹' t) x := +hg.comp' x hf + /-- `g ∘ f` is `C^n` within `s` at `x` if `g` is `C^n` at `f x` and `f` is `C^n` within `s` at `x`. -/ lemma cont_mdiff_at.comp_cont_mdiff_within_at {g : M' → M''} (x : M) @@ -949,12 +982,25 @@ lemma cont_mdiff_at.comp_cont_mdiff_within_at {g : M' → M''} (x : M) cont_mdiff_within_at I I'' n (g ∘ f) s x := hg.comp x hf (maps_to_univ _ _) +/-- `g ∘ f` is `C^∞` within `s` at `x` if `g` is `C^∞` at `f x` and +`f` is `C^∞` within `s` at `x`. -/ +lemma smooth_at.comp_smooth_within_at {g : M' → M''} (x : M) + (hg : smooth_at I' I'' g (f x)) (hf : smooth_within_at I I' f s x) : + smooth_within_at I I'' (g ∘ f) s x := +hg.comp_cont_mdiff_within_at x hf + /-- The composition of `C^n` functions at points is `C^n`. -/ lemma cont_mdiff_at.comp {g : M' → M''} (x : M) (hg : cont_mdiff_at I' I'' n g (f x)) (hf : cont_mdiff_at I I' n f x) : cont_mdiff_at I I'' n (g ∘ f) x := hg.comp x hf (maps_to_univ _ _) +/-- The composition of `C^∞` functions at points is `C^∞`. -/ +lemma smooth_at.comp {g : M' → M''} (x : M) + (hg : smooth_at I' I'' g (f x)) (hf : smooth_at I I' f x) : + smooth_at I I'' (g ∘ f) x := +hg.comp x hf + lemma cont_mdiff.comp_cont_mdiff_on {f : M → M'} {g : M' → M''} {s : set M} (hg : cont_mdiff I' I'' n g) (hf : cont_mdiff_on I I' n f s) : cont_mdiff_on I I'' n (g ∘ f) s := @@ -1157,6 +1203,29 @@ by rw [← cont_diff_on_univ, ← cont_mdiff_on_univ, alias cont_mdiff_iff_cont_diff ↔ cont_mdiff.cont_diff cont_diff.cont_mdiff +lemma cont_diff_within_at.comp_cont_mdiff_within_at + {g : F → F'} {f : M → F} {s : set M} {t : set F} {x : M} + (hg : cont_diff_within_at 𝕜 n g t (f x)) + (hf : cont_mdiff_within_at I 𝓘(𝕜, F) n f s x) (h : s ⊆ f ⁻¹' t) : + cont_mdiff_within_at I 𝓘(𝕜, F') n (g ∘ f) s x := +begin + rw cont_mdiff_within_at_iff at *, + refine ⟨hg.continuous_within_at.comp hf.1 h, _⟩, + rw [← (ext_chart_at I x).left_inv (mem_ext_chart_source I x)] at hg, + apply cont_diff_within_at.comp _ (by exact hg) hf.2 _, + exact (inter_subset_left _ _).trans (preimage_mono h) +end + +lemma cont_diff_at.comp_cont_mdiff_at {g : F → F'} {f : M → F} {x : M} + (hg : cont_diff_at 𝕜 n g (f x)) (hf : cont_mdiff_at I 𝓘(𝕜, F) n f x) : + cont_mdiff_at I 𝓘(𝕜, F') n (g ∘ f) x := +hg.comp_cont_mdiff_within_at hf subset.rfl + +lemma cont_diff.comp_cont_mdiff {g : F → F'} {f : M → F} + (hg : cont_diff 𝕜 n g) (hf : cont_mdiff I 𝓘(𝕜, F) n f) : + cont_mdiff I 𝓘(𝕜, F') n (g ∘ f) := +λ x, hg.cont_diff_at.comp_cont_mdiff_at (hf x) + end module /-! ### Smoothness of standard maps associated to the product of manifolds -/ @@ -1293,6 +1362,22 @@ lemma smooth_fst : smooth (I.prod J) I (@prod.fst M N) := cont_mdiff_fst +lemma cont_mdiff_at.fst {f : N → M × M'} {x : N} (hf : cont_mdiff_at J (I.prod I') n f x) : + cont_mdiff_at J I n (λ x, (f x).1) x := +cont_mdiff_at_fst.comp x hf + +lemma cont_mdiff.fst {f : N → M × M'} (hf : cont_mdiff J (I.prod I') n f) : + cont_mdiff J I n (λ x, (f x).1) := +cont_mdiff_fst.comp hf + +lemma smooth_at.fst {f : N → M × M'} {x : N} (hf : smooth_at J (I.prod I') f x) : + smooth_at J I (λ x, (f x).1) x := +smooth_at_fst.comp x hf + +lemma smooth.fst {f : N → M × M'} (hf : smooth J (I.prod I') f) : + smooth J I (λ x, (f x).1) := +smooth_fst.comp hf + lemma cont_mdiff_within_at_snd {s : set (M × N)} {p : M × N} : cont_mdiff_within_at (I.prod J) J n prod.snd s p := begin @@ -1332,6 +1417,22 @@ lemma smooth_snd : smooth (I.prod J) J (@prod.snd M N) := cont_mdiff_snd +lemma cont_mdiff_at.snd {f : N → M × M'} {x : N} (hf : cont_mdiff_at J (I.prod I') n f x) : + cont_mdiff_at J I' n (λ x, (f x).2) x := +cont_mdiff_at_snd.comp x hf + +lemma cont_mdiff.snd {f : N → M × M'} (hf : cont_mdiff J (I.prod I') n f) : + cont_mdiff J I' n (λ x, (f x).2) := +cont_mdiff_snd.comp hf + +lemma smooth_at.snd {f : N → M × M'} {x : N} (hf : smooth_at J (I.prod I') f x) : + smooth_at J I' (λ x, (f x).2) x := +smooth_at_snd.comp x hf + +lemma smooth.snd {f : N → M × M'} (hf : smooth J (I.prod I') f) : + smooth J I' (λ x, (f x).2) := +smooth_snd.comp hf + lemma smooth_iff_proj_smooth {f : M → M' × N'} : (smooth I (I'.prod J') f) ↔ (smooth I I' (prod.fst ∘ f)) ∧ (smooth I J' (prod.snd ∘ f)) := begin @@ -1340,6 +1441,10 @@ begin { rintro ⟨h_fst, h_snd⟩, simpa only [prod.mk.eta] using h_fst.prod_mk h_snd, } end +lemma smooth_prod_assoc : + smooth ((I.prod I').prod J) (I.prod (I'.prod J)) (λ x : (M × M') × N, (x.1.1, x.1.2, x.2)) := +smooth_fst.fst.prod_mk $ smooth_fst.snd.prod_mk smooth_snd + end projections section prod_map @@ -1474,6 +1579,21 @@ lemma continuous_linear_map.cont_mdiff (L : E →L[𝕜] F) : cont_mdiff 𝓘(𝕜, E) 𝓘(𝕜, F) n L := L.cont_diff.cont_mdiff +-- the following proof takes very long to elaborate in pure term mode +lemma cont_mdiff_at.clm_comp {g : M → F →L[𝕜] F''} {f : M → F' →L[𝕜] F} {x : M} + (hg : cont_mdiff_at I 𝓘(𝕜, F →L[𝕜] F'') n g x) (hf : cont_mdiff_at I 𝓘(𝕜, F' →L[𝕜] F) n f x) : + cont_mdiff_at I 𝓘(𝕜, F' →L[𝕜] F'') n (λ x, (g x).comp (f x)) x := +@cont_diff_at.comp_cont_mdiff_at _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ + (λ x : (F →L[𝕜] F'') × (F' →L[𝕜] F), x.1.comp x.2) (λ x, (g x, f x)) x + (by { apply cont_diff.cont_diff_at, apply is_bounded_bilinear_map.cont_diff, + exact is_bounded_bilinear_map_comp }) -- todo: simplify after #16946 + (hg.prod_mk_space hf) + +lemma cont_mdiff.clm_comp {g : M → F →L[𝕜] F''} {f : M → F' →L[𝕜] F} + (hg : cont_mdiff I 𝓘(𝕜, F →L[𝕜] F'') n g) (hf : cont_mdiff I 𝓘(𝕜, F' →L[𝕜] F) n f) : + cont_mdiff I 𝓘(𝕜, F' →L[𝕜] F'') n (λ x, (g x).comp (f x)) := +λ x, (hg x).clm_comp (hf x) + /-! ### Smoothness of standard operations -/ variables {V : Type*} [normed_add_comm_group V] [normed_space 𝕜 V] diff --git a/src/geometry/manifold/cont_mdiff_mfderiv.lean b/src/geometry/manifold/cont_mdiff_mfderiv.lean index c663fc78db4eb..d19910b660cd9 100644 --- a/src/geometry/manifold/cont_mdiff_mfderiv.lean +++ b/src/geometry/manifold/cont_mdiff_mfderiv.lean @@ -72,6 +72,16 @@ lemma cont_mdiff.mdifferentiable (hf : cont_mdiff I I' n f) (hn : 1 ≤ n) : mdifferentiable I I' f := λ x, (hf x).mdifferentiable_at hn +lemma smooth_within_at.mdifferentiable_within_at + (hf : smooth_within_at I I' f s x) : mdifferentiable_within_at I I' f s x := +hf.mdifferentiable_within_at le_top + +lemma smooth_at.mdifferentiable_at (hf : smooth_at I I' f x) : mdifferentiable_at I I' f x := +hf.mdifferentiable_at le_top + +lemma smooth_on.mdifferentiable_on (hf : smooth_on I I' f s) : mdifferentiable_on I I' f s := +hf.mdifferentiable_on le_top + lemma smooth.mdifferentiable (hf : smooth I I' f) : mdifferentiable I I' f := cont_mdiff.mdifferentiable hf le_top @@ -445,6 +455,12 @@ begin .continuous_at_of_comp_left h (mem_chart_source _ _) (h.prod hf.continuous_at.snd) end +lemma smooth_iff_target {f : N → Z.to_topological_vector_bundle_core.total_space} : + smooth J (I.prod 𝓘(𝕜, E')) f ↔ continuous (bundle.total_space.proj ∘ f) ∧ + ∀ x, smooth_at J 𝓘(𝕜, E × E') (ext_chart_at (I.prod 𝓘(𝕜, E')) (f x) ∘ f) x := +by simp_rw [smooth, smooth_at, cont_mdiff, Z.cont_mdiff_at_iff_target, forall_and_distrib, + continuous_iff_continuous_at] + lemma cont_mdiff_proj : cont_mdiff (I.prod 𝓘(𝕜, E')) I n Z.to_topological_vector_bundle_core.proj := begin diff --git a/src/geometry/manifold/local_invariant_properties.lean b/src/geometry/manifold/local_invariant_properties.lean index 4e3a8114291bf..bc49a2f0336ed 100644 --- a/src/geometry/manifold/local_invariant_properties.lean +++ b/src/geometry/manifold/local_invariant_properties.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Sébastien Gouëzel +Authors: Sébastien Gouëzel, Floris van Doorn -/ import geometry.manifold.charted_space diff --git a/src/geometry/manifold/smooth_manifold_with_corners.lean b/src/geometry/manifold/smooth_manifold_with_corners.lean index 7300bbe09e949..2c97a018a5272 100644 --- a/src/geometry/manifold/smooth_manifold_with_corners.lean +++ b/src/geometry/manifold/smooth_manifold_with_corners.lean @@ -124,7 +124,7 @@ localized "notation (name := with_top.nat.top) `∞` := (⊤ : ℕ∞)" in manif model vector space `E` over the field `𝕜`. This is all what is needed to define a smooth manifold with model space `H`, and model vector space `E`. -/ -@[nolint has_nonempty_instance] +@[ext, nolint has_nonempty_instance] structure model_with_corners (𝕜 : Type*) [nontrivially_normed_field 𝕜] (E : Type*) [normed_add_comm_group E] [normed_space 𝕜 E] (H : Type*) [topological_space H] extends local_equiv H E := @@ -228,6 +228,9 @@ I.left_inverse.right_inv_on_range @[simp, mfld_simps] protected lemma right_inv {x : E} (hx : x ∈ range I) : I (I.symm x) = x := I.right_inv_on hx +lemma preimage_image (s : set H) : I ⁻¹' (I '' s) = s := +I.injective.preimage_image s + protected lemma image_eq (s : set H) : I '' s = I.symm ⁻¹' s ∩ range I := begin refine (I.to_local_equiv.image_eq_target_inter_inv_preimage _).trans _, @@ -365,6 +368,12 @@ rfl (I : model_with_corners 𝕜 E H) (I' : model_with_corners 𝕜 E' H') : ((I.prod I').symm : _ × _ → _ × _) = prod.map I.symm I'.symm := rfl +lemma model_with_corners_self_prod : 𝓘(𝕜, E × F) = 𝓘(𝕜, E).prod 𝓘(𝕜, F) := +by { ext1, simp } + +lemma model_with_corners.range_prod : range (I.prod J) = range I ×ˢ range J := +by { simp_rw [← model_with_corners.target_eq], refl } + end model_with_corners_prod section boundaryless @@ -955,4 +964,10 @@ by simp only with mfld_simps lemma ext_chart_model_space_apply {x y : E} : ext_chart_at 𝓘(𝕜, E) x y = y := rfl +variable {𝕜} + +lemma ext_chart_at_prod (x : M × M') : + ext_chart_at (I.prod I') x = (ext_chart_at I x.1).prod (ext_chart_at I' x.2) := +by simp only with mfld_simps + end extended_charts diff --git a/src/geometry/manifold/tangent_bundle.lean b/src/geometry/manifold/tangent_bundle.lean index 88b4ba20ef87a..833b2318f7ed8 100644 --- a/src/geometry/manifold/tangent_bundle.lean +++ b/src/geometry/manifold/tangent_bundle.lean @@ -120,6 +120,39 @@ variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] instance : inhabited (basic_smooth_vector_bundle_core I M F) := ⟨trivial_basic_smooth_vector_bundle_core I M F⟩ +/-- A reformulation of `coord_change_comp`, formulated in terms of a point in `M`. +The conditions on this point a significantly simpler than in `coord_change_comp`. -/ +lemma coord_change_comp' {i j k : atlas H M} {x : M} + (hi : x ∈ i.1.source) (hj : x ∈ j.1.source) (hk : x ∈ k.1.source) (v : F) : + Z.coord_change j k (j x) (Z.coord_change i j (i x) v) = Z.coord_change i k (i x) v := +begin + rw [show j x = _, by rw [← i.1.left_inv hi]], + apply Z.coord_change_comp, + simp only [hi, hj, hk] with mfld_simps +end + +/-- A reformulation of `coord_change_self`, formulated in terms of a point in `M`. -/ +lemma coord_change_self' {i : atlas H M} {x : M} (hi : x ∈ i.1.source) (v : F) : + Z.coord_change i i (i x) v = v := +Z.coord_change_self i (i x) (i.1.maps_to hi) v + +/-- `Z.coord_change j i` is a partial inverse of `Z.coord_change i j`. -/ +lemma coord_change_comp_eq_self (i j : atlas H M) {x : H} + (hx : x ∈ (i.1.symm.trans j.1).source) (v : F) : + Z.coord_change j i (i.1.symm.trans j.1 x) (Z.coord_change i j x v) = v := +begin + rw [Z.coord_change_comp i j i x _ v, Z.coord_change_self _ _ hx.1], + simp only with mfld_simps at hx, + simp only [hx.1, hx.2] with mfld_simps +end + +/-- `Z.coord_change j i` is a partial inverse of `Z.coord_change i j`, +formulated in terms of a point in `M`. -/ +lemma coord_change_comp_eq_self' {i j : atlas H M} {x : M} + (hi : x ∈ i.1.source) (hj : x ∈ j.1.source) (v : F) : + Z.coord_change j i (j x) (Z.coord_change i j (i x) v) = v := +by rw [Z.coord_change_comp' hi hj hi v, Z.coord_change_self' hi] + lemma coord_change_continuous (i j : atlas H M) : continuous_on (Z.coord_change i j) (i.1.symm.trans j.1).source := begin @@ -182,6 +215,11 @@ def chart {e : local_homeomorph M H} (he : e ∈ atlas H M) : (Z.to_topological_vector_bundle_core.local_triv ⟨e, he⟩).to_local_homeomorph.trans (local_homeomorph.prod e (local_homeomorph.refl F)) +lemma chart_apply {x : M} (z : Z.to_topological_vector_bundle_core.total_space) : + Z.chart (chart_mem_atlas H x) z = (chart_at H x z.proj, + Z.coord_change (achart H z.proj) (achart H x) (achart H z.proj z.proj) z.2) := +rfl + @[simp, mfld_simps] lemma chart_source (e : local_homeomorph M H) (he : e ∈ atlas H M) : (Z.chart he).source = Z.to_topological_vector_bundle_core.proj ⁻¹' e.source := by { simp only [chart, mem_prod], mfld_set_tac } diff --git a/src/logic/equiv/local_equiv.lean b/src/logic/equiv/local_equiv.lean index 6b7ea9aef3b8c..e40fd0e42c045 100644 --- a/src/logic/equiv/local_equiv.lean +++ b/src/logic/equiv/local_equiv.lean @@ -691,6 +691,10 @@ lemma prod_coe_symm (e : local_equiv α β) (e' : local_equiv γ δ) : (e.prod e').symm = (e.symm.prod e'.symm) := by ext x; simp [prod_coe_symm] +@[simp, mfld_simps] lemma refl_prod_refl : + (local_equiv.refl α).prod (local_equiv.refl β) = local_equiv.refl (α × β) := +by { ext1 ⟨x, y⟩, { refl }, { rintro ⟨x, y⟩, refl }, exact univ_prod_univ } + @[simp, mfld_simps] lemma prod_trans {η : Type*} {ε : Type*} (e : local_equiv α β) (f : local_equiv β γ) (e' : local_equiv δ η) (f' : local_equiv η ε) : (e.prod e').trans (f.prod f') = (e.trans f).prod (e'.trans f') := diff --git a/src/topology/local_homeomorph.lean b/src/topology/local_homeomorph.lean index d410a89b206f0..19363b8f87e1c 100644 --- a/src/topology/local_homeomorph.lean +++ b/src/topology/local_homeomorph.lean @@ -813,6 +813,11 @@ def prod (e : local_homeomorph α β) (e' : local_homeomorph γ δ) : (e.prod e').symm = (e.symm.prod e'.symm) := rfl +@[simp] +lemma refl_prod_refl {α β : Type*} [topological_space α] [topological_space β] : + (local_homeomorph.refl α).prod (local_homeomorph.refl β) = local_homeomorph.refl (α × β) := +by { ext1 ⟨x, y⟩, { refl }, { rintro ⟨x, y⟩, refl }, exact univ_prod_univ } + @[simp, mfld_simps] lemma prod_trans {η : Type*} {ε : Type*} [topological_space η] [topological_space ε] (e : local_homeomorph α β) (f : local_homeomorph β γ) From cb8faf79918fe76e06ebba57df07eaff6c579c76 Mon Sep 17 00:00:00 2001 From: Kevin Wilson Date: Wed, 19 Oct 2022 11:59:27 +0000 Subject: [PATCH 818/828] feat(ring_theory/hahn_series): Loosen definition of additive valuation (#17057) The definitions and theorems realted to the additive valuation of Hahn series don't actually require the value group to be a group; just a cancellative monoid, as in the rest of the file. This came about while I was thinking about Bhargava's recent proof of van der Waerden's theorem and I wondered how hard proving, say, the irreducibility theorem would be in mathlib. --- src/ring_theory/hahn_series.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/ring_theory/hahn_series.lean b/src/ring_theory/hahn_series.lean index 220fff76edbd5..68ad94ae5ca85 100644 --- a/src/ring_theory/hahn_series.lean +++ b/src/ring_theory/hahn_series.lean @@ -1159,7 +1159,7 @@ end algebra section valuation -variables (Γ R) [linear_ordered_add_comm_group Γ] [ring R] [is_domain R] +variables (Γ R) [linear_ordered_cancel_add_comm_monoid Γ] [ring R] [is_domain R] /-- The additive valuation on `hahn_series Γ R`, returning the smallest index at which a Hahn Series has a nonzero coefficient, or `⊤` for the 0 series. -/ @@ -1208,7 +1208,7 @@ end end valuation lemma is_pwo_Union_support_powers - [linear_ordered_add_comm_group Γ] [ring R] [is_domain R] + [linear_ordered_cancel_add_comm_monoid Γ] [ring R] [is_domain R] {x : hahn_series Γ R} (hx : 0 < add_val Γ R x) : (⋃ n : ℕ, (x ^ n).support).is_pwo := begin @@ -1531,7 +1531,7 @@ end emb_domain section powers -variables [linear_ordered_add_comm_group Γ] [comm_ring R] [is_domain R] +variables [linear_ordered_cancel_add_comm_monoid Γ] [comm_ring R] [is_domain R] /-- The powers of an element of positive valuation form a summable family. -/ def powers (x : hahn_series Γ R) (hx : 0 < add_val Γ R x) : From 98e8d5571d9b3c1a407cb54697226561ab409862 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Wed, 19 Oct 2022 11:59:28 +0000 Subject: [PATCH 819/828] fix(topology/algebra/uniform_mul_action): remove duplicate typeclasses (#17063) --- src/topology/algebra/uniform_mul_action.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/topology/algebra/uniform_mul_action.lean b/src/topology/algebra/uniform_mul_action.lean index dc90954b31cd7..f2af832131a39 100644 --- a/src/topology/algebra/uniform_mul_action.lean +++ b/src/topology/algebra/uniform_mul_action.lean @@ -30,12 +30,12 @@ variables (R : Type u) (M : Type v) (N : Type w) (X : Type x) (Y : Type y) [uniform_space X] [uniform_space Y] /-- An additive action such that for all `c`, the map `λ x, c +ᵥ x` is uniformly continuous. -/ -class has_uniform_continuous_const_vadd [uniform_space X] [has_vadd M X] : Prop := +class has_uniform_continuous_const_vadd [has_vadd M X] : Prop := (uniform_continuous_const_vadd : ∀ (c : M), uniform_continuous ((+ᵥ) c : X → X)) /-- A multiplicative action such that for all `c`, the map `λ x, c • x` is uniformly continuous. -/ @[to_additive] -class has_uniform_continuous_const_smul [uniform_space X] [has_smul M X] : Prop := +class has_uniform_continuous_const_smul [has_smul M X] : Prop := (uniform_continuous_const_smul : ∀ (c : M), uniform_continuous ((•) c : X → X)) export has_uniform_continuous_const_vadd (uniform_continuous_const_vadd) From ccbe476c7631ead7a18e45a78bb12dd2d2f2e30d Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Wed, 19 Oct 2022 15:19:47 +0000 Subject: [PATCH 820/828] feat(algebra/squarefree): squarefree gcds of squarefree elements (#16999) --- src/algebra/squarefree.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/src/algebra/squarefree.lean b/src/algebra/squarefree.lean index d3dfa30d5e55c..3a105fba890a5 100644 --- a/src/algebra/squarefree.lean +++ b/src/algebra/squarefree.lean @@ -82,6 +82,20 @@ lemma squarefree_of_dvd_of_squarefree [comm_monoid R] squarefree x := λ a h, hsq _ (h.trans hdvd) +section squarefree_gcd_of_squarefree + +variables {α : Type*} [cancel_comm_monoid_with_zero α] [gcd_monoid α] + +lemma squarefree.gcd_right (a : α) {b : α} (hb : squarefree b) : + squarefree (gcd a b) := +squarefree_of_dvd_of_squarefree (gcd_dvd_right _ _) hb + +lemma squarefree.gcd_left {a : α} (b : α) (ha : squarefree a) : + squarefree (gcd a b) := +squarefree_of_dvd_of_squarefree (gcd_dvd_left _ _) ha + +end squarefree_gcd_of_squarefree + namespace multiplicity section comm_monoid From 7cc9c7ad820f6ce591e92b4a24f90769edd899b2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Mar=C3=ADa=20In=C3=A9s=20de=20Frutos-Fern=C3=A1ndez?= Date: Wed, 19 Oct 2022 17:53:38 +0000 Subject: [PATCH 821/828] feat(ring_theory/discrete_valuation_ring): add not_is_field (#16979) A discrete valuation ring is not a field. --- src/ring_theory/discrete_valuation_ring.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/ring_theory/discrete_valuation_ring.lean b/src/ring_theory/discrete_valuation_ring.lean index 412da64f8312d..c45237792530e 100644 --- a/src/ring_theory/discrete_valuation_ring.lean +++ b/src/ring_theory/discrete_valuation_ring.lean @@ -60,6 +60,10 @@ variables (R : Type u) [comm_ring R] [is_domain R] [discrete_valuation_ring R] lemma not_a_field : maximal_ideal R ≠ ⊥ := not_a_field' +/-- A discrete valuation ring `R` is not a field. -/ +lemma not_is_field : ¬ is_field R := +ring.not_is_field_iff_exists_prime.mpr ⟨_, not_a_field R, is_maximal.is_prime' (maximal_ideal R)⟩ + variable {R} open principal_ideal_ring From 4aa2a2e17940311e47007f087c9df229e7f12942 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 19 Oct 2022 22:10:47 +0000 Subject: [PATCH 822/828] chore(logic/basic): add move eq_true from core, as eq_true_iff (#17060) We want to delete `eq_true` and `eq_false` in core, per https://github.com/leanprover-community/lean/pull/774, as they have different meanings in Lean 4. This reintrodues `eq_true`, renamed to `eq_true_iff`, into mathlib3. Co-authored-by: Scott Morrison --- src/category_theory/limits/types.lean | 2 +- src/logic/basic.lean | 4 ++++ src/topology/order.lean | 4 ++-- src/topology/sequences.lean | 2 +- 4 files changed, 8 insertions(+), 4 deletions(-) diff --git a/src/category_theory/limits/types.lean b/src/category_theory/limits/types.lean index f02267b171972..365ae72bf946c 100644 --- a/src/category_theory/limits/types.lean +++ b/src/category_theory/limits/types.lean @@ -311,7 +311,7 @@ begin { have := congr_fun this x, have H := congr_arg ulift.down this, dsimp at H, - rwa eq_true at H }, + rwa eq_true_iff at H }, refine h.hom_ext _, intro j, ext y, erw iff_true, diff --git a/src/logic/basic.lean b/src/logic/basic.lean index ad3276f1ff2bc..0107c651c530a 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -238,6 +238,10 @@ open function theorem false_ne_true : false ≠ true | h := h.symm ▸ trivial +theorem eq_true_iff {a : Prop} : (a = true) = a := +have (a ↔ true) = a, from propext (iff_true a), +eq.subst (@iff_eq_eq a true) this + section propositional variables {a b c d e f : Prop} diff --git a/src/topology/order.lean b/src/topology/order.lean index facb9f1143ecc..c3c56da713d69 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -847,10 +847,10 @@ lemma continuous_Prop {p : α → Prop} : continuous p ↔ is_open {x | p x} := ⟨assume h : continuous p, have is_open (p ⁻¹' {true}), from is_open_singleton_true.preimage h, - by simpa [preimage, eq_true] using this, + by simpa [preimage, eq_true_iff] using this, assume h : is_open {x | p x}, continuous_generated_from $ assume s (hs : s = {true}), - by simp [hs, preimage, eq_true, h]⟩ + by simp [hs, preimage, eq_true_iff, h]⟩ lemma is_open_iff_continuous_mem {s : set α} : is_open s ↔ continuous (λ x, x ∈ s) := continuous_Prop.symm diff --git a/src/topology/sequences.lean b/src/topology/sequences.lean index c849be6c687bf..c6637f97752f5 100644 --- a/src/topology/sequences.lean +++ b/src/topology/sequences.lean @@ -149,7 +149,7 @@ begin refine ⟨λ s x hcx, _⟩, specialize h (∉ s) x, by_cases hx : x ∈ s, { exact subset_seq_closure hx }, - simp_rw [(∘), continuous_at, hx, not_false_iff, nhds_true, tendsto_pure, eq_true, + simp_rw [(∘), continuous_at, hx, not_false_iff, nhds_true, tendsto_pure, eq_true_iff, ← mem_compl_iff, eventually_mem_set, ← mem_interior_iff_mem_nhds, interior_compl] at h, rw [mem_compl_iff, imp_not_comm] at h, simp only [not_forall, not_eventually, mem_compl_iff, not_not] at h, From 97bdf517581e1ed85bb3a6e82ad0f78bf05fd4fc Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 20 Oct 2022 01:30:10 +0000 Subject: [PATCH 823/828] chore(algebra/*): mark frozen files that have been ported to mathlib4 (#17039) This is chiefly a proposal that we get started marking files as frozen, and a proposal for a specific format. Once we settle on what gets written here, we should add some CI that adds a special label to any PR modifying a frozen file. Co-authored-by: Scott Morrison --- src/algebra/abs.lean | 4 ++++ src/algebra/covariant_and_contravariant.lean | 4 ++++ src/algebra/group/basic.lean | 4 ++++ src/algebra/group/defs.lean | 4 ++++ src/logic/relator.lean | 11 +++++++++-- 5 files changed, 25 insertions(+), 2 deletions(-) diff --git a/src/algebra/abs.lean b/src/algebra/abs.lean index 3bb84fb2fd731..035d10ec6294d 100644 --- a/src/algebra/abs.lean +++ b/src/algebra/abs.lean @@ -7,6 +7,10 @@ Authors: Christopher Hoskin /-! # Absolute value +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/477 +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a notational class `has_abs` which adds the unary operator `abs` and the notation `|.|`. The concept of an absolute value occurs in lattice ordered groups and in GL and GM spaces. diff --git a/src/algebra/covariant_and_contravariant.lean b/src/algebra/covariant_and_contravariant.lean index b234d416f3974..a7918f1926195 100644 --- a/src/algebra/covariant_and_contravariant.lean +++ b/src/algebra/covariant_and_contravariant.lean @@ -12,6 +12,10 @@ import order.monotone # Covariants and contravariants +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/467 +> Any changes to this file require a corresponding PR to mathlib4. + This file contains general lemmas and instances to work with the interactions between a relation and an action on a Type. diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index c33fe507ca94d..3384d283b21fa 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -9,6 +9,10 @@ import algebra.group.defs /-! # Basic lemmas about semigroups, monoids, and groups +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/457 +> Any changes to this file require a corresponding PR to mathlib4. + This file lists various basic lemmas about semigroups, monoids, and groups. Most proofs are one-liners from the corresponding axioms. For the definitions of semigroups, monoids and groups, see `algebra/group/defs.lean`. diff --git a/src/algebra/group/defs.lean b/src/algebra/group/defs.lean index 8348d576871af..893d900b0088b 100644 --- a/src/algebra/group/defs.lean +++ b/src/algebra/group/defs.lean @@ -9,6 +9,10 @@ import tactic.basic /-! # Typeclasses for (semi)groups and monoids +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/457 +> Any changes to this file require a corresponding PR to mathlib4. + In this file we define typeclasses for algebraic structures with one binary operation. The classes are named `(add_)?(comm_)?(semigroup|monoid|group)`, where `add_` means that the class uses additive notation and `comm_` means that the class assumes that the binary diff --git a/src/logic/relator.lean b/src/logic/relator.lean index c9ea9b0afedac..de2d4441d0fa5 100644 --- a/src/logic/relator.lean +++ b/src/logic/relator.lean @@ -2,12 +2,19 @@ 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 - -Relator for functions, pairs, sums, and lists. -/ import logic.basic +/-! +# Relator for functions, pairs, sums, and lists. + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> https://github.com/leanprover-community/mathlib4/pull/385 +> Any changes to this file require a corresponding PR to mathlib4. + +-/ + namespace relator universes u₁ u₂ v₁ v₂ From 0578d23fb1e3fc1f1249a6c54d97bde7793de9c5 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 20 Oct 2022 01:30:12 +0000 Subject: [PATCH 824/828] chore: update scripts/port_status for more flexible tagging (#17056) In https://github.com/leanprover-community/mathlib/wiki/mathlib4-port-status we would like to be able to write more complicated statuses, like * `Yes mathlib4#999 1234abcd` to track commit information about the port * `No blocked by ring tactic` * `No PR'd as ...` This update to the script checks if the status starts with `Yes` or `No` to determine status. Further, when reporting good candidates for porting, it includes the comment after `No` if present. Co-authored-by: Scott Morrison --- scripts/port_status.py | 28 ++++++++++++++++++++-------- 1 file changed, 20 insertions(+), 8 deletions(-) diff --git a/scripts/port_status.py b/scripts/port_status.py index b36981b642b0a..296a2422a3116 100755 --- a/scripts/port_status.py +++ b/scripts/port_status.py @@ -13,6 +13,15 @@ def mk_label(path: Path) -> str: rel = path.relative_to(Path('src')) return str(rel.with_suffix('')).replace(os.sep, '.') +def stripPrefix(tag : str) -> str: + if tag.startswith('No: '): + return tag[4:] + else: + if tag.startswith('No'): + return tag[2:].lstrip() + else: + return tag + graph = nx.DiGraph() for path in Path('src').glob('**/*.lean'): @@ -46,20 +55,23 @@ def mk_label(path: Path) -> str: parentsDone = dict() for node in graph.nodes: ancestors = nx.ancestors(graph, node) - if all(data[imported] == 'Yes' for imported in ancestors) and data[node] == 'No': - allDone[node] = len(nx.descendants(graph, node)) + if all(data[imported].startswith('Yes') for imported in ancestors) and data[node].startswith('No'): + allDone[node] = (len(nx.descendants(graph, node)), stripPrefix(data[node])) else: - if all(data[imported] == 'Yes' for imported in graph.predecessors(node)) and data[node] == 'No': - parentsDone[node] = len(nx.descendants(graph, node)) + if all(data[imported].startswith('Yes') for imported in graph.predecessors(node)) and data[node].startswith('No'): + parentsDone[node] = (len(nx.descendants(graph, node)), stripPrefix(data[node])) print('# The following files have all dependencies ported already, and should be ready to port:') print('# Earlier items in the list are required in more places in mathlib.') -allDone = dict(sorted(allDone.items(), key=lambda item: -item[1])) -for k in allDone: - print(k) +allDone = dict(sorted(allDone.items(), key=lambda item: -item[1][0])) +for k, v in allDone.items(): + if v[1] == "": + print(k) + else: + print(k + " -- " + v[1]) print() print('# The following files have their immediate dependencies ported already, and may be ready to port:') -parentsDone = dict(sorted(parentsDone.items(), key=lambda item: -item[1])) +parentsDone = dict(sorted(parentsDone.items(), key=lambda item: -item[1][0])) for k in parentsDone: print(k) From eed79d6e62369f3c24f6ff5caf170d19c918cee4 Mon Sep 17 00:00:00 2001 From: Jakob von Raumer Date: Thu, 20 Oct 2022 04:48:09 +0000 Subject: [PATCH 825/828] chore: freeze `data.bracket` (#17077) --- src/data/bracket.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/data/bracket.lean b/src/data/bracket.lean index 4372992dcfd76..d2d4801a42eeb 100644 --- a/src/data/bracket.lean +++ b/src/data/bracket.lean @@ -5,6 +5,10 @@ Authors: Patrick Lutz, Oliver Nash -/ /-! +THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +https://github.com/leanprover-community/mathlib4/pull/480 +Any changes to this file require a corresponding PR to mathlib4. + # Bracket Notation This file provides notation which can be used for the Lie bracket, for the commutator of two From bdcb5fb7877d5f341028f489c3337526832c95a8 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 20 Oct 2022 07:06:57 +0000 Subject: [PATCH 826/828] =?UTF-8?q?feat(analysis/inner=5Fproduct=5Fspace/t?= =?UTF-8?q?wo=5Fdim):=20the=20case=20of=20`=E2=84=82`=20(#16929)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This file relates the constructions `orientation.area_form`, `orientation.right_angle_rotation`, `orientation.kahler` on an oriented two-dimensional real inner product space to their concrete interpretations over `ℂ`. --- src/analysis/inner_product_space/basic.lean | 8 + .../inner_product_space/orientation.lean | 30 ++++ src/analysis/inner_product_space/two_dim.lean | 169 +++++++++++++++++- src/data/complex/module.lean | 4 + src/linear_algebra/determinant.lean | 20 +++ src/linear_algebra/orientation.lean | 59 +++++- src/linear_algebra/ray.lean | 6 + 7 files changed, 290 insertions(+), 6 deletions(-) diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 316fb4b786889..2c9f470c3339b 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -2117,6 +2117,14 @@ omit 𝕜 instance inner_product_space.complex_to_real [inner_product_space ℂ G] : inner_product_space ℝ G := inner_product_space.is_R_or_C_to_real ℂ G +@[simp] protected lemma complex.inner (w z : ℂ) : ⟪w, z⟫_ℝ = (conj w * z).re := rfl + +/-- The inner product on an inner product space of dimension 2 can be evaluated in terms +of a complex-number representation of the space. -/ +lemma inner_map_complex [inner_product_space ℝ G] (f : G ≃ₗᵢ[ℝ] ℂ) (x y : G) : + ⟪x, y⟫_ℝ = (conj (f x) * f y).re := +by rw [← complex.inner, f.inner_map_map] + end is_R_or_C_to_real section continuous diff --git a/src/analysis/inner_product_space/orientation.lean b/src/analysis/inner_product_space/orientation.lean index b31f28dbcdd84..652b3d743b051 100644 --- a/src/analysis/inner_product_space/orientation.lean +++ b/src/analysis/inner_product_space/orientation.lean @@ -315,6 +315,36 @@ lemma abs_volume_form_apply_of_orthonormal (v : orthonormal_basis (fin n) ℝ E) |o.volume_form v| = 1 := by simpa [o.volume_form_robust' v v] using congr_arg abs v.to_basis.det_self +lemma volume_form_map {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = n)] + (φ : E ≃ₗᵢ[ℝ] F) (x : fin n → F) : + (orientation.map (fin n) φ.to_linear_equiv o).volume_form x = o.volume_form (φ.symm ∘ x) := +begin + unfreezingI { cases n }, + { refine o.eq_or_eq_neg_of_is_empty.by_cases _ _; rintros rfl; simp }, + let e : orthonormal_basis (fin n.succ) ℝ E := o.fin_orthonormal_basis n.succ_pos (fact.out _), + have he : e.to_basis.orientation = o := + (o.fin_orthonormal_basis_orientation n.succ_pos (fact.out _)), + have heφ : (e.map φ).to_basis.orientation = orientation.map (fin n.succ) φ.to_linear_equiv o, + { rw ← he, + exact (e.to_basis.orientation_map φ.to_linear_equiv) }, + rw (orientation.map (fin n.succ) φ.to_linear_equiv o).volume_form_robust (e.map φ) heφ, + rw o.volume_form_robust e he, + simp, +end + +/-- The volume form is invariant under pullback by a positively-oriented isometric automorphism. -/ +lemma volume_form_comp_linear_isometry_equiv (φ : E ≃ₗᵢ[ℝ] E) + (hφ : 0 < (φ.to_linear_equiv : E →ₗ[ℝ] E).det) (x : fin n → E) : + o.volume_form (φ ∘ x) = o.volume_form x := +begin + convert o.volume_form_map φ (φ ∘ x), + { symmetry, + rwa ← o.map_eq_iff_det_pos φ.to_linear_equiv at hφ, + rw [_i.out, fintype.card_fin] }, + { ext, + simp } +end + end volume_form end orientation diff --git a/src/analysis/inner_product_space/two_dim.lean b/src/analysis/inner_product_space/two_dim.lean index 483f6ed609686..4c77741fa89fa 100644 --- a/src/analysis/inner_product_space/two_dim.lean +++ b/src/analysis/inner_product_space/two_dim.lean @@ -32,6 +32,24 @@ product space `E`. the complex number `o.kahler x y` has modulus `∥x∥ * ∥y∥`. TODO: the argument of `o.kahler x y` is the oriented angle (`orientation.oangle`) from `x` to `y`. +## Main results + +* `orientation.right_angle_rotation_right_angle_rotation`: the identity `J (J x) = - x` + +* `orientation.nonneg_inner_and_area_form_eq_zero_iff_same_ray`: `x`, `y` are in the same ray, if + and only if `0 ≤ ⟪x, y⟫` and `ω x y = 0` + +* `orientation.kahler_mul`: the identity `o.kahler x a * o.kahler a y = ∥a∥ ^ 2 * o.kahler x y` + +* `complex.area_form`, `complex.right_angle_rotation`, `complex.kahler`: the concrete + interpretations of `area_form`, `right_angle_rotation`, `kahler` for the oriented real inner + product space `ℂ` + +* `orientation.area_form_map_complex`, `orientation.right_angle_rotation_map_complex`, + `orientation.kahler_map_complex`: given an orientation-preserving isometry from `E` to `ℂ`, + expressions for `area_form`, `right_angle_rotation`, `kahler` as the pullback of their concrete + interpretations on `ℂ` + ## Implementation notes Notation `ω` for `orientation.area_form` and `J` for `orientation.right_angle_rotation` should be @@ -128,6 +146,29 @@ begin { simpa } end +lemma area_form_map {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] + (φ : E ≃ₗᵢ[ℝ] F) (x y : F) : + (orientation.map (fin 2) φ.to_linear_equiv o).area_form x y = o.area_form (φ.symm x) (φ.symm y) := +begin + have : φ.symm ∘ ![x, y] = ![φ.symm x, φ.symm y], + { ext i, + fin_cases i; refl }, + simp [area_form_to_volume_form, volume_form_map, this], +end + +/-- The area form is invariant under pullback by a positively-oriented isometric automorphism. -/ +lemma area_form_comp_linear_isometry_equiv (φ : E ≃ₗᵢ[ℝ] E) + (hφ : 0 < (φ.to_linear_equiv : E →ₗ[ℝ] E).det) (x y : E) : + o.area_form (φ x) (φ y) = o.area_form x y := +begin + convert o.area_form_map φ (φ x) (φ y), + { symmetry, + rwa ← o.map_eq_iff_det_pos φ.to_linear_equiv at hφ, + rw [fact.out (finrank ℝ E = 2), fintype.card_fin] }, + { simp }, + { simp } +end + /-- Auxiliary construction for `orientation.right_angle_rotation`, rotation by 90 degrees in an oriented real inner product space of dimension 2. -/ @[irreducible] def right_angle_rotation_aux₁ : E →ₗ[ℝ] E := @@ -262,6 +303,45 @@ end (-o).right_angle_rotation = o.right_angle_rotation.trans (linear_isometry_equiv.neg ℝ) := linear_isometry_equiv.ext $ o.right_angle_rotation_neg_orientation +lemma right_angle_rotation_map {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] + (φ : E ≃ₗᵢ[ℝ] F) (x : F) : + (orientation.map (fin 2) φ.to_linear_equiv o).right_angle_rotation x + = φ (o.right_angle_rotation (φ.symm x)) := +begin + apply ext_inner_right ℝ, + intros y, + rw inner_right_angle_rotation_left, + transitivity ⟪J (φ.symm x), φ.symm y⟫, + { simp [o.area_form_map] }, + transitivity ⟪φ (J (φ.symm x)), φ (φ.symm y)⟫, + { rw φ.inner_map_map }, + { simp }, +end + +/-- `J` commutes with any positively-oriented isometric automorphism. -/ +lemma linear_isometry_equiv_comp_right_angle_rotation (φ : E ≃ₗᵢ[ℝ] E) + (hφ : 0 < (φ.to_linear_equiv : E →ₗ[ℝ] E).det) (x : E) : + φ (J x) = J (φ x) := +begin + convert (o.right_angle_rotation_map φ (φ x)).symm, + { simp }, + { symmetry, + rwa ← o.map_eq_iff_det_pos φ.to_linear_equiv at hφ, + rw [fact.out (finrank ℝ E = 2), fintype.card_fin] }, +end + +lemma right_angle_rotation_map' {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] + (φ : E ≃ₗᵢ[ℝ] F) : + (orientation.map (fin 2) φ.to_linear_equiv o).right_angle_rotation + = (φ.symm.trans o.right_angle_rotation).trans φ := +linear_isometry_equiv.ext $ o.right_angle_rotation_map φ + +/-- `J` commutes with any positively-oriented isometric automorphism. -/ +lemma linear_isometry_equiv_comp_right_angle_rotation' (φ : E ≃ₗᵢ[ℝ] E) + (hφ : 0 < (φ.to_linear_equiv : E →ₗ[ℝ] E).det) : + linear_isometry_equiv.trans J φ = φ.trans J := +linear_isometry_equiv.ext $ o.linear_isometry_equiv_comp_right_angle_rotation φ hφ + /-- For a nonzero vector `x` in an oriented two-dimensional real inner product space `E`, `![x, J x]` forms an (orthogonal) basis for `E`. -/ def basis_right_angle_rotation (x : E) (hx : x ≠ 0) : basis (fin 2) ℝ E := @@ -312,7 +392,7 @@ lemma inner_sq_add_area_form_sq (a b : E) : ⟪a, b⟫ ^ 2 + ω a b ^ 2 = ∥a by simpa [sq, real_inner_self_eq_norm_sq] using o.inner_mul_inner_add_area_form_mul_area_form a b b /-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ∥a∥ ^ 2 * ω x y`. (See -`orientation.inner_mul_area_form_sub` for the "applied" form.)-/ +`orientation.inner_mul_area_form_sub` for the "applied" form.) -/ lemma inner_mul_area_form_sub' (a x : E) : ⟪a, x⟫ • ω a - ω a x • @innerₛₗ ℝ _ _ _ a = ∥a∥ ^ 2 • ω x := begin @@ -371,7 +451,7 @@ end /-- A complex-valued real-bilinear map on an oriented real inner product space of dimension 2. Its real part is the inner product and its imaginary part is `orientation.area_form`. -TODO On `ℂ` with the standard orientation, `kahler w z = conj w * z`. -/ +On `ℂ` with the standard orientation, `kahler w z = conj w * z`; see `complex.kahler`. -/ def kahler : E →ₗ[ℝ] E →ₗ[ℝ] ℂ := (linear_map.llcomp ℝ E ℝ ℂ complex.of_real_clm) ∘ₗ (@innerₛₗ ℝ E _ _) + (linear_map.llcomp ℝ E ℝ ℂ ((linear_map.lsmul ℝ ℂ).flip complex.I)) ∘ₗ ω @@ -473,4 +553,89 @@ begin simp, end +lemma kahler_map {F : Type*} [inner_product_space ℝ F] [fact (finrank ℝ F = 2)] + (φ : E ≃ₗᵢ[ℝ] F) (x y : F) : + (orientation.map (fin 2) φ.to_linear_equiv o).kahler x y = o.kahler (φ.symm x) (φ.symm y) := +by simp [kahler_apply_apply, area_form_map] + +/-- The bilinear map `kahler` is invariant under pullback by a positively-oriented isometric +automorphism. -/ +lemma kahler_comp_linear_isometry_equiv (φ : E ≃ₗᵢ[ℝ] E) + (hφ : 0 < (φ.to_linear_equiv : E →ₗ[ℝ] E).det) (x y : E) : + o.kahler (φ x) (φ y) = o.kahler x y := +by simp [kahler_apply_apply, o.area_form_comp_linear_isometry_equiv φ hφ] + +end orientation + +namespace complex + +local attribute [instance] complex.finrank_real_complex_fact + +@[simp] protected lemma area_form (w z : ℂ) : complex.orientation.area_form w z = (conj w * z).im := +begin + let o := complex.orientation, + simp only [o.area_form_to_volume_form, o.volume_form_robust complex.orthonormal_basis_one_I rfl, + basis.det_apply, matrix.det_fin_two, basis.to_matrix_apply,to_basis_orthonormal_basis_one_I, + matrix.cons_val_zero, coe_basis_one_I_repr, matrix.cons_val_one, matrix.head_cons, mul_im, + conj_re, conj_im], + ring, +end + +@[simp] protected lemma right_angle_rotation (z : ℂ) : + complex.orientation.right_angle_rotation z = I * z := +begin + apply ext_inner_right ℝ, + intros w, + rw orientation.inner_right_angle_rotation_left, + simp only [complex.area_form, complex.inner, mul_re, mul_im, conj_re, conj_im, map_mul, conj_I, + neg_re, neg_im, I_re, I_im], + ring, +end + +@[simp] protected lemma kahler (w z : ℂ) : + complex.orientation.kahler w z = conj w * z := +begin + rw orientation.kahler_apply_apply, + ext1; simp, +end + +end complex + +namespace orientation + +local notation `ω` := o.area_form +local notation `J` := o.right_angle_rotation + +open complex + +/-- The area form on an oriented real inner product space of dimension 2 can be evaluated in terms +of a complex-number representation of the space. -/ +lemma area_form_map_complex (f : E ≃ₗᵢ[ℝ] ℂ) + (hf : (orientation.map (fin 2) f.to_linear_equiv o) = complex.orientation) (x y : E) : + ω x y = (conj (f x) * f y).im := +begin + rw [← complex.area_form, ← hf, o.area_form_map], + simp, +end + +/-- The rotation by 90 degrees on an oriented real inner product space of dimension 2 can be +evaluated in terms of a complex-number representation of the space. -/ +lemma right_angle_rotation_map_complex (f : E ≃ₗᵢ[ℝ] ℂ) + (hf : (orientation.map (fin 2) f.to_linear_equiv o) = complex.orientation) (x : E) : + f (J x) = I * f x := +begin + rw [← complex.right_angle_rotation, ← hf, o.right_angle_rotation_map], + simp, +end + +/-- The Kahler form on an oriented real inner product space of dimension 2 can be evaluated in terms +of a complex-number representation of the space. -/ +lemma kahler_map_complex (f : E ≃ₗᵢ[ℝ] ℂ) + (hf : (orientation.map (fin 2) f.to_linear_equiv o) = complex.orientation) (x y : E) : + o.kahler x y = conj (f x) * f y := +begin + rw [← complex.kahler, ← hf, o.kahler_map], + simp, +end + end orientation diff --git a/src/data/complex/module.lean b/src/data/complex/module.lean index fef643df9c16b..dc01517bffef9 100644 --- a/src/data/complex/module.lean +++ b/src/data/complex/module.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Alexander Bentkamp, Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp, Sébastien Gouëzel, Eric Wieser -/ +import linear_algebra.orientation import algebra.order.smul import data.complex.basic import data.fin.vec_notation @@ -164,6 +165,9 @@ by simp [← finrank_eq_dim, finrank_real_complex, bit0] circle. -/ lemma finrank_real_complex_fact : fact (finrank ℝ ℂ = 2) := ⟨finrank_real_complex⟩ +/-- The standard orientation on `ℂ`. -/ +protected noncomputable def orientation : orientation ℝ ℂ (fin 2) := complex.basis_one_I.orientation + end complex /- Register as an instance (with low priority) the fact that a complex vector space is also a real diff --git a/src/linear_algebra/determinant.lean b/src/linear_algebra/determinant.lean index 0347ab5baa19f..3a4d91fcab455 100644 --- a/src/linear_algebra/determinant.lean +++ b/src/linear_algebra/determinant.lean @@ -258,6 +258,26 @@ the determinant to be `1`. -/ linear_map.det (0 : M →ₗ[𝕜] M) = (0 : 𝕜) ^ (finite_dimensional.finrank 𝕜 M) := by simp only [← zero_smul 𝕜 (1 : M →ₗ[𝕜] M), det_smul, mul_one, monoid_hom.map_one] +lemma det_eq_one_of_subsingleton [subsingleton M] (f : M →ₗ[R] M) : (f : M →ₗ[R] M).det = 1 := +begin + have b : basis (fin 0) R M := basis.empty M, + rw ← f.det_to_matrix b, + exact matrix.det_is_empty, +end + +lemma det_eq_one_of_finrank_eq_zero {𝕜 : Type*} [field 𝕜] {M : Type*} [add_comm_group M] + [module 𝕜 M] (h : finite_dimensional.finrank 𝕜 M = 0) (f : M →ₗ[𝕜] M) : + (f : M →ₗ[𝕜] M).det = 1 := +begin + classical, + refine @linear_map.det_cases M _ 𝕜 _ _ _ (λ t, t = 1) f _ rfl, + intros s b, + haveI : is_empty s, + { rw ← fintype.card_eq_zero_iff, + exact (finite_dimensional.finrank_eq_card_basis b).symm.trans h }, + exact matrix.det_is_empty +end + /-- Conjugating a linear map by a linear equiv does not change its determinant. -/ @[simp] lemma det_conj {N : Type*} [add_comm_group N] [module A N] (f : M →ₗ[A] M) (e : M ≃ₗ[A] N) : diff --git a/src/linear_algebra/orientation.lean b/src/linear_algebra/orientation.lean index bc5c073ef61af..3a940f78fb7b9 100644 --- a/src/linear_algebra/orientation.lean +++ b/src/linear_algebra/orientation.lean @@ -76,6 +76,22 @@ by rw [orientation.map, alternating_map.dom_lcongr_refl, module.ray.map_refl] { positive_orientation := ray_of_ne_zero R (alternating_map.const_linear_equiv_of_is_empty 1) $ alternating_map.const_linear_equiv_of_is_empty.injective.ne (by simp) } +@[simp] lemma orientation.map_positive_orientation_of_is_empty [nontrivial R] [is_empty ι] + (f : M ≃ₗ[R] N) : + orientation.map ι f positive_orientation = positive_orientation := +rfl + +@[simp] lemma orientation.map_of_is_empty [is_empty ι] (x : orientation R M ι) (f : M ≃ₗ[R] M) : + orientation.map ι f x = x := +begin + induction x using module.ray.ind with g hg, + rw orientation.map_apply, + congr, + ext i, + rw alternating_map.comp_linear_map_apply, + congr, +end + end ordered_comm_semiring section ordered_comm_ring @@ -83,6 +99,11 @@ section ordered_comm_ring variables {R : Type*} [strict_ordered_comm_ring R] variables {M N : Type*} [add_comm_group M] [add_comm_group N] [module R M] [module R N] +@[simp] protected lemma orientation.map_neg {ι : Type*} [decidable_eq ι] (f : M ≃ₗ[R] N) + (x : orientation R M ι) : + orientation.map ι f (-x) = - orientation.map ι f x := +module.ray.map_neg _ x + namespace basis variables {ι : Type*} [decidable_eq ι] @@ -279,10 +300,12 @@ variables {ι : Type*} [decidable_eq ι] namespace orientation -variables [fintype ι] [finite_dimensional R M] +variables [fintype ι] [_i : finite_dimensional R M] open finite_dimensional +include _i + /-- If the index type has cardinality equal to the finite dimension, any two orientations are equal or negations. -/ lemma eq_or_eq_neg (x₁ x₂ : orientation R M ι) (h : fintype.card ι = finrank R M) : @@ -310,13 +333,27 @@ begin exact e.map_orientation_eq_det_inv_smul x f end +omit _i + /-- If the index type has cardinality equal to the finite dimension, composing an alternating map with the same linear equiv on each argument gives the same orientation if and only if the determinant is positive. -/ lemma map_eq_iff_det_pos (x : orientation R M ι) (f : M ≃ₗ[R] M) (h : fintype.card ι = finrank R M) : - orientation.map ι f x = x ↔ 0 < (f : M →ₗ[R] M).det := -by rw [map_eq_det_inv_smul _ _ h, units_inv_smul, units_smul_eq_self_iff, linear_equiv.coe_det] + orientation.map ι f x = x ↔ 0 < (f : M →ₗ[R] M).det := +begin + casesI is_empty_or_nonempty ι, + { have H : finrank R M = 0, + { refine h.symm.trans _, + convert fintype.card_of_is_empty, + apply_instance }, + simp [linear_map.det_eq_one_of_finrank_eq_zero H] }, + have H : 0 < finrank R M, + { rw ← h, + exact fintype.card_pos }, + haveI : finite_dimensional R M := finite_dimensional_of_finrank H, + rw [map_eq_det_inv_smul _ _ h, units_inv_smul, units_smul_eq_self_iff, linear_equiv.coe_det] +end /-- If the index type has cardinality equal to the finite dimension, composing an alternating map with the same linear equiv on each argument gives the negation of that orientation if and @@ -324,7 +361,21 @@ only if the determinant is negative. -/ lemma map_eq_neg_iff_det_neg (x : orientation R M ι) (f : M ≃ₗ[R] M) (h : fintype.card ι = finrank R M) : orientation.map ι f x = -x ↔ (f : M →ₗ[R] M).det < 0 := -by rw [map_eq_det_inv_smul _ _ h, units_inv_smul, units_smul_eq_neg_iff, linear_equiv.coe_det] +begin + casesI is_empty_or_nonempty ι, + { have H : finrank R M = 0, + { refine h.symm.trans _, + convert fintype.card_of_is_empty, + apply_instance }, + simp [linear_map.det_eq_one_of_finrank_eq_zero H, module.ray.ne_neg_self x] }, + have H : 0 < finrank R M, + { rw ← h, + exact fintype.card_pos }, + haveI : finite_dimensional R M := finite_dimensional_of_finrank H, + rw [map_eq_det_inv_smul _ _ h, units_inv_smul, units_smul_eq_neg_iff, linear_equiv.coe_det] +end + +include _i /-- If the index type has cardinality equal to the finite dimension, a basis with the given orientation. -/ diff --git a/src/linear_algebra/ray.lean b/src/linear_algebra/ray.lean index ac0ce2d200480..213c4b96e97bb 100644 --- a/src/linear_algebra/ray.lean +++ b/src/linear_algebra/ray.lean @@ -397,6 +397,12 @@ begin rwa [units.coe_neg, right.neg_pos_iff] end +@[simp] protected lemma map_neg (f : M ≃ₗ[R] N) (v : module.ray R M) : map f (-v) = - map f v := +begin + induction v using module.ray.ind with g hg, + simp, +end + end module.ray end strict_ordered_comm_ring From c99203f729a405bb87cda9e55cbecc091348011f Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Thu, 20 Oct 2022 08:58:46 +0000 Subject: [PATCH 827/828] feat(measure_theory/integral/interval_integral): add lemmas (#16986) * Clean up some implicit arguments * Slightly generalize `continuous_linear_map.interval_integral_comp_comm` * From the sphere eversion project Co-authored by: Patrick Massot [patrick.massot@u-psud.fr](mailto:patrick.massot@u-psud.fr) --- .../function/strongly_measurable.lean | 7 ++ .../integral/interval_integral.lean | 101 +++++++++++++----- 2 files changed, 80 insertions(+), 28 deletions(-) diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index 99974fc0ddb8a..d6b37d8ab4439 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -1628,6 +1628,13 @@ protected lemma Union [pseudo_metrizable_space β] {s : ι → set α} ae_strongly_measurable f (μ.restrict s) ∧ ae_strongly_measurable f (μ.restrict t) := by simp only [union_eq_Union, ae_strongly_measurable_Union_iff, bool.forall_bool, cond, and.comm] +lemma ae_strongly_measurable_interval_oc_iff [linear_order α] [pseudo_metrizable_space β] + {f : α → β} {a b : α} : + ae_strongly_measurable f (μ.restrict $ interval_oc a b) ↔ + ae_strongly_measurable f (μ.restrict $ Ioc a b) ∧ + ae_strongly_measurable f (μ.restrict $ Ioc b a) := +by rw [interval_oc_eq_union, ae_strongly_measurable_union_iff] + lemma smul_measure {R : Type*} [monoid R] [distrib_mul_action R ℝ≥0∞] [is_scalar_tower R ℝ≥0∞ ℝ≥0∞] (h : ae_strongly_measurable f μ) (c : R) : ae_strongly_measurable f (c • μ) := diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index 1397c54b61a15..1dd8722978420 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -179,7 +179,7 @@ variables {ι 𝕜 E F : Type*} [normed_add_comm_group E] interval `a..b` if it is integrable on both intervals `(a, b]` and `(b, a]`. One of these intervals is always empty, so this property is equivalent to `f` being integrable on `(min a b, max a b]`. -/ -def interval_integrable (f : ℝ → E) (μ : measure ℝ) (a b : ℝ) := +def interval_integrable (f : ℝ → E) (μ : measure ℝ) (a b : ℝ) : Prop := integrable_on f (Ioc a b) μ ∧ integrable_on f (Ioc b a) μ section @@ -201,8 +201,7 @@ lemma interval_integrable_iff_integrable_Ioc_of_le (hab : a ≤ b) : interval_integrable f μ a b ↔ integrable_on f (Ioc a b) μ := by rw [interval_integrable_iff, interval_oc_of_le hab] -lemma integrable_on_Icc_iff_integrable_on_Ioc' - {E : Type*} [normed_add_comm_group E] {f : ℝ → E} (ha : μ {a} ≠ ∞) : +lemma integrable_on_Icc_iff_integrable_on_Ioc' {f : ℝ → E} (ha : μ {a} ≠ ∞) : integrable_on f (Icc a b) μ ↔ integrable_on f (Ioc a b) μ := begin cases le_or_lt a b with hab hab, @@ -212,13 +211,11 @@ begin { simp [hab, hab.le] }, end -lemma integrable_on_Icc_iff_integrable_on_Ioc - {E : Type*}[normed_add_comm_group E] [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : +lemma integrable_on_Icc_iff_integrable_on_Ioc [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : integrable_on f (Icc a b) μ ↔ integrable_on f (Ioc a b) μ := integrable_on_Icc_iff_integrable_on_Ioc' (by simp) lemma integrable_on_Ioc_iff_integrable_on_Ioo' - {E : Type*} [normed_add_comm_group E] {f : ℝ → E} {a b : ℝ} (hb : μ {b} ≠ ∞) : integrable_on f (Ioc a b) μ ↔ integrable_on f (Ioo a b) μ := begin @@ -229,13 +226,11 @@ begin { simp [hab] }, end -lemma integrable_on_Ioc_iff_integrable_on_Ioo - {E : Type*} [normed_add_comm_group E] [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : +lemma integrable_on_Ioc_iff_integrable_on_Ioo [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : integrable_on f (Ioc a b) μ ↔ integrable_on f (Ioo a b) μ := integrable_on_Ioc_iff_integrable_on_Ioo' (by simp) -lemma integrable_on_Icc_iff_integrable_on_Ioo - {E : Type*} [normed_add_comm_group E] [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : +lemma integrable_on_Icc_iff_integrable_on_Ioo [has_no_atoms μ] {f : ℝ → E} {a b : ℝ} : integrable_on f (Icc a b) μ ↔ integrable_on f (Ioo a b) μ := by rw [integrable_on_Icc_iff_integrable_on_Ioc, integrable_on_Ioc_iff_integrable_on_Ioo] @@ -243,13 +238,12 @@ lemma interval_integrable_iff' [has_no_atoms μ] : interval_integrable f μ a b ↔ integrable_on f (interval a b) μ := by rw [interval_integrable_iff, interval, interval_oc, integrable_on_Icc_iff_integrable_on_Ioc] -lemma interval_integrable_iff_integrable_Icc_of_le {E : Type*} [normed_add_comm_group E] +lemma interval_integrable_iff_integrable_Icc_of_le {f : ℝ → E} {a b : ℝ} (hab : a ≤ b) {μ : measure ℝ} [has_no_atoms μ] : interval_integrable f μ a b ↔ integrable_on f (Icc a b) μ := by rw [interval_integrable_iff_integrable_Ioc_of_le hab, integrable_on_Icc_iff_integrable_on_Ioc] -lemma integrable_on_Ici_iff_integrable_on_Ioi' - {E : Type*} [normed_add_comm_group E] {f : ℝ → E} (ha : μ {a} ≠ ∞) : +lemma integrable_on_Ici_iff_integrable_on_Ioi' {f : ℝ → E} (ha : μ {a} ≠ ∞) : integrable_on f (Ici a) μ ↔ integrable_on f (Ioi a) μ := begin have : Ici a = Icc a a ∪ Ioi a := (Icc_union_Ioi_eq_Ici le_rfl).symm, @@ -257,8 +251,7 @@ begin simp [ha.lt_top] end -lemma integrable_on_Ici_iff_integrable_on_Ioi - {E : Type*} [normed_add_comm_group E] [has_no_atoms μ] {f : ℝ → E} : +lemma integrable_on_Ici_iff_integrable_on_Ioi [has_no_atoms μ] {f : ℝ → E} : integrable_on f (Ici a) μ ↔ integrable_on f (Ioi a) μ := integrable_on_Ici_iff_integrable_on_Ioi' (by simp) @@ -322,6 +315,11 @@ lemma norm (h : interval_integrable f μ a b) : interval_integrable (λ x, ∥f x∥) μ a b := ⟨h.1.norm, h.2.norm⟩ +lemma interval_integrable_norm_iff {f : ℝ → E} {μ : measure ℝ} {a b : ℝ} + (hf : ae_strongly_measurable f (μ.restrict (Ι a b))) : + interval_integrable (λ t, ∥f t∥) μ a b ↔ interval_integrable f μ a b := +by { simp_rw [interval_integrable_iff, integrable_on], exact integrable_norm_iff hf } + lemma abs {f : ℝ → ℝ} (h : interval_integrable f μ a b) : interval_integrable (λ x, |f x|) μ a b := h.norm @@ -331,18 +329,22 @@ lemma mono (hf : interval_integrable f ν a b) (h1 : [c, d] ⊆ [a, b]) (h2 : μ interval_integrable_iff.mpr $ hf.def.mono (interval_oc_subset_interval_oc_of_interval_subset_interval h1) h2 -lemma mono_set (hf : interval_integrable f μ a b) (h : [c, d] ⊆ [a, b]) : - interval_integrable f μ c d := -hf.mono h rfl.le - lemma mono_measure (hf : interval_integrable f ν a b) (h : μ ≤ ν) : interval_integrable f μ a b := hf.mono rfl.subset h +lemma mono_set (hf : interval_integrable f μ a b) (h : [c, d] ⊆ [a, b]) : + interval_integrable f μ c d := +hf.mono h rfl.le + lemma mono_set_ae (hf : interval_integrable f μ a b) (h : Ι c d ≤ᵐ[μ] Ι a b) : interval_integrable f μ c d := interval_integrable_iff.mpr $ hf.def.mono_set_ae h +lemma mono_set' (hf : interval_integrable f μ a b) (hsub : Ι c d ⊆ Ι a b) : + interval_integrable f μ c d := +hf.mono_set_ae $ eventually_of_forall hsub + lemma mono_fun [normed_add_comm_group F] {g : ℝ → F} (hf : interval_integrable f μ a b) (hgm : ae_strongly_measurable g (μ.restrict (Ι a b))) (hle : (λ x, ∥g x∥) ≤ᵐ[μ.restrict (Ι a b)] (λ x, ∥f x∥)) : interval_integrable g μ a b := @@ -395,6 +397,14 @@ lemma continuous_on_mul {f g : ℝ → ℝ} (hf : interval_integrable f μ a b) interval_integrable (λ x, g x * f x) μ a b := by simpa [mul_comm] using hf.mul_continuous_on hg +lemma const_mul {f : ℝ → ℝ} {a b : ℝ} {μ : measure ℝ} + (hf : interval_integrable f μ a b) (c : ℝ) : interval_integrable (λ x, c * f x) μ a b := +hf.continuous_on_mul continuous_on_const + +lemma mul_const {f : ℝ → ℝ} {a b : ℝ} {μ : measure ℝ} + (hf : interval_integrable f μ a b) (c : ℝ) : interval_integrable (λ x, f x * c) μ a b := +hf.mul_continuous_on continuous_on_const + lemma comp_mul_left (hf : interval_integrable f volume a b) (c : ℝ) : interval_integrable (λ x, f (c * x)) volume (a / c) (b / c) := begin @@ -511,7 +521,7 @@ variables [complete_space E] [normed_space ℝ E] /-- The interval integral `∫ x in a..b, f x ∂μ` is defined as `∫ x in Ioc a b, f x ∂μ - ∫ x in Ioc b a, f x ∂μ`. If `a ≤ b`, then it equals `∫ x in Ioc a b, f x ∂μ`, otherwise it equals `-∫ x in Ioc b a, f x ∂μ`. -/ -def interval_integral (f : ℝ → E) (a b : ℝ) (μ : measure ℝ) := +def interval_integral (f : ℝ → E) (a b : ℝ) (μ : measure ℝ) : E := ∫ x in Ioc a b, f x ∂μ - ∫ x in Ioc b a, f x ∂μ notation `∫` binders ` in ` a `..` b `, ` r:(scoped:60 f, f) ` ∂` μ:70 := interval_integral r a b μ @@ -546,6 +556,17 @@ begin { simp only [integral_of_ge (not_le.1 h).le, interval_oc_of_lt (not_le.1 h), neg_one_smul] } end +lemma norm_interval_integral_eq (f : ℝ → E) (a b : ℝ) (μ : measure ℝ) : + ∥∫ x in a..b, f x ∂μ∥ = ∥∫ x in Ι a b, f x ∂μ∥ := +begin + simp_rw [interval_integral_eq_integral_interval_oc, norm_smul], + split_ifs; simp only [norm_neg, norm_one, one_mul], +end + +lemma abs_interval_integral_eq (f : ℝ → ℝ) (a b : ℝ) (μ : measure ℝ) : + |∫ x in a..b, f x ∂μ| = |∫ x in Ι a b, f x ∂μ| := +norm_interval_integral_eq f a b μ + lemma integral_cases (f : ℝ → E) (a b) : ∫ x in a..b, f x ∂μ ∈ ({∫ x in Ι a b, f x ∂μ, -∫ x in Ι a b, f x ∂μ} : set E) := by { rw interval_integral_eq_integral_interval_oc, split_ifs; simp } @@ -557,6 +578,11 @@ by cases le_total a b with hab hab; refine integral_undef (not_imp_not.mpr integrable.integrable_on' _); simpa [hab] using not_and_distrib.mp h +lemma interval_integrable_of_integral_ne_zero {a b : ℝ} + {f : ℝ → E} {μ : measure ℝ} (h : ∫ x in a..b, f x ∂μ ≠ 0) : + interval_integrable f μ a b := +by { contrapose! h, exact interval_integral.integral_undef h } + lemma integral_non_ae_strongly_measurable (hf : ¬ ae_strongly_measurable f (μ.restrict (Ι a b))) : ∫ x in a..b, f x ∂μ = 0 := @@ -596,6 +622,14 @@ lemma norm_integral_le_integral_norm (h : a ≤ b) : ∥∫ x in a..b, f x ∂μ∥ ≤ ∫ x in a..b, ∥f x∥ ∂μ := norm_integral_le_integral_norm_Ioc.trans_eq $ by rw [interval_oc_of_le h, integral_of_le h] +lemma norm_integral_le_of_norm_le {g : ℝ → ℝ} + (h : ∀ᵐ t ∂(μ.restrict $ Ι a b), ∥f t∥ ≤ g t) + (hbound : interval_integrable g μ a b) : + ∥∫ t in a..b, f t ∂μ∥ ≤ |∫ t in a..b, g t ∂μ| := +by simp_rw [norm_interval_integral_eq, abs_interval_integral_eq, + abs_eq_self.mpr (integral_nonneg_of_ae $ h.mono $ λ t ht, (norm_nonneg _).trans ht), + norm_integral_le_of_norm_le hbound.def h] + lemma norm_integral_le_of_norm_le_const_ae {a b C : ℝ} {f : ℝ → E} (h : ∀ᵐ x, x ∈ Ι a b → ∥f x∥ ≤ C) : ∥∫ x in a..b, f x∥ ≤ C * |b - a| := @@ -661,18 +695,29 @@ lemma integral_smul_measure (c : ℝ≥0∞) : ∫ x in a..b, f x ∂(c • μ) = c.to_real • ∫ x in a..b, f x ∂μ := by simp only [interval_integral, measure.restrict_smul, integral_smul_measure, smul_sub] -variables [normed_add_comm_group F] [complete_space F] [normed_space ℝ F] +end basic + +section continuous_linear_map + +variables {a b : ℝ} {μ : measure ℝ} {f : ℝ → E} +variables [is_R_or_C 𝕜] [normed_space 𝕜 E] [normed_add_comm_group F] [normed_space 𝕜 F] + +open continuous_linear_map + +lemma _root_.continuous_linear_map.interval_integral_apply {a b : ℝ} {φ : ℝ → F →L[𝕜] E} + (hφ : interval_integrable φ μ a b) (v : F) : + (∫ x in a..b, φ x ∂μ) v = ∫ x in a..b, φ x v ∂μ := +by simp_rw [interval_integral_eq_integral_interval_oc, ← integral_apply hφ.def v, coe_smul', + pi.smul_apply] + +variables [normed_space ℝ F] [complete_space F] lemma _root_.continuous_linear_map.interval_integral_comp_comm - (L : E →L[ℝ] F) (hf : interval_integrable f μ a b) : + (L : E →L[𝕜] F) (hf : interval_integrable f μ a b) : ∫ x in a..b, L (f x) ∂μ = L (∫ x in a..b, f x ∂μ) := -begin - rw [interval_integral, interval_integral, L.integral_comp_comm, L.integral_comp_comm, L.map_sub], - exacts [hf.2, hf.1] -end - -end basic +by simp_rw [interval_integral, L.integral_comp_comm hf.1, L.integral_comp_comm hf.2, L.map_sub] +end continuous_linear_map section comp variables {a b c d : ℝ} (f : ℝ → E) From f5138420c8e0534fffca1e3d91ba4b653c5d309e Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Thu, 20 Oct 2022 11:19:51 +0000 Subject: [PATCH 828/828] chore(algebra/squarefree): rename `squarefree_of_dvd_of_squarefree` for dot notation (#17071) As suggested in #16999, this name is quite long and should get dot notation. I'm not totally convinced by the name but can't come up with a better one myself at the moment. Co-authored-by: Anne Baanen Co-authored-by: Thomas Browning --- src/algebra/squarefree.lean | 6 +++--- src/data/nat/squarefree.lean | 2 +- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/src/algebra/squarefree.lean b/src/algebra/squarefree.lean index 3a105fba890a5..c7a1d5088e629 100644 --- a/src/algebra/squarefree.lean +++ b/src/algebra/squarefree.lean @@ -77,7 +77,7 @@ lemma squarefree.of_mul_left [comm_monoid R] {m n : R} (hmn : squarefree (m * n) lemma squarefree.of_mul_right [comm_monoid R] {m n : R} (hmn : squarefree (m * n)) : squarefree n := (λ p hp, hmn p (dvd_mul_of_dvd_right hp m)) -lemma squarefree_of_dvd_of_squarefree [comm_monoid R] +lemma squarefree.squarefree_of_dvd [comm_monoid R] {x y : R} (hdvd : x ∣ y) (hsq : squarefree y) : squarefree x := λ a h, hsq _ (h.trans hdvd) @@ -88,11 +88,11 @@ variables {α : Type*} [cancel_comm_monoid_with_zero α] [gcd_monoid α] lemma squarefree.gcd_right (a : α) {b : α} (hb : squarefree b) : squarefree (gcd a b) := -squarefree_of_dvd_of_squarefree (gcd_dvd_right _ _) hb +hb.squarefree_of_dvd (gcd_dvd_right _ _) lemma squarefree.gcd_left {a : α} (b : α) (ha : squarefree a) : squarefree (gcd a b) := -squarefree_of_dvd_of_squarefree (gcd_dvd_left _ _) ha +ha.squarefree_of_dvd (gcd_dvd_left _ _) end squarefree_gcd_of_squarefree diff --git a/src/data/nat/squarefree.lean b/src/data/nat/squarefree.lean index 3a25b67b1f922..2c4e7b1fbe266 100644 --- a/src/data/nat/squarefree.lean +++ b/src/data/nat/squarefree.lean @@ -84,7 +84,7 @@ begin refine ⟨λ h, _, by { rintro ⟨hn, rfl⟩, simpa }⟩, rcases eq_or_ne n 0 with rfl | hn₀, { simpa [zero_pow hk.bot_lt] using h }, - refine ⟨squarefree_of_dvd_of_squarefree (dvd_pow_self _ hk) h, by_contradiction $ λ h₁, _⟩, + refine ⟨h.squarefree_of_dvd (dvd_pow_self _ hk), by_contradiction $ λ h₁, _⟩, have : 2 ≤ k := k.two_le_iff.mpr ⟨hk, h₁⟩, apply hn (nat.is_unit_iff.1 (h _ _)), rw ←sq,